postulates.lagda.rst 1.3 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657
  1. ..
  2. ::
  3. module language.postulates where
  4. .. _postulates:
  5. **********
  6. Postulates
  7. **********
  8. A postulate is a declaration of an element of some type without an accompanying definition. With postulates we can introduce elements in a type without actually giving the definition of the element itself.
  9. The general form of a postulate declaration is as follows:
  10. .. code-block:: agda
  11. postulate
  12. c11 ... c1i : <Type>
  13. ...
  14. cn1 ... cnj : <Type>
  15. Example: ::
  16. postulate
  17. A B : Set
  18. a : A
  19. b : B
  20. _=AB=_ : A -> B -> Set
  21. a==b : a =AB= b
  22. Introducing postulates is in general not recommended. Once postulates are introduced the consistency of the whole development is at risk, because there is nothing that prevents us from introducing an element in the empty set.
  23. ::
  24. data False : Set where
  25. postulate bottom : False
  26. A preferable way to work is to define a module parametrised by the elements we need
  27. ::
  28. module Absurd (bt : False) where
  29. -- ...
  30. module M (A B : Set) (a : A) (b : B)
  31. (_=AB=_ : A -> B -> Set) (a==b : a =AB= b) where
  32. -- ...
  33. Postulated built-ins
  34. --------------------
  35. Some :ref:`built-ins` such as `Float` and `Char` are introduced as a postulate and then given a meaning by the corresponding ``{-# BUILTIN ... #-}`` pragma.