Tree.agda 1.0 KB

12345678910111213141516171819202122232425262728293031323334353637383940414243444546
  1. {-# OPTIONS --copatterns #-}
  2. module Tree where
  3. data Bool : Set where
  4. true false : Bool
  5. record Tree (A : Set) : Set where
  6. field
  7. label : A
  8. child : Bool -> Tree A
  9. open Tree
  10. -- corecursive function defined by copattern matching
  11. alternate : {A : Set}(a b : A) -> Tree A
  12. -- deep copatterns:
  13. label (child (alternate a b) false) = b
  14. child (child (alternate a b) false) true = alternate a b
  15. child (child (alternate a b) false) false = alternate a b
  16. -- shallow copatterns
  17. child {A = A} (alternate a b) true = alternate b a
  18. label {A = A} (alternate a b) = a
  19. {- Delivers an infinite tree
  20. a
  21. b b
  22. a a a a
  23. b b b b b b b b
  24. ...
  25. -}
  26. infixr 5 _::_
  27. data List (A : Set) : Set where
  28. [] : List A
  29. _::_ : A -> List A -> List A
  30. collect : List Bool -> {A : Set} -> Tree A -> List A
  31. collect [] t = []
  32. collect (b :: l) t = label t :: collect l (child t b)
  33. test : List Bool
  34. test = collect (true :: true :: true :: []) (alternate true false)
  35. -- should give true :: false : true :: []