Operators.agda 1.9 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566
  1. -- This module introduces operators.
  2. module Introduction.Operators where
  3. -- Agda has a very flexible mechanism for defining operators, supporting infix,
  4. -- prefix, postfix and mixfix operators.
  5. data Nat : Set where
  6. zero : Nat
  7. suc : Nat -> Nat
  8. -- Any name containing underscores (_) can be used as an operator by writing
  9. -- the arguments where the underscores are. For instance, the function _+_ is
  10. -- the infix addition function. This function can be used either as a normal
  11. -- function: '_+_ zero zero', or as an operator: 'zero + zero'.
  12. _+_ : Nat -> Nat -> Nat
  13. zero + m = m
  14. suc n + m = suc (n + m)
  15. -- A fixity declaration specifies precedence level (50 in this case) and
  16. -- associativity (left associative here) of an operator. Only infix operators
  17. -- (whose names start and end with _) have associativity.
  18. infixl 50 _+_
  19. -- The only restriction on where _ can appear in a name is that there cannot be
  20. -- two underscores in sequence. For instance, we can define an if-then-else
  21. -- operator:
  22. data Bool : Set where
  23. false : Bool
  24. true : Bool
  25. if_then_else_ : {A : Set} -> Bool -> A -> A -> A
  26. if true then x else y = x
  27. if false then x else y = y
  28. -- if_then_else_ is treated as a prefix operator (ends, but doesn't begin with
  29. -- an _), so the declared precedence determines how something in an else branch
  30. -- should be parsed. For instance, with the given precedences
  31. -- if x then y else a + b
  32. -- is parsed as
  33. -- if x then y else (a + b)
  34. -- and not
  35. -- (if x then y else a) + b
  36. infix 10 if_then_else_
  37. -- In Agda there is no restriction on what characters are allowed to appear in
  38. -- an operator (as opposed to a function symbol). For instance, it is allowed
  39. -- (but not recommended) to define 'f' to be an infix operator and '+' to be a
  40. -- function symbol.
  41. module BadIdea where
  42. _f_ : Nat -> Nat -> Nat
  43. zero f zero = zero
  44. zero f suc n = suc n
  45. suc n f zero = suc n
  46. suc n f suc m = suc (n f m)
  47. + : Nat -> Nat
  48. + n = suc n