Skip to content

Commit

Permalink
Contravariance of categories of algebras (#767)
Browse files Browse the repository at this point in the history
* Define IsMonad.

* HomFunctor.

* bind as a natural transformation.

* Functor algebras (WIP).

* Functor algebras.

* Eilenberg-Moore category.

* Variance of category of algebras.

* No eta-equality for Category.

* Shorten long lines.

* Reenable eta-equality for categories.

* Uncomment involutiveOp.
  • Loading branch information
anuyts committed May 3, 2022
1 parent 5e558a7 commit 258821e
Showing 1 changed file with 27 additions and 0 deletions.
27 changes: 27 additions & 0 deletions Cubical/Categories/Instances/FunctorAlgebras.agda
Original file line number Diff line number Diff line change
Expand Up @@ -9,6 +9,7 @@ open import Cubical.Foundations.Univalence

open import Cubical.Categories.Category
open import Cubical.Categories.Functor
open import Cubical.Categories.NaturalTransformation.Base

open _≅_

Expand Down Expand Up @@ -106,3 +107,29 @@ module _ {C : Category ℓC ℓC'} (F : Functor C C) where
F-hom ForgetAlgebra = carrierHom
F-id ForgetAlgebra = refl
F-seq ForgetAlgebra algF algG = refl

module _ {C : Category ℓC ℓC'} {F G : Functor C C} (τ : NatTrans F G) where

private
module C = Category C
open Functor
open NatTrans
open AlgebraHom

AlgebrasFunctor : Functor (AlgebrasCategory G) (AlgebrasCategory F)
F-ob AlgebrasFunctor (algebra a α) = algebra a (α C.∘ N-ob τ a)
carrierHom (F-hom AlgebrasFunctor (algebraHom f isalgF)) = f
strHom (F-hom AlgebrasFunctor {algebra a α} {algebra b β} (algebraHom f isalgF)) =
(N-ob τ a C.⋆ α) C.⋆ f
≡⟨ C.⋆Assoc (N-ob τ a) α f ⟩
N-ob τ a C.⋆ (α C.⋆ f)
≡⟨ cong (N-ob τ a C.⋆_) isalgF ⟩
N-ob τ a C.⋆ (F-hom G f C.⋆ β)
≡⟨ sym (C.⋆Assoc _ _ _) ⟩
(N-ob τ a C.⋆ F-hom G f) C.⋆ β
≡⟨ cong (C._⋆ β) (sym (N-hom τ f)) ⟩
(F-hom F f C.⋆ N-ob τ b) C.⋆ β
≡⟨ C.⋆Assoc _ _ _ ⟩
F-hom F f C.⋆ (N-ob τ b C.⋆ β) ∎
F-id AlgebrasFunctor = AlgebraHom≡ F refl
F-seq AlgebrasFunctor algΦ algΨ = AlgebraHom≡ F refl

0 comments on commit 258821e

Please sign in to comment.