open import Data.Nat
open import Data.Fin
open import Data.Product
open import Data.Sum
open import Data.Unit

open import Relation.Binary.PropositionalEquality

open import Chapter1.Logic

open import Chapter2.IDesc
open import Chapter2.IDesc.Fixpoint
open import Chapter2.IDesc.InitialAlgebra
open import Chapter2.IDesc.Induction

open import Chapter5.Ornament.Examples.Lifting
open import Chapter5.Ornament

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

I : Set
I = Σ K X

u : I  K
u = proj₁

algOrn : orn D u u
algOrn = orn.mk λ { (k , x)  insert ( func.out D k  X) λ xs 
                              insert (α xs  x) λ _  
                              orn.out [ D ]^ (k , xs) }
                
algOrnD : func I I
algOrnD =  algOrn ⟧orn

{-


open Ornament.OrnamentalAlgebra D forget algOrn


-- Can't be bothered to reimplement
postulate 
  makeAlg : ∀{i} → (x : μ D i) → μ algData (i , foldID D α x)

-- Thanks to Andrea Vezzosi:
OAAO⊢ : ∀{i}{x : X i} → (t : μ algData (i , x)) → x ≡ foldID D α (forgetOrnament t)
OAAO⊢ {i} {x} = 
  induction algData motive 
            (λ { {j , .(α x)} (x , refl , t) hs → cong α (aux (D j) (x , t) hs)}) 
  where
    motive : {i' : Σ I X} → μ ⟦ D ⁺ algOrn ⟧Orn i' → Set
    motive {ij} t = proj₂ ij ≡ foldID D α (forgetOrnament t)
    aux : (E : IDesc I)
          (xs
           : Σ (⟦ E ⟧ X)
             (λ s →
                ⟦ ⟦ algOrnHelp E s ⟧Orn ⟧ (μ ⟦ D ⁺ algOrn ⟧Orn)))
          →
          □ ⟦ algOrnHelp E (proj₁ xs) ⟧Orn (proj₂ xs) motive →
          proj₁ xs ≡
          Fold.hyps D α E
          (forgetMap0 (algOrnHelp E (proj₁ xs))
           (Fold.hyps ⟦ D ⁺ algOrn ⟧Orn
            (ornAlgebra) ⟦ algOrnHelp E (proj₁ xs) ⟧Orn
            (proj₂ xs)))
    aux `1 (tt , tt) hs = refl
    aux (`X i0) (x' , con t) hs = hs
    aux (l `× r) (x' , t) hs = (cong₂ _,_ (aux l ((proj₁ x' , proj₁ t)) (proj₁ hs)) (aux r (proj₂ x' , proj₂ t) (proj₂ hs)))
    aux (l `+ r) (inj₁ x' , t) hs = cong inj₁ (aux l (x' , t) hs)
    aux (l `+ r) (inj₂ y , t) hs = cong inj₂ (aux r (y , t) hs)
    aux (`Σ S T) ((s , x') , t) hs = cong (_,_ s) (aux (T s) (x' , t) hs)
    aux (`Π S T) (x' , t) hs = ext λ s → aux (T s) (x' s , t s) (hs s) where
      postulate
        ext : ∀ {S : Set}{T : S -> Set} -> {f g : (x : S) -> T x} -> (∀ x -> f x ≡ g x) -> f ≡ g
-}