Mercurial > hg > Gears > GearsAgda
annotate hoareBinaryTree.agda @ 956:bfc7007177d0 default tip
safe and cubical compatible with no warning done
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sat, 19 Oct 2024 09:48:48 +0900 |
parents | f2a3f5707075 |
children |
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 |
933 | 10 open import Data.List hiding ( find ) |
586
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 |
716 | 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 |
716 | 50 find' key leaf st _ exit = exit leaf st |
51 find' key (node key₁ v1 tree tree₁) st next exit with <-cmp key key₁ | |
52 find' key n st _ exit | tri≈ ¬a b ¬c = exit n st | |
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) | |
55 | |
56 find : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) | |
57 → (next : (tree1 : bt A) → (stack : List (bt A)) → t) | |
58 → (exit : (tree1 : bt A) → (stack : List (bt A)) → t) → t | |
59 find key leaf st _ exit = exit leaf st | |
632 | 60 find key (node key₁ v1 tree tree₁) st next exit with <-cmp key key₁ |
716 | 61 find key n st _ exit | tri≈ ¬a refl ¬c = exit n st |
62 find {n} {_} {A} key (node key₁ v1 tree tree₁) st next _ | tri< a ¬b ¬c = next tree (tree ∷ st) | |
63 find key n@(node key₁ v1 tree tree₁) st next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st) | |
597 | 64 |
604 | 65 {-# TERMINATING #-} |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
66 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 | 67 find-loop {n} {m} {A} {t} key tree st exit = find-loop1 tree st where |
604 | 68 find-loop1 : bt A → List (bt A) → t |
716 | 69 find-loop1 tree st = find key tree st find-loop1 exit |
600 | 70 |
611 | 71 replaceNode : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → (bt A → t) → t |
632 | 72 replaceNode k v1 leaf next = next (node k v1 leaf leaf) |
73 replaceNode k v1 (node key value t t₁) next = next (node k v1 t t₁) | |
611 | 74 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
75 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 | 76 replace key value repl [] next exit = exit repl -- can't happen |
690 | 77 replace key value repl (leaf ∷ []) next exit = exit repl |
669 | 78 replace key value repl (node key₁ value₁ left right ∷ []) next exit with <-cmp key key₁ |
79 ... | tri< a ¬b ¬c = exit (node key₁ value₁ repl right ) | |
664 | 80 ... | tri≈ ¬a b ¬c = exit (node key₁ value left right ) |
669 | 81 ... | tri> ¬a ¬b c = exit (node key₁ value₁ left repl ) |
716 | 82 replace key value repl (leaf ∷ st) next exit = next key value repl st |
669 | 83 replace key value repl (node key₁ value₁ left right ∷ st) next exit with <-cmp key key₁ |
84 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st | |
604 | 85 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st |
669 | 86 ... | 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
|
87 |
604 | 88 {-# TERMINATING #-} |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
89 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
|
90 replace-loop {_} {_} {A} {t} key value tree st exit = replace-loop1 key value tree st where |
604 | 91 replace-loop1 : (key : ℕ) → (value : A) → bt A → List (bt A) → t |
92 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
|
93 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
94 insertTree : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → (next : bt A → t ) → t |
662 | 95 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 | 96 |
604 | 97 insertTest1 = insertTree leaf 1 1 (λ x → x ) |
611 | 98 insertTest2 = insertTree insertTest1 2 1 (λ x → x ) |
718 | 99 insertTest3 = insertTree insertTest2 3 3 (λ x → x ) |
100 insertTest4 = insertTree insertTest3 1 4 (λ x → x ) -- this is wrong | |
587 | 101 |
710 | 102 updateTree : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → (empty : bt A → t ) → (next : A → bt A → t ) → t |
103 updateTree {_} {_} {A} {t} tree key value empty next = find-loop key tree ( tree ∷ [] ) | |
104 $ λ t st → replaceNode key value t $ λ t1 → replace-loop key value t1 st (found? st) where | |
105 found? : List (bt A) → bt A → t | |
106 found? [] tree = empty tree -- can't happen | |
107 found? (leaf ∷ st) tree = empty tree | |
108 found? (node key value x x₁ ∷ st) tree = next value tree | |
109 | |
933 | 110 open import Data.Unit hiding ( _≟_ ) -- ; _≤?_ ; _≤_) |
605 | 111 |
620 | 112 data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where |
113 t-leaf : treeInvariant leaf | |
632 | 114 t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) |
115 t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key₁ value₁ t₁ t₂) | |
116 → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) | |
692 | 117 t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key value t₁ t₂) |
632 | 118 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) |
620 | 119 t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → (key < key₁) → (key₁ < key₂) |
120 → treeInvariant (node key value t₁ t₂) | |
121 → treeInvariant (node key₂ value₂ t₃ t₄) | |
122 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) | |
605 | 123 |
662 | 124 -- |
125 -- stack always contains original top at end | |
126 -- | |
127 data stackInvariant {n : Level} {A : Set n} (key : ℕ) : (top orig : bt A) → (stack : List (bt A)) → Set n where | |
675 | 128 s-single : {tree0 : bt A} → stackInvariant key tree0 tree0 (tree0 ∷ []) |
653 | 129 s-right : {tree tree0 tree₁ : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 130 → key₁ < key → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
653 | 131 s-left : {tree₁ tree0 tree : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 132 → key < key₁ → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree₁ tree0 (tree₁ ∷ st) |
639 | 133 |
677 | 134 data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt A ) → Set n where |
639 | 135 r-leaf : replacedTree key value leaf (node key value leaf leaf) |
136 r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) | |
137 r-right : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} | |
677 | 138 → k < key → replacedTree key value t2 t → replacedTree key value (node k v1 t1 t2) (node k v1 t1 t) |
639 | 139 r-left : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} |
687 | 140 → key < k → replacedTree key value t1 t → replacedTree key value (node k v1 t1 t2) (node k v1 t t2) |
652 | 141 |
632 | 142 add< : { i : ℕ } (j : ℕ ) → i < suc i + j |
143 add< {i} j = begin | |
144 suc i ≤⟨ m≤m+n (suc i) j ⟩ | |
145 suc i + j ∎ where open ≤-Reasoning | |
146 | |
147 treeTest1 : bt ℕ | |
692 | 148 treeTest1 = node 0 0 leaf (node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf)) |
632 | 149 treeTest2 : bt ℕ |
692 | 150 treeTest2 = node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf) |
632 | 151 |
152 treeInvariantTest1 : treeInvariant treeTest1 | |
692 | 153 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 | 154 |
639 | 155 stack-top : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
156 stack-top [] = nothing | |
157 stack-top (x ∷ s) = just x | |
606 | 158 |
639 | 159 stack-last : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
160 stack-last [] = nothing | |
161 stack-last (x ∷ []) = just x | |
162 stack-last (x ∷ s) = stack-last s | |
632 | 163 |
662 | 164 stackInvariantTest1 : stackInvariant 4 treeTest2 treeTest1 ( treeTest2 ∷ treeTest1 ∷ [] ) |
692 | 165 stackInvariantTest1 = s-right (add< 3) (s-single ) |
662 | 166 |
666 | 167 si-property0 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 stack → ¬ ( stack ≡ [] ) |
675 | 168 si-property0 (s-single ) () |
666 | 169 si-property0 (s-right x si) () |
170 si-property0 (s-left x si) () | |
665 | 171 |
666 | 172 si-property1 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 tree1 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 (tree1 ∷ stack) |
173 → tree1 ≡ tree | |
675 | 174 si-property1 (s-single ) = refl |
666 | 175 si-property1 (s-right _ si) = refl |
176 si-property1 (s-left _ si) = refl | |
662 | 177 |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
178 si-property-last : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack |
662 | 179 → stack-last stack ≡ just tree0 |
675 | 180 si-property-last key t t0 (t ∷ []) (s-single ) = refl |
666 | 181 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
|
182 ... | refl = si-property-last key x t0 (x ∷ st) si |
666 | 183 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
|
184 ... | refl = si-property-last key x t0 (x ∷ st) si |
656 | 185 |
642 | 186 ti-right : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 tree₁ repl) → treeInvariant repl |
187 ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf | |
188 ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-right x ti) = ti | |
189 ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-left x ti) = t-leaf | |
190 ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti₁ | |
191 | |
192 ti-left : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 repl tree₁ ) → treeInvariant repl | |
193 ti-left {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf | |
194 ti-left {_} {_} {_} {_} {key₁} {v1} (t-right x ti) = t-leaf | |
195 ti-left {_} {_} {_} {_} {key₁} {v1} (t-left x ti) = ti | |
196 ti-left {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti | |
197 | |
662 | 198 stackTreeInvariant : {n : Level} {A : Set n} (key : ℕ) (sub tree : bt A) → (stack : List (bt A)) |
199 → treeInvariant tree → stackInvariant key sub tree stack → treeInvariant sub | |
675 | 200 stackTreeInvariant {_} {A} key sub tree (sub ∷ []) ti (s-single ) = ti |
662 | 201 stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-right _ si ) = ti-right (si1 si) where |
202 si1 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 tree₁ sub ) tree st → treeInvariant (node key₁ v1 tree₁ sub ) | |
203 si1 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 tree₁ sub ) tree st ti si | |
204 stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-left _ si ) = ti-left ( si2 si) where | |
205 si2 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 sub tree₁ ) tree st → treeInvariant (node key₁ v1 sub tree₁ ) | |
206 si2 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 sub tree₁ ) tree st ti si | |
207 | |
639 | 208 rt-property1 : {n : Level} {A : Set n} (key : ℕ) (value : A) (tree tree1 : bt A ) → replacedTree key value tree tree1 → ¬ ( tree1 ≡ leaf ) |
209 rt-property1 {n} {A} key value .leaf .(node key value leaf leaf) r-leaf () | |
210 rt-property1 {n} {A} key value .(node key _ _ _) .(node key value _ _) r-node () | |
677 | 211 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-right x rt) = λ () |
212 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-left x rt) = λ () | |
639 | 213 |
690 | 214 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 |
215 rt-property-leaf r-leaf = refl | |
216 | |
698 | 217 rt-property-¬leaf : {n : Level} {A : Set n} {key : ℕ} {value : A} {tree : bt A} → ¬ replacedTree key value tree leaf |
218 rt-property-¬leaf () | |
219 | |
692 | 220 rt-property-key : {n : Level} {A : Set n} {key key₂ key₃ : ℕ} {value value₂ value₃ : A} {left left₁ right₂ right₃ : bt A} |
221 → replacedTree key value (node key₂ value₂ left right₂) (node key₃ value₃ left₁ right₃) → key₂ ≡ key₃ | |
222 rt-property-key r-node = refl | |
223 rt-property-key (r-right x ri) = refl | |
224 rt-property-key (r-left x ri) = refl | |
225 | |
698 | 226 nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ |
227 nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x | |
228 nat-<> : { x y : ℕ } → x < y → y < x → ⊥ | |
229 nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x | |
230 | |
231 open _∧_ | |
232 | |
233 | |
632 | 234 depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) |
235 depth-1< {i} {j} = s≤s (m≤m⊔n _ j) | |
236 | |
237 depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) | |
933 | 238 depth-2< {i} {j} = s≤s (m≤n⊔m j i) |
611 | 239 |
649 | 240 depth-3< : {i : ℕ } → suc i ≤ suc (suc i) |
241 depth-3< {zero} = s≤s ( z≤n ) | |
242 depth-3< {suc i} = s≤s (depth-3< {i} ) | |
243 | |
244 | |
634 | 245 treeLeftDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) |
246 → treeInvariant (node k v1 tree tree₁) | |
247 → treeInvariant tree | |
248 treeLeftDown {n} {A} {_} {v1} leaf leaf (t-single k1 v1) = t-leaf | |
249 treeLeftDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = t-leaf | |
250 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = ti | |
251 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti | |
252 | |
253 treeRightDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) | |
254 → treeInvariant (node k v1 tree tree₁) | |
255 → treeInvariant tree₁ | |
256 treeRightDown {n} {A} {_} {v1} .leaf .leaf (t-single _ .v1) = t-leaf | |
257 treeRightDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = ti | |
258 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = t-leaf | |
259 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti₁ | |
260 | |
704 | 261 |
718 | 262 -- record FindCond {n : Level} {A : Set n} (C : ℕ → bt A → Set n) : Set (Level.suc n) where |
263 -- field | |
264 -- c1 : {key key₁ : ℕ} {v1 : A } { tree tree₁ : bt A } → C key (node key₁ v1 tree tree₁) → key < key₁ → C key tree | |
265 -- c2 : {key key₁ : ℕ} {v1 : A } { tree tree₁ : bt A } → C key (node key₁ v1 tree tree₁) → key > key₁ → C key tree₁ | |
266 -- | |
267 -- | |
268 -- findP0 : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) | |
269 -- → {C : ℕ → bt A → Set n} → C key tree → FindCond C | |
270 -- → (next : (tree1 : bt A) → (stack : List (bt A)) → C key tree1 → bt-depth tree1 < bt-depth tree → t ) | |
271 -- → (exit : (tree1 : bt A) → (stack : List (bt A)) → C key tree1 | |
272 -- → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t | |
273 -- findP0 key leaf st Pre _ _ exit = exit leaf st Pre (case1 refl) | |
274 -- findP0 key (node key₁ v1 tree tree₁) st Pre _ next exit with <-cmp key key₁ | |
275 -- findP0 key n st Pre e _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl) | |
276 -- findP0 {n} {_} {A} key (node key₁ v1 tree tree₁) st Pre e next _ | tri< a ¬b ¬c = next tree (tree ∷ st) (FindCond.c1 e Pre a) depth-1< | |
277 -- findP0 key n@(node key₁ v1 tree tree₁) st Pre e next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st) (FindCond.c2 e Pre c) depth-2< | |
704 | 278 |
615 | 279 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A)) |
662 | 280 → treeInvariant tree ∧ stackInvariant key tree tree0 stack |
693 | 281 → (next : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) |
282 → (exit : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack | |
638 | 283 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t |
693 | 284 findP key leaf tree0 st Pre _ exit = exit leaf st Pre (case1 refl) |
632 | 285 findP key (node key₁ v1 tree tree₁) tree0 st Pre next exit with <-cmp key key₁ |
693 | 286 findP key n tree0 st Pre _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl) |
287 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
|
288 ⟪ 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
|
289 findP1 : key < key₁ → (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
664 | 290 findP1 a (x ∷ st) si = s-left a si |
693 | 291 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 | 292 |
638 | 293 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₁) |
294 replaceTree1 k v1 value (t-single .k .v1) = t-single k value | |
295 replaceTree1 k v1 value (t-right x t) = t-right x t | |
296 replaceTree1 k v1 value (t-left x t) = t-left x t | |
297 replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ | |
298 | |
649 | 299 open import Relation.Binary.Definitions |
300 | |
301 lemma3 : {i j : ℕ} → 0 ≡ i → j < i → ⊥ | |
302 lemma3 refl () | |
303 lemma5 : {i j : ℕ} → i < 1 → j < i → ⊥ | |
304 lemma5 (s≤s z≤n) () | |
700 | 305 ¬x<x : {x : ℕ} → ¬ (x < x) |
306 ¬x<x (s≤s lt) = ¬x<x lt | |
649 | 307 |
687 | 308 child-replaced : {n : Level} {A : Set n} (key : ℕ) (tree : bt A) → bt A |
309 child-replaced key leaf = leaf | |
310 child-replaced key (node key₁ value left right) with <-cmp key key₁ | |
311 ... | tri< a ¬b ¬c = left | |
312 ... | tri≈ ¬a b ¬c = node key₁ value left right | |
313 ... | tri> ¬a ¬b c = right | |
677 | 314 |
671
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
315 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
|
316 field |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
317 tree0 : bt A |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
318 ti : treeInvariant tree0 |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
319 si : stackInvariant key tree tree0 stack |
687 | 320 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
|
321 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
|
322 |
638 | 323 replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) |
324 → (tree ≡ leaf ) ∨ ( node-key tree ≡ just key ) | |
694 | 325 → (treeInvariant tree ) → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value (child-replaced key tree) tree1 → t) → t |
326 replaceNodeP k v1 leaf C P next = next (node k v1 leaf leaf) (t-single k v1 ) r-leaf | |
327 replaceNodeP k v1 (node .k value t t₁) (case2 refl) P next = next (node k v1 t t₁) (replaceTree1 k value v1 P) | |
695 | 328 (subst (λ j → replacedTree k v1 j (node k v1 t t₁) ) repl00 r-node) where |
694 | 329 repl00 : node k value t t₁ ≡ child-replaced k (node k value t t₁) |
330 repl00 with <-cmp k k | |
331 ... | tri< a ¬b ¬c = ⊥-elim (¬b refl) | |
332 ... | tri≈ ¬a b ¬c = refl | |
333 ... | tri> ¬a ¬b c = ⊥-elim (¬b refl) | |
606 | 334 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
335 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
|
336 → (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
|
337 → (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
|
338 → (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
|
339 → replacePR key value tree1 repl stack1 (λ _ _ _ → Lift n ⊤) → length stack1 < length stack → t) |
613 | 340 → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t |
675 | 341 replaceP key value {tree} repl [] Pre next exit = ⊥-elim ( si-property0 (replacePR.si Pre) refl ) -- can't happen |
342 replaceP key value {tree} repl (leaf ∷ []) Pre next exit with si-property-last _ _ _ _ (replacePR.si Pre)-- tree0 ≡ leaf | |
677 | 343 ... | refl = exit (replacePR.tree0 Pre) (node key value leaf leaf) ⟪ replacePR.ti Pre , r-leaf ⟫ |
689 | 344 replaceP key value {tree} repl (node key₁ value₁ left right ∷ []) Pre next exit with <-cmp key key₁ |
345 ... | tri< a ¬b ¬c = exit (replacePR.tree0 Pre) (node key₁ value₁ repl right ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
346 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ repl right ) | |
347 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
348 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ k right ) (node key₁ value₁ repl right )) repl02 (r-left a repl03) where | |
349 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
350 repl03 = replacePR.ri Pre | |
351 repl02 : child-replaced key (node key₁ value₁ left right) ≡ left | |
352 repl02 with <-cmp key key₁ | |
353 ... | tri< a ¬b ¬c = refl | |
354 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a a) | |
355 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a a) | |
356 ... | tri≈ ¬a b ¬c = exit (replacePR.tree0 Pre) repl ⟪ replacePR.ti Pre , repl01 ⟫ where | |
678 | 357 repl01 : replacedTree key value (replacePR.tree0 Pre) repl |
358 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
689 | 359 repl01 | refl | refl = subst (λ k → replacedTree key value k repl) repl02 (replacePR.ri Pre) where |
360 repl02 : child-replaced key (node key₁ value₁ left right) ≡ node key₁ value₁ left right | |
361 repl02 with <-cmp key key₁ | |
362 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b b) | |
363 ... | tri≈ ¬a b ¬c = refl | |
364 ... | tri> ¬a ¬b c = ⊥-elim ( ¬b b) | |
365 ... | tri> ¬a ¬b c = exit (replacePR.tree0 Pre) (node key₁ value₁ left repl ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
366 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ left repl ) | |
367 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
368 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ left k ) (node key₁ value₁ left repl )) repl02 (r-right c repl03) where | |
369 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
370 repl03 = replacePR.ri Pre | |
371 repl02 : child-replaced key (node key₁ value₁ left right) ≡ right | |
372 repl02 with <-cmp key key₁ | |
373 ... | tri< a ¬b ¬c = ⊥-elim ( ¬c c) | |
374 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬c c) | |
375 ... | tri> ¬a ¬b c = refl | |
690 | 376 replaceP {n} {_} {A} key value {tree} repl (leaf ∷ st@(tree1 ∷ st1)) Pre next exit = next key value repl st Post ≤-refl where |
377 Post : replacePR key value tree1 repl (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) | |
378 Post with replacePR.si Pre | |
379 ... | s-right {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
380 repl09 : tree1 ≡ node key₂ v1 tree₁ leaf | |
381 repl09 = si-property1 si | |
382 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
383 repl10 with si-property1 si | |
384 ... | refl = si | |
385 repl07 : child-replaced key (node key₂ v1 tree₁ leaf) ≡ leaf | |
386 repl07 with <-cmp key key₂ | |
387 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
388 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
389 ... | tri> ¬a ¬b c = refl | |
390 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
391 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 | |
392 ... | s-left {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
393 repl09 : tree1 ≡ node key₂ v1 leaf tree₁ | |
394 repl09 = si-property1 si | |
395 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
396 repl10 with si-property1 si | |
397 ... | refl = si | |
398 repl07 : child-replaced key (node key₂ v1 leaf tree₁ ) ≡ leaf | |
399 repl07 with <-cmp key key₂ | |
400 ... | tri< a ¬b ¬c = refl | |
401 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
402 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
403 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
404 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 | 405 replaceP {n} {_} {A} key value {tree} repl (node key₁ value₁ left right ∷ st@(tree1 ∷ st1)) Pre next exit with <-cmp key key₁ |
406 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st Post ≤-refl where | |
675 | 407 Post : replacePR key value tree1 (node key₁ value₁ repl right ) st (λ _ _ _ → Lift n ⊤) |
687 | 408 Post with replacePR.si Pre |
688 | 409 ... | s-right {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
410 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) | |
411 repl09 = si-property1 si | |
412 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
413 repl10 with si-property1 si | |
414 ... | refl = si | |
415 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left | |
416 repl03 with <-cmp key key₁ | |
417 ... | tri< a1 ¬b ¬c = refl | |
418 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
419 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
420 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
421 repl02 with repl09 | <-cmp key key₂ | |
422 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
689 | 423 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) |
688 | 424 ... | refl | tri> ¬a ¬b c = refl |
425 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
426 repl04 = begin | |
427 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
428 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
429 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
430 child-replaced key tree1 ∎ where open ≡-Reasoning | |
431 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
432 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
687 | 433 ... | 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 | 434 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ |
683 | 435 repl09 = si-property1 si |
436 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
437 repl10 with si-property1 si | |
438 ... | refl = si | |
687 | 439 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left |
440 repl03 with <-cmp key key₁ | |
441 ... | tri< a1 ¬b ¬c = refl | |
442 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
443 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
444 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
445 repl02 with repl09 | <-cmp key key₂ | |
446 ... | refl | tri< a ¬b ¬c = refl | |
447 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
448 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
449 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
450 repl04 = begin | |
451 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
452 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
453 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
454 child-replaced key tree1 ∎ where open ≡-Reasoning | |
455 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
456 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
705 | 457 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st Post ≤-refl where |
690 | 458 Post : replacePR key value tree1 (node key₁ value left right ) (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) |
459 Post with replacePR.si Pre | |
691 | 460 ... | 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 | 461 repl09 : tree1 ≡ node key₂ v1 tree₁ tree -- (node key₁ value₁ left right) |
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 repl07 : child-replaced key (node key₂ v1 tree₁ tree) ≡ tree | |
467 repl07 with <-cmp key key₂ | |
468 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
469 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
470 ... | tri> ¬a ¬b c = refl | |
691 | 471 repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) |
472 repl12 refl with repl09 | |
473 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node | |
474 ... | 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 | 475 repl09 : tree1 ≡ node key₂ v1 tree tree₁ |
476 repl09 = si-property1 si | |
477 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
478 repl10 with si-property1 si | |
479 ... | refl = si | |
480 repl07 : child-replaced key (node key₂ v1 tree tree₁ ) ≡ tree | |
481 repl07 with <-cmp key key₂ | |
482 ... | tri< a ¬b ¬c = refl | |
483 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
484 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
691 | 485 repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) |
486 repl12 refl with repl09 | |
487 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node | |
690 | 488 ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st Post ≤-refl where |
489 Post : replacePR key value tree1 (node key₁ value₁ left repl ) st (λ _ _ _ → Lift n ⊤) | |
490 Post with replacePR.si Pre | |
491 ... | s-right {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
492 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) | |
493 repl09 = si-property1 si | |
494 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
495 repl10 with si-property1 si | |
496 ... | refl = si | |
497 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
498 repl03 with <-cmp key key₁ | |
499 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
500 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
501 ... | tri> ¬a ¬b c = refl | |
502 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
503 repl02 with repl09 | <-cmp key key₂ | |
504 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
505 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) | |
506 ... | refl | tri> ¬a ¬b c = refl | |
507 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
508 repl04 = begin | |
509 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
510 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
511 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
512 child-replaced key tree1 ∎ where open ≡-Reasoning | |
513 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
514 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
515 ... | s-left {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
516 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ | |
517 repl09 = si-property1 si | |
518 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
519 repl10 with si-property1 si | |
520 ... | refl = si | |
521 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
522 repl03 with <-cmp key key₁ | |
523 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
524 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
525 ... | tri> ¬a ¬b c = refl | |
526 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
527 repl02 with repl09 | <-cmp key key₂ | |
528 ... | refl | tri< a ¬b ¬c = refl | |
529 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
530 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
531 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
532 repl04 = begin | |
533 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
534 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
535 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
536 child-replaced key tree1 ∎ where open ≡-Reasoning | |
537 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
538 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
644 | 539 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
540 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
|
541 → (r : Index) → (p : Invraiant r) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
542 → (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
|
543 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
|
544 ... | tri≈ ¬a b ¬c = loop r p (λ r1 p1 lt → ⊥-elim (lemma3 b lt) ) |
933 | 545 ... | tri< a ¬b ¬c = loop r p (λ r1 p1 lt1 → TerminatingLoop1 (reduce r) r r1 (m≤n⇒m≤1+n lt1) p1 lt1 ) where |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
546 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
|
547 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
|
548 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
|
549 ... | 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
|
550 ... | 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
|
551 ... | 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
|
552 |
707 | 553 record LoopTermination {n : Level} {Index : Set n } { reduce : Index → ℕ } |
554 (r : Index ) (C : Set n) : Set (Level.suc n) where | |
555 field | |
556 rd : (r1 : Index) → reduce r1 < reduce r | |
557 ci : C -- data continuation | |
558 | |
710 | 559 -- TerminatingLoopC : {l n : Level} {t : Set l} (Index : Set n ) → {C : Set n } → ( reduce : Index → ℕ) |
560 -- → (r : Index) → (P : LoopTermination r C ) | |
561 -- → (loop : (r : Index) → LoopTermination {_} {_} {reduce} r C → (next : (r1 : Index) → LoopTermination r1 C → t ) → t) → t | |
562 -- TerminatingLoopC {_} {_} {t} Index {C} reduce r P loop with <-cmp 0 (reduce r) | |
563 -- ... | tri≈ ¬a b ¬c = loop r P (λ r1 p1 → ⊥-elim (lemma3 b (LoopTermination.rd P r1))) | |
564 -- ... | tri< a ¬b ¬c = loop r P (λ r1 p1 → TerminatingLoop1 (reduce r) r r1 (≤-step (LoopTermination.rd P r1)) p1 (LoopTermination.rd P r1)) where | |
565 -- TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → {!!} → reduce r1 < reduce r → t | |
566 -- TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 {!!} (λ r2 P1 → ⊥-elim (lemma5 n≤j (LoopTermination.rd P1 r2))) | |
567 -- TerminatingLoop1 (suc j) r r1 n≤j p1 lt with <-cmp (reduce r1) (suc j) | |
568 -- ... | tri< a ¬b ¬c = TerminatingLoop1 j r r1 a p1 lt | |
569 -- ... | tri≈ ¬a b ¬c = loop r1 {!!} (λ r2 p2 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b {!!} ) p2 {!!} ) | |
570 -- ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) | |
571 -- | |
572 -- record ReplCond {n : Level} {A : Set n} (C : ℕ → bt A → List (bt A) → Set n) : Set (Level.suc n) where | |
573 -- field | |
574 -- c1 : (key : ℕ) → (repl : bt A) → (stack : List (bt A)) → C key repl stack | |
575 -- | |
576 -- replaceP0 : {n m : Level} {A : Set n} {t : Set m} | |
577 -- → (key : ℕ) → (value : A) → ( repl : bt A) | |
578 -- → (stack : List (bt A)) | |
579 -- → {C : ℕ → (repl : bt A ) → List (bt A) → Set n} → C key repl stack → ReplCond C | |
580 -- → (next : ℕ → A → (repl : bt A) → (stack1 : List (bt A)) | |
581 -- → C key repl stack → length stack1 < length stack → t) | |
582 -- → (exit : (repl : bt A) → C key repl [] → t) → t | |
583 -- replaceP0 key value repl [] Pre _ next exit = exit repl {!!} | |
584 -- replaceP0 key value repl (leaf ∷ []) Pre _ next exit = exit (node key value leaf leaf) {!!} | |
585 -- replaceP0 key value repl (node key₁ value₁ left right ∷ []) Pre e next exit with <-cmp key key₁ | |
586 -- ... | tri< a ¬b ¬c = exit (node key₁ value₁ repl right ) {!!} | |
587 -- ... | tri≈ ¬a b ¬c = exit repl {!!} | |
588 -- ... | tri> ¬a ¬b c = exit (node key₁ value₁ left repl ) {!!} | |
589 -- replaceP0 {n} {_} {A} key value repl (leaf ∷ st@(tree1 ∷ st1)) Pre e next exit = next key value repl st {!!} ≤-refl | |
590 -- replaceP0 {n} {_} {A} key value repl (node key₁ value₁ left right ∷ st@(tree1 ∷ st1)) Pre e next exit with <-cmp key key₁ | |
591 -- ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st {!!} ≤-refl | |
592 -- ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st {!!} ≤-refl | |
593 -- ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st {!!} ≤-refl | |
594 -- | |
595 -- | |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
596 open _∧_ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
597 |
615 | 598 RTtoTI0 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
599 → replacedTree key value tree repl → treeInvariant repl | |
692 | 600 RTtoTI0 .leaf .(node key value leaf leaf) key value ti r-leaf = t-single key value |
601 RTtoTI0 .(node key _ leaf leaf) .(node key value leaf leaf) key value (t-single .key _) r-node = t-single key value | |
602 RTtoTI0 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right x ti) r-node = t-right x ti | |
603 RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti | |
604 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 | 605 -- r-right case |
692 | 606 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) |
607 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 | 608 RTtoTI0 (node key₁ _ leaf right@(node key₂ _ _ _)) (node key₁ value₁ leaf right₁@(node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = |
609 t-right (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) | |
692 | 610 RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left x₁ ti) (r-right x ()) |
611 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) = | |
612 t-node x₁ x ti (t-single key value) | |
693 | 613 RTtoTI0 (node key₁ _ (node _ _ _ _) (node key₂ _ _ _)) (node key₁ _ (node _ _ _ _) (node key₃ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = |
614 t-node x₁ (subst (λ k → key₁ < k) (rt-property-key ri) x₂) ti (RTtoTI0 _ _ key value ti₁ ri) | |
701 | 615 -- r-left case |
700 | 616 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 | 617 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 |
618 RTtoTI0 (node key₃ _ (node key₂ _ _ _) leaf) (node key₃ _ (node key₁ value₁ left left₁) leaf) key value (t-left x₁ ti) (r-left x ri) = | |
619 t-left (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) -- key₁ < key₃ | |
620 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 | 621 |
622 RTtoTI1 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant repl | |
623 → replacedTree key value tree repl → treeInvariant tree | |
701 | 624 RTtoTI1 .leaf .(node key value leaf leaf) key value ti r-leaf = t-leaf |
625 RTtoTI1 (node key value₁ leaf leaf) .(node key value leaf leaf) key value (t-single .key .value) r-node = t-single key value₁ | |
626 RTtoTI1 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right x ti) r-node = t-right x ti | |
627 RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti | |
628 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₁ | |
629 -- r-right case | |
630 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₁ | |
631 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) = | |
632 t-right (subst (λ k → key₁ < k ) (sym (rt-property-key ri)) x₁) (RTtoTI1 _ _ key value ti ri) -- key₁ < key₂ | |
633 RTtoTI1 (node _ _ (node _ _ _ _) leaf) (node _ _ (node _ _ _ _) (node key value _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x r-leaf) = | |
634 t-left x₁ ti | |
635 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₁ | |
636 -- r-left case | |
637 RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ _ leaf) key value (t-left x₁ ti) (r-left x ri) = t-single key₁ value₁ | |
638 RTtoTI1 (node key₁ _ (node key₂ value₁ n n₁) leaf) (node key₁ _ (node key₃ _ _ _) leaf) key value (t-left x₁ ti) (r-left x ri) = | |
639 t-left (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) (RTtoTI1 _ _ key value ti ri) -- key₂ < key₁ | |
640 RTtoTI1 (node key₁ value₁ leaf _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x r-leaf) = t-right x₂ ti₁ | |
641 RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = | |
642 t-node (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ (RTtoTI1 _ _ key value ti ri) ti₁ -- key₂ < key₁ | |
614 | 643 |
611 | 644 insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
696 | 645 → (exit : (tree repl : bt A) → treeInvariant repl ∧ replacedTree key value tree repl → t ) → t |
693 | 646 insertTreeP {n} {m} {A} {t} tree key value P0 exit = |
647 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-single ⟫ | |
696 | 648 $ λ p P loop → findP key (proj1 p) tree (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
693 | 649 $ λ t s P C → replaceNodeP key value t C (proj1 P) |
650 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ bt A ∧ bt A ) | |
651 {λ p → replacePR key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) (λ _ _ _ → Lift n ⊤ ) } | |
696 | 652 (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ record { tree0 = tree ; ti = P0 ; si = proj2 P ; ri = R ; ci = lift tt } |
693 | 653 $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) P1 |
696 | 654 (λ key value {tree1} repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ P2 lt ) |
655 $ λ tree repl P → exit tree repl ⟪ RTtoTI0 _ _ _ _ (proj1 P) (proj2 P) , proj2 P ⟫ | |
614 | 656 |
696 | 657 insertTestP1 = insertTreeP leaf 1 1 t-leaf |
658 $ λ _ x P → insertTreeP x 2 1 (proj1 P) | |
659 $ λ _ x P → insertTreeP x 3 2 (proj1 P) | |
660 $ λ _ x P → insertTreeP x 2 2 (proj1 P) (λ _ x _ → x ) | |
694 | 661 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
662 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
|
663 top-value leaf = nothing |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
664 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
|
665 |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
666 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 | 667 field |
619 | 668 tree0 : bt A |
696 | 669 ti0 : treeInvariant tree0 |
695 | 670 ti : treeInvariant tree |
662 | 671 si : stackInvariant key tree tree0 stack |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
672 ci : C key tree stack -- data continuation |
702 | 673 |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
674 record findExt {n : Level} {A : Set n} (key : ℕ) (C : ℕ → bt A → List (bt A) → Set n) : Set (Level.suc n) where |
702 | 675 field |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
676 c1 : {key₁ : ℕ} {tree tree₁ : bt A } {st : List (bt A)} {v1 : A} |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
677 → findPR key (node key₁ v1 tree tree₁) st C → key < key₁ → C key tree (tree ∷ st) |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
678 c2 : {key₁ : ℕ} {tree tree₁ : bt A } {st : List (bt A)} {v1 : A} |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
679 → findPR key (node key₁ v1 tree tree₁) st C → key > key₁ → C key tree₁ (tree₁ ∷ st) |
618 | 680 |
695 | 681 findPP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
682 → {C : ℕ → bt A → List (bt A) → Set n } → findPR key tree stack C → findExt key C |
702 | 683 → (next : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack C → bt-depth tree1 < bt-depth tree → t ) |
684 → (exit : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack C | |
695 | 685 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t |
702 | 686 findPP key leaf st Pre _ _ exit = exit leaf st Pre (case1 refl) |
687 findPP key (node key₁ v1 tree tree₁) st Pre _ next exit with <-cmp key key₁ | |
688 findPP key n st Pre _ _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl) | |
689 findPP {n} {_} {A} key (node key₁ v1 tree tree₁) st Pre e next _ | tri< a ¬b ¬c = next tree (tree ∷ st) | |
705 | 690 record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 Pre ; ti = treeLeftDown tree tree₁ (findPR.ti Pre) ; si = findP1 a st (findPR.si Pre) ; ci = findExt.c1 e Pre a } depth-1< where |
695 | 691 findP1 : key < key₁ → (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) |
692 (findPR.tree0 Pre) st → stackInvariant key tree (findPR.tree0 Pre) (tree ∷ st) | |
693 findP1 a (x ∷ st) si = s-left a si | |
702 | 694 findPP key n@(node key₁ v1 tree tree₁) st Pre e next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st) |
695 record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 Pre ; ti = treeRightDown tree tree₁ (findPR.ti Pre) ; si = s-right c (findPR.si Pre) ; ci = findExt.c2 e Pre c } depth-2< | |
616 | 696 |
618 | 697 insertTreePP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
698 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t | |
695 | 699 insertTreePP {n} {m} {A} {t} tree key value P0 exit = |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
700 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ _ _ _ → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , tree ∷ [] ⟫ |
696 | 701 record { tree0 = tree ; ti = P0 ; ti0 = P0 ;si = s-single ; ci = lift tt } |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
702 $ λ p P loop → findPP key (proj1 p) (proj2 p) P record { c1 = λ _ _ → lift tt ; c2 = λ _ _ → lift tt } (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
695 | 703 $ λ t s P C → replaceNodeP key value t C (findPR.ti P) |
704 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ bt A ∧ bt A ) | |
705 {λ p → replacePR key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) (λ _ _ _ → Lift n ⊤ ) } | |
696 | 706 (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ record { tree0 = findPR.tree0 P ; ti = findPR.ti0 P ; si = findPR.si P ; ri = R ; ci = lift tt } |
695 | 707 $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) P1 |
708 (λ key value {tree1} repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ P2 lt ) exit | |
618 | 709 |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
710 record findPC {n : Level} {A : Set n} (value : A) (key1 : ℕ) (tree : bt A ) (stack : List (bt A)) : Set n where |
616 | 711 field |
712 tree1 : bt A | |
698 | 713 ci : replacedTree key1 value tree1 tree |
616 | 714 |
702 | 715 findPPC1 : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A ) → (stack : List (bt A)) |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
716 → findPR key tree stack (findPC value ) |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
717 → (next : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC value ) → bt-depth tree1 < bt-depth tree → t ) |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
718 → (exit : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC value ) |
702 | 719 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t |
720 findPPC1 {n} {_} {A} key value tree stack Pr next exit = findPP key tree stack Pr findext next exit where | |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
721 findext01 : {key₂ : ℕ} {tree₁ : bt A} {tree₂ : bt A} {st : List (bt A)} {v1 : A} |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
722 → (Pre : findPR key (node key₂ v1 tree₁ tree₂) st (findPC value) ) |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
723 → key < key₂ → findPC value key tree₁ (tree₁ ∷ st) |
933 | 724 findext01 Pre a with findPC.ci (findPR.ci Pre) | findPC.tree1 (findPR.ci Pre) in eqp |
725 ... | r-leaf | leaf = ⊥-elim ( nat-≤> a ≤-refl) | |
726 ... | r-node | node key value t1 t3 = ⊥-elim ( nat-≤> a ≤-refl ) | |
727 ... | r-right x t | t1 = ⊥-elim (nat-<> x a) | |
728 ... | r-left {key₂} {v1} {_} {t2} {tree₂} x ri | node key value t1 t3 = record { tree1 = t1 ; ci = subst (λ k → replacedTree _ _ k _) (lemma01 eqp) ri } where | |
729 lemma01 : {t2 : bt A} → node key₂ v1 t2 tree₂ ≡ node key value t1 t3 → t2 ≡ t1 | |
730 lemma01 refl = refl | |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
731 findext02 : {key₂ : ℕ} {tree₁ : bt A} {tree₂ : bt A} {st : List (bt A)} {v1 : A} |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
732 → (Pre : findPR key (node key₂ v1 tree₁ tree₂) st (findPC value) ) |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
733 → key > key₂ → findPC value key tree₂ (tree₂ ∷ st) |
933 | 734 findext02 Pre c with findPC.ci (findPR.ci Pre) | findPC.tree1 (findPR.ci Pre) in eqp |
735 ... | r-leaf | leaf = ⊥-elim ( nat-≤> c ≤-refl) | |
736 ... | r-node | node key value t1 t3 = ⊥-elim ( nat-≤> c ≤-refl ) | |
737 ... | r-left x t | t1 = ⊥-elim (nat-<> x c) | |
738 ... | r-right {key₂} {v1} {_} {tree₂} x ri | node key value t1 t3 = record { tree1 = t3 ; ci = subst (λ k → replacedTree _ _ k _) (lemma01 eqp) ri } where | |
739 lemma01 : {t2 : bt A} → node key₂ v1 tree₂ t2 ≡ node key value t1 t3 → t2 ≡ t3 | |
740 lemma01 refl = refl | |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
741 findext : findExt key (findPC value ) |
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
742 findext = record { c1 = findext01 ; c2 = findext02 } |
702 | 743 |
744 insertTreeSpec0 : {n : Level} {A : Set n} → (tree : bt A) → (value : A) → top-value tree ≡ just value → ⊤ | |
745 insertTreeSpec0 _ _ _ = tt | |
746 | |
700 | 747 containsTree : {n : Level} {A : Set n} → (tree tree1 : bt A) → (key : ℕ) → (value : A) → treeInvariant tree1 → replacedTree key value tree1 tree → ⊤ |
748 containsTree {n} {A} tree tree1 key value P RT = | |
617 | 749 TerminatingLoopS (bt A ∧ List (bt A) ) |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
750 {λ p → findPR key (proj1 p) (proj2 p) (findPC value ) } (λ p → bt-depth (proj1 p)) |
698 | 751 ⟪ tree , tree ∷ [] ⟫ record { tree0 = tree ; ti0 = RTtoTI0 _ _ _ _ P RT ; ti = RTtoTI0 _ _ _ _ P RT ; si = s-single |
752 ; ci = record { tree1 = tree1 ; ci = RT } } | |
702 | 753 $ λ p P loop → findPPC1 key value (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
695 | 754 $ λ t1 s1 P2 found? → insertTreeSpec0 t1 value (lemma6 t1 s1 found? P2) where |
703
23e0b9df7896
embedding invariant extentiion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
702
diff
changeset
|
755 lemma6 : (t1 : bt A) (s1 : List (bt A)) (found? : (t1 ≡ leaf) ∨ (node-key t1 ≡ just key)) (P2 : findPR key t1 s1 (findPC value )) → top-value t1 ≡ just value |
698 | 756 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 |
700 | 757 lemma8 : {tree1 t1 : bt A } → replacedTree key value tree1 t1 → node-key t1 ≡ just key → top-value t1 ≡ just value |
758 lemma8 {.leaf} {node key value .leaf .leaf} r-leaf refl = refl | |
759 lemma8 {.(node key _ t1 t2)} {node key value t1 t2} r-node refl = refl | |
760 lemma8 {.(node key value t1 _)} {node key value t1 t2} (r-right x ri) refl = ⊥-elim (¬x<x x) | |
761 lemma8 {.(node key value _ t2)} {node key value t1 t2} (r-left x ri) refl = ⊥-elim (¬x<x x) | |
698 | 762 lemma7 : (t1 : bt A) ( s1 : List (bt A) ) (tree0 tree1 : bt A) |
763 → replacedTree key value tree1 t1 → stackInvariant key t1 tree0 s1 | |
764 → ( t1 ≡ leaf ) ∨ ( node-key t1 ≡ just key) → top-value t1 ≡ just value | |
765 lemma7 .leaf (.leaf ∷ []) .leaf tree1 () s-single (case1 refl) | |
700 | 766 lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ []) .(node key value t1 t2) tree1 ri s-single (case2 x) = lemma8 ri x |
767 lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ x₁ ∷ s1) tree0 tree1 ri (s-right x si) (case2 x₂) = lemma8 ri x₂ | |
768 lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ x₁ ∷ s1) tree0 tree1 ri (s-left x si) (case2 x₂) = lemma8 ri x₂ | |
769 | |
615 | 770 |
700 | 771 containsTree1 : {n : Level} {A : Set n} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree → ⊤ |
772 containsTree1 {n} {A} tree key value ti = | |
773 insertTreeP tree key value ti | |
702 | 774 $ λ tree0 tree1 P → containsTree tree1 tree0 key value (RTtoTI1 _ _ _ _ (proj1 P) (proj2 P) ) (proj2 P) |
700 | 775 |
776 |