12-constraintFamilies.agda 1.4 KB

1234567891011121314151617181920212223242526272829303132333435363738
  1. {-# OPTIONS --universe-polymorphism #-}
  2. -- {-# OPTIONS --verbose tc.constr.findInScope:15 #-}
  3. -- {-# OPTIONS --verbose tc.term.args.ifs:15 #-}
  4. -- {-# OPTIONS --verbose tc.checkArgs:15 #-}
  5. module InstanceArguments.12-constraintFamilies where
  6. open import Level
  7. open import Function
  8. open import Data.Unit
  9. open import Data.List
  10. open import Data.Product hiding (map)
  11. open import Relation.Binary
  12. open import Relation.Binary.PropositionalEquality
  13. record ConstrainedFunctor {c} (Constraint : Set → Set c)
  14. (f : (A : Set) → Set) : Set (suc c) where
  15. field
  16. fmap : ∀ {A B : Set} → {{cA : Constraint A}} → {{cB : Constraint B}} →
  17. (A → B) → f A → f B
  18. listConstrainedFunctor : ConstrainedFunctor (const ⊤) List
  19. listConstrainedFunctor = record { fmap = map }
  20. postulate
  21. sort : {A : Set} → {{ordA : ∃ λ (_<_ : A → A → Set) → IsStrictTotalOrder {A = A} _≡_ _<_}} →
  22. List A → List A
  23. ⋯ : ∀ {a} {A : Set a} {{a : A}} → A
  24. ⋯ {{a}} = a
  25. sortedListConstrainedFunctor : ConstrainedFunctor (λ A → ∃ λ (_<_ : A → A → Set) → IsStrictTotalOrder _≡_ _<_) List
  26. sortedListConstrainedFunctor = record { fmap = λ {{x}} {{y}} → map' {{x}} {{y}} } where
  27. map' : {A B : Set} → {{ordA : ∃ λ (_<_ : A → A → Set) → (IsStrictTotalOrder _≡_ _<_)}} →
  28. {{ordB : ∃ λ (_<_ : B → B → Set) → (IsStrictTotalOrder _≡_ _<_)}} →
  29. (A → B) → List A → List B
  30. map' f = sort ∘′ map f