aboutsummaryrefslogtreecommitdiff
path: root/Category/Instance/MonoidalPreorders/Primitive.agda
blob: d00e17a3a2b3025be84108855532aa9e54eae389 (plain)
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
80
81
82
83
84
85
86
{-# OPTIONS --without-K --safe #-}

module Category.Instance.MonoidalPreorders.Primitive where

import Preorder.Primitive.MonotoneMap as MonotoneMap using (_≃_; module ≃)

open import Categories.Category using (Category)
open import Categories.Category.Helper using (categoryHelper)
open import Category.Instance.Preorders.Primitive using (Preorders)
open import Level using (Level; suc; _⊔_)
open import Preorder.Primitive.Monoidal using (MonoidalPreorder; MonoidalMonotone)
open import Relation.Binary using (IsEquivalence)

module _ {c₁ c₂ ℓ₁ ℓ₂ : Level} {A : MonoidalPreorder c₁ ℓ₁} {B : MonoidalPreorder c₂ ℓ₂} where

  -- Pointwise equality of monoidal monotone maps

  open MonoidalMonotone using (F)

  _≃_ : (f g : MonoidalMonotone A B)  Set (c₁  ℓ₂)
  f  g = F f MonotoneMap.≃ F g

  infix 4 _≃_

  ≃-isEquivalence : IsEquivalence _≃_
  ≃-isEquivalence = let open MonotoneMap.≃ in record
      { refl = λ {x}  refl {x = F x}
      ; sym = λ {f g}  sym {x = F f} {y = F g}
      ; trans = λ {f g h}  trans {i = F f} {j = F g} {k = F h}
      }

  module ≃ = IsEquivalence ≃-isEquivalence

private

  identity : {c  : Level} (A : MonoidalPreorder c )  MonoidalMonotone A A
  identity A = let open MonoidalPreorder A in record
      { F = Category.id (Preorders _ _)
      ; ε = refl
      ; ⊗-homo = λ p₁ p₂  refl {p₁  p₂}
      }

  compose
      : {c  : Level}
        {P Q R : MonoidalPreorder c }
       MonoidalMonotone Q R
       MonoidalMonotone P Q
       MonoidalMonotone P R
  compose {R = R} G F = record
      { F = let open Category (Preorders _ _) in G.F  F.F
      ; ε = trans G.ε (G.mono F.ε)
      ; ⊗-homo = λ p₁ p₂  trans (G.⊗-homo (F.map p₁) (F.map p₂)) (G.mono (F.⊗-homo p₁ p₂))
      }
    where
      module F = MonoidalMonotone F
      module G = MonoidalMonotone G
      open MonoidalPreorder R

  compose-resp-≃
      : {c  : Level}
        {A B C : MonoidalPreorder c }
        {f h : MonoidalMonotone B C}
        {g i : MonoidalMonotone A B}
       f  h
       g  i
       compose f g  compose h i
  compose-resp-≃ {C = C} {f = f} {g} {h} {i} = ∘-resp-≈ {f = F f} {F g} {F h} {F i}
    where
      open Category (Preorders _ _)
      open MonoidalMonotone using (F)

MonoidalPreorders : (c  : Level)  Category (suc (c  )) (c  ) (c  )
MonoidalPreorders c  = categoryHelper record
    { Obj = MonoidalPreorder c     ; _⇒_ = MonoidalMonotone
    ; _≈_ = _≃_
    ; id  = λ {A}  identity A
    ; _∘_ = compose
    ; assoc = λ {f = f} {g h}  ≃.refl {x = compose (compose h g) f}
    ; identityˡ = λ {f = f}  ≃.refl {x = f}
    ; identityʳ = λ {f = f}  ≃.refl {x = f}
    ; equiv = ≃-isEquivalence
    ; ∘-resp-≈ = λ {f = f} {g h i}  compose-resp-≃ {f = f} {g} {h} {i}
    }
  where
    open MonoidalMonotone using (F)