301 lines
17 KiB
Agda
301 lines
17 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 (suc; zero)
|
||
open import Data.Fin.Properties as FinProp using (suc-injective)
|
||
open import Data.List as List using (List; _∷_; [])
|
||
open import Data.List.Properties using (filter-none)
|
||
open import Data.List.Relation.Unary.Any using (here; there)
|
||
open import Data.List.Relation.Unary.All using (All; []; _∷_; map; tabulate)
|
||
open import Data.List.Membership.Propositional as ListMem using ()
|
||
open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
|
||
open import Data.Nat as Nat using ()
|
||
open import Data.Product using (Σ; _,_; _×_; proj₂)
|
||
open import Data.Product.Properties as ProdProp using ()
|
||
open import Data.Sum using (inj₁; inj₂)
|
||
open import Data.Vec as Vec using (_∷_)
|
||
open import Data.Vec.Properties using (lookup-++ˡ; ++-identityʳ; lookup-++ʳ)
|
||
open import Function using (_∘_)
|
||
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym; cong)
|
||
open import Relation.Nullary using (¬_)
|
||
|
||
open import Utils using (x∈xs⇒fx∈fxs; ∈-cartesianProduct; concat-∈)
|
||
|
||
-- All of the below helpers are to reason about what edges aren't included
|
||
-- when combinings graphs. The currenty most important use for this is proving
|
||
-- that the entry node has no incoming edges.
|
||
--
|
||
-- -------------- Begin ugly code to make this work ----------------
|
||
↑-≢ : ∀ {n m} (f₁ : Fin.Fin n) (f₂ : Fin.Fin m) → ¬ (f₁ Fin.↑ˡ m) ≡ (n Fin.↑ʳ f₂)
|
||
↑-≢ zero f₂ ()
|
||
↑-≢ (suc f₁') f₂ f₁≡f₂ = ↑-≢ f₁' f₂ (suc-injective f₁≡f₂)
|
||
|
||
idx→f∉↑ʳᵉ : ∀ {n m} (idx : Fin.Fin (n Nat.+ m)) (f : Fin.Fin n) (es₂ : List (Fin.Fin m × Fin.Fin m)) → ¬ (idx , f Fin.↑ˡ m) ListMem.∈ (n ↑ʳᵉ es₂)
|
||
idx→f∉↑ʳᵉ idx f ((idx₁ , idx₂) ∷ es') (here idx,f≡idx₁,idx₂) = ↑-≢ f idx₂ (cong proj₂ idx,f≡idx₁,idx₂)
|
||
idx→f∉↑ʳᵉ idx f (_ ∷ es₂') (there idx→f∈es₂') = idx→f∉↑ʳᵉ idx f es₂' idx→f∈es₂'
|
||
|
||
idx→f∉pair : ∀ {n m} (idx idx' : Fin.Fin (n Nat.+ m)) (f : Fin.Fin n) (inputs₂ : List (Fin.Fin m)) → ¬ (idx , f Fin.↑ˡ m) ListMem.∈ (List.map (idx' ,_) (n ↑ʳⁱ inputs₂))
|
||
idx→f∉pair idx idx' f [] ()
|
||
idx→f∉pair idx idx' f (input ∷ inputs') (here idx,f≡idx',input) = ↑-≢ f input (cong proj₂ idx,f≡idx',input)
|
||
idx→f∉pair idx idx' f (_ ∷ inputs₂') (there idx,f∈inputs₂') = idx→f∉pair idx idx' f inputs₂' idx,f∈inputs₂'
|
||
|
||
idx→f∉cart : ∀ {n m} (idx : Fin.Fin (n Nat.+ m)) (f : Fin.Fin n) (outputs₁ : List (Fin.Fin n)) (inputs₂ : List (Fin.Fin m)) → ¬ (idx , f Fin.↑ˡ m) ListMem.∈ (List.cartesianProduct (outputs₁ ↑ˡⁱ m) (n ↑ʳⁱ inputs₂))
|
||
idx→f∉cart idx f [] inputs₂ ()
|
||
idx→f∉cart {n} {m} idx f (output ∷ outputs₁') inputs₂ idx,f∈pair++cart
|
||
with ListMemProp.∈-++⁻ (List.map (output Fin.↑ˡ m ,_) (n ↑ʳⁱ inputs₂)) idx,f∈pair++cart
|
||
... | inj₁ idx,f∈pair = idx→f∉pair idx (output Fin.↑ˡ m) f inputs₂ idx,f∈pair
|
||
... | inj₂ idx,f∈cart = idx→f∉cart idx f outputs₁' inputs₂ idx,f∈cart
|
||
|
||
help : let g₁ = singleton [] in
|
||
∀ (g₂ : Graph) (idx₁ : Graph.Index g₁) (idx : Graph.Index (g₁ ↦ g₂)) →
|
||
¬ (idx , idx₁ Fin.↑ˡ Graph.size g₂) ListMem.∈ ((Graph.size g₁ ↑ʳᵉ Graph.edges g₂) List.++
|
||
(List.cartesianProduct (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂)
|
||
(Graph.size g₁ ↑ʳⁱ Graph.inputs g₂)))
|
||
help g₂ idx₁ idx idx,idx₁∈g
|
||
with ListMemProp.∈-++⁻ (Graph.size (singleton []) ↑ʳᵉ Graph.edges g₂) idx,idx₁∈g
|
||
... | inj₁ idx,idx₁∈edges₂ = idx→f∉↑ʳᵉ idx idx₁ (Graph.edges g₂) idx,idx₁∈edges₂
|
||
... | inj₂ idx,idx₁∈cart = idx→f∉cart idx idx₁ (Graph.outputs (singleton [])) (Graph.inputs g₂) idx,idx₁∈cart
|
||
|
||
helpAll : let g₁ = singleton [] in
|
||
∀ (g₂ : Graph) (idx₁ : Graph.Index g₁) →
|
||
All (λ idx → ¬ (idx , idx₁ Fin.↑ˡ Graph.size g₂) ListMem.∈ ((Graph.size g₁ ↑ʳᵉ Graph.edges g₂) List.++
|
||
(List.cartesianProduct (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂)
|
||
(Graph.size g₁ ↑ʳⁱ Graph.inputs g₂)))) (indices (g₁ ↦ g₂))
|
||
helpAll g₂ idx₁ = tabulate (λ {idx} _ → help g₂ idx₁ idx)
|
||
|
||
module _ (g : Graph) where
|
||
wrap-preds-∅ : (idx : Graph.Index (wrap g)) →
|
||
idx ListMem.∈ Graph.inputs (wrap g) → predecessors (wrap g) idx ≡ []
|
||
wrap-preds-∅ zero (here refl) =
|
||
filter-none (λ idx' → (idx' , zero) ∈?
|
||
(Graph.edges (wrap g)))
|
||
(helpAll (g ↦ singleton []) zero)
|
||
where open import Data.List.Membership.DecPropositional (ProdProp.≡-dec (FinProp._≟_ {Graph.size (wrap g)}) (FinProp._≟_ {Graph.size (wrap g)})) using (_∈?_)
|
||
-- -------------- End ugly code to make this work ----------------
|
||
|
||
|
||
module _ (g : Graph) where
|
||
wrap-input : Σ (Graph.Index (wrap g)) (λ idx → Graph.inputs (wrap g) ≡ idx ∷ [])
|
||
wrap-input = (_ , refl)
|
||
|
||
wrap-output : Σ (Graph.Index (wrap g)) (λ idx → Graph.outputs (wrap g) ≡ idx ∷ [])
|
||
wrap-output = (_ , 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-∙ˡ 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-∙ʳ tr')
|
||
|
||
EndToEndTrace-∙ˡ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ : Env} →
|
||
EndToEndTrace {g₁} ρ₁ ρ₂ →
|
||
EndToEndTrace {g₁ ∙ g₂} ρ₁ ρ₂
|
||
EndToEndTrace-∙ˡ {g₁} {g₂} etr = record
|
||
{ idx₁ = EndToEndTrace.idx₁ etr Fin.↑ˡ Graph.size g₂
|
||
; idx₁∈inputs = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂)
|
||
(EndToEndTrace.idx₁∈inputs etr))
|
||
; idx₂ = EndToEndTrace.idx₂ etr Fin.↑ˡ Graph.size g₂
|
||
; idx₂∈outputs = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂)
|
||
(EndToEndTrace.idx₂∈outputs etr))
|
||
; trace = Trace-∙ˡ (EndToEndTrace.trace etr)
|
||
}
|
||
|
||
EndToEndTrace-∙ʳ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ : Env} →
|
||
EndToEndTrace {g₂} ρ₁ ρ₂ →
|
||
EndToEndTrace {g₁ ∙ g₂} ρ₁ ρ₂
|
||
EndToEndTrace-∙ʳ {g₁} {g₂} etr = record
|
||
{ idx₁ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₁ etr
|
||
; idx₁∈inputs = ListMemProp.∈-++⁺ʳ (Graph.inputs g₁ ↑ˡⁱ Graph.size g₂)
|
||
((x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_)
|
||
(EndToEndTrace.idx₁∈inputs etr)))
|
||
; idx₂ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₂ etr
|
||
; idx₂∈outputs = ListMemProp.∈-++⁺ʳ (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂)
|
||
((x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_)
|
||
(EndToEndTrace.idx₂∈outputs etr)))
|
||
|
||
; trace = Trace-∙ʳ (EndToEndTrace.trace etr)
|
||
}
|
||
|
||
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-↦ˡ 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.∈-++⁺ʳ (Graph.edges g₁ ↑ˡᵉ Graph.size g₂)
|
||
(ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx)))
|
||
(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 {g} idx₁ idx₂ ρ₁ ρ₂ → Trace {loop g} (2 Fin.↑ʳ idx₁) (2 Fin.↑ʳ idx₂) ρ₁ ρ₂
|
||
Trace-loop {g} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂)
|
||
rewrite sym (lookup-++ʳ (List.[] ∷ List.[] ∷ Vec.[]) (Graph.nodes g) idx₁) =
|
||
Trace-single ρ₁⇒ρ₂
|
||
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 {g} ρ₁ ρ₂ → EndToEndTrace {loop g} ρ₁ ρ₂
|
||
EndToEndTrace-loop {g} etr =
|
||
let
|
||
zero→idx₁' = x∈xs⇒fx∈fxs (zero ,_)
|
||
(x∈xs⇒fx∈fxs (2 Fin.↑ʳ_)
|
||
(EndToEndTrace.idx₁∈inputs etr))
|
||
zero→idx₁ = loop-edge-help g zero→idx₁' (there (here refl))
|
||
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} ρ₁ ρ₂ →
|
||
EndToEndTrace {loop g} ρ₂ ρ₃ →
|
||
EndToEndTrace {loop g} ρ₁ ρ₃
|
||
EndToEndTrace-loop² {g} (MkEndToEndTrace zero (here refl) (suc zero) (here refl) tr₁)
|
||
(MkEndToEndTrace zero (here refl) (suc zero) (here refl) tr₂) =
|
||
let
|
||
suc→zero = loop-edge-help g (here refl)
|
||
(there (there (there (here refl))))
|
||
in
|
||
record
|
||
{ idx₁ = zero
|
||
; idx₁∈inputs = here refl
|
||
; idx₂ = suc zero
|
||
; idx₂∈outputs = here refl
|
||
; trace = tr₁ ++⟨ suc→zero ⟩ tr₂
|
||
}
|
||
|
||
EndToEndTrace-loop⁰ : ∀ {g : Graph} {ρ : Env} →
|
||
EndToEndTrace {loop g} ρ ρ
|
||
EndToEndTrace-loop⁰ {g} {ρ} =
|
||
let
|
||
zero→suc = loop-edge-help g (there (here refl))
|
||
(there (there (there (here refl))))
|
||
in
|
||
record
|
||
{ idx₁ = zero
|
||
; idx₁∈inputs = here refl
|
||
; idx₂ = suc zero
|
||
; idx₂∈outputs = here refl
|
||
; trace = Trace-single [] ++⟨ zero→suc ⟩ Trace-single []
|
||
}
|
||
|
||
infixr 5 _++_
|
||
_++_ : ∀ {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₂))
|
||
}
|
||
|
||
EndToEndTrace-singleton : ∀ {bss : List BasicStmt} {ρ₁ ρ₂ : Env} →
|
||
ρ₁ , bss ⇒ᵇˢ ρ₂ → EndToEndTrace {singleton bss} ρ₁ ρ₂
|
||
EndToEndTrace-singleton ρ₁⇒ρ₂ = record
|
||
{ idx₁ = zero
|
||
; idx₁∈inputs = here refl
|
||
; idx₂ = zero
|
||
; idx₂∈outputs = here refl
|
||
; trace = Trace-single ρ₁⇒ρ₂
|
||
}
|
||
|
||
EndToEndTrace-singleton[] : ∀ (ρ : Env) → EndToEndTrace {singleton []} ρ ρ
|
||
EndToEndTrace-singleton[] env = EndToEndTrace-singleton []
|
||
|
||
EndToEndTrace-wrap : ∀ {g : Graph} {ρ₁ ρ₂ : Env} →
|
||
EndToEndTrace {g} ρ₁ ρ₂ → EndToEndTrace {wrap g} ρ₁ ρ₂
|
||
EndToEndTrace-wrap {g} {ρ₁} {ρ₂} etr = EndToEndTrace-singleton[] ρ₁ ++ etr ++ EndToEndTrace-singleton[] ρ₂
|
||
|
||
buildCfg-sufficient : ∀ {s : Stmt} {ρ₁ ρ₂ : Env} → ρ₁ , s ⇒ˢ ρ₂ →
|
||
EndToEndTrace {buildCfg s} ρ₁ ρ₂
|
||
buildCfg-sufficient (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ₂) =
|
||
EndToEndTrace-singleton (ρ₁,bs⇒ρ₂ ∷ [])
|
||
buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ₂ ρ₂,s₂⇒ρ₃) =
|
||
buildCfg-sufficient ρ₁,s₁⇒ρ₂ ++ buildCfg-sufficient ρ₂,s₂⇒ρ₃
|
||
buildCfg-sufficient (⇒ˢ-if-true ρ₁ ρ₂ _ _ s₁ s₂ _ _ ρ₁,s₁⇒ρ₂) =
|
||
EndToEndTrace-singleton[] ρ₁ ++
|
||
(EndToEndTrace-∙ˡ (buildCfg-sufficient ρ₁,s₁⇒ρ₂)) ++
|
||
EndToEndTrace-singleton[] ρ₂
|
||
buildCfg-sufficient (⇒ˢ-if-false ρ₁ ρ₂ _ s₁ s₂ _ ρ₁,s₂⇒ρ₂) =
|
||
EndToEndTrace-singleton[] ρ₁ ++
|
||
(EndToEndTrace-∙ʳ {buildCfg s₁} (buildCfg-sufficient ρ₁,s₂⇒ρ₂)) ++
|
||
EndToEndTrace-singleton[] ρ₂
|
||
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 _) =
|
||
EndToEndTrace-loop⁰ {buildCfg s} {ρ}
|