oxidd_core/
lib.rs

1//! Collection of fundamental traits and types to represent decision diagrams
2//!
3//! # Overview
4//!
5//! One of the most central traits is [`Manager`]. The manager is responsible
6//! for storing the nodes of a decision diagram ([`InnerNode`]s and terminal
7//! nodes) and provides [`Edge`]s to identify them.
8//!
9//! From the user's perspective, [`Function`][function::Function] is very
10//! important. A function is some kind of external reference to a node and is
11//! the basis for assigning semantics to nodes and providing operations such as
12//! applying connectives of boolean logic.
13
14#![warn(missing_docs)]
15#![deny(unsafe_op_in_unsafe_fn)]
16#![allow(clippy::double_must_use)]
17// Explicitly writing out `'id` lifetimes possibly makes some code easier to
18// read.
19#![allow(clippy::needless_lifetimes)]
20// `match` syntax may be easier to read
21#![allow(clippy::manual_map)]
22
23use std::borrow::Borrow;
24use std::hash::Hash;
25
26use util::AllocResult;
27use util::Borrowed;
28use util::DropWith;
29use util::NodeSet;
30
31pub mod function;
32pub mod util;
33
34/// Manager reference
35///
36/// The methods of this trait synchronize accesses to the manager: In a
37/// concurrent setting, a manager has some kind of read/write lock, and
38/// [`Self::with_manager_shared()`] / [`Self::with_manager_exclusive()`] acquire
39/// this lock accordingly. In a sequential implementation, a
40/// [`RefCell`][std::cell::RefCell] or the like may be used instead of lock.
41pub trait ManagerRef: Clone + Eq + Hash + for<'a, 'id> From<&'a Self::Manager<'id>> {
42    /// Type of the associated manager
43    ///
44    /// For more details on why this type is generic over `'id`, see the
45    /// documentation of [`Function::Manager`][function::Function::Manager].
46    type Manager<'id>: Manager;
47
48    /// Obtain a shared manager reference
49    ///
50    /// Locking behavior: acquires the manager's lock for shared access.
51    fn with_manager_shared<F, T>(&self, f: F) -> T
52    where
53        F: for<'id> FnOnce(&Self::Manager<'id>) -> T;
54
55    /// Obtain an exclusive manager reference
56    ///
57    /// Locking behavior: acquires the manager's lock for exclusive access.
58    fn with_manager_exclusive<F, T>(&self, f: F) -> T
59    where
60        F: for<'id> FnOnce(&mut Self::Manager<'id>) -> T;
61}
62
63/// Reduction rules for decision diagrams
64///
65/// This trait is intended to be implemented on a zero-sized type. Refer to
66/// [`DiagramRules::reduce()`] for an example implementation.
67pub trait DiagramRules<E: Edge, N: InnerNode<E>, T> {
68    /// Iterator created by [`DiagramRules::cofactors()`]
69    type Cofactors<'a>: Iterator<Item = Borrowed<'a, E>>
70    where
71        E: 'a,
72        N: 'a;
73
74    /// Apply the reduction rule(s)
75    ///
76    /// Besides uniqueness of nodes (there are no two nodes with the same
77    /// children at the same level), decision diagrams typically impose some
78    /// other reduction rules. The former is provided by the [`Manager`] /
79    /// [`LevelView`]s, the latter is implemented here.
80    ///
81    /// The implementation is responsible for consuming the entire `children`
82    /// iterator and dropping unused edges.
83    ///
84    /// # Example implementation
85    ///
86    /// In binary decision diagrams (BDDs), there are no nodes with equal
87    /// children. An implementation might look like this:
88    ///
89    /// ```
90    /// # use oxidd_core::{*, util::BorrowedEdgeIter};
91    /// struct BDDRules;
92    /// impl<E: Edge, N: InnerNode<E>, T> DiagramRules<E, N, T> for BDDRules {
93    ///     type Cofactors<'a> = N::ChildrenIter<'a> where N: 'a, E: 'a;
94    ///
95    ///     fn reduce<M: Manager<Edge = E, InnerNode = N, Terminal = T>>(
96    ///         manager: &M,
97    ///         level: LevelNo,
98    ///         children: impl IntoIterator<Item = E>,
99    ///     ) -> ReducedOrNew<E, N> {
100    ///         let mut it = children.into_iter();
101    ///         let f0 = it.next().unwrap();
102    ///         let f1 = it.next().unwrap();
103    ///         debug_assert!(it.next().is_none());
104    ///
105    ///         if f0 == f1 {
106    ///             manager.drop_edge(f1);
107    ///             ReducedOrNew::Reduced(f0)
108    ///         } else {
109    ///             ReducedOrNew::New(InnerNode::new(level, [f0, f1]), Default::default())
110    ///         }
111    ///     }
112    ///
113    ///     fn cofactors(_tag: E::Tag, node: &N) -> Self::Cofactors<'_> {
114    ///         node.children()
115    ///     }
116    /// }
117    /// ```
118    ///
119    /// Note that we assume no complemented edges, hence the cofactor iterator
120    /// is just `node.children()`.
121    ///
122    /// The implementation assumes that there are no panics, otherwise it would
123    /// leak some edges. It might be a bit better to use
124    /// [`EdgeDropGuard`][util::EdgeDropGuard]s, but this would make the code
125    /// more clumsy, and our assumption is usually fair enough.
126    fn reduce<M: Manager<Edge = E, InnerNode = N, Terminal = T>>(
127        manager: &M,
128        level: LevelNo,
129        children: impl IntoIterator<Item = E>,
130    ) -> ReducedOrNew<E, N>;
131
132    /// Get the cofactors of `node` assuming an incoming edge with `tag`
133    ///
134    /// In some diagram types, this is the same as [`InnerNode::children()`].
135    /// However, in a binary decision diagram with complement edges, we need to
136    /// respect the tag of the incoming edge: If the incoming edge is
137    /// complemented, then we need to complement the outgoing edges as well.
138    fn cofactors(tag: E::Tag, node: &N) -> Self::Cofactors<'_>;
139
140    /// Get the `n`-th cofactor of `node` assuming an incoming edge with `tag`
141    ///
142    /// This is equivalent to `Self::cofactors(tag, node).nth(n).unwrap()`.
143    #[inline]
144    fn cofactor(tag: E::Tag, node: &N, n: usize) -> Borrowed<E> {
145        Self::cofactors(tag, node).nth(n).expect("out of range")
146    }
147}
148
149/// Result of the attempt to create a new node
150///
151/// Before actually creating a new node, reduction rules should be applied
152/// (see [`DiagramRules::reduce()`]). If a reduction was applied, then
153/// [`DiagramRules::reduce()`] returns the `Reduced` variant, otherwise the
154/// `New` variant.
155pub enum ReducedOrNew<E: Edge, N: InnerNode<E>> {
156    /// A reduction rule was applied
157    Reduced(E),
158    /// The node is new. After inserting it into the manager, the edge should be
159    /// tagged with the given tag.
160    New(N, E::Tag),
161}
162
163impl<E: Edge, N: InnerNode<E>> ReducedOrNew<E, N> {
164    /// Insert `self` into `manager` and `unique_table` at the given `level` if
165    /// it is `New`, otherwise return the `Reduced` edge.
166    ///
167    /// `level` must agree with the level used for creating the node, and must
168    /// be strictly above (i.e. less than) the children's levels.
169    #[must_use]
170    #[inline(always)]
171    pub fn then_insert<M>(self, manager: &M, level: LevelNo) -> AllocResult<E>
172    where
173        M: Manager<InnerNode = N, Edge = E>,
174    {
175        match self {
176            ReducedOrNew::Reduced(e) => Ok(e),
177            ReducedOrNew::New(node, tag) => {
178                debug_assert_ne!(level, LevelNo::MAX);
179                debug_assert!(node.check_level(|l| l == level));
180                debug_assert!(node.children().all(|c| {
181                    if let Node::Inner(node) = manager.get_node(&*c) {
182                        node.check_level(|l| level < l)
183                    } else {
184                        true
185                    }
186                }));
187
188                let edge = manager.level(level).get_or_insert(node)?;
189                Ok(edge.with_tag_owned(tag))
190            }
191        }
192    }
193}
194
195/// Node in a decision diagram
196///
197/// [`Eq`] and [`Hash`] should consider the children only, in particular no
198/// level information. This means that if `Self` implements [`HasLevel`],
199/// [`HasLevel::set_level()`] may be called while the node is present in a
200/// [`LevelView`]. This is not the case for [`InnerNode::set_child()`]: The user
201/// must remove the node from the [`LevelView`] before setting the children (and
202/// re-insert it afterwards).
203#[must_use]
204pub trait InnerNode<E: Edge>: Sized + Eq + Hash + DropWith<E> {
205    /// The node's arity (upper bound)
206    const ARITY: usize;
207
208    /// Iterator over children of an inner node
209    type ChildrenIter<'a>: ExactSizeIterator<Item = Borrowed<'a, E>>
210    where
211        Self: 'a,
212        E: 'a;
213
214    /// Create a new node
215    ///
216    /// Note that this does not apply any reduction rules. A node type that does
217    /// not store levels internally (does not implement [`HasLevel`]) may simply
218    /// ignore the `level` parameter.
219    ///
220    /// Panics if `children`'s length does not match the node's requirements
221    /// (typically, the length should be [`Self::ARITY`], but some node types
222    /// may deviate from that).
223    #[must_use]
224    fn new(level: LevelNo, children: impl IntoIterator<Item = E>) -> Self;
225
226    /// Returns the result of `check` applied to the node's level in case this
227    /// node type stores levels, otherwise returns `true`.
228    ///
229    /// Use [`HasLevel::level()`] if you require your nodes to store the level
230    /// number and want to get the level number.
231    fn check_level(&self, check: impl FnOnce(LevelNo) -> bool) -> bool;
232
233    /// Panics if the node types stores a level and the node's level is not
234    /// `level`
235    fn assert_level_matches(&self, level: LevelNo);
236
237    /// Get the children of this node as an iterator
238    #[must_use]
239    fn children(&self) -> Self::ChildrenIter<'_>;
240
241    /// Get the `n`-th child of this node
242    fn child(&self, n: usize) -> Borrowed<E>;
243
244    /// Set the `n`-th child of this node
245    ///
246    /// Returns the previous `n`-th child.
247    ///
248    /// Note that this function may also change the node's hash value etc., so
249    /// in case the node is stored in a hash table ([`LevelView`]), it needs to
250    /// be removed before calling this method.
251    ///
252    /// Panics if the node does not have an `n`-th child.
253    ///
254    /// # Safety
255    ///
256    /// The caller must have exclusive access to the node. In the first place,
257    /// this is granted by acquiring an exclusive manager lock
258    /// ([`Function::with_manager_exclusive()`][function::Function::with_manager_exclusive] or
259    /// [`ManagerRef::with_manager_exclusive()`]). However, exclusive access to
260    /// some nodes may be delegated to other threads (which is the reason why we
261    /// only require a shared and not a mutable manager reference). Furthermore,
262    /// there must not be a borrowed child (obtained via
263    /// [`InnerNode::children()`]).
264    #[must_use = "call `Manager::drop_edge()` if you don't need the previous edge"]
265    unsafe fn set_child(&self, n: usize, child: E) -> E;
266
267    /// Get the node's reference count
268    ///
269    /// This ignores all internal references used by the manager
270    /// implementation.
271    #[must_use]
272    fn ref_count(&self) -> usize;
273}
274
275/// Level number type
276///
277/// Levels with lower numbers are located at the top of the diagram, higher
278/// numbers at the bottom. [`LevelNo::MAX`] is reserved for terminal nodes.
279/// Adjacent levels have adjacent level numbers.
280pub type LevelNo = u32;
281
282/// Atomic version of [`LevelNo`]
283pub type AtomicLevelNo = std::sync::atomic::AtomicU32;
284
285/// Trait for nodes that have a level
286///
287/// Quasi-reduced BDDs, for instance, do not need the level information stored
288/// in their nodes, so there is no need to implement this trait.
289///
290/// Implementors should also implement [`InnerNode`]. If `Self` is [`Sync`],
291/// then the level number should be implemented using [`AtomicLevelNo`]. In
292/// particular, concurrent calls to [`Self::level()`] and [`Self::set_level()`]
293/// must not lead to data races.
294///
295/// # Safety
296///
297/// 1. A node in a [`LevelView`] with level number L has level number L (i.e.
298///    `self.level()` returns L).
299/// 2. [`InnerNode::check_level()`] with a check `c` must return
300///    `c(self.level())`. Similarly, [`InnerNode::assert_level_matches()`] must
301///    panic if the level does not match.
302///
303/// These conditions are crucial to enable concurrent level swaps as part of
304/// reordering (see the `oxidd-reorder` crate): The algorithm iterates over the
305/// nodes at the upper level and needs to know whether a node is part of the
306/// level directly below it. The procedure only has access to nodes at these two
307/// levels, hence it must rely on the level information for SAFETY.
308///
309/// Note that invariant 1 may be broken by [`HasLevel::set_level()`] and
310/// [`LevelView::swap()`]; the caller of these functions is responsible to
311/// re-establish the invariant.
312pub unsafe trait HasLevel {
313    /// Get the node's level
314    #[must_use]
315    fn level(&self) -> LevelNo;
316
317    /// Set the node's level
318    ///
319    /// # Safety
320    ///
321    /// This method may break SAFETY invariant 1 of the [`HasLevel`] trait: A
322    /// node in a [`LevelView`] with level number L has level number L (i.e.
323    /// `self.level()` returns L). The caller is responsible to re-establish the
324    /// invariant. (Make sure that the calling code is exception-safe!)
325    unsafe fn set_level(&self, level: LevelNo);
326}
327
328/// Node identifier returned by [`Edge::node_id()`]
329///
330/// The most significant bit is reserved, i.e. `NodeID`s must (normally) be less
331/// than `1 << (NodeID::BITS - 1)`
332pub type NodeID = usize;
333
334/// Edge in a decision diagram
335///
336/// Generally speaking, an edge is the directed connection between two nodes,
337/// with some kind of annotation. In a binary decision diagram with complement
338/// edges, there are the "true" edges as well as the normal and the complemented
339/// "false" edges. When considering a single edge, it usually is not so
340/// important whether this edge is a "true" or "false" edge; we can simply have
341/// distinguishable "slots" in the source node. In contrast, whether an edge is
342/// complemented or not has a much greater influence on the meaning of an edge.
343///
344/// In a decision diagram, obtaining the source of an edge is usually not so
345/// important, hence this trait does not provide such functionality. This means
346/// that an edge can (more or less) be implemented as a (tagged) pointer to the
347/// target node.
348///
349/// This trait requires implementors to also implement [`Ord`]. Edges should be
350/// considered equal if and only if they point to the same node with the same
351/// tag. Besides that, there are no further restrictions. The order implemented
352/// for [`Ord`] can be an arbitrary, fixed order (e.g. using addresses of the
353/// nodes). The main idea of this is to give the set `{f, g}` of two edges `f`
354/// and `g` a unique tuple/array representation.
355#[must_use]
356pub trait Edge: Sized + Ord + Hash {
357    /// Edge tag
358    ///
359    /// For instance, an edge tag can be used to mark an edge as complemented.
360    ///
361    /// If the decision diagram does not need any special edge tags, this can
362    /// simply be `()`.
363    type Tag: Tag;
364
365    /// Turn a reference into a borrowed handle
366    fn borrowed(&self) -> Borrowed<Self>;
367    /// Get a version of this [`Edge`] with the given tag
368    ///
369    /// Refer to [`Borrowed::edge_with_tag()`] for cases in which this method
370    /// cannot be used due to lifetime restrictions.
371    fn with_tag(&self, tag: Self::Tag) -> Borrowed<Self>;
372    /// Get a version of this [`Edge`] with the given tag
373    fn with_tag_owned(self, tag: Self::Tag) -> Self;
374
375    /// Get the [`Tag`] of this [`Edge`]
376    fn tag(&self) -> Self::Tag;
377
378    /// Returns some unique identifier for the node, e.g. for I/O purposes
379    fn node_id(&self) -> NodeID;
380}
381
382/// Trait for tags that can be attached to pointers (e.g. edges, see
383/// [`Edge::Tag`])
384///
385/// This trait is automatically implemented for types that implement [`Eq`],
386/// [`Default`], and [`Countable`].
387pub trait Tag: Sized + Copy + Eq + Default + Countable {}
388impl<T: Eq + Default + Countable> Tag for T {}
389
390/// Types whose values can be counted, i.e. there is a bijection between the
391/// values of the type and the range `0..=MAX_VALUE`.
392///
393/// This is mainly intended to be implemented on `enum`s. In most cases, you can
394/// simply derive it.
395///
396/// # Safety
397///
398/// [`Countable::as_usize()`] and [`Countable::from_usize()`] must form a
399/// bijection between the values of type `Self` and `0..=MAX_VALUE`, more
400/// formally: For all `t: Self` it must hold that
401/// `Self::from_usize(t.as_usize()) == t`.
402/// For all `u: usize` such that `t.as_usize() == u` for some `t: Self`,
403/// `Self::from_usize(u).as_usize() == u` must be true. Furthermore,
404/// `t.as_usize() <= Self::MAX_VALUE` must hold.
405///
406/// This trait is marked unsafe because violating any invariant of the above may
407/// e.g. result in out-of-bounds accesses.
408pub unsafe trait Countable: Sized + Copy {
409    /// Maximum value returned by `self.as_usize()`.
410    const MAX_VALUE: usize;
411
412    /// Convert `self` into a `usize`.
413    #[must_use]
414    fn as_usize(self) -> usize;
415    /// Convert the given `value` into an instance of `Self`.
416    ///
417    /// May panic if an invalid value is passed, or return some default value.
418    #[must_use]
419    fn from_usize(value: usize) -> Self;
420}
421
422// SAFETY: There is a bijection for all values of type `()` and `0..=MAX_VALUE`.
423unsafe impl Countable for () {
424    const MAX_VALUE: usize = 0;
425
426    #[inline]
427    fn as_usize(self) -> usize {
428        0
429    }
430
431    #[inline]
432    fn from_usize(_value: usize) -> Self {}
433}
434
435// SAFETY: There is a bijection for all values of type `bool` and `0..=1`.
436unsafe impl Countable for bool {
437    const MAX_VALUE: usize = 1;
438
439    #[inline]
440    fn as_usize(self) -> usize {
441        self as usize
442    }
443
444    #[inline]
445    fn from_usize(value: usize) -> Self {
446        value != 0
447    }
448}
449
450/// Either an inner or a terminal node
451pub enum Node<'a, M: Manager + 'a> {
452    #[allow(missing_docs)]
453    Inner(&'a M::InnerNode),
454    #[allow(missing_docs)]
455    Terminal(M::TerminalRef<'a>),
456}
457
458impl<'a, M: Manager + 'a> Clone for Node<'a, M> {
459    fn clone(&self) -> Self {
460        *self
461    }
462}
463impl<'a, M: Manager + 'a> Copy for Node<'a, M> {}
464
465impl<'a, M: Manager> Node<'a, M> {
466    /// Get the reference count of the underlying node
467    #[must_use]
468    #[inline]
469    pub fn ref_count(self) -> usize {
470        match self {
471            Node::Inner(node) => node.ref_count(),
472            Node::Terminal(_) => usize::MAX, /* TODO: should we support concrete reference counts
473                                              * for terminals? */
474        }
475    }
476}
477
478impl<'a, M: Manager> Node<'a, M>
479where
480    M::InnerNode: HasLevel,
481{
482    /// Get the level of the underlying node (`LevelNo::MAX` for terminals)
483    #[must_use]
484    #[inline]
485    pub fn level(self) -> LevelNo {
486        match self {
487            Node::Inner(node) => node.level(),
488            Node::Terminal(_) => LevelNo::MAX,
489        }
490    }
491}
492
493impl<'a, M: Manager> Node<'a, M> {
494    /// Unwrap the inner node
495    ///
496    /// Panics if `self` is a terminal.
497    #[must_use]
498    #[track_caller]
499    #[inline]
500    pub fn unwrap_inner(self) -> &'a M::InnerNode {
501        match self {
502            Node::Inner(node) => node,
503            Node::Terminal(_) => panic!("expected an inner node, but this is a terminal"),
504        }
505    }
506
507    /// Unwrap the inner node
508    ///
509    /// Panics with `msg` if `self` is a terminal.
510    #[must_use]
511    #[track_caller]
512    #[inline]
513    pub fn expect_inner(self, msg: &str) -> &'a M::InnerNode {
514        match self {
515            Node::Inner(node) => node,
516            Node::Terminal(_) => panic!("{}", msg),
517        }
518    }
519
520    /// Returns `true` if this is an inner node
521    #[inline]
522    pub fn is_inner(self) -> bool {
523        match self {
524            Node::Inner(_) => true,
525            Node::Terminal(_) => false,
526        }
527    }
528
529    /// Unwrap the terminal
530    ///
531    /// Panics if `self` is an inner node
532    #[must_use]
533    #[track_caller]
534    #[inline]
535    pub fn unwrap_terminal(&self) -> &M::Terminal {
536        match self {
537            Node::Inner(_) => panic!("expected a terminal node, but this is an inner node"),
538            Node::Terminal(ref t) => t.borrow(),
539        }
540    }
541
542    /// Unwrap the terminal
543    ///
544    /// Panics with `msg` if `self` is an inner node
545    #[must_use]
546    #[track_caller]
547    #[inline]
548    pub fn expect_terminal(&self, msg: &str) -> &M::Terminal {
549        match self {
550            Node::Inner(_) => panic!("{}", msg),
551            Node::Terminal(ref t) => t.borrow(),
552        }
553    }
554
555    /// Returns `true` if this is any terminal node
556    #[inline]
557    pub fn is_any_terminal(self) -> bool {
558        match self {
559            Node::Inner(_) => false,
560            Node::Terminal(_) => true,
561        }
562    }
563
564    /// Returns `true` if this is the given `terminal`
565    #[inline]
566    pub fn is_terminal(self, terminal: &M::Terminal) -> bool {
567        match self {
568            Node::Inner(_) => false,
569            Node::Terminal(t) => t.borrow() == terminal,
570        }
571    }
572}
573
574/// Manager for nodes in a decision diagram
575///
576/// In the basic formulation, a decision diagram is a directed acyclic graph
577/// where all inner nodes are associated with a level, and each level in turn is
578/// associated with a variable. A decision diagram can represent functions
579/// Dⁿ → T, where n is the number of variables. Every inner node has |D|
580/// outgoing edges, pointing to child nodes. The semantics of an inner node
581/// depends on the value of its associated variable. If the variable has
582/// value d ∈ D, then the node's semantics is the semantics of the child
583/// referenced by the edge corresponding to d. Every terminal node is associated
584/// with a value t ∈ T, and the node's semantics is just that value t. Some
585/// kinds of decision diagrams deviate from this formulation in one way or the
586/// other, but still fit into the framework.
587///
588/// The manager's responsibility is to store nodes and provide edges to identify
589/// them. Internally, it has a unique table to ensure uniqueness of nodes
590/// (typically, there should be no two nodes with the same children at the same
591/// level). The manager supports some kind of garbage collection: If there are
592/// no more edges pointing to a node, the node does not necessarily have to be
593/// deleted immediately. It may well be that the node revives. But from time to
594/// time it may make sense to remove all currently dead nodes. The manager also
595/// supports a levelized view on the diagram (via [`LevelView`]s).
596///
597/// Note that we are way more concerned about levels than about variables here.
598/// This is because variables can easily be handled externally. For many kinds
599/// of diagrams, we can just use the function representation of a variable and
600/// get the mapping between variables and levels "for free". This is especially
601/// nice as there are reordering operations which change the mapping between
602/// levels and variables but implicitly update the function representations
603/// accordingly.
604///
605/// # Safety
606///
607/// The implementation must ensure that every inner node only refers to nodes
608/// stored in this manager.
609///
610/// Every level view is associated with a manager and a level number.
611/// [`Manager::level()`] must always return the level view associated to this
612/// manager with the given level number.
613///
614/// If [`Manager::InnerNode`] implements [`HasLevel`], then the implementation
615/// must ensure that [`HasLevel::level()`] returns level number L for all nodes
616/// at the level view for L. Specifically this means that
617/// [`Manager::add_level()`] must check the newly created node. The invariant
618/// may only be broken by unsafe code (e.g. via [`HasLevel::set_level()`] and
619/// [`LevelView::swap()`]) and must be re-established when leaving the unsafe
620/// scope (be aware of panics!).
621pub unsafe trait Manager: Sized {
622    /// Type of edge
623    type Edge: Edge<Tag = Self::EdgeTag>;
624    /// Type of edge tags
625    type EdgeTag: Tag;
626    /// Type of inner nodes
627    type InnerNode: InnerNode<Self::Edge>;
628    /// Type of terminals
629    type Terminal: Eq + Hash;
630    /// References to [`Self::Terminal`]s
631    ///
632    /// Should either be a `&'a Self::Terminal` or just `Self::Terminal`. The
633    /// latter is useful for "static terminal managers" which don't actually
634    /// store terminal nodes but can map between edges and terminal nodes on the
635    /// fly. In this case, it would be hard to hand out node references.
636    type TerminalRef<'a>: Borrow<Self::Terminal> + Copy
637    where
638        Self: 'a;
639    /// Diagram rules, see [`DiagramRules`] for more details
640    type Rules: DiagramRules<Self::Edge, Self::InnerNode, Self::Terminal>;
641
642    /// Iterator over all terminals
643    ///
644    /// The actual items are edges pointing to terminals since this allows us to
645    /// get a [`NodeID`].
646    type TerminalIterator<'a>: Iterator<Item = Self::Edge>
647    where
648        Self: 'a;
649
650    /// Node set type, possibly more efficient than a `HashSet<NodeID>`
651    type NodeSet: NodeSet<Self::Edge>;
652
653    /// A view on a single level of the unique table.
654    type LevelView<'a>: LevelView<Self::Edge, Self::InnerNode>
655    where
656        Self: 'a;
657
658    /// Iterator over levels
659    type LevelIterator<'a>: DoubleEndedIterator<Item = Self::LevelView<'a>> + ExactSizeIterator
660    where
661        Self: 'a;
662
663    /// Get a reference to the node to which `edge` points
664    #[must_use]
665    fn get_node(&self, edge: &Self::Edge) -> Node<Self>;
666
667    /// Clone `edge`
668    #[must_use]
669    fn clone_edge(&self, edge: &Self::Edge) -> Self::Edge;
670
671    /// Drop `edge`
672    fn drop_edge(&self, edge: Self::Edge);
673
674    /// Get the count of inner nodes
675    #[must_use]
676    fn num_inner_nodes(&self) -> usize;
677
678    /// Get an approximate count of inner nodes
679    ///
680    /// For concurrent implementations, it may be much less costly to determine
681    /// an approximation of the inner node count than an accurate count
682    /// ([`Self::num_inner_nodes()`]).
683    #[must_use]
684    fn approx_num_inner_nodes(&self) -> usize {
685        self.num_inner_nodes()
686    }
687
688    /// Get the number of levels
689    #[must_use]
690    fn num_levels(&self) -> LevelNo;
691
692    /// Add a level with the given node to the unique table.
693    ///
694    /// To avoid unnecessary (un-)locking, this function takes a closure `f`
695    /// that creates a first node for the new level.
696    ///
697    /// Returns an edge for the newly created node.
698    ///
699    /// Panics if the new node's level does not match the provided level.
700    #[must_use]
701    fn add_level(&mut self, f: impl FnOnce(LevelNo) -> Self::InnerNode) -> AllocResult<Self::Edge>;
702
703    /// Get the level given by `no`
704    ///
705    /// Implementations may or may not acquire a lock here.
706    ///
707    /// Panics if `no >= self.num_levels()`.
708    #[must_use]
709    fn level(&self, no: LevelNo) -> Self::LevelView<'_>;
710
711    /// Iterate over the levels from top to bottom
712    #[must_use]
713    fn levels(&self) -> Self::LevelIterator<'_>;
714
715    /// Get an edge for the given terminal
716    ///
717    /// Locking behavior: May acquire a lock for the internal terminal unique
718    /// table. In particular, this means that calling this function while
719    /// holding a terminal iterator ([`Manager::terminals()`]) may cause a
720    /// deadlock.
721    #[must_use]
722    fn get_terminal(&self, terminal: Self::Terminal) -> AllocResult<Self::Edge>;
723
724    /// Get the number of terminals
725    ///
726    /// Should agree with the length of the iterator returned by
727    /// [`Manager::terminals()`].
728    ///
729    /// Locking behavior: May acquire a lock for the internal terminal unique
730    /// table. In particular, this means that calling this function while
731    /// holding a terminal iterator ([`Manager::terminals()`]) may cause a
732    /// deadlock.
733    #[must_use]
734    fn num_terminals(&self) -> usize;
735
736    /// Iterator over all terminals
737    ///
738    /// Locking behavior: May acquire a lock for the internal terminal unique
739    /// table.
740    #[must_use]
741    fn terminals(&self) -> Self::TerminalIterator<'_>;
742
743    /// Perform garbage collection
744    ///
745    /// This method looks for nodes that are neither referenced by a
746    /// [`Function`][function::Function] nor another node and removes them. The
747    /// method works from top to bottom, so if a node is only referenced by
748    /// nodes that can be removed, this node will be removed as well.
749    ///
750    /// Returns the number of nodes removed.
751    fn gc(&self) -> usize;
752
753    /// Prepare and postprocess a reordering operation. The reordering itself is
754    /// performed in `f`.
755    ///
756    /// Returns the value returned by `f`.
757    fn reorder<T>(&mut self, f: impl FnOnce(&mut Self) -> T) -> T;
758
759    /// Get the count of garbage collections
760    ///
761    /// This counter should monotonically increase to ensure that caches are
762    /// invalidated accordingly.
763    fn gc_count(&self) -> u64;
764
765    /// Get the count of reordering operations
766    ///
767    /// This counter should monotonically increase to ensure that caches are
768    /// invalidated accordingly.
769    fn reorder_count(&self) -> u64;
770}
771
772/// View of a single level in the manager
773///
774/// # Safety
775///
776/// Each level view is associated with a [`Manager`] M and a [`LevelNo`] L. The
777/// level view must ensure that all contained nodes and their descendants are
778/// stored in M. The edges returned by [`LevelView::get()`] and
779/// [`LevelView::get_or_insert()`] must reference nodes at this level.
780/// [`LevelView::iter()`] must yield all edges to nodes at this level (it must
781/// not hide some away).
782///
783/// [`LevelView::swap()`] conceptually removes all nodes from this level and
784/// inserts them at the other level and vice versa.
785pub unsafe trait LevelView<E: Edge, N: InnerNode<E>> {
786    /// Iterator over [`Edge`]s pointing to nodes at this level
787    type Iterator<'a>: Iterator<Item = &'a E>
788    where
789        Self: 'a,
790        E: 'a;
791
792    /// Taken level view, see [`LevelView::take()`]
793    type Taken: LevelView<E, N>;
794
795    /// Get the number of nodes on this level
796    #[must_use]
797    fn len(&self) -> usize;
798
799    /// Returns `true` iff this level contains nodes
800    #[must_use]
801    fn is_empty(&self) -> bool {
802        self.len() == 0
803    }
804
805    /// Get the level number of this level
806    #[must_use]
807    fn level_no(&self) -> LevelNo;
808
809    /// Reserve space for `additional` nodes on this level
810    fn reserve(&mut self, additional: usize);
811
812    /// Get the edge corresponding to the given node (if present)
813    #[must_use]
814    fn get(&self, node: &N) -> Option<&E>;
815
816    /// Insert the given edge into the unique table at this level, assuming that
817    /// the referenced node is already stored in the associated manager.
818    ///
819    /// Returns `true` if the edge was inserted, `false` if it was already
820    /// present.
821    ///
822    /// Panics if `edge`
823    /// - points to a terminal node,
824    /// - references a node from a different manager, or
825    /// - `N` implements [`HasLevel`] and `edge` references a node for which
826    ///   [`HasLevel::level()`] returns a different level.
827    ///
828    /// Furthermore, this function should panic if `edge` is tagged, but the
829    /// caller must not rely on that. An implementation may simply remove the
830    /// tag for optimization purposes.
831    fn insert(&mut self, edge: E) -> bool;
832
833    /// Get the edge corresponding to `level` and `node` if present, or insert
834    /// it.
835    ///
836    /// Panics if
837    /// - the children of `node` are stored in a different manager, or
838    /// - `N` implements [`HasLevel`] and [`HasLevel::level(node)`] returns a
839    ///   different level.
840    #[must_use]
841    fn get_or_insert(&mut self, node: N) -> AllocResult<E>;
842
843    /// Perform garbage collection on this level
844    ///
845    /// # Safety
846    ///
847    /// Must be called from inside the closure passed to [`Manager::reorder()`].
848    unsafe fn gc(&mut self);
849
850    /// Remove `node` from (this level of) the manager
851    ///
852    /// Returns whether the value was present at this level.
853    ///
854    /// # Safety
855    ///
856    /// Must be called from inside the closure passed to [`Manager::reorder()`].
857    unsafe fn remove(&mut self, node: &N) -> bool;
858
859    /// Move all nodes from this level to the other level and vice versa.
860    ///
861    /// # Safety
862    ///
863    /// This method does not necessarily change the level returned by
864    /// [`HasLevel::level()`] for the nodes at this or the `other` level. The
865    /// caller must ensure a consistent state using calls to
866    /// [`HasLevel::set_level()`]. (Be aware of exception safety!)
867    unsafe fn swap(&mut self, other: &mut Self);
868
869    /// Iterate over all the edges at this level
870    #[must_use]
871    fn iter(&self) -> Self::Iterator<'_>;
872
873    /// Clear this level, returning a level view containing all the previous
874    /// edges.
875    #[must_use]
876    fn take(&mut self) -> Self::Taken;
877}
878
879/// Cache for the result of apply operations
880///
881/// This trait provides methods to add computation results to the apply cache
882/// and to query the cache for these results. Just as every cache, the
883/// implementation may decide to evict results from the cache.
884pub trait ApplyCache<M: Manager, O: Copy>: DropWith<M::Edge> {
885    /// Get the result of `operation`, if cached
886    #[must_use]
887    fn get_with_numeric(
888        &self,
889        manager: &M,
890        operator: O,
891        operands: &[Borrowed<M::Edge>],
892        numeric_operands: &[u32],
893    ) -> Option<M::Edge>;
894
895    /// Add the result of `operation` to this cache
896    ///
897    /// An implementation is free to not cache any result. (This is why we use
898    /// `Borrowed<M::Edge>`, which in this case elides a few clone and drop
899    /// operations.) If the cache already contains a key equal to `operation`,
900    /// there is no need to update its value. (Again, we can elide clone and
901    /// drop operations.)
902    fn add_with_numeric(
903        &self,
904        manager: &M,
905        operator: O,
906        operands: &[Borrowed<M::Edge>],
907        numeric_operands: &[u32],
908        value: Borrowed<M::Edge>,
909    );
910
911    /// Shorthand for [`Self::get_with_numeric()`] without numeric operands
912    #[inline(always)]
913    #[must_use]
914    fn get(&self, manager: &M, operator: O, operands: &[Borrowed<M::Edge>]) -> Option<M::Edge> {
915        self.get_with_numeric(manager, operator, operands, &[])
916    }
917
918    /// Shorthand for [`Self::add_with_numeric()`] without numeric operands
919    #[inline(always)]
920    fn add(
921        &self,
922        manager: &M,
923        operator: O,
924        operands: &[Borrowed<M::Edge>],
925        value: Borrowed<M::Edge>,
926    ) {
927        self.add_with_numeric(manager, operator, operands, &[], value)
928    }
929
930    /// Remove all entries from the cache
931    fn clear(&self, manager: &M);
932}
933
934/// Apply cache container
935///
936/// Intended to be implemented by [`Manager`]s such that generic implementations
937/// of the recursive apply algorithm can simply require
938/// `M: Manager + HasApplyCache<M, O>`, where `O` is the operator type.
939pub trait HasApplyCache<M: Manager, O: Copy> {
940    /// The concrete apply cache type
941    type ApplyCache: ApplyCache<M, O>;
942
943    /// Get a shared reference to the contained apply cache
944    #[must_use]
945    fn apply_cache(&self) -> &Self::ApplyCache;
946
947    /// Get a mutable reference to the contained apply cache
948    #[must_use]
949    fn apply_cache_mut(&mut self) -> &mut Self::ApplyCache;
950}
951
952/// Worker thread pool associated with a [`Manager`]
953///
954/// A manager having its own thread pool has the advantage that it may use
955/// thread-local storage for its workers to pre-allocate some resources (e.g.,
956/// slots for nodes) and thereby reduce lock contention.
957pub trait WorkerPool: Sync {
958    /// Get the current number of threads
959    fn current_num_threads(&self) -> usize;
960
961    /// Get the recursion depth up to which operations are split
962    fn split_depth(&self) -> u32;
963
964    /// Set the recursion depth up to which operations are split
965    ///
966    /// `None` means that the implementation should automatically choose the
967    /// depth. `Some(0)` means that no operations are split.
968    fn set_split_depth(&self, depth: Option<u32>);
969
970    /// Execute `op` within the thread pool
971    ///
972    /// If this method is called from another thread pool, it may cooperatively
973    /// yield execution to that pool until `op` has finished.
974    fn install<R: Send>(&self, op: impl FnOnce() -> R + Send) -> R;
975
976    /// Execute `op_a` and `op_b` in parallel within the thread pool
977    ///
978    /// Note that the split depth has no influence on this method. Checking
979    /// whether to split an operation must be done externally.
980    fn join<RA: Send, RB: Send>(
981        &self,
982        op_a: impl FnOnce() -> RA + Send,
983        op_b: impl FnOnce() -> RB + Send,
984    ) -> (RA, RB);
985
986    /// Execute `op` on every worker in the thread pool
987    fn broadcast<R: Send>(&self, op: impl Fn(BroadcastContext) -> R + Sync) -> Vec<R>;
988}
989
990/// Context provided to workers by [`WorkerPool::broadcast()`]
991#[derive(Clone, Copy, Debug)]
992pub struct BroadcastContext {
993    /// Index of this worker (in range `0..num_threads`)
994    pub index: u32,
995
996    /// Number of threads receiving the broadcast
997    pub num_threads: u32,
998}
999
1000/// Helper trait to be implemented by [`Manager`] and [`ManagerRef`] if they
1001/// feature a [`WorkerPool`].
1002pub trait HasWorkers: Sync {
1003    /// Type of the worker pool
1004    type WorkerPool: WorkerPool;
1005
1006    /// Get the worker pool
1007    fn workers(&self) -> &Self::WorkerPool;
1008}