TypesSemantics.lagda 3.8 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118
  1. %include agda.fmt
  2. \subsection{The semantics of types}
  3. \label{types-semantics}
  4. \AgdaHide{
  5. \begin{code}
  6. {-# OPTIONS --sized-types #-}
  7. module Issue854.TypesSemantics where
  8. import Level
  9. open import Function
  10. open import Data.Empty
  11. open import Data.Unit
  12. open import Data.Sum
  13. open import Data.Product
  14. open import Data.List
  15. open import Data.List.Relation.Unary.Any
  16. open import Data.Container hiding (_∈_) renaming (μ to μ^C; ⟦_⟧ to ⟦_⟧^C; _▷_ to _◃_)
  17. open import Data.Container.Combinator using () renaming (_×_ to _×^C_)
  18. open import Data.Container.FreeMonad as FM
  19. hiding (_⋆_) renaming (_⋆C_ to _⋆^CC_)
  20. open import Category.Monad
  21. open import Relation.Binary.PropositionalEquality
  22. open import Data.List.Membership.Propositional
  23. open import Issue854.Types
  24. open import Issue854.Context
  25. open import Issue854.EilenbergMooreAlgebra
  26. \end{code}
  27. }
  28. \begin{code}
  29. mutual
  30. ⟦_⟧^VType : VType → Set
  31. ⟦ ⁅ C ⁆ ⟧^VType = Carrier ⟦ C ⟧^CType
  32. ⟦ 𝟙 ⟧^VType = ⊤
  33. ⟦ U ⊗ V ⟧^VType = ⟦ U ⟧^VType × ⟦ V ⟧^VType
  34. ⟦ 𝟘 ⟧^VType = ⊥
  35. ⟦ U ⊕ V ⟧^VType = ⟦ U ⟧^VType ⊎ ⟦ V ⟧^VType
  36. ⟦ μ Δ ⟧^VType = μ^C ⌊ Δ ⌋^Sig
  37. ⌊_⌋^Sig : Sig → Container _ _
  38. ⌊ Σ ⌋^Sig = Sh Σ ◃ Pos Σ
  39. Sh : Sig → Set
  40. Sh [] = ⊥
  41. Sh ((P , _) ∷ Σ) = ⟦ P ⟧^VType ⊎ Sh Σ
  42. Pos : (Σ : Sig) → Sh Σ → Set
  43. Pos [] ()
  44. Pos ((_ , A) ∷ Σ) (inj₁ p) = ⟦ A ⟧^VType
  45. Pos ((_ , A) ∷ Σ) (inj₂ s) = Pos Σ s
  46. ∣_∣^Sig : CType → Container _ _
  47. ∣ Σ ⋆ V ∣^Sig = ⌊ Σ ⌋^Sig
  48. ∣ V ⇒ C ∣^Sig = ∣ C ∣^Sig
  49. ∣ ⊤′ ∣^Sig = ⌊ [] ⌋^Sig
  50. ∣ B & C ∣^Sig = ∣ B ∣^Sig ×^C ∣ C ∣^Sig
  51. ⟦_⟧^CType : (C : CType) → ∣ C ∣^Sig -Alg
  52. ⟦ Σ ⋆ V ⟧^CType = ⋆-alg ⟦ V ⟧^VType
  53. ⟦ V ⇒ C ⟧^CType = ⟦ V ⟧^VType ⇒-alg ⟦ C ⟧^CType
  54. ⟦ ⊤′ ⟧^CType = ⋆-alg ⊤
  55. ⟦ B & C ⟧^CType = ⟦ B ⟧^CType ⊙-alg ⟦ C ⟧^CType
  56. -- What if we wanted to add cofree comonads/free cims? The following is
  57. -- a bit complicated, but seems to suggest that it should be possible?
  58. mutual
  59. data ∣_∣^Alg : CType → Set where
  60. alg : ∀ {C} → ∣ C ∣^Alg
  61. ∣_∣^Alg′ : (C : CType) → ∣ C ∣^Alg
  62. ∣ C ∣^Alg′ = alg
  63. El^Alg : ∀ {C} → ∣ C ∣^Alg → Set₁
  64. El^Alg {C} alg = ∣ C ∣^Sig -Alg
  65. ⟦_⟧^VType′ : VType → Set
  66. ⟦ ⁅ C ⁆ ⟧^VType′ = Carrier′ {C}{∣ C ∣^Alg′} (⟦ C ⟧^CType′ {A = ∣ C ∣^Alg′ })
  67. ⟦ 𝟙 ⟧^VType′ = ⊤
  68. ⟦ U ⊗ V ⟧^VType′ = ⟦ U ⟧^VType × ⟦ V ⟧^VType
  69. ⟦ 𝟘 ⟧^VType′ = ⊥
  70. ⟦ U ⊕ V ⟧^VType′ = ⟦ U ⟧^VType ⊎ ⟦ V ⟧^VType
  71. ⟦ μ Δ ⟧^VType′ = μ^C ⌊ Δ ⌋^Sig
  72. Carrier′ : ∀ {C} {A : ∣ C ∣^Alg} → El^Alg A → Set
  73. Carrier′ {C} {alg} A = Carrier A
  74. ⟦_⟧^CType′ : (C : CType) → {A : ∣ C ∣^Alg} → El^Alg A
  75. ⟦ Σ ⋆ V ⟧^CType′ {alg} = ⋆-alg ⟦ V ⟧^VType
  76. ⟦ V ⇒ C ⟧^CType′ {alg} = ⟦ V ⟧^VType ⇒-alg ⟦ C ⟧^CType
  77. ⟦ ⊤′ ⟧^CType′ {alg} = ⋆-alg ⊤
  78. ⟦ B & C ⟧^CType′ {alg} = ⟦ B ⟧^CType ⊙-alg ⟦ C ⟧^CType
  79. ⟦_⟧^Ctx : Ctx → Set
  80. ⟦ [] ⟧^Ctx = ⊤
  81. ⟦ Γ ▻ V ⟧^Ctx = ⟦ Γ ⟧^Ctx × ⟦ V ⟧^VType
  82. ⟦_⟧^CTypes : List CType → Set
  83. ⟦ [] ⟧^CTypes = ⊤
  84. ⟦ C ∷ Cs ⟧^CTypes = Carrier ⟦ C ⟧^CType × ⟦ Cs ⟧^CTypes
  85. ⟦_⟧^Sig : Sig → (Set → Set)
  86. ⟦ Σ ⟧^Sig = ⟦ ⌊ Σ ⌋^Sig ⟧^C
  87. sh : ∀ {Σ P A} → (P , A) ∈ Σ → ⟦ P ⟧^VType → Sh Σ
  88. sh (here refl) p = inj₁ p
  89. sh (there m) p = inj₂ (sh m p)
  90. ar : ∀ {Σ P A}(m : (P , A) ∈ Σ)(p : ⟦ P ⟧^VType) → Pos Σ (sh m p) →
  91. ⟦ A ⟧^VType
  92. ar (here refl) _ a = a
  93. ar (there m) p a = ar m p a
  94. \end{code}