cat/src/Cat/Category/Monoid.agda

51 lines
1.5 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module Cat.Category.Monoid where
open import Agda.Primitive
open import Cat.Category
open import Cat.Category.Product
open import Cat.Category.Functor
import Cat.Categories.Cat as Cat
-- TODO: Incorrect!
module _ (a b : Level) where
private
= lsuc (a b)
-- *If* the category of categories existed `_×_` would be equivalent to the
-- one brought into scope by doing:
--
-- open HasProducts (Cat.hasProducts unprovable) using (_×_)
--
-- Since it doesn't we'll make the following (definitionally equivalent) ad-hoc definition.
_×_ : {a b} Category a b Category a b Category a b
× 𝔻 = Cat.CatProduct.object 𝔻
record RawMonoidalCategory : Set where
field
category : Category a b
open Category category public
field
{{hasProducts}} : HasProducts category
empty : Object
-- aka. tensor product, monoidal product.
append : Functor (category × category) category
open HasProducts hasProducts public
record MonoidalCategory : Set where
field
raw : RawMonoidalCategory
open RawMonoidalCategory raw public
module _ {a b : Level} ( : MonoidalCategory a b) where
private
= a b
open MonoidalCategory public
record Monoid : Set where
field
carrier : Object
mempty : Arrow empty carrier
mappend : Arrow (carrier × carrier) carrier