open import 1Lab.Prelude hiding (_∘_ ; id ; _↪_)

open import Cat.Solver
open import Cat.Base

module Cat.Morphism {o h} (C : Precategory o h) where

Morphisms🔗

This module defines the three most important classes of morphisms that can be found in a category: monomorphisms, epimorphisms, and isomorphisms.

Monos🔗

A morphism is said to be monic when it is left-cancellable. A monomorphism from AA to BB, written ABA \mono B, is a monic morphism.

is-monic : Hom a b  Type _
is-monic {a = a} f =  {c}  (g h : Hom c a)  f  g  f  h  g  h

is-monic-is-prop :  {a b} (f : Hom a b)  is-prop (is-monic f)
is-monic-is-prop f x y i {c} g h p = Hom-set _ _ _ _ (x g h p) (y g h p) i

record _↪_ (a b : Ob) : Type (o  h) where
  field
    mor   : Hom a b
    monic : is-monic mor

Conversely, a morphism is said to be epic when it is right-cancellable. An epimorphism from AA to BB, written ABA \epi B, is an epic morphism.

Epis🔗

is-epic : Hom a b  Type _
is-epic {b = b} f =  {c}  (g h : Hom b c)  g  f  h  f  g  h

is-epic-is-prop :  {a b} (f : Hom a b)  is-prop (is-epic f)
is-epic-is-prop f x y i {c} g h p = Hom-set _ _ _ _ (x g h p) (y g h p) i

record _↠_ (a b : Ob) : Type (o  h) where
  field
    mor   : Hom a b
    monic : is-epic mor

Isos🔗

Maps f:ABf : A \to B and g:BAg : B \to A are inverses when we have fgf \circ g and gfg \circ f both equal to the identity. A map f:ABf : A \to B is invertible (or is an isomorphism) when there exists a gg for which ff and gg are inverses. An isomorphism ABA \cong B is a choice of map ABA \to B, together with a specified inverse.

record Inverses (f : Hom a b) (g : Hom b a) : Type h where
  field
    invl : f  g  id
    invr : g  f  id

open Inverses

record is-invertible (f : Hom a b) : Type (o  h) where
  field
    inv : Hom b a
    inverses : Inverses f inv

  open Inverses inverses public

  op : is-invertible inv
  op .inv = f
  op .inverses .Inverses.invl = invr inverses
  op .inverses .Inverses.invr = invl inverses

record _≅_ (a b : Ob) : Type (o  h) where
  field
    to       : Hom a b
    from     : Hom b a
    inverses : Inverses to from

  open Inverses inverses public

open _≅_ public

A given map is invertible in at most one way: If you have f:ABf : A \to B with two inverses g:BAg : B \to A and h:BAh : B \to A, then not only are g=hg = h equal, but the witnesses of these equalities are irrelevant.

Inverses-are-prop :  {f : Hom a b} {g : Hom b a}  is-prop (Inverses f g)
Inverses-are-prop x y i .Inverses.invl = Hom-set _ _ _ _ (x .invl) (y .invl) i
Inverses-are-prop x y i .Inverses.invr = Hom-set _ _ _ _ (x .invr) (y .invr) i

is-invertible-is-prop :  {f : Hom a b}  is-prop (is-invertible f)
is-invertible-is-prop {a = a} {b = b} {f = f} g h = p where
  module g = is-invertible g
  module h = is-invertible h

  g≡h : g.inv  h.inv
  g≡h =
    g.inv             ≡⟨ sym (idr _)  ap₂ _∘_ refl (sym h.invl) 
    g.inv  f  h.inv ≡⟨ assoc _ _ _ ·· ap₂ _∘_ g.invr refl ·· idl _ 
    h.inv             

  p : g  h
  p i .is-invertible.inv = g≡h i
  p i .is-invertible.inverses =
    is-prop→pathp  i  Inverses-are-prop {g = g≡h i}) g.inverses h.inverses i

We note that the identity morphism is always iso, and that isos compose:

id-iso : a  a
id-iso = make-iso id id (idl _) (idl _)

Inverses-∘ : {f : Hom a b} {f⁻¹ : Hom b a} {g : Hom b c} {g⁻¹ : Hom c b}
            Inverses f f⁻¹  Inverses g g⁻¹  Inverses (g  f) (f⁻¹  g⁻¹)
Inverses-∘ {f = f} {f⁻¹} {g} {g⁻¹} finv ginv = record { invl = l ; invr = r } where
  module finv = Inverses finv
  module ginv = Inverses ginv

  abstract
    l : (g  f)  f⁻¹  g⁻¹  id
    l = (g  f)  f⁻¹  g⁻¹ ≡⟨ solve C 
        g  (f  f⁻¹)  g⁻¹ ≡⟨  i  g  finv.invl i  g⁻¹) 
        g  id  g⁻¹        ≡⟨ solve C 
        g  g⁻¹             ≡⟨ ginv.invl 
        id                  

    r : (f⁻¹  g⁻¹)  g  f  id
    r = (f⁻¹  g⁻¹)  g  f ≡⟨ solve C 
        f⁻¹  (g⁻¹  g)  f ≡⟨  i  f⁻¹  ginv.invr i  f) 
        f⁻¹  id  f        ≡⟨ solve C 
        f⁻¹  f             ≡⟨ finv.invr 
        id                  

_∘Iso_ : a  b  b  c  a  c
(f ∘Iso g) .to = g .to  f .to
(f ∘Iso g) .from = f .from  g .from
(f ∘Iso g) .inverses = Inverses-∘ (f .inverses) (g .inverses)

_Iso⁻¹ : a  b  b  a
(f Iso⁻¹) .to = f .from
(f Iso⁻¹) .from = f .to
(f Iso⁻¹) .inverses .invl = f .inverses .invr
(f Iso⁻¹) .inverses .invr = f .inverses .invl

We also note that invertible morphisms are both epic and monic.

invertible→monic : is-invertible f  is-monic f
invertible→monic {f = f} invert g h p =
  g             ≡˘⟨ idl g ≡˘
  id  g        ≡˘⟨ ap (_∘ g) (is-invertible.invr invert) ≡˘
  (inv  f)  g ≡˘⟨ assoc inv f g ≡˘
  inv  (f  g) ≡⟨ ap (inv ∘_) p 
  inv  (f  h) ≡⟨ assoc inv f h 
  (inv  f)  h ≡⟨ ap (_∘ h) (is-invertible.invr invert) 
  id  h        ≡⟨ idl h 
  h 
  where
    open is-invertible invert

invertible→epic : is-invertible f  is-epic f
invertible→epic {f = f} invert g h p =
  g             ≡˘⟨ idr g ≡˘
  g  id        ≡˘⟨ ap (g ∘_) (is-invertible.invl invert) ≡˘
  g  (f  inv) ≡⟨ assoc g f inv 
  (g  f)  inv ≡⟨ ap (_∘ inv) p 
  (h  f)  inv ≡˘⟨ assoc h f inv ≡˘
  h  (f  inv) ≡⟨ ap (h  ∘_) (is-invertible.invl invert) 
  h  id        ≡⟨ idr h 
  h 
  where
    open is-invertible invert