diff --git a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda index 1e6ac96d013baf8b5c25db19c59954125cba74b6..295be28ae01e7ff86be2c1ae8ad7e886f2937e8a 100644 --- a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda +++ b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient.agda @@ -18,7 +18,7 @@ open import Cubical.Categories.Limits.BinProduct open import Cubical.Categories.Limits.BinProduct.More open import Cubical.Categories.Limits.BinCoproduct open import Cubical.Categories.Exponentials -open import Cubical.Categories.Monad.ExtensionSystem +open import Cubical.Categories.Monad.ExtensionSystem as Monad open import Cubical.Categories.Comonad.Instances.Environment open import Cubical.Categories.Monad.Strength.Cartesian.ExtensionSystem @@ -66,6 +66,7 @@ record Model â„“ â„“' : Type (â„“-suc (â„“-max â„“ â„“')) where -- TODO: should this go into notation? Linear = PKleisli ClLinear = Kleisli cat (StrongMonad→Monad binProd term monad) + clBind = G cat (StrongMonad→Monad binProd term monad) .F-hom field -- a weak model of the natural numbers, but good enough for our syntax @@ -93,4 +94,3 @@ record Model â„“ â„“' : Type (â„“-suc (â„“-max â„“ â„“')) where ℧-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/Computations.agda b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Computations.agda new file mode 100644 index 0000000000000000000000000000000000000000..0c22e04c40db4fcd75556c16ac0f1f09c77b927b --- /dev/null +++ b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Computations.agda @@ -0,0 +1,48 @@ +{-# OPTIONS --cubical #-} +module Semantics.Abstract.TermModel.Convenient.Computations where + +-- Computations form a covariant presheaf on evaluation contexts + +open import Cubical.Foundations.Prelude +open import Cubical.Foundations.Structure +open import Cubical.Categories.Category +open import Cubical.Categories.Functor +open import Cubical.Categories.Instances.Sets +open import Cubical.Categories.Constructions.BinProduct +open import Cubical.Categories.NaturalTransformation +open import Cubical.Categories.Limits.Terminal +open import Cubical.Categories.Limits.BinProduct +open import Cubical.Categories.Limits.BinCoproduct +open import Cubical.Categories.Monad.ExtensionSystem +open import Cubical.Categories.Monad.Strength.Cartesian.ExtensionSystem +open import Cubical.Categories.Comonad.Instances.Environment + +open import Semantics.Abstract.TermModel.Convenient + +private + variable + â„“ â„“' : Level + +open Category +open Functor +open NatTrans +open BinCoproduct +open BinProduct + +module _ (𓜠: Model â„“ â„“') where + open Model 𓜠+ open StrongExtensionSystem + COMP : ∀ Γ → Functor (Linear Γ) (SET â„“') + COMP Γ .F-ob a = (ClLinear [ Γ , a ]) , cat .isSetHom + COMP Γ .F-hom E M = bind E ∘⟨ cat ⟩ (cat .id ,p M) + COMP Γ .F-id = funExt λ M → congâ‚‚ (comp' cat) (ExtensionSystemFor.bind-r (monad .systems Γ)) refl ∙ ×β₂ + COMP Γ .F-seq f g = funExt (λ M → congâ‚‚ (cat ._⋆_) refl (sym (ExtensionSystemFor.bind-comp (monad .systems Γ))) + ∙ sym (cat .⋆Assoc _ _ _) ∙ congâ‚‚ (comp' cat) refl (,p-natural ∙ congâ‚‚ _,p_ ×β₠refl) ) + + COMP-Enriched : ∀ {Δ Γ a b} (γ : cat [ Δ , Γ ]) (M : ⟨ COMP Γ ⟅ a ⟆ ⟩) (E : Linear Γ [ a , b ]) + → (COMP Γ ⟪ E ⟫) M ∘⟨ cat ⟩ γ ≡ (COMP Δ ⟪ (γ ^*) ⟪ E ⟫ ⟫) (M ∘⟨ cat ⟩ γ) + COMP-Enriched γ M E = + sym (⋆Assoc cat _ _ _) + ∙ congâ‚‚ (comp' cat) refl ((,p-natural ∙ congâ‚‚ _,p_ (cat .⋆IdR _ ∙ (sym (cat .⋆IdL _) ∙ congâ‚‚ (comp' cat) refl (sym ×βâ‚)) ∙ cat .⋆Assoc _ _ _) (sym ×β₂)) ∙ sym ,p-natural) + ∙ ⋆Assoc cat _ _ _ + ∙ congâ‚‚ (seq' cat) refl (bind-natural monad) diff --git a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda index a3b950340c84e6d924afe0dffe8701a4199c9a5b..8c9c6028acc35df0bdd8fdfe117fe67191a7919b 100644 --- a/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda +++ b/formalizations/guarded-cubical/Semantics/Abstract/TermModel/Convenient/Semantics.agda @@ -23,6 +23,7 @@ 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.Computations private variable @@ -65,18 +66,18 @@ module _ (𓜠: Model â„“ â„“') where -- η (e y') ⟦ c ⇀ d ⟧e = ð“œ.lda ((ð“œ.ClLinear .id ∘⟨ ð“œ.cat ⟩ ⟦ d ⟧e) ∘⟨ ð“œ.ClLinear ⟩ ð“œ.app ∘⟨ ð“œ.Linear _ ⟩ - {!!}) + ⟦ c ⟧p ∘⟨ ð“œ.cat ⟩ ð“œ.π₂) ⟦ inj-nat ⟧e = ð“œ.inj ∘⟨ ð“œ.cat ⟩ ð“œ.σ1 ⟦ inj-arr c ⟧e = ð“œ.inj ∘⟨ ð“œ.cat ⟩ ð“œ.σ2 ∘⟨ ð“œ.cat ⟩ ⟦ c ⟧e ⟦ 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 ð“œ.|| ð“œ.℧) ∘⟨ ð“œ.ClLinear ⟩ ð“œ.prj - ⟦ inj-arr c ⟧p = {!!} -- (ð“œ.℧ ð“œ.|| ⟦ c ⟧p) ∘⟨ ð“œ.ClLinear ⟩ ð“œ.prj + ⟦ c ⇀ d ⟧p = ð“œ.ClLinear .id ∘⟨ ð“œ.cat ⟩ ⇒F ⟪ ⟦ c ⟧e , ⟦ d ⟧p' ⟫ + ⟦ inj-nat ⟧p = (ð“œ.ClLinear .id ð“œ.|| ð“œ.℧) ∘⟨ ð“œ.ClLinear ⟩ ð“œ.prj + ⟦ inj-arr c ⟧p = (ð“œ.℧ ð“œ.|| ⟦ c ⟧p) ∘⟨ ð“œ.ClLinear ⟩ ð“œ.prj - ⟦ c ⟧p' = {!!} -- ð“œ.bind .N-ob _ ⟦ c ⟧p + ⟦ c ⟧p' = ð“œ.clBind ⟦ c ⟧p ⟦_⟧ctx : Ctx → ð“œ.cat .ob ⟦ [] ⟧ctx = ð“œ.🙠@@ -85,10 +86,10 @@ module _ (𓜠: Model â„“ â„“') where -- The term syntax -- substitutions, values, ev ctxts, terms - ⟦_⟧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 ] + ⟦_⟧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 → ð“œ.ClLinear [ ⟦ Γ ⟧ctx , ⟦ S ⟧ty ] ⟦ ids ⟧S = ð“œ.cat .id ⟦ γ ∘s δ ⟧S = ⟦ γ ⟧S ∘⟨ ð“œ.cat ⟩ ⟦ δ ⟧S @@ -131,4 +132,26 @@ module _ (𓜠: Model â„“ â„“') where ⟦ 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 - ⟦ M ⟧C = {!M!} + ⟦ E [ M ]∙ ⟧C = (COMP _ 𓜠⟪ ⟦ E ⟧E ⟫) ⟦ M ⟧C + ⟦ plugId {M = M} i ⟧C = (COMP _ 𓜠.F-id i) ⟦ M ⟧C + ⟦ plugAssoc {F = F}{E = E}{M = M} i ⟧C = (COMP _ 𓜠.F-seq ⟦ E ⟧E ⟦ F ⟧E i) ⟦ M ⟧C + + ⟦ M [ γ ]c ⟧C = ⟦ M ⟧C ∘⟨ ð“œ.cat ⟩ ⟦ γ ⟧S + ⟦ substId {M = M} i ⟧C = ð“œ.cat .⋆IdL ⟦ M ⟧C i + ⟦ substAssoc {M = M}{δ = δ}{γ = γ} i ⟧C = ð“œ.cat .⋆Assoc ⟦ γ ⟧S ⟦ δ ⟧S ⟦ M ⟧C i + ⟦ substPlugDist {E = E}{M = M}{γ = γ} i ⟧C = COMP-Enriched 𓜠⟦ γ ⟧S ⟦ M ⟧C ⟦ E ⟧E i + ⟦ err ⟧C = ð“œ.℧ + ⟦ strictness {E = E} i ⟧C = ð“œ.℧-homo ⟦ E ⟧E {!i!} + + ⟦ ret ⟧C = ð“œ.Linear _ .id + ⟦ ret-β i ⟧C = {!!} + + ⟦ app ⟧C = {!!} + ⟦ fun-β i ⟧C = {!!} + + ⟦ matchNat Mz Ms ⟧C = {!!} + ⟦ matchNatβz Mz Ms i ⟧C = {!!} + ⟦ matchNatβs Mz Ms V i ⟧C = {!!} + ⟦ matchNatη i ⟧C = {!!} + + ⟦ isSetComp M N p q i j ⟧C = ð“œ.cat .isSetHom ⟦ M ⟧C ⟦ N ⟧C (cong ⟦_⟧C p) (cong ⟦_⟧C q) i j