Skip to content

Commit

Permalink
Some progress on left-id-law and such
Browse files Browse the repository at this point in the history
  • Loading branch information
gdijkstra committed May 2, 2016
1 parent b7a5454 commit f217145
Show file tree
Hide file tree
Showing 7 changed files with 237 additions and 63 deletions.
1 change: 0 additions & 1 deletion 1-hits-pf/Alg0/Limits.agda
Expand Up @@ -2,7 +2,6 @@

open import Container

-- Equality of algebra morphisms
module 1-hits-pf.Alg0.Limits (F : Container) where

open import Eq
Expand Down
46 changes: 3 additions & 43 deletions 1-hits-pf/Alg1/CatLaws.agda
@@ -1,49 +1,9 @@
{-# OPTIONS --without-K #-}

open import 1-hits-pf.Core
open import Container

-- Equality of algebra morphisms
module 1-hits-pf.Alg1.CatLaws (s : Spec) where

open import Eq
open import lib.Basics
open import 1-hits-pf.Alg1.Core s
open Spec s
open import 1-hits-pf.Alg0 F₀
open import 1-hits-pf.Alg0.FreeMonad F₀
open import 1-hits-pf.Alg1.Eq s

module _
{𝓧 𝓨 : Alg₁-obj}
(𝓯 : Alg₁-hom 𝓧 𝓨)
where

open Alg₁-obj 𝓧
open Alg₁-obj 𝓨 renaming (𝓧' to 𝓨' ; X to Y ; θ₀ to ρ₀ ; θ₁ to ρ₁)
open Alg₁-hom 𝓯

left-id₀ = Ap-idf f₀

left-id-alg₁ : Eq (∘-alg₁ (id-alg₁ 𝓨) 𝓯) 𝓯
left-id-alg₁ = alg₁-hom='
f
(∘₀ {𝓧'} {𝓨'} (id-alg₀ 𝓨') 𝓯')
f₀
(∘₁ (id-alg₁ 𝓨) 𝓯)
f₁
left-id₀
goal
where
T = Ap-∘ (idf Y ∘`) (f ∘`) θ₁ *v⊡ Ap-sq (idf _) f₁ ⊡v* sym (Ap-∘ (idf Y ∘`) (`∘ ⟦ F₁ ⟧₁ f) ρ₁)
B = Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (idf Y ∘`) ρ₁ *v⊡ Ap-sq (`∘ ⟦ F₁ ⟧₁ f) (id₁ 𝓨) ⊡v* sym (Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (`∘ idf (⟦ F₁ ⟧₀ Y)) ρ₁)
L = {!!}
R = {!!}
goal : Eq
((L *h⊡ (T ⊡v B) ⊡h* R) ⊡h* Ap (λ h₀ (star-hom₀ (alg₀-hom f h₀)) ₌∘ apply r X) left-id₀)
(Ap (λ h₀ (star-hom₀ (alg₀-hom f h₀)) ₌∘ apply l X) left-id₀ *h⊡ f₁)
goal = {!!}


right-id-alg₁ : Eq (∘-alg₁ 𝓯 (id-alg₁ 𝓧)) 𝓯
right-id-alg₁ = {!!} --Ap (alg₁-hom f) (Ap-idf f₁)
open import 1-hits-pf.Alg1.CatLaws.LeftId public
open import 1-hits-pf.Alg1.CatLaws.RightId public
open import 1-hits-pf.Alg1.CatLaws.Assoc public
34 changes: 34 additions & 0 deletions 1-hits-pf/Alg1/CatLaws/Assoc.agda
@@ -0,0 +1,34 @@
{-# OPTIONS --without-K #-}

open import 1-hits-pf.Core
open import Container

-- Equality of algebra morphisms
module 1-hits-pf.Alg1.CatLaws.Assoc (s : Spec) where

open import Eq
open import lib.Basics
open import 1-hits-pf.Alg1.Core s
open Spec s
open import 1-hits-pf.Alg0 F₀
open import FreeMonad renaming (_* to _⋆)
open import 1-hits-pf.Alg0.FreeMonad F₀
open import 1-hits-pf.Alg1.Eq s

module _
{𝓧 𝓨 𝓩 𝓦 : Alg₁-obj}
(𝓱 : Alg₁-hom 𝓩 𝓦)
(𝓰 : Alg₁-hom 𝓨 𝓩)
(𝓯 : Alg₁-hom 𝓧 𝓨)
where

open Alg₁-obj 𝓧
open Alg₁-obj 𝓨 renaming (X to Y; θ₀ to ρ₀; θ₁ to ρ₁)
open Alg₁-obj 𝓩 renaming (X to Z; θ₀ to ζ₀; θ₁ to ζ₁)
open Alg₁-obj 𝓦 renaming (X to W; θ₀ to ω₀; θ₁ to ω₁)
open Alg₁-hom 𝓱 renaming (f to h; f₀ to h₀; f₁ to h₁)
open Alg₁-hom 𝓰 renaming (f to g; f₀ to g₀; f₁ to g₁)
open Alg₁-hom 𝓯

assoc-alg₁ : Eq (∘-alg₁ (∘-alg₁ 𝓱 𝓰) 𝓯) (∘-alg₁ 𝓱 (∘-alg₁ 𝓰 𝓯))
assoc-alg₁ = {!!}
96 changes: 96 additions & 0 deletions 1-hits-pf/Alg1/CatLaws/LeftId.agda
@@ -0,0 +1,96 @@
{-# OPTIONS --without-K #-}

open import 1-hits-pf.Core
open import Container

-- Equality of algebra morphisms
module 1-hits-pf.Alg1.CatLaws.LeftId (s : Spec) where

open import Eq
open import lib.Basics
open import 1-hits-pf.Alg1.Core s
open Spec s
open import 1-hits-pf.Alg0 F₀
open import FreeMonad renaming (_* to _⋆)
open import 1-hits-pf.Alg0.FreeMonad F₀
open import 1-hits-pf.Alg1.Eq s

module _
{𝓧 𝓨 : Alg₁-obj}
(𝓯 : Alg₁-hom 𝓧 𝓨)
where

open Alg₁-obj 𝓧
open Alg₁-obj 𝓨 renaming (𝓧' to 𝓨' ; X to Y ; θ₀ to ρ₀ ; θ₁ to ρ₁)
open Alg₁-hom 𝓯

left-id₀ = Ap-idf f₀

left-id-alg₁ : Eq (∘-alg₁ (id-alg₁ 𝓨) 𝓯) 𝓯
left-id-alg₁ = alg₁-hom='
f
(∘₀ {𝓧'} {𝓨'} (id-alg₀ 𝓨') 𝓯')
f₀
(∘₁ (id-alg₁ 𝓨) 𝓯)
f₁
left-id₀
goal
where
T : Square
(idf Y ∘₌ star-hom₀ 𝓯' ₌∘ apply l X)
(f ∘₌ θ₁)
(ρ₁ ₌∘ ⟦ F₁ ⟧₁ f)
(idf Y ∘₌ star-hom₀ 𝓯' ₌∘ apply r X)
T = Ap-∘ (idf Y ∘`) (f ∘`) θ₁
*v⊡ Ap-sq (idf Y ∘`) f₁
⊡v* sym (Ap-∘ (idf Y ∘`) (`∘ ⟦ F₁ ⟧₁ f) ρ₁)

B : Square
((star-hom₀ (id-alg₀ 𝓨') ₌∘ apply l Y) ₌∘ ⟦ F₁ ⟧₁ f)
(ρ₁ ₌∘ ⟦ F₁ ⟧₁ f)
(ρ₁ ₌∘ ⟦ F₁ ⟧₁ f)
((star-hom₀ (id-alg₀ 𝓨') ₌∘ apply r Y) ₌∘ ⟦ F₁ ⟧₁ f)
B = Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (idf Y ∘`) ρ₁
*v⊡ Ap-sq (`∘ ⟦ F₁ ⟧₁ f) (id₁ 𝓨)
⊡v* sym (Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (`∘ ⟦ F₁ ⟧₁ (idf Y)) ρ₁)

lem-top :: ContHom F₁ (F₀ ⋆))
Eq ((idf Y ∘₌ star-hom₀ 𝓯') ₌∘ apply α X) (idf Y ∘₌ (star-hom₀ 𝓯' ₌∘ apply α X))
lem-top α = Ap-swap (idf Y) (apply α X) (star-hom₀ 𝓯')

lem-bot :: ContHom F₁ (F₀ ⋆))
Eq ((star-hom₀ (id-alg₀ 𝓨') ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X)
((star-hom₀ (id-alg₀ 𝓨') ₌∘ apply α Y) ₌∘ ⟦ F₁ ⟧₁ f)
lem-bot α = sym (Ap-∘ (`∘ apply α X) (`∘ ⟦ F₀ ⋆ ⟧₁ f) (star-hom₀ (id-alg₀ 𝓨')))
* Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (`∘ apply α Y) (star-hom₀ (id-alg₀ 𝓨'))

lem :: ContHom F₁ (F₀ ⋆))
Eq (star-hom₀ (∘-alg₀ (id-alg₀ 𝓨') 𝓯') ₌∘ apply α X)
(((idf Y ∘₌ star-hom₀ 𝓯') ₌∘ apply α X) *
((star-hom₀ (id-alg₀ 𝓨') ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X))
lem α =
(star-hom₀ (∘-alg₀ (id-alg₀ 𝓨') 𝓯') ₌∘ apply α X)

*⟨ Ap (λ P Ap (`∘ apply α X) P) (*-∘ (id-alg₀ 𝓨') 𝓯') ⟩ -- *-∘

((idf Y ∘₌ star-hom₀ 𝓯') * (star-hom₀ (id-alg₀ 𝓨') ₌∘ ⟦ F₀ ⋆ ⟧₁ f)) ₌∘ apply α X

*⟨ Ap-* (`∘ apply α X)
(Ap ((idf Y) ∘`) (star-hom₀ 𝓯'))
(Ap (`∘ ⟦ F₀ ⋆ ⟧₁ f) (star-hom₀ (id-alg₀ 𝓨')))
-- ap-*

((idf Y ∘₌ star-hom₀ 𝓯') ₌∘ apply α X) * ((star-hom₀ (id-alg₀ 𝓨') ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X) ∎*

id∘𝓯 : Square (star-hom₀ (∘-alg₀ (id-alg₀ 𝓨') 𝓯') ₌∘ apply l X)
(f ∘₌ θ₁) (ρ₁ ₌∘ ⟦ F₁ ⟧₁ f)
(star-hom₀ (∘-alg₀ (id-alg₀ 𝓨') 𝓯') ₌∘ apply r X)
id∘𝓯 = lem l *h⊡ ((lem-top l *h⊡ T ⊡h* sym (lem-top r))
⊡v (lem-bot l *h⊡ B ⊡h* sym (lem-bot r)))
⊡h* sym (lem r)

goal : Eq
(id∘𝓯 ⊡h* Ap (λ h₀ star-hom₀ (alg₀-hom f h₀) ₌∘ apply r X) left-id₀)
(Ap (λ h₀ star-hom₀ (alg₀-hom f h₀) ₌∘ apply l X) left-id₀ *h⊡
f₁)
goal = {!!}
30 changes: 30 additions & 0 deletions 1-hits-pf/Alg1/CatLaws/RightId.agda
@@ -0,0 +1,30 @@
{-# OPTIONS --without-K #-}

open import 1-hits-pf.Core
open import Container

-- Equality of algebra morphisms
module 1-hits-pf.Alg1.CatLaws.RightId (s : Spec) where

open import Eq
open import lib.Basics
open import 1-hits-pf.Alg1.Core s
open Spec s
open import 1-hits-pf.Alg0 F₀
open import FreeMonad renaming (_* to _⋆)
open import 1-hits-pf.Alg0.FreeMonad F₀
open import 1-hits-pf.Alg1.Eq s

module _
{𝓧 𝓨 : Alg₁-obj}
(𝓯 : Alg₁-hom 𝓧 𝓨)
where

open Alg₁-obj 𝓧
open Alg₁-obj 𝓨 renaming (𝓧' to 𝓨' ; X to Y ; θ₀ to ρ₀ ; θ₁ to ρ₁)
open Alg₁-hom 𝓯

right-id₀ = Ap-idf f₀

right-id-alg₁ : Eq (∘-alg₁ (id-alg₁ 𝓨) 𝓯) 𝓯
right-id-alg₁ = {!!}
49 changes: 30 additions & 19 deletions 1-hits-pf/Alg1/Core.agda
Expand Up @@ -70,41 +70,42 @@ module _
open Alg₁-hom 𝓯

∘₁ : is-alg₁-hom 𝓧 𝓩 (∘-alg₀ 𝓰' 𝓯')
∘₁ = L *h⊡ (T ⊡v B) ⊡h* sym R
∘₁ = lem l *h⊡ ((lem-top l *h⊡ T ⊡h* sym (lem-top r))
⊡v (lem-bot l *h⊡ B ⊡h* sym (lem-bot r)))
⊡h* sym (lem r)
where
T : Square (g ∘₌ star-hom₀ 𝓯' ₌∘ apply l X) ((g ∘ f) ∘₌ θ₁) (Ap (λ H g ∘ H ∘ ⟦ F₁ ⟧₁ f) ρ₁) (g ∘₌ star-hom₀ 𝓯' ₌∘ apply r X)
T = Ap-∘ (g ∘`) (f ∘`) θ₁
*v⊡ Ap-sq (g ∘`) f₁
⊡v* sym (Ap-∘ (g ∘`) (`∘ ⟦ F₁ ⟧₁ f) ρ₁)

B : Square ((star-hom₀ 𝓰' ₌∘ apply l Y) ₌∘ ⟦ F₁ ⟧₁ f) (Ap (λ H g ∘ H ∘ ⟦ F₁ ⟧₁ f) ρ₁) (ζ₁ ₌∘ ⟦ F₁ ⟧₁ (g ∘ f)) ((star-hom₀ 𝓰' ₌∘ apply r Y) ₌∘ ⟦ F₁ ⟧₁ f)
B = Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (g ∘`) ρ₁
*v⊡ Ap-sq (`∘ ⟦ F₁ ⟧₁ f) g₁
⊡v* sym (Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (`∘ ⟦ F₁ ⟧₁ g) ζ₁)

lem-top :: ContHom F₁ (F₀ ⋆))
Eq ((g ∘₌ star-hom₀ 𝓯') ₌∘ apply α X) (g ∘₌ (star-hom₀ 𝓯' ₌∘ apply α X))
lem-top α = Ap-swap g (apply α X) (star-hom₀ 𝓯')

lem-bot :: ContHom F₁ (F₀ ⋆))
Eq ((star-hom₀ 𝓰' ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X) ((star-hom₀ 𝓰' ₌∘ apply α Y) ₌∘ ⟦ F₁ ⟧₁ f)
lem-bot α = sym (Ap-∘ (`∘ apply α X) (`∘ ⟦ F₀ ⋆ ⟧₁ f) (star-hom₀ 𝓰'))
* Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (`∘ apply α Y) (star-hom₀ 𝓰')

lem :: ContHom F₁ (F₀ ⋆))
Eq (star-hom₀ (∘-alg₀ 𝓰' 𝓯') ₌∘ apply α X)
((g ∘₌ (star-hom₀ 𝓯' ₌∘ apply α X)) * (star-hom₀ 𝓰' ₌∘ apply α Y) ₌∘ ⟦ F₁ ⟧₁ f)
(((g ∘₌ star-hom₀ 𝓯') ₌∘ apply α X) * ((star-hom₀ 𝓰' ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X))
lem α =
(star-hom₀ (∘-alg₀ 𝓰' 𝓯') ₌∘ apply α X)

*⟨ Ap (λ P Ap (`∘ apply α X) P) (*-∘ 𝓰' 𝓯') ⟩ -- *-
*⟨ Ap (λ P Ap (`∘ apply α X) P) (*-∘ 𝓰' 𝓯') ⟩ -- *-

((g ∘₌ star-hom₀ 𝓯') * (star-hom₀ 𝓰' ₌∘ ⟦ F₀ ⋆ ⟧₁ f)) ₌∘ apply α X

*⟨ Ap-* (`∘ apply α X) (Ap (g ∘`) (star-hom₀ 𝓯')) (Ap (`∘ ⟦ F₀ ⋆ ⟧₁ f) (star-hom₀ 𝓰')) ⟩

((g ∘₌ star-hom₀ 𝓯') ₌∘ apply α X) * ((star-hom₀ 𝓰' ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X)

*⟨ Ap (λ P P * Ap (`∘ apply α X) (Ap (`∘ ⟦ F₀ ⋆ ⟧₁ f) (star-hom₀ 𝓰'))) (Ap-swap g (apply α X) (star-hom₀ 𝓯')) ⟩

(g ∘₌ (star-hom₀ 𝓯' ₌∘ apply α X)) * ((star-hom₀ 𝓰' ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X)
*⟨ Ap-* (`∘ apply α X) (Ap (g ∘`) (star-hom₀ 𝓯')) (Ap (`∘ ⟦ F₀ ⋆ ⟧₁ f) (star-hom₀ 𝓰')) ⟩ -- ap-*

*⟨ Ap (λ P Ap (g ∘`) (Ap (`∘ apply α X) (star-hom₀ 𝓯')) * P) (sym (Ap-∘ (`∘ apply α X) (`∘ ⟦ F₀ ⋆ ⟧₁ f) (star-hom₀ 𝓰'))) ⟩

(g ∘₌ (star-hom₀ 𝓯' ₌∘ apply α X)) * (star-hom₀ 𝓰' ₌∘ (apply α Y ∘ ⟦ F₁ ⟧₁ f))

*⟨ Ap (λ P Ap (g ∘`) (Ap (`∘ apply α X) (star-hom₀ 𝓯')) * P) (Ap-∘ (`∘ ⟦ F₁ ⟧₁ f) (`∘ apply α Y) (star-hom₀ 𝓰')) ⟩

(g ∘₌ (star-hom₀ 𝓯' ₌∘ apply α X)) * ((star-hom₀ 𝓰' ₌∘ apply α Y) ₌∘ ⟦ F₁ ⟧₁ f) ∎*
L = lem l
R = lem r
((g ∘₌ star-hom₀ 𝓯') ₌∘ apply α X) * ((star-hom₀ 𝓰' ₌∘ ⟦ F₀ ⋆ ⟧₁ f) ₌∘ apply α X) ∎*

∘-alg₁ : Alg₁-hom 𝓧 𝓩
∘-alg₁ = alg₁-hom (∘-alg₀ 𝓰' 𝓯') ∘₁
Expand All @@ -123,3 +124,13 @@ module _

id-alg₁ : Alg₁-hom 𝓧 𝓧
id-alg₁ = alg₁-hom (id-alg₀ 𝓧') id₁

open import Cat

Alg₁ : Cat
Alg₁ = record
{ obj = Alg₁-obj
; hom = Alg₁-hom
; comp = ∘-alg₁
; id = id-alg₁
}
44 changes: 44 additions & 0 deletions 1-hits-pf/Alg1/Limits.agda
@@ -0,0 +1,44 @@
{-# OPTIONS --without-K #-}

open import 1-hits-pf.Core

module 1-hits-pf.Alg1.Limits (s : Spec) where

open Spec s

open import Container
open import lib.Basics
open import lib.types.Sigma
open import Cat
open import Eq
open import 1-hits-pf.Alg1.Core s
open import 1-hits-pf.Alg0.Core F₀

module _
(𝓧 𝓨 : Alg₁-obj)
where

open Alg₁-obj 𝓧
open Alg₁-obj 𝓨 renaming (X to Y ; θ₀ to ρ₀ ; θ₁ to ρ₁)

product-alg₁ : Product Alg₁ 𝓧 𝓨
product-alg₁ = record
{ prod = ×-alg₁
; π₁ = π₁-alg₁
; π₂ = π₂-alg₁
}
where
×₀ : has-alg₀ (X × Y)
×₀ = λ x θ₀ (⟦ F₀ ⟧₁ fst x) , ρ₀ (⟦ F₀ ⟧₁ snd x)

×₁ : has-alg₁ (alg₀ (X × Y) ×₀) --has-alg₁
×₁ = {!!}

×-alg₁ : Alg₁-obj
×-alg₁ = alg₁ (alg₀ (X × Y) ×₀) ×₁

π₁-alg₁ : Alg₁-hom ×-alg₁ 𝓧
π₁-alg₁ = alg₁-hom (alg₀-hom fst refl) {!!}

π₂-alg₁ : Alg₁-hom ×-alg₁ 𝓨
π₂-alg₁ = alg₁-hom (alg₀-hom snd refl) {!!}

0 comments on commit f217145

Please sign in to comment.