Datatypes.agda 1.6 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101
  1. {-
  2. Types Summer School 2007
  3. Bertinoro
  4. Aug 19 - 31, 2007
  5. Agda
  6. Ulf Norell
  7. -}
  8. -- This is where the fun begins.
  9. -- Unleashing datatypes, pattern matching and recursion.
  10. module Datatypes where
  11. {-
  12. Simple datatypes.
  13. -}
  14. -- Let's define natural numbers.
  15. data Nat : Set where
  16. zero : Nat
  17. suc : Nat -> Nat
  18. -- A simple function.
  19. pred : Nat -> Nat
  20. pred zero = zero
  21. pred (suc n) = n
  22. -- Now let's do recursion.
  23. _+_ : Nat -> Nat -> Nat
  24. zero + m = m
  25. suc n + m = suc (n + m)
  26. infixl 60 _+_
  27. -- An aside on infix operators:
  28. -- Any name containing _ can be used as a mixfix operator.
  29. -- The arguments simply go in place of the _. For instance:
  30. data Bool : Set where
  31. true : Bool
  32. false : Bool
  33. if_then_else_ : {A : Set} -> Bool -> A -> A -> A
  34. if true then x else y = x
  35. -- if false then x else y = y
  36. if_then_else_ false x y = y
  37. {-
  38. Parameterised datatypes
  39. -}
  40. data List (A : Set) : Set where
  41. [] : List A
  42. _::_ : A -> List A -> List A
  43. -- The parameters are implicit arguments to the constructors.
  44. nil : (A : Set) -> List A
  45. nil A = [] {A}
  46. map : {A B : Set} -> (A -> B) -> List A -> List B
  47. map f [] = []
  48. map f (x :: xs) = f x :: map f xs
  49. {-
  50. Empty datatypes
  51. -}
  52. -- A very useful guy is the empty datatype.
  53. data False : Set where
  54. -- When pattern matching on an element of an empty type, something
  55. -- interesting happens:
  56. elim-False : {A : Set} -> False -> A
  57. elim-False () -- Look Ma, no right hand side!
  58. -- The pattern () is called an absurd pattern and matches elements
  59. -- of an empty type.
  60. {-
  61. What's next?
  62. -}
  63. -- The Curry-Howard isomorphism.
  64. -- CurryHoward.agda