changeset 695:ce6cd128595d

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 03 Dec 2021 08:14:32 +0900
parents da42fe4eda54
children 578f29a76857
files hoareBinaryTree.agda
diffstat 1 files changed, 49 insertions(+), 46 deletions(-) [+]
line wrap: on
line diff
--- a/hoareBinaryTree.agda	Thu Dec 02 15:03:21 2021 +0900
+++ b/hoareBinaryTree.agda	Fri Dec 03 08:14:32 2021 +0900
@@ -284,7 +284,7 @@
     → (treeInvariant tree ) → ((tree1 : bt A) → treeInvariant tree1 →  replacedTree key value (child-replaced key tree) tree1 → t) → t
 replaceNodeP k v1 leaf C P next = next (node k v1 leaf leaf) (t-single k v1 ) r-leaf
 replaceNodeP k v1 (node .k value t t₁) (case2 refl) P next = next (node k v1 t t₁) (replaceTree1 k value v1 P)
-     (subst (λ j → replacedTree k v1 j  (node k v1 t t₁) ) repl00 r-node) where -- (child-replaced k (node k value t t₁))
+     (subst (λ j → replacedTree k v1 j  (node k v1 t t₁) ) repl00 r-node) where 
          repl00 : node k value t t₁ ≡ child-replaced k (node k value t t₁)
          repl00 with <-cmp k k
          ... | tri< a ¬b ¬c = ⊥-elim (¬b refl)
@@ -560,67 +560,70 @@
 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
    field
      tree0 : bt A
-     ti : treeInvariant tree0
+     ti : treeInvariant tree
      si : stackInvariant key tree tree0 stack
      ci : C tree stack     -- data continuation
    
-findPP : {n m : Level} {A : Set n} {t : Set m}
-           → (key : ℕ) → (tree : bt A ) → (stack : List (bt A))
-           → (Pre :  findPR key tree stack (λ t s → Lift n ⊤))
-           → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR key tree1 stack1 (λ t s → Lift n ⊤) →  bt-depth tree1 < bt-depth tree   → t )
-           → (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
-findPP key leaf st Pre next exit = exit leaf st (case1 refl) Pre  
+findPP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A))
+           →  findPR key tree stack (λ _ _ → Lift n ⊤)
+           → (next : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (λ _ _ → Lift n ⊤) → bt-depth tree1 < bt-depth tree   → t )
+           → (exit : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (λ _ _ → Lift n ⊤)
+                 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key )  → t ) → t
+findPP key leaf st Pre _ exit = exit leaf st Pre (case1 refl)
 findPP key (node key₁ v1 tree tree₁) st Pre next exit with <-cmp key key₁
-findPP key n st P next exit | tri≈ ¬a b ¬c = exit n st (case2 {!!}) P 
-findPP {_} {_} {A} key n@(node key₁ v1 tree tree₁) st Pre next exit | tri< a ¬b ¬c =
-          next tree (n ∷ st) (record {ti = findPR.ti Pre  ; si = findPP2 st (findPR.si Pre) ; ci = lift tt} ) findPP1 where 
-    tree0 =  findPR.tree0 Pre 
-    findPP2 : (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) tree0 st →  stackInvariant key tree tree0 (node key₁ v1 tree tree₁ ∷ st)
-    findPP2 = {!!}
-    findPP1 : suc ( bt-depth tree ) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁)
-    findPP1 =  depth-1<
-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)
-    findPP2 : suc (bt-depth tree₁) ≤ suc (bt-depth tree Data.Nat.⊔ bt-depth tree₁)
-    findPP2 = depth-2<
+findPP key n st Pre _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl)
+findPP {n} {_} {A} key (node key₁ v1 tree tree₁) st  Pre next _ | tri< a ¬b ¬c = next tree (tree ∷ st)
+       record { tree0 = findPR.tree0 Pre ; ti = treeLeftDown tree tree₁ (findPR.ti Pre)  ; si =  findP1 a st (findPR.si Pre) ; ci = lift tt } depth-1< where
+   findP1 : key < key₁ → (st : List (bt A)) →  stackInvariant key (node key₁ v1 tree tree₁)
+         (findPR.tree0 Pre) st → stackInvariant key tree (findPR.tree0 Pre) (tree ∷ st)
+   findP1 a (x ∷ st) si = s-left a si
+findPP key n@(node key₁ v1 tree tree₁) st Pre next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st)
+       record { tree0 = findPR.tree0 Pre ; ti = treeRightDown tree tree₁ (findPR.ti Pre) ; si =  s-right c (findPR.si Pre) ; ci = lift tt }  depth-2<
 
 insertTreePP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree
      → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t
-insertTreePP {n} {m} {A} {t} tree key value  P exit =
-   TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ t s → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫  {!!}
-       $ λ p P loop → findPP key (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t ,  s  ⟫ P1 lt )
-       $ λ t s _ P → replaceNodeP key value t {!!} {!!}
-       $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ (bt A ∧ bt A ))
-            {λ p → treeInvariant (proj1 (proj2 p)) ∧ stackInvariant key (proj1 (proj2 p)) tree (proj1 p)  ∧ replacedTree key value (proj1 (proj2 p)) (proj2 (proj2 p)) }
-               (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ ⟪ {!!} , ⟪ {!!}  , {!!} ⟫ ⟫
-       $  λ p P1 loop → replaceP key value  (proj2 (proj2 p)) (proj1 p) {!!}
-            (λ key value repl1 stack P2 lt → loop ⟪ stack , ⟪ {!!} , repl1  ⟫ ⟫ {!!} lt )  exit 
+insertTreePP {n} {m} {A} {t} tree key value P0 exit =
+   TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ _ _ → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , tree ∷ [] ⟫
+           record { tree0 = tree ; ti = P0 ; si = s-single ; ci = lift tt }
+       $ λ p P loop → findPP key (proj1 p)  (proj2 p) P (λ t s P1 lt → loop ⟪ t ,  s  ⟫ P1 lt ) 
+       $ λ t s P C → replaceNodeP key value t C (findPR.ti P)
+       $ λ t1 P1 R → TerminatingLoopS (List (bt A) ∧ bt A ∧ bt A )
+            {λ p → replacePR key value  (proj1 (proj2 p)) (proj2 (proj2 p)) (proj1 p)  (λ _ _ _ → Lift n ⊤ ) }
+               (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ record { tree0 = tree ; ti = P0 ; si = {!!} ; ri = R ; ci = lift tt } 
+       $  λ p P1 loop → replaceP key value  (proj2 (proj2 p)) (proj1 p) P1
+            (λ key value {tree1} repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1  ⟫ ⟫ P2 lt )  exit 
 
-record findPC {n : Level} {A : Set n} (key1 : ℕ) (value1 : A) (tree : bt A ) (stack : List (bt A)) : Set n where
+record findPC {n : Level} {A : Set n} (key1 : ℕ)  (tree : bt A ) (stack : List (bt A)) : Set n where
    field
      tree1 : bt A
-     ci : replacedTree key1 value1 tree tree1
+     value : A
+     ci : replacedTree key1 value tree tree1
    
-findPPC : {n m : Level} {A : Set n} {t : Set m}
-           → (key : ℕ) → (value : A) → (tree : bt A ) → (stack : List (bt A))
-           → (Pre :  findPR key tree stack (findPC key value))
-           → (next : (tree1 : bt A) → (stack1 : List (bt A)) → findPR key tree1 stack1 (findPC key value) →  bt-depth tree1 < bt-depth tree   → t )
-           → (exit : (tree1 : bt A) → (stack1 : List (bt A)) → ( tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key)  → findPR key tree1 stack1 (findPC key value) → t) → t
-findPPC key value leaf st Pre next exit = exit leaf st (case1 refl) Pre  
-findPPC key value (node key₁ v1 tree tree₁) st Pre next exit with <-cmp key key₁
-findPPC key value n st P next exit | tri≈ ¬a b ¬c = exit n st (case2 {!!}) P 
-findPPC {_} {_} {A} key value n@(node key₁ v1 tree tree₁) st Pre next exit | tri< a ¬b ¬c =
-          next tree (n ∷ st) (record {ti = findPR.ti Pre  ; si = {!!} ; ci =  {!!} } ) {!!} 
-findPPC key value n st P next exit | tri> ¬a ¬b c = {!!}
+findPPC : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A))
+           →  findPR key tree stack (findPC key )
+           → (next : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC key ) → bt-depth tree1 < bt-depth tree   → t )
+           → (exit : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC key )
+                 → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key )  → t ) → t
+findPPC key leaf st Pre _ exit = exit leaf st Pre (case1 refl)
+findPPC key (node key₁ v1 tree tree₁) st Pre next exit with <-cmp key key₁
+findPPC key n st Pre _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl)
+findPPC {n} {_} {A} key (node key₁ v1 tree tree₁) st  Pre next _ | tri< a ¬b ¬c = next tree (tree ∷ st)
+       record { tree0 = findPR.tree0 Pre ; ti = treeLeftDown tree tree₁ (findPR.ti Pre)  ; si =  findP1 a st (findPR.si Pre) ; ci = {!!} } depth-1< where
+   findP1 : key < key₁ → (st : List (bt A)) →  stackInvariant key (node key₁ v1 tree tree₁)
+         (findPR.tree0 Pre) st → stackInvariant key tree (findPR.tree0 Pre) (tree ∷ st)
+   findP1 a (x ∷ st) si = s-left a si
+findPPC key n@(node key₁ v1 tree tree₁) st Pre next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st)
+       record { tree0 = findPR.tree0 Pre ; ti = treeRightDown tree tree₁ (findPR.ti Pre) ; si =  s-right c (findPR.si Pre) ; ci = {!!} }  depth-2<
 
 containsTree : {n m : Level} {A : Set n} {t : Set m} → (tree tree1 : bt A) → (key : ℕ) → (value : A) → treeInvariant tree1 → replacedTree key value tree1 tree  → ⊤
 containsTree {n} {m} {A} {t} tree tree1 key value P RT =
    TerminatingLoopS (bt A ∧ List (bt A) )
-     {λ p → findPR key (proj1 p) (proj2 p) (findPC key value ) } (λ p → bt-depth (proj1 p)) -- findPR key tree1 [] (findPC key value)
+     {λ p → findPR key (proj1 p) (proj2 p) (findPC key ) } (λ p → bt-depth (proj1 p)) -- findPR key tree1 [] (findPC key value)
               ⟪ tree1 , []  ⟫ record { tree0 = tree ; ti = {!!} ; si = {!!} ; ci = record { tree1 = tree ; ci = RT } }
-       $ λ p P loop → findPPC key value (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt )  
-       $ λ t1 s1 found? P2 → insertTreeSpec0 t1 value (lemma6 t1 s1 found? P2) where
-           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
-           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
+       $ λ p P loop → findPPC key (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt )  
+       $ λ t1 s1 P2 found? → insertTreeSpec0 t1 value (lemma6 t1 s1 found? P2) where
+           lemma6 : (t1 : bt A) (s1 : List (bt A)) (found? : (t1 ≡ leaf) ∨ (node-key t1 ≡ just key)) (P2 : findPR key t1 s1 (findPC key )) → top-value t1 ≡ just value
+           lemma6 t1 s1 found? P2 = lemma7 t1 s1 (findPR.tree0 P2) ( findPC.tree1  (findPR.ci P2)) ( findPC.ci  {!!} ) (findPR.si P2) found? where
               lemma7 :  (t1 : bt A) ( s1 : List (bt A) ) (tree0 tree1 : bt A) →
                  replacedTree key value t1 tree1 → stackInvariant key t1 tree0 s1  → ( t1 ≡ leaf ) ∨ ( node-key t1 ≡ just key)  →   top-value t1 ≡ just value
               lemma7 = {!!}