Skip to content

Commit 0cccb39

Browse files
authored
[ refactor ] rename ∣∣ to , ∤∤ to throughout Algebra.Definitions.RawMagma and Algebra.Properties.*.Divisibility (#2562)
* refactor: rename `∣∣` to `∥`, `∤∤` to `∦` throughout * doc: added note to `style-guide`
1 parent 908e015 commit 0cccb39

File tree

6 files changed

+176
-39
lines changed

6 files changed

+176
-39
lines changed

CHANGELOG.md

+30
Original file line numberDiff line numberDiff line change
@@ -21,6 +21,36 @@ Deprecated modules
2121
Deprecated names
2222
----------------
2323

24+
* In `Algebra.Definitions.RawMagma`:
25+
```agda
26+
_∣∣_ ↦ _∥_
27+
_∤∤_ ↦ _∦_
28+
```
29+
30+
* In `Algebra.Properties.Magma.Divisibility`:
31+
```agda
32+
∣∣-sym ↦ ∥-sym
33+
∣∣-respˡ-≈ ↦ ∥-respˡ-≈
34+
∣∣-respʳ-≈ ↦ ∥-respʳ-≈
35+
∣∣-resp-≈ ↦ ∥-resp-≈
36+
∤∤-sym -≈ ↦ ∦-sym
37+
∤∤-respˡ-≈ ↦ ∦-respˡ-≈
38+
∤∤-respʳ-≈ ↦ ∦-respʳ-≈
39+
∤∤-resp-≈ ↦ ∦-resp-≈
40+
```
41+
42+
* In `Algebra.Properties.Monoid.Divisibility`:
43+
```agda
44+
∣∣-refl ↦ ∥-refl
45+
∣∣-reflexive ↦ ∥-reflexive
46+
∣∣-isEquivalence ↦ ∥-isEquivalence
47+
```
48+
49+
* In `Algebra.Properties.Semigroup.Divisibility`:
50+
```agda
51+
∣∣-trans ↦ ∥-trans
52+
```
53+
2454
New modules
2555
-----------
2656

doc/style-guide.md

+3-2
Original file line numberDiff line numberDiff line change
@@ -571,8 +571,9 @@ word within a compound word is capitalized except for the first word.
571571

572572
* Likewise, standard infix symbols for eg, divisibility on numeric
573573
datatypes/algebraic structure, should be written `\|`, NOT the
574-
unmarked `|` character. An exception to this is the *strict*
575-
ordering relation, written using `<`, NOT `\<` as might be expected.
574+
unmarked `|` character. Ditto. mutual divisibility `\||` etc. An
575+
exception to this is the *strict* ordering relation, written using
576+
`<`, NOT `\<` as might be expected.
576577

577578
* Since v2.0, the `Algebra` hierarchy systematically introduces
578579
consistent symbolic notation for the negated versions of the usual

src/Algebra/Definitions/RawMagma.agda

+28-7
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@
1111
{-# OPTIONS --cubical-compatible --safe #-}
1212

1313
open import Algebra.Bundles.Raw using (RawMagma)
14-
open import Data.Product.Base using (_×_; ∃)
14+
open import Data.Product.Base using (_×_)
1515
open import Level using (_⊔_)
1616
open import Relation.Binary.Core using (Rel)
1717
open import Relation.Nullary.Negation.Core using (¬_)
@@ -25,7 +25,7 @@ open RawMagma M renaming (Carrier to A)
2525
------------------------------------------------------------------------
2626
-- Divisibility
2727

28-
infix 5 _∣ˡ_ _∤ˡ_ _∣ʳ_ _∤ʳ_ _∣_ _∤_ _∣∣_ _∤∤_
28+
infix 5 _∣ˡ_ _∤ˡ_ _∣ʳ_ _∤ʳ_ _∣_ _∤_ __ __
2929

3030
-- Divisibility from the left.
3131
--
@@ -34,7 +34,7 @@ infix 5 _∣ˡ_ _∤ˡ_ _∣ʳ_ _∤ʳ_ _∣_ _∤_ _∣∣_ _∤∤_
3434
-- make the use of pattern synonyms more ergonomic (see #2216 for
3535
-- further details). The record field names are not designed to be
3636
-- used explicitly and indeed aren't re-exported publicly by
37-
-- `Algebra.X.Properties.Divisibility` modules.
37+
-- `Algebra.Properties.X.Divisibility` modules.
3838

3939
record _∣ˡ_ (x y : A) : Set (a ⊔ ℓ) where
4040
constructor _,_
@@ -79,8 +79,29 @@ x ∤ y = ¬ x ∣ y
7979
-- Example: for ℕ this is equivalent to x ≡ y,
8080
-- for ℤ this is equivalent to (x ≡ y or x ≡ - y).
8181

82-
_∣∣_ : Rel A (a ⊔ ℓ)
83-
x ∣∣ y = x ∣ y × y ∣ x
82+
_∥_ : Rel A (a ⊔ ℓ)
83+
x ∥ y = x ∣ y × y ∣ x
84+
85+
_∦_ : Rel A (a ⊔ ℓ)
86+
x ∦ y = ¬ x ∥ y
87+
88+
89+
------------------------------------------------------------------------
90+
-- DEPRECATED NAMES
91+
------------------------------------------------------------------------
92+
-- Please use the new names as continuing support for the old names is
93+
-- not guaranteed.
94+
95+
-- Version 2.3
96+
97+
_∣∣_ = _∥_
98+
{-# WARNING_ON_USAGE _∣∣_
99+
"Warning: _∣∣_ was deprecated in v2.3.
100+
Please use _∥_ instead."
101+
#-}
102+
_∤∤_ = _∦_
103+
{-# WARNING_ON_USAGE _∤∤_
104+
"Warning: _∤∤_ was deprecated in v2.3.
105+
Please use _∦_ instead."
106+
#-}
84107

85-
_∤∤_ : Rel A (a ⊔ ℓ)
86-
x ∤∤ y = ¬ x ∣∣ y

src/Algebra/Properties/Magma/Divisibility.agda

+61-18
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ open Magma M
2121
-- Re-export divisibility relations publicly
2222

2323
open import Algebra.Definitions.RawMagma rawMagma public
24-
using (_∣_; _∤_; _∣∣_; _∤∤_; _∣ˡ_; _∤ˡ_; _∣ʳ_; _∤ʳ_; _,_)
24+
using (_∣_; _∤_; __; __; _∣ˡ_; _∤ˡ_; _∣ʳ_; _∤ʳ_; _,_)
2525

2626
------------------------------------------------------------------------
2727
-- Properties of divisibility
@@ -54,34 +54,34 @@ xy≈z⇒y∣z x y xy≈z = ∣-respʳ-≈ xy≈z (x∣yx y x)
5454
∤-resp-≈ = ∤-respʳ-≈ , ∤-respˡ-≈
5555

5656
------------------------------------------------------------------------
57-
-- Properties of mutual divisibility _∣∣_
57+
-- Properties of mutual divisibility __
5858

59-
∣∣-sym : Symmetric _∣∣_
60-
∣∣-sym = swap
59+
-sym : Symmetric __
60+
-sym = swap
6161

62-
∣∣-respˡ-≈ : _∣∣_ Respectsˡ _≈_
63-
∣∣-respˡ-≈ x≈z (x∣y , y∣x) = ∣-respˡ-≈ x≈z x∣y , ∣-respʳ-≈ x≈z y∣x
62+
-respˡ-≈ : __ Respectsˡ _≈_
63+
-respˡ-≈ x≈z (x∣y , y∣x) = ∣-respˡ-≈ x≈z x∣y , ∣-respʳ-≈ x≈z y∣x
6464

65-
∣∣-respʳ-≈ : _∣∣_ Respectsʳ _≈_
66-
∣∣-respʳ-≈ y≈z (x∣y , y∣x) = ∣-respʳ-≈ y≈z x∣y , ∣-respˡ-≈ y≈z y∣x
65+
-respʳ-≈ : __ Respectsʳ _≈_
66+
-respʳ-≈ y≈z (x∣y , y∣x) = ∣-respʳ-≈ y≈z x∣y , ∣-respˡ-≈ y≈z y∣x
6767

68-
∣∣-resp-≈ : _∣∣_ Respects₂ _≈_
69-
∣∣-resp-≈ = ∣∣-respʳ-≈ , ∣∣-respˡ-≈
68+
-resp-≈ : __ Respects₂ _≈_
69+
-resp-≈ = -respʳ-≈ , -respˡ-≈
7070

7171
------------------------------------------------------------------------
7272
-- Properties of mutual non-divisibility _∤∤_
7373

74-
∤∤-sym : Symmetric _∤∤_
75-
∤∤-sym x∤∤y y∣∣x = contradiction (∣∣-sym y∣∣x) x∤∤y
74+
-sym : Symmetric __
75+
-sym xy yx = contradiction (-sym yx) xy
7676

77-
∤∤-respˡ-≈ : _∤∤_ Respectsˡ _≈_
78-
∤∤-respˡ-≈ x≈y x∤∤z y∣∣z = contradiction (∣∣-respˡ-≈ (sym x≈y) y∣∣z) x∤∤z
77+
-respˡ-≈ : __ Respectsˡ _≈_
78+
-respˡ-≈ x≈y xz yz = contradiction (-respˡ-≈ (sym x≈y) yz) xz
7979

80-
∤∤-respʳ-≈ : _∤∤_ Respectsʳ _≈_
81-
∤∤-respʳ-≈ x≈y z∤∤x z∣∣y = contradiction (∣∣-respʳ-≈ (sym x≈y) z∣∣y) z∤∤x
80+
-respʳ-≈ : __ Respectsʳ _≈_
81+
-respʳ-≈ x≈y zx zy = contradiction (-respʳ-≈ (sym x≈y) zy) zx
8282

83-
∤∤-resp-≈ : _∤∤_ Respects₂ _≈_
84-
∤∤-resp-≈ = ∤∤-respʳ-≈ , ∤∤-respˡ-≈
83+
-resp-≈ : __ Respects₂ _≈_
84+
-resp-≈ = -respʳ-≈ , -respˡ-≈
8585

8686

8787
------------------------------------------------------------------------
@@ -107,3 +107,46 @@ Please use ∣-respʳ-≈ instead. "
107107
"Warning: ∣-resp was deprecated in v2.2.
108108
Please use ∣-resp-≈ instead. "
109109
#-}
110+
111+
-- Version 2.3
112+
113+
∣∣-sym = ∥-sym
114+
{-# WARNING_ON_USAGE ∣∣-sym
115+
"Warning: ∣∣-sym was deprecated in v2.3.
116+
Please use ∥-sym instead. "
117+
#-}
118+
∣∣-respˡ-≈ = ∥-respˡ-≈
119+
{-# WARNING_ON_USAGE ∣∣-respˡ-≈
120+
"Warning: ∣∣-respˡ-≈ was deprecated in v2.3.
121+
Please use ∥-respˡ-≈ instead. "
122+
#-}
123+
∣∣-respʳ-≈ = ∥-respʳ-≈
124+
{-# WARNING_ON_USAGE ∣∣-respʳ-≈
125+
"Warning: ∣∣-respʳ-≈ was deprecated in v2.3.
126+
Please use ∥-respʳ-≈ instead. "
127+
#-}
128+
∣∣-resp-≈ = ∥-resp-≈
129+
{-# WARNING_ON_USAGE ∣∣-resp-≈
130+
"Warning: ∣∣-resp-≈ was deprecated in v2.3.
131+
Please use ∥-resp-≈ instead. "
132+
#-}
133+
∤∤-sym = ∦-sym
134+
{-# WARNING_ON_USAGE ∤∤-sym
135+
"Warning: ∤∤-sym was deprecated in v2.3.
136+
Please use ∦-sym instead. "
137+
#-}
138+
∤∤-respˡ-≈ = ∦-respˡ-≈
139+
{-# WARNING_ON_USAGE ∤∤-respˡ-≈
140+
"Warning: ∤∤-respˡ-≈ was deprecated in v2.3.
141+
Please use ∦-respˡ-≈ instead. "
142+
#-}
143+
∤∤-respʳ-≈ = ∦-respʳ-≈
144+
{-# WARNING_ON_USAGE ∤∤-respʳ-≈
145+
"Warning: ∤∤-respʳ-≈ was deprecated in v2.3.
146+
Please use ∦-respʳ-≈ instead. "
147+
#-}
148+
∤∤-resp-≈ = ∦-resp-≈
149+
{-# WARNING_ON_USAGE ∤∤-resp-≈
150+
"Warning: ∤∤-resp-≈ was deprecated in v2.3.
151+
Please use ∦-resp-≈ instead. "
152+
#-}

src/Algebra/Properties/Monoid/Divisibility.agda

+36-9
Original file line numberDiff line numberDiff line change
@@ -52,15 +52,42 @@ infix 4 ε∣_
5252
------------------------------------------------------------------------
5353
-- Properties of mutual divisibiity
5454

55-
∣∣-refl : Reflexive _∣∣_
56-
∣∣-refl = ∣-refl , ∣-refl
55+
-refl : Reflexive __
56+
-refl = ∣-refl , ∣-refl
5757

58-
∣∣-reflexive : _≈_ ⇒ _∣∣_
59-
∣∣-reflexive x≈y = ∣-reflexive x≈y , ∣-reflexive (sym x≈y)
58+
-reflexive : _≈_ ⇒ __
59+
-reflexive x≈y = ∣-reflexive x≈y , ∣-reflexive (sym x≈y)
6060

61-
∣∣-isEquivalence : IsEquivalence _∣∣_
62-
∣∣-isEquivalence = record
63-
{ refl = λ {x} ∣∣-refl {x}
64-
; sym = λ {x} {y} ∣∣-sym {x} {y}
65-
; trans = λ {x} {y} {z} ∣∣-trans {x} {y} {z}
61+
-isEquivalence : IsEquivalence __
62+
-isEquivalence = record
63+
{ refl = λ {x} -refl {x}
64+
; sym = λ {x} {y} -sym {x} {y}
65+
; trans = λ {x} {y} {z} -trans {x} {y} {z}
6666
}
67+
68+
69+
------------------------------------------------------------------------
70+
-- DEPRECATED NAMES
71+
------------------------------------------------------------------------
72+
-- Please use the new names as continuing support for the old names is
73+
-- not guaranteed.
74+
75+
-- Version 2.3
76+
77+
∣∣-refl = ∥-refl
78+
{-# WARNING_ON_USAGE ∣∣-refl
79+
"Warning: ∣∣-refl was deprecated in v2.3.
80+
Please use ∥-refl instead. "
81+
#-}
82+
83+
∣∣-reflexive = ∥-reflexive
84+
{-# WARNING_ON_USAGE ∣∣-reflexive
85+
"Warning: ∣∣-reflexive was deprecated in v2.3.
86+
Please use ∥-reflexive instead. "
87+
#-}
88+
89+
∣∣-isEquivalence = ∥-isEquivalence
90+
{-# WARNING_ON_USAGE ∣∣-isEquivalence
91+
"Warning: ∣∣-isEquivalence was deprecated in v2.3.
92+
Please use ∥-isEquivalence instead. "
93+
#-}

src/Algebra/Properties/Semigroup/Divisibility.agda

+18-3
Original file line numberDiff line numberDiff line change
@@ -28,7 +28,22 @@ open import Algebra.Properties.Magma.Divisibility magma public
2828
q ∙ p , trans (assoc q p _) (trans (∙-congˡ px≈y) qy≈z)
2929

3030
------------------------------------------------------------------------
31-
-- Properties of _∣∣_
31+
-- Properties of __
3232

33-
∣∣-trans : Transitive _∣∣_
34-
∣∣-trans (x∣y , y∣x) (y∣z , z∣y) = ∣-trans x∣y y∣z , ∣-trans z∣y y∣x
33+
∥-trans : Transitive _∥_
34+
∥-trans (x∣y , y∣x) (y∣z , z∣y) = ∣-trans x∣y y∣z , ∣-trans z∣y y∣x
35+
36+
37+
------------------------------------------------------------------------
38+
-- DEPRECATED NAMES
39+
------------------------------------------------------------------------
40+
-- Please use the new names as continuing support for the old names is
41+
-- not guaranteed.
42+
43+
-- Version 2.3
44+
45+
∣∣-trans = ∥-trans
46+
{-# WARNING_ON_USAGE ∣∣-trans
47+
"Warning: ∣∣-trans was deprecated in v2.3.
48+
Please use ∥-trans instead. "
49+
#-}

0 commit comments

Comments
 (0)