| 1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
 | {-# OPTIONS --without-K --safe #-}
open import Level using (Level)
open import Category.Cocomplete.Finitely.Bundle using (FinitelyCocompleteCategory)
module Functor.Exact.Instance.Swap {o β e : Level} (π π : FinitelyCocompleteCategory o β e) where
open import Categories.Category using (_[_,_])
open import Categories.Category.BinaryProducts using (BinaryProducts)
open import Categories.Category.Product using (Product) renaming (Swap to Swapβ²)
open import Categories.Category.Cartesian using (Cartesian)
open import Categories.Diagram.Coequalizer using (IsCoequalizer)
open import Categories.Object.Initial using (IsInitial)
open import Categories.Object.Coproduct using (IsCoproduct)
open import Data.Product.Base using (_,_; projβ; projβ; swap)
open import Category.Instance.FinitelyCocompletes {o} {β} {e} using (FinitelyCocompletes-Cartesian)
open import Functor.Exact using (RightExactFunctor)
module FCC = Cartesian FinitelyCocompletes-Cartesian
open BinaryProducts (FCC.products) using (_Γ_) -- ; Οβ; Οβ; _β_; assocΛ‘)
module π = FinitelyCocompleteCategory π
module π = FinitelyCocompleteCategory π
swap-resp-β₯ : {A : π.Obj} {B : π.Obj} β IsInitial (Product π.U π.U) (A , B) β IsInitial (Product π.U π.U) (B , A)
swap-resp-β₯ {A} {B} isInitial = record
    { ! = swap !
    ; !-unique = Ξ» { (f , g) β swap (!-unique (g , f)) }
    }
  where
    open IsInitial isInitial
swap-resp-+
    : {Aβ Bβ A+Bβ : π.Obj}
    β {Aβ Bβ A+Bβ : π.Obj}
    β {iββ : π.U [ Aβ , A+Bβ ]}
    β {iββ : π.U [ Bβ , A+Bβ ]}
    β {iββ : π.U [ Aβ , A+Bβ ]}
    β {iββ : π.U [ Bβ , A+Bβ ]}
    β IsCoproduct (Product π.U π.U) (iββ , iββ) (iββ , iββ)
    β IsCoproduct (Product π.U π.U) (iββ , iββ) (iββ , iββ)
swap-resp-+ {Aβ} {Bβ} {A+Bβ} {Aβ} {Bβ} {A+Bβ} {iββ} {iββ} {iββ} {iββ} isCoproduct = record
    { [_,_] = Ξ» { X Y β swap ([ swap X , swap Y ]) }
    ; injectβ = swap injectβ
    ; injectβ = swap injectβ
    ; unique = Ξ» { ββ ββ β swap (unique (swap ββ) (swap ββ)) }
    }
  where
    open IsCoproduct isCoproduct
swap-resp-coeq
    : {Aβ Bβ Cβ : π.Obj}Β 
      {Aβ Bβ Cβ : π.Obj}
      {fβ gβ : π.U [ Aβ , Bβ ]}
      {hβ : π.U [ Bβ , Cβ ]}
      {fβ gβ : π.U [ Aβ , Bβ ]}
      {hβ : π.U [ Bβ , Cβ ]}
    β IsCoequalizer (Product π.U π.U) (fβ , fβ) (gβ , gβ) (hβ , hβ)
    β IsCoequalizer (Product π.U π.U) (fβ , fβ) (gβ , gβ) (hβ , hβ)
swap-resp-coeq isCoequalizer = record
    { equality = swap equality
    ; coequalize = Ξ» { x β swap (coequalize (swap x)) }
    ; universal = swap universal
    ; unique = Ξ» { x β swap (unique (swap x)) }
    }
  where
    open IsCoequalizer isCoequalizer
Swap : RightExactFunctor (π Γ π) (π Γ π)
Swap = record
    { F = Swapβ²
    ; isRightExact = record
        { F-resp-β₯ = swap-resp-β₯
        ; F-resp-+ = swap-resp-+
        ; F-resp-coeq = swap-resp-coeq
        }
    }
 |