Mercurial > hg > Members > Moririn
annotate hoareBinaryTree.agda @ 690:a971a954a345
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Wed, 01 Dec 2021 11:24:42 +0900 |
parents | 25f89e4bc160 |
children | ca624203b453 |
rev | line source |
---|---|
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
1 module hoareBinaryTree where |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
2 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
3 open import Level renaming (zero to Z ; suc to succ) |
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 |
588 | 21 _iso_ : {n : Level} {a : Set n} → ℕ → ℕ → Set |
22 d iso d' = (¬ (suc d ≤ d')) ∧ (¬ (suc d' ≤ d)) | |
23 | |
24 iso-intro : {n : Level} {a : Set n} {x y : ℕ} → ¬ (suc x ≤ y) → ¬ (suc y ≤ x) → _iso_ {n} {a} x y | |
25 iso-intro = λ z z₁ → record { proj1 = z ; proj2 = z₁ } | |
26 | |
590 | 27 -- |
28 -- | |
29 -- no children , having left node , having right node , having both | |
30 -- | |
597 | 31 data bt {n : Level} (A : Set n) : Set n where |
604 | 32 leaf : bt A |
33 node : (key : ℕ) → (value : A) → | |
610 | 34 (left : bt A ) → (right : bt A ) → bt A |
600 | 35 |
620 | 36 node-key : {n : Level} {A : Set n} → bt A → Maybe ℕ |
37 node-key (node key _ _ _) = just key | |
38 node-key _ = nothing | |
39 | |
40 node-value : {n : Level} {A : Set n} → bt A → Maybe A | |
41 node-value (node _ value _ _) = just value | |
42 node-value _ = nothing | |
43 | |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
44 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
|
45 bt-depth leaf = 0 |
618 | 46 bt-depth (node key value t t₁) = suc (Data.Nat._⊔_ (bt-depth t ) (bt-depth t₁ )) |
606 | 47 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
48 find : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → List (bt A) |
604 | 49 → (next : bt A → List (bt A) → t ) → (exit : bt A → List (bt A) → t ) → t |
50 find key leaf st _ exit = exit leaf st | |
632 | 51 find key (node key₁ v1 tree tree₁) st next exit with <-cmp key key₁ |
604 | 52 find key n st _ exit | tri≈ ¬a b ¬c = exit n st |
632 | 53 find key n@(node key₁ v1 tree tree₁) st next _ | tri< a ¬b ¬c = next tree (n ∷ st) |
54 find key n@(node key₁ v1 tree tree₁) st next _ | tri> ¬a ¬b c = next tree₁ (n ∷ st) | |
597 | 55 |
604 | 56 {-# TERMINATING #-} |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
57 find-loop : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → bt A → List (bt A) → (exit : bt A → List (bt A) → t) → t |
611 | 58 find-loop {n} {m} {A} {t} key tree st exit = find-loop1 tree st where |
604 | 59 find-loop1 : bt A → List (bt A) → t |
60 find-loop1 tree st = find key tree st find-loop1 exit | |
600 | 61 |
611 | 62 replaceNode : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → (bt A → t) → t |
632 | 63 replaceNode k v1 leaf next = next (node k v1 leaf leaf) |
64 replaceNode k v1 (node key value t t₁) next = next (node k v1 t t₁) | |
611 | 65 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
66 replace : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → List (bt A) → (next : ℕ → A → bt A → List (bt A) → t ) → (exit : bt A → t) → t |
669 | 67 replace key value repl [] next exit = exit repl -- can't happen |
690 | 68 replace key value repl (leaf ∷ []) next exit = exit repl |
669 | 69 replace key value repl (node key₁ value₁ left right ∷ []) next exit with <-cmp key key₁ |
70 ... | tri< a ¬b ¬c = exit (node key₁ value₁ repl right ) | |
664 | 71 ... | tri≈ ¬a b ¬c = exit (node key₁ value left right ) |
669 | 72 ... | tri> ¬a ¬b c = exit (node key₁ value₁ left repl ) |
690 | 73 replace key value repl (leaf ∷ st) next exit = next key value repl st |
669 | 74 replace key value repl (node key₁ value₁ left right ∷ st) next exit with <-cmp key key₁ |
75 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st | |
604 | 76 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st |
669 | 77 ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
78 |
604 | 79 {-# TERMINATING #-} |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
80 replace-loop : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → List (bt A) → (exit : bt A → t) → t |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
81 replace-loop {_} {_} {A} {t} key value tree st exit = replace-loop1 key value tree st where |
604 | 82 replace-loop1 : (key : ℕ) → (value : A) → bt A → List (bt A) → t |
83 replace-loop1 key value tree st = replace key value tree st replace-loop1 exit | |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
84 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
85 insertTree : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → (next : bt A → t ) → t |
662 | 86 insertTree tree key value exit = find-loop key tree ( tree ∷ [] ) $ λ t st → replaceNode key value t $ λ t1 → replace-loop key value t1 st exit |
587 | 87 |
604 | 88 insertTest1 = insertTree leaf 1 1 (λ x → x ) |
611 | 89 insertTest2 = insertTree insertTest1 2 1 (λ x → x ) |
669 | 90 insertTest3 = insertTree insertTest2 3 2 (λ x → x ) |
91 insertTest4 = insertTree insertTest3 2 2 (λ x → x ) | |
587 | 92 |
605 | 93 open import Data.Unit hiding ( _≟_ ; _≤?_ ; _≤_) |
94 | |
620 | 95 data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where |
96 t-leaf : treeInvariant leaf | |
632 | 97 t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) |
98 t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key₁ value₁ t₁ t₂) | |
99 → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) | |
100 t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key₁ < key) → treeInvariant (node key value t₁ t₂) | |
101 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) | |
620 | 102 t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → (key < key₁) → (key₁ < key₂) |
103 → treeInvariant (node key value t₁ t₂) | |
104 → treeInvariant (node key₂ value₂ t₃ t₄) | |
105 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) | |
605 | 106 |
662 | 107 -- |
108 -- stack always contains original top at end | |
109 -- | |
110 data stackInvariant {n : Level} {A : Set n} (key : ℕ) : (top orig : bt A) → (stack : List (bt A)) → Set n where | |
675 | 111 s-single : {tree0 : bt A} → stackInvariant key tree0 tree0 (tree0 ∷ []) |
653 | 112 s-right : {tree tree0 tree₁ : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 113 → key₁ < key → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
653 | 114 s-left : {tree₁ tree0 tree : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 115 → key < key₁ → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree₁ tree0 (tree₁ ∷ st) |
639 | 116 |
677 | 117 data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt A ) → Set n where |
639 | 118 r-leaf : replacedTree key value leaf (node key value leaf leaf) |
119 r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) | |
120 r-right : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} | |
677 | 121 → k < key → replacedTree key value t2 t → replacedTree key value (node k v1 t1 t2) (node k v1 t1 t) |
639 | 122 r-left : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} |
687 | 123 → key < k → replacedTree key value t1 t → replacedTree key value (node k v1 t1 t2) (node k v1 t t2) |
652 | 124 |
632 | 125 add< : { i : ℕ } (j : ℕ ) → i < suc i + j |
126 add< {i} j = begin | |
127 suc i ≤⟨ m≤m+n (suc i) j ⟩ | |
128 suc i + j ∎ where open ≤-Reasoning | |
129 | |
130 treeTest1 : bt ℕ | |
131 treeTest1 = node 1 0 leaf (node 3 1 (node 2 5 (node 4 7 leaf leaf ) leaf) (node 5 5 leaf leaf)) | |
132 treeTest2 : bt ℕ | |
133 treeTest2 = node 3 1 (node 2 5 (node 4 7 leaf leaf ) leaf) (node 5 5 leaf leaf) | |
134 | |
135 treeInvariantTest1 : treeInvariant treeTest1 | |
136 treeInvariantTest1 = t-right (m≤m+n _ 1) (t-node (add< 0) (add< 1) (t-left (add< 1) (t-single 4 7)) (t-single 5 5) ) | |
605 | 137 |
639 | 138 stack-top : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
139 stack-top [] = nothing | |
140 stack-top (x ∷ s) = just x | |
606 | 141 |
639 | 142 stack-last : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
143 stack-last [] = nothing | |
144 stack-last (x ∷ []) = just x | |
145 stack-last (x ∷ s) = stack-last s | |
632 | 146 |
662 | 147 stackInvariantTest1 : stackInvariant 4 treeTest2 treeTest1 ( treeTest2 ∷ treeTest1 ∷ [] ) |
675 | 148 stackInvariantTest1 = s-right (add< 2) (s-single ) |
662 | 149 |
666 | 150 si-property0 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 stack → ¬ ( stack ≡ [] ) |
675 | 151 si-property0 (s-single ) () |
666 | 152 si-property0 (s-right x si) () |
153 si-property0 (s-left x si) () | |
665 | 154 |
666 | 155 si-property1 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 tree1 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 (tree1 ∷ stack) |
156 → tree1 ≡ tree | |
675 | 157 si-property1 (s-single ) = refl |
666 | 158 si-property1 (s-right _ si) = refl |
159 si-property1 (s-left _ si) = refl | |
662 | 160 |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
161 si-property-last : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack |
662 | 162 → stack-last stack ≡ just tree0 |
675 | 163 si-property-last key t t0 (t ∷ []) (s-single ) = refl |
666 | 164 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
|
165 ... | refl = si-property-last key x t0 (x ∷ st) si |
666 | 166 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
|
167 ... | refl = si-property-last key x t0 (x ∷ st) si |
656 | 168 |
642 | 169 ti-right : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 tree₁ repl) → treeInvariant repl |
170 ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf | |
171 ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-right x ti) = ti | |
172 ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-left x ti) = t-leaf | |
173 ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti₁ | |
174 | |
175 ti-left : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 repl tree₁ ) → treeInvariant repl | |
176 ti-left {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf | |
177 ti-left {_} {_} {_} {_} {key₁} {v1} (t-right x ti) = t-leaf | |
178 ti-left {_} {_} {_} {_} {key₁} {v1} (t-left x ti) = ti | |
179 ti-left {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti | |
180 | |
662 | 181 stackTreeInvariant : {n : Level} {A : Set n} (key : ℕ) (sub tree : bt A) → (stack : List (bt A)) |
182 → treeInvariant tree → stackInvariant key sub tree stack → treeInvariant sub | |
675 | 183 stackTreeInvariant {_} {A} key sub tree (sub ∷ []) ti (s-single ) = ti |
662 | 184 stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-right _ si ) = ti-right (si1 si) where |
185 si1 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 tree₁ sub ) tree st → treeInvariant (node key₁ v1 tree₁ sub ) | |
186 si1 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 tree₁ sub ) tree st ti si | |
187 stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-left _ si ) = ti-left ( si2 si) where | |
188 si2 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 sub tree₁ ) tree st → treeInvariant (node key₁ v1 sub tree₁ ) | |
189 si2 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 sub tree₁ ) tree st ti si | |
190 | |
639 | 191 rt-property1 : {n : Level} {A : Set n} (key : ℕ) (value : A) (tree tree1 : bt A ) → replacedTree key value tree tree1 → ¬ ( tree1 ≡ leaf ) |
192 rt-property1 {n} {A} key value .leaf .(node key value leaf leaf) r-leaf () | |
193 rt-property1 {n} {A} key value .(node key _ _ _) .(node key value _ _) r-node () | |
677 | 194 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-right x rt) = λ () |
195 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-left x rt) = λ () | |
639 | 196 |
690 | 197 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 |
198 rt-property-leaf r-leaf = refl | |
199 | |
632 | 200 depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) |
201 depth-1< {i} {j} = s≤s (m≤m⊔n _ j) | |
202 | |
203 depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) | |
650 | 204 depth-2< {i} {j} = s≤s (m≤n⊔m j i) |
611 | 205 |
649 | 206 depth-3< : {i : ℕ } → suc i ≤ suc (suc i) |
207 depth-3< {zero} = s≤s ( z≤n ) | |
208 depth-3< {suc i} = s≤s (depth-3< {i} ) | |
209 | |
210 | |
634 | 211 treeLeftDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) |
212 → treeInvariant (node k v1 tree tree₁) | |
213 → treeInvariant tree | |
214 treeLeftDown {n} {A} {_} {v1} leaf leaf (t-single k1 v1) = t-leaf | |
215 treeLeftDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = t-leaf | |
216 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = ti | |
217 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti | |
218 | |
219 treeRightDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) | |
220 → treeInvariant (node k v1 tree tree₁) | |
221 → treeInvariant tree₁ | |
222 treeRightDown {n} {A} {_} {v1} .leaf .leaf (t-single _ .v1) = t-leaf | |
223 treeRightDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = ti | |
224 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = t-leaf | |
225 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti₁ | |
226 | |
664 | 227 nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ |
228 nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x | |
229 nat-<> : { x y : ℕ } → x < y → y < x → ⊥ | |
230 nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x | |
633 | 231 |
232 open _∧_ | |
233 | |
615 | 234 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A)) |
662 | 235 → treeInvariant tree ∧ stackInvariant key tree tree0 stack |
236 → (next : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) | |
237 → (exit : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack | |
638 | 238 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t |
239 findP key leaf tree0 st Pre _ exit = exit leaf tree0 st Pre (case1 refl) | |
632 | 240 findP key (node key₁ v1 tree tree₁) tree0 st Pre next exit with <-cmp key key₁ |
638 | 241 findP key n tree0 st Pre _ exit | tri≈ ¬a refl ¬c = exit n tree0 st Pre (case2 refl) |
664 | 242 findP {n} {_} {A} key (node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree tree0 (tree ∷ st) |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
243 ⟪ 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
|
244 findP1 : key < key₁ → (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
664 | 245 findP1 a (x ∷ st) si = s-left a si |
662 | 246 findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ tree0 (tree₁ ∷ st) ⟪ treeRightDown tree tree₁ (proj1 Pre) , s-right c (proj2 Pre) ⟫ depth-2< |
606 | 247 |
638 | 248 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₁) |
249 replaceTree1 k v1 value (t-single .k .v1) = t-single k value | |
250 replaceTree1 k v1 value (t-right x t) = t-right x t | |
251 replaceTree1 k v1 value (t-left x t) = t-left x t | |
252 replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ | |
253 | |
649 | 254 open import Relation.Binary.Definitions |
255 | |
256 lemma3 : {i j : ℕ} → 0 ≡ i → j < i → ⊥ | |
257 lemma3 refl () | |
258 lemma5 : {i j : ℕ} → i < 1 → j < i → ⊥ | |
259 lemma5 (s≤s z≤n) () | |
260 | |
687 | 261 child-replaced : {n : Level} {A : Set n} (key : ℕ) (tree : bt A) → bt A |
262 child-replaced key leaf = leaf | |
263 child-replaced key (node key₁ value left right) with <-cmp key key₁ | |
264 ... | tri< a ¬b ¬c = left | |
265 ... | tri≈ ¬a b ¬c = node key₁ value left right | |
266 ... | tri> ¬a ¬b c = right | |
677 | 267 |
671
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
268 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
|
269 field |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
270 tree0 : bt A |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
271 ti : treeInvariant tree0 |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
272 si : stackInvariant key tree tree0 stack |
687 | 273 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
|
274 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
|
275 |
638 | 276 replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) |
277 → (tree ≡ leaf ) ∨ ( node-key tree ≡ just key ) | |
278 → (treeInvariant tree ) → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value tree tree1 → t) → t | |
279 replaceNodeP k v1 leaf C P next = next (node k v1 leaf leaf) (t-single k v1 ) r-leaf | |
280 replaceNodeP k v1 (node .k value t t₁) (case2 refl) P next = next (node k v1 t t₁) (replaceTree1 k value v1 P) r-node | |
606 | 281 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
282 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
|
283 → (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
|
284 → (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
|
285 → (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
|
286 → replacePR key value tree1 repl stack1 (λ _ _ _ → Lift n ⊤) → length stack1 < length stack → t) |
613 | 287 → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t |
675 | 288 replaceP key value {tree} repl [] Pre next exit = ⊥-elim ( si-property0 (replacePR.si Pre) refl ) -- can't happen |
289 replaceP key value {tree} repl (leaf ∷ []) Pre next exit with si-property-last _ _ _ _ (replacePR.si Pre)-- tree0 ≡ leaf | |
677 | 290 ... | refl = exit (replacePR.tree0 Pre) (node key value leaf leaf) ⟪ replacePR.ti Pre , r-leaf ⟫ |
689 | 291 replaceP key value {tree} repl (node key₁ value₁ left right ∷ []) Pre next exit with <-cmp key key₁ |
292 ... | tri< a ¬b ¬c = exit (replacePR.tree0 Pre) (node key₁ value₁ repl right ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
293 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ repl right ) | |
294 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
295 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ k right ) (node key₁ value₁ repl right )) repl02 (r-left a repl03) where | |
296 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
297 repl03 = replacePR.ri Pre | |
298 repl02 : child-replaced key (node key₁ value₁ left right) ≡ left | |
299 repl02 with <-cmp key key₁ | |
300 ... | tri< a ¬b ¬c = refl | |
301 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a a) | |
302 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a a) | |
303 ... | tri≈ ¬a b ¬c = exit (replacePR.tree0 Pre) repl ⟪ replacePR.ti Pre , repl01 ⟫ where | |
678 | 304 repl01 : replacedTree key value (replacePR.tree0 Pre) repl |
305 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
689 | 306 repl01 | refl | refl = subst (λ k → replacedTree key value k repl) repl02 (replacePR.ri Pre) where |
307 repl02 : child-replaced key (node key₁ value₁ left right) ≡ node key₁ value₁ left right | |
308 repl02 with <-cmp key key₁ | |
309 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b b) | |
310 ... | tri≈ ¬a b ¬c = refl | |
311 ... | tri> ¬a ¬b c = ⊥-elim ( ¬b b) | |
312 ... | tri> ¬a ¬b c = exit (replacePR.tree0 Pre) (node key₁ value₁ left repl ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
313 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ left repl ) | |
314 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
315 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ left k ) (node key₁ value₁ left repl )) repl02 (r-right c repl03) where | |
316 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
317 repl03 = replacePR.ri Pre | |
318 repl02 : child-replaced key (node key₁ value₁ left right) ≡ right | |
319 repl02 with <-cmp key key₁ | |
320 ... | tri< a ¬b ¬c = ⊥-elim ( ¬c c) | |
321 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬c c) | |
322 ... | tri> ¬a ¬b c = refl | |
690 | 323 replaceP {n} {_} {A} key value {tree} repl (leaf ∷ st@(tree1 ∷ st1)) Pre next exit = next key value repl st Post ≤-refl where |
324 Post : replacePR key value tree1 repl (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) | |
325 Post with replacePR.si Pre | |
326 ... | s-right {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
327 repl09 : tree1 ≡ node key₂ v1 tree₁ leaf | |
328 repl09 = si-property1 si | |
329 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
330 repl10 with si-property1 si | |
331 ... | refl = si | |
332 repl07 : child-replaced key (node key₂ v1 tree₁ leaf) ≡ leaf | |
333 repl07 with <-cmp key key₂ | |
334 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
335 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
336 ... | tri> ¬a ¬b c = refl | |
337 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
338 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 | |
339 ... | s-left {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
340 repl09 : tree1 ≡ node key₂ v1 leaf tree₁ | |
341 repl09 = si-property1 si | |
342 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
343 repl10 with si-property1 si | |
344 ... | refl = si | |
345 repl07 : child-replaced key (node key₂ v1 leaf tree₁ ) ≡ leaf | |
346 repl07 with <-cmp key key₂ | |
347 ... | tri< a ¬b ¬c = refl | |
348 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
349 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
350 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
351 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 | 352 replaceP {n} {_} {A} key value {tree} repl (node key₁ value₁ left right ∷ st@(tree1 ∷ st1)) Pre next exit with <-cmp key key₁ |
353 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st Post ≤-refl where | |
675 | 354 Post : replacePR key value tree1 (node key₁ value₁ repl right ) st (λ _ _ _ → Lift n ⊤) |
687 | 355 Post with replacePR.si Pre |
688 | 356 ... | s-right {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
357 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) | |
358 repl09 = si-property1 si | |
359 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
360 repl10 with si-property1 si | |
361 ... | refl = si | |
362 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left | |
363 repl03 with <-cmp key key₁ | |
364 ... | tri< a1 ¬b ¬c = refl | |
365 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
366 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
367 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
368 repl02 with repl09 | <-cmp key key₂ | |
369 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
689 | 370 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) |
688 | 371 ... | refl | tri> ¬a ¬b c = refl |
372 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
373 repl04 = begin | |
374 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
375 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
376 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
377 child-replaced key tree1 ∎ where open ≡-Reasoning | |
378 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
379 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
687 | 380 ... | 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 | 381 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ |
683 | 382 repl09 = si-property1 si |
383 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
384 repl10 with si-property1 si | |
385 ... | refl = si | |
687 | 386 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left |
387 repl03 with <-cmp key key₁ | |
388 ... | tri< a1 ¬b ¬c = refl | |
389 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
390 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
391 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
392 repl02 with repl09 | <-cmp key key₂ | |
393 ... | refl | tri< a ¬b ¬c = refl | |
394 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
395 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
396 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
397 repl04 = begin | |
398 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
399 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
400 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
401 child-replaced key tree1 ∎ where open ≡-Reasoning | |
402 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
403 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
690 | 404 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st Post ≤-refl where -- can't happen |
405 Post : replacePR key value tree1 (node key₁ value left right ) (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) | |
406 Post with replacePR.si Pre | |
407 ... | s-right {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
408 repl09 : tree1 ≡ node key₂ v1 tree₁ tree -- (node key₁ value₁ left right) | |
409 repl09 = si-property1 si | |
410 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
411 repl10 with si-property1 si | |
412 ... | refl = si | |
413 repl07 : child-replaced key (node key₂ v1 tree₁ tree) ≡ tree | |
414 repl07 with <-cmp key key₂ | |
415 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
416 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
417 ... | tri> ¬a ¬b c = refl | |
418 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) | |
419 repl12 = {!!} | |
420 ... | s-left {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
421 repl09 : tree1 ≡ node key₂ v1 tree tree₁ | |
422 repl09 = si-property1 si | |
423 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
424 repl10 with si-property1 si | |
425 ... | refl = si | |
426 repl07 : child-replaced key (node key₂ v1 tree tree₁ ) ≡ tree | |
427 repl07 with <-cmp key key₂ | |
428 ... | tri< a ¬b ¬c = refl | |
429 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
430 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
431 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) | |
432 repl12 = {!!} | |
433 ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st Post ≤-refl where | |
434 Post : replacePR key value tree1 (node key₁ value₁ left repl ) st (λ _ _ _ → Lift n ⊤) | |
435 Post with replacePR.si Pre | |
436 ... | s-right {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
437 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) | |
438 repl09 = si-property1 si | |
439 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
440 repl10 with si-property1 si | |
441 ... | refl = si | |
442 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
443 repl03 with <-cmp key key₁ | |
444 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
445 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
446 ... | tri> ¬a ¬b c = refl | |
447 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
448 repl02 with repl09 | <-cmp key key₂ | |
449 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
450 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) | |
451 ... | refl | tri> ¬a ¬b c = refl | |
452 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
453 repl04 = begin | |
454 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
455 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
456 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
457 child-replaced key tree1 ∎ where open ≡-Reasoning | |
458 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
459 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
460 ... | s-left {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
461 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ | |
462 repl09 = si-property1 si | |
463 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
464 repl10 with si-property1 si | |
465 ... | refl = si | |
466 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
467 repl03 with <-cmp key key₁ | |
468 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
469 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
470 ... | tri> ¬a ¬b c = refl | |
471 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
472 repl02 with repl09 | <-cmp key key₂ | |
473 ... | refl | tri< a ¬b ¬c = refl | |
474 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
475 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
476 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
477 repl04 = begin | |
478 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
479 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
480 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
481 child-replaced key tree1 ∎ where open ≡-Reasoning | |
482 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
483 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
644 | 484 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
485 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
|
486 → (r : Index) → (p : Invraiant r) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
487 → (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
|
488 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
|
489 ... | 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
|
490 ... | 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
|
491 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
|
492 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
|
493 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
|
494 ... | 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
|
495 ... | 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
|
496 ... | 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
|
497 |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
498 open _∧_ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
499 |
615 | 500 RTtoTI0 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
501 → replacedTree key value tree repl → treeInvariant repl | |
502 RTtoTI0 = {!!} | |
503 | |
504 RTtoTI1 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant repl | |
505 → replacedTree key value tree repl → treeInvariant tree | |
506 RTtoTI1 = {!!} | |
614 | 507 |
611 | 508 insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
613 | 509 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t |
610 | 510 insertTreeP {n} {m} {A} {t} tree key value P exit = |
662 | 511 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant key (proj1 p) tree (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ ⟪ P , {!!} ⟫ |
615 | 512 $ λ p P loop → findP key (proj1 p) tree (proj2 p) {!!} (λ t _ s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) |
638 | 513 $ λ t _ s P C → replaceNodeP key value t C (proj1 P) |
614 | 514 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) |
662 | 515 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant key (proj1 (proj2 p)) tree (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } |
639 | 516 (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ proj1 P , ⟪ {!!} , R ⟫ ⟫ |
644 | 517 $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) {!!} |
518 (λ key value repl1 stack P2 lt → loop ⟪ stack , ⟪ {!!} , repl1 ⟫ ⟫ {!!} lt ) exit | |
614 | 519 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
520 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
|
521 top-value leaf = nothing |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
522 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
|
523 |
612 | 524 insertTreeSpec0 : {n : Level} {A : Set n} → (tree : bt A) → (value : A) → top-value tree ≡ just value → ⊤ |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
525 insertTreeSpec0 _ _ _ = tt |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
526 |
627 | 527 record findPR {n : Level} {A : Set n} (key : ℕ) (tree : bt A ) (stack : List (bt A)) (C : bt A → List (bt A) → Set n) : Set n where |
618 | 528 field |
619 | 529 tree0 : bt A |
622 | 530 ti : treeInvariant tree0 |
662 | 531 si : stackInvariant key tree tree0 stack |
631 | 532 ci : C tree stack -- data continuation |
618 | 533 |
616 | 534 findPP : {n m : Level} {A : Set n} {t : Set m} |
535 → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) | |
627 | 536 → (Pre : findPR key tree stack (λ t s → Lift n ⊤)) |
537 → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR key tree1 stack1 (λ t s → Lift n ⊤) → bt-depth tree1 < bt-depth tree → t ) | |
538 → (exit : (tree1 : bt A) → (stack1 : List (bt A)) → ( tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key) → findPR key tree1 stack1 (λ t s → Lift n ⊤) → t) → t | |
625 | 539 findPP key leaf st Pre next exit = exit leaf st (case1 refl) Pre |
632 | 540 findPP key (node key₁ v1 tree tree₁) st Pre next exit with <-cmp key key₁ |
625 | 541 findPP key n st P next exit | tri≈ ¬a b ¬c = exit n st (case2 {!!}) P |
632 | 542 findPP {_} {_} {A} key n@(node key₁ v1 tree tree₁) st Pre next exit | tri< a ¬b ¬c = |
624 | 543 next tree (n ∷ st) (record {ti = findPR.ti Pre ; si = findPP2 st (findPR.si Pre) ; ci = lift tt} ) findPP1 where |
621 | 544 tree0 = findPR.tree0 Pre |
689 | 545 findPP2 : (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (node key₁ v1 tree tree₁ ∷ st) |
623 | 546 findPP2 = {!!} |
618 | 547 findPP1 : suc ( bt-depth tree ) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) |
634 | 548 findPP1 = depth-1< |
632 | 549 findPP key n@(node key₁ v1 tree tree₁) st Pre next exit | tri> ¬a ¬b c = next tree₁ (n ∷ st) {!!} findPP2 where -- Cond n st → Cond tree₁ (n ∷ st) |
618 | 550 findPP2 : suc (bt-depth tree₁) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) |
634 | 551 findPP2 = depth-2< |
616 | 552 |
618 | 553 insertTreePP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
554 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t | |
624 | 555 insertTreePP {n} {m} {A} {t} tree key value P exit = |
627 | 556 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ t s → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ {!!} |
630 | 557 $ λ p P loop → findPP key (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
638 | 558 $ λ t s _ P → replaceNodeP key value t {!!} {!!} |
618 | 559 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) |
662 | 560 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant key (proj1 (proj2 p)) tree (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } |
639 | 561 (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ {!!} , ⟪ {!!} , R ⟫ ⟫ |
644 | 562 $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) {!!} |
563 (λ key value repl1 stack P2 lt → loop ⟪ stack , ⟪ {!!} , repl1 ⟫ ⟫ {!!} lt ) exit | |
618 | 564 |
629 | 565 record findPC {n : Level} {A : Set n} (key1 : ℕ) (value1 : A) (tree : bt A ) (stack : List (bt A)) : Set n where |
616 | 566 field |
567 tree1 : bt A | |
617 | 568 ci : replacedTree key1 value1 tree tree1 |
616 | 569 |
624 | 570 findPPC : {n m : Level} {A : Set n} {t : Set m} |
628 | 571 → (key : ℕ) → (value : A) → (tree : bt A ) → (stack : List (bt A)) |
629 | 572 → (Pre : findPR key tree stack (findPC key value)) |
573 → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR key tree1 stack1 (findPC key value) → bt-depth tree1 < bt-depth tree → t ) | |
574 → (exit : (tree1 : bt A) → (stack1 : List (bt A)) → ( tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key) → findPR key tree1 stack1 (findPC key value) → t) → t | |
575 findPPC key value leaf st Pre next exit = exit leaf st (case1 refl) Pre | |
632 | 576 findPPC key value (node key₁ v1 tree tree₁) st Pre next exit with <-cmp key key₁ |
629 | 577 findPPC key value n st P next exit | tri≈ ¬a b ¬c = exit n st (case2 {!!}) P |
632 | 578 findPPC {_} {_} {A} key value n@(node key₁ v1 tree tree₁) st Pre next exit | tri< a ¬b ¬c = |
629 | 579 next tree (n ∷ st) (record {ti = findPR.ti Pre ; si = {!!} ; ci = {!!} } ) {!!} |
580 findPPC key value n st P next exit | tri> ¬a ¬b c = {!!} | |
624 | 581 |
618 | 582 containsTree : {n m : Level} {A : Set n} {t : Set m} → (tree tree1 : bt A) → (key : ℕ) → (value : A) → treeInvariant tree1 → replacedTree key value tree1 tree → ⊤ |
615 | 583 containsTree {n} {m} {A} {t} tree tree1 key value P RT = |
617 | 584 TerminatingLoopS (bt A ∧ List (bt A) ) |
634 | 585 {λ p → findPR key (proj1 p) (proj2 p) (findPC key value ) } (λ p → bt-depth (proj1 p)) -- findPR key tree1 [] (findPC key value) |
586 ⟪ tree1 , [] ⟫ record { tree0 = tree ; ti = {!!} ; si = {!!} ; ci = record { tree1 = tree ; ci = RT } } | |
630 | 587 $ λ p P loop → findPPC key value (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
629 | 588 $ λ t1 s1 found? P2 → insertTreeSpec0 t1 value (lemma6 t1 s1 found? P2) where |
589 lemma6 : (t1 : bt A) (s1 : List (bt A)) (found? : (t1 ≡ leaf) ∨ (node-key t1 ≡ just key)) (P2 : findPR key t1 s1 (findPC key value)) → top-value t1 ≡ just value | |
590 lemma6 t1 s1 found? P2 = lemma7 t1 s1 (findPR.tree0 P2) ( findPC.tree1 (findPR.ci P2)) ( findPC.ci (findPR.ci P2)) (findPR.si P2) found? where | |
591 lemma7 : (t1 : bt A) ( s1 : List (bt A) ) (tree0 tree1 : bt A) → | |
662 | 592 replacedTree key value t1 tree1 → stackInvariant key t1 tree0 s1 → ( t1 ≡ leaf ) ∨ ( node-key t1 ≡ just key) → top-value t1 ≡ just value |
629 | 593 lemma7 = {!!} |
615 | 594 |