Compare commits

..

No commits in common. "91b5d108f61a37dda9dfa5d2667fa3da47b4c4c0" and "bbfba34e0539e02f690076e98397fc7b0e7fd0ec" have entirely different histories.

4 changed files with 59 additions and 101 deletions

View File

@ -79,13 +79,19 @@ _↦_ g₁ g₂ = record
} }
loop : Graph Graph loop : Graph Graph
loop g = record loop g = record g
{ edges = Graph.edges g List.++
List.cartesianProduct (Graph.outputs g) (Graph.inputs g)
}
optional : Graph Graph
optional g = record
{ size = 2 Nat.+ Graph.size g { size = 2 Nat.+ Graph.size g
; nodes = [] [] Graph.nodes g ; nodes = [] [] Graph.nodes g
; edges = (2 ↑ʳᵉ Graph.edges g) List.++ ; edges = (2 ↑ʳᵉ Graph.edges g) List.++
List.map (zero ,_) (2 ↑ʳⁱ Graph.inputs g) List.++ List.map (zero ,_) (2 ↑ʳⁱ Graph.inputs g) List.++
List.map (_, suc zero) (2 ↑ʳⁱ Graph.outputs g) List.++ List.map (_, suc zero) (2 ↑ʳⁱ Graph.outputs g) List.++
((suc zero , zero) (zero , suc zero) []) ((zero , suc zero) [])
; inputs = zero [] ; inputs = zero []
; outputs = (suc zero) [] ; outputs = (suc zero) []
} }
@ -116,4 +122,4 @@ buildCfg : Stmt → Graph
buildCfg bs₁ = singleton (bs₁ []) buildCfg bs₁ = singleton (bs₁ [])
buildCfg (s₁ then s₂) = buildCfg s₁ buildCfg s₂ buildCfg (s₁ then s₂) = buildCfg s₁ buildCfg s₂
buildCfg (if _ then s₁ else s₂) = singleton [] (buildCfg s₁ buildCfg s₂) singleton [] buildCfg (if _ then s₁ else s₂) = singleton [] (buildCfg s₁ buildCfg s₂) singleton []
buildCfg (while _ repeat s) = loop (buildCfg s) buildCfg (while _ repeat s) = optional (loop (buildCfg s))

View File

@ -5,18 +5,15 @@ open import Language.Semantics
open import Language.Graphs open import Language.Graphs
open import Language.Traces open import Language.Traces
open import Data.Fin as Fin using (suc; zero) open import Data.Fin as Fin using (zero)
open import Data.List as List using (List; _∷_; []) open import Data.List using (List; _∷_; [])
open import Data.List.Relation.Unary.Any using (here; there) open import Data.List.Relation.Unary.Any using (here)
open import Data.List.Membership.Propositional as ListMem using ()
open import Data.List.Membership.Propositional.Properties as ListMemProp using () open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
open import Data.Product using (Σ; _,_; _×_) open import Data.Product using (Σ; _,_; _×_)
open import Data.Vec as Vec using (_∷_)
open import Data.Vec.Properties using (lookup-++ˡ; ++-identityʳ; lookup-++ʳ) open import Data.Vec.Properties using (lookup-++ˡ; ++-identityʳ; lookup-++ʳ)
open import Function using (_∘_)
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym)
open import Utils using (x∈xs⇒fx∈fxs; ∈-cartesianProduct; concat-∈) open import Utils using (x∈xs⇒fx∈fxs; ∈-cartesianProduct)
buildCfg-input : (s : Stmt) let g = buildCfg s in Σ (Graph.Index g) (λ idx Graph.inputs g idx []) buildCfg-input : (s : Stmt) let g = buildCfg s in Σ (Graph.Index g) (λ idx Graph.inputs g idx [])
@ -109,85 +106,50 @@ Trace-↦ʳ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
(ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx))) (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx)))
(Trace-↦ʳ {g₁} {g₂} tr') (Trace-↦ʳ {g₁} {g₂} tr')
loop-edge-groups : (g : Graph) List (List (Graph.Edge (loop g)))
loop-edge-groups g =
(2 ↑ʳᵉ Graph.edges g)
(List.map (zero ,_) (2 ↑ʳⁱ Graph.inputs g))
(List.map (_, suc zero) (2 ↑ʳⁱ Graph.outputs g))
((suc zero , zero) (zero , suc zero) [])
[]
loop-edge-help : (g : Graph) {l : List (Graph.Edge (loop g))} {e : Graph.Edge (loop g)}
e ListMem.∈ l l ListMem.∈ loop-edge-groups g
e ListMem.∈ Graph.edges (loop g)
loop-edge-help g e∈l l∈ess = concat-∈ e∈l l∈ess
Trace-loop : {g : Graph} {idx₁ idx₂ : Graph.Index g} {ρ₁ ρ₂ : Env} Trace-loop : {g : Graph} {idx₁ idx₂ : Graph.Index g} {ρ₁ ρ₂ : Env}
Trace {g} idx₁ idx₂ ρ₁ ρ₂ Trace {loop g} (2 Fin.↑ʳ idx₁) (2 Fin.↑ʳ idx₂) ρ₁ ρ₂ Trace {g} idx₁ idx₂ ρ₁ ρ₂ Trace {loop g} idx₁ idx₂ ρ₁ ρ₂
Trace-loop {g} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) Trace-loop {idx₁ = idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) = Trace-single ρ₁⇒ρ₂
rewrite sym (lookup-++ʳ (List.[] List.[] Vec.[]) (Graph.nodes g) idx₁) = Trace-loop {g} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') =
Trace-single ρ₁⇒ρ₂ Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ idx₁→idx) (Trace-loop tr')
Trace-loop {g} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
rewrite sym (lookup-++ʳ (List.[] List.[] Vec.[]) (Graph.nodes g) idx₁) =
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (2 ↑ʳ_) idx₁→idx))
(Trace-loop {g} tr')
EndToEndTrace-loop : {g : Graph} {ρ₁ ρ₂ : Env} EndToEndTrace-loop : {g : Graph} {ρ₁ ρ₂ : Env}
EndToEndTrace {g} ρ₁ ρ₂ EndToEndTrace {loop g} ρ₁ ρ₂ EndToEndTrace {g} ρ₁ ρ₂ EndToEndTrace {loop g} ρ₁ ρ₂
EndToEndTrace-loop {g} etr = EndToEndTrace-loop etr = record
let { idx₁ = EndToEndTrace.idx₁ etr
zero→idx₁' = x∈xs⇒fx∈fxs (zero ,_) ; idx₁∈inputs = EndToEndTrace.idx₁∈inputs etr
(x∈xs⇒fx∈fxs (2 Fin.↑ʳ_) ; idx₂ = EndToEndTrace.idx₂ etr
(EndToEndTrace.idx₁∈inputs etr)) ; idx₂∈outputs = EndToEndTrace.idx₂∈outputs etr
zero→idx₁ = loop-edge-help g zero→idx₁' (there (here refl)) ; trace = Trace-loop (EndToEndTrace.trace etr)
idx₂→suc' = x∈xs⇒fx∈fxs (_, suc zero) }
(x∈xs⇒fx∈fxs (2 Fin.↑ʳ_)
(EndToEndTrace.idx₂∈outputs etr))
idx₂→suc = loop-edge-help g idx₂→suc' (there (there (here refl)))
in
record
{ idx₁ = zero
; idx₁∈inputs = here refl
; idx₂ = suc zero
; idx₂∈outputs = here refl
; trace =
Trace-single [] ++⟨ zero→idx₁
Trace-loop {g} (EndToEndTrace.trace etr) ++⟨ idx₂→suc
Trace-single []
}
EndToEndTrace-loop² : {g : Graph} {ρ₁ ρ₂ ρ₃ : Env} EndToEndTrace-loop² : {g : Graph} {ρ₁ ρ₂ ρ₃ : Env}
EndToEndTrace {loop g} ρ₁ ρ₂ EndToEndTrace {loop g} ρ₁ ρ₂
EndToEndTrace {loop g} ρ₂ ρ₃ EndToEndTrace {loop g} ρ₂ ρ₃
EndToEndTrace {loop g} ρ₁ ρ₃ EndToEndTrace {loop g} ρ₁ ρ₃
EndToEndTrace-loop² {g} (MkEndToEndTrace zero (here refl) (suc zero) (here refl) tr₁) EndToEndTrace-loop² {g} etr₁ etr₂ = record
(MkEndToEndTrace zero (here refl) (suc zero) (here refl) tr₂) = { idx₁ = EndToEndTrace.idx₁ etr₁
let ; idx₁∈inputs = EndToEndTrace.idx₁∈inputs etr₁
suc→zero = loop-edge-help g (here refl) ; idx₂ = EndToEndTrace.idx₂ etr₂
(there (there (there (here refl)))) ; idx₂∈outputs = EndToEndTrace.idx₂∈outputs etr₂
in ; trace =
record let
{ idx₁ = zero o∈tr₁ = EndToEndTrace.idx₂∈outputs etr₁
; idx₁∈inputs = here refl i∈tr₂ = EndToEndTrace.idx₁∈inputs etr₂
; idx₂ = suc zero oi∈es = ListMemProp.∈-++⁺ʳ (Graph.edges g) (∈-cartesianProduct o∈tr₁ i∈tr₂)
; idx₂∈outputs = here refl in
; trace = tr₁ ++⟨ suc→zero tr₂ EndToEndTrace.trace etr₁ ++⟨ oi∈es EndToEndTrace.trace etr₂
} }
EndToEndTrace-loop⁰ : {g : Graph} {ρ : Env} Trace-optional : {g : Graph} {idx₁ idx₂ : Graph.Index g} {ρ₁ ρ₂ : Env}
EndToEndTrace {loop g} ρ ρ Trace {g} idx₁ idx₂ ρ₁ ρ₂ Trace {optional g} (2 Fin.↑ʳ idx₁) (2 Fin.↑ʳ idx₂) ρ₁ ρ₂
EndToEndTrace-loop⁰ {g} {ρ} = Trace-optional = {!!}
let
zero→suc = loop-edge-help g (there (here refl)) EndToEndTrace-optional : {g : Graph} {ρ₁ ρ₂ : Env}
(there (there (there (here refl)))) EndToEndTrace {g} ρ₁ ρ₂ EndToEndTrace {optional g} ρ₁ ρ₂
in EndToEndTrace-optional = {!!}
record
{ idx₁ = zero EndToEndTrace-optional-ε : {g : Graph} {ρ : Env} EndToEndTrace {optional g} ρ ρ
; idx₁∈inputs = here refl EndToEndTrace-optional-ε = {!!}
; idx₂ = suc zero
; idx₂∈outputs = here refl
; trace = Trace-single [] ++⟨ zero→suc Trace-single []
}
infixr 5 _++_ infixr 5 _++_
_++_ : {g₁ g₂ : Graph} {ρ₁ ρ₂ ρ₃ : Env} _++_ : {g₁ g₂ : Graph} {ρ₁ ρ₂ ρ₃ : Env}
@ -211,9 +173,9 @@ _++_ {g₁} {g₂} etr₁ etr₂
(Trace-↦ʳ {g₁} {g₂} (EndToEndTrace.trace etr₂)) (Trace-↦ʳ {g₁} {g₂} (EndToEndTrace.trace etr₂))
} }
EndToEndTrace-singleton : {bss : List BasicStmt} {ρ₁ ρ₂ : Env} Trace-singleton : {bss : List BasicStmt} {ρ₁ ρ₂ : Env}
ρ₁ , bss ⇒ᵇˢ ρ₂ EndToEndTrace {singleton bss} ρ₁ ρ₂ ρ₁ , bss ⇒ᵇˢ ρ₂ EndToEndTrace {singleton bss} ρ₁ ρ₂
EndToEndTrace-singleton ρ₁⇒ρ₂ = record Trace-singleton ρ₁⇒ρ₂ = record
{ idx₁ = zero { idx₁ = zero
; idx₁∈inputs = here refl ; idx₁∈inputs = here refl
; idx₂ = zero ; idx₂ = zero
@ -221,26 +183,23 @@ EndToEndTrace-singleton ρ₁⇒ρ₂ = record
; trace = Trace-single ρ₁⇒ρ₂ ; trace = Trace-single ρ₁⇒ρ₂
} }
EndToEndTrace-singleton[] : (ρ : Env) EndToEndTrace {singleton []} ρ ρ Trace-singleton[] : (ρ : Env) EndToEndTrace {singleton []} ρ ρ
EndToEndTrace-singleton[] env = EndToEndTrace-singleton [] Trace-singleton[] env = Trace-singleton []
buildCfg-sufficient : {s : Stmt} {ρ₁ ρ₂ : Env} ρ₁ , s ⇒ˢ ρ₂ buildCfg-sufficient : {s : Stmt} {ρ₁ ρ₂ : Env} ρ₁ , s ⇒ˢ ρ₂
EndToEndTrace {buildCfg s} ρ₁ ρ₂ EndToEndTrace {buildCfg s} ρ₁ ρ₂
buildCfg-sufficient (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ) = buildCfg-sufficient (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ) =
EndToEndTrace-singleton (ρ₁,bs⇒ρ []) Trace-singleton (ρ₁,bs⇒ρ [])
buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ ρ₂,s₂⇒ρ) = buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ ρ₂,s₂⇒ρ) =
buildCfg-sufficient ρ₁,s₁⇒ρ ++ buildCfg-sufficient ρ₂,s₂⇒ρ buildCfg-sufficient ρ₁,s₁⇒ρ ++ buildCfg-sufficient ρ₂,s₂⇒ρ
buildCfg-sufficient (⇒ˢ-if-true ρ₁ ρ₂ _ _ s₁ s₂ _ _ ρ₁,s₁⇒ρ) = buildCfg-sufficient (⇒ˢ-if-true ρ₁ ρ₂ _ _ s₁ s₂ _ _ ρ₁,s₁⇒ρ) =
EndToEndTrace-singleton[] ρ₁ ++ Trace-singleton[] ρ₁ ++
(EndToEndTrace-∙ˡ (buildCfg-sufficient ρ₁,s₁⇒ρ)) ++ (EndToEndTrace-∙ˡ (buildCfg-sufficient ρ₁,s₁⇒ρ)) ++
EndToEndTrace-singleton[] ρ₂ Trace-singleton[] ρ₂
buildCfg-sufficient (⇒ˢ-if-false ρ₁ ρ₂ _ s₁ s₂ _ ρ₁,s₂⇒ρ) = buildCfg-sufficient (⇒ˢ-if-false ρ₁ ρ₂ _ s₁ s₂ _ ρ₁,s₂⇒ρ) =
EndToEndTrace-singleton[] ρ₁ ++ Trace-singleton[] ρ₁ ++
(EndToEndTrace-∙ʳ {buildCfg s₁} (buildCfg-sufficient ρ₁,s₂⇒ρ)) ++ (EndToEndTrace-∙ʳ {buildCfg s₁} (buildCfg-sufficient ρ₁,s₂⇒ρ)) ++
EndToEndTrace-singleton[] ρ₂ Trace-singleton[] ρ₂
buildCfg-sufficient (⇒ˢ-while-true ρ₁ ρ₂ ρ₃ _ _ s _ _ ρ₁,s⇒ρ ρ₂,ws⇒ρ) = buildCfg-sufficient (⇒ˢ-while-true ρ₁ ρ₂ ρ₃ _ _ s _ _ ρ₁,s⇒ρ ρ₂,ws⇒ρ) = {!!}
EndToEndTrace-loop² {buildCfg s}
(EndToEndTrace-loop {buildCfg s} (buildCfg-sufficient ρ₁,s⇒ρ))
(buildCfg-sufficient ρ₂,ws⇒ρ)
buildCfg-sufficient (⇒ˢ-while-false ρ _ s _) = buildCfg-sufficient (⇒ˢ-while-false ρ _ s _) =
EndToEndTrace-loop⁰ {buildCfg s} {ρ} EndToEndTrace-optional-ε {loop (buildCfg s)} {ρ}

View File

@ -17,7 +17,6 @@ module _ {g : Graph} where
ρ₁ , (g [ idx₁ ]) ⇒ᵇˢ ρ₂ (idx₁ , idx₂) edges ρ₁ , (g [ idx₁ ]) ⇒ᵇˢ ρ₂ (idx₁ , idx₂) edges
Trace idx₂ idx₃ ρ₂ ρ₃ Trace idx₁ idx₃ ρ₁ ρ₃ Trace idx₂ idx₃ ρ₂ ρ₃ Trace idx₁ idx₃ ρ₁ ρ₃
infixr 5 _++⟨_⟩_
_++⟨_⟩_ : {idx₁ idx₂ idx₃ idx₄ : Index} {ρ₁ ρ₂ ρ₃ : Env} _++⟨_⟩_ : {idx₁ idx₂ idx₃ idx₄ : Index} {ρ₁ ρ₂ ρ₃ : Env}
Trace idx₁ idx₂ ρ₁ ρ₂ (idx₂ , idx₃) edges Trace idx₁ idx₂ ρ₁ ρ₂ (idx₂ , idx₃) edges
Trace idx₃ idx₄ ρ₂ ρ₃ Trace idx₁ idx₄ ρ₁ ρ₃ Trace idx₃ idx₄ ρ₂ ρ₃ Trace idx₁ idx₄ ρ₁ ρ₃
@ -25,7 +24,6 @@ module _ {g : Graph} where
_++⟨_⟩_ (Trace-edge ρ₁⇒ρ₂ idx₁→idx' tr') idx₂→idx₃ tr = Trace-edge ρ₁⇒ρ₂ idx₁→idx' (tr' ++⟨ idx₂→idx₃ tr) _++⟨_⟩_ (Trace-edge ρ₁⇒ρ₂ idx₁→idx' tr') idx₂→idx₃ tr = Trace-edge ρ₁⇒ρ₂ idx₁→idx' (tr' ++⟨ idx₂→idx₃ tr)
record EndToEndTrace (ρ₁ ρ₂ : Env) : Set where record EndToEndTrace (ρ₁ ρ₂ : Env) : Set where
constructor MkEndToEndTrace
field field
idx₁ : Index idx₁ : Index
idx₁∈inputs : idx₁ inputs idx₁∈inputs : idx₁ inputs

View File

@ -3,7 +3,7 @@ module Utils where
open import Agda.Primitive using () renaming (_⊔_ to _⊔_) open import Agda.Primitive using () renaming (_⊔_ to _⊔_)
open import Data.Product as Prod using () open import Data.Product as Prod using ()
open import Data.Nat using (; suc) open import Data.Nat using (; suc)
open import Data.List using (List; cartesianProduct; []; _∷_; _++_; foldr) renaming (map to mapˡ) open import Data.List using (List; cartesianProduct; []; _∷_; _++_) renaming (map to mapˡ)
open import Data.List.Membership.Propositional using (_∈_) open import Data.List.Membership.Propositional using (_∈_)
open import Data.List.Membership.Propositional.Properties as ListMemProp using () open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
open import Data.List.Relation.Unary.All using (All; []; _∷_; map) open import Data.List.Relation.Unary.All using (All; []; _∷_; map)
@ -86,8 +86,3 @@ proj₂ (_ , v) = v
x xs y ys (x Prod., y) cartesianProduct xs ys x xs y ys (x Prod., y) cartesianProduct xs ys
∈-cartesianProduct {x = x} (here refl) y∈ys = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (x Prod.,_) y∈ys) ∈-cartesianProduct {x = x} (here refl) y∈ys = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (x Prod.,_) y∈ys)
∈-cartesianProduct {x = x} {xs = x' _} {ys = ys} (there x∈rest) y∈ys = ListMemProp.∈-++⁺ʳ (mapˡ (x' Prod.,_) ys) (∈-cartesianProduct x∈rest y∈ys) ∈-cartesianProduct {x = x} {xs = x' _} {ys = ys} (there x∈rest) y∈ys = ListMemProp.∈-++⁺ʳ (mapˡ (x' Prod.,_) ys) (∈-cartesianProduct x∈rest y∈ys)
concat-∈ : {a} {A : Set a} {x : A} {l : List A} {ls : List (List A)}
x l l ls x foldr _++_ [] ls
concat-∈ x∈l (here refl) = ListMemProp.∈-++⁺ˡ x∈l
concat-∈ {ls = l' ls'} x∈l (there l∈ls') = ListMemProp.∈-++⁺ʳ l' (concat-∈ x∈l l∈ls')