------------------------------------------------------------------------
-- The Agda standard library
--
-- M-types (the dual of W-types)
------------------------------------------------------------------------

{-# OPTIONS --cubical-compatible --sized-types #-}

module Codata.Sized.M where

open import Size
open import Level
open import Codata.Sized.Thunk using (Thunk; force)
open import Data.Product.Base hiding (map)
open import Data.Container.Core as C hiding (map)

data M {s p} (C : Container s p) (i : Size) : Set (s  p) where
  inf :  C  (Thunk (M C) i)  M C i

module _ {s p} {C : Container s p} where

  head :  {i}  M C i  Shape C
  head (inf (x , f)) = x

  tail : (x : M C )  Position C (head x)  M C 
  tail (inf (x , f)) = λ p  f p .force

-- map

module _ {s₁ s₂ p₁ p₂} {C₁ : Container s₁ p₁} {C₂ : Container s₂ p₂}
         (m : C₁  C₂) where

  map :  {i}  M C₁ i  M C₂ i
  map (inf t) = inf ( m  (C.map  t  λ where .force  map (t .force)) t))

-- unfold

module _ {s p } {C : Container s p} (open Container C)
         {S : Set } (alg : S   C  S) where

  unfold : S   {i}  M C i
  unfold seed = let (x , next) = alg seed in
                inf (x , λ p  λ where .force  unfold (next p))