Mercurial > hg > Gears > GearsAgda
view hoareBinaryTree1.agda @ 919:4d379ebc53c8
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Tue, 04 Jun 2024 17:29:33 +0900 |
parents | 1d34a752add0 |
children | 368b9be0b312 |
line wrap: on
line source
module hoareBinaryTree1 where open import Level hiding (suc ; zero ; _⊔_ ) 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 -- -- -- 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 (bt-depth t ⊔ bt-depth t₁ ) open import Data.Unit hiding ( _≟_ ) -- ; _≤?_ ; _≤_) tr< : {n : Level} {A : Set n} → (key : ℕ) → bt A → Set tr< {_} {A} key leaf = ⊤ tr< {_} {A} key (node key₁ value tr tr₁) = (key₁ < key ) ∧ tr< key tr ∧ tr< key tr₁ tr> : {n : Level} {A : Set n} → (key : ℕ) → bt A → Set tr> {_} {A} key leaf = ⊤ tr> {_} {A} key (node key₁ value tr tr₁) = (key < key₁ ) ∧ tr> key tr ∧ tr> key tr₁ -- -- 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₁ → tr> key t₁ → tr> key t₂ → 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₁ → tr< key₁ t₁ → tr< key₁ t₂ → 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₂ → tr< key₁ t₁ → tr< key₁ t₂ → tr> key₁ t₃ → tr> key₁ t₄ → 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₄)) -- -- stack always contains original top at end (path of the tree) -- data stackInvariant {n : Level} {A : Set n} (key : ℕ) : (top orig : bt A) → (stack : List (bt A)) → Set n where s-nil : {tree0 : bt A} → stackInvariant key tree0 tree0 (tree0 ∷ []) s-right : (tree tree0 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 : (tree₁ tree0 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) : (before after : 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 t2 t → replacedTree key value (node k v1 t1 t2) (node k v1 t1 t) r-left : {k : ℕ } {v1 : A} → {t t1 t2 : bt A} → key < k → replacedTree key value t1 t → replacedTree key value (node k v1 t1 t2) (node k v1 t t2) 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 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 nat-<≡ : { x : ℕ } → x < x → ⊥ nat-<≡ (s≤s lt) = nat-<≡ lt nat-≡< : { x y : ℕ } → x ≡ y → x < y → ⊥ nat-≡< refl lt = nat-<≡ lt treeTest1 : bt ℕ treeTest1 = node 0 0 leaf (node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf)) treeTest2 : bt ℕ treeTest2 = node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf) treeInvariantTest1 : treeInvariant treeTest1 treeInvariantTest1 = t-right _ _ (m≤m+n _ 2) ⟪ add< _ , ⟪ ⟪ add< _ , _ ⟫ , _ ⟫ ⟫ ⟪ add< _ , ⟪ _ , _ ⟫ ⟫ (t-node _ _ _ (add< 0) (add< 1) ⟪ add< _ , ⟪ _ , _ ⟫ ⟫ _ _ _ (t-left _ _ (add< 0) _ _ (t-single 1 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 4 treeTest2 treeTest1 ( treeTest2 ∷ treeTest1 ∷ [] ) stackInvariantTest1 = s-right _ _ _ (add< 3) (s-nil ) si-property0 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 stack → ¬ ( stack ≡ [] ) si-property0 (s-nil ) () si-property0 (s-right _ _ _ x si) () si-property0 (s-left _ _ _ x si) () si-property1 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 tree1 : bt A} → {stack : List (bt A)} → stackInvariant key tree tree0 (tree1 ∷ stack) → tree1 ≡ tree si-property1 (s-nil ) = refl si-property1 (s-right _ _ _ _ si) = refl si-property1 (s-left _ _ _ _ si) = refl si-property2 : {n : Level} {A : Set n} {key : ℕ} {tree tree0 tree1 : bt A} → (stack : List (bt A)) → stackInvariant key tree tree0 (tree1 ∷ stack) → ¬ ( just leaf ≡ stack-last stack ) si-property2 (.leaf ∷ []) (s-right _ _ tree₁ x ()) refl si-property2 (x₁ ∷ x₂ ∷ stack) (s-right _ _ tree₁ x si) eq = si-property2 (x₂ ∷ stack) si eq si-property2 (.leaf ∷ []) (s-left _ _ tree₁ x ()) refl si-property2 (x₁ ∷ x₂ ∷ stack) (s-left _ _ tree₁ x si) eq = si-property2 (x₂ ∷ stack) si eq si-property-< : {n : Level} {A : Set n} {key kp : ℕ} {value₂ : A} {tree orig tree₃ : bt A} → {stack : List (bt A)} → ¬ ( tree ≡ leaf ) → treeInvariant (node kp value₂ tree tree₃ ) → stackInvariant key tree orig (tree ∷ node kp value₂ tree tree₃ ∷ stack) → key < kp si-property-< ne (t-node _ _ _ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (s-right .(node _ _ _ _) _ .(node _ _ _ _) x s-nil) = ⊥-elim (nat-<> x₁ x₂) si-property-< ne (t-node _ _ _ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (s-right .(node _ _ _ _) _ .(node _ _ _ _) x (s-right .(node _ _ (node _ _ _ _) (node _ _ _ _)) _ tree₁ x₇ si)) = ⊥-elim (nat-<> x₁ x₂) si-property-< ne (t-node _ _ _ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (s-right .(node _ _ _ _) _ .(node _ _ _ _) x (s-left .(node _ _ (node _ _ _ _) (node _ _ _ _)) _ tree x₇ si)) = ⊥-elim (nat-<> x₁ x₂) si-property-< ne ti (s-left _ _ _ x (s-right .(node _ _ _ _) _ tree₁ x₁ si)) = x si-property-< ne ti (s-left _ _ _ x (s-left .(node _ _ _ _) _ tree₁ x₁ si)) = x si-property-< ne ti (s-left _ _ _ x s-nil) = x si-property-< ne (t-single _ _) (s-right _ _ tree₁ x si) = ⊥-elim ( ne refl ) si-property-> : {n : Level} {A : Set n} {key kp : ℕ} {value₂ : A} {tree orig tree₃ : bt A} → {stack : List (bt A)} → ¬ ( tree ≡ leaf ) → treeInvariant (node kp value₂ tree₃ tree ) → stackInvariant key tree orig (tree ∷ node kp value₂ tree₃ tree ∷ stack) → kp < key si-property-> ne ti (s-right _ _ tree₁ x s-nil) = x si-property-> ne ti (s-right _ _ tree₁ x (s-right .(node _ _ tree₁ _) _ tree₂ x₁ si)) = x si-property-> ne ti (s-right _ _ tree₁ x (s-left .(node _ _ tree₁ _) _ tree x₁ si)) = x si-property-> ne (t-node _ _ _ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (s-left _ _ _ x s-nil) = ⊥-elim (nat-<> x₁ x₂) si-property-> ne (t-node _ _ _ x₂ x₃ x₄ x₅ x₆ x₇ ti ti₁) (s-left _ _ _ x (s-right .(node _ _ _ _) _ tree₁ x₁ si)) = ⊥-elim (nat-<> x₂ x₃) si-property-> ne (t-node _ _ _ x₂ x₃ x₄ x₅ x₆ x₇ ti ti₁) (s-left _ _ _ x (s-left .(node _ _ _ _) _ tree x₁ si)) = ⊥-elim (nat-<> x₂ x₃) si-property-> ne (t-single _ _) (s-left _ _ _ x s-nil) = ⊥-elim ( ne refl ) si-property-> ne (t-single _ _) (s-left _ _ _ x (s-right .(node _ _ leaf leaf) _ tree₁ x₁ si)) = ⊥-elim ( ne refl ) si-property-> ne (t-single _ _) (s-left _ _ _ x (s-left .(node _ _ leaf leaf) _ tree x₁ si)) = ⊥-elim ( ne 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 (t ∷ []) (s-nil ) = refl si-property-last key t t0 (.t ∷ x ∷ st) (s-right _ _ _ _ si ) with si-property1 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 si ... | refl = si-property-last key x t0 (x ∷ st) si -- Diffkey : {n : Level} (A : Set n) (tree0 : bt A) → (key : ℕ) → (tree : bt A) → (stack : List (bt A)) → (si : stackInvariant key tree tree0 stack) → Set -- Diffkey A leaf key .leaf .(leaf ∷ []) s-nil = ? -- Diffkey A (node key₁ value tree0 tree1) key .(node key₁ value tree0 tree1) .(node key₁ value tree0 tree1 ∷ []) s-nil = ? -- Diffkey A tree0 key leaf .(leaf ∷ _) (s-right .leaf .tree0 tree₁ x si) = ? -- Diffkey A tree0 key (node key₁ value tree tree₂) .(node key₁ value tree tree₂ ∷ _) (s-right .(node key₁ value tree tree₂) .tree0 tree₁ x si) = ? -- Diffkey A tree0 key tree .(tree ∷ _) (s-left .tree .tree0 tree₁ x si) = ? -- si-property-ne : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack -- → length stack > 1 → ¬ ( node-key tree ≡ just key ) -- si-property-ne = ? 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 _ _ _ _) _ (r-right x rt) = λ () rt-property1 {n} {A} key value .(node _ _ _ _) _ (r-left x rt) = λ () 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 open _∧_ 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 j i) depth-3< : {i : ℕ } → suc i ≤ suc (suc i) depth-3< {zero} = s≤s ( z≤n ) depth-3< {suc i} = s≤s (depth-3< {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₁ 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 ) → (exit : (tree1 : 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 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 st Pre (case2 refl) findP {n} {_} {A} key (node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree (tree ∷ st) ⟪ treeLeftDown tree tree₁ (proj1 Pre) , findP1 a st (proj2 Pre) ⟫ depth-1< where findP1 : key < key₁ → (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) tree0 st → stackInvariant key tree tree0 (tree ∷ st) findP1 a (x ∷ st) si = s-left _ _ _ a si findP key n@(node key₁ v1 tree tree₁) tree0 st Pre next _ | tri> ¬a ¬b c = next tree₁ (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 a b t) = t-right _ _ x a b t replaceTree1 k v1 value (t-left _ _ x a b t) = t-left _ _ x a b t replaceTree1 k v1 value (t-node _ _ _ x x₁ a b c d t t₁) = t-node _ _ _ x x₁ a b c d t t₁ open import Relation.Binary.Definitions lemma3 : {i j : ℕ} → 0 ≡ i → j < i → ⊥ lemma3 refl () lemma5 : {i j : ℕ} → i < 1 → j < i → ⊥ lemma5 (s≤s z≤n) () ¬x<x : {x : ℕ} → ¬ (x < x) ¬x<x (s≤s lt) = ¬x<x lt child-replaced : {n : Level} {A : Set n} (key : ℕ) (tree : bt A) → bt A child-replaced key leaf = leaf child-replaced key (node key₁ value left right) with <-cmp key key₁ ... | tri< a ¬b ¬c = left ... | tri≈ ¬a b ¬c = node key₁ value left right ... | tri> ¬a ¬b c = right child-replaced-left : {n : Level} {A : Set n} {key key₁ : ℕ} {value : A} {left right : bt A} → key < key₁ → child-replaced key (node key₁ value left right) ≡ left child-replaced-left {n} {A} {key} {key₁} {value} {left} {right} lt = ch00 (node key₁ value left right) refl lt where ch00 : (tree : bt A) → tree ≡ node key₁ value left right → key < key₁ → child-replaced key tree ≡ left ch00 (node key₁ value tree tree₁) refl lt1 with <-cmp key key₁ ... | tri< a ¬b ¬c = refl ... | tri≈ ¬a b ¬c = ⊥-elim (¬a lt1) ... | tri> ¬a ¬b c = ⊥-elim (¬a lt1) child-replaced-right : {n : Level} {A : Set n} {key key₁ : ℕ} {value : A} {left right : bt A} → key₁ < key → child-replaced key (node key₁ value left right) ≡ right child-replaced-right {n} {A} {key} {key₁} {value} {left} {right} lt = ch00 (node key₁ value left right) refl lt where ch00 : (tree : bt A) → tree ≡ node key₁ value left right → key₁ < key → child-replaced key tree ≡ right ch00 (node key₁ value tree tree₁) refl lt1 with <-cmp key key₁ ... | tri< a ¬b ¬c = ⊥-elim (¬c lt1) ... | tri≈ ¬a b ¬c = ⊥-elim (¬c lt1) ... | tri> ¬a ¬b c = refl child-replaced-eq : {n : Level} {A : Set n} {key key₁ : ℕ} {value : A} {left right : bt A} → key₁ ≡ key → child-replaced key (node key₁ value left right) ≡ node key₁ value left right child-replaced-eq {n} {A} {key} {key₁} {value} {left} {right} lt = ch00 (node key₁ value left right) refl lt where ch00 : (tree : bt A) → tree ≡ node key₁ value left right → key₁ ≡ key → child-replaced key tree ≡ node key₁ value left right ch00 (node key₁ value tree tree₁) refl lt1 with <-cmp key key₁ ... | tri< a ¬b ¬c = ⊥-elim (¬b (sym lt1)) ... | tri≈ ¬a b ¬c = refl ... | tri> ¬a ¬b c = ⊥-elim (¬b (sym lt1)) record replacePR {n : Level} {A : Set n} (key : ℕ) (value : A) (tree repl : bt A ) (stack : List (bt A)) (C : bt A → bt A → List (bt A) → Set n) : Set n where field tree0 : bt A ti : treeInvariant tree0 si : stackInvariant key tree tree0 stack ri : replacedTree key value (child-replaced key tree ) repl ci : C tree repl stack -- data continuation record replacePR' {n : Level} {A : Set n} (key : ℕ) (value : A) (orig : bt A ) (stack : List (bt A)) : Set n where field tree repl : bt A ti : treeInvariant orig si : stackInvariant key tree orig stack ri : replacedTree key value (child-replaced key tree) repl -- treeInvariant of tree and repl is inferred from ti, si and ri. 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 (child-replaced key 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) (subst (λ j → replacedTree k v1 j (node k v1 t t₁) ) repl00 r-node) where repl00 : node k value t t₁ ≡ child-replaced k (node k value t t₁) repl00 with <-cmp k k ... | tri< a ¬b ¬c = ⊥-elim (¬b refl) ... | tri≈ ¬a b ¬c = refl ... | tri> ¬a ¬b c = ⊥-elim (¬b refl) replaceP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → {tree : bt A} ( repl : bt A) → (stack : List (bt A)) → replacePR key value tree repl stack (λ _ _ _ → Lift n ⊤) → (next : ℕ → A → {tree1 : bt A } (repl : bt A) → (stack1 : List (bt A)) → replacePR key value tree1 repl stack1 (λ _ _ _ → Lift n ⊤) → length stack1 < length stack → t) → (exit : (tree1 repl : bt A) → treeInvariant tree1 ∧ replacedTree key value tree1 repl → t) → t replaceP key value {tree} repl [] Pre next exit = ⊥-elim ( si-property0 (replacePR.si Pre) refl ) -- can't happen replaceP key value {tree} repl (leaf ∷ []) Pre next exit with si-property-last _ _ _ _ (replacePR.si Pre)-- tree0 ≡ leaf ... | refl = exit (replacePR.tree0 Pre) (node key value leaf leaf) ⟪ replacePR.ti Pre , r-leaf ⟫ replaceP key value {tree} repl (node key₁ value₁ left right ∷ []) Pre next exit with <-cmp key key₁ ... | tri< a ¬b ¬c = exit (replacePR.tree0 Pre) (node key₁ value₁ repl right ) ⟪ replacePR.ti Pre , repl01 ⟫ where repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ repl right ) repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ k right ) (node key₁ value₁ repl right )) repl02 (r-left a repl03) where repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl repl03 = replacePR.ri Pre repl02 : child-replaced key (node key₁ value₁ left right) ≡ left repl02 with <-cmp key key₁ ... | tri< a ¬b ¬c = refl ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a a) ... | tri> ¬a ¬b c = ⊥-elim ( ¬a a) ... | tri≈ ¬a b ¬c = exit (replacePR.tree0 Pre) repl ⟪ replacePR.ti Pre , repl01 ⟫ where repl01 : replacedTree key value (replacePR.tree0 Pre) repl repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) repl01 | refl | refl = subst (λ k → replacedTree key value k repl) repl02 (replacePR.ri Pre) where repl02 : child-replaced key (node key₁ value₁ left right) ≡ node key₁ value₁ left right repl02 with <-cmp key key₁ ... | tri< a ¬b ¬c = ⊥-elim ( ¬b b) ... | tri≈ ¬a b ¬c = refl ... | tri> ¬a ¬b c = ⊥-elim ( ¬b b) ... | tri> ¬a ¬b c = exit (replacePR.tree0 Pre) (node key₁ value₁ left repl ) ⟪ replacePR.ti Pre , repl01 ⟫ where repl01 : replacedTree key value (replacePR.tree0 Pre) (node key₁ value₁ left repl ) repl01 with si-property1 (replacePR.si Pre) | si-property-last _ _ _ _ (replacePR.si Pre) repl01 | refl | refl = subst (λ k → replacedTree key value (node key₁ value₁ left k ) (node key₁ value₁ left repl )) repl02 (r-right c repl03) where repl03 : replacedTree key value ( child-replaced key (node key₁ value₁ left right)) repl repl03 = replacePR.ri Pre repl02 : child-replaced key (node key₁ value₁ left right) ≡ right repl02 with <-cmp key key₁ ... | tri< a ¬b ¬c = ⊥-elim ( ¬c c) ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬c c) ... | tri> ¬a ¬b c = refl replaceP {n} {_} {A} key value {tree} repl (leaf ∷ st@(tree1 ∷ st1)) Pre next exit = next key value repl st Post ≤-refl where Post : replacePR key value tree1 repl (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) Post with replacePR.si Pre ... | s-right _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 tree₁ leaf repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl07 : child-replaced key (node key₂ v1 tree₁ leaf) ≡ leaf repl07 with <-cmp key key₂ ... | tri< a ¬b ¬c = ⊥-elim (¬c x) ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) ... | tri> ¬a ¬b c = refl repl12 : replacedTree key value (child-replaced key tree1 ) repl repl12 = subst₂ (λ j k → replacedTree key value j k ) (sym (subst (λ k → child-replaced key k ≡ leaf) (sym repl09) repl07 ) ) (sym (rt-property-leaf (replacePR.ri Pre))) r-leaf ... | s-left _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 leaf tree₁ repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl07 : child-replaced key (node key₂ v1 leaf tree₁ ) ≡ leaf repl07 with <-cmp key key₂ ... | tri< a ¬b ¬c = refl ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) ... | tri> ¬a ¬b c = ⊥-elim (¬a x) repl12 : replacedTree key value (child-replaced key tree1 ) repl repl12 = subst₂ (λ j k → replacedTree key value j k ) (sym (subst (λ k → child-replaced key k ≡ leaf) (sym repl09) repl07) ) (sym (rt-property-leaf (replacePR.ri Pre ))) r-leaf -- repl12 = subst₂ (λ j k → replacedTree key value j k ) (sym (subst (λ k → child-replaced key k ≡ leaf) (sym repl09) repl07 ) ) (sym (rt-property-leaf (replacePR.ri Pre))) r-leaf replaceP {n} {_} {A} key value {tree} repl (node key₁ value₁ left right ∷ st@(tree1 ∷ st1)) Pre next exit with <-cmp key key₁ ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st Post ≤-refl where Post : replacePR key value tree1 (node key₁ value₁ repl right ) st (λ _ _ _ → Lift n ⊤) Post with replacePR.si Pre ... | s-right _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl03 : child-replaced key (node key₁ value₁ left right) ≡ left repl03 with <-cmp key key₁ ... | tri< a1 ¬b ¬c = refl ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) ... | tri> ¬a ¬b c = ⊥-elim (¬a a) repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right repl02 with repl09 | <-cmp key key₂ ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) ... | refl | tri> ¬a ¬b c = refl repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 repl04 = begin node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ node key₁ value₁ left right ≡⟨ sym repl02 ⟩ child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ child-replaced key tree1 ∎ where open ≡-Reasoning repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) ... | s-left _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl03 : child-replaced key (node key₁ value₁ left right) ≡ left repl03 with <-cmp key key₁ ... | tri< a1 ¬b ¬c = refl ... | tri≈ ¬a b ¬c = ⊥-elim (¬a a) ... | tri> ¬a ¬b c = ⊥-elim (¬a a) repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right repl02 with repl09 | <-cmp key key₂ ... | refl | tri< a ¬b ¬c = refl ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) repl04 : node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡ child-replaced key tree1 repl04 = begin node key₁ value₁ (child-replaced key (node key₁ value₁ left right)) right ≡⟨ cong (λ k → node key₁ value₁ k right) repl03 ⟩ node key₁ value₁ left right ≡⟨ sym repl02 ⟩ child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ child-replaced key tree1 ∎ where open ≡-Reasoning repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ repl right) repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ repl right) ) repl04 (r-left a (replacePR.ri Pre)) ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st Post ≤-refl where Post : replacePR key value tree1 (node key₁ value left right ) (tree1 ∷ st1) (λ _ _ _ → Lift n ⊤) Post with replacePR.si Pre ... | s-right _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 b ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 tree₁ tree -- (node key₁ value₁ left right) repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl07 : child-replaced key (node key₂ v1 tree₁ tree) ≡ tree repl07 with <-cmp key key₂ ... | tri< a ¬b ¬c = ⊥-elim (¬c x) ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x) ... | tri> ¬a ¬b c = refl repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) repl12 refl with repl09 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node ... | s-left _ _ tree₁ {key₂} {v1} x si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 b ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 tree tree₁ repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl07 : child-replaced key (node key₂ v1 tree tree₁ ) ≡ tree repl07 with <-cmp key key₂ ... | tri< a ¬b ¬c = refl ... | tri≈ ¬a b ¬c = ⊥-elim (¬a x) ... | tri> ¬a ¬b c = ⊥-elim (¬a x) repl12 : (key ≡ key₁) → replacedTree key value (child-replaced key tree1 ) (node key₁ value left right ) repl12 refl with repl09 ... | refl = subst (λ k → replacedTree key value k (node key₁ value left right )) (sym repl07) r-node ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st Post ≤-refl where Post : replacePR key value tree1 (node key₁ value₁ left repl ) st (λ _ _ _ → Lift n ⊤) Post with replacePR.si Pre ... | s-right _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 tree₁ (node key₁ value₁ left right) repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl03 : child-replaced key (node key₁ value₁ left right) ≡ right repl03 with <-cmp key key₁ ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) ... | tri> ¬a ¬b c = refl repl02 : child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡ node key₁ value₁ left right repl02 with repl09 | <-cmp key key₂ ... | refl | tri< a ¬b ¬c = ⊥-elim (¬c lt) ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬c lt) ... | refl | tri> ¬a ¬b c = refl repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 repl04 = begin node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ node key₁ value₁ left right ≡⟨ sym repl02 ⟩ child-replaced key (node key₂ v1 tree₁ (node key₁ value₁ left right) ) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ child-replaced key tree1 ∎ where open ≡-Reasoning repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) ... | s-left _ _ tree₁ {key₂} {v1} lt si = record { tree0 = replacePR.tree0 Pre ; ti = replacePR.ti Pre ; si = repl10 ; ri = repl12 ; ci = lift tt } where repl09 : tree1 ≡ node key₂ v1 (node key₁ value₁ left right) tree₁ repl09 = si-property1 si repl10 : stackInvariant key tree1 (replacePR.tree0 Pre) (tree1 ∷ st1) repl10 with si-property1 si ... | refl = si repl03 : child-replaced key (node key₁ value₁ left right) ≡ right repl03 with <-cmp key key₁ ... | tri< a1 ¬b ¬c = ⊥-elim (¬c c) ... | tri≈ ¬a b ¬c = ⊥-elim (¬c c) ... | tri> ¬a ¬b c = refl repl02 : child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡ node key₁ value₁ left right repl02 with repl09 | <-cmp key key₂ ... | refl | tri< a ¬b ¬c = refl ... | refl | tri≈ ¬a b ¬c = ⊥-elim (¬a lt) ... | refl | tri> ¬a ¬b c = ⊥-elim (¬a lt) repl04 : node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡ child-replaced key tree1 repl04 = begin node key₁ value₁ left (child-replaced key (node key₁ value₁ left right)) ≡⟨ cong (λ k → node key₁ value₁ left k ) repl03 ⟩ node key₁ value₁ left right ≡⟨ sym repl02 ⟩ child-replaced key (node key₂ v1 (node key₁ value₁ left right) tree₁) ≡⟨ cong (λ k → child-replaced key k ) (sym repl09) ⟩ child-replaced key tree1 ∎ where open ≡-Reasoning repl12 : replacedTree key value (child-replaced key tree1 ) (node key₁ value₁ left repl) repl12 = subst (λ k → replacedTree key value k (node key₁ value₁ left repl) ) repl04 (r-right c (replacePR.ri Pre)) 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 (m≤n⇒m≤1+n 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 _∧_ ri-tr> : {n : Level} {A : Set n} → (tree repl : bt A) → (key key₁ : ℕ) → (value : A) → replacedTree key value tree repl → key₁ < key → tr> key₁ tree → tr> key₁ repl ri-tr> .leaf .(node key value leaf leaf) key key₁ value r-leaf a tgt = ⟪ a , ⟪ tt , tt ⟫ ⟫ ri-tr> .(node key _ _ _) .(node key value _ _) key key₁ value r-node a tgt = ⟪ a , ⟪ proj1 (proj2 tgt) , proj2 (proj2 tgt) ⟫ ⟫ ri-tr> .(node _ _ _ _) .(node _ _ _ _) key key₁ value (r-right x ri) a tgt = ⟪ proj1 tgt , ⟪ proj1 (proj2 tgt) , ri-tr> _ _ _ _ _ ri a (proj2 (proj2 tgt)) ⟫ ⟫ ri-tr> .(node _ _ _ _) .(node _ _ _ _) key key₁ value (r-left x ri) a tgt = ⟪ proj1 tgt , ⟪ ri-tr> _ _ _ _ _ ri a (proj1 (proj2 tgt)) , proj2 (proj2 tgt) ⟫ ⟫ ri-tr< : {n : Level} {A : Set n} → (tree repl : bt A) → (key key₁ : ℕ) → (value : A) → replacedTree key value tree repl → key < key₁ → tr< key₁ tree → tr< key₁ repl ri-tr< .leaf .(node key value leaf leaf) key key₁ value r-leaf a tgt = ⟪ a , ⟪ tt , tt ⟫ ⟫ ri-tr< .(node key _ _ _) .(node key value _ _) key key₁ value r-node a tgt = ⟪ a , ⟪ proj1 (proj2 tgt) , proj2 (proj2 tgt) ⟫ ⟫ ri-tr< .(node _ _ _ _) .(node _ _ _ _) key key₁ value (r-right x ri) a tgt = ⟪ proj1 tgt , ⟪ proj1 (proj2 tgt) , ri-tr< _ _ _ _ _ ri a (proj2 (proj2 tgt)) ⟫ ⟫ ri-tr< .(node _ _ _ _) .(node _ _ _ _) key key₁ value (r-left x ri) a tgt = ⟪ proj1 tgt , ⟪ ri-tr< _ _ _ _ _ ri a (proj1 (proj2 tgt)) , proj2 (proj2 tgt) ⟫ ⟫ <-tr> : {n : Level} {A : Set n} → {tree : bt A} → {key₁ key₂ : ℕ} → tr> key₁ tree → key₂ < key₁ → tr> key₂ tree <-tr> {n} {A} {leaf} {key₁} {key₂} tr lt = tt <-tr> {n} {A} {node key value t t₁} {key₁} {key₂} tr lt = ⟪ <-trans lt (proj1 tr) , ⟪ <-tr> (proj1 (proj2 tr)) lt , <-tr> (proj2 (proj2 tr)) lt ⟫ ⟫ >-tr< : {n : Level} {A : Set n} → {tree : bt A} → {key₁ key₂ : ℕ} → tr< key₁ tree → key₁ < key₂ → tr< key₂ tree >-tr< {n} {A} {leaf} {key₁} {key₂} tr lt = tt >-tr< {n} {A} {node key value t t₁} {key₁} {key₂} tr lt = ⟪ <-trans (proj1 tr) lt , ⟪ >-tr< (proj1 (proj2 tr)) lt , >-tr< (proj2 (proj2 tr)) lt ⟫ ⟫ RTtoTI0 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant tree → replacedTree key value tree repl → treeInvariant repl RTtoTI0 .leaf .(node key value leaf leaf) key value ti r-leaf = t-single key value RTtoTI0 .(node key _ leaf leaf) .(node key value leaf leaf) key value (t-single .key _) r-node = t-single key value RTtoTI0 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right _ _ x a b ti) r-node = t-right _ _ x a b ti RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left _ _ x a b ti) r-node = t-left _ _ x a b ti RTtoTI0 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node _ _ _ x x₁ a b c d ti ti₁) r-node = t-node _ _ _ x x₁ a b c d ti ti₁ -- r-right case RTtoTI0 (node _ _ leaf leaf) (node _ _ leaf .(node key value leaf leaf)) key value (t-single _ _) (r-right x r-leaf) = t-right _ _ x _ _ (t-single key value) RTtoTI0 (node _ _ leaf right@(node _ _ _ _)) (node key₁ value₁ leaf leaf) key value (t-right _ _ x₁ a b ti) (r-right x ri) = t-single key₁ value₁ RTtoTI0 (node key₁ _ leaf right@(node key₂ _ left₁ right₁)) (node key₁ value₁ leaf right₃@(node key₃ _ left₂ right₂)) key value (t-right key₄ key₅ x₁ a b ti) (r-right x ri) = t-right _ _ (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (rr00 ri a ) (rr02 ri b) (RTtoTI0 right right₃ key value ti ri) where rr00 : replacedTree key value (node key₂ _ left₁ right₁) (node key₃ _ left₂ right₂) → tr> key₁ left₁ → tr> key₁ left₂ rr00 r-node tb = tb rr00 (r-right x ri) tb = tb rr00 (r-left x₂ ri) tb = ri-tr> _ _ _ _ _ ri x tb rr02 : replacedTree key value (node key₂ _ left₁ right₁) (node key₃ _ left₂ right₂) → tr> key₁ right₁ → tr> key₁ right₂ rr02 r-node tb = tb rr02 (r-right x₂ ri) tb = ri-tr> _ _ _ _ _ ri x tb rr02 (r-left x ri) tb = tb RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left _ _ x₁ a b ti) (r-right x ()) RTtoTI0 (node key₁ _ (node key₃ _ _ _) leaf) (node key₁ _ (node key₃ value₃ _ _) (node key value leaf leaf)) key value (t-left _ _ x₁ a b ti) (r-right x r-leaf) = t-node _ _ _ x₁ x a b tt tt ti (t-single key value) RTtoTI0 (node key₁ _ (node _ _ left₀ right₀) (node key₂ _ left₁ right₁)) (node key₁ _ (node _ _ left₂ right₂) (node key₃ _ left₃ right₃)) key value (t-node _ _ _ x₁ x₂ a b c d ti ti₁) (r-right x ri) = t-node _ _ _ x₁ (subst (λ k → key₁ < k ) (rt-property-key ri) x₂) a b (rr00 ri c) (rr02 ri d) ti (RTtoTI0 _ _ key value ti₁ ri) where rr00 : replacedTree key value (node key₂ _ _ _) (node key₃ _ _ _) → tr> key₁ left₁ → tr> key₁ left₃ rr00 r-node tb = tb rr00 (r-right x₃ ri) tb = tb rr00 (r-left x₃ ri) tb = ri-tr> _ _ _ _ _ ri x tb rr02 : replacedTree key value (node key₂ _ _ _) (node key₃ _ _ _) → tr> key₁ right₁ → tr> key₁ right₃ rr02 r-node tb = tb rr02 (r-right x₃ ri) tb = ri-tr> _ _ _ _ _ ri x tb rr02 (r-left x₃ ri) tb = tb -- r-left case RTtoTI0 .(node _ _ leaf leaf) .(node _ _ (node key value leaf leaf) leaf) key value (t-single _ _) (r-left x r-leaf) = t-left _ _ x tt tt (t-single _ _ ) RTtoTI0 .(node _ _ leaf (node _ _ _ _)) (node key₁ value₁ (node key value leaf leaf) (node _ _ _ _)) key value (t-right _ _ x₁ a b ti) (r-left x r-leaf) = t-node _ _ _ x x₁ tt tt a b (t-single key value) ti RTtoTI0 (node key₃ _ (node key₂ _ left₁ right₁) leaf) (node key₃ _ (node key₁ value₁ left₂ right₂) leaf) key value (t-left _ _ x₁ a b ti) (r-left x ri) = t-left _ _ (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (rr00 ri a) (rr02 ri b) (RTtoTI0 _ _ key value ti ri) where -- key₁ < key₃ rr00 : replacedTree key value (node key₂ _ left₁ right₁) (node key₁ _ left₂ right₂) → tr< key₃ left₁ → tr< key₃ left₂ rr00 r-node tb = tb rr00 (r-right x₂ ri) tb = tb rr00 (r-left x₂ ri) tb = ri-tr< _ _ _ _ _ ri x tb rr02 : replacedTree key value (node key₂ _ left₁ right₁) (node key₁ _ left₂ right₂) → tr< key₃ right₁ → tr< key₃ right₂ rr02 r-node tb = tb rr02 (r-right x₃ ri) tb = ri-tr< _ _ _ _ _ ri x tb rr02 (r-left x₃ ri) tb = tb RTtoTI0 (node key₁ _ (node key₂ _ left₂ right₂) (node key₃ _ left₃ right₃)) (node key₁ _ (node key₄ _ left₄ right₄) (node key₅ _ left₅ right₅)) key value (t-node _ _ _ x₁ x₂ a b c d ti ti₁) (r-left x ri) = t-node _ _ _ (subst (λ k → k < key₁ ) (rt-property-key ri) x₁) x₂ (rr00 ri a) (rr02 ri b) c d (RTtoTI0 _ _ key value ti ri) ti₁ where rr00 : replacedTree key value (node key₂ _ left₂ right₂) (node key₄ _ left₄ right₄) → tr< key₁ left₂ → tr< key₁ left₄ rr00 r-node tb = tb rr00 (r-right x₃ ri) tb = tb rr00 (r-left x₃ ri) tb = ri-tr< _ _ _ _ _ ri x tb rr02 : replacedTree key value (node key₂ _ left₂ right₂) (node key₄ _ left₄ right₄) → tr< key₁ right₂ → tr< key₁ right₄ rr02 r-node tb = tb rr02 (r-right x₃ ri) tb = ri-tr< _ _ _ _ _ ri x tb rr02 (r-left x₃ ri) tb = tb -- RTtoTI1 : {n : Level} {A : Set n} → (tree repl : bt A) → (key : ℕ) → (value : A) → treeInvariant repl -- → replacedTree key value tree repl → treeInvariant tree -- RTtoTI1 .leaf .(node key value leaf leaf) key value ti r-leaf = t-leaf -- RTtoTI1 (node key value₁ leaf leaf) .(node key value leaf leaf) key value (t-single .key .value) r-node = t-single key value₁ -- RTtoTI1 .(node key _ leaf (node _ _ _ _)) .(node key value leaf (node _ _ _ _)) key value (t-right _ _ x a b ti) r-node = t-right _ _ x a b ti -- RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left _ _ x a b ti) r-node = t-left _ _ x a b ti -- RTtoTI1 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node _ _ _ x x₁ a b c d ti ti₁) r-node = t-node _ _ _ x x₁ a b c d ti ti₁ -- -- r-right case -- RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ leaf (node _ _ _ _)) key value (t-right _ _ x₁ a b ti) (r-right x r-leaf) = t-single key₁ value₁ -- RTtoTI1 (node key₁ value₁ leaf (node key₂ value₂ t2 t3)) (node key₁ _ leaf (node key₃ _ _ _)) key value (t-right _ _ x₁ a b ti) (r-right x ri) = -- t-right _ _ (subst (λ k → key₁ < k ) (sym (rt-property-key ri)) x₁) ? ? (RTtoTI1 _ _ key value ti ri) -- key₁ < key₂ -- RTtoTI1 (node _ _ (node _ _ _ _) leaf) (node _ _ (node _ _ _ _) (node key value _ _)) key value (t-node _ _ _ x₁ x₂ a b c d ti ti₁) (r-right x r-leaf) = -- t-left _ _ x₁ ? ? ti -- RTtoTI1 (node key₄ _ (node key₃ _ _ _) (node key₁ value₁ n n₁)) (node key₄ _ (node key₃ _ _ _) (node key₂ _ _ _)) key value (t-node _ _ _ x₁ x₂ a b c d ti ti₁) (r-right x ri) = t-node _ _ _ x₁ (subst (λ k → key₄ < k ) (sym (rt-property-key ri)) x₂) a b ? ? ti (RTtoTI1 _ _ key value ti₁ ri) -- key₄ < key₁ -- -- r-left case -- RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ _ leaf) key value (t-left _ _ x₁ a b ti) (r-left x ri) = t-single key₁ value₁ -- RTtoTI1 (node key₁ _ (node key₂ value₁ n n₁) leaf) (node key₁ _ (node key₃ _ _ _) leaf) key value (t-left _ _ x₁ a b ti) (r-left x ri) = -- t-left _ _ (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) ? ? (RTtoTI1 _ _ key value ti ri) -- key₂ < key₁ -- RTtoTI1 (node key₁ value₁ leaf _) (node key₁ _ _ _) key value (t-node _ _ _ x₁ x₂ a b c d ti ti₁) (r-left x r-leaf) = t-right _ _ x₂ c d ti₁ -- RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node _ _ _ x₁ x₂ a b c d ti ti₁) (r-left x ri) = -- t-node _ _ _ (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ ? ? c d (RTtoTI1 _ _ key value ti ri) ti₁ -- key₂ < key₁ insertTreeP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree → (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-nil ⟫ $ λ 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 } $ λ 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 ) $ λ tree repl P → exit tree repl ⟪ RTtoTI0 _ _ _ _ (proj1 P) (proj2 P) , proj2 P ⟫ insertTestP1 = insertTreeP leaf 1 1 t-leaf $ λ _ x0 P0 → insertTreeP x0 2 1 (proj1 P0) $ λ _ x1 P1 → insertTreeP x1 3 2 (proj1 P1) $ λ _ x2 P2 → insertTreeP x2 2 2 (proj1 P2) (λ _ x P → x ) 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 -- is realy inserted? -- other element is preserved? -- deletion? data Color : Set where Red : Color Black : Color RB→bt : {n : Level} (A : Set n) → (bt (Color ∧ A)) → bt A RB→bt {n} A leaf = leaf RB→bt {n} A (node key ⟪ C , value ⟫ tr t1) = (node key value (RB→bt A tr) (RB→bt A t1)) color : {n : Level} {A : Set n} → (bt (Color ∧ A)) → Color color leaf = Black color (node key ⟪ C , value ⟫ rb rb₁) = C to-red : {n : Level} {A : Set n} → (tree : bt (Color ∧ A)) → bt (Color ∧ A) to-red leaf = leaf to-red (node key ⟪ _ , value ⟫ t t₁) = (node key ⟪ Red , value ⟫ t t₁) to-black : {n : Level} {A : Set n} → (tree : bt (Color ∧ A)) → bt (Color ∧ A) to-black leaf = leaf to-black (node key ⟪ _ , value ⟫ t t₁) = (node key ⟪ Black , value ⟫ t t₁) black-depth : {n : Level} {A : Set n} → (tree : bt (Color ∧ A) ) → ℕ black-depth leaf = 1 black-depth (node key ⟪ Red , value ⟫ t t₁) = black-depth t ⊔ black-depth t₁ black-depth (node key ⟪ Black , value ⟫ t t₁) = suc (black-depth t ⊔ black-depth t₁ ) zero≢suc : { m : ℕ } → zero ≡ suc m → ⊥ zero≢suc () suc≢zero : {m : ℕ } → suc m ≡ zero → ⊥ suc≢zero () data RBtreeInvariant {n : Level} {A : Set n} : (tree : bt (Color ∧ A)) → Set n where rb-leaf : RBtreeInvariant leaf rb-red : (key : ℕ) → (value : A) → {left right : bt (Color ∧ A)} → color left ≡ Black → color right ≡ Black → black-depth left ≡ black-depth right → RBtreeInvariant left → RBtreeInvariant right → RBtreeInvariant (node key ⟪ Red , value ⟫ left right) rb-black : (key : ℕ) → (value : A) → {left right : bt (Color ∧ A)} → black-depth left ≡ black-depth right → RBtreeInvariant left → RBtreeInvariant right → RBtreeInvariant (node key ⟪ Black , value ⟫ left right) RightDown : {n : Level} {A : Set n} → bt (Color ∧ A) → bt (Color ∧ A) RightDown leaf = leaf RightDown (node key ⟪ c , value ⟫ t1 t2) = t2 LeftDown : {n : Level} {A : Set n} → bt (Color ∧ A) → bt (Color ∧ A) LeftDown leaf = leaf LeftDown (node key ⟪ c , value ⟫ t1 t2 ) = t1 RBtreeLeftDown : {n : Level} {A : Set n} {key : ℕ} {value : A} {c : Color} → (left right : bt (Color ∧ A)) → RBtreeInvariant (node key ⟪ c , value ⟫ left right) → RBtreeInvariant left RBtreeLeftDown left right (rb-red _ _ x x₁ x₂ rb rb₁) = rb RBtreeLeftDown left right (rb-black _ _ x rb rb₁) = rb RBtreeRightDown : {n : Level} {A : Set n} { key : ℕ} {value : A} {c : Color} → (left right : bt (Color ∧ A)) → RBtreeInvariant (node key ⟪ c , value ⟫ left right) → RBtreeInvariant right RBtreeRightDown left right (rb-red _ _ x x₁ x₂ rb rb₁) = rb₁ RBtreeRightDown left right (rb-black _ _ x rb rb₁) = rb₁ RBtreeEQ : {n : Level} {A : Set n} {key : ℕ} {value : A} {c : Color} → {left right : bt (Color ∧ A)} → RBtreeInvariant (node key ⟪ c , value ⟫ left right) → black-depth left ≡ black-depth right RBtreeEQ (rb-red _ _ x x₁ x₂ rb rb₁) = x₂ RBtreeEQ (rb-black _ _ x rb rb₁) = x RBtreeToBlack : {n : Level} {A : Set n} → (tree : bt (Color ∧ A)) → RBtreeInvariant tree → RBtreeInvariant (to-black tree) RBtreeToBlack leaf rb-leaf = rb-leaf RBtreeToBlack (node key ⟪ Red , value ⟫ left right) (rb-red _ _ x x₁ x₂ rb rb₁) = rb-black key value x₂ rb rb₁ RBtreeToBlack (node key ⟪ Black , value ⟫ left right) (rb-black _ _ x rb rb₁) = rb-black key value x rb rb₁ RBtreeToBlackColor : {n : Level} {A : Set n} → (tree : bt (Color ∧ A)) → RBtreeInvariant tree → color (to-black tree) ≡ Black RBtreeToBlackColor leaf rb-leaf = refl RBtreeToBlackColor (node key ⟪ Red , value ⟫ left right) (rb-red _ _ x x₁ x₂ rb rb₁) = refl RBtreeToBlackColor (node key ⟪ Black , value ⟫ left right) (rb-black _ _ x rb rb₁) = refl RBtreeParentColorBlack : {n : Level} {A : Set n} → (left right : bt (Color ∧ A)) { value : A} {key : ℕ} { c : Color} → RBtreeInvariant (node key ⟪ c , value ⟫ left right) → (color left ≡ Red) ∨ (color right ≡ Red) → c ≡ Black RBtreeParentColorBlack leaf leaf (rb-red _ _ x₁ x₂ x₃ x₄ x₅) (case1 ()) RBtreeParentColorBlack leaf leaf (rb-red _ _ x₁ x₂ x₃ x₄ x₅) (case2 ()) RBtreeParentColorBlack (node key ⟪ Red , proj4 ⟫ left left₁) right (rb-red _ _ () x₁ x₂ rb rb₁) (case1 x₃) RBtreeParentColorBlack (node key ⟪ Black , proj4 ⟫ left left₁) right (rb-red _ _ x x₁ x₂ rb rb₁) (case1 ()) RBtreeParentColorBlack left (node key ⟪ Red , proj4 ⟫ right right₁) (rb-red _ _ x () x₂ rb rb₁) (case2 x₃) RBtreeParentColorBlack left (node key ⟪ Black , proj4 ⟫ right right₁) (rb-red _ _ x x₁ x₂ rb rb₁) (case2 ()) RBtreeParentColorBlack left right (rb-black _ _ x rb rb₁) x₃ = refl RBtreeChildrenColorBlack : {n : Level} {A : Set n} → (left right : bt (Color ∧ A)) { value : Color ∧ A} {key : ℕ} → RBtreeInvariant (node key value left right) → proj1 value ≡ Red → (color left ≡ Black) ∧ (color right ≡ Black) RBtreeChildrenColorBlack left right (rb-red _ _ x x₁ x₂ rbi rbi₁) refl = ⟪ x , x₁ ⟫ -- -- findRBT exit with replaced node -- case-eq node value is replaced, just do replacedTree and rebuild rb-invariant -- case-leaf insert new single node -- case1 if parent node is black, just do replacedTree and rebuild rb-invariant -- case2 if parent node is red, increase blackdepth, do rotatation -- findRBT : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt (Color ∧ A) ) → (stack : List (bt (Color ∧ A))) → RBtreeInvariant tree ∧ stackInvariant key tree tree0 stack → (next : (tree1 : bt (Color ∧ A) ) → (stack : List (bt (Color ∧ A))) → RBtreeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) → (exit : (tree1 : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) → RBtreeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t findRBT key leaf tree0 stack rb0 next exit = exit leaf stack rb0 (case1 refl) findRBT key (node key₁ value left right) tree0 stack rb0 next exit with <-cmp key key₁ findRBT key (node key₁ value left right) tree0 stack rb0 next exit | tri< a ¬b ¬c = next left (left ∷ stack) ⟪ RBtreeLeftDown left right (_∧_.proj1 rb0) , s-left _ _ _ a (_∧_.proj2 rb0) ⟫ depth-1< findRBT key n tree0 stack rb0 _ exit | tri≈ ¬a refl ¬c = exit n stack rb0 (case2 refl) findRBT key (node key₁ value left right) tree0 stack rb0 next exit | tri> ¬a ¬b c = next right (right ∷ stack) ⟪ RBtreeRightDown left right (_∧_.proj1 rb0), s-right _ _ _ c (_∧_.proj2 rb0) ⟫ depth-2< findTest : {n m : Level} {A : Set n } {t : Set m } → (key : ℕ) → (tree0 : bt (Color ∧ A)) → RBtreeInvariant tree0 → (exit : (tree1 : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) → RBtreeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t findTest {n} {m} {A} {t} k tr0 rb0 exit = TerminatingLoopS (bt (Color ∧ A) ∧ List (bt (Color ∧ A))) {λ p → RBtreeInvariant (proj1 p) ∧ stackInvariant k (proj1 p) tr0 (proj2 p) } (λ p → bt-depth (proj1 p)) ⟪ tr0 , tr0 ∷ [] ⟫ ⟪ rb0 , s-nil ⟫ $ λ p RBP loop → findRBT k (proj1 p) tr0 (proj2 p) RBP (λ t1 s1 P2 lt1 → loop ⟪ t1 , s1 ⟫ P2 lt1 ) $ λ tr1 st P2 O → exit tr1 st P2 O testRBTree0 : bt (Color ∧ ℕ) testRBTree0 = node 8 ⟪ Black , 800 ⟫ (node 5 ⟪ Red , 500 ⟫ (node 2 ⟪ Black , 200 ⟫ leaf leaf) (node 6 ⟪ Black , 600 ⟫ leaf leaf)) (node 10 ⟪ Red , 1000 ⟫ (leaf) (node 15 ⟪ Black , 1500 ⟫ (node 14 ⟪ Red , 1400 ⟫ leaf leaf) leaf)) -- testRBI0 : RBtreeInvariant testRBTree0 -- testRBI0 = rb-node-black (add< 2) (add< 1) refl (rb-node-red (add< 2) (add< 0) refl (rb-single 2 200) (rb-single 6 600)) (rb-right-red (add< 4) refl (rb-left-black (add< 0) refl (rb-single 14 1400) )) -- findRBTreeTest : result -- findRBTreeTest = findTest 14 testRBTree0 testRBI0 -- $ λ tr s P O → (record {tree = tr ; stack = s ; ti = (proj1 P) ; si = (proj2 P)}) -- create replaceRBTree with rotate data replacedRBTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt (Color ∧ A) ) → Set n where -- no rotation case rbr-leaf : replacedRBTree key value leaf (node key ⟪ Red , value ⟫ leaf leaf) rbr-node : {value₁ : A} → {ca : Color } → {t t₁ : bt (Color ∧ A)} → replacedRBTree key value (node key ⟪ ca , value₁ ⟫ t t₁) (node key ⟪ ca , value ⟫ t t₁) rbr-right : {k : ℕ } {v1 : A} → {ca : Color} → {t t1 t2 : bt (Color ∧ A)} → color t2 ≡ color t → k < key → replacedRBTree key value t2 t → replacedRBTree key value (node k ⟪ ca , v1 ⟫ t1 t2) (node k ⟪ ca , v1 ⟫ t1 t) rbr-left : {k : ℕ } {v1 : A} → {ca : Color} → {t t1 t2 : bt (Color ∧ A)} → color t1 ≡ color t → key < k → replacedRBTree key value t1 t → replacedRBTree key value (node k ⟪ ca , v1 ⟫ t1 t2) (node k ⟪ ca , v1 ⟫ t t2) -- k < key → key < k -- in all other case, color of replaced node is changed from Black to Red -- case1 parent is black rbr-black-right : {t t₁ t₂ : bt (Color ∧ A)} {value₁ : A} {key₁ : ℕ} → color t₂ ≡ Red → key₁ < key → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node key₁ ⟪ Black , value₁ ⟫ t t₁) (node key₁ ⟪ Black , value₁ ⟫ t t₂) rbr-black-left : {t t₁ t₂ : bt (Color ∧ A)} {value₁ : A} {key₁ : ℕ} → color t₂ ≡ Red → key < key₁ → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node key₁ ⟪ Black , value₁ ⟫ t₁ t) (node key₁ ⟪ Black , value₁ ⟫ t₂ t) -- case2 both parent and uncle are red (should we check uncle color?), flip color and up rbr-flip-ll : {t t₁ t₂ uncle : bt (Color ∧ A)} {kg kp : ℕ} {vg vp : A} → color t₂ ≡ Red → color uncle ≡ Red → key < kp → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node kg ⟪ Black , vg ⟫ (node kp ⟪ Red , vp ⟫ t₁ t) uncle) (node kg ⟪ Red , vg ⟫ (node kp ⟪ Black , vp ⟫ t₂ t) (to-black uncle)) rbr-flip-lr : {t t₁ t₂ uncle : bt (Color ∧ A)} {kg kp : ℕ} {vg vp : A} → color t₂ ≡ Red → color uncle ≡ Red → kp < key → key < kg → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node kg ⟪ Black , vg ⟫ (node kp ⟪ Red , vp ⟫ t t₁) uncle) (node kg ⟪ Red , vg ⟫ (node kp ⟪ Black , vp ⟫ t t₂) (to-black uncle)) rbr-flip-rl : {t t₁ t₂ uncle : bt (Color ∧ A)} {kg kp : ℕ} {vg vp : A} → color t₂ ≡ Red → color uncle ≡ Red → kg < key → key < kp → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Red , vp ⟫ t₁ t)) (node kg ⟪ Red , vg ⟫ (to-black uncle) (node kp ⟪ Black , vp ⟫ t₂ t)) rbr-flip-rr : {t t₁ t₂ uncle : bt (Color ∧ A)} {kg kp : ℕ} {vg vp : A} → color t₂ ≡ Red → color uncle ≡ Red → kp < key → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Red , vp ⟫ t t₁)) (node kg ⟪ Red , vg ⟫ (to-black uncle) (node kp ⟪ Black , vp ⟫ t t₂)) -- case6 the node is outer, rotate grand rbr-rotate-ll : {t t₁ t₂ uncle : bt (Color ∧ A)} {kg kp : ℕ} {vg vp : A} → color t₂ ≡ Red → key < kp → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node kg ⟪ Black , vg ⟫ (node kp ⟪ Red , vp ⟫ t₁ t) uncle) (node kp ⟪ Black , vp ⟫ t₂ (node kg ⟪ Red , vg ⟫ t uncle)) rbr-rotate-rr : {t t₁ t₂ uncle : bt (Color ∧ A)} {kg kp : ℕ} {vg vp : A} → color t₂ ≡ Red → kp < key → replacedRBTree key value t₁ t₂ → replacedRBTree key value (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Red , vp ⟫ t t₁)) (node kp ⟪ Black , vp ⟫ (node kg ⟪ Red , vg ⟫ uncle t) t₂ ) -- case56 the node is inner, make it outer and rotate grand rbr-rotate-lr : {t t₁ uncle : bt (Color ∧ A)} (t₂ t₃ : bt (Color ∧ A)) (kg kp kn : ℕ) {vg vp vn : A} → color t₃ ≡ Red → kp < key → key < kg → replacedRBTree key value t₁ (node kn ⟪ Red , vn ⟫ t₂ t₃) → replacedRBTree key value (node kg ⟪ Black , vg ⟫ (node kp ⟪ Red , vp ⟫ t t₁) uncle) (node kn ⟪ Black , vn ⟫ (node kp ⟪ Red , vp ⟫ t t₂) (node kg ⟪ Red , vg ⟫ t₃ uncle)) rbr-rotate-rl : {t t₁ uncle : bt (Color ∧ A)} (t₂ t₃ : bt (Color ∧ A)) (kg kp kn : ℕ) {vg vp vn : A} → color t₃ ≡ Red → kg < key → key < kp → replacedRBTree key value t₁ (node kn ⟪ Red , vn ⟫ t₂ t₃) → replacedRBTree key value (node kg ⟪ Black , vg ⟫ uncle (node kp ⟪ Red , vp ⟫ t₁ t)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , vg ⟫ uncle t₂) (node kp ⟪ Red , vp ⟫ t₃ t)) -- -- Parent Grand Relation -- should we require stack-invariant? -- data ParentGrand {n : Level} {A : Set n} (self : bt A) : (parent uncle grand : bt A) → Set n where s2-s1p2 : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } → parent ≡ node kp vp self n1 → grand ≡ node kg vg parent n2 → ParentGrand self parent n2 grand s2-1sp2 : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } → parent ≡ node kp vp n1 self → grand ≡ node kg vg parent n2 → ParentGrand self parent n2 grand s2-s12p : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } → parent ≡ node kp vp self n1 → grand ≡ node kg vg n2 parent → ParentGrand self parent n2 grand s2-1s2p : {kp kg : ℕ} {vp vg : A} → {n1 n2 : bt A} {parent grand : bt A } → parent ≡ node kp vp n1 self → grand ≡ node kg vg n2 parent → ParentGrand self parent n2 grand record PG {n : Level } (A : Set n) (self : bt A) (stack : List (bt A)) : Set n where field parent grand uncle : bt A pg : ParentGrand self parent uncle grand rest : List (bt A) stack=gp : stack ≡ ( self ∷ parent ∷ grand ∷ rest ) -- -- RBI : Invariant on InsertCase2 -- color repl ≡ Red ∧ black-depth repl ≡ suc (black-depth tree) -- data RBI-state {n : Level} {A : Set n} (key : ℕ) (value : A) : (tree repl : bt (Color ∧ A) ) → (stak : List (bt (Color ∧ A))) → Set n where rebuild : {tree repl : bt (Color ∧ A) } {stack : List (bt (Color ∧ A))} → color tree ≡ color repl → black-depth repl ≡ black-depth tree → (rotated : replacedRBTree key value tree repl) → RBI-state key value tree repl stack -- one stage up rotate : {tree repl : bt (Color ∧ A) } {stack : List (bt (Color ∧ A))} → color repl ≡ Red → black-depth repl ≡ black-depth tree → ¬ (tree ≡ leaf) → (rotated : replacedRBTree key value tree repl) → RBI-state key value tree repl stack -- two stages up top-black : {tree repl : bt (Color ∧ A) } → {stack : List (bt (Color ∧ A))} → stack ≡ tree ∷ [] → (rotated : replacedRBTree key value tree repl ∨ replacedRBTree key value (to-black tree) repl) → RBI-state key value tree repl stack record RBI {n : Level} {A : Set n} (key : ℕ) (value : A) (orig : bt (Color ∧ A) ) (stack : List (bt (Color ∧ A))) : Set n where field tree repl : bt (Color ∧ A) origti : treeInvariant orig origrb : RBtreeInvariant orig treerb : RBtreeInvariant tree -- tree node te be replaced replrb : RBtreeInvariant repl si : stackInvariant key tree orig stack state : RBI-state key value tree repl stack tr>-to-black : {n : Level} {A : Set n} {key : ℕ} {tree : bt (Color ∧ A)} → tr> key tree → tr> key (to-black tree) tr>-to-black {n} {A} {key} {leaf} tr = tt tr>-to-black {n} {A} {key} {node key₁ value tree tree₁} tr = tr tr<-to-black : {n : Level} {A : Set n} {key : ℕ} {tree : bt (Color ∧ A)} → tr< key tree → tr< key (to-black tree) tr<-to-black {n} {A} {key} {leaf} tr = tt tr<-to-black {n} {A} {key} {node key₁ value tree tree₁} tr = tr to-black-eq : {n : Level} {A : Set n} (tree : bt (Color ∧ A)) → color tree ≡ Red → suc (black-depth tree) ≡ black-depth (to-black tree) to-black-eq {n} {A} (leaf) () to-black-eq {n} {A} (node key₁ ⟪ Red , proj4 ⟫ tree tree₁) eq = refl to-black-eq {n} {A} (node key₁ ⟪ Black , proj4 ⟫ tree tree₁) () red-children-eq : {n : Level} {A : Set n} {tree left right : bt (Color ∧ A)} → {key : ℕ} → {value : A} → {c : Color} → tree ≡ node key ⟪ c , value ⟫ left right → c ≡ Red → RBtreeInvariant tree → (black-depth tree ≡ black-depth left ) ∧ (black-depth tree ≡ black-depth right) red-children-eq {n} {A} {.(node key₁ ⟪ Red , value₁ ⟫ left right)} {left} {right} {.key₁} {.value₁} {Red} refl eq₁ rb2@(rb-red key₁ value₁ x x₁ x₂ rb rb₁) = ⟪ ( begin black-depth left ⊔ black-depth right ≡⟨ cong (λ k → black-depth left ⊔ k) (sym (RBtreeEQ rb2)) ⟩ black-depth left ⊔ black-depth left ≡⟨ ⊔-idem _ ⟩ black-depth left ∎ ) , ( begin black-depth left ⊔ black-depth right ≡⟨ cong (λ k → k ⊔ black-depth right ) (RBtreeEQ rb2) ⟩ black-depth right ⊔ black-depth right ≡⟨ ⊔-idem _ ⟩ black-depth right ∎ ) ⟫ where open ≡-Reasoning red-children-eq {n} {A} {tree} {left} {right} {key} {value} {Black} eq () rb black-children-eq : {n : Level} {A : Set n} {tree left right : bt (Color ∧ A)} → {key : ℕ} → {value : A} → {c : Color} → tree ≡ node key ⟪ c , value ⟫ left right → c ≡ Black → RBtreeInvariant tree → (black-depth tree ≡ suc (black-depth left) ) ∧ (black-depth tree ≡ suc (black-depth right)) black-children-eq {n} {A} {.(node key₁ ⟪ Black , value₁ ⟫ left right)} {left} {right} {.key₁} {.value₁} {Black} refl eq₁ rb2@(rb-black key₁ value₁ x rb rb₁) = ⟪ ( begin suc (black-depth left ⊔ black-depth right) ≡⟨ cong (λ k → suc (black-depth left ⊔ k)) (sym (RBtreeEQ rb2)) ⟩ suc (black-depth left ⊔ black-depth left) ≡⟨ ⊔-idem _ ⟩ suc (black-depth left) ∎ ) , ( begin suc (black-depth left ⊔ black-depth right) ≡⟨ cong (λ k → suc (k ⊔ black-depth right)) (RBtreeEQ rb2) ⟩ suc (black-depth right ⊔ black-depth right) ≡⟨ ⊔-idem _ ⟩ suc (black-depth right) ∎ ) ⟫ where open ≡-Reasoning black-children-eq {n} {A} {tree} {left} {right} {key} {value} {Red} eq () rb ⊔-succ : {m n : ℕ} → suc (m ⊔ n) ≡ suc m ⊔ suc n ⊔-succ {m} {n} = refl RB-repl→eq : {n : Level} {A : Set n} → (tree repl : bt (Color ∧ A)) → {key : ℕ} → {value : A} → RBtreeInvariant tree → replacedRBTree key value tree repl → black-depth tree ≡ black-depth repl RB-repl→eq = ? -- RB-repl→eq {n} {A} .leaf .(node _ ⟪ Red , _ ⟫ leaf leaf) rbt rbr-leaf = refl -- RB-repl→eq {n} {A} (node _ ⟪ Red , _ ⟫ _ _) .(node _ ⟪ Red , _ ⟫ _ _) rbt rbr-node = refl -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ _ _) .(node _ ⟪ Black , _ ⟫ _ _) rbt rbr-node = refl -- RB-repl→eq {n} {A} (node _ ⟪ Red , _ ⟫ left _) .(node _ ⟪ Red , _ ⟫ left _) (rb-red _ _ x₂ x₃ x₄ rbt rbt₁) (rbr-right x x₁ t) = -- cong (λ k → black-depth left ⊔ k ) (RB-repl→eq _ _ rbt₁ t) -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ left _) .(node _ ⟪ Black , _ ⟫ left _) (rb-black _ _ x₂ rbt rbt₁) (rbr-right x x₁ t) = -- cong (λ k → suc (black-depth left ⊔ k)) (RB-repl→eq _ _ rbt₁ t) -- RB-repl→eq {n} {A} (node _ ⟪ Red , _ ⟫ _ right) .(node _ ⟪ Red , _ ⟫ _ right) (rb-red _ _ x₂ x₃ x₄ rbt rbt₁) (rbr-left x x₁ t) = cong (λ k → k ⊔ black-depth right) (RB-repl→eq _ _ rbt t) -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ _ right) .(node _ ⟪ Black , _ ⟫ _ right) (rb-black _ _ x₂ rbt rbt₁) (rbr-left x x₁ t) = cong (λ k → suc (k ⊔ black-depth right)) (RB-repl→eq _ _ rbt t) -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ t₁ _) .(node _ ⟪ Black , _ ⟫ t₁ _) (rb-black _ _ x₂ rbt rbt₁) (rbr-black-right x x₁ t) = cong (λ k → suc (black-depth t₁ ⊔ k)) (RB-repl→eq _ _ rbt₁ t) -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ _ t₁) .(node _ ⟪ Black , _ ⟫ _ t₁) (rb-black _ _ x₂ rbt rbt₁) (rbr-black-left x x₁ t) = cong (λ k → suc (k ⊔ black-depth t₁)) (RB-repl→eq _ _ rbt t) -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ t₁ t₂) t₃) .(node _ ⟪ Red , _ ⟫ (node _ ⟪ Black , _ ⟫ t₄ t₂) (to-black t₃)) (rb-black _ _ x₃ (rb-red _ _ x₄ x₅ x₆ rbt rbt₂) rbt₁) (rbr-flip-ll {_} {_} {t₄} x x₁ x₂ t) = begin -- suc (black-depth t₁ ⊔ black-depth t₂ ⊔ black-depth t₃) ≡⟨ cong (λ k → suc (k ⊔ black-depth t₂ ⊔ black-depth t₃)) (RB-repl→eq _ _ rbt t) ⟩ -- suc (black-depth t₄ ⊔ black-depth t₂) ⊔ suc (black-depth t₃) ≡⟨ cong (λ k → suc (black-depth t₄ ⊔ black-depth t₂) ⊔ k ) ( to-black-eq t₃ x₁ ) ⟩ -- suc (black-depth t₄ ⊔ black-depth t₂) ⊔ black-depth (to-black t₃) ∎ -- where open ≡-Reasoning -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ t₁ t₂) t₃) .(node _ ⟪ Red , _ ⟫ (node _ ⟪ Black , _ ⟫ t₁ t₄) (to-black t₃)) (rb-black _ _ x₄ (rb-red _ _ x₅ x₆ x₇ rbt rbt₂) rbt₁) (rbr-flip-lr {_} {_} {t₄} x x₁ x₂ x₃ t) = begin -- suc (black-depth t₁ ⊔ black-depth t₂) ⊔ suc (black-depth t₃) ≡⟨ cong (λ k → suc (black-depth t₁ ⊔ black-depth t₂) ⊔ k ) ( to-black-eq t₃ x₁ ) ⟩ -- suc (black-depth t₁ ⊔ black-depth t₂) ⊔ black-depth (to-black t₃) ≡⟨ cong (λ k → suc (black-depth t₁ ⊔ k ) ⊔ black-depth (to-black t₃)) (RB-repl→eq _ _ rbt₂ t) ⟩ -- suc (black-depth t₁ ⊔ black-depth t₄) ⊔ black-depth (to-black t₃) ∎ -- where open ≡-Reasoning -- RB-repl→eq {n} {A} (node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Red , _ ⟫ (to-black t₄) (node _ ⟪ Black , _ ⟫ t₃ t₁)) (rb-black _ _ x₄ rbt (rb-red _ _ x₅ x₆ x₇ rbt₁ rbt₂)) (rbr-flip-rl {t₁} {t₂} {t₃} {t₄} x x₁ x₂ x₃ t) = begin -- suc (black-depth t₄ ⊔ (black-depth t₂ ⊔ black-depth t₁)) ≡⟨ cong (λ k → suc (black-depth t₄ ⊔ ( k ⊔ black-depth t₁)) ) (RB-repl→eq _ _ rbt₁ t) ⟩ -- suc (black-depth t₄ ⊔ (black-depth t₃ ⊔ black-depth t₁)) ≡⟨ cong (λ k → k ⊔ suc (black-depth t₃ ⊔ black-depth t₁)) ( to-black-eq t₄ x₁ ) ⟩ -- black-depth (to-black t₄) ⊔ suc (black-depth t₃ ⊔ black-depth t₁) ∎ -- where open ≡-Reasoning -- RB-repl→eq {n} {A} .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Red , _ ⟫ (to-black _) (node _ ⟪ Black , _ ⟫ _ _)) (rb-black _ _ x₃ rbt (rb-red _ _ x₄ x₅ x₆ rbt₁ rbt₂)) (rbr-flip-rr {t₁} {t₂} {t₃} {t₄} x x₁ x₂ t) = begin -- suc (black-depth t₄ ⊔ (black-depth t₁ ⊔ black-depth t₂)) ≡⟨ cong (λ k → suc (black-depth t₄ ⊔ (black-depth t₁ ⊔ k ))) ( RB-repl→eq _ _ rbt₂ t) ⟩ -- suc (black-depth t₄ ⊔ (black-depth t₁ ⊔ black-depth t₃)) ≡⟨ cong (λ k → k ⊔ suc (black-depth t₁ ⊔ black-depth t₃)) ( to-black-eq t₄ x₁ ) ⟩ -- black-depth (to-black t₄) ⊔ suc (black-depth t₁ ⊔ black-depth t₃) ∎ -- where open ≡-Reasoning -- RB-repl→eq {n} {A} .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) (rb-black _ _ x₂ (rb-red _ _ x₃ x₄ x₅ rbt rbt₂) rbt₁) (rbr-rotate-ll {t₁} {t₂} {t₃} {t₄} x x₁ t) = begin -- suc (black-depth t₂ ⊔ black-depth t₁ ⊔ black-depth t₄) ≡⟨ cong suc ( ⊔-assoc (black-depth t₂) (black-depth t₁) (black-depth t₄)) ⟩ -- suc (black-depth t₂ ⊔ (black-depth t₁ ⊔ black-depth t₄)) ≡⟨ cong (λ k → suc (k ⊔ (black-depth t₁ ⊔ black-depth t₄)) ) (RB-repl→eq _ _ rbt t) ⟩ -- suc (black-depth t₃ ⊔ (black-depth t₁ ⊔ black-depth t₄)) ∎ -- where open ≡-Reasoning -- RB-repl→eq {n} {A} .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) (rb-black _ _ x₂ rbt (rb-red _ _ x₃ x₄ x₅ rbt₁ rbt₂)) (rbr-rotate-rr {t₁} {t₂} {t₃} {t₄} x x₁ t) = begin -- suc (black-depth t₄ ⊔ (black-depth t₁ ⊔ black-depth t₂)) ≡⟨ cong (λ k → suc (black-depth t₄ ⊔ (black-depth t₁ ⊔ k ))) ( RB-repl→eq _ _ rbt₂ t) ⟩ -- suc (black-depth t₄ ⊔ (black-depth t₁ ⊔ black-depth t₃)) ≡⟨ cong suc (sym ( ⊔-assoc (black-depth t₄) (black-depth t₁) (black-depth t₃))) ⟩ -- suc (black-depth t₄ ⊔ black-depth t₁ ⊔ black-depth t₃) ∎ -- where open ≡-Reasoning -- RB-repl→eq {n} {A} .(node kg ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ _) _) .(node kn ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ t₂) (node kg ⟪ Red , _ ⟫ t₃ _)) (rb-black .kg _ x₃ (rb-red .kp _ x₄ x₅ x₆ rbt rbt₂) rbt₁) (rbr-rotate-lr {t₀} {t₁} {uncle} t₂ t₃ kg kp kn x x₁ x₂ t) = begin -- suc (black-depth t₀ ⊔ black-depth t₁ ⊔ black-depth uncle) ≡⟨ cong suc ( ⊔-assoc (black-depth t₀) (black-depth t₁) (black-depth uncle)) ⟩ -- suc (black-depth t₀ ⊔ (black-depth t₁ ⊔ black-depth uncle)) ≡⟨ cong (λ k → suc (black-depth t₀ ⊔ (k ⊔ black-depth uncle))) (RB-repl→eq _ _ rbt₂ t) ⟩ -- suc (black-depth t₀ ⊔ ((black-depth t₂ ⊔ black-depth t₃) ⊔ black-depth uncle)) ≡⟨ cong (λ k → suc (black-depth t₀ ⊔ k )) ( ⊔-assoc (black-depth t₂) (black-depth t₃) (black-depth uncle)) ⟩ -- suc (black-depth t₀ ⊔ (black-depth t₂ ⊔ (black-depth t₃ ⊔ black-depth uncle))) ≡⟨ cong suc (sym ( ⊔-assoc (black-depth t₀) (black-depth t₂) (black-depth t₃ ⊔ black-depth uncle))) ⟩ -- suc (black-depth t₀ ⊔ black-depth t₂ ⊔ (black-depth t₃ ⊔ black-depth uncle)) ∎ -- where open ≡-Reasoning -- RB-repl→eq {n} {A} .(node kg ⟪ Black , _ ⟫ _ (node kp ⟪ Red , _ ⟫ _ _)) .(node kn ⟪ Black , _ ⟫ (node kg ⟪ Red , _ ⟫ _ t₂) (node kp ⟪ Red , _ ⟫ t₃ _)) (rb-black .kg _ x₃ rbt (rb-red .kp _ x₄ x₅ x₆ rbt₁ rbt₂)) (rbr-rotate-rl {t₀} {t₁} {uncle} t₂ t₃ kg kp kn x x₁ x₂ t) = begin -- suc (black-depth uncle ⊔ (black-depth t₁ ⊔ black-depth t₀)) ≡⟨ cong (λ k → suc (black-depth uncle ⊔ (k ⊔ black-depth t₀))) (RB-repl→eq _ _ rbt₁ t) ⟩ -- suc (black-depth uncle ⊔ ((black-depth t₂ ⊔ black-depth t₃) ⊔ black-depth t₀)) ≡⟨ cong (λ k → suc (black-depth uncle ⊔ k)) ( ⊔-assoc (black-depth t₂) (black-depth t₃) (black-depth t₀)) ⟩ -- suc (black-depth uncle ⊔ (black-depth t₂ ⊔ (black-depth t₃ ⊔ black-depth t₀))) ≡⟨ cong suc (sym ( ⊔-assoc (black-depth uncle) (black-depth t₂) (black-depth t₃ ⊔ black-depth t₀))) ⟩ -- suc (black-depth uncle ⊔ black-depth t₂ ⊔ (black-depth t₃ ⊔ black-depth t₀)) ∎ -- where open ≡-Reasoning RB-repl→ti> : {n : Level} {A : Set n} → (tree repl : bt (Color ∧ A)) → (key key₁ : ℕ) → (value : A) → replacedRBTree key value tree repl → key₁ < key → tr> key₁ tree → tr> key₁ repl RB-repl→ti> = ? -- RB-repl→ti> .leaf .(node key ⟪ Red , value ⟫ leaf leaf) key key₁ value rbr-leaf lt tr = ⟪ lt , ⟪ tt , tt ⟫ ⟫ -- RB-repl→ti> .(node key ⟪ _ , _ ⟫ _ _) .(node key ⟪ _ , value ⟫ _ _) key key₁ value (rbr-node ) lt tr = tr -- RB-repl→ti> .(node _ ⟪ _ , _ ⟫ _ _) .(node _ ⟪ _ , _ ⟫ _ _) key key₁ value (rbr-right _ x rbt) lt tr -- = ⟪ proj1 tr , ⟪ proj1 (proj2 tr) , RB-repl→ti> _ _ _ _ _ rbt lt (proj2 (proj2 tr)) ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ _ , _ ⟫ _ _) .(node _ ⟪ _ , _ ⟫ _ _) key key₁ value (rbr-left _ x rbt) lt tr -- = ⟪ proj1 tr , ⟪ RB-repl→ti> _ _ _ _ _ rbt lt (proj1 (proj2 tr)) , proj2 (proj2 tr) ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ _ _) .(node _ ⟪ Black , _ ⟫ _ _) key key₁ value (rbr-black-right x _ rbt) lt tr -- = ⟪ proj1 tr , ⟪ proj1 (proj2 tr) , RB-repl→ti> _ _ _ _ _ rbt lt (proj2 (proj2 tr)) ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ _ _) .(node _ ⟪ Black , _ ⟫ _ _) key key₁ value (rbr-black-left x _ rbt) lt tr -- = ⟪ proj1 tr , ⟪ RB-repl→ti> _ _ _ _ _ rbt lt (proj1 (proj2 tr)) , proj2 (proj2 tr) ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) .(node _ ⟪ Red , _ ⟫ (node _ ⟪ Black , _ ⟫ _ _) (to-black _)) key key₁ value (rbr-flip-ll x _ _ rbt) lt tr -- = ⟪ proj1 tr , ⟪ ⟪ proj1 (proj1 (proj2 tr)) , ⟪ RB-repl→ti> _ _ _ _ _ rbt lt (proj1 (proj2 (proj1 (proj2 tr)))) -- , proj2 (proj2 (proj1 (proj2 tr))) ⟫ ⟫ , tr>-to-black (proj2 (proj2 tr)) ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) .(node _ ⟪ Red , _ ⟫ (node _ ⟪ Black , _ ⟫ _ _) (to-black _)) key key₁ value -- (rbr-flip-lr x _ _ _ rbt) lt ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ , tr5 ⟫ ⟫ = ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , RB-repl→ti> _ _ _ _ _ rbt lt tr7 ⟫ ⟫ , tr>-to-black tr5 ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Red , _ ⟫ (to-black _) (node _ ⟪ Black , _ ⟫ _ _)) key key₁ value -- (rbr-flip-rl x _ _ _ rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ tr3 , ⟪ tr>-to-black tr5 , ⟪ tr4 , ⟪ RB-repl→ti> _ _ _ _ _ rbt lt tr6 , tr7 ⟫ ⟫ ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Red , _ ⟫ (to-black _) (node _ ⟪ Black , _ ⟫ _ _)) key key₁ value -- (rbr-flip-rr x _ _ rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ tr3 , ⟪ tr>-to-black tr5 , ⟪ tr4 , ⟪ tr6 , RB-repl→ti> _ _ _ _ _ rbt lt tr7 ⟫ ⟫ ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) key key₁ value -- (rbr-rotate-ll x lt2 rbt) lt ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ , tr5 ⟫ ⟫ = ⟪ tr4 , ⟪ RB-repl→ti> _ _ _ _ _ rbt lt tr6 , ⟪ tr3 , ⟪ tr7 , tr5 ⟫ ⟫ ⟫ ⟫ -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) key key₁ value -- (rbr-rotate-rr x lt2 rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ tr4 , ⟪ ⟪ tr3 , ⟪ tr5 , tr6 ⟫ ⟫ , RB-repl→ti> _ _ _ _ _ rbt lt tr7 ⟫ ⟫ -- RB-repl→ti> (node kg ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ _) _) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node _ ⟪ Red , _ ⟫ _ _)) key key₁ value -- (rbr-rotate-lr left right _ _ kn _ _ _ rbt) lt ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ , tr5 ⟫ ⟫ = ⟪ rr00 , ⟪ ⟪ tr4 , ⟪ tr6 , proj1 (proj2 rr01) ⟫ ⟫ , ⟪ tr3 , ⟪ proj2 (proj2 rr01) , tr5 ⟫ ⟫ ⟫ ⟫ where -- rr01 : (key₁ < kn) ∧ tr> key₁ left ∧ tr> key₁ right -- rr01 = RB-repl→ti> _ _ _ _ _ rbt lt tr7 -- rr00 : key₁ < kn -- rr00 = proj1 rr01 -- RB-repl→ti> .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node _ ⟪ Red , _ ⟫ _ _)) key key₁ value -- (rbr-rotate-rl left right kg kp kn _ _ _ rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ rr00 , ⟪ ⟪ tr3 , ⟪ tr5 , proj1 (proj2 rr01) ⟫ ⟫ , ⟪ tr4 , ⟪ proj2 (proj2 rr01) , tr7 ⟫ ⟫ ⟫ ⟫ where -- rr01 : (key₁ < kn) ∧ tr> key₁ left ∧ tr> key₁ right -- rr01 = RB-repl→ti> _ _ _ _ _ rbt lt tr6 -- rr00 : key₁ < kn -- rr00 = proj1 rr01 RB-repl→ti< : {n : Level} {A : Set n} → (tree repl : bt (Color ∧ A)) → (key key₁ : ℕ) → (value : A) → replacedRBTree key value tree repl → key < key₁ → tr< key₁ tree → tr< key₁ repl RB-repl→ti< = ? -- RB-repl→ti< .leaf .(node key ⟪ Red , value ⟫ leaf leaf) key key₁ value rbr-leaf lt tr = ⟪ lt , ⟪ tt , tt ⟫ ⟫ -- RB-repl→ti< .(node key ⟪ _ , _ ⟫ _ _) .(node key ⟪ _ , value ⟫ _ _) key key₁ value (rbr-node ) lt tr = tr -- RB-repl→ti< .(node _ ⟪ _ , _ ⟫ _ _) .(node _ ⟪ _ , _ ⟫ _ _) key key₁ value (rbr-right _ x rbt) lt tr -- = ⟪ proj1 tr , ⟪ proj1 (proj2 tr) , RB-repl→ti< _ _ _ _ _ rbt lt (proj2 (proj2 tr)) ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ _ , _ ⟫ _ _) .(node _ ⟪ _ , _ ⟫ _ _) key key₁ value (rbr-left _ x rbt) lt tr -- = ⟪ proj1 tr , ⟪ RB-repl→ti< _ _ _ _ _ rbt lt (proj1 (proj2 tr)) , proj2 (proj2 tr) ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ _ _) .(node _ ⟪ Black , _ ⟫ _ _) key key₁ value (rbr-black-right x _ rbt) lt tr -- = ⟪ proj1 tr , ⟪ proj1 (proj2 tr) , RB-repl→ti< _ _ _ _ _ rbt lt (proj2 (proj2 tr)) ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ _ _) .(node _ ⟪ Black , _ ⟫ _ _) key key₁ value (rbr-black-left x _ rbt) lt tr -- = ⟪ proj1 tr , ⟪ RB-repl→ti< _ _ _ _ _ rbt lt (proj1 (proj2 tr)) , proj2 (proj2 tr) ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) .(node _ ⟪ Red , _ ⟫ (node _ ⟪ Black , _ ⟫ _ _) (to-black _)) key key₁ value (rbr-flip-ll x _ _ rbt) lt tr -- = ⟪ proj1 tr , ⟪ ⟪ proj1 (proj1 (proj2 tr)) , ⟪ RB-repl→ti< _ _ _ _ _ rbt lt (proj1 (proj2 (proj1 (proj2 tr)))) -- , proj2 (proj2 (proj1 (proj2 tr))) ⟫ ⟫ , tr<-to-black (proj2 (proj2 tr)) ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) .(node _ ⟪ Red , _ ⟫ (node _ ⟪ Black , _ ⟫ _ _) (to-black _)) key key₁ value -- (rbr-flip-lr x _ _ _ rbt) lt ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ , tr5 ⟫ ⟫ = ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , RB-repl→ti< _ _ _ _ _ rbt lt tr7 ⟫ ⟫ , tr<-to-black tr5 ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Red , _ ⟫ (to-black _) (node _ ⟪ Black , _ ⟫ _ _)) key key₁ value -- (rbr-flip-rl x _ _ _ rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ tr3 , ⟪ tr<-to-black tr5 , ⟪ tr4 , ⟪ RB-repl→ti< _ _ _ _ _ rbt lt tr6 , tr7 ⟫ ⟫ ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Red , _ ⟫ (to-black _) (node _ ⟪ Black , _ ⟫ _ _)) key key₁ value -- (rbr-flip-rr x _ _ rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ tr3 , ⟪ tr<-to-black tr5 , ⟪ tr4 , ⟪ tr6 , RB-repl→ti< _ _ _ _ _ rbt lt tr7 ⟫ ⟫ ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) key key₁ value -- (rbr-rotate-ll x lt2 rbt) lt ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ , tr5 ⟫ ⟫ = ⟪ tr4 , ⟪ RB-repl→ti< _ _ _ _ _ rbt lt tr6 , ⟪ tr3 , ⟪ tr7 , tr5 ⟫ ⟫ ⟫ ⟫ -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) _) key key₁ value -- (rbr-rotate-rr x lt2 rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ tr4 , ⟪ ⟪ tr3 , ⟪ tr5 , tr6 ⟫ ⟫ , RB-repl→ti< _ _ _ _ _ rbt lt tr7 ⟫ ⟫ -- RB-repl→ti< (node kg ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ _) _) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node _ ⟪ Red , _ ⟫ _ _)) key key₁ value -- (rbr-rotate-lr left right _ _ kn _ _ _ rbt) lt ⟪ tr3 , ⟪ ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ , tr5 ⟫ ⟫ = ⟪ rr00 , ⟪ ⟪ tr4 , ⟪ tr6 , proj1 (proj2 rr01) ⟫ ⟫ , ⟪ tr3 , ⟪ proj2 (proj2 rr01) , tr5 ⟫ ⟫ ⟫ ⟫ where -- rr01 : (kn < key₁ ) ∧ tr< key₁ left ∧ tr< key₁ right -- rr01 = RB-repl→ti< _ _ _ _ _ rbt lt tr7 -- rr00 : kn < key₁ -- rr00 = proj1 rr01 -- RB-repl→ti< .(node _ ⟪ Black , _ ⟫ _ (node _ ⟪ Red , _ ⟫ _ _)) .(node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node _ ⟪ Red , _ ⟫ _ _)) key key₁ value -- (rbr-rotate-rl left right kg kp kn _ _ _ rbt) lt ⟪ tr3 , ⟪ tr5 , ⟪ tr4 , ⟪ tr6 , tr7 ⟫ ⟫ ⟫ ⟫ = ⟪ rr00 , ⟪ ⟪ tr3 , ⟪ tr5 , proj1 (proj2 rr01) ⟫ ⟫ , ⟪ tr4 , ⟪ proj2 (proj2 rr01) , tr7 ⟫ ⟫ ⟫ ⟫ where -- rr01 : (kn < key₁ ) ∧ tr< key₁ left ∧ tr< key₁ right -- rr01 = RB-repl→ti< _ _ _ _ _ rbt lt tr6 -- rr00 : kn < key₁ -- rr00 = proj1 rr01 RB-repl→ti : {n : Level} {A : Set n} → (tree repl : bt (Color ∧ A) ) → (key : ℕ ) → (value : A) → treeInvariant tree → replacedRBTree key value tree repl → treeInvariant repl RB-repl→ti = ? -- RB-repl→ti .leaf .(node key ⟪ Red , value ⟫ leaf leaf) key value ti rbr-leaf = t-single key ⟪ Red , value ⟫ -- RB-repl→ti .(node key ⟪ _ , _ ⟫ leaf leaf) .(node key ⟪ _ , value ⟫ leaf leaf) key value (t-single .key .(⟪ _ , _ ⟫)) (rbr-node ) = t-single key ⟪ _ , value ⟫ -- RB-repl→ti .(node key ⟪ _ , _ ⟫ leaf (node key₁ _ _ _)) .(node key ⟪ _ , value ⟫ leaf (node key₁ _ _ _)) key value -- (t-right .key key₁ x x₁ x₂ ti) (rbr-node ) = t-right key key₁ x x₁ x₂ ti -- RB-repl→ti .(node key ⟪ _ , _ ⟫ (node key₁ _ _ _) leaf) .(node key ⟪ _ , value ⟫ (node key₁ _ _ _) leaf) key value -- (t-left key₁ .key x x₁ x₂ ti) (rbr-node ) = t-left key₁ key x x₁ x₂ ti -- RB-repl→ti .(node key ⟪ _ , _ ⟫ (node key₁ _ _ _) (node key₂ _ _ _)) .(node key ⟪ _ , value ⟫ (node key₁ _ _ _) (node key₂ _ _ _)) key value -- (t-node key₁ .key key₂ x x₁ x₂ x₃ x₄ x₅ ti ti₁) (rbr-node ) = t-node key₁ key key₂ x x₁ x₂ x₃ x₄ x₅ ti ti₁ -- RB-repl→ti (node key₁ ⟪ ca , v1 ⟫ leaf leaf) (node key₁ ⟪ ca , v1 ⟫ leaf tree@(node key₂ value₁ t t₁)) key value -- (t-single key₁ ⟪ ca , v1 ⟫) (rbr-right _ x trb) = t-right _ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₁ < key₂ ) ∧ tr> key₁ t ∧ tr> key₁ t₁ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x tt -- RB-repl→ti (node _ ⟪ _ , _ ⟫ leaf (node key₁ _ _ _)) (node key₂ ⟪ ca , v1 ⟫ leaf (node key₃ value₁ t t₁)) key value -- (t-right _ key₁ x₁ x₂ x₃ ti) (rbr-right _ x trb) = t-right _ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ ti trb) where -- rr00 : (key₂ < key₃) ∧ tr> key₂ t ∧ tr> key₂ t₁ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x ⟪ x₁ , ⟪ x₂ , x₃ ⟫ ⟫ -- RB-repl→ti .(node key₂ ⟪ ca , v1 ⟫ (node key₁ value₁ t t₁) leaf) (node key₂ ⟪ ca , v1 ⟫ (node key₁ value₁ t t₁) (node key₃ value₂ t₂ t₃)) key value -- (t-left key₁ _ x₁ x₂ x₃ ti) (rbr-right _ x trb) = t-node _ _ _ x₁ (proj1 rr00) x₂ x₃ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) ti (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₂ < key₃) ∧ tr> key₂ t₂ ∧ tr> key₂ t₃ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x tt -- RB-repl→ti .(node key₃ ⟪ ca , v1 ⟫ (node key₁ v2 t₁ t₂) (node key₂ _ _ _)) (node key₃ ⟪ ca , v1 ⟫ (node key₁ v2 t₁ t₂) (node key₄ value₁ t₃ t₄)) key value -- (t-node key₁ _ key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-right _ x trb) = t-node _ _ _ x₁ -- (proj1 rr00) x₃ x₄ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) ti (RB-repl→ti _ _ _ _ ti₁ trb) where -- rr00 : (key₃ < key₄) ∧ tr> key₃ t₃ ∧ tr> key₃ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x ⟪ x₂ , ⟪ x₅ , x₆ ⟫ ⟫ -- RB-repl→ti .(node key₁ ⟪ _ , _ ⟫ leaf leaf) (node key₁ ⟪ _ , _ ⟫ (node key₂ value₁ left left₁) leaf) key value -- (t-single _ .(⟪ _ , _ ⟫)) (rbr-left _ x trb) = t-left _ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₂ < key₁) ∧ tr< key₁ left ∧ tr< key₁ left₁ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x tt -- RB-repl→ti .(node key₂ ⟪ _ , _ ⟫ leaf (node key₁ _ t₁ t₂)) (node key₂ ⟪ _ , _ ⟫ (node key₃ value₁ t t₃) (node key₁ _ t₁ t₂)) key value -- (t-right _ key₁ x₁ x₂ x₃ ti) (rbr-left _ x trb) = t-node _ _ _ (proj1 rr00) x₁ (proj1 (proj2 rr00))(proj2 (proj2 rr00)) x₂ x₃ rr01 ti where -- rr00 : (key₃ < key₂) ∧ tr< key₂ t ∧ tr< key₂ t₃ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x tt -- rr01 : treeInvariant (node key₃ value₁ t t₃) -- rr01 = RB-repl→ti _ _ _ _ t-leaf trb -- RB-repl→ti .(node _ ⟪ _ , _ ⟫ (node key₁ _ _ _) leaf) (node key₃ ⟪ _ , _ ⟫ (node key₂ value₁ t t₁) leaf) key value -- (t-left key₁ _ x₁ x₂ x₃ ti) (rbr-left _ x trb) = t-left key₂ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ ti trb) where -- rr00 : (key₂ < key₃) ∧ tr< key₃ t ∧ tr< key₃ t₁ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x ⟪ x₁ , ⟪ x₂ , x₃ ⟫ ⟫ -- RB-repl→ti .(node key₃ ⟪ _ , _ ⟫ (node key₁ _ _ _) (node key₂ _ t₁ t₂)) (node key₃ ⟪ _ , _ ⟫ (node key₄ value₁ t t₃) (node key₂ _ t₁ t₂)) key value -- (t-node key₁ _ key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-left _ x trb) = t-node _ _ _ (proj1 rr00) x₂ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) x₅ x₆ (RB-repl→ti _ _ _ _ ti trb) ti₁ where -- rr00 : (key₄ < key₃) ∧ tr< key₃ t ∧ tr< key₃ t₃ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x ⟪ x₁ , ⟪ x₃ , x₄ ⟫ ⟫ -- RB-repl→ti .(node x₁ ⟪ Black , c ⟫ leaf leaf) (node x₁ ⟪ Black , c ⟫ leaf (node key₁ value₁ t t₁)) key value -- (t-single x₂ .(⟪ Black , c ⟫)) (rbr-black-right x x₄ trb) = t-right _ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (x₁ < key₁) ∧ tr> x₁ t ∧ tr> x₁ t₁ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x₄ tt -- RB-repl→ti .(node key₂ ⟪ Black , _ ⟫ leaf (node key₁ _ _ _)) (node key₂ ⟪ Black , _ ⟫ leaf (node key₃ value₁ t₂ t₃)) key value -- (t-right _ key₁ x₁ x₂ x₃ ti) (rbr-black-right x x₄ trb) = t-right _ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ ti trb) where -- rr00 : (key₂ < key₃) ∧ tr> key₂ t₂ ∧ tr> key₂ t₃ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x₄ ⟪ x₁ , ⟪ x₂ , x₃ ⟫ ⟫ -- RB-repl→ti .(node key₂ ⟪ Black , _ ⟫ (node key₁ _ _ _) leaf) (node key₂ ⟪ Black , _ ⟫ (node key₁ _ _ _) (node key₃ value₁ t₂ t₃)) key value (t-left key₁ _ x₁ x₂ x₃ ti) (rbr-black-right x x₄ trb) = t-node _ _ _ x₁ (proj1 rr00) x₂ x₃ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) ti (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₂ < key₃) ∧ tr> key₂ t₂ ∧ tr> key₂ t₃ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x₄ tt -- RB-repl→ti .(node key₃ ⟪ Black , _ ⟫ (node key₁ _ _ _) (node key₂ _ _ _)) (node key₃ ⟪ Black , _ ⟫ (node key₁ _ _ _) (node key₄ value₁ t₂ t₃)) key value -- (t-node key₁ _ key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-black-right x x₇ trb) = t-node _ _ _ x₁ (proj1 rr00) x₃ x₄ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) ti (RB-repl→ti _ _ _ _ ti₁ trb) where -- rr00 : (key₃ < key₄) ∧ tr> key₃ t₂ ∧ tr> key₃ t₃ -- rr00 = RB-repl→ti> _ _ _ _ _ trb x₇ ⟪ x₂ , ⟪ x₅ , x₆ ⟫ ⟫ -- RB-repl→ti .(node key₂ ⟪ Black , _ ⟫ leaf leaf) (node key₂ ⟪ Black , _ ⟫ (node key₁ value₁ t t₁) .leaf) key value -- (t-single .key₂ .(⟪ Black , _ ⟫)) (rbr-black-left x x₇ trb) = t-left _ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₁ < key₂) ∧ tr< key₂ t ∧ tr< key₂ t₁ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x₇ tt -- RB-repl→ti .(node key₂ ⟪ Black , _ ⟫ leaf (node key₁ _ _ _)) (node key₂ ⟪ Black , _ ⟫ (node key₃ value₁ t t₁) .(node key₁ _ _ _)) key value -- (t-right .key₂ key₁ x₁ x₂ x₃ ti) (rbr-black-left x x₇ trb) = t-node _ _ _ (proj1 rr00) x₁ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) x₂ x₃ (RB-repl→ti _ _ _ _ t-leaf trb) ti where -- rr00 : (key₃ < key₂) ∧ tr< key₂ t ∧ tr< key₂ t₁ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x₇ tt -- RB-repl→ti .(node key₂ ⟪ Black , _ ⟫ (node key₁ _ _ _) leaf) (node key₂ ⟪ Black , _ ⟫ (node key₃ value₁ t t₁) .leaf) key value -- (t-left key₁ .key₂ x₁ x₂ x₃ ti) (rbr-black-left x x₇ trb) = t-left _ _ (proj1 rr00) (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RB-repl→ti _ _ _ _ ti trb) where -- rr00 : (key₃ < key₂) ∧ tr< key₂ t ∧ tr< key₂ t₁ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x₇ ⟪ x₁ , ⟪ x₂ , x₃ ⟫ ⟫ -- RB-repl→ti .(node key₂ ⟪ Black , _ ⟫ (node key₁ _ _ _) (node key₃ _ _ _)) (node key₂ ⟪ Black , _ ⟫ (node key₄ value₁ t t₁) .(node key₃ _ _ _)) key value -- (t-node key₁ .key₂ key₃ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-black-left x x₇ trb) = t-node _ _ _ (proj1 rr00) x₂ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) x₅ x₆ (RB-repl→ti _ _ _ _ ti trb) ti₁ where -- rr00 : (key₄ < key₂) ∧ tr< key₂ t ∧ tr< key₂ t₁ -- rr00 = RB-repl→ti< _ _ _ _ _ trb x₇ ⟪ x₁ , ⟪ x₃ , x₄ ⟫ ⟫ -- RB-repl→ti .(node key₂ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ t₁) leaf) (node key₂ ⟪ Red , value₁ ⟫ (node key₁ ⟪ Black , value₂ ⟫ t t₁) .(to-black leaf)) key value -- (t-left _ .key₂ x₁ x₂ x₃ ti) (rbr-flip-ll x _ lt trb) = t-left _ _ x₁ rr00 x₃ (RTtoTI0 _ _ _ _ rr02 r-node ) where -- rr00 : tr< key₂ t -- rr00 = RB-repl→ti< _ _ _ _ _ trb (<-trans lt x₁) x₂ -- rr02 : treeInvariant (node key₁ ⟪ Red , value₂ ⟫ t t₁) -- rr02 = RB-repl→ti _ _ _ _ ti (rbr-left _ lt trb) -- RB-repl→ti (node key₂ ⟪ Black , _ ⟫ (node key₁ ⟪ Red , _ ⟫ t₀ t₁) (node key₃ ⟪ c1 , v1 ⟫ left right)) (node key₂ ⟪ Red , value₁ ⟫ (node _ ⟪ Black , value₂ ⟫ t t₁) (node key₃ ⟪ Black , v1 ⟫ left right)) key value -- (t-node _ .key₂ key₃ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-flip-ll x _ lt trb) = t-node _ _ _ x₁ x₂ rr00 x₄ x₅ x₆ (RTtoTI0 _ _ _ _ rr02 r-node ) (RTtoTI0 _ _ _ _ ti₁ r-node ) where -- rr00 : tr< key₂ t -- rr00 = RB-repl→ti< _ _ _ _ _ trb (<-trans lt x₁) x₃ -- rr02 : treeInvariant (node key₁ ⟪ Red , value₂ ⟫ t t₁) -- rr02 = RB-repl→ti _ _ _ _ ti (rbr-left _ lt trb) -- RB-repl→ti (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ left right) leaf) (node key₂ ⟪ Red , v1 ⟫ (node key₃ ⟪ Black , v2 ⟫ left right₁) leaf) key value -- (t-left _ _ x₁ x₂ x₃ ti) (rbr-flip-lr x _ lt lt2 trb) = t-left _ _ x₁ x₂ rr00 (RTtoTI0 _ _ _ _ rr02 r-node ) where -- rr00 : tr< key₂ right₁ -- rr00 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- rr02 : treeInvariant (node key₃ ⟪ Red , v2 ⟫ left right₁ ) -- rr02 = RB-repl→ti _ _ _ _ ti (rbr-right _ lt trb) -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ (node key₂ ⟪ Red , v2 ⟫ t t₁) (node key₃ ⟪ c3 , v3 ⟫ t₂ t₃)) (node key₁ ⟪ Red , _ ⟫ (node _ ⟪ Black , _ ⟫ t t₄) .(to-black (node key₃ ⟪ c3 , _ ⟫ _ _))) key value -- (t-node _ _ key₃ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-flip-lr x _ lt lt2 trb) = t-node _ _ _ x₁ x₂ x₃ rr00 x₅ x₆ (RTtoTI0 _ _ _ _ rr02 r-node ) (RTtoTI0 _ _ _ _ ti₁ r-node ) where -- rr00 : tr< key₁ t₄ -- rr00 = RB-repl→ti< _ _ _ _ _ trb lt2 x₄ -- rr02 : treeInvariant (node key₂ ⟪ Red , v2 ⟫ t t₄) -- rr02 = RB-repl→ti _ _ _ _ ti (rbr-right _ lt trb) -- RB-repl→ti (node _ ⟪ Black , _ ⟫ leaf (node _ ⟪ Red , _ ⟫ t t₁)) (node key₁ ⟪ Red , v1 ⟫ .(to-black leaf) (node key₂ ⟪ Black , v2 ⟫ t₂ t₁)) key value -- (t-right _ _ x₁ x₂ x₃ ti) (rbr-flip-rl x _ lt lt2 trb) = t-right _ _ x₁ rr00 x₃ (RTtoTI0 _ _ _ _ rr02 r-node ) where -- rr00 : tr> key₁ t₂ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt x₂ -- rr02 : treeInvariant (node key₂ ⟪ Red , v2 ⟫ t₂ t₁) -- rr02 = RB-repl→ti _ _ _ _ ti (rbr-left _ lt2 trb) -- RB-repl→ti (node _ ⟪ Black , v1 ⟫ (node key₂ ⟪ c2 , v2 ⟫ t t₁) (node _ ⟪ Red , v3 ⟫ t₂ t₃)) (node key₁ ⟪ Red , _ ⟫ .(to-black (node key₂ ⟪ c2 , _ ⟫ _ _)) (node key₃ ⟪ Black , _ ⟫ t₄ t₃)) key value -- (t-node key₂ _ _ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-flip-rl x _ lt lt2 trb) = t-node key₂ _ _ x₁ x₂ x₃ x₄ rr00 x₆ (RTtoTI0 _ _ _ _ ti r-node ) (RTtoTI0 _ _ _ _ rr02 r-node ) where -- rr00 : tr> key₁ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt x₅ -- rr02 : treeInvariant (node key₃ ⟪ Red , v3 ⟫ t₄ t₃) -- rr02 = RB-repl→ti _ _ _ _ ti₁ (rbr-left _ lt2 trb) -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ leaf (node key₂ ⟪ Red , v2 ⟫ t t₁)) (node _ ⟪ Red , _ ⟫ .(to-black leaf) (node _ ⟪ Black , v2 ⟫ t t₂)) key value -- (t-right _ _ x₁ x₂ x₃ ti) (rbr-flip-rr x _ lt trb) = t-right _ _ x₁ x₂ rr00 (RTtoTI0 _ _ _ _ rr02 r-node ) where -- rr00 : tr> key₁ t₂ -- rr00 = RB-repl→ti> _ _ _ _ _ trb (<-trans x₁ lt ) x₃ -- rr02 : treeInvariant (node key₂ ⟪ Red , v2 ⟫ t t₂) -- rr02 = RB-repl→ti _ _ _ _ ti (rbr-right _ lt trb) -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ (node key₂ ⟪ c2 , v2 ⟫ t t₁) (node key₃ ⟪ Red , c3 ⟫ t₂ t₃)) (node _ ⟪ Red , _ ⟫ .(to-black (node key₂ ⟪ c2 , _ ⟫ _ _)) (node _ ⟪ Black , c3 ⟫ t₂ t₄)) key value -- (t-node key₂ _ _ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (rbr-flip-rr x _ lt trb) = t-node key₂ _ _ x₁ x₂ x₃ x₄ x₅ rr00 (RTtoTI0 _ _ _ _ ti r-node ) (RTtoTI0 _ _ _ _ rr02 r-node ) where -- rr00 : tr> key₁ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb (<-trans x₂ lt) x₆ -- rr02 : treeInvariant (node key₃ ⟪ Red , c3 ⟫ t₂ t₄) -- rr02 = RB-repl→ti _ _ _ _ ti₁ (rbr-right _ lt trb) -- RB-repl→ti {_} {A} (node k1 ⟪ Black , c1 ⟫ (node k2 ⟪ Red , c2 ⟫ .leaf .leaf) leaf) (node _ ⟪ Black , _ ⟫ (node key₁ value₁ t₂ t₃) (node _ ⟪ Red , _ ⟫ .leaf leaf)) key value (t-left _ _ x₁ x₂ x₃ -- (t-single .k2 .(⟪ Red , c2 ⟫))) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr10) x₁ (proj1 (proj2 rr10)) (proj2 (proj2 rr10)) tt tt (RB-repl→ti _ _ _ _ t-leaf trb) (t-single _ _ ) where -- rr10 : (key₁ < k2 ) ∧ tr< k2 t₂ ∧ tr< k2 t₃ -- rr10 = RB-repl→ti< _ _ _ _ _ trb lt tt -- RB-repl→ti {_} {A} (node k1 ⟪ Black , c1 ⟫ (node k2 ⟪ Red , c2 ⟫ .leaf .(node key₂ _ _ _)) leaf) (node _ ⟪ Black , _ ⟫ (node key₁ value₁ t₂ t₃) (node _ ⟪ Red , _ ⟫ (node key₂ value₂ t₁ t₄) leaf)) key value -- (t-left _ _ x₁ x₂ x₃ (t-right .k2 key₂ x₄ x₅ x₆ ti)) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr10) x₁ (proj1 (proj2 rr10)) (proj2 (proj2 rr10)) ⟪ x₄ , ⟪ x₅ , x₆ ⟫ ⟫ tt rr05 rr04 where -- rr10 : (key₁ < k2 ) ∧ tr< k2 t₂ ∧ tr< k2 t₃ -- rr10 = RB-repl→ti< _ _ _ _ _ trb lt tt -- rr04 : treeInvariant (node k1 ⟪ Red , c1 ⟫ (node key₂ value₂ t₁ t₄) leaf) -- rr04 = RTtoTI0 _ _ _ _ (t-left key₂ _ {_} {⟪ Red , c1 ⟫} {t₁} {t₄} (proj1 x₃) (proj1 (proj2 x₃)) (proj2 (proj2 x₃)) ti) r-node -- rr05 : treeInvariant (node key₁ value₁ t₂ t₃) -- rr05 = RB-repl→ti _ _ _ _ t-leaf trb -- RB-repl→ti {_} {A} (node k1 ⟪ Black , c1 ⟫ (node k2 ⟪ Red , c2 ⟫ (node key₂ value₂ t₁ t₄) .leaf) leaf) (node _ ⟪ Black , _ ⟫ (node key₁ value₁ t₂ t₃) (node _ ⟪ Red , _ ⟫ .leaf leaf)) key value -- (t-left _ _ x₁ x₂ x₃ (t-left key₂ .k2 x₄ x₅ x₆ ti)) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr10) x₁ (proj1 (proj2 rr10)) (proj2 (proj2 rr10)) tt tt (RB-repl→ti _ _ _ _ ti trb) (t-single _ _) where -- rr10 : (key₁ < k2 ) ∧ tr< k2 t₂ ∧ tr< k2 t₃ -- rr10 = RB-repl→ti< _ _ _ _ _ trb lt ⟪ x₄ , ⟪ x₅ , x₆ ⟫ ⟫ -- RB-repl→ti {_} {A} (node k1 ⟪ Black , c1 ⟫ (node k2 ⟪ Red , c2 ⟫ (node key₂ value₃ left right) (node key₃ value₂ t₄ t₅)) leaf) (node _ ⟪ Black , _ ⟫ (node key₁ value₁ t₂ t₃) (node _ ⟪ Red , _ ⟫ .(node key₃ _ _ _) leaf)) key value -- (t-left _ _ x₁ x₂ x₃ (t-node key₂ .k2 key₃ x₄ x₅ x₆ x₇ x₈ x₉ ti ti₁)) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr10) x₁ (proj1 (proj2 rr10)) (proj2 (proj2 rr10)) ⟪ x₅ , ⟪ x₈ , x₉ ⟫ ⟫ tt rr05 rr04 where -- rr06 : key < k2 -- rr06 = lt -- rr10 : (key₁ < k2) ∧ tr< k2 t₂ ∧ tr< k2 t₃ -- rr10 = RB-repl→ti< _ _ _ _ _ trb rr06 ⟪ x₄ , ⟪ x₆ , x₇ ⟫ ⟫ -- rr04 : treeInvariant (node k1 ⟪ Red , c1 ⟫ (node key₃ value₂ t₄ t₅) leaf) -- rr04 = RTtoTI0 _ _ _ _ (t-left _ _ (proj1 x₃) (proj1 (proj2 x₃)) (proj2 (proj2 x₃)) ti₁ ) (r-left (proj1 x₃) r-node) -- rr05 : treeInvariant (node key₁ value₁ t₂ t₃) -- rr05 = RB-repl→ti _ _ _ _ ti trb -- RB-repl→ti (node key₁ ⟪ Black , c1 ⟫ (node key₂ ⟪ Red , c2 ⟫ .leaf .leaf) (node key₃ v3 t₂ t₃)) (node _ ⟪ Black , _ ⟫ (node key₄ value₁ t₄ t₅) (node _ ⟪ Red , _ ⟫ .leaf (node key₃ _ _ _))) key value -- (t-node _ _ key₃ x₁ x₂ x₃ x₄ x₅ x₆ (t-single .key₂ .(⟪ Red , c2 ⟫)) ti₁) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr00) x₁ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) tt ⟪ <-trans x₁ x₂ , ⟪ <-tr> x₅ x₁ , <-tr> x₆ x₁ ⟫ ⟫ rr02 rr03 where -- rr00 : (key₄ < key₂) ∧ tr< key₂ t₄ ∧ tr< key₂ t₅ -- rr00 = RB-repl→ti< _ _ _ _ _ trb lt tt -- rr02 : treeInvariant (node key₄ value₁ t₄ t₅) -- rr02 = RB-repl→ti _ _ _ _ t-leaf trb -- rr03 : treeInvariant (node key₁ ⟪ Red , c1 ⟫ leaf (node key₃ v3 t₂ t₃)) -- rr03 = RTtoTI0 _ _ _ _ (t-right _ _ {v3} {_} x₂ x₅ x₆ ti₁) r-node -- RB-repl→ti (node key₁ ⟪ Black , c1 ⟫ (node key₂ ⟪ Red , c2 ⟫ leaf (node key₅ _ _ _)) (node key₃ v3 t₂ t₃)) (node _ ⟪ Black , _ ⟫ (node key₄ value₁ t₄ t₅) (node _ ⟪ Red , _ ⟫ (node key₅ value₂ t₁ t₆) (node key₃ _ _ _))) key value -- (t-node _ _ key₃ x₁ x₂ x₃ x₄ x₅ x₆ (t-right .key₂ key₅ x₇ x₈ x₉ ti) ti₁) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr00) x₁ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) ⟪ x₇ , ⟪ x₈ , x₉ ⟫ ⟫ ⟪ <-trans x₁ x₂ , ⟪ <-tr> x₅ x₁ , <-tr> x₆ x₁ ⟫ ⟫ rr02 rr03 where -- rr00 : (key₄ < key₂) ∧ tr< key₂ t₄ ∧ tr< key₂ t₅ -- rr00 = RB-repl→ti< _ _ _ _ _ trb lt tt -- rr02 : treeInvariant (node key₄ value₁ t₄ t₅) -- rr02 = RB-repl→ti _ _ _ _ t-leaf trb -- rr03 : treeInvariant (node key₁ ⟪ Red , c1 ⟫ (node key₅ value₂ t₁ t₆) (node key₃ v3 t₂ t₃)) -- rr03 = RTtoTI0 _ _ _ _ (t-node _ _ _ {_} {v3} {_} {_} {_} {_} {_} (proj1 x₄) x₂ (proj1 (proj2 x₄)) (proj2 (proj2 x₄)) x₅ x₆ ti ti₁ ) r-node -- RB-repl→ti (node key₁ ⟪ Black , c1 ⟫ (node key₂ ⟪ Red , c2 ⟫ .(node key₅ _ _ _) .leaf) (node key₃ v3 t₂ t₃)) (node _ ⟪ Black , _ ⟫ (node key₄ value₁ t₄ t₅) (node _ ⟪ Red , _ ⟫ .leaf (node key₃ _ _ _))) key value (t-node _ _ key₃ x₁ x₂ x₃ x₄ x₅ x₆ -- (t-left key₅ .key₂ x₇ x₈ x₉ ti) ti₁) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr00) x₁ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) tt ⟪ <-trans x₁ x₂ , ⟪ <-tr> x₅ x₁ , <-tr> x₆ x₁ ⟫ ⟫ rr02 rr04 where -- rr00 : (key₄ < key₂) ∧ tr< key₂ t₄ ∧ tr< key₂ t₅ -- rr00 = RB-repl→ti< _ _ _ _ _ trb lt ⟪ x₇ , ⟪ x₈ , x₉ ⟫ ⟫ -- rr02 : treeInvariant (node key₄ value₁ t₄ t₅) -- rr02 = RB-repl→ti _ _ _ _ ti trb -- rr03 : treeInvariant (node key₁ ⟪ Red , c1 ⟫ (node key₅ _ _ _) (node key₃ v3 t₂ t₃)) -- rr03 = RTtoTI0 _ _ _ _ (t-node _ _ _ {_} {v3} {_} {_} {_} {_} {_} (proj1 x₃) x₂ (proj1 (proj2 x₃)) (proj2 (proj2 x₃)) x₅ x₆ ti ti₁) r-node -- rr04 : treeInvariant (node key₁ ⟪ Red , c1 ⟫ leaf (node key₃ v3 t₂ t₃)) -- rr04 = RTtoTI0 _ _ _ _ (t-right _ _ {v3} {_} x₂ x₅ x₆ ti₁) r-node -- RB-repl→ti {_} {A} (node key₁ ⟪ Black , c1 ⟫ (node key₂ ⟪ Red , c2 ⟫ .(node key₅ _ _ _) (node key₆ value₆ t₆ t₇)) (node key₃ v3 t₂ t₃)) (node _ ⟪ Black , _ ⟫ (node key₄ value₁ t₄ t₅) (node _ ⟪ Red , _ ⟫ .(node key₆ _ _ _) (node key₃ _ _ _))) key value -- (t-node _ _ key₃ x₁ x₂ x₃ x₄ x₅ x₆ (t-node key₅ .key₂ key₆ x₇ x₈ x₉ x₁₀ x₁₁ x₁₂ ti ti₂) ti₁) (rbr-rotate-ll x lt trb) = t-node _ _ _ (proj1 rr00) x₁ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) ⟪ x₈ , ⟪ x₁₁ , x₁₂ ⟫ ⟫ ⟪ <-trans x₁ x₂ , ⟪ rr05 , <-tr> x₆ x₁ ⟫ ⟫ rr02 rr03 where -- rr00 : (key₄ < key₂) ∧ tr< key₂ t₄ ∧ tr< key₂ t₅ -- rr00 = RB-repl→ti< _ _ _ _ _ trb lt ⟪ x₇ , ⟪ x₉ , x₁₀ ⟫ ⟫ -- rr02 : treeInvariant (node key₄ value₁ t₄ t₅) -- rr02 = RB-repl→ti _ _ _ _ ti trb -- rr03 : treeInvariant (node key₁ ⟪ Red , c1 ⟫ (node key₆ value₆ t₆ t₇) (node key₃ v3 t₂ t₃)) -- rr03 = RTtoTI0 _ _ _ _(t-node _ _ _ {_} {value₁} {_} {_} {_} {_} {_} (proj1 x₄) x₂ (proj1 (proj2 x₄)) (proj2 (proj2 x₄)) x₅ x₆ ti₂ ti₁) r-node -- rr05 : tr> key₂ t₂ -- rr05 = <-tr> x₅ x₁ -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ leaf (node key₂ ⟪ Red , v2 ⟫ leaf leaf)) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₃ value₁ t₃ t₄)) key value -- (t-right .key₁ .key₂ x₁ x₂ x₃ (t-single .key₂ .(⟪ Red , _ ⟫))) (rbr-rotate-rr x lt trb) -- = t-node _ _ _ x₁ (proj1 rr00) tt tt (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (t-single _ _) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₂ < key₃) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt tt -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ leaf (node key₂ ⟪ Red , v2 ⟫ leaf (node key₃ value₃ t t₁))) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₄ value₁ t₃ t₄)) key value -- (t-right .key₁ .key₂ x₁ x₂ x₃ (t-right .key₂ key₃ x₄ x₅ x₆ ti)) (rbr-rotate-rr x lt trb) -- = t-node _ _ _ x₁ (proj1 rr00) tt tt (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (t-single _ _ ) (RB-repl→ti _ _ _ _ ti trb) where -- rr00 : (key₂ < key₄) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt ⟪ x₄ , ⟪ x₅ , x₆ ⟫ ⟫ -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ leaf (node key₂ ⟪ Red , v2 ⟫ (node key₃ _ _ _) leaf)) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₄ value₁ t₃ t₄)) key value -- (t-right .key₁ .key₂ x₁ x₂ x₃ (t-left key₃ .key₂ x₄ x₅ x₆ ti)) (rbr-rotate-rr x lt trb) -- = t-node _ _ _ x₁ (proj1 rr00) tt ⟪ x₄ , ⟪ x₅ , x₆ ⟫ ⟫ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (t-right _ _ (proj1 x₂) (proj1 (proj2 x₂)) (proj2 (proj2 x₂)) ti) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₂ < key₄) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt tt -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ leaf (node key₂ ⟪ Red , v2 ⟫ (node key₃ _ _ _) (node key₄ _ _ _))) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₅ value₁ t₃ t₄)) key value -- (t-right .key₁ .key₂ x₁ x₂ x₃ (t-node key₃ .key₂ key₄ x₄ x₅ x₆ x₇ x₈ x₉ ti ti₁)) (rbr-rotate-rr x lt trb) = t-node _ _ _ x₁ (proj1 rr00) tt ⟪ x₄ , ⟪ x₆ , x₇ ⟫ ⟫ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (t-right _ _ (proj1 x₂) (proj1 (proj2 x₂)) (proj2 (proj2 x₂)) ti) (RB-repl→ti _ _ _ _ ti₁ trb) where -- rr00 : (key₂ < key₅) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt ⟪ x₅ , ⟪ x₈ , x₉ ⟫ ⟫ -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ .(node key₃ _ _ _) (node key₂ ⟪ Red , v2 ⟫ .leaf .leaf)) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₄ value₁ t₃ t₄)) key value -- (t-node key₃ .key₁ .key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti (t-single .key₂ .(⟪ Red , v2 ⟫))) (rbr-rotate-rr x lt trb) -- = t-node _ _ _ x₂ (proj1 rr00) ⟪ <-trans x₁ x₂ , ⟪ >-tr< x₃ x₂ , >-tr< x₄ x₂ ⟫ ⟫ tt (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (t-left _ _ x₁ x₃ x₄ ti) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₂ < key₄) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt tt -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ (node key₃ v3 t t₁) (node key₂ ⟪ Red , v2 ⟫ leaf (node key₄ _ _ _))) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₅ value₁ t₃ t₄)) key value -- (t-node key₃ .key₁ .key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti (t-right .key₂ key₄ x₇ x₈ x₉ ti₁)) (rbr-rotate-rr x lt trb) -- = t-node _ _ _ x₂ (proj1 rr00) ⟪ <-trans x₁ x₂ , ⟪ >-tr< x₃ x₂ , >-tr< x₄ x₂ ⟫ ⟫ tt (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (t-left _ _ x₁ x₃ x₄ ti) (RB-repl→ti _ _ _ _ ti₁ trb) where -- rr00 : (key₂ < key₅) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt ⟪ x₇ , ⟪ x₈ , x₉ ⟫ ⟫ -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ (node key₃ _ _ _) (node key₂ ⟪ Red , v2 ⟫ (node key₄ _ _ _) leaf)) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₅ value₁ t₃ t₄)) key value -- (t-node key₃ key₁ key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti (t-left key₄ key₂ x₇ x₈ x₉ ti₁)) (rbr-rotate-rr x lt trb) -- = t-node _ _ _ x₂ (proj1 rr00) ⟪ <-trans x₁ x₂ , ⟪ >-tr< x₃ x₂ , >-tr< x₄ x₂ ⟫ ⟫ ⟪ x₇ , ⟪ x₈ , x₉ ⟫ ⟫ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (t-node _ _ _ x₁ (proj1 x₅) x₃ x₄ (proj1 (proj2 x₅)) (proj2 (proj2 x₅)) ti ti₁) (RB-repl→ti _ _ _ _ t-leaf trb) where -- rr00 : (key₂ < key₅) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt tt -- RB-repl→ti (node key₁ ⟪ Black , v1 ⟫ (node key₃ _ _ _) (node key₂ ⟪ Red , v2 ⟫ (node key₄ _ left right) (node key₅ _ _ _))) (node _ ⟪ Black , _ ⟫ (node _ ⟪ Red , _ ⟫ _ _) (node key₆ value₁ t₃ t₄)) key value -- (t-node key₃ key₁ key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti (t-node key₄ key₂ key₅ x₇ x₈ x₉ x₁₀ x₁₁ x₁₂ ti₁ ti₂)) (rbr-rotate-rr x lt trb) -- = t-node _ _ _ x₂ (proj1 rr00) ⟪ <-trans x₁ x₂ , ⟪ >-tr< x₃ x₂ , >-tr< x₄ x₂ ⟫ ⟫ ⟪ x₇ , ⟪ x₉ , x₁₀ ⟫ ⟫ (proj1 (proj2 rr00)) (proj2 (proj2 rr00)) (RTtoTI0 _ _ _ _ (t-node _ _ _ {_} {value₁} x₁ (proj1 x₅) x₃ x₄ (proj1 (proj2 x₅)) (proj2 (proj2 x₅)) ti ti₁ ) r-node ) -- (RB-repl→ti _ _ _ _ ti₂ trb) where -- rr00 : (key₂ < key₆) ∧ tr> key₂ t₃ ∧ tr> key₂ t₄ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt ⟪ x₈ , ⟪ x₁₁ , x₁₂ ⟫ ⟫ -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ leaf leaf) .leaf) (node kn ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ leaf) (node kg ⟪ Red , _ ⟫ leaf _)) .kn _ (t-left .kp .kg x x₁ x₂ ti) (rbr-rotate-lr .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = t-node _ _ _ lt1 lt2 tt tt tt tt (t-single _ _) (t-single _ _) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ (node key₁ value₁ t t₁) leaf) .leaf) (node kn ⟪ Black , v3 ⟫ (node kp ⟪ Red , _ ⟫ _ leaf) (node kg ⟪ Red , _ ⟫ leaf _)) .kn .v3 (t-left .kp .kg x x₁ x₂ (t-left .key₁ .kp x₃ x₄ x₅ ti)) (rbr-rotate-lr .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = -- t-node _ _ _ lt1 lt2 ⟪ <-trans x₃ lt1 , ⟪ >-tr< x₄ lt1 , >-tr< x₅ lt1 ⟫ ⟫ tt tt tt (t-left _ _ x₃ x₄ x₅ ti) (t-single _ _) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .leaf (node key₁ .(⟪ Red , _ ⟫) .leaf .leaf)) .leaf) (node .key₁ ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ leaf) (node kg ⟪ Red , _ ⟫ leaf _)) .key₁ _ (t-left .kp .kg x x₁ x₂ (t-right .kp .key₁ x₃ x₄ x₅ ti)) (rbr-rotate-lr .leaf .leaf kg kp .key₁ _ lt1 lt2 (rbr-node )) = t-node _ _ _ lt1 lt2 tt tt tt tt (t-single _ _) (t-single _ _) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .leaf (node key₁ value₁ t₁ t₂)) .leaf) (node kn ⟪ Black , value₃ ⟫ (node kp ⟪ Red , _ ⟫ _ leaf) (node kg ⟪ Red , _ ⟫ (node key₂ value₂ t₄ t₅) t₆)) key value (t-left .kp .kg x x₁ x₂ (t-right .kp .key₁ x₃ x₄ x₅ ti)) (rbr-rotate-lr .leaf .(node key₂ value₂ t₄ t₅) kg kp kn _ lt1 lt2 trb) = -- t-node _ _ _ (proj1 rr00) (proj1 rr01) tt tt rr03 tt (t-single _ _) (t-left _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) (treeRightDown _ _ ( RB-repl→ti _ _ _ _ ti trb))) where -- rr00 : (kp < kn) ∧ ⊤ ∧ ((kp < key₂) ∧ tr> kp t₄ ∧ tr> kp t₅ ) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₃ , ⟪ x₄ , x₅ ⟫ ⟫ -- rr01 : (kn < kg) ∧ ⊤ ∧ ((key₂ < kg ) ∧ tr< kg t₄ ∧ tr< kg t₅ ) -- tr< kg (node key₂ value₂ t₄ t₅) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₂ -- rr02 = proj2 (proj2 rr01) -- rr03 : (kn < key₂) ∧ tr> kn t₄ ∧ tr> kn t₅ -- rr03 with RB-repl→ti _ _ _ _ ti trb -- ... | t-right .kn .key₂ x x₁ x₂ t = ⟪ x , ⟪ x₁ , x₂ ⟫ ⟫ -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .leaf (node key₁ value₁ t₁ t₂)) .leaf) (node kn ⟪ Black , value₃ ⟫ (node kp ⟪ Red , _ ⟫ _ (node key₂ value₂ t₃ t₅)) (node kg ⟪ Red , _ ⟫ leaf _)) key value (t-left .kp .kg x x₁ x₂ (t-right .kp .key₁ x₃ x₄ x₅ ti)) (rbr-rotate-lr .(node key₂ value₂ t₃ t₅) .leaf kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti trb -- ... | t-left .key₂ .kn x₆ x₇ x₈ t = -- t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ tt tt (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) t) (t-single _ _) where -- rr00 : (kp < kn) ∧ ((kp < key₂) ∧ tr> kp t₃ ∧ tr> kp t₅) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₃ , ⟪ x₄ , x₅ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₂ < kg) ∧ tr< kg t₃ ∧ tr< kg t₅) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₂ -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .leaf (node key₁ value₁ t₁ t₂)) .leaf) (node kn ⟪ Black , value₄ ⟫ (node kp ⟪ Red , _ ⟫ _ (node key₂ value₂ t₃ t₅)) (node kg ⟪ Red , _ ⟫ (node key₃ value₃ t₄ t₆) _)) key value (t-left .kp .kg x x₁ x₂ (t-right .kp .key₁ x₃ x₄ x₅ ti)) (rbr-rotate-lr .(node key₂ value₂ t₃ t₅) .(node key₃ value₃ t₄ t₆) kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti trb -- ... | t-node .key₂ .kn .key₃ x₆ x₇ x₈ x₉ x₁₀ x₁₁ t t₇ = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₆ , ⟪ x₈ , x₉ ⟫ ⟫ ⟪ x₇ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) t) (t-left _ _ (proj1 rr03) (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) t₇) where -- rr00 : (kp < kn) ∧ ((kp < key₂) ∧ tr> kp t₃ ∧ tr> kp t₅) ∧ ((kp < key₃) ∧ tr> kp t₄ ∧ tr> kp t₆ ) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₃ , ⟪ x₄ , x₅ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₂ < kg) ∧ tr< kg t₃ ∧ tr< kg t₅) ∧ ((key₃ < kg) ∧ tr< kg t₄ ∧ tr< kg t₆ ) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₂ -- rr03 = proj2 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ (node key₂ value₂ t₅ t₆) (node key₁ value₁ t₁ t₂)) leaf) (node kn ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ t₃) (node kg ⟪ Red , _ ⟫ t₄ _)) key value (t-left .kp .kg x x₁ x₂ (t-node key₂ .kp .key₁ x₃ x₄ x₅ x₆ x₇ x₈ ti ti₁)) (rbr-rotate-lr t₃ t₄ kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti₁ trb -- ... | t-single .kn .(⟪ Red , _ ⟫) = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₃ (proj1 rr00) , ⟪ >-tr< x₅ (proj1 rr00) , >-tr< x₆ (proj1 rr00) ⟫ ⟫ tt tt tt (t-left _ _ x₃ x₅ x₆ ti) (t-single _ _) where -- rr00 : (kp < kn) ∧ ⊤ ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₄ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr01 : (kn < kg) ∧ ⊤ ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₂ -- ... | t-right .kn key₃ {v1} {v3} {t₇} {t₈} x₉ x₁₀ x₁₁ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₃ (proj1 rr00) , ⟪ >-tr< x₅ (proj1 rr00) , >-tr< x₆ (proj1 rr00) ⟫ ⟫ tt ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt (t-left _ _ x₃ x₅ x₆ ti) (t-left _ _ (proj1 rr03) (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) (treeRightDown _ _ (RB-repl→ti _ _ _ _ ti₁ trb))) where -- rr00 : (kp < kn) ∧ ⊤ ∧ ((kp < key₃) ∧ tr> kp t₇ ∧ tr> kp t₈) -- tr> kp (node key₃ v3 t₇ t₈) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₄ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ⊤ ∧ ((key₃ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) -- tr< kg (node key₃ v3 t₇ t₈) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₂ -- rr03 = proj2 (proj2 rr01) -- ... | t-left key₃ .kn {v1} {v3} {t₇} {t₈} x₉ x₁₀ x₁₁ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₃ (proj1 rr00) , ⟪ >-tr< x₅ (proj1 rr00) , >-tr< x₆ (proj1 rr00) ⟫ ⟫ ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt tt (t-node key₂ kp key₃ x₃ (proj1 rr02) x₅ x₆ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti (treeLeftDown _ _ (RB-repl→ti _ _ _ _ ti₁ trb))) (t-single _ _) where -- rr00 : (kp < kn) ∧ ((kp < key₃) ∧ tr> kp t₇ ∧ tr> kp t₈) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₄ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₃ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₂ -- ... | t-node key₃ .kn key₄ {v0} {v1} {v2} {t₇} {t₈} {t₉} {t₁₀} x₉ x₁₀ x₁₁ x₁₂ x₁₃ x₁₄ t t₃ = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₃ (proj1 rr00) , ⟪ >-tr< x₅ (proj1 rr00) , >-tr< x₆ (proj1 rr00) ⟫ ⟫ ⟪ x₉ , ⟪ x₁₁ , x₁₂ ⟫ ⟫ ⟪ x₁₀ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ tt (t-node _ _ _ x₃ (proj1 rr02) x₅ x₆ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti (treeLeftDown _ _ (RB-repl→ti _ _ _ _ ti₁ trb))) (t-left _ _ (proj1 rr03) (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) t₃) where -- rr00 : (kp < kn) ∧ ((kp < key₃) ∧ tr> kp t₇ ∧ tr> kp t₈) ∧ ((kp < key₄) ∧ tr> kp t₉ ∧ tr> kp t₁₀) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₄ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₃ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) ∧ ((key₄ < kg) ∧ tr< kg t₉ ∧ tr< kg t₁₀) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₂ -- rr03 = proj2 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .leaf leaf) (node key₂ value₂ t₅ t₆)) (node kn ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ .leaf) (node kg ⟪ Red , _ ⟫ .leaf _)) .kn _ (t-node .kp .kg key₂ x x₁ x₂ x₃ x₄ x₅ (t-single .kp .(⟪ Red , v2 ⟫)) ti₁) (rbr-rotate-lr .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = t-node _ _ _ lt1 lt2 tt tt tt ⟪ <-trans lt2 x₁ , ⟪ <-tr> x₄ lt2 , <-tr> x₅ lt2 ⟫ ⟫ (t-single _ _) (t-right _ _ x₁ x₄ x₅ ti₁) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .(node key _ _ _) leaf) (node key₂ value₂ t₅ t₆)) (node kn ⟪ Black , _ ⟫ (node kp ⟪ Red , _ ⟫ _ .leaf) (node kg ⟪ Red , _ ⟫ .leaf _)) .kn _ (t-node .kp .kg key₂ x x₁ x₂ x₃ x₄ x₅ (t-left key .kp x₆ x₇ x₈ ti) ti₁) (rbr-rotate-lr .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = t-node _ _ _ lt1 lt2 ⟪ <-trans x₆ lt1 , ⟪ >-tr< x₇ lt1 , >-tr< x₈ lt1 ⟫ ⟫ tt tt ⟪ <-trans lt2 x₁ , ⟪ <-tr> x₄ lt2 , <-tr> x₅ lt2 ⟫ ⟫ (t-left _ _ x₆ x₇ x₈ ti) (t-right _ _ x₁ x₄ x₅ ti₁) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .leaf (node key₁ value₁ t₁ t₂)) .(node key₂ _ _ _)) (node kn ⟪ Black , value₃ ⟫ (node kp ⟪ Red , _ ⟫ _ t₃) (node kg ⟪ Red , _ ⟫ t₄ _)) key value (t-node .kp .kg key₂ x x₁ x₂ x₃ x₄ x₅ (t-right .kp .key₁ x₆ x₇ x₈ ti) ti₁) (rbr-rotate-lr t₃ t₄ kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti trb -- ... | t-single .kn .(⟪ Red , value₃ ⟫) = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt tt tt ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-single _ _) (t-right _ _ x₁ x₄ x₅ ti₁) where -- rr00 : (kp < kn) ∧ ⊤ ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr01 : (kn < kg) ∧ ⊤ ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- ... | t-right .kn key₃ {v1} {v3} {t₇} {t₈} x₉ x₁₀ x₁₁ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt tt ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-single _ _) (t-node _ _ _ (proj1 rr03) x₁ (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) x₄ x₅ (treeRightDown _ _ (RB-repl→ti _ _ _ _ ti trb)) ti₁) where -- rr00 : (kp < kn) ∧ ⊤ ∧ ((kp < key₃) ∧ tr> kp t₇ ∧ tr> kp t₈) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ⊤ ∧ ((key₃ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- rr03 = proj2 (proj2 rr01) -- ... | t-left key₃ .kn {v1} {v3} {t₇} {t₈} x₉ x₁₀ x₁₁ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) (treeLeftDown _ _ (RB-repl→ti _ _ _ _ ti trb))) (t-right _ _ x₁ x₄ x₅ ti₁) where -- rr00 : (kp < kn) ∧ ((kp < key₃) ∧ tr> kp t₇ ∧ tr> kp t₈) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₃ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- rr03 = proj1 (proj2 rr01) -- ... | t-node key₃ .kn key₄ {v0} {v1} {v2} {t₇} {t₈} {t₉} {t₁₀} x₉ x₁₀ x₁₁ x₁₂ x₁₃ x₁₄ t t₃ = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₉ , ⟪ x₁₁ , x₁₂ ⟫ ⟫ ⟪ x₁₀ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) (treeLeftDown _ _ (RB-repl→ti _ _ _ _ ti trb))) (t-node _ _ _ (proj1 rr03) x₁ (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) x₄ x₅ t₃ ti₁) where -- rr00 : (kp < kn) ∧ ((kp < key₃) ∧ tr> kp t₇ ∧ tr> kp t₈) ∧ ((kp < key₄) ∧ tr> kp t₉ ∧ tr> kp t₁₀) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₃ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) ∧ ((key₄ < kg) ∧ tr< kg t₉ ∧ tr< kg t₁₀) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- rr03 = proj2 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , v1 ⟫ (node kp ⟪ Red , v2 ⟫ .(node key₃ _ _ _) (node key₁ value₁ t₁ t₂)) .(node key₂ _ _ _)) (node kn ⟪ Black , value₃ ⟫ (node kp ⟪ Red , _ ⟫ _ t₃) (node kg ⟪ Red , _ ⟫ t₄ _)) key value (t-node .kp .kg key₂ x x₁ x₂ x₃ x₄ x₅ (t-node key₃ .kp .key₁ x₆ x₇ x₈ x₉ x₁₀ x₁₁ ti ti₂) ti₁) (rbr-rotate-lr t₃ t₄ kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti₂ trb -- ... | t-single .kn .(⟪ Red , value₃ ⟫) = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₆ (proj1 rr00) , ⟪ >-tr< x₈ (proj1 rr00) , >-tr< x₉ (proj1 rr00) ⟫ ⟫ tt tt ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-left _ _ x₆ x₈ x₉ ti) (t-right _ _ x₁ x₄ x₅ ti₁) where -- rr00 : (kp < kn) ∧ ⊤ ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₇ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ -- rr01 : (kn < kg) ∧ ⊤ ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- ... | t-right .kn key₄ {v1} {v3} {t₇} {t₈} x₁₂ x₁₃ x₁₄ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₆ (proj1 rr00) , ⟪ >-tr< x₈ (proj1 rr00) , >-tr< x₉ (proj1 rr00) ⟫ ⟫ tt ⟪ x₁₂ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-left _ _ x₆ x₈ x₉ ti) (t-node _ _ _ (proj1 rr03) x₁ (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) x₄ x₅ (treeRightDown _ _ (RB-repl→ti _ _ _ _ ti₂ trb)) ti₁ ) where -- rr00 : (kp < kn) ∧ ⊤ ∧ ((kp < key₄) ∧ tr> kp t₇ ∧ tr> kp t₈) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₇ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ -- rr02 = proj2 (proj2 rr00) -- rr01 : (kn < kg) ∧ ⊤ ∧ ((key₄ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- rr03 = proj2 (proj2 rr01) -- ... | t-left key₄ .kn {v1} {v3} {t₇} {t₈} x₁₂ x₁₃ x₁₄ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₆ (proj1 rr00) , ⟪ >-tr< x₈ (proj1 rr00) , >-tr< x₉ (proj1 rr00) ⟫ ⟫ ⟪ x₁₂ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ tt ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-node _ _ _ x₆ (proj1 rr02) x₈ x₉ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti (treeLeftDown _ _ (RB-repl→ti _ _ _ _ ti₂ trb)) ) (t-right _ _ x₁ x₄ x₅ ti₁) where -- rr00 : (kp < kn) ∧ ((kp < key₄) ∧ tr> kp t₇ ∧ tr> kp t₈) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₇ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₄ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- rr03 = proj1 (proj2 rr01) -- ... | t-node key₄ .kn key₅ {v0} {v1} {v2} {t₇} {t₈} {t₉} {t₁₀}x₁₂ x₁₃ x₁₄ x₁₅ x₁₆ x₁₇ t t₃ = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x₆ (proj1 rr00) , ⟪ >-tr< x₈ (proj1 rr00) , >-tr< x₉ (proj1 rr00) ⟫ ⟫ ⟪ x₁₂ , ⟪ x₁₄ , x₁₅ ⟫ ⟫ ⟪ x₁₃ , ⟪ x₁₆ , x₁₇ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₁ , ⟪ <-tr> x₄ (proj1 rr01) , <-tr> x₅ (proj1 rr01) ⟫ ⟫ (t-node _ _ _ x₆ (proj1 rr02) x₈ x₉ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti t ) (t-node _ _ _ (proj1 rr04) x₁ (proj1 (proj2 rr04)) (proj2 (proj2 rr04)) x₄ x₅ (treeRightDown _ _ (RB-repl→ti _ _ _ _ ti₂ trb)) ti₁ ) where -- rr00 : (kp < kn) ∧ ((kp < key₄) ∧ tr> kp t₇ ∧ tr> kp t₈) ∧ ((kp < key₅) ∧ tr> kp t₉ ∧ tr> kp t₁₀) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 ⟪ x₇ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ -- rr02 = proj1 (proj2 rr00) -- rr05 = proj2 (proj2 rr00) -- rr01 : (kn < kg) ∧ ((key₄ < kg) ∧ tr< kg t₇ ∧ tr< kg t₈) ∧ ((key₅ < kg) ∧ tr< kg t₉ ∧ tr< kg t₁₀) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 x₃ -- rr03 = proj1 (proj2 rr01) -- rr04 = proj2 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .leaf leaf)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf leaf) (node kp ⟪ Red , _ ⟫ leaf _)) .kn .vn (t-right .kg .kp x x₁ x₂ ti) (rbr-rotate-rl .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = t-node _ _ _ lt1 lt2 tt tt tt tt (t-single _ _) (t-single _ _) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node kn ⟪ Red , _ ⟫ leaf leaf) leaf)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf leaf) (node kp ⟪ Red , _ ⟫ leaf _)) .kn .vn (t-right .kg .kp x x₁ x₂ ti) (rbr-rotate-rl .leaf .leaf kg kp kn _ lt1 lt2 (rbr-node )) = t-node _ _ _ lt1 lt2 tt tt tt tt (t-single _ _) (t-single _ _) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .leaf (node key₁ value₁ t₅ t₆))) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf leaf) (node kp ⟪ Red , _ ⟫ leaf _)) .kn .vn (t-right .kg .kp x x₁ x₂ (t-right .kp .key₁ x₃ x₄ x₅ ti)) (rbr-rotate-rl .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = t-node _ _ _ lt1 lt2 tt tt tt ⟪ <-trans lt2 x₃ , ⟪ <-tr> x₄ lt2 , <-tr> x₅ lt2 ⟫ ⟫ (t-single _ _) (t-right _ _ x₃ x₄ x₅ ti) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₂ _ _ _) (node key₁ value₁ t₅ t₆))) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf leaf) (node kp ⟪ Red , _ ⟫ leaf _)) key value (t-right .kg .kp x x₁ x₂ (t-node key₂ .kp .key₁ x₃ x₄ x₅ x₆ x₇ x₈ ti ti₁)) (rbr-rotate-rl .leaf .leaf kg kp kn _ lt1 lt2 trb) = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt tt tt ⟪ <-trans (proj1 rr01) x₄ , ⟪ <-tr> x₇ (proj1 rr01) , <-tr> x₈ (proj1 rr01) ⟫ ⟫ (t-single _ _) (t-right _ _ x₄ x₇ x₈ ti₁) where -- rr00 : (kg < kn) ∧ ⊤ ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₁ -- rr01 : (kn < kp) ∧ ⊤ ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₃ , ⟪ x₅ , x₆ ⟫ ⟫ -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₂ _ _ _) .leaf)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf (node key₁ value₁ t₂ t₃)) (node kp ⟪ Red , _ ⟫ leaf _)) key value (t-right .kg .kp x x₁ x₂ (t-left key₂ .kp x₃ x₄ x₅ ti)) (rbr-rotate-rl .(node key₁ value₁ t₂ t₃) .leaf kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti trb -- ... | t-left .key₁ .kn x₆ x₇ x₈ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ tt tt (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) t) (t-single _ _) where -- rr00 : (kg < kn) ∧ ((kg < key₁) ∧ tr> kg t₂ ∧ tr> kg t₃) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₁ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₁ < kp) ∧ tr< kp t₂ ∧ tr< kp t₃) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₃ , ⟪ x₄ , x₅ ⟫ ⟫ -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₂ _ _ _) .(node key₃ _ _ _))) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf (node key₁ value₁ t₂ t₃)) (node kp ⟪ Red , _ ⟫ leaf _)) key value (t-right .kg .kp x x₁ x₂ (t-node key₂ .kp key₃ x₃ x₄ x₅ x₆ x₇ x₈ ti ti₁)) (rbr-rotate-rl .(node key₁ value₁ t₂ t₃) .leaf kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti trb -- ... | t-left .key₁ .kn x₉ x₁₀ x₁₁ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt ⟪ <-trans (proj1 rr01) x₄ , ⟪ <-tr> x₇ (proj1 rr01) , <-tr> x₈ (proj1 rr01) ⟫ ⟫ (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) t) (t-right _ _ x₄ x₇ x₈ ti₁) where -- rr00 : (kg < kn) ∧ ((kg < key₁) ∧ tr> kg t₂ ∧ tr> kg t₃) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₁ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₁ < kp) ∧ tr< kp t₂ ∧ tr< kp t₃) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₃ , ⟪ x₅ , x₆ ⟫ ⟫ -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .leaf .leaf)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .(node key₁ _ _ _) .leaf) (node kp ⟪ Red , _ ⟫ leaf _)) .kn .vn (t-node key₁ .kg .kp x x₁ x₂ x₃ x₄ x₅ ti (t-single .kp .(⟪ Red , vp ⟫))) (rbr-rotate-rl .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = t-node _ _ _ lt1 lt2 ⟪ <-trans x lt1 , ⟪ >-tr< x₂ lt1 , >-tr< x₃ lt1 ⟫ ⟫ tt tt tt (t-left _ _ x x₂ x₃ ti) (t-single _ _) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .leaf .(node key₂ _ _ _))) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .(node key₁ _ _ _) .leaf) (node kp ⟪ Red , _ ⟫ leaf _)) .kn .vn (t-node key₁ .kg .kp x x₁ x₂ x₃ x₄ x₅ ti (t-right .kp key₂ x₆ x₇ x₈ ti₁)) (rbr-rotate-rl .leaf .leaf kg kp kn _ lt1 lt2 rbr-leaf) = t-node _ _ _ lt1 lt2 ⟪ <-trans x lt1 , ⟪ >-tr< x₂ lt1 , >-tr< x₃ lt1 ⟫ ⟫ tt tt ⟪ <-trans lt2 x₆ , ⟪ <-tr> x₇ lt2 , <-tr> x₈ lt2 ⟫ ⟫ (t-left _ _ x x₂ x₃ ti) (t-right _ _ x₆ x₇ x₈ ti₁) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₂ _ _ _) .leaf)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .(node key₁ _ _ _) t₂) (node kp ⟪ Red , _ ⟫ leaf _)) key value (t-node key₁ .kg .kp x x₁ x₂ x₃ x₄ x₅ ti (t-left key₂ .kp x₆ x₇ x₈ ti₁)) (rbr-rotate-rl t₂ .leaf kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti₁ trb -- ... | t-single .kn .(⟪ Red , vn ⟫) = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ tt tt tt (t-left _ _ x x₂ x₃ ti) (t-single _ _) where -- rr00 : (kg < kn) ∧ ⊤ ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr01 : (kn < kp) ∧ ⊤ ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- ... | t-left key₃ .kn {v1} {v3} {t₇} {t₃} x₉ x₁₀ x₁₁ ti₀ = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt tt (t-node _ _ _ x (proj1 rr02) x₂ x₃ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti ti₀) (t-single _ _) where -- rr00 : (kg < kn) ∧ ((kg < key₃) ∧ tr> kg t₇ ∧ tr> kg t₃) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₃ < kp) ∧ tr< kp t₇ ∧ tr< kp t₃) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr03 = proj1 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₂ _ _ _) .(node key₃ _ _ _))) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .(node key₁ _ _ _) t₂) (node kp ⟪ Red , _ ⟫ leaf _)) key value (t-node key₁ .kg .kp x x₁ x₂ x₃ x₄ x₅ ti (t-node key₂ .kp key₃ x₆ x₇ x₈ x₉ x₁₀ x₁₁ ti₁ ti₂)) (rbr-rotate-rl t₂ .leaf kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti₁ trb -- ... | t-single .kn .(⟪ Red , vn ⟫) = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ tt tt ⟪ <-trans (proj1 rr01) x₇ , ⟪ <-tr> x₁₀ (proj1 rr01) , <-tr> x₁₁ (proj1 rr01) ⟫ ⟫ (t-left _ _ x x₂ x₃ ti) (t-right _ _ x₇ x₁₀ x₁₁ ti₂) where -- rr00 : (kg < kn) ∧ ⊤ ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kp) ∧ ⊤ ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₈ , x₉ ⟫ ⟫ -- ... | t-left key₄ .kn {v1} {v3} {t₇} {t₃} x₁₂ x₁₃ x₁₄ ti₀ = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ ⟪ x₁₂ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ tt ⟪ <-trans (proj1 rr01) x₇ , ⟪ <-tr> x₁₀ (proj1 rr01) , <-tr> x₁₁ (proj1 rr01) ⟫ ⟫ (t-node _ _ _ x (proj1 rr02) x₂ x₃ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti ti₀) (t-right _ _ x₇ x₁₀ x₁₁ ti₂) where -- rr00 : (kg < kn) ∧ ((kg < key₄) ∧ tr> kg t₇ ∧ tr> kg t₃) ∧ ⊤ -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr02 = proj1 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₄ < kp) ∧ tr< kp t₇ ∧ tr< kp t₃) ∧ ⊤ -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₈ , x₉ ⟫ ⟫ -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₂ _ _ _) .leaf)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf t₂) (node kp ⟪ Red , _ ⟫ (node key₁ value₁ t₃ t₄) _)) key value (t-right .kg .kp x x₁ x₂ (t-left key₂ .kp x₃ x₄ x₅ ti)) (rbr-rotate-rl t₂ .(node key₁ value₁ t₃ t₄) kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti trb -- ... | t-right .kn .key₁ x₆ x₇ x₈ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt tt ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ tt (t-single _ _) (t-left _ _ (proj1 rr03) (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) t) where -- rr00 : (kg < kn) ∧ ⊤ ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₁ -- rr02 = proj2 (proj2 rr00) -- rr01 : (kn < kp) ∧ ⊤ ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₃ , ⟪ x₄ , x₅ ⟫ ⟫ -- rr03 = proj2 (proj2 rr01) -- ... | t-node key₃ .kn .key₁ {v0} {v1} {v2} {t₇} {t₈} {t₁₀} x₆ x₇ x₈ x₉ x₁₀ x₁₁ t t₁ = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₆ , ⟪ x₈ , x₉ ⟫ ⟫ ⟪ x₇ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) t) (t-left _ _ (proj1 rr03) (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) t₁) where -- rr00 : (kg < kn) ∧ ((kg < key₃) ∧ tr> kg t₇ ∧ tr> kg t₈) ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₁ -- rr02 = proj1 (proj2 rr00) -- rr04 = proj2 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₃ < kp) ∧ tr< kp t₇ ∧ tr< kp t₈) ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₃ , ⟪ x₄ , x₅ ⟫ ⟫ -- rr03 = proj2 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₂ _ _ _) .(node key₃ _ _ _))) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .leaf t₂) (node kp ⟪ Red , _ ⟫ (node key₁ value₁ t₃ t₄) _)) key value (t-right .kg .kp x x₁ x₂ (t-node key₂ .kp key₃ x₃ x₄ x₅ x₆ x₇ x₈ ti ti₁)) (rbr-rotate-rl t₂ .(node key₁ value₁ t₃ t₄) kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti trb -- ... | t-right .kn .key₁ x₉ x₁₀ x₁₁ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt tt ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₄ , ⟪ <-tr> x₇ (proj1 rr01) , <-tr> x₈ (proj1 rr01) ⟫ ⟫ (t-single _ _) (t-node _ _ _ (proj1 rr03) x₄ (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) x₇ x₈ t ti₁ ) where -- rr00 : (kg < kn) ∧ ⊤ ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₁ -- rr01 : (kn < kp) ∧ ⊤ ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₃ , ⟪ x₅ , x₆ ⟫ ⟫ -- rr03 = proj2 (proj2 rr01) -- ... | t-node key₄ .kn .key₁ {v0} {v1} {v2} {t₇} {t₈} {t₁₀} x₉ x₁₀ x₁₁ x₁₂ x₁₃ x₁₄ t t₁ = t-node _ _ _ (proj1 rr00) (proj1 rr01) tt ⟪ x₉ , ⟪ x₁₁ , x₁₂ ⟫ ⟫ ⟪ x₁₀ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₄ , ⟪ <-tr> x₇ (proj1 rr01) , <-tr> x₈ (proj1 rr01) ⟫ ⟫ (t-right _ _ (proj1 rr02) (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) t) (t-node _ _ _ (proj1 rr04) x₄ (proj1 (proj2 rr04)) (proj2 (proj2 rr04)) x₇ x₈ t₁ ti₁ ) where -- rr00 : (kg < kn) ∧ ((kg < key₄) ∧ tr> kg t₇ ∧ tr> kg t₈) ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₁ -- rr02 = proj1 (proj2 rr00) -- rr05 = proj2 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₄ < kp) ∧ tr< kp t₇ ∧ tr< kp t₈) ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₃ , ⟪ x₅ , x₆ ⟫ ⟫ -- rr03 = proj1 (proj2 rr01) -- rr04 = proj2 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₃ _ _ _) .leaf)) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .(node key₂ _ _ _) t₂) (node kp ⟪ Red , _ ⟫ (node key₁ value₁ t₃ t₄) _)) key value (t-node key₂ .kg .kp x x₁ x₂ x₃ x₄ x₅ ti (t-left key₃ .kp x₆ x₇ x₈ ti₁)) (rbr-rotate-rl t₂ .(node key₁ value₁ t₃ t₄) kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti₁ trb -- ... | t-right .kn .key₁ x₉ x₁₀ x₁₁ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ tt ⟪ x₉ , ⟪ x₁₀ , x₁₁ ⟫ ⟫ tt (t-left _ _ x x₂ x₃ ti ) (t-left _ _ (proj1 rr03) (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) t) where -- rr00 : (kg < kn) ∧ ⊤ ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr02 = proj2 (proj2 rr00) -- rr01 : (kn < kp) ∧ ⊤ ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr03 = proj2 (proj2 rr01) -- ... | t-node key₄ .kn .key₁ {v0} {v1} {v2} {t₇} {t₈} {t₁₀} x₉ x₁₀ x₁₁ x₁₂ x₁₃ x₁₄ t t₁ = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ ⟪ x₉ , ⟪ x₁₁ , x₁₂ ⟫ ⟫ ⟪ x₁₀ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ tt (t-node _ _ _ x (proj1 rr02) x₂ x₃ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti t) (t-left _ _ (proj1 rr04) (proj1 (proj2 rr04)) (proj2 (proj2 rr04)) t₁) where -- rr00 : (kg < kn) ∧ ((kg < key₄) ∧ tr> kg t₇ ∧ tr> kg t₈) ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr02 = proj1 (proj2 rr00) -- rr05 = proj2 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₄ < kp) ∧ tr< kp t₇ ∧ tr< kp t₈) ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₇ , x₈ ⟫ ⟫ -- rr03 = proj1 (proj2 rr01) -- rr04 = proj2 (proj2 rr01) -- RB-repl→ti (node kg ⟪ Black , vg ⟫ _ (node kp ⟪ Red , vp ⟫ .(node key₃ _ _ _) .(node key₄ _ _ _))) (node kn ⟪ Black , vn ⟫ (node kg ⟪ Red , _ ⟫ .(node key₂ _ _ _) t₂) (node kp ⟪ Red , _ ⟫ (node key₁ value₁ t₃ t₄) _)) key value (t-node key₂ .kg .kp x x₁ x₂ x₃ x₄ x₅ ti (t-node key₃ .kp key₄ x₆ x₇ x₈ x₉ x₁₀ x₁₁ ti₁ ti₂)) (rbr-rotate-rl t₂ .(node key₁ value₁ t₃ t₄) kg kp kn _ lt1 lt2 trb) with RB-repl→ti _ _ _ _ ti₁ trb -- ... | t-right .kn .key₁ x₁₂ x₁₃ x₁₄ t = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ tt ⟪ x₁₂ , ⟪ x₁₃ , x₁₄ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₇ , ⟪ <-tr> x₁₀ (proj1 rr01) , <-tr> x₁₁ (proj1 rr01) ⟫ ⟫ (t-left _ _ x x₂ x₃ ti) (t-node _ _ _ (proj1 rr03) x₇ (proj1 (proj2 rr03)) (proj2 (proj2 rr03)) x₁₀ x₁₁ t ti₂ ) where -- rr00 : (kg < kn) ∧ ⊤ ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr02 = proj2 (proj2 rr00) -- rr01 : (kn < kp) ∧ ⊤ ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₈ , x₉ ⟫ ⟫ -- rr03 = proj2 (proj2 rr01) -- ... | t-node key₅ .kn .key₁ {v0} {v1} {v2} {t₇} {t₈} {t₁₀} x₁₂ x₁₃ x₁₄ x₁₅ x₁₆ x₁₇ t t₁ = t-node _ _ _ (proj1 rr00) (proj1 rr01) ⟪ <-trans x (proj1 rr00) , ⟪ >-tr< x₂ (proj1 rr00) , >-tr< x₃ (proj1 rr00) ⟫ ⟫ ⟪ x₁₂ , ⟪ x₁₄ , x₁₅ ⟫ ⟫ ⟪ x₁₃ , ⟪ x₁₆ , x₁₇ ⟫ ⟫ ⟪ <-trans (proj1 rr01) x₇ , ⟪ <-tr> x₁₀ (proj1 rr01) , <-tr> x₁₁ (proj1 rr01) ⟫ ⟫ (t-node _ _ _ x (proj1 rr02) x₂ x₃ (proj1 (proj2 rr02)) (proj2 (proj2 rr02)) ti t ) (t-node _ _ _ (proj1 rr04) x₇ (proj1 (proj2 rr04)) (proj2 (proj2 rr04)) x₁₀ x₁₁ t₁ ti₂ ) where -- rr00 : (kg < kn) ∧ ((kg < key₅) ∧ tr> kg t₇ ∧ tr> kg t₈) ∧ ((kg < key₁) ∧ tr> kg t₃ ∧ tr> kg t₄) -- rr00 = RB-repl→ti> _ _ _ _ _ trb lt1 x₄ -- rr02 = proj1 (proj2 rr00) -- rr05 = proj2 (proj2 rr00) -- rr01 : (kn < kp) ∧ ((key₅ < kp) ∧ tr< kp t₇ ∧ tr< kp t₈) ∧ ((key₁ < kp) ∧ tr< kp t₃ ∧ tr< kp t₄) -- rr01 = RB-repl→ti< _ _ _ _ _ trb lt2 ⟪ x₆ , ⟪ x₈ , x₉ ⟫ ⟫ -- rr03 = proj1 (proj2 rr01) -- rr04 = proj2 (proj2 rr01) -- -- if we consider tree invariant, this may be much simpler and faster -- stackToPG : {n : Level} {A : Set n} → {key : ℕ } → (tree orig : bt A ) → (stack : List (bt A)) → stackInvariant key tree orig stack → ( stack ≡ orig ∷ [] ) ∨ ( stack ≡ tree ∷ orig ∷ [] ) ∨ PG A tree stack stackToPG {n} {A} {key} tree .tree .(tree ∷ []) s-nil = case1 refl stackToPG {n} {A} {key} tree .(node _ _ _ tree) .(tree ∷ node _ _ _ tree ∷ []) (s-right _ _ _ x s-nil) = case2 (case1 refl) stackToPG {n} {A} {key} tree .(node k2 v2 t5 (node k1 v1 t2 tree)) (tree ∷ node _ _ _ tree ∷ .(node k2 v2 t5 (node k1 v1 t2 tree) ∷ [])) (s-right tree (node k2 v2 t5 (node k1 v1 t2 tree)) t2 {k1} {v1} x (s-right (node k1 v1 t2 tree) (node k2 v2 t5 (node k1 v1 t2 tree)) t5 {k2} {v2} x₁ s-nil)) = case2 (case2 record { parent = node k1 v1 t2 tree ; grand = _ ; pg = s2-1s2p refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig (tree ∷ node _ _ _ tree ∷ .(node k2 v2 t5 (node k1 v1 t2 tree) ∷ _)) (s-right tree orig t2 {k1} {v1} x (s-right (node k1 v1 t2 tree) orig t5 {k2} {v2} x₁ (s-right _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 t2 tree ; grand = _ ; pg = s2-1s2p refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig (tree ∷ node _ _ _ tree ∷ .(node k2 v2 t5 (node k1 v1 t2 tree) ∷ _)) (s-right tree orig t2 {k1} {v1} x (s-right (node k1 v1 t2 tree) orig t5 {k2} {v2} x₁ (s-left _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 t2 tree ; grand = _ ; pg = s2-1s2p refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree .(node k2 v2 (node k1 v1 t1 tree) t2) .(tree ∷ node k1 v1 t1 tree ∷ node k2 v2 (node k1 v1 t1 tree) t2 ∷ []) (s-right _ _ t1 {k1} {v1} x (s-left _ _ t2 {k2} {v2} x₁ s-nil)) = case2 (case2 record { parent = node k1 v1 t1 tree ; grand = _ ; pg = s2-1sp2 refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 t1 tree ∷ node k2 v2 (node k1 v1 t1 tree) t2 ∷ _) (s-right _ _ t1 {k1} {v1} x (s-left _ _ t2 {k2} {v2} x₁ (s-right _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 t1 tree ; grand = _ ; pg = s2-1sp2 refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 t1 tree ∷ node k2 v2 (node k1 v1 t1 tree) t2 ∷ _) (s-right _ _ t1 {k1} {v1} x (s-left _ _ t2 {k2} {v2} x₁ (s-left _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 t1 tree ; grand = _ ; pg = s2-1sp2 refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree .(node _ _ tree _) .(tree ∷ node _ _ tree _ ∷ []) (s-left _ _ t1 {k1} {v1} x s-nil) = case2 (case1 refl) stackToPG {n} {A} {key} tree .(node _ _ _ (node k1 v1 tree t1)) .(tree ∷ node k1 v1 tree t1 ∷ node _ _ _ (node k1 v1 tree t1) ∷ []) (s-left _ _ t1 {k1} {v1} x (s-right _ _ _ x₁ s-nil)) = case2 (case2 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s12p refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ _ (node k1 v1 tree t1) ∷ _) (s-left _ _ t1 {k1} {v1} x (s-right _ _ _ x₁ (s-right _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s12p refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ _ (node k1 v1 tree t1) ∷ _) (s-left _ _ t1 {k1} {v1} x (s-right _ _ _ x₁ (s-left _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s12p refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree .(node _ _ (node k1 v1 tree t1) _) .(tree ∷ node k1 v1 tree t1 ∷ node _ _ (node k1 v1 tree t1) _ ∷ []) (s-left _ _ t1 {k1} {v1} x (s-left _ _ _ x₁ s-nil)) = case2 (case2 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s1p2 refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ (node k1 v1 tree t1) _ ∷ _) (s-left _ _ t1 {k1} {v1} x (s-left _ _ _ x₁ (s-right _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s1p2 refl refl ; rest = _ ; stack=gp = refl } ) stackToPG {n} {A} {key} tree orig .(tree ∷ node k1 v1 tree t1 ∷ node _ _ (node k1 v1 tree t1) _ ∷ _) (s-left _ _ t1 {k1} {v1} x (s-left _ _ _ x₁ (s-left _ _ _ x₂ si))) = case2 (case2 record { parent = node k1 v1 tree t1 ; grand = _ ; pg = s2-s1p2 refl refl ; rest = _ ; stack=gp = refl } ) stackCase1 : {n : Level} {A : Set n} → {key : ℕ } → {tree orig : bt A } → {stack : List (bt A)} → stackInvariant key tree orig stack → stack ≡ orig ∷ [] → tree ≡ orig stackCase1 s-nil refl = refl pg-prop-1 : {n : Level} (A : Set n) → (tree orig : bt A ) → (stack : List (bt A)) → (pg : PG A tree stack) → (¬ PG.grand pg ≡ leaf ) ∧ (¬ PG.parent pg ≡ leaf) pg-prop-1 {_} A tree orig stack pg with PG.pg pg ... | s2-s1p2 refl refl = ⟪ (λ () ) , ( λ () ) ⟫ ... | s2-1sp2 refl refl = ⟪ (λ () ) , ( λ () ) ⟫ ... | s2-s12p refl refl = ⟪ (λ () ) , ( λ () ) ⟫ ... | s2-1s2p refl refl = ⟪ (λ () ) , ( λ () ) ⟫ popStackInvariant : {n : Level} {A : Set n} → (rest : List ( bt (Color ∧ A))) → ( tree parent orig : bt (Color ∧ A)) → (key : ℕ) → stackInvariant key tree orig ( tree ∷ parent ∷ rest ) → stackInvariant key parent orig (parent ∷ rest ) popStackInvariant rest tree parent orig key (s-right .tree .orig tree₁ x si) = sc00 where sc00 : stackInvariant key parent orig (parent ∷ rest ) sc00 with si-property1 si ... | refl = si popStackInvariant rest tree parent orig key (s-left .tree .orig tree₁ x si) = sc00 where sc00 : stackInvariant key parent orig (parent ∷ rest ) sc00 with si-property1 si ... | refl = si siToTreeinvariant : {n : Level} {A : Set n} → (rest : List ( bt (Color ∧ A))) → ( tree orig : bt (Color ∧ A)) → (key : ℕ) → treeInvariant orig → stackInvariant key tree orig ( tree ∷ rest ) → treeInvariant tree siToTreeinvariant .[] tree .tree key ti s-nil = ti siToTreeinvariant .(node _ _ leaf leaf ∷ []) .leaf .(node _ _ leaf leaf) key (t-single _ _) (s-right .leaf .(node _ _ leaf leaf) .leaf x s-nil) = t-leaf siToTreeinvariant .(node _ _ leaf (node key₁ _ _ _) ∷ []) .(node key₁ _ _ _) .(node _ _ leaf (node key₁ _ _ _)) key (t-right _ key₁ x₁ x₂ x₃ ti) (s-right .(node key₁ _ _ _) .(node _ _ leaf (node key₁ _ _ _)) .leaf x s-nil) = ti siToTreeinvariant .(node _ _ (node key₁ _ _ _) leaf ∷ []) .leaf .(node _ _ (node key₁ _ _ _) leaf) key (t-left key₁ _ x₁ x₂ x₃ ti) (s-right .leaf .(node _ _ (node key₁ _ _ _) leaf) .(node key₁ _ _ _) x s-nil) = t-leaf siToTreeinvariant .(node _ _ (node key₁ _ _ _) (node key₂ _ _ _) ∷ []) .(node key₂ _ _ _) .(node _ _ (node key₁ _ _ _) (node key₂ _ _ _)) key (t-node key₁ _ key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (s-right .(node key₂ _ _ _) .(node _ _ (node key₁ _ _ _) (node key₂ _ _ _)) .(node key₁ _ _ _) x s-nil) = ti₁ siToTreeinvariant .(node _ _ tree₁ tree ∷ _) tree orig key ti (s-right .tree .orig tree₁ x si2@(s-right .(node _ _ tree₁ tree) .orig tree₂ x₁ si)) with siToTreeinvariant _ _ _ _ ti si2 ... | t-single _ _ = t-leaf ... | t-right _ key₁ x₂ x₃ x₄ ti₁ = ti₁ ... | t-left key₁ _ x₂ x₃ x₄ ti₁ = t-leaf ... | t-node key₁ _ key₂ x₂ x₃ x₄ x₅ x₆ x₇ ti₁ ti₂ = ti₂ siToTreeinvariant .(node _ _ tree₁ tree ∷ _) tree orig key ti (s-right .tree .orig tree₁ x si2@(s-left .(node _ _ tree₁ tree) .orig tree₂ x₁ si)) with siToTreeinvariant _ _ _ _ ti si2 ... | t-single _ _ = t-leaf ... | t-right _ key₁ x₂ x₃ x₄ ti₁ = ti₁ ... | t-left key₁ _ x₂ x₃ x₄ ti₁ = t-leaf ... | t-node key₁ _ key₂ x₂ x₃ x₄ x₅ x₆ x₇ ti₁ ti₂ = ti₂ siToTreeinvariant .(node _ _ leaf leaf ∷ []) .leaf .(node _ _ leaf leaf) key (t-single _ _) (s-left .leaf .(node _ _ leaf leaf) .leaf x s-nil) = t-leaf siToTreeinvariant .(node _ _ leaf (node key₁ _ _ _) ∷ []) .leaf .(node _ _ leaf (node key₁ _ _ _)) key (t-right _ key₁ x₁ x₂ x₃ ti) (s-left .leaf .(node _ _ leaf (node key₁ _ _ _)) .(node key₁ _ _ _) x s-nil) = t-leaf siToTreeinvariant .(node _ _ (node key₁ _ _ _) leaf ∷ []) .(node key₁ _ _ _) .(node _ _ (node key₁ _ _ _) leaf) key (t-left key₁ _ x₁ x₂ x₃ ti) (s-left .(node key₁ _ _ _) .(node _ _ (node key₁ _ _ _) leaf) .leaf x s-nil) = ti siToTreeinvariant .(node _ _ (node key₁ _ _ _) (node key₂ _ _ _) ∷ []) .(node key₁ _ _ _) .(node _ _ (node key₁ _ _ _) (node key₂ _ _ _)) key (t-node key₁ _ key₂ x₁ x₂ x₃ x₄ x₅ x₆ ti ti₁) (s-left .(node key₁ _ _ _) .(node _ _ (node key₁ _ _ _) (node key₂ _ _ _)) .(node key₂ _ _ _) x s-nil) = ti siToTreeinvariant .(node _ _ tree tree₁ ∷ _) tree orig key ti (s-left .tree .orig tree₁ x si2@(s-right .(node _ _ tree tree₁) .orig tree₂ x₁ si)) with siToTreeinvariant _ _ _ _ ti si2 ... | t-single _ _ = t-leaf ... | t-right _ key₁ x₂ x₃ x₄ t = t-leaf ... | t-left key₁ _ x₂ x₃ x₄ t = t ... | t-node key₁ _ key₂ x₂ x₃ x₄ x₅ x₆ x₇ t t₁ = t siToTreeinvariant .(node _ _ tree tree₁ ∷ _) tree orig key ti (s-left .tree .orig tree₁ x si2@(s-left .(node _ _ tree tree₁) .orig tree₂ x₁ si)) with siToTreeinvariant _ _ _ _ ti si2 ... | t-single _ _ = t-leaf ... | t-right _ key₁ x₂ x₃ x₄ t = t-leaf ... | t-left key₁ _ x₂ x₃ x₄ t = t ... | t-node key₁ _ key₂ x₂ x₃ x₄ x₅ x₆ x₇ t t₁ = t PGtoRBinvariant1 : {n : Level} {A : Set n} → (tree orig : bt (Color ∧ A) ) → {key : ℕ } → (rb : RBtreeInvariant orig) → (stack : List (bt (Color ∧ A))) → (si : stackInvariant key tree orig stack ) → RBtreeInvariant tree PGtoRBinvariant1 tree .tree rb .(tree ∷ []) s-nil = rb PGtoRBinvariant1 tree orig rb (tree ∷ rest) (s-right .tree .orig tree₁ x si) with PGtoRBinvariant1 _ orig rb _ si ... | rb-red _ value x₁ x₂ x₃ t t₁ = t₁ ... | rb-black _ value x₁ t t₁ = t₁ PGtoRBinvariant1 tree orig rb (tree ∷ rest) (s-left .tree .orig tree₁ x si) with PGtoRBinvariant1 _ orig rb _ si ... | rb-red _ value x₁ x₂ x₃ t t₁ = t ... | rb-black _ value x₁ t t₁ = t RBI-child-replaced : {n : Level} {A : Set n} (tr : bt (Color ∧ A)) (key : ℕ) → RBtreeInvariant tr → RBtreeInvariant (child-replaced key tr) RBI-child-replaced {n} {A} leaf key rbi = rbi RBI-child-replaced {n} {A} (node key₁ value tr tr₁) key rbi with <-cmp key key₁ ... | tri< a ¬b ¬c = RBtreeLeftDown _ _ rbi ... | tri≈ ¬a b ¬c = rbi ... | tri> ¬a ¬b c = RBtreeRightDown _ _ rbi -- -- create RBT invariant after findRBT, continue to replaceRBT -- replaceRBTNode : {n m : Level} {A : Set n } {t : Set m } → (key : ℕ) (value : A) → (tree0 : bt (Color ∧ A)) → RBtreeInvariant tree0 → treeInvariant tree0 → (tree1 : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) → stackInvariant key tree1 tree0 stack → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → (exit : (stack : List ( bt (Color ∧ A))) (r : RBI key value tree0 stack ) → t ) → t replaceRBTNode {n} {m} {A} {t} key value orig rbi ti leaf stack si (case1 refl) exit with stackToPG leaf orig stack si ... | case1 x = ? ... | case2 (case1 x) = ? ... | case2 (case2 pg) = rp00 where rb00 : stackInvariant key leaf orig (leaf ∷ PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) rb00 = subst (λ k → stackInvariant key leaf orig k) (PG.stack=gp pg) si treerb : RBtreeInvariant (PG.parent pg) treerb = PGtoRBinvariant1 _ orig rbi _ (popStackInvariant _ _ _ _ _ rb00) pti : treeInvariant (PG.parent pg) pti = siToTreeinvariant _ _ _ _ ti (popStackInvariant _ _ _ _ _ rb00) rp00 : t rp00 with PG.pg pg ... | s2-s1p2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ with proj1 vp in pcolor ... | Red = rotateCase1 where rotateCase1 : t rotateCase1 = exit (PG.grand pg ∷ PG.rest pg) record { tree = PG.grand pg ; repl = to-red ( node kg vg (to-black (node kp vp (node key ⟪ Red , value ⟫ leaf leaf) n1)) n2 ) ; origti = ti ; origrb = rbi ; treerb = PGtoRBinvariant1 _ orig rbi _ (popStackInvariant _ _ _ _ _ (popStackInvariant _ _ _ _ _ rb00)) ; replrb = rb-red _ ? ? ? ? ? ? ; si = popStackInvariant _ _ _ _ _ (popStackInvariant _ _ _ _ _ rb00) ; state = rotate ? ? ? ? } ... | Black = buildCase1 where rb01 : bt (Color ∧ A) rb01 = node kp vp (node key ⟪ Red , value ⟫ leaf leaf) n1 rb03 : 1 ≡ black-depth n1 rb03 = RBtreeEQ (subst (λ k → RBtreeInvariant k) x treerb) rb02 : RBtreeInvariant (node kp ⟪ Black , proj2 vp ⟫ (node key ⟪ Red , value ⟫ leaf leaf) n1) rb02 = rb-black kp (proj2 vp) rb03 (rb-red _ _ refl refl refl rb-leaf rb-leaf) (RBtreeRightDown _ _ (subst (λ k → RBtreeInvariant k) x treerb)) rb04 : proj1 vp ≡ Black → ⟪ Black , proj2 vp ⟫ ≡ vp rb04 refl = refl rb05 : black-depth rb01 ≡ black-depth (PG.parent pg) rb05 = begin black-depth (node kp vp (node key ⟪ Red , value ⟫ leaf leaf) n1) ≡⟨ cong (λ k → black-depth (node kp k _ n1)) (sym (rb04 pcolor)) ⟩ black-depth (node kp ⟪ Black , proj2 vp ⟫ (node key ⟪ Red , value ⟫ leaf leaf) n1) ≡⟨ refl ⟩ black-depth (node kp ⟪ Black , proj2 vp ⟫ leaf n1 ) ≡⟨ cong (λ k → black-depth (node kp k leaf n1)) (rb04 pcolor) ⟩ black-depth (node kp vp leaf n1) ≡⟨ cong black-depth (sym x) ⟩ black-depth (PG.parent pg) ∎ where open ≡-Reasoning rb07 : key < kp rb07 = si-property-< ? (subst (λ k → treeInvariant k) x pti) (subst₂ (λ j k → stackInvariant key j orig k) refl (trans (PG.stack=gp pg) ? ) si) rb06 : replacedRBTree key value (PG.parent pg) (node kp vp (node key ⟪ Red , value ⟫ leaf leaf) n1) rb06 = subst₂ (λ j k → replacedRBTree key value k (node kp j (node key ⟪ Red , value ⟫ leaf leaf) n1)) (rb04 pcolor) (trans (cong (λ k → node kp k leaf n1) (rb04 pcolor)) (sym x)) ( rbr-black-left refl rb07 rbr-leaf ) buildCase1 : t buildCase1 = exit (PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) record { tree = PG.parent pg ; repl = rb01 ; origti = ti ; origrb = rbi ; treerb = PGtoRBinvariant1 _ orig rbi _ (popStackInvariant _ _ _ _ _ rb00) ; replrb = subst (λ k → RBtreeInvariant k) (cong (λ k → node kp k (node key ⟪ Red , value ⟫ leaf leaf) n1) (rb04 pcolor)) rb02 ; si = popStackInvariant _ _ _ _ _ rb00 ; state = rebuild (cong color x) rb05 rb06 } rp00 | s2-1sp2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = {!!} rp00 | s2-s12p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = {!!} rp00 | s2-1s2p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = {!!} replaceRBTNode {n} {m} {A} key value orig rbi ti tree@(node key₁ value₁ left right) stack si (case2 P) exit = exit ? record { tree = tree ; repl = node key₁ ⟪ proj1 value₁ , value ⟫ left right ; origti = ti ; origrb = rbi ; treerb = PGtoRBinvariant1 tree orig rbi stack si ; replrb = ? ; si = si ; state = rebuild ? ? ? } -- -- rotate and rebuild replaceTree and rb-invariant -- replaceRBP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (orig : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) → (r : RBI key value orig stack ) → (next : (stack1 : List (bt (Color ∧ A))) → (r : RBI key value orig stack1 ) → length stack1 < length stack → t ) → (exit : (stack1 : List (bt (Color ∧ A))) → stack1 ≡ (orig ∷ []) → RBI key value orig stack1 → t ) → t replaceRBP {n} {m} {A} {t} key value orig stack r next exit = replaceRBP1 where -- we have no grand parent -- eq : stack₁ ≡ RBI.tree r ∷ orig ∷ [] -- change parent color ≡ Black and exit -- one level stack, orig is parent of repl repl = RBI.repl r insertCase4 : (tr0 : bt (Color ∧ A)) → tr0 ≡ orig → (eq : stack ≡ RBI.tree r ∷ orig ∷ []) → (rot : replacedRBTree key value (RBI.tree r) repl) → ( color repl ≡ Red ) ∨ (color (RBI.tree r) ≡ color repl) → stackInvariant key (RBI.tree r) orig stack → t insertCase4 leaf eq1 eq rot col si = ⊥-elim (rb04 eq eq1 si) where -- can't happen rb04 : {stack : List ( bt ( Color ∧ A))} → stack ≡ RBI.tree r ∷ orig ∷ [] → leaf ≡ orig → stackInvariant key (RBI.tree r) orig stack → ⊥ rb04 refl refl (s-right tree leaf tree₁ x si) = si-property2 _ (s-right tree leaf tree₁ x si) refl rb04 refl refl (s-left tree₁ leaf tree x si) = si-property2 _ (s-left tree₁ leaf tree x si) refl insertCase4 tr0@(node key₁ value₁ left right) refl eq rot col si with <-cmp key key₁ ... | tri< a ¬b ¬c = rb07 col where rb04 : stackInvariant key (RBI.tree r) orig stack → stack ≡ RBI.tree r ∷ orig ∷ [] → tr0 ≡ orig → left ≡ RBI.tree r rb04 (s-left tree₁ .(node key₁ value₁ left right) tree {key₂} x s-nil) refl refl = refl rb04 (s-right tree .(node key₁ _ tree₁ tree) tree₁ x s-nil) refl refl with si-property1 si ... | refl = ⊥-elim ( nat-<> x a ) rb06 : black-depth repl ≡ black-depth right rb06 = begin black-depth repl ≡⟨ sym (RB-repl→eq _ _ (RBI.treerb r) rot) ⟩ black-depth (RBI.tree r) ≡⟨ cong black-depth (sym (rb04 si eq refl)) ⟩ black-depth left ≡⟨ (RBtreeEQ (RBI.origrb r)) ⟩ black-depth right ∎ where open ≡-Reasoning rb08 : (color (RBI.tree r) ≡ color repl) → RBtreeInvariant (node key₁ value₁ repl right) rb08 ceq with proj1 value₁ in coeq ... | Red = subst (λ k → RBtreeInvariant (node key₁ k repl right)) (cong (λ k → ⟪ k , _ ⟫) (sym coeq) ) (rb-red _ (proj2 value₁) rb09 (proj2 (RBtreeChildrenColorBlack _ _ (RBI.origrb r) coeq)) rb06 (RBI.replrb r) (RBtreeRightDown _ _ (RBI.origrb r))) where rb09 : color repl ≡ Black rb09 = trans (trans (sym ceq) (sym (cong color (rb04 si eq refl) ))) (proj1 (RBtreeChildrenColorBlack _ _ (RBI.origrb r) coeq)) ... | Black = subst (λ k → RBtreeInvariant (node key₁ k repl right)) (cong (λ k → ⟪ k , _ ⟫) (sym coeq) ) (rb-black _ (proj2 value₁) rb06 (RBI.replrb r) (RBtreeRightDown _ _ (RBI.origrb r))) rb07 : ( color repl ≡ Red ) ∨ (color (RBI.tree r) ≡ color repl) → t rb07 (case2 cc ) = exit (orig ∷ []) refl record { tree = orig ; repl = node key₁ value₁ repl right ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = RBI.origrb r ; replrb = rb08 cc ; si = s-nil ; state = top-black refl (case1 (rbr-left (trans (cong color (rb04 si eq refl)) cc) a (subst (λ k → replacedRBTree key value k repl) (sym (rb04 si eq refl)) rot))) } rb07 (case1 repl-red) = exit (orig ∷ []) refl record { tree = orig ; repl = to-black (node key₁ value₁ repl right) ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = RBI.origrb r ; replrb = rb-black _ _ rb06 (RBI.replrb r) (RBtreeRightDown _ _ (RBI.origrb r)) ; si = s-nil ; state = top-black refl (case2 (rbr-black-left repl-red a (subst (λ k → replacedRBTree key value k repl) (sym (rb04 si eq refl)) rot))) } ... | tri≈ ¬a b ¬c = ⊥-elim ( rb06 _ eq (RBI.si r) b ) where -- can't happen rb06 : (stack : List (bt (Color ∧ A))) → stack ≡ RBI.tree r ∷ node key₁ value₁ left right ∷ [] → stackInvariant key (RBI.tree r) (node key₁ value₁ left right) stack → key ≡ key₁ → ⊥ rb06 (.right ∷ node key₁ value₁ left right ∷ []) refl (s-right .right .(node key₁ value₁ left right) .left x s-nil) eq = ⊥-elim ( nat-≡< (sym eq) x) rb06 (.left ∷ node key₁ value₁ left right ∷ []) refl (s-left .left .(node key₁ value₁ left right) .right x s-nil) eq = ⊥-elim ( nat-≡< eq x) ... | tri> ¬a ¬b c = rb07 col where rb04 : stackInvariant key (RBI.tree r) orig stack → stack ≡ RBI.tree r ∷ orig ∷ [] → tr0 ≡ orig → right ≡ RBI.tree r rb04 (s-right tree .(node key₁ _ tree₁ tree) tree₁ x s-nil) refl refl = refl rb04 (s-left tree₁ .(node key₁ value₁ left right) tree {key₂} x si) refl refl with si-property1 si ... | refl = ⊥-elim ( nat-<> x c ) rb06 : black-depth repl ≡ black-depth left rb06 = begin black-depth repl ≡⟨ sym (RB-repl→eq _ _ (RBI.treerb r) rot) ⟩ black-depth (RBI.tree r) ≡⟨ cong black-depth (sym (rb04 si eq refl)) ⟩ black-depth right ≡⟨ (sym (RBtreeEQ (RBI.origrb r))) ⟩ black-depth left ∎ where open ≡-Reasoning rb08 : (color (RBI.tree r) ≡ color repl) → RBtreeInvariant (node key₁ value₁ left repl ) rb08 ceq with proj1 value₁ in coeq ... | Red = subst (λ k → RBtreeInvariant (node key₁ k left repl )) (cong (λ k → ⟪ k , _ ⟫) (sym coeq) ) (rb-red _ (proj2 value₁) (proj1 (RBtreeChildrenColorBlack _ _ (RBI.origrb r) coeq)) rb09 (sym rb06) (RBtreeLeftDown _ _ (RBI.origrb r))(RBI.replrb r)) where rb09 : color repl ≡ Black rb09 = trans (trans (sym ceq) (sym (cong color (rb04 si eq refl) ))) (proj2 (RBtreeChildrenColorBlack _ _ (RBI.origrb r) coeq)) ... | Black = subst (λ k → RBtreeInvariant (node key₁ k left repl )) (cong (λ k → ⟪ k , _ ⟫) (sym coeq) ) (rb-black _ (proj2 value₁) (sym rb06) (RBtreeLeftDown _ _ (RBI.origrb r)) (RBI.replrb r)) rb07 : ( color repl ≡ Red ) ∨ (color (RBI.tree r) ≡ color repl) → t rb07 (case2 cc ) = exit (orig ∷ []) refl record { tree = orig ; repl = node key₁ value₁ left repl ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = RBI.origrb r ; replrb = rb08 cc ; si = s-nil ; state = top-black refl (case1 (rbr-right (trans (cong color (rb04 si eq refl)) cc) c (subst (λ k → replacedRBTree key value k repl) (sym (rb04 si eq refl)) rot))) } rb07 (case1 repl-red ) = exit (orig ∷ []) refl record { tree = orig ; repl = to-black (node key₁ value₁ left repl) ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = RBI.origrb r ; replrb = rb-black _ _ (sym rb06) (RBtreeLeftDown _ _ (RBI.origrb r)) (RBI.replrb r) ; si = s-nil ; state = top-black refl (case2 (rbr-black-right repl-red c (subst (λ k → replacedRBTree key value k repl) (sym (rb04 si eq refl)) rot))) } rebuildCase : (ceq : color (RBI.tree r) ≡ color repl) → (bdepth-eq : black-depth repl ≡ black-depth (RBI.tree r)) → (rot : replacedRBTree key value (RBI.tree r) repl ) → t rebuildCase ceq bdepth-eq rot with stackToPG (RBI.tree r) orig stack (RBI.si r) ... | case1 x = exit stack x r where -- single node case rb00 : stack ≡ orig ∷ [] → orig ≡ RBI.tree r rb00 refl = si-property1 (RBI.si r) ... | case2 (case1 x) = insertCase4 orig refl x rot (case2 ceq) (RBI.si r) -- one level stack, orig is parent of repl ... | case2 (case2 pg) = rebuildCase1 pg where rb00 : (pg : PG (Color ∧ A) (RBI.tree r) stack) → stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) rb00 pg = subst (λ k → stackInvariant key (RBI.tree r) orig k) (PG.stack=gp pg) (RBI.si r) treerb : (pg : PG (Color ∧ A) (RBI.tree r) stack) → RBtreeInvariant (PG.parent pg) treerb pg = PGtoRBinvariant1 _ orig (RBI.origrb r) _ (popStackInvariant _ _ _ _ _ (rb00 pg)) rb08 : (pg : PG (Color ∧ A) (RBI.tree r) stack) → treeInvariant (PG.parent pg) rb08 pg = siToTreeinvariant _ _ _ _ (RBI.origti r) (popStackInvariant _ _ _ _ _ (rb00 pg)) rebuildCase1 : (PG (Color ∧ A) (RBI.tree r) stack) → t rebuildCase1 pg with PG.pg pg ... | s2-s1p2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | s2-1sp2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | s2-s12p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | s2-1s2p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = next (PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) record { tree = PG.parent pg ; repl = rb01 ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = treerb pg ; replrb = rb02 ; si = popStackInvariant _ _ _ _ _ (rb00 pg) ; state = rebuild (cong color x) rb06 (subst (λ k → replacedRBTree key value k rb01) (sym x) (rbr-right ceq rb04 rot)) } (subst₂ (λ j k → j < length k ) refl (sym (PG.stack=gp pg)) ≤-refl) where rb01 : bt (Color ∧ A) rb01 = node kp vp n1 repl rb03 : black-depth n1 ≡ black-depth repl rb03 = trans (RBtreeEQ (subst (λ k → RBtreeInvariant k) x (treerb pg))) ((RB-repl→eq _ _ (RBI.treerb r) rot)) rb02 : RBtreeInvariant rb01 rb02 with subst (λ k → RBtreeInvariant k) x (treerb pg) ... | rb-red .kp value cx cx₁ ex₂ t t₁ = rb-red kp value cx (trans (sym ceq) cx₁) rb03 t (RBI.replrb r) ... | rb-black .kp value ex t t₁ = rb-black kp value rb03 t (RBI.replrb r) rb05 : treeInvariant (node kp vp n1 (RBI.tree r)) rb05 = subst (λ k → treeInvariant k) x (rb08 pg) rb04 : kp < key rb04 = si-property-> ? rb05 (subst (λ k → stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ k ∷ PG.grand pg ∷ PG.rest pg)) x (rb00 pg)) rb06 : black-depth rb01 ≡ black-depth (PG.parent pg) rb06 = trans (rb07 vp) ( cong black-depth (sym x) ) where rb07 : (vp : Color ∧ A) → black-depth (node kp vp n1 repl) ≡ black-depth (node kp vp n1 (RBI.tree r)) rb07 ⟪ Black , proj4 ⟫ = cong (λ k → suc (black-depth n1 ⊔ k)) bdepth-eq rb07 ⟪ Red , proj4 ⟫ = cong (λ k → (black-depth n1 ⊔ k)) bdepth-eq -- -- both parent and uncle are Red, rotate then goto rebuild -- insertCase5 : (repl1 : bt (Color ∧ A)) → repl1 ≡ repl → (pg : PG (Color ∧ A) (RBI.tree r) stack) → (rot : replacedRBTree key value (RBI.tree r) repl) → color repl ≡ Red → color (PG.uncle pg) ≡ Black → color (PG.parent pg) ≡ Red → t insertCase5 leaf eq pg rot repl-red uncle-black pcolor = ⊥-elim ( rb00 repl repl-red (cong color (sym eq))) where rb00 : (repl : bt (Color ∧ A)) → color repl ≡ Red → color repl ≡ Black → ⊥ rb00 (node key ⟪ Black , proj4 ⟫ repl repl₁) () eq1 rb00 (node key ⟪ Red , proj4 ⟫ repl repl₁) eq () insertCase5 (node rkey vr rp-left rp-right) eq pg rot repl-red uncle-black pcolor = insertCase51 where rb00 : stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) rb00 = subst (λ k → stackInvariant key (RBI.tree r) orig k) (PG.stack=gp pg) (RBI.si r) rb15 : suc (length (PG.rest pg)) < length stack rb15 = begin suc (suc (length (PG.rest pg))) ≤⟨ <-trans (n<1+n _) (n<1+n _) ⟩ length (RBI.tree r ∷ PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) ≡⟨ cong length (sym (PG.stack=gp pg)) ⟩ length stack ∎ where open ≤-Reasoning rb02 : RBtreeInvariant (PG.grand pg) rb02 = PGtoRBinvariant1 _ orig (RBI.origrb r) _ (popStackInvariant _ _ _ _ _ (popStackInvariant _ _ _ _ _ rb00)) rb09 : RBtreeInvariant (PG.parent pg) rb09 = PGtoRBinvariant1 _ orig (RBI.origrb r) _ (popStackInvariant _ _ _ _ _ rb00) rb08 : treeInvariant (PG.parent pg) rb08 = siToTreeinvariant _ _ _ _ (RBI.origti r) (popStackInvariant _ _ _ _ _ rb00) insertCase51 : t insertCase51 with PG.pg pg ... | s2-s1p2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = next (PG.grand pg ∷ PG.rest pg) record { tree = PG.grand pg ; repl = rb01 ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = rb02 ; replrb = subst (λ k → RBtreeInvariant k) rb10 (rb-black _ _ rb18 (RBI.replrb r) (rb-red _ _ rb16 uncle-black rb19 rb06 rb05)) ; si = popStackInvariant _ _ _ _ _ (popStackInvariant _ _ _ _ _ rb00) ; state = rebuild (trans (cong color x₁) (cong proj1 (sym rb14))) rb17 rb11 } rb15 where -- x₁ : PG.grand pg ≡ node kg vg (PG.parent pg) (PG.uncle pg) -- x : PG.parent pg ≡ node kp vp (RBI.tree r) n1 -- repl : ode rkey value rp-reft rp-right rb01 : bt (Color ∧ A) rb01 = to-black (node kp vp (node rkey vr rp-left rp-right) (to-red (node kg vg n1 (PG.uncle pg)))) rb04 : key < kp rb04 = si-property-< ? (subst (λ k → treeInvariant k) x rb08) (subst (λ k → stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ k ∷ PG.grand pg ∷ PG.rest pg)) x (rb00)) rb16 : color n1 ≡ Black rb16 = proj2 (RBtreeChildrenColorBlack _ _ (subst (λ k → RBtreeInvariant k) x rb09) (trans (cong color (sym x)) pcolor)) rb13 : ⟪ Red , proj2 vp ⟫ ≡ vp rb13 with subst (λ k → color k ≡ Red ) x pcolor ... | refl = refl rb14 : ⟪ Black , proj2 vg ⟫ ≡ vg rb14 with RBtreeParentColorBlack _ _ (subst (λ k → RBtreeInvariant k) x₁ rb02) (case1 pcolor) ... | refl = refl rb03 : replacedRBTree key value (node kg _ (node kp ⟪ Red , proj2 vp ⟫ (RBI.tree r) n1) (PG.uncle pg)) (node kp ⟪ Black , proj2 vp ⟫ repl (node kg ⟪ Red , proj2 vg ⟫ n1 (PG.uncle pg))) rb03 = rbr-rotate-ll repl-red rb04 rot rb10 : node kp ⟪ Black , proj2 vp ⟫ repl (node kg ⟪ Red , proj2 vg ⟫ n1 (PG.uncle pg)) ≡ rb01 rb10 = begin to-black (node kp vp repl (to-red (node kg vg n1 (PG.uncle pg)))) ≡⟨ cong (λ k → node _ _ k _) (sym eq) ⟩ rb01 ∎ where open ≡-Reasoning rb12 : node kg ⟪ Black , proj2 vg ⟫ (node kp ⟪ Red , proj2 vp ⟫ (RBI.tree r) n1) (PG.uncle pg) ≡ PG.grand pg rb12 = begin node kg ⟪ Black , proj2 vg ⟫ (node kp ⟪ Red , proj2 vp ⟫ (RBI.tree r) n1) (PG.uncle pg) ≡⟨ cong₂ (λ j k → node kg j (node kp k (RBI.tree r) n1) (PG.uncle pg) ) rb14 rb13 ⟩ node kg vg _ (PG.uncle pg) ≡⟨ cong (λ k → node _ _ k _) (sym x) ⟩ node kg vg (PG.parent pg) (PG.uncle pg) ≡⟨ sym x₁ ⟩ PG.grand pg ∎ where open ≡-Reasoning rb11 : replacedRBTree key value (PG.grand pg) rb01 rb11 = subst₂ (λ j k → replacedRBTree key value j k) rb12 rb10 rb03 rb05 : RBtreeInvariant (PG.uncle pg) rb05 = RBtreeRightDown _ _ (subst (λ k → RBtreeInvariant k) x₁ rb02) rb06 : RBtreeInvariant n1 rb06 = RBtreeRightDown _ _ (subst (λ k → RBtreeInvariant k) x rb09) rb19 : black-depth n1 ≡ black-depth (PG.uncle pg) rb19 = trans (sym ( proj2 (red-children-eq x (sym (cong proj1 rb13)) rb09) )) (RBtreeEQ (subst (λ k → RBtreeInvariant k) x₁ rb02)) rb18 : black-depth repl ≡ black-depth n1 ⊔ black-depth (PG.uncle pg) rb18 = begin black-depth repl ≡⟨ sym (RB-repl→eq _ _ (RBI.treerb r) rot) ⟩ black-depth (RBI.tree r) ≡⟨ RBtreeEQ (subst (λ k → RBtreeInvariant k) x rb09) ⟩ black-depth n1 ≡⟨ sym (⊔-idem (black-depth n1)) ⟩ black-depth n1 ⊔ black-depth n1 ≡⟨ cong (λ k → black-depth n1 ⊔ k) rb19 ⟩ black-depth n1 ⊔ black-depth (PG.uncle pg) ∎ where open ≡-Reasoning rb17 : suc (black-depth (node rkey vr rp-left rp-right) ⊔ (black-depth n1 ⊔ black-depth (PG.uncle pg))) ≡ black-depth (PG.grand pg) rb17 = begin suc (black-depth (node rkey vr rp-left rp-right) ⊔ (black-depth n1 ⊔ black-depth (PG.uncle pg))) ≡⟨ cong₂ (λ j k → suc (black-depth j ⊔ k)) eq (sym rb18) ⟩ suc (black-depth repl ⊔ black-depth repl) ≡⟨ ⊔-idem _ ⟩ suc (black-depth repl ) ≡⟨ cong suc (sym (RB-repl→eq _ _ (RBI.treerb r) rot)) ⟩ suc (black-depth (RBI.tree r) ) ≡⟨ cong suc (sym (proj1 (red-children-eq x (cong proj1 (sym rb13)) rb09))) ⟩ suc (black-depth (PG.parent pg) ) ≡⟨ sym (proj1 (black-children-eq refl (cong proj1 (sym rb14)) (subst (λ k → RBtreeInvariant k) x₁ rb02))) ⟩ black-depth (node kg vg (PG.parent pg) (PG.uncle pg)) ≡⟨ cong black-depth (sym x₁) ⟩ black-depth (PG.grand pg) ∎ where open ≡-Reasoning ... | s2-1sp2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | s2-s12p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | s2-1s2p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? replaceRBP1 : t replaceRBP1 with RBI.state r ... | rebuild ceq bdepth-eq rot = rebuildCase ceq bdepth-eq rot ... | top-black eq rot = exit stack (trans eq (cong (λ k → k ∷ []) rb00)) r where rb00 : RBI.tree r ≡ orig rb00 with si-property-last _ _ _ _ (subst (λ k → stackInvariant key (RBI.tree r) orig k) (eq) (RBI.si r)) ... | refl = refl ... | rotate repl-red pdepth-eq ¬leaf rot with stackToPG (RBI.tree r) orig stack (RBI.si r) ... | case1 eq = exit stack eq r -- no stack, replace top node ... | case2 (case1 eq) = insertCase4 orig refl eq rot (case1 repl-red) (RBI.si r) -- one level stack, orig is parent of repl ... | case2 (case2 pg) with color (PG.parent pg) in pcolor ... | Black = insertCase1 where -- Parent is Black, make color repl ≡ color tree then goto rebuildCase rb00 : (pg : PG (Color ∧ A) (RBI.tree r) stack) → stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) rb00 pg = subst (λ k → stackInvariant key (RBI.tree r) orig k) (PG.stack=gp pg) (RBI.si r) treerb : (pg : PG (Color ∧ A) (RBI.tree r) stack) → RBtreeInvariant (PG.parent pg) treerb pg = PGtoRBinvariant1 _ orig (RBI.origrb r) _ (popStackInvariant _ _ _ _ _ (rb00 pg)) rb08 : (pg : PG (Color ∧ A) (RBI.tree r) stack) → treeInvariant (PG.parent pg) rb08 pg = siToTreeinvariant _ _ _ _ (RBI.origti r) (popStackInvariant _ _ _ _ _ (rb00 pg)) insertCase1 : t insertCase1 with PG.pg pg ... | s2-s1p2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = next (PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) record { tree = PG.parent pg ; repl = rb01 ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = treerb pg ; replrb = rb06 ; si = popStackInvariant _ _ _ _ _ (rb00 pg) ; state = rebuild (cong color x) (rb05 (trans (sym (cong color x)) pcolor)) (subst (λ k → replacedRBTree key value k (node kp vp repl n1) ) (sym x) (rb02 rb04 )) } (subst₂ (λ j k → j < length k ) refl (sym (PG.stack=gp pg)) ≤-refl) where rb01 : bt (Color ∧ A) rb01 = node kp vp repl n1 rb03 : key < kp rb03 = si-property-< ¬leaf (subst (λ k → treeInvariant k) x (rb08 pg)) (subst (λ k → stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ k ∷ PG.grand pg ∷ PG.rest pg)) x (rb00 pg)) rb04 : ⟪ Black , proj2 vp ⟫ ≡ vp rb04 with subst (λ k → color k ≡ Black) x pcolor ... | refl = refl rb02 : ⟪ Black , proj2 vp ⟫ ≡ vp → replacedRBTree key value (node kp vp (RBI.tree r) n1) (node kp vp repl n1) rb02 eq = subst (λ k → replacedRBTree key value (node kp k (RBI.tree r) n1) (node kp k repl n1)) eq (rbr-black-left repl-red rb03 rot ) rb07 : black-depth repl ≡ black-depth n1 rb07 = begin black-depth repl ≡⟨ sym (RB-repl→eq _ _ (RBI.treerb r) rot) ⟩ black-depth (RBI.tree r) ≡⟨ RBtreeEQ (subst (λ k → RBtreeInvariant k) x (treerb pg)) ⟩ black-depth n1 ∎ where open ≡-Reasoning rb05 : proj1 vp ≡ Black → black-depth rb01 ≡ black-depth (PG.parent pg) rb05 refl = begin suc (black-depth repl ⊔ black-depth n1) ≡⟨ cong suc (cong (λ k → k ⊔ black-depth n1) pdepth-eq) ⟩ suc (black-depth (RBI.tree r) ⊔ black-depth n1) ≡⟨ refl ⟩ black-depth (node kp vp (RBI.tree r) n1) ≡⟨ cong black-depth (sym x) ⟩ black-depth (PG.parent pg) ∎ where open ≡-Reasoning rb06 : RBtreeInvariant rb01 rb06 = subst (λ k → RBtreeInvariant (node kp k repl n1)) rb04 ( rb-black _ _ rb07 (RBI.replrb r) (RBtreeRightDown _ _ (subst (λ k → RBtreeInvariant k) x (treerb pg)))) ... | s2-1sp2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | s2-s12p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | s2-1s2p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = ? ... | Red with PG.uncle pg in uneq ... | leaf = insertCase5 repl refl pg rot repl-red (cong color uneq) pcolor ... | node key₁ ⟪ Black , value₁ ⟫ t₁ t₂ = insertCase5 repl refl pg rot repl-red (cong color uneq) pcolor ... | node key₁ ⟪ Red , value₁ ⟫ t₁ t₂ with PG.pg pg -- insertCase2 : uncle and parent are Red, flip color and go up by two level ... | s2-s1p2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = next (PG.grand pg ∷ (PG.rest pg)) record { tree = PG.grand pg ; repl = to-red (node kg vg (to-black (node kp vp repl n1)) (to-black (PG.uncle pg))) ; origti = RBI.origti r ; origrb = RBI.origrb r ; treerb = rb01 ; replrb = rb-red _ _ refl (RBtreeToBlackColor _ rb02) rb12 (rb-black _ _ rb13 (RBI.replrb r) rb04) (RBtreeToBlack _ rb02) ; si = popStackInvariant _ _ _ _ _ (popStackInvariant _ _ _ _ _ rb00) ; state = rotate refl rb17 ? (subst₂ (λ j k → replacedRBTree key value j k ) (sym rb09) refl (rbr-flip-ll repl-red (rb05 refl uneq) rb06 rot)) } rb15 where rb00 : stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) rb00 = subst (λ k → stackInvariant key (RBI.tree r) orig k) (PG.stack=gp pg) (RBI.si r) rb01 : RBtreeInvariant (PG.grand pg) rb01 = PGtoRBinvariant1 _ orig (RBI.origrb r) _ (popStackInvariant _ _ _ _ _ (popStackInvariant _ _ _ _ _ rb00)) rb02 : RBtreeInvariant (PG.uncle pg) rb02 = RBtreeRightDown _ _ (subst (λ k → RBtreeInvariant k) x₁ rb01) rb03 : RBtreeInvariant (PG.parent pg) rb03 = RBtreeLeftDown _ _ (subst (λ k → RBtreeInvariant k) x₁ rb01) rb04 : RBtreeInvariant n1 rb04 = RBtreeRightDown _ _ (subst (λ k → RBtreeInvariant k) x rb03) rb05 : { tree : bt (Color ∧ A) } → tree ≡ PG.uncle pg → tree ≡ node key₁ ⟪ Red , value₁ ⟫ t₁ t₂ → color (PG.uncle pg) ≡ Red rb05 refl refl = refl rb08 : treeInvariant (PG.parent pg) rb08 = siToTreeinvariant _ _ _ _ (RBI.origti r) (popStackInvariant _ _ _ _ _ rb00) rb07 : treeInvariant (node kp vp (RBI.tree r) n1) rb07 = subst (λ k → treeInvariant k) x (siToTreeinvariant _ _ _ _ (RBI.origti r) (popStackInvariant _ _ _ _ _ rb00)) rb06 : key < kp rb06 = si-property-< ¬leaf rb07 (subst (λ k → stackInvariant key (RBI.tree r) orig (RBI.tree r ∷ k ∷ PG.grand pg ∷ PG.rest pg)) x rb00) rb10 : vg ≡ ⟪ Black , proj2 vg ⟫ rb10 with RBtreeParentColorBlack _ _ (subst (λ k → RBtreeInvariant k) x₁ rb01) (case1 pcolor) ... | refl = refl rb11 : vp ≡ ⟪ Red , proj2 vp ⟫ rb11 with subst (λ k → color k ≡ Red) x pcolor ... | refl = refl rb09 : PG.grand pg ≡ node kg ⟪ Black , proj2 vg ⟫ (node kp ⟪ Red , proj2 vp ⟫ (RBI.tree r) n1) (PG.uncle pg) rb09 = begin PG.grand pg ≡⟨ x₁ ⟩ node kg vg (PG.parent pg) (PG.uncle pg) ≡⟨ cong (λ k → node kg vg k (PG.uncle pg)) x ⟩ node kg vg (node kp vp (RBI.tree r) n1) (PG.uncle pg) ≡⟨ cong₂ (λ j k → node kg j (node kp k (RBI.tree r) n1) (PG.uncle pg)) rb10 rb11 ⟩ node kg ⟪ Black , proj2 vg ⟫ (node kp ⟪ Red , proj2 vp ⟫ (RBI.tree r) n1) (PG.uncle pg) ∎ where open ≡-Reasoning rb13 : black-depth repl ≡ black-depth n1 rb13 = begin black-depth repl ≡⟨ sym (RB-repl→eq _ _ (RBI.treerb r) rot) ⟩ black-depth (RBI.tree r) ≡⟨ RBtreeEQ (subst (λ k → RBtreeInvariant k) x rb03) ⟩ black-depth n1 ∎ where open ≡-Reasoning rb12 : suc (black-depth repl ⊔ black-depth n1) ≡ black-depth (to-black (PG.uncle pg)) rb12 = begin suc (black-depth repl ⊔ black-depth n1) ≡⟨ cong (λ k → suc (k ⊔ black-depth n1)) (sym (RB-repl→eq _ _ (RBI.treerb r) rot)) ⟩ suc (black-depth (RBI.tree r) ⊔ black-depth n1) ≡⟨ cong (λ k → suc (k ⊔ black-depth n1)) (RBtreeEQ (subst (λ k → RBtreeInvariant k) x rb03)) ⟩ suc (black-depth n1 ⊔ black-depth n1) ≡⟨ ⊔-idem _ ⟩ suc (black-depth n1 ) ≡⟨ cong suc (sym (proj2 (red-children-eq x (cong proj1 rb11) rb03 ))) ⟩ suc (black-depth (PG.parent pg)) ≡⟨ cong suc (RBtreeEQ (subst (λ k → RBtreeInvariant k) x₁ rb01)) ⟩ suc (black-depth (PG.uncle pg)) ≡⟨ to-black-eq (PG.uncle pg) (cong color uneq ) ⟩ black-depth (to-black (PG.uncle pg)) ∎ where open ≡-Reasoning rb17 : suc (black-depth repl ⊔ black-depth n1) ⊔ black-depth (to-black (PG.uncle pg)) ≡ black-depth (PG.grand pg) rb17 = begin suc (black-depth repl ⊔ black-depth n1) ⊔ black-depth (to-black (PG.uncle pg)) ≡⟨ cong (λ k → k ⊔ black-depth (to-black (PG.uncle pg))) rb12 ⟩ black-depth (to-black (PG.uncle pg)) ⊔ black-depth (to-black (PG.uncle pg)) ≡⟨ ⊔-idem _ ⟩ black-depth (to-black (PG.uncle pg)) ≡⟨ sym (to-black-eq (PG.uncle pg) (cong color uneq )) ⟩ suc (black-depth (PG.uncle pg)) ≡⟨ sym ( proj2 (black-children-eq x₁ (cong proj1 rb10) rb01 )) ⟩ black-depth (PG.grand pg) ∎ where open ≡-Reasoning rb15 : suc (length (PG.rest pg)) < length stack rb15 = begin suc (suc (length (PG.rest pg))) ≤⟨ <-trans (n<1+n _) (n<1+n _) ⟩ length (RBI.tree r ∷ PG.parent pg ∷ PG.grand pg ∷ PG.rest pg) ≡⟨ cong length (sym (PG.stack=gp pg)) ⟩ length stack ∎ where open ≤-Reasoning ... | s2-1sp2 {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = {!!} ... | s2-s12p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = {!!} ... | s2-1s2p {kp} {kg} {vp} {vg} {n1} {n2} x x₁ = {!!}