Mercurial > hg > Gears > GearsAgda
changeset 781:68904fdaab71
te
line wrap: on
line diff
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/#btree.agda# Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,541 @@ +module btree 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 ( _≟_ ; _≤?_ ; _≤_) + +data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where + t-leaf : treeInvariant leaf + t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) + t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key₁ value₁ t₁ t₂) + → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) + t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) + t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → key < key₁ → key₁ < key₂ + → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₂ value₂ t₃ t₄) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) + +-- +-- 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 + +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) +treeTest3 : bt ℕ +treeTest3 = node 2 5 (node 1 7 leaf leaf ) leaf +treeTest4 : bt ℕ +treeTest4 = node 5 5 leaf leaf +treeTest5 : bt ℕ +treeTest5 = node 1 7 leaf leaf + + +treeInvariantTest1 : treeInvariant treeTest1 +treeInvariantTest1 = t-right (m≤m+n _ 2) (t-node (add< 0) (add< 1) (t-left (add< 0) (t-single 1 7)) (t-single 5 5) ) + +treeInvariantTest2 : treeInvariant treeTest2 +treeInvariantTest2 = t-node (add< 0) (add< 1) (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) + +stackInvariantTest111 : stackInvariant 4 treeTest4 treeTest1 ( treeTest4 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +stackInvariantTest111 = s-right (add< 0) (s-right (add< 3) (s-nil)) + +stackInvariantTest11 : stackInvariant 4 treeTest4 treeTest1 ( treeTest4 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +stackInvariantTest11 = s-right (add< 0) (s-right (add< 3) (s-nil)) --treeTest4から見てみぎ、みぎnil + + +stackInvariantTest2' : stackInvariant 2 treeTest3 treeTest1 (treeTest3 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +stackInvariantTest2' = s-left (add< 0) (s-right (add< 1) (s-nil)) + +--stackInvariantTest121 : stackInvariant 2 treeTest5 treeTest1 ( treeTest5 ∷ treeTest3 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +--stackInvariantTest121 = s-left (_) (s-left (add< 0) (s-right (add< 1) (s-nil))) -- 2<2が示せない + +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-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 + +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 + +nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ +nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x +nat-<> : { x y : ℕ } → x < y → y < x → ⊥ +nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x + +open _∧_ + + +depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) +depth-1< {i} {j} = s≤s (m≤m⊔n _ j) + +depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) +depth-2< {i} {j} = s≤s + +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) --leafである +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) --探しているkeyと一致 +findP {n} {_} {A} key (node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree (tree ∷ st) --keyが求めているkey1より小さい。今いるツリーの左側をstackに積む。 +-- ⟪ treeLeftDown tree tree₁ (proj1 Pre) , s-left a (proj2 Pre)⟫ depth-1< --これでも通った。 + ⟪ 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 t) = t-right x t +replaceTree1 k v1 value (t-left x t) = t-left x t +replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ + +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 + +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 + +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 +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 (≤-step lt1) p1 lt1 ) where + TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → Invraiant r1 → reduce r1 < reduce r → t + TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 p1 (λ r2 p1 lt1 → ⊥-elim (lemma5 n≤j lt1)) + TerminatingLoop1 (suc j) r r1 n≤j p1 lt with <-cmp (reduce r1) (suc j) + ... | tri< a ¬b ¬c = TerminatingLoop1 j r r1 a p1 lt + ... | tri≈ ¬a b ¬c = loop r1 p1 (λ r2 p2 lt1 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b lt1 ) p2 lt1 ) + ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) +{- +open _∧_ + +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 ti) r-node = t-right x ti +RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI0 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ 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₁ ti) (r-right x ri) = t-single key₁ value₁ +RTtoTI0 (node key₁ _ leaf right@(node key₂ _ _ _)) (node key₁ value₁ leaf right₁@(node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = + t-right (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) +RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left x₁ 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₁ ti) (r-right x r-leaf) = + t-node x₁ x ti (t-single key value) +RTtoTI0 (node key₁ _ (node _ _ _ _) (node key₂ _ _ _)) (node key₁ _ (node _ _ _ _) (node key₃ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = + t-node x₁ (subst (λ k → key₁ < k) (rt-property-key ri) x₂) ti (RTtoTI0 _ _ key value ti₁ ri) +-- 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 (t-single _ _ ) +RTtoTI0 .(node _ _ leaf (node _ _ _ _)) (node key₁ value₁ (node key value leaf leaf) (node _ _ _ _)) key value (t-right x₁ ti) (r-left x r-leaf) = t-node x x₁ (t-single key value) ti +RTtoTI0 (node key₃ _ (node key₂ _ _ _) leaf) (node key₃ _ (node key₁ value₁ left left₁) leaf) key value (t-left x₁ ti) (r-left x ri) = + t-left (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) -- key₁ < key₃ +RTtoTI0 (node key₁ _ (node key₂ _ _ _) (node _ _ _ _)) (node key₁ _ (node key₃ _ _ _) (node _ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = t-node (subst (λ k → k < key₁ ) (rt-property-key ri) x₁) x₂ (RTtoTI0 _ _ key value ti ri) ti₁ + +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 ti) r-node = t-right x ti +RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI1 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ ti ti₁ +-- r-right case +RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ leaf (node _ _ _ _)) key value (t-right x₁ 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₁ 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₂ 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₂ ti ti₁) (r-right x ri) = t-node x₁ (subst (λ k → key₄ < k ) (sym (rt-property-key ri)) x₂) 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₁ 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₁ 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₂ ti ti₁) (r-left x r-leaf) = t-right x₂ ti₁ +RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = + t-node (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ (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 +-} \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/#hoareBinar# Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,2 @@ + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/COMMIT_EDITMSG Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +7/7
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/FETCH_HEAD Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +baeb8f1be6e05a9a3a38e619a9c9e0496e8da3e1 branch 'master' of github.com:e205718/GearsAgda
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/HEAD Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +ref: refs/heads/master
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/ORIG_HEAD Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +0487cd0857ea20b08f8080576621afa1ab1c17c0
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/config Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,13 @@ +[core] + repositoryformatversion = 0 + filemode = true + bare = false + logallrefupdates = true + ignorecase = true + precomposeunicode = true +[remote "origin"] + url = git@github.com:e205718/GearsAgda.git + fetch = +refs/heads/*:refs/remotes/origin/* +[branch "master"] + remote = origin + merge = refs/heads/master
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/description Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +Unnamed repository; edit this file 'description' to name the repository.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/applypatch-msg.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,15 @@ +#!/bin/sh +# +# An example hook script to check the commit log message taken by +# applypatch from an e-mail message. +# +# The hook should exit with non-zero status after issuing an +# appropriate message if it wants to stop the commit. The hook is +# allowed to edit the commit message file. +# +# To enable this hook, rename this file to "applypatch-msg". + +. git-sh-setup +commitmsg="$(git rev-parse --git-path hooks/commit-msg)" +test -x "$commitmsg" && exec "$commitmsg" ${1+"$@"} +:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/commit-msg.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,24 @@ +#!/bin/sh +# +# An example hook script to check the commit log message. +# Called by "git commit" with one argument, the name of the file +# that has the commit message. The hook should exit with non-zero +# status after issuing an appropriate message if it wants to stop the +# commit. The hook is allowed to edit the commit message file. +# +# To enable this hook, rename this file to "commit-msg". + +# Uncomment the below to add a Signed-off-by line to the message. +# Doing this in a hook is a bad idea in general, but the prepare-commit-msg +# hook is more suited to it. +# +# SOB=$(git var GIT_AUTHOR_IDENT | sed -n 's/^\(.*>\).*$/Signed-off-by: \1/p') +# grep -qs "^$SOB" "$1" || echo "$SOB" >> "$1" + +# This example catches duplicate Signed-off-by lines. + +test "" = "$(grep '^Signed-off-by: ' "$1" | + sort | uniq -c | sed -e '/^[ ]*1[ ]/d')" || { + echo >&2 Duplicate Signed-off-by lines. + exit 1 +}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/fsmonitor-watchman.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,114 @@ +#!/usr/bin/perl + +use strict; +use warnings; +use IPC::Open2; + +# An example hook script to integrate Watchman +# (https://facebook.github.io/watchman/) with git to speed up detecting +# new and modified files. +# +# The hook is passed a version (currently 1) and a time in nanoseconds +# formatted as a string and outputs to stdout all files that have been +# modified since the given time. Paths must be relative to the root of +# the working tree and separated by a single NUL. +# +# To enable this hook, rename this file to "query-watchman" and set +# 'git config core.fsmonitor .git/hooks/query-watchman' +# +my ($version, $time) = @ARGV; + +# Check the hook interface version + +if ($version == 1) { + # convert nanoseconds to seconds + $time = int $time / 1000000000; +} else { + die "Unsupported query-fsmonitor hook version '$version'.\n" . + "Falling back to scanning...\n"; +} + +my $git_work_tree; +if ($^O =~ 'msys' || $^O =~ 'cygwin') { + $git_work_tree = Win32::GetCwd(); + $git_work_tree =~ tr/\\/\//; +} else { + require Cwd; + $git_work_tree = Cwd::cwd(); +} + +my $retry = 1; + +launch_watchman(); + +sub launch_watchman { + + my $pid = open2(\*CHLD_OUT, \*CHLD_IN, 'watchman -j --no-pretty') + or die "open2() failed: $!\n" . + "Falling back to scanning...\n"; + + # In the query expression below we're asking for names of files that + # changed since $time but were not transient (ie created after + # $time but no longer exist). + # + # To accomplish this, we're using the "since" generator to use the + # recency index to select candidate nodes and "fields" to limit the + # output to file names only. Then we're using the "expression" term to + # further constrain the results. + # + # The category of transient files that we want to ignore will have a + # creation clock (cclock) newer than $time_t value and will also not + # currently exist. + + my $query = <<" END"; + ["query", "$git_work_tree", { + "since": $time, + "fields": ["name"], + "expression": ["not", ["allof", ["since", $time, "cclock"], ["not", "exists"]]] + }] + END + + print CHLD_IN $query; + close CHLD_IN; + my $response = do {local $/; <CHLD_OUT>}; + + die "Watchman: command returned no output.\n" . + "Falling back to scanning...\n" if $response eq ""; + die "Watchman: command returned invalid output: $response\n" . + "Falling back to scanning...\n" unless $response =~ /^\{/; + + my $json_pkg; + eval { + require JSON::XS; + $json_pkg = "JSON::XS"; + 1; + } or do { + require JSON::PP; + $json_pkg = "JSON::PP"; + }; + + my $o = $json_pkg->new->utf8->decode($response); + + if ($retry > 0 and $o->{error} and $o->{error} =~ m/unable to resolve root .* directory (.*) is not watched/) { + print STDERR "Adding '$git_work_tree' to watchman's watch list.\n"; + $retry--; + qx/watchman watch "$git_work_tree"/; + die "Failed to make watchman watch '$git_work_tree'.\n" . + "Falling back to scanning...\n" if $? != 0; + + # Watchman will always return all files on the first query so + # return the fast "everything is dirty" flag to git and do the + # Watchman query just to get it over with now so we won't pay + # the cost in git to look up each individual file. + print "/\0"; + eval { launch_watchman() }; + exit 0; + } + + die "Watchman: $o->{error}.\n" . + "Falling back to scanning...\n" if $o->{error}; + + binmode STDOUT, ":utf8"; + local $, = "\0"; + print @{$o->{files}}; +}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/post-update.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,8 @@ +#!/bin/sh +# +# An example hook script to prepare a packed repository for use over +# dumb transports. +# +# To enable this hook, rename this file to "post-update". + +exec git update-server-info
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/pre-applypatch.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,14 @@ +#!/bin/sh +# +# An example hook script to verify what is about to be committed +# by applypatch from an e-mail message. +# +# The hook should exit with non-zero status after issuing an +# appropriate message if it wants to stop the commit. +# +# To enable this hook, rename this file to "pre-applypatch". + +. git-sh-setup +precommit="$(git rev-parse --git-path hooks/pre-commit)" +test -x "$precommit" && exec "$precommit" ${1+"$@"} +:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/pre-commit.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,49 @@ +#!/bin/sh +# +# An example hook script to verify what is about to be committed. +# Called by "git commit" with no arguments. The hook should +# exit with non-zero status after issuing an appropriate message if +# it wants to stop the commit. +# +# To enable this hook, rename this file to "pre-commit". + +if git rev-parse --verify HEAD >/dev/null 2>&1 +then + against=HEAD +else + # Initial commit: diff against an empty tree object + against=$(git hash-object -t tree /dev/null) +fi + +# If you want to allow non-ASCII filenames set this variable to true. +allownonascii=$(git config --bool hooks.allownonascii) + +# Redirect output to stderr. +exec 1>&2 + +# Cross platform projects tend to avoid non-ASCII filenames; prevent +# them from being added to the repository. We exploit the fact that the +# printable range starts at the space character and ends with tilde. +if [ "$allownonascii" != "true" ] && + # Note that the use of brackets around a tr range is ok here, (it's + # even required, for portability to Solaris 10's /usr/bin/tr), since + # the square bracket bytes happen to fall in the designated range. + test $(git diff --cached --name-only --diff-filter=A -z $against | + LC_ALL=C tr -d '[ -~]\0' | wc -c) != 0 +then + cat <<\EOF +Error: Attempt to add a non-ASCII file name. + +This can cause problems if you want to work with people on other platforms. + +To be portable it is advisable to rename the file. + +If you know what you are doing you can disable this check using: + + git config hooks.allownonascii true +EOF + exit 1 +fi + +# If there are whitespace errors, print the offending file names and fail. +exec git diff-index --check --cached $against --
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/pre-merge-commit.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,13 @@ +#!/bin/sh +# +# An example hook script to verify what is about to be committed. +# Called by "git merge" with no arguments. The hook should +# exit with non-zero status after issuing an appropriate message to +# stderr if it wants to stop the merge commit. +# +# To enable this hook, rename this file to "pre-merge-commit". + +. git-sh-setup +test -x "$GIT_DIR/hooks/pre-commit" && + exec "$GIT_DIR/hooks/pre-commit" +:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/pre-push.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,53 @@ +#!/bin/sh + +# An example hook script to verify what is about to be pushed. Called by "git +# push" after it has checked the remote status, but before anything has been +# pushed. If this script exits with a non-zero status nothing will be pushed. +# +# This hook is called with the following parameters: +# +# $1 -- Name of the remote to which the push is being done +# $2 -- URL to which the push is being done +# +# If pushing without using a named remote those arguments will be equal. +# +# Information about the commits which are being pushed is supplied as lines to +# the standard input in the form: +# +# <local ref> <local sha1> <remote ref> <remote sha1> +# +# This sample shows how to prevent push of commits where the log message starts +# with "WIP" (work in progress). + +remote="$1" +url="$2" + +z40=0000000000000000000000000000000000000000 + +while read local_ref local_sha remote_ref remote_sha +do + if [ "$local_sha" = $z40 ] + then + # Handle delete + : + else + if [ "$remote_sha" = $z40 ] + then + # New branch, examine all commits + range="$local_sha" + else + # Update to existing branch, examine new commits + range="$remote_sha..$local_sha" + fi + + # Check for WIP commit + commit=`git rev-list -n 1 --grep '^WIP' "$range"` + if [ -n "$commit" ] + then + echo >&2 "Found WIP commit in $local_ref, not pushing" + exit 1 + fi + fi +done + +exit 0
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/pre-rebase.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,169 @@ +#!/bin/sh +# +# Copyright (c) 2006, 2008 Junio C Hamano +# +# The "pre-rebase" hook is run just before "git rebase" starts doing +# its job, and can prevent the command from running by exiting with +# non-zero status. +# +# The hook is called with the following parameters: +# +# $1 -- the upstream the series was forked from. +# $2 -- the branch being rebased (or empty when rebasing the current branch). +# +# This sample shows how to prevent topic branches that are already +# merged to 'next' branch from getting rebased, because allowing it +# would result in rebasing already published history. + +publish=next +basebranch="$1" +if test "$#" = 2 +then + topic="refs/heads/$2" +else + topic=`git symbolic-ref HEAD` || + exit 0 ;# we do not interrupt rebasing detached HEAD +fi + +case "$topic" in +refs/heads/??/*) + ;; +*) + exit 0 ;# we do not interrupt others. + ;; +esac + +# Now we are dealing with a topic branch being rebased +# on top of master. Is it OK to rebase it? + +# Does the topic really exist? +git show-ref -q "$topic" || { + echo >&2 "No such branch $topic" + exit 1 +} + +# Is topic fully merged to master? +not_in_master=`git rev-list --pretty=oneline ^master "$topic"` +if test -z "$not_in_master" +then + echo >&2 "$topic is fully merged to master; better remove it." + exit 1 ;# we could allow it, but there is no point. +fi + +# Is topic ever merged to next? If so you should not be rebasing it. +only_next_1=`git rev-list ^master "^$topic" ${publish} | sort` +only_next_2=`git rev-list ^master ${publish} | sort` +if test "$only_next_1" = "$only_next_2" +then + not_in_topic=`git rev-list "^$topic" master` + if test -z "$not_in_topic" + then + echo >&2 "$topic is already up to date with master" + exit 1 ;# we could allow it, but there is no point. + else + exit 0 + fi +else + not_in_next=`git rev-list --pretty=oneline ^${publish} "$topic"` + /usr/bin/perl -e ' + my $topic = $ARGV[0]; + my $msg = "* $topic has commits already merged to public branch:\n"; + my (%not_in_next) = map { + /^([0-9a-f]+) /; + ($1 => 1); + } split(/\n/, $ARGV[1]); + for my $elem (map { + /^([0-9a-f]+) (.*)$/; + [$1 => $2]; + } split(/\n/, $ARGV[2])) { + if (!exists $not_in_next{$elem->[0]}) { + if ($msg) { + print STDERR $msg; + undef $msg; + } + print STDERR " $elem->[1]\n"; + } + } + ' "$topic" "$not_in_next" "$not_in_master" + exit 1 +fi + +<<\DOC_END + +This sample hook safeguards topic branches that have been +published from being rewound. + +The workflow assumed here is: + + * Once a topic branch forks from "master", "master" is never + merged into it again (either directly or indirectly). + + * Once a topic branch is fully cooked and merged into "master", + it is deleted. If you need to build on top of it to correct + earlier mistakes, a new topic branch is created by forking at + the tip of the "master". This is not strictly necessary, but + it makes it easier to keep your history simple. + + * Whenever you need to test or publish your changes to topic + branches, merge them into "next" branch. + +The script, being an example, hardcodes the publish branch name +to be "next", but it is trivial to make it configurable via +$GIT_DIR/config mechanism. + +With this workflow, you would want to know: + +(1) ... if a topic branch has ever been merged to "next". Young + topic branches can have stupid mistakes you would rather + clean up before publishing, and things that have not been + merged into other branches can be easily rebased without + affecting other people. But once it is published, you would + not want to rewind it. + +(2) ... if a topic branch has been fully merged to "master". + Then you can delete it. More importantly, you should not + build on top of it -- other people may already want to + change things related to the topic as patches against your + "master", so if you need further changes, it is better to + fork the topic (perhaps with the same name) afresh from the + tip of "master". + +Let's look at this example: + + o---o---o---o---o---o---o---o---o---o "next" + / / / / + / a---a---b A / / + / / / / + / / c---c---c---c B / + / / / \ / + / / / b---b C \ / + / / / / \ / + ---o---o---o---o---o---o---o---o---o---o---o "master" + + +A, B and C are topic branches. + + * A has one fix since it was merged up to "next". + + * B has finished. It has been fully merged up to "master" and "next", + and is ready to be deleted. + + * C has not merged to "next" at all. + +We would want to allow C to be rebased, refuse A, and encourage +B to be deleted. + +To compute (1): + + git rev-list ^master ^topic next + git rev-list ^master next + + if these match, topic has not merged in next at all. + +To compute (2): + + git rev-list master..topic + + if this is empty, it is fully merged to "master". + +DOC_END
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/pre-receive.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,24 @@ +#!/bin/sh +# +# An example hook script to make use of push options. +# The example simply echoes all push options that start with 'echoback=' +# and rejects all pushes when the "reject" push option is used. +# +# To enable this hook, rename this file to "pre-receive". + +if test -n "$GIT_PUSH_OPTION_COUNT" +then + i=0 + while test "$i" -lt "$GIT_PUSH_OPTION_COUNT" + do + eval "value=\$GIT_PUSH_OPTION_$i" + case "$value" in + echoback=*) + echo "echo from the pre-receive-hook: ${value#*=}" >&2 + ;; + reject) + exit 1 + esac + i=$((i + 1)) + done +fi
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/prepare-commit-msg.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,42 @@ +#!/bin/sh +# +# An example hook script to prepare the commit log message. +# Called by "git commit" with the name of the file that has the +# commit message, followed by the description of the commit +# message's source. The hook's purpose is to edit the commit +# message file. If the hook fails with a non-zero status, +# the commit is aborted. +# +# To enable this hook, rename this file to "prepare-commit-msg". + +# This hook includes three examples. The first one removes the +# "# Please enter the commit message..." help message. +# +# The second includes the output of "git diff --name-status -r" +# into the message, just before the "git status" output. It is +# commented because it doesn't cope with --amend or with squashed +# commits. +# +# The third example adds a Signed-off-by line to the message, that can +# still be edited. This is rarely a good idea. + +COMMIT_MSG_FILE=$1 +COMMIT_SOURCE=$2 +SHA1=$3 + +/usr/bin/perl -i.bak -ne 'print unless(m/^. Please enter the commit message/..m/^#$/)' "$COMMIT_MSG_FILE" + +# case "$COMMIT_SOURCE,$SHA1" in +# ,|template,) +# /usr/bin/perl -i.bak -pe ' +# print "\n" . `git diff --cached --name-status -r` +# if /^#/ && $first++ == 0' "$COMMIT_MSG_FILE" ;; +# *) ;; +# esac + +# SOB=$(git var GIT_COMMITTER_IDENT | sed -n 's/^\(.*>\).*$/Signed-off-by: \1/p') +# git interpret-trailers --in-place --trailer "$SOB" "$COMMIT_MSG_FILE" +# if test -z "$COMMIT_SOURCE" +# then +# /usr/bin/perl -i.bak -pe 'print "\n" if !$first_line++' "$COMMIT_MSG_FILE" +# fi
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/push-to-checkout.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,78 @@ +#!/bin/sh + +# An example hook script to update a checked-out tree on a git push. +# +# This hook is invoked by git-receive-pack(1) when it reacts to git +# push and updates reference(s) in its repository, and when the push +# tries to update the branch that is currently checked out and the +# receive.denyCurrentBranch configuration variable is set to +# updateInstead. +# +# By default, such a push is refused if the working tree and the index +# of the remote repository has any difference from the currently +# checked out commit; when both the working tree and the index match +# the current commit, they are updated to match the newly pushed tip +# of the branch. This hook is to be used to override the default +# behaviour; however the code below reimplements the default behaviour +# as a starting point for convenient modification. +# +# The hook receives the commit with which the tip of the current +# branch is going to be updated: +commit=$1 + +# It can exit with a non-zero status to refuse the push (when it does +# so, it must not modify the index or the working tree). +die () { + echo >&2 "$*" + exit 1 +} + +# Or it can make any necessary changes to the working tree and to the +# index to bring them to the desired state when the tip of the current +# branch is updated to the new commit, and exit with a zero status. +# +# For example, the hook can simply run git read-tree -u -m HEAD "$1" +# in order to emulate git fetch that is run in the reverse direction +# with git push, as the two-tree form of git read-tree -u -m is +# essentially the same as git switch or git checkout that switches +# branches while keeping the local changes in the working tree that do +# not interfere with the difference between the branches. + +# The below is a more-or-less exact translation to shell of the C code +# for the default behaviour for git's push-to-checkout hook defined in +# the push_to_deploy() function in builtin/receive-pack.c. +# +# Note that the hook will be executed from the repository directory, +# not from the working tree, so if you want to perform operations on +# the working tree, you will have to adapt your code accordingly, e.g. +# by adding "cd .." or using relative paths. + +if ! git update-index -q --ignore-submodules --refresh +then + die "Up-to-date check failed" +fi + +if ! git diff-files --quiet --ignore-submodules -- +then + die "Working directory has unstaged changes" +fi + +# This is a rough translation of: +# +# head_has_history() ? "HEAD" : EMPTY_TREE_SHA1_HEX +if git cat-file -e HEAD 2>/dev/null +then + head=HEAD +else + head=$(git hash-object -t tree --stdin </dev/null) +fi + +if ! git diff-index --quiet --cached --ignore-submodules $head -- +then + die "Working directory has staged changes" +fi + +if ! git read-tree -u -m "$commit" +then + die "Could not update working tree to new HEAD" +fi
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/hooks/update.sample Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,128 @@ +#!/bin/sh +# +# An example hook script to block unannotated tags from entering. +# Called by "git receive-pack" with arguments: refname sha1-old sha1-new +# +# To enable this hook, rename this file to "update". +# +# Config +# ------ +# hooks.allowunannotated +# This boolean sets whether unannotated tags will be allowed into the +# repository. By default they won't be. +# hooks.allowdeletetag +# This boolean sets whether deleting tags will be allowed in the +# repository. By default they won't be. +# hooks.allowmodifytag +# This boolean sets whether a tag may be modified after creation. By default +# it won't be. +# hooks.allowdeletebranch +# This boolean sets whether deleting branches will be allowed in the +# repository. By default they won't be. +# hooks.denycreatebranch +# This boolean sets whether remotely creating branches will be denied +# in the repository. By default this is allowed. +# + +# --- Command line +refname="$1" +oldrev="$2" +newrev="$3" + +# --- Safety check +if [ -z "$GIT_DIR" ]; then + echo "Don't run this script from the command line." >&2 + echo " (if you want, you could supply GIT_DIR then run" >&2 + echo " $0 <ref> <oldrev> <newrev>)" >&2 + exit 1 +fi + +if [ -z "$refname" -o -z "$oldrev" -o -z "$newrev" ]; then + echo "usage: $0 <ref> <oldrev> <newrev>" >&2 + exit 1 +fi + +# --- Config +allowunannotated=$(git config --bool hooks.allowunannotated) +allowdeletebranch=$(git config --bool hooks.allowdeletebranch) +denycreatebranch=$(git config --bool hooks.denycreatebranch) +allowdeletetag=$(git config --bool hooks.allowdeletetag) +allowmodifytag=$(git config --bool hooks.allowmodifytag) + +# check for no description +projectdesc=$(sed -e '1q' "$GIT_DIR/description") +case "$projectdesc" in +"Unnamed repository"* | "") + echo "*** Project description file hasn't been set" >&2 + exit 1 + ;; +esac + +# --- Check types +# if $newrev is 0000...0000, it's a commit to delete a ref. +zero="0000000000000000000000000000000000000000" +if [ "$newrev" = "$zero" ]; then + newrev_type=delete +else + newrev_type=$(git cat-file -t $newrev) +fi + +case "$refname","$newrev_type" in + refs/tags/*,commit) + # un-annotated tag + short_refname=${refname##refs/tags/} + if [ "$allowunannotated" != "true" ]; then + echo "*** The un-annotated tag, $short_refname, is not allowed in this repository" >&2 + echo "*** Use 'git tag [ -a | -s ]' for tags you want to propagate." >&2 + exit 1 + fi + ;; + refs/tags/*,delete) + # delete tag + if [ "$allowdeletetag" != "true" ]; then + echo "*** Deleting a tag is not allowed in this repository" >&2 + exit 1 + fi + ;; + refs/tags/*,tag) + # annotated tag + if [ "$allowmodifytag" != "true" ] && git rev-parse $refname > /dev/null 2>&1 + then + echo "*** Tag '$refname' already exists." >&2 + echo "*** Modifying a tag is not allowed in this repository." >&2 + exit 1 + fi + ;; + refs/heads/*,commit) + # branch + if [ "$oldrev" = "$zero" -a "$denycreatebranch" = "true" ]; then + echo "*** Creating a branch is not allowed in this repository" >&2 + exit 1 + fi + ;; + refs/heads/*,delete) + # delete branch + if [ "$allowdeletebranch" != "true" ]; then + echo "*** Deleting a branch is not allowed in this repository" >&2 + exit 1 + fi + ;; + refs/remotes/*,commit) + # tracking branch + ;; + refs/remotes/*,delete) + # delete tracking branch + if [ "$allowdeletebranch" != "true" ]; then + echo "*** Deleting a tracking branch is not allowed in this repository" >&2 + exit 1 + fi + ;; + *) + # Anything else (is there anything else?) + echo "*** Update hook: unknown type of update to ref $refname of type $newrev_type" >&2 + exit 1 + ;; +esac + +# --- Finished +exit 0
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/info/exclude Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,6 @@ +# git ls-files --others --exclude-from=.git/info/exclude +# Lines that start with '#' are comments. +# For a project mostly in C, the following would be a good set of +# exclude patterns (uncomment them if you want to use them): +# *.[oa] +# *~
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/logs/HEAD Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,11 @@ +0000000000000000000000000000000000000000 ac55a36987a9273979425ebe2b0d7434f57b1590 e205718 <e205718@ie.u-ryukyu.ac.jp> 1683789831 +0900 clone: from git@github.com:e205718/GearsAgda.git +ac55a36987a9273979425ebe2b0d7434f57b1590 8da83e4c94275b301ebeadd753645bda2dceefdc e205718 <e205718@ie.u-ryukyu.ac.jp> 1685239824 +0900 commit: work +8da83e4c94275b301ebeadd753645bda2dceefdc fd86b13af379a711d7ef0b07d0b2460e8088e14f e205718 <e205718@ie.u-ryukyu.ac.jp> 1685776571 +0900 commit: s +fd86b13af379a711d7ef0b07d0b2460e8088e14f 926d3938256a6b2afc28f042373194a692a3db85 e205718 <e205718@ie.u-ryukyu.ac.jp> 1686213376 +0900 commit: from amane +926d3938256a6b2afc28f042373194a692a3db85 de4c763a3b65d439157a28bca6e574e7e6bd6f48 e205718 <e205718@ie.u-ryukyu.ac.jp> 1686560989 +0900 commit: amane +de4c763a3b65d439157a28bca6e574e7e6bd6f48 e5d4bfc497171d8a659cb1ddd695af3e05a92105 e205718 <e205718@ie.u-ryukyu.ac.jp> 1687159862 +0900 commit: 6/19 +e5d4bfc497171d8a659cb1ddd695af3e05a92105 69d6285765c99fd5afe14cafd957ccbccc7d348e e205718 <e205718@ie.u-ryukyu.ac.jp> 1687778424 +0900 commit: remove black depth +69d6285765c99fd5afe14cafd957ccbccc7d348e 46c240f5b42cb03db61137e92d2f67892b410362 e205718 <e205718@ie.u-ryukyu.ac.jp> 1687779007 +0900 commit: d +46c240f5b42cb03db61137e92d2f67892b410362 0487cd0857ea20b08f8080576621afa1ab1c17c0 e205718 <e205718@ie.u-ryukyu.ac.jp> 1688718982 +0900 commit: 7/7 +0487cd0857ea20b08f8080576621afa1ab1c17c0 0487cd0857ea20b08f8080576621afa1ab1c17c0 e205718 <e205718@ie.u-ryukyu.ac.jp> 1688719191 +0900 reset: moving to HEAD +0487cd0857ea20b08f8080576621afa1ab1c17c0 2edd8b993212ee33e85ac9b10b3a413a8f4dced7 e205718 <e205718@ie.u-ryukyu.ac.jp> 1688719228 +0900 merge origin/master: Merge made by the 'ort' strategy.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/logs/refs/heads/master Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,10 @@ +0000000000000000000000000000000000000000 ac55a36987a9273979425ebe2b0d7434f57b1590 e205718 <e205718@ie.u-ryukyu.ac.jp> 1683789831 +0900 clone: from git@github.com:e205718/GearsAgda.git +ac55a36987a9273979425ebe2b0d7434f57b1590 8da83e4c94275b301ebeadd753645bda2dceefdc e205718 <e205718@ie.u-ryukyu.ac.jp> 1685239824 +0900 commit: work +8da83e4c94275b301ebeadd753645bda2dceefdc fd86b13af379a711d7ef0b07d0b2460e8088e14f e205718 <e205718@ie.u-ryukyu.ac.jp> 1685776571 +0900 commit: s +fd86b13af379a711d7ef0b07d0b2460e8088e14f 926d3938256a6b2afc28f042373194a692a3db85 e205718 <e205718@ie.u-ryukyu.ac.jp> 1686213376 +0900 commit: from amane +926d3938256a6b2afc28f042373194a692a3db85 de4c763a3b65d439157a28bca6e574e7e6bd6f48 e205718 <e205718@ie.u-ryukyu.ac.jp> 1686560989 +0900 commit: amane +de4c763a3b65d439157a28bca6e574e7e6bd6f48 e5d4bfc497171d8a659cb1ddd695af3e05a92105 e205718 <e205718@ie.u-ryukyu.ac.jp> 1687159862 +0900 commit: 6/19 +e5d4bfc497171d8a659cb1ddd695af3e05a92105 69d6285765c99fd5afe14cafd957ccbccc7d348e e205718 <e205718@ie.u-ryukyu.ac.jp> 1687778424 +0900 commit: remove black depth +69d6285765c99fd5afe14cafd957ccbccc7d348e 46c240f5b42cb03db61137e92d2f67892b410362 e205718 <e205718@ie.u-ryukyu.ac.jp> 1687779007 +0900 commit: d +46c240f5b42cb03db61137e92d2f67892b410362 0487cd0857ea20b08f8080576621afa1ab1c17c0 e205718 <e205718@ie.u-ryukyu.ac.jp> 1688718982 +0900 commit: 7/7 +0487cd0857ea20b08f8080576621afa1ab1c17c0 2edd8b993212ee33e85ac9b10b3a413a8f4dced7 e205718 <e205718@ie.u-ryukyu.ac.jp> 1688719228 +0900 merge origin/master: Merge made by the 'ort' strategy.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/logs/refs/remotes/origin/HEAD Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +0000000000000000000000000000000000000000 ac55a36987a9273979425ebe2b0d7434f57b1590 e205718 <e205718@ie.u-ryukyu.ac.jp> 1683789831 +0900 clone: from git@github.com:e205718/GearsAgda.git
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/logs/refs/remotes/origin/master Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,9 @@ +ac55a36987a9273979425ebe2b0d7434f57b1590 8da83e4c94275b301ebeadd753645bda2dceefdc e205718 <e205718@ie.u-ryukyu.ac.jp> 1685239838 +0900 update by push +8da83e4c94275b301ebeadd753645bda2dceefdc fd86b13af379a711d7ef0b07d0b2460e8088e14f e205718 <e205718@ie.u-ryukyu.ac.jp> 1685776624 +0900 update by push +fd86b13af379a711d7ef0b07d0b2460e8088e14f 926d3938256a6b2afc28f042373194a692a3db85 e205718 <e205718@ie.u-ryukyu.ac.jp> 1686213386 +0900 update by push +926d3938256a6b2afc28f042373194a692a3db85 de4c763a3b65d439157a28bca6e574e7e6bd6f48 e205718 <e205718@ie.u-ryukyu.ac.jp> 1686560995 +0900 update by push +de4c763a3b65d439157a28bca6e574e7e6bd6f48 e5d4bfc497171d8a659cb1ddd695af3e05a92105 e205718 <e205718@ie.u-ryukyu.ac.jp> 1687159869 +0900 update by push +e5d4bfc497171d8a659cb1ddd695af3e05a92105 69d6285765c99fd5afe14cafd957ccbccc7d348e e205718 <e205718@ie.u-ryukyu.ac.jp> 1687778432 +0900 update by push +69d6285765c99fd5afe14cafd957ccbccc7d348e 46c240f5b42cb03db61137e92d2f67892b410362 e205718 <e205718@ie.u-ryukyu.ac.jp> 1687779013 +0900 update by push +46c240f5b42cb03db61137e92d2f67892b410362 baeb8f1be6e05a9a3a38e619a9c9e0496e8da3e1 e205718 <e205718@ie.u-ryukyu.ac.jp> 1688719065 +0900 fetch: fast-forward +baeb8f1be6e05a9a3a38e619a9c9e0496e8da3e1 2edd8b993212ee33e85ac9b10b3a413a8f4dced7 e205718 <e205718@ie.u-ryukyu.ac.jp> 1688719257 +0900 update by push
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/04/87cd0857ea20b08f8080576621afa1ab1c17c0 Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,3 @@ +xK +0@]ًu2M2*|Jm)ɢW{eжR4GLbȄBRClj+jI`0fXG{/.LЗܦ2ӶއzƇq8>l\љ +q? \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/35/ecdc8b1245db239dcb3237aaaccb114bb55a3b Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,3 @@ +xU10E}JlAX`-bc"'EV )lF[ӎ(\.hy`@W +tWȦ7 +]Im 9,*?3Fj=K?XV!$[of&E \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/52/fc61663c194230a5807e6605e9834e68855aad Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,2 @@ +x-ʻ 0@jO) 8#8LOZu*Q +hrT8kv&P;D/FA/|! \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/5f/589889a6cf31c5c32eb9b915a625cb3330bda2 Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +xKOR060```f`d7?!LƳ+^2rJqT=ت7'(;!nuenMu:бpқ~GhcVeM͓( Uosjw,/vo|w>-yuNJuJ'u(2=4k4,R95W~Sܿv^>Xw^p}SVdL|$C,95-:&X ;رZImBgETU8e?_ƿڝ5 \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/69/d6285765c99fd5afe14cafd957ccbccc7d348e Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,3 @@ +xK + @/ qB)#'MSJOm<Ъx=J=%E`e +ǁ'sR;$6 9MT]\@9筱pĀ;\U@zF~@gF \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/78/6af45fcfa479b23eafbf8abfe6d46ca9c3347c Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,2 @@ +xKOR025```f`d]&P>ۮpAڪnOK~c8T=ت7AKE) ߦ +2{N]?qMA=vռ,J}h*zqDd΄=oP=8bsV\ne\xqXUlT)eKq ^k"d]ع)iZo|?Ro;YX^h](CC+NO^^< Cmi \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/82/b91ab401f24b107f9c2b9a1a199c03e355fa9a Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +x+)JMU043f040031Q)JO+I(/,Ɉ/K-LL-KdGZgxVj$33W)bCUmoŵlCe||AiI||r1آlYy_;7:ᅡkBD \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/8d/a83e4c94275b301ebeadd753645bda2dceefdc Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,2 @@ +xA +0@Q9Ebt2 x$b-%Ho/'p7~YejNZiG>#2H YP)RW0'1H(%Ț3S?dLcVX\_'K=ػTvF{{`wܼ:E.@ \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/90/5ab3ca700bca01f3cd5c2b065caec485445a52 Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +x <H`]tu3d+ f^g+a! \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/92/6d3938256a6b2afc28f042373194a692a3db85 Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,2 @@ +xA + @Ѯ=0Q(WuiID}̓u]6&zќj6B%qaH$ڹɻZOhZLB*A 8Wţ?ŀ#"Ӹc1qMޠ".rU۶j^-+|D \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/93/7e41a35cb2e20b8066f540cb188da2b42a3e86 Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +xm1K1ɶJ*ԥԓv"jQk.\L0OpwK͗PnySڠ8.&,<n{c3ΤҐdCX70JNL|30wNuqCp|bUTTv@'@Qo#ͣXu&܀u=ά2#i|I'W^T7iư3jX .5ٶemNjUEI4艹qRc73@jf5M \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/ce/173ce78c80bbaf25fb5ae91d3c963aad3da94c Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +xKOR0d0b(H,N-/)f`0bHKNLHE \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/objects/e5/d4bfc497171d8a659cb1ddd695af3e05a92105 Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,2 @@ +xM +0@a9Ebd"^e2bREo'pm<^yj9h>2R5HTq.LjM^MW /kDHLA@0T}ٴ8E}6Oޟ{z0٘>ld?sg4B \ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/packed-refs Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,2 @@ +# pack-refs with: peeled fully-peeled sorted +ac55a36987a9273979425ebe2b0d7434f57b1590 refs/remotes/origin/master
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/refs/heads/master Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +2edd8b993212ee33e85ac9b10b3a413a8f4dced7
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/refs/remotes/origin/HEAD Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +ref: refs/remotes/origin/master
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/.git/refs/remotes/origin/master Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,1 @@ +2edd8b993212ee33e85ac9b10b3a413a8f4dced7
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/AgdaLink.txt Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,15 @@ +http://web.student.chalmers.se/groups/datx02-dtp/report.pdf + + + +https://www.google.co.jp/url?sa=t&rct=j&q=&esrc=s&source=web&cd=1&ved=0ahUKEwj-4dbpzM_YAhXDGpQKHbHXAjcQFggnMAA&url=http%3A%2F%2Fhome.iitk.ac.in%2F~shrutib%2FCS395a%2Freport.pdf&usg=AOvVaw1Qp_3vb2fO-RkdfEGT0Fun + + +https://akaposi.github.io/proplogic.pdf + + + +Book: +Verified Functional Programming in Agda +url: https://www.amazon.co.jp/dp/B01K0MK318/ref=dp-kindle-redirect?_encoding=UTF8&btkr=1 +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/ModelChecking.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,269 @@ +module ModelChecking where + +open import Level renaming (zero to Z ; suc to succ) + +open import Data.Nat hiding (compare) +open import Data.Nat.Properties as NatProp +open import Data.Maybe +-- open import Data.Maybe.Properties +open import Data.Empty +open import Data.List +-- open import Data.Tree.Binary +-- 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 +open import Data.Unit hiding ( _≟_ ; _≤?_ ; _≤_) +open import Relation.Binary.Definitions + +record AtomicNat : Set where + field + ptr : ℕ -- memory pointer ( unique id of DataGear ) + value : ℕ + +init-AtomicNat : AtomicNat +init-AtomicNat = record { ptr = 0 ; value = 0 } + +-- +-- single process implemenation +-- + +f_set : {n : Level } {t : Set n} → AtomicNat → (value : ℕ) → ( AtomicNat → t ) → t +f_set a v next = next record a { value = v } + +record Phil : Set where + field + ptr : ℕ + left right : AtomicNat + +init-Phil : Phil +init-Phil = record { ptr = 0 ; left = init-AtomicNat ; right = init-AtomicNat } + +pickup_rfork : {n : Level} {t : Set n} → Phil → ( Phil → t ) → t +pickup_lfork : {n : Level} {t : Set n} → Phil → ( Phil → t ) → t +eating : {n : Level} {t : Set n} → Phil → ( Phil → t ) → t +putdown_rfork : {n : Level} {t : Set n} → Phil → ( Phil → t ) → t +putdown_lfork : {n : Level} {t : Set n} → Phil → ( Phil → t ) → t +thinking : {n : Level} {t : Set n} → Phil → ( Phil → t ) → t + +pickup_rfork p next = f_set (Phil.right p) (Phil.ptr p) ( λ f → pickup_lfork record p { right = f } next ) +pickup_lfork p next = f_set (Phil.left p) (Phil.ptr p) ( λ f → eating record p { left = f } next ) +eating p next = putdown_rfork p next +putdown_rfork p next = f_set (Phil.right p) 0 ( λ f → putdown_lfork record p { right = f } next ) +putdown_lfork p next = f_set (Phil.left p) 0 ( λ f → thinking record p { left = f } next ) +thinking p next = next p + +run : Phil +run = pickup_rfork record { ptr = 1 ; left = record { ptr = 2 ; value = 0 } ; right = record { ptr = 3 ; value = 0 } } $ λ p → p + +-- +-- Coda Gear +-- + +data Code : Set where + C_nop : Code + C_cas_int : Code + C_putdown_rfork : Code + C_putdown_lfork : Code + C_thinking : Code + C_pickup_rfork : Code + C_pickup_lfork : Code + C_eating : Code + +-- +-- all possible arguments in -APIs +-- +record AtomicNat-API : Set where + field + next : Code + fail : Code + value : ℕ + impl : AtomicNat + +record Phil-API : Set where + field + next : Code + impl : Phil + +data Error : Set where + E_panic : Error + E_cas_fail : Error + +-- +-- Data Gear +-- +-- waiting / pointer / available +-- +data Data : Set where + -- D_AtomicNat : AtomicNat → ℕ → Data + D_AtomicNat : AtomicNat → Data + D_Phil : Phil → Data + D_Error : Error → Data + +-- data API : Set where +-- A_AtomicNat : AtomicNat-API → API +-- A_Phil : Phil-API → API + +open import hoareBinaryTree + +record Context : Set where + field + next : Code + fail : Code + + -- -API list (frame in Gears Agda ) --- a Data of API + -- api : List API + -- c_Phil-API : Maybe Phil-API + c_Phil-API : Phil-API + c_AtomicNat-API : AtomicNat-API + + mbase : ℕ + memory : bt Data + error : Data + fail_next : Code + +-- +-- second level stub +-- +AtomicInt_cas_stub : {n : Level} {t : Set n} → Context → ( Code → Context → t ) → t +AtomicInt_cas_stub {_} {t} c next = updateTree (Context.memory c) (AtomicNat.ptr ai) ( D_AtomicNat (record ai { value = AtomicNat-API.value api } )) + ( λ bt → next ( Context.fail c ) c ) -- segmentation fault ( null pointer ) + $ λ prev bt → f_cas prev bt where + api : AtomicNat-API + api = Context.c_AtomicNat-API c + ai : AtomicNat + ai = AtomicNat-API.impl api + f_cas : Data → bt Data → t + f_cas (D_AtomicNat prev) bt with <-cmp ( AtomicNat.value ai ) ( AtomicNat.value prev ) + ... | tri≈ ¬a b ¬c = next (AtomicNat-API.next api) ( record c { memory = bt ; c_AtomicNat-API = record api { impl = prev } } ) -- update memory + ... | tri< a₁ ¬b ¬c = next (AtomicNat-API.fail api) c --- cleaer API + ... | tri> ¬a ¬b _ = next (AtomicNat-API.fail api) c + f_cas a bt = next ( Context.fail c ) c -- type error / panic + +-- putdown_rfork : Phil → (? → t) → t +-- putdown_rfork p next = goto p->lfork->cas( 0 , putdown_lfork, putdown_lfork ) , next + +Phil_putdown_rfork_sub : {n : Level} {t : Set n} → Context → ( Code → Context → t ) → t +Phil_putdown_rfork_sub c next = next C_cas_int record c { + c_AtomicNat-API = record { impl = Phil.right phil ; value = 0 ; fail = C_putdown_lfork ; next = C_putdown_lfork } } where + phil : Phil + phil = Phil-API.impl ( Context.c_Phil-API c ) + +Phil_putdown_lfork_sub : {n : Level} {t : Set n} → Context → ( Code → Context → t ) → t +Phil_putdown_lfork_sub c next = next C_cas_int record c { + c_AtomicNat-API = record { impl = Phil.left phil ; value = 0 ; fail = C_thinking ; next = C_thinking } } where + phil : Phil + phil = Phil-API.impl ( Context.c_Phil-API c ) + +Phil_thiking : {n : Level} {t : Set n} → Context → ( Code → Context → t ) → t +Phil_thiking c next = next C_pickup_rfork c + +Phil_pickup_lfork_sub : {n : Level} {t : Set n} → Context → ( Code → Context → t ) → t +Phil_pickup_lfork_sub c next = next C_cas_int record c { + c_AtomicNat-API = record { impl = Phil.left phil ; value = Phil.ptr phil ; fail = C_pickup_lfork ; next = C_pickup_rfork } } where + phil : Phil + phil = Phil-API.impl ( Context.c_Phil-API c ) + +Phil_pickup_rfork_sub : {n : Level} {t : Set n} → Context → ( Code → Context → t ) → t +Phil_pickup_rfork_sub c next = next C_cas_int record c { + c_AtomicNat-API = record { impl = Phil.left phil ; value = Phil.ptr phil ; fail = C_pickup_rfork ; next = C_eating } } where + phil : Phil + phil = Phil-API.impl ( Context.c_Phil-API c ) + +Phil_eating_sub : {n : Level} {t : Set n} → Context → ( Code → Context → t ) → t +Phil_eating_sub c next = next C_putdown_rfork c + +code_table : {n : Level} {t : Set n} → Code → Context → ( Code → Context → t) → t +code_table C_nop = λ c next → next C_nop c +code_table C_cas_int = AtomicInt_cas_stub +code_table C_putdown_rfork = Phil_putdown_rfork_sub +code_table C_putdown_lfork = Phil_putdown_lfork_sub +code_table C_thinking = Phil_thiking +code_table C_pickup_rfork = Phil_pickup_lfork_sub +code_table C_pickup_lfork = Phil_pickup_rfork_sub +code_table C_eating = Phil_eating_sub + +step : {n : Level} {t : Set n} → List Context → (List Context → t) → t +step {n} {t} [] next0 = next0 [] +step {n} {t} (p ∷ ps) next0 = code_table (Context.next p) p ( λ code np → next0 (update-context ps np ++ ( record np { next = code } ∷ [] ))) where + update-context : List Context → Context → List Context + update-context [] _ = [] + update-context (c1 ∷ t) np = record c1 { memory = Context.memory np ; mbase = Context.mbase np } ∷ t + +init-context : Context +init-context = record { + next = C_nop + ; fail = C_nop + ; c_Phil-API = record { next = C_nop ; impl = init-Phil } + ; c_AtomicNat-API = record { next = C_nop ; fail = C_nop ; value = 0 ; impl = init-AtomicNat } + ; mbase = 0 + ; memory = leaf + ; error = D_Error E_panic + ; fail_next = C_nop + } + +alloc-data : {n : Level} {t : Set n} → ( c : Context ) → ( Context → ℕ → t ) → t +alloc-data {n} {t} c next = next record c { mbase = suc ( Context.mbase c ) } mnext where + mnext = suc ( Context.mbase c ) + +new-data : {n : Level} {t : Set n} → ( c : Context ) → (ℕ → Data ) → ( Context → ℕ → t ) → t +new-data c x next = alloc-data c $ λ c1 n → insertTree (Context.memory c1) n (x n) ( λ bt → next record c1 { memory = bt } n ) + +init-AtomicNat0 : {n : Level} {t : Set n} → Context → ( Context → ℕ → t ) → t +init-AtomicNat0 c1 next = new-data c1 ( λ ptr → D_AtomicNat (a ptr) ) $ λ c2 ptr → next c2 ptr where + a : ℕ → AtomicNat + a ptr = record { ptr = ptr ; value = 0 } + +init-Phil-0 : (ps : List Context) → (left right : AtomicNat ) → List Context +init-Phil-0 ps left right = new-data (c1 ps) ( λ ptr → D_Phil (p ptr) ) $ λ c ptr → record c { c_Phil-API = record { next = C_thinking ; impl = p ptr }} ∷ ps where + p : ℕ → Phil + p ptr = record init-Phil { ptr = ptr ; left = left ; right = right } + c1 : List Context → Context + c1 [] = init-context + c1 (c2 ∷ ct) = c2 + +init-AtomicNat1 : {n : Level} {t : Set n} → Context → ( Context → AtomicNat → t ) → t +init-AtomicNat1 c1 next = new-data c1 ( λ ptr → D_AtomicNat (a ptr) ) $ λ c2 ptr → next c2 (a ptr) where + a : ℕ → AtomicNat + a ptr = record { ptr = ptr ; value = 0 } + +init-Phil-1 : (c1 : Context) → Context +init-Phil-1 c1 = record c1 { memory = mem2 $ mem1 $ mem ; mbase = n + 3 } where + n : ℕ + n = Context.mbase c1 + left = record { ptr = suc n ; value = 0} + right = record { ptr = suc (suc n) ; value = 0} + mem : bt Data + mem = insertTree ( Context.memory c1) (suc n) (D_AtomicNat left) + $ λ t → t + mem1 : bt Data → bt Data + mem1 t = insertTree t (suc (suc n)) (D_AtomicNat right ) + $ λ t → t + mem2 : bt Data → bt Data + mem2 t = insertTree t (n + 3) (D_Phil record { ptr = n + 3 ; left = left ; right = right }) + $ λ t → t + +test-i0 : bt Data +test-i0 = Context.memory (init-Phil-1 init-context) + +make-phil : ℕ → Context +make-phil zero = init-context +make-phil (suc n) = init-Phil-1 ( make-phil n ) + +test-i5 : bt Data +test-i5 = Context.memory (make-phil 5) + +-- loop exexution + +-- concurrnt execution + +-- state db ( binary tree of processes ) + +-- depth first execution + +-- verify temporal logic poroerries + + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/RedBlackTree.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,242 @@ +module RedBlackTree where + + +open import Level hiding (zero) + +open import Data.Nat hiding (compare) +open import Data.Nat.Properties as NatProp +open import Data.Maybe +-- open import Data.Bool +open import Data.Empty + +open import Relation.Binary +open import Relation.Binary.PropositionalEquality + +open import logic + +open import stack + +record TreeMethods {n m : Level } {a : Set n } {t : Set m } (treeImpl : Set n ) : Set (m Level.⊔ n) where + field + putImpl : treeImpl → a → (treeImpl → t) → t + getImpl : treeImpl → (treeImpl → Maybe a → t) → t +open TreeMethods + +record Tree {n m : Level } {a : Set n } {t : Set m } (treeImpl : Set n ) : Set (m Level.⊔ n) where + field + tree : treeImpl + treeMethods : TreeMethods {n} {m} {a} {t} treeImpl + putTree : a → (Tree treeImpl → t) → t + putTree d next = putImpl (treeMethods ) tree d (\t1 → next (record {tree = t1 ; treeMethods = treeMethods} )) + getTree : (Tree treeImpl → Maybe a → t) → t + getTree next = getImpl (treeMethods ) tree (\t1 d → next (record {tree = t1 ; treeMethods = treeMethods} ) d ) + +open Tree + +data Color {n : Level } : Set n where + Red : Color + Black : Color + + +record Node {n : Level } (a : Set n) : Set n where + inductive + field + key : ℕ + value : a + right : Maybe (Node a ) + left : Maybe (Node a ) + color : Color {n} +open Node + +record RedBlackTree {n m : Level } {t : Set m} (a : Set n) : Set (m Level.⊔ n) where + field + root : Maybe (Node a ) + nodeStack : SingleLinkedStack (Node a ) + +open RedBlackTree + +open SingleLinkedStack + +compTri : ( x y : ℕ ) -> Tri ( x < y ) ( x ≡ y ) ( x > y ) +compTri = IsStrictTotalOrder.compare (Relation.Binary.StrictTotalOrder.isStrictTotalOrder <-strictTotalOrder) + where open import Relation.Binary + +-- put new node at parent node, and rebuild tree to the top +-- +{-# TERMINATING #-} +replaceNode : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a ) → Node a → (RedBlackTree {n} {m} {t} a → t) → t +replaceNode {n} {m} {t} {a} tree s n0 next = popSingleLinkedStack s ( + \s parent → replaceNode1 s parent) + module ReplaceNode where + replaceNode1 : SingleLinkedStack (Node a) → Maybe ( Node a ) → t + replaceNode1 s nothing = next ( record tree { root = just (record n0 { color = Black}) } ) + replaceNode1 s (just n1) with compTri (key n1) (key n0) + replaceNode1 s (just n1) | tri< lt ¬eq ¬gt = replaceNode {n} {m} {t} {a} tree s ( record n1 { value = value n0 ; left = left n0 ; right = right n0 } ) next + replaceNode1 s (just n1) | tri≈ ¬lt eq ¬gt = replaceNode {n} {m} {t} {a} tree s ( record n1 { left = just n0 } ) next + replaceNode1 s (just n1) | tri> ¬lt ¬eq gt = replaceNode {n} {m} {t} {a} tree s ( record n1 { right = just n0 } ) next + + +rotateRight : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → + (RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a ) → Maybe (Node a) → Maybe (Node a) → t) → t +rotateRight {n} {m} {t} {a} tree s n0 parent rotateNext = getSingleLinkedStack s (\ s n0 → rotateRight1 {n} {m} {t} {a} tree s n0 parent rotateNext) + where + rotateRight1 : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → + (RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → t) → t + rotateRight1 {n} {m} {t} {a} tree s n0 parent rotateNext with n0 + ... | nothing = rotateNext tree s nothing n0 + ... | just n1 with parent + ... | nothing = rotateNext tree s (just n1 ) n0 + ... | just parent1 with left parent1 + ... | nothing = rotateNext tree s (just n1) nothing + ... | just leftParent with compTri (key n1) (key leftParent) + rotateRight1 {n} {m} {t} {a} tree s n0 parent rotateNext | just n1 | just parent1 | just leftParent | tri< a₁ ¬b ¬c = rotateNext tree s (just n1) parent + rotateRight1 {n} {m} {t} {a} tree s n0 parent rotateNext | just n1 | just parent1 | just leftParent | tri≈ ¬a b ¬c = rotateNext tree s (just n1) parent + rotateRight1 {n} {m} {t} {a} tree s n0 parent rotateNext | just n1 | just parent1 | just leftParent | tri> ¬a ¬b c = rotateNext tree s (just n1) parent + + +rotateLeft : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → + (RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → t) → t +rotateLeft {n} {m} {t} {a} tree s n0 parent rotateNext = getSingleLinkedStack s (\ s n0 → rotateLeft1 tree s n0 parent rotateNext) + where + rotateLeft1 : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → + (RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → t) → t + rotateLeft1 {n} {m} {t} {a} tree s n0 parent rotateNext with n0 + ... | nothing = rotateNext tree s nothing n0 + ... | just n1 with parent + ... | nothing = rotateNext tree s (just n1) nothing + ... | just parent1 with right parent1 + ... | nothing = rotateNext tree s (just n1) nothing + ... | just rightParent with compTri (key n1) (key rightParent) + rotateLeft1 {n} {m} {t} {a} tree s n0 parent rotateNext | just n1 | just parent1 | just rightParent | tri< a₁ ¬b ¬c = rotateNext tree s (just n1) parent + rotateLeft1 {n} {m} {t} {a} tree s n0 parent rotateNext | just n1 | just parent1 | just rightParent | tri≈ ¬a b ¬c = rotateNext tree s (just n1) parent + rotateLeft1 {n} {m} {t} {a} tree s n0 parent rotateNext | just n1 | just parent1 | just rightParent | tri> ¬a ¬b c = rotateNext tree s (just n1) parent + +{-# TERMINATING #-} +insertCase5 : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Node a → Node a → (RedBlackTree {n} {m} {t} a → t) → t +insertCase5 {n} {m} {t} {a} tree s n0 parent grandParent next = pop2SingleLinkedStack s (\ s parent grandParent → insertCase51 tree s n0 parent grandParent next) + where + insertCase51 : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → Maybe (Node a) → (RedBlackTree {n} {m} {t} a → t) → t + insertCase51 {n} {m} {t} {a} tree s n0 parent grandParent next with n0 + ... | nothing = next tree + ... | just n1 with parent | grandParent + ... | nothing | _ = next tree + ... | _ | nothing = next tree + ... | just parent1 | just grandParent1 with left parent1 | left grandParent1 + ... | nothing | _ = next tree + ... | _ | nothing = next tree + ... | just leftParent1 | just leftGrandParent1 + with compTri (key n1) (key leftParent1) | compTri (key leftParent1) (key leftGrandParent1) + ... | tri≈ ¬a b ¬c | tri≈ ¬a1 b1 ¬c1 = rotateRight tree s n0 parent (\ tree s n0 parent → insertCase5 tree s n0 parent1 grandParent1 next) + ... | _ | _ = rotateLeft tree s n0 parent (\ tree s n0 parent → insertCase5 tree s n0 parent1 grandParent1 next) + +insertCase4 : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Node a → Node a → Node a → (RedBlackTree {n} {m} {t} a → t) → t +insertCase4 {n} {m} {t} {a} tree s n0 parent grandParent next + with (right parent) | (left grandParent) +... | nothing | _ = insertCase5 tree s (just n0) parent grandParent next +... | _ | nothing = insertCase5 tree s (just n0) parent grandParent next +... | just rightParent | just leftGrandParent with compTri (key n0) (key rightParent) | compTri (key parent) (key leftGrandParent) -- (key n0) (key rightParent) | (key parent) (key leftGrandParent) +... | tri≈ ¬a b ¬c | tri≈ ¬a1 b1 ¬c1 = popSingleLinkedStack s (\ s n1 → rotateLeft tree s (left n0) (just grandParent) (\ tree s n0 parent → insertCase5 tree s n0 rightParent grandParent next)) +... | _ | _ = insertCase41 tree s n0 parent grandParent next + where + insertCase41 : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Node a → Node a → Node a → (RedBlackTree {n} {m} {t} a → t) → t + insertCase41 {n} {m} {t} {a} tree s n0 parent grandParent next + with (left parent) | (right grandParent) + ... | nothing | _ = insertCase5 tree s (just n0) parent grandParent next + ... | _ | nothing = insertCase5 tree s (just n0) parent grandParent next + ... | just leftParent | just rightGrandParent with compTri (key n0) (key leftParent) | compTri (key parent) (key rightGrandParent) + ... | tri≈ ¬a b ¬c | tri≈ ¬a1 b1 ¬c1 = popSingleLinkedStack s (\ s n1 → rotateRight tree s (right n0) (just grandParent) (\ tree s n0 parent → insertCase5 tree s n0 leftParent grandParent next)) + ... | _ | _ = insertCase5 tree s (just n0) parent grandParent next + +colorNode : {n : Level } {a : Set n} → Node a → Color → Node a +colorNode old c = record old { color = c } + +{-# TERMINATING #-} +insertNode : {n m : Level } {t : Set m } {a : Set n} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Node a → (RedBlackTree {n} {m} {t} a → t) → t +insertNode {n} {m} {t} {a} tree s n0 next = get2SingleLinkedStack s (insertCase1 n0) + where + insertCase1 : Node a → SingleLinkedStack (Node a) → Maybe (Node a) → Maybe (Node a) → t -- placed here to allow mutual recursion + insertCase3 : SingleLinkedStack (Node a) → Node a → Node a → Node a → t + insertCase3 s n0 parent grandParent with left grandParent | right grandParent + ... | nothing | nothing = insertCase4 tree s n0 parent grandParent next + ... | nothing | just uncle = insertCase4 tree s n0 parent grandParent next + ... | just uncle | _ with compTri (key uncle ) (key parent ) + insertCase3 s n0 parent grandParent | just uncle | _ | tri≈ ¬a b ¬c = insertCase4 tree s n0 parent grandParent next + insertCase3 s n0 parent grandParent | just uncle | _ | tri< a ¬b ¬c with color uncle + insertCase3 s n0 parent grandParent | just uncle | _ | tri< a ¬b ¬c | Red = pop2SingleLinkedStack s ( \s p0 p1 → insertCase1 ( + record grandParent { color = Red ; left = just ( record parent { color = Black } ) ; right = just ( record uncle { color = Black } ) }) s p0 p1 ) + insertCase3 s n0 parent grandParent | just uncle | _ | tri< a ¬b ¬c | Black = insertCase4 tree s n0 parent grandParent next + insertCase3 s n0 parent grandParent | just uncle | _ | tri> ¬a ¬b c with color uncle + insertCase3 s n0 parent grandParent | just uncle | _ | tri> ¬a ¬b c | Red = pop2SingleLinkedStack s ( \s p0 p1 → insertCase1 ( record grandParent { color = Red ; left = just ( record parent { color = Black } ) ; right = just ( record uncle { color = Black } ) }) s p0 p1 ) + insertCase3 s n0 parent grandParent | just uncle | _ | tri> ¬a ¬b c | Black = insertCase4 tree s n0 parent grandParent next + insertCase2 : SingleLinkedStack (Node a) → Node a → Node a → Node a → t + insertCase2 s n0 parent grandParent with color parent + ... | Black = replaceNode tree s n0 next + ... | Red = insertCase3 s n0 parent grandParent + insertCase1 n0 s nothing nothing = next tree + insertCase1 n0 s nothing (just grandParent) = next tree + insertCase1 n0 s (just parent) nothing = replaceNode tree s (colorNode n0 Black) next + insertCase1 n0 s (just parent) (just grandParent) = insertCase2 s n0 parent grandParent + +---- +-- find node potition to insert or to delete, the path will be in the stack +-- +findNode : {n m : Level } {a : Set n} {t : Set m} → RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → (Node a) → (Node a) → (RedBlackTree {n} {m} {t} a → SingleLinkedStack (Node a) → Node a → t) → t +findNode {n} {m} {a} {t} tree s n0 n1 next = pushSingleLinkedStack s n1 (\ s → findNode1 s n1) + module FindNode where + findNode2 : SingleLinkedStack (Node a) → (Maybe (Node a)) → t + findNode2 s nothing = next tree s n0 + findNode2 s (just n) = findNode tree s n0 n next + findNode1 : SingleLinkedStack (Node a) → (Node a) → t + findNode1 s n1 with (compTri (key n0) (key n1)) + findNode1 s n1 | tri< a ¬b ¬c = popSingleLinkedStack s ( \s _ → next tree s (record n1 {key = key n1 ; value = value n0 } ) ) + findNode1 s n1 | tri≈ ¬a b ¬c = findNode2 s (right n1) + findNode1 s n1 | tri> ¬a ¬b c = findNode2 s (left n1) + -- ... | EQ = popSingleLinkedStack s ( \s _ → next tree s (record n1 {ey =ey n1 ; value = value n0 } ) ) + -- ... | GT = findNode2 s (right n1) + -- ... | LT = findNode2 s (left n1) + + + + +leafNode : {n : Level } { a : Set n } → a → ℕ → (Node a) +leafNode v k1 = record {key = k1 ; value = v ; right = nothing ; left = nothing ; color = Red } + +putRedBlackTree : {n m : Level} {t : Set m} {a : Set n} → RedBlackTree {n} {m} {t} a → a → (key1 : ℕ) → (RedBlackTree {n} {m} {t} a → t) → t +putRedBlackTree {n} {m} {t} {a} tree val1 key1 next with (root tree) +putRedBlackTree {n} {m} {t} {a} tree val1 key1 next | nothing = next (record tree {root = just (leafNode val1 key1 ) }) +putRedBlackTree {n} {m} {t} {a} tree val1 key1 next | just n2 = clearSingleLinkedStack (nodeStack tree) (λ s → findNode tree s (leafNode val1 key1) n2 (λ tree1 s n1 → insertNode tree1 s n1 next)) + + +-- getRedBlackTree : {n m : Level } {t : Set m} {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} {A} a → → (RedBlackTree {n} {m} {t} {A} a → (Maybe (Node a)) → t) → t +-- getRedBlackTree {_} {_} {t} {a} {k} tree1 cs = checkNode (root tree) +-- module GetRedBlackTree where -- http://agda.readthedocs.io/en/v2.5.2/language/let-and-where.html +-- search : Node a → t +-- checkNode : Maybe (Node a) → t +-- checkNode nothing = cs tree nothing +-- checkNode (just n) = search n +-- search n with compTri1 (key n) +-- search n | tri< a ¬b ¬c = checkNode (left n) +-- search n | tri≈ ¬a b ¬c = cs tree (just n) +-- search n | tri> ¬a ¬b c = checkNode (right n) + + + +-- putUnblanceTree : {n m : Level } {a : Set n} {k : ℕ} {t : Set m} → RedBlackTree {n} {m} {t} {A} a → → a → (RedBlackTree {n} {m} {t} {A} a → t) → t +-- putUnblanceTree {n} {m} {A} {a} {k} {t} tree1 value next with (root tree) +-- ... | nothing = next (record tree {root = just (leafNode1 value) }) +-- ... | just n2 = clearSingleLinkedStack (nodeStack tree) (λ s → findNode tree s (leafNode1 value) n2 (λ tree1 s n1 → replaceNode tree1 s n1 next)) + +createEmptyRedBlackTreeℕ : {n m : Level} {t : Set m} (a : Set n) + → RedBlackTree {n} {m} {t} a +createEmptyRedBlackTreeℕ a = record { + root = nothing + ; nodeStack = emptySingleLinkedStack + } + + +test : {m : Level} (t : Set) → RedBlackTree {Level.zero} {Level.zero} ℕ +test t = createEmptyRedBlackTreeℕ {Level.zero} {Level.zero} {t} ℕ + +-- ts = createEmptyRedBlackTreeℕ {ℕ} {?} {!!} 0 + +-- tes = putRedBlackTree {_} {_} {_} (createEmptyRedBlackTreeℕ {_} {_} {_} 3 3) 2 2 (λ t → t)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Todo.txt Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,91 @@ +Wed May 4 22:07:32 JST 2022 + + Context memory に DataGear を Binary Tree で入れる。List で良いのだが + Libaryの Data.Tree.Binary には insert はない ( まぁ、alloc / read / update くらいで map で書けるとか?) + 共有データは、memory に入れる + free はしない(?) + +Sun May 1 15:04:57 JST 2022 + + Model checking + + goto 先の番号を stub に書くのは変 + + process : context が別々 + thread : context 一緒、goto 先が異なる + context は共有 実行は codeGear は atomic + shared file descriptor など + + single phils direct connection single thread + no shared data + + multi process phils separate process + shared context + + multi threaded phils shared process + serate next + atomic codeGear execution + + +Thu May 17 15:26:56 JST 2018 + + findNode -> replaceNode -> getRedBlackTree だが + + findNode -> P0 -> replaceNode -> P1 -> getRedBlackTree + + という形で証明しても良い。一挙に証明するのは,可能だろうけど、良くないはず。 + +Sun May 6 17:54:50 JST 2018 + + do1 a $ \b -> do2 b next を、do1 と do2 に分離することはできる? + + +Mon Apr 30 17:15:16 JST 2018 + + Stack の初期化を別にするだけだと、置き換えの条件に到達した時に、Stack が empty になるのを保証できない + やはり、 Stack + Current Tree = Original Tree という不変式を入れないとだめらしい + +Mon Mar 26 17:43:06 JST 2018 + + Decidable を使って Compare の場合分けを行う + Decidable を使うと Eq から x ≡ y の証明を取り出すことができる + 場合分けには Trichotomous を使う + compareTri を完成させる Done + +Fri Jan 5 16:43:26 JST 2018 + + unbalanced binary search tree の動作を調べる + + RedBlackTree の put を完成させる + + RedBlackTree の Deletion を完成させる + + unbalanced binary search tree と同様の動作をする + + 木の深さの最小と最大の差が2倍を超えない + + CodeGear/DataGear が構成する圏を定義する + + goto を定義して meta 計算を可能にする + + DataSegment をすべて含む sum 型を定義しmetaDataSegmentとする + + 実行環境をcontextとして定義しgotoと合わせて並列実行をモデル化する + + Monad の合成に必要な規則を上の圏上に定義する + + synchronizedQueue の仕様をCTLを使って定義する + + Gearsで記述したsynchornizedQueueを検証する + + gotoを用いてモデル検査と証明の組み合わせを実現する + + +Wed Aug 27 17:52:00 JST 2019 + + 別で定義した TriCotomos や \=? などの Relation の関数を + Agdaで定義してあるものに置き換える,まとめる + + HoareLogic をベースにした SingleLinkedStack の作成 + + HoareLogic ベースの Tree の証明
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/btree.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,547 @@ +module btree 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 ( _≟_ ; _≤?_ ; _≤_) + +data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where + t-leaf : treeInvariant leaf + t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) + t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key₁ value₁ t₁ t₂) + → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) + t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) + t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → key < key₁ → key₁ < key₂ + → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₂ value₂ t₃ t₄) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) + +-- +-- 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 + +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) +treeTest3 : bt ℕ +treeTest3 = node 2 5 (node 1 7 leaf leaf ) leaf +treeTest4 : bt ℕ +treeTest4 = node 5 5 leaf leaf +treeTest5 : bt ℕ +treeTest5 = node 1 7 leaf leaf + + +treeInvariantTest1 : treeInvariant treeTest1 +treeInvariantTest1 = t-right (m≤m+n _ 2) (t-node (add< 0) (add< 1) (t-left (add< 0) (t-single 1 7)) (t-single 5 5) ) + +treeInvariantTest2 : treeInvariant treeTest2 +treeInvariantTest2 = t-node (add< 0) (add< 1) (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) + +stackInvariantTest111 : stackInvariant 4 treeTest4 treeTest1 ( treeTest4 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +stackInvariantTest111 = s-right (add< 0) (s-right (add< 3) (s-nil)) + +stackInvariantTest11 : stackInvariant 4 treeTest4 treeTest1 ( treeTest4 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +stackInvariantTest11 = s-right (add< 0) (s-right (add< 3) (s-nil)) --treeTest4から見てみぎ、みぎnil + + +stackInvariantTest2' : stackInvariant 2 treeTest3 treeTest1 (treeTest3 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +stackInvariantTest2' = s-left (add< 0) (s-right (add< 1) (s-nil)) + +--stackInvariantTest121 : stackInvariant 2 treeTest5 treeTest1 ( treeTest5 ∷ treeTest3 ∷ treeTest2 ∷ treeTest1 ∷ [] ) +--stackInvariantTest121 = s-left (_) (s-left (add< 0) (s-right (add< 1) (s-nil))) -- 2<2が示せない + +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-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 + +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 + +nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ +nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x +nat-<> : { x y : ℕ } → x < y → y < x → ⊥ +nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x + +open _∧_ + + +depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) +depth-1< {i} {j} = s≤s (m≤m⊔n _ j) + +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) --leafである +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) --探しているkeyと一致 +findP {n} {_} {A} key (node key₁ v1 tree tree₁) tree0 st Pre next _ | tri< a ¬b ¬c = next tree (tree ∷ st) --keyが求めているkey1より小さい。今いるツリーの左側をstackに積む。 +-- ⟪ treeLeftDown tree tree₁ (proj1 Pre) , s-left a (proj2 Pre)⟫ depth-1< --これでも通った。 + ⟪ 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 t) = t-right x t +replaceTree1 k v1 value (t-left x t) = t-left x t +replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ + +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 + +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 + +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 (replacePR)を定める必要があるが、siの値のよって影響されるため、場合分けをする。 + --siとriが変化するから、(nextとすると)場合分けで定義し直す必要がある。 + 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₂ {!!} {!!} {!!} {!!} + 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 +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 (≤-step lt1) p1 lt1 ) where + TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → Invraiant r1 → reduce r1 < reduce r → t + TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 p1 (λ r2 p1 lt1 → ⊥-elim (lemma5 n≤j lt1)) + TerminatingLoop1 (suc j) r r1 n≤j p1 lt with <-cmp (reduce r1) (suc j) + ... | tri< a ¬b ¬c = TerminatingLoop1 j r r1 a p1 lt + ... | tri≈ ¬a b ¬c = loop r1 p1 (λ r2 p2 lt1 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b lt1 ) p2 lt1 ) + ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) +{- +open _∧_ + +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 ti) r-node = t-right x ti +RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI0 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ 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₁ ti) (r-right x ri) = t-single key₁ value₁ +RTtoTI0 (node key₁ _ leaf right@(node key₂ _ _ _)) (node key₁ value₁ leaf right₁@(node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = + t-right (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) +RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left x₁ 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₁ ti) (r-right x r-leaf) = + t-node x₁ x ti (t-single key value) +RTtoTI0 (node key₁ _ (node _ _ _ _) (node key₂ _ _ _)) (node key₁ _ (node _ _ _ _) (node key₃ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = + t-node x₁ (subst (λ k → key₁ < k) (rt-property-key ri) x₂) ti (RTtoTI0 _ _ key value ti₁ ri) +-- 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 (t-single _ _ ) +RTtoTI0 .(node _ _ leaf (node _ _ _ _)) (node key₁ value₁ (node key value leaf leaf) (node _ _ _ _)) key value (t-right x₁ ti) (r-left x r-leaf) = t-node x x₁ (t-single key value) ti +RTtoTI0 (node key₃ _ (node key₂ _ _ _) leaf) (node key₃ _ (node key₁ value₁ left left₁) leaf) key value (t-left x₁ ti) (r-left x ri) = + t-left (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) -- key₁ < key₃ +RTtoTI0 (node key₁ _ (node key₂ _ _ _) (node _ _ _ _)) (node key₁ _ (node key₃ _ _ _) (node _ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = t-node (subst (λ k → k < key₁ ) (rt-property-key ri) x₁) x₂ (RTtoTI0 _ _ key value ti ri) ti₁ + +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 ti) r-node = t-right x ti +RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI1 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ ti ti₁ +-- r-right case +RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ leaf (node _ _ _ _)) key value (t-right x₁ 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₁ 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₂ 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₂ ti ti₁) (r-right x ri) = t-node x₁ (subst (λ k → key₄ < k ) (sym (rt-property-key ri)) x₂) 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₁ 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₁ 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₂ ti ti₁) (r-left x r-leaf) = t-right x₂ ti₁ +RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = + t-node (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ (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 +-}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/etc/trichotomos-ex.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,205 @@ +module trichotomos-ex where + +open import Level hiding (zero) renaming ( suc to succ ) +open import Data.Empty +open import Data.Nat +open import Relation.Nullary +open import Algebra +open import Data.Nat.Properties as NatProp +open import Relation.Binary +open import Relation.Binary.PropositionalEquality +open import Relation.Binary.Core +open import Function + +data Bool {n : Level } : Set n where + True : Bool + False : Bool + +record _∧_ {n : Level } (a : Set n) (b : Set n): Set n where + field + pi1 : a + pi2 : b + +data _∨_ {n : Level } (a : Set n) (b : Set n): Set n where + pi1 : a -> a ∨ b + pi2 : b -> a ∨ b + + +data Maybe {n : Level } (a : Set n) : Set n where + Nothing : Maybe a + Just : a -> Maybe a + +data Color {n : Level } : Set n where + Red : Color + Black : Color + +data CompareResult {n : Level } : Set n where + LT : CompareResult + GT : CompareResult + EQ : CompareResult + +record Node {n : Level } (a k : Set n) : Set n where + inductive + field + key : k + value : a + right : Maybe (Node a k) + left : Maybe (Node a k) + color : Color {n} +open Node + +compareℕ : ℕ → ℕ → CompareResult {Level.zero} +compareℕ x y with Data.Nat.compare x y +... | less _ _ = LT +... | equal _ = EQ +... | greater _ _ = GT + +compareT : ℕ → ℕ → CompareResult {Level.zero} +compareT x y with IsStrictTotalOrder.compare (Relation.Binary.StrictTotalOrder.isStrictTotalOrder strictTotalOrder) x y +... | tri≈ _ _ _ = EQ +... | tri< _ _ _ = LT +... | tri> _ _ _ = GT + + +compare2 : (x y : ℕ ) → CompareResult {Level.zero} +compare2 zero zero = EQ +compare2 (suc _) zero = GT +compare2 zero (suc _) = LT +compare2 (suc x) (suc y) = compare2 x y + +contraposition : {m : Level } {A B : Set m} → (B → A) → (¬ A → ¬ B) +contraposition f = λ x y → x (f y) + +compareTri1 : (n : ℕ) → zero <′ suc n +compareTri1 zero = ≤′-refl +compareTri1 (suc n) = ≤′-step ( compareTri1 n ) + +compareTri2 : (n m : ℕ) → n ≤′ m → suc n ≤′ suc m +compareTri2 _ _ ≤′-refl = ≤′-refl +compareTri2 n (suc m) ( ≤′-step p ) = ≤′-step { suc n } ( compareTri2 n m p ) + +<′dec : Set +<′dec = ∀ m n → Dec ( m ≤′ n ) + +compareTri6 : ∀ m {n} → ¬ m ≤′ n → ¬ suc m ≤′ suc n + +is≤′ : <′dec +is≤′ zero zero = yes ≤′-refl +is≤′ zero (suc n) = yes (lemma n) + where + lemma : (n : ℕ) → zero ≤′ suc n + lemma zero = ≤′-step ≤′-refl + lemma (suc n) = ≤′-step (lemma n) +is≤′ (suc m) (zero) = no ( λ () ) +is≤′ (suc m) (suc n) with is≤′ m n +... | yes p = yes ( compareTri2 _ _ p ) +... | no p = no ( compareTri6 _ p ) + +compareTri20 : {n : ℕ} → ¬ suc n ≤′ zero +compareTri20 () + + +compareTri21 : (n m : ℕ) → suc n ≤′ suc m → n ≤′ m +compareTri21 _ _ ≤′-refl = ≤′-refl +compareTri21 (suc n) zero ( ≤′-step p ) = compareTri23 (suc n) ( ≤′-step p ) p + where + compareTri23 : (n : ℕ) → suc n ≤′ suc zero → suc n ≤′ zero → n ≤′ zero + compareTri23 zero ( ≤′-step p ) _ = ≤′-refl + compareTri23 zero ≤′-refl _ = ≤′-refl + compareTri23 (suc n1) ( ≤′-step p ) () +compareTri21 n (suc m) ( ≤′-step p ) = ≤′-step (compareTri21 _ _ p) +compareTri21 zero zero ( ≤′-step p ) = ≤′-refl + + +compareTri3 : ∀ m {n} → ¬ m ≡ suc (m + n) +compareTri3 zero () +compareTri3 (suc m) eq = compareTri3 m (cong pred eq) + +compareTri4' : ∀ m {n} → ¬ n ≡ m → ¬ suc n ≡ suc m +compareTri4' m {n} with n ≟ m +... | yes refl = λ x y → x refl +... | no p = λ x y → p ( cong pred y ) + +compareTri4 : ∀ m {n} → ¬ n ≡ m → ¬ suc n ≡ suc m +compareTri4 m = contraposition ( λ x → cong pred x ) + +-- compareTri6 : ∀ m {n} → ¬ m ≤′ n → ¬ suc m ≤′ suc n +compareTri6 m {n} = λ x y → x (compareTri21 _ _ y) + +compareTri5 : ∀ m {n} → ¬ m <′ n → ¬ suc m <′ suc n +compareTri5 m {n} = compareTri6 (suc m) + +compareTri : Trichotomous _≡_ _<′_ +compareTri zero zero = tri≈ ( λ ()) refl ( λ ()) +compareTri zero (suc n) = tri< ( compareTri1 n ) ( λ ()) ( λ ()) +compareTri (suc n) zero = tri> ( λ ()) ( λ ()) ( compareTri1 n ) +compareTri (suc n) (suc m) with compareTri n m +... | tri< p q r = tri< (compareTri2 (suc n) m p ) (compareTri4 _ q) ( compareTri5 _ r ) +... | tri≈ p refl r = tri≈ (compareTri5 _ p) refl ( compareTri5 _ r ) +... | tri> p q r = tri> ( compareTri5 _ p ) (compareTri4 _ q) (compareTri2 (suc m) n r ) + +compareDecTest : (n n1 : Node ℕ ℕ) → ( key n ≡ key n1 ) ∨ ( ¬ key n ≡ key n1 ) +compareDecTest n n1 with (key n) ≟ (key n1) +... | yes p = pi1 p +... | no ¬p = pi2 ¬p + + +putTest1Lemma2 : (k : ℕ) → compare2 k k ≡ EQ +putTest1Lemma2 zero = refl +putTest1Lemma2 (suc k) = putTest1Lemma2 k + +putTest1Lemma1 : (x y : ℕ) → compareℕ x y ≡ compare2 x y +putTest1Lemma1 zero zero = refl +putTest1Lemma1 (suc m) zero = refl +putTest1Lemma1 zero (suc n) = refl +putTest1Lemma1 (suc m) (suc n) with Data.Nat.compare m n +putTest1Lemma1 (suc .m) (suc .(Data.Nat.suc m + k)) | less m k = lemma1 m + where + lemma1 : (m : ℕ) → LT ≡ compare2 m (ℕ.suc (m + k)) + lemma1 zero = refl + lemma1 (suc y) = lemma1 y +putTest1Lemma1 (suc .m) (suc .m) | equal m = lemma1 m + where + lemma1 : (m : ℕ) → EQ ≡ compare2 m m + lemma1 zero = refl + lemma1 (suc y) = lemma1 y +putTest1Lemma1 (suc .(Data.Nat.suc m + k)) (suc .m) | greater m k = lemma1 m + where + lemma1 : (m : ℕ) → GT ≡ compare2 (ℕ.suc (m + k)) m + lemma1 zero = refl + lemma1 (suc y) = lemma1 y + +putTest1Lemma3 : (k : ℕ) → compareℕ k k ≡ EQ +putTest1Lemma3 k = trans (putTest1Lemma1 k k) ( putTest1Lemma2 k ) + +compareLemma1 : {x y : ℕ} → compare2 x y ≡ EQ → x ≡ y +compareLemma1 {zero} {zero} refl = refl +compareLemma1 {zero} {suc _} () +compareLemma1 {suc _} {zero} () +compareLemma1 {suc x} {suc y} eq = cong ( λ z → ℕ.suc z ) ( compareLemma1 ( trans lemma2 eq ) ) + where + lemma2 : compare2 (ℕ.suc x) (ℕ.suc y) ≡ compare2 x y + lemma2 = refl + + +compTri : ( x y : ℕ ) -> Tri (x < y) ( x ≡ y ) ( x > y ) +compTri = IsStrictTotalOrder.compare (Relation.Binary.StrictTotalOrder.isStrictTotalOrder strictTotalOrder) + where open import Relation.Binary + +checkT : {m : Level } (n : Maybe (Node ℕ ℕ)) → ℕ → Bool {m} +checkT Nothing _ = False +checkT (Just n) x with compTri (value n) x +... | tri≈ _ _ _ = True +... | _ = False + +checkEQ : {m : Level } ( x : ℕ ) -> ( n : Node ℕ ℕ ) -> (value n ) ≡ x -> checkT {m} (Just n) x ≡ True +checkEQ x n refl with compTri (value n) x +... | tri≈ _ refl _ = refl +... | tri> _ neq gt = ⊥-elim (neq refl) +... | tri< lt neq _ = ⊥-elim (neq refl) + +checkEQ' : {m : Level } ( x y : ℕ ) -> ( eq : x ≡ y ) -> ( x ≟ y ) ≡ yes eq +checkEQ' x y refl with x ≟ y +... | yes refl = refl +... | no neq = ⊥-elim ( neq refl ) +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/fig/tree-invariant.svg Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,371 @@ +<?xml version="1.0" encoding="UTF-8" standalone="no"?> +<!DOCTYPE svg PUBLIC "-//W3C//DTD SVG 1.1//EN" "http://www.w3.org/Graphics/SVG/1.1/DTD/svg11.dtd"> +<svg xmlns="http://www.w3.org/2000/svg" xmlns:xl="http://www.w3.org/1999/xlink" xmlns:dc="http://purl.org/dc/elements/1.1/" version="1.1" viewBox="-244 -179 810 1558" width="810" height="1558"> + <defs> + <font-face font-family="Helvetica Neue" font-size="16" panose-1="2 0 5 3 0 0 0 2 0 4" units-per-em="1000" underline-position="-100" underline-thickness="50" slope="0" x-height="517" cap-height="714" ascent="951.9958" descent="-212.99744" font-weight="400"> + <font-face-src> + <font-face-name name="HelveticaNeue"/> + </font-face-src> + </font-face> + <marker orient="auto" overflow="visible" markerUnits="strokeWidth" id="FilledArrow_Marker" stroke-linejoin="miter" stroke-miterlimit="10" viewBox="-1 -3 7 6" markerWidth="7" markerHeight="6" color="black"> + <g> + <path d="M 4.8 0 L 0 -1.8 L 0 1.8 Z" fill="currentColor" stroke="currentColor" stroke-width="1"/> + </g> + </marker> + <font-face font-family="Helvetica Neue" font-size="14" panose-1="2 0 5 3 0 0 0 2 0 4" units-per-em="1000" underline-position="-100" underline-thickness="50" slope="0" x-height="517" cap-height="714" ascent="951.9958" descent="-212.99744" font-weight="400"> + <font-face-src> + <font-face-name name="HelveticaNeue"/> + </font-face-src> + </font-face> + </defs> + <metadata> Produced by OmniGraffle 7.18.5\n2021-11-03 06:50:44 +0000</metadata> + <g id="Canvas_1" stroke-opacity="1" stroke-dasharray="none" stroke="none" fill="none" fill-opacity="1"> + <title>Canvas 1</title> + <rect fill="white" x="-244" y="-179" width="810" height="1558"/> + <g id="Canvas_1_Layer_1"> + <title>Layer 1</title> + <g id="Graphic_123"> + <rect x="-238" y="414" width="803" height="254" fill="white"/> + <rect x="-238" y="414" width="803" height="254" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="1"/> + </g> + <g id="Graphic_119"> + <rect x="231" y="954.25" width="323" height="238.75" fill="white"/> + <rect x="231" y="954.25" width="323" height="238.75" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="1"/> + </g> + <g id="Graphic_118"> + <rect x="-43" y="675" width="603" height="254" fill="white"/> + <rect x="-43" y="675" width="603" height="254" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="1"/> + </g> + <g id="Graphic_3"> + <circle cx="401.5" cy="-48.5" r="15.5000247674794" fill="black"/> + <circle cx="401.5" cy="-48.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + <text transform="translate(394.1 -57.724)" fill="white"> + <tspan font-family="Helvetica Neue" font-size="16" font-weight="400" fill="white" x="2.952" y="15">5</tspan> + </text> + </g> + <g id="Graphic_5"> + <circle cx="355.5" cy="33.5" r="15.5000247674794" fill="black"/> + <circle cx="355.5" cy="33.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + <text transform="translate(348.1 24.276)" fill="black"> + <tspan font-family="Helvetica Neue" font-size="16" font-weight="400" fill="black" x="2.952" y="15">3</tspan> + </text> + </g> + <g id="Graphic_6"> + <circle cx="309.5" cy="115.5" r="15.5000247674794" fill="black"/> + <circle cx="309.5" cy="115.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_7"> + <circle cx="458.5" cy="33.5" r="15.5000247674794" fill="black"/> + <circle cx="458.5" cy="33.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_8"> + <circle cx="515.5" cy="115.5" r="15.5000247674794" fill="black"/> + <circle cx="515.5" cy="115.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_9"> + <circle cx="406.5" cy="115.5" r="15.5000247674794" fill="black"/> + <circle cx="406.5" cy="115.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_10"> + <line x1="393.4253" y1="-34.10597" x2="369.88604" y2="7.855327" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_11"> + <line x1="347.4253" y1="47.89403" x2="323.88604" y2="89.85533" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_12"> + <line x1="364.2162" y1="47.51426" x2="390.97086" y2="90.53158" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_13"> + <line x1="410.9193" y1="-34.949395" x2="441.71773" y2="9.357086" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_14"> + <line x1="467.9193" y1="47.050605" x2="498.71773" y2="91.35709" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_27"> + <circle cx="396.5" cy="212.5" r="15.5000247674794" fill="black"/> + <circle cx="396.5" cy="212.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_26"> + <circle cx="350.5" cy="294.5" r="15.5000247674794" fill="black"/> + <circle cx="350.5" cy="294.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_25"> + <circle cx="304.5" cy="376.5" r="15.5000247674794" fill="black"/> + <circle cx="304.5" cy="376.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_24"> + <circle cx="453.5" cy="294.5" r="15.5000247674794" fill="black"/> + <circle cx="453.5" cy="294.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_23"> + <circle cx="510.5" cy="376.5" r="15.5000247674794" fill="black"/> + <circle cx="510.5" cy="376.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_22"> + <circle cx="401.5" cy="376.5" r="15.5000247674794" fill="black"/> + <circle cx="401.5" cy="376.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_21"> + <line x1="388.4253" y1="226.89403" x2="364.88604" y2="268.85533" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_20"> + <line x1="342.4253" y1="308.89403" x2="318.88604" y2="350.85533" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_19"> + <line x1="359.2162" y1="308.51426" x2="385.97086" y2="351.5316" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_18"> + <line x1="405.9193" y1="226.0506" x2="436.71773" y2="270.3571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_17"> + <line x1="462.9193" y1="308.0506" x2="493.71773" y2="352.3571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_32"> + <circle cx="45" cy="301" r="15.5000247674794" fill="black"/> + <circle cx="45" cy="301" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_31"> + <circle cx="-1" cy="383" r="15.5000247674794" fill="black"/> + <circle cx="-1" cy="383" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_30"> + <circle cx="96" cy="383" r="15.5000247674794" fill="black"/> + <circle cx="96" cy="383" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_29"> + <line x1="36.9253" y1="315.39403" x2="13.386036" y2="357.35533" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_28"> + <line x1="53.716185" y1="315.01426" x2="80.47086" y2="358.0316" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_48"> + <circle cx="396.5" cy="448.5" r="15.5000247674794" fill="black"/> + <circle cx="396.5" cy="448.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_47"> + <circle cx="350.5" cy="530.5" r="15.5000247674794" fill="black"/> + <circle cx="350.5" cy="530.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_46"> + <circle cx="304.5" cy="612.5" r="15.5000247674794" fill="black"/> + <circle cx="304.5" cy="612.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_45"> + <circle cx="453.5" cy="530.5" r="15.5000247674794" fill="black"/> + <circle cx="453.5" cy="530.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_44"> + <circle cx="510.5" cy="612.5" r="15.5000247674794" fill="black"/> + <circle cx="510.5" cy="612.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_43"> + <circle cx="401.5" cy="612.5" r="15.5000247674794" fill="black"/> + <circle cx="401.5" cy="612.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_42"> + <line x1="388.4253" y1="462.89403" x2="364.88604" y2="504.8553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_41"> + <line x1="342.4253" y1="544.894" x2="318.88604" y2="586.8553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_40"> + <line x1="359.2162" y1="544.51426" x2="385.97086" y2="587.5316" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_39"> + <line x1="405.9193" y1="462.0506" x2="436.71773" y2="506.3571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_38"> + <line x1="462.9193" y1="544.0506" x2="493.71773" y2="588.3571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_37"> + <circle cx="58.5" cy="511.25" r="15.5000247674794" fill="black"/> + <circle cx="58.5" cy="511.25" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_36"> + <circle cx="12.5" cy="593.25" r="15.5000247674794" fill="black"/> + <circle cx="12.5" cy="593.25" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_35"> + <circle cx="109.5" cy="593.25" r="15.5000247674794" fill="black"/> + <circle cx="109.5" cy="593.25" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_34"> + <line x1="50.4253" y1="525.644" x2="26.886036" y2="567.6053" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_33"> + <line x1="67.216185" y1="525.26426" x2="93.97086" y2="568.2816" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_49"> + <circle cx="-142.5" cy="593.25" r="15.5000247674794" fill="black"/> + <circle cx="-142.5" cy="593.25" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_83"> + <circle cx="375.5" cy="721.5" r="15.5000247674794" fill="black"/> + <circle cx="375.5" cy="721.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_82"> + <circle cx="329.5" cy="803.5" r="15.5000247674794" fill="black"/> + <circle cx="329.5" cy="803.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_81"> + <circle cx="283.5" cy="885.5" r="15.5000247674794" fill="black"/> + <circle cx="283.5" cy="885.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_80"> + <circle cx="432.5" cy="803.5" r="15.5000247674794" fill="black"/> + <circle cx="432.5" cy="803.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_79"> + <circle cx="489.5" cy="885.5" r="15.5000247674793" fill="black"/> + <circle cx="489.5" cy="885.5" r="15.5000247674793" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_78"> + <circle cx="380.5" cy="885.5" r="15.5000247674793" fill="black"/> + <circle cx="380.5" cy="885.5" r="15.5000247674793" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_77"> + <line x1="367.4253" y1="735.894" x2="343.88604" y2="777.8553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_76"> + <line x1="321.4253" y1="817.894" x2="297.88604" y2="859.8553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_75"> + <line x1="338.2162" y1="817.5143" x2="364.97086" y2="860.5316" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_74"> + <line x1="384.91932" y1="735.0506" x2="415.71773" y2="779.3571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_73"> + <line x1="441.9193" y1="817.0506" x2="472.71773" y2="861.3571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_72"> + <circle cx="56" cy="803.5" r="15.5000247674794" fill="black"/> + <circle cx="56" cy="803.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_71"> + <circle cx="10" cy="885.5" r="15.5000247674793" fill="black"/> + <circle cx="10" cy="885.5" r="15.5000247674793" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_70"> + <circle cx="107" cy="885.5" r="15.5000247674794" fill="black"/> + <circle cx="107" cy="885.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_69"> + <line x1="47.9253" y1="817.894" x2="24.386036" y2="859.8553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_68"> + <line x1="64.716185" y1="817.5143" x2="91.47086" y2="860.5316" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_67"> + <circle cx="-153.5" cy="879.5" r="15.5000247674794" fill="white"/> + <circle cx="-153.5" cy="879.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_99"> + <circle cx="323.5" cy="1063" r="15.5000247674794" fill="black"/> + <circle cx="323.5" cy="1063" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_98"> + <circle cx="277.5" cy="1145" r="15.5000247674794" fill="black"/> + <circle cx="277.5" cy="1145" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_97"> + <circle cx="426.5" cy="1063" r="15.5000247674794" fill="black"/> + <circle cx="426.5" cy="1063" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_96"> + <circle cx="483.5" cy="1145" r="15.5000247674794" fill="black"/> + <circle cx="483.5" cy="1145" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_95"> + <circle cx="374.5" cy="1145" r="15.5000247674794" fill="black"/> + <circle cx="374.5" cy="1145" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_94"> + <line x1="369.5" y1="981" x2="336.9073" y2="1039.1" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_93"> + <line x1="315.4253" y1="1077.394" x2="291.88604" y2="1119.3553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_92"> + <line x1="332.2162" y1="1077.0143" x2="358.97086" y2="1120.0316" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_91"> + <line x1="369.5" y1="981" x2="410.85947" y2="1040.4996" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_90"> + <line x1="435.9193" y1="1076.5506" x2="466.71773" y2="1120.8571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_89"> + <circle cx="-151.5" cy="1015.5" r="15.5000247674794" fill="black"/> + <circle cx="-151.5" cy="1015.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_88"> + <circle cx="-197.5" cy="1097.5" r="15.5000247674794" fill="black"/> + <circle cx="-197.5" cy="1097.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_87"> + <circle cx="-100.5" cy="1097.5" r="15.5000247674795" fill="white"/> + <circle cx="-100.5" cy="1097.5" r="15.5000247674795" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_86"> + <line x1="-159.5747" y1="1029.894" x2="-183.11396" y2="1071.8553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_85"> + <line x1="-142.78382" y1="1029.5143" x2="-116.02914" y2="1072.5316" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_115"> + <circle cx="-181.5" cy="1280" r="15.5000247674795" fill="black"/> + <circle cx="-181.5" cy="1280" r="15.5000247674795" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_114"> + <circle cx="-227.5" cy="1362" r="15.5000247674794" fill="black"/> + <circle cx="-227.5" cy="1362" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_113"> + <circle cx="-78.5" cy="1280" r="15.5000247674795" fill="black"/> + <circle cx="-78.5" cy="1280" r="15.5000247674795" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_112"> + <circle cx="-21.5" cy="1362" r="15.5000247674794" fill="black"/> + <circle cx="-21.5" cy="1362" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_111"> + <circle cx="-124.5" cy="1362" r="15.5000247674794" fill="white"/> + <circle cx="-124.5" cy="1362" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_110"> + <line x1="-135.5" y1="1198" x2="-168.09272" y2="1256.1" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_109"> + <line x1="-189.5747" y1="1294.394" x2="-213.11396" y2="1336.3553" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_108"> + <line x1="-172.08068" y1="1293.5506" x2="-141.28227" y2="1337.8571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_107"> + <line x1="-135.5" y1="1198" x2="-94.14053" y2="1257.4996" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Line_106"> + <line x1="-69.08068" y1="1293.5506" x2="-38.28227" y2="1337.8571" marker-end="url(#FilledArrow_Marker)" stroke="black" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_116"> + <circle cx="369.5" cy="984.5" r="15.5000247674794" fill="black"/> + <circle cx="369.5" cy="984.5" r="15.5000247674794" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_117"> + <circle cx="-136.5" cy="1200.5" r="15.5000247674795" fill="black"/> + <circle cx="-136.5" cy="1200.5" r="15.5000247674795" stroke="gray" stroke-linecap="round" stroke-linejoin="round" stroke-width="2"/> + </g> + <g id="Graphic_120"> + <text transform="translate(360.91 1307)" fill="black"> + <tspan font-family="Helvetica Neue" font-size="14" font-weight="400" fill="black" x="20383695e-20" y="13">[]</tspan> + </text> + </g> + <g id="Graphic_122"> + <text transform="translate(397.874 -174)" fill="black"> + <tspan font-family="Helvetica Neue" font-size="14" font-weight="400" fill="black" x="20383695e-20" y="13">[]</tspan> + </text> + </g> + </g> + </g> +</svg>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/hoareBinaryTree.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,772 @@ +module hoareBinaryTree where + +open import Level renaming (zero to Z ; suc to succ) + +open import Data.Nat hiding (compare) +open import Data.Nat.Properties as NatProp +open import Data.Maybe +-- open import Data.Maybe.Properties +open import Data.Empty +open import Data.List +open import Data.Product + +open import Function as F hiding (const) + +open import Relation.Binary +open import Relation.Binary.PropositionalEquality +open import Relation.Nullary +open import logic + + +_iso_ : {n : Level} {a : Set n} → ℕ → ℕ → Set +d iso d' = (¬ (suc d ≤ d')) ∧ (¬ (suc d' ≤ d)) + +iso-intro : {n : Level} {a : Set n} {x y : ℕ} → ¬ (suc x ≤ y) → ¬ (suc y ≤ x) → _iso_ {n} {a} x y +iso-intro = λ z z₁ → record { proj1 = z ; proj2 = z₁ } + +-- +-- +-- no children , having left node , having right node , having both +-- +data bt {n : Level} (A : Set n) : Set n where + leaf : bt A + node : (key : ℕ) → (value : A) → + (left : bt A ) → (right : bt A ) → bt A + +node-key : {n : Level} {A : Set n} → bt A → Maybe ℕ +node-key (node key _ _ _) = just key +node-key _ = nothing + +node-value : {n : Level} {A : Set n} → bt A → Maybe A +node-value (node _ value _ _) = just value +node-value _ = nothing + +bt-depth : {n : Level} {A : Set n} → (tree : bt A ) → ℕ +bt-depth leaf = 0 +bt-depth (node key value t t₁) = suc (Data.Nat._⊔_ (bt-depth t ) (bt-depth t₁ )) + +find' : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → List (bt A) + → (next : bt A → List (bt A) → t ) → (exit : bt A → List (bt A) → t ) → t +find' key leaf st _ exit = exit leaf st +find' key (node key₁ v1 tree tree₁) st next exit with <-cmp key key₁ +find' key n st _ exit | tri≈ ¬a b ¬c = exit n st +find' key n@(node key₁ v1 tree tree₁) st next _ | tri< a ¬b ¬c = next tree (n ∷ st) +find' key n@(node key₁ v1 tree tree₁) st next _ | tri> ¬a ¬b c = next tree₁ (n ∷ st) + +find : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) + → (next : (tree1 : bt A) → (stack : List (bt A)) → t) + → (exit : (tree1 : bt A) → (stack : List (bt A)) → t) → t +find key leaf st _ exit = exit leaf st +find key (node key₁ v1 tree tree₁) st next exit with <-cmp key key₁ +find key n st _ exit | tri≈ ¬a refl ¬c = exit n st +find {n} {_} {A} key (node key₁ v1 tree tree₁) st next _ | tri< a ¬b ¬c = next tree (tree ∷ st) +find key n@(node key₁ v1 tree tree₁) st next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st) + +{-# TERMINATING #-} +find-loop : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → bt A → List (bt A) → (exit : bt A → List (bt A) → t) → t +find-loop {n} {m} {A} {t} key tree st exit = find-loop1 tree st where + find-loop1 : bt A → List (bt A) → t + find-loop1 tree st = find key tree st find-loop1 exit + +replaceNode : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → (bt A → t) → t +replaceNode k v1 leaf next = next (node k v1 leaf leaf) +replaceNode k v1 (node key value t t₁) next = next (node k v1 t t₁) + +replace : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → List (bt A) → (next : ℕ → A → bt A → List (bt A) → t ) → (exit : bt A → t) → t +replace key value repl [] next exit = exit repl -- can't happen +replace key value repl (leaf ∷ []) next exit = exit repl +replace key value repl (node key₁ value₁ left right ∷ []) next exit with <-cmp key key₁ +... | tri< a ¬b ¬c = exit (node key₁ value₁ repl right ) +... | tri≈ ¬a b ¬c = exit (node key₁ value left right ) +... | tri> ¬a ¬b c = exit (node key₁ value₁ left repl ) +replace key value repl (leaf ∷ st) next exit = next key value repl st +replace key value repl (node key₁ value₁ left right ∷ st) next exit with <-cmp key key₁ +... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st +... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st +... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st + +{-# TERMINATING #-} +replace-loop : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → bt A → List (bt A) → (exit : bt A → t) → t +replace-loop {_} {_} {A} {t} key value tree st exit = replace-loop1 key value tree st where + replace-loop1 : (key : ℕ) → (value : A) → bt A → List (bt A) → t + replace-loop1 key value tree st = replace key value tree st replace-loop1 exit + +insertTree : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → (next : bt A → t ) → t +insertTree tree key value exit = find-loop key tree ( tree ∷ [] ) $ λ t st → replaceNode key value t $ λ t1 → replace-loop key value t1 st exit + +insertTest1 = insertTree leaf 1 1 (λ x → x ) +insertTest2 = insertTree insertTest1 2 1 (λ x → x ) +insertTest3 = insertTree insertTest2 3 3 (λ x → x ) +insertTest4 = insertTree insertTest3 1 4 (λ x → x ) -- this is wrong + +updateTree : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → (empty : bt A → t ) → (next : A → bt A → t ) → t +updateTree {_} {_} {A} {t} tree key value empty next = find-loop key tree ( tree ∷ [] ) + $ λ t st → replaceNode key value t $ λ t1 → replace-loop key value t1 st (found? st) where + found? : List (bt A) → bt A → t + found? [] tree = empty tree -- can't happen + found? (leaf ∷ st) tree = empty tree + found? (node key value x x₁ ∷ st) tree = next value tree + +open import Data.Unit hiding ( _≟_ ; _≤?_ ; _≤_) + +data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where + t-leaf : treeInvariant leaf + t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) + t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key₁ value₁ t₁ t₂) + → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) + t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → (key < key₁) → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) + t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → (key < key₁) → (key₁ < key₂) + → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₂ value₂ t₃ t₄) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) + +-- +-- stack always contains original top at end +-- +data stackInvariant {n : Level} {A : Set n} (key : ℕ) : (top orig : bt A) → (stack : List (bt A)) → Set n where + s-single : {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 + +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) (t-node (add< 0) (add< 1) (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-single ) + +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-single ) () +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-single ) = refl +si-property1 (s-right _ si) = refl +si-property1 (s-left _ si) = refl + +si-property-last : {n : Level} {A : Set n} (key : ℕ) (tree tree0 : bt A) → (stack : List (bt A)) → stackInvariant key tree tree0 stack + → stack-last stack ≡ just tree0 +si-property-last key t t0 (t ∷ []) (s-single ) = 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 + +ti-right : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 tree₁ repl) → treeInvariant repl +ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf +ti-right {_} {_} {.leaf} {_} {key₁} {v1} (t-right x ti) = ti +ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-left x ti) = t-leaf +ti-right {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti₁ + +ti-left : {n : Level} {A : Set n} {tree₁ repl : bt A} → {key₁ : ℕ} → {v1 : A} → treeInvariant (node key₁ v1 repl tree₁ ) → treeInvariant repl +ti-left {_} {_} {.leaf} {_} {key₁} {v1} (t-single .key₁ .v1) = t-leaf +ti-left {_} {_} {_} {_} {key₁} {v1} (t-right x ti) = t-leaf +ti-left {_} {_} {_} {_} {key₁} {v1} (t-left x ti) = ti +ti-left {_} {_} {.(node _ _ _ _)} {_} {key₁} {v1} (t-node x x₁ ti ti₁) = ti + +stackTreeInvariant : {n : Level} {A : Set n} (key : ℕ) (sub tree : bt A) → (stack : List (bt A)) + → treeInvariant tree → stackInvariant key sub tree stack → treeInvariant sub +stackTreeInvariant {_} {A} key sub tree (sub ∷ []) ti (s-single ) = ti +stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-right _ si ) = ti-right (si1 si) where + si1 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 tree₁ sub ) tree st → treeInvariant (node key₁ v1 tree₁ sub ) + si1 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 tree₁ sub ) tree st ti si +stackTreeInvariant {_} {A} key sub tree (sub ∷ st) ti (s-left _ si ) = ti-left ( si2 si) where + si2 : {tree₁ : bt A} → {key₁ : ℕ} → {v1 : A} → stackInvariant key (node key₁ v1 sub tree₁ ) tree st → treeInvariant (node key₁ v1 sub tree₁ ) + si2 {tree₁ } {key₁ } {v1 } si = stackTreeInvariant key (node key₁ v1 sub tree₁ ) tree st ti si + +rt-property1 : {n : Level} {A : Set n} (key : ℕ) (value : A) (tree tree1 : bt A ) → replacedTree key value tree tree1 → ¬ ( tree1 ≡ leaf ) +rt-property1 {n} {A} key value .leaf .(node key value leaf leaf) r-leaf () +rt-property1 {n} {A} key value .(node key _ _ _) .(node key value _ _) r-node () +rt-property1 {n} {A} key value .(node _ _ _ _) _ (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 + +nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ +nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x +nat-<> : { x y : ℕ } → x < y → y < x → ⊥ +nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x + +open _∧_ + + +depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) +depth-1< {i} {j} = s≤s (m≤m⊔n _ j) + +depth-2< : {i j : ℕ} → suc i ≤ suc (j Data.Nat.⊔ i ) +depth-2< {i} {j} = s≤s {! !} + +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₁ + + +-- record FindCond {n : Level} {A : Set n} (C : ℕ → bt A → Set n) : Set (Level.suc n) where +-- field +-- c1 : {key key₁ : ℕ} {v1 : A } { tree tree₁ : bt A } → C key (node key₁ v1 tree tree₁) → key < key₁ → C key tree +-- c2 : {key key₁ : ℕ} {v1 : A } { tree tree₁ : bt A } → C key (node key₁ v1 tree tree₁) → key > key₁ → C key tree₁ +-- +-- +-- findP0 : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) +-- → {C : ℕ → bt A → Set n} → C key tree → FindCond C +-- → (next : (tree1 : bt A) → (stack : List (bt A)) → C key tree1 → bt-depth tree1 < bt-depth tree → t ) +-- → (exit : (tree1 : bt A) → (stack : List (bt A)) → C key tree1 +-- → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t +-- findP0 key leaf st Pre _ _ exit = exit leaf st Pre (case1 refl) +-- findP0 key (node key₁ v1 tree tree₁) st Pre _ next exit with <-cmp key key₁ +-- findP0 key n st Pre e _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl) +-- findP0 {n} {_} {A} key (node key₁ v1 tree tree₁) st Pre e next _ | tri< a ¬b ¬c = next tree (tree ∷ st) (FindCond.c1 e Pre a) depth-1< +-- findP0 key n@(node key₁ v1 tree tree₁) st Pre e next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st) (FindCond.c2 e Pre c) depth-2< + +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 t) = t-right x t +replaceTree1 k v1 value (t-left x t) = t-left x t +replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ + +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 + +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 + +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 +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 (≤-step lt1) p1 lt1 ) where + TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → Invraiant r1 → reduce r1 < reduce r → t + TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 p1 (λ r2 p1 lt1 → ⊥-elim (lemma5 n≤j lt1)) + TerminatingLoop1 (suc j) r r1 n≤j p1 lt with <-cmp (reduce r1) (suc j) + ... | tri< a ¬b ¬c = TerminatingLoop1 j r r1 a p1 lt + ... | tri≈ ¬a b ¬c = loop r1 p1 (λ r2 p2 lt1 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b lt1 ) p2 lt1 ) + ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) + +record LoopTermination {n : Level} {Index : Set n } { reduce : Index → ℕ } + (r : Index ) (C : Set n) : Set (Level.suc n) where + field + rd : (r1 : Index) → reduce r1 < reduce r + ci : C -- data continuation + +-- TerminatingLoopC : {l n : Level} {t : Set l} (Index : Set n ) → {C : Set n } → ( reduce : Index → ℕ) +-- → (r : Index) → (P : LoopTermination r C ) +-- → (loop : (r : Index) → LoopTermination {_} {_} {reduce} r C → (next : (r1 : Index) → LoopTermination r1 C → t ) → t) → t +-- TerminatingLoopC {_} {_} {t} Index {C} reduce r P loop with <-cmp 0 (reduce r) +-- ... | tri≈ ¬a b ¬c = loop r P (λ r1 p1 → ⊥-elim (lemma3 b (LoopTermination.rd P r1))) +-- ... | tri< a ¬b ¬c = loop r P (λ r1 p1 → TerminatingLoop1 (reduce r) r r1 (≤-step (LoopTermination.rd P r1)) p1 (LoopTermination.rd P r1)) where +-- TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → {!!} → reduce r1 < reduce r → t +-- TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 {!!} (λ r2 P1 → ⊥-elim (lemma5 n≤j (LoopTermination.rd P1 r2))) +-- 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 {!!} (λ r2 p2 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b {!!} ) p2 {!!} ) +-- ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) +-- +-- record ReplCond {n : Level} {A : Set n} (C : ℕ → bt A → List (bt A) → Set n) : Set (Level.suc n) where +-- field +-- c1 : (key : ℕ) → (repl : bt A) → (stack : List (bt A)) → C key repl stack +-- +-- replaceP0 : {n m : Level} {A : Set n} {t : Set m} +-- → (key : ℕ) → (value : A) → ( repl : bt A) +-- → (stack : List (bt A)) +-- → {C : ℕ → (repl : bt A ) → List (bt A) → Set n} → C key repl stack → ReplCond C +-- → (next : ℕ → A → (repl : bt A) → (stack1 : List (bt A)) +-- → C key repl stack → length stack1 < length stack → t) +-- → (exit : (repl : bt A) → C key repl [] → t) → t +-- replaceP0 key value repl [] Pre _ next exit = exit repl {!!} +-- replaceP0 key value repl (leaf ∷ []) Pre _ next exit = exit (node key value leaf leaf) {!!} +-- replaceP0 key value repl (node key₁ value₁ left right ∷ []) Pre e next exit with <-cmp key key₁ +-- ... | tri< a ¬b ¬c = exit (node key₁ value₁ repl right ) {!!} +-- ... | tri≈ ¬a b ¬c = exit repl {!!} +-- ... | tri> ¬a ¬b c = exit (node key₁ value₁ left repl ) {!!} +-- replaceP0 {n} {_} {A} key value repl (leaf ∷ st@(tree1 ∷ st1)) Pre e next exit = next key value repl st {!!} ≤-refl +-- replaceP0 {n} {_} {A} key value repl (node key₁ value₁ left right ∷ st@(tree1 ∷ st1)) Pre e next exit with <-cmp key key₁ +-- ... | tri< a ¬b ¬c = next key value (node key₁ value₁ repl right ) st {!!} ≤-refl +-- ... | tri≈ ¬a b ¬c = next key value (node key₁ value left right ) st {!!} ≤-refl +-- ... | tri> ¬a ¬b c = next key value (node key₁ value₁ left repl ) st {!!} ≤-refl +-- +-- +open _∧_ + +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 ti) r-node = t-right x ti +RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI0 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ 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₁ ti) (r-right x ri) = t-single key₁ value₁ +RTtoTI0 (node key₁ _ leaf right@(node key₂ _ _ _)) (node key₁ value₁ leaf right₁@(node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = + t-right (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) +RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left x₁ 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₁ ti) (r-right x r-leaf) = + t-node x₁ x ti (t-single key value) +RTtoTI0 (node key₁ _ (node _ _ _ _) (node key₂ _ _ _)) (node key₁ _ (node _ _ _ _) (node key₃ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = + t-node x₁ (subst (λ k → key₁ < k) (rt-property-key ri) x₂) ti (RTtoTI0 _ _ key value ti₁ ri) +-- 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 (t-single _ _ ) +RTtoTI0 .(node _ _ leaf (node _ _ _ _)) (node key₁ value₁ (node key value leaf leaf) (node _ _ _ _)) key value (t-right x₁ ti) (r-left x r-leaf) = t-node x x₁ (t-single key value) ti +RTtoTI0 (node key₃ _ (node key₂ _ _ _) leaf) (node key₃ _ (node key₁ value₁ left left₁) leaf) key value (t-left x₁ ti) (r-left x ri) = + t-left (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) -- key₁ < key₃ +RTtoTI0 (node key₁ _ (node key₂ _ _ _) (node _ _ _ _)) (node key₁ _ (node key₃ _ _ _) (node _ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = t-node (subst (λ k → k < key₁ ) (rt-property-key ri) x₁) x₂ (RTtoTI0 _ _ key value ti ri) ti₁ + +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 ti) r-node = t-right x ti +RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI1 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ ti ti₁ +-- r-right case +RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ leaf (node _ _ _ _)) key value (t-right x₁ 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₁ 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₂ 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₂ ti ti₁) (r-right x ri) = t-node x₁ (subst (λ k → key₄ < k ) (sym (rt-property-key ri)) x₂) 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₁ 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₁ 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₂ ti ti₁) (r-left x r-leaf) = t-right x₂ ti₁ +RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = + t-node (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ (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-single ⟫ + $ λ 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 + $ λ _ x P → insertTreeP x 2 1 (proj1 P) + $ λ _ x P → insertTreeP x 3 2 (proj1 P) + $ λ _ x P → insertTreeP x 2 2 (proj1 P) (λ _ x _ → 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 + +record findPR {n : Level} {A : Set n} (key : ℕ) (tree : bt A ) (stack : List (bt A)) (C : ℕ → bt A → List (bt A) → Set n) : Set n where + field + tree0 : bt A + ti0 : treeInvariant tree0 + ti : treeInvariant tree + si : stackInvariant key tree tree0 stack + ci : C key tree stack -- data continuation + +record findExt {n : Level} {A : Set n} (key : ℕ) (C : ℕ → bt A → List (bt A) → Set n) : Set (Level.suc n) where + field + c1 : {key₁ : ℕ} {tree tree₁ : bt A } {st : List (bt A)} {v1 : A} + → findPR key (node key₁ v1 tree tree₁) st C → key < key₁ → C key tree (tree ∷ st) + c2 : {key₁ : ℕ} {tree tree₁ : bt A } {st : List (bt A)} {v1 : A} + → findPR key (node key₁ v1 tree tree₁) st C → key > key₁ → C key tree₁ (tree₁ ∷ st) + +findPP : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree : bt A ) → (stack : List (bt A)) + → {C : ℕ → bt A → List (bt A) → Set n } → findPR key tree stack C → findExt key C + → (next : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack C → bt-depth tree1 < bt-depth tree → t ) + → (exit : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack C + → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t +findPP key leaf st Pre _ _ exit = exit leaf st Pre (case1 refl) +findPP key (node key₁ v1 tree tree₁) st Pre _ next exit with <-cmp key key₁ +findPP key n st Pre _ _ exit | tri≈ ¬a refl ¬c = exit n st Pre (case2 refl) +findPP {n} {_} {A} key (node key₁ v1 tree tree₁) st Pre e next _ | tri< a ¬b ¬c = next tree (tree ∷ st) + record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 Pre ; ti = treeLeftDown tree tree₁ (findPR.ti Pre) ; si = findP1 a st (findPR.si Pre) ; ci = findExt.c1 e Pre a } depth-1< where + findP1 : key < key₁ → (st : List (bt A)) → stackInvariant key (node key₁ v1 tree tree₁) + (findPR.tree0 Pre) st → stackInvariant key tree (findPR.tree0 Pre) (tree ∷ st) + findP1 a (x ∷ st) si = s-left a si +findPP key n@(node key₁ v1 tree tree₁) st Pre e next _ | tri> ¬a ¬b c = next tree₁ (tree₁ ∷ st) + record { tree0 = findPR.tree0 Pre ; ti0 = findPR.ti0 Pre ; ti = treeRightDown tree tree₁ (findPR.ti Pre) ; si = s-right c (findPR.si Pre) ; ci = findExt.c2 e Pre c } depth-2< + +insertTreePP : {n m : Level} {A : Set n} {t : Set m} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree + → (exit : (tree repl : bt A) → treeInvariant tree ∧ replacedTree key value tree repl → t ) → t +insertTreePP {n} {m} {A} {t} tree key value P0 exit = + TerminatingLoopS (bt A ∧ List (bt A) ) {λ p → findPR key (proj1 p) (proj2 p) (λ _ _ _ → Lift n ⊤) } (λ p → bt-depth (proj1 p)) ⟪ tree , tree ∷ [] ⟫ + record { tree0 = tree ; ti = P0 ; ti0 = P0 ;si = s-single ; ci = lift tt } + $ λ p P loop → findPP key (proj1 p) (proj2 p) P record { c1 = λ _ _ → lift tt ; c2 = λ _ _ → lift tt } (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) + $ λ t s P C → replaceNodeP key value t C (findPR.ti 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 = findPR.tree0 P ; ti = findPR.ti0 P ; si = findPR.si 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 ) exit + +record findPC {n : Level} {A : Set n} (value : A) (key1 : ℕ) (tree : bt A ) (stack : List (bt A)) : Set n where + field + tree1 : bt A + ci : replacedTree key1 value tree1 tree + +findPPC1 : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (value : A) → (tree : bt A ) → (stack : List (bt A)) + → findPR key tree stack (findPC value ) + → (next : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC value ) → bt-depth tree1 < bt-depth tree → t ) + → (exit : (tree1 : bt A) → (stack : List (bt A)) → findPR key tree1 stack (findPC value ) + → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t +findPPC1 {n} {_} {A} key value tree stack Pr next exit = findPP key tree stack Pr findext next exit where + findext01 : {key₂ : ℕ} {tree₁ : bt A} {tree₂ : bt A} {st : List (bt A)} {v1 : A} + → (Pre : findPR key (node key₂ v1 tree₁ tree₂) st (findPC value) ) + → key < key₂ → findPC value key tree₁ (tree₁ ∷ st) + findext01 Pre a with findPC.ci (findPR.ci Pre) | findPC.tree1 (findPR.ci Pre) | inspect findPC.tree1 (findPR.ci Pre) + ... | r-leaf | leaf | record { eq = refl } = ⊥-elim ( nat-≤> a ≤-refl) + ... | r-node | node key value t1 t3 | record { eq = refl } = ⊥-elim ( nat-≤> a ≤-refl ) + ... | r-right x t | t1 | t2 = ⊥-elim (nat-<> x a) + ... | r-left x ri | node key value t1 t3 | record { eq = refl } = record { tree1 = t1 ; ci = ri } + findext02 : {key₂ : ℕ} {tree₁ : bt A} {tree₂ : bt A} {st : List (bt A)} {v1 : A} + → (Pre : findPR key (node key₂ v1 tree₁ tree₂) st (findPC value) ) + → key > key₂ → findPC value key tree₂ (tree₂ ∷ st) + findext02 Pre c with findPC.ci (findPR.ci Pre) | findPC.tree1 (findPR.ci Pre) | inspect findPC.tree1 (findPR.ci Pre) + ... | r-leaf | leaf | record { eq = refl } = ⊥-elim ( nat-≤> c ≤-refl) + ... | r-node | node key value t1 t3 | record { eq = refl } = ⊥-elim ( nat-≤> c ≤-refl ) + ... | r-left x t | t1 | t2 = ⊥-elim (nat-<> x c) + ... | r-right x ri | node key value t1 t3 | record { eq = refl } = record { tree1 = t3 ; ci = ri } + findext : findExt key (findPC value ) + findext = record { c1 = findext01 ; c2 = findext02 } + +insertTreeSpec0 : {n : Level} {A : Set n} → (tree : bt A) → (value : A) → top-value tree ≡ just value → ⊤ +insertTreeSpec0 _ _ _ = tt + +containsTree : {n : Level} {A : Set n} → (tree tree1 : bt A) → (key : ℕ) → (value : A) → treeInvariant tree1 → replacedTree key value tree1 tree → ⊤ +containsTree {n} {A} tree tree1 key value P RT = + TerminatingLoopS (bt A ∧ List (bt A) ) + {λ p → findPR key (proj1 p) (proj2 p) (findPC value ) } (λ p → bt-depth (proj1 p)) + ⟪ tree , tree ∷ [] ⟫ record { tree0 = tree ; ti0 = RTtoTI0 _ _ _ _ P RT ; ti = RTtoTI0 _ _ _ _ P RT ; si = s-single + ; ci = record { tree1 = tree1 ; ci = RT } } + $ λ p P loop → findPPC1 key value (proj1 p) (proj2 p) P (λ t s P1 lt → loop ⟪ t , s ⟫ P1 lt ) + $ λ t1 s1 P2 found? → insertTreeSpec0 t1 value (lemma6 t1 s1 found? P2) where + lemma6 : (t1 : bt A) (s1 : List (bt A)) (found? : (t1 ≡ leaf) ∨ (node-key t1 ≡ just key)) (P2 : findPR key t1 s1 (findPC value )) → top-value t1 ≡ just value + lemma6 t1 s1 found? P2 = lemma7 t1 s1 (findPR.tree0 P2) ( findPC.tree1 (findPR.ci P2)) (findPC.ci (findPR.ci P2)) (findPR.si P2) found? where + lemma8 : {tree1 t1 : bt A } → replacedTree key value tree1 t1 → node-key t1 ≡ just key → top-value t1 ≡ just value + lemma8 {.leaf} {node key value .leaf .leaf} r-leaf refl = refl + lemma8 {.(node key _ t1 t2)} {node key value t1 t2} r-node refl = refl + lemma8 {.(node key value t1 _)} {node key value t1 t2} (r-right x ri) refl = ⊥-elim (¬x<x x) + lemma8 {.(node key value _ t2)} {node key value t1 t2} (r-left x ri) refl = ⊥-elim (¬x<x x) + lemma7 : (t1 : bt A) ( s1 : List (bt A) ) (tree0 tree1 : bt A) + → replacedTree key value tree1 t1 → stackInvariant key t1 tree0 s1 + → ( t1 ≡ leaf ) ∨ ( node-key t1 ≡ just key) → top-value t1 ≡ just value + lemma7 .leaf (.leaf ∷ []) .leaf tree1 () s-single (case1 refl) + lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ []) .(node key value t1 t2) tree1 ri s-single (case2 x) = lemma8 ri x + lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ x₁ ∷ s1) tree0 tree1 ri (s-right x si) (case2 x₂) = lemma8 ri x₂ + lemma7 (node key value t1 t2) (.(node key value t1 t2) ∷ x₁ ∷ s1) tree0 tree1 ri (s-left x si) (case2 x₂) = lemma8 ri x₂ + + +containsTree1 : {n : Level} {A : Set n} → (tree : bt A) → (key : ℕ) → (value : A) → treeInvariant tree → ⊤ +containsTree1 {n} {A} tree key value ti = + insertTreeP tree key value ti + $ λ tree0 tree1 P → containsTree tree1 tree0 key value (RTtoTI1 _ _ _ _ (proj1 P) (proj2 P) ) (proj2 P) + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/hoareBinaryTree1.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,791 @@ +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 ( _≟_ ; _≤?_ ; _≤_) + +data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where + t-leaf : treeInvariant leaf + t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) + t-right : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key₁ value₁ t₁ t₂) + → treeInvariant (node key value leaf (node key₁ value₁ t₁ t₂)) + t-left : {key key₁ : ℕ} → {value value₁ : A} → {t₁ t₂ : bt A} → key < key₁ → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) leaf ) + t-node : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt A} → key < key₁ → key₁ < key₂ + → treeInvariant (node key value t₁ t₂) + → treeInvariant (node key₂ value₂ t₃ t₄) + → treeInvariant (node key₁ value₁ (node key value t₁ t₂) (node key₂ value₂ t₃ t₄)) + +-- +-- 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 + +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) (t-node (add< 0) (add< 1) (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-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 + +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 + +nat-≤> : { x y : ℕ } → x ≤ y → y < x → ⊥ +nat-≤> (s≤s x<y) (s≤s y<x) = nat-≤> x<y y<x +nat-<> : { x y : ℕ } → x < y → y < x → ⊥ +nat-<> (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x + +open _∧_ + + +depth-1< : {i j : ℕ} → suc i ≤ suc (i Data.Nat.⊔ j ) +depth-1< {i} {j} = s≤s (m≤m⊔n _ j) + +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 t) = t-right x t +replaceTree1 k v1 value (t-left x t) = t-left x t +replaceTree1 k v1 value (t-node x x₁ t t₁) = t-node x x₁ t t₁ + +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 + +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 + +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 +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 (≤-step lt1) p1 lt1 ) where + TerminatingLoop1 : (j : ℕ) → (r r1 : Index) → reduce r1 < suc j → Invraiant r1 → reduce r1 < reduce r → t + TerminatingLoop1 zero r r1 n≤j p1 lt = loop r1 p1 (λ r2 p1 lt1 → ⊥-elim (lemma5 n≤j lt1)) + TerminatingLoop1 (suc j) r r1 n≤j p1 lt with <-cmp (reduce r1) (suc j) + ... | tri< a ¬b ¬c = TerminatingLoop1 j r r1 a p1 lt + ... | tri≈ ¬a b ¬c = loop r1 p1 (λ r2 p2 lt1 → TerminatingLoop1 j r1 r2 (subst (λ k → reduce r2 < k ) b lt1 ) p2 lt1 ) + ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> c n≤j ) + +open _∧_ + +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 ti) r-node = t-right x ti +RTtoTI0 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI0 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ 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₁ ti) (r-right x ri) = t-single key₁ value₁ +RTtoTI0 (node key₁ _ leaf right@(node key₂ _ _ _)) (node key₁ value₁ leaf right₁@(node key₃ _ _ _)) key value (t-right x₁ ti) (r-right x ri) = + t-right (subst (λ k → key₁ < k ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) +RTtoTI0 (node key₁ _ (node _ _ _ _) leaf) (node key₁ _ (node key₃ value left right) leaf) key value₁ (t-left x₁ 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₁ ti) (r-right x r-leaf) = + t-node x₁ x ti (t-single key value) +RTtoTI0 (node key₁ _ (node _ _ _ _) (node key₂ _ _ _)) (node key₁ _ (node _ _ _ _) (node key₃ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-right x ri) = + t-node x₁ (subst (λ k → key₁ < k) (rt-property-key ri) x₂) ti (RTtoTI0 _ _ key value ti₁ ri) +-- 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 (t-single _ _ ) +RTtoTI0 .(node _ _ leaf (node _ _ _ _)) (node key₁ value₁ (node key value leaf leaf) (node _ _ _ _)) key value (t-right x₁ ti) (r-left x r-leaf) = t-node x x₁ (t-single key value) ti +RTtoTI0 (node key₃ _ (node key₂ _ _ _) leaf) (node key₃ _ (node key₁ value₁ left left₁) leaf) key value (t-left x₁ ti) (r-left x ri) = + t-left (subst (λ k → k < key₃ ) (rt-property-key ri) x₁) (RTtoTI0 _ _ key value ti ri) -- key₁ < key₃ +RTtoTI0 (node key₁ _ (node key₂ _ _ _) (node _ _ _ _)) (node key₁ _ (node key₃ _ _ _) (node _ _ _ _)) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = t-node (subst (λ k → k < key₁ ) (rt-property-key ri) x₁) x₂ (RTtoTI0 _ _ key value ti ri) ti₁ + +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 ti) r-node = t-right x ti +RTtoTI1 .(node key _ (node _ _ _ _) leaf) .(node key value (node _ _ _ _) leaf) key value (t-left x ti) r-node = t-left x ti +RTtoTI1 .(node key _ (node _ _ _ _) (node _ _ _ _)) .(node key value (node _ _ _ _) (node _ _ _ _)) key value (t-node x x₁ ti ti₁) r-node = t-node x x₁ ti ti₁ +-- r-right case +RTtoTI1 (node key₁ value₁ leaf leaf) (node key₁ _ leaf (node _ _ _ _)) key value (t-right x₁ 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₁ 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₂ 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₂ ti ti₁) (r-right x ri) = t-node x₁ (subst (λ k → key₄ < k ) (sym (rt-property-key ri)) x₂) 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₁ 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₁ 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₂ ti ti₁) (r-left x r-leaf) = t-right x₂ ti₁ +RTtoTI1 (node key₁ value₁ (node key₂ value₂ n n₁) _) (node key₁ _ _ _) key value (t-node x₁ x₂ ti ti₁) (r-left x ri) = + t-node (subst (λ k → k < key₁ ) (sym (rt-property-key ri)) x₁) x₂ (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 + +data RBtreeInvariant {n : Level} {A : Set n} : (tree : bt (Color ∧ A)) → (bd : ℕ) → Set n where + rb-leaf : RBtreeInvariant leaf 0 + rb-single : (key : ℕ) → (value : A) → (c : Color) → RBtreeInvariant (node key ⟪ c , value ⟫ leaf leaf) 1 + rb-right-red : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ → {d : ℕ } + → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ t t₁) d + → RBtreeInvariant (node key ⟪ Red , value ⟫ leaf (node key₁ ⟪ Black , value₁ ⟫ t t₁)) d + rb-right-black : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ → {d : ℕ } {c : Color} + → RBtreeInvariant (node key₁ ⟪ c , value₁ ⟫ t t₁) d + → RBtreeInvariant (node key ⟪ Black , value ⟫ leaf (node key₁ ⟪ c , value₁ ⟫ t t₁)) (suc d) --この時点でbalanceしてる必要はない + rb-left-red : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ → {d : ℕ } + → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ t t₁) d + → RBtreeInvariant (node key₁ ⟪ Red , value ⟫ (node key ⟪ Black , value₁ ⟫ t t₁) leaf ) d + rb-left-black : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ → {d : ℕ } {c : Color} + → RBtreeInvariant (node key₁ ⟪ c , value₁ ⟫ t t₁) d + → RBtreeInvariant (node key₁ ⟪ Black , value ⟫ (node key ⟪ c , value₁ ⟫ t t₁) leaf) (suc d) + rb-node-red : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt (Color ∧ A)} → key < key₁ → key₁ < key₂ → {d : ℕ} + → RBtreeInvariant (node key ⟪ Black , value ⟫ t₁ t₂) d + → RBtreeInvariant (node key₂ ⟪ Black , value₂ ⟫ t₃ t₄) d + → RBtreeInvariant (node key₁ ⟪ Red , value₁ ⟫ (node key ⟪ Black , value ⟫ t₁ t₂) (node key₂ ⟪ Black , value₂ ⟫ t₃ t₄)) d + rb-node-black : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt (Color ∧ A)} → key < key₁ → key₁ < key₂ → {d : ℕ} + → {c c₁ : Color} + → RBtreeInvariant (node key ⟪ c , value ⟫ t₁ t₂) d + → RBtreeInvariant (node key₂ ⟪ c₁ , value₂ ⟫ t₃ t₄) d + → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ (node key ⟪ c , value ⟫ t₁ t₂) (node key₂ ⟪ c₁ , value₂ ⟫ t₃ t₄)) (suc d) + + +-- This one can be very difficult +-- data replacedRBTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt (Color ∧ A) ) → Set n where +-- rb-leaf : replacedRBTree key value leaf (node key ⟪ Black , value ⟫ leaf leaf) + +color : {n : Level} (A : Set n) → (rb : bt (Color ∧ A)) → Color +color {n} A leaf = Red +color {n} A (node key ⟪ c , v ⟫ rb rb₁) = c + +RB→bt : {n : Level} (A : Set n) → (rb : bt (Color ∧ A)) → bt A +RB→bt {n} A leaf = leaf +RB→bt {n} A (node key ⟪ c , value ⟫ rb rb₁) = node key value (RB→bt A rb) (RB→bt A rb₁) + +data ParentGrand {n : Level} {A : Set n} (self : bt A) : (parent uncle grand : bt A) → Set n where +--self parent grand n の並び +-- n2 is uncle + 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 + +-- with color +data rotatedTree {n : Level} {A : Set n} : (before after : bt A ) → Set n where + rr-node : {t : bt A} → rotatedTree t t + -- a b + -- b c d a + -- d e e c + rr-right : {ka kb : ℕ } {va vb : A} → {c c₁ d d₁ e e₁ : bt A} + → ka < kb + → rotatedTree d d₁ → rotatedTree e e₁ → rotatedTree c c₁ + → rotatedTree (node ka va (node kb vb d e) c) (node kb vb d₁ (node ka va e₁ c₁) ) + -- b a + -- d a b c + -- e c d e + rr-left : {ka kb : ℕ } {va vb : A} → {c c₁ d d₁ e e₁ : bt A} + → ka < kb + → rotatedTree d d₁ → rotatedTree e e₁ → rotatedTree c c₁ + → rotatedTree (node kb vb d (node ka va e c) ) (node ka va (node kb vb d₁ e₁) c₁) + +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 ) + +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 } ) + + +PGtoRBinvariant : {n : Level} {A : Set n} → {key d0 ds dp dg : ℕ } → (tree orig : bt (Color ∧ A) ) + → RBtreeInvariant orig d0 + → (stack : List (bt (Color ∧ A))) → (pg : PG (Color ∧ A) tree stack ) + → RBtreeInvariant tree ds ∧ RBtreeInvariant (PG.parent pg) dp ∧ RBtreeInvariant (PG.grand pg) dg +PGtoRBinvariant = {!!} + +findRBT : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt (Color ∧ A) ) → (stack : List (bt (Color ∧ A))) → {d d0 : ℕ} + → RBtreeInvariant tree0 d0 + → RBtreeInvariant tree d ∧ stackInvariant key tree tree0 stack + → (next : (tree1 : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) → {d1 : ℕ} → RBtreeInvariant tree1 d1 ∧ stackInvariant key tree1 tree0 stack → bt-depth tree1 < bt-depth tree → t ) + → (exit : (tree1 : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) → {d1 : ℕ} → RBtreeInvariant tree1 d1 ∧ stackInvariant key tree1 tree0 stack + → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t +findRBT = {!!} + +rotateLeft : {n m : Level} {A : Set n} {t : Set m} + → (key : ℕ) → (value : A) → {d0 : ℕ} + → (orig tree rot repl : bt (Color ∧ A)) {d0 : ℕ} + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant repl dr + → (stack : List (bt (Color ∧ A))) → (si : stackInvariant key tree orig stack ) + → rotatedTree tree rot + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key rot) repl + → (next : {key2 d2 : ℕ} {c2 : Color} + → (to tr rot rr : bt (Color ∧ A)) + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant rr dr + → (stack1 : List (bt (Color ∧ A))) → stackInvariant key tr to stack1 + → rotatedTree tr rot + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key rot) rr + → length stack1 < length stack → t ) + → (exit : {k1 d1 : ℕ} {c1 : Color} → (rot repl : bt (Color ∧ A) ) + → {d : ℕ} → RBtreeInvariant repl d + → (ri : rotatedTree orig rot ) → {c : Color} → replacedTree key ⟪ c , value ⟫ rot repl → t ) → t +rotateLeft {n} {m} {A} {t} key value orig tree rot repl rbo rbt rbr stack si ri rr next exit = rotateLeft1 where + rotateLeft1 : t + rotateLeft1 with stackToPG tree orig stack si + ... | case1 x = {!!} + ... | case2 (case1 x) = {!!} + ... | case2 (case2 pg) = {!!} + +rotateRight : {n m : Level} {A : Set n} {t : Set m} + → (key : ℕ) → (value : A) → {d0 : ℕ} + → (orig tree rot repl : bt (Color ∧ A)) {d0 : ℕ} + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant repl dr + → (stack : List (bt (Color ∧ A))) → (si : stackInvariant key tree orig stack ) + → rotatedTree tree rot + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key rot) repl + → (next : {key2 d2 : ℕ} {c2 : Color} + → (to tr rot rr : bt (Color ∧ A)) + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant rr dr + → (stack1 : List (bt (Color ∧ A))) → stackInvariant key tr to stack1 + → rotatedTree tr rot + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key rot) rr + → length stack1 < length stack → t ) + → (exit : {k1 d1 : ℕ} {c1 : Color} → (rot repl : bt (Color ∧ A) ) + → {d : ℕ} → RBtreeInvariant repl d + → (ri : rotatedTree orig rot ) → {c : Color} → replacedTree key ⟪ c , value ⟫ rot repl → t ) → t +rotateRight {n} {m} {A} {t} key value orig tree rot repl rbo rbt rbr stack si ri rr next exit = rotateRight1 where + rotateRight1 : t + rotateRight1 with stackToPG tree orig stack si + ... | case1 x = {!!} + ... | case2 (case1 x) = {!!} + ... | case2 (case2 pg) = {!!} + +insertCase5 : {n m : Level} {A : Set n} {t : Set m} + → (key : ℕ) → (value : A) → {d0 : ℕ} + → (orig tree rot repl : bt (Color ∧ A)) {d0 : ℕ} + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant repl dr + → (stack : List (bt (Color ∧ A))) → (si : stackInvariant key tree orig stack ) + → rotatedTree tree rot + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key rot) repl + → (next : {key2 d2 : ℕ} {c2 : Color} + → (to tr rot rr : bt (Color ∧ A)) + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant rr dr + → (stack1 : List (bt (Color ∧ A))) → stackInvariant key tr to stack1 + → rotatedTree tr rot + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key rot) rr + → length stack1 < length stack → t ) + → (exit : {k1 d1 : ℕ} {c1 : Color} → (rot repl : bt (Color ∧ A) ) + → {d : ℕ} → RBtreeInvariant repl d + → (ri : rotatedTree orig rot ) → {c : Color} → replacedTree key ⟪ c , value ⟫ rot repl → t ) → t +insertCase5 {n} {m} {A} {t} key value orig tree rot repl rbo rbt rbr stack si ri rr next exit = insertCase51 where + insertCase51 : t + insertCase51 with stackToPG tree orig stack si + ... | case1 eq = {!!} + ... | case2 (case1 eq ) = {!!} + ... | case2 (case2 pg) with PG.pg pg + ... | s2-s1p2 x x₁ = rotateRight {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit + -- x : PG.parent pg ≡ node kp vp tree n1 + -- x₁ : PG.grand pg ≡ node kg vg (PG.parent pg) (PG.uncle pg) + ... | s2-1sp2 x x₁ = rotateLeft {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit + ... | s2-s12p x x₁ = rotateLeft {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit + ... | s2-1s2p x x₁ = rotateLeft {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} next exit + -- = insertCase2 tree (PG.parent pg) (PG.uncle pg) (PG.grand pg) stack si (PG.pg pg) + + +-- if we have replacedNode on RBTree, we have RBtreeInvariant + +replaceRBP : {n m : Level} {A : Set n} {t : Set m} + → (key : ℕ) → (value : A) → {d0 : ℕ} + → (orig tree repl : bt (Color ∧ A)) {d0 : ℕ} + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant repl dr + → (stack : List (bt (Color ∧ A))) → (si : stackInvariant key tree orig stack ) + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key tree) repl + → (next : {key2 d2 : ℕ} {c2 : Color} -- rotating case + → (to tr rot rr : bt (Color ∧ A)) + → RBtreeInvariant orig d0 + → {d : ℕ} → RBtreeInvariant tree d → {dr : ℕ} → RBtreeInvariant rr dr + → (stack1 : List (bt (Color ∧ A))) → stackInvariant key tr to stack1 + → rotatedTree tr rot + → {c : Color} → replacedTree key ⟪ c , value ⟫ (child-replaced key rot) rr + → length stack1 < length stack → t ) + → (exit : {k1 d1 : ℕ} {c1 : Color} → (rot repl : bt (Color ∧ A) ) + → {d : ℕ} → RBtreeInvariant repl d + → (ri : rotatedTree orig rot ) → {c : Color} → replacedTree key ⟪ c , value ⟫ rot repl → t ) → t +replaceRBP {n} {m} {A} {t} key value orig tree repl rbio rbit rbir stack si ri next exit = insertCase1 where + insertCase2 : (tree parent uncle grand : bt (Color ∧ A)) + → (stack : List (bt (Color ∧ A))) → (si : stackInvariant key tree orig stack ) + → (pg : ParentGrand tree parent uncle grand ) → t + insertCase2 tree leaf uncle grand stack si (s2-s1p2 () x₁) + insertCase2 tree leaf uncle grand stack si (s2-1sp2 () x₁) + insertCase2 tree leaf uncle grand stack si (s2-s12p () x₁) + insertCase2 tree leaf uncle grand stack si (s2-1s2p () x₁) + insertCase2 tree parent@(node kp ⟪ Red , _ ⟫ _ _) uncle grand stack si pg = next {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} + insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) leaf grand stack si pg = {!!} + insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Red , _ ⟫ _ _ ) grand stack si pg = next {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} {!!} + insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-s1p2 x x₁) = + insertCase5 key value orig tree {!!} repl rbio {!!} {!!} stack si {!!} ri {!!} {!!} next exit + -- tree is left of parent, parent is left of grand + -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp tree n1 + -- grand ≡ node kg vg (node kp ⟪ Black , proj3 ⟫ left right) (node ku ⟪ Black , proj4 ⟫ left₁ right₁) + insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-1sp2 x x₁) = + rotateLeft key value orig tree {!!} repl rbio {!!} {!!} stack si {!!} ri {!!} {!!} + (λ a b c d e f h i j k l m → insertCase5 key value a b c d {!!} {!!} h i j k l {!!} {!!} next exit ) exit + -- tree is right of parent, parent is left of grand rotateLeft + -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp n1 tree + -- grand ≡ node kg vg (node kp ⟪ Black , proj3 ⟫ left right) (node ku ⟪ Black , proj4 ⟫ left₁ right₁) + insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-s12p x x₁) = + rotateRight key value orig tree {!!} repl rbio {!!} {!!} stack si {!!} ri {!!} {!!} + (λ a b c d e f h i j k l m → insertCase5 key value a b c d {!!} {!!} h i j k l {!!} {!!} next exit ) exit + -- tree is left of parent, parent is right of grand, rotateRight + -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp tree n1 + -- grand ≡ node kg vg (node ku ⟪ Black , proj4 ⟫ left₁ right₁) (node kp ⟪ Black , proj3 ⟫ left right) + insertCase2 tree parent@(node kp ⟪ Black , _ ⟫ _ _) (node ku ⟪ Black , _ ⟫ _ _) grand stack si (s2-1s2p x x₁) = + insertCase5 key value orig tree {!!} repl rbio {!!} {!!} stack si {!!} ri {!!} {!!} next exit + -- tree is right of parent, parent is right of grand + -- node kp ⟪ Black , proj3 ⟫ left right ≡ node kp₁ vp n1 tree + -- grand ≡ node kg vg (node ku ⟪ Black , proj4 ⟫ left₁ right₁) (node kp ⟪ Black , proj3 ⟫ left right) + insertCase1 : t + insertCase1 with stackToPG tree orig stack si + ... | case1 eq = {!!} + ... | case2 (case1 eq ) = {!!} + ... | case2 (case2 pg) = insertCase2 tree (PG.parent pg) (PG.uncle pg) (PG.grand pg) stack si (PG.pg pg) + +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/logic.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,92 @@ +module logic where + +open import Level + +open import Relation.Nullary +open import Relation.Binary hiding (_⇔_) +open import Relation.Binary.PropositionalEquality + +open import Data.Empty +open import Data.Nat hiding (_⊔_) + + +data Bool : Set where + true : Bool + false : Bool + +record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where + constructor ⟪_,_⟫ + field + proj1 : A + proj2 : B + +data _∨_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where + case1 : A → A ∨ B + case2 : B → A ∨ B + +_⇔_ : {n m : Level } → ( A : Set n ) ( B : Set m ) → Set (n ⊔ m) +_⇔_ A B = ( A → B ) ∧ ( B → A ) + +contra-position : {n m : Level } {A : Set n} {B : Set m} → (A → B) → ¬ B → ¬ A +contra-position {n} {m} {A} {B} f ¬b a = ¬b ( f a ) + +double-neg : {n : Level } {A : Set n} → A → ¬ ¬ A +double-neg A notnot = notnot A + +double-neg2 : {n : Level } {A : Set n} → ¬ ¬ ¬ A → ¬ A +double-neg2 notnot A = notnot ( double-neg A ) + +de-morgan : {n : Level } {A B : Set n} → A ∧ B → ¬ ( (¬ A ) ∨ (¬ B ) ) +de-morgan {n} {A} {B} and (case1 ¬A) = ⊥-elim ( ¬A ( _∧_.proj1 and )) +de-morgan {n} {A} {B} and (case2 ¬B) = ⊥-elim ( ¬B ( _∧_.proj2 and )) + +dont-or : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → ¬ A → B +dont-or {A} {B} (case1 a) ¬A = ⊥-elim ( ¬A a ) +dont-or {A} {B} (case2 b) ¬A = b + +dont-orb : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → ¬ B → A +dont-orb {A} {B} (case2 b) ¬B = ⊥-elim ( ¬B b ) +dont-orb {A} {B} (case1 a) ¬B = a + + +infixr 130 _∧_ +infixr 140 _∨_ +infixr 150 _⇔_ + +_/\_ : Bool → Bool → Bool +true /\ true = true +_ /\ _ = false + +_<B?_ : ℕ → ℕ → Bool +ℕ.zero <B? x = true +ℕ.suc x <B? ℕ.zero = false +ℕ.suc x <B? ℕ.suc xx = x <B? xx + +-- _<BT_ : ℕ → ℕ → Set +-- ℕ.zero <BT ℕ.zero = ⊤ +-- ℕ.zero <BT ℕ.suc b = ⊤ +-- ℕ.suc a <BT ℕ.zero = ⊥ +-- ℕ.suc a <BT ℕ.suc b = a <BT b + + +_≟B_ : Decidable {A = Bool} _≡_ +true ≟B true = yes refl +false ≟B false = yes refl +true ≟B false = no λ() +false ≟B true = no λ() + +_\/_ : Bool → Bool → Bool +false \/ false = false +_ \/ _ = true + +not_ : Bool → Bool +not true = false +not false = true + +_<=>_ : Bool → Bool → Bool +true <=> true = true +false <=> false = true +_ <=> _ = false + +infixr 130 _\/_ +infixr 140 _/\_
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/queue.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,122 @@ +open import Level renaming (suc to succ ; zero to Zero ) +module Queue where + +open import Relation.Binary.PropositionalEquality +open import Relation.Binary.Core +open import Data.Nat + +data Maybe {n : Level } (a : Set n) : Set n where + Nothing : Maybe a + Just : a -> Maybe a + +data Bool {n : Level }: Set n where + True : Bool + False : Bool + + +record QueueMethods {n m : Level} (a : Set n) {t : Set m} (queueImpl : Set n) : Set (m Level.⊔ n) where + field + put : queueImpl -> a -> (queueImpl -> t) -> t + take : queueImpl -> (queueImpl -> Maybe a -> t) -> t + clear : queueImpl -> (queueImpl -> t) -> t +open QueueMethods + +record Queue {n m : Level} (a : Set n) {t : Set m} (qu : Set n) : Set (m Level.⊔ n) where + field + queue : qu + queueMethods : QueueMethods {n} {m} a {t} qu + putQueue : a -> (Queue a qu -> t) -> t + putQueue a next = put (queueMethods) (queue) a (\q1 -> next record {queue = q1 ; queueMethods = queueMethods}) + takeQueue : (Queue a qu -> Maybe a -> t) -> t + takeQueue next = take (queueMethods) (queue) (\q1 d1 -> next record {queue = q1 ; queueMethods = queueMethods} d1) + clearQueue : (Queue a qu -> t) -> t + clearQueue next = clear (queueMethods) (queue) (\q1 -> next record {queue = q1 ; queueMethods = queueMethods}) +open Queue + + + +record Element {n : Level} (a : Set n) : Set n where + inductive + constructor cons + field + datum : a -- `data` is reserved by Agda. + next : Maybe (Element a) +open Element + + +record SingleLinkedQueue {n : Level} (a : Set n) : Set n where + field + top : Maybe (Element a) + last : Maybe (Element a) +open SingleLinkedQueue + + +{-# TERMINATING #-} +reverseElement : {n : Level} {a : Set n} -> Element a -> Maybe (Element a) -> Element a +reverseElement el Nothing with next el +... | Just el1 = reverseElement el1 (Just rev) + where + rev = cons (datum el) Nothing +... | Nothing = el +reverseElement el (Just el0) with next el +... | Just el1 = reverseElement el1 (Just (cons (datum el) (Just el0))) +... | Nothing = (cons (datum el) (Just el0)) + + +{-# TERMINATING #-} +putSingleLinkedQueue : {n m : Level} {t : Set m} {a : Set n} -> SingleLinkedQueue a -> a -> (Code : SingleLinkedQueue a -> t) -> t +putSingleLinkedQueue queue a cs with top queue +... | Just te = cs queue1 + where + re = reverseElement te Nothing + ce = cons a (Just re) + commit = reverseElement ce Nothing + queue1 = record queue {top = Just commit} +... | Nothing = cs queue1 + where + cel = record {datum = a ; next = Nothing} + queue1 = record {top = Just cel ; last = Just cel} + +{-# TERMINATING #-} +takeSingleLinkedQueue : {n m : Level} {t : Set m} {a : Set n} -> SingleLinkedQueue a -> (Code : SingleLinkedQueue a -> (Maybe a) -> t) -> t +takeSingleLinkedQueue queue cs with (top queue) +... | Just te = cs record {top = (next te) ; last = Just (lastElement te)} (Just (datum te)) + where + lastElement : {n : Level} {a : Set n} -> Element a -> Element a + lastElement el with next el + ... | Just el1 = lastElement el1 + ... | Nothing = el +... | Nothing = cs queue Nothing + +clearSingleLinkedQueue : {n m : Level} {t : Set m} {a : Set n} -> SingleLinkedQueue a -> (SingleLinkedQueue a -> t) -> t +clearSingleLinkedQueue queue cs = cs (record {top = Nothing ; last = Nothing}) + + +emptySingleLinkedQueue : {n : Level} {a : Set n} -> SingleLinkedQueue a +emptySingleLinkedQueue = record {top = Nothing ; last = Nothing} + +singleLinkedQueueSpec : {n m : Level } {t : Set m } {a : Set n} -> QueueMethods {n} {m} a {t} (SingleLinkedQueue a) +singleLinkedQueueSpec = record { + put = putSingleLinkedQueue + ; take = takeSingleLinkedQueue + ; clear = clearSingleLinkedQueue + } + + +createSingleLinkedQueue : {n m : Level} {t : Set m} {a : Set n} -> Queue {n} {m} a {t} (SingleLinkedQueue a) +createSingleLinkedQueue = record { + queue = emptySingleLinkedQueue ; + queueMethods = singleLinkedQueueSpec + } + + +check1 = putSingleLinkedQueue emptySingleLinkedQueue 1 (\q1 -> putSingleLinkedQueue q1 2 (\q2 -> putSingleLinkedQueue q2 3 (\q3 -> putSingleLinkedQueue q3 4 (\q4 -> putSingleLinkedQueue q4 5 (\q5 -> q5))))) + + +check2 = putSingleLinkedQueue emptySingleLinkedQueue 1 (\q1 -> putSingleLinkedQueue q1 2 (\q2 -> putSingleLinkedQueue q2 3 (\q3 -> putSingleLinkedQueue q3 4 (\q4 -> takeSingleLinkedQueue q4 (\q d -> q))))) + + +test01 : {n : Level } {a : Set n} -> SingleLinkedQueue a -> Maybe a -> Bool {n} +test01 queue _ with (top queue) +... | (Just _) = True +... | Nothing = False
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/redBlackTreeHoare.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,292 @@ +module redBlackTreeHoare where + + +open import Level hiding (zero) + +open import Data.Nat hiding (compare) +open import Data.Nat.Properties as NatProp +open import Data.Maybe +open import Data.Bool hiding ( _<_ ) +open import Data.Empty + +open import Relation.Binary +open import Relation.Binary.PropositionalEquality + +open import stack + +record TreeMethods {n m : Level } {a : Set n } {t : Set m } (treeImpl : Set n ) : Set (m Level.⊔ n) where + field + putImpl : treeImpl → a → (treeImpl → t) → t + getImpl : treeImpl → (treeImpl → Maybe a → t) → t +open TreeMethods + +record Tree {n m : Level } {a : Set n } {t : Set m } (treeImpl : Set n ) : Set (m Level.⊔ n) where + field + tree : treeImpl + treeMethods : TreeMethods {n} {m} {a} {t} treeImpl + putTree : a → (Tree treeImpl → t) → t + putTree d next = putImpl (treeMethods ) tree d (\t1 → next (record {tree = t1 ; treeMethods = treeMethods} )) + getTree : (Tree treeImpl → Maybe a → t) → t + getTree next = getImpl (treeMethods ) tree (\t1 d → next (record {tree = t1 ; treeMethods = treeMethods} ) d ) + +open Tree + +data Color {n : Level } : Set n where + Red : Color + Black : Color + + +record Node {n : Level } (a : Set n) (k : ℕ) : Set n where + inductive + field + key : ℕ + value : a + right : Maybe (Node a k) + left : Maybe (Node a k) + color : Color {n} +open Node + +record RedBlackTree {n m : Level } {t : Set m} (a : Set n) (k : ℕ) : Set (m Level.⊔ n) where + field + root : Maybe (Node a k) + nodeStack : SingleLinkedStack (Node a k) + -- compare : k → k → Tri A B C + -- <-cmp + +open RedBlackTree + +open SingleLinkedStack + +compTri : ( x y : ℕ ) -> Tri ( x < y ) ( x ≡ y ) ( x > y ) +compTri = IsStrictTotalOrder.compare (Relation.Binary.StrictTotalOrder.isStrictTotalOrder <-strictTotalOrder) + where open import Relation.Binary + +-- put new node at parent node, and rebuild tree to the top +-- +{-# TERMINATING #-} -- https://agda.readthedocs.io/en/v2.5.3/language/termination-checking.html +replaceNode : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Node a k → (RedBlackTree {n} {m} {t} a k → t) → t +replaceNode {n} {m} {t} {a} {k} tree s n0 next = popSingleLinkedStack s ( + \s parent → replaceNode1 s parent) + module ReplaceNode where + replaceNode1 : SingleLinkedStack (Node a k) → Maybe ( Node a k ) → t + replaceNode1 s nothing = next ( record tree { root = just (record n0 { color = Black}) } ) + replaceNode1 s (just n1) with compTri (key n1) (key n0) + replaceNode1 s (just n1) | tri< lt ¬eq ¬gt = replaceNode {n} {m} {t} {a} {k} tree s ( record n1 { value = value n0 ; left = left n0 ; right = right n0 } ) next + replaceNode1 s (just n1) | tri≈ ¬lt eq ¬gt = replaceNode {n} {m} {t} {a} {k} tree s ( record n1 { left = just n0 } ) next + replaceNode1 s (just n1) | tri> ¬lt ¬eq gt = replaceNode {n} {m} {t} {a} {k} tree s ( record n1 { right = just n0 } ) next + + +rotateRight : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → + (RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → t) → t +rotateRight {n} {m} {t} {a} {k} tree s n0 parent rotateNext = getSingleLinkedStack s (\ s n0 → rotateRight1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext) + where + rotateRight1 : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → + (RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → t) → t + rotateRight1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext with n0 + ... | nothing = rotateNext tree s nothing n0 + ... | just n1 with parent + ... | nothing = rotateNext tree s (just n1 ) n0 + ... | just parent1 with left parent1 + ... | nothing = rotateNext tree s (just n1) nothing + ... | just leftParent with compTri (key n1) (key leftParent) + rotateRight1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext | just n1 | just parent1 | just leftParent | tri< a₁ ¬b ¬c = rotateNext tree s (just n1) parent + rotateRight1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext | just n1 | just parent1 | just leftParent | tri≈ ¬a b ¬c = rotateNext tree s (just n1) parent + rotateRight1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext | just n1 | just parent1 | just leftParent | tri> ¬a ¬b c = rotateNext tree s (just n1) parent + + +rotateLeft : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → + (RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → t) → t +rotateLeft {n} {m} {t} {a} {k} tree s n0 parent rotateNext = getSingleLinkedStack s (\ s n0 → rotateLeft1 tree s n0 parent rotateNext) + where + rotateLeft1 : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → + (RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → t) → t + rotateLeft1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext with n0 + ... | nothing = rotateNext tree s nothing n0 + ... | just n1 with parent + ... | nothing = rotateNext tree s (just n1) nothing + ... | just parent1 with right parent1 + ... | nothing = rotateNext tree s (just n1) nothing + ... | just rightParent with compTri (key n1) (key rightParent) + rotateLeft1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext | just n1 | just parent1 | just rightParent | tri< a₁ ¬b ¬c = rotateNext tree s (just n1) parent + rotateLeft1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext | just n1 | just parent1 | just rightParent | tri≈ ¬a b ¬c = rotateNext tree s (just n1) parent + rotateLeft1 {n} {m} {t} {a} {k} tree s n0 parent rotateNext | just n1 | just parent1 | just rightParent | tri> ¬a ¬b c = rotateNext tree s (just n1) parent + -- ... | EQ = rotateNext tree s (just n1) parent + -- ... | _ = rotateNext tree s (just n1) parent + +{-# TERMINATING #-} +insertCase5 : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Node a k → Node a k → (RedBlackTree {n} {m} {t} a k → t) → t +insertCase5 {n} {m} {t} {a} {k} tree s n0 parent grandParent next = pop2SingleLinkedStack s (\ s parent grandParent → insertCase51 tree s n0 parent grandParent next) + where + insertCase51 : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → Maybe (Node a k) → (RedBlackTree {n} {m} {t} a k → t) → t + insertCase51 {n} {m} {t} {a} {k} tree s n0 parent grandParent next with n0 + ... | nothing = next tree + ... | just n1 with parent | grandParent + ... | nothing | _ = next tree + ... | _ | nothing = next tree + ... | just parent1 | just grandParent1 with left parent1 | left grandParent1 + ... | nothing | _ = next tree + ... | _ | nothing = next tree + ... | just leftParent1 | just leftGrandParent1 + with compTri (key n1) (key leftParent1) | compTri (key leftParent1) (key leftGrandParent1) + ... | tri≈ ¬a b ¬c | tri≈ ¬a1 b1 ¬c1 = rotateRight tree s n0 parent (\ tree s n0 parent → insertCase5 tree s n0 parent1 grandParent1 next) + ... | _ | _ = rotateLeft tree s n0 parent (\ tree s n0 parent → insertCase5 tree s n0 parent1 grandParent1 next) + -- ... | EQ | EQ = rotateRight tree s n0 parent (\ tree s n0 parent → insertCase5 tree s n0 parent1 grandParent1 next) + -- ... | _ | _ = rotateLeft tree s n0 parent (\ tree s n0 parent → insertCase5 tree s n0 parent1 grandParent1 next) + +insertCase4 : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Node a k → Node a k → Node a k → (RedBlackTree {n} {m} {t} a k → t) → t +insertCase4 {n} {m} {t} {a} {k} tree s n0 parent grandParent next + with (right parent) | (left grandParent) +... | nothing | _ = insertCase5 tree s (just n0) parent grandParent next +... | _ | nothing = insertCase5 tree s (just n0) parent grandParent next +... | just rightParent | just leftGrandParent with compTri (key n0) (key rightParent) | compTri (key parent) (key leftGrandParent) -- (key n0) (key rightParent) | (key parent) (key leftGrandParent) +-- ... | EQ | EQ = popSingleLinkedStack s (\ s n1 → rotateLeft tree s (left n0) (just grandParent) +-- (\ tree s n0 parent → insertCase5 tree s n0 rightParent grandParent next)) +-- ... | _ | _ = insertCase41 tree s n0 parent grandParent next +... | tri≈ ¬a b ¬c | tri≈ ¬a1 b1 ¬c1 = popSingleLinkedStack s (\ s n1 → rotateLeft tree s (left n0) (just grandParent) (\ tree s n0 parent → insertCase5 tree s n0 rightParent grandParent next)) +... | _ | _ = insertCase41 tree s n0 parent grandParent next + where + insertCase41 : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Node a k → Node a k → Node a k → (RedBlackTree {n} {m} {t} a k → t) → t + insertCase41 {n} {m} {t} {a} {k} tree s n0 parent grandParent next + with (left parent) | (right grandParent) + ... | nothing | _ = insertCase5 tree s (just n0) parent grandParent next + ... | _ | nothing = insertCase5 tree s (just n0) parent grandParent next + ... | just leftParent | just rightGrandParent with compTri (key n0) (key leftParent) | compTri (key parent) (key rightGrandParent) + ... | tri≈ ¬a b ¬c | tri≈ ¬a1 b1 ¬c1 = popSingleLinkedStack s (\ s n1 → rotateRight tree s (right n0) (just grandParent) (\ tree s n0 parent → insertCase5 tree s n0 leftParent grandParent next)) + ... | _ | _ = insertCase5 tree s (just n0) parent grandParent next + -- ... | EQ | EQ = popSingleLinkedStack s (\ s n1 → rotateRight tree s (right n0) (just grandParent) + -- (\ tree s n0 parent → insertCase5 tree s n0 leftParent grandParent next)) + -- ... | _ | _ = insertCase5 tree s (just n0) parent grandParent next + +colorNode : {n : Level } {a : Set n} {k : ℕ} → Node a k → Color → Node a k +colorNode old c = record old { color = c } + +{-# TERMINATING #-} +insertNode : {n m : Level } {t : Set m } {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Node a k → (RedBlackTree {n} {m} {t} a k → t) → t +insertNode {n} {m} {t} {a} {k} tree s n0 next = get2SingleLinkedStack s (insertCase1 n0) + where + insertCase1 : Node a k → SingleLinkedStack (Node a k) → Maybe (Node a k) → Maybe (Node a k) → t -- placed here to allow mutual recursion + -- http://agda.readthedocs.io/en/v2.5.2/language/mutual-recursion.html + insertCase3 : SingleLinkedStack (Node a k) → Node a k → Node a k → Node a k → t + insertCase3 s n0 parent grandParent with left grandParent | right grandParent + ... | nothing | nothing = insertCase4 tree s n0 parent grandParent next + ... | nothing | just uncle = insertCase4 tree s n0 parent grandParent next + ... | just uncle | _ with compTri ( key uncle ) ( key parent ) + insertCase3 s n0 parent grandParent | just uncle | _ | tri≈ ¬a b ¬c = insertCase4 tree s n0 parent grandParent next + insertCase3 s n0 parent grandParent | just uncle | _ | tri< a ¬b ¬c with color uncle + insertCase3 s n0 parent grandParent | just uncle | _ | tri< a ¬b ¬c | Red = pop2SingleLinkedStack s ( \s p0 p1 → insertCase1 ( + record grandParent { color = Red ; left = just ( record parent { color = Black } ) ; right = just ( record uncle { color = Black } ) }) s p0 p1 ) + insertCase3 s n0 parent grandParent | just uncle | _ | tri< a ¬b ¬c | Black = insertCase4 tree s n0 parent grandParent next + insertCase3 s n0 parent grandParent | just uncle | _ | tri> ¬a ¬b c with color uncle + insertCase3 s n0 parent grandParent | just uncle | _ | tri> ¬a ¬b c | Red = pop2SingleLinkedStack s ( \s p0 p1 → insertCase1 ( record grandParent { color = Red ; left = just ( record parent { color = Black } ) ; right = just ( record uncle { color = Black } ) }) s p0 p1 ) + insertCase3 s n0 parent grandParent | just uncle | _ | tri> ¬a ¬b c | Black = insertCase4 tree s n0 parent grandParent next + -- ... | EQ = insertCase4 tree s n0 parent grandParent next + -- ... | _ with color uncle + -- ... | Red = pop2SingleLinkedStack s ( \s p0 p1 → insertCase1 ( + -- record grandParent { color = Red ; left = just ( record parent { color = Black } ) ; right = just ( record uncle { color = Black } ) }) s p0 p1 ) + -- ... | Black = insertCase4 tree s n0 parent grandParent next --!! + insertCase2 : SingleLinkedStack (Node a k) → Node a k → Node a k → Node a k → t + insertCase2 s n0 parent grandParent with color parent + ... | Black = replaceNode tree s n0 next + ... | Red = insertCase3 s n0 parent grandParent + insertCase1 n0 s nothing nothing = next tree + insertCase1 n0 s nothing (just grandParent) = next tree + insertCase1 n0 s (just parent) nothing = replaceNode tree s (colorNode n0 Black) next + insertCase1 n0 s (just parent) (just grandParent) = insertCase2 s n0 parent grandParent + + +---- +-- find node potition to insert or to delete, the path will be in the stack +-- +findNode : {n m : Level } {a : Set n} {k : ℕ} {t : Set m} → RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → (Node a k) → (Node a k) → (RedBlackTree {n} {m} {t} a k → SingleLinkedStack (Node a k) → Node a k → t) → t +findNode {n} {m} {a} {k} {t} tree s n0 n1 next = pushSingleLinkedStack s n1 (\ s → findNode1 s n1) + module FindNode where + findNode2 : SingleLinkedStack (Node a k) → (Maybe (Node a k)) → t + findNode2 s nothing = next tree s n0 + findNode2 s (just n) = findNode tree s n0 n next + findNode1 : SingleLinkedStack (Node a k) → (Node a k) → t + findNode1 s n1 with (compTri (key n0) (key n1)) + findNode1 s n1 | tri< a ¬b ¬c = popSingleLinkedStack s ( \s _ → next tree s (record n1 { key = key n1 ; value = value n0 } ) ) + findNode1 s n1 | tri≈ ¬a b ¬c = findNode2 s (right n1) + findNode1 s n1 | tri> ¬a ¬b c = findNode2 s (left n1) + -- ... | EQ = popSingleLinkedStack s ( \s _ → next tree s (record n1 { key = key n1 ; value = value n0 } ) ) + -- ... | GT = findNode2 s (right n1) + -- ... | LT = findNode2 s (left n1) + + + + +leafNode : {n : Level } { a : Set n } → a → (k : ℕ) → (Node a k) +leafNode v k1 = record { key = k1 ; value = v ; right = nothing ; left = nothing ; color = Red } + +putRedBlackTree : {n m : Level} {t : Set m} {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} a k → {!!} → {!!} → (RedBlackTree {n} {m} {t} a k → t) → t +putRedBlackTree {n} {m} {t} {a} {k} tree val k1 next with (root tree) +putRedBlackTree {n} {m} {t} {a} {k} tree val k1 next | nothing = next (record tree {root = just (leafNode {!!} {!!}) }) +putRedBlackTree {n} {m} {t} {a} {k} tree val k1 next | just n2 = clearSingleLinkedStack (nodeStack tree) (λ s → findNode tree s (leafNode {!!} {!!}) n2 (λ tree1 s n1 → insertNode tree1 s n1 next)) +-- putRedBlackTree {n} {m} {t} {a} {k} tree value k1 next with (root tree) +-- ... | nothing = next (record tree {root = just (leafNode k1 value) }) +-- ... | just n2 = clearSingleLinkedStack (nodeStack tree) (\ s → findNode tree s (leafNode k1 value) n2 (\ tree1 s n1 → insertNode tree1 s n1 next)) + + +-- getRedBlackTree : {n m : Level } {t : Set m} {a : Set n} {k : ℕ} → RedBlackTree {n} {m} {t} {A} a k → k → (RedBlackTree {n} {m} {t} {A} a k → (Maybe (Node a k)) → t) → t +-- getRedBlackTree {_} {_} {t} {a} {k} tree k1 cs = checkNode (root tree) +-- module GetRedBlackTree where -- http://agda.readthedocs.io/en/v2.5.2/language/let-and-where.html +-- search : Node a k → t +-- checkNode : Maybe (Node a k) → t +-- checkNode nothing = cs tree nothing +-- checkNode (just n) = search n +-- search n with compTri k1 (key n) +-- search n | tri< a ¬b ¬c = checkNode (left n) +-- search n | tri≈ ¬a b ¬c = cs tree (just n) +-- search n | tri> ¬a ¬b c = checkNode (right n) + + + +-- compareT : {A B C : Set } → ℕ → ℕ → Tri A B C +-- compareT x y with IsStrictTotalOrder.compare (Relation.Binary.StrictTotalOrder.isStrictTotalOrder <-strictTotalOrder) x y +-- compareT x y | tri< a ¬b ¬c = tri< {!!} {!!} {!!} +-- compareT x y | tri≈ ¬a b ¬c = {!!} +-- compareT x y | tri> ¬a ¬b c = {!!} +-- -- ... | tri≈ a b c = {!!} +-- -- ... | tri< a b c = {!!} +-- -- ... | tri> a b c = {!!} + +-- compare2 : (x y : ℕ ) → CompareResult {Level.zero} +-- compare2 zero zero = EQ +-- compare2 (suc _) zero = GT +-- compare2 zero (suc _) = LT +-- compare2 (suc x) (suc y) = compare2 x y + +-- -- putUnblanceTree : {n m : Level } {a : Set n} {k : ℕ} {t : Set m} → RedBlackTree {n} {m} {t} {A} a k → k → a → (RedBlackTree {n} {m} {t} {A} a k → t) → t +-- -- putUnblanceTree {n} {m} {A} {a} {k} {t} tree k1 value next with (root tree) +-- -- ... | nothing = next (record tree {root = just (leafNode k1 value) }) +-- -- ... | just n2 = clearSingleLinkedStack (nodeStack tree) (λ s → findNode tree s (leafNode k1 value) n2 (λ tree1 s n1 → replaceNode tree1 s n1 next)) + +-- -- checkT : {m : Level } (n : Maybe (Node ℕ ℕ)) → ℕ → Bool +-- -- checkT nothing _ = false +-- -- checkT (just n) x with compTri (value n) x +-- -- ... | tri≈ _ _ _ = true +-- -- ... | _ = false + +-- -- checkEQ : {m : Level } ( x : ℕ ) -> ( n : Node ℕ ℕ ) -> (value n ) ≡ x -> checkT {m} (just n) x ≡ true +-- -- checkEQ x n refl with compTri (value n) x +-- -- ... | tri≈ _ refl _ = refl +-- -- ... | tri> _ neq gt = ⊥-elim (neq refl) +-- -- ... | tri< lt neq _ = ⊥-elim (neq refl) + + +createEmptyRedBlackTreeℕ : {n m : Level} {t : Set m} (a : Set n) (b : ℕ) + → RedBlackTree {n} {m} {t} a b +createEmptyRedBlackTreeℕ a b = record { + root = nothing + ; nodeStack = emptySingleLinkedStack + -- ; nodeComp = λ x x₁ → {!!} + + } + +-- ( x y : ℕ ) -> Tri ( x < y ) ( x ≡ y ) ( x > y ) + +-- test = (λ x → (createEmptyRedBlackTreeℕ x x) + +-- ts = createEmptyRedBlackTreeℕ {ℕ} {?} {!!} 0 + +-- tes = putRedBlackTree {_} {_} {_} (createEmptyRedBlackTreeℕ {_} {_} {_} 3 3) 2 2 (λ t → t)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/stack.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,155 @@ +open import Level renaming (suc to succ ; zero to Zero ) +module stack where + +open import Relation.Binary.PropositionalEquality +open import Relation.Binary.Core +open import Data.Nat hiding (compare) +open import Data.Maybe + +ex : 1 + 2 ≡ 3 +ex = refl + +-- data Bool {n : Level } : Set n where +-- True : Bool +-- False : Bool + +-- record _∧_ {n : Level } (a : Set n) (b : Set n): Set n where +-- field +-- pi1 : a +-- pi2 : b + +-- data _∨_ {n : Level } (a : Set n) (b : Set n): Set n where +-- pi1 : a -> a ∨ b +-- pi2 : b -> a ∨ b + +-- data Maybe {n : Level } (a : Set n) : Set n where +-- nothing : Maybe a +-- just : a -> Maybe a + +record StackMethods {n m : Level } (a : Set n ) {t : Set m } (stackImpl : Set n ) : Set (m Level.⊔ n) where + field + push : stackImpl -> a -> (stackImpl -> t) -> t + pop : stackImpl -> (stackImpl -> Maybe a -> t) -> t + pop2 : stackImpl -> (stackImpl -> Maybe a -> Maybe a -> t) -> t + get : stackImpl -> (stackImpl -> Maybe a -> t) -> t + get2 : stackImpl -> (stackImpl -> Maybe a -> Maybe a -> t) -> t + clear : stackImpl -> (stackImpl -> t) -> t +open StackMethods + +record Stack {n m : Level } (a : Set n ) {t : Set m } (si : Set n ) : Set (m Level.⊔ n) where + field + stack : si + stackMethods : StackMethods {n} {m} a {t} si + pushStack : a -> (Stack a si -> t) -> t + pushStack d next = push (stackMethods ) (stack ) d (\s1 -> next (record {stack = s1 ; stackMethods = stackMethods } )) + popStack : (Stack a si -> Maybe a -> t) -> t + popStack next = pop (stackMethods ) (stack ) (\s1 d1 -> next (record {stack = s1 ; stackMethods = stackMethods }) d1 ) + pop2Stack : (Stack a si -> Maybe a -> Maybe a -> t) -> t + pop2Stack next = pop2 (stackMethods ) (stack ) (\s1 d1 d2 -> next (record {stack = s1 ; stackMethods = stackMethods }) d1 d2) + getStack : (Stack a si -> Maybe a -> t) -> t + getStack next = get (stackMethods ) (stack ) (\s1 d1 -> next (record {stack = s1 ; stackMethods = stackMethods }) d1 ) + get2Stack : (Stack a si -> Maybe a -> Maybe a -> t) -> t + get2Stack next = get2 (stackMethods ) (stack ) (\s1 d1 d2 -> next (record {stack = s1 ; stackMethods = stackMethods }) d1 d2) + clearStack : (Stack a si -> t) -> t + clearStack next = clear (stackMethods ) (stack ) (\s1 -> next (record {stack = s1 ; stackMethods = stackMethods } )) + +open Stack + +{-- +data Element {n : Level } (a : Set n) : Set n where + cons : a -> Maybe (Element a) -> Element a + + +datum : {n : Level } {a : Set n} -> Element a -> a +datum (cons a _) = a + +next : {n : Level } {a : Set n} -> Element a -> Maybe (Element a) +next (cons _ n) = n +--} + + +-- cannot define recrusive record definition. so use linked list with maybe. +record Element {l : Level} (a : Set l) : Set l where + inductive + constructor cons + field + datum : a -- `data` is reserved by Agda. + next : Maybe (Element a) + +open Element + + +record SingleLinkedStack {n : Level } (a : Set n) : Set n where + field + top : Maybe (Element a) +open SingleLinkedStack + +pushSingleLinkedStack : {n m : Level } {t : Set m } {Data : Set n} -> SingleLinkedStack Data -> Data -> (Code : SingleLinkedStack Data -> t) -> t +pushSingleLinkedStack stack datum next = next stack1 + where + element = cons datum (top stack) + stack1 = record {top = just element} + + +popSingleLinkedStack : {n m : Level } {t : Set m } {a : Set n} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> t) -> t +popSingleLinkedStack stack cs with (top stack) +... | nothing = cs stack nothing +... | just d = cs stack1 (just data1) + where + data1 = datum d + stack1 = record { top = (next d) } + +pop2SingleLinkedStack : {n m : Level } {t : Set m } {a : Set n} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> (Maybe a) -> t) -> t +pop2SingleLinkedStack {n} {m} {t} {a} stack cs with (top stack) +... | nothing = cs stack nothing nothing +... | just d = pop2SingleLinkedStack' {n} {m} stack cs + where + pop2SingleLinkedStack' : {n m : Level } {t : Set m } -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> (Maybe a) -> t) -> t + pop2SingleLinkedStack' stack cs with (next d) + ... | nothing = cs stack nothing nothing + ... | just d1 = cs (record {top = (next d1)}) (just (datum d)) (just (datum d1)) + + +getSingleLinkedStack : {n m : Level } {t : Set m } {a : Set n} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> t) -> t +getSingleLinkedStack stack cs with (top stack) +... | nothing = cs stack nothing +... | just d = cs stack (just data1) + where + data1 = datum d + +get2SingleLinkedStack : {n m : Level } {t : Set m } {a : Set n} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> (Maybe a) -> t) -> t +get2SingleLinkedStack {n} {m} {t} {a} stack cs with (top stack) +... | nothing = cs stack nothing nothing +... | just d = get2SingleLinkedStack' {n} {m} stack cs + where + get2SingleLinkedStack' : {n m : Level} {t : Set m } -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> (Maybe a) -> t) -> t + get2SingleLinkedStack' stack cs with (next d) + ... | nothing = cs stack nothing nothing + ... | just d1 = cs stack (just (datum d)) (just (datum d1)) + +clearSingleLinkedStack : {n m : Level } {t : Set m } {a : Set n} -> SingleLinkedStack a -> (SingleLinkedStack a -> t) -> t +clearSingleLinkedStack stack next = next (record {top = nothing}) + + +emptySingleLinkedStack : {n : Level } {a : Set n} -> SingleLinkedStack a +emptySingleLinkedStack = record {top = nothing} + +----- +-- Basic stack implementations are specifications of a Stack +-- +singleLinkedStackSpec : {n m : Level } {t : Set m } {a : Set n} -> StackMethods {n} {m} a {t} (SingleLinkedStack a) +singleLinkedStackSpec = record { + push = pushSingleLinkedStack + ; pop = popSingleLinkedStack + ; pop2 = pop2SingleLinkedStack + ; get = getSingleLinkedStack + ; get2 = get2SingleLinkedStack + ; clear = clearSingleLinkedStack + } + +createSingleLinkedStack : {n m : Level } {t : Set m } {a : Set n} -> Stack {n} {m} a {t} (SingleLinkedStack a) +createSingleLinkedStack = record { + stack = emptySingleLinkedStack ; + stackMethods = singleLinkedStackSpec + } +
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/stackTest.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,144 @@ +open import Level renaming (suc to succ ; zero to Zero ) +module stackTest where + +open import stack + +open import Relation.Binary.PropositionalEquality +open import Relation.Binary.Core +open import Data.Nat +open import Function + + +open SingleLinkedStack +open Stack + +---- +-- +-- proof of properties ( concrete cases ) +-- + +test01 : {n : Level } {a : Set n} -> SingleLinkedStack a -> Maybe a -> Bool {n} +test01 stack _ with (top stack) +... | (Just _) = True +... | Nothing = False + + +test02 : {n : Level } {a : Set n} -> SingleLinkedStack a -> Bool +test02 stack = popSingleLinkedStack stack test01 + +test03 : {n : Level } {a : Set n} -> a -> Bool +test03 v = pushSingleLinkedStack emptySingleLinkedStack v test02 + +-- after a push and a pop, the stack is empty +lemma : {n : Level} {A : Set n} {a : A} -> test03 a ≡ False +lemma = refl + +testStack01 : {n m : Level } {a : Set n} -> a -> Bool {m} +testStack01 v = pushStack createSingleLinkedStack v ( + \s -> popStack s (\s1 d1 -> True)) + +-- after push 1 and 2, pop2 get 1 and 2 + +testStack02 : {m : Level } -> ( Stack ℕ (SingleLinkedStack ℕ) -> Bool {m} ) -> Bool {m} +testStack02 cs = pushStack createSingleLinkedStack 1 ( + \s -> pushStack s 2 cs) + + +testStack031 : (d1 d2 : ℕ ) -> Bool {Zero} +testStack031 2 1 = True +testStack031 _ _ = False + +testStack032 : (d1 d2 : Maybe ℕ) -> Bool {Zero} +testStack032 (Just d1) (Just d2) = testStack031 d1 d2 +testStack032 _ _ = False + +testStack03 : {m : Level } -> Stack ℕ (SingleLinkedStack ℕ) -> ((Maybe ℕ) -> (Maybe ℕ) -> Bool {m} ) -> Bool {m} +testStack03 s cs = pop2Stack s ( + \s d1 d2 -> cs d1 d2 ) + +testStack04 : Bool +testStack04 = testStack02 (\s -> testStack03 s testStack032) + +testStack05 : testStack04 ≡ True +testStack05 = refl + +testStack06 : {m : Level } -> Maybe (Element ℕ) +testStack06 = pushStack createSingleLinkedStack 1 ( + \s -> pushStack s 2 (\s -> top (stack s))) + +testStack07 : {m : Level } -> Maybe (Element ℕ) +testStack07 = pushSingleLinkedStack emptySingleLinkedStack 1 ( + \s -> pushSingleLinkedStack s 2 (\s -> top s)) + +testStack08 = pushSingleLinkedStack emptySingleLinkedStack 1 + $ \s -> pushSingleLinkedStack s 2 + $ \s -> pushSingleLinkedStack s 3 + $ \s -> pushSingleLinkedStack s 4 + $ \s -> pushSingleLinkedStack s 5 + $ \s -> top s + +------ +-- +-- proof of properties with indefinite state of stack +-- +-- this should be proved by properties of the stack inteface, not only by the implementation, +-- and the implementation have to provides the properties. +-- +-- we cannot write "s ≡ s3", since level of the Set does not fit , but use stack s ≡ stack s3 is ok. +-- anyway some implementations may result s != s3 +-- + +stackInSomeState : {l m : Level } {D : Set l} {t : Set m } (s : SingleLinkedStack D ) -> Stack {l} {m} D {t} ( SingleLinkedStack D ) +stackInSomeState s = record { stack = s ; stackMethods = singleLinkedStackSpec } + +push->push->pop2 : {l : Level } {D : Set l} (x y : D ) (s : SingleLinkedStack D ) -> + pushStack ( stackInSomeState s ) x ( \s1 -> pushStack s1 y ( \s2 -> pop2Stack s2 ( \s3 y1 x1 -> (Just x ≡ x1 ) ∧ (Just y ≡ y1 ) ) )) +push->push->pop2 {l} {D} x y s = record { pi1 = refl ; pi2 = refl } + + +-- id : {n : Level} {A : Set n} -> A -> A +-- id a = a + +-- push a, n times + +n-push : {n : Level} {A : Set n} {a : A} -> ℕ -> SingleLinkedStack A -> SingleLinkedStack A +n-push zero s = s +n-push {l} {A} {a} (suc n) s = pushSingleLinkedStack (n-push {l} {A} {a} n s) a (\s -> s ) + +n-pop : {n : Level}{A : Set n} {a : A} -> ℕ -> SingleLinkedStack A -> SingleLinkedStack A +n-pop zero s = s +n-pop {_} {A} {a} (suc n) s = popSingleLinkedStack (n-pop {_} {A} {a} n s) (\s _ -> s ) + +open ≡-Reasoning + +push-pop-equiv : {n : Level} {A : Set n} {a : A} (s : SingleLinkedStack A) -> (popSingleLinkedStack (pushSingleLinkedStack s a (\s -> s)) (\s _ -> s) ) ≡ s +push-pop-equiv s = refl + +push-and-n-pop : {n : Level} {A : Set n} {a : A} (n : ℕ) (s : SingleLinkedStack A) -> n-pop {_} {A} {a} (suc n) (pushSingleLinkedStack s a id) ≡ n-pop {_} {A} {a} n s +push-and-n-pop zero s = refl +push-and-n-pop {_} {A} {a} (suc n) s = begin + n-pop {_} {A} {a} (suc (suc n)) (pushSingleLinkedStack s a id) + ≡⟨ refl ⟩ + popSingleLinkedStack (n-pop {_} {A} {a} (suc n) (pushSingleLinkedStack s a id)) (\s _ -> s) + ≡⟨ cong (\s -> popSingleLinkedStack s (\s _ -> s )) (push-and-n-pop n s) ⟩ + popSingleLinkedStack (n-pop {_} {A} {a} n s) (\s _ -> s) + ≡⟨ refl ⟩ + n-pop {_} {A} {a} (suc n) s + ∎ + + +n-push-pop-equiv : {n : Level} {A : Set n} {a : A} (n : ℕ) (s : SingleLinkedStack A) -> (n-pop {_} {A} {a} n (n-push {_} {A} {a} n s)) ≡ s +n-push-pop-equiv zero s = refl +n-push-pop-equiv {_} {A} {a} (suc n) s = begin + n-pop {_} {A} {a} (suc n) (n-push (suc n) s) + ≡⟨ refl ⟩ + n-pop {_} {A} {a} (suc n) (pushSingleLinkedStack (n-push n s) a (\s -> s)) + ≡⟨ push-and-n-pop n (n-push n s) ⟩ + n-pop {_} {A} {a} n (n-push n s) + ≡⟨ n-push-pop-equiv n s ⟩ + s + ∎ + + +n-push-pop-equiv-empty : {n : Level} {A : Set n} {a : A} -> (n : ℕ) -> n-pop {_} {A} {a} n (n-push {_} {A} {a} n emptySingleLinkedStack) ≡ emptySingleLinkedStack +n-push-pop-equiv-empty n = n-push-pop-equiv n emptySingleLinkedStack
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/work.agda Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,365 @@ +module work 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 + +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 leaf = nothing +node-key (node key value tree tree₁) = just key + +node-value : {n : Level} {A : Set n} → bt A → Maybe A +node-value leaf = nothing +node-value (node key value tree tree₁) = just value + +bt-depth : {n : Level} {A : Set n} → (tree : bt A) → ℕ +bt-depth leaf = 0 +bt-depth (node key value tree tree₁) = suc (bt-depth tree ⊔ bt-depth tree₁) +--一番下のleaf =0から戻るたびにsucをしていく +treeTest1 : bt ℕ +treeTest1 = node 0 0 leaf (node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf)) + +-- 0 0 +-- / \ +-- leaf 3 1 +-- / \ +-- 2 5 2 +-- / \ +-- 1 leaf 3 +-- / \ +-- leaf leaf 4 + +treeTest2 : bt ℕ +treeTest2 = node 3 1 (node 2 5 (node 1 7 leaf leaf ) leaf) (node 5 5 leaf leaf) + +testdb : ℕ +testdb = bt-depth treeTest1 -- 4 + +import Data.Unit hiding ( _≟_ ; _≤?_ ; _≤_) + +data treeInvariant {n : Level} {A : Set n} : (tree : bt A) → Set n where + t-leaf : treeInvariant leaf + + t-single : (key : ℕ) → (value : A) → treeInvariant (node key value leaf leaf) + + t-left : {key key1 : ℕ} → {value value1 : A} → {t1 t2 : bt A} → key < key1 + → treeInvariant (node key value t1 t2) + → treeInvariant (node key1 value1 (node key value t1 t2) leaf) + + t-right : {key key1 : ℕ} → {value value1 : A} → {t1 t2 : bt A} → key < key1 + → treeInvariant (node key1 value1 t1 t2) + → treeInvariant (node key value leaf (node key1 value1 t1 t2)) + + t-node : {key key1 key2 : ℕ}→ {value value1 value2 : A} → {t1 t2 t3 t4 : bt A} → key1 < key → key < key2 + → treeInvariant (node key1 value1 t1 t2) + → treeInvariant (node key2 value2 t3 t4) + → treeInvariant (node key value (node key1 value1 t1 t2) (node key2 value2 t3 t4)) + +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 : {key1 : ℕ } → {value : A } → {tree0 t1 t2 : bt A } → {st : List (bt A)} + → key1 < key + → stackInvariant key (node key1 value t1 t2) tree0 st + → stackInvariant key t2 tree0 (t2 ∷ st) + + s-left : {key1 : ℕ } → {value : A } → {tree0 t1 t2 : bt A } → {st : List (bt A)} + → key < key1 + → stackInvariant key (node key1 value t1 t2) tree0 st + → stackInvariant key t1 tree0 (t1 ∷ 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 : {value1 : A} → {left right : bt A} → replacedTree key value (node key value left right) (node key value1 left right) + + -- key is the repl's key , so need comp key and key1 + r-left : {key1 : ℕ} {value1 : A }→ {left right repl : bt A} → key < key1 + → replacedTree key value left repl → replacedTree key value (node key1 value1 left right) (node key1 value1 repl right) + + r-right : {key1 : ℕ } {value1 : A} → {left right repl : bt A} → key1 < key + → replacedTree key value right repl → replacedTree key value (node key1 value1 left right) (node key1 value1 left repl) + +{- +RTtoTI0 : {n : Level} {A : Set n } → (key : ℕ ) → (value : A) → (tree repl : bt A) + → treeInvariant tree → replacedTree key value tree repl → treeInvariant repl +RTtoTI0 key value leaf (node key value leaf leaf) tr r-leaf = t-single key value +RTtoTI0 key value (node key₁ value₁ tree tree₁) (node key₂ value₂ repl repl₁) (t-node x x₁ s s₁) r-node = t-node x x₁ s s₁ +-} +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} {key : ℕ} {value : A} → (tleft tright : bt A) + → treeInvariant (node key value tleft tright) + → treeInvariant tleft +treeLeftDown leaf leaf (t-single key value) = t-leaf +treeLeftDown leaf (node key value t1 t2) (t-right x ti) = t-leaf +treeLeftDown (node key value t t₁) leaf (t-left x ti) = ti +treeLeftDown (node key value t t₁) (node key₁ value₁ t1 t2) (t-node x x1 ti ti2 ) = ti + +treeRightDown : {n : Level} {A : Set n} {key : ℕ} {value : A} → (tleft tright : bt A) + → treeInvariant (node key value tleft tright) + → treeInvariant tright +treeRightDown leaf leaf (t-single key value) = t-leaf +treeRightDown leaf (node key value t1 t2) (t-right x ti) = ti + +treeRightDown (node key value t t₁) leaf (t-left x ti) = t-leaf +treeRightDown (node key value t t₁) (node key₁ value₁ t1 t2) (t-node x x1 ti ti2 ) = ti2 + + +findP : {n m : Level} {A : Set n} {t : Set n} → (tkey : ℕ) → (top orig : bt A) → (st : List (bt A)) + → (treeInvariant top) + → stackInvariant tkey top orig st + → (next : (newtop : bt A) → (stack : List (bt A)) + → (treeInvariant newtop) + → (stackInvariant tkey newtop orig stack) + → bt-depth newtop < bt-depth top → t) + → (exit : (newtop : bt A) → (stack : List (bt A)) + → (treeInvariant newtop) + → (stackInvariant tkey newtop orig stack) --need new stack ? + → (newtop ≡ leaf) ∨ (node-key newtop ≡ just tkey) → t) + → t +findP tkey leaf orig st ti si next exit = exit leaf st ti si (case1 refl) +findP tkey (node key value tl tr) orig st ti si next exit with <-cmp tkey key +findP tkey top orig st ti si next exit | tri≈ ¬a refl ¬c = exit top st ti si (case2 refl) +findP tkey (node key value tl tr) orig st ti si next exit | tri< a ¬b ¬c = next tl (tl ∷ st) (treeLeftDown tl tr ti) (s-left a si) (s≤s (m≤m⊔n (bt-depth tl) (bt-depth tr))) + +findP tkey (node key value tl tr) orig st ti si next exit | tri> ¬a ¬b c = next tr (tr ∷ st) (treeRightDown tl tr ti) (s-right c si) (s≤s (m≤n⊔m (bt-depth tl) (bt-depth tr))) + + +--RBT +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 + +black-depth : {n : Level} {A : Set n} → (tree : bt (Color ∧ A) ) → ℕ +black-depth leaf = 0 +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₁ ) + +data RBtreeInvariant {n : Level} {A : Set n} : (tree : bt (Color ∧ A)) → Set n where + rb-leaf : RBtreeInvariant leaf + rb-single : (key : ℕ) → (value : A) → RBtreeInvariant (node key ⟪ Black , value ⟫ leaf leaf) + rb-right-red : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ + → black-depth t ≡ black-depth t₁ + → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ t t₁) + → RBtreeInvariant (node key ⟪ Red , value ⟫ leaf (node key₁ ⟪ Black , value₁ ⟫ t t₁)) + rb-right-black : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ → {c : Color} + → black-depth t ≡ black-depth t₁ + → RBtreeInvariant (node key₁ ⟪ c , value₁ ⟫ t t₁) + → RBtreeInvariant (node key ⟪ Black , value ⟫ leaf (node key₁ ⟪ c , value₁ ⟫ t t₁)) + rb-left-red : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ + → black-depth t ≡ black-depth t₁ + → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ t t₁) + → RBtreeInvariant (node key ⟪ Red , value ⟫ (node key₁ ⟪ Black , value₁ ⟫ t t₁) leaf ) + rb-left-black : {key key₁ : ℕ} → {value value₁ : A} → {t t₁ : bt (Color ∧ A)} → key < key₁ → {c : Color} + → black-depth t ≡ black-depth t₁ + → RBtreeInvariant (node key₁ ⟪ c , value₁ ⟫ t t₁) + → RBtreeInvariant (node key ⟪ Black , value ⟫ (node key₁ ⟪ c , value₁ ⟫ t t₁) leaf) + rb-node-red : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt (Color ∧ A)} → key < key₁ → key₁ < key₂ + → black-depth t₁ ≡ black-depth t₂ + → RBtreeInvariant (node key ⟪ Black , value ⟫ t₁ t₂) + → black-depth t₃ ≡ black-depth t₄ + → RBtreeInvariant (node key₂ ⟪ Black , value₂ ⟫ t₃ t₄) + → RBtreeInvariant (node key₁ ⟪ Red , value₁ ⟫ (node key ⟪ Black , value ⟫ t₁ t₂) (node key₂ ⟪ Black , value₂ ⟫ t₃ t₄)) + rb-node-black : {key key₁ key₂ : ℕ} → {value value₁ value₂ : A} → {t₁ t₂ t₃ t₄ : bt (Color ∧ A)} → key < key₁ → key₁ < key₂ + → {c c₁ : Color} + → black-depth t₁ ≡ black-depth t₂ + → RBtreeInvariant (node key ⟪ c , value ⟫ t₁ t₂) + → black-depth t₃ ≡ black-depth t₄ + → RBtreeInvariant (node key₂ ⟪ c₁ , value₂ ⟫ t₃ t₄) + → RBtreeInvariant (node key₁ ⟪ Black , value₁ ⟫ (node key ⟪ c , value ⟫ t₁ t₂) (node key₂ ⟪ c₁ , value₂ ⟫ t₃ t₄)) + + +data rotatedTree {n : Level} {A : Set n} : (before after : bt A) → Set n where + rtt-node : {t : bt A } → rotatedTree t t + -- a b + -- b c d a + -- d e e c + -- + rtt-right : {ka kb kc kd ke : ℕ} {va vb vc vd ve : A} → {c d e c1 d1 e1 : bt A} → {ctl ctr dtl dtr etl etr : bt A} + --kd < kb < ke < ka< kc + → {ctl1 ctr1 dtl1 dtr1 etl1 etr1 : bt A} + → kd < kb → kb < ke → ke < ka → ka < kc + → rotatedTree (node ke ve etl etr) (node ke ve etl1 etr1) + → rotatedTree (node kd vd dtl dtr) (node kd vd dtl1 dtr1) + → rotatedTree (node kc vc ctl ctr) (node kc vc ctl1 ctr1) + → rotatedTree (node ka va (node kb vb (node kd vd dtl dtr) (node ke ve etl etr)) (node kc vc ctl ctr)) + (node kb vb (node kd vd dtl1 dtr1) (node ka va (node ke ve etl1 etr1) (node kc vc ctl1 ctr1))) + + rtt-left : {ka kb kc kd ke : ℕ} {va vb vc vd ve : A} → {c d e c1 d1 e1 : bt A} → {ctl ctr dtl dtr etl etr : bt A} + --kd < kb < ke < ka< kc + → {ctl1 ctr1 dtl1 dtr1 etl1 etr1 : bt A} -- after child + → kd < kb → kb < ke → ke < ka → ka < kc + → rotatedTree (node ke ve etl etr) (node ke ve etl1 etr1) + → rotatedTree (node kd vd dtl dtr) (node kd vd dtl1 dtr1) + → rotatedTree (node kc vc ctl ctr) (node kc vc ctl1 ctr1) + → rotatedTree (node kb vb (node kd vd dtl1 dtr1) (node ka va (node ke ve etl1 etr1) (node kc vc ctl1 ctr1))) + (node ka va (node kb vb (node kd vd dtl dtr) (node ke ve etl etr)) (node kc vc ctl ctr)) + +RBtreeLeftDown : {n : Level} {A : Set n} {key : ℕ} {value : A} {c : Color} + → (tleft tright : bt (Color ∧ A)) + → RBtreeInvariant (node key ⟪ c , value ⟫ tleft tright) + → RBtreeInvariant tleft +RBtreeLeftDown leaf leaf (rb-single k1 v) = rb-leaf +RBtreeLeftDown leaf (node key ⟪ Black , value ⟫ t1 t2 ) (rb-right-red x bde rbti) = rb-leaf +RBtreeLeftDown leaf (node key ⟪ Black , value ⟫ t1 t2 ) (rb-right-black x bde rbti) = rb-leaf +RBtreeLeftDown leaf (node key ⟪ Red , value ⟫ t1 t2 ) (rb-right-black x bde rbti)= rb-leaf +RBtreeLeftDown (node key ⟪ Black , value ⟫ t t₁) leaf (rb-left-black x bde ti) = ti +RBtreeLeftDown (node key ⟪ Black , value ⟫ t t₁) leaf (rb-left-red x bde ti)= ti +RBtreeLeftDown (node key ⟪ Red , value ⟫ t t₁) leaf (rb-left-black x bde ti) = ti +RBtreeLeftDown (node key ⟪ Black , value ⟫ t t₁) (node key₁ ⟪ Black , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = til +RBtreeLeftDown (node key ⟪ Black , value ⟫ t t₁) (node key₁ ⟪ Black , value1 ⟫ t1 t2) (rb-node-red x x1 bde1 til bde2 tir) = til +RBtreeLeftDown (node key ⟪ Red , value ⟫ t t₁) (node key₁ ⟪ Black , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = til +RBtreeLeftDown (node key ⟪ Black , value ⟫ t t₁) (node key₁ ⟪ Red , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = til +RBtreeLeftDown (node key ⟪ Red , value ⟫ t t₁) (node key₁ ⟪ Red , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = til + +RBtreeRightDown : {n : Level} {A : Set n} { key : ℕ} {value : A} {c : Color} + → (tleft tright : bt (Color ∧ A)) + → RBtreeInvariant (node key ⟪ c , value ⟫ tleft tright) + → RBtreeInvariant tright +RBtreeRightDown leaf leaf (rb-single k1 v1 ) = rb-leaf +RBtreeRightDown leaf (node key ⟪ Black , value ⟫ t1 t2 ) (rb-right-red x bde rbti) = rbti +RBtreeRightDown leaf (node key ⟪ Black , value ⟫ t1 t2 ) (rb-right-black x bde rbti) = rbti +RBtreeRightDown leaf (node key ⟪ Red , value ⟫ t1 t2 ) (rb-right-black x bde rbti)= rbti +RBtreeRightDown (node key ⟪ Black , value ⟫ t t₁) leaf (rb-left-black x bde ti) = rb-leaf +RBtreeRightDown (node key ⟪ Black , value ⟫ t t₁) leaf (rb-left-red x bde ti) = rb-leaf +RBtreeRightDown (node key ⟪ Red , value ⟫ t t₁) leaf (rb-left-black x bde ti) = rb-leaf +RBtreeRightDown (node key ⟪ Black , value ⟫ t t₁) (node key₁ ⟪ Black , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = tir +RBtreeRightDown (node key ⟪ Black , value ⟫ t t₁) (node key₁ ⟪ Black , value1 ⟫ t1 t2) (rb-node-red x x1 bde1 til bde2 tir) = tir +RBtreeRightDown (node key ⟪ Red , value ⟫ t t₁) (node key₁ ⟪ Black , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = tir +RBtreeRightDown (node key ⟪ Black , value ⟫ t t₁) (node key₁ ⟪ Red , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = tir +RBtreeRightDown (node key ⟪ Red , value ⟫ t t₁) (node key₁ ⟪ Red , value1 ⟫ t1 t2) (rb-node-black x x1 bde1 til bde2 tir) = tir + +findRBT : {n m : Level} {A : Set n} {t : Set m} → (key : ℕ) → (tree tree0 : bt (Color ∧ A) ) + → (stack : List (bt (Color ∧ A))) + → treeInvariant tree ∧ stackInvariant key tree tree0 stack + → RBtreeInvariant tree + → (next : (tree1 : bt (Color ∧ A) ) → (stack : List (bt (Color ∧ A))) + → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack + → RBtreeInvariant tree1 + → bt-depth tree1 < bt-depth tree → t ) + → (exit : (tree1 : bt (Color ∧ A)) → (stack : List (bt (Color ∧ A))) + → treeInvariant tree1 ∧ stackInvariant key tree1 tree0 stack + → RBtreeInvariant tree1 + → (tree1 ≡ leaf ) ∨ ( node-key tree1 ≡ just key ) → t ) → t +findRBT key leaf tree0 stack ti rb0 next exit = exit leaf stack ti rb0 (case1 refl) +findRBT key n@(node key₁ value left right) tree0 stack ti rb0 next exit with <-cmp key key₁ +findRBT key (node key₁ value left right) tree0 stack ti rb0 next exit | tri< a ¬b ¬c + = next left (left ∷ stack) ⟪ treeLeftDown left right (_∧_.proj1 ti) , s-left a (_∧_.proj2 ti) ⟫ (RBtreeLeftDown left right rb0) depth-1< +findRBT key n tree0 stack ti rb0 _ exit | tri≈ ¬a refl ¬c = exit n stack ti rb0 (case2 refl) +findRBT key (node key₁ value left right) tree0 stack ti rb0 next exit | tri> ¬a ¬b c + = next right (right ∷ stack) ⟪ treeRightDown left right (_∧_.proj1 ti), s-right c (_∧_.proj2 ti) ⟫ (RBtreeRightDown left right rb0) depth-2< + +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 + + +data replacedRBTree {n : Level} {A : Set n} (key : ℕ) (value : A) : (before after : bt (Color ∧ A) ) → Set n where + rbr-leaf : {ca cb : Color} → replacedRBTree key value leaf (node key ⟪ cb , value ⟫ leaf leaf) + rbr-node : {value₁ : A} → {ca cb : Color } → {t t₁ : bt (Color ∧ A)} + → replacedRBTree key value (node key ⟪ ca , value₁ ⟫ t t₁) (node key ⟪ cb , value ⟫ t t₁) + rbr-right : {k : ℕ } {v1 : A} → {ca cb : Color} → {t t1 t2 : bt (Color ∧ A)} + → k < key → replacedRBTree key value t2 t → replacedRBTree key value (node k ⟪ ca , v1 ⟫ t1 t2) (node k ⟪ cb , v1 ⟫ t1 t) + rbr-left : {k : ℕ } {v1 : A} → {ca cb : Color} → {t t1 t2 : bt (Color ∧ A)} + → k < key → replacedRBTree key value t1 t → replacedRBTree key value (node k ⟪ ca , v1 ⟫ t1 t2) (node k ⟪ cb , v1 ⟫ t t2) + +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 ) + +record RBI {n : Level} {A : Set n} (key : ℕ) (value : A) (orig repl : bt (Color ∧ A) ) (stack : List (bt (Color ∧ A))) : Set n where + field + od d rd : ℕ + tree rot : bt (Color ∧ A) + origti : treeInvariant orig + origrb : RBtreeInvariant orig + treerb : RBtreeInvariant tree + replrb : RBtreeInvariant repl + d=rd : ( d ≡ rd ) ∨ ((suc d ≡ rd ) ∧ (color tree ≡ Red)) + si : stackInvariant key tree orig stack + rotated : rotatedTree tree rot + ri : replacedRBTree key value (child-replaced key rot ) repl + + +{- +rbi-case1 : {n : Level} {A : Set n} → {key : ℕ} → {value : A} → (parent repl : bt (Color ∧ A) ) + → RBtreeInvariant parent + → RBtreeInvariant repl + → {left right : bt (Color ∧ A)} → parent ≡ node key ⟪ Black , value ⟫ left right + → (color right ≡ Red → RBtreeInvariant (node key ⟪ Black , value ⟫ left repl ) ) + ∧ (color left ≡ Red → RBtreeInvariant (node key ⟪ Black , value ⟫ repl right ) ) +rbi-case1 {n} {A} {key} (node key1 ⟪ Black , value1 ⟫ l r) leaf rbip rbir left right x = {!!} +rbi-case1 {n} {A} {key} parent (node key₁ value₁ tree1 tree2) rbi rb2 x = {!!} + +-} +blackdepth≡ : {n : Level } {A : Set n} → {C : Color} {key : ℕ} {value : A} → (tree1 tree2 : bt (Color ∧ A)) + → RBtreeInvariant tree1 + → RBtreeInvariant tree2 + → RBtreeInvariant (node key ⟪ C , value ⟫ tree1 tree2) + → black-depth tree1 ≡ black-depth tree2 + +blackdepth≡ leaf leaf ri1 ri2 rip = refl +blackdepth≡ leaf (node key value t2 t3) ri1 ri2 rip = {!!} --rip kara mitibiki daseru RBinvariant kara toreruka +blackdepth≡ (node key value t1 t3) leaf ri1 ri2 rip = {!!} +blackdepth≡ (node key value t1 t3) (node key₁ value₁ t2 t4) ri1 ri2 rip = {!!} + +rbi-case1 : {n : Level} {A : Set n} → {key : ℕ} → {value : A} → (parent repl : bt (Color ∧ A) ) + → RBtreeInvariant parent + → RBtreeInvariant repl + → (left right : bt (Color ∧ A)) → parent ≡ node key ⟪ Black , value ⟫ left right + → RBtreeInvariant left + → RBtreeInvariant right + → (color right ≡ Red → RBtreeInvariant (node key ⟪ Black , value ⟫ left repl ) ) ∧ (color left ≡ Red → RBtreeInvariant (node key ⟪ Black , value ⟫ repl right ) ) +rbi-case1 {n} {A} {key} (node key1 ⟪ Black , value1 ⟫ l r) leaf rbip rbir (node key3 ⟪ Red , val3 ⟫ la ra) (node key4 ⟪ Red , val4 ⟫ lb rb) pa li ri + = ⟪ {!!} rb-left-black {!!} {!!} li , (λ x → rb-right-black {!!} {!!} ri) ⟫ + +--⟪ rb-left-black {!!} {!!} (RBtreeLeftDown left right rbip ) , {!!} ⟫ +--rbi-case1 {n} {A} {key} parent (node key₁ value₁ tree1 tree2) rbi rb2 x = {!!}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/work.agda~ Mon Jul 10 19:59:14 2023 +0900 @@ -0,0 +1,22 @@ +module work 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 + +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