view hoareBinaryTree.agda @ 608:8df36383ced0

fix RedBlackTree
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Wed, 03 Nov 2021 18:28:31 +0900
parents 61a0491a627b
children 79418701a283
line wrap: on
line source

module hoareBinaryTree where

open import Level renaming (zero to Z ; suc to succ)

open import Data.Nat hiding (compare)
open import Data.Nat.Properties as NatProp
open import Data.Maybe
-- open import Data.Maybe.Properties
open import Data.Empty
open import Data.List
open import Data.Product

open import Function as F hiding (const)

open import Relation.Binary
open import Relation.Binary.PropositionalEquality
open import Relation.Nullary
open import logic


_iso_ : {n : Level} {a : Set n} → ℕ → ℕ → Set
d iso d' = (¬ (suc d ≤ d')) ∧ (¬ (suc d' ≤ d))

iso-intro : {n : Level} {a : Set n} {x y : ℕ} → ¬ (suc x ≤ y) → ¬ (suc y ≤ x) → _iso_ {n} {a} x y
iso-intro = λ z z₁ → record { proj1 = z ; proj2 = z₁ }

--
--
--  no children , having left node , having right node , having both
--
data bt {n : Level} (A : Set n) : Set n where
  leaf : bt A
  node :  (key : ℕ) → (value : A) →
    (left : bt A ) → (write : bt A ) → bt A

bt-length : {n : Level} {A : Set n} → (tree : bt A ) → ℕ
bt-length leaf = 0
bt-length (node key value t t₁) = Data.Nat._⊔_ (bt-length t ) (bt-length t₁ )

find : {n : Level} {A : Set n} {t : Set n} → (key : ℕ) → (tree : bt A ) → List (bt A)
           → (next : bt A → List (bt A) → t ) → (exit : bt A → List (bt A) → t ) → t
find key leaf st _ exit = exit leaf st
find key (node key₁ v tree tree₁) st next exit with <-cmp key key₁
find key n st _ exit | tri≈ ¬a b ¬c = exit n st
find key n@(node key₁ v tree tree₁) st next _ | tri< a ¬b ¬c = next tree (n ∷ st)
find key n@(node key₁ v tree tree₁) st next _ | tri> ¬a ¬b c = next tree₁ (n ∷ st)

{-# TERMINATING #-}
find-loop : {n : Level} {A : Set n} {t : Set n} → (key : ℕ) → bt A → List (bt A)  → (exit : bt A → List (bt A) → t) → t
find-loop {_} {A} {t} key tree st exit = find-loop1 tree st where
    find-loop1 : bt A → List (bt A) → t
    find-loop1 tree st = find key tree st find-loop1  exit

replace : {n : Level} {A : Set n} {t : Set n} → (key : ℕ) → (value : A) → bt A → List (bt A) → (next : ℕ → A → bt A → List (bt A) → t ) → (exit : bt A → t) → t
replace key value tree [] next exit = exit tree
replace key value tree (leaf ∷ st) next exit = next key value tree st 
replace key value tree (node key₁ value₁ left right ∷ st) next exit with <-cmp key key₁
... | tri< a ¬b ¬c = next key value (node key₁ value₁ tree right ) st
... | tri≈ ¬a b ¬c = next key value (node key₁ value  left right ) st
... | tri> ¬a ¬b c = next key value (node key₁ value₁ left tree ) st

{-# TERMINATING #-}
replace-loop : {n : Level} {A : Set n} {t : Set n} → (key : ℕ) → (value : A) → bt A → List (bt A)  → (exit : bt A → t) → t
replace-loop {_} {A} {t} key value tree st exit = replace-loop1 key value tree st where
    replace-loop1 : (key : ℕ) → (value : A) → bt A → List (bt A) → t
    replace-loop1 key value tree st = replace key value tree st replace-loop1  exit

insertTree : {n : Level} {A : Set n} {t : Set n} → (tree : bt A) → (key : ℕ) → (value : A) → (next : bt A → t ) → t
insertTree tree key value exit = find-loop key tree [] ( λ t st → replace-loop key value t st exit ) 

insertTest1 = insertTree leaf 1 1 (λ x → x )

open import Data.Unit hiding ( _≟_ ;  _≤?_ ; _≤_)

treeInvariant : {n : Level} {A : Set n} → (tree : bt A) → Set
treeInvariant leaf = ⊤
treeInvariant (node key value leaf leaf) = ⊤
treeInvariant (node key value leaf n@(node key₁ value₁ t₁ t₂)) =  (key < key₁) ∧ treeInvariant n 
treeInvariant (node key value n@(node key₁ value₁ t t₁) leaf) =  treeInvariant n ∧ (key < key₁) 
treeInvariant (node key value n@(node key₁ value₁ t t₁) m@(node key₂ value₂ t₂ t₃)) = treeInvariant n ∧  (key < key₁) ∧ (key₁ < key₂) ∧ treeInvariant m

treeInvariantTest1  = treeInvariant (node 3 0 leaf (node 1 1 leaf (node 3 5 leaf leaf)))

stackInvariant : {n : Level} {A : Set n} → (stack  : List (bt A)) → Set n
stackInvariant {_} {A} [] = Lift _ ⊤
stackInvariant {_} {A} (leaf ∷ stack) = Lift _ ⊤
stackInvariant {_} {A} (node key value leaf leaf ∷ []) = Lift _ ⊤
stackInvariant {_} {A} (node key value _ (node _ _ _ _) ∷ []) = Lift _ ⊥
stackInvariant {_} {A} (node key value (node _ _ _ _) _ ∷ []) = Lift _ ⊥
stackInvariant {_} {A} (x ∷ node key value leaf leaf ∷ tail ) = Lift _ ⊥
stackInvariant {_} {A} (x ∷ tail @ (node key value leaf tree ∷ _) ) = (tree ≡ x) ∧ stackInvariant tail
stackInvariant {_} {A} (x ∷ tail @ (node key value tree leaf ∷ _) ) = (tree ≡ x) ∧ stackInvariant tail
stackInvariant {_} {A} (x ∷ tail @ (node key value left right ∷ _  )) = ( (left ≡ x) ∧ stackInvariant tail) ∨ ( (right ≡ x) ∧ stackInvariant tail)
stackInvariant {_} {A} s = Lift _ ⊥

findP : {n : Level} {A : Set n} {t : Set n} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A))
           →  treeInvariant tree  → stackInvariant stack  
           → (next : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant stack → bt-length tree1 < bt-length tree   → t )
           → (exit : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant stack → t ) → t
findP = {!!}

replaceP : {n : Level} {A : Set n} {t : Set n}
     → (key : ℕ) → (value : A) → (tree : bt A) → (stack : List (bt A)) → treeInvariant tree ∧ stackInvariant stack 
     → (next : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant stack → bt-length tree1 < bt-length tree   → t )
     → (exit : (tree1 : bt A) → treeInvariant tree1 → t) → t
replaceP = {!!}