From 00a0ce9082b4cc1389815defcc806efd4a9b80f4 Mon Sep 17 00:00:00 2001 From: Greg Brown Date: Mon, 18 Apr 2022 15:05:24 +0100 Subject: Do a big refactor. - Replace the decidable predicates on expressions and statements with separate data types. - Reorganise the Hoare logic semantics to remove unnecessary definitions. - Make liberal use of modules to group related definitions together. - Unify the types for denotational and Hoare logic semantics. - Make bits an abstraction of array types. --- src/Helium/Semantics/Core.agda | 209 +++++++++++++++++++++++++++++++++++++++++ 1 file changed, 209 insertions(+) create mode 100644 src/Helium/Semantics/Core.agda (limited to 'src/Helium/Semantics/Core.agda') diff --git a/src/Helium/Semantics/Core.agda b/src/Helium/Semantics/Core.agda new file mode 100644 index 0000000..688f6f6 --- /dev/null +++ b/src/Helium/Semantics/Core.agda @@ -0,0 +1,209 @@ +------------------------------------------------------------------------ +-- Agda Helium +-- +-- Base definitions for semantics +------------------------------------------------------------------------ + +{-# OPTIONS --safe --without-K #-} + +open import Helium.Data.Pseudocode.Algebra using (RawPseudocode) + +module Helium.Semantics.Core + {b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃} + (rawPseudocode : RawPseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃) + where + +private + open module C = RawPseudocode rawPseudocode + +open import Algebra.Core using (Op₁) +open import Data.Bool as Bool using (Bool) +open import Data.Fin as Fin using (Fin; zero; suc) +open import Data.Fin.Patterns +import Data.Fin.Properties as Finₚ +open import Data.Integer as 𝕀 using () renaming (ℤ to 𝕀) +open import Data.Nat as ℕ using (ℕ; suc) +import Data.Nat.Properties as ℕₚ +open import Data.Product as × using (_×_; _,_) +open import Data.Sign using (Sign) +open import Data.Unit using (⊤) +open import Data.Vec as Vec using (Vec; []; _∷_; _++_; lookup; map; take; drop) +open import Data.Vec.Relation.Binary.Pointwise.Extensional using (Pointwise; decidable) +open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_) +open import Function +open import Helium.Data.Pseudocode.Core +open import Level hiding (suc) +open import Relation.Binary +import Relation.Binary.Construct.On as On +open import Relation.Binary.PropositionalEquality +open import Relation.Nullary using (¬_) +open import Relation.Nullary.Decidable.Core using (map′) + +private + variable + a : Level + A : Set a + t t′ t₁ t₂ : Type + m n : ℕ + Γ Δ Σ ts : Vec Type m + + ℓ = b₁ ⊔ i₁ ⊔ r₁ + ℓ₁ = ℓ ⊔ b₂ ⊔ i₂ ⊔ r₂ + ℓ₂ = i₁ ⊔ i₃ ⊔ r₁ ⊔ r₃ + + Sign⇒- : Sign → Op₁ A → Op₁ A + Sign⇒- Sign.+ f = id + Sign⇒- Sign.- f = f + +open ℕₚ.≤-Reasoning + +𝕀⇒ℤ : 𝕀 → ℤ +𝕀⇒ℤ z = Sign⇒- (𝕀.sign z) ℤ.-_ (𝕀.∣ z ∣ ℤ′.×′ 1ℤ) + +𝕀⇒ℝ : 𝕀 → ℝ +𝕀⇒ℝ z = Sign⇒- (𝕀.sign z) ℝ.-_ (𝕀.∣ z ∣ ℝ′.×′ 1ℝ) + +castVec : .(eq : m ≡ n) → Vec A m → Vec A n +castVec {m = .0} {0} eq [] = [] +castVec {m = .suc m} {suc n} eq (x ∷ xs) = x ∷ castVec (ℕₚ.suc-injective eq) xs + +⟦_⟧ₜ : Type → Set ℓ +⟦_⟧ₜ′ : Vec Type n → Set ℓ + +⟦ bool ⟧ₜ = Lift ℓ Bool +⟦ int ⟧ₜ = Lift ℓ ℤ +⟦ fin n ⟧ₜ = Lift ℓ (Fin n) +⟦ real ⟧ₜ = Lift ℓ ℝ +⟦ bit ⟧ₜ = Lift ℓ Bit +⟦ tuple ts ⟧ₜ = ⟦ ts ⟧ₜ′ +⟦ array t n ⟧ₜ = Vec ⟦ t ⟧ₜ n + +⟦ [] ⟧ₜ′ = Lift ℓ ⊤ +⟦ t ∷ [] ⟧ₜ′ = ⟦ t ⟧ₜ +⟦ t ∷ t₁ ∷ ts ⟧ₜ′ = ⟦ t ⟧ₜ × ⟦ t₁ ∷ ts ⟧ₜ′ + +fetch : ∀ (i : Fin n) Γ → ⟦ Γ ⟧ₜ′ → ⟦ lookup Γ i ⟧ₜ +fetch 0F (t ∷ []) x = x +fetch 0F (t ∷ t₁ ∷ Γ) (x , xs) = x +fetch (suc i) (t ∷ t₁ ∷ Γ) (x , xs) = fetch i (t₁ ∷ Γ) xs + +updateAt : ∀ (i : Fin n) Γ → ⟦ lookup Γ i ⟧ₜ → ⟦ Γ ⟧ₜ′ → ⟦ Γ ⟧ₜ′ +updateAt 0F (t ∷ []) v x = v +updateAt 0F (t ∷ t₁ ∷ Γ) v (x , xs) = v , xs +updateAt (suc i) (t ∷ t₁ ∷ Γ) v (x , xs) = x , updateAt i (t₁ ∷ Γ) v xs + +cons′ : ∀ (ts : Vec Type n) → ⟦ t ⟧ₜ → ⟦ tuple ts ⟧ₜ → ⟦ tuple (t ∷ ts) ⟧ₜ +cons′ [] x xs = x +cons′ (_ ∷ _) x xs = x , xs + +head′ : ∀ (ts : Vec Type n) → ⟦ tuple (t ∷ ts) ⟧ₜ → ⟦ t ⟧ₜ +head′ [] x = x +head′ (_ ∷ _) (x , xs) = x + +tail′ : ∀ (ts : Vec Type n) → ⟦ tuple (t ∷ ts) ⟧ₜ → ⟦ tuple ts ⟧ₜ +tail′ [] x = _ +tail′ (_ ∷ _) (x , xs) = xs + +_≈_ : ⦃ HasEquality t ⦄ → Rel ⟦ t ⟧ₜ ℓ₁ +_≈_ ⦃ bool ⦄ = Lift ℓ₁ ∘₂ _≡_ on lower +_≈_ ⦃ int ⦄ = Lift ℓ₁ ∘₂ ℤ._≈_ on lower +_≈_ ⦃ fin ⦄ = Lift ℓ₁ ∘₂ _≡_ on lower +_≈_ ⦃ real ⦄ = Lift ℓ₁ ∘₂ ℝ._≈_ on lower +_≈_ ⦃ bit ⦄ = Lift ℓ₁ ∘₂ 𝔹._≈_ on lower +_≈_ ⦃ array ⦄ = Pointwise _≈_ + +_<_ : ⦃ Ordered t ⦄ → Rel ⟦ t ⟧ₜ ℓ₂ +_<_ ⦃ int ⦄ = Lift ℓ₂ ∘₂ ℤ._<_ on lower +_<_ ⦃ fin ⦄ = Lift ℓ₂ ∘₂ Fin._<_ on lower +_<_ ⦃ real ⦄ = Lift ℓ₂ ∘₂ ℝ._<_ on lower + +≈-dec : ⦃ hasEq : HasEquality t ⦄ → Decidable (_≈_ ⦃ hasEq ⦄) +≈-dec ⦃ bool ⦄ = map′ lift lower ∘₂ On.decidable lower _≡_ Bool._≟_ +≈-dec ⦃ int ⦄ = map′ lift lower ∘₂ On.decidable lower ℤ._≈_ _≟ᶻ_ +≈-dec ⦃ fin ⦄ = map′ lift lower ∘₂ On.decidable lower _≡_ Fin._≟_ +≈-dec ⦃ real ⦄ = map′ lift lower ∘₂ On.decidable lower ℝ._≈_ _≟ʳ_ +≈-dec ⦃ bit ⦄ = map′ lift lower ∘₂ On.decidable lower 𝔹._≈_ _≟ᵇ₁_ +≈-dec ⦃ array ⦄ = decidable ≈-dec + +<-dec : ⦃ ordered : Ordered t ⦄ → Decidable (_<_ ⦃ ordered ⦄) +<-dec ⦃ int ⦄ = map′ lift lower ∘₂ On.decidable lower ℤ._<_ _<ᶻ?_ +<-dec ⦃ fin ⦄ = map′ lift lower ∘₂ On.decidable lower Fin._<_ Fin._