Path operations; modify and alter #
This develops the necessary theorems to construct the modify and alter functions on RBSet
using path operations for in-place modification of an RBTree.
path balance #
Asserts that property p holds on the root of the tree, if any.
Equations
- Batteries.RBNode.OnRoot p x = match x with | Batteries.RBNode.nil => True | Batteries.RBNode.node c l x r => p x
Instances For
Same as fill but taking its arguments in a pair for easier composition with zoom.
Equations
- Batteries.RBNode.Path.fill' x = match x with | (t, path) => path.fill t
Instances For
The balance invariant for a path. path.Balanced c₀ n₀ c n means that path is a red-black tree
with balance invariant c₀, n₀, but it has a "hole" where a tree with balance invariant c, n
has been removed. The defining property is Balanced.fill: if path.Balanced c₀ n₀ c n and you
fill the hole with a tree satisfying t.Balanced c n, then (path.fill t).Balanced c₀ n₀ .
- root: ∀ {c₀ : Batteries.RBColor} {n₀ : Nat} {α : Type u_1},
  Batteries.RBNode.Path.Balanced c₀ n₀ Batteries.RBNode.Path.root c₀ n₀The root of the tree is c₀, n₀-balanced by assumption.
- redL: ∀ {c₀ : Batteries.RBColor} {n₀ : Nat} {α : Type u_1} {y : Batteries.RBNode α} {n : Nat}
  {parent : Batteries.RBNode.Path α} {v : α},
  y.Balanced Batteries.RBColor.black n →
    Batteries.RBNode.Path.Balanced c₀ n₀ parent Batteries.RBColor.red n →
      Batteries.RBNode.Path.Balanced c₀ n₀ (Batteries.RBNode.Path.left Batteries.RBColor.red parent v y)
        Batteries.RBColor.black nDescend into the left subtree of a red node. 
- redR: ∀ {c₀ : Batteries.RBColor} {n₀ : Nat} {α : Type u_1} {x : Batteries.RBNode α} {n : Nat} {v : α}
  {parent : Batteries.RBNode.Path α},
  x.Balanced Batteries.RBColor.black n →
    Batteries.RBNode.Path.Balanced c₀ n₀ parent Batteries.RBColor.red n →
      Batteries.RBNode.Path.Balanced c₀ n₀ (Batteries.RBNode.Path.right Batteries.RBColor.red x v parent)
        Batteries.RBColor.black nDescend into the right subtree of a red node. 
- blackL: ∀ {c₀ : Batteries.RBColor} {n₀ : Nat} {α : Type u_1} {y : Batteries.RBNode α} {c₂ : Batteries.RBColor} {n : Nat}
  {parent : Batteries.RBNode.Path α} {v : α} {c₁ : Batteries.RBColor},
  y.Balanced c₂ n →
    Batteries.RBNode.Path.Balanced c₀ n₀ parent Batteries.RBColor.black (n + 1) →
      Batteries.RBNode.Path.Balanced c₀ n₀ (Batteries.RBNode.Path.left Batteries.RBColor.black parent v y) c₁ nDescend into the left subtree of a black node. 
- blackR: ∀ {c₀ : Batteries.RBColor} {n₀ : Nat} {α : Type u_1} {x : Batteries.RBNode α} {c₁ : Batteries.RBColor} {n : Nat} {v : α}
  {parent : Batteries.RBNode.Path α} {c₂ : Batteries.RBColor},
  x.Balanced c₁ n →
    Batteries.RBNode.Path.Balanced c₀ n₀ parent Batteries.RBColor.black (n + 1) →
      Batteries.RBNode.Path.Balanced c₀ n₀ (Batteries.RBNode.Path.right Batteries.RBColor.black x v parent) c₂ nDescend into the right subtree of a black node. 
Instances For
The defining property of a balanced path: If path is a c₀,n₀ tree with a c,n hole,
then filling the hole with a c,n tree yields a c₀,n₀ tree.
The property of a path returned by t.zoom cut. Each of the parents visited along the path have
the appropriate ordering relation to the cut.
Equations
- Batteries.RBNode.Path.Zoomed cut Batteries.RBNode.Path.root = True
- Batteries.RBNode.Path.Zoomed cut (Batteries.RBNode.Path.left c parent x_1 r) = (cut x_1 = Ordering.lt ∧ Batteries.RBNode.Path.Zoomed cut parent)
- Batteries.RBNode.Path.Zoomed cut (Batteries.RBNode.Path.right c l x_1 parent) = (cut x_1 = Ordering.gt ∧ Batteries.RBNode.Path.Zoomed cut parent)
Instances For
path.RootOrdered cmp v is true if v would be able to fit into the hole
without violating the ordering invariant.
Equations
- Batteries.RBNode.Path.RootOrdered cmp Batteries.RBNode.Path.root x = True
- Batteries.RBNode.Path.RootOrdered cmp (Batteries.RBNode.Path.left c parent x_2 r) x = (Batteries.RBNode.cmpLT cmp x x_2 ∧ Batteries.RBNode.Path.RootOrdered cmp parent x)
- Batteries.RBNode.Path.RootOrdered cmp (Batteries.RBNode.Path.right c l x_2 parent) x = (Batteries.RBNode.cmpLT cmp x_2 x ∧ Batteries.RBNode.Path.RootOrdered cmp parent x)
Instances For
The ordering invariant for a Path.
Equations
- One or more equations did not get rendered due to their size.
- Batteries.RBNode.Path.Ordered cmp Batteries.RBNode.Path.root = True
Instances For
alter #
The alter function preserves the ordering invariants.
The alter function preserves the balance invariants.
The modify function preserves the ordering invariants.
The modify function preserves the balance invariants.
A sufficient condition for ModifyWF is that the new element compares equal to the original.
O(log n). In-place replace the corresponding to key k.
This takes the element out of the tree while f runs,
so it uses the element linearly if t is unshared.
Equations
- t.modify k f = Batteries.RBSet.modifyP t (fun (x : α × β) => cmp k x.fst) fun (x : α × β) => match x with | (a, b) => (a, f b)
Instances For
O(log n). alterP cut f t simultaneously handles inserting, erasing and replacing an element
using a function f : Option α → Option α. It is passed the result of t.findP? cut
and can either return none to remove the element or some a to replace/insert
the element with a (which must have the same ordering properties as the original element).
The element is used linearly if t is unshared.
The AlterWF assumption is required because f may change
the ordering properties of the element, which would break the invariants.
Equations
- t.alter k f = Batteries.RBSet.alterP t (fun (x : α × β) => cmp k x.fst) (Batteries.RBMap.alter.adapt k f)