From 687f7031131ac12bd382be831114661be785dd0d Mon Sep 17 00:00:00 2001 From: Greg Brown Date: Sun, 13 Feb 2022 13:51:43 +0000 Subject: Finish definition of denotational semantics. --- src/Helium/Data/Pseudocode/Core.agda | 84 ++++++----- src/Helium/Data/Pseudocode/Types.agda | 264 ---------------------------------- 2 files changed, 47 insertions(+), 301 deletions(-) (limited to 'src/Helium/Data/Pseudocode') diff --git a/src/Helium/Data/Pseudocode/Core.agda b/src/Helium/Data/Pseudocode/Core.agda index 7b21824..a63cc39 100644 --- a/src/Helium/Data/Pseudocode/Core.agda +++ b/src/Helium/Data/Pseudocode/Core.agda @@ -8,9 +8,10 @@ module Helium.Data.Pseudocode.Core where -open import Data.Bool using (Bool) +open import Data.Bool using (Bool; true; false) open import Data.Fin using (Fin; #_) -open import Data.Nat as ℕ using (ℕ; suc) +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.Vec using (Vec; []; _∷_; lookup; map) open import Data.Vec.Relation.Unary.All using (All; []; _∷_; reduce; all?) @@ -29,7 +30,6 @@ data Type : Set where fin : (n : ℕ) → Type real : Type bits : (n : ℕ) → Type - enum : (n : ℕ) → Type tuple : ∀ n → Vec Type n → Type array : Type → (n : ℕ) → Type @@ -42,7 +42,6 @@ data HasEquality : Type → Set where fin : ∀ {n} → HasEquality (fin n) real : HasEquality real bits : ∀ {n} → HasEquality (bits n) - enum : ∀ {n} → HasEquality (enum n) hasEquality? : Decidable HasEquality hasEquality? unit = no (λ ()) @@ -51,7 +50,6 @@ 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 (λ ()) @@ -66,7 +64,6 @@ 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 (λ ()) @@ -95,8 +92,7 @@ data Literal : Type → Set where _′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) - _′a : ∀ {n t} → Vec (Literal t) n → Literal (array t n) + _′a : ∀ {n t} → Literal t → Literal (array t n) --- Expressions, references, statements, functions and procedures @@ -111,7 +107,8 @@ module Code {o} (Σ : Vec Type o) where infix 8 -_ infixr 7 _^_ - infixl 6 _*_ _/_ _and_ + infixl 6 _*_ _and_ _>>_ + -- infixl 6 _/_ infixl 5 _+_ _or_ _&&_ _||_ _∶_ infix 4 _≟_ _>_ : Expression Γ int → ℕ → Expression Γ int rnd : Expression Γ real → Expression Γ int - get : ℕ → Expression Γ int → Expression Γ bit + -- get : ℕ → Expression Γ int → Expression Γ bit fin : ∀ {k ms n} → (All (Fin) ms → Fin n) → Expression Γ (tuple k (map fin ms)) → Expression Γ (fin n) asInt : ∀ {n} → Expression Γ (fin n) → Expression Γ int tup : ∀ {m ts} → All (Expression Γ) ts → Expression Γ (tuple m ts) - head : ∀ {m t ts} → Expression Γ (tuple (suc m) (t ∷ ts)) → Expression Γ t - tail : ∀ {m t ts} → Expression Γ (tuple (suc m) (t ∷ ts)) → Expression Γ (tuple m ts) call : ∀ {t m Δ} → Function Δ t → Expression Γ (tuple 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? (get x e) = no λ () + -- canAssign? (get x e) = no λ () canAssign? (fin x e) = no λ () canAssign? (asInt e) = no λ () canAssign? (tup es) = map′ tup (λ { (tup es) → es }) (canAssignAll? es) - canAssign? (head e) = map′ head (λ { (head e) → e }) (canAssign? e) - canAssign? (tail e) = map′ tail (λ { (tail e) → e }) (canAssign? e) canAssign? (call x e) = no λ () canAssign? (if e then e₁ else e₂) = no λ () @@ -222,13 +211,34 @@ module Code {o} (Σ : Vec Type o) where _∙end : Statement Γ unit → Procedure Γ declare : ∀ {t} → Expression Γ t → Procedure (t ∷ Γ) → Procedure Γ - infixl 6 _<<_ _>>_ + infixl 6 _<<_ infixl 5 _-_ + + slice′ : ∀ {n Γ i j t} → Expression {n} Γ (asType t (i ℕ.+ j)) → Expression Γ (fin (suc j)) → Expression Γ (asType t i) + slice′ {i = i} e₁ e₂ = slice (cast (+-comm i _) e₁) e₂ + _-_ : ∀ {n Γ t₁ t₂} {isNumeric₁ : True (isNumeric? t₁)} {isNumeric₂ : True (isNumeric? t₂)} → Expression {n} Γ t₁ → Expression Γ t₂ → Expression Γ (combineNumeric t₁ t₂ {isNumeric₁} {isNumeric₂}) _-_ {isNumeric₂ = isNumeric₂} x y = x + (-_ {isNumeric = isNumeric₂} y) - _<<_ : ∀ {n Γ} → Expression {n} Γ int → Expression Γ int → Expression Γ int + _<<_ : ∀ {n Γ} → Expression {n} Γ int → ℕ → Expression Γ int x << n = rnd (x * lit (2 ′r) ^ n) - _>>_ : ∀ {n Γ} → Expression {n} Γ int → Expression Γ int → Expression Γ int - x >> n = rnd (x * lit (2 ′r) ^ - n) + get : ∀ {n Γ} → ℕ → Expression {n} Γ int → Expression Γ bit + get i x = if x - x >> suc i << suc i >_ - _>>_ : ℤ → ℕ → ℤ - x >> zero = x - x >> suc n = (x div (1ℤ << suc n)) {1<> 1) - - uint : ∀ {n} → Bits n → ℤ - uint x = ℤ′.sum λ i → if hasBit i x then 1ℤ << toℕ i else 0ℤ - - sint : ∀ {n} → Bits n → ℤ - sint {zero} x = 0ℤ - sint {suc n} x = uint x ℤ.+ ℤ.- (if hasBit (fromℕ n) x then 1ℤ << suc n else 0ℤ) - record Pseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃ : Set (ℓsuc (b₁ ⊔ b₂ ⊔ i₁ ⊔ i₂ ⊔ i₃ ⊔ r₁ ⊔ r₂ ⊔ r₃)) where field @@ -222,191 +146,3 @@ record Pseudocode b₁ b₂ i₁ i₂ i₃ r₁ r₂ r₃ : ; _/1 = _/1 ; ⌊_⌋ = ⌊_⌋ } - - open RawPseudocode rawPseudocode public - using - ( 2ℤ; cast; getᵇ; setᵇ; sliceᵇ; updateᵇ; hasBit - ; _div_; _mod_; _<<_; uint; sint - ) - - private - -- FIXME: move almost all of these proofs into a separate module - a-a : ∀ {a} → 0ℤ ℤ.< a → 0ℤ ℤ.> ℤ.- a - 0-a {a} 0a⇒0<-a : ∀ {a} → 0ℤ ℤ.> a → 0ℤ ℤ.< ℤ.- a - 0>a⇒0<-a {a} 0>a = begin-strict - 0ℤ ≈˘⟨ ℤ.-‿inverseʳ a ⟩ - a ℤ.- a <⟨ ℤ.+-invariantʳ _ 0>a ⟩ - 0ℤ ℤ.- a ≈⟨ ℤ.+-identityˡ _ ⟩ - ℤ.- a ∎ - where open ℤ-Reasoning - - 0<-a⇒0>a : ∀ {a} → 0ℤ ℤ.< ℤ.- a → 0ℤ ℤ.> a - 0<-a⇒0>a {a} 0<-a = begin-strict - a ≈˘⟨ ℤ.+-identityʳ a ⟩ - a ℤ.+ 0ℤ <⟨ ℤ.+-invariantˡ a 0<-a ⟩ - a ℤ.- a ≈⟨ ℤ.-‿inverseʳ a ⟩ - 0ℤ ∎ - where open ℤ-Reasoning - - 0>-a⇒0 ℤ.- a → 0ℤ ℤ.< a - 0>-a⇒0-a = begin-strict - 0ℤ ≈˘⟨ ℤ.-‿inverseʳ a ⟩ - a ℤ.- a <⟨ ℤ.+-invariantˡ a 0>-a ⟩ - a ℤ.+ 0ℤ ≈⟨ ℤ.+-identityʳ a ⟩ - a ∎ - where open ℤ-Reasoning - - 0>a+0ab : ∀ {a b} → 0ℤ ℤ.> a → 0ℤ ℤ.< b → 0ℤ ℤ.> a ℤ.* b - 0>a+0ab {a} {b} 0>a 0a $ begin-strict - 0ℤ <⟨ ℤ.0a⇒0<-a 0>a) 0b⇒0>ab : ∀ {a b} → 0ℤ ℤ.< a → 0ℤ ℤ.> b → 0ℤ ℤ.> a ℤ.* b - 0b⇒0>ab {a} {b} 0b = 0<-a⇒0>a $ begin-strict - 0ℤ <⟨ ℤ.0a⇒0<-a 0>b) ⟩ - a ℤ.* ℤ.- b ≈⟨ a*-b≈-ab a b ⟩ - ℤ.- (a ℤ.* b) ∎ - where open ℤ-Reasoning - - 0>a+0>b⇒0 a → 0ℤ ℤ.> b → 0ℤ ℤ.< a ℤ.* b - 0>a+0>b⇒0a 0>b = begin-strict - 0ℤ <⟨ ℤ.0a⇒0<-a 0>a) (0>a⇒0<-a 0>b) ⟩ - ℤ.- a ℤ.* ℤ.- b ≈⟨ -a*b≈-ab a (ℤ.- b) ⟩ - ℤ.- (a ℤ.* ℤ.- b) ≈⟨ ℤ.-‿cong $ a*-b≈-ab a b ⟩ - ℤ.- (ℤ.- (a ℤ.* b)) ≈⟨ -‿idem (a ℤ.* b) ⟩ - a ℤ.* b ∎ - where open ℤ-Reasoning - - a≉0+b≉0⇒ab≉0 : ∀ {a b} → a ℤ.≉ 0ℤ → b ℤ.≉ 0ℤ → a ℤ.* b ℤ.≉ 0ℤ - a≉0+b≉0⇒ab≉0 {a} {b} a≉0 b≉0 ab≈0 with ℤ.compare a 0ℤ | ℤ.compare b 0ℤ - ... | tri< a<0 _ _ | tri< b<0 _ _ = ℤ.irrefl (ℤ.Eq.sym ab≈0) $ 0>a+0>b⇒0 _ _ b>0 = ℤ.irrefl ab≈0 $ 0>a+0ab a<0 b>0 - ... | tri≈ _ a≈0 _ | _ = a≉0 a≈0 - ... | tri> _ _ a>0 | tri< b<0 _ _ = ℤ.irrefl ab≈0 $ 0b⇒0>ab a>0 b<0 - ... | tri> _ _ a>0 | tri≈ _ b≈0 _ = b≉0 b≈0 - ... | tri> _ _ a>0 | tri> _ _ b>0 = ℤ.irrefl (ℤ.Eq.sym ab≈0) $ ℤ.00 b>0 - - ab≈0⇒a≈0⊎b≈0 : ∀ {a b} → a ℤ.* b ℤ.≈ 0ℤ → a ℤ.≈ 0ℤ ⊎ b ℤ.≈ 0ℤ - ab≈0⇒a≈0⊎b≈0 {a} {b} ab≈0 with a ℤ.≟ 0ℤ | b ℤ.≟ 0ℤ - ... | yes a≈0 | _ = inj₁ a≈0 - ... | no a≉0 | yes b≈0 = inj₂ b≈0 - ... | no a≉0 | no b≉0 = ⊥-elim (a≉0+b≉0⇒ab≉0 a≉0 b≉0 ab≈0) - - 2a< _ _ 0>1 = begin-strict - 0ℤ ≈˘⟨ ℤ.zeroʳ (ℤ.- 1ℤ) ⟩ - ℤ.- 1ℤ ℤ.* 0ℤ <⟨ aa⇒0<-a 0>1) (0>a⇒0<-a 0>1) ⟩ - ℤ.- 1ℤ ℤ.* ℤ.- 1ℤ ≈⟨ -a*b≈-ab 1ℤ (ℤ.- 1ℤ) ⟩ - ℤ.- (1ℤ ℤ.* ℤ.- 1ℤ) ≈⟨ ℤ.-‿cong $ ℤ.*-identityˡ (ℤ.- 1ℤ) ⟩ - ℤ.- (ℤ.- 1ℤ) ≈⟨ -‿idem 1ℤ ⟩ - 1ℤ ∎ - where open ℤ-Reasoning - - 0<2 : 0ℤ ℤ.< 2ℤ - 0<2 = begin-strict - 0ℤ ≈˘⟨ ℤ.+-identity² ⟩ - 0ℤ ℤ.+ 0ℤ <⟨ ℤ.+-invariantˡ 0ℤ 0<1 ⟩ - 0ℤ ℤ.+ 1ℤ <⟨ ℤ.+-invariantʳ 1ℤ 0<1 ⟩ - 2ℤ ∎ - where open ℤ-Reasoning - - 1<