cat/src/Cat/Categories/Opposite.agda

97 lines
3.6 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.

{-# OPTIONS --cubical #-}
module Cat.Categories.Opposite where
open import Cat.Prelude
open import Cat.Equivalence
open import Cat.Category
-- | The opposite category
--
-- The opposite category is the category where the direction of the arrows are
-- flipped.
module _ {a b : Level} where
module _ ( : Category a b) where
private
module _ where
module = Category
opRaw : RawCategory a b
RawCategory.Object opRaw = .Object
RawCategory.Arrow opRaw = flip .Arrow
RawCategory.identity opRaw = .identity
RawCategory._<<<_ opRaw = ._>>>_
open RawCategory opRaw
isPreCategory : IsPreCategory opRaw
IsPreCategory.isAssociative isPreCategory = sym .isAssociative
IsPreCategory.isIdentity isPreCategory = swap .isIdentity
IsPreCategory.arrowsAreSets isPreCategory = .arrowsAreSets
open IsPreCategory isPreCategory
module _ {A B : .Object} where
open Σ (toIso _ _ (.univalent {A} {B}))
renaming (fst to idToIso* ; snd to inv*)
open AreInverses {f = .idToIso A B} {idToIso*} inv*
shuffle : A B A .≊ B
shuffle (f , g , inv) = g , f , inv
shuffle~ : A .≊ B A B
shuffle~ (f , g , inv) = g , f , inv
lem : (p : A B) idToIso A B p shuffle~ (.idToIso A B p)
lem p = isoEq refl
isoToId* : A B A B
isoToId* = idToIso* shuffle
inv : AreInverses (idToIso A B) isoToId*
inv =
( funExt (λ x begin
(isoToId* idToIso A B) x
≡⟨⟩
(idToIso* shuffle idToIso A B) x
≡⟨ cong (λ φ φ x) (cong (λ φ idToIso* shuffle φ) (funExt lem))
(idToIso* shuffle shuffle~ .idToIso A B) x
≡⟨⟩
(idToIso* .idToIso A B) x
≡⟨ (λ i verso-recto i x)
x )
, funExt (λ x begin
(idToIso A B idToIso* shuffle) x
≡⟨ cong (λ φ φ x) (cong (λ φ φ idToIso* shuffle) (funExt lem))
(shuffle~ .idToIso A B idToIso* shuffle) x
≡⟨ cong (λ φ φ x) (cong (λ φ shuffle~ φ shuffle) recto-verso)
(shuffle~ shuffle) x
≡⟨⟩
x )
)
isCategory : IsCategory opRaw
IsCategory.isPreCategory isCategory = isPreCategory
IsCategory.univalent isCategory
= univalenceFromIsomorphism (isoToId* , inv)
opposite : Category a b
Category.raw opposite = opRaw
Category.isCategory opposite = isCategory
-- As demonstrated here a side-effect of having no-eta-equality on constructors
-- means that we need to pick things apart to show that things are indeed
-- definitionally equal. I.e; a thing that would normally be provable in one
-- line now takes 13!! Admittedly it's a simple proof.
module _ { : Category a b} where
open Category
private
-- Since they really are definitionally equal we just need to pick apart
-- the data-type.
rawInv : Category.raw (opposite (opposite )) raw
RawCategory.Object (rawInv _) = Object
RawCategory.Arrow (rawInv _) = Arrow
RawCategory.identity (rawInv _) = identity
RawCategory._<<<_ (rawInv _) = _<<<_
oppositeIsInvolution : opposite (opposite )
oppositeIsInvolution = Category≡ rawInv