2022-03-02 22:02:51 +00:00
|
|
|
|
{-# OPTIONS --rewriting #-}
|
|
|
|
|
|
2022-02-09 23:14:29 +00:00
|
|
|
|
module Luau.Heap where
|
|
|
|
|
|
2022-03-02 22:02:51 +00:00
|
|
|
|
open import Agda.Builtin.Equality using (_≡_; refl)
|
|
|
|
|
open import FFI.Data.Maybe using (Maybe; just; nothing)
|
|
|
|
|
open import FFI.Data.Vector using (Vector; length; snoc; empty; lookup-snoc-not)
|
|
|
|
|
open import Luau.Addr using (Addr; _≡ᴬ_)
|
2022-02-09 23:14:29 +00:00
|
|
|
|
open import Luau.Var using (Var)
|
2022-03-02 22:02:51 +00:00
|
|
|
|
open import Luau.Syntax using (Block; Expr; Annotated; FunDec; nil; function_is_end)
|
|
|
|
|
open import Properties.Equality using (_≢_; trans)
|
|
|
|
|
open import Properties.Remember using (remember; _,_)
|
|
|
|
|
open import Properties.Dec using (yes; no)
|
|
|
|
|
|
|
|
|
|
-- Heap-allocated objects
|
|
|
|
|
data Object (a : Annotated) : Set where
|
2022-02-09 23:14:29 +00:00
|
|
|
|
|
2022-03-02 22:02:51 +00:00
|
|
|
|
function_is_end : FunDec a → Block a → Object a
|
2022-02-09 23:14:29 +00:00
|
|
|
|
|
2022-02-12 01:03:26 +00:00
|
|
|
|
Heap : Annotated → Set
|
2022-03-02 22:02:51 +00:00
|
|
|
|
Heap a = Vector (Object a)
|
2022-02-09 23:14:29 +00:00
|
|
|
|
|
2022-03-02 22:02:51 +00:00
|
|
|
|
data _≡_⊕_↦_ {a} : Heap a → Heap a → Addr → Object a → Set where
|
2022-02-09 23:14:29 +00:00
|
|
|
|
|
|
|
|
|
defn : ∀ {H val} →
|
|
|
|
|
|
|
|
|
|
-----------------------------------
|
|
|
|
|
(snoc H val) ≡ H ⊕ (length H) ↦ val
|
|
|
|
|
|
2022-03-02 22:02:51 +00:00
|
|
|
|
_[_] : ∀ {a} → Heap a → Addr → Maybe (Object a)
|
2022-02-12 01:03:26 +00:00
|
|
|
|
_[_] = FFI.Data.Vector.lookup
|
2022-02-09 23:14:29 +00:00
|
|
|
|
|
2022-02-12 01:03:26 +00:00
|
|
|
|
∅ : ∀ {a} → Heap a
|
|
|
|
|
∅ = empty
|
2022-02-09 23:14:29 +00:00
|
|
|
|
|
2022-03-02 22:02:51 +00:00
|
|
|
|
data AllocResult a (H : Heap a) (V : Object a) : Set where
|
2022-02-12 01:03:26 +00:00
|
|
|
|
ok : ∀ b H′ → (H′ ≡ H ⊕ b ↦ V) → AllocResult a H V
|
2022-02-09 23:14:29 +00:00
|
|
|
|
|
2022-02-12 01:03:26 +00:00
|
|
|
|
alloc : ∀ {a} H V → AllocResult a H V
|
2022-02-09 23:14:29 +00:00
|
|
|
|
alloc H V = ok (length H) (snoc H V) defn
|
|
|
|
|
|
2022-02-12 01:03:26 +00:00
|
|
|
|
next : ∀ {a} → Heap a → Addr
|
2022-02-09 23:14:29 +00:00
|
|
|
|
next = length
|
|
|
|
|
|
2022-03-02 22:02:51 +00:00
|
|
|
|
allocated : ∀ {a} → Heap a → Object a → Heap a
|
2022-02-09 23:14:29 +00:00
|
|
|
|
allocated = snoc
|
|
|
|
|
|
2022-03-02 22:02:51 +00:00
|
|
|
|
lookup-not-allocated : ∀ {a} {H H′ : Heap a} {b c O} → (H′ ≡ H ⊕ b ↦ O) → (c ≢ b) → (H [ c ] ≡ H′ [ c ])
|
|
|
|
|
lookup-not-allocated {H = H} {O = O} defn p = lookup-snoc-not O H p
|