From aacd555894fb559365774ddfa899656b95205e4e Mon Sep 17 00:00:00 2001 From: Chloe Brown Date: Tue, 20 Jul 2021 15:10:50 +0100 Subject: Complete Section 2 - Structure. --- src/Wasm/Expression/Values.agda | 118 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 118 insertions(+) create mode 100644 src/Wasm/Expression/Values.agda (limited to 'src/Wasm/Expression/Values.agda') diff --git a/src/Wasm/Expression/Values.agda b/src/Wasm/Expression/Values.agda new file mode 100644 index 0000000..d6bce60 --- /dev/null +++ b/src/Wasm/Expression/Values.agda @@ -0,0 +1,118 @@ +{-# OPTIONS --without-K --safe #-} + +------------------------------------------------------------------------ +-- 2.2 Values + +module Wasm.Expression.Values where + +open import Data.Char using (Char; toℕ) +open import Data.Empty using (⊥) +open import Data.Fin using (Fin) +open import Data.Nat using (ℕ; zero; suc; _^_; pred) +open import Wasm.Expression.Utilities using (BitWidth; 32Bit; 64Bit) + +infix 8 +_ -_ + +------------------------------------------------------------------------ +-- 2.2.1 Bytes + +Byte : Set +Byte = Fin 256 + +------------------------------------------------------------------------ +-- 2.2.2 Integers + +UnsignedInteger : ℕ → Set +UnsignedInteger n = Fin (2 ^ n) + +data SignedInteger : ℕ → Set where + +_ : ∀ {n} → Fin n → SignedInteger (suc n) + -[1+_] : ∀ {n} → Fin n → SignedInteger (suc n) + +UninterpretedInteger : ℕ → Set +UninterpretedInteger = UnsignedInteger + +------------------------------------------------------------------------ +-- 2.2.3 Floating-Point + +signif : BitWidth → ℕ +signif 32Bit = 23 +signif 64Bit = 52 + +expon : BitWidth → ℕ +expon 32Bit = 8 +expon 64Bit = 11 + +data FloatMag : BitWidth → Set where + nan : ∀ {w} → (n : Fin (pred (2 ^ (signif w)))) → FloatMag w + ∞ : ∀ {w} → FloatMag w + normal : ∀ {w} → (e : Fin (pred (pred (2 ^ (expon w))))) → (t : Fin (2 ^ (signif w))) → FloatMag w + subnormal : ∀ {w} → (t : Fin (pred (2 ^ (signif w)))) → FloatMag w + zero : ∀ {w} → FloatMag w + +data Float : BitWidth → Set where + +_ : ∀ {w} → FloatMag w → Float w + -_ : ∀ {w} → FloatMag w → Float w + +------------------------------------------------------------------------ +-- 2.2.4 Names + +module _ where + open import Data.Bool using (Bool; true; false; if_then_else_) + open import Data.List using (List; []; _∷_; length; map; sum) + open import Data.Nat using (_≤ᵇ_; _<_; z≤n; s≤s) + open import Data.Nat.Properties using (≤-step) + open import Data.Product using (_×_; _,_; proj₂; map₂) + open import Data.String using (String; toList) + + halfEven? : ℕ → Bool × ℕ + halfEven? zero = true , 0 + halfEven? (suc zero) = false , 0 + halfEven? (suc (suc n)) = map₂ suc (halfEven? n) + + n≢0⇒halfEven?