2018-02-25 13:28:01 +00:00
|
|
|
|
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)
|
|
|
|
|
|
2018-03-05 10:13:37 +00:00
|
|
|
|
-- *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
|
2018-03-20 13:58:27 +00:00
|
|
|
|
ℂ × 𝔻 = Cat.CatProduct.object ℂ 𝔻
|
2018-02-25 13:28:01 +00:00
|
|
|
|
|
|
|
|
|
record RawMonoidalCategory : Set ℓ where
|
|
|
|
|
field
|
|
|
|
|
category : Category ℓa ℓb
|
|
|
|
|
open Category category public
|
|
|
|
|
field
|
|
|
|
|
{{hasProducts}} : HasProducts category
|
2018-03-05 12:52:41 +00:00
|
|
|
|
empty : Object
|
2018-02-25 13:28:01 +00:00
|
|
|
|
-- aka. tensor product, monoidal product.
|
2018-03-05 12:52:41 +00:00
|
|
|
|
append : Functor (category × category) category
|
|
|
|
|
open HasProducts hasProducts public
|
2018-02-25 13:28:01 +00:00
|
|
|
|
|
|
|
|
|
record MonoidalCategory : Set ℓ where
|
|
|
|
|
field
|
|
|
|
|
raw : RawMonoidalCategory
|
|
|
|
|
open RawMonoidalCategory raw public
|
|
|
|
|
|
|
|
|
|
module _ {ℓa ℓb : Level} (ℂ : MonoidalCategory ℓa ℓb) where
|
|
|
|
|
private
|
|
|
|
|
ℓ = ℓa ⊔ ℓb
|
|
|
|
|
|
2018-03-05 12:52:41 +00:00
|
|
|
|
open MonoidalCategory ℂ public
|
|
|
|
|
|
2018-02-25 13:28:01 +00:00
|
|
|
|
record Monoid : Set ℓ where
|
|
|
|
|
field
|
2018-03-05 12:52:41 +00:00
|
|
|
|
carrier : Object
|
|
|
|
|
mempty : Arrow empty carrier
|
|
|
|
|
mappend : Arrow (carrier × carrier) carrier
|