agda-spa/Language/Properties.agda

55 lines
2.8 KiB
Agda
Raw Normal View History

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.List.Membership.Propositional.Properties as ListMemProp using ()
open import Data.Product using (Σ; _,_)
open import Data.Vec.Properties using (lookup-++ˡ; ++-identityʳ; lookup-++ʳ)
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym)
open import Utils using (x∈xs⇒fx∈fxs)
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)
Trace-∙ˡ : (g₁ g₂ : Graph) {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env}
Trace {g₁} idx₁ idx₂ ρ₁ ρ₂
Trace {g₁ g₂} (idx₁ Fin.↑ˡ Graph.size g₂) (idx₂ Fin.↑ˡ Graph.size g₂) ρ₁ ρ₂
Trace-∙ˡ g₁ g₂ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂)
rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
Trace-single ρ₁⇒ρ₂
Trace-∙ˡ g₁ g₂ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (_↑ˡ Graph.size g₂) idx₁→idx))
(Trace-∙ˡ g₁ g₂ tr')
Trace-∙ʳ : (g₁ g₂ : Graph) {idx₁ idx₂ : Graph.Index g₂} {ρ₁ ρ₂ : Env}
Trace {g₂} idx₁ idx₂ ρ₁ ρ₂
Trace {g₁ g₂} (Graph.size g₁ Fin.↑ʳ idx₁) (Graph.size g₁ Fin.↑ʳ idx₂) ρ₁ ρ₂
Trace-∙ʳ g₁ g₂ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂)
rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
Trace-single ρ₁⇒ρ₂
Trace-∙ʳ g₁ g₂ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ʳ _ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx))
(Trace-∙ʳ g₁ g₂ tr')