Mercurial > hg > Gears > GearsAgda
comparison hoareBinaryTree.agda @ 634:189cf03bda5f
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Fri, 12 Nov 2021 16:09:01 +0900 |
parents | 119f340c0b10 |
children | e30dcd03c07f |
comparison
equal
deleted
inserted
replaced
633:119f340c0b10 | 634:189cf03bda5f |
---|---|
136 depth-1< {i} {j} = s≤s (m≤m⊔n _ j) | 136 depth-1< {i} {j} = s≤s (m≤m⊔n _ j) |
137 | 137 |
138 depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) | 138 depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) |
139 depth-2< {i} {j} = s≤s (m≤n⊔m _ i) | 139 depth-2< {i} {j} = s≤s (m≤n⊔m _ i) |
140 | 140 |
141 lemma11 : {n : Level} {A : Set n} {v1 : A} → (key key₁ : ℕ) → (tree tree₁ : bt A ) | 141 treeLeftDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) |
142 → key < key₁ | 142 → treeInvariant (node k v1 tree tree₁) |
143 → treeInvariant tree | |
144 treeLeftDown {n} {A} {_} {v1} leaf leaf (t-single k1 v1) = t-leaf | |
145 treeLeftDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = t-leaf | |
146 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = ti | |
147 treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti | |
148 | |
149 treeRightDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) | |
150 → treeInvariant (node k v1 tree tree₁) | |
151 → treeInvariant tree₁ | |
152 treeRightDown {n} {A} {_} {v1} .leaf .leaf (t-single _ .v1) = t-leaf | |
153 treeRightDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = ti | |
154 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = t-leaf | |
155 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti₁ | |
156 | |
157 siConsLeft : {n : Level } {A : Set n} (key key₁ : ℕ) → { v1 : A } (tree tree₁ tree0 : bt A ) (st : List (bt A)) | |
158 → key < key₁ → stackInvariant key (node key₁ v1 tree tree₁) tree0 st | |
143 → treeInvariant (node key₁ v1 tree tree₁) | 159 → treeInvariant (node key₁ v1 tree tree₁) |
144 → treeInvariant tree | 160 → stackInvariant key tree tree0 (node key₁ v1 tree tree₁ ∷ st) |
145 lemma11 = {!!} | 161 siConsLeft {n} {A} k k1 {v1} t t1 t0 st k<k1 ti si = {!!} |
146 | 162 |
147 -- stackInvariant key (node key₁ v1 tree tree₁) tree0 st | 163 -- stackInvariant key (node key₁ v1 tree tree₁) tree0 st |
148 -- → stackInvariant key tree tree0 (node key₁ v1 tree tree₁ ∷ st) | 164 -- → stackInvariant key tree tree0 (node key₁ v1 tree tree₁ ∷ st) |
149 | 165 |
150 open _∧_ | 166 open _∧_ |
154 → (next : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) | 170 → (next : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) |
155 → (exit : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → t ) → t | 171 → (exit : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → t ) → t |
156 findP key leaf tree0 st Pre _ exit = exit leaf tree0 st Pre | 172 findP key leaf tree0 st Pre _ exit = exit leaf tree0 st Pre |
157 findP key (node key₁ v1 tree tree₁) tree0 st Pre next exit with <-cmp key key₁ | 173 findP key (node key₁ v1 tree tree₁) tree0 st Pre next exit with <-cmp key key₁ |
158 findP key n tree0 st Pre _ exit | tri≈ ¬a b ¬c = exit n tree0 st Pre | 174 findP key n tree0 st Pre _ exit | tri≈ ¬a b ¬c = exit n tree0 st Pre |
159 findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree tree0 (n ∷ st) ⟪ lemma11 {!!} {!!} {!!} {!!} {!!} (proj1 Pre) , {!!} ⟫ depth-1< | 175 findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree tree0 (n ∷ st) ⟪ treeLeftDown tree tree₁ (proj1 Pre) , findP1 a (proj2 Pre) ⟫ depth-1< where |
176 findP1 : key < key₁ → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (node key₁ v1 tree tree₁ ∷ st) | |
177 findP1 a si = siConsLeft key key₁ {v1} tree tree₁ tree0 st a si (proj1 Pre) | |
160 findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ tree0 (n ∷ st) {!!} depth-2< | 178 findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ tree0 (n ∷ st) {!!} depth-2< |
161 | 179 |
162 | 180 |
163 replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) → (treeInvariant tree ) | 181 replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) → (treeInvariant tree ) |
164 → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value tree tree1 → t) → t | 182 → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value tree tree1 → t) → t |
246 next tree (n ∷ st) (record {ti = findPR.ti Pre ; si = findPP2 st (findPR.si Pre) ; ci = lift tt} ) findPP1 where | 264 next tree (n ∷ st) (record {ti = findPR.ti Pre ; si = findPP2 st (findPR.si Pre) ; ci = lift tt} ) findPP1 where |
247 tree0 = findPR.tree0 Pre | 265 tree0 = findPR.tree0 Pre |
248 findPP2 : (st : List (bt A)) → stackInvariant key {!!} tree0 st → stackInvariant key {!!} tree0 (node key₁ v1 tree tree₁ ∷ st) | 266 findPP2 : (st : List (bt A)) → stackInvariant key {!!} tree0 st → stackInvariant key {!!} tree0 (node key₁ v1 tree tree₁ ∷ st) |
249 findPP2 = {!!} | 267 findPP2 = {!!} |
250 findPP1 : suc ( bt-depth tree ) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) | 268 findPP1 : suc ( bt-depth tree ) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) |
251 findPP1 = {!!} | 269 findPP1 = depth-1< |
252 findPP key n@(node key₁ v1 tree tree₁) st Pre next exit | tri> ¬a ¬b c = next tree₁ (n ∷ st) {!!} findPP2 where -- Cond n st → Cond tree₁ (n ∷ st) | 270 findPP key n@(node key₁ v1 tree tree₁) st Pre next exit | tri> ¬a ¬b c = next tree₁ (n ∷ st) {!!} findPP2 where -- Cond n st → Cond tree₁ (n ∷ st) |
253 findPP2 : suc (bt-depth tree₁) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) | 271 findPP2 : suc (bt-depth tree₁) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) |
254 findPP2 = {!!} | 272 findPP2 = depth-2< |
255 | 273 |
256 insertTreePP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree | 274 insertTreePP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
257 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t | 275 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t |
258 insertTreePP {n} {m} {A} {t} tree key value P exit = | 276 insertTreePP {n} {m} {A} {t} tree key value P exit = |
259 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ t s → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ {!!} | 277 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ t s → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ {!!} |
283 findPPC key value n st P next exit | tri> ¬a ¬b c = {!!} | 301 findPPC key value n st P next exit | tri> ¬a ¬b c = {!!} |
284 | 302 |
285 containsTree : {n m : Level} {A : Set n} {t : Set m} → (tree tree1 : bt A) → (key : ℕ) → (value : A) → treeInvariant tree1 → replacedTree key value tree1 tree → ⊤ | 303 containsTree : {n m : Level} {A : Set n} {t : Set m} → (tree tree1 : bt A) → (key : ℕ) → (value : A) → treeInvariant tree1 → replacedTree key value tree1 tree → ⊤ |
286 containsTree {n} {m} {A} {t} tree tree1 key value P RT = | 304 containsTree {n} {m} {A} {t} tree tree1 key value P RT = |
287 TerminatingLoopS (bt A ∧ List (bt A) ) | 305 TerminatingLoopS (bt A ∧ List (bt A) ) |
288 {λ p → findPR key (proj1 p) (proj2 p) (findPC key value ) } (λ p → bt-depth (proj1 p)) | 306 {λ p → findPR key (proj1 p) (proj2 p) (findPC key value ) } (λ p → bt-depth (proj1 p)) -- findPR key tree1 [] (findPC key value) |
289 ⟪ tree1 , [] ⟫ {!!} | 307 ⟪ tree1 , [] ⟫ record { tree0 = tree ; ti = {!!} ; si = {!!} ; ci = record { tree1 = tree ; ci = RT } } |
290 $ λ p P loop → findPPC key value (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) | 308 $ λ p P loop → findPPC key value (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) |
291 $ λ t1 s1 found? P2 → insertTreeSpec0 t1 value (lemma6 t1 s1 found? P2) where | 309 $ λ t1 s1 found? P2 → insertTreeSpec0 t1 value (lemma6 t1 s1 found? P2) where |
292 lemma6 : (t1 : bt A) (s1 : List (bt A)) (found? : (t1 ≡ leaf) ∨ (node-key t1 ≡ just key)) (P2 : findPR key t1 s1 (findPC key value)) → top-value t1 ≡ just value | 310 lemma6 : (t1 : bt A) (s1 : List (bt A)) (found? : (t1 ≡ leaf) ∨ (node-key t1 ≡ just key)) (P2 : findPR key t1 s1 (findPC key value)) → top-value t1 ≡ just value |
293 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 | 311 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 |
294 lemma7 : (t1 : bt A) ( s1 : List (bt A) ) (tree0 tree1 : bt A) → | 312 lemma7 : (t1 : bt A) ( s1 : List (bt A) ) (tree0 tree1 : bt A) → |