UniversePolymorphicFunctor.agda 2.5 KB

1234567891011121314151617181920212223242526272829303132333435363738394041424344454647484950515253545556575859606162636465666768697071727374757677
  1. {-# OPTIONS --universe-polymorphism #-}
  2. module UniversePolymorphicFunctor where
  3. open import Agda.Primitive renaming (lsuc to suc)
  4. record IsEquivalence {a ℓ} {A : Set a}
  5. (_≈_ : A → A → Set ℓ) : Set (a ⊔ ℓ) where
  6. field
  7. refl : ∀ {x} → x ≈ x
  8. sym : ∀ {i j} → i ≈ j → j ≈ i
  9. trans : ∀ {i j k} → i ≈ j → j ≈ k → i ≈ k
  10. record Setoid c ℓ : Set (suc (c ⊔ ℓ)) where
  11. infix 4 _≈_
  12. field
  13. Carrier : Set c
  14. _≈_ : Carrier → Carrier → Set ℓ
  15. isEquivalence : IsEquivalence _≈_
  16. open IsEquivalence isEquivalence public
  17. infixr 0 _⟶_
  18. record _⟶_ {f₁ f₂ t₁ t₂}
  19. (From : Setoid f₁ f₂) (To : Setoid t₁ t₂) :
  20. Set (f₁ ⊔ f₂ ⊔ t₁ ⊔ t₂) where
  21. infixl 5 _⟨$⟩_
  22. field
  23. _⟨$⟩_ : Setoid.Carrier From → Setoid.Carrier To
  24. cong : ∀ {x y} →
  25. Setoid._≈_ From x y → Setoid._≈_ To (_⟨$⟩_ x) (_⟨$⟩_ y)
  26. open _⟶_ public
  27. id : ∀ {a₁ a₂} {A : Setoid a₁ a₂} → A ⟶ A
  28. id = record { _⟨$⟩_ = λ x → x; cong = λ x≈y → x≈y }
  29. infixr 9 _∘_
  30. _∘_ : ∀ {a₁ a₂} {A : Setoid a₁ a₂}
  31. {b₁ b₂} {B : Setoid b₁ b₂}
  32. {c₁ c₂} {C : Setoid c₁ c₂} →
  33. B ⟶ C → A ⟶ B → A ⟶ C
  34. f ∘ g = record
  35. { _⟨$⟩_ = λ x → f ⟨$⟩ (g ⟨$⟩ x)
  36. ; cong = λ x≈y → cong f (cong g x≈y)
  37. }
  38. _⇨_ : ∀ {f₁ f₂ t₁ t₂} → Setoid f₁ f₂ → Setoid t₁ t₂ → Setoid _ _
  39. From ⇨ To = record
  40. { Carrier = From ⟶ To
  41. ; _≈_ = λ f g → ∀ {x y} → x ≈₁ y → f ⟨$⟩ x ≈₂ g ⟨$⟩ y
  42. ; isEquivalence = record
  43. { refl = λ {f} → cong f
  44. ; sym = λ f∼g x∼y → To.sym (f∼g (From.sym x∼y))
  45. ; trans = λ f∼g g∼h x∼y → To.trans (f∼g From.refl) (g∼h x∼y)
  46. }
  47. }
  48. where
  49. open module From = Setoid From using () renaming (_≈_ to _≈₁_)
  50. open module To = Setoid To using () renaming (_≈_ to _≈₂_)
  51. record Functor {f₁ f₂ f₃ f₄}
  52. (F : Setoid f₁ f₂ → Setoid f₃ f₄) :
  53. Set (suc (f₁ ⊔ f₂) ⊔ f₃ ⊔ f₄) where
  54. field
  55. map : ∀ {A B} → (A ⇨ B) ⟶ (F A ⇨ F B)
  56. identity : ∀ {A} →
  57. let open Setoid (F A ⇨ F A) in
  58. map ⟨$⟩ id ≈ id
  59. composition : ∀ {A B C} (f : B ⟶ C) (g : A ⟶ B) →
  60. let open Setoid (F A ⇨ F C) in
  61. map ⟨$⟩ (f ∘ g) ≈ (map ⟨$⟩ f) ∘ (map ⟨$⟩ g)