Use different graph operations to implement construction
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
@@ -5,48 +5,24 @@ open import Language.Semantics
|
||||
open import Language.Graphs
|
||||
open import Language.Traces
|
||||
|
||||
open import MonotonicState _⊆_ ⊆-trans renaming (MonotonicState to MonotonicGraphFunction)
|
||||
open import Utils using (_⊗_; _,_)
|
||||
open Relaxable {{...}}
|
||||
open import Data.Fin as Fin using (zero)
|
||||
open import Data.List using (_∷_; [])
|
||||
open import Data.Product using (Σ; _,_)
|
||||
|
||||
open import Data.Fin using (zero)
|
||||
open import Data.List using (List; _∷_; [])
|
||||
open import Data.Vec using (_∷_; [])
|
||||
open import Data.Vec.Properties using (cast-is-id; lookup-++ˡ; lookup-++ʳ)
|
||||
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym; trans; subst)
|
||||
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl)
|
||||
|
||||
relax-preserves-[]≡ : ∀ (g₁ g₂ : Graph) (g₁⊆g₂ : g₁ ⊆ g₂) (idx : Graph.Index g₁) →
|
||||
g₁ [ idx ] ≡ g₂ [ relax g₁⊆g₂ idx ]
|
||||
relax-preserves-[]≡ g₁ g₂ (Mk-⊆ n refl newNodes nsg₂≡nsg₁++newNodes _) idx
|
||||
rewrite cast-is-id refl (Graph.nodes g₂)
|
||||
with refl ← nsg₂≡nsg₁++newNodes = sym (lookup-++ˡ (Graph.nodes g₁) _ _)
|
||||
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)
|
||||
|
||||
instance
|
||||
NodeEqualsMonotonic : ∀ {bss : List BasicStmt} →
|
||||
MonotonicPredicate (λ g n → g [ n ] ≡ bss)
|
||||
NodeEqualsMonotonic = record
|
||||
{ relaxPredicate = λ g₁ g₂ idx g₁⊆g₂ g₁[idx]≡bss →
|
||||
trans (sym (relax-preserves-[]≡ g₁ g₂ g₁⊆g₂ idx)) g₁[idx]≡bss
|
||||
}
|
||||
|
||||
pushBasicBlock-works : ∀ (bss : List BasicStmt) → Always (λ g idx → g [ idx ] ≡ bss) (pushBasicBlock bss)
|
||||
pushBasicBlock-works bss = MkAlways (λ g → lookup-++ʳ (Graph.nodes g) (bss ∷ []) zero)
|
||||
|
||||
TransformsEnv : ∀ (ρ₁ ρ₂ : Env) → DependentPredicate (Graph.Index ⊗ Graph.Index)
|
||||
TransformsEnv ρ₁ ρ₂ g (idx₁ , idx₂) = Trace {g} idx₁ idx₂ ρ₁ ρ₂
|
||||
|
||||
instance
|
||||
TransformsEnvMonotonic : ∀ {ρ₁ ρ₂ : Env} → MonotonicPredicate (TransformsEnv ρ₁ ρ₂)
|
||||
TransformsEnvMonotonic = {!!}
|
||||
|
||||
buildCfg-sufficient : ∀ {ρ₁ ρ₂ : Env} {s : Stmt} → ρ₁ , s ⇒ˢ ρ₂ → Always (TransformsEnv ρ₁ ρ₂) (buildCfg s)
|
||||
buildCfg-sufficient {ρ₁} {ρ₂} {⟨ bs ⟩} (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ₂) =
|
||||
pushBasicBlock-works (bs ∷ [])
|
||||
map-reason
|
||||
(λ g idx g[idx]≡[bs] → Trace-single (subst (ρ₁ ,_⇒ᵇˢ ρ₂)
|
||||
(sym g[idx]≡[bs])
|
||||
(ρ₁,bs⇒ρ₂ ∷ [])))
|
||||
buildCfg-sufficient {ρ₁} {ρ₂} {s₁ then s₂} (⇒ˢ-then ρ₁ ρ ρ₂ s₁ s₂ ρ₁,s₁⇒ρ₂ ρ₂,s₂⇒ρ₃) =
|
||||
(buildCfg-sufficient ρ₁,s₁⇒ρ₂ ⟨⊗⟩-reason buildCfg-sufficient ρ₂,s₂⇒ρ₃)
|
||||
update-reason {!!}
|
||||
map-reason {!!}
|
||||
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)
|
||||
|
||||
Reference in New Issue
Block a user