forked from agda/agda-stdlib
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathProperties.agda
147 lines (116 loc) · 6.66 KB
/
Properties.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
------------------------------------------------------------------------
-- The Agda standard library
--
-- Properties of the membership relation for AVL Maps identifying values up to
-- propositional equality.
------------------------------------------------------------------------
{-# OPTIONS --cubical-compatible --safe #-}
open import Relation.Binary using (StrictTotalOrder)
module Data.Tree.AVL.Map.Membership.Propositional.Properties
{a ℓ₁ ℓ₂} (strictTotalOrder : StrictTotalOrder a ℓ₁ ℓ₂)
where
open import Data.Bool.Base using (true; false)
open import Data.Maybe.Base using (just; nothing; is-just)
open import Data.Product as Prod using (_×_; ∃-syntax; _,_; proj₁; proj₂)
open import Data.Product.Relation.Binary.Pointwise.NonDependent renaming (Pointwise to _×ᴿ_)
open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
open import Function.Base using (_∘_; _∘′_)
open import Level using (Level)
open import Relation.Binary using (Transitive; Symmetric; _Respectsˡ_)
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Construct.Intersection using (_∩_)
open import Relation.Binary.PropositionalEquality
using (_≡_; cong) renaming (refl to ≡-refl; sym to ≡-sym; trans to ≡-trans)
open import Relation.Nullary using (Reflects; ¬_; yes; no)
open import Relation.Nullary.Negation using (contradiction)
open StrictTotalOrder strictTotalOrder renaming (Carrier to Key) hiding (trans)
open Eq using (_≉_; refl; sym; trans)
open import Data.Tree.AVL strictTotalOrder using (tree)
open import Data.Tree.AVL.Indexed strictTotalOrder using (key)
import Data.Tree.AVL.Indexed.Relation.Unary.Any strictTotalOrder as IAny
import Data.Tree.AVL.Indexed.Relation.Unary.Any.Properties strictTotalOrder as IAnyₚ
open import Data.Tree.AVL.Key strictTotalOrder using (⊥⁺<[_]<⊤⁺)
open import Data.Tree.AVL.Map strictTotalOrder
open import Data.Tree.AVL.Map.Relation.Unary.Any strictTotalOrder as Mapₚ using (Any)
open import Data.Tree.AVL.Map.Membership.Propositional strictTotalOrder
open import Data.Tree.AVL.Relation.Unary.Any strictTotalOrder as Any using (tree)
private
variable
v p q : Level
V : Set v
m : Map V
k k′ : Key
x x′ y y′ : V
kx : Key × V
≈ₖᵥ-trans : Transitive (_≈ₖᵥ_ {V = V})
≈ₖᵥ-trans {i = i} {k = k} = ×-transitive Eq.trans ≡-trans {i = i} {k = k}
≈ₖᵥ-sym : Symmetric (_≈ₖᵥ_ {V = V})
≈ₖᵥ-sym {x = x} {y = y} = ×-symmetric sym ≡-sym {x} {y}
∈ₖᵥ-Respectsˡ : _∈ₖᵥ_ {V = V} Respectsˡ _≈ₖᵥ_
∈ₖᵥ-Respectsˡ x~y = Any.map (≈ₖᵥ-trans (≈ₖᵥ-sym x~y))
∈ₖᵥ-empty : kx ∉ₖᵥ empty
∈ₖᵥ-empty (tree ())
------------------------------------------------------------------------
-- singleton
∈ₖᵥ-singleton⁺ : (k , x) ∈ₖᵥ singleton k x
∈ₖᵥ-singleton⁺ = tree (IAnyₚ.singleton⁺ _ _ _ (refl , ≡-refl))
∈ₖᵥ-singleton⁻ : (k , x) ∈ₖᵥ singleton k′ x′ → k ≈ k′ × x ≡ x′
∈ₖᵥ-singleton⁻ (tree p) = IAnyₚ.singleton⁻ _ _ _ p
private
≈-lookup : (p : (k , x) ∈ₖᵥ m) → k ≈ Any.lookupKey p
≈-lookup (tree p) = proj₁ (IAnyₚ.lookup-result p)
------------------------------------------------------------------------
-- insert
∈ₖᵥ-insert⁺ : k ≉ k′ → (k , x) ∈ₖᵥ m → (k , x) ∈ₖᵥ insert k′ x′ m
∈ₖᵥ-insert⁺ {k′ = k′} {m = m@(tree t)} k≉k′ (tree p) =
tree (IAnyₚ.insert⁺ _ _ _ (⊥⁺<[ _ ]<⊤⁺) p k′≉key-p)
where
k′≉key-p : k′ ≉ Any.lookupKey (tree p)
k′≉key-p k′≈key-p = k≉k′ (Eq.trans (≈-lookup (tree p)) (Eq.sym k′≈key-p))
∈ₖᵥ-insert⁺⁺ : (k , x) ∈ₖᵥ insert k x m
∈ₖᵥ-insert⁺⁺ {k = k} {m = tree t} with IAny.any? ((k ≟_) ∘ key) t
... | yes k∈ = tree (IAnyₚ.Any-insert-just _ _ _ _ (λ k′ → _, ≡-refl) k∈)
... | no ¬k∈ = tree (IAnyₚ.Any-insert-nothing _ _ _ _ (refl , ≡-refl) ¬k∈)
private
≈ₖᵥ-Resp : k ≈ k′ → kx ≈ₖᵥ (k′ , x) → kx ≈ₖᵥ (k , x)
≈ₖᵥ-Resp = (λ{ k′≈l (k≈l , x≡) → (trans k≈l (sym k′≈l) , x≡)})
∈ₖᵥ-insert⁻ : (k , x) ∈ₖᵥ insert k′ x′ m → (k ≈ k′ × x ≡ x′) ⊎ (k ≉ k′ × (k , x) ∈ₖᵥ m)
∈ₖᵥ-insert⁻ {m = tree t} (tree kx∈insert)
with IAnyₚ.insert⁻ ≈ₖᵥ-Resp _ _ t (⊥⁺<[ _ ]<⊤⁺) kx∈insert
... | inj₁ p = inj₁ p
... | inj₂ p = let k′≉p , k≈p , _ = IAnyₚ.lookup-result p
k≉k′ = λ k≈k′ → k′≉p (trans (sym k≈k′) k≈p)
in inj₂ (k≉k′ , tree (IAny.map proj₂ p))
------------------------------------------------------------------------
-- lookup
∈ₖᵥ-lookup⁺ : (k , x) ∈ₖᵥ m → lookup m k ≡ just x
∈ₖᵥ-lookup⁺ {k = k} {m = tree t} (tree kx∈m)
with IAnyₚ.lookup⁺ t k (⊥⁺<[ _ ]<⊤⁺) kx∈m | IAnyₚ.lookup-result kx∈m
... | inj₁ p≉k | k≈p , x≡p = contradiction (sym k≈p) p≉k
... | inj₂ (p≈k , eq) | k≈p , x≡p = ≡-trans eq (cong just (≡-sym x≡p))
∈ₖᵥ-lookup⁻ : lookup m k ≡ just x → (k , x) ∈ₖᵥ m
∈ₖᵥ-lookup⁻ {m = tree t} {k = k} {x = x} eq
= tree (IAny.map (Prod.map sym ≡-sym) (IAnyₚ.lookup⁻ t k x (⊥⁺<[ _ ]<⊤⁺) eq))
∈ₖᵥ-lookup-nothing⁺ : (∀ x → (k , x) ∉ₖᵥ m) → lookup m k ≡ nothing
∈ₖᵥ-lookup-nothing⁺ {k = k} {m = m@(tree t)} k∉m with lookup m k in eq
... | nothing = ≡-refl
... | just x = contradiction (∈ₖᵥ-lookup⁻ eq) (k∉m x)
∈ₖᵥ-lookup-nothing⁻ : lookup m k ≡ nothing → (k , x) ∉ₖᵥ m
∈ₖᵥ-lookup-nothing⁻ eq kx∈m with ≡-trans (≡-sym eq) (∈ₖᵥ-lookup⁺ kx∈m)
... | ()
------------------------------------------------------------------------
-- member
∈ₖᵥ-member : (k , x) ∈ₖᵥ m → member k m ≡ true
∈ₖᵥ-member = cong is-just ∘ ∈ₖᵥ-lookup⁺
∉ₖᵥ-member : (∀ x → (k , x) ∉ₖᵥ m) → member k m ≡ false
∉ₖᵥ-member = cong is-just ∘ ∈ₖᵥ-lookup-nothing⁺
member-∈ₖᵥ : member k m ≡ true → ∃[ x ] (k , x) ∈ₖᵥ m
member-∈ₖᵥ {k = k} {m = m} ≡true with lookup m k in eq
... | just x = x , ∈ₖᵥ-lookup⁻ eq
member-∉ₖᵥ : member k m ≡ false → ∀ x → (k , x) ∉ₖᵥ m
member-∉ₖᵥ {k = k} {m = m} ≡false x with lookup m k in eq
... | nothing = ∈ₖᵥ-lookup-nothing⁻ eq
member-Reflects-∈ₖᵥ : Reflects (∃[ x ] (k , x) ∈ₖᵥ m) (member k m)
member-Reflects-∈ₖᵥ {k = k} {m = m} with lookup m k in eq
... | just x = Reflects.ofʸ (x , ∈ₖᵥ-lookup⁻ eq)
... | nothing = Reflects.ofⁿ (∈ₖᵥ-lookup-nothing⁻ eq ∘ proj₂)