Mercurial > hg > Members > Moririn
view hoareBinaryTree.agda @ 612:57d6c594da08
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Fri, 05 Nov 2021 09:35:20 +0900 |
parents | db42d1033857 |
children | eeb9eb38e5e2 |
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 ) → (right : bt A ) → bt A bt-depth : {n : Level} {A : Set n} → (tree : bt A ) → ℕ bt-depth leaf = 0 bt-depth (node key value t t₁) = Data.Nat._⊔_ (bt-depth t ) (bt-depth t₁ ) find : {n m : Level} {A : Set n} {t : Set m} → (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 m : Level} {A : Set n} {t : Set m} → (key : ℕ) → bt A → List (bt A) → (exit : bt A → List (bt A) → t) → t find-loop {n} {m} {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 replaceNode : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → (bt A → t) → t replaceNode k v leaf next = next (node k v leaf leaf) replaceNode k v (node key value t t₁) next = next (node k v t t₁) 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 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 m : Level} {A : Set n} {t : Set m} → (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 m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → (next : bt A → t ) → t insertTree tree key value exit = find-loop key tree [] $ λ t st → replaceNode key value t $ λ t1 → replace-loop key value t1 st exit insertTest1 = insertTree leaf 1 1 (λ x → x ) insertTest2 = insertTree insertTest1 2 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} → (tree : bt A) → (stack : List (bt A)) → Set n stackInvariant {_} {A} _ [] = Lift _ ⊤ stackInvariant {_} {A} tree (tree1 ∷ [] ) = tree1 ≡ tree stackInvariant {_} {A} tree (x ∷ tail @ (node key value leaf right ∷ _) ) = (right ≡ x) ∧ stackInvariant tree tail stackInvariant {_} {A} tree (x ∷ tail @ (node key value left leaf ∷ _) ) = (left ≡ x) ∧ stackInvariant tree tail stackInvariant {_} {A} tree (x ∷ tail @ (node key value left right ∷ _ )) = ( (left ≡ x) ∧ stackInvariant tree tail) ∨ ( (right ≡ x) ∧ stackInvariant tree tail) stackInvariant {_} {A} tree s = Lift _ ⊥ rstackInvariant : {n : Level} {A : Set n} → (tree : bt A) → (stack : List (bt A)) → Set n rstackInvariant {_} {A} _ [] = Lift _ ⊤ rstackInvariant {_} {A} tree (tree1 ∷ [] ) = tree1 ≡ tree rstackInvariant {_} {A} tree (node key value leaf right ∷ x ∷ y ) = (right ≡ x) ∧ rstackInvariant tree (x ∷ y) rstackInvariant {_} {A} tree (node key value left leaf ∷ x ∷ y ) = (left ≡ x) ∧ rstackInvariant tree (x ∷ y) rstackInvariant {_} {A} tree (node key value left right ∷ x ∷ y ) = ( (left ≡ x) ∧ rstackInvariant tree (x ∷ y)) ∨ ( (right ≡ x) ∧ rstackInvariant tree (x ∷ y)) rstackInvariant {_} {A} tree s = Lift _ ⊥ findP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) → treeInvariant tree ∧ stackInvariant tree stack → (next : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree1 stack → bt-depth tree1 < bt-depth tree → t ) → (exit : (tree1 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree1 stack → t ) → t findP key leaf st Pre _ exit = exit leaf st {!!} findP key (node key₁ v tree tree₁) st Pre next exit with <-cmp key key₁ findP key n st Pre _ exit | tri≈ ¬a b ¬c = exit n st {!!} findP key n@(node key₁ v tree tree₁) st Pre next _ | tri< a ¬b ¬c = next tree (n ∷ st) {!!} {!!} findP key n@(node key₁ v tree tree₁) st Pre next _ | tri> ¬a ¬b c = next tree₁ (n ∷ st) {!!} {!!} replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) → (treeInvariant tree ) → ((tree : bt A) → treeInvariant tree → (rstack : List (bt A)) → rstackInvariant tree rstack → t) → t replaceNodeP k v leaf P next = next (node k v leaf leaf) {!!} {!!} {!!} replaceNodeP k v (node key value t t₁) P next = next (node k v t t₁) {!!} {!!} {!!} replaceP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree repl : bt A) → (stack rstack : List (bt A)) → treeInvariant tree ∧ stackInvariant tree stack ∧ rstackInvariant repl rstack → (next : ℕ → A → (tree1 : bt A) → (stack rstack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant tree1 stack ∧ rstackInvariant repl rstack → bt-depth tree1 < bt-depth tree → t ) → (exit : (tree1 repl : bt A) → (rstack : List (bt A)) → treeInvariant tree1 ∧ rstackInvariant repl rstack → t) → t replaceP key value tree repl [] rs Pre next exit = exit tree repl {!!} {!!} replaceP key value tree repl (leaf ∷ st) rs Pre next exit = next key value tree st {!!} {!!} {!!} replaceP key value tree repl (node key₁ value₁ left right ∷ st) rs Pre 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 {!!} {!!} {!!} open import Relation.Binary.Definitions nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x lemma3 : {i j : ℕ} → 0 ≡ i → j < i → ⊥ lemma3 refl () lemma5 : {i j : ℕ} → i < 1 → j < i → ⊥ lemma5 (s≤s z≤n) () TerminatingLoopS : {l m : Level} {t : Set l} (Index : Set m ) → {Invraiant : Index → Set m } → ( reduce : Index → ℕ) → (r : Index) → (p : Invraiant r) → (loop : (r : Index) → Invraiant r → (next : (r1 : Index) → Invraiant r1 → reduce r1 < reduce r → t ) → t) → t TerminatingLoopS {_} {_} {t} Index {Invraiant} reduce r p loop with <-cmp 0 (reduce r) ... | tri≈ ¬a b ¬c = loop r p (λ r1 p1 lt → ⊥-elim (lemma3 b lt) ) ... | tri< a ¬b ¬c = loop r p (λ r1 p1 lt1 → TerminatingLoop1 (reduce r) r r1 (≤-step lt1) p1 lt1 ) where TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → Invraiant r1 → reduce r1 < reduce r → t TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 p1 (λ r2 p1 lt1 → ⊥-elim (lemma5 n≤j lt1)) TerminatingLoop1 (suc j) r r1 n≤j p1 lt with <-cmp (reduce r1) (suc j) ... | tri< a ¬b ¬c = TerminatingLoop1 j r r1 a p1 lt ... | tri≈ ¬a b ¬c = loop r1 p1 (λ r2 p2 lt1 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b lt1 ) p2 lt1 ) ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) open _∧_ insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree → (exit : (tree repl : bt A) → (rstack : List (bt A)) → treeInvariant tree ∧ rstackInvariant repl rstack → t ) → t insertTreeP {n} {m} {A} {t} tree key value P exit = TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant (proj1 p) (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ ⟪ P , lift tt ⟫ $ λ p P loop → findP key (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) $ λ t s P → replaceNodeP key value t (proj1 P) $ λ t1 P1 r R → TerminatingLoopS (bt A ∧ List (bt A) ∧ List (bt A)) {λ p → treeInvariant (proj1 p) ∧ stackInvariant (proj1 p) (proj1 (proj2 p)) ∧ rstackInvariant t1 (proj2 (proj2 p))} (λ p → bt-depth (proj1 p)) ⟪ t , ⟪ s , r ⟫ ⟫ ⟪ proj1 P , ⟪ proj2 P , R ⟫ ⟫ $ λ p P1 loop → replaceP key value (proj1 p) t1 (proj1 (proj2 p)) (proj2 (proj2 p)) P1 (λ k v t s s1 P2 lt → loop ⟪ t , ⟪ s , s1 ⟫ ⟫ P2 lt ) exit top-value : {n : Level} {A : Set n} → (tree : bt A) → Maybe A top-value leaf = nothing top-value (node key value tree tree₁) = just value insertTreeSpec0 : {n : Level} {A : Set n} → (tree : bt A) → (value : A) → top-value tree ≡ just value → ⊤ insertTreeSpec0 _ _ _ = tt insertTreeSpec1 : {n : Level} {A : Set n} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree → ⊤ insertTreeSpec1 {n} {A} tree key value P = insertTreeP tree key value P (λ (tree₁ repl : bt A) (rstack : List (bt A)) (P1 : treeInvariant tree₁ ∧ rstackInvariant repl rstack ) → insertTreeSpec0 tree₁ value (lemma1 {!!} ) ) where lemma1 : (tree₁ : bt A) → top-value tree₁ ≡ just value lemma1 = {!!}