diff --git a/ElgotAlgebra.agda b/ElgotAlgebra.agda index d57c496..3f90e72 100644 --- a/ElgotAlgebra.agda +++ b/ElgotAlgebra.agda @@ -171,7 +171,7 @@ module _ {Cπ’ž : CocartesianCategory o β„“ e} where -- products and exponentials of elgot-algebras --* - -- product elgot-algebras + -- if the carriers of the algebra form a product, so do the algebras Product-Elgot-Algebra : βˆ€ {EA EB : Elgot-Algebra} β†’ Product π’ž (Elgot-Algebra.A EA) (Elgot-Algebra.A EB) β†’ Elgot-Algebra Product-Elgot-Algebra {EA} {EB} p = record { A = AΓ—B @@ -232,16 +232,42 @@ module _ {Cπ’ž : CocartesianCategory o β„“ e} where ((Ο€β‚‚ ∘ ⟨ ((π₁ +₁ idC) ∘ g)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ g)#ᡇ ⟩) ∘ h) β‰ˆβŸ¨ assoc ⟩ Ο€β‚‚ ∘ ⟨ ((π₁ +₁ idC) ∘ g)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ g)#ᡇ ⟩ ∘ h ∎ ) - ; #-Folding = _ -- TODO - ; #-resp-β‰ˆ = _ -- TODO + ; #-Folding = Ξ» {X} {Y} {f} {h} β†’ ⟨⟩-congβ‚‚ (foldingΛ‘ {X} {Y}) (foldingΚ³ {X} {Y}) + ; #-resp-β‰ˆ = Ξ» fg β†’ ⟨⟩-congβ‚‚ (#ᡃ-resp-β‰ˆ (∘-resp-β‰ˆΚ³ fg)) (#ᡇ-resp-β‰ˆ (∘-resp-β‰ˆΚ³ fg)) } where - open Elgot-Algebra EA using (A) renaming (_# to _#ᡃ; #-Fixpoint to #ᡃ-Fixpoint; #-Uniformity to #ᡃ-Uniformity) - open Elgot-Algebra EB using () renaming (A to B; _# to _#ᡇ; #-Fixpoint to #ᡇ-Fixpoint; #-Uniformity to #ᡇ-Uniformity) + open Elgot-Algebra EA using (A) renaming (_# to _#ᡃ; #-Fixpoint to #ᡃ-Fixpoint; #-Uniformity to #ᡃ-Uniformity; #-Folding to #ᡃ-Folding; #-resp-β‰ˆ to #ᡃ-resp-β‰ˆ) + open Elgot-Algebra EB using () renaming (A to B; _# to _#ᡇ; #-Fixpoint to #ᡇ-Fixpoint; #-Uniformity to #ᡇ-Uniformity; #-Folding to #ᡇ-Folding; #-resp-β‰ˆ to #ᡇ-resp-β‰ˆ) open Product π’ž p open HomReasoning open Equiv - + foldingΛ‘ : βˆ€ {X} {Y} {f} {h} β†’ (((π₁ +₁ idC) ∘ (⟨ ((π₁ +₁ idC) ∘ f)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ f)#ᡇ ⟩ +₁ h))#ᡃ) β‰ˆ ((π₁ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , iβ‚‚ ∘ h ])#ᡃ + foldingΛ‘ {X} {Y} {f} {h} = begin + ((π₁ +₁ idC) ∘ (⟨ ((π₁ +₁ idC) ∘ f)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ f)#ᡇ ⟩ +₁ h))#ᡃ β‰ˆβŸ¨ #ᡃ-resp-β‰ˆ +β‚βˆ˜+₁ ⟩ + ((π₁ ∘ ⟨ ((π₁ +₁ idC) ∘ f)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ f)#ᡇ ⟩ +₁ idC ∘ h)#ᡃ) β‰ˆβŸ¨ #ᡃ-resp-β‰ˆ (+₁-congβ‚‚ project₁ identityΛ‘) ⟩ + ((((π₁ +₁ idC) ∘ f)#ᡃ +₁ h)#ᡃ) β‰ˆβŸ¨ #ᡃ-Folding ⟩ + ([ (idC +₁ i₁) ∘ ((π₁ +₁ idC) ∘ f) , iβ‚‚ ∘ h ] #ᡃ) β‰ˆβŸ¨ #ᡃ-resp-β‰ˆ ([]-congΚ³ sym-assoc) ⟩ + ([ ((idC +₁ i₁) ∘ (π₁ +₁ idC)) ∘ f , iβ‚‚ ∘ h ] #ᡃ) β‰ˆβŸ¨ #ᡃ-resp-β‰ˆ ([]-congΚ³ (∘-resp-β‰ˆΛ‘ +β‚βˆ˜+₁)) ⟩ + ([ ((idC ∘ π₁ +₁ i₁ ∘ idC)) ∘ f , iβ‚‚ ∘ h ] #ᡃ) β‰ˆβŸ¨ #ᡃ-resp-β‰ˆ ([]-congΚ³ (∘-resp-β‰ˆΛ‘ (+₁-congβ‚‚ identityΛ‘ identityΚ³))) ⟩ + ([ ((π₁ +₁ i₁)) ∘ f , iβ‚‚ ∘ h ] #ᡃ) β‰ˆβŸ¨ sym (#ᡃ-resp-β‰ˆ ([]-congβ‚‚ (∘-resp-β‰ˆΛ‘ (+₁-congβ‚‚ identityΚ³ identityΛ‘)) (∘-resp-β‰ˆΛ‘ identityΚ³))) ⟩ + (([ (π₁ ∘ idC +₁ idC ∘ i₁) ∘ f , (iβ‚‚ ∘ idC) ∘ h ])#ᡃ) β‰ˆβŸ¨ sym (#ᡃ-resp-β‰ˆ ([]-congβ‚‚ (∘-resp-β‰ˆΛ‘ +β‚βˆ˜+₁) (∘-resp-β‰ˆΛ‘ +β‚βˆ˜iβ‚‚))) ⟩ + (([ ((π₁ +₁ idC) ∘ (idC +₁ i₁)) ∘ f , ((π₁ +₁ idC) ∘ iβ‚‚) ∘ h ])#ᡃ) β‰ˆβŸ¨ #ᡃ-resp-β‰ˆ ([]-congβ‚‚ assoc assoc) ⟩ + (([ (π₁ +₁ idC) ∘ (idC +₁ i₁) ∘ f , (π₁ +₁ idC) ∘ iβ‚‚ ∘ h ])#ᡃ) β‰ˆβŸ¨ sym (#ᡃ-resp-β‰ˆ ∘[]) ⟩ + ((π₁ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , iβ‚‚ ∘ h ])#ᡃ ∎ + foldingΚ³ : βˆ€ {X} {Y} {f} {h} β†’ ((Ο€β‚‚ +₁ idC) ∘ (⟨ ((π₁ +₁ idC) ∘ f)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ f)#ᡇ ⟩ +₁ h))#ᡇ β‰ˆ ((Ο€β‚‚ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , iβ‚‚ ∘ h ])#ᡇ + foldingΚ³ {X} {Y} {f} {h} = begin + ((Ο€β‚‚ +₁ idC) ∘ (⟨ ((π₁ +₁ idC) ∘ f)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ f)#ᡇ ⟩ +₁ h))#ᡇ β‰ˆβŸ¨ #ᡇ-resp-β‰ˆ +β‚βˆ˜+₁ ⟩ + ((Ο€β‚‚ ∘ ⟨ ((π₁ +₁ idC) ∘ f)#ᡃ , ((Ο€β‚‚ +₁ idC) ∘ f)#ᡇ ⟩ +₁ idC ∘ h)#ᡇ) β‰ˆβŸ¨ #ᡇ-resp-β‰ˆ (+₁-congβ‚‚ projectβ‚‚ identityΛ‘) ⟩ + ((((Ο€β‚‚ +₁ idC) ∘ f)#ᡇ +₁ h)#ᡇ) β‰ˆβŸ¨ #ᡇ-Folding ⟩ + [ (idC +₁ i₁) ∘ ((Ο€β‚‚ +₁ idC) ∘ f) , iβ‚‚ ∘ h ] #ᡇ β‰ˆβŸ¨ #ᡇ-resp-β‰ˆ ([]-congΚ³ sym-assoc) ⟩ + ([ ((idC +₁ i₁) ∘ (Ο€β‚‚ +₁ idC)) ∘ f , iβ‚‚ ∘ h ] #ᡇ) β‰ˆβŸ¨ #ᡇ-resp-β‰ˆ ([]-congΚ³ (∘-resp-β‰ˆΛ‘ +β‚βˆ˜+₁)) ⟩ + ([ ((idC ∘ Ο€β‚‚ +₁ i₁ ∘ idC)) ∘ f , iβ‚‚ ∘ h ] #ᡇ) β‰ˆβŸ¨ #ᡇ-resp-β‰ˆ ([]-congΚ³ (∘-resp-β‰ˆΛ‘ (+₁-congβ‚‚ identityΛ‘ identityΚ³))) ⟩ + ([ ((Ο€β‚‚ +₁ i₁)) ∘ f , iβ‚‚ ∘ h ] #ᡇ) β‰ˆβŸ¨ sym (#ᡇ-resp-β‰ˆ ([]-congβ‚‚ (∘-resp-β‰ˆΛ‘ (+₁-congβ‚‚ identityΚ³ identityΛ‘)) (∘-resp-β‰ˆΛ‘ identityΚ³))) ⟩ + (([ (Ο€β‚‚ ∘ idC +₁ idC ∘ i₁) ∘ f , (iβ‚‚ ∘ idC) ∘ h ])#ᡇ) β‰ˆβŸ¨ sym (#ᡇ-resp-β‰ˆ ([]-congβ‚‚ (∘-resp-β‰ˆΛ‘ +β‚βˆ˜+₁) (∘-resp-β‰ˆΛ‘ +β‚βˆ˜iβ‚‚))) ⟩ + (([ ((Ο€β‚‚ +₁ idC) ∘ (idC +₁ i₁)) ∘ f , ((Ο€β‚‚ +₁ idC) ∘ iβ‚‚) ∘ h ])#ᡇ) β‰ˆβŸ¨ #ᡇ-resp-β‰ˆ ([]-congβ‚‚ assoc assoc) ⟩ + (([ (Ο€β‚‚ +₁ idC) ∘ (idC +₁ i₁) ∘ f , (Ο€β‚‚ +₁ idC) ∘ iβ‚‚ ∘ h ])#ᡇ) β‰ˆβŸ¨ sym (#ᡇ-resp-β‰ˆ ∘[]) ⟩ + ((Ο€β‚‚ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , iβ‚‚ ∘ h ])#ᡇ ∎ + --* -- Here follows the proof of equivalence for unguarded and Id-guarded Elgot-Algebras --*