summaryrefslogtreecommitdiff
path: root/src/Helium/Semantics/Axiomatic/Core.agda
blob: dd03a55cd873451952ee77febfcfe91cf2b56527 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
------------------------------------------------------------------------
-- Agda Helium
--
-- Base definitions for the axiomatic semantics
------------------------------------------------------------------------

{-# OPTIONS --safe --without-K #-}

open import Helium.Data.Pseudocode.Types using (RawPseudocode)

module Helium.Semantics.Axiomatic.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 Data.Bool as Bool using (Bool)
open import Data.Fin as Fin using (Fin; zero; suc)
open import Data.Fin.Patterns
open import Data.Nat as ℕ using (ℕ; suc)
import Data.Nat.Induction as Natᵢ
import Data.Nat.Properties as ℕₚ
open import Data.Product as × using (_×_; _,_; uncurry)
open import Data.Sum using (_⊎_)
open import Data.Unit using (⊤)
open import Data.Vec as Vec using (Vec; []; _∷_; _++_; lookup)
open import Data.Vec.Relation.Unary.All as All using (All; []; _∷_)
open import Function using (_on_)
open import Helium.Data.Pseudocode.Core
open import Helium.Data.Pseudocode.Properties
import Induction.WellFounded as Wf
open import Level using (_⊔_; Lift; lift)
import Relation.Binary.Construct.On as On
open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; cong₂)
open import Relation.Nullary using (Dec; does; yes; no)
open import Relation.Nullary.Decidable.Core using (True; toWitness; map′)
open import Relation.Nullary.Product using (_×-dec_)
open import Relation.Unary using (_⊆_)

private
  variable
    t t′     : Type
    m n      : ℕ
    Γ Δ Σ ts : Vec Type m

⟦_⟧ₜ  : Type → Set (b₁ ⊔ i₁ ⊔ r₁)
⟦_⟧ₜ′ : Vec Type n → Set (b₁ ⊔ i₁ ⊔ r₁)

⟦ bool ⟧ₜ       = Lift (b₁ ⊔ i₁ ⊔ r₁) Bool
⟦ int ⟧ₜ        = Lift (b₁ ⊔ r₁) ℤ
⟦ fin n ⟧ₜ      = Lift (b₁ ⊔ i₁ ⊔ r₁) (Fin n)
⟦ real ⟧ₜ       = Lift (b₁ ⊔ i₁) ℝ
⟦ bit ⟧ₜ        = Lift (i₁ ⊔ r₁) Bit
⟦ bits n ⟧ₜ     = Vec ⟦ bit ⟧ₜ n
⟦ tuple n ts ⟧ₜ = ⟦ ts ⟧ₜ′ 
⟦ array t n ⟧ₜ  = Vec ⟦ t ⟧ₜ n

⟦ [] ⟧ₜ′     = Lift (b₁ ⊔ i₁ ⊔ r₁) ⊤
⟦ t ∷ ts ⟧ₜ′ = ⟦ t ⟧ₜ × ⟦ ts ⟧ₜ′

fetch : ∀ i → ⟦ Γ ⟧ₜ′ → ⟦ lookup Γ i ⟧ₜ
fetch {Γ = _ ∷ _} 0F      (x , _)  = x
fetch {Γ = _ ∷ _} (suc i) (_ , xs) = fetch i xs

Transform : Vec Type m → Type → Set (b₁ ⊔ i₁ ⊔ r₁)
Transform ts t = ⟦ ts ⟧ₜ′ → ⟦ t ⟧ₜ

Predicate : Vec Type m → Set (b₁ ⊔ i₁ ⊔ r₁)
Predicate ts =  ⟦ ts ⟧ₜ′ → Bool