changeset 696:578f29a76857

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 03 Dec 2021 11:14:19 +0900
parents ce6cd128595d
children e5140faf1602
files hoareBinaryTree.agda
diffstat 1 files changed, 20 insertions(+), 16 deletions(-) [+]
line wrap: on
line diff
--- a/hoareBinaryTree.agda	Fri Dec 03 08:14:32 2021 +0900
+++ b/hoareBinaryTree.agda	Fri Dec 03 11:14:19 2021 +0900
@@ -88,7 +88,7 @@
 insertTest1 = insertTree leaf 1 1 (λ x → x )
 insertTest2 = insertTree insertTest1 2 1 (λ x → x )
 insertTest3 = insertTree insertTest2 3 2 (λ x → x )
-insertTest4 = insertTree insertTest3 2 2 (λ x → x )
+insertTest4 = insertTree insertTest3 2 2 (λ x → x ) -- this is wrong
 
 open import Data.Unit hiding ( _≟_ ;  _≤?_ ; _≤_)
 
@@ -534,21 +534,22 @@
 RTtoTI1  = {!!}
 
 insertTreeP : {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
+     → (exit : (tree repl : bt A) → treeInvariant repl ∧ replacedTree key value tree repl → t ) → t
 insertTreeP {n} {m} {A} {t} tree key value P0 exit =
    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-single ⟫
-       $ λ p P loop → findP key (proj1 p)  tree (proj2 p) P (λ t s P1 lt → loop ⟪ t ,  s  ⟫ P1 lt ) -- treeInvariant t ∧ stackInvariant key t tree s
+       $ λ p P loop → findP key (proj1 p)  tree (proj2 p) P (λ t s P1 lt → loop ⟪ t ,  s  ⟫ P1 lt ) 
        $ λ t s P C → replaceNodeP key value t C (proj1 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 = proj2 P ; ri = R ; ci = lift tt } -- replacedTree key value (child-replaced key t) t1
+               (λ p → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ record { tree0 = tree ; ti = P0 ; si = proj2 P ; 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 
+            (λ key value {tree1} repl1 stack P2 lt → loop ⟪ stack , ⟪ tree1 , repl1  ⟫ ⟫ P2 lt )
+       $ λ tree repl P →  exit tree repl ⟪ RTtoTI0 _ _ _ _ (proj1 P) (proj2 P) , proj2 P ⟫   
 
-insertTestP1 = insertTreeP leaf 1 1 _ (λ _ x _ → x )
-insertTestP2 = insertTreeP insertTestP1 2 1 _ (λ _ x _  → x )
-insertTestP3 = insertTreeP insertTestP2 3 2 _ (λ _ x _  → x )
-insertTestP4 = insertTreeP insertTestP3 2 2 _ (λ _ x _  → x )
+insertTestP1 = insertTreeP leaf 1 1 t-leaf
+  $ λ _ x P → insertTreeP x 2 1 (proj1 P) 
+  $ λ _ x P → insertTreeP x 3 2 (proj1 P)  
+  $ λ _ x P → insertTreeP x 2 2 (proj1 P) (λ _ x _  → x )
 
 top-value : {n : Level} {A : Set n} → (tree : bt A) →  Maybe A 
 top-value leaf = nothing
@@ -560,6 +561,7 @@
 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
+     ti0 : treeInvariant tree0
      ti : treeInvariant tree
      si : stackInvariant key tree tree0 stack
      ci : C tree stack     -- data continuation
@@ -573,23 +575,23 @@
 findPP key (node key₁ v1 tree tree₁) st Pre next exit with <-cmp key key₁
 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
+       record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 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<
+       record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 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 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 }
+           record { tree0 = tree ; ti = P0 ; ti0 = 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 → length (proj1 p)) ⟪ s , ⟪ t , t1 ⟫ ⟫ record { tree0 = findPR.tree0 P ; ti = findPR.ti0 P ; si = findPR.si P ; 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 
 
@@ -608,18 +610,20 @@
 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
+       record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 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<
+       record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 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 ) } (λ p → bt-depth (proj1 p)) -- findPR key tree1 [] (findPC key value)
-              ⟪ tree1 , []  ⟫ record { tree0 = tree ; ti = {!!} ; si = {!!} ; ci = record { tree1 = tree ; ci = RT } }
+              ⟪ tree1 , []  ⟫ record { tree0 = tree ; ti0 = {!!} ; ti = {!!} ; si = {!!} ; ci = record { tree1 = tree ; ci = RT } }
        $ λ 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