@@ -21,7 +21,7 @@ open import Relation.Binary.Reasoning.Syntax
2121
2222
2323module 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
2727open IsPreorder isPreorder
@@ -32,24 +32,24 @@ open IsPreorder isPreorder
3232infix 4 _IsRelatedTo_
3333
3434data _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
3636 equals : (x≈y : x ≈ y) → x IsRelatedTo y
3737
38- start : _IsRelatedTo_ ⇒ _∼ _
38+ start : _IsRelatedTo_ ⇒ _≲ _
3939start (equals x≈y) = reflexive x≈y
40- start (nonstrict x∼ y) = x∼ y
40+ start (nonstrict x≲ y) = x≲ y
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 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)
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 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)
5353
5454stop : Reflexive _IsRelatedTo_
5555stop = equals Eq.refl
@@ -81,6 +81,21 @@ equalitySubRelation = record
8181open begin-syntax _IsRelatedTo_ start public
8282open begin-equality-syntax _IsRelatedTo_ equalitySubRelation public
8383open ≡-syntax _IsRelatedTo_ ≡-go public
84- open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public
8584open ≈-syntax _IsRelatedTo_ _IsRelatedTo_ ≈-go Eq.sym public
85+ open ≲-syntax _IsRelatedTo_ _IsRelatedTo_ ≲-go public
8686open 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