StreamProc.agda 2.0 KB

1234567891011121314151617181920212223242526272829303132333435363738394041424344454647484950515253545556575859606162636465666768697071
  1. {- Agda can check termination of Stream transducer operations.
  2. (Created: Andreas Abel, 2008-12-01
  3. at Agda Intensive Meeting 9 in Sendai, Japan.
  4. I acknowledge the support by AIST and JST.)
  5. Stream transducers have been described in:
  6. N. Ghani, P. Hancock, and D. Pattinson,
  7. Continuous functions on final coalgebras.
  8. In Proc. CMCS 2006, Electr. Notes in Theoret. Comp. Sci., 2006.
  9. They have been modelled by mixed equi-(co)inductive sized types in
  10. A. Abel,
  11. Mixed Inductive/Coinductive Types and Strong Normalization.
  12. In APLAS 2007, LNCS 4807.
  13. Here we model them by mutual data/codata and mutual recursion/corecursion.
  14. Cf. examples/StreamEating.agda
  15. -}
  16. module StreamProc where
  17. open import Common.Coinduction
  18. data Stream (A : Set) : Set where
  19. cons : A -> ∞ (Stream A) -> Stream A
  20. -- Stream Transducer: Trans A B
  21. -- intended semantics: Stream A -> Stream B
  22. mutual
  23. data Trans (A B : Set) : Set where
  24. 〈_〉 : ∞ (Trans' A B) -> Trans A B
  25. data Trans' (A B : Set) : Set where
  26. get : (A -> Trans' A B) -> Trans' A B
  27. put : B -> Trans A B -> Trans' A B
  28. out : forall {A B} -> Trans A B -> Trans' A B
  29. out 〈 p 〉 = p
  30. -- evaluating a stream transducer ("stream eating")
  31. mutual
  32. -- eat is defined by corecursion into Stream B
  33. eat : forall {A B} -> Trans A B -> Stream A -> Stream B
  34. eat 〈 sp 〉 as ~ eat' sp as
  35. -- eat' is defined by a local recursion on Trans' A B
  36. eat' : forall {A B} -> Trans' A B -> Stream A -> Stream B
  37. eat' (get f) (cons a as) = eat' (f a) as
  38. eat' (put b sp) as = cons b (eat sp as)
  39. -- composing two stream transducers
  40. mutual
  41. -- comb is defined by corecursion into Trans A B
  42. comb : forall {A B C} -> Trans A B -> Trans B C -> Trans A C
  43. comb 〈 p1 〉 〈 p2 〉 ~ 〈 comb' p1 p2 〉
  44. -- comb' preforms a local lexicographic recursion on (Trans' B C, Trans' A B)
  45. comb' : forall {A B C} -> Trans' A B -> Trans' B C -> Trans' A C
  46. comb' (put b p1) (get f) = comb' (out p1) (f b)
  47. comb' (put b p1) (put c p2) = put c (comb p1 p2)
  48. comb' (get f) p2 = get (\ a -> comb' (f a) p2)