Mercurial > hg > Members > Moririn
annotate hoareBinaryTree1.agda @ 740:9ff79715588e
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 23 Apr 2023 14:29:09 +0900 |
parents | 3443703a68cc |
children | c44edea35126 |
rev | line source |
---|---|
722 | 1 module hoareBinaryTree1 where |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
2 |
727 | 3 open import Level hiding (suc ; zero ; _⊔_ ) |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
4 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
5 open import Data.Nat hiding (compare) |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
6 open import Data.Nat.Properties as NatProp |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
7 open import Data.Maybe |
588 | 8 -- open import Data.Maybe.Properties |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
9 open import Data.Empty |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
10 open import Data.List |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
11 open import Data.Product |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
12 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
13 open import Function as F hiding (const) |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
14 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
15 open import Relation.Binary |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
16 open import Relation.Binary.PropositionalEquality |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
17 open import Relation.Nullary |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
18 open import logic |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
19 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
20 |
590 | 21 -- |
22 -- | |
23 -- no children , having left node , having right node , having both | |
24 -- | |
597 | 25 data bt {n : Level} (A : Set n) : Set n where |
604 | 26 leaf : bt A |
27 node : (key : ℕ) → (value : A) → | |
610 | 28 (left : bt A ) → (right : bt A ) → bt A |
600 | 29 |
620 | 30 node-key : {n : Level} {A : Set n} → bt A → Maybe ℕ |
31 node-key (node key _ _ _) = just key | |
32 node-key _ = nothing | |
33 | |
34 node-value : {n : Level} {A : Set n} → bt A → Maybe A | |
35 node-value (node _ value _ _) = just value | |
36 node-value _ = nothing | |
37 | |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
38 bt-depth : {n : Level} {A : Set n} → (tree : bt A ) → ℕ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
39 bt-depth leaf = 0 |
727 | 40 bt-depth (node key value t t₁) = suc (bt-depth t ⊔ bt-depth t₁ ) |
606 | 41 |
605 | 42 open import Data.Unit hiding ( _≟_ ; _≤?_ ; _≤_) |
43 | |
620 | 44 data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where |
45 t-leaf : treeInvariant leaf | |
632 | 46 t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) |
47 t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key₁ value₁ t₁ t₂) | |
48 → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) | |
692 | 49 t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key value t₁ t₂) |
632 | 50 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) |
620 | 51 t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → (key < key₁) → (key₁ < key₂) |
52 → treeInvariant (node key value t₁ t₂) | |
53 → treeInvariant (node key₂ value₂ t₃ t₄) | |
54 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) | |
605 | 55 |
662 | 56 -- |
57 -- stack always contains original top at end | |
58 -- | |
59 data stackInvariant {n : Level} {A : Set n} (key : ℕ) : (top orig : bt A) → (stack : List (bt A)) → Set n where | |
729 | 60 s-nil : {tree0 : bt A} → stackInvariant key tree0 tree0 (tree0 ∷ []) |
653 | 61 s-right : {tree tree0 tree₁ : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 62 → key₁ < key → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
653 | 63 s-left : {tree₁ tree0 tree : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} |
662 | 64 → key < key₁ → stackInvariant key (node key₁ v1 tree₁ tree) tree0 st → stackInvariant key tree₁ tree0 (tree₁ ∷ st) |
639 | 65 |
677 | 66 data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt A ) → Set n where |
639 | 67 r-leaf : replacedTree key value leaf (node key value leaf leaf) |
68 r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) | |
69 r-right : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} | |
677 | 70 → k < key → replacedTree key value t2 t → replacedTree key value (node k v1 t1 t2) (node k v1 t1 t) |
639 | 71 r-left : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} |
687 | 72 → key < k → replacedTree key value t1 t → replacedTree key value (node k v1 t1 t2) (node k v1 t t2) |
652 | 73 |
632 | 74 add< : { i : ℕ } (j : ℕ ) → i < suc i + j |
75 add< {i} j = begin | |
76 suc i ≤⟨ m≤m+n (suc i) j ⟩ | |
77 suc i + j ∎ where open ≤-Reasoning | |
78 | |
79 treeTest1 : bt ℕ | |
692 | 80 treeTest1 = node 0 0 leaf (node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf)) |
632 | 81 treeTest2 : bt ℕ |
692 | 82 treeTest2 = node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf) |
632 | 83 |
84 treeInvariantTest1 : treeInvariant treeTest1 | |
692 | 85 treeInvariantTest1 = t-right (m≤m+n _ 2) (t-node (add< 0) (add< 1) (t-left (add< 0) (t-single 1 7)) (t-single 5 5) ) |
605 | 86 |
639 | 87 stack-top : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
88 stack-top [] = nothing | |
89 stack-top (x ∷ s) = just x | |
606 | 90 |
639 | 91 stack-last : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) |
92 stack-last [] = nothing | |
93 stack-last (x ∷ []) = just x | |
94 stack-last (x ∷ s) = stack-last s | |
632 | 95 |
662 | 96 stackInvariantTest1 : stackInvariant 4 treeTest2 treeTest1 ( treeTest2 ∷ treeTest1 ∷ [] ) |
729 | 97 stackInvariantTest1 = s-right (add< 3) (s-nil ) |
662 | 98 |
666 | 99 si-property0 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 stack → ¬ ( stack ≡ [] ) |
729 | 100 si-property0 (s-nil ) () |
666 | 101 si-property0 (s-right x si) () |
102 si-property0 (s-left x si) () | |
665 | 103 |
666 | 104 si-property1 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 tree1 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 (tree1 ∷ stack) |
105 → tree1 ≡ tree | |
729 | 106 si-property1 (s-nil ) = refl |
666 | 107 si-property1 (s-right _ si) = refl |
108 si-property1 (s-left _ si) = refl | |
662 | 109 |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
110 si-property-last : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack |
662 | 111 → stack-last stack ≡ just tree0 |
729 | 112 si-property-last key t t0 (t ∷ []) (s-nil ) = refl |
666 | 113 si-property-last key t t0 (.t ∷ x ∷ st) (s-right _ si ) with si-property1 si |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
114 ... | refl = si-property-last key x t0 (x ∷ st) si |
666 | 115 si-property-last key t t0 (.t ∷ x ∷ st) (s-left _ si ) with si-property1 si |
663
cf5095488bbd
stack contains original tree at end always
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
116 ... | refl = si-property-last key x t0 (x ∷ st) si |
656 | 117 |
642 | 118 ti-right : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 tree₁ repl) → treeInvariant repl |
119 ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf | |
120 ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-right x ti) = ti | |
121 ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-left x ti) = t-leaf | |
122 ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti₁ | |
123 | |
124 ti-left : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 repl tree₁ ) → treeInvariant repl | |
125 ti-left {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf | |
126 ti-left {_} {_} {_} {_} {key₁} {v1} (t-right x ti) = t-leaf | |
127 ti-left {_} {_} {_} {_} {key₁} {v1} (t-left x ti) = ti | |
128 ti-left {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti | |
129 | |
662 | 130 stackTreeInvariant : {n : Level} {A : Set n} (key : ℕ) (sub tree : bt A) → (stack : List (bt A)) |
131 → treeInvariant tree → stackInvariant key sub tree stack → treeInvariant sub | |
729 | 132 stackTreeInvariant {_} {A} key sub tree (sub ∷ []) ti (s-nil ) = ti |
662 | 133 stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-right _ si ) = ti-right (si1 si) where |
134 si1 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 tree₁ sub ) tree st → treeInvariant (node key₁ v1 tree₁ sub ) | |
135 si1 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 tree₁ sub ) tree st ti si | |
136 stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-left _ si ) = ti-left ( si2 si) where | |
137 si2 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 sub tree₁ ) tree st → treeInvariant (node key₁ v1 sub tree₁ ) | |
138 si2 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 sub tree₁ ) tree st ti si | |
139 | |
639 | 140 rt-property1 : {n : Level} {A : Set n} (key : ℕ) (value : A) (tree tree1 : bt A ) → replacedTree key value tree tree1 → ¬ ( tree1 ≡ leaf ) |
141 rt-property1 {n} {A} key value .leaf .(node key value leaf leaf) r-leaf () | |
142 rt-property1 {n} {A} key value .(node key _ _ _) .(node key value _ _) r-node () | |
677 | 143 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-right x rt) = λ () |
144 rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-left x rt) = λ () | |
639 | 145 |
690 | 146 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 |
147 rt-property-leaf r-leaf = refl | |
148 | |
698 | 149 rt-property-¬leaf : {n : Level} {A : Set n} {key : ℕ} {value : A} {tree : bt A} → ¬ replacedTree key value tree leaf |
150 rt-property-¬leaf () | |
151 | |
692 | 152 rt-property-key : {n : Level} {A : Set n} {key key₂ key₃ : ℕ} {value value₂ value₃ : A} {left left₁ right₂ right₃ : bt A} |
153 → replacedTree key value (node key₂ value₂ left right₂) (node key₃ value₃ left₁ right₃) → key₂ ≡ key₃ | |
154 rt-property-key r-node = refl | |
155 rt-property-key (r-right x ri) = refl | |
156 rt-property-key (r-left x ri) = refl | |
157 | |
698 | 158 nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ |
159 nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x | |
160 nat-<> : { x y : ℕ } → x < y → y < x → ⊥ | |
161 nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x | |
162 | |
163 open _∧_ | |
164 | |
165 | |
632 | 166 depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) |
167 depth-1< {i} {j} = s≤s (m≤m⊔n _ j) | |
168 | |
169 depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) | |
650 | 170 depth-2< {i} {j} = s≤s (m≤n⊔m j i) |
611 | 171 |
649 | 172 depth-3< : {i : ℕ } → suc i ≤ suc (suc i) |
173 depth-3< {zero} = s≤s ( z≤n ) | |
174 depth-3< {suc i} = s≤s (depth-3< {i} ) | |
175 | |
176 | |
634 | 177 treeLeftDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) |
178 → treeInvariant (node k v1 tree tree₁) | |
179 → treeInvariant tree | |
180 treeLeftDown {n} {A} {_} {v1} leaf leaf (t-single k1 v1) = t-leaf | |
181 treeLeftDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = t-leaf | |
182 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = ti | |
183 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti | |
184 | |
185 treeRightDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) | |
186 → treeInvariant (node k v1 tree tree₁) | |
187 → treeInvariant tree₁ | |
188 treeRightDown {n} {A} {_} {v1} .leaf .leaf (t-single _ .v1) = t-leaf | |
189 treeRightDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = ti | |
190 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = t-leaf | |
191 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti₁ | |
192 | |
615 | 193 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A)) |
662 | 194 → treeInvariant tree ∧ stackInvariant key tree tree0 stack |
693 | 195 → (next : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) |
196 → (exit : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack | |
638 | 197 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t |
693 | 198 findP key leaf tree0 st Pre _ exit = exit leaf st Pre (case1 refl) |
632 | 199 findP key (node key₁ v1 tree tree₁) tree0 st Pre next exit with <-cmp key key₁ |
693 | 200 findP key n tree0 st Pre _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl) |
201 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
|
202 ⟪ 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
|
203 findP1 : key < key₁ → (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (tree ∷ st) |
664 | 204 findP1 a (x ∷ st) si = s-left a si |
693 | 205 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 | 206 |
638 | 207 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₁) |
208 replaceTree1 k v1 value (t-single .k .v1) = t-single k value | |
209 replaceTree1 k v1 value (t-right x t) = t-right x t | |
210 replaceTree1 k v1 value (t-left x t) = t-left x t | |
211 replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ | |
212 | |
649 | 213 open import Relation.Binary.Definitions |
214 | |
215 lemma3 : {i j : ℕ} → 0 ≡ i → j < i → ⊥ | |
216 lemma3 refl () | |
217 lemma5 : {i j : ℕ} → i < 1 → j < i → ⊥ | |
218 lemma5 (s≤s z≤n) () | |
700 | 219 ¬x<x : {x : ℕ} → ¬ (x < x) |
220 ¬x<x (s≤s lt) = ¬x<x lt | |
649 | 221 |
687 | 222 child-replaced : {n : Level} {A : Set n} (key : ℕ) (tree : bt A) → bt A |
223 child-replaced key leaf = leaf | |
224 child-replaced key (node key₁ value left right) with <-cmp key key₁ | |
225 ... | tri< a ¬b ¬c = left | |
226 ... | tri≈ ¬a b ¬c = node key₁ value left right | |
227 ... | tri> ¬a ¬b c = right | |
677 | 228 |
671
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
229 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
|
230 field |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
231 tree0 : bt A |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
232 ti : treeInvariant tree0 |
b5fde9727830
use record invariant for replace
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
233 si : stackInvariant key tree tree0 stack |
687 | 234 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
|
235 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
|
236 |
638 | 237 replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) |
238 → (tree ≡ leaf ) ∨ ( node-key tree ≡ just key ) | |
694 | 239 → (treeInvariant tree ) → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value (child-replaced key tree) tree1 → t) → t |
240 replaceNodeP k v1 leaf C P next = next (node k v1 leaf leaf) (t-single k v1 ) r-leaf | |
241 replaceNodeP k v1 (node .k value t t₁) (case2 refl) P next = next (node k v1 t t₁) (replaceTree1 k value v1 P) | |
695 | 242 (subst (λ j → replacedTree k v1 j (node k v1 t t₁) ) repl00 r-node) where |
694 | 243 repl00 : node k value t t₁ ≡ child-replaced k (node k value t t₁) |
244 repl00 with <-cmp k k | |
245 ... | tri< a ¬b ¬c = ⊥-elim (¬b refl) | |
246 ... | tri≈ ¬a b ¬c = refl | |
247 ... | tri> ¬a ¬b c = ⊥-elim (¬b refl) | |
606 | 248 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
249 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
|
250 → (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
|
251 → (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
|
252 → (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
|
253 → replacePR key value tree1 repl stack1 (λ _ _ _ → Lift n ⊤) → length stack1 < length stack → t) |
613 | 254 → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t |
675 | 255 replaceP key value {tree} repl [] Pre next exit = ⊥-elim ( si-property0 (replacePR.si Pre) refl ) -- can't happen |
256 replaceP key value {tree} repl (leaf ∷ []) Pre next exit with si-property-last _ _ _ _ (replacePR.si Pre)-- tree0 ≡ leaf | |
677 | 257 ... | refl = exit (replacePR.tree0 Pre) (node key value leaf leaf) ⟪ replacePR.ti Pre , r-leaf ⟫ |
689 | 258 replaceP key value {tree} repl (node key₁ value₁ left right ∷ []) Pre next exit with <-cmp key key₁ |
259 ... | tri< a ¬b ¬c = exit (replacePR.tree0 Pre) (node key₁ value₁ repl right ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
260 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ repl right ) | |
261 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
262 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ k right ) (node key₁ value₁ repl right )) repl02 (r-left a repl03) where | |
263 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
264 repl03 = replacePR.ri Pre | |
265 repl02 : child-replaced key (node key₁ value₁ left right) ≡ left | |
266 repl02 with <-cmp key key₁ | |
267 ... | tri< a ¬b ¬c = refl | |
268 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a a) | |
269 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a a) | |
270 ... | tri≈ ¬a b ¬c = exit (replacePR.tree0 Pre) repl ⟪ replacePR.ti Pre , repl01 ⟫ where | |
678 | 271 repl01 : replacedTree key value (replacePR.tree0 Pre) repl |
272 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
689 | 273 repl01 | refl | refl = subst (λ k → replacedTree key value k repl) repl02 (replacePR.ri Pre) where |
274 repl02 : child-replaced key (node key₁ value₁ left right) ≡ node key₁ value₁ left right | |
275 repl02 with <-cmp key key₁ | |
276 ... | tri< a ¬b ¬c = ⊥-elim ( ¬b b) | |
277 ... | tri≈ ¬a b ¬c = refl | |
278 ... | tri> ¬a ¬b c = ⊥-elim ( ¬b b) | |
279 ... | tri> ¬a ¬b c = exit (replacePR.tree0 Pre) (node key₁ value₁ left repl ) ⟪ replacePR.ti Pre , repl01 ⟫ where | |
280 repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ left repl ) | |
281 repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) | |
282 repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ left k ) (node key₁ value₁ left repl )) repl02 (r-right c repl03) where | |
283 repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl | |
284 repl03 = replacePR.ri Pre | |
285 repl02 : child-replaced key (node key₁ value₁ left right) ≡ right | |
286 repl02 with <-cmp key key₁ | |
287 ... | tri< a ¬b ¬c = ⊥-elim ( ¬c c) | |
288 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬c c) | |
289 ... | tri> ¬a ¬b c = refl | |
690 | 290 replaceP {n} {_} {A} key value {tree} repl (leaf ∷ st@(tree1 ∷ st1)) Pre next exit = next key value repl st Post ≤-refl where |
291 Post : replacePR key value tree1 repl (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) | |
292 Post with replacePR.si Pre | |
293 ... | s-right {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
294 repl09 : tree1 ≡ node key₂ v1 tree₁ leaf | |
295 repl09 = si-property1 si | |
296 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
297 repl10 with si-property1 si | |
298 ... | refl = si | |
299 repl07 : child-replaced key (node key₂ v1 tree₁ leaf) ≡ leaf | |
300 repl07 with <-cmp key key₂ | |
301 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
302 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
303 ... | tri> ¬a ¬b c = refl | |
304 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
305 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 | |
306 ... | s-left {_} {_} {tree₁} {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
307 repl09 : tree1 ≡ node key₂ v1 leaf tree₁ | |
308 repl09 = si-property1 si | |
309 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
310 repl10 with si-property1 si | |
311 ... | refl = si | |
312 repl07 : child-replaced key (node key₂ v1 leaf tree₁ ) ≡ leaf | |
313 repl07 with <-cmp key key₂ | |
314 ... | tri< a ¬b ¬c = refl | |
315 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
316 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
317 repl12 : replacedTree key value (child-replaced key tree1 ) repl | |
318 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 | 319 replaceP {n} {_} {A} key value {tree} repl (node key₁ value₁ left right ∷ st@(tree1 ∷ st1)) Pre next exit with <-cmp key key₁ |
320 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st Post ≤-refl where | |
675 | 321 Post : replacePR key value tree1 (node key₁ value₁ repl right ) st (λ _ _ _ → Lift n ⊤) |
687 | 322 Post with replacePR.si Pre |
688 | 323 ... | s-right {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where |
324 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) | |
325 repl09 = si-property1 si | |
326 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
327 repl10 with si-property1 si | |
328 ... | refl = si | |
329 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left | |
330 repl03 with <-cmp key key₁ | |
331 ... | tri< a1 ¬b ¬c = refl | |
332 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
333 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
334 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
335 repl02 with repl09 | <-cmp key key₂ | |
336 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
689 | 337 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) |
688 | 338 ... | refl | tri> ¬a ¬b c = refl |
339 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
340 repl04 = begin | |
341 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
342 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
343 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
344 child-replaced key tree1 ∎ where open ≡-Reasoning | |
345 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
346 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
687 | 347 ... | 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 | 348 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ |
683 | 349 repl09 = si-property1 si |
350 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
351 repl10 with si-property1 si | |
352 ... | refl = si | |
687 | 353 repl03 : child-replaced key (node key₁ value₁ left right) ≡ left |
354 repl03 with <-cmp key key₁ | |
355 ... | tri< a1 ¬b ¬c = refl | |
356 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) | |
357 ... | tri> ¬a ¬b c = ⊥-elim (¬a a) | |
358 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
359 repl02 with repl09 | <-cmp key key₂ | |
360 ... | refl | tri< a ¬b ¬c = refl | |
361 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
362 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
363 repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 | |
364 repl04 = begin | |
365 node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ | |
366 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
367 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
368 child-replaced key tree1 ∎ where open ≡-Reasoning | |
369 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) | |
370 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) | |
705 | 371 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st Post ≤-refl where |
690 | 372 Post : replacePR key value tree1 (node key₁ value left right ) (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) |
373 Post with replacePR.si Pre | |
691 | 374 ... | 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 | 375 repl09 : tree1 ≡ node key₂ v1 tree₁ tree -- (node key₁ value₁ left right) |
376 repl09 = si-property1 si | |
377 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
378 repl10 with si-property1 si | |
379 ... | refl = si | |
380 repl07 : child-replaced key (node key₂ v1 tree₁ tree) ≡ tree | |
381 repl07 with <-cmp key key₂ | |
382 ... | tri< a ¬b ¬c = ⊥-elim (¬c x) | |
383 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) | |
384 ... | tri> ¬a ¬b c = refl | |
691 | 385 repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) |
386 repl12 refl with repl09 | |
387 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node | |
388 ... | 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 | 389 repl09 : tree1 ≡ node key₂ v1 tree tree₁ |
390 repl09 = si-property1 si | |
391 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
392 repl10 with si-property1 si | |
393 ... | refl = si | |
394 repl07 : child-replaced key (node key₂ v1 tree tree₁ ) ≡ tree | |
395 repl07 with <-cmp key key₂ | |
396 ... | tri< a ¬b ¬c = refl | |
397 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) | |
398 ... | tri> ¬a ¬b c = ⊥-elim (¬a x) | |
691 | 399 repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) |
400 repl12 refl with repl09 | |
401 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node | |
690 | 402 ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st Post ≤-refl where |
403 Post : replacePR key value tree1 (node key₁ value₁ left repl ) st (λ _ _ _ → Lift n ⊤) | |
404 Post with replacePR.si Pre | |
405 ... | s-right {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
406 repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) | |
407 repl09 = si-property1 si | |
408 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
409 repl10 with si-property1 si | |
410 ... | refl = si | |
411 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
412 repl03 with <-cmp key key₁ | |
413 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
414 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
415 ... | tri> ¬a ¬b c = refl | |
416 repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right | |
417 repl02 with repl09 | <-cmp key key₂ | |
418 ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) | |
419 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) | |
420 ... | refl | tri> ¬a ¬b c = refl | |
421 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
422 repl04 = begin | |
423 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
424 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
425 child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
426 child-replaced key tree1 ∎ where open ≡-Reasoning | |
427 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
428 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
429 ... | s-left {_} {_} {tree₁} {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where | |
430 repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ | |
431 repl09 = si-property1 si | |
432 repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) | |
433 repl10 with si-property1 si | |
434 ... | refl = si | |
435 repl03 : child-replaced key (node key₁ value₁ left right) ≡ right | |
436 repl03 with <-cmp key key₁ | |
437 ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) | |
438 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) | |
439 ... | tri> ¬a ¬b c = refl | |
440 repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right | |
441 repl02 with repl09 | <-cmp key key₂ | |
442 ... | refl | tri< a ¬b ¬c = refl | |
443 ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) | |
444 ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) | |
445 repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 | |
446 repl04 = begin | |
447 node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ | |
448 node key₁ value₁ left right ≡⟨ sym repl02 ⟩ | |
449 child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ | |
450 child-replaced key tree1 ∎ where open ≡-Reasoning | |
451 repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) | |
452 repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) | |
644 | 453 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
454 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
|
455 → (r : Index) → (p : Invraiant r) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
456 → (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
|
457 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
|
458 ... | 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
|
459 ... | 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
|
460 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
|
461 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
|
462 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
|
463 ... | 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
|
464 ... | 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
|
465 ... | 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
|
466 |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
467 open _∧_ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
468 |
615 | 469 RTtoTI0 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
470 → replacedTree key value tree repl → treeInvariant repl | |
692 | 471 RTtoTI0 .leaf .(node key value leaf leaf) key value ti r-leaf = t-single key value |
472 RTtoTI0 .(node key _ leaf leaf) .(node key value leaf leaf) key value (t-single .key _) r-node = t-single key value | |
473 RTtoTI0 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right x ti) r-node = t-right x ti | |
474 RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti | |
475 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 | 476 -- r-right case |
692 | 477 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) |
478 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 | 479 RTtoTI0 (node key₁ _ leaf right@(node key₂ _ _ _)) (node key₁ value₁ leaf right₁@(node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = |
480 t-right (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) | |
692 | 481 RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left x₁ ti) (r-right x ()) |
482 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) = | |
483 t-node x₁ x ti (t-single key value) | |
693 | 484 RTtoTI0 (node key₁ _ (node _ _ _ _) (node key₂ _ _ _)) (node key₁ _ (node _ _ _ _) (node key₃ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = |
485 t-node x₁ (subst (λ k → key₁ < k) (rt-property-key ri) x₂) ti (RTtoTI0 _ _ key value ti₁ ri) | |
701 | 486 -- r-left case |
700 | 487 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 | 488 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 |
489 RTtoTI0 (node key₃ _ (node key₂ _ _ _) leaf) (node key₃ _ (node key₁ value₁ left left₁) leaf) key value (t-left x₁ ti) (r-left x ri) = | |
490 t-left (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) -- key₁ < key₃ | |
491 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 | 492 |
493 RTtoTI1 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant repl | |
494 → replacedTree key value tree repl → treeInvariant tree | |
701 | 495 RTtoTI1 .leaf .(node key value leaf leaf) key value ti r-leaf = t-leaf |
496 RTtoTI1 (node key value₁ leaf leaf) .(node key value leaf leaf) key value (t-single .key .value) r-node = t-single key value₁ | |
497 RTtoTI1 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right x ti) r-node = t-right x ti | |
498 RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti | |
499 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₁ | |
500 -- r-right case | |
501 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₁ | |
502 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) = | |
503 t-right (subst (λ k → key₁ < k ) (sym (rt-property-key ri)) x₁) (RTtoTI1 _ _ key value ti ri) -- key₁ < key₂ | |
504 RTtoTI1 (node _ _ (node _ _ _ _) leaf) (node _ _ (node _ _ _ _) (node key value _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x r-leaf) = | |
505 t-left x₁ ti | |
506 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₁ | |
507 -- r-left case | |
508 RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ _ leaf) key value (t-left x₁ ti) (r-left x ri) = t-single key₁ value₁ | |
509 RTtoTI1 (node key₁ _ (node key₂ value₁ n n₁) leaf) (node key₁ _ (node key₃ _ _ _) leaf) key value (t-left x₁ ti) (r-left x ri) = | |
510 t-left (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) (RTtoTI1 _ _ key value ti ri) -- key₂ < key₁ | |
511 RTtoTI1 (node key₁ value₁ leaf _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x r-leaf) = t-right x₂ ti₁ | |
512 RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = | |
513 t-node (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ (RTtoTI1 _ _ key value ti ri) ti₁ -- key₂ < key₁ | |
614 | 514 |
611 | 515 insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
696 | 516 → (exit : (tree repl : bt A) → treeInvariant repl ∧ replacedTree key value tree repl → t ) → t |
693 | 517 insertTreeP {n} {m} {A} {t} tree key value P0 exit = |
729 | 518 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant key (proj1 p) tree (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , tree ∷ [] ⟫ ⟪ P0 , s-nil ⟫ |
696 | 519 $ λ p P loop → findP key (proj1 p) tree (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
693 | 520 $ λ t s P C → replaceNodeP key value t C (proj1 P) |
521 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ bt A ∧ bt A ) | |
522 {λ p → replacePR key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) (λ _ _ _ → Lift n ⊤ ) } | |
696 | 523 (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ record { tree0 = tree ; ti = P0 ; si = proj2 P ; ri = R ; ci = lift tt } |
693 | 524 $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) P1 |
696 | 525 (λ key value {tree1} repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ P2 lt ) |
526 $ λ tree repl P → exit tree repl ⟪ RTtoTI0 _ _ _ _ (proj1 P) (proj2 P) , proj2 P ⟫ | |
614 | 527 |
696 | 528 insertTestP1 = insertTreeP leaf 1 1 t-leaf |
722 | 529 $ λ _ x0 P0 → insertTreeP x0 2 1 (proj1 P0) |
530 $ λ _ x1 P1 → insertTreeP x1 3 2 (proj1 P1) | |
727 | 531 $ λ _ x2 P2 → insertTreeP x2 2 2 (proj1 P2) (λ _ x P → x ) |
694 | 532 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
533 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
|
534 top-value leaf = nothing |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
535 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
|
536 |
722 | 537 -- is realy inserted? |
538 | |
539 -- other element is preserved? | |
702 | 540 |
722 | 541 -- deletion? |
616 | 542 |
722 | 543 data Color : Set where |
544 Red : Color | |
545 Black : Color | |
618 | 546 |
723 | 547 data RBTree {n : Level} (A : Set n) : (key : ℕ) → Color → (b-depth : ℕ) → Set n where |
548 rb-leaf : (key : ℕ) → RBTree A key Black 0 | |
549 rb-single : (key : ℕ) → (value : A) → (c : Color) → RBTree A key c 1 | |
725 | 550 t-right-red : (key : ℕ) {key₁ : ℕ} → (value : A) → key < key₁ → {d : ℕ } → RBTree A key₁ Black d → RBTree A key Red d |
551 t-right-black : (key : ℕ) {key₁ : ℕ} → (value : A) → key < key₁ → {c : Color} → {d : ℕ }→ RBTree A key₁ c d | |
723 | 552 → RBTree A key Black (suc d) |
725 | 553 t-left-red : (key₁ : ℕ) { key : ℕ} → (value : A) → key < key₁ → {d : ℕ} → RBTree A key Black d |
723 | 554 → RBTree A key₁ Red d |
725 | 555 t-left-black : (key₁ : ℕ) {key : ℕ} → (value : A) → key < key₁ → {c : Color} → {d : ℕ} → RBTree A key c d |
723 | 556 → RBTree A key₁ Black (suc d) |
725 | 557 t-node-red : (key₁ : ℕ) { key key₂ : ℕ} → (value : A) → key < key₁ → key₁ < key₂ → {d : ℕ} |
723 | 558 → RBTree A key Black d |
559 → RBTree A key₂ Black d | |
560 → RBTree A key₁ Red d | |
725 | 561 t-node-black : (key₁ : ℕ) {key key₂ : ℕ} → (value : A) → key < key₁ → key₁ < key₂ → {c c1 : Color} {d : ℕ} |
723 | 562 → RBTree A key c d |
563 → RBTree A key₂ c1 d | |
564 → RBTree A key₁ Black (suc d) | |
565 | |
738 | 566 color : {n : Level} (A : Set n) {key d : ℕ} {c : Color } → (rb : RBTree A key c d ) → Color |
567 color {n} A {k} {d} {c} rb = c | |
568 | |
739
3443703a68cc
it is no good to develop all invariant at once
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
738
diff
changeset
|
569 rb-key : {n : Level} (A : Set n) {key d : ℕ} {c : Color } → (rb : RBTree A key c d ) → ℕ |
3443703a68cc
it is no good to develop all invariant at once
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
738
diff
changeset
|
570 rb-key {n} A {k} {d} {c} rb = k |
3443703a68cc
it is no good to develop all invariant at once
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
738
diff
changeset
|
571 |
723 | 572 RB→bt : {n : Level} (A : Set n) {key d : ℕ} {c : Color } → (rb : RBTree A key c d ) → bt A |
573 RB→bt {n} A (rb-leaf _) = leaf | |
574 RB→bt {n} A (rb-single key value _) = node key value leaf leaf | |
725 | 575 RB→bt {n} A (t-right-red key value x rb) = node key value leaf (RB→bt A rb) |
576 RB→bt {n} A (t-right-black key value x rb) = node key value leaf (RB→bt A rb) | |
577 RB→bt {n} A (t-left-red key value x rb) = node key value (RB→bt A rb) leaf | |
578 RB→bt {n} A (t-left-black key value x rb) = node key value (RB→bt A rb) leaf | |
579 RB→bt {n} A (t-node-red key value x x₁ rb rb₁) = node key value (RB→bt A rb) (RB→bt A rb₁) | |
580 RB→bt {n} A (t-node-black key value x x₁ rb rb₁) = node key value (RB→bt A rb) (RB→bt A rb₁) | |
723 | 581 |
724 | 582 rbt-depth : {n : Level} (A : Set n) {key : ℕ} {c : Color} {d : ℕ} → RBTree A key c d → ℕ |
583 rbt-depth A (rb-leaf _) = zero | |
584 rbt-depth A (rb-single _ value _) = 1 | |
725 | 585 rbt-depth A (t-right-red _ value x ab) = suc ( rbt-depth A ab) |
586 rbt-depth A (t-right-black _ value x ab) = suc ( rbt-depth A ab) | |
587 rbt-depth A (t-left-red _ value x ab) = suc ( rbt-depth A ab) | |
588 rbt-depth A (t-left-black _ value x ab) = suc ( rbt-depth A ab) | |
589 rbt-depth A (t-node-red _ value x x₁ ab ab₁) = suc (Data.Nat._⊔_ (rbt-depth A ab ) (rbt-depth A ab₁ )) | |
590 rbt-depth A (t-node-black _ value x x₁ ab ab₁) = suc (Data.Nat._⊔_ (rbt-depth A ab ) (rbt-depth A ab₁ )) | |
723 | 591 |
725 | 592 rbt-key : {n : Level} (A : Set n) {key : ℕ} {c : Color} {d : ℕ} → RBTree A key c d → Maybe ℕ |
593 rbt-key {n} A (rb-leaf _) = nothing | |
594 rbt-key {n} A (rb-single key value _) = just key | |
595 rbt-key {n} A (t-right-red key value x rb) = just key | |
596 rbt-key {n} A (t-right-black key value x rb) = just key | |
597 rbt-key {n} A (t-left-red key value x rb) = just key | |
598 rbt-key {n} A (t-left-black key value x rb) = just key | |
599 rbt-key {n} A (t-node-red key value x x₁ rb rb₁) = just key | |
600 rbt-key {n} A (t-node-black key value x x₁ rb rb₁) = just key | |
723 | 601 |
737 | 602 data ParentGrand {n : Level} {A : Set n} (self : bt A) : (parent grand : bt A) → Set n where |
740 | 603 s2-s1p2 : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } |
604 → parent ≡ node kp vp self n1 → grand ≡ node kg vg parent n2 → ParentGrand self parent grand | |
605 s2-1sp2 : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } | |
606 → parent ≡ node kp vp n1 self → grand ≡ node kg vg parent n2 → ParentGrand self parent grand | |
607 s2-s12p : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } | |
608 → parent ≡ node kp vp self n1 → grand ≡ node kg vg n2 parent → ParentGrand self parent grand | |
609 s2-1s2p : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } | |
610 → parent ≡ node kp vp n1 self → grand ≡ node kg vg n2 parent → ParentGrand self parent grand | |
734 | 611 |
736 | 612 data rotatedTree {n : Level} {A : Set n} : (before after : bt A ) → Set n where |
613 rr-node : {t : bt A} → rotatedTree t t | |
614 rr-right : {ka kb : ℕ } {va vb : A} → {ta tb tc ta1 tb1 tc1 : bt A} | |
615 → ka < kb | |
616 → rotatedTree ta ta1 → rotatedTree tb tb1 → rotatedTree tc tc1 | |
617 → rotatedTree (node ka va (node kb vb ta tb) tc) (node kb vb ta1 (node ka va tb1 tc1) ) | |
618 rr-left : {ka kb : ℕ } {va vb : A} → {ta tb tc ta1 tb1 tc1 : bt A} | |
619 → ka < kb | |
620 → rotatedTree ta ta1 → rotatedTree tb tb1 → rotatedTree tc tc1 | |
621 → rotatedTree (node kb vb ta (node ka va tb tc) ) (node ka va (node kb vb ta1 tb1) tc1) | |
734 | 622 |
740 | 623 record PG {n : Level } (A : Set n) (stack : List (bt A)) : Set n where |
624 field | |
625 self parent grand : bt A | |
626 pg : ParentGrand self parent grand | |
627 rest : List (bt A) | |
628 stack=gp : stack ≡ ( self ∷ parent ∷ grand ∷ rest ) | |
629 | |
630 stackToPG : {n : Level} {A : Set n} → {key : ℕ } → (tree orig : bt A ) | |
631 → (stack : List (bt A)) → stackInvariant key tree orig stack | |
632 → ( stack ≡ orig ∷ [] ) ∨ ( stack ≡ tree ∷ orig ∷ [] ) ∨ PG A stack | |
633 stackToPG {n} {A} {key} tree .tree .(tree ∷ []) s-nil = case1 refl | |
634 stackToPG {n} {A} {key} tree .(node _ _ _ tree) .(tree ∷ node _ _ _ tree ∷ []) (s-right x s-nil) = case2 (case1 refl) | |
635 stackToPG {n} {A} {key} tree orig (tree ∷ node _ _ _ tree ∷ grand ∷ rest) (s-right {t0} {t1} {t2} {k1} {v1} x (s-right {t3} {t4} {t5} {k2} {v2} x₁ si)) = case2 (case2 | |
636 record { self = tree ; parent = node k1 v1 t2 tree ; grand = grand ; pg = s2-1s2p refl ? ; rest = rest ; stack=gp = refl } ) | |
637 stackToPG {n} {A} {key} tree orig .(tree ∷ node _ _ _ tree ∷ _) (s-right x (s-left x₁ si)) = ? | |
638 stackToPG {n} {A} {key} tree orig .(tree ∷ _) (s-left x si) = ? | |
639 | |
640 | |
736 | 641 rbsi-len : {n : Level} {A : Set n} {orig parent grand : bt A} |
737 | 642 → ParentGrand orig parent grand → ℕ |
736 | 643 rbsi-len {n} {A} {s} {p} {g} st = ? |
734 | 644 |
737 | 645 findRBP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) {key1 d d1 : ℕ} → {c c1 : Color} → (tree : RBTree A key c d ) (orig : RBTree A key1 c1 d1 ) |
646 → (stack : List (bt A)) → stackInvariant key (RB→bt A tree) (RB→bt A orig) stack | |
647 → (next : {key0 d0 : ℕ} {c0 : Color} → (tree1 : RBTree A key0 c0 d0 ) → (stack : List (bt A)) → stackInvariant key (RB→bt A tree1) (RB→bt A orig) stack → rbt-depth A tree1 < rbt-depth A tree → t ) | |
648 → (exit : {key0 d0 : ℕ} {c0 : Color} → (tree1 : RBTree A key0 c0 d0 ) → (stack : List (bt A)) → stackInvariant key (RB→bt A tree1) (RB→bt A orig) stack | |
649 → (rbt-depth A tree1 ≡ 0 ) ∨ ( rbt-key A tree1 ≡ just key ) → t ) → t | |
650 findRBP {n} {m} {A} {t} key {key1} tree orig st si next exit = ? | |
723 | 651 |
731 | 652 rotateRight : ? |
653 rotateRight = ? | |
733 | 654 |
731 | 655 rotateLeft : ? |
656 rotateLeft = ? | |
657 | |
729 | 658 insertCase5 : {n m : Level} {A : Set n} {t : Set m} |
659 → (key : ℕ) → (value : A) → {key0 key1 key2 d0 d1 d2 : ℕ} {c0 c1 c2 : Color} | |
660 → (orig : RBTree A key1 c1 d1 ) → (tree : RBTree A key1 c1 d1 ) ( repl : RBTree A key2 c2 d2 ) | |
737 | 661 → (si : ParentGrand ? ? ?) |
736 | 662 → (ri : rotatedTree (RB→bt A tree) (RB→bt A repl)) |
729 | 663 → (next : ℕ → A → {k1 k2 d1 d2 : ℕ} {c1 c2 : Color} → (tree1 : RBTree A k1 c1 d1 ) (repl1 : RBTree A k2 c2 d2 ) |
737 | 664 → (si1 : ParentGrand ? ? ?) |
736 | 665 → (ri : rotatedTree (RB→bt A tree1) (RB→bt A repl1)) |
666 → rbsi-len si1 < rbsi-len si → t ) | |
729 | 667 → (exit : {k1 k2 d1 d2 : ℕ} {c1 c2 : Color} (tree1 : RBTree A k1 c1 d1 ) → (repl1 : RBTree A k2 c2 d2 ) |
736 | 668 → (ri : rotatedTree (RB→bt A orig) (RB→bt A repl1)) |
729 | 669 → t ) → t |
670 insertCase5 {n} {m} {A} {t} key value orig tree repl si ri next exit = ? where | |
737 | 671 insertCase51 : (key1 : ℕ) (si : ParentGrand ? ? ? ) → t |
729 | 672 insertCase51 = ? |
723 | 673 |
724 | 674 replaceRBP : {n m : Level} {A : Set n} {t : Set m} |
737 | 675 → (key : ℕ) → (value : A) → {key0 key1 d0 d1 : ℕ} {c0 c1 : Color} |
676 → (orig : RBTree A key0 c0 d0 ) → (tree : RBTree A key1 c1 d1 ) | |
677 → (stack : List (bt A)) → (si : stackInvariant key (RB→bt A tree) (RB→bt A orig) stack ) | |
738 | 678 → (next : {key2 d2 : ℕ} {c2 : Color} → (tree2 : RBTree A key2 c2 d2 ) |
679 → {tr to : bt A} → RB→bt A tree2 ≡ tr → RB→bt A orig ≡ to | |
680 → (stack1 : List (bt A)) → stackInvariant key tr to stack1 | |
737 | 681 → length stack1 < length stack → t ) |
682 → (exit : {k1 d1 : ℕ} {c1 : Color} → (repl1 : RBTree A k1 c1 d1 ) → (rot : bt A ) | |
738 | 683 → (ri : rotatedTree (RB→bt A orig) rot ) → replacedTree key value rot (RB→bt A repl1) → t ) → t |
740 | 684 replaceRBP {n} {m} {A} {t} key value {_} {key1} orig tree stack si next exit = insertCase1 stack _ _ refl refl si where |
738 | 685 insertCase2 : {k0 k1 d0 d1 d2 : ℕ} {c0 c1 c2 : Color} → (tree : RBTree A k0 c0 d0) |
686 → (parent : RBTree A k1 c1 d1) → (grand : RBTree A key1 c2 d2) | |
687 → (stack : List (bt A)) → (tr to : bt A) → RB→bt A grand ≡ tr → RB→bt A orig ≡ to → (si : stackInvariant key tr to stack ) | |
688 → (pg : ParentGrand (RB→bt A tree) (RB→bt A parent) tr ) → t | |
739
3443703a68cc
it is no good to develop all invariant at once
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
738
diff
changeset
|
689 insertCase2 tree parent grand stack tr to treq toeq si pg = ? |
740 | 690 insertCase1 : (stack : List (bt A)) → (tr to : bt A) → RB→bt A tree ≡ tr → RB→bt A orig ≡ to → (si : stackInvariant key tr to stack ) → t |
691 insertCase1 = ? | |
724 | 692 |
740 | 693 |