29 lines
1.2 KiB
Agda
29 lines
1.2 KiB
Agda
module Language.Properties where
|
|
|
|
open import Language.Base
|
|
open import Language.Semantics
|
|
open import Language.Graphs
|
|
open import Language.Traces
|
|
|
|
open import Data.Fin as Fin using (zero)
|
|
open import Data.List using (_∷_; [])
|
|
open import Data.Product using (Σ; _,_)
|
|
|
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl)
|
|
|
|
buildCfg-input : ∀ (s : Stmt) → let g = buildCfg s in Σ (Graph.Index g) (λ idx → Graph.inputs g ≡ idx ∷ [])
|
|
buildCfg-input ⟨ bs₁ ⟩ = (zero , refl)
|
|
buildCfg-input (s₁ then s₂)
|
|
with (idx , p) ← buildCfg-input s₁ rewrite p = (_ , refl)
|
|
buildCfg-input (if _ then s₁ else s₂) = (zero , refl)
|
|
buildCfg-input (while _ repeat s)
|
|
with (idx , p) ← buildCfg-input s rewrite p = (_ , refl)
|
|
|
|
buildCfg-output : ∀ (s : Stmt) → let g = buildCfg s in Σ (Graph.Index g) (λ idx → Graph.outputs g ≡ idx ∷ [])
|
|
buildCfg-output ⟨ bs₁ ⟩ = (zero , refl)
|
|
buildCfg-output (s₁ then s₂)
|
|
with (idx , p) ← buildCfg-output s₂ rewrite p = (_ , refl)
|
|
buildCfg-output (if _ then s₁ else s₂) = (_ , refl)
|
|
buildCfg-output (while _ repeat s)
|
|
with (idx , p) ← buildCfg-output s rewrite p = (_ , refl)
|