Mercurial > hg > Gears > GearsAgda
annotate hoareBinaryTree1.agda @ 778:4d71d0894cfa
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 22 May 2023 19:06:08 +0900 |
parents | 8e5159a02b76 |
children | 7e5dfe642064 0b791ae19543 |
rev | line source |
---|---|
722 | 1 module hoareBinaryTree1 where |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
2 |
727 | 3 open import Level hiding (suc ; zero ; _⊔_ ) |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
4 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
5 open import Data.Nat hiding (compare) |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
6 open import Data.Nat.Properties as NatProp |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
7 open import Data.Maybe |
588 | 8 -- open import Data.Maybe.Properties |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
9 open import Data.Empty |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
10 open import Data.List |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
11 open import Data.Product |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
12 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
13 open import Function as F hiding (const) |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
14 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
15 open import Relation.Binary |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
16 open import Relation.Binary.PropositionalEquality |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
17 open import Relation.Nullary |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
18 open import logic |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
19 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
20 |
590 | 21 -- |
22 -- | |
23 -- no children , having left node , having right node , having both | |
24 -- | |
597 | 25 data bt {n : Level} (A : Set n) : Set n where |
604 | 26 leaf : bt A |
27 node : (key : ℕ) → (value : A) → | |
610 | 28 (left : bt A ) → (right : bt A ) → bt A |
600 | 29 |
620 | 30 node-key : {n : Level} {A : Set n} → bt A → Maybe ℕ |
31 node-key (node key _ _ _) = just key | |
32 node-key _ = nothing | |
33 | |
34 node-value : {n : Level} {A : Set n} → bt A → Maybe A | |
35 node-value (node _ value _ _) = just value | |
36 node-value _ = nothing | |
37 | |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
38 bt-depth : {n : Level} {A : Set n} → (tree : bt A ) → ℕ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
39 bt-depth leaf = 0 |
727 | 40 bt-depth (node key value t t₁) = suc (bt-depth t ⊔ bt-depth t₁ ) |
606 | 41 |
605 | 42 open import Data.Unit hiding ( _≟_ ; _≤?_ ; _≤_) |
43 | |
620 | 44 data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where |
45 t-leaf : treeInvariant leaf | |
632 | 46 t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) |
745 | 47 t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key₁ value₁ t₁ t₂) |
632 | 48 → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) |
745 | 49 t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key value t₁ t₂) |
632 | 50 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) |
745 | 51 t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → key < key₁ → key₁ < key₂ |
620 | 52 → treeInvariant (node key value t₁ t₂) |
53 → treeInvariant (node key₂ value₂ t₃ t₄) | |
54 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) | |
605 | 55 |
662 | 56 -- |
745 | 57 -- stack always contains original top at end (path of the tree) |
662 | 58 -- |
59 data stackInvariant {n : Level} {A : Set n} (key : ℕ) : (top orig : bt A) → (stack : List (bt A)) → Set n where | |
729 | 60 s-nil : {tree0 : bt A} → stackInvariant key tree0 tree0 (tree0 ∷ []) |
771 | 61 s-right : (tree tree0 tree₁ : bt A) → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 62 → key₁ < key → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
771 | 63 s-left : (tree₁ tree0 tree : bt A) → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 64 → key < key₁ → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree₁ tree0 (tree₁ ∷ st) |
639 | 65 |
677 | 66 data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt A ) → Set n where |
639 | 67 r-leaf : replacedTree key value leaf (node key value leaf leaf) |
68 r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) | |
69 r-right : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} | |
677 | 70 → k < key → replacedTree key value t2 t → replacedTree key value (node k v1 t1 t2) (node k v1 t1 t) |
639 | 71 r-left : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} |
687 | 72 → key < k → replacedTree key value t1 t → replacedTree key value (node k v1 t1 t2) (node k v1 t t2) |
652 | 73 |
632 | 74 add< : { i : ℕ } (j : ℕ ) → i < suc i + j |
75 add< {i} j = begin | |
76 suc i ≤⟨ m≤m+n (suc i) j ⟩ | |
77 suc i + j ∎ where open ≤-Reasoning | |
78 | |
79 treeTest1 : bt ℕ | |
692 | 80 treeTest1 = node 0 0 leaf (node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf)) |
632 | 81 treeTest2 : bt ℕ |
692 | 82 treeTest2 = node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf) |
632 | 83 |
84 treeInvariantTest1 : treeInvariant treeTest1 | |
692 | 85 treeInvariantTest1 = t-right (m≤m+n _ 2) (t-node (add< 0) (add< 1) (t-left (add< 0) (t-single 1 7)) (t-single 5 5) ) |
605 | 86 |
639 | 87 stack-top : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
88 stack-top [] = nothing | |
89 stack-top (x ∷ s) = just x | |
606 | 90 |
639 | 91 stack-last : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
92 stack-last [] = nothing | |
93 stack-last (x ∷ []) = just x | |
94 stack-last (x ∷ s) = stack-last s | |
632 | 95 |
662 | 96 stackInvariantTest1 : stackInvariant 4 treeTest2 treeTest1 ( treeTest2 ∷ treeTest1 ∷ [] ) |
771 | 97 stackInvariantTest1 = s-right _ _ _ (add< 3) (s-nil ) |
662 | 98 |
666 | 99 si-property0 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 stack → ¬ ( stack ≡ [] ) |
729 | 100 si-property0 (s-nil ) () |
771 | 101 si-property0 (s-right _ _ _ x si) () |
102 si-property0 (s-left _ _ _ x si) () | |
665 | 103 |
666 | 104 si-property1 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 tree1 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 (tree1 ∷ stack) |
105 → tree1 ≡ tree | |
729 | 106 si-property1 (s-nil ) = refl |
771 | 107 si-property1 (s-right _ _ _ _ si) = refl |
108 si-property1 (s-left _ _ _ _ si) = refl | |
662 | 109 |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
110 si-property-last : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack |
662 | 111 → stack-last stack ≡ just tree0 |
729 | 112 si-property-last key t t0 (t ∷ []) (s-nil ) = refl |
771 | 113 si-property-last key t t0 (.t ∷ x ∷ st) (s-right _ _ _ _ si ) with si-property1 si |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
114 ... | refl = si-property-last key x t0 (x ∷ st) si |
771 | 115 si-property-last key t t0 (.t ∷ x ∷ st) (s-left _ _ _ _ si ) with si-property1 si |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
116 ... | refl = si-property-last key x t0 (x ∷ st) si |
656 | 117 |
639 | 118 rt-property1 : {n : Level} {A : Set n} (key : ℕ) (value : A) (tree tree1 : bt A ) → replacedTree key value tree tree1 → ¬ ( tree1 ≡ leaf ) |
119 rt-property1 {n} {A} key value .leaf .(node key value leaf leaf) r-leaf () | |
120 rt-property1 {n} {A} key value .(node key _ _ _) .(node key value _ _) r-node () | |
677 | 121 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-right x rt) = λ () |
122 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-left x rt) = λ () | |
639 | 123 |
690 | 124 rt-property-leaf : {n : Level} {A : Set n} {key : ℕ} {value : A} {repl : bt A} → replacedTree key value leaf repl → repl ≡ node key value leaf leaf |
125 rt-property-leaf r-leaf = refl | |
126 | |
698 | 127 rt-property-¬leaf : {n : Level} {A : Set n} {key : ℕ} {value : A} {tree : bt A} → ¬ replacedTree key value tree leaf |
128 rt-property-¬leaf () | |
129 | |
692 | 130 rt-property-key : {n : Level} {A : Set n} {key key₂ key₃ : ℕ} {value value₂ value₃ : A} {left left₁ right₂ right₃ : bt A} |
131 → replacedTree key value (node key₂ value₂ left right₂) (node key₃ value₃ left₁ right₃) → key₂ ≡ key₃ | |
132 rt-property-key r-node = refl | |
133 rt-property-key (r-right x ri) = refl | |
134 rt-property-key (r-left x ri) = refl | |
135 | |
698 | 136 nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ |
137 nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x | |
138 nat-<> : { x y : ℕ } → x < y → y < x → ⊥ | |
139 nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x | |
140 | |
141 open _∧_ | |
142 | |
143 | |
632 | 144 depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) |
145 depth-1< {i} {j} = s≤s (m≤m⊔n _ j) | |
146 | |
147 depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) | |
650 | 148 depth-2< {i} {j} = s≤s (m≤n⊔m j i) |
611 | 149 |
649 | 150 depth-3< : {i : ℕ } → suc i ≤ suc (suc i) |
151 depth-3< {zero} = s≤s ( z≤n ) | |
152 depth-3< {suc i} = s≤s (depth-3< {i} ) | |
153 | |
154 | |
634 | 155 treeLeftDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) |
156 → treeInvariant (node k v1 tree tree₁) | |
157 → treeInvariant tree | |
158 treeLeftDown {n} {A} {_} {v1} leaf leaf (t-single k1 v1) = t-leaf | |
159 treeLeftDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = t-leaf | |
160 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = ti | |
161 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti | |
162 | |
163 treeRightDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) | |
164 → treeInvariant (node k v1 tree tree₁) | |
165 → treeInvariant tree₁ | |
166 treeRightDown {n} {A} {_} {v1} .leaf .leaf (t-single _ .v1) = t-leaf | |
167 treeRightDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = ti | |
168 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = t-leaf | |
169 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti₁ | |
170 | |
615 | 171 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A)) |
662 | 172 → treeInvariant tree ∧ stackInvariant key tree tree0 stack |
693 | 173 → (next : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) |
174 → (exit : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack | |
638 | 175 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t |
693 | 176 findP key leaf tree0 st Pre _ exit = exit leaf st Pre (case1 refl) |
632 | 177 findP key (node key₁ v1 tree tree₁) tree0 st Pre next exit with <-cmp key key₁ |
693 | 178 findP key n tree0 st Pre _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl) |
179 findP {n} {_} {A} key (node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree (tree ∷ st) | |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
180 ⟪ treeLeftDown tree tree₁ (proj1 Pre) , findP1 a st (proj2 Pre) ⟫ depth-1< where |
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
181 findP1 : key < key₁ → (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
771 | 182 findP1 a (x ∷ st) si = s-left _ _ _ a si |
183 findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st) ⟪ treeRightDown tree tree₁ (proj1 Pre) , s-right _ _ _ c (proj2 Pre) ⟫ depth-2< | |
606 | 184 |
638 | 185 replaceTree1 : {n : Level} {A : Set n} {t t₁ : bt A } → ( k : ℕ ) → (v1 value : A ) → treeInvariant (node k v1 t t₁) → treeInvariant (node k value t t₁) |
186 replaceTree1 k v1 value (t-single .k .v1) = t-single k value | |
187 replaceTree1 k v1 value (t-right x t) = t-right x t | |
188 replaceTree1 k v1 value (t-left x t) = t-left x t | |
189 replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ | |
190 | |
649 | 191 open import Relation.Binary.Definitions |
192 | |
193 lemma3 : {i j : ℕ} → 0 ≡ i → j < i → ⊥ | |
194 lemma3 refl () | |
195 lemma5 : {i j : ℕ} → i < 1 → j < i → ⊥ | |
196 lemma5 (s≤s z≤n) () | |
700 | 197 ¬x<x : {x : ℕ} → ¬ (x < x) |
198 ¬x<x (s≤s lt) = ¬x<x lt | |
649 | 199 |
687 | 200 child-replaced : {n : Level} {A : Set n} (key : ℕ) (tree : bt A) → bt A |
201 child-replaced key leaf = leaf | |
202 child-replaced key (node key₁ value left right) with <-cmp key key₁ | |
203 ... | tri< a ¬b ¬c = left | |
204 ... | tri≈ ¬a b ¬c = node key₁ value left right | |
205 ... | tri> ¬a ¬b c = right | |
677 | 206 |
671
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
207 record replacePR {n : Level} {A : Set n} (key : ℕ) (value : A) (tree repl : bt A ) (stack : List (bt A)) (C : bt A → bt A → List (bt A) → Set n) : Set n where |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
208 field |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
209 tree0 : bt A |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
210 ti : treeInvariant tree0 |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
211 si : stackInvariant key tree tree0 stack |
687 | 212 ri : replacedTree key value (child-replaced key tree ) repl |
671
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
213 ci : C tree repl stack -- data continuation |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
214 |
772 | 215 record replacePR' {n : Level} {A : Set n} (key : ℕ) (value : A) (orig : bt A ) (stack : List (bt A)) : Set n where |
216 field | |
217 tree repl : bt A | |
218 ti : treeInvariant orig | |
219 si : stackInvariant key tree orig stack | |
220 ri : replacedTree key value (child-replaced key tree) repl | |
221 -- treeInvariant of tree and repl is inferred from ti, si and ri. | |
222 | |
638 | 223 replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) |
224 → (tree ≡ leaf ) ∨ ( node-key tree ≡ just key ) | |
694 | 225 → (treeInvariant tree ) → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value (child-replaced key tree) tree1 → t) → t |
226 replaceNodeP k v1 leaf C P next = next (node k v1 leaf leaf) (t-single k v1 ) r-leaf | |
227 replaceNodeP k v1 (node .k value t t₁) (case2 refl) P next = next (node k v1 t t₁) (replaceTree1 k value v1 P) | |
695 | 228 (subst (λ j → replacedTree k v1 j (node k v1 t t₁) ) repl00 r-node) where |
694 | 229 repl00 : node k value t t₁ ≡ child-replaced k (node k value t t₁) |
230 repl00 with <-cmp k k | |
231 ... | tri< a ¬b ¬c = ⊥-elim (¬b refl) | |
232 ... | tri≈ ¬a b ¬c = refl | |
233 ... | tri> ¬a ¬b c = ⊥-elim (¬b refl) | |
606 | 234 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
235 replaceP : {n m : Level} {A : Set n} {t : Set m} |
671
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
236 → (key : ℕ) → (value : A) → {tree : bt A} ( repl : bt A) |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
237 → (stack : List (bt A)) → replacePR key value tree repl stack (λ _ _ _ → Lift n ⊤) |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
238 → (next : ℕ → A → {tree1 : bt A } (repl : bt A) → (stack1 : List (bt A)) |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
239 → replacePR key value tree1 repl stack1 (λ _ _ _ → Lift n ⊤) → length stack1 < length stack → t) |
613 | 240 → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t |
675 | 241 replaceP key value {tree} repl [] Pre next exit = ⊥-elim ( si-property0 (replacePR.si Pre) refl ) -- can't happen |
242 replaceP key value {tree} repl (leaf ∷ []) Pre next exit with si-property-last _ _ _ _ (replacePR.si Pre)-- tree0 ≡ leaf | |
677 | 243 ... | refl = exit (replacePR.tree0 Pre) (node key value leaf leaf) ⟪ replacePR.ti Pre , r-leaf ⟫ |
689 | 244 replaceP key value {tree} repl (node key₁ value₁ left right ∷ []) Pre next exit with <-cmp key key₁ |
245 ... | tri< a ¬b ¬c = exit (replacePR.tree0 Pre) (node key₁ value₁ repl right ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
246 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ repl right ) | |
247 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
248 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ k right ) (node key₁ value₁ repl right )) repl02 (r-left a repl03) where | |
249 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
250 repl03 = replacePR.ri Pre | |
251 repl02 : child-replaced key (node key₁ value₁ left right) ≡ left | |
252 repl02 with <-cmp key key₁ | |
253 ... | tri< a ¬b ¬c = refl | |
254 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a a) | |
255 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a a) | |
256 ... | tri≈ ¬a b ¬c = exit (replacePR.tree0 Pre) repl ⟪ replacePR.ti Pre , repl01 ⟫ where | |
678 | 257 repl01 : replacedTree key value (replacePR.tree0 Pre) repl |
258 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
689 | 259 repl01 | refl | refl = subst (λ k → replacedTree key value k repl) repl02 (replacePR.ri Pre) where |
260 repl02 : child-replaced key (node key₁ value₁ left right) ≡ node key₁ value₁ left right | |
261 repl02 with <-cmp key key₁ | |
262 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b b) | |
263 ... | tri≈ ¬a b ¬c = refl | |
264 ... | tri> ¬a ¬b c = ⊥-elim ( ¬b b) | |
265 ... | tri> ¬a ¬b c = exit (replacePR.tree0 Pre) (node key₁ value₁ left repl ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
266 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ left repl ) | |
267 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
268 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ left k ) (node key₁ value₁ left repl )) repl02 (r-right c repl03) where | |
269 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
270 repl03 = replacePR.ri Pre | |
271 repl02 : child-replaced key (node key₁ value₁ left right) ≡ right | |
272 repl02 with <-cmp key key₁ | |
273 ... | tri< a ¬b ¬c = ⊥-elim ( ¬c c) | |
274 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬c c) | |
275 ... | tri> ¬a ¬b c = refl | |
690 | 276 replaceP {n} {_} {A} key value {tree} repl (leaf ∷ st@(tree1 ∷ st1)) Pre next exit = next key value repl st Post ≤-refl where |
277 Post : replacePR key value tree1 repl (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) | |
278 Post with replacePR.si Pre | |
771 | 279 ... | s-right _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
690 | 280 repl09 : tree1 ≡ node key₂ v1 tree₁ leaf |
281 repl09 = si-property1 si | |
282 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
283 repl10 with si-property1 si | |
284 ... | refl = si | |
285 repl07 : child-replaced key (node key₂ v1 tree₁ leaf) ≡ leaf | |
286 repl07 with <-cmp key key₂ | |
287 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
288 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
289 ... | tri> ¬a ¬b c = refl | |
290 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
291 repl12 = subst₂ (λ j k → replacedTree key value j k ) (sym (subst (λ k → child-replaced key k ≡ leaf) (sym repl09) repl07 ) ) (sym (rt-property-leaf (replacePR.ri Pre))) r-leaf | |
771 | 292 ... | s-left _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
690 | 293 repl09 : tree1 ≡ node key₂ v1 leaf tree₁ |
294 repl09 = si-property1 si | |
295 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
296 repl10 with si-property1 si | |
297 ... | refl = si | |
298 repl07 : child-replaced key (node key₂ v1 leaf tree₁ ) ≡ leaf | |
299 repl07 with <-cmp key key₂ | |
300 ... | tri< a ¬b ¬c = refl | |
301 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
302 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
303 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
304 repl12 = subst₂ (λ j k → replacedTree key value j k ) (sym (subst (λ k → child-replaced key k ≡ leaf) (sym repl09) repl07 ) ) (sym (rt-property-leaf (replacePR.ri Pre))) r-leaf | |
683 | 305 replaceP {n} {_} {A} key value {tree} repl (node key₁ value₁ left right ∷ st@(tree1 ∷ st1)) Pre next exit with <-cmp key key₁ |
306 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st Post ≤-refl where | |
675 | 307 Post : replacePR key value tree1 (node key₁ value₁ repl right ) st (λ _ _ _ → Lift n ⊤) |
687 | 308 Post with replacePR.si Pre |
771 | 309 ... | s-right _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
688 | 310 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) |
311 repl09 = si-property1 si | |
312 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
313 repl10 with si-property1 si | |
314 ... | refl = si | |
315 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left | |
316 repl03 with <-cmp key key₁ | |
317 ... | tri< a1 ¬b ¬c = refl | |
318 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
319 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
320 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
321 repl02 with repl09 | <-cmp key key₂ | |
322 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
689 | 323 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) |
688 | 324 ... | refl | tri> ¬a ¬b c = refl |
325 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
326 repl04 = begin | |
327 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
328 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
329 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
330 child-replaced key tree1 ∎ where open ≡-Reasoning | |
331 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
332 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
771 | 333 ... | s-left _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
688 | 334 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ |
683 | 335 repl09 = si-property1 si |
336 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
337 repl10 with si-property1 si | |
338 ... | refl = si | |
687 | 339 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left |
340 repl03 with <-cmp key key₁ | |
341 ... | tri< a1 ¬b ¬c = refl | |
342 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
343 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
344 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
345 repl02 with repl09 | <-cmp key key₂ | |
346 ... | refl | tri< a ¬b ¬c = refl | |
347 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
348 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
349 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
350 repl04 = begin | |
351 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
352 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
353 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
354 child-replaced key tree1 ∎ where open ≡-Reasoning | |
355 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
356 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
705 | 357 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st Post ≤-refl where |
690 | 358 Post : replacePR key value tree1 (node key₁ value left right ) (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) |
359 Post with replacePR.si Pre | |
771 | 360 ... | s-right _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 b ; ci = lift tt } where |
690 | 361 repl09 : tree1 ≡ node key₂ v1 tree₁ tree -- (node key₁ value₁ left right) |
362 repl09 = si-property1 si | |
363 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
364 repl10 with si-property1 si | |
365 ... | refl = si | |
366 repl07 : child-replaced key (node key₂ v1 tree₁ tree) ≡ tree | |
367 repl07 with <-cmp key key₂ | |
368 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
369 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
370 ... | tri> ¬a ¬b c = refl | |
691 | 371 repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) |
372 repl12 refl with repl09 | |
373 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node | |
771 | 374 ... | s-left _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 b ; ci = lift tt } where |
690 | 375 repl09 : tree1 ≡ node key₂ v1 tree tree₁ |
376 repl09 = si-property1 si | |
377 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
378 repl10 with si-property1 si | |
379 ... | refl = si | |
380 repl07 : child-replaced key (node key₂ v1 tree tree₁ ) ≡ tree | |
381 repl07 with <-cmp key key₂ | |
382 ... | tri< a ¬b ¬c = refl | |
383 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
384 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
691 | 385 repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) |
386 repl12 refl with repl09 | |
387 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node | |
690 | 388 ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st Post ≤-refl where |
389 Post : replacePR key value tree1 (node key₁ value₁ left repl ) st (λ _ _ _ → Lift n ⊤) | |
390 Post with replacePR.si Pre | |
771 | 391 ... | s-right _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
690 | 392 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) |
393 repl09 = si-property1 si | |
394 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
395 repl10 with si-property1 si | |
396 ... | refl = si | |
397 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
398 repl03 with <-cmp key key₁ | |
399 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
400 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
401 ... | tri> ¬a ¬b c = refl | |
402 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
403 repl02 with repl09 | <-cmp key key₂ | |
404 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
405 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) | |
406 ... | refl | tri> ¬a ¬b c = refl | |
407 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
408 repl04 = begin | |
409 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
410 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
411 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
412 child-replaced key tree1 ∎ where open ≡-Reasoning | |
413 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
414 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
771 | 415 ... | s-left _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
690 | 416 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ |
417 repl09 = si-property1 si | |
418 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
419 repl10 with si-property1 si | |
420 ... | refl = si | |
421 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
422 repl03 with <-cmp key key₁ | |
423 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
424 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
425 ... | tri> ¬a ¬b c = refl | |
426 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
427 repl02 with repl09 | <-cmp key key₂ | |
428 ... | refl | tri< a ¬b ¬c = refl | |
429 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
430 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
431 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
432 repl04 = begin | |
433 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
434 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
435 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
436 child-replaced key tree1 ∎ where open ≡-Reasoning | |
437 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
438 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
644 | 439 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
440 TerminatingLoopS : {l m : Level} {t : Set l} (Index : Set m ) → {Invraiant : Index → Set m } → ( reduce : Index → ℕ) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
441 → (r : Index) → (p : Invraiant r) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
442 → (loop : (r : Index) → Invraiant r → (next : (r1 : Index) → Invraiant r1 → reduce r1 < reduce r → t ) → t) → t |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
443 TerminatingLoopS {_} {_} {t} Index {Invraiant} reduce r p loop with <-cmp 0 (reduce r) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
444 ... | tri≈ ¬a b ¬c = loop r p (λ r1 p1 lt → ⊥-elim (lemma3 b lt) ) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
445 ... | tri< a ¬b ¬c = loop r p (λ r1 p1 lt1 → TerminatingLoop1 (reduce r) r r1 (≤-step lt1) p1 lt1 ) where |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
446 TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → Invraiant r1 → reduce r1 < reduce r → t |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
447 TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 p1 (λ r2 p1 lt1 → ⊥-elim (lemma5 n≤j lt1)) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
448 TerminatingLoop1 (suc j) r r1 n≤j p1 lt with <-cmp (reduce r1) (suc j) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
449 ... | tri< a ¬b ¬c = TerminatingLoop1 j r r1 a p1 lt |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
450 ... | tri≈ ¬a b ¬c = loop r1 p1 (λ r2 p2 lt1 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b lt1 ) p2 lt1 ) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
451 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
452 |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
453 open _∧_ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
454 |
615 | 455 RTtoTI0 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
456 → replacedTree key value tree repl → treeInvariant repl | |
692 | 457 RTtoTI0 .leaf .(node key value leaf leaf) key value ti r-leaf = t-single key value |
458 RTtoTI0 .(node key _ leaf leaf) .(node key value leaf leaf) key value (t-single .key _) r-node = t-single key value | |
459 RTtoTI0 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right x ti) r-node = t-right x ti | |
460 RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti | |
461 RTtoTI0 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ ti ti₁ | |
701 | 462 -- r-right case |
692 | 463 RTtoTI0 (node _ _ leaf leaf) (node _ _ leaf .(node key value leaf leaf)) key value (t-single _ _) (r-right x r-leaf) = t-right x (t-single key value) |
464 RTtoTI0 (node _ _ leaf right@(node _ _ _ _)) (node key₁ value₁ leaf leaf) key value (t-right x₁ ti) (r-right x ri) = t-single key₁ value₁ | |
693 | 465 RTtoTI0 (node key₁ _ leaf right@(node key₂ _ _ _)) (node key₁ value₁ leaf right₁@(node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = |
466 t-right (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) | |
692 | 467 RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left x₁ ti) (r-right x ()) |
468 RTtoTI0 (node key₁ _ (node key₃ _ _ _) leaf) (node key₁ _ (node key₃ value₃ _ _) (node key value leaf leaf)) key value (t-left x₁ ti) (r-right x r-leaf) = | |
469 t-node x₁ x ti (t-single key value) | |
693 | 470 RTtoTI0 (node key₁ _ (node _ _ _ _) (node key₂ _ _ _)) (node key₁ _ (node _ _ _ _) (node key₃ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = |
471 t-node x₁ (subst (λ k → key₁ < k) (rt-property-key ri) x₂) ti (RTtoTI0 _ _ key value ti₁ ri) | |
701 | 472 -- r-left case |
700 | 473 RTtoTI0 .(node _ _ leaf leaf) .(node _ _ (node key value leaf leaf) leaf) key value (t-single _ _) (r-left x r-leaf) = t-left x (t-single _ _ ) |
701 | 474 RTtoTI0 .(node _ _ leaf (node _ _ _ _)) (node key₁ value₁ (node key value leaf leaf) (node _ _ _ _)) key value (t-right x₁ ti) (r-left x r-leaf) = t-node x x₁ (t-single key value) ti |
475 RTtoTI0 (node key₃ _ (node key₂ _ _ _) leaf) (node key₃ _ (node key₁ value₁ left left₁) leaf) key value (t-left x₁ ti) (r-left x ri) = | |
476 t-left (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) -- key₁ < key₃ | |
477 RTtoTI0 (node key₁ _ (node key₂ _ _ _) (node _ _ _ _)) (node key₁ _ (node key₃ _ _ _) (node _ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = t-node (subst (λ k → k < key₁ ) (rt-property-key ri) x₁) x₂ (RTtoTI0 _ _ key value ti ri) ti₁ | |
615 | 478 |
479 RTtoTI1 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant repl | |
480 → replacedTree key value tree repl → treeInvariant tree | |
701 | 481 RTtoTI1 .leaf .(node key value leaf leaf) key value ti r-leaf = t-leaf |
482 RTtoTI1 (node key value₁ leaf leaf) .(node key value leaf leaf) key value (t-single .key .value) r-node = t-single key value₁ | |
483 RTtoTI1 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right x ti) r-node = t-right x ti | |
484 RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti | |
485 RTtoTI1 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ ti ti₁ | |
486 -- r-right case | |
487 RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ leaf (node _ _ _ _)) key value (t-right x₁ ti) (r-right x r-leaf) = t-single key₁ value₁ | |
488 RTtoTI1 (node key₁ value₁ leaf (node key₂ value₂ t2 t3)) (node key₁ _ leaf (node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = | |
489 t-right (subst (λ k → key₁ < k ) (sym (rt-property-key ri)) x₁) (RTtoTI1 _ _ key value ti ri) -- key₁ < key₂ | |
490 RTtoTI1 (node _ _ (node _ _ _ _) leaf) (node _ _ (node _ _ _ _) (node key value _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x r-leaf) = | |
491 t-left x₁ ti | |
492 RTtoTI1 (node key₄ _ (node key₃ _ _ _) (node key₁ value₁ n n₁)) (node key₄ _ (node key₃ _ _ _) (node key₂ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = t-node x₁ (subst (λ k → key₄ < k ) (sym (rt-property-key ri)) x₂) ti (RTtoTI1 _ _ key value ti₁ ri) -- key₄ < key₁ | |
493 -- r-left case | |
494 RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ _ leaf) key value (t-left x₁ ti) (r-left x ri) = t-single key₁ value₁ | |
495 RTtoTI1 (node key₁ _ (node key₂ value₁ n n₁) leaf) (node key₁ _ (node key₃ _ _ _) leaf) key value (t-left x₁ ti) (r-left x ri) = | |
496 t-left (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) (RTtoTI1 _ _ key value ti ri) -- key₂ < key₁ | |
497 RTtoTI1 (node key₁ value₁ leaf _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x r-leaf) = t-right x₂ ti₁ | |
498 RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = | |
499 t-node (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ (RTtoTI1 _ _ key value ti ri) ti₁ -- key₂ < key₁ | |
614 | 500 |
611 | 501 insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
696 | 502 → (exit : (tree repl : bt A) → treeInvariant repl ∧ replacedTree key value tree repl → t ) → t |
693 | 503 insertTreeP {n} {m} {A} {t} tree key value P0 exit = |
729 | 504 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant key (proj1 p) tree (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , tree ∷ [] ⟫ ⟪ P0 , s-nil ⟫ |
696 | 505 $ λ p P loop → findP key (proj1 p) tree (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
693 | 506 $ λ t s P C → replaceNodeP key value t C (proj1 P) |
507 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ bt A ∧ bt A ) | |
508 {λ p → replacePR key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) (λ _ _ _ → Lift n ⊤ ) } | |
696 | 509 (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ record { tree0 = tree ; ti = P0 ; si = proj2 P ; ri = R ; ci = lift tt } |
693 | 510 $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) P1 |
696 | 511 (λ key value {tree1} repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ P2 lt ) |
512 $ λ tree repl P → exit tree repl ⟪ RTtoTI0 _ _ _ _ (proj1 P) (proj2 P) , proj2 P ⟫ | |
614 | 513 |
696 | 514 insertTestP1 = insertTreeP leaf 1 1 t-leaf |
722 | 515 $ λ _ x0 P0 → insertTreeP x0 2 1 (proj1 P0) |
516 $ λ _ x1 P1 → insertTreeP x1 3 2 (proj1 P1) | |
727 | 517 $ λ _ x2 P2 → insertTreeP x2 2 2 (proj1 P2) (λ _ x P → x ) |
694 | 518 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
519 top-value : {n : Level} {A : Set n} → (tree : bt A) → Maybe A |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
520 top-value leaf = nothing |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
521 top-value (node key value tree tree₁) = just value |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
522 |
722 | 523 -- is realy inserted? |
524 | |
525 -- other element is preserved? | |
702 | 526 |
722 | 527 -- deletion? |
616 | 528 |
722 | 529 data Color : Set where |
530 Red : Color | |
531 Black : Color | |
618 | 532 |
746 | 533 data RBtreeInvariant {n : Level} {A : Set n} : (tree : bt (Color ∧ A)) → (bd : ℕ) → Set n where |
534 rb-leaf : RBtreeInvariant leaf 0 | |
758 | 535 rb-single : (key : ℕ) → (value : A) → RBtreeInvariant (node key ⟪ Black , value ⟫ leaf leaf) 1 |
773 | 536 rb-right-red : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} |
758 | 537 → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ t t₁) 1 |
538 → RBtreeInvariant (node key ⟪ Red , value ⟫ leaf (node key₁ ⟪ Black , value₁ ⟫ t t₁)) 1 | |
773 | 539 rb-right-black : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → {c : Color} |
758 | 540 → RBtreeInvariant (node key₁ ⟪ c , value₁ ⟫ t t₁) 1 |
541 → RBtreeInvariant (node key ⟪ Black , value ⟫ leaf (node key₁ ⟪ c , value₁ ⟫ t t₁)) 1 | |
773 | 542 rb-left-red : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} |
758 | 543 → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ t t₁) 1 |
544 → RBtreeInvariant (node key₁ ⟪ Red , value ⟫ (node key ⟪ Black , value₁ ⟫ t t₁) leaf ) 1 | |
773 | 545 rb-left-black : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → {c : Color} |
758 | 546 → RBtreeInvariant (node key₁ ⟪ c , value₁ ⟫ t t₁) 1 |
547 → RBtreeInvariant (node key₁ ⟪ Black , value ⟫ (node key ⟪ c , value₁ ⟫ t t₁) leaf) 1 | |
773 | 548 rb-node-red : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt (Color ∧ A)} → (d : ℕ) |
746 | 549 → RBtreeInvariant (node key ⟪ Black , value ⟫ t₁ t₂) d |
550 → RBtreeInvariant (node key₂ ⟪ Black , value₂ ⟫ t₃ t₄) d | |
551 → RBtreeInvariant (node key₁ ⟪ Red , value₁ ⟫ (node key ⟪ Black , value ⟫ t₁ t₂) (node key₂ ⟪ Black , value₂ ⟫ t₃ t₄)) d | |
773 | 552 rb-node-black : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt (Color ∧ A)} → (d : ℕ) |
746 | 553 → {c c₁ : Color} |
554 → RBtreeInvariant (node key ⟪ c , value ⟫ t₁ t₂) d | |
555 → RBtreeInvariant (node key₂ ⟪ c₁ , value₂ ⟫ t₃ t₄) d | |
556 → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ (node key ⟪ c , value ⟫ t₁ t₂) (node key₂ ⟪ c₁ , value₂ ⟫ t₃ t₄)) (suc d) | |
745 | 557 |
758 | 558 tesr-rb-0 : RBtreeInvariant {_} {ℕ} leaf 0 |
559 tesr-rb-0 = rb-leaf | |
560 | |
561 tesr-rb-1 : RBtreeInvariant {_} {ℕ} (node 10 ⟪ Black , 10 ⟫ leaf leaf) 1 | |
562 tesr-rb-1 = rb-single 10 10 | |
563 | |
564 tesr-rb-2 : RBtreeInvariant {_} {ℕ} (node 10 ⟪ Red , 10 ⟫ (node 5 ⟪ Black , 5 ⟫ leaf leaf) leaf) 1 | |
773 | 565 tesr-rb-2 = rb-left-red ( rb-single 10 5) |
746 | 566 |
773 | 567 color : {n : Level} {A : Set n} → (rb : bt (Color ∧ A)) → Color |
568 color leaf = Red | |
569 color (node key ⟪ c , v ⟫ rb rb₁) = c | |
723 | 570 |
747 | 571 RB→bt : {n : Level} (A : Set n) → (rb : bt (Color ∧ A)) → bt A |
746 | 572 RB→bt {n} A leaf = leaf |
747 | 573 RB→bt {n} A (node key ⟪ c , value ⟫ rb rb₁) = node key value (RB→bt A rb) (RB→bt A rb₁) |
723 | 574 |
765 | 575 data replacedRBTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt (Color ∧ A) ) → Set n where |
576 rbr-leaf : {ca cb : Color} → replacedRBTree key value leaf (node key ⟪ cb , value ⟫ leaf leaf) | |
577 rbr-node : {value₁ : A} → {ca cb : Color } → {t t₁ : bt (Color ∧ A)} → replacedRBTree key value (node key ⟪ ca , value₁ ⟫ t t₁) (node key ⟪ cb , value ⟫ t t₁) | |
578 rbr-right : {k : ℕ } {v1 : A} → {ca cb : Color} → {t t1 t2 : bt (Color ∧ A)} | |
579 → k < key → replacedRBTree key value t2 t → replacedRBTree key value (node k ⟪ ca , v1 ⟫ t1 t2) (node k ⟪ cb , v1 ⟫ t1 t) | |
580 rbr-left : {k : ℕ } {v1 : A} → {ca cb : Color} → {t t1 t2 : bt (Color ∧ A)} | |
581 → k < key → replacedRBTree key value t1 t → replacedRBTree key value (node k ⟪ ca , v1 ⟫ t1 t2) (node k ⟪ cb , v1 ⟫ t t2) | |
582 | |
754 | 583 data ParentGrand {n : Level} {A : Set n} (self : bt A) : (parent uncle grand : bt A) → Set n where |
740 | 584 s2-s1p2 : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } |
754 | 585 → parent ≡ node kp vp self n1 → grand ≡ node kg vg parent n2 → ParentGrand self parent n2 grand |
740 | 586 s2-1sp2 : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } |
754 | 587 → parent ≡ node kp vp n1 self → grand ≡ node kg vg parent n2 → ParentGrand self parent n2 grand |
740 | 588 s2-s12p : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } |
754 | 589 → parent ≡ node kp vp self n1 → grand ≡ node kg vg n2 parent → ParentGrand self parent n2 grand |
740 | 590 s2-1s2p : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } |
754 | 591 → parent ≡ node kp vp n1 self → grand ≡ node kg vg n2 parent → ParentGrand self parent n2 grand |
734 | 592 |
746 | 593 -- with color |
736 | 594 data rotatedTree {n : Level} {A : Set n} : (before after : bt A ) → Set n where |
595 rr-node : {t : bt A} → rotatedTree t t | |
746 | 596 -- a b |
597 -- b c d a | |
598 -- d e e c | |
599 rr-right : {ka kb : ℕ } {va vb : A} → {c c₁ d d₁ e e₁ : bt A} | |
736 | 600 → ka < kb |
746 | 601 → rotatedTree d d₁ → rotatedTree e e₁ → rotatedTree c c₁ |
747 | 602 → rotatedTree (node ka va (node kb vb d e) c) (node kb vb d₁ (node ka va e₁ c₁) ) |
745 | 603 -- b a |
604 -- d a b c | |
746 | 605 -- e c d e |
606 rr-left : {ka kb : ℕ } {va vb : A} → {c c₁ d d₁ e e₁ : bt A} | |
736 | 607 → ka < kb |
746 | 608 → rotatedTree d d₁ → rotatedTree e e₁ → rotatedTree c c₁ |
609 → rotatedTree (node kb vb d (node ka va e c) ) (node ka va (node kb vb d₁ e₁) c₁) | |
734 | 610 |
749 | 611 record PG {n : Level } (A : Set n) (self : bt A) (stack : List (bt A)) : Set n where |
740 | 612 field |
754 | 613 parent grand uncle : bt A |
614 pg : ParentGrand self parent uncle grand | |
740 | 615 rest : List (bt A) |
616 stack=gp : stack ≡ ( self ∷ parent ∷ grand ∷ rest ) | |
617 | |
775 | 618 record RBI {n : Level} {A : Set n} (key : ℕ) (value : A) (orig repl : bt (Color ∧ A) ) (stack : List (bt (Color ∧ A))) : Set n where |
769 | 619 field |
772 | 620 od d rd : ℕ |
775 | 621 tree rot : bt (Color ∧ A) |
773 | 622 origti : treeInvariant orig |
772 | 623 origrb : RBtreeInvariant orig od |
624 treerb : RBtreeInvariant tree d | |
625 replrb : RBtreeInvariant repl rd | |
773 | 626 d=rd : ( d ≡ rd ) ∨ ( (suc d ≡ rd ) ∧ (color tree ≡ Red)) |
772 | 627 si : stackInvariant key tree orig stack |
628 rotated : rotatedTree tree rot | |
629 ri : replacedRBTree key value (child-replaced key rot ) repl | |
773 | 630 |
631 -- r (b , b) inserting a node into black node does not change the black depth, but color may change | |
632 -- → b (b , r ) ∨ b (r , b) | |
633 -- b (b , b) | |
634 -- → b (b , r ) ∨ b (r , b) | |
635 -- b (r , b) inserting to right → b (r , r ) | |
636 -- b (r , b) inserting to left may increse black depth, need rotation | |
637 -- find b in left and move the b to the right (one down or two down) | |
638 -- | |
639 rbi-case1 : {n : Level} {A : Set n} → {key d : ℕ} → {value : A} → (parent repl : bt (Color ∧ A) ) | |
640 → RBtreeInvariant parent (suc d) | |
641 → RBtreeInvariant repl d | |
642 → {left right : bt (Color ∧ A) } → parent ≡ node key ⟪ Black , value ⟫ left right | |
643 → (color right ≡ Red → RBtreeInvariant (node key ⟪ Black , value ⟫ left repl ) (suc d)) | |
644 ∧ (color left ≡ Red → RBtreeInvariant (node key ⟪ Black , value ⟫ repl right ) (suc d)) | |
645 rbi-case1 {n} {A} {key} = ? | |
646 | |
774 | 647 rbi-case31 : {n : Level} {A : Set n} → {kg kp d dp : ℕ} → {vg vp : A} → {cg : Color} → (tree grand parent repl : bt (Color ∧ A) ) |
773 | 648 → RBtreeInvariant grand dp |
649 → RBtreeInvariant repl d | |
650 → {uncle left right : bt (Color ∧ A) } | |
651 → grand ≡ node kg ⟪ cg , vg ⟫ uncle parent | |
652 → parent ≡ node kp ⟪ Red , vp ⟫ left right | |
653 → color uncle ≡ Red | |
654 → (color left ≡ Red → RBtreeInvariant (node kg ⟪ Red , vg ⟫ uncle (node kp ⟪ Black , vp ⟫ repl right )) dp) | |
655 ∧ (color right ≡ Red → RBtreeInvariant (node kg ⟪ Red , vg ⟫ uncle (node kp ⟪ Black , vp ⟫ left repl )) dp) | |
774 | 656 rbi-case31 {n} {A} {key} = ? |
657 | |
658 -- | |
659 -- case4 increase the black depth | |
660 -- | |
661 rbi-case41 : {n : Level} {A : Set n} → {kg kp d dp : ℕ} → {vg vp : A} → {cg : Color} → (tree grand parent repl : bt (Color ∧ A) ) | |
662 → RBtreeInvariant grand dp | |
663 → RBtreeInvariant repl d | |
664 → {uncle left right : bt (Color ∧ A) } | |
665 → grand ≡ node kg ⟪ cg , vg ⟫ uncle parent | |
666 → parent ≡ node kp ⟪ Red , vp ⟫ left right | |
667 → color uncle ≡ Black | |
668 → (color left ≡ Red → RBtreeInvariant (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Black , vp ⟫ repl right )) (suc dp) ) | |
669 ∧ (color right ≡ Red → RBtreeInvariant (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Black , vp ⟫ left repl )) (suc dp) ) | |
670 rbi-case41 {n} {A} {key} = ? | |
671 | |
672 rbi-case51 : {n : Level} {A : Set n} → {kg kp d dp : ℕ} → {vg vp : A} → {cg : Color} → (tree grand parent repl : bt (Color ∧ A) ) | |
673 → RBtreeInvariant grand dp | |
674 → RBtreeInvariant repl d | |
675 → {uncle left right : bt (Color ∧ A) } | |
676 → grand ≡ node kg ⟪ cg , vg ⟫ uncle parent | |
677 → parent ≡ node kp ⟪ Red , vp ⟫ left right | |
678 → color uncle ≡ Black | |
679 → (color left ≡ Red → RBtreeInvariant (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Black , vp ⟫ repl right )) (suc dp) ) | |
680 ∧ (color right ≡ Red → RBtreeInvariant (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Black , vp ⟫ left repl )) (suc dp) ) | |
681 rbi-case51 {n} {A} {key} = ? | |
773 | 682 |
683 --... | Black = record { | |
684 -- d = ? ; od = RBI.od rbi ; rd = ? | |
685 -- ; tree = ? ; rot = ? ; repl = ? | |
686 -- ; treerb = ? ; replrb = ? | |
687 -- ; d=rd = ? ; si = ? ; rotated = ? ; ri = ? | |
688 -- ; origti = RBI.origti rbi ; origrb = RBI.origrb rbi | |
689 -- } | |
690 --... | Red = ? | |
770 | 691 |
740 | 692 stackToPG : {n : Level} {A : Set n} → {key : ℕ } → (tree orig : bt A ) |
693 → (stack : List (bt A)) → stackInvariant key tree orig stack | |
749 | 694 → ( stack ≡ orig ∷ [] ) ∨ ( stack ≡ tree ∷ orig ∷ [] ) ∨ PG A tree stack |
740 | 695 stackToPG {n} {A} {key} tree .tree .(tree ∷ []) s-nil = case1 refl |
771 | 696 stackToPG {n} {A} {key} tree .(node _ _ _ tree) .(tree ∷ node _ _ _ tree ∷ []) (s-right _ _ _ x s-nil) = case2 (case1 refl) |
697 stackToPG {n} {A} {key} tree .(node k2 v2 t5 (node k1 v1 t2 tree)) (tree ∷ node _ _ _ tree ∷ .(node k2 v2 t5 (node k1 v1 t2 tree) ∷ [])) (s-right tree (node k2 v2 t5 (node k1 v1 t2 tree)) t2 {k1} {v1} x (s-right (node k1 v1 t2 tree) (node k2 v2 t5 (node k1 v1 t2 tree)) t5 {k2} {v2} x₁ s-nil)) = case2 (case2 | |
749 | 698 record { parent = node k1 v1 t2 tree ; grand = _ ; pg = s2-1s2p refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 699 stackToPG {n} {A} {key} tree orig (tree ∷ node _ _ _ tree ∷ .(node k2 v2 t5 (node k1 v1 t2 tree) ∷ _)) (s-right tree orig t2 {k1} {v1} x (s-right (node k1 v1 t2 tree) orig t5 {k2} {v2} x₁ (s-right _ _ _ x₂ si))) = case2 (case2 |
749 | 700 record { parent = node k1 v1 t2 tree ; grand = _ ; pg = s2-1s2p refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 701 stackToPG {n} {A} {key} tree orig (tree ∷ node _ _ _ tree ∷ .(node k2 v2 t5 (node k1 v1 t2 tree) ∷ _)) (s-right tree orig t2 {k1} {v1} x (s-right (node k1 v1 t2 tree) orig t5 {k2} {v2} x₁ (s-left _ _ _ x₂ si))) = case2 (case2 |
749 | 702 record { parent = node k1 v1 t2 tree ; grand = _ ; pg = s2-1s2p refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 703 stackToPG {n} {A} {key} tree .(node k2 v2 (node k1 v1 t1 tree) t2) .(tree ∷ node k1 v1 t1 tree ∷ node k2 v2 (node k1 v1 t1 tree) t2 ∷ []) (s-right _ _ t1 {k1} {v1} x (s-left _ _ t2 {k2} {v2} x₁ s-nil)) = case2 (case2 |
749 | 704 record { parent = node k1 v1 t1 tree ; grand = _ ; pg = s2-1sp2 refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 705 stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 t1 tree ∷ node k2 v2 (node k1 v1 t1 tree) t2 ∷ _) (s-right _ _ t1 {k1} {v1} x (s-left _ _ t2 {k2} {v2} x₁ (s-right _ _ _ x₂ si))) = case2 (case2 |
749 | 706 record { parent = node k1 v1 t1 tree ; grand = _ ; pg = s2-1sp2 refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 707 stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 t1 tree ∷ node k2 v2 (node k1 v1 t1 tree) t2 ∷ _) (s-right _ _ t1 {k1} {v1} x (s-left _ _ t2 {k2} {v2} x₁ (s-left _ _ _ x₂ si))) = case2 (case2 |
749 | 708 record { parent = node k1 v1 t1 tree ; grand = _ ; pg = s2-1sp2 refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 709 stackToPG {n} {A} {key} tree .(node _ _ tree _) .(tree ∷ node _ _ tree _ ∷ []) (s-left _ _ t1 {k1} {v1} x s-nil) = case2 (case1 refl) |
710 stackToPG {n} {A} {key} tree .(node _ _ _ (node k1 v1 tree t1)) .(tree ∷ node k1 v1 tree t1 ∷ node _ _ _ (node k1 v1 tree t1) ∷ []) (s-left _ _ t1 {k1} {v1} x (s-right _ _ _ x₁ s-nil)) = case2 (case2 | |
749 | 711 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s12p refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 712 stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ _ (node k1 v1 tree t1) ∷ _) (s-left _ _ t1 {k1} {v1} x (s-right _ _ _ x₁ (s-right _ _ _ x₂ si))) = case2 (case2 |
749 | 713 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s12p refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 714 stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ _ (node k1 v1 tree t1) ∷ _) (s-left _ _ t1 {k1} {v1} x (s-right _ _ _ x₁ (s-left _ _ _ x₂ si))) = case2 (case2 |
749 | 715 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s12p refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 716 stackToPG {n} {A} {key} tree .(node _ _ (node k1 v1 tree t1) _) .(tree ∷ node k1 v1 tree t1 ∷ node _ _ (node k1 v1 tree t1) _ ∷ []) (s-left _ _ t1 {k1} {v1} x (s-left _ _ _ x₁ s-nil)) = case2 (case2 |
749 | 717 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s1p2 refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 718 stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ (node k1 v1 tree t1) _ ∷ _) (s-left _ _ t1 {k1} {v1} x (s-left _ _ _ x₁ (s-right _ _ _ x₂ si))) = case2 (case2 |
749 | 719 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s1p2 refl refl ; rest = _ ; stack=gp = refl } ) |
771 | 720 stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ (node k1 v1 tree t1) _ ∷ _) (s-left _ _ t1 {k1} {v1} x (s-left _ _ _ x₁ (s-left _ _ _ x₂ si))) = case2 (case2 |
749 | 721 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s1p2 refl refl ; rest = _ ; stack=gp = refl } ) |
740 | 722 |
760 | 723 stackCase1 : {n : Level} {A : Set n} → {key : ℕ } → {tree orig : bt A } |
724 → {stack : List (bt A)} → stackInvariant key tree orig stack | |
725 → stack ≡ orig ∷ [] → tree ≡ orig | |
726 stackCase1 s-nil refl = refl | |
727 | |
749 | 728 PGtoRBinvariant : {n : Level} {A : Set n} → {key d0 ds dp dg : ℕ } → (tree orig : bt (Color ∧ A) ) |
729 → RBtreeInvariant orig d0 | |
730 → (stack : List (bt (Color ∧ A))) → (pg : PG (Color ∧ A) tree stack ) | |
731 → RBtreeInvariant tree ds ∧ RBtreeInvariant (PG.parent pg) dp ∧ RBtreeInvariant (PG.grand pg) dg | |
758 | 732 PGtoRBinvariant = {!!} |
734 | 733 |
778 | 734 findRBT : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt (Color ∧ A) ) |
735 → (stack : List (bt (Color ∧ A))) | |
736 → treeInvariant tree0 ∧ stackInvariant key tree tree0 stack | |
776 | 737 → {d : ℕ} → RBtreeInvariant tree0 d |
778 | 738 → (exit : (tree1 : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) |
739 → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack | |
776 | 740 → {d1 : ℕ} → RBtreeInvariant tree1 d1 |
741 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t | |
778 | 742 findRBT {_} {_} {A} {t} key tree0 leaf stack ti {d} rb0 exit = ? |
743 findRBT {_} {_} {A} {t} key tree0 (node key₁ value left right) stack ti {d} rb0 exit with <-cmp key key₁ | |
744 ... | tri< a ¬b ¬c = ? | |
745 ... | tri≈ ¬a b ¬c = ? | |
746 ... | tri> ¬a ¬b c = ? | |
723 | 747 |
754 | 748 rotateLeft : {n m : Level} {A : Set n} {t : Set m} |
775 | 749 → (key : ℕ) → (value : A) |
750 → (orig tree : bt (Color ∧ A)) | |
751 → (stack : List (bt (Color ∧ A))) | |
752 → (r : RBI key value orig tree stack ) | |
776 | 753 → (next : (current : bt (Color ∧ A)) → (stack1 : List (bt (Color ∧ A))) |
775 | 754 → (r : RBI key value orig current stack1 ) |
754 | 755 → length stack1 < length stack → t ) |
775 | 756 → (exit : (repl : bt (Color ∧ A) ) → (stack1 : List (bt (Color ∧ A))) |
757 → stack1 ≡ (orig ∷ []) | |
758 → RBI key value orig repl stack1 | |
759 → t ) → t | |
760 rotateLeft {n} {m} {A} {t} key value = ? where | |
756 | 761 rotateLeft1 : t |
776 | 762 rotateLeft1 with stackToPG ? ? ? ? |
763 ... | case1 x = ? -- {!!} {!!} {!!} {!!} rr | |
758 | 764 ... | case2 (case1 x) = {!!} |
765 ... | case2 (case2 pg) = {!!} | |
754 | 766 |
767 rotateRight : {n m : Level} {A : Set n} {t : Set m} | |
775 | 768 → (key : ℕ) → (value : A) |
769 → (orig tree : bt (Color ∧ A)) | |
770 → (stack : List (bt (Color ∧ A))) | |
771 → (r : RBI key value orig tree stack ) | |
776 | 772 → (next : (current : bt (Color ∧ A)) → (stack1 : List (bt (Color ∧ A))) |
773 → (r : RBI key value orig current stack1 ) | |
754 | 774 → length stack1 < length stack → t ) |
775 | 775 → (exit : (repl : bt (Color ∧ A) ) → (stack1 : List (bt (Color ∧ A))) |
776 → stack1 ≡ (orig ∷ []) | |
777 → RBI key value orig repl stack1 | |
778 → t ) → t | |
779 rotateRight {n} {m} {A} {t} key value = ? where | |
756 | 780 rotateRight1 : t |
775 | 781 rotateRight1 with stackToPG ? ? ? ? |
758 | 782 ... | case1 x = {!!} |
783 ... | case2 (case1 x) = {!!} | |
784 ... | case2 (case2 pg) = {!!} | |
754 | 785 |
729 | 786 insertCase5 : {n m : Level} {A : Set n} {t : Set m} |
775 | 787 → (key : ℕ) → (value : A) |
788 → (orig tree : bt (Color ∧ A)) | |
789 → (stack : List (bt (Color ∧ A))) | |
790 → (r : RBI key value orig tree stack ) | |
776 | 791 → (next : (tree1 : (bt (Color ∧ A))) → (stack1 : List (bt (Color ∧ A))) |
775 | 792 → (r : RBI key value orig tree1 stack1 ) |
751 | 793 → length stack1 < length stack → t ) |
775 | 794 → (exit : (repl : bt (Color ∧ A) ) → (stack1 : List (bt (Color ∧ A))) |
795 → stack1 ≡ (orig ∷ []) | |
796 → RBI key value orig repl stack1 | |
797 → t ) → t | |
798 insertCase5 {n} {m} {A} {t} key value = ? where | |
755 | 799 insertCase51 : t |
776 | 800 insertCase51 with stackToPG ? ? ? ? |
758 | 801 ... | case1 eq = {!!} |
802 ... | case2 (case1 eq ) = {!!} | |
755 | 803 ... | case2 (case2 pg) with PG.pg pg |
775 | 804 ... | s2-s1p2 x x₁ = ? -- rotateRight {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit |
755 | 805 -- x : PG.parent pg ≡ node kp vp tree n1 |
806 -- x₁ : PG.grand pg ≡ node kg vg (PG.parent pg) (PG.uncle pg) | |
775 | 807 ... | s2-1sp2 x x₁ = ? -- rotateLeft {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit |
808 ... | s2-s12p x x₁ = ? -- rotateLeft {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit | |
809 ... | s2-1s2p x x₁ = ? -- rotateLeft {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit | |
755 | 810 -- = insertCase2 tree (PG.parent pg) (PG.uncle pg) (PG.grand pg) stack si (PG.pg pg) |
751 | 811 |
723 | 812 |
749 | 813 -- if we have replacedNode on RBTree, we have RBtreeInvariant |
814 | |
724 | 815 replaceRBP : {n m : Level} {A : Set n} {t : Set m} |
759 | 816 → (key : ℕ) → (value : A) |
777 | 817 → (orig repl : bt (Color ∧ A)) |
770 | 818 → (stack : List (bt (Color ∧ A))) |
777 | 819 → (r : RBI key value orig repl stack ) |
820 → (next : (repl1 : (bt (Color ∧ A))) → (stack1 : List (bt (Color ∧ A))) | |
821 → (r : RBI key value orig repl1 stack1 ) | |
770 | 822 → length stack1 < length stack → t ) |
823 → (exit : (repl : bt (Color ∧ A) ) → (stack1 : List (bt (Color ∧ A))) | |
824 → stack1 ≡ (orig ∷ []) | |
776 | 825 → RBI key value orig repl stack1 |
770 | 826 → t ) → t |
777 | 827 replaceRBP {n} {m} {A} {t} key value orig repl stack r next exit = ? where |
754 | 828 insertCase2 : (tree parent uncle grand : bt (Color ∧ A)) |
749 | 829 → (stack : List (bt (Color ∧ A))) → (si : stackInvariant key tree orig stack ) |
754 | 830 → (pg : ParentGrand tree parent uncle grand ) → t |
831 insertCase2 tree leaf uncle grand stack si (s2-s1p2 () x₁) | |
832 insertCase2 tree leaf uncle grand stack si (s2-1sp2 () x₁) | |
833 insertCase2 tree leaf uncle grand stack si (s2-s12p () x₁) | |
834 insertCase2 tree leaf uncle grand stack si (s2-1s2p () x₁) | |
775 | 835 insertCase2 tree parent@(node kp ⟪ Red , _ ⟫ _ _) uncle grand stack si pg = ? -- next {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} |
758 | 836 insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) leaf grand stack si pg = {!!} |
775 | 837 insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Red , _ ⟫ _ _ ) grand stack si pg = ? -- next {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} |
838 insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-s1p2 x x₁) = ? | |
839 -- insertCase5 key value orig tree {!!} rbio {!!} {!!} stack si {!!} ri {!!} {!!} next exit | |
754 | 840 -- tree is left of parent, parent is left of grand |
841 -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp tree n1 | |
842 -- grand ≡ node kg vg (node kp ⟪ Black , proj3 ⟫ left right) (node ku ⟪ Black , proj4 ⟫ left₁ right₁) | |
775 | 843 insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-1sp2 x x₁) = ? |
844 -- rotateLeft key value orig tree {!!} repl rbio {!!} {!!} stack si {!!} ri {!!} {!!} | |
845 -- (λ a b c d e f h i j k l m → insertCase5 key value a b c d {!!} {!!} h i j k l {!!} {!!} next exit ) exit | |
754 | 846 -- tree is right of parent, parent is left of grand rotateLeft |
847 -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp n1 tree | |
848 -- grand ≡ node kg vg (node kp ⟪ Black , proj3 ⟫ left right) (node ku ⟪ Black , proj4 ⟫ left₁ right₁) | |
775 | 849 insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-s12p x x₁) = ? |
850 -- rotateRight key value orig tree {!!} repl rbio {!!} {!!} stack si {!!} ri {!!} {!!} | |
851 -- (λ a b c d e f h i j k l m → insertCase5 key value a b c d {!!} {!!} h i j k l {!!} {!!} next exit ) exit | |
754 | 852 -- tree is left of parent, parent is right of grand, rotateRight |
853 -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp tree n1 | |
854 -- grand ≡ node kg vg (node ku ⟪ Black , proj4 ⟫ left₁ right₁) (node kp ⟪ Black , proj3 ⟫ left right) | |
775 | 855 insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-1s2p x x₁) = ? |
856 -- insertCase5 key value orig tree {!!} repl rbio {!!} {!!} stack si {!!} ri {!!} {!!} next exit | |
754 | 857 -- tree is right of parent, parent is right of grand |
858 -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp n1 tree | |
859 -- grand ≡ node kg vg (node ku ⟪ Black , proj4 ⟫ left₁ right₁) (node kp ⟪ Black , proj3 ⟫ left right) | |
749 | 860 insertCase1 : t |
777 | 861 insertCase1 with stackToPG (RBI.tree r) orig stack (RBI.si r) |
862 ... | case1 eq = exit repl stack eq record { -- exit rot repl rbir (subst (λ k → rotatedTree k rot) (stackCase1 si eq) roti) ri | |
863 od = ? ; d = ? ; rd = ? | |
864 ; tree = RBI.tree r ; rot = ? | |
865 ; origti = ? | |
866 ; origrb = ? | |
867 ; treerb = ? | |
868 ; replrb = ? | |
869 ; d=rd = ? | |
870 ; si = ? | |
871 ; rotated = ? | |
872 ; ri = ? } | |
760 | 873 ... | case2 (case1 eq ) = ? where |
767
9cdf272ca38c
... we need another invariant on depth of black
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
766
diff
changeset
|
874 insertCase12 : (to : bt (Color ∧ A)) → {d : ℕ} → RBtreeInvariant orig d → to ≡ orig |
775 | 875 → (rr : bt (Color ∧ A)) → {dr : ℕ} → RBtreeInvariant ? dr → rr ≡ ? |
876 → {stack : List (bt (Color ∧ A))} → (si : stackInvariant key ? to stack ) | |
877 → stack ≡ ? ∷ to ∷ [] → t | |
768 | 878 insertCase12 = ? |
760 | 879 -- exit rot repl rbir ? ? |
775 | 880 ... | case2 (case2 pg) = ? -- insertCase2 tree (PG.parent pg) (PG.uncle pg) (PG.grand pg) stack si (PG.pg pg) |
724 | 881 |
740 | 882 |
775 | 883 |