Mercurial > hg > Gears > GearsAgda
comparison hoareBinaryTree.agda @ 621:6861bcb9c54d
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 08 Nov 2021 16:36:26 +0900 |
parents | fe8c2d82c05c |
children | a1849f24fa66 |
comparison
equal
deleted
inserted
replaced
620:fe8c2d82c05c | 621:6861bcb9c54d |
---|---|
96 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) | 96 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) |
97 | 97 |
98 treeInvariantTest1 : treeInvariant (node 3 0 leaf (node 1 1 leaf (node 3 5 leaf leaf))) | 98 treeInvariantTest1 : treeInvariant (node 3 0 leaf (node 1 1 leaf (node 3 5 leaf leaf))) |
99 treeInvariantTest1 = {!!} | 99 treeInvariantTest1 = {!!} |
100 | 100 |
101 data stackInvariant {n : Level} {A : Set n} : (tree0 : bt A) → (stack : List (bt A)) → Set n where | 101 data stackInvariant {n : Level} {A : Set n} : (tree tree0 : bt A) → (stack : List (bt A)) → Set n where |
102 s-nil : stackInvariant leaf [] | 102 s-nil : stackInvariant leaf leaf [] |
103 s-single : (tree : bt A) → stackInvariant tree (tree ∷ [] ) | 103 s-single : (tree : bt A) → stackInvariant tree tree (tree ∷ [] ) |
104 s-right : (tree : bt A) → {key : ℕ } → {value : A } { left : bt A} → stackInvariant (node key value left tree ) (tree ∷ node key value left tree ∷ []) | |
105 s-left : (tree : bt A) → {key : ℕ } → {value : A } { right : bt A} → stackInvariant (node key value tree right) (tree ∷ node key value tree right ∷ []) | |
106 s-< : (tree0 tree : bt A) → {key : ℕ } → {value : A } { left : bt A} → {st : List (bt A)} | 104 s-< : (tree0 tree : bt A) → {key : ℕ } → {value : A } { left : bt A} → {st : List (bt A)} |
107 → stackInvariant tree0 (tree ∷ st ) → stackInvariant tree0 ((node key value left tree ) ∷ tree ∷ st ) | 105 → stackInvariant (node key value left tree ) tree0 (node key value left tree ∷ st ) → stackInvariant tree tree0 (tree ∷ node key value left tree ∷ st ) |
108 s-> : (tree0 tree : bt A) → {key : ℕ } → {value : A } { right : bt A} → {st : List (bt A)} | 106 s-> : (tree0 tree : bt A) → {key : ℕ } → {value : A } { right : bt A} → {st : List (bt A)} |
109 → stackInvariant tree0 (tree ∷ st ) → stackInvariant tree0 ((node key value tree right ) ∷ tree ∷ st ) | 107 → stackInvariant (node key value tree right ) tree0 (node key value tree right ∷ st ) → stackInvariant tree tree0 (tree ∷ node key value tree right ∷ st ) |
110 | 108 |
111 data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (tree tree1 : bt A ) → Set n where | 109 data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (tree tree1 : bt A ) → Set n where |
112 r-leaf : replacedTree key value leaf (node key value leaf leaf) | 110 r-leaf : replacedTree key value leaf (node key value leaf leaf) |
113 r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) | 111 r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) |
114 r-right : {k : ℕ } {v : A} → {t t1 t2 : bt A} | 112 r-right : {k : ℕ } {v : A} → {t t1 t2 : bt A} |
115 → k > key → ( replacedTree key value t1 t2 → replacedTree key value (node k v t t1) (node k v t t2) ) | 113 → k > key → ( replacedTree key value t1 t2 → replacedTree key value (node k v t t1) (node k v t t2) ) |
116 r-left : {k : ℕ } {v : A} → {t t1 t2 : bt A} | 114 r-left : {k : ℕ } {v : A} → {t t1 t2 : bt A} |
117 → k < key → ( replacedTree key value t1 t2 → replacedTree key value (node k v t1 t) (node k v t2 t) ) | 115 → k < key → ( replacedTree key value t1 t2 → replacedTree key value (node k v t1 t) (node k v t2 t) ) |
118 | 116 |
119 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A)) | 117 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A)) |
120 → treeInvariant tree ∧ stackInvariant tree0 stack | 118 → treeInvariant tree ∧ stackInvariant tree tree0 stack |
121 → (next : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree0 stack → bt-depth tree1 < bt-depth tree → t ) | 119 → (next : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) |
122 → (exit : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree0 stack → t ) → t | 120 → (exit : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree1 tree0 stack → t ) → t |
123 findP key leaf tree0 st Pre _ exit = exit leaf tree0 st {!!} | 121 findP key leaf tree0 st Pre _ exit = exit leaf tree0 st {!!} |
124 findP key (node key₁ v tree tree₁) tree0 st Pre next exit with <-cmp key key₁ | 122 findP key (node key₁ v tree tree₁) tree0 st Pre next exit with <-cmp key key₁ |
125 findP key n tree0 st Pre _ exit | tri≈ ¬a b ¬c = exit n tree0 st {!!} | 123 findP key n tree0 st Pre _ exit | tri≈ ¬a b ¬c = exit n tree0 st {!!} |
126 findP key n@(node key₁ v tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree tree0 (n ∷ st) {!!} {!!} | 124 findP key n@(node key₁ v tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree tree0 (n ∷ st) {!!} {!!} |
127 findP key n@(node key₁ v tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ tree0 (n ∷ st) {!!} {!!} | 125 findP key n@(node key₁ v tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ tree0 (n ∷ st) {!!} {!!} |
130 → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value tree tree1 → t) → t | 128 → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value tree tree1 → t) → t |
131 replaceNodeP k v leaf P next = next (node k v leaf leaf) {!!} {!!} | 129 replaceNodeP k v leaf P next = next (node k v leaf leaf) {!!} {!!} |
132 replaceNodeP k v (node key value t t₁) P next = next (node k v t t₁) {!!} {!!} | 130 replaceNodeP k v (node key value t t₁) P next = next (node k v t t₁) {!!} {!!} |
133 | 131 |
134 replaceP : {n m : Level} {A : Set n} {t : Set m} | 132 replaceP : {n m : Level} {A : Set n} {t : Set m} |
135 → (key : ℕ) → (value : A) → (tree repl : bt A) → (stack : List (bt A)) → treeInvariant tree ∧ stackInvariant tree stack ∧ replacedTree key value tree repl | 133 → (key : ℕ) → (value : A) → (tree repl : bt A) → (stack : List (bt A)) → treeInvariant tree ∧ stackInvariant repl tree stack ∧ replacedTree key value tree repl |
136 → (next : ℕ → A → (tree1 repl : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree1 stack ∧ replacedTree key value tree1 repl → bt-depth tree1 < bt-depth tree → t ) | 134 → (next : ℕ → A → (tree1 repl : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant repl tree1 stack ∧ replacedTree key value tree1 repl → bt-depth tree1 < bt-depth tree → t ) |
137 → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t | 135 → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t |
138 replaceP key value tree repl [] Pre next exit = exit tree repl {!!} | 136 replaceP key value tree repl [] Pre next exit = exit tree repl {!!} |
139 replaceP key value tree repl (leaf ∷ st) Pre next exit = next key value tree {!!} st {!!} {!!} | 137 replaceP key value tree repl (leaf ∷ st) Pre next exit = next key value tree {!!} st {!!} {!!} |
140 replaceP key value tree repl (node key₁ value₁ left right ∷ st) Pre next exit with <-cmp key key₁ | 138 replaceP key value tree repl (node key₁ value₁ left right ∷ st) Pre next exit with <-cmp key key₁ |
141 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ tree right ) {!!} st {!!} {!!} | 139 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ tree right ) {!!} st {!!} {!!} |
175 RTtoTI1 = {!!} | 173 RTtoTI1 = {!!} |
176 | 174 |
177 insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree | 175 insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
178 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t | 176 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t |
179 insertTreeP {n} {m} {A} {t} tree key value P exit = | 177 insertTreeP {n} {m} {A} {t} tree key value P exit = |
180 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant tree (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ ⟪ P , {!!} ⟫ | 178 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant (proj1 p) tree (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ ⟪ P , {!!} ⟫ |
181 $ λ p P loop → findP key (proj1 p) tree (proj2 p) {!!} (λ t _ s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) | 179 $ λ p P loop → findP key (proj1 p) tree (proj2 p) {!!} (λ t _ s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) |
182 $ λ t _ s P → replaceNodeP key value t (proj1 P) | 180 $ λ t _ s P → replaceNodeP key value t (proj1 P) |
183 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) | 181 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) |
184 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant (proj1 (proj2 p)) (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } | 182 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant (proj1 (proj2 p)) tree (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } |
185 (λ p → bt-depth (proj1 (proj2 p))) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ proj1 P , ⟪ {!!} , R ⟫ ⟫ | 183 (λ p → bt-depth (proj1 (proj2 p))) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ proj1 P , ⟪ {!!} , R ⟫ ⟫ |
186 $ λ p P1 loop → replaceP key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) P1 | 184 $ λ p P1 loop → replaceP key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) {!!} |
187 (λ key value tree1 repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ P2 lt ) exit | 185 (λ key value tree1 repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ {!!} lt ) exit |
188 | 186 |
189 top-value : {n : Level} {A : Set n} → (tree : bt A) → Maybe A | 187 top-value : {n : Level} {A : Set n} → (tree : bt A) → Maybe A |
190 top-value leaf = nothing | 188 top-value leaf = nothing |
191 top-value (node key value tree tree₁) = just value | 189 top-value (node key value tree tree₁) = just value |
192 | 190 |
195 | 193 |
196 record findPR {n : Level} {A : Set n} (tree : bt A ) (stack : List (bt A)) : Set n where | 194 record findPR {n : Level} {A : Set n} (tree : bt A ) (stack : List (bt A)) : Set n where |
197 field | 195 field |
198 tree0 : bt A | 196 tree0 : bt A |
199 ti : treeInvariant tree | 197 ti : treeInvariant tree |
200 si : stackInvariant tree0 stack | 198 si : stackInvariant tree tree0 stack |
201 | 199 |
202 findPP : {n m : Level} {A : Set n} {t : Set m} | 200 findPP : {n m : Level} {A : Set n} {t : Set m} |
203 → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) | 201 → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) |
204 → (Pre : bt A → List (bt A) → findPR tree stack ) | 202 → (Pre : findPR tree stack ) |
205 → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR tree1 stack1 → bt-depth tree1 < bt-depth tree → t ) | 203 → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR tree1 stack1 → bt-depth tree1 < bt-depth tree → t ) |
206 → (exit : (tree1 : bt A) → (stack1 : List (bt A)) → findPR tree1 stack1 → t) → t | 204 → (exit : (tree1 : bt A) → (stack1 : List (bt A)) → findPR tree1 stack1 → t) → t |
207 findPP key leaf st Pre next exit = exit leaf st (Pre leaf st ) | 205 findPP key leaf st Pre next exit = exit leaf st Pre |
208 findPP key (node key₁ v tree tree₁) st Pre next exit with <-cmp key key₁ | 206 findPP key (node key₁ v tree tree₁) st Pre next exit with <-cmp key key₁ |
209 findPP key n st P next exit | tri≈ ¬a b ¬c = exit n st (P n st) | 207 findPP key n st P next exit | tri≈ ¬a b ¬c = exit n st P |
210 findPP {_} {_} {A} key n@(node key₁ v tree tree₁) st Pre next exit | tri< a ¬b ¬c = | 208 findPP {_} {_} {A} key n@(node key₁ v tree tree₁) st Pre next exit | tri< a ¬b ¬c = |
211 next tree (n ∷ st) (record {ti = findPP0 tree tree₁ (findPR.ti (Pre n st)) ; si = findPP2 st (findPR.si (Pre n st))} ) findPP1 where | 209 next tree (n ∷ st) (record {ti = findPP0 tree tree₁ (findPR.ti Pre ) ; si = findPP2 st (findPR.si Pre)} ) findPP1 where |
212 tree0 = findPR.tree0 (Pre n st) | 210 tree0 = findPR.tree0 Pre |
213 findPP0 : (tree tree₁ : bt A) → treeInvariant ( node key₁ v tree tree₁ ) → treeInvariant tree | 211 findPP0 : (tree tree₁ : bt A) → treeInvariant ( node key₁ v tree tree₁ ) → treeInvariant tree |
214 findPP0 leaf t x = {!!} | 212 findPP0 leaf t x = {!!} |
215 findPP0 (node key value tree tree₁) leaf x = proj1 {!!} | 213 findPP0 (node key value tree tree₁) leaf x = proj1 {!!} |
216 findPP0 (node key value tree tree₁) (node key₁ value₁ t t₁) x = proj1 {!!} | 214 findPP0 (node key value tree tree₁) (node key₁ value₁ t t₁) x = proj1 {!!} |
217 findPP2 : (st : List (bt A)) → stackInvariant tree0 st → stackInvariant tree0 (node key₁ v tree tree₁ ∷ st) | 215 findPP2 : (st : List (bt A)) → stackInvariant {!!} tree0 st → stackInvariant {!!} tree0 (node key₁ v tree tree₁ ∷ st) |
218 findPP2 [] = {!!} | 216 findPP2 [] = {!!} |
219 findPP2 (leaf ∷ st) x = {!!} | 217 findPP2 (leaf ∷ st) x = {!!} |
220 findPP2 (node key value leaf leaf ∷ st) x = {!!} | 218 findPP2 (node key value leaf leaf ∷ st) x = {!!} |
221 findPP2 (node key value leaf (node key₁ value₁ x₂ x₃) ∷ st) x = {!!} | 219 findPP2 (node key value leaf (node key₁ value₁ x₂ x₃) ∷ st) x = {!!} |
222 findPP2 (node key value (node key₁ value₁ x₁ x₃) leaf ∷ st) x = {!!} | 220 findPP2 (node key value (node key₁ value₁ x₁ x₃) leaf ∷ st) x = {!!} |
232 insertTreePP {n} {m} {A} {t} tree key value P exit = | 230 insertTreePP {n} {m} {A} {t} tree key value P exit = |
233 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR (proj1 p) (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ {!!} | 231 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR (proj1 p) (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ {!!} |
234 $ λ p P loop → findPP key (proj1 p) (proj2 p) {!!} (λ t s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) | 232 $ λ p P loop → findPP key (proj1 p) (proj2 p) {!!} (λ t s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) |
235 $ λ t s P → replaceNodeP key value t {!!} | 233 $ λ t s P → replaceNodeP key value t {!!} |
236 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) | 234 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) |
237 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant (proj1 (proj2 p)) (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } | 235 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant (proj1 (proj2 p)) tree (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } |
238 (λ p → bt-depth (proj1 (proj2 p))) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ {!!} , ⟪ {!!} , R ⟫ ⟫ | 236 (λ p → bt-depth (proj1 (proj2 p))) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ {!!} , ⟪ {!!} , R ⟫ ⟫ |
239 $ λ p P1 loop → replaceP key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) P1 | 237 $ λ p P1 loop → replaceP key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) {!!} |
240 (λ key value tree1 repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ P2 lt ) exit | 238 (λ key value tree1 repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ {!!} lt ) exit |
241 | 239 |
242 -- findP key tree stack = findPP key tree stack {findPR} → record { ti = tree-invariant tree ; si stack-invariant tree stack } → | 240 -- findP key tree stack = findPP key tree stack {findPR} → record { ti = tree-invariant tree ; si stack-invariant tree stack } → |
243 | 241 |
244 record findP-contains {n : Level} {A : Set n} (tree : bt A ) (stack : List (bt A)) : Set n where | 242 record findP-contains {n : Level} {A : Set n} (tree : bt A ) (stack : List (bt A)) : Set n where |
245 field | 243 field |