Empty.agda 1.2 KB

123456789101112131415161718192021222324252627282930313233343536373839404142
  1. -- This module explains how to combine elimination of empty types with pattern
  2. -- match style definitions without running into problems with decidability.
  3. module Introduction.Data.Empty where
  4. -- First we introduce an empty and a singleton type.
  5. data Zero : Set where
  6. data One : Set where
  7. one : One
  8. -- There is a special pattern () which matches any element of an (obviously)
  9. -- empty type. If there is a ()-pattern in a left-hand side the right-hand side
  10. -- can be omitted.
  11. elim-Zero : {A : Set} -> Zero -> A
  12. elim-Zero ()
  13. data _×_ (A B : Set) : Set where
  14. pair : A -> B -> A × B
  15. -- The algorithm for checking if a type is empty is very naive. In this example
  16. -- you cannot replace pair () _ with () because the type checker cannot see
  17. -- that Zero × B is empty.
  18. elim-EmptyPair : {A B : Set} -> Zero × B -> A
  19. elim-EmptyPair (pair () _)
  20. data Nat : Set where
  21. zero : Nat
  22. suc : Nat -> Nat
  23. -- For some empty types finite unfolding is not enough.
  24. ConstZero : Nat -> Set
  25. ConstZero zero = Zero
  26. ConstZero (suc n) = ConstZero n
  27. -- We can still define the elimination function but we have to do it
  28. -- recursively over the n.
  29. elim-ConstZero : (n : Nat) -> ConstZero n -> {A : Set} -> A
  30. elim-ConstZero zero ()
  31. elim-ConstZero (suc n) x = elim-ConstZero n x