Records.agda 2.6 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128
  1. {-
  2. Agda Implementors' Meeting VI
  3. Göteborg
  4. May 24 - 30, 2007
  5. Hello Agda!
  6. Ulf Norell
  7. -}
  8. -- Records are labeled sigma types.
  9. module Records where
  10. open import Naturals
  11. open import Bool
  12. {-
  13. A very simple record.
  14. -}
  15. record Point : Set where
  16. field x : Nat
  17. y : Nat
  18. -- A record can be seen as a one constructor datatype. In this case:
  19. data Point' : Set where
  20. mkPoint : (x : Nat)(y : Nat) -> Point'
  21. -- There are a few differences, though:
  22. -- To construct a record you use the syntax record { ..; x = e; .. }
  23. origin : Point
  24. origin = record { x = 0; y = 0 }
  25. -- instead of
  26. origin' : Point'
  27. origin' = mkPoint 0 0
  28. -- What's more interesting is that you get projection functions
  29. -- for free when you declare a record. More precisely, you get a module
  30. -- parameterised over a record, containing functions corresponding to the
  31. -- fields. In the Point example you get:
  32. {-
  33. module Point (p : Point) where
  34. x : Nat
  35. y : Nat
  36. -}
  37. -- So Point.x : Point -> Nat is the projection function for the field x.
  38. getX : Point -> Nat
  39. getX = Point.x
  40. -- A nifty thing with having the projection functions in a module is that
  41. -- you can apply the module to a record value, in effect opening the record.
  42. sum : Point -> Nat
  43. sum p = x + y
  44. where
  45. open module Pp = Point p
  46. -- The final difference between records and datatypes is that we have
  47. -- η-equality on records.
  48. data _==_ {A : Set}(x : A) : A -> Set where
  49. refl : x == x
  50. η-Point : (p : Point) -> p == record { x = Point.x p; y = Point.y p }
  51. η-Point p = refl
  52. {-
  53. The empty record
  54. -}
  55. -- One interesting benefit of this is that we get a unit type with
  56. -- η-equality.
  57. record True : Set where
  58. tt : True
  59. tt = record{}
  60. -- Now, since any element of True is equal to tt, metavariables of
  61. -- type True will simply disappear. The following cute example exploits
  62. -- this:
  63. data False : Set where
  64. NonZero : Nat -> Set
  65. NonZero zero = False
  66. NonZero (suc _) = True
  67. -- We make the proof that m is non-zero implicit.
  68. _/_ : (n m : Nat){p : NonZero m} -> Nat
  69. (n / zero) {}
  70. zero / suc m = zero
  71. suc n / suc m = div (suc n) (suc m) m
  72. where
  73. div : Nat -> Nat -> Nat -> Nat
  74. div zero zero c = suc zero
  75. div zero (suc y) c = zero
  76. div (suc x) zero c = suc (div x c c)
  77. div (suc x) (suc y) c = div x y c
  78. -- Now, as long as we're dividing by things which are obviously
  79. -- NonZero we can completely ignore the proof.
  80. five = 17 / 3
  81. {-
  82. A dependent record
  83. -}
  84. -- Of course, records can be dependent, and have parameters.
  85. record ∃ {A : Set}(P : A -> Set) : Set where
  86. field
  87. witness : A
  88. proof : P witness