Implicit.agda 3.7 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111
  1. -- This module introduces implicit arguments.
  2. module Introduction.Implicit where
  3. -- In Agda you can omit things that the type checker can figure out for itself.
  4. -- This is a crucial feature in a monomorphic language, since you would
  5. -- otherwise be overwhelmed by type arguments.
  6. -- Let's revisit the identity function from 'Introduction.Basics'.
  7. id' : (A : Set) -> A -> A
  8. id' A x = x
  9. -- Since Agda is monomorphic we have to take the type A as an argument. So when
  10. -- using the identity function we have to provide the type explicitly.
  11. data Nat : Set where
  12. zero : Nat
  13. suc : Nat -> Nat
  14. one : Nat
  15. one = id' Nat (suc zero)
  16. -- Always having to provide the type argument to the identity function would of
  17. -- course be very tedious, and seemingly unnecessary. We would expect the type
  18. -- checker to be able to figure out what it should be by looking at the second
  19. -- argument. And indeed the type checker can do this.
  20. -- One way of indicating that the type checker will have figure something out
  21. -- for itself is to write an underscore (_) instead of the term. So when
  22. -- applying the identity function we can write
  23. two : Nat
  24. two = id' _ (suc one)
  25. -- Now the type checker will try to figure out what the _ should be, and in
  26. -- this case it will have no problems doing so. If it should fail to infer the
  27. -- value of an _ it will issue an error message.
  28. -- In the case of the identity function we expect the type argument to always
  29. -- be inferrable so it would be nice if we could avoid having to write the _.
  30. -- This can be achieved by declaring the first argument as an implicit
  31. -- argument.
  32. id : {A : Set} -> A -> A -- implicit arguments are enclosed in curly braces
  33. id x = x -- now we don't have to mention A in the left-hand side
  34. three : Nat
  35. three = id (suc two)
  36. -- If, for some reason, an implicit argument cannot be inferred it can be given
  37. -- explicitly by enclosing it in curly braces :
  38. four : Nat
  39. four = id {Nat} (suc three)
  40. -- To summarise we give a bunch of possible variants of the identity function
  41. -- and its use.
  42. -- Various definitions of the identity function. Definitions 0 through 3 are
  43. -- all equivalent, as are definitions 4 through 6.
  44. id0 : (A : Set) -> A -> A
  45. id0 A x = x
  46. id1 : (A : Set) -> A -> A
  47. id1 _ x = x -- in left-hand sides _ means "don't care"
  48. id2 : (A : Set) -> A -> A
  49. id2 = \A x -> x
  50. id3 = \(A : Set)(x : A) -> x -- the type signature can be omitted for definitions
  51. -- of the form x = e if the type of e can be
  52. -- inferred.
  53. id4 : {A : Set} -> A -> A
  54. id4 x = x
  55. id5 : {A : Set} -> A -> A
  56. id5 {A} x = x
  57. id6 : {A : Set} -> A -> A
  58. id6 = \x -> x
  59. id7 = \{A : Set}(x : A) -> x
  60. -- id8 : {A : Set} -> A -> A
  61. -- id8 = \{A} x -> x -- this doesn't work since the type checker assumes
  62. -- that the implicit A has been has been omitted in
  63. -- the left-hand side (as in id6).
  64. -- Various uses of the identity function.
  65. zero0 = id0 Nat zero
  66. zero1 = id0 _ zero -- in right-hand sides _ means "go figure"
  67. zero2 = id4 zero
  68. zero3 = id4 {Nat} zero
  69. zero4 = id4 {_} zero -- This is equivalent to zero2, but it can be useful if
  70. -- a function has two implicit arguments and you need
  71. -- to provide the second one (when you provide an
  72. -- implicit argument explicitly it is assumed to be the
  73. -- left-most one).
  74. -- In this module we have looked at implicit arguments as a substitute for
  75. -- polymorphism. The implicit argument mechanism is more general than that and
  76. -- not limited to inferring the values of type arguments. For more information
  77. -- on implicit arguments see, for instance
  78. -- 'Introduction.Data.ByRecursion'