catlog/dbl/
category.rs

1/*! Virtual double categories.
2
3# Background
4
5A [*virtual double
6category*](https://ncatlab.org/nlab/show/virtual+double+category) (VDC) is like
7a double category, except that there is no external composition operation on
8proarrows or cells. Rather, a cell has a domain that is a path of proarrows (a
9"virtual" composite). The name "virtual double category" was introduced by
10[Cruttwell and Shulman](crate::refs::GeneralizedMulticategories) but the concept
11has gone by many other names, notably *fc-multicategory* ([Leinster
122004](crate::refs::HigherOperads)).
13
14*Composites* of proarrows in a VDC, if they exist, are represented by cells
15satisfying a universal property ([Cruttwell-Shulman
162008](crate::refs::GeneralizedMulticategories), Section 5). In our usage of
17virtual double categories as double theories, we will assume that *units*
18(nullary composites) exist. We will not assume that any other composites exist,
19though often they do.
20
21Virtual double categories have pros and cons compared with ordinary double
22categories. We prefer VDCs in `catlog` because pastings of cells are much
23simpler in a VDC than in a double category: a pasting diagram in VDC is a
24well-typed [tree](super::tree) of cells, rather than a kind of planar string
25diagram, and the notorious
26[pinwheel](https://ncatlab.org/nlab/show/double+category#Unbiased) obstruction
27to composition in a double category does not arise in VDCs.
28
29# Examples
30
31A [double theory](super::theory) is "just" a unital virtual double category, so
32any double theory in the standard library is an example of a VDC. For testing
33purposes, this module provides several minimal examples of VDCs implemented
34directly, namely ["walking"](https://ncatlab.org/nlab/show/walking+structure)
35categorical structures that can be interpreted in any VDC:
36
37- the [walking category](WalkingCategory)
38- the [walking functor](WalkingFunctor)
39- the [walking bimodule](WalkingBimodule) or profunctor
40
41The walking category and bimodule can be seen as discrete double theories, while
42the walking functor is a simple double theory, but here they are implemented at
43the type level rather than as instances of general data structures.
44 */
45
46use derive_more::From;
47use ref_cast::RefCast;
48use std::ops::Range;
49
50use super::graph::{EdgeGraph, VDblGraph};
51use super::tree::DblTree;
52use crate::one::path::Path;
53
54/** A virtual double category (VDC).
55
56See the [module-level docs](super::category) for background on VDCs.
57 */
58pub trait VDblCategory {
59    /// Type of objects in the VDC.
60    type Ob: Eq + Clone;
61
62    /// Type of arrows (tight morphisms) in the VDC.
63    type Arr: Eq + Clone;
64
65    /// Type of proarrows (loose morphisms) in the VDC.
66    type Pro: Eq + Clone;
67
68    /// Type of cells in the VDC;
69    type Cell: Eq + Clone;
70
71    /// Does the object belong to the VDC?
72    fn has_ob(&self, ob: &Self::Ob) -> bool;
73
74    /// Does the arrow belong to the VDC?
75    fn has_arrow(&self, arr: &Self::Arr) -> bool;
76
77    /// Does the proarrow belong to the VDC?
78    fn has_proarrow(&self, pro: &Self::Pro) -> bool;
79
80    /// Does the cell belong to the VDC?
81    fn has_cell(&self, cell: &Self::Cell) -> bool;
82
83    /// Gets the domain of an arrow.
84    fn dom(&self, f: &Self::Arr) -> Self::Ob;
85
86    /// Gets the codomain of an arrow.
87    fn cod(&self, f: &Self::Arr) -> Self::Ob;
88
89    /// Gets the source of a proarrow.
90    fn src(&self, m: &Self::Pro) -> Self::Ob;
91
92    /// Gets the target of a proarrow.
93    fn tgt(&self, m: &Self::Pro) -> Self::Ob;
94
95    /// Gets the domain of a cell, a path of proarrows.
96    fn cell_dom(&self, cell: &Self::Cell) -> Path<Self::Ob, Self::Pro>;
97
98    /// Gets the codomain of a cell, a single proarrow.
99    fn cell_cod(&self, cell: &Self::Cell) -> Self::Pro;
100
101    /// Gets the source of a cell, an arrow.
102    fn cell_src(&self, cell: &Self::Cell) -> Self::Arr;
103
104    /// Gets the target of a cell, an edge.
105    fn cell_tgt(&self, cell: &Self::Cell) -> Self::Arr;
106
107    /** Gets the arity of a cell.
108
109    The default implementation returns the length of the cell's domain.
110     */
111    fn arity(&self, cell: &Self::Cell) -> usize {
112        self.cell_dom(cell).len()
113    }
114
115    /// Composes a path of arrows in the VDC.
116    fn compose(&self, path: Path<Self::Ob, Self::Arr>) -> Self::Arr;
117
118    /// Composes a pair of arrows with compatible (co)domains.
119    fn compose2(&self, f: Self::Arr, g: Self::Arr) -> Self::Arr {
120        self.compose(Path::pair(f, g))
121    }
122
123    /// Constructs the identity arrow at an object.
124    fn id(&self, x: Self::Ob) -> Self::Arr {
125        self.compose(Path::empty(x))
126    }
127
128    /// Composes a tree of cells in the VDC.
129    fn compose_cells(&self, tree: DblTree<Self::Arr, Self::Pro, Self::Cell>) -> Self::Cell;
130
131    /// Composes a two-layer pasting of cells.
132    fn compose_cells2(
133        &self,
134        αs: impl IntoIterator<Item = Self::Cell>,
135        β: Self::Cell,
136    ) -> Self::Cell {
137        self.compose_cells(DblTree::two_level(αs, β))
138    }
139
140    /// Constructs the identity cell on a proarrow.
141    fn id_cell(&self, m: Self::Pro) -> Self::Cell {
142        self.compose_cells(DblTree::empty(m))
143    }
144}
145
146/** A virtual double category with some or all chosen composites.
147
148Like anything defined by a universal property, composites in a VDC are not
149strictly unique if they exist but they *are* unique up to unique isomorphism. As
150often when working with (co)limits, this trait assumes that a *choice* of
151composites has been made whenever they are exist. We do not attempt to
152"recognize" whether an arbitrary cell has the relevant universal property.
153 */
154pub trait VDCWithComposites: VDblCategory {
155    /** Does the path of proarrows have a chosen composite?
156
157    The default implementation checks whether [`composite`](Self::composite)
158    returns something.
159    */
160    fn has_composite(&self, path: &Path<Self::Ob, Self::Pro>) -> bool {
161        self.composite(path.clone()).is_some()
162    }
163
164    /** Does the object have a chosen unit?
165
166    The default implementation checks whether [`unit`](Self::unit) returns
167    something.
168     */
169    fn has_unit(&self, x: &Self::Ob) -> bool {
170        self.unit(x.clone()).is_some()
171    }
172
173    /** Gets the chosen cell witnessing a composite of proarrows, if there is one.
174
175    Such a cell is called an **extension or **opcartesian** cell.
176     */
177    fn composite_ext(&self, path: Path<Self::Ob, Self::Pro>) -> Option<Self::Cell>;
178
179    /// Gets the chosen cell witnessing a composite of two proarrows, if there is one.
180    fn composite2_ext(&self, m: Self::Pro, n: Self::Pro) -> Option<Self::Cell> {
181        self.composite_ext(Path::pair(m, n))
182    }
183
184    /** Gets the chosen composite for a path of proarrows, if there is one.
185
186    The default implementation returns the codomain of the extension cell from
187    [`composite_ext`](Self::composite_ext).
188     */
189    fn composite(&self, path: Path<Self::Ob, Self::Pro>) -> Option<Self::Pro> {
190        self.composite_ext(path).map(|α| self.cell_cod(&α))
191    }
192
193    /// Gets the chosen composite for a pair of consecutive proarrows, if there is one.
194    fn composite2(&self, m: Self::Pro, n: Self::Pro) -> Option<Self::Pro> {
195        self.composite(Path::pair(m, n))
196    }
197
198    /** Gets the chosen extension cell for an object, if there is one.
199
200    Such a cell is an [extension](Self::composite_ext) or opcartesian cell
201    in the nullary case.
202     */
203    fn unit_ext(&self, x: Self::Ob) -> Option<Self::Cell> {
204        self.composite_ext(Path::empty(x))
205    }
206
207    /** Gets the chosen unit for an object, if there is one.
208
209    The default implementation returns the codomain of the extension cell from
210    [`unit_ext`](Self::unit_ext).
211     */
212    fn unit(&self, x: Self::Ob) -> Option<Self::Pro> {
213        self.unit_ext(x).map(|α| self.cell_cod(&α))
214    }
215
216    /** Factorizes a cell through a composite of proarrows.
217
218    The subpath of the domain path at the given range is replaced with the
219    composite of that subpath, if the composite exists. This is the universal
220    property of the composite.
221    */
222    fn through_composite(&self, cell: Self::Cell, range: Range<usize>) -> Option<Self::Cell>;
223
224    /** Factorizes a cell through the unit proarrow for an object.
225
226    A unit proarrow is inserted into the domain path at the given index, if the
227    unit exists. This is the universal property of the unit.
228     */
229    fn through_unit(&self, cell: Self::Cell, index: usize) -> Option<Self::Cell> {
230        self.through_composite(cell, index..index)
231    }
232}
233
234/// The underlying [virtual double graph](VDblGraph) of a VDC.
235#[derive(From, RefCast)]
236#[repr(transparent)]
237pub struct UnderlyingDblGraph<VDC: VDblCategory>(pub VDC);
238
239impl<VDC: VDblCategory> VDblGraph for UnderlyingDblGraph<VDC> {
240    type V = VDC::Ob;
241    type E = VDC::Arr;
242    type ProE = VDC::Pro;
243    type Sq = VDC::Cell;
244
245    fn has_vertex(&self, v: &Self::V) -> bool {
246        self.0.has_ob(v)
247    }
248    fn has_edge(&self, e: &Self::E) -> bool {
249        self.0.has_arrow(e)
250    }
251    fn has_proedge(&self, p: &Self::ProE) -> bool {
252        self.0.has_proarrow(p)
253    }
254    fn has_square(&self, sq: &Self::Sq) -> bool {
255        self.0.has_cell(sq)
256    }
257
258    fn dom(&self, e: &Self::E) -> Self::V {
259        self.0.dom(e)
260    }
261    fn cod(&self, e: &Self::E) -> Self::V {
262        self.0.cod(e)
263    }
264    fn src(&self, p: &Self::ProE) -> Self::V {
265        self.0.src(p)
266    }
267    fn tgt(&self, p: &Self::ProE) -> Self::V {
268        self.0.tgt(p)
269    }
270
271    fn square_dom(&self, sq: &Self::Sq) -> Path<Self::V, Self::ProE> {
272        self.0.cell_dom(sq)
273    }
274    fn square_cod(&self, sq: &Self::Sq) -> Self::ProE {
275        self.0.cell_cod(sq)
276    }
277    fn square_src(&self, sq: &Self::Sq) -> Self::E {
278        self.0.cell_src(sq)
279    }
280    fn square_tgt(&self, sq: &Self::Sq) -> Self::E {
281        self.0.cell_tgt(sq)
282    }
283    fn arity(&self, sq: &Self::Sq) -> usize {
284        self.0.arity(sq)
285    }
286}
287
288/** The VDC freely generated by a virtual double graph.
289
290A virtual double graph freely generated a virtual double category, generalizing
291how a graph [freely generates](crate::one::category::FreeCategory) a category.
292This is not, however, the most general way to freely generate a VDC. A "virtual
293double computad" freely generates a VDC but allows the generating cells to have
294sources and targets that are *paths* of generating arrows.
295 */
296#[derive(From, RefCast)]
297#[repr(transparent)]
298pub struct FreeVDblCategory<G: VDblGraph>(pub G);
299
300impl<G: VDblGraph> VDblCategory for FreeVDblCategory<G>
301where
302    G::ProE: std::fmt::Debug,
303{
304    type Ob = G::V;
305    type Arr = Path<G::V, G::E>;
306    type Pro = G::ProE;
307    type Cell = DblTree<G::E, G::ProE, G::Sq>;
308
309    fn has_ob(&self, ob: &Self::Ob) -> bool {
310        self.0.has_vertex(ob)
311    }
312    fn has_arrow(&self, path: &Self::Arr) -> bool {
313        path.contained_in(EdgeGraph::ref_cast(&self.0))
314    }
315    fn has_proarrow(&self, pro: &Self::Pro) -> bool {
316        self.0.has_proedge(pro)
317    }
318    fn has_cell(&self, tree: &Self::Cell) -> bool {
319        tree.contained_in(&self.0)
320    }
321
322    fn dom(&self, path: &Self::Arr) -> Self::Ob {
323        path.src(EdgeGraph::ref_cast(&self.0))
324    }
325    fn cod(&self, path: &Self::Arr) -> Self::Ob {
326        path.tgt(EdgeGraph::ref_cast(&self.0))
327    }
328    fn src(&self, m: &Self::Pro) -> Self::Ob {
329        self.0.src(m)
330    }
331    fn tgt(&self, m: &Self::Pro) -> Self::Ob {
332        self.0.tgt(m)
333    }
334
335    fn cell_dom(&self, tree: &Self::Cell) -> Path<Self::Ob, Self::Pro> {
336        tree.dom(&self.0)
337    }
338    fn cell_cod(&self, tree: &Self::Cell) -> Self::Pro {
339        tree.cod(&self.0)
340    }
341    fn cell_src(&self, tree: &Self::Cell) -> Self::Arr {
342        tree.src(&self.0)
343    }
344    fn cell_tgt(&self, tree: &Self::Cell) -> Self::Arr {
345        tree.tgt(&self.0)
346    }
347    fn arity(&self, tree: &Self::Cell) -> usize {
348        tree.arity(&self.0)
349    }
350
351    fn compose(&self, path: Path<Self::Ob, Self::Arr>) -> Self::Arr {
352        path.flatten_in(EdgeGraph::ref_cast(&self.0))
353            .expect("Path of paths should be valid before flattening")
354    }
355    fn compose_cells(&self, tree: DblTree<Self::Arr, Self::Pro, Self::Cell>) -> Self::Cell {
356        tree.flatten_in(&self.0)
357    }
358}
359
360/** The walking category as a VDC.
361
362The walking category is the simplest example of a virtual double category that
363has units (and in fact all composites). Specifically, the **walking category**
364is the unital VDC freely generated by a single object, here called `()`.
365
366The concept of a category can be interpreted in any virtual double category: a
367**category object** in a VDC is a functor from the walking category into that
368VDC. In particular, a category object in spans is a category in the ordinary
369sense.
370 */
371pub struct WalkingCategory();
372
373impl VDblCategory for WalkingCategory {
374    type Ob = ();
375    type Arr = ();
376    type Pro = ();
377    type Cell = usize;
378
379    fn has_ob(&self, _: &Self::Ob) -> bool {
380        true
381    }
382    fn has_arrow(&self, _: &Self::Arr) -> bool {
383        true
384    }
385    fn has_proarrow(&self, _: &Self::Pro) -> bool {
386        true
387    }
388    fn has_cell(&self, _: &Self::Cell) -> bool {
389        true
390    }
391
392    fn dom(&self, _: &Self::Arr) -> Self::Ob {}
393    fn cod(&self, _: &Self::Arr) -> Self::Ob {}
394    fn src(&self, _: &Self::Pro) -> Self::Ob {}
395    fn tgt(&self, _: &Self::Pro) -> Self::Ob {}
396
397    fn cell_dom(&self, n: &Self::Cell) -> Path<Self::Ob, Self::Pro> {
398        Path::repeat_n((), (), *n)
399    }
400    fn cell_cod(&self, _: &Self::Cell) -> Self::Pro {}
401    fn cell_src(&self, _: &Self::Cell) -> Self::Arr {}
402    fn cell_tgt(&self, _: &Self::Cell) -> Self::Arr {}
403
404    fn compose(&self, _: Path<Self::Ob, Self::Arr>) -> Self::Arr {}
405    fn compose_cells(&self, tree: DblTree<Self::Arr, Self::Pro, Self::Cell>) -> Self::Cell {
406        tree.dom(UnderlyingDblGraph::ref_cast(self)).len()
407    }
408}
409
410impl VDCWithComposites for WalkingCategory {
411    /// In the walking category, every cell is an extension.
412    fn composite_ext(&self, path: Path<Self::Ob, Self::Pro>) -> Option<Self::Cell> {
413        Some(path.len())
414    }
415
416    fn through_composite(&self, n: Self::Cell, range: Range<usize>) -> Option<Self::Cell> {
417        Some(n + 1 - range.len())
418    }
419}
420
421#[allow(non_snake_case)]
422pub mod WalkingBimodule {
423    /*! The walking bimodule as a VDC.
424
425    The **walking bimodule**, also known as the **walking profunctor**, is the
426    unital virtual double category freely generated by a pair of objects, here
427    called [`Left`](Ob::Left) and [`Right`](Ob::Right), and a single proarrow
428    between them, here called [`Middle`](Pro::Middle). In fact, this VDC has all
429    composites.
430    */
431    use super::super::graph::ProedgeGraph;
432    use super::*;
433
434    /// Struct representing the walking bimodule.
435    pub struct Main();
436
437    /// Type of objects in the walking bimodule.
438    #[derive(Clone, Copy, Debug, PartialEq, Eq)]
439    pub enum Ob {
440        /// Object representing the source of the bimodule.
441        Left,
442        /// Object representing the target of the bimodule.
443        Right,
444    }
445
446    /// Type of proarrows in the walking bimodule.
447    #[derive(Clone, Copy, Debug, PartialEq, Eq)]
448    pub enum Pro {
449        /// Unit proarrow on the object [`Left`](Ob::Left).
450        Left,
451        /// Generating proarrow from [`Left`](Ob::Left) to [`Right`](Ob::Right).
452        Middle,
453        /// Unit proarrow on the object [`Right`](Ob::Right).
454        Right,
455    }
456
457    impl Pro {
458        fn src(self) -> Ob {
459            match self {
460                Pro::Left => Ob::Left,
461                Pro::Middle => Ob::Left,
462                Pro::Right => Ob::Right,
463            }
464        }
465
466        fn tgt(self) -> Ob {
467            match self {
468                Pro::Left => Ob::Left,
469                Pro::Middle => Ob::Right,
470                Pro::Right => Ob::Right,
471            }
472        }
473    }
474
475    impl VDblCategory for Main {
476        type Ob = Ob;
477        type Arr = Ob;
478        type Pro = Pro;
479        type Cell = Path<Ob, Pro>;
480
481        fn has_ob(&self, _: &Self::Ob) -> bool {
482            true
483        }
484        fn has_arrow(&self, _: &Self::Arr) -> bool {
485            true
486        }
487        fn has_proarrow(&self, _: &Self::Pro) -> bool {
488            true
489        }
490        fn has_cell(&self, path: &Path<Ob, Pro>) -> bool {
491            path.contained_in(ProedgeGraph::ref_cast(UnderlyingDblGraph::ref_cast(self)))
492        }
493
494        fn dom(&self, f: &Self::Arr) -> Self::Ob {
495            *f
496        }
497        fn cod(&self, f: &Self::Arr) -> Self::Ob {
498            *f
499        }
500        fn src(&self, m: &Self::Pro) -> Self::Ob {
501            m.src()
502        }
503        fn tgt(&self, m: &Self::Pro) -> Self::Ob {
504            m.tgt()
505        }
506
507        fn cell_dom(&self, path: &Path<Ob, Pro>) -> Path<Self::Ob, Self::Pro> {
508            path.clone()
509        }
510        fn cell_cod(&self, path: &Path<Ob, Pro>) -> Self::Pro {
511            assert!(self.has_cell(path));
512            match path {
513                Path::Id(Ob::Left) => Pro::Left,
514                Path::Id(Ob::Right) => Pro::Right,
515                Path::Seq(pros) => {
516                    *pros.iter().find(|m| **m == Pro::Middle).unwrap_or_else(|| pros.first())
517                }
518            }
519        }
520        fn cell_src(&self, path: &Path<Ob, Pro>) -> Self::Arr {
521            path.src(ProedgeGraph::ref_cast(UnderlyingDblGraph::ref_cast(self)))
522        }
523        fn cell_tgt(&self, path: &Path<Ob, Pro>) -> Self::Arr {
524            path.tgt(ProedgeGraph::ref_cast(UnderlyingDblGraph::ref_cast(self)))
525        }
526
527        fn compose(&self, path: Path<Self::Ob, Self::Arr>) -> Self::Arr {
528            match path {
529                Path::Id(x) => x,
530                Path::Seq(arrows) => *arrows.first(),
531            }
532        }
533        fn compose_cells(&self, tree: DblTree<Self::Arr, Self::Pro, Self::Cell>) -> Self::Cell {
534            let path = tree.dom(UnderlyingDblGraph::ref_cast(self));
535            assert!(self.has_cell(&path));
536            path
537        }
538    }
539
540    impl VDCWithComposites for Main {
541        /// In the walking bimodule, every cell is an extension.
542        fn composite_ext(&self, path: Path<Self::Ob, Self::Pro>) -> Option<Self::Cell> {
543            Some(path)
544        }
545
546        fn through_composite(&self, path: Self::Cell, range: Range<usize>) -> Option<Self::Cell> {
547            let graph = ProedgeGraph::ref_cast(UnderlyingDblGraph::ref_cast(self));
548            Some(path.replace_subpath(graph, range, |subpath| self.cell_cod(&subpath).into()))
549        }
550    }
551}
552
553#[allow(non_snake_case)]
554pub mod WalkingFunctor {
555    /*! The walking functor as a VDC.
556
557    The **walking functor** is the unital virtual double category freely
558    generated by a pair of objects, here called [`Zero`](Ob::Zero) and
559    [`One`](Ob::One), and a single arrow between them.
560     */
561    use super::super::graph::{EdgeGraph, ProedgeGraph};
562    use super::*;
563
564    /// Struct representing the walking functor.
565    pub struct Main();
566
567    /// Type of objects in the walking functor.
568    #[derive(Clone, Copy, Debug, PartialEq, Eq)]
569    pub enum Ob {
570        /// Object representing the domain of the functor.
571        Zero,
572        /// Object representing the codomain of the functor.
573        One,
574    }
575
576    /// Type of arrows in the walking functor.
577    #[derive(Clone, Copy, Debug, PartialEq, Eq)]
578    pub enum Arr {
579        /// Identity arrow on the object [`Zero`](Ob::Zero).
580        Zero,
581        /// Identity arrow on the object [`One`](Ob::One).
582        One,
583        /// Generating arrow from [`Zero`](Ob::Zero) to [`One`](Ob::One).
584        Arrow,
585    }
586
587    impl Arr {
588        fn dom(self) -> Ob {
589            match self {
590                Arr::Zero => Ob::Zero,
591                Arr::Arrow => Ob::Zero,
592                Arr::One => Ob::One,
593            }
594        }
595
596        fn cod(self) -> Ob {
597            match self {
598                Arr::Zero => Ob::Zero,
599                Arr::Arrow => Ob::One,
600                Arr::One => Ob::One,
601            }
602        }
603    }
604
605    /// Type of cells in the walking functor.
606    #[derive(Clone, Copy, Debug, PartialEq, Eq)]
607    pub enum Cell {
608        /// Cell bounded by identity arrows on the object [`Zero`](Ob::Zero).
609        Zero(usize),
610        /// Cell bounded by identity arrows on the object [`One`](Ob::One).
611        One(usize),
612        /// Cell bounded by the generating arrow.
613        Arrow(usize),
614    }
615
616    impl Cell {
617        fn with_src_and_tgt(f: Arr, n: usize) -> Self {
618            match f {
619                Arr::Zero => Cell::Zero(n),
620                Arr::Arrow => Cell::Arrow(n),
621                Arr::One => Cell::One(n),
622            }
623        }
624
625        fn dom(self) -> Path<Ob, Ob> {
626            let (ob, n) = match self {
627                Cell::Zero(n) => (Ob::Zero, n),
628                Cell::Arrow(n) => (Ob::Zero, n),
629                Cell::One(n) => (Ob::One, n),
630            };
631            Path::repeat_n(ob, ob, n)
632        }
633
634        fn cod(self) -> Ob {
635            match self {
636                Cell::Zero(_) => Ob::Zero,
637                Cell::Arrow(_) => Ob::One,
638                Cell::One(_) => Ob::One,
639            }
640        }
641
642        fn src(self) -> Arr {
643            match self {
644                Cell::Zero(_) => Arr::Zero,
645                Cell::Arrow(_) => Arr::Arrow,
646                Cell::One(_) => Arr::One,
647            }
648        }
649
650        fn tgt(self) -> Arr {
651            self.src()
652        }
653
654        fn map<F>(self, f: F) -> Self
655        where
656            F: FnOnce(usize) -> usize,
657        {
658            match self {
659                Cell::Zero(n) => Cell::Zero(f(n)),
660                Cell::Arrow(n) => Cell::Arrow(f(n)),
661                Cell::One(n) => Cell::One(f(n)),
662            }
663        }
664    }
665
666    impl VDblCategory for Main {
667        type Ob = Ob;
668        type Arr = Arr;
669        type Pro = Ob;
670        type Cell = Cell;
671
672        fn has_ob(&self, _: &Self::Ob) -> bool {
673            true
674        }
675        fn has_arrow(&self, _: &Self::Arr) -> bool {
676            true
677        }
678        fn has_proarrow(&self, _: &Self::Pro) -> bool {
679            true
680        }
681        fn has_cell(&self, _: &Self::Cell) -> bool {
682            true
683        }
684
685        fn dom(&self, f: &Self::Arr) -> Self::Ob {
686            f.dom()
687        }
688        fn cod(&self, f: &Self::Arr) -> Self::Ob {
689            f.cod()
690        }
691        fn src(&self, m: &Self::Pro) -> Self::Ob {
692            *m
693        }
694        fn tgt(&self, m: &Self::Pro) -> Self::Ob {
695            *m
696        }
697
698        fn cell_dom(&self, cell: &Self::Cell) -> Path<Self::Ob, Self::Pro> {
699            cell.dom()
700        }
701        fn cell_cod(&self, cell: &Self::Cell) -> Self::Pro {
702            cell.cod()
703        }
704        fn cell_src(&self, cell: &Self::Cell) -> Self::Arr {
705            cell.src()
706        }
707        fn cell_tgt(&self, cell: &Self::Cell) -> Self::Arr {
708            cell.tgt()
709        }
710
711        fn compose(&self, path: Path<Self::Ob, Self::Arr>) -> Self::Arr {
712            assert!(path.contained_in(EdgeGraph::ref_cast(UnderlyingDblGraph::ref_cast(self))));
713            match path {
714                Path::Id(Ob::Zero) => Arr::Zero,
715                Path::Id(Ob::One) => Arr::One,
716                Path::Seq(arrows) => {
717                    *arrows.iter().find(|f| **f == Arr::Arrow).unwrap_or_else(|| arrows.first())
718                }
719            }
720        }
721        fn compose_cells(&self, tree: DblTree<Self::Arr, Self::Pro, Self::Cell>) -> Self::Cell {
722            let graph = UnderlyingDblGraph::ref_cast(self);
723            let n = tree.arity(graph);
724            let (f, g) = (self.compose(tree.src(graph)), self.compose(tree.tgt(graph)));
725            assert_eq!(f, g, "Cells in walking functor have the same source and target");
726            Cell::with_src_and_tgt(f, n)
727        }
728    }
729
730    impl VDCWithComposites for Main {
731        fn composite_ext(&self, path: Path<Self::Ob, Self::Pro>) -> Option<Self::Cell> {
732            let graph = ProedgeGraph::ref_cast(UnderlyingDblGraph::ref_cast(self));
733            let (x, y) = (path.src(graph), path.tgt(graph));
734            assert_eq!(x, y, "Paths in walking functor have the same source and target");
735            Some(Cell::with_src_and_tgt(self.id(x), path.len()))
736        }
737
738        fn through_composite(&self, cell: Self::Cell, range: Range<usize>) -> Option<Self::Cell> {
739            Some(cell.map(|n| n + 1 - range.len()))
740        }
741    }
742}
743
744#[cfg(test)]
745mod tests {
746    use super::*;
747
748    #[test]
749    fn walking_category() {
750        let vdc = WalkingCategory();
751        assert!(vdc.has_unit(&()));
752        assert_eq!(vdc.unit(()), Some(()));
753        assert_eq!(vdc.unit_ext(()), Some(0));
754        assert_eq!(vdc.through_unit(2, 1), Some(3));
755        assert_eq!(vdc.cell_dom(&0), Path::empty(()));
756        assert_eq!(vdc.cell_dom(&2), Path::pair((), ()));
757    }
758
759    #[test]
760    fn walking_bimodule() {
761        use WalkingBimodule::{Ob, Pro};
762
763        let vdc = WalkingBimodule::Main();
764        assert!(vdc.has_unit(&Ob::Left));
765        assert_eq!(vdc.unit(Ob::Left), Some(Pro::Left));
766        let ext = vdc.unit_ext(Ob::Left).unwrap();
767        assert_eq!(vdc.cell_dom(&ext), Path::empty(Ob::Left));
768        assert_eq!(vdc.cell_cod(&ext), Pro::Left);
769
770        let path = Path::from_vec(vec![Pro::Left, Pro::Middle, Pro::Right]).unwrap();
771        assert!(vdc.has_cell(&Path::single(Pro::Middle)));
772        assert!(vdc.has_cell(&path));
773        assert!(!vdc.has_cell(&Path::pair(Pro::Left, Pro::Right)));
774        assert_eq!(vdc.composite(Path::empty(Ob::Left)), Some(Pro::Left));
775        assert_eq!(vdc.composite(Path::empty(Ob::Right)), Some(Pro::Right));
776        assert_eq!(vdc.composite(path), Some(Pro::Middle));
777
778        assert_eq!(
779            vdc.through_unit(Pro::Middle.into(), 0),
780            Some(Path::pair(Pro::Left, Pro::Middle))
781        );
782        assert_eq!(
783            vdc.through_unit(Pro::Middle.into(), 1),
784            Some(Path::pair(Pro::Middle, Pro::Right))
785        );
786    }
787
788    #[test]
789    fn walking_functor() {
790        use WalkingFunctor::{Arr, Cell, Ob};
791
792        let vdc = WalkingFunctor::Main();
793        let cell = Cell::Arrow(2);
794        assert_eq!(vdc.cell_dom(&cell), Path::pair(Ob::Zero, Ob::Zero));
795        assert_eq!(vdc.cell_cod(&cell), Ob::One);
796        assert_eq!(vdc.cell_src(&cell), Arr::Arrow);
797        assert_eq!(vdc.cell_tgt(&cell), Arr::Arrow);
798
799        let ext = vdc.composite_ext(Path::pair(Ob::Zero, Ob::Zero)).unwrap();
800        assert_eq!(vdc.cell_src(&ext), Arr::Zero);
801        assert_eq!(vdc.cell_tgt(&ext), Arr::Zero);
802        let new_cell = vdc.compose_cells2(vec![ext, ext], cell);
803        assert_eq!(vdc.cell_dom(&new_cell).len(), 4);
804
805        assert_eq!(vdc.through_unit(Cell::Arrow(2), 1), Some(Cell::Arrow(3)));
806    }
807}