forked from agda/agda-stdlib
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathTransitive.agda
158 lines (117 loc) · 5.12 KB
/
Transitive.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
------------------------------------------------------------------------
-- The Agda standard library
--
-- Transitive closures
------------------------------------------------------------------------
{-# OPTIONS --cubical-compatible --safe #-}
module Relation.Binary.Construct.Closure.Transitive where
open import Function.Base
open import Function.Bundles using (_⇔_; mk⇔)
open import Induction.WellFounded
open import Level
open import Relation.Binary.Core using (Rel; _=[_]⇒_; _⇒_)
open import Relation.Binary.Definitions
using (Reflexive; Symmetric; Transitive)
open import Relation.Binary.PropositionalEquality.Core as P using (_≡_)
private
variable
a ℓ ℓ₁ ℓ₂ : Level
A B : Set a
------------------------------------------------------------------------
-- Definition
infixr 5 _∷_
infix 4 TransClosure
data TransClosure {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
[_] : ∀ {x y} (x∼y : x ∼ y) → TransClosure _∼_ x y
_∷_ : ∀ {x y z} (x∼y : x ∼ y) (y∼⁺z : TransClosure _∼_ y z) → TransClosure _∼_ x z
syntax TransClosure R x y = x ⟨ R ⟩⁺ y
------------------------------------------------------------------------
-- Operations
module _ {_∼_ : Rel A ℓ} where
private
_∼⁺_ = TransClosure _∼_
infixr 5 _∷ʳ_
_∷ʳ_ : ∀ {x y z} → (x∼⁺y : x ∼⁺ y) (y∼z : y ∼ z) → x ∼⁺ z
[ x∼y ] ∷ʳ y∼z = x∼y ∷ [ y∼z ]
(x∼y ∷ x∼⁺y) ∷ʳ y∼z = x∼y ∷ (x∼⁺y ∷ʳ y∼z)
infixr 5 _++_
_++_ : ∀ {x y z} → x ∼⁺ y → y ∼⁺ z → x ∼⁺ z
[ x∼y ] ++ y∼⁺z = x∼y ∷ y∼⁺z
(x∼y ∷ y∼⁺z) ++ z∼⁺u = x∼y ∷ (y∼⁺z ++ z∼⁺u)
------------------------------------------------------------------------
-- Properties
module _ (_∼_ : Rel A ℓ) where
private
_∼⁺_ = TransClosure _∼_
module ∼⊆∼⁺ = Subrelation {_<₂_ = _∼⁺_} [_]
reflexive : Reflexive _∼_ → Reflexive _∼⁺_
reflexive refl = [ refl ]
symmetric : Symmetric _∼_ → Symmetric _∼⁺_
symmetric sym [ x∼y ] = [ sym x∼y ]
symmetric sym (x∼y ∷ y∼⁺z) = symmetric sym y∼⁺z ∷ʳ sym x∼y
transitive : Transitive _∼⁺_
transitive = _++_
accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x
accessible⁻ = ∼⊆∼⁺.accessible
wellFounded⁻ : WellFounded _∼⁺_ → WellFounded _∼_
wellFounded⁻ = ∼⊆∼⁺.wellFounded
accessible : ∀ {x} → Acc _∼_ x → Acc _∼⁺_ x
accessible acc[x] = acc (wf-acc acc[x])
where
wf-acc : ∀ {x} → Acc _∼_ x → WfRec _∼⁺_ (Acc _∼⁺_) x
wf-acc (acc rec) _ [ y∼x ] = acc (wf-acc (rec _ y∼x))
wf-acc acc[x] _ (y∼z ∷ z∼⁺x) = acc-inverse (wf-acc acc[x] _ z∼⁺x) _ [ y∼z ]
wellFounded : WellFounded _∼_ → WellFounded _∼⁺_
wellFounded wf x = accessible (wf x)
------------------------------------------------------------------------
-- Alternative definition of transitive closure
infix 4 Plus
syntax Plus R x y = x [ R ]⁺ y
data Plus {A : Set a} (_∼_ : Rel A ℓ) : Rel A (a ⊔ ℓ) where
[_] : ∀ {x y} (x∼y : x ∼ y) → x [ _∼_ ]⁺ y
_∼⁺⟨_⟩_ : ∀ x {y z} (x∼⁺y : x [ _∼_ ]⁺ y) (y∼⁺z : y [ _∼_ ]⁺ z) →
x [ _∼_ ]⁺ z
module _ {_∼_ : Rel A ℓ} where
[]-injective : ∀ {x y p q} → (x [ _∼_ ]⁺ y ∋ [ p ]) ≡ [ q ] → p ≡ q
[]-injective P.refl = P.refl
-- See also ∼⁺⟨⟩-injectiveˡ and ∼⁺⟨⟩-injectiveʳ in
-- Relation.Binary.Construct.Closure.Transitive.WithK.
-- "Equational" reasoning notation. Example:
--
-- lemma =
-- x ∼⁺⟨ [ lemma₁ ] ⟩
-- y ∼⁺⟨ lemma₂ ⟩∎
-- z ∎
finally : ∀ {_∼_ : Rel A ℓ} x y → x [ _∼_ ]⁺ y → x [ _∼_ ]⁺ y
finally _ _ = id
syntax finally x y x∼⁺y = x ∼⁺⟨ x∼⁺y ⟩∎ y ∎
infixr 2 _∼⁺⟨_⟩_
infix 3 finally
-- Map.
map : {_R₁_ : Rel A ℓ} {_R₂_ : Rel B ℓ₂} {f : A → B} →
_R₁_ =[ f ]⇒ _R₂_ → Plus _R₁_ =[ f ]⇒ Plus _R₂_
map R₁⇒R₂ [ xRy ] = [ R₁⇒R₂ xRy ]
map R₁⇒R₂ (x ∼⁺⟨ xR⁺z ⟩ zR⁺y) =
_ ∼⁺⟨ map R₁⇒R₂ xR⁺z ⟩ map R₁⇒R₂ zR⁺y
-- Plus and TransClosure are equivalent.
equivalent : ∀ {_∼_ : Rel A ℓ} {x y} →
Plus _∼_ x y ⇔ TransClosure _∼_ x y
equivalent {_∼_ = _∼_} = mk⇔ complete sound
where
complete : Plus _∼_ ⇒ TransClosure _∼_
complete [ x∼y ] = [ x∼y ]
complete (x ∼⁺⟨ x∼⁺y ⟩ y∼⁺z) = complete x∼⁺y ++ complete y∼⁺z
sound : TransClosure _∼_ ⇒ Plus _∼_
sound [ x∼y ] = [ x∼y ]
sound (x∼y ∷ y∼⁺z) = _ ∼⁺⟨ [ x∼y ] ⟩ sound y∼⁺z
------------------------------------------------------------------------
-- Deprecations
------------------------------------------------------------------------
-- Please use the new names as continuing support for the old names is
-- not guaranteed.
-- v1.5
Plus′ = TransClosure
{-# WARNING_ON_USAGE Plus′
"Warning: Plus′ was deprecated in v1.5.
Please use TransClosure instead."
#-}