Modules.agda 2.9 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146
  1. {-
  2. Types Summer School 2007
  3. Bertinoro
  4. Aug 19 - 31, 2007
  5. Agda
  6. Ulf Norell
  7. -}
  8. -- Let's have a closer look at the module system
  9. module Modules where
  10. {-
  11. Importing and opening modules
  12. -}
  13. -- You can import a module defined in a different file.
  14. import Nat
  15. -- This will bring the module into scope and allows you to
  16. -- access its contents using qualified names.
  17. plusTwo : Nat.Nat -> Nat.Nat
  18. plusTwo n = Nat._+_ n 2
  19. -- To bring everything from a module into scope you can open
  20. -- the module.
  21. open Nat
  22. z : Nat
  23. z = zero
  24. -- There's also a short-hand to import and open at the same time
  25. open import Bool
  26. _&&_ : Bool -> Bool -> Bool
  27. x && y = if x then y else false
  28. -- Sometimes it's nice to be able to control what is brought
  29. -- into scope when you open a module. There are three modifiers
  30. -- that affect this: using, hiding and renaming.
  31. module DifferentWaysOfOpeningNat where
  32. -- nothing but Nat
  33. open Nat using (Nat)
  34. -- everything but zero
  35. open Nat hiding (zero)
  36. -- everything, but zero and suc under different names
  37. open Nat renaming (zero to ZZ; suc to S_S)
  38. two : Nat
  39. two = S S zero S S
  40. -- you can combine using or hiding with renaming, but not using
  41. -- with hiding (for obvious reasons).
  42. -- To re-export something opened use the public modifier.
  43. module A where
  44. open Nat public using (Nat)
  45. N = A.Nat -- now Nat is a visible name in module A
  46. {-
  47. Parameterised modules
  48. -}
  49. -- A very useful feature is parameterised modules.
  50. data Vec (A : Set) : Nat -> Set where
  51. [] : Vec A 0
  52. _::_ : {n : Nat} -> A -> Vec A n -> Vec A (suc n)
  53. infixr 40 _::_
  54. module Sort {A : Set}(_≤_ : A -> A -> Bool) where
  55. insert : {n : Nat} -> A -> Vec A n -> Vec A (suc n)
  56. insert x [] = x :: []
  57. insert x (y :: ys) = if x ≤ y
  58. then x :: y :: ys
  59. else y :: insert x ys
  60. sort : {n : Nat} -> Vec A n -> Vec A n
  61. sort [] = []
  62. sort (x :: xs) = insert x (sort xs)
  63. _≤_ : Nat -> Nat -> Bool
  64. zero ≤ m = true
  65. suc n ≤ zero = false
  66. suc n ≤ suc m = n ≤ m
  67. -- When used directly, functions from parameterised modules
  68. -- take the parameters as extra arguments.
  69. test = Sort.sort _≤_ (6 :: 2 :: 0 :: 4 :: [])
  70. -- But, you can also apply the entire module to its arguments.
  71. -- Let's open the new module while we're at it.
  72. open module SortNat = Sort _≤_
  73. test' = sort (3 :: 2 :: 4 :: 0 :: [])
  74. {-
  75. Local definitions
  76. -}
  77. data _==_ {A : Set}(x : A) : A -> Set where
  78. refl : x == x
  79. subst : {A : Set}(C : A -> Set){x y : A} -> x == y -> C x -> C y
  80. subst C refl cx = cx
  81. cong : {A B : Set}(f : A -> B){x y : A} -> x == y -> f x == f y
  82. cong f refl = refl
  83. lem₁ : (n : Nat) -> n + 0 == n
  84. lem₁ zero = refl
  85. lem₁ (suc n) = cong suc (lem₁ n)
  86. lem₂ : (n m : Nat) -> n + suc m == suc n + m
  87. lem₂ n zero = refl
  88. lem₂ n (suc m) = cong suc (lem₂ n m)
  89. {-
  90. What's next?
  91. -}
  92. -- The final thing on the agenda is records.
  93. -- Move on to: Records.agda