changeset 698:28e0f7f4777d

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 04 Dec 2021 19:47:07 +0900
parents e5140faf1602
children 59f80c1049e9
files hoareBinaryTree.agda
diffstat 1 files changed, 51 insertions(+), 37 deletions(-) [+]
line wrap: on
line diff
--- a/hoareBinaryTree.agda	Sat Dec 04 14:47:03 2021 +0900
+++ b/hoareBinaryTree.agda	Sat Dec 04 19:47:07 2021 +0900
@@ -197,12 +197,23 @@
 rt-property-leaf : {n : Level} {A : Set n} {key : ℕ} {value : A} {repl : bt A} → replacedTree key value leaf repl → repl ≡ node key value leaf leaf
 rt-property-leaf r-leaf = refl
 
+rt-property-¬leaf : {n : Level} {A : Set n} {key : ℕ} {value : A} {tree : bt A} → ¬ replacedTree key value tree leaf 
+rt-property-¬leaf ()
+
 rt-property-key : {n : Level} {A : Set n} {key key₂ key₃ : ℕ} {value value₂ value₃ : A} {left left₁ right₂ right₃ : bt A}
     →  replacedTree key value (node key₂ value₂ left right₂) (node key₃ value₃ left₁ right₃) → key₂ ≡ key₃
 rt-property-key r-node = refl
 rt-property-key (r-right x ri) = refl
 rt-property-key (r-left x ri) = refl
 
+nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥
+nat-≤>  (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x
+nat-<> : { x y : ℕ } → x < y → y < x → ⊥
+nat-<>  (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x
+
+open _∧_
+
+
 depth-1< : {i j : ℕ} →   suc i ≤ suc (i Data.Nat.⊔ j )
 depth-1< {i} {j} = s≤s (m≤m⊔n _ j)
 
@@ -230,13 +241,6 @@
 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = t-leaf
 treeRightDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti₁
 
-nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥
-nat-≤>  (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x
-nat-<> : { x y : ℕ } → x < y → y < x → ⊥
-nat-<>  (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x
-
-open _∧_
-
 findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt A ) → (stack : List (bt A))
            →  treeInvariant tree ∧ stackInvariant key tree tree0 stack  
            → (next : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree   → t )
@@ -595,45 +599,55 @@
        $  λ 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 : ℕ)  (tree : bt A ) (stack : List (bt A)) : Set n where
+record findPC {n : Level} {A : Set n} (key1 : ℕ) (value : A) (tree : bt A ) (stack : List (bt A)) : Set n where
    field
      tree1 : bt A
-     value : A
-     ci : replacedTree key1 value tree tree1
+     ci : replacedTree key1 value tree1 tree
    
-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 )
+findPPC : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A ) → (stack : List (bt A))
+           →  findPR key tree stack (findPC key value )
+           → (next : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC key value ) → bt-depth tree1 < bt-depth tree   → t )
+           → (exit : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC key value )
                  → (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)
+findPPC key value leaf st Pre _ exit = exit leaf st Pre (case1 refl)
+findPPC key value (node key₁ v1 tree tree₁) st Pre next exit with <-cmp key key₁
+findPPC key value n st Pre _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl)
+findPPC {n} {_} {A} key value (node key₁ v1 tree tree₁) st  Pre next _ | tri< a ¬b ¬c = next tree (tree ∷ st)
        record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 Pre  ; ti = treeLeftDown tree tree₁ (findPR.ti Pre)  ; si =  s-left a (findPR.si Pre)
-          ; ci = {!!} } depth-1< where
-   findP2 : findPC key tree (tree ∷ st)
-   findP2 with findPC.ci (findPR.ci Pre) | findPC.tree1 (findPR.ci Pre)
-   findP2 | r-node | leaf = {!!}
-   findP2 | r-node | node key value t t₁ = record { tree1 = t ; value = findPC.value (findPR.ci Pre)  ; ci =  {!!} }
-   findP2 | (r-right x ri) | t = ⊥-elim (nat-<> x a)
-   findP2 | (r-left x ri) | node key value t t₁ = record { tree1 = t ; value = findPC.value (findPR.ci Pre)  ; ci =  {!!} }
-   findP2 | r-left x ri | leaf = {!!}
-   -- findP2 (r-left x ri) = subst₂ (λ j k →  replacedTree key (findPC.value (findPR.ci Pre)) j k ) {!!} {!!} ri 
-findPPC key n@(node key₁ v1 tree tree₁) st Pre next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st)
+          ; ci = findP2 } depth-1< where
+   findP2 : findPC key value tree (tree ∷ st)
+   findP2 with findPC.ci (findPR.ci Pre) | findPC.tree1 (findPR.ci Pre) | inspect  findPC.tree1 (findPR.ci Pre) 
+   findP2 | r-node | leaf | _ = ⊥-elim ( nat-≤> a ≤-refl )
+   findP2 | r-node | node key value t t₁ | _ = ⊥-elim ( nat-≤> a ≤-refl )
+   findP2 | (r-right x ri) | t | _ = ⊥-elim (nat-<> x a)
+   findP2 | (r-left x ri) | node key value t t₁ | record { eq = refl } = record { tree1 = t ; ci =  ri }
+   findP2 | r-left x ri | leaf | record { eq = () }
+   findP2 | r-leaf | leaf | record { eq = eq } = ⊥-elim ( nat-≤> a ≤-refl )
+findPPC key value n@(node key₁ v1 tree tree₁) st Pre next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st)
        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<
+          ; ci = findP2 }  depth-2< where
+   findP2 : findPC key value tree₁ (tree₁ ∷ st)
+   findP2 with findPC.ci (findPR.ci Pre) | findPC.tree1 (findPR.ci Pre) | inspect  findPC.tree1 (findPR.ci Pre) 
+   findP2 | r-node | node key value ti ti₁ | eq = ⊥-elim ( nat-≤> c ≤-refl )
+   findP2 | r-left x ri | ti | eq = ⊥-elim ( nat-<> x c )
+   findP2 | r-right x ri | node key value t t₁ | record { eq = refl } = record { tree1 = t₁ ; ci =  ri }
+   findP2 | r-leaf | leaf | record { eq = eq } = ⊥-elim ( nat-≤> c ≤-refl )
 
 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 ; 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 )  
+     {λ p → findPR key (proj1 p) (proj2 p) (findPC key value ) } (λ p → bt-depth (proj1 p)) 
+              ⟪ tree , tree ∷ []  ⟫ record { tree0 = tree ; ti0 = RTtoTI0 _ _ _ _ P RT ; ti = RTtoTI0 _ _ _ _ P RT ; si = s-single
+                    ; ci = record { tree1 = tree1 ; ci = RT } }
+       $ λ p P loop → findPPC key value (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 = {!!}
+           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
+              lemma7 :  (t1 : bt A) ( s1 : List (bt A) ) (tree0 tree1 : bt A)
+                 → replacedTree key  value tree1 t1 → stackInvariant key t1 tree0 s1
+                 → ( t1 ≡ leaf ) ∨ ( node-key t1 ≡ just key)  →   top-value t1 ≡ just value
+              lemma7 .leaf (.leaf ∷ []) .leaf tree1 () s-single (case1 refl)
+              lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ []) .(node key value t1 t2) tree1 ri s-single (case2 x) = {!!}
+              lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ x₁ ∷ s1) tree0 tree1 ri (s-right x si) found =  ?
+              lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ x₁ ∷ s1) tree0 tree1 ri (s-left x si) found =  ?