Define end-to-end path concatenation and prove one more case
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
parent
316e56f2bc
commit
4fee16413a
|
@ -13,7 +13,7 @@ open import Data.Product using (Σ; _,_; _×_)
|
||||||
open import Data.Vec.Properties using (lookup-++ˡ; ++-identityʳ; lookup-++ʳ)
|
open import Data.Vec.Properties using (lookup-++ˡ; ++-identityʳ; lookup-++ʳ)
|
||||||
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)
|
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 ∷ [])
|
||||||
|
@ -83,6 +83,27 @@ Trace-loop g₁ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) = Trace-single ρ
|
||||||
Trace-loop g₁ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') =
|
Trace-loop g₁ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') =
|
||||||
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ idx₁→idx) (Trace-loop g₁ tr')
|
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ idx₁→idx) (Trace-loop g₁ tr')
|
||||||
|
|
||||||
|
_++_ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ ρ₃ : Env} →
|
||||||
|
EndToEndTrace {g₁} ρ₁ ρ₂ → EndToEndTrace {g₂} ρ₂ ρ₃ →
|
||||||
|
EndToEndTrace {g₁ ↦ g₂} ρ₁ ρ₃
|
||||||
|
_++_ {g₁} {g₂} etr₁ etr₂
|
||||||
|
= record
|
||||||
|
{ idx₁ = EndToEndTrace.idx₁ etr₁ Fin.↑ˡ Graph.size g₂
|
||||||
|
; idx₁∈inputs = x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂) (EndToEndTrace.idx₁∈inputs etr₁)
|
||||||
|
; idx₂ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₂ etr₂
|
||||||
|
; idx₂∈outputs = x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_) (EndToEndTrace.idx₂∈outputs etr₂)
|
||||||
|
; trace =
|
||||||
|
let
|
||||||
|
o∈tr₁ = x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂) (EndToEndTrace.idx₂∈outputs etr₁)
|
||||||
|
i∈tr₂ = x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_) (EndToEndTrace.idx₁∈inputs etr₂)
|
||||||
|
oi∈es = ListMemProp.∈-++⁺ʳ (Graph.edges g₁ ↑ˡᵉ Graph.size g₂)
|
||||||
|
(ListMemProp.∈-++⁺ʳ (Graph.size g₁ ↑ʳᵉ Graph.edges g₂)
|
||||||
|
(∈-cartesianProduct o∈tr₁ i∈tr₂))
|
||||||
|
in
|
||||||
|
(Trace-↦ˡ g₁ g₂ (EndToEndTrace.trace etr₁)) ++⟨ oi∈es ⟩
|
||||||
|
(Trace-↦ʳ g₁ g₂ (EndToEndTrace.trace etr₂))
|
||||||
|
}
|
||||||
|
|
||||||
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⇒ρ₂) =
|
||||||
|
@ -93,3 +114,5 @@ buildCfg-sufficient (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ₂) =
|
||||||
; idx₂∈outputs = here refl
|
; idx₂∈outputs = here refl
|
||||||
; trace = Trace-single (ρ₁,bs⇒ρ₂ ∷ [])
|
; trace = Trace-single (ρ₁,bs⇒ρ₂ ∷ [])
|
||||||
}
|
}
|
||||||
|
buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ₂ ρ₂,s₂⇒ρ₃) =
|
||||||
|
buildCfg-sufficient ρ₁,s₁⇒ρ₂ ++ buildCfg-sufficient ρ₂,s₂⇒ρ₃
|
||||||
|
|
Loading…
Reference in New Issue
Block a user