├── .github ├── dependabot.yml └── workflows │ └── ci-ubuntu.yml ├── .gitignore ├── CITATION.cff ├── LICENSE ├── Makefile ├── README.md ├── README └── KanComplexes.md ├── agda-categories.agda-lib ├── src ├── Categories │ ├── 2-Category.agda │ ├── 2-Functor.agda │ ├── Adjoint.agda │ ├── Adjoint │ │ ├── AFT.agda │ │ ├── AFT │ │ │ └── SolutionSet.agda │ │ ├── Alternatives.agda │ │ ├── Compose.agda │ │ ├── Construction │ │ │ ├── CoEilenbergMoore.agda │ │ │ ├── CoKleisli.agda │ │ │ ├── EilenbergMoore.agda │ │ │ └── Kleisli.agda │ │ ├── Equivalence.agda │ │ ├── Equivalents.agda │ │ ├── Instance │ │ │ ├── 0-Truncation.agda │ │ │ ├── 01-Truncation.agda │ │ │ ├── BaseChange.agda │ │ │ ├── Core.agda │ │ │ ├── PathsOf.agda │ │ │ ├── PosetCore.agda │ │ │ ├── Slice.agda │ │ │ ├── StrictCore.agda │ │ │ └── StrictDiscrete.agda │ │ ├── Mate.agda │ │ ├── Monadic.agda │ │ ├── Monadic │ │ │ ├── Crude.agda │ │ │ └── Properties.agda │ │ ├── Parametric.agda │ │ ├── Properties.agda │ │ ├── RAPL.agda │ │ ├── Relative.agda │ │ ├── TwoSided.agda │ │ └── TwoSided │ │ │ └── Compose.agda │ ├── Bicategory.agda │ ├── Bicategory │ │ ├── Bigroupoid.agda │ │ ├── Construction │ │ │ ├── 1-Category.agda │ │ │ ├── LaxSlice.agda │ │ │ ├── Spans.agda │ │ │ └── Spans │ │ │ │ └── Properties.agda │ │ ├── Extras.agda │ │ ├── Instance │ │ │ ├── Cats.agda │ │ │ └── EnrichedCats.agda │ │ ├── Monad.agda │ │ ├── Monad │ │ │ └── Properties.agda │ │ ├── Object │ │ │ ├── Product.agda │ │ │ └── Terminal.agda │ │ └── Opposite.agda │ ├── Category.agda │ ├── Category │ │ ├── BicartesianClosed.agda │ │ ├── BinaryProducts.agda │ │ ├── CMonoidEnriched.agda │ │ ├── Cartesian.agda │ │ ├── Cartesian │ │ │ ├── Bundle.agda │ │ │ ├── Monoidal.agda │ │ │ ├── Properties.agda │ │ │ └── SymmetricMonoidal.agda │ │ ├── CartesianClosed.agda │ │ ├── CartesianClosed │ │ │ ├── Bundle.agda │ │ │ ├── Canonical.agda │ │ │ ├── Locally.agda │ │ │ ├── Locally │ │ │ │ └── Properties.agda │ │ │ └── Properties.agda │ │ ├── Closed.agda │ │ ├── Cocartesian.agda │ │ ├── Cocartesian │ │ │ └── Bundle.agda │ │ ├── Cocomplete.agda │ │ ├── Cocomplete │ │ │ ├── Finitely.agda │ │ │ ├── Finitely │ │ │ │ └── Properties.agda │ │ │ ├── Properties.agda │ │ │ └── Properties │ │ │ │ └── Construction.agda │ │ ├── Complete.agda │ │ ├── Complete │ │ │ ├── Finitely.agda │ │ │ ├── Finitely │ │ │ │ └── Properties.agda │ │ │ ├── Properties.agda │ │ │ └── Properties │ │ │ │ ├── Construction.agda │ │ │ │ └── SolutionSet.agda │ │ ├── Concrete.agda │ │ ├── Concrete │ │ │ └── Properties.agda │ │ ├── Construction │ │ │ ├── 0-Groupoid.agda │ │ │ ├── Adjoints.agda │ │ │ ├── Arrow.agda │ │ │ ├── CoEilenbergMoore.agda │ │ │ ├── CoKleisli.agda │ │ │ ├── Cocones.agda │ │ │ ├── Comma.agda │ │ │ ├── Cones.agda │ │ │ ├── Coproduct.agda │ │ │ ├── Core.agda │ │ │ ├── Cowedges.agda │ │ │ ├── DaggerFunctors.agda │ │ │ ├── EilenbergMoore.agda │ │ │ ├── Elements.agda │ │ │ ├── EnrichedFunctors.agda │ │ │ ├── F-Algebras.agda │ │ │ ├── F-Coalgebras.agda │ │ │ ├── Fin.agda │ │ │ ├── Functors.agda │ │ │ ├── Grothendieck.agda │ │ │ ├── GroupAsCategory.agda │ │ │ ├── Groups.agda │ │ │ ├── KanComplex.agda │ │ │ ├── KaroubiEnvelope.agda │ │ │ ├── KaroubiEnvelope │ │ │ │ └── Properties.agda │ │ │ ├── Kleisli.agda │ │ │ ├── LT-Models.agda │ │ │ ├── MonoidAsCategory.agda │ │ │ ├── MonoidalFunctors.agda │ │ │ ├── Monoids.agda │ │ │ ├── ObjectRestriction.agda │ │ │ ├── Path.agda │ │ │ ├── PathCategory.agda │ │ │ ├── Presheaves.agda │ │ │ ├── Properties │ │ │ │ ├── CoEilenbergMoore.agda │ │ │ │ ├── CoKleisli.agda │ │ │ │ ├── Comma.agda │ │ │ │ ├── EilenbergMoore.agda │ │ │ │ ├── Functors.agda │ │ │ │ ├── Kleisli.agda │ │ │ │ ├── Presheaves.agda │ │ │ │ └── Presheaves │ │ │ │ │ ├── Cartesian.agda │ │ │ │ │ ├── CartesianClosed.agda │ │ │ │ │ ├── Complete.agda │ │ │ │ │ └── FromCartesianCCC.agda │ │ │ ├── Pullbacks.agda │ │ │ ├── SetoidDiscrete.agda │ │ │ ├── Spans.agda │ │ │ ├── StrictDiscrete.agda │ │ │ ├── SymmetricMonoidalFunctors.agda │ │ │ ├── Thin.agda │ │ │ ├── TwistedArrow.agda │ │ │ ├── Wedges.agda │ │ │ └── mu-Bialgebras.agda │ │ ├── Core.agda │ │ ├── Dagger.agda │ │ ├── Dagger │ │ │ ├── Construction │ │ │ │ ├── DaggerFunctors.agda │ │ │ │ └── Discrete.agda │ │ │ └── Instance │ │ │ │ └── Rels.agda │ │ ├── Diagram │ │ │ └── Span.agda │ │ ├── Discrete.agda │ │ ├── Distributive.agda │ │ ├── Distributive │ │ │ └── Properties.agda │ │ ├── Duality.agda │ │ ├── Equivalence.agda │ │ ├── Equivalence │ │ │ ├── Preserves.agda │ │ │ └── Properties.agda │ │ ├── Exact.agda │ │ ├── Extensive.agda │ │ ├── Extensive │ │ │ ├── Bundle.agda │ │ │ └── Properties │ │ │ │ └── Distributive.agda │ │ ├── Finite.agda │ │ ├── Finite │ │ │ ├── Fin.agda │ │ │ └── Fin │ │ │ │ ├── Construction │ │ │ │ ├── Discrete.agda │ │ │ │ └── Poset.agda │ │ │ │ └── Instance │ │ │ │ ├── Parallel.agda │ │ │ │ ├── Span.agda │ │ │ │ └── Triangle.agda │ │ ├── Groupoid.agda │ │ ├── Groupoid │ │ │ └── Properties.agda │ │ ├── Helper.agda │ │ ├── Indiscrete.agda │ │ ├── Instance │ │ │ ├── Cartesians.agda │ │ │ ├── Cats.agda │ │ │ ├── DagCats.agda │ │ │ ├── EmptySet.agda │ │ │ ├── FamilyOfSetoids.agda │ │ │ ├── FinCatShapes.agda │ │ │ ├── FinSetoids.agda │ │ │ ├── Globe.agda │ │ │ ├── Groupoids.agda │ │ │ ├── KanComplexes.agda │ │ │ ├── LawvereTheories.agda │ │ │ ├── Monoidals.agda │ │ │ ├── Nat.agda │ │ │ ├── One.agda │ │ │ ├── PartialFunctions.agda │ │ │ ├── PointedSets.agda │ │ │ ├── Posets.agda │ │ │ ├── Preds.agda │ │ │ ├── Properties │ │ │ │ ├── Cats.agda │ │ │ │ ├── Posets.agda │ │ │ │ ├── Setoids.agda │ │ │ │ └── Setoids │ │ │ │ │ ├── CCC.agda │ │ │ │ │ ├── Choice.agda │ │ │ │ │ ├── Cocomplete.agda │ │ │ │ │ ├── Complete.agda │ │ │ │ │ ├── Exact.agda │ │ │ │ │ ├── Extensive.agda │ │ │ │ │ ├── LCCC.agda │ │ │ │ │ └── Limits │ │ │ │ │ └── Canonical.agda │ │ │ ├── Quivers.agda │ │ │ ├── Rels.agda │ │ │ ├── Setoids.agda │ │ │ ├── Sets.agda │ │ │ ├── Simplex.agda │ │ │ ├── SimplicialSet.agda │ │ │ ├── SimplicialSet │ │ │ │ └── Properties.agda │ │ │ ├── SingletonSet.agda │ │ │ ├── Span.agda │ │ │ ├── StrictCats.agda │ │ │ ├── StrictGroupoids.agda │ │ │ ├── Zero.agda │ │ │ └── Zero │ │ │ │ ├── Core.agda │ │ │ │ └── Properties.agda │ │ ├── Inverse.agda │ │ ├── Lift.agda │ │ ├── Lift │ │ │ └── Properties.agda │ │ ├── Monoidal.agda │ │ ├── Monoidal │ │ │ ├── Braided.agda │ │ │ ├── Braided │ │ │ │ └── Properties.agda │ │ │ ├── Bundle.agda │ │ │ ├── Closed.agda │ │ │ ├── Closed │ │ │ │ ├── IsClosed.agda │ │ │ │ └── IsClosed │ │ │ │ │ ├── Diagonal.agda │ │ │ │ │ ├── Dinatural.agda │ │ │ │ │ ├── Identity.agda │ │ │ │ │ ├── L.agda │ │ │ │ │ └── Pentagon.agda │ │ │ ├── CompactClosed.agda │ │ │ ├── Construction │ │ │ │ ├── Endofunctors.agda │ │ │ │ ├── Minus2.agda │ │ │ │ ├── Product.agda │ │ │ │ └── Reverse.agda │ │ │ ├── Core.agda │ │ │ ├── Instance │ │ │ │ ├── Cats.agda │ │ │ │ ├── One.agda │ │ │ │ ├── Rels.agda │ │ │ │ ├── Setoids.agda │ │ │ │ ├── Sets.agda │ │ │ │ └── StrictCats.agda │ │ │ ├── Interchange.agda │ │ │ ├── Interchange │ │ │ │ ├── Braided.agda │ │ │ │ └── Symmetric.agda │ │ │ ├── Properties.agda │ │ │ ├── Reasoning.agda │ │ │ ├── Rigid.agda │ │ │ ├── Star-Autonomous.agda │ │ │ ├── Symmetric.agda │ │ │ ├── Symmetric │ │ │ │ └── Properties.agda │ │ │ ├── Traced.agda │ │ │ └── Utilities.agda │ │ ├── Product.agda │ │ ├── Product │ │ │ └── Properties.agda │ │ ├── Regular.agda │ │ ├── Restriction.agda │ │ ├── Restriction │ │ │ ├── Construction │ │ │ │ └── Trivial.agda │ │ │ ├── Instance │ │ │ │ └── PartialFunctions.agda │ │ │ ├── Properties.agda │ │ │ └── Properties │ │ │ │ └── Poset.agda │ │ ├── RigCategory.agda │ │ ├── Site.agda │ │ ├── Slice.agda │ │ ├── Slice │ │ │ └── Properties.agda │ │ ├── Species.agda │ │ ├── Species │ │ │ └── Constructions.agda │ │ ├── SubCategory.agda │ │ ├── Topos.agda │ │ ├── Unbundled.agda │ │ └── Unbundled │ │ │ ├── Properties.agda │ │ │ └── Utilities.agda │ ├── CoYoneda.agda │ ├── Comonad.agda │ ├── Comonad │ │ ├── Construction │ │ │ └── CoKleisli.agda │ │ ├── Morphism.agda │ │ └── Relative.agda │ ├── Diagram │ │ ├── Cocone.agda │ │ ├── Cocone │ │ │ └── Properties.agda │ │ ├── Coend.agda │ │ ├── Coend │ │ │ ├── Colimit.agda │ │ │ └── Properties.agda │ │ ├── Coequalizer.agda │ │ ├── Coequalizer │ │ │ └── Properties.agda │ │ ├── Colimit.agda │ │ ├── Colimit │ │ │ ├── DualProperties.agda │ │ │ ├── Lan.agda │ │ │ └── Properties.agda │ │ ├── Cone.agda │ │ ├── Cone │ │ │ └── Properties.agda │ │ ├── Cowedge.agda │ │ ├── Cowedge │ │ │ └── Properties.agda │ │ ├── Duality.agda │ │ ├── Empty.agda │ │ ├── End.agda │ │ ├── End │ │ │ ├── Fubini.agda │ │ │ ├── Instance │ │ │ │ └── NaturalTransformations.agda │ │ │ ├── Instances │ │ │ │ └── NaturalTransformation.agda │ │ │ ├── Limit.agda │ │ │ ├── Parameterized.agda │ │ │ └── Properties.agda │ │ ├── Equalizer.agda │ │ ├── Equalizer │ │ │ ├── Indexed.agda │ │ │ ├── Limit.agda │ │ │ └── Properties.agda │ │ ├── Finite.agda │ │ ├── KernelPair.agda │ │ ├── Limit.agda │ │ ├── Limit │ │ │ ├── Properties.agda │ │ │ └── Ran.agda │ │ ├── Pullback.agda │ │ ├── Pullback │ │ │ ├── Limit.agda │ │ │ └── Properties.agda │ │ ├── Pushout.agda │ │ ├── Pushout │ │ │ └── Properties.agda │ │ ├── ReflexivePair.agda │ │ ├── SubobjectClassifier.agda │ │ ├── Wedge.agda │ │ └── Wedge │ │ │ └── Properties.agda │ ├── Double.agda │ ├── Double │ │ └── Core.agda │ ├── Enriched │ │ ├── Category.agda │ │ ├── Category │ │ │ ├── Opposite.agda │ │ │ └── Underlying.agda │ │ ├── Functor.agda │ │ ├── NaturalTransformation.agda │ │ ├── NaturalTransformation │ │ │ └── NaturalIsomorphism.agda │ │ └── Over │ │ │ ├── One.agda │ │ │ └── Setoids.agda │ ├── FreeObjects │ │ └── Free.agda │ ├── Functor.agda │ ├── Functor │ │ ├── Algebra.agda │ │ ├── Bialgebra.agda │ │ ├── Bifunctor.agda │ │ ├── Bifunctor │ │ │ └── Properties.agda │ │ ├── Cartesian.agda │ │ ├── Cartesian │ │ │ └── Properties.agda │ │ ├── CartesianClosed.agda │ │ ├── Coalgebra.agda │ │ ├── Construction │ │ │ ├── Constant.agda │ │ │ ├── Diagonal.agda │ │ │ ├── FromDiscrete.agda │ │ │ ├── LiftAlgebras.agda │ │ │ ├── LiftCoalgebras.agda │ │ │ ├── LiftSetoids.agda │ │ │ ├── Limit.agda │ │ │ ├── ObjectRestriction.agda │ │ │ ├── PathsOf.agda │ │ │ ├── SubCategory.agda │ │ │ ├── SubCategory │ │ │ │ └── Properties.agda │ │ │ └── Zero.agda │ │ ├── Core.agda │ │ ├── Dagger.agda │ │ ├── DistributiveLaw.agda │ │ ├── Duality.agda │ │ ├── Equivalence.agda │ │ ├── Fibration.agda │ │ ├── Groupoid.agda │ │ ├── Hom.agda │ │ ├── Hom │ │ │ ├── Properties.agda │ │ │ └── Properties │ │ │ │ ├── Contra.agda │ │ │ │ └── Covariant.agda │ │ ├── IdentityOnObjects.agda │ │ ├── Instance │ │ │ ├── 0-Truncation.agda │ │ │ ├── 01-Truncation.agda │ │ │ ├── ConnectedComponents.agda │ │ │ ├── Core.agda │ │ │ ├── Discrete.agda │ │ │ ├── SetoidDiscrete.agda │ │ │ ├── StrictCore.agda │ │ │ ├── Twisted.agda │ │ │ └── UnderlyingQuiver.agda │ │ ├── Limits.agda │ │ ├── Monoidal.agda │ │ ├── Monoidal │ │ │ ├── Braided.agda │ │ │ ├── Construction │ │ │ │ └── Product.agda │ │ │ ├── PointwiseTensor.agda │ │ │ ├── Properties.agda │ │ │ ├── Symmetric.agda │ │ │ └── Tensor.agda │ │ ├── Power.agda │ │ ├── Power │ │ │ ├── Functorial.agda │ │ │ └── NaturalTransformation.agda │ │ ├── Presheaf.agda │ │ ├── Profunctor.agda │ │ ├── Profunctor │ │ │ ├── Cograph.agda │ │ │ └── FormalComposite.agda │ │ ├── Properties.agda │ │ ├── Representable.agda │ │ ├── Restriction.agda │ │ ├── Slice.agda │ │ └── Slice │ │ │ └── BaseChange.agda │ ├── GlobularSet.agda │ ├── Kan.agda │ ├── Kan │ │ └── Duality.agda │ ├── Minus2-Category.agda │ ├── Minus2-Category │ │ ├── Construction │ │ │ └── Indiscrete.agda │ │ ├── Instance │ │ │ └── One.agda │ │ └── Properties.agda │ ├── Monad.agda │ ├── Monad │ │ ├── Commutative.agda │ │ ├── Construction │ │ │ └── Kleisli.agda │ │ ├── Duality.agda │ │ ├── Graded.agda │ │ ├── Idempotent.agda │ │ ├── Morphism.agda │ │ ├── Relative.agda │ │ ├── Strong.agda │ │ └── Strong │ │ │ └── Properties.agda │ ├── Morphism.agda │ ├── Morphism │ │ ├── Cartesian.agda │ │ ├── Duality.agda │ │ ├── Extremal.agda │ │ ├── Extremal │ │ │ └── Properties.agda │ │ ├── HeterogeneousEquality.agda │ │ ├── HeterogeneousIdentity.agda │ │ ├── HeterogeneousIdentity │ │ │ └── Properties.agda │ │ ├── Idempotent.agda │ │ ├── Idempotent │ │ │ └── Bundles.agda │ │ ├── IsoEquiv.agda │ │ ├── Isomorphism.agda │ │ ├── Lifts.agda │ │ ├── Lifts │ │ │ └── Properties.agda │ │ ├── Normal.agda │ │ ├── Notation.agda │ │ ├── Properties.agda │ │ ├── Reasoning.agda │ │ ├── Reasoning │ │ │ ├── Core.agda │ │ │ └── Iso.agda │ │ ├── Regular.agda │ │ ├── Regular │ │ │ └── Properties.agda │ │ └── Universal.agda │ ├── Multi │ │ └── Category │ │ │ └── Indexed.agda │ ├── NaturalTransformation.agda │ ├── NaturalTransformation │ │ ├── Core.agda │ │ ├── Dinatural.agda │ │ ├── Equivalence.agda │ │ ├── Extranatural.agda │ │ ├── Hom.agda │ │ ├── Monoidal.agda │ │ ├── Monoidal │ │ │ ├── Braided.agda │ │ │ └── Symmetric.agda │ │ ├── NaturalIsomorphism.agda │ │ ├── NaturalIsomorphism │ │ │ ├── Equivalence.agda │ │ │ ├── Functors.agda │ │ │ ├── Monoidal.agda │ │ │ ├── Monoidal │ │ │ │ ├── Braided.agda │ │ │ │ └── Symmetric.agda │ │ │ └── Properties.agda │ │ └── Properties.agda │ ├── Object │ │ ├── Biproduct.agda │ │ ├── Cokernel.agda │ │ ├── Coproduct.agda │ │ ├── Coproduct │ │ │ ├── Indexed.agda │ │ │ └── Indexed │ │ │ │ └── Properties.agda │ │ ├── Duality.agda │ │ ├── Exponential.agda │ │ ├── Group.agda │ │ ├── Initial.agda │ │ ├── Initial │ │ │ └── Colimit.agda │ │ ├── InternalRelation.agda │ │ ├── Kernel.agda │ │ ├── Kernel │ │ │ └── Properties.agda │ │ ├── Monoid.agda │ │ ├── NaturalNumbers.agda │ │ ├── NaturalNumbers │ │ │ ├── Parametrized.agda │ │ │ ├── Parametrized │ │ │ │ └── Properties │ │ │ │ │ └── F-Algebras.agda │ │ │ └── Properties │ │ │ │ ├── F-Algebras.agda │ │ │ │ └── Parametrized.agda │ │ ├── Product.agda │ │ ├── Product │ │ │ ├── Construction.agda │ │ │ ├── Core.agda │ │ │ ├── Indexed.agda │ │ │ ├── Indexed │ │ │ │ └── Properties.agda │ │ │ ├── Limit.agda │ │ │ └── Morphisms.agda │ │ ├── StrictInitial.agda │ │ ├── Subobject.agda │ │ ├── Subobject │ │ │ └── Properties.agda │ │ ├── Terminal.agda │ │ ├── Terminal │ │ │ └── Limit.agda │ │ └── Zero.agda │ ├── Pseudofunctor.agda │ ├── Pseudofunctor │ │ ├── Composition.agda │ │ ├── Hom.agda │ │ ├── Identity.agda │ │ └── Instance │ │ │ └── EnrichedUnderlying.agda │ ├── Tactic │ │ └── Category.agda │ ├── Theory │ │ ├── Lawvere.agda │ │ └── Lawvere │ │ │ └── Instance │ │ │ ├── Identity.agda │ │ │ └── Triv.agda │ ├── Utils │ │ ├── EqReasoning.agda │ │ └── Product.agda │ ├── Yoneda.agda │ └── Yoneda │ │ ├── Continuous.agda │ │ └── Properties.agda ├── Data │ ├── Quiver.agda │ └── Quiver │ │ ├── Morphism.agda │ │ └── Paths.agda ├── Function │ └── Construct │ │ └── Setoid.agda └── Relation │ └── Binary │ └── PropositionalEquality │ └── Subst │ └── Properties.agda └── travis ├── everything.sh ├── index.agda └── index.sh /.github/dependabot.yml: -------------------------------------------------------------------------------- 1 | version: 2 2 | updates: 3 | - package-ecosystem: "github-actions" 4 | directory: "/" 5 | schedule: 6 | # Check for updates to GitHub Actions every weekday 7 | interval: "daily" 8 | labels: 9 | - "dependencies" 10 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | .DS_Store 3 | src/MAlonzo 4 | Everything.agda 5 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2019 Agda Github Community 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | .PHONY: test Everything.agda clean 2 | 3 | OTHEROPTS = --auto-inline -Werror 4 | 5 | RTSARGS = +RTS -M6G -A128M -RTS ${OTHEROPTS} 6 | 7 | test: Everything.agda 8 | agda ${RTSARGS} -i. Everything.agda 9 | 10 | html: Everything.agda 11 | agda ${RTSARGS} --html -i. Everything.agda 12 | 13 | Everything.agda: 14 | git ls-tree --full-tree -r --name-only HEAD | grep '^src/[^\.]*.agda' | sed -e 's|^src/[/]*|import |' -e 's|/|.|g' -e 's/.agda//' -e '/import Everything/d' | LC_COLLATE='C' sort > Everything.agda 15 | 16 | clean: 17 | find . -name '*.agdai' -exec rm \{\} \; 18 | 19 | profile: Everything.agda 20 | agda ${RTSARGS} -v profile:7 -v profile.definitions:15 -i. Everything.agda 21 | -------------------------------------------------------------------------------- /README/KanComplexes.md: -------------------------------------------------------------------------------- 1 | # Kan Complexes 2 | 3 | The classical definition of Kan Complex is a bit tricky in Agda, as 4 | they explicitly don't specify a choice of filler for horns, merely their 5 | existence. Instead, we have opted to implement *Algebraic Kan Complexes*, 6 | which are Kan Complexes with a distinguished choice of filler. This does 7 | mean that the Category of Kan Complexes needs morphisms to preserve these 8 | distinguished fillers. 9 | 10 | ## Implementation Notes 11 | * The inequality in the definition of Horn seems like it could become a bit awkward down the line? 12 | As equality of `Fin` is decidable, it isn't *too* bad, but does pose the potential to be annoying. 13 | * The unfolded equality in ∂Δ[_] and Λ[_,_] does better than just lifting the equality from Δ, but it still falls on it's face from time to time 14 | 15 | ## References 16 | * https://ncatlab.org/nlab/show/Kan+complex 17 | * https://ncatlab.org/nlab/show/weak+Kan+complex 18 | * https://ncatlab.org/nlab/show/algebraic+Kan+complex 19 | -------------------------------------------------------------------------------- /agda-categories.agda-lib: -------------------------------------------------------------------------------- 1 | name: agda-categories 2 | depend: standard-library-2.1 3 | include: src/ 4 | -------------------------------------------------------------------------------- /src/Categories/2-Category.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.2-Category where 4 | 5 | open import Level 6 | 7 | open import Categories.Category.Monoidal.Instance.StrictCats using (module Product) 8 | open import Categories.Enriched.Category using (Category) 9 | 10 | 2-Category : (o ℓ e t : Level) → Set (suc (o ⊔ ℓ ⊔ e ⊔ t)) 11 | 2-Category o ℓ e t = Category (Product.Cats-Monoidal {o} {ℓ} {e}) t 12 | -------------------------------------------------------------------------------- /src/Categories/2-Functor.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.2-Functor where 4 | 5 | open import Level 6 | 7 | open import Categories.Category.Monoidal.Instance.StrictCats using (module Product) 8 | open import Categories.2-Category using (2-Category) 9 | open import Categories.Enriched.Functor using (Functor) 10 | 11 | 2-Functor : ∀ {o ℓ e c d} → 12 | 2-Category o ℓ e c → 2-Category o ℓ e d → Set (o ⊔ ℓ ⊔ e ⊔ c ⊔ d) 13 | 2-Functor C D = Functor (Product.Cats-Monoidal) C D 14 | -------------------------------------------------------------------------------- /src/Categories/Adjoint/AFT/SolutionSet.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Functor 5 | 6 | module Categories.Adjoint.AFT.SolutionSet {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} 7 | (F : Functor C D) where 8 | 9 | open import Level 10 | 11 | private 12 | module C = Category C 13 | module D = Category D 14 | module F = Functor F 15 | open D 16 | open F 17 | 18 | record SolutionSet i : Set (o ⊔ ℓ ⊔ o′ ⊔ ℓ′ ⊔ e′ ⊔ (suc i)) where 19 | field 20 | I : Set i 21 | S : I → C.Obj 22 | S₀ : ∀ {A X} → X ⇒ F₀ A → I 23 | S₁ : ∀ {A X} (f : X ⇒ F₀ A) → S (S₀ f) C.⇒ A 24 | ϕ : ∀ {A X} (f : X ⇒ F₀ A) → X ⇒ F₀ (S (S₀ f)) 25 | commute : ∀ {A X} (f : X ⇒ F₀ A) → F₁ (S₁ f) ∘ ϕ f ≈ f 26 | 27 | 28 | record SolutionSet′ : Set (o ⊔ ℓ ⊔ o′ ⊔ ℓ′ ⊔ e′) where 29 | field 30 | S₀ : ∀ {A X} → X ⇒ F₀ A → C.Obj 31 | S₁ : ∀ {A X} (f : X ⇒ F₀ A) → S₀ f C.⇒ A 32 | ϕ : ∀ {A X} (f : X ⇒ F₀ A) → X ⇒ F₀ (S₀ f) 33 | commute : ∀ {A X} (f : X ⇒ F₀ A) → F₁ (S₁ f) ∘ ϕ f ≈ f 34 | 35 | SolutionSet⇒SolutionSet′ : ∀ {i} → SolutionSet i → SolutionSet′ 36 | SolutionSet⇒SolutionSet′ s = record 37 | { S₀ = λ f → S (S₀ f) 38 | ; S₁ = S₁ 39 | ; ϕ = ϕ 40 | ; commute = commute 41 | } 42 | where open SolutionSet s 43 | 44 | SolutionSet′⇒SolutionSet : ∀ i → SolutionSet′ → SolutionSet (o ⊔ i) 45 | SolutionSet′⇒SolutionSet i s = record 46 | { I = Lift i C.Obj 47 | ; S = lower 48 | ; S₀ = λ f → lift (S₀ f) 49 | ; S₁ = S₁ 50 | ; ϕ = ϕ 51 | ; commute = commute 52 | } 53 | where open SolutionSet′ s 54 | -------------------------------------------------------------------------------- /src/Categories/Adjoint/Instance/BaseChange.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category) 4 | 5 | module Categories.Adjoint.Instance.BaseChange {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Categories.Adjoint using (_⊣_) 8 | open import Categories.Adjoint.Compose using (_∘⊣_) 9 | open import Categories.Adjoint.Instance.Slice using (TotalSpace⊣ConstantFamily) 10 | open import Categories.Category.Slice C using (Slice) 11 | open import Categories.Category.Slice.Properties C using (pullback⇒product; slice-slice≃slice) 12 | open import Categories.Category.Equivalence.Properties using (module C≅D) 13 | open import Categories.Diagram.Pullback C using (Pullback) 14 | open import Categories.Functor.Slice.BaseChange C using (BaseChange!; BaseChange*) 15 | 16 | open Category C 17 | 18 | module _ {A B : Obj} (f : B ⇒ A) (pullback : ∀ {C} {h : C ⇒ A} → Pullback f h) where 19 | 20 | !⊣* : BaseChange! f ⊣ BaseChange* f pullback 21 | !⊣* = C≅D.L⊣R (slice-slice≃slice f) ∘⊣ TotalSpace⊣ConstantFamily (Slice A) (pullback⇒product pullback) 22 | -------------------------------------------------------------------------------- /src/Categories/Adjoint/Monadic.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Monadic Adjunctions 4 | -- https://ncatlab.org/nlab/show/monadic+adjunction 5 | module Categories.Adjoint.Monadic where 6 | 7 | open import Level 8 | 9 | open import Categories.Adjoint 10 | open import Categories.Adjoint.Properties 11 | open import Categories.Category 12 | open import Categories.Category.Equivalence 13 | open import Categories.Functor 14 | open import Categories.Monad 15 | 16 | open import Categories.Category.Construction.EilenbergMoore 17 | open import Categories.Category.Construction.Properties.EilenbergMoore 18 | 19 | private 20 | variable 21 | o ℓ e : Level 22 | 𝒞 𝒟 : Category o ℓ e 23 | 24 | -- An adjunction is monadic if the adjunction "comes from" the induced monad in some sense. 25 | record IsMonadicAdjunction {L : Functor 𝒞 𝒟} {R : Functor 𝒟 𝒞} (adjoint : L ⊣ R) : Set (levelOfTerm 𝒞 ⊔ levelOfTerm 𝒟) where 26 | private 27 | T : Monad 𝒞 28 | T = adjoint⇒monad adjoint 29 | 30 | field 31 | Comparison⁻¹ : Functor (EilenbergMoore T) 𝒟 32 | comparison-equiv : WeakInverse (ComparisonF adjoint) Comparison⁻¹ 33 | -------------------------------------------------------------------------------- /src/Categories/Bicategory/Monad.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- A Monad in a Bicategory. 4 | -- For the more elementary version of Monads, see 'Categories.Monad'. 5 | module Categories.Bicategory.Monad where 6 | 7 | open import Level 8 | open import Data.Product using (_,_) 9 | 10 | open import Categories.Bicategory 11 | import Categories.Bicategory.Extras as Bicat 12 | open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism) 13 | 14 | 15 | record Monad {o ℓ e t} (𝒞 : Bicategory o ℓ e t) : Set (o ⊔ ℓ ⊔ e ⊔ t) where 16 | open Bicat 𝒞 17 | 18 | field 19 | C : Obj 20 | T : C ⇒₁ C 21 | η : id₁ ⇒₂ T 22 | μ : (T ⊚₀ T) ⇒₂ T 23 | 24 | assoc : μ ∘ᵥ (T ▷ μ) ∘ᵥ associator.from ≈ (μ ∘ᵥ (μ ◁ T)) 25 | sym-assoc : μ ∘ᵥ (μ ◁ T) ∘ᵥ associator.to ≈ (μ ∘ᵥ (T ▷ μ)) 26 | identityˡ : μ ∘ᵥ (T ▷ η) ∘ᵥ unitorʳ.to ≈ id₂ 27 | identityʳ : μ ∘ᵥ (η ◁ T) ∘ᵥ unitorˡ.to ≈ id₂ 28 | -------------------------------------------------------------------------------- /src/Categories/Bicategory/Object/Product.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Bicategory using (Bicategory) 3 | 4 | module Categories.Bicategory.Object.Product {o ℓ e t} (𝒞 : Bicategory o ℓ e t) where 5 | 6 | open import Level 7 | 8 | open Bicategory 𝒞 9 | open import Categories.Category using (_[_,_]) 10 | open import Categories.Morphism using (_≅_) 11 | open import Categories.Morphism.HeterogeneousEquality 12 | open import Categories.Morphism.Notation using (_[_≅_]) 13 | 14 | record Product (A B : Obj) : Set (o ⊔ ℓ ⊔ e ⊔ t) where 15 | infix 10 ⟨_,_⟩₁ ⟨_,_⟩₂ 16 | field 17 | A×B : Obj 18 | πa : A×B ⇒₁ A 19 | πb : A×B ⇒₁ B 20 | ⟨_,_⟩₁ : ∀ {Γ} → Γ ⇒₁ A → Γ ⇒₁ B → Γ ⇒₁ A×B 21 | ⟨_,_⟩₂ : ∀ {Γ}{fa ga : Γ ⇒₁ A}{fb gb : Γ ⇒₁ B} 22 | → fa ⇒₂ ga → fb ⇒₂ gb → ⟨ fa , fb ⟩₁ ⇒₂ ⟨ ga , gb ⟩₁ 23 | 24 | β₁a : ∀ {Γ} f g → hom Γ A [ πa ∘₁ ⟨ f , g ⟩₁ ≅ f ] 25 | β₁b : ∀ {Γ} f g → hom Γ B [ πb ∘₁ ⟨ f , g ⟩₁ ≅ g ] 26 | β₂a : ∀ {Γ}{fa ga fb gb}(αa : hom Γ A [ fa , ga ])(αb : hom Γ B [ fb , gb ]) 27 | → Along β₁a _ _ , β₁a _ _ [ πa ▷ ⟨ αa , αb ⟩₂ ≈ αa ] 28 | β₂b : ∀ {Γ}{fa ga fb gb}(αa : hom Γ A [ fa , ga ])(αb : hom Γ B [ fb , gb ]) 29 | → Along β₁b _ _ , β₁b _ _ [ πb ▷ ⟨ αa , αb ⟩₂ ≈ αb ] 30 | 31 | η₁ : ∀ {Γ} p → hom Γ A×B [ p ≅ ⟨ πa ∘₁ p , πb ∘₁ p ⟩₁ ] 32 | η₂ : ∀ {Γ}{p p'}(ϕ : hom Γ A×B [ p , p' ]) 33 | → Along (η₁ _) , (η₁ _) [ ϕ ≈ ⟨ πa ▷ ϕ , πb ▷ ϕ ⟩₂ ] 34 | -------------------------------------------------------------------------------- /src/Categories/Bicategory/Object/Terminal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Bicategory using (Bicategory) 3 | 4 | module Categories.Bicategory.Object.Terminal {o ℓ e t} (𝒞 : Bicategory o ℓ e t) where 5 | 6 | open Bicategory 𝒞 7 | open import Level 8 | open import Categories.Category using (_[_,_]) 9 | open import Categories.Morphism.HeterogeneousEquality using (Along_,_[_≈_]) 10 | open import Categories.Morphism.Notation using (_[_≅_]) 11 | open import Categories.Morphism using (_≅_) 12 | 13 | record IsTerminal (⊤ : Obj) : Set (o ⊔ ℓ ⊔ e ⊔ t) where 14 | field 15 | !₁ : {A : Obj} → (A ⇒₁ ⊤) 16 | !₂ : {A : Obj} → !₁ {A} ⇒₂ !₁ 17 | 18 | η₁ : ∀ {A} f → hom A ⊤ [ f ≅ !₁ ] 19 | η₂ : ∀ {A}{f g}(α : hom A ⊤ [ f , g ]) 20 | → Along η₁ _ , η₁ _ [ α ≈ !₂ ] 21 | 22 | -------------------------------------------------------------------------------- /src/Categories/Category.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category where 3 | 4 | open import Level 5 | 6 | -- The main definitions are in: 7 | open import Categories.Category.Core public 8 | 9 | private 10 | variable 11 | o ℓ e : Level 12 | 13 | -- Convenience functions for working over multiple categories at once: 14 | -- C [ x , y ] (for x y objects of C) - Hom_C(x , y) 15 | -- C [ f ≈ g ] (for f g arrows of C) - that f and g are equivalent arrows 16 | -- C [ f ∘ g ] (for f g composable arrows of C) - composition in C 17 | infix 10 _[_,_] _[_≈_] _[_∘_] 18 | 19 | _[_,_] : (C : Category o ℓ e) → (X : Category.Obj C) → (Y : Category.Obj C) → Set ℓ 20 | _[_,_] = Category._⇒_ 21 | 22 | _[_≈_] : (C : Category o ℓ e) → ∀ {X Y} (f g : C [ X , Y ]) → Set e 23 | _[_≈_] = Category._≈_ 24 | 25 | _[_∘_] : (C : Category o ℓ e) → ∀ {X Y Z} (f : C [ Y , Z ]) → (g : C [ X , Y ]) → C [ X , Z ] 26 | _[_∘_] = Category._∘_ 27 | 28 | module Definitions (𝓒 : Category o ℓ e) where 29 | open Category 𝓒 30 | 31 | CommutativeSquare : {A B C D : Obj} → (f : A ⇒ B) (g : A ⇒ C) (h : B ⇒ D) (i : C ⇒ D) → Set _ 32 | CommutativeSquare f g h i = h ∘ f ≈ i ∘ g 33 | 34 | -- Combinators for commutative diagram 35 | -- The idea is to use the combinators to write commutations in a more readable way. 36 | -- It starts with [_⇒_]⟨_≈_⟩, and within the third and fourth places, use _⇒⟨_⟩_ to 37 | -- connect morphisms with the intermediate object specified. 38 | module Commutation (𝓒 : Category o ℓ e) where 39 | open Category 𝓒 40 | 41 | infix 1 [_⇒_]⟨_≈_⟩ 42 | [_⇒_]⟨_≈_⟩ : ∀ (A B : Obj) → A ⇒ B → A ⇒ B → Set _ 43 | [ A ⇒ B ]⟨ f ≈ g ⟩ = f ≈ g 44 | 45 | infixl 2 connect 46 | connect : ∀ {A C : Obj} (B : Obj) → A ⇒ B → B ⇒ C → A ⇒ C 47 | connect B f g = g ∘ f 48 | 49 | syntax connect B f g = f ⇒⟨ B ⟩ g 50 | -------------------------------------------------------------------------------- /src/Categories/Category/BicartesianClosed.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core using (Category) 4 | 5 | module Categories.Category.BicartesianClosed {o ℓ e} (𝒞 : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.CartesianClosed 𝒞 10 | open import Categories.Category.Cocartesian 𝒞 11 | 12 | record BicartesianClosed : Set (levelOfTerm 𝒞) where 13 | field 14 | cartesianClosed : CartesianClosed 15 | cocartesian : Cocartesian 16 | -------------------------------------------------------------------------------- /src/Categories/Category/Cartesian.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category) 4 | 5 | -- Defines the following properties of a Category: 6 | -- Cartesian -- a Cartesian category is a category with all products 7 | 8 | -- (for the induced Monoidal structure, see Cartesian.Monoidal) 9 | 10 | module Categories.Category.Cartesian {o ℓ e} (𝒞 : Category o ℓ e) where 11 | 12 | open import Level using (levelOfTerm) 13 | open import Data.Nat using (ℕ; zero; suc) 14 | 15 | open Category 𝒞 16 | open HomReasoning 17 | 18 | open import Categories.Category.BinaryProducts 𝒞 using (BinaryProducts; module BinaryProducts) 19 | open import Categories.Object.Terminal 𝒞 using (Terminal) 20 | 21 | private 22 | variable 23 | A B C D W X Y Z : Obj 24 | f f′ g g′ h i : A ⇒ B 25 | 26 | -- Cartesian monoidal category 27 | record Cartesian : Set (levelOfTerm 𝒞) where 28 | field 29 | terminal : Terminal 30 | products : BinaryProducts 31 | open BinaryProducts products using (_×_) 32 | 33 | power : Obj → ℕ → Obj 34 | power A 0 = Terminal.⊤ terminal 35 | power A 1 = A 36 | power A (suc (suc n)) = A × power A (suc n) 37 | -------------------------------------------------------------------------------- /src/Categories/Category/Cartesian/Bundle.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Bundled version of a Cartesian Category 4 | module Categories.Category.Cartesian.Bundle where 5 | 6 | open import Level 7 | 8 | open import Categories.Category.Core using (Category) 9 | open import Categories.Category.Cartesian using (Cartesian) 10 | open import Categories.Category.Cartesian.Monoidal using (module CartesianMonoidal) 11 | open import Categories.Category.Monoidal using (MonoidalCategory) 12 | 13 | record CartesianCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 14 | field 15 | U : Category o ℓ e -- U for underlying 16 | cartesian : Cartesian U 17 | 18 | open Category U public 19 | open Cartesian cartesian public 20 | 21 | monoidalCategory : MonoidalCategory o ℓ e 22 | monoidalCategory = record 23 | { U = U 24 | ; monoidal = CartesianMonoidal.monoidal cartesian 25 | } 26 | -------------------------------------------------------------------------------- /src/Categories/Category/CartesianClosed/Bundle.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Bundled version of a Cartesian Closed Category 4 | module Categories.Category.CartesianClosed.Bundle where 5 | 6 | open import Level 7 | 8 | open import Categories.Category.Core using (Category) 9 | open import Categories.Category.CartesianClosed using (CartesianClosed) 10 | open import Categories.Category.Cartesian.Bundle 11 | 12 | record CartesianClosedCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 13 | field 14 | U : Category o ℓ e -- U for underlying 15 | cartesianClosed : CartesianClosed U 16 | 17 | open Category U public 18 | 19 | cartesianCategory : CartesianCategory o ℓ e 20 | cartesianCategory = record 21 | { U = U 22 | ; cartesian = CartesianClosed.cartesian cartesianClosed 23 | } 24 | -------------------------------------------------------------------------------- /src/Categories/Category/CartesianClosed/Locally/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Category.CartesianClosed.Locally 5 | 6 | module Categories.Category.CartesianClosed.Locally.Properties {o ℓ e} {C : Category o ℓ e} 7 | (LCCC : Locally C) where 8 | 9 | open import Categories.Category.CartesianClosed 10 | open import Categories.Category.Slice 11 | open import Categories.Category.Slice.Properties 12 | open import Categories.Functor 13 | open import Categories.Functor.Slice 14 | open import Categories.Functor.Slice.BaseChange 15 | 16 | private 17 | module C = Category C 18 | open C 19 | open Locally LCCC 20 | variable 21 | A B : Obj 22 | 23 | module _ (f : A ⇒ B) where 24 | open CartesianClosed (sliceCCC B) 25 | 26 | private 27 | C/A = Slice C A 28 | C/B = Slice C B 29 | C/B/f = Slice C/B (sliceobj f) 30 | 31 | fObj : SliceObj C B 32 | fObj = sliceobj f 33 | 34 | i : Slice⇒ C (sliceTerm.⊤ B) (fObj ^ fObj) 35 | i = λg (sliceProd.π₂ B) 36 | 37 | J : Functor C/A C/B/f 38 | J = slice⇒slice-slice C f 39 | 40 | I : Functor (Slice C/B (fObj ^ fObj)) C/B 41 | I = TotalSpace C/B ∘F BaseChange* C/B i (slice-pullbacks i _) 42 | 43 | K : Functor C/B/f (Slice C/B (fObj ^ fObj)) 44 | K = Base-F C/B (fObj ⇨-) 45 | 46 | -- this functor should be the right adjoint functor of (BaseChange* C pullbacks f). 47 | BaseChange⁎ : Functor C/A C/B 48 | BaseChange⁎ = I ∘F K ∘F J 49 | -------------------------------------------------------------------------------- /src/Categories/Category/Cocartesian/Bundle.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Bundled version of a Cocartesian Category 4 | module Categories.Category.Cocartesian.Bundle where 5 | 6 | open import Level 7 | 8 | open import Categories.Category.Core using (Category) 9 | open import Categories.Category.Cocartesian using (Cocartesian) 10 | 11 | record CocartesianCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 12 | field 13 | U : Category o ℓ e -- U for underlying 14 | cocartesian : Cocartesian U 15 | 16 | open Category U public 17 | open Cocartesian cocartesian public 18 | -------------------------------------------------------------------------------- /src/Categories/Category/Cocomplete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Cocomplete where 4 | 5 | open import Level 6 | 7 | open import Categories.Category using (Category) 8 | open import Categories.Functor using (Functor) 9 | open import Categories.Diagram.Colimit using (Colimit) 10 | 11 | Cocomplete : (o ℓ e : Level) {o′ ℓ′ e′ : Level} (C : Category o′ ℓ′ e′) → Set _ 12 | Cocomplete o ℓ e C = ∀ {J : Category o ℓ e} (F : Functor J C) → Colimit F 13 | -------------------------------------------------------------------------------- /src/Categories/Category/Cocomplete/Finitely.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Cocomplete.Finitely {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.Cocartesian C 10 | open import Categories.Diagram.Coequalizer C 11 | open import Categories.Diagram.Pushout C 12 | open import Categories.Diagram.Pushout.Properties C 13 | 14 | open Category C 15 | 16 | record FinitelyCocomplete : Set (levelOfTerm C) where 17 | field 18 | cocartesian : Cocartesian 19 | coequalizer : ∀ {A B} (f g : A ⇒ B) → Coequalizer f g 20 | 21 | module cocartesian = Cocartesian cocartesian 22 | module coequalizer {A B} (f g : A ⇒ B) = Coequalizer (coequalizer f g) 23 | open cocartesian public 24 | 25 | pushout : ∀ {X Y Z} (f : X ⇒ Y) (g : X ⇒ Z) → Pushout f g 26 | pushout f g = Coproduct×Coequalizer⇒Pushout coproduct (coequalizer _ _) 27 | 28 | module pushout {X Y Z} (f : X ⇒ Y) (g : X ⇒ Z) = Pushout (pushout f g) 29 | -------------------------------------------------------------------------------- /src/Categories/Category/Cocomplete/Finitely/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Category.Cocomplete.Finitely 5 | 6 | module Categories.Category.Cocomplete.Finitely.Properties {o ℓ e} {C : Category o ℓ e} (finite : FinitelyCocomplete C) where 7 | 8 | open import Categories.Category.Duality C 9 | private 10 | module C = Category C 11 | open C 12 | 13 | open import Categories.Category.Finite.Fin 14 | open import Categories.Category.Complete.Finitely C.op 15 | open import Categories.Category.Complete.Finitely.Properties (FinitelyCocomplete⇒coFinitelyComplete finite) 16 | 17 | open import Categories.Diagram.Colimit 18 | open import Categories.Diagram.Limit 19 | open import Categories.Diagram.Duality C 20 | open import Categories.Functor 21 | 22 | finiteColimit : (shape : FinCatShape) (F : Functor (FinCategory shape) C) → Colimit F 23 | finiteColimit shape F = coLimit⇒Colimit (finiteLimit shape.op F.op) 24 | where module shape = FinCatShape shape 25 | module F = Functor F 26 | -------------------------------------------------------------------------------- /src/Categories/Category/Cocomplete/Properties/Construction.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Cocomplete.Properties.Construction {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.Construction.Arrow using (Morphism) 10 | open import Categories.Category.Cocomplete using (Cocomplete) 11 | open import Categories.Diagram.Coequalizer C using (Coequalizer) 12 | open import Categories.Diagram.Colimit.Properties using (build-colim) 13 | open import Categories.Functor using (Functor) 14 | open import Categories.Object.Coproduct.Indexed C using (AllCoproductsOf) 15 | open import Categories.Object.Coproduct.Indexed.Properties C using (lowerAllCoproductsOf) 16 | 17 | private 18 | variable 19 | o′ ℓ′ e′ : Level 20 | 21 | module _ (coprods : AllCoproductsOf (o′ ⊔ ℓ′)) (coequalizer : ∀ {A B} (f g : C [ A , B ]) → Coequalizer f g) where 22 | 23 | AllCoproducts×Coequalizer⇒Cocomplete : Cocomplete o′ ℓ′ e′ C 24 | AllCoproducts×Coequalizer⇒Cocomplete F = build-colim {OP = OP} {MP = MP} (coequalizer _ _) 25 | where 26 | open Functor F 27 | OP = lowerAllCoproductsOf ℓ′ coprods F₀ 28 | MP = lowerAllCoproductsOf o′ coprods λ f → F₀ (Morphism.dom f) 29 | -------------------------------------------------------------------------------- /src/Categories/Category/Complete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Complete where 4 | 5 | open import Level 6 | 7 | open import Categories.Category 8 | open import Categories.Category.Construction.Cones 9 | open import Categories.Functor 10 | open import Categories.Diagram.Cone.Properties 11 | open import Categories.Diagram.Limit using (Limit) 12 | 13 | Complete : (o ℓ e : Level) {o′ ℓ′ e′ : Level} (C : Category o′ ℓ′ e′) → Set _ 14 | Complete o ℓ e C = ∀ {J : Category o ℓ e} (F : Functor J C) → Limit F 15 | 16 | -- a functor between diagrams corresponds to a morphism between limits 17 | module _ {o ℓ e o′ ℓ′ e′} {C : Category o′ ℓ′ e′} (Com : Complete o ℓ e C) 18 | {J J′ : Category o ℓ e} (F : Functor J′ J) (G : Functor J C) where 19 | 20 | private 21 | module C = Category C 22 | module J = Category J 23 | module J′ = Category J′ 24 | module F = Functor F 25 | module G = Functor G 26 | 27 | F⇒arr : Cones (G ∘F F) [ F-map-Coneʳ F (Limit.limit (Com G)) , Limit.limit (Com (G ∘F F)) ] 28 | F⇒arr = Limit.rep-cone (Com (G ∘F F)) (F-map-Coneʳ F (Limit.limit (Com G))) 29 | -------------------------------------------------------------------------------- /src/Categories/Category/Complete/Finitely.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Complete.Finitely {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.BinaryProducts C using (BinaryProducts) 10 | open import Categories.Category.Cartesian C using (Cartesian) 11 | open import Categories.Diagram.Equalizer C using (Equalizer) 12 | open import Categories.Diagram.Pullback C using (Pullback; Product×Equalizer⇒Pullback) 13 | 14 | open Category C 15 | 16 | record FinitelyComplete : Set (levelOfTerm C) where 17 | field 18 | cartesian : Cartesian 19 | equalizer : ∀ {A B} (f g : A ⇒ B) → Equalizer f g 20 | 21 | module equalizer {A B} (f g : A ⇒ B) = Equalizer (equalizer f g) 22 | open Cartesian cartesian using (products) 23 | 24 | pullback : ∀ {X Y Z} (f : X ⇒ Z) (g : Y ⇒ Z) → Pullback f g 25 | pullback f g = Product×Equalizer⇒Pullback (BinaryProducts.product products) (equalizer _ _) 26 | 27 | module pullback {X Y Z} (f : X ⇒ Z) (g : Y ⇒ Z) = Pullback (pullback f g) 28 | -------------------------------------------------------------------------------- /src/Categories/Category/Complete/Properties/Construction.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Complete.Properties.Construction {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.Construction.Arrow using (Morphism) 10 | open import Categories.Category.Complete using (Complete) 11 | open import Categories.Diagram.Equalizer C using (Equalizer) 12 | open import Categories.Diagram.Limit.Properties using (build-lim) 13 | open import Categories.Functor using (Functor) 14 | open import Categories.Object.Product.Indexed C using (AllProductsOf) 15 | open import Categories.Object.Product.Indexed.Properties C using (lowerAllProductsOf) 16 | 17 | private 18 | variable 19 | o′ ℓ′ e′ : Level 20 | 21 | module _ (prods : AllProductsOf (o′ ⊔ ℓ′)) (equalizer : ∀ {A B} (f g : C [ A , B ]) → Equalizer f g) where 22 | 23 | AllProducts×Equalizer⇒Complete : Complete o′ ℓ′ e′ C 24 | AllProducts×Equalizer⇒Complete F = build-lim {OP = OP} {MP = MP} (equalizer _ _) 25 | where 26 | open Functor F 27 | OP = lowerAllProductsOf ℓ′ prods F₀ 28 | MP = lowerAllProductsOf o′ prods λ f → F₀ (Morphism.cod f) 29 | -------------------------------------------------------------------------------- /src/Categories/Category/Concrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Concrete where 4 | 5 | open import Level 6 | 7 | open import Categories.Category.Core using (Category) 8 | open import Categories.Category.Instance.Setoids using (Setoids) 9 | open import Categories.Functor.Core using (Functor) 10 | open import Categories.Functor.Representable using (Representable) 11 | open import Categories.Functor.Properties using (Faithful) 12 | 13 | -- A Concrete Category is a category along with a faithful 14 | -- functor to Setoid. 15 | -- [Normally Set, but that doesn't work so well here] 16 | 17 | private 18 | variable 19 | o ℓ e : Level 20 | 21 | record Concrete (C : Category o ℓ e) (ℓ′ e′ : Level) : Set (o ⊔ ℓ ⊔ e ⊔ suc (ℓ′ ⊔ e′)) where 22 | field 23 | concretize : Functor C (Setoids ℓ′ e′) 24 | conc-faithful : Faithful concretize 25 | 26 | -- Because of the use of the Hom functor, some levels collapse 27 | record RepresentablyConcrete (C : Category o ℓ e) : Set (o ⊔ suc (e ⊔ ℓ)) where 28 | open Concrete 29 | field 30 | conc : Concrete C ℓ e 31 | representable : Representable (concretize conc) 32 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/0-Groupoid.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Relation.Binary using (Setoid) 4 | 5 | -- Lifts a setoid into a 1-groupoid 6 | 7 | module Categories.Category.Construction.0-Groupoid 8 | {c ℓ} e (S : Setoid c ℓ) where 9 | 10 | open import Level using (Lift) 11 | open import Data.Unit using (⊤) 12 | open import Function using (flip) 13 | 14 | open import Categories.Category.Groupoid using (Groupoid) 15 | 16 | open Setoid S 17 | 18 | 0-Groupoid : Groupoid c ℓ e 19 | 0-Groupoid = record 20 | { category = record 21 | { Obj = Carrier 22 | ; _⇒_ = _≈_ 23 | ; _≈_ = λ _ _ → Lift e ⊤ 24 | ; id = refl 25 | ; _∘_ = flip trans 26 | } 27 | ; isGroupoid = record { _⁻¹ = sym } 28 | } 29 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Cowedges.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core using (Category) 4 | open import Categories.Functor.Bifunctor using (Bifunctor) 5 | 6 | module Categories.Category.Construction.Cowedges {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} 7 | (F : Bifunctor (Category.op C) C D) where 8 | 9 | open import Level 10 | 11 | open import Categories.Category.Core using (Category) 12 | open import Categories.Diagram.Cowedge F 13 | 14 | Cowedges : Category (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) e′ 15 | Cowedges = record 16 | { Obj = Cowedge 17 | ; _⇒_ = Cowedge-Morphism 18 | ; _≈_ = λ M N → u M ≈ u N 19 | ; id = Cowedge-id 20 | ; _∘_ = Cowedge-Morphism-∘ 21 | ; assoc = assoc 22 | ; sym-assoc = sym-assoc 23 | ; identityˡ = identityˡ 24 | ; identityʳ = identityʳ 25 | ; identity² = identity² 26 | ; equiv = record { refl = Equiv.refl ; sym = Equiv.sym ; trans = Equiv.trans } 27 | ; ∘-resp-≈ = ∘-resp-≈ 28 | } 29 | where 30 | open Cowedge-Morphism 31 | open Category D 32 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/DaggerFunctors.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --without-K #-} 2 | 3 | module Categories.Category.Construction.DaggerFunctors where 4 | 5 | open import Categories.Category.Core using (Category) 6 | open import Categories.Category.Construction.Functors using (Functors) 7 | open import Categories.Category.SubCategory using (FullSubCategory) 8 | open import Categories.Category.Dagger using (DaggerCategory) 9 | open import Categories.Functor.Dagger using (DaggerFunctor) 10 | 11 | open import Level using (Level; _⊔_) 12 | 13 | private 14 | variable 15 | o ℓ e o′ ℓ′ e′ : Level 16 | 17 | DaggerFunctors : DaggerCategory o ℓ e → DaggerCategory o′ ℓ′ e′ → Category (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) (o ⊔ ℓ ⊔ ℓ′ ⊔ e′) (o ⊔ e′) 18 | DaggerFunctors C D = FullSubCategory (Functors (DaggerCategory.C C) (DaggerCategory.C D)) {I = DaggerFunctor C D} DaggerFunctor.functor 19 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Fin.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Data.Nat.Base using (ℕ) 4 | 5 | module Categories.Category.Construction.Fin (n : ℕ) where 6 | 7 | open import Level 8 | open import Data.Fin.Properties 9 | 10 | open import Categories.Category 11 | open import Categories.Category.Construction.Thin 0ℓ (≤-poset n) 12 | 13 | Fin : Category 0ℓ 0ℓ 0ℓ 14 | Fin = Thin 15 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/GroupAsCategory.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --without-K #-} 2 | open import Algebra.Bundles using (Group) 3 | 4 | module Categories.Category.Construction.GroupAsCategory o {c ℓ} (G : Group c ℓ) where 5 | 6 | open import Level 7 | 8 | open import Categories.Category.Groupoid 9 | 10 | open Group G 11 | 12 | -- A group is a groupoid with a single object 13 | open import Categories.Category.Construction.MonoidAsCategory o monoid 14 | renaming (MonoidAsCategory to GroupAsCategory) public 15 | 16 | GroupIsGroupoid : IsGroupoid GroupAsCategory 17 | GroupIsGroupoid = record 18 | { iso = record 19 | { isoˡ = inverseˡ _ 20 | ; isoʳ = inverseʳ _ 21 | } 22 | } 23 | 24 | GroupAsGroupoid : Groupoid o c ℓ 25 | GroupAsGroupoid = record { isGroupoid = GroupIsGroupoid } 26 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Groups.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --without-K #-} 2 | 3 | -- The category of group objects with a cartesian category 4 | 5 | open import Categories.Category 6 | open import Categories.Category.Cartesian 7 | 8 | module Categories.Category.Construction.Groups {o ℓ e} {𝒞 : Category o ℓ e} (C : Cartesian 𝒞) where 9 | 10 | open import Level 11 | 12 | open import Categories.Category.BinaryProducts 𝒞 13 | open import Categories.Morphism.Reasoning 𝒞 14 | open import Categories.Object.Group C 15 | open import Categories.Object.Product.Morphisms 𝒞 16 | 17 | open Category 𝒞 18 | open Cartesian C 19 | open BinaryProducts products 20 | open Equiv 21 | open HomReasoning 22 | open Group using (μ) 23 | open Group⇒ 24 | 25 | Groups : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e 26 | Groups = record 27 | { Obj = Group 28 | ; _⇒_ = Group⇒ 29 | ; _≈_ = λ f g → arr f ≈ arr g 30 | ; id = record 31 | { arr = id 32 | ; preserves-μ = identityˡ ○ introʳ (id×id product) 33 | ; preserves-η = identityˡ 34 | ; preserves-ι = id-comm-sym 35 | } 36 | ; _∘_ = λ f g → record 37 | { arr = arr f ∘ arr g 38 | ; preserves-μ = glue (preserves-μ f) (preserves-μ g) ○ ∘-resp-≈ʳ ⁂∘⁂ 39 | ; preserves-η = glueTrianglesˡ (preserves-η f) (preserves-η g) 40 | ; preserves-ι = glue (preserves-ι f) (preserves-ι g) 41 | } 42 | ; assoc = assoc 43 | ; sym-assoc = sym-assoc 44 | ; identityˡ = identityˡ 45 | ; identityʳ = identityʳ 46 | ; identity² = identity² 47 | ; equiv = record 48 | { refl = refl 49 | ; sym = sym 50 | ; trans = trans 51 | } 52 | ; ∘-resp-≈ = ∘-resp-≈ 53 | } 54 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/KaroubiEnvelope.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category; _[_≈_]) 4 | 5 | -- Karoubi Envelopes. These are the free completions of categories under split idempotents 6 | module Categories.Category.Construction.KaroubiEnvelope {o ℓ e} (𝒞 : Category o ℓ e) where 7 | 8 | open import Level 9 | 10 | open import Categories.Morphism.Idempotent.Bundles 𝒞 11 | open import Categories.Morphism.Reasoning 𝒞 12 | 13 | private 14 | module 𝒞 = Category 𝒞 15 | open 𝒞.HomReasoning 16 | open 𝒞.Equiv 17 | 18 | open Idempotent 19 | open Idempotent⇒ 20 | 21 | KaroubiEnvelope : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e 22 | KaroubiEnvelope = record 23 | { Obj = Idempotent 24 | ; _⇒_ = Idempotent⇒ 25 | ; _≈_ = λ f g → 𝒞 [ Idempotent⇒.hom f ≈ Idempotent⇒.hom g ] 26 | ; id = id 27 | ; _∘_ = _∘_ 28 | ; assoc = 𝒞.assoc 29 | ; sym-assoc = 𝒞.sym-assoc 30 | ; identityˡ = λ {I} {J} {f} → absorbˡ f 31 | ; identityʳ = λ {I} {J} {f} → absorbʳ f 32 | ; identity² = λ {I} → idempotent I 33 | ; equiv = record 34 | { refl = refl 35 | ; sym = sym 36 | ; trans = trans 37 | } 38 | ; ∘-resp-≈ = 𝒞.∘-resp-≈ 39 | } 40 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/MonoidAsCategory.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --without-K #-} 2 | open import Algebra.Bundles using (Monoid) 3 | 4 | module Categories.Category.Construction.MonoidAsCategory o {c ℓ} (M : Monoid c ℓ) where 5 | 6 | open import Data.Unit.Polymorphic 7 | open import Level 8 | 9 | open import Categories.Category.Core 10 | 11 | open Monoid M 12 | 13 | -- A monoid is a category with one object 14 | MonoidAsCategory : Category o c ℓ 15 | MonoidAsCategory = record 16 | { Obj = ⊤ 17 | ; assoc = assoc _ _ _ 18 | ; sym-assoc = sym (assoc _ _ _) 19 | ; identityˡ = identityˡ _ 20 | ; identityʳ = identityʳ _ 21 | ; identity² = identityˡ _ 22 | ; equiv = isEquivalence 23 | ; ∘-resp-≈ = ∙-cong 24 | } 25 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Monoids.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category.Core 3 | open import Categories.Category.Monoidal.Core 4 | 5 | -- This module defines the category of monoids internal to a given monoidal 6 | -- category. 7 | 8 | module Categories.Category.Construction.Monoids {o ℓ e} {𝒞 : Category o ℓ e} (C : Monoidal 𝒞) where 9 | 10 | open import Level 11 | 12 | open import Categories.Functor using (Functor) 13 | open import Categories.Morphism.Reasoning 𝒞 14 | open import Categories.Object.Monoid C 15 | 16 | open Category 𝒞 17 | open Monoidal C 18 | open HomReasoning 19 | open Monoid using (η; μ) 20 | open Monoid⇒ 21 | 22 | Monoids : Category (o ⊔ ℓ ⊔ e) (ℓ ⊔ e) e 23 | Monoids = record 24 | { Obj = Monoid 25 | ; _⇒_ = Monoid⇒ 26 | ; _≈_ = λ f g → arr f ≈ arr g 27 | ; id = record 28 | { arr = id 29 | ; preserves-μ = identityˡ ○ introʳ (Functor.identity ⊗) 30 | ; preserves-η = identityˡ 31 | } 32 | ; _∘_ = λ f g → record 33 | { arr = arr f ∘ arr g 34 | ; preserves-μ = glue (preserves-μ f) (preserves-μ g) ○ ∘-resp-≈ʳ (⟺ (Functor.homomorphism ⊗)) 35 | ; preserves-η = glueTrianglesˡ (preserves-η f) (preserves-η g) 36 | } 37 | ; assoc = assoc 38 | ; sym-assoc = sym-assoc 39 | ; identityˡ = identityˡ 40 | ; identityʳ = identityʳ 41 | ; identity² = identity² 42 | -- We cannot define equiv = equiv here, because _⇒_ of this category is a 43 | -- different level to the _⇒_ of 𝒞. 44 | ; equiv = record 45 | { refl = refl 46 | ; sym = sym 47 | ; trans = trans 48 | } 49 | ; ∘-resp-≈ = ∘-resp-≈ 50 | } where open Equiv 51 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/ObjectRestriction.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Given a predicate on the objects of a Category C, build another category 4 | -- with just those that satisfy a pr 5 | 6 | module Categories.Category.Construction.ObjectRestriction where 7 | 8 | open import Level 9 | open import Data.Product using (Σ; proj₁) 10 | open import Relation.Unary using (Pred) 11 | open import Function using (_on_) renaming (id to id→) 12 | 13 | open import Categories.Category.Core 14 | 15 | private 16 | variable 17 | o ℓ e ℓ′ : Level 18 | 19 | ObjectRestriction : (C : Category o ℓ e) → Pred (Category.Obj C) ℓ′ → Category (o ⊔ ℓ′) ℓ e 20 | ObjectRestriction C R = record 21 | { Obj = Σ Obj R 22 | ; _⇒_ = _⇒_ on proj₁ 23 | ; _≈_ = _≈_ 24 | ; id = id 25 | ; _∘_ = _∘_ 26 | ; assoc = assoc 27 | ; sym-assoc = sym-assoc 28 | ; identityˡ = identityˡ 29 | ; identityʳ = identityʳ 30 | ; identity² = identity² 31 | ; equiv = equiv 32 | ; ∘-resp-≈ = ∘-resp-≈ 33 | } 34 | where open Category C 35 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Path.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Level 3 | 4 | open import Categories.Category 5 | 6 | module Categories.Category.Construction.Path {o ℓ e : Level} (C : Category o ℓ e) where 7 | 8 | open import Function using (flip) 9 | open import Relation.Binary hiding (_⇒_) 10 | open import Relation.Binary.Construct.Closure.Transitive 11 | 12 | open Category C 13 | 14 | -- Defining the Path Category 15 | ∘-tc : {A B : Obj} → A [ _⇒_ ]⁺ B → A ⇒ B 16 | ∘-tc [ f ] = f 17 | ∘-tc (_ ∼⁺⟨ f⁺ ⟩ f⁺′) = ∘-tc f⁺′ ∘ ∘-tc f⁺ 18 | 19 | infix 4 _≈⁺_ 20 | _≈⁺_ : {A B : Obj} → (i j : A [ _⇒_ ]⁺ B) → Set e 21 | f⁺ ≈⁺ g⁺ = ∘-tc f⁺ ≈ ∘-tc g⁺ 22 | 23 | Path : Category o (o ⊔ ℓ) e 24 | Path = record 25 | { Obj = Obj 26 | ; _⇒_ = λ A B → A [ _⇒_ ]⁺ B 27 | ; _≈_ = _≈⁺_ 28 | ; id = [ id ] 29 | ; _∘_ = flip (_ ∼⁺⟨_⟩_) 30 | ; assoc = assoc 31 | ; sym-assoc = sym-assoc 32 | ; identityˡ = identityˡ 33 | ; identityʳ = identityʳ 34 | ; identity² = identity² 35 | ; equiv = record 36 | { refl = Equiv.refl 37 | ; sym = Equiv.sym 38 | ; trans = Equiv.trans 39 | } 40 | ; ∘-resp-≈ = ∘-resp-≈ 41 | } 42 | where open HomReasoning 43 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/PathCategory.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level 4 | open import Data.Quiver using (Quiver) 5 | 6 | -- The Category of (free) paths over a Quiver 7 | module Categories.Category.Construction.PathCategory {o ℓ e} (G : Quiver o ℓ e) where 8 | 9 | open import Function.Base using (_$_) 10 | open import Relation.Binary.Construct.Closure.ReflexiveTransitive 11 | open import Relation.Binary.Construct.Closure.ReflexiveTransitive.Properties hiding (trans) 12 | open import Data.Quiver.Paths 13 | 14 | open import Categories.Category.Core using (Category) 15 | 16 | open Quiver G 17 | private module P = Paths G 18 | open P 19 | 20 | PathCategory : Category o (o ⊔ ℓ) (o ⊔ ℓ ⊔ e) 21 | PathCategory = record 22 | { Obj = Obj 23 | ; _⇒_ = Star _⇒_ 24 | ; _≈_ = _≈*_ 25 | ; id = ε 26 | ; _∘_ = _▻▻_ 27 | ; assoc = λ {_ _ _ _} {f g h} → sym $ ≡⇒≈* $ ◅◅-assoc f g h 28 | ; sym-assoc = λ {_ _ _ _} {f g h} → ≡⇒≈* $ ◅◅-assoc f g h 29 | ; identityˡ = λ {_ _ f} → ◅◅-identityʳ f 30 | ; identityʳ = refl 31 | ; identity² = refl 32 | ; equiv = isEquivalence 33 | ; ∘-resp-≈ = resp 34 | } 35 | where 36 | resp : ∀ {A B C} {f h : Star _⇒_ B C} {g i : Star _⇒_ A B} → 37 | f ≈* h → g ≈* i → (f ▻▻ g) ≈* (h ▻▻ i) 38 | resp eq ε = eq 39 | resp eq (eq₁ ◅ eq₂) = eq₁ ◅ (resp eq eq₂) 40 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Presheaves.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Construction.Presheaves where 3 | 4 | -- The Category of Presheaves over a Category C, i.e. 5 | -- the Functor Category [ C.op , Setoids ] 6 | -- Again, the levels are made explicit to show the generality and constraints. 7 | 8 | -- CoPreasheaves are defined here as well, for convenience 9 | -- The main reson to name them is that some things (like CoYoneda) 10 | -- look more natural/symmetric then. 11 | 12 | open import Level 13 | 14 | open import Categories.Category 15 | open import Categories.Category.Construction.Functors 16 | open import Categories.Category.Instance.Setoids using (Setoids) 17 | 18 | Presheaves′ : ∀ o′ ℓ′ {o ℓ e : Level} → Category o ℓ e → 19 | Category (o ⊔ ℓ ⊔ e ⊔ suc (o′ ⊔ ℓ′)) (o ⊔ ℓ ⊔ o′ ⊔ ℓ′) (o ⊔ o′ ⊔ ℓ′) 20 | Presheaves′ o′ ℓ′ C = Functors (Category.op C) (Setoids o′ ℓ′) 21 | 22 | Presheaves : ∀ {o ℓ e o′ ℓ′ : Level} → Category o ℓ e → 23 | Category (o ⊔ ℓ ⊔ e ⊔ suc (o′ ⊔ ℓ′)) (o ⊔ ℓ ⊔ o′ ⊔ ℓ′) (o ⊔ o′ ⊔ ℓ′) 24 | Presheaves {o} {ℓ} {e} {o′} {ℓ′} C = Presheaves′ o′ ℓ′ C 25 | 26 | CoPresheaves′ : ∀ o′ ℓ′ {o ℓ e : Level} → Category o ℓ e → 27 | Category (o ⊔ ℓ ⊔ e ⊔ suc (o′ ⊔ ℓ′)) (o ⊔ ℓ ⊔ o′ ⊔ ℓ′) (o ⊔ o′ ⊔ ℓ′) 28 | CoPresheaves′ o′ ℓ′ C = Functors C (Setoids o′ ℓ′) 29 | 30 | CoPresheaves : ∀ {o ℓ e o′ ℓ′ : Level} → Category o ℓ e → 31 | Category (o ⊔ ℓ ⊔ e ⊔ suc (o′ ⊔ ℓ′)) (o ⊔ ℓ ⊔ o′ ⊔ ℓ′) (o ⊔ o′ ⊔ ℓ′) 32 | CoPresheaves {o} {ℓ} {e} {o′} {ℓ′} C = CoPresheaves′ o′ ℓ′ C 33 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Properties/Functors.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Construction.Properties.Functors where 4 | 5 | open import Categories.Category.Complete.Properties 6 | using ( Functors-Complete 7 | ; evalF-Continuous) 8 | public 9 | open import Categories.Category.Cocomplete.Properties 10 | using (Functors-Cocomplete) 11 | public 12 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Properties/Presheaves.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Construction.Properties.Presheaves where 4 | 5 | open import Categories.Category.Construction.Properties.Presheaves.Cartesian 6 | using (module IsCartesian) 7 | public 8 | 9 | open import Categories.Category.Construction.Properties.Presheaves.CartesianClosed 10 | using (module IsCartesianClosed) 11 | public 12 | 13 | open import Categories.Category.Construction.Properties.Presheaves.FromCartesianCCC 14 | using (module FromCartesianCCC) 15 | public 16 | 17 | open import Categories.Category.Construction.Properties.Presheaves.Complete 18 | using (Presheaves-Complete; Presheaves-Cocomplete) 19 | public 20 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Properties/Presheaves/Complete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Construction.Properties.Presheaves.Complete {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Categories.Category.Complete 8 | open import Categories.Category.Complete.Finitely 9 | open import Categories.Category.Complete.Properties 10 | open import Categories.Category.Cocomplete 11 | open import Categories.Category.Cocomplete.Finitely 12 | open import Categories.Category.Cocomplete.Properties 13 | open import Categories.Category.Construction.Presheaves 14 | open import Categories.Category.Instance.Setoids 15 | open import Categories.Category.Instance.Properties.Setoids 16 | 17 | private 18 | module C = Category C 19 | 20 | Presheaves-Complete : ∀ o′ ℓ′ e′ ℓ″ e″ → Complete o′ ℓ′ e′ (Presheaves C) 21 | Presheaves-Complete o′ ℓ′ e′ ℓ″ e″ = Functors-Complete C.op (Setoids-Complete o′ ℓ′ e′ ℓ″ e″) 22 | 23 | Presheaves-FinitelyComplete : ∀ o′ ℓ′ e′ ℓ″ e″ → FinitelyComplete (Presheaves C) 24 | Presheaves-FinitelyComplete o′ ℓ′ e′ ℓ″ e″ = Complete⇒FinitelyComplete (Presheaves C) (Presheaves-Complete o′ ℓ′ e′ ℓ″ e″) 25 | 26 | Presheaves-Cocomplete : ∀ o′ ℓ′ e′ ℓ″ e″ → Cocomplete o′ ℓ′ e′ (Presheaves C) 27 | Presheaves-Cocomplete o′ ℓ′ e′ ℓ″ e″ = Functors-Cocomplete C.op (Setoids-Cocomplete o′ ℓ′ e′ ℓ″ e″) 28 | 29 | Presheaves-FinitelyCocomplete : ∀ o′ ℓ′ e′ ℓ″ e″ → FinitelyCocomplete (Presheaves C) 30 | Presheaves-FinitelyCocomplete o′ ℓ′ e′ ℓ″ e″ = Cocomplete⇒FinitelyCocomplete (Presheaves C) (Presheaves-Cocomplete o′ ℓ′ e′ ℓ″ e″) 31 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/SetoidDiscrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Construction.SetoidDiscrete where 3 | 4 | open import Data.Unit using (⊤; tt) 5 | open import Function using (flip) 6 | open import Level using (Lift; lift) 7 | open import Relation.Binary using (Setoid; IsEquivalence) 8 | 9 | open import Categories.Category.Core using (Category) 10 | open import Categories.Category.Discrete using (DiscreteCategory) 11 | open import Categories.Morphism using (_≅_; ≅-isEquivalence) 12 | {- 13 | This is a better version of Discrete, which is more in line with 14 | the rest of this library, and makes a Category out of a Setoid. 15 | It is still 'wrong' as the equivalence is completely uninformative. 16 | -} 17 | 18 | Discrete : ∀ {o ℓ e} (A : Setoid o ℓ) → DiscreteCategory o ℓ e 19 | Discrete {o} {ℓ} {e} A = record 20 | { category = C 21 | ; isDiscrete = record 22 | { isGroupoid = record { _⁻¹ = sym } 23 | } 24 | } 25 | where 26 | open Setoid A 27 | C : Category o ℓ e 28 | C = record 29 | { Obj = Carrier 30 | ; _⇒_ = _≈_ 31 | ; _≈_ = λ _ _ → Lift e ⊤ 32 | ; id = Setoid.refl A 33 | ; _∘_ = flip trans 34 | } 35 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Spans.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Construction.Spans {o ℓ e} (𝒞 : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.Diagram.Span 𝒞 10 | open import Categories.Morphism.Reasoning 𝒞 11 | 12 | open Category 𝒞 13 | open HomReasoning 14 | open Equiv 15 | 16 | open Span 17 | open Span⇒ 18 | 19 | Spans : Obj → Obj → Category (o ⊔ ℓ) (o ⊔ ℓ ⊔ e) e 20 | Spans X Y = record 21 | { Obj = Span X Y 22 | ; _⇒_ = Span⇒ 23 | ; _≈_ = λ f g → arr f ≈ arr g 24 | ; id = id-span⇒ 25 | ; _∘_ = _∘ₛ_ 26 | ; assoc = assoc 27 | ; sym-assoc = sym-assoc 28 | ; identityˡ = identityˡ 29 | ; identityʳ = identityʳ 30 | ; identity² = identity² 31 | ; equiv = record 32 | { refl = refl 33 | ; sym = sym 34 | ; trans = trans 35 | } 36 | ; ∘-resp-≈ = ∘-resp-≈ 37 | } 38 | 39 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/StrictDiscrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Construction.StrictDiscrete where 3 | 4 | -- This is not 'the' Discrete Category construction, but one of them. 5 | -- In this case, what is built is a Category, but it's actually Strict (thus 6 | -- the name). 7 | open import Level 8 | open import Data.Unit 9 | open import Function 10 | open import Relation.Binary.PropositionalEquality as ≡ 11 | 12 | open import Categories.Category 13 | open import Categories.Functor 14 | 15 | Discrete : ∀ {a} (A : Set a) → Category a a a 16 | Discrete A = record 17 | { Obj = A 18 | ; _⇒_ = _≡_ 19 | ; _≈_ = _≡_ 20 | ; id = refl 21 | ; _∘_ = flip ≡.trans 22 | ; assoc = λ {_ _ _ _ g} → sym (trans-assoc g) 23 | ; sym-assoc = λ {_ _ _ _ g} → trans-assoc g 24 | ; identityˡ = λ {_ _ f} → trans-reflʳ f 25 | ; identityʳ = refl 26 | ; identity² = refl 27 | ; equiv = isEquivalence 28 | ; ∘-resp-≈ = λ where 29 | refl refl → refl 30 | } 31 | 32 | module _ {a o ℓ e} {A : Set a} (C : Category o ℓ e) where 33 | open Category C renaming (id to one) 34 | 35 | module _ (f : A → Obj) where 36 | 37 | lift-func : Functor (Discrete A) C 38 | lift-func = record 39 | { F₀ = f 40 | ; F₁ = λ { refl → one } 41 | ; identity = Equiv.refl 42 | ; homomorphism = λ { {_} {_} {_} {refl} {refl} → Equiv.sym identity² } 43 | ; F-resp-≈ = λ { {_} {_} {refl} refl → Equiv.refl } 44 | } 45 | -------------------------------------------------------------------------------- /src/Categories/Category/Construction/Wedges.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core using (Category) 4 | open import Categories.Functor.Bifunctor using (Bifunctor) 5 | 6 | module Categories.Category.Construction.Wedges {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} 7 | (F : Bifunctor (Category.op C) C D) where 8 | 9 | open import Level 10 | 11 | open import Categories.Category.Core using (Category) 12 | open import Categories.Diagram.Wedge F 13 | 14 | Wedges : Category (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) e′ 15 | Wedges = record 16 | { Obj = Wedge 17 | ; _⇒_ = Wedge-Morphism 18 | ; _≈_ = λ M N → u M ≈ u N 19 | ; id = Wedge-id 20 | ; _∘_ = Wedge-Morphism-∘ 21 | ; assoc = assoc 22 | ; sym-assoc = sym-assoc 23 | ; identityˡ = identityˡ 24 | ; identityʳ = identityʳ 25 | ; identity² = identity² 26 | ; equiv = record { refl = Equiv.refl ; sym = Equiv.sym ; trans = Equiv.trans } 27 | ; ∘-resp-≈ = ∘-resp-≈ 28 | } 29 | where 30 | open Wedge-Morphism 31 | open Category D 32 | -------------------------------------------------------------------------------- /src/Categories/Category/Dagger.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Dagger where 3 | 4 | open import Level using (_⊔_; suc) 5 | open import Relation.Unary using (Pred) 6 | 7 | open import Categories.Category.Core using (Category) 8 | open import Categories.Functor.Core using (Functor) 9 | open import Categories.Morphism using (Iso) 10 | 11 | record HasDagger {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 12 | open Category C 13 | infix 10 _† 14 | 15 | field 16 | _† : ∀ {A B} → A ⇒ B → B ⇒ A 17 | †-identity : ∀ {A} → id {A} † ≈ id 18 | †-homomorphism : ∀ {A B C} {f : A ⇒ B} {g : B ⇒ C} → (g ∘ f) † ≈ f † ∘ g † 19 | †-resp-≈ : ∀ {A B} {f g : A ⇒ B} → f ≈ g → f † ≈ g † 20 | †-involutive : ∀ {A B} (f : A ⇒ B) → f † † ≈ f 21 | 22 | †-Functor : Functor op C 23 | †-Functor = record 24 | { F₀ = λ A → A 25 | ; F₁ = _† 26 | ; identity = †-identity 27 | ; homomorphism = †-homomorphism 28 | ; F-resp-≈ = †-resp-≈ 29 | } 30 | 31 | isUnitary : ∀ {A B} → Pred (A ⇒ B) e 32 | isUnitary f = Iso C f (f †) 33 | 34 | isSelfAdjoint : ∀ {A} → Pred (A ⇒ A) e 35 | isSelfAdjoint f = f † ≈ f 36 | 37 | record DaggerCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 38 | field 39 | C : Category o ℓ e 40 | hasDagger : HasDagger C 41 | 42 | open Category C public 43 | open HasDagger hasDagger public 44 | -------------------------------------------------------------------------------- /src/Categories/Category/Dagger/Construction/Discrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Dagger.Construction.Discrete where 3 | 4 | open import Relation.Binary.PropositionalEquality 5 | 6 | open import Categories.Category.Dagger 7 | open import Categories.Category.Construction.StrictDiscrete 8 | 9 | Discrete-HasDagger : ∀ {a} (A : Set a) → HasDagger (Discrete A) 10 | Discrete-HasDagger A = record 11 | { _† = sym 12 | ; †-identity = refl 13 | ; †-homomorphism = λ { {f = refl} {g = refl} → refl } 14 | ; †-resp-≈ = cong sym 15 | ; †-involutive = λ { refl → refl} 16 | } 17 | 18 | Discrete-DaggerCategory : ∀ {a} (A : Set a) → DaggerCategory a a a 19 | Discrete-DaggerCategory A = record { hasDagger = Discrete-HasDagger A } 20 | -------------------------------------------------------------------------------- /src/Categories/Category/Dagger/Instance/Rels.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Dagger.Instance.Rels where 3 | 4 | open import Data.Product 5 | open import Function 6 | open import Relation.Binary.PropositionalEquality 7 | open import Level 8 | 9 | open import Categories.Category.Dagger 10 | open import Categories.Category.Instance.Rels 11 | 12 | RelsHasDagger : ∀ {o ℓ} → HasDagger (Rels o ℓ) 13 | RelsHasDagger = record 14 | { _† = flip 15 | ; †-identity = (lift ∘ sym ∘ lower) , (lift ∘ sym ∘ lower) 16 | ; †-homomorphism = (map₂ swap) , (map₂ swap) 17 | ; †-resp-≈ = λ p → (proj₁ p) , (proj₂ p) -- it's the implicits that need flipped 18 | ; †-involutive = λ _ → id , id 19 | } 20 | 21 | RelsDagger : ∀ o ℓ → DaggerCategory (suc o) (suc (o ⊔ ℓ)) (o ⊔ ℓ) 22 | RelsDagger o ℓ = record 23 | { C = Rels o ℓ 24 | ; hasDagger = RelsHasDagger 25 | } 26 | -------------------------------------------------------------------------------- /src/Categories/Category/Discrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Discrete where 3 | 4 | -- Discrete Category. 5 | -- https://ncatlab.org/nlab/show/discrete+category 6 | -- says: 7 | -- A category is discrete if it is both a groupoid and a preorder. That is, 8 | -- every morphism should be invertible, any two parallel morphisms should be equal. 9 | -- The idea is that in a discrete category, no two distinct (nonisomorphic) objects 10 | -- are connectable by any path (morphism), and an object connects to itself only through 11 | -- its identity morphism. 12 | 13 | open import Level using (Level; suc; _⊔_) 14 | 15 | open import Categories.Category using (Category) 16 | open import Categories.Category.Groupoid using (IsGroupoid) 17 | 18 | record IsDiscrete {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 19 | open Category C using (Obj; _⇒_; _≈_) 20 | field 21 | isGroupoid : IsGroupoid C 22 | preorder : {A B : Obj} → (f g : A ⇒ B) → f ≈ g 23 | 24 | record DiscreteCategory (o ℓ e : Level) : Set (suc (o ⊔ ℓ ⊔ e)) where 25 | field 26 | category : Category o ℓ e 27 | isDiscrete : IsDiscrete category 28 | 29 | open IsDiscrete isDiscrete public 30 | -------------------------------------------------------------------------------- /src/Categories/Category/Exact.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Exact category (https://ncatlab.org/nlab/show/exact+category) 4 | -- is a regular category 5 | -- in which every internal equivalence is a kernel pair 6 | 7 | module Categories.Category.Exact where 8 | 9 | open import Level 10 | 11 | open import Categories.Category.Core 12 | open import Categories.Diagram.Pullback 13 | open import Categories.Category.Cocartesian 14 | open import Categories.Object.Coproduct 15 | open import Categories.Morphism 16 | open import Categories.Diagram.Coequalizer 17 | open import Categories.Diagram.KernelPair 18 | 19 | open import Categories.Category.Regular 20 | open import Categories.Morphism.Regular 21 | 22 | open import Categories.Object.InternalRelation 23 | 24 | record Exact {o ℓ e : Level} (𝒞 : Category o ℓ e) : Set (suc (o ⊔ ℓ ⊔ e)) where 25 | open Category 𝒞 26 | open Pullback 27 | open Coequalizer 28 | open Equivalence 29 | 30 | field 31 | regular : Regular 𝒞 32 | quotient : ∀ {X : Obj} (E : Equivalence 𝒞 X) → Coequalizer 𝒞 (R.p₁ E) (R.p₂ E) 33 | effective : ∀ {X : Obj} (E : Equivalence 𝒞 X) → IsPullback 𝒞 (R.p₁ E) (R.p₂ E) 34 | (arr (quotient E)) (arr (quotient E)) 35 | -------------------------------------------------------------------------------- /src/Categories/Category/Extensive/Bundle.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Bundled version of an extensive (distributive) Category 4 | module Categories.Category.Extensive.Bundle where 5 | 6 | open import Level 7 | 8 | open import Categories.Category.Core using (Category) 9 | open import Categories.Category.Extensive using (Extensive) 10 | open import Categories.Category.Distributive using (Distributive) 11 | open import Categories.Category.Cartesian using (Cartesian) 12 | open import Categories.Category.Extensive.Properties.Distributive using (Extensive×Cartesian⇒Distributive) 13 | 14 | record ExtensiveCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 15 | field 16 | U : Category o ℓ e -- U for underlying 17 | extensive : Extensive U 18 | 19 | open Category U public 20 | open Extensive extensive public 21 | 22 | -- An extensive category with finite products 23 | record ExtensiveDistributiveCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 24 | field 25 | U : Category o ℓ e -- U for underlying 26 | extensive : Extensive U 27 | cartesian : Cartesian U 28 | 29 | open Category U public 30 | open Extensive extensive public 31 | open Distributive (Extensive×Cartesian⇒Distributive U extensive cartesian) public 32 | -------------------------------------------------------------------------------- /src/Categories/Category/Finite.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Finite where 4 | 5 | open import Level 6 | 7 | open import Categories.Adjoint.Equivalence 8 | open import Categories.Category.Core 9 | open import Categories.Category.Finite.Fin 10 | 11 | -- definition of a finite category 12 | -- 13 | -- the idea is to require a functor from C to a category generated from a finite shape 14 | -- is the right adjoint. 15 | -- 16 | -- Question: it seems to me right adjoint is enough, though the original plan is to 17 | -- use adjoint equivalence. intuitively, the shape category is an "overapproximation" 18 | -- of C, which is a very strong constraint. so requiring adjoint equivalence sounds an 19 | -- unnecessarily stronger constraint. is adjoint equivalence necessary? 20 | -- 21 | -- Answer: probably yes. adjoint equivalence seems necessary as the notion needs to 22 | -- show that shapes are preserved. 23 | -- 24 | -- c.f. Categories.Adjoint.Equivalence.Properties.⊣equiv-preserves-diagram 25 | record Finite {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 26 | field 27 | shape : FinCatShape 28 | 29 | open FinCatShape public renaming (size to ∣Obj∣) 30 | 31 | shapeCat : Category _ _ _ 32 | shapeCat = FinCategory shape 33 | 34 | -- 35 | -- /------------\ 36 | -- < - \ 37 | -- C | S 38 | -- \ - ^ 39 | -- \------------/ 40 | -- 41 | field 42 | ⊣equiv : ⊣Equivalence shapeCat C 43 | 44 | module ⊣equiv = ⊣Equivalence ⊣equiv 45 | open ⊣equiv public 46 | -------------------------------------------------------------------------------- /src/Categories/Category/Groupoid.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Groupoid where 3 | 4 | open import Level using (Level; suc; _⊔_) 5 | 6 | open import Categories.Category 7 | import Categories.Morphism 8 | 9 | record IsGroupoid {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 10 | open Category C public 11 | open Definitions C public 12 | 13 | open Categories.Morphism C 14 | 15 | infix 10 _⁻¹ 16 | 17 | field 18 | _⁻¹ : ∀ {A B} → A ⇒ B → B ⇒ A 19 | iso : ∀ {A B} {f : A ⇒ B} → Iso f (f ⁻¹) 20 | 21 | module iso {A B f} = Iso (iso {A} {B} {f}) 22 | 23 | equiv-obj : ∀ {A B} → A ⇒ B → A ≅ B 24 | equiv-obj f = record 25 | { from = f 26 | ; to = _ 27 | ; iso = iso 28 | } 29 | 30 | -- this definition doesn't seem to 'carry its weight' 31 | equiv-obj-sym : ∀ {A B} → A ⇒ B → B ≅ A 32 | equiv-obj-sym f = ≅.sym (equiv-obj f) 33 | 34 | -- A groupoid is a category that has a groupoid structure 35 | 36 | record Groupoid (o ℓ e : Level) : Set (suc (o ⊔ ℓ ⊔ e)) where 37 | field 38 | category : Category o ℓ e 39 | isGroupoid : IsGroupoid category 40 | 41 | open IsGroupoid isGroupoid public 42 | -------------------------------------------------------------------------------- /src/Categories/Category/Groupoid/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category.Groupoid 3 | 4 | module Categories.Category.Groupoid.Properties {o ℓ e} (G : Groupoid o ℓ e) where 5 | 6 | import Categories.Morphism as Morphism 7 | import Categories.Morphism.Properties as MorphismProps 8 | import Categories.Morphism.Reasoning as MR 9 | 10 | open Groupoid G 11 | open Morphism category 12 | open MorphismProps category 13 | open HomReasoning 14 | open MR category 15 | 16 | private 17 | variable 18 | A B C : Obj 19 | 20 | mono : {f : A ⇒ B} → Mono f 21 | mono = Iso⇒Mono iso 22 | 23 | epi : {f : A ⇒ B} → Epi f 24 | epi = Iso⇒Epi iso 25 | 26 | id-inverse : id {A = A} ⁻¹ ≈ id 27 | id-inverse = ⟺ identityˡ ○ iso.isoʳ 28 | 29 | ⁻¹-involutive : {f : A ⇒ B} → f ⁻¹ ⁻¹ ≈ f 30 | ⁻¹-involutive {f = f} = begin 31 | f ⁻¹ ⁻¹ ≈⟨ introʳ iso.isoˡ ⟩ 32 | f ⁻¹ ⁻¹ ∘ f ⁻¹ ∘ f ≈⟨ sym-assoc ○ elimˡ iso.isoˡ ⟩ 33 | f ∎ 34 | 35 | ⁻¹-commute : {f : A ⇒ B} {g : C ⇒ A} → (f ∘ g) ⁻¹ ≈ g ⁻¹ ∘ f ⁻¹ 36 | ⁻¹-commute {f = f} {g} = epi _ _ ( begin 37 | (f ∘ g) ⁻¹ ∘ f ∘ g ≈⟨ iso.isoˡ ⟩ 38 | id ≈˘⟨ iso.isoˡ ⟩ 39 | g ⁻¹ ∘ g ≈˘⟨ cancelInner iso.isoˡ ⟩ 40 | (g ⁻¹ ∘ f ⁻¹) ∘ f ∘ g ∎ ) 41 | -------------------------------------------------------------------------------- /src/Categories/Category/Helper.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Helper where 3 | 4 | open import Level 5 | open import Relation.Binary using (Rel; IsEquivalence) 6 | 7 | open import Categories.Category.Core using (Category) 8 | 9 | -- Since we add extra proofs in the definition of `Category` (i.e. `sym-assoc` and 10 | -- `identity²`), we might still want to construct a `Category` in its originally 11 | -- easier manner. Thus, this redundant definition is here to ease the construction. 12 | private 13 | record CategoryHelper (o ℓ e : Level) : Set (suc (o ⊔ ℓ ⊔ e)) where 14 | infix 4 _≈_ _⇒_ 15 | infixr 9 _∘_ 16 | 17 | field 18 | Obj : Set o 19 | _⇒_ : Rel Obj ℓ 20 | _≈_ : ∀ {A B} → Rel (A ⇒ B) e 21 | 22 | id : ∀ {A} → (A ⇒ A) 23 | _∘_ : ∀ {A B C} → (B ⇒ C) → (A ⇒ B) → (A ⇒ C) 24 | 25 | field 26 | assoc : ∀ {A B C D} {f : A ⇒ B} {g : B ⇒ C} {h : C ⇒ D} → (h ∘ g) ∘ f ≈ h ∘ (g ∘ f) 27 | identityˡ : ∀ {A B} {f : A ⇒ B} → id ∘ f ≈ f 28 | identityʳ : ∀ {A B} {f : A ⇒ B} → f ∘ id ≈ f 29 | equiv : ∀ {A B} → IsEquivalence (_≈_ {A} {B}) 30 | ∘-resp-≈ : ∀ {A B C} {f h : B ⇒ C} {g i : A ⇒ B} → f ≈ h → g ≈ i → f ∘ g ≈ h ∘ i 31 | 32 | categoryHelper : ∀ {o ℓ e} → CategoryHelper o ℓ e → Category o ℓ e 33 | categoryHelper C = record 34 | { Obj = Obj 35 | ; _⇒_ = _⇒_ 36 | ; _≈_ = _≈_ 37 | ; id = id 38 | ; _∘_ = _∘_ 39 | ; assoc = assoc 40 | ; sym-assoc = sym assoc 41 | ; identityˡ = identityˡ 42 | ; identityʳ = identityʳ 43 | ; identity² = identityˡ 44 | ; equiv = equiv 45 | ; ∘-resp-≈ = ∘-resp-≈ 46 | } 47 | where open CategoryHelper C 48 | module _ {A B} where 49 | open IsEquivalence (equiv {A} {B}) public 50 | -------------------------------------------------------------------------------- /src/Categories/Category/Indiscrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Indiscrete where 3 | 4 | -- Category where all arrows are inhabited by a single element 5 | open import Level 6 | open import Data.Unit 7 | 8 | open import Categories.Category 9 | 10 | open import Relation.Binary.PropositionalEquality as ≡ 11 | 12 | Indiscrete : ∀ {o ℓ} (A : Set o) → Category o ℓ ℓ 13 | Indiscrete {_} {ℓ} A = record 14 | { Obj = A 15 | ; _⇒_ = λ _ _ → Lift ℓ ⊤ 16 | ; _≈_ = _≡_ 17 | ; id = _ 18 | ; _∘_ = _ 19 | ; assoc = refl 20 | ; sym-assoc = refl 21 | ; identityˡ = refl 22 | ; identityʳ = refl 23 | ; identity² = refl 24 | ; equiv = isEquivalence 25 | ; ∘-resp-≈ = λ where 26 | refl refl → refl 27 | } 28 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Cartesians.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Cartesians where 4 | 5 | open import Level 6 | 7 | open import Categories.Category.Core using (Category) 8 | open import Categories.Category.Helper 9 | open import Categories.Category.Cartesian.Bundle using (CartesianCategory) 10 | open import Categories.Functor.Cartesian 11 | open import Categories.Functor.Cartesian.Properties 12 | open import Categories.NaturalTransformation.NaturalIsomorphism 13 | 14 | module _ o ℓ e where 15 | 16 | Cartesians : Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 17 | Cartesians = categoryHelper record 18 | { Obj = CartesianCategory o ℓ e 19 | ; _⇒_ = CartesianF 20 | ; _≈_ = λ F G → CF.F F ≃ CF.F G 21 | ; id = idF-CartesianF _ 22 | ; _∘_ = ∘-CartesianF 23 | ; assoc = λ {_ _ _ _ F G H} → associator (CF.F F) (CF.F G) (CF.F H) 24 | ; identityˡ = unitorˡ 25 | ; identityʳ = unitorʳ 26 | ; equiv = record 27 | { refl = ≃.refl 28 | ; sym = ≃.sym 29 | ; trans = ≃.trans 30 | } 31 | ; ∘-resp-≈ = _ⓘₕ_ 32 | } 33 | where module CF = CartesianF 34 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Cats.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.Cats where 3 | 4 | -- The (large) category of (small) categories. 5 | -- Even though Agda can figure out the levels, it is worth making them explicit, 6 | -- to see the large level jumps involved. 7 | 8 | open import Level 9 | open import Categories.Category using (Category) 10 | open import Categories.Functor using (Functor; id; _∘F_) 11 | open import Categories.NaturalTransformation.NaturalIsomorphism 12 | using (NaturalIsomorphism; associator; sym-associator; unitorˡ; unitorʳ; unitor²; isEquivalence; _ⓘₕ_) 13 | private 14 | variable 15 | o ℓ e : Level 16 | C D E : Category o ℓ e 17 | F G H I : Functor C D 18 | 19 | Cats : ∀ o ℓ e → Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 20 | Cats o ℓ e = record 21 | { Obj = Category o ℓ e 22 | ; _⇒_ = Functor 23 | ; _≈_ = NaturalIsomorphism 24 | ; id = id 25 | ; _∘_ = _∘F_ 26 | ; assoc = λ {_ _ _ _ F G H} → associator F G H 27 | ; sym-assoc = λ {_ _ _ _ F G H} → sym-associator F G H 28 | ; identityˡ = unitorˡ 29 | ; identityʳ = unitorʳ 30 | ; identity² = unitor² 31 | ; equiv = isEquivalence 32 | ; ∘-resp-≈ = _ⓘₕ_ 33 | } 34 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/DagCats.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --without-K #-} 2 | 3 | module Categories.Category.Instance.DagCats where 4 | 5 | open import Categories.Category.Core using (Category) 6 | open import Categories.Category.Dagger using (DaggerCategory) 7 | open import Categories.Functor.Dagger using (DaggerFunctor; id; _∘F†_) 8 | open import Categories.NaturalTransformation.NaturalIsomorphism 9 | 10 | open import Function.Base using (_on_) 11 | open import Level using (suc; _⊔_) 12 | 13 | DagCats : ∀ o ℓ e → Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 14 | DagCats o ℓ e = record 15 | { Obj = DaggerCategory o ℓ e 16 | ; _⇒_ = DaggerFunctor 17 | ; _≈_ = NaturalIsomorphism on functor 18 | ; id = id 19 | ; _∘_ = _∘F†_ 20 | ; assoc = λ {_ _ _ _ F G H} → associator (functor F) (functor G) (functor H) 21 | ; sym-assoc = λ {_ _ _ _ F G H} → sym-associator (functor F) (functor G) (functor H) 22 | ; identityˡ = unitorˡ 23 | ; identityʳ = unitorʳ 24 | ; identity² = unitor² 25 | ; equiv = record 26 | { refl = refl 27 | ; sym = sym 28 | ; trans = trans 29 | } 30 | ; ∘-resp-≈ = _ⓘₕ_ 31 | } where open DaggerFunctor 32 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/EmptySet.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level 4 | 5 | -- This is really a degenerate version of Categories.Category.Instance.Zero 6 | -- Here EmptySet is not given an explicit name, it is an alias for Lift o ⊥ 7 | module Categories.Category.Instance.EmptySet where 8 | 9 | open import Data.Empty.Polymorphic using (⊥; ⊥-elim) 10 | 11 | open import Relation.Binary using (Setoid) 12 | 13 | open import Categories.Category.Instance.Sets 14 | open import Categories.Category.Instance.Setoids 15 | import Categories.Object.Initial as Init 16 | 17 | module _ {o : Level} where 18 | open Init (Sets o) 19 | 20 | EmptySet-⊥ : Initial 21 | EmptySet-⊥ .Initial.⊥ = ⊥ 22 | EmptySet-⊥ .Initial.⊥-is-initial .IsInitial.! () 23 | 24 | module _ {c ℓ : Level} where 25 | open Init (Setoids c ℓ) 26 | 27 | EmptySetoid : Setoid c ℓ 28 | EmptySetoid = record 29 | { Carrier = ⊥ 30 | ; _≈_ = λ () 31 | ; isEquivalence = record { refl = λ { {()} } ; sym = λ { {()} } ; trans = λ { {()} } } 32 | } 33 | 34 | EmptySetoid-⊥ : Initial 35 | EmptySetoid-⊥ = record 36 | { ⊥ = EmptySetoid 37 | ; ⊥-is-initial = record 38 | { ! = record 39 | { to = λ () 40 | ; cong = λ { {()} } 41 | } 42 | ; !-unique = λ { _ {()} } 43 | } 44 | } 45 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/FinSetoids.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.FinSetoids where 3 | 4 | -- Category of Finite Setoids, as a sub-category of Setoids 5 | 6 | open import Level 7 | 8 | open import Data.Fin.Base using (Fin) 9 | open import Data.Nat.Base using (ℕ) 10 | open import Data.Product using (Σ) 11 | open import Function.Bundles using (Inverse) 12 | open import Relation.Unary using (Pred) 13 | open import Relation.Binary.Bundles using (Setoid; module Setoid) 14 | import Relation.Binary.PropositionalEquality as ≡ 15 | 16 | open import Categories.Category.Core using (Category) 17 | open import Categories.Category.Construction.ObjectRestriction 18 | open import Categories.Category.Instance.Setoids 19 | 20 | -- The predicate that will be used 21 | IsFiniteSetoid : {c ℓ : Level} → Pred (Setoid c ℓ) (c ⊔ ℓ) 22 | IsFiniteSetoid X = Σ ℕ (λ n → Inverse X (≡.setoid (Fin n))) 23 | 24 | -- The actual Category 25 | FinSetoids : (c ℓ : Level) → Category (suc (c ⊔ ℓ)) (c ⊔ ℓ) (c ⊔ ℓ) 26 | FinSetoids c ℓ = ObjectRestriction (Setoids c ℓ) IsFiniteSetoid 27 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/LawvereTheories.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.LawvereTheories where 3 | 4 | -- Category of Lawvere Theories 5 | 6 | open import Level 7 | 8 | open import Categories.Category.Core using (Category) 9 | open import Categories.Functor.Cartesian using (CartesianF) 10 | open import Categories.NaturalTransformation.NaturalIsomorphism 11 | using (_≃_; associator; sym-associator; unitorˡ; unitorʳ; unitor²; refl; sym; trans; _ⓘₕ_) 12 | open import Categories.Theory.Lawvere using (LawvereTheory; LT-Hom; LT-id; LT-∘) 13 | 14 | LawvereTheories : (ℓ e : Level) → Category (suc (ℓ ⊔ e)) (ℓ ⊔ e) (ℓ ⊔ e) 15 | LawvereTheories ℓ e = record 16 | { Obj = LawvereTheory ℓ e 17 | ; _⇒_ = LT-Hom 18 | ; _≈_ = λ H₁ H₂ → cartF.F H₁ ≃ cartF.F H₂ 19 | ; id = LT-id 20 | ; _∘_ = LT-∘ 21 | ; assoc = λ { {f = f} {g} {h} → associator (cartF.F f) (cartF.F g) (cartF.F h) } 22 | ; sym-assoc = λ { {f = f} {g} {h} → sym-associator (cartF.F f) (cartF.F g) (cartF.F h) } 23 | ; identityˡ = unitorˡ 24 | ; identityʳ = unitorʳ 25 | ; identity² = unitor² 26 | ; equiv = record { refl = refl ; sym = sym ; trans = trans } 27 | ; ∘-resp-≈ = _ⓘₕ_ 28 | } 29 | where open LT-Hom 30 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Monoidals.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Monoidals where 4 | 5 | open import Level 6 | 7 | open import Categories.Category 8 | open import Categories.Category.Helper 9 | open import Categories.Category.Monoidal 10 | open import Categories.Functor.Monoidal 11 | open import Categories.Functor.Monoidal.Properties 12 | open import Categories.NaturalTransformation.NaturalIsomorphism.Monoidal 13 | 14 | module _ o ℓ e where 15 | 16 | Monoidals : Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 17 | Monoidals = categoryHelper record 18 | { Obj = MonoidalCategory o ℓ e 19 | ; _⇒_ = MonoidalFunctor 20 | ; _≈_ = Lax._≃_ 21 | ; id = idF-Monoidal _ 22 | ; _∘_ = ∘-Monoidal 23 | -- NOTE: these η-expanded versions typecheck much faster... 24 | ; assoc = λ {_ _ _ _ F G H} → associator {F = F} {G} {H} 25 | ; identityˡ = λ {_ _ F} → unitorˡ {F = F} 26 | ; identityʳ = λ {_ _ F} → unitorʳ {F = F} 27 | ; equiv = isEquivalence 28 | ; ∘-resp-≈ = _ⓘₕ_ 29 | } 30 | where open Lax 31 | 32 | StrongMonoidals : Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 33 | StrongMonoidals = categoryHelper record 34 | { Obj = MonoidalCategory o ℓ e 35 | ; _⇒_ = StrongMonoidalFunctor 36 | ; _≈_ = Strong._≃_ 37 | ; id = idF-StrongMonoidal _ 38 | ; _∘_ = ∘-StrongMonoidal 39 | -- NOTE: these η-expanded versions typecheck much faster... 40 | ; assoc = λ {_ _ _ _ F G H} → associator {F = F} {G} {H} 41 | ; identityˡ = λ {_ _ F} → unitorˡ {F = F} 42 | ; identityʳ = λ {_ _ F} → unitorʳ {F = F} 43 | ; equiv = isEquivalence 44 | ; ∘-resp-≈ = _ⓘₕ_ 45 | } 46 | where open Strong 47 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/One.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level 4 | 5 | module Categories.Category.Instance.One where 6 | 7 | open import Data.Unit using (⊤; tt) 8 | 9 | open import Categories.Category 10 | open import Categories.Functor 11 | open import Categories.Category.Instance.Cats 12 | import Categories.Object.Terminal as Term 13 | 14 | module _ {o ℓ e : Level} where 15 | open Term (Cats o ℓ e) 16 | 17 | One : Category o ℓ e 18 | One = record 19 | { Obj = Lift o ⊤ 20 | ; _⇒_ = λ _ _ → Lift ℓ ⊤ 21 | ; _≈_ = λ _ _ → Lift e ⊤ 22 | } 23 | 24 | One-⊤ : Terminal 25 | One-⊤ = record { ⊤ = One ; ⊤-is-terminal = record { ! = record { F₀ = λ _ → lift tt } } } 26 | 27 | -- not only is One terminal, it can be shifted anywhere else. Stronger version of ! 28 | shift : {o ℓ e : Level} (o′ ℓ′ e′ : Level) → Functor (One {o} {ℓ} {e}) (One {o′} {ℓ′} {e′}) 29 | shift o′ ℓ′ e′ = _ -- so obvious, Agda can fill it all in automatically! 30 | 31 | One0 : Category 0ℓ 0ℓ 0ℓ 32 | One0 = One {0ℓ} {0ℓ} {0ℓ} 33 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/PointedSets.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.PointedSets where 3 | 4 | -- Category of Pointed Sets 5 | 6 | open import Level 7 | open import Relation.Binary 8 | open import Function using (_∘′_; _$_) renaming (id to idf) 9 | open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; refl) 10 | open import Data.Product 11 | 12 | open import Categories.Category 13 | open import Categories.Category.Instance.Sets 14 | open import Categories.Functor using (Functor) 15 | 16 | PointedSets : ∀ o → Category (suc o) o o 17 | PointedSets o = record 18 | { Obj = Σ (Set o) idf 19 | ; _⇒_ = λ c d → Σ (proj₁ c → proj₁ d) (λ f → f (proj₂ c) ≡ proj₂ d) 20 | ; _≈_ = λ f g → ∀ x → proj₁ f x ≡ proj₁ g x 21 | ; id = idf , refl 22 | ; _∘_ = λ { (f , f-pres-pt) (g , g-pres-pt) → f ∘′ g , ≡.trans (≡.cong f g-pres-pt) f-pres-pt} 23 | ; assoc = λ _ → refl 24 | ; sym-assoc = λ _ → refl 25 | ; identityˡ = λ _ → refl 26 | ; identityʳ = λ _ → refl 27 | ; identity² = λ _ → refl 28 | ; equiv = record { refl = λ _ → refl ; sym = λ pf x → ≡.sym $ pf x ; trans = λ i≡j j≡k x → ≡.trans (i≡j x) (j≡k x) } 29 | ; ∘-resp-≈ = λ {_} {_} {_} {_} {h} {g} f≡h g≡i x → ≡.trans (f≡h (proj₁ g x)) (≡.cong (proj₁ h) $ g≡i x) 30 | } 31 | 32 | Underlying : {o : Level} → Functor (PointedSets o) (Sets o) 33 | Underlying = record 34 | { F₀ = proj₁ 35 | ; F₁ = proj₁ 36 | ; identity = λ _ → refl 37 | ; homomorphism = λ _ → refl 38 | ; F-resp-≈ = λ f≈g → f≈g 39 | } 40 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Preds.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Preds where 4 | 5 | -- Category of (Agda) predicates, aka (indexed types, families of functions, pointwise equality with implicit value) 6 | -- Note the (explicit) levels in each 7 | 8 | open import Function renaming (id to idf) using (_∘_) 9 | open import Level using (suc) 10 | open import Relation.Binary.PropositionalEquality using (_≡_; cong; refl; sym; trans) 11 | open import Relation.Unary using (_⇒_; Pred; IUniversal) 12 | 13 | open import Categories.Category using (Category) 14 | 15 | Preds : ∀ o → Set o → Category (suc o) o o 16 | Preds o A = record 17 | { Obj = Pred A o 18 | ; _⇒_ = λ P Q → ∀[ P ⇒ Q ] 19 | ; _≈_ = λ {P} f g → (x : A) → (y : P x) → f y ≡ g y 20 | ; id = idf 21 | ; _∘_ = λ α β → α ∘ β 22 | ; assoc = λ _ _ → refl 23 | ; sym-assoc = λ _ _ → refl 24 | ; identityˡ = λ _ _ → refl 25 | ; identityʳ = λ _ _ → refl 26 | ; identity² = λ _ _ → refl 27 | ; equiv = record 28 | { refl = λ _ _ → refl 29 | ; sym = λ α a x → sym (α a x) 30 | ; trans = λ α β a x → trans (α a x) (β a x) 31 | } 32 | ; ∘-resp-≈ = λ {_} {_} {_} {f} {_} {_} {i} α β a x → trans (cong f (β a x)) (α a (i x)) 33 | } 34 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Properties/Setoids.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Properties.Setoids where 4 | 5 | open import Categories.Category.Instance.Properties.Setoids.Complete 6 | using (Setoids-Complete) 7 | public 8 | 9 | open import Categories.Category.Instance.Properties.Setoids.Cocomplete 10 | using (Setoids-Cocomplete) 11 | public 12 | 13 | open import Categories.Category.Instance.Properties.Setoids.LCCC 14 | using (Setoids-LCCC) 15 | public 16 | 17 | open import Categories.Category.Instance.Properties.Setoids.CCC 18 | using (Setoids-CCC) 19 | public 20 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Quivers.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Quivers where 4 | 5 | -- The Category of Quivers 6 | 7 | open import Level using (Level; suc; _⊔_) 8 | open import Relation.Binary.PropositionalEquality.Core using (refl) 9 | open import Data.Quiver using (Quiver) 10 | open import Data.Quiver.Morphism using (Morphism; id; _∘_; _≃_; ≃-Equivalence; ≃-resp-∘) 11 | 12 | open import Categories.Category.Core using (Category) 13 | 14 | private 15 | variable 16 | o ℓ e o′ ℓ′ e′ : Level 17 | 18 | Quivers : ∀ o ℓ e → Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 19 | Quivers o ℓ e = record 20 | { Obj = Quiver o ℓ e 21 | ; _⇒_ = Morphism 22 | ; _≈_ = _≃_ 23 | ; id = id 24 | ; _∘_ = _∘_ 25 | ; assoc = λ {_ _ _ G} → record { F₀≡ = refl ; F₁≡ = Equiv.refl G } 26 | ; sym-assoc = λ {_ _ _ G} → record { F₀≡ = refl ; F₁≡ = Equiv.refl G } 27 | ; identityˡ = λ {_ G} → record { F₀≡ = refl ; F₁≡ = Equiv.refl G } 28 | ; identityʳ = λ {_ G} → record { F₀≡ = refl ; F₁≡ = Equiv.refl G } 29 | ; identity² = λ {G} → record { F₀≡ = refl ; F₁≡ = Equiv.refl G } 30 | ; equiv = ≃-Equivalence 31 | ; ∘-resp-≈ = ≃-resp-∘ 32 | } 33 | where open Quiver using (module Equiv) 34 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Rels.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.Rels where 3 | 4 | open import Data.Product 5 | open import Function hiding (_⇔_) 6 | open import Level 7 | open import Relation.Binary 8 | open import Relation.Binary.Construct.Composition 9 | open import Relation.Binary.PropositionalEquality 10 | 11 | open import Categories.Category.Core 12 | 13 | -- the category whose objects are sets and whose morphisms are binary relations. 14 | Rels : ∀ o ℓ → Category (suc o) (suc (o ⊔ ℓ)) (o ⊔ ℓ) 15 | Rels o ℓ = record 16 | { Obj = Set o 17 | ; _⇒_ = λ A B → REL A B (o ⊔ ℓ) 18 | ; _≈_ = λ L R → L ⇔ R 19 | ; id = λ x y → Lift ℓ (x ≡ y) 20 | ; _∘_ = λ L R → R ; L 21 | ; assoc = (λ { (b , fxb , c , gbc , hcy) → c , ((b , (fxb , gbc)) , hcy)}) 22 | , λ { (c , (b , fxb , gbc) , hcy) → b , fxb , c , gbc , hcy} 23 | ; sym-assoc = (λ { (c , (b , fxb , gbc) , hcy) → b , fxb , c , gbc , hcy}) 24 | , (λ { (b , fxb , c , gbc , hcy) → c , ((b , (fxb , gbc)) , hcy)}) 25 | ; identityˡ = (λ { (b , fxb , lift refl) → fxb}) , λ {_} {y} fxy → y , fxy , lift refl 26 | ; identityʳ = (λ { (a , lift refl , fxy) → fxy}) , λ {x} {_} fxy → x , lift refl , fxy 27 | ; identity² = (λ { (_ , lift p , lift q) → lift (trans p q)}) , λ { (lift refl) → _ , lift refl , lift refl } 28 | ; equiv = record 29 | { refl = id , id 30 | ; sym = swap 31 | ; trans = λ { (p₁ , p₂) (q₁ , q₂) → (q₁ ∘′ p₁) , p₂ ∘′ q₂} 32 | } 33 | ; ∘-resp-≈ = λ f⇔h g⇔i → 34 | (λ { (b , gxb , fky) → b , proj₁ g⇔i gxb , proj₁ f⇔h fky }) , 35 | λ { (b , ixb , hby) → b , proj₂ g⇔i ixb , proj₂ f⇔h hby } 36 | } 37 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Setoids.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.Setoids where 3 | 4 | -- Category of Setoids, aka (Setoid, _⟶_, Setoid ≈) 5 | -- Note the (explicit) levels in each 6 | 7 | open import Level using (suc; _⊔_) 8 | open import Relation.Binary.Bundles using (Setoid) 9 | open import Function.Bundles using (Func; _⟨$⟩_) 10 | open import Function.Base using (_$_) 11 | import Function.Construct.Composition as Comp 12 | import Function.Construct.Identity as Id 13 | import Function.Construct.Setoid as S 14 | 15 | open import Categories.Category.Core using (Category) 16 | 17 | open Func 18 | open Setoid 19 | 20 | Setoids : ∀ c ℓ → Category (suc (c ⊔ ℓ)) (c ⊔ ℓ) (c ⊔ ℓ) 21 | Setoids c ℓ = record 22 | { Obj = Setoid c ℓ 23 | ; _⇒_ = Func 24 | ; _≈_ = λ {A} {B} f g → _≈_ (S.setoid A B) f g 25 | ; id = Id.function _ 26 | ; _∘_ = λ f g → Comp.function g f 27 | ; assoc = λ {_} {_} {_} {D} → refl D 28 | ; sym-assoc = λ {_} {_} {_} {D} → refl D 29 | ; identityˡ = λ {_} {B} → refl B 30 | ; identityʳ = λ {_} {B} → refl B 31 | ; identity² = λ {A} → refl A 32 | ; equiv = λ {A} {B} → isEquivalence (S.setoid A B) 33 | ; ∘-resp-≈ = λ {_} {_} {C} {f} {h} {g} {i} f≈h g≈i {x} → trans C f≈h (cong h g≈i) 34 | } 35 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Sets.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.Sets where 3 | 4 | -- Category of (Agda) Sets, aka (types, functions, pointwise equality with implicit value) 5 | -- Note the (explicit) levels in each 6 | 7 | open import Level 8 | open import Relation.Binary 9 | open import Function using (_∘′_) renaming (id to idf) 10 | open import Relation.Binary.PropositionalEquality as ≡ using (_≡_; _≗_) 11 | 12 | open import Categories.Category 13 | 14 | Sets : ∀ o → Category (suc o) o o 15 | Sets o = record 16 | { Obj = Set o 17 | ; _⇒_ = λ c d → c → d 18 | ; _≈_ = _≗_ 19 | ; id = idf 20 | ; _∘_ = _∘′_ 21 | ; assoc = λ _ → ≡.refl 22 | ; sym-assoc = λ _ → ≡.refl 23 | ; identityˡ = λ _ → ≡.refl 24 | ; identityʳ = λ _ → ≡.refl 25 | ; identity² = λ _ → ≡.refl 26 | ; equiv = record 27 | { refl = λ _ → ≡.refl 28 | ; sym = λ eq x → ≡.sym (eq x) 29 | ; trans = λ eq₁ eq₂ x → ≡.trans (eq₁ x) (eq₂ x) 30 | } 31 | ; ∘-resp-≈ = resp 32 | } 33 | where resp : ∀ {A B C : Set o} {f h : B → C} {g i : A → B} → 34 | (f ≗ h) → (g ≗ i) → f ∘′ g ≗ h ∘′ i 35 | resp {h = h} eq₁ eq₂ x = ≡.trans (eq₁ _) (≡.cong h (eq₂ x)) 36 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/SimplicialSet.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.SimplicialSet where 4 | 5 | open import Level 6 | 7 | open import Categories.Category 8 | open import Categories.Category.Instance.Simplex 9 | open import Categories.Category.Construction.Presheaves 10 | 11 | SimplicialSet : ∀ o ℓ → Category (suc (o ⊔ ℓ)) (o ⊔ ℓ) (o ⊔ ℓ) 12 | SimplicialSet o ℓ = Presheaves {o′ = o} {ℓ′ = ℓ} Δ 13 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/SingletonSet.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level 4 | 5 | -- This is really a degenerate version of Categories.Category.Instance.One 6 | -- Here SingletonSet is not given an explicit name, it is an alias for ⊤ 7 | module Categories.Category.Instance.SingletonSet where 8 | 9 | open import Data.Unit.Polymorphic using (⊤; tt) 10 | open import Relation.Binary using (Setoid) 11 | open import Relation.Binary.PropositionalEquality using (refl) 12 | 13 | open import Categories.Category.Instance.Sets 14 | open import Categories.Category.Instance.Setoids 15 | import Categories.Object.Terminal as Term 16 | 17 | module _ {o : Level} where 18 | open Term (Sets o) 19 | 20 | SingletonSet-⊤ : Terminal 21 | SingletonSet-⊤ = record { ⊤ = ⊤ ; ⊤-is-terminal = record { !-unique = λ _ _ → refl } } 22 | 23 | module _ {c ℓ : Level} where 24 | open Term (Setoids c ℓ) 25 | 26 | SingletonSetoid : Setoid c ℓ 27 | SingletonSetoid = record { Carrier = ⊤ ; _≈_ = λ _ _ → ⊤ } 28 | 29 | SingletonSetoid-⊤ : Terminal 30 | SingletonSetoid-⊤ = record { ⊤ = SingletonSetoid } 31 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/StrictCats.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.StrictCats where 3 | 4 | -- The (large) 'strict' category of (small) categories. 5 | -- The difference here is that _≈_ is not |NaturalIsomorphism| but |_≈F_| 6 | 7 | open import Level 8 | open import Relation.Binary.PropositionalEquality using (refl) 9 | 10 | open import Categories.Category using (Category) 11 | open import Categories.Functor using (Functor; id; _∘F_) 12 | open import Categories.Functor.Equivalence 13 | 14 | private 15 | variable 16 | o ℓ e : Level 17 | 18 | StrictCats : ∀ o ℓ e → Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 19 | StrictCats o ℓ e = record 20 | { Obj = Category o ℓ e 21 | ; _⇒_ = Functor 22 | ; _≈_ = _≡F_ 23 | ; id = id 24 | ; _∘_ = _∘F_ 25 | ; assoc = λ {_ _ _ _ F G H} → ≡F-assoc {F = F} {G} {H} 26 | ; sym-assoc = λ {_ _ _ _ F G H} → ≡F-sym-assoc {F = F} {G} {H} 27 | ; identityˡ = ≡F-identityˡ 28 | ; identityʳ = ≡F-identityʳ 29 | ; identity² = ≡F-identity² 30 | ; equiv = ≡F-equiv 31 | ; ∘-resp-≈ = ∘F-resp-≡F 32 | } 33 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/StrictGroupoids.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Instance.StrictGroupoids where 3 | 4 | -- The 'strict' category of groupoids. 5 | -- The difference here is that _≈_ is not |NaturalIsomorphism| but |_≡F_| 6 | 7 | open import Level 8 | open import Relation.Binary.PropositionalEquality using (refl) 9 | 10 | open import Categories.Category using (Category) 11 | open import Categories.Category.Groupoid using (Groupoid) 12 | open import Categories.Category.Instance.Groupoids using (F-resp-⁻¹) 13 | open import Categories.Functor using (Functor; id; _∘F_) 14 | open import Categories.Functor.Equivalence 15 | 16 | private 17 | variable 18 | o ℓ e : Level 19 | 20 | open Groupoid using (category) 21 | 22 | StrictGroupoids : ∀ o ℓ e → Category (suc (o ⊔ ℓ ⊔ e)) (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) 23 | StrictGroupoids o ℓ e = record 24 | { Obj = Groupoid o ℓ e 25 | ; _⇒_ = λ G H → Functor (category G) (category H) 26 | ; _≈_ = _≡F_ 27 | ; id = id 28 | ; _∘_ = _∘F_ 29 | ; assoc = λ {_ _ _ _ F G H} → ≡F-assoc {F = F} {G} {H} 30 | ; sym-assoc = λ {_ _ _ _ F G H} → ≡F-sym-assoc {F = F} {G} {H} 31 | ; identityˡ = ≡F-identityˡ 32 | ; identityʳ = ≡F-identityʳ 33 | ; identity² = ≡F-identity² 34 | ; equiv = ≡F-equiv 35 | ; ∘-resp-≈ = ∘F-resp-≡F 36 | } 37 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Zero.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Zero {o ℓ e} where 4 | 5 | open import Categories.Category.Instance.Zero.Core {o} {ℓ} {e} public 6 | open import Categories.Category.Instance.Zero.Properties {o} {ℓ} {e} public 7 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Zero/Core.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Zero.Core {o ℓ e} where 4 | 5 | open import Categories.Category.Core using (Category) 6 | open import Data.Empty.Polymorphic using (⊥) 7 | 8 | open Category 9 | 10 | -- A level-polymorphic empty category 11 | Zero : Category o ℓ e 12 | Zero .Obj = ⊥ 13 | -------------------------------------------------------------------------------- /src/Categories/Category/Instance/Zero/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Instance.Zero.Properties {o ℓ e} where 4 | 5 | open import Data.Empty.Polymorphic using (⊥-elim) 6 | 7 | open import Categories.Category using (Category) 8 | open import Categories.Functor using (Functor) 9 | open import Categories.Category.Instance.Cats using (Cats) 10 | open import Categories.Category.Instance.Zero.Core using (Zero) 11 | open import Categories.NaturalTransformation.NaturalIsomorphism using (niHelper) 12 | open import Categories.Object.Initial (Cats o ℓ e) using (Initial; IsInitial) 13 | 14 | open Initial 15 | open IsInitial 16 | open Functor 17 | 18 | -- Unlike for ⊤ being Terminal, Agda can't deduce these, need to be explicit 19 | Zero-⊥ : Initial 20 | Zero-⊥ .⊥ = Zero 21 | Zero-⊥ .⊥-is-initial .! .F₀ () 22 | Zero-⊥ .⊥-is-initial .!-unique f = niHelper record 23 | { η = λ() 24 | ; η⁻¹ = λ() 25 | ; commute = λ{ {()} } 26 | ; iso = λ() 27 | } 28 | -------------------------------------------------------------------------------- /src/Categories/Category/Inverse.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Inverse where 3 | 4 | open import Level using (Level; suc; _⊔_) 5 | 6 | open import Categories.Category 7 | open import Data.Product 8 | import Categories.Morphism 9 | 10 | record pseudo-iso {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 11 | open Category C 12 | open Definitions C 13 | open Categories.Morphism C 14 | 15 | infix 10 _⁻¹ 16 | field 17 | _⁻¹ : ∀ {A B} → (f : A ⇒ B) → B ⇒ A 18 | pseudo-iso₁ : ∀ {A B} {f : A ⇒ B} → f ∘ f ⁻¹ ∘ f ≈ f 19 | pseudo-iso₂ : ∀ {A B} {f : A ⇒ B} → f ⁻¹ ∘ f ∘ f ⁻¹ ≈ f ⁻¹ 20 | 21 | 22 | record Inverse {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 23 | open Category C 24 | open Definitions C 25 | open Categories.Morphism C 26 | open pseudo-iso 27 | 28 | field 29 | piso : pseudo-iso C 30 | unique : ∀ {p : pseudo-iso C} {A B} → (f : A ⇒ B) → _⁻¹ piso f ≈ _⁻¹ p f 31 | -------------------------------------------------------------------------------- /src/Categories/Category/Lift.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Lift where 4 | 5 | open import Level 6 | open import Categories.Category 7 | open import Categories.Functor using (Functor) 8 | 9 | liftC : ∀ {o ℓ e} o′ ℓ′ e′ → Category o ℓ e → Category (o ⊔ o′) (ℓ ⊔ ℓ′) (e ⊔ e′) 10 | liftC o′ ℓ′ e′ C = record 11 | { Obj = Lift o′ Obj 12 | ; _⇒_ = λ X Y → Lift ℓ′ (lower X ⇒ lower Y) 13 | ; _≈_ = λ f g → Lift e′ (lower f ≈ lower g) 14 | ; id = lift id 15 | ; _∘_ = λ f g → lift (lower f ∘ lower g) 16 | ; assoc = lift assoc 17 | ; sym-assoc = lift sym-assoc 18 | ; identityˡ = lift identityˡ 19 | ; identityʳ = lift identityʳ 20 | ; identity² = lift identity² 21 | ; equiv = record 22 | { refl = lift Equiv.refl 23 | ; sym = λ eq → lift (Equiv.sym (lower eq)) 24 | ; trans = λ eq eq′ → lift (Equiv.trans (lower eq) (lower eq′)) 25 | } 26 | ; ∘-resp-≈ = λ eq eq′ → lift (∘-resp-≈ (lower eq) (lower eq′)) 27 | } 28 | where open Category C 29 | 30 | liftF : ∀ {o ℓ e} o′ ℓ′ e′ (C : Category o ℓ e) → Functor C (liftC o′ ℓ′ e′ C) 31 | liftF o′ ℓ′ e′ C = record 32 | { F₀ = lift 33 | ; F₁ = lift 34 | ; identity = lift refl 35 | ; homomorphism = lift refl 36 | ; F-resp-≈ = lift 37 | } 38 | where open Category C 39 | open Equiv 40 | 41 | unliftF : ∀ {o ℓ e} o′ ℓ′ e′ (C : Category o ℓ e) → Functor (liftC o′ ℓ′ e′ C) C 42 | unliftF o′ ℓ′ e′ C = record 43 | { F₀ = lower 44 | ; F₁ = lower 45 | ; identity = refl 46 | ; homomorphism = refl 47 | ; F-resp-≈ = lower 48 | } 49 | where open Category C 50 | open Equiv 51 | -------------------------------------------------------------------------------- /src/Categories/Category/Lift/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Lift.Properties where 4 | 5 | open import Level 6 | 7 | open import Categories.Category.Core 8 | open import Categories.Category.Equivalence 9 | open import Categories.Category.Lift 10 | open import Categories.NaturalTransformation.NaturalIsomorphism 11 | import Categories.Morphism.Reasoning.Core as MR 12 | 13 | unliftF-liftF-weakInverse : ∀ {o ℓ e} o′ ℓ′ e′ (C : Category o ℓ e) → WeakInverse (unliftF o′ ℓ′ e′ C) (liftF o′ ℓ′ e′ C) 14 | unliftF-liftF-weakInverse o′ ℓ′ e′ C = record 15 | { F∘G≈id = niHelper record 16 | { η = λ _ → id 17 | ; η⁻¹ = λ _ → id 18 | ; commute = λ f → id-comm-sym 19 | ; iso = λ _ → record 20 | { isoˡ = identity² 21 | ; isoʳ = identity² 22 | } 23 | } 24 | ; G∘F≈id = niHelper record 25 | { η = λ _ → lift id 26 | ; η⁻¹ = λ _ → lift id 27 | ; commute = λ _ → lift id-comm-sym 28 | ; iso = λ _ → record 29 | { isoˡ = lift identity² 30 | ; isoʳ = lift identity² 31 | } 32 | } 33 | } 34 | where 35 | open Category C 36 | open MR C 37 | 38 | liftC-strongEquivalence : ∀ {o ℓ e} o′ ℓ′ e′ (C : Category o ℓ e) → StrongEquivalence (liftC o′ ℓ′ e′ C) C 39 | liftC-strongEquivalence o′ ℓ′ e′ C = record { weak-inverse = unliftF-liftF-weakInverse o′ ℓ′ e′ C } 40 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Definition of Monoidal Category 4 | 5 | -- Big design decision that differs from the previous version: 6 | -- Do not go through "Functor.Power" to encode variables and work 7 | -- at the level of NaturalIsomorphisms, instead work at the object/morphism 8 | -- level, via the more direct _⊗₀_ _⊗₁_ _⊗- -⊗_. 9 | -- The original design needed quite a few contortions to get things working, 10 | -- but these are simply not needed when working directly with the morphisms. 11 | -- 12 | -- Smaller design decision: export some items with long names 13 | -- (unitorˡ, unitorʳ and associator), but internally work with the more classical 14 | -- short greek names (λ, ρ and α respectively). 15 | 16 | module Categories.Category.Monoidal where 17 | 18 | open import Categories.Category.Monoidal.Core public 19 | open import Categories.Category.Monoidal.Bundle public 20 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal/Bundle.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Bundled version of Monoidal Category 4 | module Categories.Category.Monoidal.Bundle where 5 | 6 | open import Level 7 | 8 | open import Categories.Category.Core using (Category) 9 | open import Categories.Category.Monoidal.Core using (Monoidal) 10 | open import Categories.Category.Monoidal.Braided using (Braided) 11 | open import Categories.Category.Monoidal.Symmetric using (Symmetric) 12 | 13 | record MonoidalCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 14 | field 15 | U : Category o ℓ e 16 | monoidal : Monoidal U 17 | 18 | open Category U public 19 | open Monoidal monoidal public 20 | 21 | record BraidedMonoidalCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 22 | field 23 | U : Category o ℓ e 24 | monoidal : Monoidal U 25 | braided : Braided monoidal 26 | 27 | monoidalCategory : MonoidalCategory o ℓ e 28 | monoidalCategory = record { U = U ; monoidal = monoidal } 29 | 30 | open Category U public 31 | open Braided braided public 32 | 33 | record SymmetricMonoidalCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 34 | field 35 | U : Category o ℓ e 36 | monoidal : Monoidal U 37 | symmetric : Symmetric monoidal 38 | 39 | open Category U public 40 | open Symmetric symmetric public 41 | 42 | braidedMonoidalCategory : BraidedMonoidalCategory o ℓ e 43 | braidedMonoidalCategory = record 44 | { U = U 45 | ; monoidal = monoidal 46 | ; braided = braided 47 | } 48 | 49 | open BraidedMonoidalCategory braidedMonoidalCategory public 50 | using (monoidalCategory) 51 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal/CompactClosed.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Category.Monoidal.Core using (Monoidal) 5 | open import Categories.Category.Monoidal.Symmetric 6 | open import Data.Sum 7 | 8 | module Categories.Category.Monoidal.CompactClosed {o ℓ e} {C : Category o ℓ e} (M : Monoidal C) where 9 | 10 | open import Level 11 | 12 | open import Categories.Functor.Bifunctor 13 | open import Categories.NaturalTransformation.NaturalIsomorphism 14 | open import Categories.Category.Monoidal.Rigid 15 | 16 | open Category C 17 | open Commutation C 18 | 19 | record CompactClosed : Set (levelOfTerm M) where 20 | field 21 | symmetric : Symmetric M 22 | rigid : LeftRigid M ⊎ RightRigid M 23 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal/Construction/Minus2.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Monoidal.Construction.Minus2 where 4 | 5 | -- Any -2-Category is Monoidal. Of course, One is Monoidal, but 6 | -- we don't need to shrink to do this, it can be done directly. 7 | -- The assumptions in the construction of a -2-Category are all 8 | -- needed to make things work properly. 9 | 10 | open import Data.Product using (proj₁; proj₂) 11 | 12 | open import Categories.Minus2-Category 13 | open import Categories.Category.Monoidal 14 | import Categories.Morphism as M 15 | 16 | -- Doing it manually is just as easy as going through Cartesian here 17 | -2-Monoidal : ∀ {o ℓ e} → (C : -2-Category {o} {ℓ} {e}) → Monoidal (-2-Category.cat C) 18 | -2-Monoidal C = record 19 | { ⊗ = record 20 | { F₀ = proj₁ 21 | ; F₁ = proj₁ 22 | ; identity = Hom-Conn 23 | ; homomorphism = Hom-Conn 24 | ; F-resp-≈ = λ _ → Hom-Conn 25 | } 26 | ; unit = proj₁ Obj-Contr 27 | ; unitorˡ = λ {X} → proj₂ Obj-Contr X 28 | ; unitorʳ = M.≅.refl cat 29 | ; associator = M.≅.refl cat 30 | ; unitorˡ-commute-from = Hom-Conn 31 | ; unitorˡ-commute-to = Hom-Conn 32 | ; unitorʳ-commute-from = Hom-Conn 33 | ; unitorʳ-commute-to = Hom-Conn 34 | ; assoc-commute-from = Hom-Conn 35 | ; assoc-commute-to = Hom-Conn 36 | ; triangle = Hom-Conn 37 | ; pentagon = Hom-Conn 38 | } 39 | where 40 | open -2-Category C 41 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal/Instance/Cats.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- The category of Cats is Monoidal 4 | 5 | module Categories.Category.Monoidal.Instance.Cats where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.BinaryProducts using (BinaryProducts) 10 | open import Categories.Category.Cartesian using (Cartesian) 11 | open import Categories.Category.Cartesian.Monoidal using (module CartesianMonoidal) 12 | open import Categories.Category.Instance.Cats using (Cats) 13 | open import Categories.Category.Instance.One using (One-⊤) 14 | open import Categories.Category.Monoidal using (Monoidal) 15 | open import Categories.Category.Product using (Product; πˡ; πʳ; _※_) 16 | open import Categories.Category.Product.Properties using (project₁; project₂; unique) 17 | 18 | -- Cats is a Monoidal Category with Product as Bifunctor 19 | module Product {o ℓ e : Level} where 20 | private 21 | C = Cats o ℓ e 22 | 23 | Cats-has-all : BinaryProducts C 24 | Cats-has-all = record { product = λ {A} {B} → record 25 | { A×B = Product A B 26 | ; π₁ = πˡ 27 | ; π₂ = πʳ 28 | ; ⟨_,_⟩ = _※_ 29 | ; project₁ = λ {_} {h} {i} → project₁ {i = h} {j = i} 30 | ; project₂ = λ {_} {h} {i} → project₂ {i = h} {j = i} 31 | ; unique = unique 32 | } } 33 | 34 | Cats-is : Cartesian C 35 | Cats-is = record { terminal = One-⊤ ; products = Cats-has-all } 36 | 37 | Cats-Monoidal : Monoidal C 38 | Cats-Monoidal = CartesianMonoidal.monoidal Cats-is 39 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal/Instance/One.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Monoidal.Instance.One where 4 | 5 | open import Level 6 | open import Data.Unit using (⊤; tt) 7 | 8 | open import Categories.Category 9 | open import Categories.Category.Instance.One 10 | open import Categories.Category.Monoidal 11 | open import Categories.Functor.Bifunctor 12 | open import Categories.Morphism using (_≅_) 13 | 14 | -- That One is monoidal is so easy to prove that Agda can do it all on its own! 15 | One-Monoidal : {o ℓ e : Level} → Monoidal (One {o} {ℓ} {e}) 16 | One-Monoidal = _ 17 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal/Star-Autonomous.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category) 4 | open import Categories.Category.Monoidal.Core using (Monoidal) 5 | open import Categories.Category.Monoidal.Symmetric using (Symmetric) 6 | 7 | module Categories.Category.Monoidal.Star-Autonomous {o ℓ e} {C : Category o ℓ e} (M : Monoidal C) where 8 | 9 | open import Level 10 | 11 | open import Categories.Category.Product using (_⁂_; assocˡ) 12 | open import Categories.Functor using (Functor; _∘F_; id) 13 | open import Categories.Functor.Properties using (FullyFaithful) 14 | open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_) 15 | open import Categories.Functor.Hom 16 | 17 | open Category C renaming (op to Cᵒᵖ) hiding (id) 18 | open Monoidal M 19 | open Functor ⊗ renaming (op to ⊗ₒₚ) 20 | open Hom C 21 | 22 | record Star-Autonomous : Set (levelOfTerm M) where 23 | field 24 | symmetric : Symmetric M 25 | Star : Functor Cᵒᵖ C 26 | 27 | open Functor Star renaming (op to Starₒₚ) public 28 | 29 | field 30 | FF-Star : FullyFaithful Star 31 | A**≃A : (Star ∘F Starₒₚ) ≃ id 32 | 𝒞[A⊗B,C*]≃𝒞[A,B⊗C*] : Hom[-,-] ∘F (⊗ₒₚ ⁂ Star) ≃ Hom[-,-] ∘F (id ⁂ (Star ∘F ⊗ₒₚ)) ∘F assocˡ _ _ _ 33 | -------------------------------------------------------------------------------- /src/Categories/Category/Monoidal/Symmetric/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category) 4 | open import Categories.Category.Monoidal using (Monoidal) 5 | open import Categories.Category.Monoidal.Symmetric using (Symmetric) 6 | 7 | module Categories.Category.Monoidal.Symmetric.Properties 8 | {o ℓ e} {C : Category o ℓ e} {M : Monoidal C} (SM : Symmetric M) where 9 | 10 | open import Data.Product using (_,_) 11 | 12 | import Categories.Category.Monoidal.Braided.Properties as BraidedProperties 13 | open import Categories.Morphism.Reasoning C 14 | 15 | open Category C 16 | open HomReasoning 17 | open Symmetric SM 18 | 19 | -- Shorthands for the braiding 20 | 21 | open BraidedProperties braided public using (module Shorthands) 22 | 23 | -- Extra properties of the braiding in a symmetric monoidal category 24 | 25 | braiding-selfInverse : ∀ {X Y} → braiding.⇐.η (X , Y) ≈ braiding.⇒.η (Y , X) 26 | braiding-selfInverse = introʳ commutative ○ cancelˡ (braiding.iso.isoˡ _) 27 | 28 | inv-commutative : ∀ {X Y} → braiding.⇐.η (X , Y) ∘ braiding.⇐.η (Y , X) ≈ id 29 | inv-commutative = ∘-resp-≈ braiding-selfInverse braiding-selfInverse ○ commutative 30 | -------------------------------------------------------------------------------- /src/Categories/Category/Regular.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Category.Regular where 4 | 5 | -- https://ncatlab.org/nlab/show/regular+category 6 | -- https://en.wikipedia.org/wiki/Regular_category 7 | 8 | open import Level 9 | 10 | open import Categories.Category.Core 11 | open import Categories.Category.Complete.Finitely using (FinitelyComplete) 12 | open import Categories.Diagram.Coequalizer 13 | open import Categories.Diagram.KernelPair 14 | open import Categories.Diagram.Pullback 15 | open import Categories.Morphism.Regular 16 | 17 | record Regular {o ℓ e : Level} (𝒞 : Category o ℓ e) : Set (suc (o ⊔ ℓ ⊔ e)) where 18 | open Category 𝒞 19 | open Pullback 20 | 21 | field 22 | finitely-complete : FinitelyComplete 𝒞 23 | coeq-of-kernelpairs : {A B : Obj} (f : A ⇒ B) (kp : KernelPair 𝒞 f) → Coequalizer 𝒞 (p₁ kp) (p₂ kp) 24 | pullback-of-regularepi-is-regularepi : {A B C : Obj} (f : B ⇒ A) {g : C ⇒ A} → RegularEpi 𝒞 f → (p : Pullback 𝒞 f g) → RegularEpi 𝒞 (p₂ p) 25 | -------------------------------------------------------------------------------- /src/Categories/Category/Restriction/Construction/Trivial.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Any category can be made into a trivial restriction category 4 | 5 | module Categories.Category.Restriction.Construction.Trivial where 6 | 7 | open import Level using (Level) 8 | 9 | open import Categories.Category.Core using (Category) 10 | open import Categories.Category.Restriction using (Restriction) 11 | open import Categories.Morphism.Reasoning.Core using (id-comm-sym) 12 | 13 | private 14 | variable 15 | o ℓ e : Level 16 | 17 | Trivial : (C : Category o ℓ e) → Restriction C 18 | Trivial C = record 19 | { _↓ = λ _ → id 20 | ; pidʳ = identityʳ 21 | ; ↓-comm = Equiv.refl 22 | ; ↓-denestʳ = Equiv.sym identity² 23 | ; ↓-skew-comm = id-comm-sym C 24 | ; ↓-cong = λ _ → Equiv.refl 25 | } 26 | where open Category C 27 | -------------------------------------------------------------------------------- /src/Categories/Category/Restriction/Properties/Poset.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core using (Category) 4 | open import Categories.Category.Restriction using (Restriction) 5 | open import Relation.Binary.Bundles using (Poset) 6 | open import Relation.Binary.Structures using (IsPartialOrder) 7 | 8 | import Categories.Morphism.Reasoning as MR 9 | 10 | -- Every restriction category induces a partial order (the restriction order) on hom-sets 11 | 12 | module Categories.Category.Restriction.Properties.Poset {o ℓ e} {𝒞 : Category o ℓ e} (R : Restriction 𝒞) where 13 | open Category 𝒞 14 | open Restriction R 15 | open Equiv 16 | open HomReasoning 17 | open MR 𝒞 using (pullʳ; pullˡ) 18 | 19 | poset : (A B : Obj) → Poset ℓ e e 20 | poset A B = record 21 | { Carrier = A ⇒ B 22 | ; _≈_ = _≈_ 23 | ; _≤_ = λ f g → f ≈ g ∘ f ↓ 24 | ; isPartialOrder = record 25 | { isPreorder = record 26 | { isEquivalence = equiv 27 | ; reflexive = λ {x} {y} x≈y → begin 28 | x ≈˘⟨ pidʳ ⟩ 29 | x ∘ x ↓ ≈⟨ x≈y ⟩∘⟨refl ⟩ 30 | y ∘ x ↓ ∎ 31 | ; trans = λ {i} {j} {k} i≈j∘i↓ j≈k∘j↓ → begin 32 | i ≈⟨ i≈j∘i↓ ⟩ 33 | j ∘ i ↓ ≈⟨ j≈k∘j↓ ⟩∘⟨refl ⟩ 34 | (k ∘ j ↓) ∘ i ↓ ≈⟨ pullʳ (sym ↓-denestʳ) ⟩ 35 | k ∘ (j ∘ i ↓) ↓ ≈⟨ refl⟩∘⟨ ↓-cong (sym i≈j∘i↓) ⟩ 36 | k ∘ i ↓ ∎ 37 | } 38 | ; antisym = λ {i} {j} i≈j∘i↓ j≈i∘j↓ → begin 39 | i ≈⟨ i≈j∘i↓ ⟩ 40 | j ∘ i ↓ ≈⟨ j≈i∘j↓ ⟩∘⟨refl ⟩ 41 | (i ∘ j ↓) ∘ i ↓ ≈⟨ pullʳ ↓-comm ⟩ 42 | i ∘ i ↓ ∘ j ↓ ≈⟨ pullˡ pidʳ ⟩ 43 | i ∘ j ↓ ≈˘⟨ j≈i∘j↓ ⟩ 44 | j ∎ 45 | } 46 | } 47 | -------------------------------------------------------------------------------- /src/Categories/Category/Site.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Site {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | open import Data.Product using (Σ; _,_; ∃₂) 9 | 10 | open Category C 11 | 12 | private 13 | variable 14 | X Y Z : Obj 15 | 16 | record Coverage {i} j {I : Obj → Set i} 17 | (covering₀ : ∀ {X} → I X → Obj) 18 | (covering₁ : ∀ {X} (i : I X) → covering₀ i ⇒ X) : Set (i ⊔ suc j ⊔ o ⊔ ℓ ⊔ e) where 19 | field 20 | J : ∀ (g : Y ⇒ Z) → Set j 21 | universal₀ : ∀ {g : Y ⇒ Z} → J g → Obj 22 | universal₁ : ∀ {g : Y ⇒ Z} (j : J g) → universal₀ j ⇒ Y 23 | commute : ∀ {g : Y ⇒ Z} (j : J g) → ∃₂ (λ i k → g ∘ universal₁ j ≈ covering₁ i ∘ k) 24 | 25 | record Site i j : Set (suc i ⊔ suc j ⊔ o ⊔ ℓ ⊔ e) where 26 | field 27 | I : Obj → Set i 28 | covering₀ : ∀ {X} → I X → Obj 29 | covering₁ : ∀ {X} (i : I X) → covering₀ i ⇒ X 30 | coverage : Coverage j covering₀ covering₁ 31 | 32 | module coverage = Coverage coverage 33 | open coverage public 34 | -------------------------------------------------------------------------------- /src/Categories/Category/Slice.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category.Core 3 | 4 | -- slice category (https://ncatlab.org/nlab/show/over+category) 5 | -- TODO: Forgetful Functor from Slice to 𝒞 6 | module Categories.Category.Slice {o ℓ e} (𝒞 : Category o ℓ e) where 7 | 8 | open Category 𝒞 9 | open HomReasoning 10 | open Equiv 11 | 12 | open import Level 13 | open import Function.Base using (_$_) 14 | open import Relation.Binary.Core using (Rel) 15 | 16 | open import Categories.Morphism.Reasoning 𝒞 17 | 18 | record SliceObj (X : Obj) : Set (o ⊔ ℓ) where 19 | constructor sliceobj 20 | field 21 | {Y} : Obj 22 | arr : Y ⇒ X 23 | 24 | private 25 | variable 26 | A : Obj 27 | X Y Z : SliceObj A 28 | 29 | record Slice⇒ {A : Obj} (X Y : SliceObj A) : Set (ℓ ⊔ e) where 30 | constructor slicearr 31 | private 32 | module X = SliceObj X 33 | module Y = SliceObj Y 34 | field 35 | {h} : X.Y ⇒ Y.Y 36 | △ : Y.arr ∘ h ≈ X.arr 37 | 38 | Slice : Obj → Category _ _ _ 39 | Slice A = record 40 | { Obj = SliceObj A 41 | ; _⇒_ = Slice⇒ 42 | ; _≈_ = λ where 43 | (slicearr {f} _) (slicearr {g} _) → f ≈ g 44 | ; id = slicearr identityʳ 45 | ; _∘_ = _∘′_ 46 | ; assoc = assoc 47 | ; sym-assoc = sym-assoc 48 | ; identityˡ = identityˡ 49 | ; identityʳ = identityʳ 50 | ; identity² = identity² 51 | ; equiv = record -- must be expanded to get levels to work out 52 | { refl = refl 53 | ; sym = sym 54 | ; trans = trans 55 | } 56 | ; ∘-resp-≈ = ∘-resp-≈ 57 | } 58 | where _∘′_ : Slice⇒ Y Z → Slice⇒ X Y → Slice⇒ X Z 59 | _∘′_ {Y = sliceobj y} {Z = sliceobj z} {X = sliceobj x} (slicearr {g} △) (slicearr {f} △′) = slicearr $ begin 60 | z ∘ g ∘ f ≈⟨ pullˡ △ ⟩ 61 | y ∘ f ≈⟨ △′ ⟩ 62 | x ∎ 63 | -------------------------------------------------------------------------------- /src/Categories/Category/Species.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Species where 3 | 4 | -- The Category of Species, as the Functor category from Core (FinSetoids) to Setoids. 5 | -- Setoids used here because that's what fits best in this setting. 6 | -- The constructions of the theory of Species are in Species.Construction 7 | 8 | open import Level 9 | 10 | open import Categories.Category.Core using (Category) 11 | open import Categories.Category.Construction.Functors 12 | open import Categories.Category.Construction.Core using (Core) 13 | open import Categories.Category.Instance.FinSetoids using (FinSetoids) 14 | open import Categories.Category.Instance.Setoids using (Setoids) 15 | 16 | private 17 | variable 18 | o ℓ o′ ℓ′ : Level 19 | 20 | -- note how Species, as a category, raises levels. 21 | Species : (o ℓ o′ ℓ′ : Level) → Category (suc (o ⊔ ℓ ⊔ o′ ⊔ ℓ′)) (suc (o ⊔ ℓ) ⊔ (o′ ⊔ ℓ′)) (suc (o ⊔ ℓ) ⊔ o′ ⊔ ℓ′) 22 | Species o ℓ o′ ℓ′ = Functors (Core (FinSetoids o ℓ)) (Setoids o′ ℓ′) 23 | -------------------------------------------------------------------------------- /src/Categories/Category/Topos.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Category.Topos {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Category.CartesianClosed C 10 | open import Categories.Category.Complete.Finitely C 11 | open import Categories.Diagram.Equalizer C 12 | open import Categories.Diagram.SubobjectClassifier C 13 | 14 | import Categories.Category.Complete.Finitely.Properties as Fₚ 15 | 16 | open Category C 17 | 18 | record Topos : Set (levelOfTerm C) where 19 | field 20 | cartesianClosed : CartesianClosed 21 | subobjectClassifier : SubobjectClassifier 22 | equalizer : ∀ {A B} (f g : A ⇒ B) → Equalizer f g 23 | 24 | finitelyComplete : FinitelyComplete 25 | finitelyComplete = record 26 | { cartesian = CartesianClosed.cartesian cartesianClosed 27 | ; equalizer = equalizer 28 | } 29 | 30 | open FinitelyComplete finitelyComplete using (module equalizer; pullback) public 31 | 32 | open Fₚ finitelyComplete using (finiteLimit) public 33 | -------------------------------------------------------------------------------- /src/Categories/Category/Unbundled.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Unbundled where 3 | 4 | -- This is basically identical to Category, except that the 5 | -- Obj type is a parameter rather than a field. 6 | 7 | open import Level 8 | open import Function.Base using (flip) 9 | 10 | open import Relation.Binary using (Rel; IsEquivalence) 11 | 12 | record Category {o : Level} (Obj : Set o) (ℓ e : Level) : Set (suc (o ⊔ ℓ ⊔ e)) where 13 | eta-equality 14 | infix 4 _≈_ _⇒_ 15 | infixr 9 _∘_ 16 | 17 | field 18 | _⇒_ : Rel Obj ℓ 19 | _≈_ : ∀ {A B} → Rel (A ⇒ B) e 20 | 21 | id : ∀ {A} → (A ⇒ A) 22 | _∘_ : ∀ {A B C} → (B ⇒ C) → (A ⇒ B) → (A ⇒ C) 23 | 24 | field 25 | assoc : ∀ {A B C D} {f : A ⇒ B} {g : B ⇒ C} {h : C ⇒ D} → (h ∘ g) ∘ f ≈ h ∘ (g ∘ f) 26 | -- We add a symmetric proof of associativity so that the opposite category of the 27 | -- opposite category is definitionally equal to the original category. See how 28 | -- `op` is implemented. 29 | sym-assoc : ∀ {A B C D} {f : A ⇒ B} {g : B ⇒ C} {h : C ⇒ D} → h ∘ (g ∘ f) ≈ (h ∘ g) ∘ f 30 | identityˡ : ∀ {A B} {f : A ⇒ B} → id ∘ f ≈ f 31 | identityʳ : ∀ {A B} {f : A ⇒ B} → f ∘ id ≈ f 32 | -- We add a proof of "neutral" identity proof, in order to ensure the opposite of 33 | -- constant functor is definitionally equal to itself. 34 | identity² : ∀ {A} → id ∘ id {A} ≈ id {A} 35 | equiv : ∀ {A B} → IsEquivalence (_≈_ {A} {B}) 36 | ∘-resp-≈ : ∀ {A B C} {f h : B ⇒ C} {g i : A ⇒ B} → f ≈ h → g ≈ i → f ∘ g ≈ h ∘ i 37 | -------------------------------------------------------------------------------- /src/Categories/Category/Unbundled/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Category.Unbundled.Properties where 3 | 4 | -- The Obj-unbundled Category is equivalent (as a type) to the 5 | -- usual kind. Quite straightforward and because of η, the proofs are just refl. 6 | 7 | open import Data.Product using (Σ; _,_) 8 | open import Level 9 | open import Function.Bundles using (_↔_; mk↔ₛ′) 10 | open import Relation.Binary.PropositionalEquality using (refl) 11 | 12 | open import Categories.Category.Core using (Category) 13 | open import Categories.Category.Unbundled renaming (Category to Unb-Cat) 14 | 15 | private 16 | variable 17 | o ℓ e : Level 18 | 19 | unpack : Category o ℓ e → Σ (Set o) (λ Obj → Unb-Cat Obj ℓ e) 20 | unpack C = C.Obj , record { C } 21 | where module C = Category C 22 | 23 | unpack′ : (C : Category o ℓ e) → Unb-Cat (Category.Obj C) ℓ e 24 | unpack′ C = record { C } 25 | where module C = Category C 26 | 27 | pack : Σ (Set o) (λ Obj → Unb-Cat Obj ℓ e) → Category o ℓ e 28 | pack (o , uc) = record { Obj = o; UC } 29 | where module UC = Unb-Cat uc 30 | 31 | pack′ : {Obj : Set o} → Unb-Cat Obj ℓ e → Category o ℓ e 32 | pack′ {Obj = o} uc = record { Obj = o; UC } 33 | where module UC = Unb-Cat uc 34 | 35 | equiv : (Category o ℓ e) ↔ (Σ (Set o) (λ Obj → Unb-Cat Obj ℓ e)) 36 | equiv = mk↔ₛ′ unpack pack (λ _ → refl) λ _ → refl 37 | -------------------------------------------------------------------------------- /src/Categories/Comonad.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Comonad where 3 | 4 | open import Level 5 | 6 | open import Categories.Category using (Category) 7 | open import Categories.Functor using (Functor; Endofunctor; _∘F_) renaming (id to idF) 8 | open import Categories.NaturalTransformation renaming (id to idN) 9 | open import Categories.NaturalTransformation.NaturalIsomorphism hiding (_≃_) 10 | open import Categories.NaturalTransformation.Equivalence 11 | open NaturalIsomorphism 12 | 13 | record Comonad {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 14 | field 15 | F : Endofunctor C 16 | ε : NaturalTransformation F idF 17 | δ : NaturalTransformation F (F ∘F F) 18 | 19 | module F = Functor F 20 | module ε = NaturalTransformation ε 21 | module δ = NaturalTransformation δ 22 | 23 | open Category C 24 | open Functor F 25 | 26 | field 27 | assoc : ∀ {X : Obj} → δ.η (F₀ X) ∘ δ.η X ≈ F₁ (δ.η X) ∘ δ.η X 28 | sym-assoc : ∀ {X : Obj} → F₁ (δ.η X) ∘ δ.η X ≈ δ.η (F₀ X) ∘ δ.η X 29 | identityˡ : ∀ {X : Obj} → F₁ (ε.η X) ∘ δ.η X ≈ id 30 | identityʳ : ∀ {X : Obj} → ε.η (F₀ X) ∘ δ.η X ≈ id 31 | 32 | module _ {o ℓ e} (C : Category o ℓ e) where 33 | open Comonad 34 | open Category C hiding (id) 35 | id : Comonad C 36 | id .F = idF 37 | id .ε = idN 38 | id .δ = unitor² .F⇐G 39 | id .assoc = Equiv.refl 40 | id .sym-assoc = Equiv.refl 41 | id .identityˡ = identity² 42 | id .identityʳ = identity² 43 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Cocone.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category 3 | open import Categories.Functor hiding (id) 4 | 5 | -- Cocone over a Functor F (from shape category J into category C) 6 | 7 | module Categories.Diagram.Cocone 8 | {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} (F : Functor J C) where 9 | 10 | open Category C 11 | open Functor F 12 | 13 | open import Level 14 | 15 | record Coapex (N : Obj) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 16 | field 17 | ψ : (X : Category.Obj J) → F₀ X ⇒ N 18 | commute : ∀ {X Y} (f : J [ X , Y ]) → ψ Y ∘ F₁ f ≈ ψ X 19 | 20 | record Cocone : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′) where 21 | field 22 | {N} : Obj 23 | coapex : Coapex N 24 | 25 | open Coapex coapex public 26 | 27 | open Coapex 28 | open Cocone 29 | 30 | record Cocone⇒ (c c′ : Cocone) : Set (ℓ ⊔ e ⊔ o′) where 31 | field 32 | arr : N c ⇒ N c′ 33 | commute : ∀ {X} → arr ∘ ψ c X ≈ ψ c′ X 34 | 35 | open Cocone⇒ 36 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Coend.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core using (Category) 4 | open import Categories.Functor.Bifunctor using (Bifunctor) 5 | 6 | module Categories.Diagram.Coend {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} 7 | (F : Bifunctor (Category.op C) C D) where 8 | 9 | private 10 | module C = Category C 11 | module D = Category D 12 | open D 13 | open HomReasoning 14 | open Equiv 15 | variable 16 | A B : Obj 17 | f g : A ⇒ B 18 | 19 | open import Level 20 | 21 | open import Categories.Diagram.Cowedge F 22 | open import Categories.Functor 23 | open import Categories.Functor.Construction.Constant 24 | open import Categories.NaturalTransformation.Dinatural 25 | open import Categories.Morphism.Reasoning D 26 | 27 | open Functor F 28 | 29 | record Coend : Set (levelOfTerm F) where 30 | field 31 | cowedge : Cowedge 32 | 33 | module cowedge = Cowedge cowedge 34 | open cowedge public 35 | open Cowedge 36 | 37 | field 38 | factor : (W : Cowedge) → cowedge.E ⇒ E W 39 | universal : ∀ {W : Cowedge} {A} → factor W ∘ cowedge.dinatural.α A ≈ dinatural.α W A 40 | unique : ∀ {W : Cowedge} {g : cowedge.E ⇒ E W} → (∀ {A} → g ∘ cowedge.dinatural.α A ≈ dinatural.α W A) → factor W ≈ g 41 | 42 | η-id : factor cowedge ≈ D.id 43 | η-id = unique identityˡ 44 | 45 | unique′ :(∀ {A} → f ∘ cowedge.dinatural.α A ≈ g ∘ cowedge.dinatural.α A) → f ≈ g 46 | unique′ {f = f} {g = g} eq = ⟺ (unique {W = Cowedge-∘ cowedge f} refl) ○ unique (⟺ eq) 47 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Colimit/DualProperties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Functor hiding (id) 5 | 6 | module Categories.Diagram.Colimit.DualProperties 7 | {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} where 8 | 9 | open import Categories.Category.Instance.Setoids 10 | open import Categories.Diagram.Duality C 11 | open import Categories.Functor.Hom 12 | open import Categories.Category.Construction.Cocones as Coc 13 | open import Categories.Diagram.Limit as Lim 14 | open import Categories.Diagram.Limit.Properties 15 | open import Categories.Diagram.Colimit as Col 16 | open import Categories.NaturalTransformation 17 | open import Categories.NaturalTransformation.NaturalIsomorphism 18 | open import Categories.Morphism C 19 | open import Categories.Morphism.Duality C 20 | 21 | private 22 | module C = Category C 23 | module J = Category J 24 | open C 25 | open Hom C 26 | open HomReasoning 27 | 28 | -- natural isomorphisms respects limits 29 | module _ {F G : Functor J C} (F≃G : F ≃ G) where 30 | private 31 | module F = Functor F 32 | module G = Functor G 33 | open NaturalIsomorphism F≃G 34 | 35 | ≃-resp-colim : Colimit F → Colimit G 36 | ≃-resp-colim Cf = coLimit⇒Colimit (≃-resp-lim op′ (Colimit⇒coLimit Cf)) 37 | 38 | ≃⇒Cocone⇒ : ∀ (Cf : Colimit F) (Cg : Colimit G) → Cocones G [ Colimit.colimit Cg , Colimit.colimit (≃-resp-colim Cf) ] 39 | ≃⇒Cocone⇒ Cf Cg = coCone⇒⇒Cocone⇒ (≃⇒Cone⇒ op′ (Colimit⇒coLimit Cf) (Colimit⇒coLimit Cg)) 40 | 41 | ≃⇒colim≅ : ∀ {F G : Functor J C} (F≃G : F ≃ G) (Cf : Colimit F) (Cg : Colimit G) → Colimit.coapex Cf ≅ Colimit.coapex Cg 42 | ≃⇒colim≅ F≃G Cf Cg = op-≅⇒≅ (≃⇒lim≅ (NaturalIsomorphism.op′ F≃G) (Colimit⇒coLimit Cf) (Colimit⇒coLimit Cg)) 43 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Colimit/Lan.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Diagram.Colimit.Lan where 4 | 5 | open import Level 6 | open import Categories.Category 7 | open import Categories.Category.Cocomplete 8 | open import Categories.Diagram.Duality 9 | open import Categories.Diagram.Limit.Ran 10 | open import Categories.Functor 11 | open import Categories.Kan 12 | open import Categories.Kan.Duality 13 | 14 | private 15 | variable 16 | o ℓ e : Level 17 | C D E : Category o ℓ e 18 | 19 | module _ {o ℓ e o′ ℓ′ e′} {C : Category o′ ℓ′ e′} {D : Category o ℓ e} 20 | (F : Functor C D) (X : Functor C E) (Com : Cocomplete (o′ ⊔ ℓ) (ℓ′ ⊔ e) e′ E) where 21 | private 22 | module F = Functor F 23 | module X = Functor X 24 | 25 | colimit-is-lan : Lan F X 26 | colimit-is-lan = coRan⇒Lan (limit-is-ran F.op X.op λ G → Colimit⇒coLimit E (Com (Functor.op G))) 27 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Cone.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category 3 | open import Categories.Functor hiding (id) 4 | 5 | -- Cone over a Functor F (from shape category J into category C) 6 | 7 | module Categories.Diagram.Cone 8 | {o ℓ e} {o′ ℓ′ e′} {C : Category o ℓ e} {J : Category o′ ℓ′ e′} (F : Functor J C) where 9 | 10 | open Category C 11 | open Functor F 12 | 13 | open import Level 14 | 15 | record Apex (N : Obj) : Set (ℓ′ ⊔ o′ ⊔ e ⊔ ℓ) where 16 | field 17 | ψ : (X : Category.Obj J) → (N ⇒ F₀ X) 18 | commute : ∀ {X Y} (f : J [ X , Y ]) → F₁ f ∘ ψ X ≈ ψ Y 19 | 20 | record Cone : Set (o ⊔ ℓ′ ⊔ o′ ⊔ e ⊔ ℓ) where 21 | field 22 | {N} : Obj 23 | apex : Apex N 24 | 25 | open Apex apex public 26 | 27 | open Apex 28 | open Cone 29 | 30 | record Cone⇒ (c c′ : Cone) : Set (ℓ ⊔ e ⊔ o′) where 31 | field 32 | arr : N c ⇒ N c′ 33 | commute : ∀ {X} → ψ c′ X ∘ arr ≈ ψ c X 34 | 35 | open Cone⇒ 36 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Empty.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level 4 | 5 | open import Categories.Category.Core using (Category) 6 | 7 | module Categories.Diagram.Empty {o ℓ e} (C : Category o ℓ e) (o′ ℓ′ e′ : Level) where 8 | 9 | open import Categories.Category.Instance.Zero {o′} {ℓ′} {e′} using (Zero) 10 | open import Categories.Functor.Core using (Functor) 11 | 12 | open Functor 13 | 14 | empty : Functor Zero C 15 | empty .F₀ () 16 | -------------------------------------------------------------------------------- /src/Categories/Diagram/End.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Functor.Bifunctor 5 | 6 | module Categories.Diagram.End {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} 7 | (F : Bifunctor (Category.op C) C D) where 8 | 9 | private 10 | module D = Category D 11 | open D 12 | open HomReasoning 13 | open Equiv 14 | variable 15 | A B : Obj 16 | f g : A ⇒ B 17 | 18 | open import Level 19 | 20 | open import Categories.Diagram.Wedge F 21 | open import Categories.NaturalTransformation.Dinatural 22 | 23 | record End : Set (levelOfTerm F) where 24 | field 25 | wedge : Wedge 26 | 27 | module wedge = Wedge wedge 28 | open wedge public 29 | open Wedge 30 | 31 | field 32 | factor : (W : Wedge) → E W ⇒ wedge.E 33 | universal : ∀ {W : Wedge} {A} → wedge.dinatural.α A ∘ factor W ≈ dinatural.α W A 34 | unique : ∀ {W : Wedge} {g : E W ⇒ wedge.E} → (∀ {A} → wedge.dinatural.α A ∘ g ≈ dinatural.α W A) → factor W ≈ g 35 | 36 | η-id : factor wedge ≈ D.id 37 | η-id = unique identityʳ 38 | 39 | unique′ :(∀ {A} → wedge.dinatural.α A ∘ f ≈ wedge.dinatural.α A ∘ g) → f ≈ g 40 | unique′ {f = f} {g = g} eq = ⟺ (unique {W = Wedge-∘ wedge f} refl) ○ unique (⟺ eq) 41 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Equalizer/Indexed.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | -- this module characterizes a category of all equalizer indexed by I. 6 | -- this notion formalizes a category with all equalizer up to certain cardinal. 7 | module Categories.Diagram.Equalizer.Indexed {o ℓ e} (C : Category o ℓ e) where 8 | 9 | open import Level 10 | 11 | open Category C 12 | 13 | record IndexedEqualizerOf {i} {I : Set i} {A B : Obj} (M : I → A ⇒ B) : Set (i ⊔ o ⊔ e ⊔ ℓ) where 14 | field 15 | E : Obj 16 | arr : E ⇒ A 17 | 18 | -- a reference morphism 19 | ref : E ⇒ B 20 | equality : ∀ i → M i ∘ arr ≈ ref 21 | equalize : ∀ {X} (h : X ⇒ A) (r : X ⇒ B) → (∀ i → M i ∘ h ≈ r) → X ⇒ E 22 | universal : ∀ {X} (h : X ⇒ A) (r : X ⇒ B) (eq : ∀ i → M i ∘ h ≈ r) → h ≈ arr ∘ equalize h r eq 23 | unique : ∀ {X} {l : X ⇒ E} (h : X ⇒ A) (r : X ⇒ B) (eq : ∀ i → M i ∘ h ≈ r) → h ≈ arr ∘ l → l ≈ equalize h r eq 24 | 25 | record IndexedEqualizer {i} (I : Set i) : Set (i ⊔ o ⊔ e ⊔ ℓ) where 26 | field 27 | A B : Obj 28 | M : I → A ⇒ B 29 | equalizerOf : IndexedEqualizerOf M 30 | 31 | open IndexedEqualizerOf equalizerOf public 32 | 33 | AllEqualizers : ∀ i → Set (o ⊔ ℓ ⊔ e ⊔ suc i) 34 | AllEqualizers i = (I : Set i) → IndexedEqualizer I 35 | 36 | AllEqualizersOf : ∀ i → Set (o ⊔ ℓ ⊔ e ⊔ suc i) 37 | AllEqualizersOf i = ∀ {I : Set i} {A B : Obj} (M : I → A ⇒ B) → IndexedEqualizerOf M 38 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Finite.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Diagram.Finite where 4 | 5 | open import Level 6 | 7 | open import Categories.Category 8 | open import Categories.Category.Finite renaming (Finite to FiniteC) 9 | open import Categories.Functor 10 | 11 | private 12 | variable 13 | o ℓ e : Level 14 | J C : Category o ℓ e 15 | 16 | record Finite (F : Functor J C) : Set (levelOfTerm F) where 17 | field 18 | finite : FiniteC J 19 | 20 | open FiniteC finite public 21 | -------------------------------------------------------------------------------- /src/Categories/Diagram/KernelPair.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core 4 | 5 | -- Kernel Pair - a Pullback of a morphism along itself 6 | -- https://ncatlab.org/nlab/show/kernel+pair 7 | module Categories.Diagram.KernelPair {o ℓ e} (𝒞 : Category o ℓ e) where 8 | 9 | open import Level 10 | 11 | open import Categories.Diagram.Pullback 𝒞 12 | 13 | open Category 𝒞 14 | 15 | private 16 | variable 17 | A B : Obj 18 | 19 | -- Make it a pure synonym 20 | KernelPair : (f : A ⇒ B) → Set (o ⊔ ℓ ⊔ e) 21 | KernelPair f = Pullback f f 22 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Pushout.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core using (Category) 4 | 5 | module Categories.Diagram.Pushout {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open Category C 8 | open HomReasoning 9 | 10 | open import Level 11 | 12 | private 13 | variable 14 | A B E X Y Z : Obj 15 | f g h j : A ⇒ B 16 | 17 | record IsPushout {Q : Obj} (f : X ⇒ Y) (g : X ⇒ Z) (i₁ : Y ⇒ Q) (i₂ : Z ⇒ Q) : Set (o ⊔ ℓ ⊔ e) where 18 | 19 | field 20 | commute : i₁ ∘ f ≈ i₂ ∘ g 21 | universal : {h₁ : Y ⇒ B} {h₂ : Z ⇒ B} → h₁ ∘ f ≈ h₂ ∘ g → Q ⇒ B 22 | universal∘i₁≈h₁ : {h₁ : Y ⇒ E} {h₂ : Z ⇒ E} {eq : h₁ ∘ f ≈ h₂ ∘ g} → 23 | universal eq ∘ i₁ ≈ h₁ 24 | universal∘i₂≈h₂ : {h₁ : Y ⇒ E} {h₂ : Z ⇒ E} {eq : h₁ ∘ f ≈ h₂ ∘ g} → 25 | universal eq ∘ i₂ ≈ h₂ 26 | unique-diagram : h ∘ i₁ ≈ j ∘ i₁ → h ∘ i₂ ≈ j ∘ i₂ → h ≈ j 27 | 28 | unique : {h₁ : Y ⇒ E} {h₂ : Z ⇒ E} {eq : h₁ ∘ f ≈ h₂ ∘ g} → 29 | j ∘ i₁ ≈ h₁ → j ∘ i₂ ≈ h₂ → j ≈ universal eq 30 | unique j∘i₁≈h₁ j∘i₂≈h₂ = 31 | unique-diagram 32 | (j∘i₁≈h₁ ○ ⟺ universal∘i₁≈h₁) 33 | (j∘i₂≈h₂ ○ ⟺ universal∘i₂≈h₂) 34 | 35 | record Pushout (f : X ⇒ Y) (g : X ⇒ Z) : Set (o ⊔ ℓ ⊔ e) where 36 | 37 | field 38 | {Q} : Obj 39 | i₁ : Y ⇒ Q 40 | i₂ : Z ⇒ Q 41 | isPushout : IsPushout f g i₁ i₂ 42 | 43 | open IsPushout isPushout public 44 | -------------------------------------------------------------------------------- /src/Categories/Diagram/ReflexivePair.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core 4 | 5 | -- Reflexive pairs and reflexive coequalizers 6 | -- https://ncatlab.org/nlab/show/reflexive+coequalizer 7 | module Categories.Diagram.ReflexivePair {o ℓ e} (𝒞 : Category o ℓ e) where 8 | 9 | open import Level 10 | 11 | open import Categories.Diagram.Coequalizer 𝒞 12 | 13 | open Category 𝒞 14 | open HomReasoning 15 | open Equiv 16 | 17 | private 18 | variable 19 | A B R : Obj 20 | 21 | -- A reflexive pair can be thought of as a vast generalization of a reflexive relation. 22 | -- To see this, consider the case in 'Set' where 'R ⊆ A × A', and 'f' and 'g' are the projections. 23 | -- Then, our morphism 's' would have to look something like the diagonal morphism due to the 24 | -- restriction it is a section of both 'f' and 'g'. 25 | record IsReflexivePair (f g : R ⇒ A) (s : A ⇒ R) : Set e where 26 | field 27 | sectionₗ : f ∘ s ≈ id 28 | sectionᵣ : g ∘ s ≈ id 29 | 30 | section : f ∘ s ≈ g ∘ s 31 | section = sectionₗ ○ ⟺ sectionᵣ 32 | 33 | record ReflexivePair (f g : R ⇒ A) : Set (ℓ ⊔ e) where 34 | field 35 | s : A ⇒ R 36 | isReflexivePair : IsReflexivePair f g s 37 | 38 | open IsReflexivePair isReflexivePair public 39 | -------------------------------------------------------------------------------- /src/Categories/Diagram/SubobjectClassifier.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | -- the usual notion requires C to have finite limits. 6 | -- this definition is a generalization by omitting that requirement as the definition 7 | -- itself does not involve any limit. 8 | module Categories.Diagram.SubobjectClassifier {o ℓ e} (C : Category o ℓ e) where 9 | 10 | open Category C 11 | 12 | open import Level 13 | 14 | open import Categories.Object.Terminal C 15 | open import Categories.Morphism C 16 | open import Categories.Diagram.Pullback C 17 | 18 | record SubobjectClassifier : Set (o ⊔ ℓ ⊔ e) where 19 | field 20 | Ω : Obj 21 | terminal : Terminal 22 | 23 | module terminal = Terminal terminal 24 | open terminal 25 | 26 | field 27 | true : ⊤ ⇒ Ω 28 | universal : ∀ {X Y} {f : X ⇒ Y} → Mono f → Y ⇒ Ω 29 | pullback : ∀ {X Y} {f : X ⇒ Y} {mono : Mono f} → IsPullback f ! (universal mono) true 30 | unique : ∀ {X Y} {f : X ⇒ Y} {g} {mono : Mono f} → IsPullback f ! g true → universal mono ≈ g 31 | -------------------------------------------------------------------------------- /src/Categories/Diagram/Wedge.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Functor.Bifunctor 5 | 6 | module Categories.Diagram.Wedge {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} 7 | (F : Bifunctor (Category.op C) C D) where 8 | 9 | private 10 | module C = Category C 11 | module D = Category D 12 | open D 13 | open HomReasoning 14 | variable 15 | A : Obj 16 | 17 | open import Level 18 | 19 | open import Categories.Functor hiding (id) 20 | open import Categories.Functor.Construction.Constant 21 | open import Categories.NaturalTransformation.Dinatural 22 | 23 | open Functor F 24 | 25 | record Wedge : Set (levelOfTerm F) where 26 | field 27 | E : Obj 28 | dinatural : DinaturalTransformation (const E) F 29 | 30 | module dinatural = DinaturalTransformation dinatural 31 | 32 | Wedge-∘ : (W : Wedge) → A ⇒ Wedge.E W → Wedge 33 | Wedge-∘ {A = A} W f = record 34 | { E = A 35 | ; dinatural = extranaturalʳ (λ X → dinatural.α X ∘ f) 36 | (sym-assoc ○ ∘-resp-≈ˡ (extranatural-commʳ dinatural) ○ assoc) 37 | } 38 | where open Wedge W 39 | 40 | record Wedge-Morphism (W₁ W₂ : Wedge) : Set (levelOfTerm F) where 41 | private 42 | module W₁ = Wedge W₁ 43 | module W₂ = Wedge W₂ 44 | open DinaturalTransformation 45 | field 46 | u : W₁.E ⇒ W₂.E 47 | commute : ∀ {C} → W₂.dinatural.α C ∘ u ≈ W₁.dinatural.α C 48 | 49 | Wedge-id : ∀ {W} → Wedge-Morphism W W 50 | Wedge-id {W} = record { u = D.id ; commute = D.identityʳ } 51 | 52 | Wedge-Morphism-∘ : {A B C : Wedge} → Wedge-Morphism B C → Wedge-Morphism A B → Wedge-Morphism A C 53 | Wedge-Morphism-∘ M N = record { u = u M ∘ u N ; commute = sym-assoc ○ (∘-resp-≈ˡ (commute M) ○ commute N) } 54 | where 55 | open Wedge-Morphism 56 | open HomReasoning 57 | -------------------------------------------------------------------------------- /src/Categories/Double.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Double where 3 | 4 | -- The main definitions are in: 5 | open import Categories.Double.Core public 6 | -------------------------------------------------------------------------------- /src/Categories/Enriched/Over/One.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Enriched.Over.One where 4 | 5 | open import Level 6 | 7 | open import Categories.Category.Monoidal.Instance.One 8 | open import Categories.Enriched.Category 9 | 10 | TruthValue : (o ℓ e t : Level) → Set (o ⊔ ℓ ⊔ e ⊔ suc t) 11 | TruthValue o ℓ e t = Category (One-Monoidal {o} {ℓ} {e}) t 12 | -------------------------------------------------------------------------------- /src/Categories/Functor/Algebra.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Functor.Algebra where 3 | 4 | -- Algebra for a Functor 5 | 6 | open import Level 7 | open import Function using (_$_) 8 | 9 | open import Categories.Category using (Category) 10 | open import Categories.Functor using (Functor; Endofunctor) 11 | 12 | private 13 | variable 14 | o ℓ e : Level 15 | 16 | module _ {C : Category o ℓ e} where 17 | 18 | module _ (F : Endofunctor C) where 19 | open Category C 20 | open Functor F 21 | F-Algebra-on : Obj → Set ℓ 22 | F-Algebra-on A = F₀ A ⇒ A 23 | 24 | record F-Algebra (F : Endofunctor C) : Set (o ⊔ ℓ) where 25 | open Category C 26 | field 27 | A : Obj 28 | α : F-Algebra-on F A 29 | 30 | to-Algebra : {A : Category.Obj C} → {F : Endofunctor C} → (F-Algebra-on F A) → (F-Algebra F) 31 | to-Algebra {A = A} α = record {A = A; α = α} 32 | 33 | open F-Algebra 34 | 35 | -- Given an F-Algebra F, one can apply F to it to obtain an new 'iterated' F-Algebra 36 | iterate : {F : Endofunctor C} → F-Algebra F → F-Algebra F 37 | iterate {F} c = record { A = Functor.F₀ F $ A c ; α = Functor.F₁ F $ α c } 38 | 39 | module _ {F : Endofunctor C} (X Y : F-Algebra F) where 40 | open Category C using (_⇒_; _∘_; _≈_) 41 | open Functor F using (F₁) 42 | private 43 | module X = F-Algebra X 44 | module Y = F-Algebra Y 45 | is-F-Algebra-Morphism : (X.A ⇒ Y.A) → (Set e) 46 | is-F-Algebra-Morphism f = f ∘ X.α ≈ Y.α ∘ F₁ f 47 | 48 | record F-Algebra-Morphism : Set (ℓ ⊔ e) where 49 | field 50 | f : X.A ⇒ Y.A 51 | commutes : is-F-Algebra-Morphism f 52 | -------------------------------------------------------------------------------- /src/Categories/Functor/Coalgebra.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Functor.Coalgebra where 3 | 4 | -- Co-algebras of a Functor 5 | open import Level 6 | open import Function using (_$_) 7 | 8 | open import Categories.Category using (Category) 9 | open import Categories.Functor using (Functor; Endofunctor) 10 | 11 | private 12 | variable 13 | o ℓ e : Level 14 | 15 | module _ {C : Category o ℓ e} where 16 | 17 | F-Coalgebra-on : Endofunctor C → C .Category.Obj → Set ℓ 18 | F-Coalgebra-on F A = A ⇒ F₀ A where open Category C; open Functor F 19 | 20 | record F-Coalgebra (F : Endofunctor C) : Set (o ⊔ ℓ) where 21 | open Category C 22 | field 23 | A : Obj 24 | α : F-Coalgebra-on F A 25 | 26 | to-Coalgebra : {A : Category.Obj C} → {F : Endofunctor C} → (F-Coalgebra-on F A) → (F-Coalgebra F) 27 | to-Coalgebra {A = A} α = record {A = A; α = α} 28 | 29 | open F-Coalgebra 30 | 31 | -- Given a F-Coalgebra F, one can apply F to it to obtain an new 'iterated' F-Coalgebra 32 | iterate : {F : Endofunctor C} → F-Coalgebra F → F-Coalgebra F 33 | iterate {F = F} F-alg = record { A = Functor.F₀ F $ A F-alg ; α = Functor.F₁ F $ α F-alg } 34 | 35 | module _ {F : Endofunctor C} (X Y : F-Coalgebra F) where 36 | open Category C using (_⇒_; _∘_; _≈_) 37 | open Functor F using (F₁) 38 | private 39 | module X = F-Coalgebra X 40 | module Y = F-Coalgebra Y 41 | is-F-Coalgebra-Morphism : (X.A ⇒ Y.A) → (Set e) 42 | is-F-Coalgebra-Morphism f = Y.α ∘ f ≈ F₁ f ∘ X.α 43 | 44 | record F-Coalgebra-Morphism : Set (ℓ ⊔ e) where 45 | field 46 | f : X.A ⇒ Y.A 47 | commutes : is-F-Coalgebra-Morphism f 48 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/Constant.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Functor.Construction.Constant where 3 | 4 | open import Level 5 | 6 | open import Categories.Category 7 | open import Categories.Category.Instance.One 8 | open import Categories.Category.Product 9 | open import Categories.Functor renaming (id to idF) 10 | open import Categories.NaturalTransformation using (NaturalTransformation) 11 | import Categories.Morphism.Reasoning as MR 12 | 13 | private 14 | variable 15 | o ℓ e : Level 16 | C D : Category o ℓ e 17 | 18 | const : (d : Category.Obj D) → Functor C D 19 | const {D = D} d = record 20 | { F₀ = λ _ → d 21 | ; F₁ = λ _ → id 22 | ; identity = refl 23 | ; homomorphism = sym identity² 24 | ; F-resp-≈ = λ _ → refl 25 | } 26 | where open Category D 27 | open Equiv 28 | 29 | const! : (d : Category.Obj D) → Functor (One {0ℓ} {0ℓ} {0ℓ}) D 30 | const! = const 31 | 32 | constˡ : (c : Category.Obj C) → Functor D (Product C D) 33 | constˡ c = const c ※ idF 34 | 35 | constʳ : (c : Category.Obj C) → Functor D (Product D C) 36 | constʳ c = idF ※ (const c) 37 | 38 | constNat : ∀ {A B} → Category._⇒_ D A B → NaturalTransformation (const {D = D} {C = C} A) (const B) 39 | constNat {D = D} f = record 40 | { η = λ _ → f 41 | ; commute = λ _ → MR.id-comm D 42 | ; sym-commute = λ _ → MR.id-comm-sym D 43 | } 44 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/Diagonal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Functor.Construction.Diagonal where 3 | 4 | -- A variety of Diagonal functors 5 | 6 | open import Level 7 | open import Data.Product using (_,_) 8 | 9 | open import Categories.Category 10 | open import Categories.Functor 11 | open import Categories.Category.Product 12 | open import Categories.Category.Construction.Functors 13 | open import Categories.Functor.Construction.Constant 14 | open import Categories.NaturalTransformation using (ntHelper) 15 | 16 | import Categories.Functor.Power as Power 17 | 18 | private 19 | variable 20 | o ℓ e o′ ℓ′ e′ : Level 21 | 22 | Δ : (C : Category o ℓ e) → Functor C (Product C C) 23 | Δ C = record 24 | { F₀ = λ x → x , x 25 | ; F₁ = λ f → f , f 26 | ; identity = refl , refl 27 | ; homomorphism = refl , refl 28 | ; F-resp-≈ = λ x → x , x 29 | } 30 | where 31 | open Category C 32 | open Equiv 33 | 34 | Δ′ : (I : Set) → (C : Category o ℓ e) → Functor C (Power.Exp C I) 35 | Δ′ I C = record 36 | { F₀ = λ x _ → x 37 | ; F₁ = λ f _ → f 38 | ; identity = λ _ → refl 39 | ; homomorphism = λ _ → refl 40 | ; F-resp-≈ = λ x _ → x 41 | } 42 | where open Category.Equiv C 43 | 44 | ΔF : {C : Category o ℓ e} (I : Category o′ ℓ′ e′) → Functor C (Functors I C) 45 | ΔF {C = C} I = record 46 | { F₀ = const 47 | ; F₁ = λ f → ntHelper record { η = λ _ → f; commute = λ _ → C.identityʳ ○ ⟺ C.identityˡ } 48 | ; identity = refl 49 | ; homomorphism = refl 50 | ; F-resp-≈ = λ x → x 51 | } 52 | where 53 | module C = Category C 54 | open C.Equiv 55 | open C.HomReasoning using (_○_; ⟺) 56 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/FromDiscrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category.Core using (Category) 3 | 4 | -- You can transform functions out of discrete 5 | -- categories into functors. 6 | module Categories.Functor.Construction.FromDiscrete {o ℓ e} (𝒞 : Category o ℓ e) where 7 | 8 | open import Relation.Binary.PropositionalEquality.Core as ≡ 9 | 10 | open import Categories.Category.Construction.StrictDiscrete using (Discrete) 11 | open import Categories.Functor.Core using (Functor) 12 | 13 | open Category 𝒞 14 | open Equiv 15 | 16 | FromDiscrete : ∀ {o} {A : Set o} → (A → Obj) → Functor (Discrete A) 𝒞 17 | FromDiscrete {o} {A = A} select = record 18 | { F₀ = select 19 | ; F₁ = λ { ≡.refl → id } 20 | ; identity = Equiv.refl 21 | ; homomorphism = λ { {_} {_} {_} {≡.refl} {≡.refl} → Equiv.sym identity² } 22 | ; F-resp-≈ = λ { ≡.refl → Equiv.refl } 23 | } 24 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/LiftSetoids.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Functor.Construction.LiftSetoids where 4 | 5 | open import Level using (Level; _⊔_; Lift; lift) 6 | open import Relation.Binary.Bundles using (Setoid) 7 | open import Function.Bundles using (Func; _⟨$⟩_) 8 | open import Function using (_$_) renaming (id to idf) 9 | 10 | open import Categories.Category.Instance.Setoids using (Setoids) 11 | open import Categories.Functor.Core using (Functor) 12 | 13 | open Func 14 | 15 | private 16 | variable 17 | c ℓ : Level 18 | 19 | -- Use pattern-matching (instead of explicit calls to lower) to minimize the 20 | -- number of needed parens, and also make it syntactically apparent that 21 | -- this is indeed just a Lift. 22 | LiftedSetoid : ∀ c′ ℓ′ → Setoid c ℓ → Setoid (c ⊔ c′) (ℓ ⊔ ℓ′) 23 | LiftedSetoid c′ ℓ′ S = record 24 | { Carrier = Lift c′ Carrier 25 | ; _≈_ = λ where (lift x) (lift y) → Lift ℓ′ $ x ≈ y 26 | ; isEquivalence = record 27 | { refl = lift refl 28 | ; sym = λ where (lift eq) → lift $ sym eq 29 | ; trans = λ where (lift eq) (lift eq′) → lift $ trans eq eq′ 30 | } 31 | } 32 | where open Setoid S 33 | 34 | LiftSetoids : ∀ c′ ℓ′ → Functor (Setoids c ℓ) (Setoids (c ⊔ c′) (ℓ ⊔ ℓ′)) 35 | LiftSetoids c′ ℓ′ = record 36 | { F₀ = LiftedSetoid c′ ℓ′ 37 | ; F₁ = λ f → record 38 | { to = λ where (lift x) → lift $ f ⟨$⟩ x 39 | ; cong = λ where (lift eq) → lift $ cong f eq 40 | } 41 | ; identity = λ {A} → lift $ refl A 42 | ; homomorphism = λ {_} {_} {Z} → lift $ refl Z 43 | ; F-resp-≈ = λ f≈g → lift f≈g 44 | } 45 | where 46 | open Setoid 47 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/ObjectRestriction.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- The category build from ObjectRestriction can be done so Functorially 4 | 5 | module Categories.Functor.Construction.ObjectRestriction where 6 | 7 | open import Level using (Level) 8 | open import Data.Product using (proj₁; _,_) 9 | open import Relation.Unary using (Pred) 10 | open import Function using () renaming (id to id→) 11 | 12 | open import Categories.Category.Core using (Category) 13 | open import Categories.Category.Construction.ObjectRestriction 14 | open import Categories.Functor.Core using (Functor) 15 | open import Categories.Functor.Properties using (Faithful; Full) 16 | 17 | private 18 | variable 19 | o ℓ e ℓ′ : Level 20 | C : Category o ℓ e 21 | 22 | RestrictionFunctor : {ℓ′ : Level} → (C : Category o ℓ e) → (E : Pred (Category.Obj C) ℓ′) → Functor (ObjectRestriction C E) C 23 | RestrictionFunctor C _ = record 24 | { F₀ = proj₁ 25 | ; F₁ = id→ 26 | ; identity = Equiv.refl 27 | ; homomorphism = Equiv.refl 28 | ; F-resp-≈ = id→ 29 | } 30 | where 31 | open Category C 32 | 33 | RF-Faithful : {E : Pred (Category.Obj C) ℓ′} → Faithful (RestrictionFunctor C E) 34 | RF-Faithful = id→ 35 | 36 | RF-Full : {E : Pred (Category.Obj C) ℓ′} → Full (RestrictionFunctor C E) 37 | RF-Full {C = C} f = f , Category.Equiv.refl C 38 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/SubCategory.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Functor.Construction.SubCategory {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Categories.Category.SubCategory C 8 | 9 | open Category C 10 | open Equiv 11 | 12 | open import Level 13 | open import Function.Base using () renaming (id to id→) 14 | open import Data.Product 15 | 16 | open import Categories.Functor using (Functor) 17 | 18 | private 19 | variable 20 | ℓ′ i : Level 21 | I : Set i 22 | U : I → Obj 23 | 24 | Sub : ∀ (sub : SubCat {i} {ℓ′} I) → Functor (SubCategory sub) C 25 | Sub (record {U = U}) = record 26 | { F₀ = U 27 | ; F₁ = proj₁ 28 | ; identity = refl 29 | ; homomorphism = refl 30 | ; F-resp-≈ = id→ 31 | } 32 | 33 | FullSub : Functor (FullSubCategory U) C 34 | FullSub {U = U} = record 35 | { F₀ = U 36 | ; F₁ = id→ 37 | ; identity = refl 38 | ; homomorphism = refl 39 | ; F-resp-≈ = id→ 40 | } 41 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/SubCategory/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Functor.Construction.SubCategory.Properties {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open Category C 8 | open Equiv 9 | 10 | open import Data.Product using (_,_) 11 | open import Level 12 | open import Function.Base using () renaming (id to id→) 13 | open import Function.Bundles using (Surjection) 14 | 15 | open import Categories.Category.SubCategory C 16 | open import Categories.Functor.Construction.SubCategory C 17 | open import Categories.Functor.Properties 18 | 19 | private 20 | variable 21 | ℓ′ i : Level 22 | I : Set i 23 | U : I → Obj 24 | 25 | SubFaithful : ∀ (sub : SubCat {i} {ℓ′} I) → Faithful (Sub sub) 26 | SubFaithful _ x≈y = x≈y 27 | 28 | FullSubFaithful : Faithful (FullSub {U = U}) 29 | FullSubFaithful = id→ 30 | 31 | FullSubFull : Full (FullSub {U = U}) 32 | FullSubFull f = f , refl 33 | -------------------------------------------------------------------------------- /src/Categories/Functor/Construction/Zero.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Functor.Construction.Zero where 3 | 4 | -- The Zero functor maps everything to the initial object of a 5 | -- category (when it exists). Note quite const. 6 | 7 | open import Level 8 | 9 | open import Categories.Category 10 | open import Categories.Functor using (Functor) 11 | open import Categories.Object.Initial 12 | 13 | private 14 | variable 15 | o ℓ e : Level 16 | C D : Category o ℓ e 17 | 18 | Zero : Initial D → Functor C D 19 | Zero {D = D} init = record 20 | { F₀ = λ _ → ⊥ 21 | ; F₁ = λ _ → id 22 | ; identity = Equiv.refl 23 | ; homomorphism = Equiv.sym identity² 24 | ; F-resp-≈ = λ _ → Equiv.refl 25 | } 26 | where 27 | open Initial init 28 | open Category D 29 | -------------------------------------------------------------------------------- /src/Categories/Functor/Core.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Functor.Core where 6 | 7 | open import Level 8 | open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) 9 | 10 | private 11 | variable 12 | o ℓ e o′ ℓ′ e′ : Level 13 | 14 | record Functor (C : Category o ℓ e) (D : Category o′ ℓ′ e′) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 15 | eta-equality 16 | private module C = Category C 17 | private module D = Category D 18 | 19 | field 20 | F₀ : C.Obj → D.Obj 21 | F₁ : ∀ {A B} (f : C [ A , B ]) → D [ F₀ A , F₀ B ] 22 | 23 | identity : ∀ {A} → D [ F₁ (C.id {A}) ≈ D.id ] 24 | homomorphism : ∀ {X Y Z} {f : C [ X , Y ]} {g : C [ Y , Z ]} → 25 | D [ F₁ (C [ g ∘ f ]) ≈ D [ F₁ g ∘ F₁ f ] ] 26 | F-resp-≈ : ∀ {A B} {f g : C [ A , B ]} → C [ f ≈ g ] → D [ F₁ f ≈ F₁ g ] 27 | 28 | -- nice shorthands 29 | ₀ = F₀ 30 | ₁ = F₁ 31 | 32 | op : Functor C.op D.op 33 | op = record 34 | { F₀ = F₀ 35 | ; F₁ = F₁ 36 | ; identity = identity 37 | ; homomorphism = homomorphism 38 | ; F-resp-≈ = F-resp-≈ 39 | } 40 | 41 | private 42 | op-involutive : {C : Category o ℓ e} {D : Category o′ ℓ′ e′} → (F : Functor C D) → Functor.op (Functor.op F) ≡ F 43 | op-involutive _ = ≡.refl 44 | -------------------------------------------------------------------------------- /src/Categories/Functor/Dagger.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --without-K #-} 2 | 3 | module Categories.Functor.Dagger where 4 | 5 | open import Categories.Category.Dagger using (DaggerCategory) 6 | open import Categories.Functor using (Functor; _∘F_) renaming (id to idF) 7 | 8 | open import Level using (Level; _⊔_) 9 | 10 | private 11 | variable 12 | o ℓ e o′ ℓ′ e′ o″ ℓ″ e″ : Level 13 | 14 | record DaggerFunctor (C : DaggerCategory o ℓ e) (D : DaggerCategory o′ ℓ′ e′) : Set (o ⊔ ℓ ⊔ e ⊔ o′ ⊔ ℓ′ ⊔ e′) where 15 | private 16 | module C = DaggerCategory C 17 | module D = DaggerCategory D 18 | field 19 | functor : Functor C.C D.C 20 | 21 | open Functor functor public 22 | 23 | field 24 | F-resp-† : ∀ {X Y} {f : X C.⇒ Y} → F₁ f D.† D.≈ F₁ (f C.†) 25 | 26 | id : ∀ {C : DaggerCategory o ℓ e} → DaggerFunctor C C 27 | id {C = C} = record 28 | { functor = idF 29 | ; F-resp-† = DaggerCategory.Equiv.refl C 30 | } 31 | 32 | _∘F†_ : ∀ {C : DaggerCategory o ℓ e} {D : DaggerCategory o′ ℓ′ e′} {E : DaggerCategory o″ ℓ″ e″} 33 | → DaggerFunctor D E → DaggerFunctor C D → DaggerFunctor C E 34 | _∘F†_ {E = E} F G = record 35 | { functor = F.functor ∘F G.functor 36 | ; F-resp-† = DaggerCategory.Equiv.trans E F.F-resp-† (F.F-resp-≈ G.F-resp-†) 37 | } 38 | where 39 | module F = DaggerFunctor F 40 | module G = DaggerFunctor G 41 | -------------------------------------------------------------------------------- /src/Categories/Functor/DistributiveLaw.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Functor.DistributiveLaw where 3 | 4 | open import Level 5 | open import Categories.Category using (Category) 6 | open import Categories.Functor using (Endofunctor; _∘F_) 7 | open import Categories.NaturalTransformation using (NaturalTransformation) 8 | 9 | private 10 | variable 11 | o ℓ e : Level 12 | 13 | DistributiveLaw : {C : Category o ℓ e} → (T F : Endofunctor C) → Set (o ⊔ ℓ ⊔ e) 14 | DistributiveLaw {C = C} T F = NaturalTransformation (T ∘F F) (F ∘F T) where open Category C 15 | -------------------------------------------------------------------------------- /src/Categories/Functor/Fibration.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Functor 5 | 6 | -- Street fibration, which is the version of fibration that respects the principle of equivalence. 7 | -- https://ncatlab.org/nlab/show/Grothendieck+fibration#StreetFibration 8 | module Categories.Functor.Fibration {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} (F : Functor C D) where 9 | 10 | open import Level 11 | 12 | open import Categories.Morphism D using (_≅_) 13 | open import Categories.Morphism.Cartesian using (Cartesian) 14 | 15 | private 16 | module C = Category C 17 | module D = Category D 18 | open Functor F 19 | 20 | record Fibration : Set (levelOfTerm F) where 21 | field 22 | universal₀ : ∀ {A B} (f : A D.⇒ F₀ B) → C.Obj 23 | universal₁ : ∀ {A B} (f : A D.⇒ F₀ B) → universal₀ f C.⇒ B 24 | iso : ∀ {A B} (f : A D.⇒ F₀ B) → F₀ (universal₀ f) ≅ A 25 | 26 | module iso {A B} (f : A D.⇒ F₀ B) = _≅_ (iso f) 27 | 28 | field 29 | commute : ∀ {A B} (f : A D.⇒ F₀ B) → f D.∘ iso.from f D.≈ F₁ (universal₁ f) 30 | cartesian : ∀ {A B} (f : A D.⇒ F₀ B) → Cartesian F (universal₁ f) 31 | 32 | module cartesian {A B} (f : A D.⇒ F₀ B) = Cartesian (cartesian f) 33 | -------------------------------------------------------------------------------- /src/Categories/Functor/Groupoid.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Functor.Groupoid where 4 | 5 | open import Level 6 | 7 | open import Categories.Category 8 | open import Categories.Category.Groupoid using (Groupoid; IsGroupoid) 9 | open import Categories.Functor 10 | import Categories.Morphism.Reasoning as MR 11 | 12 | private 13 | variable 14 | o ℓ e : Level 15 | C D : Category o ℓ e 16 | 17 | -- functor respects groupoid inverse 18 | module _ (GC : IsGroupoid C) (GD : IsGroupoid D) (F : Functor C D) where 19 | private 20 | module C = IsGroupoid GC 21 | module D = IsGroupoid GD 22 | open Functor F 23 | open D 24 | open HomReasoning 25 | open MR D 26 | 27 | F-resp-inv : ∀ {A B} (f : A C.⇒ B) → F₁ (f C.⁻¹) ≈ (F₁ f) ⁻¹ 28 | F-resp-inv f = begin 29 | F₁ (f C.⁻¹) ≈⟨ introˡ D.iso.isoˡ ⟩ 30 | ((F₁ f) ⁻¹ ∘ F₁ f) ∘ F₁ (f C.⁻¹) ≈˘⟨ pushʳ homomorphism ⟩ 31 | (F₁ f) ⁻¹ ∘ F₁ (f C.∘ f C.⁻¹) ≈⟨ elimʳ (F-resp-≈ C.iso.isoʳ ○ identity) ⟩ 32 | (F₁ f) ⁻¹ ∎ 33 | 34 | -- Nevertheless, create a |GrFunctor| that is typed over Groupoid instead of Category 35 | GrFunctor : (A B : Groupoid o ℓ e) → Set _ 36 | GrFunctor A B = Functor (category A) (category B) 37 | where open Groupoid 38 | -------------------------------------------------------------------------------- /src/Categories/Functor/Hom/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Functor.Hom.Properties {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Categories.Functor.Hom.Properties.Covariant C public 8 | open import Categories.Functor.Hom.Properties.Contra C public 9 | -------------------------------------------------------------------------------- /src/Categories/Functor/IdentityOnObjects.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | 4 | -- Identity-on-Objects Functor for (Object-)Unbundled Categories 5 | 6 | module Categories.Functor.IdentityOnObjects where 7 | 8 | open import Data.Product using (_,_) 9 | open import Function using () renaming (id to id→) 10 | open import Level 11 | 12 | open import Categories.Category.Unbundled using (Category) 13 | open import Categories.Category.Unbundled.Properties using (pack′) 14 | open import Categories.Category.Unbundled.Utilities using (module Equiv) 15 | open import Categories.Functor.Core using (Functor) 16 | 17 | private 18 | variable 19 | o ℓ e ℓ′ e′ : Level 20 | 21 | record IdentityOnObjects {Obj : Set o} (C : Category Obj ℓ e) (D : Category Obj ℓ′ e′) : Set (o ⊔ ℓ ⊔ e ⊔ ℓ′ ⊔ e′) where 22 | eta-equality 23 | private module C = Category C 24 | private module D = Category D 25 | 26 | field 27 | F₁ : ∀ {A B} → (A C.⇒ B) → A D.⇒ B 28 | 29 | identity : ∀ {A} → F₁ (C.id {A}) D.≈ D.id 30 | homomorphism : ∀ {X Y Z} {f : X C.⇒ Y} {g : Y C.⇒ Z} → 31 | F₁ (g C.∘ f) D.≈ F₁ g D.∘ F₁ f 32 | F-resp-≈ : ∀ {A B} {f g : A C.⇒ B} → f C.≈ g → F₁ f D.≈ F₁ g 33 | 34 | IOO⇒Functor : {Ob : Set o} {C : Category Ob ℓ e} {D : Category Ob ℓ′ e′} → 35 | (F : IdentityOnObjects C D) → Functor (pack′ C) (pack′ D) 36 | IOO⇒Functor F = record { F₀ = id→; IOO } 37 | where module IOO = IdentityOnObjects F 38 | 39 | id-IOO : {Obj : Set o} {C : Category Obj ℓ e} → IdentityOnObjects C C 40 | id-IOO {C = C} = record { F₁ = id→ ; identity = refl ; homomorphism = refl ; F-resp-≈ = id→ } 41 | where open Equiv C 42 | -------------------------------------------------------------------------------- /src/Categories/Functor/Instance/ConnectedComponents.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Functor.Instance.ConnectedComponents where 4 | 5 | open import Level 6 | open import Function using () renaming (_∘′_ to _∙_) 7 | open import Relation.Binary.Construct.Closure.SymmetricTransitive as ST using (Plus⇔; minimal) 8 | 9 | open import Categories.Category 10 | open import Categories.Category.Instance.Cats using (Cats) 11 | open import Categories.Category.Instance.Setoids using (Setoids) 12 | open import Categories.Functor hiding (id) 13 | open import Categories.NaturalTransformation.NaturalIsomorphism using (NaturalIsomorphism) 14 | 15 | -- this is the left adjoint to Disc : Setoids ⇒ Cats 16 | Π₀ : ∀ {o ℓ e} → Functor (Cats o ℓ e) (Setoids o (o ⊔ ℓ)) 17 | Π₀ = record 18 | { F₀ = λ C → ST.setoid (Category._⇒_ C) (Category.id C) 19 | ; F₁ = λ F → record 20 | { to = Functor.F₀ F 21 | ; cong = ST.gmap (Functor.F₀ F) (Functor.F₁ F) 22 | } 23 | ; identity = λ {A} → Plus⇔.forth (Category.id A) 24 | ; homomorphism = λ {_ _ Z F G x} → Plus⇔.forth (Category.id Z) 25 | ; F-resp-≈ = λ {A} {B} {f} {g} α {a} → Plus⇔.forth (NaturalIsomorphism.⇒.η α a) 26 | } 27 | -------------------------------------------------------------------------------- /src/Categories/Functor/Instance/Twisted.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category; module Definitions) 4 | 5 | -- Definition of the "Twisted" Functor between certain Functor Categories 6 | module Categories.Functor.Instance.Twisted {o ℓ e o′ ℓ′ e′} (C : Category o ℓ e) (D : Category o′ ℓ′ e′)where 7 | 8 | open import Level 9 | open import Data.Product using (_,_) 10 | 11 | open import Categories.Category.Construction.Functors using (Functors; product) 12 | open import Categories.Category.Construction.TwistedArrow using (TwistedArrow; Morphism; Morphism⇒; Codomain) 13 | open import Categories.Category.Product using () renaming (Product to _×ᶜ_) 14 | open import Categories.Functor using (Functor; _∘F_) 15 | open import Categories.Functor.Bifunctor using (appʳ) 16 | open import Categories.Functor.Limits using (Continuous) 17 | open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) 18 | open import Categories.NaturalTransformation.NaturalIsomorphism using (_≃_; _ⓘʳ_) 19 | 20 | private 21 | module C = Category C 22 | module D = Category D 23 | 24 | open Morphism 25 | open Morphism⇒ 26 | 27 | -- precomposition with the forgetful functor 28 | Twist : Functor (C.op ×ᶜ C) D → Functor (TwistedArrow C) D 29 | Twist F = F ∘F Codomain C 30 | 31 | Twist′ : Functor (C.op ×ᶜ C) D → Functor (Category.op (TwistedArrow C.op)) D 32 | Twist′ F = F ∘F (Functor.op (Codomain C.op)) 33 | 34 | -- precomposition is functorial 35 | Twisted : Functor (Functors (C.op ×ᶜ C) D) (Functors (TwistedArrow C) D) 36 | Twisted = appʳ product (Codomain C) 37 | 38 | Twistⁿⁱ : ∀ {F G : Functor (C.op ×ᶜ C) D } → (F ≃ G) → Twist F ≃ Twist G 39 | Twistⁿⁱ α = α ⓘʳ Codomain C 40 | -------------------------------------------------------------------------------- /src/Categories/Functor/Presheaf.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Functor.Presheaf where 4 | 5 | open import Categories.Category 6 | open import Categories.Functor 7 | 8 | Presheaf : ∀ {o ℓ e} {o′ ℓ′ e′} (C : Category o ℓ e) (V : Category o′ ℓ′ e′) → Set _ 9 | Presheaf C V = Functor C.op V 10 | where module C = Category C 11 | -------------------------------------------------------------------------------- /src/Categories/Functor/Representable.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Functor.Representable where 3 | 4 | -- A Presheaf (into Setoids) is representation if it is naturally isomorphic to a Hom functor 5 | -- over a particular object A of the base category. 6 | open import Level 7 | 8 | open import Categories.Category using (Category) 9 | open import Categories.Category.Instance.Setoids 10 | open import Categories.Functor using (Functor) 11 | open import Categories.Functor.Hom 12 | open import Categories.Functor.Presheaf 13 | open import Categories.NaturalTransformation.NaturalIsomorphism 14 | 15 | record Representable {o ℓ e} {C : Category o ℓ e} (F : Presheaf C (Setoids ℓ e)) : Set (o ⊔ suc ℓ ⊔ suc e) where 16 | open Category C 17 | open Hom C 18 | field 19 | A : Obj 20 | Iso : NaturalIsomorphism F Hom[-, A ] 21 | -------------------------------------------------------------------------------- /src/Categories/Functor/Slice/BaseChange.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category) 4 | 5 | module Categories.Functor.Slice.BaseChange {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Categories.Category.Slice C using (Slice) 8 | open import Categories.Category.Slice.Properties C using (pullback⇒product; slice-slice⇒slice; slice⇒slice-slice) 9 | open import Categories.Functor using (Functor; _∘F_) 10 | open import Categories.Functor.Slice using (TotalSpace; ConstantFamily) 11 | open import Categories.Diagram.Pullback C using (Pullback) 12 | 13 | open Category C 14 | 15 | module _ {A B : Obj} (f : B ⇒ A) where 16 | 17 | -- Any morphism induces a functor between slices. 18 | BaseChange! : Functor (Slice B) (Slice A) 19 | BaseChange! = TotalSpace (Slice A) ∘F slice⇒slice-slice f 20 | 21 | -- Any morphism which admits pullbacks induces a functor the other way between slices. 22 | -- This is adjoint to BaseChange!: see Categories.Adjoint.Instance.BaseChange. 23 | BaseChange* : (∀ {C} {h : C ⇒ A} → Pullback f h) → Functor (Slice A) (Slice B) 24 | BaseChange* pullback = slice-slice⇒slice f ∘F ConstantFamily (Slice A) (pullback⇒product pullback) 25 | -------------------------------------------------------------------------------- /src/Categories/GlobularSet.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.GlobularSet where 3 | 4 | -- Globular sets are defined in a Categorical context, but 5 | -- should they be inside the Categories hierarchy? 6 | 7 | open import Level 8 | open import Data.Unit using (⊤) 9 | open import Relation.Binary.PropositionalEquality using (refl) 10 | 11 | open import Categories.Category 12 | open import Categories.Category.Instance.Globe 13 | open import Categories.Category.Instance.Sets 14 | open import Categories.Functor 15 | open import Categories.Functor.Presheaf 16 | 17 | private 18 | variable 19 | o ℓ e : Level 20 | 21 | GlobularSet : (o : Level) → Set (suc o) 22 | GlobularSet o = Presheaf Globe (Sets o) 23 | 24 | -- TODO? make universe polymorphic with polymorphic ⊤ 25 | Trivial : GlobularSet zero 26 | Trivial = record 27 | { F₀ = λ _ → ⊤ 28 | ; F₁ = λ _ x → x 29 | ; identity = λ _ → refl 30 | ; homomorphism = λ _ → refl 31 | ; F-resp-≈ = λ _ _ → refl 32 | } 33 | 34 | GlobularObject : Category o ℓ e → Set _ 35 | GlobularObject C = Functor (Category.op Globe) C 36 | -------------------------------------------------------------------------------- /src/Categories/Kan/Duality.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Kan.Duality where 4 | 5 | open import Level 6 | open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) 7 | 8 | open import Categories.Category 9 | open import Categories.Functor 10 | open import Categories.NaturalTransformation 11 | open import Categories.Kan 12 | 13 | private 14 | variable 15 | o ℓ e : Level 16 | C D E : Category o ℓ e 17 | F G : Functor C D 18 | 19 | module _ {F : Functor C D} {G : Functor C E} where 20 | private 21 | module F = Functor F 22 | module G = Functor G 23 | 24 | coLan⇒Ran : Lan F.op G.op → Ran F G 25 | coLan⇒Ran lan = record 26 | { R = L.op 27 | ; ε = η.op 28 | ; δ = λ M α → NaturalTransformation.op (σ (Functor.op M) (NaturalTransformation.op α)) 29 | ; δ-unique = λ δ′ eq → σ-unique (NaturalTransformation.op δ′) eq 30 | ; commutes = λ M α → commutes (Functor.op M) (NaturalTransformation.op α) 31 | } 32 | where open Lan lan 33 | 34 | coRan⇒Lan : Ran F.op G.op → Lan F G 35 | coRan⇒Lan ran = record 36 | { L = R.op 37 | ; η = ε.op 38 | ; σ = λ M α → NaturalTransformation.op (δ (Functor.op M) (NaturalTransformation.op α)) 39 | ; σ-unique = λ σ′ eq → δ-unique (NaturalTransformation.op σ′) eq 40 | ; commutes = λ M α → commutes (Functor.op M) (NaturalTransformation.op α) 41 | } 42 | where open Ran ran 43 | 44 | private 45 | module _ {F : Functor C D} {G : Functor C E} where 46 | module F = Functor F 47 | module G = Functor G 48 | 49 | coLan⟺Ran : (R : Ran F.op G.op) → coLan⇒Ran (coRan⇒Lan R) ≡ R 50 | coLan⟺Ran _ = ≡.refl 51 | 52 | coRan⟺Lan : (L : Lan F.op G.op) → coRan⇒Lan (coLan⇒Ran L) ≡ L 53 | coRan⟺Lan _ = ≡.refl 54 | -------------------------------------------------------------------------------- /src/Categories/Minus2-Category.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- 'Traditionally', meaning in nLab and in 4 | -- "Lectures on n-Categories and Cohomology" by Baez and Shulman 5 | -- https://arxiv.org/abs/math/0608420 6 | -- (-2)-Categories are defined to be just a single value, with trivial Hom 7 | 8 | -- But that's hardly a definition of a class of things, it's a definition of 9 | -- a single structure! What we want is the definition of a class which turns 10 | -- out to be (essentially) unique. Rather like the reals are (essentially) the 11 | -- only ordered complete archimedean field. 12 | 13 | -- So we will take a -2-Category to be a full-fledge Category, but where 14 | -- 1. |Obj| is (Categorically) contractible 15 | -- 2. |Hom| is connected (all arrows are equal) 16 | -- Note that we don't need to say anything at all about ≈ 17 | 18 | module Categories.Minus2-Category where 19 | 20 | open import Level 21 | open import Categories.Category 22 | open import Data.Product using (Σ) 23 | import Categories.Morphism as M 24 | 25 | private 26 | variable 27 | o ℓ e : Level 28 | 29 | record -2-Category : Set (suc (o ⊔ ℓ ⊔ e)) where 30 | field 31 | cat : Category o ℓ e 32 | open Category cat public 33 | open M cat using (_≅_) 34 | 35 | field 36 | Obj-Contr : Σ Obj (λ x → (y : Obj) → x ≅ y) 37 | Hom-Conn : {x y : Obj} {f g : x ⇒ y} → f ≈ g 38 | -------------------------------------------------------------------------------- /src/Categories/Minus2-Category/Construction/Indiscrete.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- An inhabited Indiscrete category is a -2-Category 4 | module Categories.Minus2-Category.Construction.Indiscrete where 5 | 6 | open import Level 7 | open import Data.Unit using (tt) 8 | open import Data.Product using (_,_) 9 | open import Relation.Binary.PropositionalEquality using (refl) 10 | 11 | open import Categories.Minus2-Category 12 | open import Categories.Category.Indiscrete 13 | open import Categories.Morphism using (_≅_) 14 | 15 | InhIndIs-2 : ∀ {o ℓ} → (X : Set o) → (x : X) → -2-Category {o} {ℓ} {ℓ} 16 | InhIndIs-2 X x = record 17 | { cat = Indiscrete X 18 | ; Obj-Contr = x , λ y → record 19 | { from = lift tt 20 | ; to = lift tt 21 | ; iso = record 22 | { isoˡ = refl 23 | ; isoʳ = refl 24 | } 25 | } 26 | ; Hom-Conn = refl 27 | } 28 | -------------------------------------------------------------------------------- /src/Categories/Minus2-Category/Instance/One.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- The One Category is also a -2-Category 4 | module Categories.Minus2-Category.Instance.One where 5 | 6 | open import Categories.Minus2-Category 7 | open import Categories.Category.Instance.One 8 | 9 | -- Proof is trivial 10 | ⊤-is-2-Category : ∀ {o ℓ e} → -2-Category {o} {ℓ} {e} 11 | ⊤-is-2-Category = record { cat = One } 12 | -------------------------------------------------------------------------------- /src/Categories/Minus2-Category/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Minus2-Category.Properties where 4 | 5 | -- All -2-Categories are equivalent to One 6 | 7 | open import Level 8 | open import Data.Product using (Σ; _,_; proj₁; proj₂) 9 | open import Data.Unit using (⊤; tt) 10 | 11 | open import Categories.Minus2-Category 12 | open import Categories.Category 13 | import Categories.Morphism as M 14 | open import Categories.Category.Monoidal 15 | open import Categories.Category.Instance.One 16 | open import Categories.Category.Equivalence hiding (refl) 17 | open import Categories.NaturalTransformation using (ntHelper) 18 | 19 | private 20 | variable 21 | o ℓ e : Level 22 | 23 | shrink-them-all : (X : -2-Category {o} {ℓ} {e}) → StrongEquivalence (-2-Category.cat X) (One {o} {ℓ} {e}) 24 | shrink-them-all X = record 25 | { F = record 26 | { F₀ = λ _ → lift tt 27 | ; F₁ = λ _ → lift tt 28 | } 29 | ; G = record 30 | { F₀ = λ _ → proj₁ Obj-Contr 31 | ; F₁ = λ _ → M._≅_.from (proj₂ Obj-Contr (proj₁ Obj-Contr)) 32 | ; identity = Hom-Conn 33 | ; homomorphism = Hom-Conn 34 | ; F-resp-≈ = λ _ → Hom-Conn 35 | } 36 | ; weak-inverse = record 37 | { F∘G≈id = _ 38 | ; G∘F≈id = record 39 | { F⇒G = ntHelper (record 40 | { η = λ y → M._≅_.from (proj₂ Obj-Contr y) 41 | ; commute = λ _ → Hom-Conn 42 | }) 43 | ; F⇐G = ntHelper (record 44 | { η = λ y → M._≅_.to (proj₂ Obj-Contr y) 45 | ; commute = λ _ → Hom-Conn 46 | }) 47 | ; iso = λ Z → record 48 | { isoˡ = M._≅_.isoˡ (proj₂ Obj-Contr Z) 49 | ; isoʳ = M._≅_.isoʳ (proj₂ Obj-Contr Z) 50 | } 51 | } 52 | } 53 | } 54 | where 55 | open -2-Category X 56 | open Category cat 57 | -------------------------------------------------------------------------------- /src/Categories/Monad.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Monad where 3 | 4 | open import Level 5 | 6 | open import Categories.Category using (Category) 7 | open import Categories.Functor using (Functor; Endofunctor; _∘F_) renaming (id to idF) 8 | open import Categories.NaturalTransformation renaming (id to idN) 9 | open import Categories.NaturalTransformation.NaturalIsomorphism hiding (_≃_) 10 | open import Categories.NaturalTransformation.Equivalence 11 | open NaturalIsomorphism 12 | 13 | record Monad {o ℓ e} (C : Category o ℓ e) : Set (o ⊔ ℓ ⊔ e) where 14 | field 15 | F : Endofunctor C 16 | η : NaturalTransformation idF F 17 | μ : NaturalTransformation (F ∘F F) F 18 | 19 | module F = Functor F 20 | module η = NaturalTransformation η 21 | module μ = NaturalTransformation μ 22 | 23 | open Category C 24 | open F 25 | 26 | field 27 | assoc : ∀ {X : Obj} → μ.η X ∘ F₁ (μ.η X) ≈ μ.η X ∘ μ.η (F₀ X) 28 | sym-assoc : ∀ {X : Obj} → μ.η X ∘ μ.η (F₀ X) ≈ μ.η X ∘ F₁ (μ.η X) 29 | identityˡ : ∀ {X : Obj} → μ.η X ∘ F₁ (η.η X) ≈ id 30 | identityʳ : ∀ {X : Obj} → μ.η X ∘ η.η (F₀ X) ≈ id 31 | 32 | module _ {o ℓ e} (C : Category o ℓ e) where 33 | open Monad 34 | open Category C hiding (id) 35 | id : Monad C 36 | id .F = idF 37 | id .η = idN 38 | id .μ = unitor² .F⇒G 39 | id .assoc = Equiv.refl 40 | id .sym-assoc = Equiv.refl 41 | id .identityˡ = identity² 42 | id .identityʳ = identity² 43 | -------------------------------------------------------------------------------- /src/Categories/Monad/Commutative.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- Commutative Monad on a braided monoidal category 4 | -- https://ncatlab.org/nlab/show/commutative+monad 5 | 6 | module Categories.Monad.Commutative where 7 | 8 | open import Level 9 | open import Data.Product using (_,_) 10 | 11 | open import Categories.Category.Core using (Category) 12 | open import Categories.Category.Monoidal using (Monoidal) 13 | open import Categories.Category.Monoidal.Braided using (Braided) 14 | open import Categories.Monad using (Monad) 15 | open import Categories.Monad.Strong using (StrongMonad; RightStrength; Strength) 16 | import Categories.Monad.Strong.Properties as StrongProps 17 | 18 | private 19 | variable 20 | o ℓ e : Level 21 | 22 | module _ {C : Category o ℓ e} {V : Monoidal C} (B : Braided V) where 23 | record Commutative (LSM : StrongMonad V) : Set (o ⊔ ℓ ⊔ e) where 24 | open Category C using (_⇒_; _∘_; _≈_) 25 | open Braided B using (_⊗₀_) 26 | open StrongMonad LSM using (M; strength) 27 | open StrongProps.Left.Shorthands strength 28 | 29 | rightStrength : RightStrength V M 30 | rightStrength = StrongProps.Strength⇒RightStrength B strength 31 | 32 | open StrongProps.Right.Shorthands rightStrength 33 | 34 | field 35 | commutes : ∀ {X Y} → M.μ.η (X ⊗₀ Y) ∘ M.F.₁ τ ∘ σ ≈ M.μ.η (X ⊗₀ Y) ∘ M.F.₁ σ ∘ τ 36 | 37 | record CommutativeMonad : Set (o ⊔ ℓ ⊔ e) where 38 | field 39 | strongMonad : StrongMonad V 40 | commutative : Commutative strongMonad 41 | 42 | open StrongMonad strongMonad public 43 | open Commutative commutative public 44 | 45 | -------------------------------------------------------------------------------- /src/Categories/Monad/Duality.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Monad.Duality {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl) 8 | 9 | open import Categories.Functor.Core using (Functor) 10 | open import Categories.NaturalTransformation.Core using (NaturalTransformation) 11 | open import Categories.Monad 12 | open import Categories.Comonad 13 | 14 | private 15 | module C = Category C 16 | open C 17 | open HomReasoning 18 | 19 | coMonad⇒Comonad : Monad C.op → Comonad C 20 | coMonad⇒Comonad M = record 21 | { F = Functor.op F 22 | ; ε = NaturalTransformation.op η 23 | ; δ = NaturalTransformation.op μ 24 | ; assoc = M.sym-assoc 25 | ; sym-assoc = M.assoc 26 | ; identityˡ = M.identityˡ 27 | ; identityʳ = M.identityʳ 28 | } 29 | where module M = Monad M 30 | open M using (F; η; μ) 31 | 32 | Comonad⇒coMonad : Comonad C → Monad C.op 33 | Comonad⇒coMonad M = record 34 | { F = Functor.op F 35 | ; η = NaturalTransformation.op ε 36 | ; μ = NaturalTransformation.op δ 37 | ; assoc = M.sym-assoc 38 | ; sym-assoc = M.assoc 39 | ; identityˡ = M.identityˡ 40 | ; identityʳ = M.identityʳ 41 | } 42 | where module M = Comonad M 43 | open M using (F; ε; δ) 44 | 45 | 46 | module MonadDualityConversionProperties where 47 | private 48 | coMonad⇔Comonad : ∀ (coMonad : Monad C.op) → 49 | Comonad⇒coMonad (coMonad⇒Comonad coMonad) ≡ coMonad 50 | coMonad⇔Comonad _ = refl 51 | 52 | Comonad⇔coMonad : ∀ (M : Comonad C) → coMonad⇒Comonad (Comonad⇒coMonad M) ≡ M 53 | Comonad⇔coMonad _ = refl 54 | -------------------------------------------------------------------------------- /src/Categories/Monad/Idempotent.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level 4 | open import Categories.Category 5 | open import Categories.Monad 6 | 7 | module Categories.Monad.Idempotent {o ℓ e} {C : Category o ℓ e} (M : Monad C) where 8 | 9 | open import Categories.NaturalTransformation 10 | open import Categories.NaturalTransformation.Equivalence 11 | open import Categories.Functor 12 | 13 | private 14 | module M = Monad M 15 | 16 | open NaturalTransformation 17 | 18 | record Idempotent : Set (o ⊔ ℓ ⊔ e) where 19 | field 20 | Fη≡ηF : ∀ X → C [ η (M.F ∘ˡ M.η) X ≈ η (M.η ∘ʳ M.F) X ] 21 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Cartesian.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Morphism.Cartesian where 4 | 5 | open import Level 6 | 7 | open import Categories.Category 8 | open import Categories.Functor 9 | 10 | private 11 | variable 12 | o ℓ e : Level 13 | C D : Category o ℓ e 14 | 15 | record Cartesian (F : Functor C D) {X Y} (f : C [ X , Y ]) : Set (levelOfTerm F) where 16 | private 17 | module C = Category C 18 | module D = Category D 19 | open Functor F 20 | open D 21 | 22 | field 23 | universal : ∀ {A} {u : F₀ A ⇒ F₀ X} (h : C [ A , Y ]) → 24 | F₁ f ∘ u ≈ F₁ h → C [ A , X ] 25 | commute : ∀ {A} {u : F₀ A ⇒ F₀ X} {h : C [ A , Y ]} 26 | (eq : F₁ f ∘ u ≈ F₁ h) → 27 | C [ C [ f ∘ universal h eq ] ≈ h ] 28 | compat : ∀ {A} {u : F₀ A ⇒ F₀ X} {h : C [ A , Y ]} 29 | (eq : F₁ f ∘ u ≈ F₁ h) → 30 | F₁ (universal h eq) ≈ u 31 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Duality.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Morphism.Duality {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Relation.Binary.PropositionalEquality using (_≡_; refl) 8 | 9 | open Category C 10 | 11 | import Categories.Morphism as M 12 | private 13 | module Op = M op 14 | open M C 15 | 16 | open import Categories.Morphism.Properties C 17 | 18 | private 19 | variable 20 | A B X Y : Obj 21 | f g h : A ⇒ B 22 | 23 | Mono⇒op-Epi : Mono f → Op.Epi f 24 | Mono⇒op-Epi mono = mono 25 | 26 | Epi⇒op-Mono : Epi f → Op.Mono f 27 | Epi⇒op-Mono epi = epi 28 | 29 | Iso⇒op-Iso : Iso f g → Op.Iso g f 30 | Iso⇒op-Iso iso = record 31 | { isoˡ = isoˡ 32 | ; isoʳ = isoʳ 33 | } 34 | where open Iso iso 35 | 36 | op-Iso⇒Iso : Op.Iso g f → Iso f g 37 | op-Iso⇒Iso iso = record 38 | { isoˡ = isoˡ 39 | ; isoʳ = isoʳ 40 | } 41 | where open Op.Iso iso 42 | 43 | ≅⇒op-≅ : A ≅ B → A Op.≅ B 44 | ≅⇒op-≅ A≅B = record 45 | { from = to 46 | ; to = from 47 | ; iso = Iso⇒op-Iso iso 48 | } 49 | where open _≅_ A≅B 50 | 51 | op-≅⇒≅ : A Op.≅ B → A ≅ B 52 | op-≅⇒≅ A≅B = record 53 | { from = to 54 | ; to = from 55 | ; iso = op-Iso⇒Iso iso 56 | } 57 | where open Op._≅_ A≅B 58 | 59 | 60 | module MorphismDualityConversionProperties where 61 | private 62 | op-Iso-involutive : ∀(iso : Iso f g) → op-Iso⇒Iso (Iso⇒op-Iso iso) ≡ iso 63 | op-Iso-involutive _ = refl 64 | 65 | op-≅-involutive : ∀(A′ B′ : Obj) → (A′≅B′ : A′ ≅ B′) 66 | → op-≅⇒≅ (≅⇒op-≅ A′≅B′) ≡ A′≅B′ 67 | op-≅-involutive _ _ _ = refl 68 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Extremal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | {- 4 | Extremal Mono and Epimorphisms. 5 | 6 | https://ncatlab.org/nlab/show/extremal+epimorphism 7 | -} 8 | 9 | open import Categories.Category.Core 10 | 11 | module Categories.Morphism.Extremal {o ℓ e} (𝒞 : Category o ℓ e) where 12 | 13 | open import Level 14 | open import Categories.Morphism 𝒞 15 | 16 | open Category 𝒞 17 | 18 | IsExtremalEpi : ∀ {A B} {f : A ⇒ B} → Epi f → Set (o ⊔ ℓ ⊔ e) 19 | IsExtremalEpi {A = A} {B = B} {f = f} epi = 20 | ∀ {X} {i : X ⇒ B} {g : A ⇒ X} → Mono i → f ≈ i ∘ g → IsIso i 21 | 22 | IsExtremalMono : ∀ {A B} {f : A ⇒ B} → Mono f → Set (o ⊔ ℓ ⊔ e) 23 | IsExtremalMono {A = A} {B = B} {f = f} mono = 24 | ∀ {X} {g : X ⇒ B} {i : A ⇒ X} → Epi i → f ≈ g ∘ i → IsIso i 25 | 26 | record ExtremalEpi {A B} (f : A ⇒ B) : Set (o ⊔ ℓ ⊔ e) where 27 | field 28 | epi : Epi f 29 | extremal : IsExtremalEpi epi 30 | 31 | 32 | record ExtremalMono {A B} (f : A ⇒ B) : Set (o ⊔ ℓ ⊔ e) where 33 | field 34 | mono : Mono f 35 | extremal : IsExtremalMono mono 36 | -------------------------------------------------------------------------------- /src/Categories/Morphism/HeterogeneousEquality.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category using (Category; module Definitions) 3 | 4 | module Categories.Morphism.HeterogeneousEquality where 5 | 6 | open import Level 7 | 8 | open import Categories.Category using (_[_,_]; _[_≈_]) 9 | open import Categories.Morphism using (_≅_) 10 | open import Categories.Morphism.Notation using (_[_≅_]) 11 | 12 | private 13 | variable 14 | o ℓ e : Level 15 | 16 | 17 | Along_,_[_≈_] : {C : Category o ℓ e}{A A' B B' : Category.Obj C} 18 | → (i : C [ A ≅ A' ])(j : C [ B ≅ B' ])(f : C [ A , B ])(f' : C [ A' , B' ]) → Set e 19 | Along_,_[_≈_] {C = C} i j f f' = C [ _≅_.from j ∘ f ≈ f' ∘ _≅_.from i ] 20 | where open Category C 21 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Idempotent/Bundles.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category using (Category; _[_,_]; _[_∘_]; _[_≈_]) 4 | 5 | -- Bundled versions of Idempotents, as well as maps between idempotents. 6 | module Categories.Morphism.Idempotent.Bundles {o ℓ e} (𝒞 : Category o ℓ e) where 7 | 8 | open import Level 9 | 10 | import Categories.Morphism.Idempotent 𝒞 as Idem 11 | open import Categories.Morphism.Reasoning 𝒞 12 | 13 | private 14 | module 𝒞 = Category 𝒞 15 | open 𝒞.HomReasoning 16 | open 𝒞.Equiv 17 | 18 | -------------------------------------------------------------------------------- 19 | -- Bundled Idempotents, and maps between them 20 | 21 | record Idempotent : Set (o ⊔ ℓ ⊔ e) where 22 | field 23 | {obj} : 𝒞.Obj 24 | isIdempotent : Idem.Idempotent obj 25 | 26 | open Idem.Idempotent isIdempotent public 27 | 28 | open Idempotent 29 | 30 | record Idempotent⇒ (I J : Idempotent) : Set (ℓ ⊔ e) where 31 | private 32 | module I = Idempotent I 33 | module J = Idempotent J 34 | field 35 | hom : 𝒞 [ I.obj , J.obj ] 36 | absorbˡ : 𝒞 [ 𝒞 [ J.idem ∘ hom ] ≈ hom ] 37 | absorbʳ : 𝒞 [ 𝒞 [ hom ∘ I.idem ] ≈ hom ] 38 | 39 | open Idempotent⇒ 40 | 41 | -------------------------------------------------------------------------------- 42 | -- Identity and Composition of maps between Idempotents 43 | 44 | id : ∀ {I} → Idempotent⇒ I I 45 | id {I} = record 46 | { hom = idem I 47 | ; absorbˡ = idempotent I 48 | ; absorbʳ = idempotent I 49 | } 50 | 51 | _∘_ : ∀ {I J K} → (f : Idempotent⇒ J K) → (g : Idempotent⇒ I J) → Idempotent⇒ I K 52 | _∘_ {I} {J} {K} f g = record 53 | { hom = 𝒞 [ f.hom ∘ g.hom ] 54 | ; absorbˡ = pullˡ f.absorbˡ 55 | ; absorbʳ = pullʳ g.absorbʳ 56 | } 57 | where 58 | module f = Idempotent⇒ f 59 | module g = Idempotent⇒ g 60 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Normal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core 4 | open import Categories.Object.Zero 5 | 6 | -- Normal Mono and Epimorphisms 7 | -- https://ncatlab.org/nlab/show/normal+monomorphism 8 | module Categories.Morphism.Normal {o ℓ e} (𝒞 : Category o ℓ e) (𝒞-Zero : Zero 𝒞) where 9 | 10 | open import Level 11 | 12 | open import Categories.Object.Kernel 𝒞-Zero 13 | open import Categories.Object.Kernel.Properties 𝒞-Zero 14 | open import Categories.Morphism 𝒞 15 | 16 | open Category 𝒞 17 | 18 | record IsNormalMonomorphism {A K : Obj} (k : K ⇒ A) : Set (o ⊔ ℓ ⊔ e) where 19 | field 20 | {B} : Obj 21 | arr : A ⇒ B 22 | isKernel : IsKernel k arr 23 | 24 | open IsKernel isKernel public 25 | 26 | mono : Mono k 27 | mono = Kernel-Mono (IsKernel⇒Kernel isKernel) 28 | 29 | record NormalMonomorphism (K A : Obj) : Set (o ⊔ ℓ ⊔ e) where 30 | field 31 | mor : K ⇒ A 32 | isNormalMonomorphism : IsNormalMonomorphism mor 33 | 34 | open IsNormalMonomorphism isNormalMonomorphism public 35 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Notation.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | {- 4 | Useful notation for Epimorphisms, Mononmorphisms, and isomorphisms 5 | -} 6 | module Categories.Morphism.Notation where 7 | 8 | open import Level 9 | 10 | open import Categories.Category.Core 11 | open import Categories.Morphism 12 | 13 | private 14 | variable 15 | o ℓ e : Level 16 | 17 | _[_↣_] : (𝒞 : Category o ℓ e) → (A B : Category.Obj 𝒞) → Set (o ⊔ ℓ ⊔ e) 18 | 𝒞 [ A ↣ B ] = _↣_ 𝒞 A B 19 | 20 | _[_↠_] : (𝒞 : Category o ℓ e) → (A B : Category.Obj 𝒞) → Set (o ⊔ ℓ ⊔ e) 21 | 𝒞 [ A ↠ B ] = _↠_ 𝒞 A B 22 | 23 | _[_≅_] : (𝒞 : Category o ℓ e) → (A B : Category.Obj 𝒞) → Set (ℓ ⊔ e) 24 | 𝒞 [ A ≅ B ] = _≅_ 𝒞 A B 25 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Reasoning.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | -- Reasoning facilities about morphism equivalences (not necessarily 'squares') 6 | 7 | module Categories.Morphism.Reasoning {o ℓ e} (C : Category o ℓ e) where 8 | 9 | -- some items are defined in sub-modules 10 | open import Categories.Morphism.Reasoning.Core C public 11 | open import Categories.Morphism.Reasoning.Iso C public 12 | 13 | open Category C 14 | open Definitions C 15 | open HomReasoning 16 | 17 | -- create a commutative square from an equivalence 18 | toSquare : ∀ {A B} {f g : A ⇒ B} → f ≈ g → CommutativeSquare f id id g 19 | toSquare {_} {_} {f} {g} f≈g = begin 20 | id ∘ f ≈⟨ identityˡ ⟩ 21 | f ≈⟨ f≈g ⟩ 22 | g ≈˘⟨ identityʳ ⟩ 23 | g ∘ id ∎ 24 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Regular.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | {- 4 | Properties regarding Morphisms of a category: 5 | - Regular Monomorphism 6 | - Regular Epimorphism 7 | 8 | https://ncatlab.org/nlab/show/regular+epimorphism 9 | 10 | These are defined here rather than in Morphism, as this 11 | might cause import cycles (and make the dependency graph 12 | very odd). 13 | -} 14 | open import Categories.Category.Core 15 | 16 | module Categories.Morphism.Regular {o ℓ e} (𝒞 : Category o ℓ e) where 17 | 18 | open import Level 19 | 20 | open import Categories.Morphism 𝒞 21 | open import Categories.Diagram.Equalizer 𝒞 22 | open import Categories.Diagram.Coequalizer 𝒞 23 | 24 | open Category 𝒞 25 | 26 | private 27 | variable 28 | A B : Obj 29 | f : A ⇒ B 30 | 31 | record RegularMono (f : A ⇒ B) : Set (o ⊔ ℓ ⊔ e) where 32 | field 33 | { C } : Obj 34 | g : B ⇒ C 35 | h : B ⇒ C 36 | equalizer : IsEqualizer f h g 37 | 38 | record RegularEpi (f : A ⇒ B) : Set (o ⊔ ℓ ⊔ e) where 39 | field 40 | { C } : Obj 41 | h : C ⇒ A 42 | g : C ⇒ A 43 | coequalizer : IsCoequalizer h g f 44 | 45 | RegularMono⇒Mono : RegularMono f → Mono f 46 | RegularMono⇒Mono regular = IsEqualizer⇒Mono equalizer 47 | where 48 | open RegularMono regular 49 | 50 | RegularEpi⇒Epi : RegularEpi f → Epi f 51 | RegularEpi⇒Epi regular = IsCoequalizer⇒Epi coequalizer 52 | where 53 | open RegularEpi regular 54 | 55 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Regular/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category.Core 4 | 5 | module Categories.Morphism.Regular.Properties {o ℓ e} (𝒞 : Category o ℓ e) where 6 | 7 | open import Categories.Morphism 𝒞 8 | open import Categories.Morphism.Regular 𝒞 9 | open import Categories.Diagram.Equalizer 𝒞 10 | open import Categories.Diagram.Equalizer.Properties 𝒞 11 | open import Categories.Diagram.Coequalizer.Properties 𝒞 12 | 13 | open Category 𝒞 14 | 15 | private 16 | variable 17 | A B : Obj 18 | f g : A ⇒ B 19 | 20 | Section⇒RegularMono : f SectionOf g → RegularMono f 21 | Section⇒RegularMono {f = f} {g = g} g∘f≈id = record 22 | { g = id 23 | ; h = f ∘ g 24 | ; equalizer = section-equalizer g∘f≈id 25 | } 26 | 27 | Retract⇒RegularEpi : f RetractOf g → RegularEpi f 28 | Retract⇒RegularEpi {f = f} {g = g} f∘g≈id = record 29 | { h = g ∘ f 30 | ; g = id 31 | ; coequalizer = retract-coequalizer f∘g≈id 32 | } 33 | -------------------------------------------------------------------------------- /src/Categories/Morphism/Universal.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.Morphism.Universal where 4 | 5 | open import Level 6 | open import Categories.Category 7 | open import Categories.Category.Construction.Comma 8 | open import Categories.Functor 9 | open import Categories.Object.Initial 10 | 11 | record UniversalMorphism {o ℓ e o′ ℓ′ e′} {C : Category o ℓ e} {D : Category o′ ℓ′ e′} 12 | (X : Category.Obj C) (F : Functor D C) : Set (e ⊔ ℓ ⊔ o′ ⊔ ℓ′ ⊔ e′) where 13 | X↙F : Category _ _ _ 14 | X↙F = X ↙ F 15 | 16 | private 17 | module X↙F = Category X↙F 18 | 19 | field 20 | initial : Initial X↙F 21 | 22 | module initial = Initial initial 23 | open initial public 24 | -------------------------------------------------------------------------------- /src/Categories/NaturalTransformation.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.NaturalTransformation where 3 | 4 | -- all the important stuff about NaturalTransformation are defined in .Core 5 | 6 | open import Categories.NaturalTransformation.Core public 7 | -------------------------------------------------------------------------------- /src/Categories/NaturalTransformation/Equivalence.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | -- define a less-than-great equivalence on natural transformations 6 | module Categories.NaturalTransformation.Equivalence {o ℓ e o′ ℓ′ e′} 7 | {C : Category o ℓ e} {D : Category o′ ℓ′ e′} where 8 | 9 | open import Level 10 | open import Relation.Binary using (Rel; IsEquivalence; Setoid) 11 | 12 | open import Categories.Functor using (Functor) 13 | open import Categories.NaturalTransformation.Core using (NaturalTransformation) 14 | 15 | module _ {F G : Functor C D} where 16 | infix 4 _≃_ 17 | open Category.Equiv D 18 | 19 | -- This ad hoc equivalence for NaturalTransformation should really be 'modification' 20 | -- (yep, tricategories!). What is below is only part of the definition of a 'modification'. TODO 21 | _≃_ : Rel (NaturalTransformation F G) (o ⊔ e′) 22 | _≃_ X Y = ∀ {x} → D [ NaturalTransformation.η X x ≈ NaturalTransformation.η Y x ] 23 | 24 | ≃-isEquivalence : IsEquivalence _≃_ 25 | ≃-isEquivalence = record 26 | { refl = refl 27 | ; sym = λ f → sym f -- need to eta-expand to get things to line up properly 28 | ; trans = λ f g → trans f g 29 | } 30 | 31 | ≃-setoid : ∀ (F G : Functor C D) → Setoid (o ⊔ ℓ ⊔ ℓ′ ⊔ e′) (o ⊔ e′) 32 | ≃-setoid F G = record 33 | { Carrier = NaturalTransformation F G 34 | ; _≈_ = _≃_ 35 | ; isEquivalence = ≃-isEquivalence 36 | } 37 | -------------------------------------------------------------------------------- /src/Categories/NaturalTransformation/Hom.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Level 3 | open import Categories.Category using (Category) 4 | 5 | module Categories.NaturalTransformation.Hom {o ℓ e : Level} (C : Category o ℓ e) where 6 | 7 | open import Categories.Category.Instance.Setoids 8 | open import Categories.Functor.Hom using (module Hom; Hom[_][-,_]; Hom[_][_,-]; Hom[_][-,-]) 9 | open import Categories.NaturalTransformation using (NaturalTransformation; ntHelper) renaming (id to idN) 10 | 11 | import Categories.Morphism.Reasoning as MR 12 | 13 | open Category C 14 | open HomReasoning 15 | open MR C 16 | open NaturalTransformation 17 | private 18 | module CE = Category.Equiv C 19 | module C = Category C 20 | 21 | Hom[A,C]⇒Hom[B,C] : {A B : Obj} → (A ⇒ B) → NaturalTransformation Hom[ C ][-, A ] Hom[ C ][-, B ] 22 | Hom[A,C]⇒Hom[B,C] {A} A⇒B = ntHelper record 23 | { η = λ X → record { to = λ X⇒A → A⇒B ∘ X⇒A ; cong = ∘-resp-≈ʳ } 24 | ; commute = λ f {g} → begin 25 | A⇒B ∘ id ∘ g ∘ f ≈⟨ pullˡ id-comm ⟩ 26 | (id ∘ A⇒B) ∘ g ∘ f ≈⟨ pullʳ sym-assoc ⟩ 27 | id ∘ (A⇒B ∘ g) ∘ f ∎ 28 | } 29 | 30 | Hom[C,A]⇒Hom[C,B] : {A B : Obj} → (B ⇒ A) → NaturalTransformation Hom[ C ][ A ,-] Hom[ C ][ B ,-] 31 | Hom[C,A]⇒Hom[C,B] {A} B⇒A = ntHelper record 32 | { η = λ X → record { to = λ A⇒X → A⇒X ∘ B⇒A ; cong = ∘-resp-≈ˡ } 33 | ; commute = λ f {g} → begin 34 | (f ∘ g ∘ id) ∘ B⇒A ≈⟨ pullʳ (pullʳ id-comm-sym) ⟩ 35 | f ∘ g ∘ B⇒A ∘ id ≈⟨ (refl⟩∘⟨ sym-assoc) ⟩ 36 | f ∘ (g ∘ B⇒A) ∘ id ∎ 37 | } 38 | -------------------------------------------------------------------------------- /src/Categories/NaturalTransformation/NaturalIsomorphism/Equivalence.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.NaturalTransformation.NaturalIsomorphism.Equivalence where 3 | 4 | -- a certain notion of equivalence between Natural Isomorphisms. 5 | 6 | open import Level 7 | open import Data.Product using (_×_; _,_; map; zip) 8 | open import Relation.Binary.Structures using (IsEquivalence) 9 | 10 | open import Categories.Category.Core 11 | open import Categories.Functor.Core using (Functor) 12 | open import Categories.NaturalTransformation.NaturalIsomorphism hiding (_≃_) 13 | open import Categories.NaturalTransformation.Equivalence 14 | open NaturalIsomorphism 15 | 16 | private 17 | variable 18 | o ℓ e o′ ℓ′ e′ : Level 19 | C D : Category o ℓ e 20 | 21 | infix 4 _≅_ 22 | _≅_ : ∀ {F G : Functor C D} → (α β : NaturalIsomorphism F G) → Set _ 23 | α ≅ β = F⇒G α ≃ F⇒G β × F⇐G α ≃ F⇐G β 24 | 25 | ≅-isEquivalence : ∀ {F G : Functor C D} → IsEquivalence (_≅_ {F = F} {G = G}) 26 | ≅-isEquivalence {D = D} {F = F} {G = G} = record 27 | { refl = H.refl , H.refl 28 | ; sym = map (λ z → H.sym z) (λ z → H.sym z) -- eta expansion needed 29 | ; trans = zip (λ a b → H.trans a b) λ a b → H.trans a b -- ditto 30 | } 31 | where module H = Category.Equiv D 32 | -------------------------------------------------------------------------------- /src/Categories/NaturalTransformation/NaturalIsomorphism/Functors.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Categories.NaturalTransformation.NaturalIsomorphism.Functors where 4 | 5 | open import Level 6 | 7 | open import Categories.Category 8 | open import Categories.Category.Construction.Functors 9 | open import Categories.Functor 10 | open import Categories.NaturalTransformation.NaturalIsomorphism 11 | 12 | import Categories.Morphism as Mor 13 | 14 | private 15 | variable 16 | o ℓ e : Level 17 | C D : Category o ℓ e 18 | 19 | -- isomorphism in Functors category is the same as natural isomorphism 20 | module _ {F G : Functor C D} where 21 | open Mor (Functors C D) 22 | 23 | Functors-iso⇒NI : F ≅ G → NaturalIsomorphism F G 24 | Functors-iso⇒NI F≅G = record 25 | { F⇒G = from 26 | ; F⇐G = to 27 | ; iso = λ X → record 28 | { isoˡ = isoˡ 29 | ; isoʳ = isoʳ 30 | } 31 | } 32 | where open Mor._≅_ F≅G 33 | 34 | NI⇒Functors-iso : NaturalIsomorphism F G → F ≅ G 35 | NI⇒Functors-iso α = record 36 | { from = F⇒G 37 | ; to = F⇐G 38 | ; iso = record 39 | { isoˡ = isoˡ (iso _) 40 | ; isoʳ = isoʳ (iso _) 41 | } 42 | } 43 | where open NaturalIsomorphism α 44 | open Mor.Iso 45 | -------------------------------------------------------------------------------- /src/Categories/Object/Cokernel.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | open import Categories.Object.Zero 5 | 6 | -- Cokernels of morphisms. 7 | -- https://ncatlab.org/nlab/show/cokernel 8 | module Categories.Object.Cokernel {o ℓ e} {𝒞 : Category o ℓ e} (𝟎 : Zero 𝒞) where 9 | 10 | open import Level 11 | 12 | open import Categories.Morphism 𝒞 13 | open import Categories.Morphism.Reasoning 𝒞 14 | hiding (glue) 15 | 16 | open Category 𝒞 17 | open Zero 𝟎 18 | 19 | open HomReasoning 20 | 21 | private 22 | variable 23 | A B X : Obj 24 | f h i j k : A ⇒ B 25 | 26 | record IsCokernel {A B K} (f : A ⇒ B) (k : B ⇒ K) : Set (o ⊔ ℓ ⊔ e) where 27 | field 28 | commute : k ∘ f ≈ zero⇒ 29 | universal : ∀ {X} {h : B ⇒ X} → h ∘ f ≈ zero⇒ → K ⇒ X 30 | factors : ∀ {eq : h ∘ f ≈ zero⇒} → h ≈ universal eq ∘ k 31 | unique : ∀ {eq : h ∘ f ≈ zero⇒} → h ≈ i ∘ k → i ≈ universal eq 32 | 33 | universal-resp-≈ : ∀ {eq : h ∘ f ≈ zero⇒} {eq′ : i ∘ f ≈ zero⇒} → 34 | h ≈ i → universal eq ≈ universal eq′ 35 | universal-resp-≈ h≈i = unique (⟺ h≈i ○ factors) 36 | 37 | universal-∘ : h ∘ k ∘ f ≈ zero⇒ 38 | universal-∘ {h = h} = begin 39 | h ∘ k ∘ f ≈⟨ refl⟩∘⟨ commute ⟩ 40 | h ∘ zero⇒ ≈⟨ zero-∘ˡ h ⟩ 41 | zero⇒ ∎ 42 | 43 | record Cokernel {A B} (f : A ⇒ B) : Set (o ⊔ ℓ ⊔ e) where 44 | field 45 | {cokernel} : Obj 46 | cokernel⇒ : B ⇒ cokernel 47 | isCokernel : IsCokernel f cokernel⇒ 48 | 49 | open IsCokernel isCokernel public 50 | 51 | IsCokernel⇒Cokernel : IsCokernel f k → Cokernel f 52 | IsCokernel⇒Cokernel {k = k} isCokernel = record 53 | { cokernel⇒ = k 54 | ; isCokernel = isCokernel 55 | } 56 | -------------------------------------------------------------------------------- /src/Categories/Object/Coproduct/Indexed.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Coproduct.Indexed {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Level 8 | 9 | open import Categories.Morphism.Reasoning.Core C 10 | 11 | open Category C 12 | open Equiv 13 | 14 | record IndexedCoproductOf {i} {I : Set i} (P : I → Obj) : Set (i ⊔ o ⊔ e ⊔ ℓ) where 15 | field 16 | X : Obj 17 | 18 | ι : ∀ i → P i ⇒ X 19 | [_] : ∀ {Y} → (∀ i → P i ⇒ Y) → X ⇒ Y 20 | 21 | commute : ∀ {Y} {f : ∀ i → P i ⇒ Y} {i} → [ f ] ∘ ι i ≈ f i 22 | unique : ∀ {Y} {h : X ⇒ Y} {f : ∀ i → P i ⇒ Y} → (∀ {i} → h ∘ ι i ≈ f i) → [ f ] ≈ h 23 | 24 | η : ∀ {Y} (h : X ⇒ Y) → [ (λ i → h ∘ ι i) ] ≈ h 25 | η h = unique refl 26 | 27 | ∘[] : ∀ {Y Z} (f : ∀ i → P i ⇒ Y) (g : Y ⇒ Z) → g ∘ [ f ] ≈ [ (λ i → g ∘ f i) ] 28 | ∘[] f g = sym (unique (pullʳ commute)) 29 | 30 | []-cong : ∀ {Y} {f g : ∀ i → P i ⇒ Y} → (∀ {i} → f i ≈ g i) → [ f ] ≈ [ g ] 31 | []-cong eq = unique (trans commute (sym eq)) 32 | 33 | unique′ : ∀ {Y} {h h′ : X ⇒ Y} → (∀ {i} → h′ ∘ ι i ≈ h ∘ ι i) → h′ ≈ h 34 | unique′ f = trans (sym (unique f)) (η _) 35 | 36 | AllCoproductsOf : ∀ i → Set (o ⊔ ℓ ⊔ e ⊔ suc i) 37 | AllCoproductsOf i = ∀ {I : Set i} (P : I → Obj) → IndexedCoproductOf P 38 | -------------------------------------------------------------------------------- /src/Categories/Object/Coproduct/Indexed/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Coproduct.Indexed.Properties {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Function.Base using () renaming (_∘_ to _∙_) 8 | open import Level 9 | 10 | open import Categories.Object.Coproduct.Indexed C using (IndexedCoproductOf; AllCoproductsOf) 11 | 12 | lowerAllCoproductsOf : ∀ {i} j → AllCoproductsOf (i ⊔ j) → AllCoproductsOf i 13 | lowerAllCoproductsOf j coprod P = record 14 | { X = X 15 | ; ι = ι ∙ lift 16 | ; [_] = λ f → [ f ∙ lower ] 17 | ; commute = commute 18 | ; unique = λ eq → unique eq 19 | } 20 | where open IndexedCoproductOf (coprod {Lift j _} (P ∙ lower)) 21 | -------------------------------------------------------------------------------- /src/Categories/Object/Group.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --safe --without-K #-} 2 | ------------------------------------------------------------------------ 3 | -- Group objects in a cartesian category. 4 | ------------------------------------------------------------------------ 5 | 6 | open import Categories.Category 7 | open import Categories.Category.Cartesian 8 | 9 | module Categories.Object.Group {o ℓ e} {𝒞 : Category o ℓ e} (C : Cartesian 𝒞) where 10 | 11 | open import Level 12 | 13 | open import Categories.Category.BinaryProducts 𝒞 using (BinaryProducts) 14 | open import Categories.Category.Cartesian.Monoidal 15 | open import Categories.Object.Monoid (CartesianMonoidal.monoidal C) 16 | open import Categories.Object.Terminal 𝒞 17 | 18 | open Category 𝒞 19 | open Cartesian C 20 | module Π = BinaryProducts products 21 | open BinaryProducts products using (_×_; _⁂_; ⟨_,_⟩) 22 | open Terminal terminal 23 | 24 | record IsGroup (G : Obj) : Set (ℓ ⊔ e) where 25 | -- any group object is also a monoid object 26 | field 27 | isMonoid : IsMonoid G 28 | 29 | open IsMonoid isMonoid public 30 | 31 | field 32 | -- inverse operation 33 | ι : G ⇒ G 34 | -- ι is in fact an inverse 35 | inverseˡ : η ∘ ! ≈ μ ∘ ⟨ ι , id ⟩ 36 | inverseʳ : η ∘ ! ≈ μ ∘ ⟨ id , ι ⟩ 37 | 38 | record Group : Set (o ⊔ ℓ ⊔ e) where 39 | field 40 | Carrier : Obj 41 | isGroup : IsGroup Carrier 42 | 43 | open IsGroup isGroup public 44 | 45 | monoid : Monoid 46 | monoid = record { isMonoid = isMonoid } 47 | 48 | open Group 49 | 50 | record Group⇒ (G H : Group) : Set (ℓ ⊔ e) where 51 | field 52 | arr : Carrier G ⇒ Carrier H 53 | preserves-μ : arr ∘ μ G ≈ μ H ∘ (arr ⁂ arr) 54 | preserves-η : arr ∘ η G ≈ η H 55 | preserves-ι : arr ∘ ι G ≈ ι H ∘ arr 56 | -------------------------------------------------------------------------------- /src/Categories/Object/Initial/Colimit.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Initial.Colimit {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Categories.Category.Construction.Cocones using (Cocone; Cocone⇒) 8 | open import Categories.Category.Instance.Zero using (Zero) 9 | open import Categories.Diagram.Colimit using (Colimit) 10 | open import Categories.Functor.Core using (Functor) 11 | open import Categories.Object.Initial C 12 | 13 | private 14 | open module C = Category C 15 | 16 | module _ {o′ ℓ′ e′} {F : Functor (Zero {o′} {ℓ′} {e′}) C} where 17 | colimit⇒⊥ : Colimit F → Initial 18 | colimit⇒⊥ L = record 19 | { ⊥ = coapex 20 | ; ⊥-is-initial = record 21 | { ! = rep record 22 | { coapex = record 23 | { ψ = λ () 24 | ; commute = λ { {()} } 25 | } 26 | } 27 | ; !-unique = λ f → initial.!-unique record 28 | { arr = f 29 | ; commute = λ { {()} } 30 | } 31 | } 32 | } 33 | where open Colimit L 34 | 35 | module _ {o′ ℓ′ e′} {F : Functor (Zero {o′} {ℓ′} {e′}) C} where 36 | 37 | ⊥⇒colimit : Initial → Colimit F 38 | ⊥⇒colimit t = record 39 | { initial = record 40 | { ⊥ = record 41 | { N = ⊥ 42 | ; coapex = record 43 | { ψ = λ () 44 | ; commute = λ { {()} } 45 | } 46 | } 47 | ; ⊥-is-initial = record 48 | { ! = λ {K} → 49 | let open Cocone F K 50 | in record 51 | { arr = ! 52 | ; commute = λ { {()} } 53 | } 54 | ; !-unique = λ f → 55 | let module f = Cocone⇒ F f 56 | in !-unique f.arr 57 | } 58 | } 59 | } 60 | where open Initial t 61 | -------------------------------------------------------------------------------- /src/Categories/Object/Monoid.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category.Core 3 | open import Categories.Category.Monoidal.Core 4 | 5 | module Categories.Object.Monoid {o ℓ e} {𝒞 : Category o ℓ e} (C : Monoidal 𝒞) where 6 | 7 | open import Level 8 | 9 | -- a monoid object is a generalization of the idea from algebra of a monoid, 10 | -- extended into any monoidal category 11 | 12 | open Category 𝒞 13 | open Monoidal C 14 | 15 | record IsMonoid (M : Obj) : Set (ℓ ⊔ e) where 16 | field 17 | μ : M ⊗₀ M ⇒ M 18 | η : unit ⇒ M 19 | 20 | field 21 | assoc : μ ∘ μ ⊗₁ id ≈ μ ∘ id ⊗₁ μ ∘ associator.from 22 | identityˡ : unitorˡ.from ≈ μ ∘ η ⊗₁ id 23 | identityʳ : unitorʳ.from ≈ μ ∘ id ⊗₁ η 24 | 25 | record Monoid : Set (o ⊔ ℓ ⊔ e) where 26 | field 27 | Carrier : Obj 28 | isMonoid : IsMonoid Carrier 29 | 30 | open IsMonoid isMonoid public 31 | 32 | open Monoid 33 | 34 | record Monoid⇒ (M M′ : Monoid) : Set (ℓ ⊔ e) where 35 | field 36 | arr : Carrier M ⇒ Carrier M′ 37 | preserves-μ : arr ∘ μ M ≈ μ M′ ∘ arr ⊗₁ arr 38 | preserves-η : arr ∘ η M ≈ η M′ 39 | -------------------------------------------------------------------------------- /src/Categories/Object/Product.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | open import Categories.Category 3 | 4 | --(Binary) product of objects and morphisms 5 | 6 | module Categories.Object.Product {o ℓ e} (C : Category o ℓ e) where 7 | 8 | open import Categories.Object.Product.Core C public 9 | open import Categories.Object.Product.Morphisms C public 10 | -------------------------------------------------------------------------------- /src/Categories/Object/Product/Indexed.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | -- this module characterizes a category of all products indexed by I. 6 | -- this notion formalizes a category with all products up to certain cardinal. 7 | module Categories.Object.Product.Indexed {o ℓ e} (C : Category o ℓ e) where 8 | 9 | open import Level 10 | 11 | open import Categories.Morphism.Reasoning C 12 | 13 | open Category C 14 | open Equiv 15 | open HomReasoning 16 | 17 | record IndexedProductOf {i} {I : Set i} (P : I → Obj) : Set (i ⊔ o ⊔ e ⊔ ℓ) where 18 | field 19 | -- the product 20 | X : Obj 21 | 22 | π : ∀ i → X ⇒ P i 23 | ⟨_⟩ : ∀ {Y} → (∀ i → Y ⇒ P i) → Y ⇒ X 24 | 25 | commute : ∀ {Y} {f : ∀ i → Y ⇒ P i} {i} → π i ∘ ⟨ f ⟩ ≈ f i 26 | unique : ∀ {Y} {h : Y ⇒ X} {f : ∀ i → Y ⇒ P i} → (∀ {i} → π i ∘ h ≈ f i) → ⟨ f ⟩ ≈ h 27 | 28 | η : ∀ {Y} (h : Y ⇒ X) → ⟨ (λ i → π i ∘ h) ⟩ ≈ h 29 | η h = unique refl 30 | 31 | ⟨⟩∘ : ∀ {Y Z} (f : ∀ i → Y ⇒ P i) (g : Z ⇒ Y) → ⟨ f ⟩ ∘ g ≈ ⟨ (λ i → f i ∘ g) ⟩ 32 | ⟨⟩∘ f g = ⟺ (unique (pullˡ commute)) 33 | 34 | ⟨⟩-cong : ∀ {Y} {f g : ∀ i → Y ⇒ P i} → (eq : ∀ {i} → f i ≈ g i) → ⟨ f ⟩ ≈ ⟨ g ⟩ 35 | ⟨⟩-cong eq = unique (trans commute (⟺ eq)) 36 | 37 | unique′ : ∀ {Y} {h h′ : Y ⇒ X} → (∀ {i} → π i ∘ h′ ≈ π i ∘ h) → h′ ≈ h 38 | unique′ f = trans (⟺ (unique f)) (η _) 39 | 40 | AllProductsOf : ∀ i → Set (o ⊔ ℓ ⊔ e ⊔ suc i) 41 | AllProductsOf i = ∀ {I : Set i} (P : I → Obj) → IndexedProductOf P 42 | -------------------------------------------------------------------------------- /src/Categories/Object/StrictInitial.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Level 4 | 5 | open import Categories.Category.Core using (Category) 6 | 7 | module Categories.Object.StrictInitial {o ℓ e} (C : Category o ℓ e) where 8 | 9 | open Category C 10 | open import Categories.Morphism C using (Epi; _≅_) 11 | open import Categories.Object.Initial C 12 | 13 | record IsStrictInitial (⊥ : Obj) : Set (o ⊔ ℓ ⊔ e) where 14 | field 15 | is-initial : IsInitial ⊥ 16 | open IsInitial is-initial public 17 | 18 | field 19 | is-strict : ∀ {A} → A ⇒ ⊥ → A ≅ ⊥ 20 | 21 | open IsStrictInitial 22 | 23 | record StrictInitial : Set (o ⊔ ℓ ⊔ e) where 24 | field 25 | ⊥ : Obj 26 | is-strict-initial : IsStrictInitial ⊥ 27 | 28 | initial : Initial 29 | initial .Initial.⊥ = ⊥ 30 | initial .Initial.⊥-is-initial = is-strict-initial .is-initial 31 | 32 | open Initial initial public 33 | -------------------------------------------------------------------------------- /src/Categories/Object/Terminal/Limit.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | module Categories.Object.Terminal.Limit {o ℓ e} (C : Category o ℓ e) where 6 | 7 | open import Categories.Category.Instance.Zero using (Zero) 8 | open import Categories.Diagram.Limit 9 | open import Categories.Functor.Core 10 | open import Categories.Object.Terminal C 11 | 12 | import Categories.Category.Construction.Cones as Co 13 | 14 | private 15 | module C = Category C 16 | open C 17 | 18 | module _ {o′ ℓ′ e′} {F : Functor (Zero {o′} {ℓ′} {e′}) C} where 19 | limit⇒⊤ : Limit F → Terminal 20 | limit⇒⊤ L = record 21 | { ⊤ = apex 22 | ; ⊤-is-terminal = record 23 | { ! = rep record 24 | { apex = record 25 | { ψ = λ () 26 | ; commute = λ { {()} } 27 | } 28 | } 29 | ; !-unique = λ f → terminal.!-unique record 30 | { arr = f 31 | ; commute = λ { {()} } 32 | } 33 | } 34 | } 35 | where open Limit L 36 | 37 | module _ {o′ ℓ′ e′} {F : Functor (Zero {o′} {ℓ′} {e′}) C} where 38 | ⊤⇒limit : Terminal → Limit F 39 | ⊤⇒limit t = record 40 | { terminal = record 41 | { ⊤ = record 42 | { N = ⊤ 43 | ; apex = record 44 | { ψ = λ () 45 | ; commute = λ { {()} } 46 | } 47 | } 48 | ; ⊤-is-terminal = record 49 | { ! = λ {K} → 50 | let open Co.Cone F K 51 | in record 52 | { arr = ! 53 | ; commute = λ { {()} } 54 | } 55 | ; !-unique = λ f → 56 | let module f = Co.Cone⇒ F f 57 | in !-unique f.arr 58 | } 59 | } 60 | } 61 | where open Terminal t 62 | -------------------------------------------------------------------------------- /src/Categories/Object/Zero.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Categories.Category 4 | 5 | -- a zero object is both terminal and initial. 6 | module Categories.Object.Zero {o ℓ e} (C : Category o ℓ e) where 7 | 8 | open import Level 9 | 10 | open import Categories.Object.Terminal C 11 | open import Categories.Object.Initial C 12 | 13 | open import Categories.Morphism C 14 | open import Categories.Morphism.Reasoning C 15 | 16 | open Category C 17 | open HomReasoning 18 | 19 | record IsZero (Z : Obj) : Set (o ⊔ ℓ ⊔ e) where 20 | field 21 | isInitial : IsInitial Z 22 | isTerminal : IsTerminal Z 23 | 24 | open IsInitial isInitial public 25 | renaming 26 | ( ! to ¡ 27 | ; !-unique to ¡-unique 28 | ; !-unique₂ to ¡-unique₂ 29 | ) 30 | open IsTerminal isTerminal public 31 | 32 | zero⇒ : ∀ {A B : Obj} → A ⇒ B 33 | zero⇒ = ¡ ∘ ! 34 | 35 | zero-∘ˡ : ∀ {X Y Z} → (f : Y ⇒ Z) → f ∘ zero⇒ {X} ≈ zero⇒ 36 | zero-∘ˡ f = pullˡ (⟺ (¡-unique (f ∘ ¡))) 37 | 38 | zero-∘ʳ : ∀ {X Y Z} → (f : X ⇒ Y) → zero⇒ {Y} {Z} ∘ f ≈ zero⇒ 39 | zero-∘ʳ f = pullʳ (⟺ (!-unique (! ∘ f))) 40 | 41 | record Zero : Set (o ⊔ ℓ ⊔ e) where 42 | field 43 | 𝟘 : Obj 44 | isZero : IsZero 𝟘 45 | 46 | open IsZero isZero public 47 | 48 | terminal : Terminal 49 | terminal = record { ⊤-is-terminal = isTerminal } 50 | 51 | initial : Initial 52 | initial = record { ⊥-is-initial = isInitial } 53 | 54 | open Zero 55 | 56 | ¡-Mono : ∀ {A} {z : Zero} → Mono (¡ z {A}) 57 | ¡-Mono {z = z} = from-⊤-is-Mono {t = terminal z} (¡ z) 58 | 59 | !-Epi : ∀ {A} {z : Zero} → Epi (! z {A}) 60 | !-Epi {z = z} = to-⊥-is-Epi {i = initial z} (! z) 61 | -------------------------------------------------------------------------------- /src/Categories/Theory/Lawvere/Instance/Identity.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- The 'Identity' instance, with all of Setoids as models 4 | 5 | module Categories.Theory.Lawvere.Instance.Identity where 6 | 7 | open import Data.Fin using (splitAt) 8 | open import Data.Sum using ([_,_]′) 9 | open import Data.Unit.Polymorphic using (⊤; tt) 10 | open import Level 11 | open import Relation.Binary.PropositionalEquality 12 | using (_≡_; refl; isEquivalence) 13 | 14 | open import Categories.Category.Cartesian.Bundle using (CartesianCategory) 15 | open import Categories.Category.Core using (Category) 16 | open import Categories.Category.Instance.Nat 17 | open import Categories.Category.Unbundled.Properties using (unpack′) 18 | open import Categories.Functor.IdentityOnObjects using (id-IOO) 19 | open import Categories.Object.Product using (Product) 20 | open import Categories.Theory.Lawvere using (LawvereTheory) 21 | 22 | Identity : LawvereTheory 0ℓ 0ℓ 23 | Identity = record 24 | { L = unpack′ (Category.op Nat) 25 | ; T = CartesianCategory.cartesian Natop-Cartesian 26 | ; I = id-IOO 27 | ; CartF = record 28 | { F-resp-⊤ = record { ! = λ () ; !-unique = λ _ () } 29 | ; F-resp-× = λ {m} {n} → record { P m n } 30 | } 31 | } 32 | where 33 | module P m n = Product Natop (Natop-Product m n) 34 | -------------------------------------------------------------------------------- /src/Categories/Theory/Lawvere/Instance/Triv.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- The 'Trivial' instance, with a single arrow between objects 4 | 5 | module Categories.Theory.Lawvere.Instance.Triv where 6 | 7 | open import Data.Nat using (_*_) 8 | open import Data.Unit.Polymorphic using (⊤; tt) 9 | open import Relation.Binary.PropositionalEquality 10 | using (_≡_; refl; isEquivalence) 11 | 12 | open import Categories.Theory.Lawvere using (LawvereTheory) 13 | 14 | Triv : ∀ {ℓ} → LawvereTheory ℓ ℓ 15 | Triv = record 16 | { L = record 17 | { _⇒_ = λ _ _ → ⊤ 18 | ; _≈_ = _≡_ 19 | ; assoc = refl 20 | ; sym-assoc = refl 21 | ; identityˡ = refl 22 | ; identityʳ = refl 23 | ; identity² = refl 24 | ; equiv = isEquivalence 25 | ; ∘-resp-≈ = λ _ _ → refl 26 | } 27 | ; T = record 28 | { terminal = record { ⊤ = 0 ; ⊤-is-terminal = record { ! = tt ; !-unique = λ _ → refl } } 29 | ; products = record 30 | { product = λ {m} {n} → record 31 | { A×B = m * n 32 | ; project₁ = refl 33 | ; project₂ = refl 34 | ; unique = λ _ _ → refl 35 | } 36 | } 37 | } 38 | ; I = record 39 | { F₁ = λ _ → tt 40 | ; identity = refl 41 | ; homomorphism = refl 42 | ; F-resp-≈ = λ _ → refl 43 | } 44 | ; CartF = record 45 | { F-resp-⊤ = record { ! = tt ; !-unique = λ _ → refl } 46 | ; F-resp-× = record { ⟨_,_⟩ = λ _ _ → tt ; project₁ = refl ; project₂ = refl ; unique = λ _ _ → refl } 47 | } 48 | } 49 | -------------------------------------------------------------------------------- /src/Categories/Utils/Product.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | module Categories.Utils.Product where 3 | 4 | open import Level 5 | open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂) 6 | open import Relation.Binary.PropositionalEquality 7 | 8 | -- "very dependent" versions of map and zipWith 9 | map⁎ : ∀ {a b p q} {A : Set a} {B : A → Set b} {P : A → Set p} {Q : {x : A} → P x → B x → Set q} → 10 | (f : (x : A) → B x) → (∀ {x} → (y : P x) → Q y (f x)) → (v : Σ A P) → Σ (B (proj₁ v)) (Q (proj₂ v)) 11 | map⁎ f g (x , y) = (f x , g y) 12 | 13 | map⁎′ : ∀ {a b p q} {A : Set a} {B : A → Set b} {P : Set p} {Q : P → Set q} → (f : (x : A) → B x) → ((x : P) → Q x) → (v : A × P) → B (proj₁ v) × Q (proj₂ v) 14 | map⁎′ f g (x , y) = (f x , g y) 15 | 16 | zipWith : ∀ {a b c p q r s} {A : Set a} {B : Set b} {C : Set c} {P : A → Set p} {Q : B → Set q} {R : C → Set r} {S : (x : C) → R x → Set s} (_∙_ : A → B → C) → (_∘_ : ∀ {x y} → P x → Q y → R (x ∙ y)) → (_*_ : (x : C) → (y : R x) → S x y) → (x : Σ A P) → (y : Σ B Q) → S (proj₁ x ∙ proj₁ y) (proj₂ x ∘ proj₂ y) 17 | zipWith _∙_ _∘_ _*_ (a , p) (b , q) = (a ∙ b) * (p ∘ q) 18 | syntax zipWith f g h = f -< h >- g 19 | -------------------------------------------------------------------------------- /src/Data/Quiver.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Data.Quiver where 4 | 5 | -- A Quiver, also known as a multidigraph, is the "underlying graph" of 6 | -- a category. Note how a Quiver has a *setoid* of edges. 7 | 8 | open import Level 9 | open import Relation.Binary using (Rel; IsEquivalence; Setoid) 10 | import Relation.Binary.Reasoning.Setoid as EqR 11 | 12 | -- a Quiver has vertices Obj and edges _⇒_, where edges form a setoid over _≈_. 13 | record Quiver o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where 14 | infix 4 _≈_ _⇒_ 15 | 16 | field 17 | Obj : Set o 18 | _⇒_ : Rel Obj ℓ 19 | _≈_ : ∀ {A B} → Rel (A ⇒ B) e 20 | equiv : ∀ {A B} → IsEquivalence (_≈_ {A} {B}) 21 | 22 | setoid : {A B : Obj} → Setoid _ _ 23 | setoid {A} {B} = record 24 | { Carrier = A ⇒ B 25 | ; _≈_ = _≈_ 26 | ; isEquivalence = equiv 27 | } 28 | 29 | module Equiv {A B : Obj} = IsEquivalence (equiv {A} {B}) 30 | module EdgeReasoning {A B : Obj} = EqR (setoid {A} {B}) 31 | -------------------------------------------------------------------------------- /src/Function/Construct/Setoid.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | -- was not ported from old function hierarchy 4 | 5 | module Function.Construct.Setoid where 6 | 7 | open import Function.Bundles using (Func; _⟨$⟩_) 8 | import Function.Construct.Composition as Comp 9 | open import Level using (Level) 10 | open import Relation.Binary.Bundles using (Setoid) 11 | 12 | 13 | private 14 | variable 15 | a₁ a₂ b₁ b₂ c₁ c₂ : Level 16 | 17 | setoid : Setoid a₁ a₂ → Setoid b₁ b₂ → Setoid _ _ 18 | setoid From To = record 19 | { Carrier = Func From To 20 | ; _≈_ = λ f g → ∀ {x} → f ⟨$⟩ x To.≈ g ⟨$⟩ x 21 | ; isEquivalence = record 22 | { refl = To.refl 23 | ; sym = λ f≈g → To.sym f≈g 24 | ; trans = λ f≈g g≈h → To.trans f≈g g≈h 25 | } 26 | } 27 | where 28 | module To = Setoid To 29 | 30 | -- This doesn't really belong here, it should be in Function.Construct.Composition but that's in stdlib 31 | -- so will need to be contributed to there first. 32 | infixr 9 _∙_ 33 | _∙_ : {A : Setoid a₁ a₂} {B : Setoid b₁ b₂} {C : Setoid c₁ c₂} → Func B C → Func A B → Func A C 34 | f ∙ g = Comp.function g f 35 | 36 | -------------------------------------------------------------------------------- /travis/everything.sh: -------------------------------------------------------------------------------- 1 | #/bin/sh 2 | 3 | find src/ -name '[^\.]*.agda' \ 4 | | sed -e 's|^src/[/]*|import |' -e 's|/|.|g' -e 's/.agda//' -e '/import Everything/d' \ 5 | | LC_COLLATE='C' sort \ 6 | > Everything.agda 7 | cat Everything.agda >> index.agda 8 | -------------------------------------------------------------------------------- /travis/index.agda: -------------------------------------------------------------------------------- 1 | module index where 2 | 3 | -- For a brief presentation of every single module, head over to 4 | import Everything 5 | 6 | -- Otherwise, here is an exhaustive, stern list of all the available modules. 7 | -------------------------------------------------------------------------------- /travis/index.sh: -------------------------------------------------------------------------------- 1 | set -eu 2 | set -o pipefail 3 | for file in $( find src -name "*.agda" | sort ); do 4 | if [[ ! $( head -n 10 $file | grep -m 1 "This module is DEPRECATED" ) ]]; then 5 | i=$( echo $file | sed 's/src\/\(.*\)\.agda/\1/' | sed 's/\//\./g' ) 6 | echo "import $i" >> index.agda; 7 | fi 8 | done 9 | --------------------------------------------------------------------------------