- nil: {α : Type u} → Std.BinomialHeap.Imp.HeapNode α
An empty forest, which has depth
0
. - node: {α : Type u} → α → Std.BinomialHeap.Imp.HeapNode α → Std.BinomialHeap.Imp.HeapNode α → Std.BinomialHeap.Imp.HeapNode α
A forest of rank
r + 1
consists of a roota
, a forestchild
of rankr
elements greater thana
, and another forestsibling
of rankr
.
A HeapNode
is one of the internal nodes of the binomial heap.
It is always a perfect binary tree, with the depth of the tree stored in the Heap
.
However the interpretation of the two pointers is different: we view the child
as going to the first child of this node, and sibling
goes to the next sibling
of this tree. So it actually encodes a forest where each node has children
node.child
, node.child.sibling
, node.child.sibling.sibling
, etc.
Each edge in this forest denotes a le a b
relation that has been checked, so
the root is smaller than everything else under it.
Instances For
The "real size" of the node, counting up how many values of type α
are stored.
This is O(n)
and is intended mainly for specification purposes.
For a well formed HeapNode
the size is always 2^n - 1
where n
is the depth.
Equations
- Std.BinomialHeap.Imp.HeapNode.realSize Std.BinomialHeap.Imp.HeapNode.nil = 0
- Std.BinomialHeap.Imp.HeapNode.realSize (Std.BinomialHeap.Imp.HeapNode.node a a_1 a_2) = Std.BinomialHeap.Imp.HeapNode.realSize a_1 + 1 + Std.BinomialHeap.Imp.HeapNode.realSize a_2
Instances For
A node containing a single element a
.
Instances For
O(log n)
. The rank, or the number of trees in the forest.
It is also the depth of the forest.
Equations
- Std.BinomialHeap.Imp.HeapNode.rank Std.BinomialHeap.Imp.HeapNode.nil = 0
- Std.BinomialHeap.Imp.HeapNode.rank (Std.BinomialHeap.Imp.HeapNode.node a a_1 a_2) = Std.BinomialHeap.Imp.HeapNode.rank a_2 + 1
Instances For
- nil: {α : Type u} → Std.BinomialHeap.Imp.Heap α
An empty heap.
- cons: {α : Type u} → Nat → α → Std.BinomialHeap.Imp.HeapNode α → Std.BinomialHeap.Imp.Heap α → Std.BinomialHeap.Imp.Heap α
A Heap
is the top level structure in a binomial heap.
It consists of a forest of HeapNode
s with strictly increasing ranks.
Instances For
O(n)
. The "real size" of the heap, counting up how many values of type α
are stored.
This is intended mainly for specification purposes.
Prefer Heap.size
, which is the same for well formed heaps.
Equations
- Std.BinomialHeap.Imp.Heap.realSize Std.BinomialHeap.Imp.Heap.nil = 0
- Std.BinomialHeap.Imp.Heap.realSize (Std.BinomialHeap.Imp.Heap.cons a a_1 a_2 a_3) = Std.BinomialHeap.Imp.HeapNode.realSize a_2 + 1 + Std.BinomialHeap.Imp.Heap.realSize a_3
Instances For
O(log n)
. The number of elements in the heap.
Equations
- Std.BinomialHeap.Imp.Heap.size Std.BinomialHeap.Imp.Heap.nil = 0
- Std.BinomialHeap.Imp.Heap.size (Std.BinomialHeap.Imp.Heap.cons a a_1 a_2 a_3) = 1 <<< a + Std.BinomialHeap.Imp.Heap.size a_3
Instances For
O(1)
. Is the heap empty?
Instances For
O(1)
. The heap containing a single value a
.
Instances For
O(1)
. Auxiliary for Heap.merge
: Is the minimum rank in Heap
strictly larger than n
?
Instances For
O(log n)
. The number of trees in the forest.
Equations
- Std.BinomialHeap.Imp.Heap.length Std.BinomialHeap.Imp.Heap.nil = 0
- Std.BinomialHeap.Imp.Heap.length (Std.BinomialHeap.Imp.Heap.cons a a_1 a_2 a_3) = Std.BinomialHeap.Imp.Heap.length a_3 + 1
Instances For
O(1)
. Auxiliary for Heap.merge
: combines two heap nodes of the same rank
into one with the next larger rank.
Instances For
Merge two forests of binomial trees. The forests are assumed to be ordered
by rank and merge
maintains this invariant.
Equations
- One or more equations did not get rendered due to their size.
- Std.BinomialHeap.Imp.Heap.merge le Std.BinomialHeap.Imp.Heap.nil x = x
- Std.BinomialHeap.Imp.Heap.merge le x Std.BinomialHeap.Imp.Heap.nil = x
Instances For
O(log n)
. Convert a HeapNode
to a Heap
by reversing the order of the nodes
along the sibling
spine.
Instances For
Computes s.toHeap ++ res
tail-recursively, assuming n = s.rank
.
Equations
- Std.BinomialHeap.Imp.HeapNode.toHeap.go Std.BinomialHeap.Imp.HeapNode.nil x x = x
- Std.BinomialHeap.Imp.HeapNode.toHeap.go (Std.BinomialHeap.Imp.HeapNode.node a c s) x x = Std.BinomialHeap.Imp.HeapNode.toHeap.go s (x - 1) (Std.BinomialHeap.Imp.Heap.cons (x - 1) a c x)
Instances For
O(log n)
. Get the smallest element in the heap, including the passed in value a
.
Equations
- Std.BinomialHeap.Imp.Heap.headD le a Std.BinomialHeap.Imp.Heap.nil = a
- Std.BinomialHeap.Imp.Heap.headD le a (Std.BinomialHeap.Imp.Heap.cons a_1 a_2 a_3 a_4) = Std.BinomialHeap.Imp.Heap.headD le (if le a a_2 = true then a else a_2) a_4
Instances For
O(log n)
. Get the smallest element in the heap, if it has an element.
Instances For
- before : Std.BinomialHeap.Imp.Heap α → Std.BinomialHeap.Imp.Heap α
The list of elements prior to the minimum element, encoded as a "difference list".
- val : α
The minimum element.
- node : Std.BinomialHeap.Imp.HeapNode α
The children of the minimum element.
- next : Std.BinomialHeap.Imp.Heap α
The forest after the minimum element.
The return type of FindMin
, which encodes various quantities needed to
reconstruct the tree in deleteMin
.
Instances For
O(log n)
. Find the minimum element, and return a data structure FindMin
with information
needed to reconstruct the rest of the binomial heap.
Equations
- One or more equations did not get rendered due to their size.
- Std.BinomialHeap.Imp.Heap.findMin le k Std.BinomialHeap.Imp.Heap.nil x = x
Instances For
O(log n)
. Find and remove the the minimum element from the binomial heap.
Instances For
O(log n)
. Get the tail of the binomial heap after removing the minimum element.
Instances For
O(log n)
. Remove the minimum element of the heap.
Instances For
O(n log n)
. Monadic fold over the elements of a heap in increasing order,
by repeatedly pulling the minimum element out of the heap.
Equations
- One or more equations did not get rendered due to their size.
Instances For
O(n log n)
. Fold over the elements of a heap in increasing order,
by repeatedly pulling the minimum element out of the heap.
Instances For
O(n log n)
. Convert the heap to an array in increasing order.
Instances For
O(n log n)
. Convert the heap to a list in increasing order.
Instances For
O(n)
. Fold a monadic function over the tree structure to accumulate a value.
Equations
- One or more equations did not get rendered due to their size.
- Std.BinomialHeap.Imp.HeapNode.foldTreeM nil join Std.BinomialHeap.Imp.HeapNode.nil = pure nil
Instances For
O(n)
. Fold a monadic function over the tree structure to accumulate a value.
Equations
- One or more equations did not get rendered due to their size.
- Std.BinomialHeap.Imp.Heap.foldTreeM nil join Std.BinomialHeap.Imp.Heap.nil = pure nil
Instances For
O(n)
. Fold a function over the tree structure to accumulate a value.
Instances For
O(n)
. Convert the heap to a list in arbitrary order.
Instances For
O(n)
. Convert the heap to an array in arbitrary order.
Instances For
The well formedness predicate for a heap node. It asserts that:
- If
a
is added at the top to make the forest into a tree, the resulting tree is ale
-min-heap (ifle
is well-behaved) - When interpreting
child
andsibling
as left and right children of a binary tree, it is a perfect binary tree with depthr
Equations
- One or more equations did not get rendered due to their size.
- Std.BinomialHeap.Imp.HeapNode.WF le a Std.BinomialHeap.Imp.HeapNode.nil x = (x = 0)
Instances For
The well formedness predicate for a binomial heap. It asserts that:
- It consists of a list of well formed trees with the specified ranks
- The ranks are in strictly increasing order, and all are at least
n
Equations
- Std.BinomialHeap.Imp.Heap.WF le n Std.BinomialHeap.Imp.Heap.nil = True
- Std.BinomialHeap.Imp.Heap.WF le n (Std.BinomialHeap.Imp.Heap.cons a a_1 a_2 a_3) = (n ≤ a ∧ Std.BinomialHeap.Imp.HeapNode.WF le a_1 a_2 a ∧ Std.BinomialHeap.Imp.Heap.WF le (a + 1) a_3)
Instances For
- rank : Nat
The rank of the minimum element
- before : ∀ {s_1 : Std.BinomialHeap.Imp.Heap α}, Std.BinomialHeap.Imp.Heap.WF le s.rank s_1 → Std.BinomialHeap.Imp.Heap.WF le 0 (Std.BinomialHeap.Imp.FindMin.before res s_1)
- node : Std.BinomialHeap.Imp.HeapNode.WF le res.val res.node s.rank
- next : Std.BinomialHeap.Imp.Heap.WF le (s.rank + 1) res.next
The well formedness predicate for a FindMin
value.
This is not actually a predicate, as it contains an additional data value
rank
corresponding to the rank of the returned node, which is omitted from findMin
.
Instances For
The conditions under which findMin
is well-formed.
Equations
- One or more equations did not get rendered due to their size.
- Std.BinomialHeap.Imp.Heap.WF.findMin h_2 hr hk = hr
Instances For
A binomial heap is a data structure which supports the following primary operations:
insert : α → BinomialHeap α → BinomialHeap α
: add an element to the heapdeleteMin : BinomialHeap α → Option (α × BinomialHeap α)
: remove the minimum element from the heapmerge : BinomialHeap α → BinomialHeap α → BinomialHeap α
: combine two heaps
The first two operations are known as a "priority queue", so this could be called
a "mergeable priority queue". The standard choice for a priority queue is a binary heap,
which supports insert
and deleteMin
in O(log n)
, but merge
is O(n)
.
With a BinomialHeap
, all three operations are O(log n)
.
Instances For
O(1)
. Make a new empty binomial heap.
Instances For
O(1)
. Make a new empty binomial heap.
Instances For
O(1)
. Is the heap empty?
Instances For
O(log n)
. The number of elements in the heap.
Instances For
O(1)
. Make a new heap containing a
.
Instances For
O(log n)
. Merge the contents of two heaps.
Instances For
O(log n)
. Add element a
to the given heap h
.
Instances For
O(n log n)
. Construct a heap from a list by inserting all the elements.
Instances For
O(n log n)
. Construct a heap from a list by inserting all the elements.
Instances For
O(log n)
. Remove and return the minimum element from the heap.
Instances For
O(n log n)
. Implementation of for x in (b : BinomialHeap α le) ...
notation,
which iterates over the elements in the heap in increasing order.
Instances For
O(log n)
. Returns the smallest element in the heap, or none
if the heap is empty.
Instances For
O(log n)
. Returns the smallest element in the heap, or panics if the heap is empty.
Instances For
O(log n)
. Returns the smallest element in the heap, or default
if the heap is empty.
Instances For
O(log n)
. Removes the smallest element from the heap, or none
if the heap is empty.
Instances For
O(log n)
. Removes the smallest element from the heap, if possible.
Instances For
O(n log n)
. Monadic fold over the elements of a heap in increasing order,
by repeatedly pulling the minimum element out of the heap.
Instances For
O(n log n)
. Fold over the elements of a heap in increasing order,
by repeatedly pulling the minimum element out of the heap.
Instances For
O(n log n)
. Convert the heap to a list in increasing order.
Instances For
O(n log n)
. Convert the heap to an array in increasing order.
Instances For
O(n)
. Convert the heap to a list in arbitrary order.
Instances For
O(n)
. Convert the heap to an array in arbitrary order.