Skip to content

Commit 604b1f4

Browse files
authored
[ DRY ] redefine Algebra.Definitions.Medial in terms of Interchangable (#2764)
* [ DRY ] redefine `Algebra.Definitions.Medial` in terms of `Interchangable`, plus knock-ons * reset: `CHANGELOG` for v2.4
1 parent aef1785 commit 604b1f4

File tree

6 files changed

+40
-20
lines changed

6 files changed

+40
-20
lines changed

CHANGELOG.md

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -32,6 +32,11 @@ Deprecated modules
3232
Deprecated names
3333
----------------
3434

35+
* In `Algebra.Properties.CommutativeSemigroup`:
36+
```agda
37+
interchange ↦ medial
38+
```
39+
3540
New modules
3641
-----------
3742

src/Algebra/Definitions.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -212,7 +212,7 @@ Flexible : Op₂ A → Set _
212212
Flexible _∙_ = x y ((x ∙ y) ∙ x) ≈ (x ∙ (y ∙ x))
213213

214214
Medial : Op₂ A Set _
215-
Medial _∙_ = x y u z ((x ∙ y) ∙ (u ∙ z)) ≈ ((x ∙ u) ∙ (y ∙ z))
215+
Medial _∙_ = Interchangable _∙_ _∙_
216216

217217
LeftSemimedial : Op₂ A Set _
218218
LeftSemimedial _∙_ = x y z ((x ∙ x) ∙ (y ∙ z)) ≈ ((x ∙ y) ∙ (x ∙ z))

src/Algebra/Properties/CommutativeSemigroup.agda

Lines changed: 19 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -9,10 +9,10 @@
99
open import Algebra.Bundles using (CommutativeSemigroup)
1010

1111
module Algebra.Properties.CommutativeSemigroup
12-
{a ℓ} (CS : CommutativeSemigroup a ℓ)
12+
{a ℓ} (commutativeSemigroup : CommutativeSemigroup a ℓ)
1313
where
1414

15-
open CommutativeSemigroup CS
15+
open CommutativeSemigroup commutativeSemigroup
1616
open import Algebra.Definitions _≈_
1717
open import Relation.Binary.Reasoning.Setoid setoid
1818
open import Data.Product.Base using (_,_)
@@ -25,8 +25,8 @@ open import Algebra.Properties.Semigroup semigroup public
2525
------------------------------------------------------------------------
2626
-- Properties
2727

28-
interchange : Interchangable _∙_ _∙_
29-
interchange a b c d = begin
28+
medial : Medial _∙_
29+
medial a b c d = begin
3030
(a ∙ b) ∙ (c ∙ d) ≈⟨ assoc a b (c ∙ d) ⟩
3131
a ∙ (b ∙ (c ∙ d)) ≈⟨ ∙-congˡ (assoc b c d) ⟨
3232
a ∙ ((b ∙ c) ∙ d) ≈⟨ ∙-congˡ (∙-congʳ (comm b c)) ⟩
@@ -171,3 +171,18 @@ middleSemimedial x y z = begin
171171

172172
semimedial : Semimedial _∙_
173173
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+
#-}

src/Algebra/Properties/CommutativeSemigroup/Divisibility.agda

Lines changed: 8 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -7,17 +7,17 @@
77
{-# OPTIONS --cubical-compatible --safe #-}
88

99
open import Algebra using (CommutativeSemigroup)
10-
open import Data.Product.Base using (_,_)
11-
import Relation.Binary.Reasoning.Setoid as ≈-Reasoning
1210

1311
module Algebra.Properties.CommutativeSemigroup.Divisibility
14-
{a ℓ} (CS : CommutativeSemigroup a ℓ)
12+
{a ℓ} (commutativeSemigroup : CommutativeSemigroup a ℓ)
1513
where
1614

17-
open CommutativeSemigroup CS
18-
open import Algebra.Properties.CommutativeSemigroup CS
19-
using (interchange; x∙yz≈xz∙y; x∙yz≈y∙xz)
20-
open ≈-Reasoning setoid
15+
open import Data.Product.Base using (_,_)
16+
17+
open CommutativeSemigroup commutativeSemigroup
18+
open import Algebra.Properties.CommutativeSemigroup commutativeSemigroup
19+
using (medial; x∙yz≈xz∙y; x∙yz≈y∙xz)
20+
open import Relation.Binary.Reasoning.Setoid setoid
2121

2222
------------------------------------------------------------------------
2323
-- Re-export the contents of divisibility over semigroups
@@ -42,7 +42,7 @@ x∣y∧z∣x/y⇒xz∣y {x} {y} {z} (x/y , x/y∙x≈y) (p , pz≈x/y) = p , (b
4242

4343
∙-cong-∣ : {x y a b} x ∣ y a ∣ b x ∙ a ∣ y ∙ b
4444
∙-cong-∣ {x} {y} {a} {b} (p , px≈y) (q , qa≈b) = p ∙ q , (begin
45-
(p ∙ q) ∙ (x ∙ a) ≈⟨ interchange p q x a ⟩
45+
(p ∙ q) ∙ (x ∙ a) ≈⟨ medial p q x a ⟩
4646
(p ∙ x) ∙ (q ∙ a) ≈⟨ ∙-cong px≈y qa≈b ⟩
4747
y ∙ b ∎)
4848

src/Algebra/Properties/IdempotentCommutativeMonoid.agda

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -9,16 +9,16 @@
99
open import Algebra.Bundles using (IdempotentCommutativeMonoid)
1010

1111
module Algebra.Properties.IdempotentCommutativeMonoid
12-
{c ℓ} (M : IdempotentCommutativeMonoid c ℓ) where
12+
{c ℓ} (idempotentCommutativeMonoid : IdempotentCommutativeMonoid c ℓ) where
1313

14-
open IdempotentCommutativeMonoid M
14+
open IdempotentCommutativeMonoid idempotentCommutativeMonoid
1515

1616
open import Algebra.Consequences.Setoid setoid
1717
using (comm∧distrˡ⇒distrʳ; comm∧distrˡ⇒distr)
1818
open import Algebra.Definitions _≈_
1919
using (_DistributesOverˡ_; _DistributesOverʳ_; _DistributesOver_)
2020
open import Algebra.Properties.CommutativeSemigroup commutativeSemigroup
21-
using (interchange)
21+
using (medial)
2222
open import Relation.Binary.Reasoning.Setoid setoid
2323

2424

@@ -28,7 +28,7 @@ open import Relation.Binary.Reasoning.Setoid setoid
2828
∙-distrˡ-∙ : _∙_ DistributesOverˡ _∙_
2929
∙-distrˡ-∙ a b c = begin
3030
a ∙ (b ∙ c) ≈⟨ ∙-congʳ (idem a) ⟨
31-
(a ∙ a) ∙ (b ∙ c) ≈⟨ interchange _ _ _ _ ⟩
31+
(a ∙ a) ∙ (b ∙ c) ≈⟨ medial _ _ _ _ ⟩
3232
(a ∙ b) ∙ (a ∙ c) ∎
3333

3434
∙-distrʳ-∙ : _∙_ DistributesOverʳ _∙_

src/Algebra/Solver/CommutativeMonoid/Normal.agda

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -13,7 +13,7 @@ open import Algebra.Bundles using (CommutativeMonoid)
1313
module Algebra.Solver.CommutativeMonoid.Normal {c ℓ} (M : CommutativeMonoid c ℓ) where
1414

1515
import Algebra.Properties.CommutativeSemigroup as CSProperties
16-
using (interchange)
16+
using (medial)
1717
import Algebra.Properties.Monoid.Mult as MultProperties
1818
using (_×_; ×-homo-1; ×-homo-+)
1919
open import Data.Fin.Base using (Fin; zero; suc)
@@ -27,7 +27,7 @@ import Relation.Nullary.Decidable as Dec using (map)
2727

2828
open CommutativeMonoid M
2929
open MultProperties monoid using (_×_; ×-homo-1; ×-homo-+)
30-
open CSProperties commutativeSemigroup using (interchange)
30+
open CSProperties commutativeSemigroup using (medial)
3131
open ≈-Reasoning setoid
3232

3333
private
@@ -117,7 +117,7 @@ comp-correct [] [] _ = sym (identityˡ _)
117117
comp-correct (l ∷ v) (m ∷ w) (a ∷ ρ) = begin
118118
((l + m) × a) ∙ ⟦ v • w ⟧⇓ ρ ≈⟨ ∙-congʳ (×-homo-+ a l m) ⟩
119119
(l × a) ∙ (m × a) ∙ ⟦ v • w ⟧⇓ ρ ≈⟨ ∙-congˡ (comp-correct v w ρ) ⟩
120-
(l × a) ∙ (m × a) ∙ (⟦ v ⟧⇓ ρ ∙ ⟦ w ⟧⇓ ρ) ≈⟨ interchange _ _ _ _ ⟩
120+
(l × a) ∙ (m × a) ∙ (⟦ v ⟧⇓ ρ ∙ ⟦ w ⟧⇓ ρ) ≈⟨ medial _ _ _ _ ⟩
121121
⟦ l ∷ v ⟧⇓ (a ∷ ρ) ∙ ⟦ m ∷ w ⟧⇓ (a ∷ ρ) ∎
122122

123123
------------------------------------------------------------------------

0 commit comments

Comments
 (0)