Skip to content

Commit a3235b9

Browse files
MatthewDaggittandreasabel
authored andcommitted
Rename preorder ~ relation reasoning combinators (#2156)
1 parent 3fe4163 commit a3235b9

File tree

10 files changed

+61
-30
lines changed

10 files changed

+61
-30
lines changed

src/Codata/Musical/Colist.agda

+4-2
Original file line numberDiff line numberDiff line change
@@ -190,10 +190,12 @@ module ⊑-Reasoning {a} {A : Set a} where
190190
module ⊆-Reasoning {A : Set a} where
191191
private module Base = PreR (⊆-Preorder A)
192192

193-
open Base public hiding (step-∼)
193+
open Base public
194+
hiding (step-≲; step-∼)
195+
renaming (≲-go to ⊆-go)
194196

195197
open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x begin x)
196-
open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public
198+
open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ -go public
197199

198200

199201
-- take returns a prefix.

src/Data/Integer/Divisibility/Signed.agda

+4-2
Original file line numberDiff line numberDiff line change
@@ -121,9 +121,11 @@ open _∣_ using (quotient) public
121121
module ∣-Reasoning where
122122
private module Base = PreorderReasoning ∣-preorder
123123

124-
open Base public hiding (step-∼; step-≈; step-≈˘)
124+
open Base public
125+
hiding (step-≲; step-∼; step-≈; step-≈˘)
126+
renaming (≲-go to ∣-go)
125127

126-
open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ -go public
128+
open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ -go public
127129

128130
------------------------------------------------------------------------
129131
-- Other properties of _∣_

src/Data/List/Relation/Binary/BagAndSetEquality.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -96,8 +96,8 @@ module ⊆-Reasoning {A : Set a} where
9696
private module Base = PreorderReasoning (⊆-preorder A)
9797

9898
open Base public
99-
hiding (step-≈; step-≈˘; step-∼)
100-
renaming (-go to ⊆-go)
99+
hiding (step-≈; step-≈˘; step-∼; step-≲)
100+
renaming (-go to ⊆-go)
101101

102102
open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x begin x) public
103103
open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ ⊆-go public

src/Data/List/Relation/Binary/Subset/Setoid/Properties.agda

+5-3
Original file line numberDiff line numberDiff line change
@@ -117,11 +117,13 @@ module ⊆-Reasoning (S : Setoid a ℓ) where
117117

118118
private module Base = PreorderReasoning (⊆-preorder S)
119119

120-
open Base public hiding (step-∼; step-≈; step-≈˘)
120+
open Base public
121+
hiding (step-≲; step-≈; step-≈˘; step-∼)
122+
renaming (≲-go to ⊆-go; ≈-go to ≋-go)
121123

122124
open begin-membership-syntax _IsRelatedTo_ _∈_ (λ x Base.begin x) public
123-
open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public
124-
open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ Base.≈-go public
125+
open ⊆-syntax _IsRelatedTo_ _IsRelatedTo_ -go public
126+
open ≋-syntax _IsRelatedTo_ _IsRelatedTo_ -go public
125127

126128
------------------------------------------------------------------------
127129
-- Relationship with other binary relations

src/Data/Nat/Divisibility.agda

+4-2
Original file line numberDiff line numberDiff line change
@@ -120,9 +120,11 @@ suc n ∣? m = Dec.map (m%n≡0⇔n∣m m (suc n)) (m % suc n ≟ 0)
120120
module ∣-Reasoning where
121121
private module Base = PreorderReasoning ∣-preorder
122122

123-
open Base public hiding (step-≈; step-≈˘; step-∼)
123+
open Base public
124+
hiding (step-≈; step-≈˘; step-∼; step-≲)
125+
renaming (≲-go to ∣-go)
124126

125-
open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public
127+
open ∣-syntax _IsRelatedTo_ _IsRelatedTo_ -go public
126128

127129
------------------------------------------------------------------------
128130
-- Simple properties of _∣_

src/Data/Rational/Properties.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -164,7 +164,7 @@ drop-*≡* (*≡* eq) = eq
164164
... | refl = refl
165165

166166
≃-sym : Symmetric _≃_
167-
≃-sym = ≡⇒≃ ∘′ sym ∘′ ≃⇒≡
167+
≃-sym = ≡⇒≃ ∘′ sym ∘′ ≃⇒≡
168168

169169
------------------------------------------------------------------------
170170
-- Properties of ↥
@@ -753,7 +753,7 @@ module ≤-Reasoning where
753753

754754
≃-go : Trans _≃_ _IsRelatedTo_ _IsRelatedTo_
755755
≃-go = Triple.≈-go ∘′ ≃⇒≡
756-
756+
757757
open ≃-syntax _IsRelatedTo_ _IsRelatedTo_ ≃-go (λ {x y} ≃-sym {x} {y}) public
758758

759759
------------------------------------------------------------------------

src/Data/Rational/Unnormalised/Properties.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -164,7 +164,7 @@ p≃0⇒↥p≡0 p (*≡* eq) = begin
164164
↥ p ℤ.* 1ℤ ≡⟨ eq ⟩
165165
0ℤ ∎
166166
where open ≡-Reasoning
167-
167+
168168
↥p≡↥q≡0⇒p≃q : p q ↥ p ≡ 0ℤ ↥ q ≡ 0ℤ p ≃ q
169169
↥p≡↥q≡0⇒p≃q p q ↥p≡0 ↥q≡0 = ≃-trans (↥p≡0⇒p≃0 p ↥p≡0) (≃-sym (↥p≡0⇒p≃0 _ ↥q≡0))
170170

src/Relation/Binary/Construct/Closure/ReflexiveTransitive/Properties.agda

+5-3
Original file line numberDiff line numberDiff line change
@@ -115,7 +115,9 @@ module _ {i t} {I : Set i} (T : Rel I t) where
115115
module StarReasoning {i t} {I : Set i} (T : Rel I t) where
116116
private module Base = PreorderReasoning (preorder T)
117117

118-
open Base public hiding (step-≈; step-∼)
118+
open Base public
119+
hiding (step-≈; step-∼; step-≲)
120+
renaming (≲-go to ⟶-go)
119121

120-
open ⟶-syntax _IsRelatedTo_ _IsRelatedTo_ (Base.∼-go ∘ (_◅ ε)) public
121-
open ⟶*-syntax _IsRelatedTo_ _IsRelatedTo_ Base.∼-go public
122+
open ⟶-syntax _IsRelatedTo_ _IsRelatedTo_ (-go ∘ (_◅ ε)) public
123+
open ⟶*-syntax _IsRelatedTo_ _IsRelatedTo_ -go public

src/Relation/Binary/Reasoning/Base/Double.agda

+24-9
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ open import Relation.Binary.Reasoning.Syntax
2121

2222

2323
module Relation.Binary.Reasoning.Base.Double {a ℓ₁ ℓ₂} {A : Set a}
24-
{_≈_ : Rel A ℓ₁} {__ : Rel A ℓ₂} (isPreorder : IsPreorder _≈_ __)
24+
{_≈_ : Rel A ℓ₁} {__ : Rel A ℓ₂} (isPreorder : IsPreorder _≈_ __)
2525
where
2626

2727
open IsPreorder isPreorder
@@ -32,24 +32,24 @@ open IsPreorder isPreorder
3232
infix 4 _IsRelatedTo_
3333

3434
data _IsRelatedTo_ (x y : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where
35-
nonstrict : (xy : x y) x IsRelatedTo y
35+
nonstrict : (xy : x y) x IsRelatedTo y
3636
equals : (x≈y : x ≈ y) x IsRelatedTo y
3737

38-
start : _IsRelatedTo_ ⇒ __
38+
start : _IsRelatedTo_ ⇒ __
3939
start (equals x≈y) = reflexive x≈y
40-
start (nonstrict xy) = xy
40+
start (nonstrict xy) = xy
4141

4242
≡-go : Trans _≡_ _IsRelatedTo_ _IsRelatedTo_
4343
≡-go x≡y (equals y≈z) = equals (case x≡y of λ where P.refl y≈z)
4444
≡-go x≡y (nonstrict y≤z) = nonstrict (case x≡y of λ where P.refl y≤z)
4545

46-
-go : Trans __ _IsRelatedTo_ _IsRelatedTo_
47-
-go xy (equals y≈z) = nonstrict (∼-respʳ-≈ y≈z xy)
48-
-go xy (nonstrict yz) = nonstrict (trans xy yz)
46+
-go : Trans __ _IsRelatedTo_ _IsRelatedTo_
47+
-go xy (equals y≈z) = nonstrict (∼-respʳ-≈ y≈z xy)
48+
-go xy (nonstrict yz) = nonstrict (trans xy yz)
4949

5050
≈-go : Trans _≈_ _IsRelatedTo_ _IsRelatedTo_
5151
≈-go x≈y (equals y≈z) = equals (Eq.trans x≈y y≈z)
52-
≈-go x≈y (nonstrict yz) = nonstrict (∼-respˡ-≈ (Eq.sym x≈y) yz)
52+
≈-go x≈y (nonstrict yz) = nonstrict (∼-respˡ-≈ (Eq.sym x≈y) yz)
5353

5454
stop : Reflexive _IsRelatedTo_
5555
stop = equals Eq.refl
@@ -81,6 +81,21 @@ equalitySubRelation = record
8181
open begin-syntax _IsRelatedTo_ start public
8282
open begin-equality-syntax _IsRelatedTo_ equalitySubRelation public
8383
open ≡-syntax _IsRelatedTo_ ≡-go public
84-
open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public
8584
open ≈-syntax _IsRelatedTo_ _IsRelatedTo_ ≈-go Eq.sym public
85+
open ≲-syntax _IsRelatedTo_ _IsRelatedTo_ ≲-go public
8686
open end-syntax _IsRelatedTo_ stop public
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.0
96+
97+
open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ≲-go public
98+
{-# WARNING_ON_USAGE step-∼
99+
"Warning: step-∼ and _∼⟨_⟩_ syntax was deprecated in v2.0.
100+
Please use step-≲ and _≲⟨_⟩_ instead. "
101+
#-}

src/Relation/Binary/Reasoning/Syntax.agda

+10-4
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Syntax for the building blocks of equational reasoning modules
4+
-- Syntax for the building blocks of equational reasoning modules
55
------------------------------------------------------------------------
66

77
{-# OPTIONS --cubical-compatible --safe #-}
@@ -26,7 +26,7 @@ open import Relation.Binary.PropositionalEquality.Core as P
2626
-- Codata/Guarded/Stream/Relation/Binary/Pointwise
2727
-- Codata/Sized/Stream/Bisimilarity
2828
-- Function/Reasoning
29-
29+
3030
module Relation.Binary.Reasoning.Syntax where
3131

3232
private
@@ -170,14 +170,20 @@ module _
170170
forward : (x : A) {y z} S y z R x y T x z
171171
forward x yRz x∼y = step {x} x∼y yRz
172172

173-
174-
-- Preorder syntax
173+
-- Arbitrary relation syntax
175174
module ∼-syntax where
176175
infixr 2 step-∼
177176
step-∼ = forward
178177
syntax step-∼ x yRz x∼y = x ∼⟨ x∼y ⟩ yRz
179178

180179

180+
-- Preorder syntax
181+
module ≲-syntax where
182+
infixr 2 step-≲
183+
step-≲ = forward
184+
syntax step-≲ x yRz x≲y = x ≲⟨ x≲y ⟩ yRz
185+
186+
181187
-- Partial order syntax
182188
module ≤-syntax where
183189
infixr 2 step-≤

0 commit comments

Comments
 (0)