{-# OPTIONS --guardedness --safe #-}

module Cubical.Codata.M.AsLimit.Container where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Equiv using (_≃_)
open import Cubical.Foundations.Function using (_∘_)

open import Cubical.Data.Unit
open import Cubical.Data.Prod
open import Cubical.Data.Nat as  using ( ; suc ; _+_ )

open import Cubical.Foundations.Transport

open import Cubical.Foundations.Univalence
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Function

open import Cubical.Data.Sum

open import Cubical.Foundations.Structure

open import Cubical.Codata.M.AsLimit.helper

-------------------------------------
-- Container and Container Functor --
-------------------------------------

-- Σ[ A ∈ (Type ℓ) ] (A → Type ℓ)
Container :   -> Type (ℓ-suc )
Container  = TypeWithStr   x  x  Type )

-- Polynomial functor (P₀ , P₁)  defined over a container
-- https://ncatlab.org/nlab/show/polynomial+functor

-- P₀ object part of functor
P₀ :  {} (S : Container ) -> Type  -> Type 
P₀ (A , B) X  = Σ[ a  A ] (B a -> X)

-- P₁ morphism part of functor
P₁ :  {} {S : Container } {X Y} (f : X -> Y) -> P₀ S X -> P₀ S Y
P₁ {S = S} f = λ { (a , g) ->  a , f  g }

-----------------------
-- Chains and Limits --
-----------------------

record Chain  : Type (ℓ-suc ) where
  constructor _,,_
  field
    X :  -> Type 
    π : {n : } -> X (suc n) -> X n

open Chain public

limit-of-chain :  {} -> Chain   Type 
limit-of-chain (x ,, pi) = Σ[ x  ((n : )  x n) ] ((n : )  pi {n = n} (x (suc n))  x n)

shift-chain :  {} -> Chain  -> Chain 
shift-chain = λ X,π -> ((λ x  X X,π (suc x)) ,, λ {n}  π X,π {suc n})

------------------------------------------------------
-- Limit type of a Container , and Shift of a Limit --
------------------------------------------------------

Wₙ :  {} -> Container  ->  -> Type 
Wₙ S 0 = Lift Unit
Wₙ S (suc n) = P₀ S (Wₙ S n)

πₙ :  {} -> (S : Container ) -> {n : } -> Wₙ S (suc n) -> Wₙ S n
πₙ {} S {0} _ = lift tt
πₙ S {suc n} = P₁ (πₙ S {n})

sequence :  {} -> Container  -> Chain 
X (sequence S) n = Wₙ S n
π (sequence S) {n} = πₙ S {n}

PX,Pπ :  {} (S : Container ) -> Chain 
PX,Pπ {} S =
   n  P₀ S (X (sequence S) n)) ,,
   {n : } x  P₁  z  z) (π (sequence S) {n = suc n} x ))

-----------------------------------
-- M type is limit of a sequence --
-----------------------------------

M :  {} -> Container   Type 
M = limit-of-chain  sequence