aboutsummaryrefslogtreecommitdiff
path: root/Category/Instance/Preorder/Primitive/Preorders.agda
blob: 9832376947cee1b339f00f01b79ce19035a12889 (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
{-# OPTIONS --without-K --safe #-}

module Category.Instance.Preorder.Primitive.Preorders where

open import Categories.Category using (Category)
open import Categories.Category.Helper using (categoryHelper)
open import Function using (id; _∘_)
open import Level using (Level; _⊔_; suc)
open import Preorder.Primitive using (Preorder; module Isomorphism)
open import Preorder.Primitive.MonotoneMap using (MonotoneMap; _≃_; ≃-isEquivalence; module ≃)

-- The category of primitive preorders and monotone maps

private

  module _ {c  : Level} where

    identity : {A : Preorder c }  MonotoneMap A A
    identity = record
        { map = id
        ; mono = id
        }

    compose : {A B C : Preorder c }  MonotoneMap B C  MonotoneMap A B  MonotoneMap A C
    compose f g = record
        { map = f.map  g.map
        ; mono = f.mono  g.mono
        }
      where
        module f = MonotoneMap f
        module g = MonotoneMap g

    compose-resp-≃
        : {A B C : Preorder c }
          {f h : MonotoneMap B C}
          {g i : MonotoneMap A B}
         f  h
         g  i
         compose f g  compose h i
    compose-resp-≃ {C = C} {h = h} {g} f≃h g≃i x = ≅.trans (f≃h (map g x)) (map-resp-≅ h (g≃i x))
      where
        open MonotoneMap using (map; mono; map-resp-≅)
        open Preorder C
        open Isomorphism C

Preorders : (c  : Level)  Category (suc (c  )) (c  ) (c  )
Preorders c  = categoryHelper record
    { Obj = Preorder c     ; _⇒_ = MonotoneMap
    ; _≈_ = _≃_
    ; id = identity
    ; _∘_ = 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}
    }