open import Level
open import Categories.FreeObjects.Free
open import Categories.Functor.Core
open import Categories.Functor.Algebra
open import Categories.Functor.Coalgebra
open import Categories.Object.Terminal
open import Categories.Object.Exponential
open import Categories.Category.Core using (Category)
open import Categories.Category.Distributive using (Distributive)
open import Categories.Category.CartesianClosed using (CartesianClosed)
open import Categories.Category.Cocartesian using (Cocartesian)
open import Categories.Category.Cartesian using (Cartesian)
open import Categories.Category.BinaryProducts using (BinaryProducts)

import Categories.Morphism as M
import Categories.Morphism.Reasoning as MR

In a CCC Free Elgot Algebras are Automatically Stable.

module Algebra.Elgot.Properties {o  e} {C : Category o  e} (distributive : Distributive C) (expos :  {A B}  Exponential C A B) where
  open Category C
  open HomReasoning
  open Equiv
  open MR C
  open Distributive distributive
  open import Categories.Category.Distributive.Properties distributive
  open Cartesian cartesian
  open BinaryProducts products hiding (η) renaming (unique to ⟨⟩-unique)
  open Cocartesian cocartesian
  private
    ccc : CartesianClosed C
    ccc = record { cartesian = cartesian ; exp = expos }

  open CartesianClosed ccc hiding (exp; cartesian)

  open import Category.Construction.ElgotAlgebras cocartesian
  open import Category.Construction.ElgotAlgebras.Exponentials distributive expos
  open import Algebra.Elgot cocartesian
  open import Algebra.Elgot.Free cocartesian
  open import Algebra.Elgot.Stable distributive

  open Elgot-Algebra-Morphism renaming (h to <_>)

  free-isStableˡ :  {Y} (fe : FreeElgotAlgebra Y)  IsStableFreeElgotAlgebraˡ fe
  IsStableFreeElgotAlgebraˡ.[_,_]♯ˡ (free-isStableˡ {Y} fe) {X} A f = (eval′  (< FreeObject._* fe {Exponential-Elgot-Algebra {A} {X}} (λg f) >  id))
    where open FreeObject fe
  IsStableFreeElgotAlgebraˡ.♯ˡ-law (free-isStableˡ {Y} fe) {X} {A} f = sym (begin 
    (eval′  (< (λg f)* >  id))  (η  id) ≈⟨ pullʳ ⁂∘⁂  
    eval′  ((< (λg f)* >  η)  id  id)   ≈⟨ refl⟩∘⟨ (⁂-cong₂ (*-lift (λg f)) identity²)  
    eval′  (λg f  id)                     ≈⟨ β′  
    f                                       )
    where open FreeObject fe
  IsStableFreeElgotAlgebraˡ.♯ˡ-preserving (free-isStableˡ {Y} fe) {X} {B} f {Z} h = begin 
    (eval′  (< (λg f)* >  id))  (h   id)                                              ≈⟨ pullʳ ⁂∘⁂  
    eval′  (< (λg f)* >  h   id  id)                                                  ≈⟨ refl⟩∘⟨ (⁂-cong₂ (preserves ((λg f)*)) identity²)  
    eval′  (λg (((eval′ +₁ id)  distributeʳ⁻¹  ((< (λg f)* > +₁ id)  h  id)) #ᵇ)  id) ≈⟨ β′  
    ((eval′ +₁ id)  distributeʳ⁻¹  ((< (λg f)* > +₁ id)  h  id)) #ᵇ                     ≈˘⟨ #ᵇ-resp-≈ (refl⟩∘⟨ refl⟩∘⟨ (⁂∘⁂  ⁂-cong₂ refl identity²))  
    ((eval′ +₁ id)  distributeʳ⁻¹  ((< (λg f)* > +₁ id)  id)  (h  id)) #ᵇ              ≈⟨ #ᵇ-resp-≈ (refl⟩∘⟨ pullˡ (sym (distributeʳ⁻¹-natural id < (λg f) * > id)))  
    ((eval′ +₁ id)  ((< (λg f)* >  id +₁ id  id)  distributeʳ⁻¹)  (h  id)) #ᵇ         ≈⟨ #ᵇ-resp-≈ (refl⟩∘⟨ assoc  pullˡ (+₁∘+₁  +₁-cong₂ refl (elimʳ (⟨⟩-unique id-comm id-comm))))  
    ((eval′  (< (λg f)* >  id) +₁ id)  distributeʳ⁻¹  (h  id)) #ᵇ                      
    where 
      open FreeObject fe renaming (FX to KY)
      open Elgot-Algebra B using () renaming (_# to _#ᵇ; #-resp-≈ to #ᵇ-resp-≈)
      open Elgot-Algebra KY using () renaming (_# to _#ʸ)
  IsStableFreeElgotAlgebraˡ.♯ˡ-unique (free-isStableˡ {Y} fe) {X} {B} f g eq₁ eq₂ = λ-inj (begin
    λg g ≈⟨ *-uniq (λg f) (record { h = λg g ; preserves = λ {D} {h}  begin 
      λg g  (h ) ≈⟨ subst  
      λg (g  (h   id)) ≈⟨ λ-cong (eq₂ h) 
      λg (((g +₁ id)  distributeʳ⁻¹  (h  id)) #ᵇ) ≈˘⟨ λ-cong (#ᵇ-resp-≈ (pullˡ (+₁∘+₁  +₁-cong₂ β′ (elimʳ (⟨⟩-unique id-comm id-comm))))) 
      λg (((eval′ +₁ id)  (λg g  id +₁ id  id)  distributeʳ⁻¹  (h  id)) #ᵇ) ≈˘⟨ λ-cong (#ᵇ-resp-≈ (refl⟩∘⟨ (pullˡ (sym (distributeʳ⁻¹-natural id (λg g) id))  assoc))) 
      λg (((eval′ +₁ id)  distributeʳ⁻¹  ((λg g +₁ id)  id)  (h  id)) #ᵇ) ≈⟨ λ-cong (#ᵇ-resp-≈ (refl⟩∘⟨ refl⟩∘⟨ (⁂∘⁂  ⁂-cong₂ refl identity²))) 
      λg (((eval′ +₁ id)  distributeʳ⁻¹  ((λg g +₁ id)  h  id)) #ᵇ)  }) (subst  λ-cong (sym eq₁)) 
    < (λg f)* > ≈˘⟨ η-exp′ 
    λg (eval′  (< (λg f)* >  id)) )
    where 
      open FreeObject fe renaming (FX to KY)
      open Elgot-Algebra B using () renaming (_# to _#ᵇ; #-resp-≈ to #ᵇ-resp-≈)
      open Elgot-Algebra KY using () renaming (_# to _#ʸ)

  free-isStable :  {Y} (fe : FreeElgotAlgebra Y)  IsStableFreeElgotAlgebra fe
  free-isStable {Y} fe = isStableˡ⇒isStable fe (free-isStableˡ fe)