Skip to content

Commit 3880dec

Browse files
authored
[add] Update PartialSetoid reasoning (#2689)
* Update PartialSetoid reasoning * Add changelog entry for this PR * Fix Changelog link * Move Changelog entry
1 parent 8b8946a commit 3880dec

File tree

2 files changed

+32
-23
lines changed

2 files changed

+32
-23
lines changed

CHANGELOG.md

+6
Original file line numberDiff line numberDiff line change
@@ -22,6 +22,12 @@ Bug-fixes
2222
This has been deprecated in favor or `rightInverse`, and a corrected (and
2323
correctly-named) function `leftInverse` has been added.
2424

25+
* The implementation of `_IsRelatedTo_` in `Relation.Binary.Reasoning.Base.Partial`
26+
has been modified to correctly support equational reasoning at the beginning and the end.
27+
The detail of this issue is described in [#2677](https://github.com/agda/agda-stdlib/pull/2677). Since the names of constructors
28+
of `_IsRelatedTo_` are changed and the reduction behaviour of reasoning steps
29+
are changed, this modification is non-backwards compatible.
30+
2531
Non-backwards compatible changes
2632
--------------------------------
2733

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

+26-23
Original file line numberDiff line numberDiff line change
@@ -22,52 +22,55 @@ module Relation.Binary.Reasoning.Base.Partial
2222
------------------------------------------------------------------------
2323
-- Definition of "related to"
2424

25-
-- This seemingly unnecessary type is used to make it possible to
26-
-- infer arguments even if the underlying equality evaluates.
25+
-- This type allows us to track whether reasoning steps
26+
-- include _∼_ or not.
2727

2828
infix 4 _IsRelatedTo_
2929

3030
data _IsRelatedTo_ : A A Set (a ⊔ ℓ) where
31-
singleStep : x x IsRelatedTo x
32-
multiStep : {x y} (x∼y : x ∼ y) x IsRelatedTo y
31+
reflexive : {x y} x ≡ y x IsRelatedTo y
32+
relTo : {x y} (x∼y : x ∼ y) x IsRelatedTo y
33+
34+
≡-go : Trans _≡_ _IsRelatedTo_ _IsRelatedTo_
35+
≡-go x≡y (reflexive y≡z) = reflexive (case x≡y of λ where ≡.refl y≡z)
36+
≡-go x≡y (relTo y∼z) = relTo (case x≡y of λ where ≡.refl y∼z)
3337

3438
∼-go : Trans _∼_ _IsRelatedTo_ _IsRelatedTo_
35-
∼-go x∼y (singleStep y) = multiStep x∼y
36-
∼-go x∼y (multiStep y∼z) = multiStep (trans x∼y y∼z)
39+
∼-go x∼y (reflexive y≡z) = relTo (case y≡z of λ where ≡.refl x∼y)
40+
∼-go x∼y (relTo y∼z) = relTo (trans x∼y y∼z)
3741

3842
stop : Reflexive _IsRelatedTo_
39-
stop = singleStep _
43+
stop = reflexive ≡.refl
4044

4145
------------------------------------------------------------------------
4246
-- Types that are used to ensure that the final relation proved by the
4347
-- chain of reasoning can be converted into the required relation.
4448

45-
data IsMultiStep {x y} : x IsRelatedTo y Set (a ⊔ ℓ) where
46-
isMultiStep : x∼y IsMultiStep (multiStep x∼y)
49+
data IsRelTo {x y} : x IsRelatedTo y Set (a ⊔ ℓ) where
50+
isRelTo : x∼y IsRelTo (relTo x∼y)
4751

48-
IsMultiStep? : {x y} (x∼y : x IsRelatedTo y) Dec (IsMultiStep x∼y)
49-
IsMultiStep? (multiStep x<y) = yes (isMultiStep x<y)
50-
IsMultiStep? (singleStep _) = no λ()
52+
IsRelTo? : {x y} (x∼y : x IsRelatedTo y) Dec (IsRelTo x∼y)
53+
IsRelTo? (relTo x∼y) = yes (isRelTo x∼y)
54+
IsRelTo? (reflexive _) = no λ()
5155

52-
extractMultiStep : {x y} {x∼y : x IsRelatedTo y} IsMultiStep x∼y x ∼ y
53-
extractMultiStep (isMultiStep x≈y) = xy
56+
extractRelTo : {x y} {x∼y : x IsRelatedTo y} IsRelTo x∼y x ∼ y
57+
extractRelTo (isRelTo x∼y) = xy
5458

55-
multiStepSubRelation : SubRelation _IsRelatedTo_ _ _
56-
multiStepSubRelation = record
57-
{ IsS = IsMultiStep
58-
; IsS? = IsMultiStep?
59-
; extract = extractMultiStep
59+
relToSubRelation : SubRelation _IsRelatedTo_ _ _
60+
relToSubRelation = record
61+
{ IsS = IsRelTo
62+
; IsS? = IsRelTo?
63+
; extract = extractRelTo
6064
}
6165

6266
------------------------------------------------------------------------
6367
-- Reasoning combinators
6468

65-
open begin-subrelation-syntax _IsRelatedTo_ multiStepSubRelation public
66-
open ≡-noncomputing-syntax _IsRelatedTo_ public
69+
open begin-subrelation-syntax _IsRelatedTo_ relToSubRelation public
70+
open ≡-syntax _IsRelatedTo_ ≡-go public
6771
open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public
6872
open end-syntax _IsRelatedTo_ stop public
6973

70-
7174
------------------------------------------------------------------------
7275
-- DEPRECATED NAMES
7376
------------------------------------------------------------------------
@@ -79,7 +82,7 @@ open end-syntax _IsRelatedTo_ stop public
7982
infix 3 _∎⟨_⟩
8083

8184
_∎⟨_⟩ : x x ∼ x x IsRelatedTo x
82-
_ ∎⟨ x∼x ⟩ = multiStep x∼x
85+
_ ∎⟨ x∼x ⟩ = relTo x∼x
8386
{-# WARNING_ON_USAGE _∎⟨_⟩
8487
"Warning: _∎⟨_⟩ was deprecated in v1.6.
8588
Please use _∎ instead if used in a chain, otherwise simply provide

0 commit comments

Comments
 (0)