catlog/dbl/
tree.rs

1/*! Double trees: pasting diagrams in virtual double categories.
2
3A *double tree* (nonstandard term) is the data structure for a [pasting
4diagram](https://ncatlab.org/nlab/show/pasting+diagram) in a virtual double
5category. In other words, a double tree specifies, in the most general and
6unbiased form, a composite of cells in a virtual double category.
7 */
8
9use derive_more::From;
10use ego_tree::Tree;
11use std::collections::VecDeque;
12
13use super::graph::VDblGraph;
14use super::tree_algorithms::*;
15use crate::one::path::Path;
16
17/** A node in a [double tree](DblTree).
18
19To be more precise, this enum is the type of a *value* carried by a node in a
20double tree.
21 */
22#[derive(Clone, Debug, PartialEq, Eq)]
23pub enum DblNode<E, ProE, Sq> {
24    /// A generic cell, given by a square in a virtual double graph.
25    Cell(Sq),
26
27    /** The identity cell on a pro-edge in a virtual double graph.
28
29    Any node with an identity as its value should be a leaf node. While not
30    logically required, we enforce this invariant to obtain a normal form for
31    pastings in VDCs.
32     */
33    Id(ProE),
34
35    /** An edge dangling from a nullary cell.
36
37    In a well-formed double tree, a spine node can be a child only of a nullary
38    cell or of another spine node. Spines represent the operation of
39    precomposing a nullary cell with an arrow to obtain another nullary cell, a
40    degenerate case of composition in a virtual double category.
41     */
42    Spine(E),
43}
44
45impl<E, ProE, Sq> DblNode<E, ProE, Sq> {
46    /// Is the node a generic cell?
47    pub fn is_cell(&self) -> bool {
48        matches!(*self, DblNode::Cell(_))
49    }
50
51    /// Is the node an identity?
52    pub fn is_id(&self) -> bool {
53        matches!(*self, DblNode::Id(_))
54    }
55
56    /// Is the node a spine?
57    pub fn is_spine(&self) -> bool {
58        matches!(*self, DblNode::Spine(_))
59    }
60
61    /// Domain of node in the given virtual double graph.
62    pub fn dom<V>(
63        &self,
64        graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>,
65    ) -> Path<V, ProE>
66    where
67        ProE: Clone,
68    {
69        match self {
70            DblNode::Cell(sq) => graph.square_dom(sq),
71            DblNode::Id(p) => p.clone().into(),
72            DblNode::Spine(e) => Path::empty(graph.dom(e)),
73        }
74    }
75
76    /// Codomain of node in the given virtual double graph.
77    pub fn cod<V>(
78        &self,
79        graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>,
80    ) -> Path<V, ProE>
81    where
82        ProE: Clone,
83    {
84        match self {
85            DblNode::Cell(sq) => graph.square_cod(sq).into(),
86            DblNode::Id(p) => p.clone().into(),
87            DblNode::Spine(e) => Path::empty(graph.cod(e)),
88        }
89    }
90
91    /// Source of node in the given virtual double graph.
92    pub fn src<V>(&self, graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>) -> Path<V, E>
93    where
94        E: Clone,
95    {
96        match self {
97            DblNode::Cell(sq) => graph.square_src(sq).into(),
98            DblNode::Id(p) => Path::empty(graph.src(p)),
99            DblNode::Spine(e) => e.clone().into(),
100        }
101    }
102
103    /// Target of node in the given virtual double graph.
104    pub fn tgt<V>(&self, graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>) -> Path<V, E>
105    where
106        E: Clone,
107    {
108        match self {
109            DblNode::Cell(sq) => graph.square_tgt(sq).into(),
110            DblNode::Id(p) => Path::empty(graph.tgt(p)),
111            DblNode::Spine(e) => e.clone().into(),
112        }
113    }
114
115    /// Arity of node in the given virtual double graph.
116    pub fn arity(&self, graph: &impl VDblGraph<E = E, ProE = ProE, Sq = Sq>) -> usize {
117        match self {
118            DblNode::Cell(sq) => graph.arity(sq),
119            DblNode::Id(_) => 1,
120            DblNode::Spine(_) => 0,
121        }
122    }
123
124    /// Is the node contained in the given virtual double graph?
125    pub fn contained_in(&self, graph: &impl VDblGraph<E = E, ProE = ProE, Sq = Sq>) -> bool {
126        match self {
127            DblNode::Cell(sq) => graph.has_square(sq),
128            DblNode::Id(p) => graph.has_proedge(p),
129            DblNode::Spine(e) => graph.has_edge(e),
130        }
131    }
132}
133
134/** A double tree, or pasting diagram in a virtual double category.
135
136As the name suggests, the underlying data structure of a double tree is a
137[`Tree`] whose [nodes](DblNode) represent cells (or occasionally arrows) in the
138pasting diagram. Not just any underlying tree constitutes a valid pasting. For
139example, the domains/codomains and sources/targets of the cells must compatible,
140and [spines](DblNode::Spine) can only appear in certain configurations.
141Moreover, among the valid trees, invariants are maintained to ensure a normal
142form among equivalent representations of the same pasting.
143
144TODO: Enforce invariant with identities when `graft`-ing.
145*/
146#[derive(Clone, Debug, From, PartialEq, Eq)]
147pub struct DblTree<E, ProE, Sq>(pub Tree<DblNode<E, ProE, Sq>>);
148
149impl<E, ProE, Sq> DblTree<E, ProE, Sq> {
150    /// Constructs the empty or identity tree.
151    pub fn empty(p: ProE) -> Self {
152        Tree::new(DblNode::Id(p)).into()
153    }
154
155    /// Constructs a tree with a single square, its root.
156    pub fn single(sq: Sq) -> Self {
157        Tree::new(DblNode::Cell(sq)).into()
158    }
159
160    /// Constructs a linear tree from a sequence of squares.
161    pub fn linear(iter: impl IntoIterator<Item = Sq>) -> Option<Self> {
162        DblTree::from_nodes(iter.into_iter().map(DblNode::Cell))
163    }
164
165    /// Constructs a tree with a single spine node.
166    pub fn spine(e: E) -> Self {
167        Tree::new(DblNode::Spine(e)).into()
168    }
169
170    /// Constructs a tree from a non-empty path of edges.
171    pub fn spines<V>(path: Path<V, E>) -> Option<Self> {
172        DblTree::from_nodes(path.into_iter().map(DblNode::Spine))
173    }
174
175    /// Constructs a linear tree from a sequence of node values.
176    pub fn from_nodes(iter: impl IntoIterator<Item = DblNode<E, ProE, Sq>>) -> Option<Self> {
177        let mut values: Vec<_> = iter.into_iter().collect();
178        let value = values.pop()?;
179        let mut tree = Tree::new(value);
180        let mut node_id = tree.root().id();
181        for value in values.into_iter().rev() {
182            node_id = tree.get_mut(node_id).unwrap().append(value).id();
183        }
184        Some(tree.into())
185    }
186
187    /// Constructs a tree of a height two.
188    pub fn two_level(leaves: impl IntoIterator<Item = Sq>, base: Sq) -> Self {
189        Self::graft(leaves.into_iter().map(DblTree::single), base)
190    }
191
192    /// Constructs a tree by grafting trees as subtrees onto a base cell.
193    pub fn graft(subtrees: impl IntoIterator<Item = Self>, base: Sq) -> Self {
194        let mut tree = Tree::new(DblNode::Cell(base));
195        for subtree in subtrees {
196            tree.root_mut().append_subtree(subtree.0);
197        }
198        tree.into()
199    }
200
201    /** The size of the tree.
202
203    The **size** of a double tree is the number of non-identity nodes in it.
204     */
205    pub fn size(&self) -> usize {
206        self.0.values().filter(|dn| dn.is_cell()).count()
207    }
208
209    /** Is the tree empty?
210
211    A double tree is **empty** if its sole node, the root, is an identity.
212    */
213    pub fn is_empty(&self) -> bool {
214        let root = self.0.root();
215        let root_is_id = root.value().is_id();
216        assert!(!(root_is_id && root.has_children()), "Identity node should not have children");
217        root_is_id
218    }
219
220    /// Gets the root of the double tree.
221    pub fn root(&self) -> &DblNode<E, ProE, Sq> {
222        self.0.root().value()
223    }
224
225    /// Iterates over the leaves of the double tree.
226    pub fn leaves(&self) -> impl Iterator<Item = &DblNode<E, ProE, Sq>> {
227        self.0.root().descendants().filter_map(|node| {
228            if node.has_children() {
229                None
230            } else {
231                Some(node.value())
232            }
233        })
234    }
235
236    /** Iterates over nodes along the source (left) boundary of the double tree.
237
238    *Warning*: iteration proceeds from the tree's root to its left-most leaf,
239    which is the opposite order of the path of edges.
240     */
241    pub fn src_nodes(&self) -> impl Iterator<Item = &DblNode<E, ProE, Sq>> {
242        let mut maybe_node = Some(self.0.root());
243        std::iter::from_fn(move || {
244            let prev = maybe_node;
245            maybe_node = maybe_node.and_then(|node| node.first_child());
246            prev.map(|node| node.value())
247        })
248    }
249
250    /** Iterates over nodes along the target (right) boundary of the double tree.
251
252    *Warning*: iteration proceeds from the tree's root to its right-most leaf,
253    which is the opposite order of the path of edges.
254     */
255    pub fn tgt_nodes(&self) -> impl Iterator<Item = &DblNode<E, ProE, Sq>> {
256        let mut maybe_node = Some(self.0.root());
257        std::iter::from_fn(move || {
258            let prev = maybe_node;
259            maybe_node = maybe_node.and_then(|node| node.last_child());
260            prev.map(|node| node.value())
261        })
262    }
263
264    /// Domain of the tree in the given virtual double graph.
265    pub fn dom<V>(
266        &self,
267        graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>,
268    ) -> Path<V, ProE>
269    where
270        ProE: Clone,
271    {
272        Path::collect(self.leaves().map(|dn| dn.dom(graph))).unwrap().flatten()
273    }
274
275    /// Codomain of the tree in the given virtual double graph.
276    pub fn cod(&self, graph: &impl VDblGraph<E = E, ProE = ProE, Sq = Sq>) -> ProE
277    where
278        ProE: Clone,
279    {
280        self.root()
281            .cod(graph)
282            .only()
283            .expect("The root of a double tree should not be a spine")
284    }
285
286    /// Source of the tree in the given virtual double graph.
287    pub fn src<V>(&self, graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>) -> Path<V, E>
288    where
289        E: Clone,
290    {
291        let mut edges: Vec<_> = self.src_nodes().map(|dn| dn.src(graph)).collect();
292        edges.reverse();
293        Path::from_vec(edges).unwrap().flatten()
294    }
295
296    /// Target of the tree in the given virtual double graph.
297    pub fn tgt<V>(&self, graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>) -> Path<V, E>
298    where
299        E: Clone,
300    {
301        let mut edges: Vec<_> = self.tgt_nodes().map(|dn| dn.tgt(graph)).collect();
302        edges.reverse();
303        Path::from_vec(edges).unwrap().flatten()
304    }
305
306    /// Arity of the composite cell specified by the tree.
307    pub fn arity(&self, graph: &impl VDblGraph<E = E, ProE = ProE, Sq = Sq>) -> usize {
308        self.leaves().map(|dn| dn.arity(graph)).sum()
309    }
310
311    /** Is the double tree contained in the given virtual double graph?
312
313    This includes checking whether the double tree is well-typed, i.e., that the
314    domains and codomains, and sources and targets, of the cells are compatible.
315     */
316    pub fn contained_in(&self, graph: &impl VDblGraph<E = E, ProE = ProE, Sq = Sq>) -> bool
317    where
318        E: Eq + Clone,
319        ProE: Eq + Clone,
320    {
321        let mut traverse = self.0.bfs();
322        while let Some(node) = traverse.next() {
323            let dn = node.value();
324            // The cell itself is contained in the graph.
325            if !dn.contained_in(graph) {
326                return false;
327            }
328            // Source and target edges are compatible.
329            if !traverse
330                .peek_at_same_level()
331                .is_none_or(|next| dn.tgt(graph) == next.value().src(graph))
332            {
333                return false;
334            }
335            // Domain and codomain pro-edges are compatible.
336            if node.has_children() {
337                let codomains = node.children().map(|child| child.value().cod(graph));
338                if Path::collect(codomains).unwrap().flatten() != dn.dom(graph) {
339                    return false;
340                }
341            }
342        }
343        true
344    }
345
346    /// Is the double tree isomorphic to another?
347    pub fn is_isomorphic_to(&self, other: &Self) -> bool
348    where
349        E: Eq,
350        ProE: Eq,
351        Sq: Eq,
352    {
353        self.0.is_isomorphic_to(&other.0)
354    }
355
356    /// Maps over the edges and squares of the tree.
357    pub fn map<CodE, CodSq>(
358        self,
359        mut fn_e: impl FnMut(E) -> CodE,
360        mut fn_sq: impl FnMut(Sq) -> CodSq,
361    ) -> DblTree<CodE, ProE, CodSq> {
362        self.0
363            .map(|dn| match dn {
364                DblNode::Cell(sq) => DblNode::Cell(fn_sq(sq)),
365                DblNode::Spine(e) => DblNode::Spine(fn_e(e)),
366                DblNode::Id(m) => DblNode::Id(m),
367            })
368            .into()
369    }
370}
371
372impl<V, E, ProE, Sq> DblNode<Path<V, E>, ProE, DblTree<E, ProE, Sq>> {
373    /// Flattens a node containing another tree.
374    fn flatten(self) -> DblTree<E, ProE, Sq> {
375        match self {
376            DblNode::Cell(tree) => tree,
377            DblNode::Id(m) => DblTree::empty(m),
378            DblNode::Spine(path) => {
379                DblTree::spines(path).expect("Spine should be a non-empty path")
380            }
381        }
382    }
383}
384
385impl<V, E, ProE, Sq> DblTree<Path<V, E>, ProE, DblTree<E, ProE, Sq>>
386where
387    V: Clone,
388    E: Clone,
389    ProE: Clone + Eq + std::fmt::Debug,
390    Sq: Clone,
391{
392    /// Flattens a tree of trees into a single tree.
393    pub fn flatten_in(
394        &self,
395        graph: &impl VDblGraph<V = V, E = E, ProE = ProE, Sq = Sq>,
396    ) -> DblTree<E, ProE, Sq> {
397        // Initialize flattened tree using the root of the outer tree.
398        let outer_root = self.0.root();
399        let mut tree = outer_root.value().clone().flatten().0;
400
401        // We'll iterate over the outer tree in breadth-first order.
402        let mut outer_nodes = self.0.bfs();
403        outer_nodes.next();
404
405        // In parallel order, we'll iterate over the roots of the subtrees added
406        // to the flattened tree.
407        let mut queue = VecDeque::new();
408        if outer_root.has_children() {
409            queue.push_back(tree.root().id());
410        }
411
412        while let Some(node_id) = queue.pop_front() {
413            let leaf_ids: Vec<_> = tree
414                .get(node_id)
415                .unwrap()
416                .descendants()
417                .filter_map(|node| {
418                    if node.has_children() {
419                        None
420                    } else {
421                        Some(node.id())
422                    }
423                })
424                .collect();
425            for leaf_id in leaf_ids {
426                let mut leaf = tree.get_mut(leaf_id).unwrap();
427                for m in leaf.value().dom(graph) {
428                    let outer_node =
429                        outer_nodes.next().expect("Outer tree should have enough nodes");
430
431                    let inner_tree = outer_node.value().clone().flatten();
432                    assert_eq!(m, inner_tree.cod(graph), "(Co)domains should be compatible");
433
434                    let subtree_id = leaf.append_subtree(inner_tree.0).id();
435                    if outer_node.has_children() {
436                        queue.push_back(subtree_id);
437                    }
438                }
439            }
440        }
441
442        assert!(outer_nodes.next().is_none(), "Outer tree should not have extra nodes");
443        tree.into()
444    }
445}
446
447#[cfg(test)]
448mod tests {
449    use ego_tree::tree;
450    use nonempty::nonempty;
451
452    use super::super::category::{WalkingBimodule as Bimod, WalkingFunctor as Funct, *};
453    use super::*;
454
455    #[test]
456    fn tree_dom_cod() {
457        let bimod = Bimod::Main();
458        let graph = UnderlyingDblGraph(Bimod::Main());
459        let path = Path::Seq(nonempty![Bimod::Pro::Left, Bimod::Pro::Middle, Bimod::Pro::Right]);
460        let mid = bimod.composite_ext(path).unwrap();
461        let tree = DblTree::two_level(
462            vec![bimod.id_cell(Bimod::Pro::Left), mid.clone(), bimod.id_cell(Bimod::Pro::Right)],
463            mid.clone(),
464        );
465        let tree_alt = tree!(
466            mid.clone() => {
467                bimod.id_cell(Bimod::Pro::Left), mid.clone(), bimod.id_cell(Bimod::Pro::Right)
468            }
469        );
470        let tree_alt = DblTree(tree_alt.map(DblNode::Cell));
471        assert_eq!(tree, tree_alt);
472        assert!(tree.contained_in(&graph));
473
474        assert_eq!(tree.leaves().count(), 3);
475        assert_eq!(tree.arity(&graph), 5);
476        assert_eq!(
477            tree.dom(&graph),
478            Path::Seq(nonempty![
479                Bimod::Pro::Left,
480                Bimod::Pro::Left,
481                Bimod::Pro::Middle,
482                Bimod::Pro::Right,
483                Bimod::Pro::Right
484            ])
485        );
486        assert_eq!(tree.cod(&graph), Bimod::Pro::Middle);
487
488        // Trees with incompatible (co)domains.
489        let tree = tree!(
490            mid.clone() => {
491                bimod.id_cell(Bimod::Pro::Left), mid.clone()
492            }
493        );
494        assert!(!DblTree(tree.map(DblNode::Cell)).contained_in(&graph));
495        let tree = tree!(
496            mid.clone() => {
497                bimod.id_cell(Bimod::Pro::Right), mid.clone(), bimod.id_cell(Bimod::Pro::Left)
498            }
499        );
500        assert!(!DblTree(tree.map(DblNode::Cell)).contained_in(&graph));
501    }
502
503    #[test]
504    fn tree_src_tgt() {
505        let funct = Funct::Main();
506        let graph = UnderlyingDblGraph(Funct::Main());
507        let f = Funct::Arr::Arrow;
508        let unit1 = funct.unit_ext(Funct::Ob::One).unwrap();
509        let tree = DblTree::from_nodes(vec![DblNode::Spine(f), DblNode::Cell(unit1)]).unwrap();
510        let tree_alt = DblTree(tree!(
511            DblNode::Cell(unit1) => { DblNode::Spine(f) }
512        ));
513        assert_eq!(tree, tree_alt);
514        assert!(tree.contained_in(&graph));
515
516        assert_eq!(tree.src_nodes().count(), 2);
517        assert_eq!(tree.tgt_nodes().count(), 2);
518        assert_eq!(tree.src(&graph), Path::pair(f, Funct::Arr::One));
519        assert_eq!(tree.tgt(&graph), Path::pair(f, Funct::Arr::One));
520        assert!(tree.dom(&graph).is_empty());
521
522        // Trees with incompatible sources and targets.
523        let tree = DblTree(tree!(
524            DblNode::Cell(funct.composite2_ext(Funct::Ob::One, Funct::Ob::One).unwrap()) => {
525                DblNode::Cell(unit1) => { DblNode::Spine(Funct::Arr::One) },
526                DblNode::Cell(unit1) => { DblNode::Spine(f) },
527            }
528        ));
529        assert!(!tree.contained_in(&graph));
530    }
531
532    #[test]
533    fn flatten_tree() {
534        let bimod = Bimod::Main();
535        let graph = UnderlyingDblGraph(Bimod::Main());
536        let path = Path::Seq(nonempty![Bimod::Pro::Left, Bimod::Pro::Middle, Bimod::Pro::Right]);
537        let unitl = bimod.unit_ext(Bimod::Ob::Left).unwrap();
538        let unitr = bimod.unit_ext(Bimod::Ob::Right).unwrap();
539        let mid = bimod.composite_ext(path).unwrap();
540        let tree = tree!(
541            mid.clone() => {
542                bimod.id_cell(Bimod::Pro::Left) => {
543                    unitl.clone(),
544                },
545                mid => {
546                    unitl, bimod.id_cell(Bimod::Pro::Middle), unitr.clone(),
547                },
548                bimod.id_cell(Bimod::Pro::Right) => {
549                    unitr,
550                }
551            }
552        );
553        let tree = DblTree(tree.map(DblNode::Cell));
554        assert_eq!(tree.dom(&graph), Path::single(Bimod::Pro::Middle));
555        assert_eq!(tree.cod(&graph), Bimod::Pro::Middle);
556
557        // Degenerate case: the outer tree is a singleton.
558        let outer = DblTree::single(tree.clone());
559        assert_eq!(outer.flatten_in(&graph), tree);
560
561        // Degenerate case: all inner trees are singletons.
562        let outer = tree.clone().map(Path::single, DblTree::single);
563        let result = outer.flatten_in(&graph);
564        assert!(result.is_isomorphic_to(&tree));
565    }
566}