------------------------------------------------------------------------ -- Agda Helium -- -- Definition of instructions using the Armv8-M pseudocode. ------------------------------------------------------------------------ {-# OPTIONS --safe --without-K #-} module Helium.Instructions.Base where open import Data.Bool as Bool using (true; false) open import Data.Fin as Fin using (Fin; Fin′; zero; suc; toℕ) open import Data.Nat as ℕ using (ℕ; zero; suc) import Data.Nat.Properties as ℕₚ open import Data.Sum using ([_,_]′; inj₂) open import Data.Vec as Vec using (Vec; []; _∷_) open import Data.Vec.Relation.Unary.All using (All; []; _∷_) open import Function using (_$_) open import Helium.Data.Pseudocode.Core as Core public hiding (module Code) import Helium.Instructions.Core as Instr import Relation.Binary.PropositionalEquality as P open import Relation.Nullary.Decidable.Core using (True) --- Types beat : Type beat = fin 4 elmtMask : Type elmtMask = bits 4 --- State State : Vec Type _ State = array (bits 32) 32 -- S ∷ array (bits 32) 16 -- R ∷ bits 16 -- VPR-P0 ∷ bits 8 -- VPR-mask ∷ bit -- FPSCR-QC ∷ bool -- _AdvanceVPTState ∷ beat -- _BeatId ∷ [] open Core.Code State public --- References -- Direct from State S : ∀ {n Γ} → Expression {n} Γ (array (bits 32) 32) S = state 0 R : ∀ {n Γ} → Expression {n} Γ (array (bits 32) 16) R = state 1 VPR-P0 : ∀ {n Γ} → Expression {n} Γ (bits 16) VPR-P0 = state 2 VPR-mask : ∀ {n Γ} → Expression {n} Γ (bits 8) VPR-mask = state 3 FPSCR-QC : ∀ {n Γ} → Expression {n} Γ bit FPSCR-QC = state 4 AdvanceVPTState : ∀ {n Γ} → Expression {n} Γ bool AdvanceVPTState = state 5 BeatId : ∀ {n Γ} → Expression {n} Γ beat BeatId = state 6 -- Indirect group : ∀ {n Γ t k} m → Expression {n} Γ (asType t (k ℕ.* suc m)) → Expression Γ (array (asType t k) (suc m)) group {k = k} zero x = [ cast (P.trans (ℕₚ.*-comm k 1) (ℕₚ.+-comm k 0)) x ] group {k = k} (suc m) x = group m (slice x′ (lit (Fin.fromℕ k ′f))) ∶ [ slice (cast (ℕₚ.+-comm k _) x′) (lit (zero ′f)) ] where x′ = cast (P.trans (ℕₚ.*-comm k _) (P.cong (k ℕ.+_) (ℕₚ.*-comm _ k))) x join : ∀ {n Γ t k m} → Expression {n} Γ (array (asType t k) (suc m)) → Expression Γ (asType t (k ℕ.* suc m)) join {k = k} {zero} x = cast (P.trans (ℕₚ.+-comm 0 k) (ℕₚ.*-comm 1 k)) (unbox x) join {k = k} {suc m} x = cast eq (join (slice x (lit (Fin.fromℕ 1 ′f))) ∶ unbox (slice {i = suc m} (cast (ℕₚ.+-comm 1 _) x) (lit (zero ′f)))) where eq = P.trans (P.cong (k ℕ.+_) (ℕₚ.*-comm k (suc m))) (ℕₚ.*-comm (suc (suc m)) k) index : ∀ {n Γ t m} → Expression {n} Γ (asType t (suc m)) → Expression Γ (fin (suc m)) → Expression Γ (elemType t) index {m = m} x i = unbox (slice (cast (ℕₚ.+-comm 1 m) x) i) Q : ∀ {n Γ} → Expression {n} Γ (array (array (bits 32) 4) 8) Q = group 7 S elem : ∀ {n Γ t k} m → Expression {n} Γ (asType t (k ℕ.* m)) → Expression Γ (fin k) → Expression Γ (asType t m) elem {k = zero} m x i = abort i elem {k = suc k} zero x i = cast (ℕₚ.*-comm k 0) x elem {k = suc k} (suc m) x i = index (group k (cast (ℕₚ.*-comm (suc k) (suc m)) x)) i --- Other utiliies hasBit : ∀ {n Γ m} → Expression {n} Γ (bits (suc m)) → Expression Γ (fin (suc m)) → Expression Γ bool hasBit {n} x i = index x i ≟ lit (true ′x) sliceⁱ : ∀ {n Γ m} → ℕ → Expression {n} Γ int → Expression Γ (bits m) sliceⁱ {m = zero} n i = lit ([] ′xs) sliceⁱ {m = suc m} n i = sliceⁱ (suc n) i ∶ [ get n i ] --- Functions Int : ∀ {n} → Function (bits n ∷ bool ∷ []) int Int = skip ∙return (if var 1 then uint (var 0) else sint (var 0)) -- arguments swapped, pred n SignedSatQ : ∀ n → Function (int ∷ []) (tuple 2 (bits (suc n) ∷ bool ∷ [])) SignedSatQ n = declare (lit (true ′b)) ( if max >3 (tup ((var 3) ∷ []))) then FPSCR-QC ≔ lit (true ′x) else skip))) where open Instr.VecOp₂ d e*esize>>3 : All Fin (toℕ elements ∷ []) → Fin 4 e*esize>>3 (x ∷ []) = helper size x where helper : ∀ size → Fin′ (Instr.Size.elements size) → Fin 4 helper Instr.8bit i = Fin.combine i (zero {0}) helper Instr.16bit i = Fin.combine i (zero {1}) helper Instr.32bit i = Fin.combine i zero vqdmulh : Instr.VQDMulH → Procedure [] vqdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0 * var 1 >> toℕ esize) where open Instr.VecOp₂ d using (esize) vqrdmulh : Instr.VQRDMulH → Procedure [] vqrdmulh d = vqr?dmulh d (skip ∙return lit (2 ′i) * var 0 * var 1 + lit (1 ′i) << toℕ esize-1 >> toℕ esize) where open Instr.VecOp₂ d using (esize; esize-1)