------------------------------------------------------------------------ -- Agda Helium -- -- Definition of the Armv8-M pseudocode. ------------------------------------------------------------------------ {-# OPTIONS --safe --without-K #-} module Helium.Data.Pseudocode.Core where open import Data.Bool using (Bool; true; false) open import Data.Fin using (Fin; #_) open import Data.Nat as ℕ using (ℕ; zero; suc) open import Data.Nat.Properties using (+-comm) open import Data.Product using (∃; _,_; proj₂; uncurry) open import Data.Sum using ([_,_]′; inj₁; inj₂) open import Data.Vec using (Vec; []; _∷_; lookup; map) open import Data.Vec.Relation.Unary.All using (All; []; _∷_; reduce) open import Data.Vec.Relation.Unary.Any using (Any; here; there) 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; False; toWitness; fromWitness; map′) open import Relation.Nullary.Product using (_×-dec_) open import Relation.Nullary.Sum using (_⊎-dec_) open import Relation.Unary using (Decidable) --- The set of types and boolean properties of them data Type : Set where bool : Type int : Type fin : (n : ℕ) → Type real : Type bit : Type bits : (n : ℕ) → Type tuple : ∀ n → Vec Type n → Type array : Type → (n : ℕ) → Type data HasEquality : Type → Set where bool : HasEquality bool int : HasEquality int fin : ∀ {n} → HasEquality (fin n) real : HasEquality real bit : HasEquality bit bits : ∀ {n} → HasEquality (bits n) hasEquality? : Decidable HasEquality hasEquality? bool = yes bool hasEquality? int = yes int hasEquality? (fin n) = yes fin hasEquality? real = yes real hasEquality? bit = yes bit hasEquality? (bits n) = yes bits 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? bool = no (λ ()) isNumeric? int = yes int isNumeric? (fin n) = no (λ ()) isNumeric? real = yes real isNumeric? bit = no (λ ()) isNumeric? (bits 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 : Bool → Literal bit _′xs : ∀ {n} → Vec Bool n → Literal (bits n) _′a : ∀ {n t} → Literal t → Literal (array t 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 AssignsState {n Γ} : ∀ {t e} → CanAssign {n} {Γ} {t} e → Set assignsState? : ∀ {n Γ t e} → Decidable (AssignsState {n} {Γ} {t} {e}) assignsStateAny? : ∀ {n Γ m ts es} → Decidable {A = All (CanAssign ∘ proj₂) (reduce {P = Expression {n} Γ} (λ {x} e → x , e) {m} {ts} es)} (Any (AssignsState ∘ proj₂) ∘ reduce (λ {x} a → x , a)) data Statement {n} (Γ : Vec Type n) : Set data ChangesState {n Γ} : Statement {n} Γ → Set changesState? : ∀ {n Γ} → Decidable (ChangesState {n} {Γ}) data Function {n} (Γ : Vec Type n) (ret : Type) : Set data Procedure {n} (Γ : Vec Type n) : Set infix 8 -_ infixr 7 _^_ infixl 6 _*_ _and_ _>>_ -- infixl 6 _/_ infixl 5 _+_ _or_ _&&_ _||_ _∶_ infix 4 _≟_ _>_ : Expression Γ int → ℕ → Expression Γ int rnd : Expression Γ real → Expression Γ int fin : ∀ {k ms n} → (All (Fin) ms → Fin n) → Expression Γ (tuple k (map fin ms)) → Expression Γ (fin n) asInt : ∀ {m} → Expression Γ (fin m) → Expression Γ int tup : ∀ {m ts} → All (Expression Γ) ts → Expression Γ (tuple m ts) call : ∀ {t m Δ} → Function Δ t → All (Expression Γ) {m} Δ → Expression Γ t if_then_else_ : ∀ {t} → Expression Γ bool → Expression Γ t → Expression Γ t → Expression Γ t data CanAssign {n} {Γ} where state : ∀ i {i> e₁) = no λ () canAssign? (rnd e) = no λ () canAssign? (fin x e) = no λ () canAssign? (asInt e) = no λ () canAssign? (tup es) = map′ tup (λ { (tup es) → es }) (canAssignAll? es) canAssign? (call x e) = no λ () canAssign? (if e then e₁ else e₂) = no λ () canAssignAll? [] = yes [] canAssignAll? (e ∷ es) = map′ (uncurry _∷_) (λ { (e ∷ es) → e , es }) (canAssign? e ×-dec canAssignAll? es) data AssignsState where state : ∀ i {i> suc i << suc i