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 node-key : {n : Level} {A : Set n} → bt A → Maybe ℕ node-key (node key _ _ _) = just key node-key _ = nothing node-value : {n : Level} {A : Set n} → bt A → Maybe A node-value (node _ value _ _) = just value node-value _ = nothing bt-depth : {n : Level} {A : Set n} → (tree : bt A ) → ℕ bt-depth leaf = 0 bt-depth (node key value t t₁) = suc (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₁ v1 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₁ v1 tree tree₁) st next _ | tri< a ¬b ¬c = next tree (n ∷ st) find key n@(node key₁ v1 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 v1 leaf next = next (node k v1 leaf leaf) replaceNode k v1 (node key value t t₁) next = next (node k v1 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 ∷ []) next exit = exit (node key value leaf leaf) replace key value tree (leaf ∷ leaf ∷ st) next exit = exit (node key value leaf leaf) replace key value tree (leaf ∷ node key₁ value₁ left right ∷ st) next exit with <-cmp key key₁ ... | tri< a ¬b ¬c = next key value (node key₁ value₁ (node key value leaf leaf) 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 (node key value leaf leaf) ) 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 ( _≟_ ; _≤?_ ; _≤_) data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where t-leaf : treeInvariant leaf t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key₁ value₁ t₁ t₂) → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key₁ < key) → treeInvariant (node key value t₁ t₂) → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → (key < key₁) → (key₁ < key₂) → treeInvariant (node key value t₁ t₂) → treeInvariant (node key₂ value₂ t₃ t₄) → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) data stackInvariant {n : Level} {A : Set n} (key : ℕ) : (tree tree0 : bt A) → (stack : List (bt A)) → Set n where s-single : (tree : bt A) → stackInvariant key tree tree (tree ∷ [] ) s-right : {tree0 tree tree₁ : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} → key₁ < key → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree₁ tree0 (tree₁ ∷ st) s-left : {tree0 tree tree₁ : bt A} → {key₁ : ℕ } → {v1 : A } → {st : List (bt A)} → key < key₁ → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (tree ∷ st) data replacedTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (tree tree1 : bt A ) → Set n where r-leaf : replacedTree key value leaf (node key value leaf leaf) r-node : {value₁ : A} → {t t₁ : bt A} → replacedTree key value (node key value₁ t t₁) (node key value t t₁) r-right : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} → k > key → replacedTree key value t1 t2 → replacedTree key value (node k v1 t t1) (node k v1 t t2) r-left : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} → k < key → replacedTree key value t1 t2 → replacedTree key value (node k v1 t1 t) (node k v1 t2 t) add< : { i : ℕ } (j : ℕ ) → i < suc i + j add< {i} j = begin suc i ≤⟨ m≤m+n (suc i) j ⟩ suc i + j ∎ where open ≤-Reasoning treeTest1 : bt ℕ treeTest1 = node 1 0 leaf (node 3 1 (node 2 5 (node 4 7 leaf leaf ) leaf) (node 5 5 leaf leaf)) treeTest2 : bt ℕ treeTest2 = node 3 1 (node 2 5 (node 4 7 leaf leaf ) leaf) (node 5 5 leaf leaf) treeInvariantTest1 : treeInvariant treeTest1 treeInvariantTest1 = t-right (m≤m+n _ 1) (t-node (add< 0) (add< 1) (t-left (add< 1) (t-single 4 7)) (t-single 5 5) ) stack-top : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) stack-top [] = nothing stack-top (x ∷ s) = just x stack-last : {n : Level} {A : Set n} (stack : List (bt A)) → Maybe (bt A) stack-last [] = nothing stack-last (x ∷ []) = just x stack-last (x ∷ s) = stack-last s stackInvariantTest1 : stackInvariant 2 treeTest2 treeTest1 ( treeTest2 ∷ treeTest1 ∷ [] ) stackInvariantTest1 = s-right (add< 0) (s-single treeTest1 ) si-property1 : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack → stack-top stack ≡ just tree si-property1 key t t0 (x ∷ .[]) (s-single .x) = refl si-property1 key t t0 (t ∷ st) (s-right _ si) = refl si-property1 key t t0 (t ∷ st) (s-left _ si) = refl si-property-last : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack → stack-last stack ≡ just tree0 si-property-last key t t0 (x ∷ []) (s-single .x) = refl si-property-last key t t0 (.t ∷ x ∷ st) (s-right _ si) with si-property1 key _ _ (x ∷ st) si ... | refl = si-property-last key x t0 (x ∷ st) si si-property-last key t t0 (.t ∷ x ∷ st) (s-left _ si) with si-property1 key _ _ (x ∷ st) si ... | refl = si-property-last key x t0 (x ∷ st) si ti-right : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 tree₁ repl) → treeInvariant repl ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-right x ti) = ti ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-left x ti) = t-leaf ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti₁ ti-left : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 repl tree₁ ) → treeInvariant repl ti-left {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf ti-left {_} {_} {_} {_} {key₁} {v1} (t-right x ti) = t-leaf ti-left {_} {_} {_} {_} {key₁} {v1} (t-left x ti) = ti ti-left {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti stackTreeInvariant : {n : Level} {A : Set n} (key : ℕ) (repl tree : bt A) → (stack : List (bt A)) → treeInvariant tree → stackInvariant key repl tree stack → treeInvariant repl stackTreeInvariant key repl .repl .(repl ∷ []) ti (s-single .repl) = ti stackTreeInvariant {_} {A} key repl tree (repl ∷ st) ti (s-right _ si) = ti-right (si1 si) where si1 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 tree₁ repl) tree st → treeInvariant (node key₁ v1 tree₁ repl) si1 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 tree₁ repl) tree st ti si stackTreeInvariant {_} {A} key repl tree (repl ∷ st) ti (s-left _ si) = ti-left ( si2 si ) where si2 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 repl tree₁ ) tree st → treeInvariant (node key₁ v1 repl tree₁ ) si2 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 repl tree₁ ) tree st ti si rt-property1 : {n : Level} {A : Set n} (key : ℕ) (value : A) (tree tree1 : bt A ) → replacedTree key value tree tree1 → ¬ ( tree1 ≡ leaf ) rt-property1 {n} {A} key value .leaf .(node key value leaf leaf) r-leaf () rt-property1 {n} {A} key value .(node key _ _ _) .(node key value _ _) r-node () rt-property1 {n} {A} key value .(node _ _ _ _) .(node _ _ _ _) (r-right x rt) () rt-property1 {n} {A} key value .(node _ _ _ _) .(node _ _ _ _) (r-left x rt) () depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) depth-1< {i} {j} = s≤s (m≤m⊔n _ j) depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) depth-2< {i} {j} = s≤s (m≤n⊔m _ i) treeLeftDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) → treeInvariant (node k v1 tree tree₁) → treeInvariant tree treeLeftDown {n} {A} {_} {v1} leaf leaf (t-single k1 v1) = t-leaf treeLeftDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = t-leaf treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .leaf (t-left x ti) = ti treeLeftDown {n} {A} {_} {v1} .(node _ _ _ _) .(node _ _ _ _) (t-node x x₁ ti ti₁) = ti treeRightDown : {n : Level} {A : Set n} {k : ℕ} {v1 : A} → (tree tree₁ : bt A ) → treeInvariant (node k v1 tree tree₁) → treeInvariant tree₁ treeRightDown {n} {A} {_} {v1} .leaf .leaf (t-single _ .v1) = t-leaf treeRightDown {n} {A} {_} {v1} .leaf .(node _ _ _ _) (t-right x ti) = ti 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₁ 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 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) → (exit : (tree1 tree0 : bt A) → (stack : List (bt A)) → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t findP key leaf tree0 st Pre _ exit = exit leaf tree0 st Pre (case1 refl) findP key (node key₁ v1 tree tree₁) tree0 st Pre next exit with <-cmp key key₁ findP key n tree0 st Pre _ exit | tri≈ ¬a refl ¬c = exit n tree0 st Pre (case2 refl) findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree tree0 (tree ∷ st) ⟪ treeLeftDown tree tree₁ (proj1 Pre) , findP1 a (proj2 Pre) ⟫ depth-1< where findP1 : key < key₁ → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (tree ∷ st) findP1 a si = s-left a si findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ tree0 (tree₁ ∷ st) ⟪ treeRightDown tree tree₁ (proj1 Pre) , s-right c (proj2 Pre) ⟫ depth-2< replaceTree1 : {n : Level} {A : Set n} {t t₁ : bt A } → ( k : ℕ ) → (v1 value : A ) → treeInvariant (node k v1 t t₁) → treeInvariant (node k value t t₁) replaceTree1 k v1 value (t-single .k .v1) = t-single k value replaceTree1 k v1 value (t-right x t) = t-right x t replaceTree1 k v1 value (t-left x t) = t-left x t replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ replaceNodeP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A) → (tree ≡ leaf ) ∨ ( node-key tree ≡ just key ) → (treeInvariant tree ) → ((tree1 : bt A) → treeInvariant tree1 → replacedTree key value 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) r-node replaceP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → {tree0 tree : bt A} ( repl : bt A) → (stack : List (bt A)) → treeInvariant tree0 ∧ stackInvariant key tree tree0 stack ∧ replacedTree key value tree repl → (next : ℕ → A → {tree0 tree1 : bt A } (repl : bt A) → (stack1 : List (bt A)) → treeInvariant tree0 ∧ stackInvariant key tree1 tree0 stack1 ∧ replacedTree key value tree1 repl → length stack1 < length stack → t) → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t replaceP key value {tree0} {tree} repl [] Pre next exit with proj1 (proj2 Pre) ... | () replaceP {_} {_} {A} key value {tree0} {tree} repl (leaf ∷ []) Pre next exit = exit tree0 repl ⟪ proj1 Pre , subst (λ k → replacedTree key value k repl ) (repl4 (proj1 (proj2 Pre))) (proj2 (proj2 Pre)) ⟫ where repl4 : stackInvariant key tree tree0 (leaf ∷ []) → tree ≡ tree0 repl4 (s-single .leaf) = refl replaceP key value {tree0} {tree} repl (leaf ∷ leaf ∷ st) Pre next exit = {!!} -- can't happen replaceP key value {tree0} {tree} repl (leaf ∷ node key₁ value₁ left right ∷ st) Pre next exit with <-cmp key key₁ ... | tri< a ¬b ¬c = next key value (node key₁ value₁ left repl ) st {!!} {!!} ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right) st {!!} {!!} ... | tri> ¬a ¬b c = next key value (node key₁ value₁ repl right) st {!!} {!!} replaceP key value {tree0} {tree} repl (node key₁ value₁ left right ∷ st) Pre next exit with <-cmp key key₁ ... | tri> ¬a ¬b c = next key value (node key₁ value₁ repl right ) st {!!} ≤-refl ... | tri≈ ¬a b ¬c = next key value (node key value left right ) st {!!} ≤-refl where -- this case won't happen ... | tri< a ¬b ¬c = next key value (node key₁ value₁ left repl) st ⟪ proj1 Pre , ⟪ {!!} , r-right a (proj2 (proj2 Pre)) ⟫ ⟫ ≤-refl where repl2 : stackInvariant key tree tree0 (node key₁ value₁ left right ∷ st) → stackInvariant key (node key₁ value₁ left tree) tree0 st repl2 (s-single .(node key₁ value₁ left right)) = {!!} repl2 (s-right _ si) = {!!} repl2 (s-left _ si) = {!!} open import Relation.Binary.Definitions nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ nat-≤> (s≤s x x ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) open _∧_ RTtoTI0 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant tree → replacedTree key value tree repl → treeInvariant repl RTtoTI0 = {!!} RTtoTI1 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant repl → replacedTree key value tree repl → treeInvariant tree 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 insertTreeP {n} {m} {A} {t} tree key value P exit = TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → treeInvariant (proj1 p) ∧ stackInvariant key (proj1 p) tree (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tree , [] ⟫ ⟪ P , {!!} ⟫ $ λ p P loop → findP key (proj1 p) tree (proj2 p) {!!} (λ t _ s P1 lt → loop ⟪ t , s ⟫ {!!} lt ) $ λ t _ s P C → replaceNodeP key value t C (proj1 P) $ λ 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 ⟫ ⟫ ⟪ proj1 P , ⟪ {!!} , R ⟫ ⟫ $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) {!!} (λ key value repl1 stack P2 lt → loop ⟪ stack , ⟪ {!!} , repl1 ⟫ ⟫ {!!} 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 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 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 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 {!!} tree0 st → stackInvariant key {!!} 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< 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 ⟫ ⟫ ⟪ {!!} , ⟪ {!!} , R ⟫ ⟫ $ λ p P1 loop → replaceP key value (proj2 (proj2 p)) (proj1 p) {!!} (λ key value repl1 stack P2 lt → loop ⟪ stack , ⟪ {!!} , repl1 ⟫ ⟫ {!!} lt ) exit record findPC {n : Level} {A : Set n} (key1 : ℕ) (value1 : A) (tree : bt A ) (stack : List (bt A)) : Set n where field tree1 : bt A ci : replacedTree key1 value1 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 = {!!} 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) ⟪ 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 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 = {!!}