{-# OPTIONS --without-K --safe #-}
open import Categories.Category
module Categories.Morphism.IsoEquiv {o β e} (π : Category o β e) where
open import Level
open import Function using (flip; _on_)
open import Relation.Binary hiding (_β_)
import Relation.Binary.Construct.On as On
open import Categories.Morphism π
open Category π
private
variable
A B C : Obj
to-unique : β {fβ fβ : A β B} {gβ gβ} β
Iso fβ gβ β Iso fβ gβ β fβ β fβ β gβ β gβ
to-unique {_} {_} {fβ} {fβ} {gβ} {gβ} isoβ isoβ hyp = begin
gβ βΛβ¨ identityΛ‘ β©
id β gβ βΛβ¨ β-resp-βΛ‘ Isoβ.isoΛ‘ β©
(gβ β fβ) β gβ βΛβ¨ β-resp-βΛ‘ (β-resp-βΚ³ hyp) β©
(gβ β fβ) β gβ ββ¨ assoc β©
gβ β (fβ β gβ) ββ¨ β-resp-βΚ³ Isoβ.isoΚ³ β©
gβ β id ββ¨ identityΚ³ β©
gβ β
where
open HomReasoning
module Isoβ = Iso isoβ
module Isoβ = Iso isoβ
from-unique : β {fβ fβ : A β B} {gβ gβ} β
Iso fβ gβ β Iso fβ gβ β gβ β gβ β fβ β fβ
from-unique isoβ isoβ hyp = to-unique isoββ»ΒΉ isoββ»ΒΉ hyp
where
isoββ»ΒΉ = record { isoΛ‘ = Iso.isoΚ³ isoβ ; isoΚ³ = Iso.isoΛ‘ isoβ }
isoββ»ΒΉ = record { isoΛ‘ = Iso.isoΚ³ isoβ ; isoΚ³ = Iso.isoΛ‘ isoβ }
infix 4 _β_
record _β_ (i j : A β
B) : Set e where
constructor β_β
open _β
_
field from-β : from i β from j
to-β : to i β to j
to-β = to-unique (iso i) (iso j) from-β
open _β_
module _ {A B : Obj} where
open Equiv
β-refl : Reflexive (_β_ {A} {B})
β-refl = β refl β
β-sym : Symmetric (_β_ {A} {B})
β-sym = Ξ» where β eq β β β sym eq β
β-trans : Transitive (_β_ {A} {B})
β-trans = Ξ» where β eqβ β β eqβ β β β trans eqβ eqβ β
β-isEquivalence : IsEquivalence (_β_ {A} {B})
β-isEquivalence = record
{ refl = β-refl
; sym = β-sym
; trans = β-trans
}
β-setoid : β {A B : Obj} β Setoid _ _
β-setoid {A} {B} = record
{ Carrier = A β
B
; _β_ = _β_
; isEquivalence = β-isEquivalence
}
infix 4 _ββ²_
_ββ²_ : Rel (A β
B) e
_ββ²_ = _β_ on _β
_.from
ββ²-isEquivalence : IsEquivalence (_ββ²_ {A} {B})
ββ²-isEquivalence = On.isEquivalence _β
_.from equiv
ββ²-setoid : β {A B : Obj} β Setoid _ _
ββ²-setoid {A} {B} = record
{ Carrier = A β
B
; _β_ = _ββ²_
; isEquivalence = ββ²-isEquivalence
}
ββββ² : β {i j : A β
B} β i β j β i ββ² j
ββββ² eq = from-β eq
ββ²ββ : β {i j : A β
B} β i ββ² j β i β j
ββ²ββ {_} {_} {i} {j} eq = β eq β