vers_vecs/trees/bp/
mod.rs

1//! A succinct tree data structure backed by the balanced parenthesis representation.
2//! The tree supports navigation operations between parent, child, and sibling nodes in `O(log n)`
3//! time, as well as subtree size, level-order, and ancestor queries in `O(log n)` time.
4//! The tree is succinct (ideally sublinear space overhead) and pointer-less.
5
6use crate::bit_vec::fast_rs_vec::SelectIntoIter;
7use crate::trees::mmt::MinMaxTree;
8use crate::trees::{IsAncestor, LevelTree, SubtreeSize, Tree};
9use crate::{BitVec, RsVec};
10use std::cmp::{max, min};
11use std::iter::FusedIterator;
12
13/// The default block size for the tree, used in several const generics
14const DEFAULT_BLOCK_SIZE: usize = 512;
15
16const OPEN_PAREN: u64 = 1;
17const CLOSE_PAREN: u64 = 0;
18
19mod builder;
20// re-export the builders toplevel
21pub use builder::BpBuilder;
22
23#[cfg(feature = "bp_u16_lookup")]
24mod lookup;
25#[cfg(feature = "bp_u16_lookup")]
26use lookup::{process_block_bwd, process_block_fwd, LOOKUP_BLOCK_SIZE};
27
28#[cfg(not(feature = "bp_u16_lookup"))]
29mod lookup_query;
30#[cfg(not(feature = "bp_u16_lookup"))]
31use lookup_query::{process_block_bwd, process_block_fwd, LOOKUP_BLOCK_SIZE};
32
33/// A succinct tree data structure based on balanced parenthesis expressions.
34/// A tree with `n` nodes is encoded in a bit vector using `2n` bits plus the rank/select overhead
35/// of the [`RsVec`] implementation.
36/// Additionally, a small pointerless heap data structure stores
37/// additional meta information required to perform most tree operations.
38///
39/// The tree is thus pointer-less and succinct.
40/// It supports tree navigation operations between parent, child, and sibling nodes, both in
41/// depth-first search order and in level order.
42/// All operations run in `O(log n)` time with small overheads.
43///
44/// ## Lookup Table
45/// The tree internally uses a lookup table for subqueries on blocks of bits.
46/// The lookup table requires 4 KiB of memory and is compiled into the binary.
47/// If the `bp_u16_lookup` feature is enabled, a larger lookup table is used, which requires 128 KiB of
48/// memory, but answers queries faster.
49///
50/// ## Block Size
51/// The tree has a block size of 512 bits by default, which can be changed by setting the
52/// `BLOCK_SIZE` generic parameter.
53/// This block size is expected to be a good choice for most applications,
54/// as it will fit a cache line.
55///
56/// If you want to tune the parameter,
57/// the block size should be chosen based on the expected size of the tree and the available memory.
58/// Smaller block sizes increase the size of the supporting data structure but reduce the time
59/// complexity of some operations by a constant amount.
60/// Larger block sizes are best combined with the `bp_u16_lookup` feature to keep the query time
61/// low.
62/// In any case, benchmarking for the specific use case is recommended for tuning.
63///
64/// ## Unbalanced Parentheses
65/// The tree is implemented in a way to theoretically support unbalanced parenthesis expressions
66/// (which encode invalid trees) without panicking.
67/// However, some operations may behave erratically if the parenthesis expression isn't balanced.
68/// Generally, operations specify if they require a balanced tree.
69///
70/// The results of the operations are unspecified,
71/// meaning no guarantees are made about the stability of the results across versions
72/// (except the operations not panicking).
73/// However, for research purposes, this behavior can be useful and should yield expected results
74/// in most cases.
75///
76/// Only the basic operations like [`fwd_search`] and [`bwd_search`],
77/// as well as the tree navigation operations
78/// (defined by the traits [`Tree`], [`IsAncestor`], [`LevelTree`], and [`SubtreeSize`]),
79/// are included in this guarantee.
80/// Additional operations like iterators may panic if the tree is unbalanced (this is documented per
81/// operation).
82///
83/// # Examples
84///
85/// The high-level approach to building a tree is to use the [`BpBuilder`] to construct the tree
86/// using depth-first traversal of all its nodes.
87/// ```rust
88/// use vers_vecs::{BitVec, BpBuilder, BpTree, TreeBuilder, Tree};
89///
90/// let mut builder = BpBuilder::<512>::new();
91///
92/// // build the tree by depth-first traversal
93/// builder.enter_node();
94/// builder.enter_node();
95/// builder.enter_node();
96/// builder.leave_node();
97/// builder.enter_node();
98/// builder.leave_node();
99/// builder.leave_node();
100/// builder.enter_node();
101/// builder.leave_node();
102/// builder.leave_node();
103///
104/// let tree = builder.build().unwrap();
105/// let root = tree.root().unwrap();
106/// assert_eq!(root, 0);
107/// assert_eq!(tree.first_child(root), Some(1));
108/// assert_eq!(tree.next_sibling(1), Some(7));
109/// assert_eq!(tree.next_sibling(7), None);
110///
111/// assert_eq!(root, 0);
112/// assert_eq!(tree.depth(2), 2);
113/// assert_eq!(tree.depth(7), 1);
114/// ```
115///
116/// Alternatively, the tree can be constructed from a [`BitVec`] containing the parenthesis
117/// expression directly.
118/// This is also how trees with unbalanced parenthesis expressions can be constructed.
119///
120/// ```rust
121/// use vers_vecs::{BitVec, BpTree, Tree};
122/// let bv = BitVec::pack_sequence_u8(&[0b1101_0111, 0b0010_0100], 8);
123/// let tree = BpTree::<4>::from_bit_vector(bv);
124///
125/// let nodes = tree.dfs_iter().collect::<Vec<_>>();
126/// assert_eq!(nodes, vec![0, 1, 2, 4, 6, 7, 10, 13]);
127/// ```
128///
129/// [`RsVec`]: RsVec
130/// [`fwd_search`]: BpTree::fwd_search
131/// [`bwd_search`]: BpTree::bwd_search
132/// [`Tree`]: Tree
133/// [`IsAncestor`]: IsAncestor
134/// [`LevelTree`]: LevelTree
135/// [`SubtreeSize`]: SubtreeSize
136/// [`BpBuilder`]: BpBuilder
137/// [`BitVec`]: BitVec
138#[derive(Clone, Debug)]
139#[cfg_attr(feature = "serde", derive(serde::Serialize, serde::Deserialize))]
140pub struct BpTree<const BLOCK_SIZE: usize = DEFAULT_BLOCK_SIZE> {
141    vec: RsVec,
142    min_max_tree: MinMaxTree,
143}
144
145impl<const BLOCK_SIZE: usize> BpTree<BLOCK_SIZE> {
146    /// Construct a new `BpTree` from a given bit vector.
147    #[must_use]
148    pub fn from_bit_vector(bv: BitVec) -> Self {
149        let min_max_tree = MinMaxTree::excess_tree(&bv, BLOCK_SIZE);
150        let vec = bv.into();
151        Self { vec, min_max_tree }
152    }
153
154    /// Search for a position where the excess relative to the starting `index` is `relative_excess`.
155    /// Returns `None` if no such position exists.
156    /// The initial position is never considered in the search.
157    /// Searches forward in the bit vector.
158    ///
159    /// # Arguments
160    /// - `index`: The starting index.
161    /// - `relative_excess`: The desired relative excess value.
162    pub fn fwd_search(&self, index: usize, mut relative_excess: i64) -> Option<usize> {
163        // check for greater than or equal length minus one, because the last element
164        // won't ever have a result from fwd_search
165        if index >= (self.vec.len() - 1) {
166            return None;
167        }
168
169        let block_index = (index + 1) / BLOCK_SIZE;
170        self.fwd_search_block(index, block_index, &mut relative_excess)
171            .map_or_else(
172                |()| {
173                    // find the block that contains the desired relative excess
174                    let block = self.min_max_tree.fwd_search(block_index, relative_excess);
175
176                    // check the result block for the exact position
177                    block.and_then(|(block, mut relative_excess)| {
178                        self.fwd_search_block(block * BLOCK_SIZE - 1, block, &mut relative_excess)
179                            .ok()
180                    })
181                },
182                Some,
183            )
184    }
185
186    /// Perform the forward search within one block. If this doesn't yield a result, the caller must
187    /// continue the search in the min-max-tree.
188    ///
189    /// Returns Ok(index) if an index with the desired relative excess is found, or None(excess)
190    /// with the excess at the end of the current block if no index with the desired relative excess
191    /// is found.
192    #[inline(always)]
193    fn fwd_search_block(
194        &self,
195        start_index: usize,
196        block_index: usize,
197        relative_excess: &mut i64,
198    ) -> Result<usize, ()> {
199        let block_boundary = min((block_index + 1) * BLOCK_SIZE, self.vec.len());
200
201        // the boundary at which we can start with table lookups
202        let lookup_boundary = min(
203            (start_index + 1).div_ceil(LOOKUP_BLOCK_SIZE as usize) * LOOKUP_BLOCK_SIZE as usize,
204            block_boundary,
205        );
206        for i in start_index + 1..lookup_boundary {
207            let bit = self.vec.get_unchecked(i);
208            *relative_excess -= if bit == 1 { 1 } else { -1 };
209
210            if *relative_excess == 0 {
211                return Ok(i);
212            }
213        }
214
215        // the boundary up to which we can use table lookups
216        let upper_lookup_boundary = max(
217            lookup_boundary,
218            (block_boundary / LOOKUP_BLOCK_SIZE as usize) * LOOKUP_BLOCK_SIZE as usize,
219        );
220
221        for i in (lookup_boundary..upper_lookup_boundary).step_by(LOOKUP_BLOCK_SIZE as usize) {
222            if let Ok(idx) = process_block_fwd(
223                self.vec
224                    .get_bits_unchecked(i, LOOKUP_BLOCK_SIZE as usize)
225                    .try_into()
226                    .unwrap(),
227                relative_excess,
228            ) {
229                return Ok(i + idx as usize);
230            }
231        }
232
233        // if the upper_lookup_boundary isn't the block_boundary (which happens in non-full blocks, i.e. the last
234        // block in the vector)
235        for i in upper_lookup_boundary..block_boundary {
236            let bit = self.vec.get_unchecked(i);
237            *relative_excess -= if bit == 1 { 1 } else { -1 };
238
239            if *relative_excess == 0 {
240                return Ok(i);
241            }
242        }
243
244        Err(())
245    }
246
247    /// Search for a position where the excess relative to the starting `index` is `relative_excess`.
248    /// Returns `None` if no such position exists.
249    /// The initial position is never considered in the search.
250    /// Searches backward in the bit vector.
251    ///
252    /// # Arguments
253    /// - `index`: The starting index.
254    /// - `relative_excess`: The desired relative excess value.
255    pub fn bwd_search(&self, index: usize, mut relative_excess: i64) -> Option<usize> {
256        if index >= self.vec.len() {
257            return None;
258        }
259
260        // if the index is 0, we cant have a valid result anyway, and this would overflow the
261        // subtraction below, so we report None
262        if index == 0 {
263            return None;
264        }
265
266        // calculate the block we start searching in. It starts at index - 1, so we don't accidentally
267        // search the mM tree and immediately find `index` as the position
268        let block_index = (index - 1) / BLOCK_SIZE;
269
270        // check the current block
271        self.bwd_search_block(index, block_index, &mut relative_excess)
272            .map_or_else(
273                |()| {
274                    // find the block that contains the desired relative excess
275                    let block = self.min_max_tree.bwd_search(block_index, relative_excess);
276
277                    // check the result block for the exact position
278                    block.and_then(|(block, mut relative_excess)| {
279                        self.bwd_search_block((block + 1) * BLOCK_SIZE, block, &mut relative_excess)
280                            .ok()
281                    })
282                },
283                Some,
284            )
285    }
286
287    /// Perform the backward search within one block. If this doesn't yield a result, the caller must
288    /// continue the search in the min-max-tree.
289    ///
290    /// Returns Ok(index) if an index with the desired relative excess is found, or None(excess)
291    /// with the excess at the end of the current block if no index with the desired relative excess
292    /// is found.
293    #[inline(always)]
294    fn bwd_search_block(
295        &self,
296        start_index: usize,
297        block_index: usize,
298        relative_excess: &mut i64,
299    ) -> Result<usize, ()> {
300        let block_boundary = min(block_index * BLOCK_SIZE, self.vec.len());
301
302        // the boundary at which we can start with table lookups
303        let lookup_boundary = max(
304            ((start_index - 1) / LOOKUP_BLOCK_SIZE as usize) * LOOKUP_BLOCK_SIZE as usize,
305            block_boundary,
306        );
307        for i in (lookup_boundary..start_index).rev() {
308            let bit = self.vec.get_unchecked(i);
309            *relative_excess += if bit == 1 { 1 } else { -1 };
310
311            if *relative_excess == 0 {
312                return Ok(i);
313            }
314        }
315
316        for i in (block_boundary..lookup_boundary)
317            .step_by(LOOKUP_BLOCK_SIZE as usize)
318            .rev()
319        {
320            if let Ok(idx) = process_block_bwd(
321                self.vec
322                    .get_bits_unchecked(i, LOOKUP_BLOCK_SIZE as usize)
323                    .try_into()
324                    .unwrap(),
325                relative_excess,
326            ) {
327                return Ok(i + idx as usize);
328            }
329        }
330
331        Err(())
332    }
333
334    /// Find the position of the matching closing parenthesis for the opening parenthesis at `index`.
335    /// If the bit at `index` is not an opening parenthesis, the result is meaningless.
336    /// If there is no matching closing parenthesis, `None` is returned.
337    #[must_use]
338    pub fn close(&self, index: usize) -> Option<usize> {
339        if index >= self.vec.len() {
340            return None;
341        }
342
343        self.fwd_search(index, -1)
344    }
345
346    /// Find the position of the matching opening parenthesis for the closing parenthesis at `index`.
347    /// If the bit at `index` is not a closing parenthesis, the result is meaningless.
348    /// If there is no matching opening parenthesis, `None` is returned.
349    #[must_use]
350    pub fn open(&self, index: usize) -> Option<usize> {
351        if index >= self.vec.len() {
352            return None;
353        }
354
355        self.bwd_search(index, -1)
356    }
357
358    /// Find the position of the opening parenthesis that encloses the position `index`.
359    /// This works regardless of whether the bit at `index` is an opening or closing parenthesis.
360    /// If there is no enclosing parenthesis, `None` is returned.
361    #[must_use]
362    pub fn enclose(&self, index: usize) -> Option<usize> {
363        if index >= self.vec.len() {
364            return None;
365        }
366
367        self.bwd_search(
368            index,
369            if self.vec.get_unchecked(index) == 1 {
370                -1
371            } else {
372                -2
373            },
374        )
375    }
376
377    /// Get the excess of open parentheses up to and including the position `index`.
378    /// The excess is the number of open parentheses minus the number of closing parentheses.
379    /// If `index` is out of bounds, the total excess of the parentheses expression is returned.
380    #[must_use]
381    pub fn excess(&self, index: usize) -> i64 {
382        debug_assert!(index < self.vec.len(), "Index out of bounds");
383        self.vec.rank1(index + 1) as i64 - self.vec.rank0(index + 1) as i64
384    }
385
386    /// Iterate over the nodes of the tree.
387    /// The iterator yields the nodes in depth-first (pre-)order.
388    /// This method is an alias for [`dfs_iter`].
389    ///
390    /// If the tree is unbalanced, the iterator returns the node handles in the order they appear in
391    /// the parenthesis expression, and it will return handles that don't have a matching closing
392    /// parenthesis.
393    ///
394    /// [`dfs_iter`]: BpTree::dfs_iter
395    pub fn iter(
396        &self,
397    ) -> impl Iterator<Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle> + use<'_, BLOCK_SIZE> {
398        self.dfs_iter()
399    }
400
401    /// Iterate over the nodes of the tree in depth-first (pre-)order.
402    /// This is the most efficient way to iterate over all nodes of the tree.
403    ///
404    /// If the tree is unbalanced, the iterator returns the node handles in the order they appear in
405    /// the parenthesis expression, and it will return handles that don't have a matching closing
406    /// parenthesis.
407    pub fn dfs_iter(
408        &self,
409    ) -> impl Iterator<Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle> + use<'_, BLOCK_SIZE> {
410        self.vec.iter1()
411    }
412
413    /// Iterate over the nodes of a valid tree in depth-first (post-)order.
414    /// This is slower than the pre-order iteration.
415    ///
416    /// # Panics
417    /// The iterator may panic at any point if the parenthesis expression is unbalanced.
418    pub fn dfs_post_iter(
419        &self,
420    ) -> impl Iterator<Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle> + use<'_, BLOCK_SIZE> {
421        self.vec.iter0().map(|n| self.open(n).unwrap())
422    }
423
424    /// Iterate over a subtree rooted at `node` in depth-first (pre-)order.
425    /// The iteration starts with the node itself.
426    ///
427    /// Calling this method on an invalid node handle, or an unbalanced parenthesis expression,
428    /// will produce an iterator over an unspecified subset of nodes.
429    pub fn subtree_iter(
430        &self,
431        node: <BpTree<BLOCK_SIZE> as Tree>::NodeHandle,
432    ) -> impl Iterator<Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle> + use<'_, BLOCK_SIZE> {
433        debug_assert!(
434            self.vec.get(node) == Some(OPEN_PAREN),
435            "Node handle is invalid"
436        );
437
438        let index = self.vec.rank1(node);
439        let close = self.close(node).unwrap_or(node);
440        let subtree_size = self.vec.rank1(close) - index;
441
442        self.vec.iter1().skip(index).take(subtree_size)
443    }
444
445    /// Iterate over a subtree rooted at `node` in depth-first (post-)order.
446    /// This is slower than the pre-order iteration.
447    /// The iteration ends with the node itself.
448    ///
449    /// # Panics
450    /// Calling this method on an invalid node handle, or an unbalanced parenthesis expression,
451    /// will produce an iterator over an unspecified subset of nodes, or panic either during
452    /// construction or iteration.
453    pub fn subtree_post_iter(
454        &self,
455        node: <BpTree<BLOCK_SIZE> as Tree>::NodeHandle,
456    ) -> impl Iterator<Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle> + use<'_, BLOCK_SIZE> {
457        debug_assert!(
458            self.vec.get(node) == Some(OPEN_PAREN),
459            "Node handle is invalid"
460        );
461
462        let index = self.vec.rank0(node);
463        let close = self.close(node).unwrap_or(node);
464        let subtree_size = self.vec.rank0(close) + 1 - index;
465
466        self.vec
467            .iter0()
468            .skip(index)
469            .take(subtree_size)
470            .map(|n| self.open(n).unwrap())
471    }
472
473    /// Iterate over the children of a node in the tree.
474    /// The iterator yields the children in the order they appear in the parenthesis expression.
475    /// If the node is a leaf, the iterator is empty.
476    /// If the node is not a valid node handle, or the tree is unbalanced,
477    /// the iterator will produce an unspecified subset of the tree's nodes.
478    pub fn children(
479        &self,
480        node: <BpTree<BLOCK_SIZE> as Tree>::NodeHandle,
481    ) -> impl Iterator<Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle> + use<'_, BLOCK_SIZE> {
482        debug_assert!(
483            self.vec.get(node) == Some(OPEN_PAREN),
484            "Node handle is invalid"
485        );
486
487        ChildrenIter::<BLOCK_SIZE, true>::new(self, node)
488    }
489
490    /// Iterate over the children of a node in the tree in reverse order.
491    /// The iterator yields the children in the reverse order they appear in the parenthesis expression.
492    /// If the node is a leaf, the iterator is empty.
493    /// If the node is not a valid node handle, or the tree is unbalanced,
494    /// the iterator will produce an unspecified subset of the tree's nodes.
495    pub fn rev_children(
496        &self,
497        node: <BpTree<BLOCK_SIZE> as Tree>::NodeHandle,
498    ) -> impl Iterator<Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle> + use<'_, BLOCK_SIZE> {
499        debug_assert!(
500            self.vec.get(node) == Some(OPEN_PAREN),
501            "Node handle is invalid"
502        );
503
504        ChildrenIter::<BLOCK_SIZE, false>::new(self, node)
505    }
506
507    /// Transform the tree into a [`RsVec`] containing the balanced parenthesis expression.
508    /// This consumes the tree and returns the underlying bit vector with the rank and select
509    /// support structure.
510    /// The remaining min-max-tree support structure of the `BpTree` is discarded.
511    /// Since the tree is innately immutable, this is the only way to access the underlying bit
512    /// vector for potential modification.
513    /// Modification requires turning the `RsVec` back into a `BitVec`, discarding the rank and select
514    /// support structure, however.
515    ///
516    /// # Examples
517    /// ```rust
518    /// use vers_vecs::{BitVec, RsVec, BpTree, Tree};
519    ///
520    /// let bv = BitVec::pack_sequence_u8(&[0b1101_0111, 0b0010_0100], 8);
521    /// let tree = BpTree::<4>::from_bit_vector(bv);
522    /// assert_eq!(tree.size(), 8);
523    ///
524    /// let rs_vec = tree.into_parentheses_vec();
525    /// let mut bv = rs_vec.into_bit_vec();
526    ///
527    /// bv.flip_bit(15);
528    /// bv.append_bits(0, 2);
529    /// let tree = BpTree::<4>::from_bit_vector(bv);
530    /// assert_eq!(tree.size(), 9);
531    /// ```
532    #[must_use]
533    pub fn into_parentheses_vec(self) -> RsVec {
534        self.vec
535    }
536
537    /// Returns the number of bytes used on the heap for this tree. This does not include
538    /// allocated space that is not used (e.g. by the allocation behavior of `Vec`).
539    #[must_use]
540    pub fn heap_size(&self) -> usize {
541        self.vec.heap_size() + self.min_max_tree.heap_size()
542    }
543}
544
545impl<const BLOCK_SIZE: usize> Tree for BpTree<BLOCK_SIZE> {
546    type NodeHandle = usize;
547
548    fn root(&self) -> Option<Self::NodeHandle> {
549        if self.vec.is_empty() {
550            None
551        } else {
552            Some(0)
553        }
554    }
555
556    fn parent(&self, node: Self::NodeHandle) -> Option<Self::NodeHandle> {
557        debug_assert!(
558            self.vec.get(node) == Some(OPEN_PAREN),
559            "Node handle is invalid"
560        );
561
562        self.enclose(node)
563    }
564
565    fn first_child(&self, node: Self::NodeHandle) -> Option<Self::NodeHandle> {
566        debug_assert!(
567            self.vec.get(node) == Some(OPEN_PAREN),
568            "Node handle is invalid"
569        );
570
571        if let Some(bit) = self.vec.get(node + 1) {
572            if bit == OPEN_PAREN {
573                return Some(node + 1);
574            }
575        }
576
577        None
578    }
579
580    fn next_sibling(&self, node: Self::NodeHandle) -> Option<Self::NodeHandle> {
581        debug_assert!(
582            self.vec.get(node) == Some(OPEN_PAREN),
583            "Node handle is invalid"
584        );
585        self.close(node).and_then(|i| {
586            self.vec
587                .get(i + 1)
588                .and_then(|bit| if bit == OPEN_PAREN { Some(i + 1) } else { None })
589        })
590    }
591
592    fn previous_sibling(&self, node: Self::NodeHandle) -> Option<Self::NodeHandle> {
593        debug_assert!(
594            self.vec.get(node) == Some(OPEN_PAREN),
595            "Node handle is invalid"
596        );
597        if node == 0 {
598            None
599        } else {
600            self.vec.get(node - 1).and_then(|bit| {
601                if bit == CLOSE_PAREN {
602                    self.open(node - 1)
603                } else {
604                    None
605                }
606            })
607        }
608    }
609
610    fn last_child(&self, node: Self::NodeHandle) -> Option<Self::NodeHandle> {
611        debug_assert!(
612            self.vec.get(node) == Some(OPEN_PAREN),
613            "Node handle is invalid"
614        );
615        self.vec.get(node + 1).and_then(|bit| {
616            if bit == OPEN_PAREN {
617                if let Some(i) = self.close(node) {
618                    self.open(i - 1)
619                } else {
620                    None
621                }
622            } else {
623                None
624            }
625        })
626    }
627
628    fn node_index(&self, node: Self::NodeHandle) -> usize {
629        debug_assert!(
630            self.vec.get(node) == Some(OPEN_PAREN),
631            "Node handle is invalid"
632        );
633        self.vec.rank1(node)
634    }
635
636    fn node_handle(&self, index: usize) -> Self::NodeHandle {
637        self.vec.select1(index)
638    }
639
640    fn is_leaf(&self, node: Self::NodeHandle) -> bool {
641        debug_assert!(
642            self.vec.get(node) == Some(OPEN_PAREN),
643            "Node handle is invalid"
644        );
645        self.vec.get(node + 1) == Some(CLOSE_PAREN)
646    }
647
648    fn depth(&self, node: Self::NodeHandle) -> u64 {
649        debug_assert!(
650            self.vec.get(node) == Some(OPEN_PAREN),
651            "Node handle is invalid"
652        );
653        let excess: u64 = self.excess(node).try_into().unwrap_or(0);
654        excess.saturating_sub(1)
655    }
656
657    fn size(&self) -> usize {
658        self.vec.rank1(self.vec.len())
659    }
660
661    fn is_empty(&self) -> bool {
662        self.vec.is_empty()
663    }
664}
665
666impl<const BLOCK_SIZE: usize> IsAncestor for BpTree<BLOCK_SIZE> {
667    fn is_ancestor(
668        &self,
669        ancestor: Self::NodeHandle,
670        descendant: Self::NodeHandle,
671    ) -> Option<bool> {
672        debug_assert!(
673            self.vec.get(ancestor) == Some(OPEN_PAREN),
674            "Node handle is invalid"
675        );
676        debug_assert!(
677            self.vec.get(descendant) == Some(OPEN_PAREN),
678            "Node handle is invalid"
679        );
680
681        self.close(ancestor)
682            .map(|closing| ancestor <= descendant && descendant < closing)
683    }
684}
685
686impl<const BLOCK_SIZE: usize> LevelTree for BpTree<BLOCK_SIZE> {
687    fn level_ancestor(&self, node: Self::NodeHandle, level: u64) -> Option<Self::NodeHandle> {
688        if level == 0 {
689            return Some(node);
690        }
691
692        #[allow(clippy::cast_possible_wrap)]
693        // if the level exceeds 2^63, we accept that the result is wrong
694        self.bwd_search(node, -(level as i64))
695    }
696
697    fn level_next(&self, node: Self::NodeHandle) -> Option<Self::NodeHandle> {
698        self.fwd_search(self.close(node)?, 1)
699    }
700
701    fn level_prev(&self, node: Self::NodeHandle) -> Option<Self::NodeHandle> {
702        self.open(self.bwd_search(node, 1)?)
703    }
704
705    fn level_leftmost(&self, level: u64) -> Option<Self::NodeHandle> {
706        // fwd_search doesn't support returning the input position
707        if level == 0 {
708            return Some(0);
709        }
710
711        #[allow(clippy::cast_possible_wrap)]
712        // if the level exceeds 2^63, we accept that the result is wrong
713        self.fwd_search(0, level as i64)
714    }
715
716    fn level_rightmost(&self, level: u64) -> Option<Self::NodeHandle> {
717        #[allow(clippy::cast_possible_wrap)]
718        // if the level exceeds 2^63, we accept that the result is wrong
719        self.open(self.bwd_search(self.size() * 2 - 1, level as i64)?)
720    }
721}
722
723impl<const BLOCK_SIZE: usize> SubtreeSize for BpTree<BLOCK_SIZE> {
724    fn subtree_size(&self, node: Self::NodeHandle) -> Option<usize> {
725        debug_assert!(
726            self.vec.get(node) == Some(OPEN_PAREN),
727            "Node handle is invalid"
728        );
729
730        self.close(node)
731            .map(|c| self.vec.rank1(c) - self.vec.rank1(node))
732    }
733}
734
735impl<const BLOCK_SIZE: usize> IntoIterator for BpTree<BLOCK_SIZE> {
736    type Item = <BpTree<BLOCK_SIZE> as Tree>::NodeHandle;
737    type IntoIter = SelectIntoIter<false>;
738
739    fn into_iter(self) -> Self::IntoIter {
740        self.vec.into_iter1()
741    }
742}
743
744impl<const BLOCK_SIZE: usize> From<BitVec> for BpTree<BLOCK_SIZE> {
745    fn from(bv: BitVec) -> Self {
746        Self::from_bit_vector(bv)
747    }
748}
749
750impl<const BLOCK_SIZE: usize> From<BpTree<BLOCK_SIZE>> for BitVec {
751    fn from(value: BpTree<BLOCK_SIZE>) -> Self {
752        value.into_parentheses_vec().into_bit_vec()
753    }
754}
755
756impl<const BLOCK_SIZE: usize> From<BpTree<BLOCK_SIZE>> for RsVec {
757    fn from(value: BpTree<BLOCK_SIZE>) -> Self {
758        value.into_parentheses_vec()
759    }
760}
761
762/// An iterator over the children of a node.
763/// Calls to `next` return the next child node handle in the order they appear in the parenthesis
764/// expression.
765struct ChildrenIter<'a, const BLOCK_SIZE: usize, const FORWARD: bool> {
766    tree: &'a BpTree<BLOCK_SIZE>,
767    current_sibling: Option<usize>,
768}
769
770impl<'a, const BLOCK_SIZE: usize, const FORWARD: bool> ChildrenIter<'a, BLOCK_SIZE, FORWARD> {
771    fn new(tree: &'a BpTree<BLOCK_SIZE>, node: usize) -> Self {
772        Self {
773            tree,
774            current_sibling: if FORWARD {
775                tree.first_child(node)
776            } else {
777                tree.last_child(node)
778            },
779        }
780    }
781}
782
783impl<const BLOCK_SIZE: usize, const FORWARD: bool> Iterator
784    for ChildrenIter<'_, BLOCK_SIZE, FORWARD>
785{
786    type Item = usize;
787
788    fn next(&mut self) -> Option<Self::Item> {
789        let current = self.current_sibling?;
790        let next = if FORWARD {
791            self.tree.next_sibling(current)
792        } else {
793            self.tree.previous_sibling(current)
794        };
795        self.current_sibling = next;
796        Some(current)
797    }
798}
799
800impl<const BLOCK_SIZE: usize, const FORWARD: bool> FusedIterator
801    for ChildrenIter<'_, BLOCK_SIZE, FORWARD>
802{
803}
804
805#[cfg(test)]
806mod tests;