@@ -21,7 +21,7 @@ open import Relation.Binary.Reasoning.Syntax
21
21
22
22
23
23
module Relation.Binary.Reasoning.Base.Double {a ℓ₁ ℓ₂} {A : Set a}
24
- {_≈_ : Rel A ℓ₁} {_∼ _ : Rel A ℓ₂} (isPreorder : IsPreorder _≈_ _∼ _)
24
+ {_≈_ : Rel A ℓ₁} {_≲ _ : Rel A ℓ₂} (isPreorder : IsPreorder _≈_ _≲ _)
25
25
where
26
26
27
27
open IsPreorder isPreorder
@@ -32,24 +32,24 @@ open IsPreorder isPreorder
32
32
infix 4 _IsRelatedTo_
33
33
34
34
data _IsRelatedTo_ (x y : A) : Set (a ⊔ ℓ₁ ⊔ ℓ₂) where
35
- nonstrict : (x∼ y : x ∼ y) → x IsRelatedTo y
35
+ nonstrict : (x≲ y : x ≲ y) → x IsRelatedTo y
36
36
equals : (x≈y : x ≈ y) → x IsRelatedTo y
37
37
38
- start : _IsRelatedTo_ ⇒ _∼ _
38
+ start : _IsRelatedTo_ ⇒ _≲ _
39
39
start (equals x≈y) = reflexive x≈y
40
- start (nonstrict x∼ y) = x∼ y
40
+ start (nonstrict x≲ y) = x≲ y
41
41
42
42
≡-go : Trans _≡_ _IsRelatedTo_ _IsRelatedTo_
43
43
≡-go x≡y (equals y≈z) = equals (case x≡y of λ where P.refl → y≈z)
44
44
≡-go x≡y (nonstrict y≤z) = nonstrict (case x≡y of λ where P.refl → y≤z)
45
45
46
- ∼ -go : Trans _∼ _ _IsRelatedTo_ _IsRelatedTo_
47
- ∼ -go x∼ y (equals y≈z) = nonstrict (∼-respʳ-≈ y≈z x∼ y)
48
- ∼ -go x∼ y (nonstrict y∼ z) = nonstrict (trans x∼ y y∼ z)
46
+ ≲ -go : Trans _≲ _ _IsRelatedTo_ _IsRelatedTo_
47
+ ≲ -go x≲ y (equals y≈z) = nonstrict (∼-respʳ-≈ y≈z x≲ y)
48
+ ≲ -go x≲ y (nonstrict y≲ z) = nonstrict (trans x≲ y y≲ z)
49
49
50
50
≈-go : Trans _≈_ _IsRelatedTo_ _IsRelatedTo_
51
51
≈-go x≈y (equals y≈z) = equals (Eq.trans x≈y y≈z)
52
- ≈-go x≈y (nonstrict y∼ z) = nonstrict (∼-respˡ-≈ (Eq.sym x≈y) y∼ z)
52
+ ≈-go x≈y (nonstrict y≲ z) = nonstrict (∼-respˡ-≈ (Eq.sym x≈y) y≲ z)
53
53
54
54
stop : Reflexive _IsRelatedTo_
55
55
stop = equals Eq.refl
@@ -81,6 +81,21 @@ equalitySubRelation = record
81
81
open begin-syntax _IsRelatedTo_ start public
82
82
open begin-equality-syntax _IsRelatedTo_ equalitySubRelation public
83
83
open ≡-syntax _IsRelatedTo_ ≡-go public
84
- open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public
85
84
open ≈-syntax _IsRelatedTo_ _IsRelatedTo_ ≈-go Eq.sym public
85
+ open ≲-syntax _IsRelatedTo_ _IsRelatedTo_ ≲-go public
86
86
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
+ #-}
0 commit comments