@@ -22,52 +22,55 @@ module Relation.Binary.Reasoning.Base.Partial
22
22
------------------------------------------------------------------------
23
23
-- Definition of "related to"
24
24
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 .
27
27
28
28
infix 4 _IsRelatedTo_
29
29
30
30
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)
33
37
34
38
∼-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)
37
41
38
42
stop : Reflexive _IsRelatedTo_
39
- stop = singleStep _
43
+ stop = reflexive ≡.refl
40
44
41
45
------------------------------------------------------------------------
42
46
-- Types that are used to ensure that the final relation proved by the
43
47
-- chain of reasoning can be converted into the required relation.
44
48
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)
47
51
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 λ ()
51
55
52
- extractMultiStep : ∀ {x y} {x∼y : x IsRelatedTo y} → IsMultiStep x∼y → x ∼ y
53
- extractMultiStep (isMultiStep x≈ y) = x≈ y
56
+ extractRelTo : ∀ {x y} {x∼y : x IsRelatedTo y} → IsRelTo x∼y → x ∼ y
57
+ extractRelTo (isRelTo x∼ y) = x∼ y
54
58
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
60
64
}
61
65
62
66
------------------------------------------------------------------------
63
67
-- Reasoning combinators
64
68
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
67
71
open ∼-syntax _IsRelatedTo_ _IsRelatedTo_ ∼-go public
68
72
open end-syntax _IsRelatedTo_ stop public
69
73
70
-
71
74
------------------------------------------------------------------------
72
75
-- DEPRECATED NAMES
73
76
------------------------------------------------------------------------
@@ -79,7 +82,7 @@ open end-syntax _IsRelatedTo_ stop public
79
82
infix 3 _∎⟨_⟩
80
83
81
84
_∎⟨_⟩ : ∀ x → x ∼ x → x IsRelatedTo x
82
- _ ∎⟨ x∼x ⟩ = multiStep x∼x
85
+ _ ∎⟨ x∼x ⟩ = relTo x∼x
83
86
{-# WARNING_ON_USAGE _∎⟨_⟩
84
87
"Warning: _∎⟨_⟩ was deprecated in v1.6.
85
88
Please use _∎ instead if used in a chain, otherwise simply provide
0 commit comments