From 66fa99f84c918ad3a7680a6df141367c291ceaee Mon Sep 17 00:00:00 2001 From: Greg Brown Date: Fri, 21 Jan 2022 15:20:58 +0000 Subject: Add pseudocode as a data type. --- src/Helium/Data/Pseudocode/Core.agda | 220 +++++++++++++++++++++++++++++++++++ 1 file changed, 220 insertions(+) create mode 100644 src/Helium/Data/Pseudocode/Core.agda (limited to 'src/Helium/Data/Pseudocode/Core.agda') diff --git a/src/Helium/Data/Pseudocode/Core.agda b/src/Helium/Data/Pseudocode/Core.agda new file mode 100644 index 0000000..68b5449 --- /dev/null +++ b/src/Helium/Data/Pseudocode/Core.agda @@ -0,0 +1,220 @@ +------------------------------------------------------------------------ +-- Agda Helium +-- +-- Definition of the Armv8-M pseudocode. +------------------------------------------------------------------------ + +{-# OPTIONS --safe --without-K #-} + +module Helium.Data.Pseudocode.Core where + +open import Data.Bool using (Bool) +open import Data.Fin using (Fin; #_) +open import Data.Nat as ℕ using (ℕ; suc) +open import Data.Product using (∃; _,_; proj₂; uncurry) +open import Data.Vec using (Vec; []; _∷_; lookup; map) +open import Data.Vec.Relation.Unary.All using (All; []; _∷_; reduce; all?) +open import Function as F using (_∘_; _∘′_; _∋_) +open import Relation.Binary.PropositionalEquality using (_≡_; refl) +open import Relation.Nullary using (yes; no) +open import Relation.Nullary.Decidable.Core using (True; map′) +open import Relation.Nullary.Product using (_×-dec_) +open import Relation.Unary using (Decidable) + +--- The set of types and boolean properties of them +data Type : Set where + unit : Type + bool : Type + int : Type + fin : (n : ℕ) → Type + real : Type + bits : (n : ℕ) → Type + enum : (n : ℕ) → Type + tuple : ∀ n → Vec Type n → Type + array : Type → (n : ℕ) → Type + +bit : Type +bit = bits 1 + +data HasEquality : Type → Set where + bool : HasEquality bool + int : HasEquality int + fin : ∀ {n} → HasEquality (fin n) + real : HasEquality real + bits : ∀ {n} → HasEquality (bits n) + enum : ∀ {n} → HasEquality (enum n) + +hasEquality? : Decidable HasEquality +hasEquality? unit = no (λ ()) +hasEquality? bool = yes bool +hasEquality? int = yes int +hasEquality? (fin n) = yes fin +hasEquality? real = yes real +hasEquality? (bits n) = yes bits +hasEquality? (enum n) = yes enum +hasEquality? (tuple n x) = no (λ ()) +hasEquality? (array t n) = no (λ ()) + +data IsNumeric : Type → Set where + int : IsNumeric int + real : IsNumeric real + +isNumeric? : Decidable IsNumeric +isNumeric? unit = no (λ ()) +isNumeric? bool = no (λ ()) +isNumeric? int = yes int +isNumeric? real = yes real +isNumeric? (fin n) = no (λ ()) +isNumeric? (bits n) = no (λ ()) +isNumeric? (enum n) = no (λ ()) +isNumeric? (tuple n x) = no (λ ()) +isNumeric? (array t n) = no (λ ()) + +combineNumeric : ∀ t₁ t₂ → {isNumeric₁ : True (isNumeric? t₁)} → {isNumeric₂ : True (isNumeric? t₂)} → Type +combineNumeric int int = int +combineNumeric int real = real +combineNumeric real _ = real + +data Sliced : Set where + bits : Sliced + array : Type → Sliced + +asType : Sliced → ℕ → Type +asType bits n = bits n +asType (array t) n = array t n + +elemType : Sliced → Type +elemType bits = bit +elemType (array t) = t + +--- Literals + +data Literal : Type → Set where + _′b : Bool → Literal bool + _′i : ℕ → Literal int + _′f : ∀ {n} → Fin n → Literal (fin n) + _′r : ℕ → Literal real + _′x : ∀ {n} → Vec Bool n → Literal (bits n) + _′e : ∀ {n} → Fin n → Literal (enum n) + +--- Expressions, references, statements, functions and procedures + +module Code {o} (Σ : Vec Type o) where + data Expression {n} (Γ : Vec Type n) : Type → Set + data CanAssign {n} {Γ} : ∀ {t} → Expression {n} Γ t → Set + canAssign? : ∀ {n Γ t} → Decidable (CanAssign {n} {Γ} {t}) + canAssignAll? : ∀ {n Γ m ts} → Decidable {A = All (Expression {n} Γ) {m} ts} (All (CanAssign ∘ proj₂) ∘ (reduce (λ {x} e → x , e))) + data Statement {n} (Γ : Vec Type n) (ret : Type) : Set + data Function {n} (Γ : Vec Type n) (ret : Type) : Set + data Procedure {n} (Γ : Vec Type n) : Set + + infix 8 -_ + infixr 7 _^_ + infixl 6 _*_ _/_ _and_ + infixl 5 _+_ _or_ _&&_ _||_ _∶_ + infix 4 _≟_ _