summaryrefslogtreecommitdiff
path: root/src/Helium/Semantics/Denotational/Core.agda
diff options
context:
space:
mode:
authorGreg Brown <greg.brown@cl.cam.ac.uk>2022-01-12 18:32:48 +0000
committerGreg Brown <greg.brown@cl.cam.ac.uk>2022-01-12 18:32:48 +0000
commit9901eac3a64249b58789588385a10df9802c9f42 (patch)
tree4f25e0ac2ceaaec2cc077738f8cec275e815b22b /src/Helium/Semantics/Denotational/Core.agda
parentd5f3e7bc675a07bd04c746512c6f1b0b1250b55e (diff)
Eliminate even more state from the denotational semantics.
Diffstat (limited to 'src/Helium/Semantics/Denotational/Core.agda')
-rw-r--r--src/Helium/Semantics/Denotational/Core.agda87
1 files changed, 36 insertions, 51 deletions
diff --git a/src/Helium/Semantics/Denotational/Core.agda b/src/Helium/Semantics/Denotational/Core.agda
index d92c1b2..9b6a3c8 100644
--- a/src/Helium/Semantics/Denotational/Core.agda
+++ b/src/Helium/Semantics/Denotational/Core.agda
@@ -18,7 +18,7 @@ open import Data.Nat using (ℕ; zero; suc)
import Data.Nat.Properties as ℕₚ
open import Data.Product hiding (_<*>_; _,′_)
open import Data.Product.Nary.NonDependent
-open import Data.Sum using (_⊎_; inj₁; inj₂; [_,_]′)
+open import Data.Sum using (_⊎_; inj₁; inj₂; fromInj₂; [_,_]′)
open import Data.Unit using (⊤)
open import Level renaming (suc to ℓsuc) hiding (zero)
open import Function using (_∘_; _∘₂_; _|>_)
@@ -30,65 +30,43 @@ private
ℓ ℓ₁ ℓ₂ : Level
τ τ′ : Set ℓ
- mapAll : ∀ {m ls} {Γ : Sets m ls} {l l′}
- (f : ∀ {ℓ} → Set ℓ → Set (l ℓ))
- (g : ∀ {ℓ} → Set ℓ → Set (l′ ℓ))
- (h : ∀ {a} {A : Set a} → f A → g A) →
- Product⊤ m (smap l f m Γ) →
- Product⊤ m (smap l′ g m Γ)
- mapAll {zero} f g h xs = xs
- mapAll {suc m} f g h (x , xs) = h x , mapAll f g h xs
-
update : ∀ {n ls} {Γ : Sets n ls} i → Projₙ Γ i → Product⊤ n Γ → Product⊤ n Γ
update zero y (_ , xs) = y , xs
update (suc i) y (x , xs) = x , update i y xs
-PureExpr : ∀ n {ls} → Sets n ls → Set ℓ → Set (ℓ ⊔ ⨆ n ls ⊔ ℓ′)
-PureExpr n Γ τ = (σ : State) → (ρ : Product⊤ n Γ) → τ
-
Expr : ∀ n {ls} → Sets n ls → Set ℓ → Set (ℓ ⊔ ⨆ n ls ⊔ ℓ′)
-Expr n Γ τ = (σ : State) → (ρ : Product⊤ n Γ) → State × τ
-
+Expr n Γ τ = (σ : State) → (ρ : Product⊤ n Γ) → τ
record Reference n {ls} (Γ : Sets n ls) (τ : Set ℓ) : Set (ℓ ⊔ ⨆ n ls ⊔ ℓ′) where
field
- get : PureExpr n Γ τ
- set : τ → (σ : State) → (ρ : Product⊤ n Γ) → State × Product⊤ n Γ
+ get : Expr n Γ τ
+ set : τ → Expr n Γ (State × Product⊤ n Γ)
Function : ∀ n {ls} → Sets n ls → Set ℓ → Set (ℓ ⊔ ⨆ n ls ⊔ ℓ′)
-Function = Expr
+Function n Γ τ = Expr n Γ τ
Procedure : ∀ n {ls} → Sets n ls → Set (⨆ n ls ⊔ ℓ′)
-Procedure n Γ = Function n Γ ⊤
+Procedure n Γ = Expr n Γ State
-Block : ∀ n {ls} → Sets n ls → Set ℓ → Set (ℓ ⊔ ⨆ n ls ⊔ ℓ′)
-Block n Γ τ = (σ : State) → (ρ : Product⊤ n Γ) → State × (Product⊤ n Γ ⊎ τ)
+Statement : ∀ n {ls} → Sets n ls → Set ℓ → Set (ℓ ⊔ ⨆ n ls ⊔ ℓ′)
+Statement n Γ τ = Expr n Γ (State × (Product⊤ n Γ ⊎ τ))
-- Expressions
-pure′ : ∀ {n ls} {Γ : Sets n ls} → τ → PureExpr n Γ τ
-pure′ v σ ρ = v
-
pure : ∀ {n ls} {Γ : Sets n ls} → τ → Expr n Γ τ
-pure v σ ρ = σ , v
-
-↓_ : ∀ {n ls} {Γ : Sets n ls} → PureExpr n Γ τ → Expr n Γ τ
-(↓ e) σ ρ = σ , e σ ρ
+pure v σ ρ = v
_<*>_ : ∀ {n ls} {Γ : Sets n ls} → Expr n Γ (τ → τ′) → Expr n Γ τ → Expr n Γ τ′
-_<*>_ f e σ ρ = f σ ρ |> λ (σ , f) → map₂ f (e σ ρ)
+_<*>_ f e σ ρ = f σ ρ |> λ f → f (e σ ρ)
-!_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → PureExpr n Γ τ
+!_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Expr n Γ τ
! r = Reference.get r
-↓!_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Expr n Γ τ
-↓! r = ↓ ! r
-
call : ∀ {m n ls₁ ls₂} {Γ : Sets m ls₁} {Δ : Sets n ls₂} → Function m Γ τ → Expr n Δ (Product m Γ) → Expr n Δ τ
-call f e σ ρ = e σ ρ |> map₂ (toProduct⊤ _) |> uncurry f
+call f e σ ρ = e σ ρ |> toProduct⊤ _ |> f σ
declare : ∀ {n ls} {Γ : Sets n ls} → Expr n Γ τ → Expr (suc n) (τ , Γ) τ′ → Expr n Γ τ′
-declare e s σ ρ = e σ ρ |> map₂ (_, ρ) |> uncurry s
+declare e s σ ρ = e σ ρ |> _, ρ |> s σ
-- References
@@ -98,7 +76,7 @@ var i = record
; set = λ v → curry (map₂ (update i v))
}
-!#_ : ∀ {n ls} {Γ : Sets n ls} m {m<n : True (suc m ℕₚ.≤? n)} → PureExpr n Γ (Projₙ Γ (#_ m {n} {m<n}))
+!#_ : ∀ {n ls} {Γ : Sets n ls} m {m<n : True (suc m ℕₚ.≤? n)} → Expr n Γ (Projₙ Γ (#_ m {n} {m<n}))
(!# m) {m<n} = ! var (#_ m {m<n = m<n})
_,′_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Reference n Γ τ′ → Reference n Γ (τ × τ′)
@@ -107,37 +85,44 @@ _,′_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Reference n Γ τ
; set = λ (x , y) σ ρ → uncurry (Reference.set &y y) (Reference.set &x x σ ρ)
}
--- Blocks
+-- Statements
-infixr 1 _∙_
+infixl 1 _∙_ _∙return_
+infix 1 _∙end
infixl 2 if_then_else_
infix 4 _≔_ _⟵_
-skip : ∀ {n ls} {Γ : Sets n ls} → Block n Γ τ
+skip : ∀ {n ls} {Γ : Sets n ls} → Statement n Γ τ
skip σ ρ = σ , inj₁ ρ
-invoke : ∀ {m n ls₁ ls₂} {Γ : Sets m ls₁} {Δ : Sets n ls₂} → Procedure m Γ → Expr n Δ (Product m Γ) → Block n Δ τ
-invoke f e σ ρ = call f e σ ρ |> map₂ (λ _ → inj₁ ρ)
+invoke : ∀ {m n ls₁ ls₂} {Γ : Sets m ls₁} {Δ : Sets n ls₂} → Procedure m Γ → Expr n Δ (Product m Γ) → Statement n Δ τ
+invoke f e σ ρ = call f e σ ρ |> _, inj₁ ρ
-_≔_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Expr n Γ τ → Block n Γ τ′
-(&x ≔ e) σ ρ = e σ ρ |> λ (σ , x) → Reference.set &x x σ ρ |> map₂ inj₁
+_≔_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Expr n Γ τ → Statement n Γ τ′
+(&x ≔ e) σ ρ = e σ ρ |> λ x → Reference.set &x x σ ρ |> map₂ inj₁
-_⟵_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Op₁ τ → Block n Γ τ′
-&x ⟵ e = &x ≔ ⦇ e (↓ (! &x)) ⦈
+_⟵_ : ∀ {n ls} {Γ : Sets n ls} → Reference n Γ τ → Op₁ τ → Statement n Γ τ′
+&x ⟵ e = &x ≔ ⦇ e (! &x) ⦈
-if_then_else_ : ∀ {n ls} {Γ : Sets n ls} → Expr n Γ Bool → Block n Γ τ → Block n Γ τ → Block n Γ τ
-(if e then b₁ else b₂) σ ρ = e σ ρ |> λ (σ , b) → Bool.if b then b₁ σ ρ else b₂ σ ρ
+if_then_else_ : ∀ {n ls} {Γ : Sets n ls} → Expr n Γ Bool → Statement n Γ τ → Statement n Γ τ → Statement n Γ τ
+(if e then b₁ else b₂) σ ρ = e σ ρ |> (Bool.if_then b₁ σ ρ else b₂ σ ρ)
-for : ∀ {n ls} {Γ : Sets n ls} m → Block (suc n) (Fin m , Γ) τ → Block n Γ τ
+for : ∀ {n ls} {Γ : Sets n ls} m → Statement (suc n) (Fin m , Γ) τ → Statement n Γ τ
for zero b σ ρ = σ , inj₁ ρ
for (suc m) b σ ρ with b σ (zero , ρ)
... | σ′ , inj₂ x = σ′ , inj₂ x
... | σ′ , inj₁ (_ , ρ′) = for m b′ σ′ ρ′
where
- b′ : Block (suc _) (Fin m , _) _
+ b′ : Statement (suc _) (Fin m , _) _
b′ σ (i , ρ) with b σ (suc i , ρ)
... | σ′ , inj₂ x = σ′ , inj₂ x
... | σ′ , inj₁ (_ , ρ′) = σ′ , inj₁ (i , ρ′)
-_∙_ : ∀ {n ls} {Γ : Sets n ls} → Block n Γ τ → Expr n Γ τ → Expr n Γ τ
-(b ∙ e) σ τ = b σ τ |> λ (σ , ρ⊎x) → [ e σ , σ ,_ ]′ ρ⊎x
+_∙_ : ∀ {n ls} {Γ : Sets n ls} → Statement n Γ τ → Statement n Γ τ → Statement n Γ τ
+(b₁ ∙ b₂) σ ρ = b₁ σ ρ |> λ (σ , ρ⊎x) → [ b₂ σ , (σ ,_) ∘ inj₂ ]′ ρ⊎x
+
+_∙end : ∀ {n ls} {Γ : Sets n ls} → Statement n Γ ⊤ → Procedure n Γ
+_∙end s σ ρ = s σ ρ |> proj₁
+
+_∙return_ : ∀ {n ls} {Γ : Sets n ls} → Statement n Γ τ → Expr n Γ τ → Function n Γ τ
+(s ∙return e) σ ρ = s σ ρ |> λ (σ , ρ⊎x) → fromInj₂ (e σ) ρ⊎x