cat/src/Cat/Category.agda

602 lines
22 KiB
Agda
Raw Normal View History

2018-02-25 14:21:38 +00:00
-- | Univalent categories
--
-- This module defines:
--
-- Categories
-- ==========
--
-- Types
-- ------
--
-- Object, Arrow
--
-- Data
-- ----
-- identity; the identity arrow
2018-02-25 14:21:38 +00:00
-- _∘_; function composition
--
-- Laws
-- ----
--
-- associativity, identity, arrows form sets, univalence.
--
-- Lemmas
-- ------
--
-- Propositionality for all laws about the category.
2018-02-02 14:33:54 +00:00
{-# OPTIONS --allow-unsolved-metas --cubical #-}
2017-11-10 15:00:00 +00:00
module Cat.Category where
2017-11-10 15:00:00 +00:00
2018-03-21 13:39:56 +00:00
open import Cat.Prelude
2018-04-05 08:41:56 +00:00
import Function
2018-03-22 13:27:16 +00:00
------------------
2018-02-25 14:21:38 +00:00
-- * Categories --
2018-03-22 13:27:16 +00:00
------------------
2018-02-25 14:21:38 +00:00
-- | Raw categories
--
-- This record desribes the data that a category consist of as well as some laws
-- about these. The laws defined are the types the propositions - not the
-- witnesses to them!
record RawCategory (a b : Level) : Set (lsuc (a b)) where
no-eta-equality
field
Object : Set a
Arrow : Object Object Set b
identity : {A : Object} Arrow A A
_∘_ : {A B C : Object} Arrow B C Arrow A B Arrow A C
2018-03-06 14:52:22 +00:00
infixl 10 _∘_ _>>>_
2018-02-25 14:21:38 +00:00
-- | Operations on data
domain : {a b : Object} Arrow a b Object
domain {a} _ = a
codomain : {a b : Object} Arrow a b Object
codomain {b = b} _ = b
_>>>_ : {A B C : Object} (Arrow A B) (Arrow B C) Arrow A C
f >>> g = g f
2018-02-25 14:21:38 +00:00
-- | Laws about the data
2018-03-08 00:09:40 +00:00
-- FIXME It seems counter-intuitive that the normal-form is on the
-- right-hand-side.
IsAssociative : Set (a b)
IsAssociative = {A B C D} {f : Arrow A B} {g : Arrow B C} {h : Arrow C D}
h (g f) (h g) f
IsIdentity : ({A : Object} Arrow A A) Set (a b)
IsIdentity id = {A B : Object} {f : Arrow A B}
id f f × f id f
2018-02-23 09:35:42 +00:00
ArrowsAreSets : Set (a b)
ArrowsAreSets = {A B : Object} isSet (Arrow A B)
IsInverseOf : {A B} (Arrow A B) (Arrow B A) Set b
IsInverseOf = λ f g g f identity × f g identity
Isomorphism : {A B} (f : Arrow A B) Set b
Isomorphism {A} {B} f = Σ[ g Arrow B A ] IsInverseOf f g
_≅_ : (A B : Object) Set b
_≅_ A B = Σ[ f Arrow A B ] (Isomorphism f)
module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set b
Epimorphism {X} f = (g₀ g₁ : Arrow B X) g₀ f g₁ f g₀ g₁
Monomorphism : {X : Object} (f : Arrow A B) Set b
Monomorphism {X} f = (g₀ g₁ : Arrow X A) f g₀ f g₁ g₀ g₁
2018-02-21 11:59:31 +00:00
IsInitial : Object Set (a b)
IsInitial I = {X : Object} isContr (Arrow I X)
IsTerminal : Object Set (a b)
2018-02-20 17:15:07 +00:00
IsTerminal T = {X : Object} isContr (Arrow X T)
2018-02-21 11:59:31 +00:00
Initial : Set (a b)
Initial = Σ Object IsInitial
Terminal : Set (a b)
2018-02-21 11:59:31 +00:00
Terminal = Σ Object IsTerminal
-- | Univalence is indexed by a raw category as well as an identity proof.
module Univalence (isIdentity : IsIdentity identity) where
2018-03-22 13:27:16 +00:00
-- | The identity isomorphism
idIso : (A : Object) A A
idIso A = identity , identity , isIdentity
2018-03-22 13:27:16 +00:00
-- | Extract an isomorphism from an equality
--
-- [HoTT §9.1.4]
id-to-iso : (A B : Object) A B A B
id-to-iso A B eq = transp (\ i A eq i) (idIso A)
Univalent : Set (a b)
Univalent = {A B : Object} isEquiv (A B) (A B) (id-to-iso A B)
2018-03-22 13:27:16 +00:00
-- A perhaps more readable version of univalence:
Univalent≃ = {A B : Object} (A B) (A B)
-- | Equivalent formulation of univalence.
Univalent[Contr] : Set _
Univalent[Contr] = A isContr (Σ[ X Object ] A X)
2018-03-21 13:39:56 +00:00
2018-04-04 09:27:03 +00:00
private
module _ (A : Object)
2018-04-04 09:27:03 +00:00
-- It may be that we need something weaker than this, in that there
-- may be some other lemmas available to us.
-- For instance, `need0` should be available to us when we prove `need1`.
2018-04-05 08:41:56 +00:00
(need0 : (s : Σ Object (A ≅_)) (open Σ s renaming (fst to Y) using ()) A Y)
(need2 : (iso : A A)
2018-04-05 08:41:56 +00:00
(open Σ iso renaming (fst to f ; snd to iso-f))
(open Σ iso-f renaming (fst to f~ ; snd to areInv))
2018-04-04 10:01:29 +00:00
(identity , identity) (f , f~)
) where
2018-04-04 09:27:03 +00:00
c : Σ Object (A ≅_)
c = A , idIso A
module _ (y : Σ Object (A ≅_)) where
2018-04-05 08:41:56 +00:00
open Σ y renaming (fst to Y ; snd to isoY)
2018-04-04 09:27:03 +00:00
q : A Y
2018-04-04 10:01:29 +00:00
q = need0 y
2018-04-04 09:27:03 +00:00
-- Some error with primComp
isoAY : A Y
isoAY = {!id-to-iso A Y q!}
lem : PathP (λ i A q i) (idIso A) isoY
lem = d* isoAY
where
D : (Y : Object) (A Y) Set _
D Y p = (A≅Y : A Y) PathP (λ i A p i) (idIso A) A≅Y
d : D A refl
d A≅Y i = a0 i , a1 i , a2 i
where
2018-04-05 08:41:56 +00:00
open Σ A≅Y renaming (fst to f ; snd to iso-f)
open Σ iso-f renaming (fst to f~ ; snd to areInv)
2018-04-04 10:01:29 +00:00
aaa : (identity , identity) (f , f~)
aaa = need2 A≅Y
2018-04-04 09:27:03 +00:00
a0 : identity f
2018-04-04 10:01:29 +00:00
a0 i = fst (aaa i)
2018-04-04 09:27:03 +00:00
a1 : identity f~
2018-04-04 10:01:29 +00:00
a1 i = snd (aaa i)
2018-04-04 09:27:03 +00:00
-- we do have this!
2018-04-04 10:01:29 +00:00
-- I just need to rearrange the proofs a bit.
2018-04-04 09:27:03 +00:00
postulate
prop : {A B} (fg : Arrow A B × Arrow B A) isProp (IsInverseOf (fst fg) (snd fg))
a2 : PathP (λ i IsInverseOf (a0 i) (a1 i)) isIdentity areInv
2018-04-04 10:01:29 +00:00
a2 = lemPropF prop aaa
2018-04-04 09:27:03 +00:00
d* : D Y q
d* = pathJ D d Y q
p : (A , idIso A) (Y , isoY)
2018-04-04 10:01:29 +00:00
p i = q i , lem i
2018-04-04 09:27:03 +00:00
univ-lem : isContr (Σ Object (A ≅_))
univ-lem = c , p
univalence-lemma
: ( {A} (s : Σ Object (_≅_ A)) A fst s)
( {A} (iso : A A) (identity , identity) (fst iso , fst (snd iso)))
Univalent[Contr]
univalence-lemma s u A = univ-lem A s u
-- From: Thierry Coquand <Thierry.Coquand@cse.gu.se>
-- Date: Wed, Mar 21, 2018 at 3:12 PM
--
-- This is not so straight-forward so you can assume it
postulate from[Contr] : Univalent[Contr] Univalent
2018-02-25 14:21:38 +00:00
-- | The mere proposition of being a category.
--
-- Also defines a few lemmas:
--
-- iso-is-epi : Isomorphism f → Epimorphism {X = X} f
-- iso-is-mono : Isomorphism f → Monomorphism {X = X} f
--
-- Sans `univalent` this would be what is referred to as a pre-category in
-- [HoTT].
record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where
2018-02-25 13:37:28 +00:00
open RawCategory public
field
2018-02-23 11:43:49 +00:00
isAssociative : IsAssociative
isIdentity : IsIdentity identity
2018-02-23 11:51:44 +00:00
arrowsAreSets : ArrowsAreSets
open Univalence isIdentity public
field
univalent : Univalent
2018-02-25 14:21:38 +00:00
leftIdentity : {A B : Object} {f : Arrow A B} identity f f
leftIdentity {A} {B} {f} = fst (isIdentity {A = A} {B} {f})
rightIdentity : {A B : Object} {f : Arrow A B} f identity f
rightIdentity {A} {B} {f} = snd (isIdentity {A = A} {B} {f})
2018-03-22 13:27:16 +00:00
------------
-- Lemmas --
------------
-- | Relation between iso- epi- and mono- morphisms.
2018-02-25 13:44:03 +00:00
module _ {A B : Object} {X : Object} (f : Arrow A B) where
2018-03-22 13:27:16 +00:00
iso→epi : Isomorphism f Epimorphism {X = X} f
iso→epi (f- , left-inv , right-inv) g₀ g₁ eq = begin
g₀ ≡⟨ sym rightIdentity
g₀ identity ≡⟨ cong (_∘_ g₀) (sym right-inv)
2018-02-25 13:44:03 +00:00
g₀ (f f-) ≡⟨ isAssociative
(g₀ f) f- ≡⟨ cong (λ φ φ f-) eq
(g₁ f) f- ≡⟨ sym isAssociative
g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv
g₁ identity ≡⟨ rightIdentity
2018-02-25 13:44:03 +00:00
g₁
2018-03-22 13:27:16 +00:00
iso→mono : Isomorphism f Monomorphism {X = X} f
iso→mono (f- , left-inv , right-inv) g₀ g₁ eq =
2018-02-25 13:44:03 +00:00
begin
g₀ ≡⟨ sym leftIdentity
identity g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv)
2018-02-25 13:44:03 +00:00
(f- f) g₀ ≡⟨ sym isAssociative
f- (f g₀) ≡⟨ cong (_∘_ f-) eq
f- (f g₁) ≡⟨ isAssociative
(f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv
identity g₁ ≡⟨ leftIdentity
2018-02-25 13:44:03 +00:00
g₁
2018-03-22 13:27:16 +00:00
iso→epi×mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
iso→epi×mono iso = iso→epi iso , iso→mono iso
-- | The formulation of univalence expressed with _≃_ is trivially admissable -
-- just "forget" the equivalence.
univalent≃ : Univalent≃
univalent≃ = _ , univalent
module _ {A B : Object} where
open import Cat.Equivalence using (module Equiv)
iso-to-id : (A B) (A B)
iso-to-id = fst (Equiv≃.toIso _ _ univalent)
-- | All projections are propositions.
module Propositionality where
2018-02-20 16:59:48 +00:00
propIsAssociative : isProp IsAssociative
2018-02-23 11:51:44 +00:00
propIsAssociative x y i = arrowsAreSets _ _ x y i
2018-02-20 16:59:48 +00:00
propIsIdentity : {f : {A} Arrow A A} isProp (IsIdentity f)
propIsIdentity a b i
2018-02-23 11:51:44 +00:00
= arrowsAreSets _ _ (fst a) (fst b) i
, arrowsAreSets _ _ (snd a) (snd b) i
2018-02-20 16:59:48 +00:00
propArrowIsSet : isProp ( {A B} isSet (Arrow A B))
propArrowIsSet a b i = isSetIsProp a b i
propIsInverseOf : {A B f g} isProp (IsInverseOf {A} {B} f g)
propIsInverseOf x y = λ i
let
h : fst x fst y
2018-02-23 11:51:44 +00:00
h = arrowsAreSets _ _ (fst x) (fst y)
2018-02-20 16:59:48 +00:00
hh : snd x snd y
2018-02-23 11:51:44 +00:00
hh = arrowsAreSets _ _ (snd x) (snd y)
2018-02-20 16:59:48 +00:00
in h i , hh i
module _ {A B : Object} {f : Arrow A B} where
isoIsProp : isProp (Isomorphism f)
isoIsProp a@(g , η , ε) a'@(g' , η' , ε') =
lemSig (λ g propIsInverseOf) a a' geq
where
geq : g g'
geq = begin
g ≡⟨ sym rightIdentity
g identity ≡⟨ cong (λ φ g φ) (sym ε')
g (f g') ≡⟨ isAssociative
(g f) g' ≡⟨ cong (λ φ φ g') η
identity g' ≡⟨ leftIdentity
g'
2018-02-20 16:59:48 +00:00
propUnivalent : isProp Univalent
2018-03-22 13:27:16 +00:00
propUnivalent a b i = propPi (λ iso propIsContr) a b i
2018-02-20 16:59:48 +00:00
propIsTerminal : T isProp (IsTerminal T)
propIsTerminal T x y i {X} = res X i
where
module _ (X : Object) where
2018-04-05 08:41:56 +00:00
open Σ (x {X}) renaming (fst to fx ; snd to cx)
open Σ (y {X}) renaming (fst to fy ; snd to cy)
fp : fx fy
fp = cx fy
prop : (x : Arrow X T) isProp ( f x f)
prop x = propPi (λ y arrowsAreSets x y)
cp : (λ i f fp i f) [ cx cy ]
cp = lemPropF prop fp
res : (fx , cx) (fy , cy)
res i = fp i , cp i
2018-03-29 12:31:03 +00:00
-- | Terminal objects are propositional - a.k.a uniqueness of terminal
-- | objects.
--
-- Having two terminal objects induces an isomorphism between them - and
-- because of univalence this is equivalent to equality.
2018-03-28 22:07:49 +00:00
propTerminal : isProp Terminal
2018-03-29 12:26:47 +00:00
propTerminal Xt Yt = res
where
2018-04-05 08:41:56 +00:00
open Σ Xt renaming (fst to X ; snd to Xit)
open Σ Yt renaming (fst to Y ; snd to Yit)
open Σ (Xit {Y}) renaming (fst to Y→X) using ()
open Σ (Yit {X}) renaming (fst to X→Y) using ()
2018-03-29 12:26:47 +00:00
open import Cat.Equivalence hiding (_≅_)
-- Need to show `left` and `right`, what we know is that the arrows are
-- unique. Well, I know that if I compose these two arrows they must give
-- the identity, since also the identity is the unique such arrow (by X
-- and Y both being terminal objects.)
Xprop : isProp (Arrow X X)
Xprop f g = trans (sym (snd Xit f)) (snd Xit g)
Yprop : isProp (Arrow Y Y)
Yprop f g = trans (sym (snd Yit f)) (snd Yit g)
left : Y→X X→Y identity
2018-03-29 12:26:47 +00:00
left = Xprop _ _
right : X→Y Y→X identity
2018-03-29 12:26:47 +00:00
right = Yprop _ _
iso : X Y
iso = X→Y , Y→X , left , right
fromIso : X Y X Y
fromIso = fst (Equiv≃.toIso (X Y) (X Y) univalent)
p0 : X Y
p0 = fromIso iso
p1 : (λ i IsTerminal (p0 i)) [ Xit Yit ]
p1 = lemPropF propIsTerminal p0
res : Xt Yt
res i = p0 i , p1 i
2018-03-28 22:07:49 +00:00
-- Merely the dual of the above statement.
propIsInitial : I isProp (IsInitial I)
propIsInitial I x y i {X} = res X i
where
module _ (X : Object) where
2018-04-05 08:41:56 +00:00
open Σ (x {X}) renaming (fst to fx ; snd to cx)
open Σ (y {X}) renaming (fst to fy ; snd to cy)
fp : fx fy
fp = cx fy
prop : (x : Arrow I X) isProp ( f x f)
prop x = propPi (λ y arrowsAreSets x y)
cp : (λ i f fp i f) [ cx cy ]
cp = lemPropF prop fp
res : (fx , cx) (fy , cy)
res i = fp i , cp i
2018-03-28 22:07:49 +00:00
propInitial : isProp Initial
2018-04-04 09:27:03 +00:00
propInitial Xi Yi = res
2018-03-29 12:31:03 +00:00
where
2018-04-05 08:41:56 +00:00
open Σ Xi renaming (fst to X ; snd to Xii)
open Σ Yi renaming (fst to Y ; snd to Yii)
open Σ (Xii {Y}) renaming (fst to Y→X) using ()
open Σ (Yii {X}) renaming (fst to X→Y) using ()
2018-03-29 12:31:03 +00:00
open import Cat.Equivalence hiding (_≅_)
-- Need to show `left` and `right`, what we know is that the arrows are
-- unique. Well, I know that if I compose these two arrows they must give
-- the identity, since also the identity is the unique such arrow (by X
-- and Y both being terminal objects.)
Xprop : isProp (Arrow X X)
Xprop f g = trans (sym (snd Xii f)) (snd Xii g)
Yprop : isProp (Arrow Y Y)
Yprop f g = trans (sym (snd Yii f)) (snd Yii g)
left : Y→X X→Y identity
2018-03-29 12:31:03 +00:00
left = Yprop _ _
right : X→Y Y→X identity
2018-03-29 12:31:03 +00:00
right = Xprop _ _
iso : X Y
iso = Y→X , X→Y , right , left
fromIso : X Y X Y
fromIso = fst (Equiv≃.toIso (X Y) (X Y) univalent)
p0 : X Y
p0 = fromIso iso
p1 : (λ i IsInitial (p0 i)) [ Xii Yii ]
p1 = lemPropF propIsInitial p0
res : Xi Yi
res i = p0 i , p1 i
-- | Propositionality of being a category
module _ {a b : Level} ( : RawCategory a b) where
open RawCategory
open Univalence
2018-02-20 15:42:56 +00:00
private
module _ (x y : IsCategory ) where
2018-02-20 15:42:56 +00:00
module X = IsCategory x
module Y = IsCategory y
2018-02-20 16:59:48 +00:00
-- In a few places I use the result of propositionality of the various
-- projections of `IsCategory` - Here I arbitrarily chose to use this
2018-02-20 16:59:48 +00:00
-- result from `x : IsCategory C`. I don't know which (if any) possibly
-- adverse effects this may have.
module Prop = X.Propositionality
isIdentity : (λ _ IsIdentity identity) [ X.isIdentity Y.isIdentity ]
isIdentity = Prop.propIsIdentity X.isIdentity Y.isIdentity
U : {a : IsIdentity identity}
(λ _ IsIdentity identity) [ X.isIdentity a ]
2018-02-23 09:35:42 +00:00
(b : Univalent a)
Set _
U eqwal univ =
2018-02-23 09:35:42 +00:00
(λ i Univalent (eqwal i))
[ X.univalent univ ]
P : (y : IsIdentity identity)
(λ _ IsIdentity identity) [ X.isIdentity y ] Set _
P y eq = (univ : Univalent y) U eq univ
p : (b' : Univalent X.isIdentity)
2018-02-23 11:49:41 +00:00
(λ _ Univalent X.isIdentity) [ X.univalent b' ]
p univ = Prop.propUnivalent X.univalent univ
helper : P Y.isIdentity isIdentity
helper = pathJ P p Y.isIdentity isIdentity
2018-02-23 11:49:41 +00:00
eqUni : U isIdentity Y.univalent
eqUni = helper Y.univalent
done : x y
IsCategory.isAssociative (done i) = Prop.propIsAssociative X.isAssociative Y.isAssociative i
IsCategory.isIdentity (done i) = isIdentity i
IsCategory.arrowsAreSets (done i) = Prop.propArrowIsSet X.arrowsAreSets Y.arrowsAreSets i
IsCategory.univalent (done i) = eqUni i
propIsCategory : isProp (IsCategory )
2018-02-20 15:42:56 +00:00
propIsCategory = done
2018-02-25 14:21:38 +00:00
-- | Univalent categories
--
-- Just bundles up the data with witnesses inhabiting the propositions.
2018-02-05 13:47:15 +00:00
record Category (a b : Level) : Set (lsuc (a b)) where
field
raw : RawCategory a b
2018-02-05 13:47:15 +00:00
{{isCategory}} : IsCategory raw
2018-02-21 12:37:07 +00:00
open IsCategory isCategory public
-- The fact that being a category is a mere proposition gives rise to this
-- equality principle for categories.
module _ {a b : Level} { 𝔻 : Category a b} where
private
module = Category
module 𝔻 = Category 𝔻
module _ (rawEq : .raw 𝔻.raw) where
private
isCategoryEq : (λ i IsCategory (rawEq i)) [ .isCategory 𝔻.isCategory ]
isCategoryEq = lemPropF propIsCategory rawEq
Category≡ : 𝔻
Category≡ i = record
{ raw = rawEq i
; isCategory = isCategoryEq i
}
2018-03-08 00:09:40 +00:00
2018-02-25 14:21:38 +00:00
-- | Syntax for arrows- and composition in a given category.
2018-02-21 12:37:07 +00:00
module _ {a b : Level} ( : Category a b) where
open Category
_[_,_] : (A : Object) (B : Object) Set b
2018-02-20 15:25:49 +00:00
_[_,_] = Arrow
2017-11-10 15:00:00 +00:00
2018-02-20 15:25:49 +00:00
_[_∘_] : {A B C : Object} (g : Arrow B C) (f : Arrow A B) Arrow A C
_[_∘_] = _∘_
2017-11-10 15:00:00 +00:00
2018-02-25 14:21:38 +00:00
-- | The opposite category
--
-- The opposite category is the category where the direction of the arrows are
-- flipped.
module Opposite {a b : Level} where
module _ ( : Category a b) where
private
module = Category
2018-02-25 14:21:38 +00:00
opRaw : RawCategory a b
RawCategory.Object opRaw = .Object
RawCategory.Arrow opRaw = Function.flip .Arrow
RawCategory.identity opRaw = .identity
RawCategory._∘_ opRaw = Function.flip ._∘_
open RawCategory opRaw
2018-02-25 14:21:38 +00:00
isIdentity : IsIdentity identity
isIdentity = swap .isIdentity
open Univalence isIdentity
module _ {A B : .Object} where
open import Cat.Equivalence as Equivalence hiding (_≅_)
k : Equivalence.Isomorphism (.id-to-iso A B)
k = Equiv≃.toIso _ _ .univalent
2018-04-05 08:41:56 +00:00
open Σ k renaming (fst to f ; snd to inv)
open AreInverses inv
_⊙_ = Function._∘_
infixr 9 _⊙_
-- f : A .≅ B → A ≡ B
flipDem : A B A .≅ B
flipDem (f , g , inv) = g , f , inv
flopDem : A .≅ B A B
flopDem (f , g , inv) = g , f , inv
flipInv : {x} (flipDem flopDem) x x
flipInv = refl
-- Shouldn't be necessary to use `arrowsAreSets` here, but we have it,
-- so why not?
lem : (p : A B) id-to-iso A B p flopDem (.id-to-iso A B p)
lem p i = l≡r i
where
l = id-to-iso A B p
r = flopDem (.id-to-iso A B p)
2018-04-05 08:41:56 +00:00
open Σ l renaming (fst to l-obv ; snd to l-areInv)
open Σ l-areInv renaming (fst to l-invs ; snd to l-iso)
open Σ l-iso renaming (fst to l-l ; snd to l-r)
open Σ r renaming (fst to r-obv ; snd to r-areInv)
open Σ r-areInv renaming (fst to r-invs ; snd to r-iso)
open Σ r-iso renaming (fst to r-l ; snd to r-r)
l-obv≡r-obv : l-obv r-obv
l-obv≡r-obv = refl
l-invs≡r-invs : l-invs r-invs
l-invs≡r-invs = refl
l-l≡r-l : l-l r-l
l-l≡r-l = .arrowsAreSets _ _ l-l r-l
l-r≡r-r : l-r r-r
l-r≡r-r = .arrowsAreSets _ _ l-r r-r
l≡r : l r
l≡r i = l-obv≡r-obv i , l-invs≡r-invs i , l-l≡r-l i , l-r≡r-r i
ff : A B A B
ff = f flipDem
-- inv : AreInverses (.id-to-iso A B) f
invv : AreInverses (id-to-iso A B) ff
-- recto-verso : .id-to-iso A B ∘ f ≡ idFun (A .≅ B)
invv = record
{ verso-recto = funExt (λ x begin
(ff id-to-iso A B) x ≡⟨⟩
(f flipDem id-to-iso A B) x ≡⟨ cong (λ φ φ x) (cong (λ φ f flipDem φ) (funExt lem))
(f flipDem flopDem .id-to-iso A B) x ≡⟨⟩
(f .id-to-iso A B) x ≡⟨ (λ i verso-recto i x)
x )
; recto-verso = funExt (λ x begin
(id-to-iso A B f flipDem) x ≡⟨ cong (λ φ φ x) (cong (λ φ φ f flipDem) (funExt lem))
(flopDem .id-to-iso A B f flipDem) x ≡⟨ cong (λ φ φ x) (cong (λ φ flopDem φ flipDem) recto-verso)
(flopDem flipDem) x ≡⟨⟩
x )
}
h : Equivalence.Isomorphism (id-to-iso A B)
h = ff , invv
univalent : isEquiv (A B) (A B)
(Univalence.id-to-iso (swap .isIdentity) A B)
univalent = Equiv≃.fromIso _ _ h
isCategory : IsCategory opRaw
IsCategory.isAssociative isCategory = sym .isAssociative
IsCategory.isIdentity isCategory = isIdentity
IsCategory.arrowsAreSets isCategory = .arrowsAreSets
IsCategory.univalent isCategory = univalent
2018-02-25 14:21:38 +00:00
opposite : Category a b
Category.raw opposite = opRaw
Category.isCategory opposite = isCategory
2018-02-25 14:21:38 +00:00
-- 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
2018-02-25 14:21:38 +00:00
RawCategory._∘_ (rawInv _) = _∘_
oppositeIsInvolution : opposite (opposite )
2018-03-08 00:09:40 +00:00
oppositeIsInvolution = Category≡ rawInv
2018-02-25 14:21:38 +00:00
2018-02-25 14:23:33 +00:00
open Opposite public