Parameterised.agda 1.7 KB

1234567891011121314151617181920212223242526272829303132333435363738394041424344454647484950515253545556575859606162636465
  1. -- This module introduces parameterised modules.
  2. module Introduction.Modules.Parameterised where
  3. -- First some familiar datatypes.
  4. data Bool : Set where
  5. false : Bool
  6. true : Bool
  7. data List (A : Set) : Set where
  8. nil : List A
  9. _::_ : A -> List A -> List A
  10. infixr 15 _::_ -- see 'Introduction.Operators' for information on infix
  11. -- declarations
  12. -- Agda supports parameterised modules. A parameterised module is declared by
  13. -- giving the parameters after the module name.
  14. module Sorting {A : Set}(_<_ : A -> A -> Bool) where
  15. insert : A -> List A -> List A
  16. insert x nil = x :: nil
  17. insert x (y :: ys) = ins (x < y)
  18. where
  19. ins : Bool -> List A -- local functions can do pattern matching and
  20. ins true = x :: y :: ys -- be recursive
  21. ins false = y :: insert x ys
  22. sort : List A -> List A
  23. sort nil = nil
  24. sort (x :: xs) = insert x (sort xs)
  25. -- Before a parameterised module can be used it has to be instantiated. So, we
  26. -- need something to instantiate it with.
  27. data Nat : Set where
  28. zero : Nat
  29. suc : Nat -> Nat
  30. _<_ : Nat -> Nat -> Bool
  31. zero < zero = false
  32. zero < suc _ = true
  33. suc _ < zero = false
  34. suc n < suc m = n < m
  35. -- To instantiate a module you define a new module in terms of the
  36. -- parameterised module. Module instantiation also supports the using, hiding
  37. -- and renaming modifiers.
  38. module SortNat = Sorting _<_
  39. sort' : {A : Set}(_<_ : A -> A -> Bool) -> List A -> List A
  40. sort' less = Sort'.sort
  41. where
  42. module Sort' = Sorting less
  43. -- Now the instantiated module can be opened and we can use the sorting
  44. -- function.
  45. open SortNat
  46. test = sort (suc zero :: zero :: suc (suc zero) :: nil)