open import Level

open import Data.Nat hiding (fold)
open import Data.Fin hiding (fold ; lift )
open import Data.Product
open import Data.Sum
open import Data.Unit

open import Relation.Binary.PropositionalEquality

open import Chapter2.Logic
open import Chapter2.IProp

open import Chapter6.IDesc
open import Chapter6.IDesc.Fixpoint
open import Chapter6.IDesc.InitialAlgebra
open import Chapter6.IDesc.Lifting
open import Chapter6.IDesc.Induction

open import Chapter10.Ornament.Examples.Lifting hiding ([_]^) renaming ([_]^h to [_]^ho)
open import Chapter10.Ornament

module Chapter10.AlgebraicOrnament.Coherence
           { : Level}
           {K : Set }
           {X : K  Set }
           (D : func  K K)
           (α :  D ⟧func X  X) where

open import Chapter10.AlgebraicOrnament
open import Chapter10.Ornament.Algebra {u = proj₁} (algOrn D α)
open import Chapter10.Ornament.CartesianMorphism

-- Thanks to Andrea Vezzosi for the proof!

coherentOrn : ∀{k x}  
                (t : μ (algOrnD D α) (k , x))  
                 x  fold D α (forgetOrnament t)
coherentOrn = induction (algOrnD D α) P
                         {i}{xs}  step {i} {xs})

    where P : {kx : Σ K X}  μ (algOrnD D α) kx  Set 
          P {(k , x)} t =  x  fold D α (forgetOrnament t)

          step' : (kx : Σ K X)
                  (D' : IDesc  K)
                  (α' :  D'  X  X (proj₁ kx)) 
                  let algOrnD' : IDesc  (Σ K X)
                      algOrnD' =  Desc.algOrnD (proj₁ kx) (proj₂ kx) D' α' in
                  (xs :  algOrnD'  (μ (algOrnD D α))) 
                  [  algOrnD' ]^h  {kx} t   proj₂ kx  fold D α  (forgetOrnament t)) xs 
                   proj₁ xs  Fold.hyps D α D' 
                                           (forgetOrnNat (algOrn D α) ([ D' ]^ho (proj₁ xs))
                                             (Fold.hyps (algOrnD D α) ornAlgebra  [ D' ]^ho (proj₁ xs) ⟧Orn  (proj₂ (proj₂ xs))))
          step' (k , x) (`var i) α' (xs , q , xxs) ih = ih
          step' (k , x) `1  α' (lift tt , q , lift tt) (lift tt) = pf refl
          step' (k , x) (D₁  D₂)  α' 
                ((xs₁ , xs₂) , q  , xxs₁ , xxs₂) (ih₁ , ih₂) = 
                cong₂  x y  (x , y)) <$>
                  step' (k , x) D₁  xs₁  α' (xs₁ , xs₂)) (xs₁ , q , xxs₁) ih₁  
                  step' (k , x) D₂  xs₂  α' (xs₁ , xs₂)) (xs₂ , q , xxs₂) ih₂
          step' (k , x) ( n T)  α' ((s , xs) , q , xxs) ih = 
                cong⊢  x  s , x) 
                      (step' (k , x) (T s)  xs  α' (s , xs)) (xs , q , xxs) ih)
          step' (k , x) ( S T)  α' ((s , xs) , q , xxs) ih = 
                cong⊢  x  s , x) 
                      (step' (k , x) (T s)  xs  α' (s , xs)) (xs , q , xxs) ih)
          step' (k , .(α' f)) ( S T)  α' (f , refl , xxs) ih = 
                extensionality  s  
                  step' (k , α' f) (T s)  xs  α' f ) (f s , refl , (xxs s)) (ih s))

          step : DAlg (algOrnD D α) P
          step {(k , .(α xs))} {(xs , refl , xxs)} ih =
            cong⊢ α (step' (k , α xs) 
                           (func.out D k) (α {k}) 
                           ((xs , refl ,  xxs)) ih)