From f01a54b4d27e2be9d6fd57878786f4dbe23e01a2 Mon Sep 17 00:00:00 2001 From: Max New <maxsnew@gmail.com> Date: Fri, 2 Jun 2023 16:26:27 -0400 Subject: [PATCH] update to extensionsystem version of comonads/strong monads --- .../Abstract/TermModel/Convenient.agda | 46 ++++++++------- .../Abstract/TermModel/Convenient/Linear.agda | 32 ----------- .../Convenient/Linear/Properties.agda | 38 ------------- .../TermModel/Convenient/Semantics.agda | 56 +++++++++---------- 4 files changed, 48 insertions(+), 124 deletions(-) delete mode 100644 formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear.agda delete mode 100644 formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear/Properties.agda diff --git a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda index 5be8841..1e6ac96 100644 --- a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda +++ b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda @@ -13,12 +13,14 @@ open import Cubical.Categories.Functor open import Cubical.Categories.Functors.Constant open import Cubical.Categories.NaturalTransformation as NT hiding (NatTrans) open import Cubical.Categories.Limits.Terminal +open import Cubical.Categories.Limits.Terminal.More open import Cubical.Categories.Limits.BinProduct open import Cubical.Categories.Limits.BinProduct.More open import Cubical.Categories.Limits.BinCoproduct -open import Cubical.Categories.Monad.Base open import Cubical.Categories.Exponentials -open import Cubical.Categories.Monad.Strength.Cartesian +open import Cubical.Categories.Monad.ExtensionSystem +open import Cubical.Categories.Comonad.Instances.Environment +open import Cubical.Categories.Monad.Strength.Cartesian.ExtensionSystem private variable @@ -28,7 +30,6 @@ open Category open Functor open BinCoproduct open BinProduct -open IsMonad record Model â„“ â„“' : Type (â„“-suc (â„“-max â„“ â„“')) where field @@ -39,21 +40,15 @@ record Model â„“ â„“' : Type (â„“-suc (â„“-max â„“ â„“')) where exponentials : Exponentials cat binProd binCoprod : BinCoproducts cat -- with a strong monad - monad : Monad cat - strength : Strength binProd monad + monad : StrongExtensionSystem binProd -- TODO: rename Notation and make similar modules for terminal, coprod open Notation cat binProd public open ExpNotation cat binProd exponentials public - - 🙠= term .fst - - !t : ∀ {a} → cat [ a , 🙠] - !t = terminalArrow cat term _ - - ðŸ™Î· : ∀ {a} → (f : cat [ a , 🙠]) → f ≡ !t - ðŸ™Î· f = sym (terminalArrowUnique cat {T = term} f) + open TerminalNotation cat term public + open StrongMonadNotation binProd monad public + open EnvNotation binProd public _+_ : (a b : cat .ob) → cat .ob a + b = binCoprod a b .binCoprodOb @@ -65,13 +60,12 @@ record Model â„“ â„“' : Type (â„“-suc (â„“-max â„“ â„“')) where _||_ : {a b c : cat .ob} → cat [ a , c ] → cat [ b , c ] → cat [ a + b , c ] f1 || f2 = binCoprod _ _ .univProp f1 f2 .fst .fst - T = monad .fst - - ret : ∀ {a} → cat [ a , T ⟅ a ⟆ ] - ret = monad .snd .η .NT.NatTrans.N-ob _ - _⇀_ : (a b : cat .ob) → cat .ob - a ⇀ b = a ⇒ T ⟅ b ⟆ + a ⇀ b = a ⇒ T b + + -- TODO: should this go into notation? + Linear = PKleisli + ClLinear = Kleisli cat (StrongMonad→Monad binProd term monad) field -- a weak model of the natural numbers, but good enough for our syntax @@ -84,15 +78,19 @@ record Model â„“ â„“' : Type (â„“-suc (â„“-max â„“ â„“')) where -- at this point we will model injection and projection as -- arbitrary morphisms - inj : cat [ nat + (dyn ⇀ dyn) , dyn ] - prj : cat [ dyn , T ⟅ nat + (dyn ⇀ dyn) ⟆ ] + inj : cat [ nat + (dyn ⇀ dyn) , dyn ] + prj : ClLinear [ dyn , nat + (dyn ⇀ dyn) ] -- and the error of course! -- err : 1 ⇀ A -- naturality says err ≡ T f ∘ err -- not sure if that's exactly right... - err : NT.NatTrans (Constant _ _ ðŸ™) T + err : ∀ {a} → cat [ 🙠, a ] + ℧ : ∀ {Γ a} → cat [ Γ , T a ] + ℧ = err ∘⟨ cat ⟩ !t - ℧ : ∀ {Γ d} → cat [ Γ , T ⟅ d ⟆ ] - ℧ {Γ} = NT.NatTrans.N-ob err _ ∘⟨ cat ⟩ terminalArrow cat term Γ + field + ℧-homo : ∀ {Γ a b} → (f : Linear Γ [ a , b ]) + -- define this explicitly as a profunctor? + → bind f ∘⟨ cat ⟩ ((cat .id ,p ℧)) ≡ ℧ diff --git a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear.agda b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear.agda deleted file mode 100644 index 942529e..0000000 --- a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear.agda +++ /dev/null @@ -1,32 +0,0 @@ -{-# OPTIONS --lossy-unification #-} -module Semantics.Abstract.TermModel.Convenient.Linear where - -open import Cubical.Foundations.Prelude -open import Cubical.Categories.Category -open import Cubical.Categories.Functor -open import Cubical.Categories.Monad.Strength.Cartesian -open import Cubical.Categories.Monad.Kleisli -open import Cubical.Categories.DistributiveLaw.ComonadOverMonad.BiKleisli.Base -open import Cubical.Categories.DistributiveLaw.ComonadOverMonad.BiKleisli.Morphism -open import Cubical.Categories.Comonad.Instances.Environment - -open import Semantics.Abstract.TermModel.Convenient - -private - variable - â„“ â„“' : Level - -module StrengthNotation (𓜠: Model â„“ â„“') where - open Model 𓜠- open Category cat - Linear : ∀ (Γ : ob) → Category _ _ - Linear Γ = BiKleisli (Env Γ (binProd Γ)) monad (strength-law binProd monad strength Γ) - - ClosedLinear : Category _ _ - ClosedLinear = Kleisli monad - - wkClosed : ∀ Γ → Functor ClosedLinear (Linear Γ) - wkClosed = wkF binProd monad strength - - _^* : ∀ {Δ Γ}(γ : Hom[ Δ , Γ ]) → Functor (Linear Γ) (Linear Δ) - γ ^* = change-of-comonad (change-of-base _ _ strength γ) diff --git a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear/Properties.agda b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear/Properties.agda deleted file mode 100644 index 9e8312b..0000000 --- a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Linear/Properties.agda +++ /dev/null @@ -1,38 +0,0 @@ -module Semantics.Abstract.TermModel.Convenient.Linear.Properties where - -open import Cubical.Foundations.Prelude -open import Cubical.Categories.Category -open import Cubical.Categories.Functor -open import Cubical.Categories.Monad.Strength.Cartesian -open import Cubical.Categories.Monad.Kleisli -open import Cubical.Categories.DistributiveLaw.ComonadOverMonad.BiKleisli.Base -open import Cubical.Categories.DistributiveLaw.ComonadOverMonad.BiKleisli.Morphism -open import Cubical.Categories.DistributiveLaw.ComonadOverMonad.BiKleisli.Morphism.Properties -open import Cubical.Categories.Comonad.Instances.Environment - -open import Semantics.Abstract.TermModel.Convenient -open import Semantics.Abstract.TermModel.Convenient.Linear -private - variable - â„“ â„“' : Level - -module StrengthLemmas (𓜠: Model â„“ â„“') where - open Model 𓜠- open Category cat - open Functor - open StrengthNotation 𓜠- - private - variable - Γ Δ a b c : ob - γ δ : Hom[ Δ , Γ ] - E : Linear Γ [ a , b ] - id^* : (id ^*) ⟪ E ⟫ ≡ E - id^* {E = E} = congâ‚‚ _∘_ refl (×pF .F-id) ∙ ⋆IdL _ - - comp^* : ((γ ∘ δ) ^*) ⟪ E ⟫ ≡ ((δ ^*) ⟪ ((γ ^*) ⟪ E ⟫) ⟫) - comp^* {E = E} = - congâ‚‚ _∘_ refl (congâ‚‚ _,p_ (sym (⋆Assoc _ _ _) ∙ congâ‚‚ _∘_ refl (sym ×βâ‚) ∙ ⋆Assoc _ _ _) - (sym ×β₂ ∙ congâ‚‚ _∘_ (sym (⋆IdR _)) refl) - ∙ sym ,p-natural) - ∙ (⋆Assoc _ _ _) diff --git a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda index cc851e2..a3b9503 100644 --- a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda +++ b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda @@ -23,8 +23,6 @@ open import Cubical.Data.List hiding ([_]) open import Syntax.Types open import Syntax.Terms open import Semantics.Abstract.TermModel.Convenient -open import Semantics.Abstract.TermModel.Convenient.Linear -open import Semantics.Abstract.TermModel.Convenient.Linear.Properties private variable @@ -48,10 +46,7 @@ private module _ (𓜠: Model â„“ â„“') where module 𓜠= Model 𓜠- module T = IsMonad (ð“œ.monad .snd) ⇒F = ExponentialF ð“œ.cat ð“œ.binProd ð“œ.exponentials - open StrengthNotation 𓜠- open StrengthLemmas 𓜠⟦_⟧ty : Ty → ð“œ.cat .ob ⟦ nat ⟧ty = ð“œ.nat @@ -59,8 +54,8 @@ module _ (𓜠: Model â„“ â„“') where ⟦ S ⇀ T ⟧ty = ⟦ S ⟧ty ð“œ.⇀ ⟦ T ⟧ty ⟦_⟧e : S ⊑ R → ð“œ.cat [ ⟦ S ⟧ty , ⟦ R ⟧ty ] - ⟦_⟧p : S ⊑ R → ð“œ.cat [ ⟦ R ⟧ty , ð“œ.T ⟅ ⟦ S ⟧ty ⟆ ] - ⟦_⟧p' : S ⊑ R → ð“œ.cat [ ð“œ.T ⟅ ⟦ R ⟧ty ⟆ , ð“œ.T ⟅ ⟦ S ⟧ty ⟆ ] + ⟦_⟧p : S ⊑ R → ð“œ.ClLinear [ ⟦ R ⟧ty , ⟦ S ⟧ty ] + ⟦_⟧p' : S ⊑ R → ð“œ.cat [ ð“œ.T ⟦ R ⟧ty , ð“œ.T ⟦ S ⟧ty ] ⟦ nat ⟧e = ð“œ.cat .id ⟦ dyn ⟧e = ð“œ.cat .id @@ -68,20 +63,20 @@ module _ (𓜠: Model â„“ â„“') where -- λ f . λ x . x' <- p x; -- y' <- app(f,x'); -- η (e y') - ⟦ c ⇀ d ⟧e = ð“œ.lda ((ð“œ.ret ∘⟨ ð“œ.cat ⟩ ⟦ d ⟧e) ∘⟨ ClosedLinear ⟩ - ð“œ.app ∘⟨ Linear _ ⟩ - wkClosed _ ⟪ ⟦ c ⟧p ⟫) + ⟦ c ⇀ d ⟧e = ð“œ.lda ((ð“œ.ClLinear .id ∘⟨ ð“œ.cat ⟩ ⟦ d ⟧e) ∘⟨ ð“œ.ClLinear ⟩ + ð“œ.app ∘⟨ ð“œ.Linear _ ⟩ + {!!}) ⟦ inj-nat ⟧e = ð“œ.inj ∘⟨ ð“œ.cat ⟩ ð“œ.σ1 ⟦ inj-arr c ⟧e = ð“œ.inj ∘⟨ ð“œ.cat ⟩ ð“œ.σ2 ∘⟨ ð“œ.cat ⟩ ⟦ c ⟧e - ⟦ nat ⟧p = ð“œ.ret - ⟦ dyn ⟧p = ð“œ.ret + ⟦ nat ⟧p = ð“œ.ClLinear .id + ⟦ dyn ⟧p = ð“œ.ClLinear .id -- = η ∘ (⟦ c ⟧e ⇒ ⟦ d ⟧p') - ⟦ c ⇀ d ⟧p = ð“œ.ret ∘⟨ ð“œ.cat ⟩ ⇒F ⟪ ⟦ c ⟧e , ⟦ d ⟧p' ⟫ - ⟦ inj-nat ⟧p = (ð“œ.ret ð“œ.|| ð“œ.℧) ∘⟨ ClosedLinear ⟩ ð“œ.prj - ⟦ inj-arr c ⟧p = (ð“œ.℧ ð“œ.|| ⟦ c ⟧p) ∘⟨ ClosedLinear ⟩ ð“œ.prj + ⟦ c ⇀ d ⟧p = {!!} -- ð“œ.ret ∘⟨ ð“œ.cat ⟩ ⇒F ⟪ ⟦ c ⟧e , ⟦ d ⟧p' ⟫ + ⟦ inj-nat ⟧p = {!!} -- (ð“œ.ret ð“œ.|| ð“œ.℧) ∘⟨ ð“œ.ClLinear ⟩ ð“œ.prj + ⟦ inj-arr c ⟧p = {!!} -- (ð“œ.℧ ð“œ.|| ⟦ c ⟧p) ∘⟨ ð“œ.ClLinear ⟩ ð“œ.prj - ⟦ c ⟧p' = T.bind .N-ob _ ⟦ c ⟧p + ⟦ c ⟧p' = {!!} -- ð“œ.bind .N-ob _ ⟦ c ⟧p ⟦_⟧ctx : Ctx → ð“œ.cat .ob ⟦ [] ⟧ctx = ð“œ.🙠@@ -92,8 +87,8 @@ module _ (𓜠: Model â„“ â„“') where ⟦_⟧S : Subst Δ Γ → ð“œ.cat [ ⟦ Δ ⟧ctx , ⟦ Γ ⟧ctx ] ⟦_⟧V : Val Γ S → ð“œ.cat [ ⟦ Γ ⟧ctx , ⟦ S ⟧ty ] - ⟦_⟧E : EvCtx Γ R S → Linear ⟦ Γ ⟧ctx [ ⟦ R ⟧ty , ⟦ S ⟧ty ] - ⟦_⟧C : Comp Γ S → ð“œ.cat [ ⟦ Γ ⟧ctx , ð“œ.T ⟅ ⟦ S ⟧ty ⟆ ] + ⟦_⟧E : EvCtx Γ R S → ð“œ.Linear ⟦ Γ ⟧ctx [ ⟦ R ⟧ty , ⟦ S ⟧ty ] + ⟦_⟧C : Comp Γ S → ð“œ.cat [ ⟦ Γ ⟧ctx , ð“œ.T ⟦ S ⟧ty ] ⟦ ids ⟧S = ð“œ.cat .id ⟦ γ ∘s δ ⟧S = ⟦ γ ⟧S ∘⟨ ð“œ.cat ⟩ ⟦ δ ⟧S @@ -120,19 +115,20 @@ module _ (𓜠: Model â„“ â„“') where ⟦ up S⊑T ⟧V = ⟦ S⊑T .ty-prec ⟧e ∘⟨ ð“œ.cat ⟩ ð“œ.π₂ ⟦ isSetVal V V' p q i j ⟧V = ð“œ.cat .isSetHom ⟦ V ⟧V ⟦ V' ⟧V (cong ⟦_⟧V p) (cong ⟦_⟧V q) i j - ⟦ ∙E ⟧E = Linear _ .id - ⟦ E ∘E F ⟧E = ⟦ E ⟧E ∘⟨ Linear _ ⟩ ⟦ F ⟧E - ⟦ ∘IdL {E = E} i ⟧E = Linear _ .⋆IdR ⟦ E ⟧E i - ⟦ ∘IdR {E = E} i ⟧E = Linear _ .⋆IdL ⟦ E ⟧E i - ⟦ ∘Assoc {E = E}{F = F}{F' = F'} i ⟧E = Linear _ .⋆Assoc ⟦ F' ⟧E ⟦ F ⟧E ⟦ E ⟧E i - ⟦ E [ γ ]e ⟧E = (⟦ γ ⟧S ^*) ⟪ ⟦ E ⟧E ⟫ - ⟦ substId {E = E} i ⟧E = id^* {E = ⟦ E ⟧E} i - ⟦ substAssoc {E = E}{γ = γ}{δ = δ} i ⟧E = comp^* {γ = ⟦ γ ⟧S} {δ = ⟦ δ ⟧S} {E = ⟦ E ⟧E} i - ⟦ ∙substDist {γ = γ} i ⟧E = (⟦ γ ⟧S ^*) .F-id i - ⟦ ∘substDist {E = E}{F = F}{γ = γ} i ⟧E = (⟦ γ ⟧S ^*) .F-seq ⟦ F ⟧E ⟦ E ⟧E i + ⟦ ∙E ⟧E = ð“œ.Linear _ .id + ⟦ E ∘E F ⟧E = ⟦ E ⟧E ∘⟨ ð“œ.Linear _ ⟩ ⟦ F ⟧E + ⟦ ∘IdL {E = E} i ⟧E = ð“œ.Linear _ .⋆IdR ⟦ E ⟧E i + ⟦ ∘IdR {E = E} i ⟧E = ð“œ.Linear _ .⋆IdL ⟦ E ⟧E i + ⟦ ∘Assoc {E = E}{F = F}{F' = F'} i ⟧E = ð“œ.Linear _ .⋆Assoc ⟦ F' ⟧E ⟦ F ⟧E ⟦ E ⟧E i + ⟦ E [ γ ]e ⟧E = (ð“œ.pull ⟦ γ ⟧S) ⟪ ⟦ E ⟧E ⟫ + ⟦ substId {E = E} i ⟧E = ð“œ.id^* i ⟪ ⟦ E ⟧E ⟫ + ⟦ substAssoc {E = E}{γ = γ}{δ = δ} i ⟧E = ð“œ.comp^* ⟦ γ ⟧S ⟦ δ ⟧S i ⟪ ⟦ E ⟧E ⟫ + ⟦ ∙substDist {γ = γ} i ⟧E = (ð“œ.pull ⟦ γ ⟧S) .F-id i + + ⟦ ∘substDist {E = E}{F = F}{γ = γ} i ⟧E = ð“œ.pull ⟦ γ ⟧S .F-seq ⟦ F ⟧E ⟦ E ⟧E i ⟦ bind M ⟧E = ⟦ M ⟧C ⟦ ret-η i ⟧E = {!!} - ⟦ dn S⊑T ⟧E = wkClosed ð“œ.🙠⟪ ⟦ S⊑T .ty-prec ⟧p ⟫ + ⟦ dn S⊑T ⟧E = ⟦ S⊑T .ty-prec ⟧p ∘⟨ ð“œ.cat ⟩ ð“œ.π₂ ⟦ isSetEvCtx E F p q i j ⟧E = ð“œ.cat .isSetHom ⟦ E ⟧E ⟦ F ⟧E (cong ⟦_⟧E p) (cong ⟦_⟧E q) i j - ⟦_⟧C = {!!} + ⟦ M ⟧C = {!M!} -- GitLab