Mercurial > hg > Gears > GearsAgda
annotate hoareBinaryTree.agda @ 628:ec2506b532ba
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 08 Nov 2021 23:44:24 +0900 |
parents | 967547859521 |
children | 7a19d4b43795 |
rev | line source |
---|---|
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
1 module hoareBinaryTree where |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
2 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
3 open import Level renaming (zero to Z ; suc to succ) |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
4 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
5 open import Data.Nat hiding (compare) |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
6 open import Data.Nat.Properties as NatProp |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
7 open import Data.Maybe |
588 | 8 -- open import Data.Maybe.Properties |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
9 open import Data.Empty |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
10 open import Data.List |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
11 open import Data.Product |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
12 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
13 open import Function as F hiding (const) |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
14 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
15 open import Relation.Binary |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
16 open import Relation.Binary.PropositionalEquality |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
17 open import Relation.Nullary |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
18 open import logic |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
19 |
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
20 |
588 | 21 _iso_ : {n : Level} {a : Set n} → ℕ → ℕ → Set |
22 d iso d' = (¬ (suc d ≤ d')) ∧ (¬ (suc d' ≤ d)) | |
23 | |
24 iso-intro : {n : Level} {a : Set n} {x y : ℕ} → ¬ (suc x ≤ y) → ¬ (suc y ≤ x) → _iso_ {n} {a} x y | |
25 iso-intro = λ z z₁ → record { proj1 = z ; proj2 = z₁ } | |
26 | |
590 | 27 -- |
28 -- | |
29 -- no children , having left node , having right node , having both | |
30 -- | |
597 | 31 data bt {n : Level} (A : Set n) : Set n where |
604 | 32 leaf : bt A |
33 node : (key : ℕ) → (value : A) → | |
610 | 34 (left : bt A ) → (right : bt A ) → bt A |
600 | 35 |
620 | 36 node-key : {n : Level} {A : Set n} → bt A → Maybe ℕ |
37 node-key (node key _ _ _) = just key | |
38 node-key _ = nothing | |
39 | |
40 node-value : {n : Level} {A : Set n} → bt A → Maybe A | |
41 node-value (node _ value _ _) = just value | |
42 node-value _ = nothing | |
43 | |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
44 bt-depth : {n : Level} {A : Set n} → (tree : bt A ) → ℕ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
45 bt-depth leaf = 0 |
618 | 46 bt-depth (node key value t t₁) = suc (Data.Nat._⊔_ (bt-depth t ) (bt-depth t₁ )) |
606 | 47 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
48 find : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → List (bt A) |
604 | 49 → (next : bt A → List (bt A) → t ) → (exit : bt A → List (bt A) → t ) → t |
50 find key leaf st _ exit = exit leaf st | |
51 find key (node key₁ v 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₁ v tree tree₁) st next _ | tri< a ¬b ¬c = next tree (n ∷ st) | |
54 find key n@(node key₁ v tree tree₁) st next _ | tri> ¬a ¬b c = next tree₁ (n ∷ st) | |
597 | 55 |
604 | 56 {-# TERMINATING #-} |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
57 find-loop : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → bt A → List (bt A) → (exit : bt A → List (bt A) → t) → t |
611 | 58 find-loop {n} {m} {A} {t} key tree st exit = find-loop1 tree st where |
604 | 59 find-loop1 : bt A → List (bt A) → t |
60 find-loop1 tree st = find key tree st find-loop1 exit | |
600 | 61 |
611 | 62 replaceNode : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → (bt A → t) → t |
63 replaceNode k v leaf next = next (node k v leaf leaf) | |
64 replaceNode k v (node key value t t₁) next = next (node k v t t₁) | |
65 | |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
66 replace : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → List (bt A) → (next : ℕ → A → bt A → List (bt A) → t ) → (exit : bt A → t) → t |
604 | 67 replace key value tree [] next exit = exit tree |
68 replace key value tree (leaf ∷ st) next exit = next key value tree st | |
69 replace key value tree (node key₁ value₁ left right ∷ st) next exit with <-cmp key key₁ | |
70 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ tree right ) st | |
71 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st | |
72 ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left tree ) st | |
586
0ddfa505d612
isolate search function problem, and add hoareBinaryTree.agda.
ryokka
parents:
diff
changeset
|
73 |
604 | 74 {-# TERMINATING #-} |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
75 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
|
76 replace-loop {_} {_} {A} {t} key value tree st exit = replace-loop1 key value tree st where |
604 | 77 replace-loop1 : (key : ℕ) → (value : A) → bt A → List (bt A) → t |
78 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
|
79 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
80 insertTree : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → (next : bt A → t ) → t |
611 | 81 insertTree tree key value exit = find-loop key tree [] $ λ t st → replaceNode key value t $ λ t1 → replace-loop key value t1 st exit |
587 | 82 |
604 | 83 insertTest1 = insertTree leaf 1 1 (λ x → x ) |
611 | 84 insertTest2 = insertTree insertTest1 2 1 (λ x → x ) |
587 | 85 |
605 | 86 open import Data.Unit hiding ( _≟_ ; _≤?_ ; _≤_) |
87 | |
620 | 88 data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where |
89 t-leaf : treeInvariant leaf | |
90 t-single : {key : ℕ} → {value : A} → treeInvariant (node key value leaf leaf) | |
91 t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key₁ value₁ t₁ t₂) → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) | |
92 t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key₁ < key) → treeInvariant (node key value₁ t₁ t₂) → treeInvariant (node key₁ value₁ (node key value₁ t₁ t₂) leaf ) | |
93 t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → (key < key₁) → (key₁ < key₂) | |
94 → treeInvariant (node key value t₁ t₂) | |
95 → treeInvariant (node key₂ value₂ t₃ t₄) | |
96 → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) | |
605 | 97 |
620 | 98 treeInvariantTest1 : treeInvariant (node 3 0 leaf (node 1 1 leaf (node 3 5 leaf leaf))) |
99 treeInvariantTest1 = {!!} | |
605 | 100 |
627 | 101 data stackInvariant {n : Level} {A : Set n} (key0 : ℕ) : (tree tree0 : bt A) → (stack : List (bt A)) → Set n where |
102 s-nil : stackInvariant key0 leaf leaf [] | |
103 s-single : (tree : bt A) → stackInvariant key0 tree tree (tree ∷ [] ) | |
620 | 104 s-< : (tree0 tree : bt A) → {key : ℕ } → {value : A } { left : bt A} → {st : List (bt A)} |
627 | 105 → key < key0 → stackInvariant key0(node key value left tree ) tree0 (node key value left tree ∷ st ) → stackInvariant key0 tree tree0 (tree ∷ node key value left tree ∷ st ) |
621 | 106 s-> : (tree0 tree : bt A) → {key : ℕ } → {value : A } { right : bt A} → {st : List (bt A)} |
627 | 107 → key0 < key → stackInvariant key0(node key value tree right ) tree0 (node key value tree right ∷ st ) → stackInvariant key0 tree tree0 (tree ∷ node key value tree right ∷ st ) |
606 | 108 |
613 | 109 data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (tree tree1 : bt A ) → Set n where |
615 | 110 r-leaf : replacedTree key value leaf (node key value leaf leaf) |
111 r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) | |
112 r-right : {k : ℕ } {v : A} → {t t1 t2 : bt A} | |
627 | 113 → k > key → replacedTree key value t1 t2 → replacedTree key value (node k v t t1) (node k v t t2) |
615 | 114 r-left : {k : ℕ } {v : A} → {t t1 t2 : bt A} |
627 | 115 → k < key → replacedTree key value t1 t2 → replacedTree key value (node k v t1 t) (node k v t2 t) |
611 | 116 |
615 | 117 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A)) |
627 | 118 → treeInvariant tree ∧ stackInvariant key tree tree0 stack |
119 → (next : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) | |
120 → (exit : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → t ) → t | |
615 | 121 findP key leaf tree0 st Pre _ exit = exit leaf tree0 st {!!} |
122 findP key (node key₁ v tree tree₁) tree0 st Pre next exit with <-cmp key key₁ | |
123 findP key n tree0 st Pre _ exit | tri≈ ¬a b ¬c = exit n tree0 st {!!} | |
124 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) {!!} {!!} | |
606 | 126 |
611 | 127 replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) → (treeInvariant tree ) |
613 | 128 → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value tree tree1 → t) → t |
129 replaceNodeP k v leaf P next = next (node k v leaf leaf) {!!} {!!} | |
130 replaceNodeP k v (node key value t t₁) P next = next (node k v t t₁) {!!} {!!} | |
606 | 131 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
132 replaceP : {n m : Level} {A : Set n} {t : Set m} |
627 | 133 → (key : ℕ) → (value : A) → (tree repl : bt A) → (stack : List (bt A)) → treeInvariant tree ∧ stackInvariant key repl tree stack ∧ replacedTree key value tree repl |
134 → (next : ℕ → A → (tree1 repl : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key repl tree1 stack ∧ replacedTree key value tree1 repl → bt-depth tree1 < bt-depth tree → t ) | |
613 | 135 → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t |
136 replaceP key value tree repl [] Pre next exit = exit tree repl {!!} | |
614 | 137 replaceP key value tree repl (leaf ∷ st) Pre next exit = next key value tree {!!} st {!!} {!!} |
613 | 138 replaceP key value tree repl (node key₁ value₁ left right ∷ st) Pre next exit with <-cmp key key₁ |
614 | 139 ... | tri< a ¬b ¬c = next key value (node key₁ value₁ tree right ) {!!} st {!!} {!!} |
140 ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) {!!} st {!!} {!!} | |
141 ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left tree ) {!!} st {!!} {!!} | |
606 | 142 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
143 open import Relation.Binary.Definitions |
606 | 144 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
145 nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
146 nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
147 lemma3 : {i j : ℕ} → 0 ≡ i → j < i → ⊥ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
148 lemma3 refl () |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
149 lemma5 : {i j : ℕ} → i < 1 → j < i → ⊥ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
150 lemma5 (s≤s z≤n) () |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
151 |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
152 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
|
153 → (r : Index) → (p : Invraiant r) |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
154 → (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
|
155 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
|
156 ... | 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
|
157 ... | 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
|
158 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
|
159 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
|
160 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
|
161 ... | 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
|
162 ... | 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
|
163 ... | 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
|
164 |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
165 open _∧_ |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
166 |
615 | 167 RTtoTI0 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
168 → replacedTree key value tree repl → treeInvariant repl | |
169 RTtoTI0 = {!!} | |
170 | |
171 RTtoTI1 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant repl | |
172 → replacedTree key value tree repl → treeInvariant tree | |
173 RTtoTI1 = {!!} | |
614 | 174 |
611 | 175 insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
613 | 176 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t |
610 | 177 insertTreeP {n} {m} {A} {t} tree key value P exit = |
627 | 178 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant key (proj1 p) tree (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ ⟪ P , {!!} ⟫ |
615 | 179 $ λ p P loop → findP key (proj1 p) tree (proj2 p) {!!} (λ t _ s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) |
180 $ λ t _ s P → replaceNodeP key value t (proj1 P) | |
614 | 181 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) |
627 | 182 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant key (proj1 (proj2 p)) tree (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } |
615 | 183 (λ p → bt-depth (proj1 (proj2 p))) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ proj1 P , ⟪ {!!} , R ⟫ ⟫ |
621 | 184 $ λ p P1 loop → replaceP key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) {!!} |
185 (λ key value tree1 repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ {!!} lt ) exit | |
614 | 186 |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
187 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
|
188 top-value leaf = nothing |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
189 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
|
190 |
612 | 191 insertTreeSpec0 : {n : Level} {A : Set n} → (tree : bt A) → (value : A) → top-value tree ≡ just value → ⊤ |
609
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
192 insertTreeSpec0 _ _ _ = tt |
79418701a283
add test and speciication
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
606
diff
changeset
|
193 |
627 | 194 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 | 195 field |
619 | 196 tree0 : bt A |
622 | 197 ti : treeInvariant tree0 |
627 | 198 si : stackInvariant key tree tree0 stack |
623 | 199 ci : C tree stack |
618 | 200 |
616 | 201 findPP : {n m : Level} {A : Set n} {t : Set m} |
202 → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) | |
627 | 203 → (Pre : findPR key tree stack (λ t s → Lift n ⊤)) |
204 → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR key tree1 stack1 (λ t s → Lift n ⊤) → bt-depth tree1 < bt-depth tree → t ) | |
205 → (exit : (tree1 : bt A) → (stack1 : List (bt A)) → ( tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key) → findPR key tree1 stack1 (λ t s → Lift n ⊤) → t) → t | |
625 | 206 findPP key leaf st Pre next exit = exit leaf st (case1 refl) Pre |
624 | 207 findPP key (node key₁ v tree tree₁) st Pre next exit with <-cmp key key₁ |
625 | 208 findPP key n st P next exit | tri≈ ¬a b ¬c = exit n st (case2 {!!}) P |
624 | 209 findPP {_} {_} {A} key n@(node key₁ v tree tree₁) st Pre next exit | tri< a ¬b ¬c = |
210 next tree (n ∷ st) (record {ti = findPR.ti Pre ; si = findPP2 st (findPR.si Pre) ; ci = lift tt} ) findPP1 where | |
621 | 211 tree0 = findPR.tree0 Pre |
627 | 212 findPP2 : (st : List (bt A)) → stackInvariant key {!!} tree0 st → stackInvariant key {!!} tree0 (node key₁ v tree tree₁ ∷ st) |
623 | 213 findPP2 = {!!} |
618 | 214 findPP1 : suc ( bt-depth tree ) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) |
215 findPP1 = {!!} | |
624 | 216 findPP key n@(node key₁ v tree tree₁) st Pre next exit | tri> ¬a ¬b c = next tree₁ (n ∷ st) {!!} findPP2 where -- Cond n st → Cond tree₁ (n ∷ st) |
618 | 217 findPP2 : suc (bt-depth tree₁) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁) |
218 findPP2 = {!!} | |
616 | 219 |
618 | 220 insertTreePP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree |
221 → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t | |
624 | 222 insertTreePP {n} {m} {A} {t} tree key value P exit = |
627 | 223 TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ t s → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ {!!} |
618 | 224 $ λ p P loop → findPP key (proj1 p) (proj2 p) {!!} (λ t s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) |
625 | 225 $ λ t s _ P → replaceNodeP key value t {!!} |
618 | 226 $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A )) |
627 | 227 {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant key (proj1 (proj2 p)) tree (proj1 p) ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) } |
618 | 228 (λ p → bt-depth (proj1 (proj2 p))) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ {!!} , ⟪ {!!} , R ⟫ ⟫ |
621 | 229 $ λ p P1 loop → replaceP key value (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p) {!!} |
230 (λ key value tree1 repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1 ⟫ ⟫ {!!} lt ) exit | |
618 | 231 |
628 | 232 record findP-contains {n : Level} {A : Set n} (key1 : ℕ) (value1 : A) (tree : bt A ) (stack : List (bt A)) : Set n where |
616 | 233 field |
234 tree1 : bt A | |
617 | 235 ci : replacedTree key1 value1 tree tree1 |
616 | 236 |
624 | 237 findPPC : {n m : Level} {A : Set n} {t : Set m} |
628 | 238 → (key : ℕ) → (value : A) → (tree : bt A ) → (stack : List (bt A)) |
239 → (Pre : findPR key tree stack (findP-contains key value)) | |
240 → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR key tree1 stack1 (findP-contains key value) → bt-depth tree1 < bt-depth tree → t ) | |
241 → (exit : (tree1 : bt A) → (stack1 : List (bt A)) → ( tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key) → findPR key tree1 stack1 (findP-contains key value) → t) → t | |
624 | 242 findPPC = {!!} |
243 | |
618 | 244 containsTree : {n m : Level} {A : Set n} {t : Set m} → (tree tree1 : bt A) → (key : ℕ) → (value : A) → treeInvariant tree1 → replacedTree key value tree1 tree → ⊤ |
615 | 245 containsTree {n} {m} {A} {t} tree tree1 key value P RT = |
617 | 246 TerminatingLoopS (bt A ∧ List (bt A) ) |
628 | 247 {λ p → findPR key (proj1 p) (proj2 p) (findP-contains key value ) } (λ p → bt-depth (proj1 p)) |
624 | 248 ⟪ tree1 , [] ⟫ {!!} |
628 | 249 $ λ p P loop → findPPC key value (proj1 p) (proj2 p) {!!} (λ t s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) |
250 $ λ t1 s1 found? P2 → insertTreeSpec0 t1 value (lemma7 {!!} (findPR.si P2 ) found? ) where | |
626 | 251 lemma7 : {key : ℕ } {value1 : A } {t1 tree : bt A } { s1 : List (bt A) } → |
627 | 252 replacedTree key value1 tree t1 → stackInvariant key t1 tree s1 → ( t1 ≡ leaf ) ∨ ( node-key t1 ≡ just key) → top-value t1 ≡ just value |
628 | 253 lemma7 = {!!} |
615 | 254 |