Compare commits

...

2 Commits

Author SHA1 Message Date
Danila Fedorin 91b5d108f6 Simplify proofs about 'loop' using concatenation lemma
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
2024-04-29 21:28:21 -07:00
Danila Fedorin c574ca9c56 Prove that graphs build by buildCfg are sufficient
That is, if we have a (semantic) trace, we can
find a corresponding path through the CFG.

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
2024-04-29 20:57:43 -07:00
4 changed files with 101 additions and 59 deletions

View File

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

View File

@ -5,15 +5,18 @@ 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 (List; _∷_; [])
open import Data.List.Relation.Unary.Any using (here)
open import Data.Fin as Fin using (suc; zero)
open import Data.List as List using (List; _∷_; [])
open import Data.List.Relation.Unary.Any using (here; there)
open import Data.List.Membership.Propositional as ListMem using ()
open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
open import Data.Product using (Σ; _,_; _×_)
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)
open import Utils using (x∈xs⇒fx∈fxs; ∈-cartesianProduct)
open import Utils using (x∈xs⇒fx∈fxs; ∈-cartesianProduct; concat-∈)
buildCfg-input : ∀ (s : Stmt) → let g = buildCfg s in Σ (Graph.Index g) (λ idx → Graph.inputs g ≡ idx ∷ [])
@ -106,50 +109,85 @@ Trace-↦ʳ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
(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} idx₁ idx₂ ρ₁ ρ₂
Trace-loop {idx₁ = idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) = Trace-single ρ₁⇒ρ₂
Trace-loop {g} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') =
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ idx₁→idx) (Trace-loop tr')
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 etr = record
{ idx₁ = EndToEndTrace.idx₁ etr
; idx₁∈inputs = EndToEndTrace.idx₁∈inputs etr
; idx₂ = EndToEndTrace.idx₂ etr
; idx₂∈outputs = EndToEndTrace.idx₂∈outputs etr
; trace = Trace-loop (EndToEndTrace.trace etr)
}
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} etr₁ etr₂ = record
{ idx₁ = EndToEndTrace.idx₁ etr₁
; idx₁∈inputs = EndToEndTrace.idx₁∈inputs etr₁
; idx₂ = EndToEndTrace.idx₂ etr₂
; idx₂∈outputs = EndToEndTrace.idx₂∈outputs etr₂
; trace =
let
o∈tr₁ = EndToEndTrace.idx₂∈outputs etr₁
i∈tr₂ = EndToEndTrace.idx₁∈inputs etr₂
oi∈es = ListMemProp.∈-++⁺ʳ (Graph.edges g) (∈-cartesianProduct o∈tr₁ i∈tr₂)
in
EndToEndTrace.trace etr₁ ++⟨ oi∈es ⟩ EndToEndTrace.trace etr₂
}
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₂
}
Trace-optional : ∀ {g : Graph} {idx₁ idx₂ : Graph.Index g} {ρ₁ ρ₂ : Env} →
Trace {g} idx₁ idx₂ ρ₁ ρ₂ → Trace {optional g} (2 Fin.↑ʳ idx₁) (2 Fin.↑ʳ idx₂) ρ₁ ρ₂
Trace-optional = {!!}
EndToEndTrace-optional : ∀ {g : Graph} {ρ₁ ρ₂ : Env} →
EndToEndTrace {g} ρ₁ ρ₂ → EndToEndTrace {optional g} ρ₁ ρ₂
EndToEndTrace-optional = {!!}
EndToEndTrace-optional-ε : ∀ {g : Graph} {ρ : Env} → EndToEndTrace {optional g} ρ ρ
EndToEndTrace-optional-ε = {!!}
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} →
@ -173,9 +211,9 @@ _++_ {g₁} {g₂} etr₁ etr₂
(Trace-↦ʳ {g₁} {g₂} (EndToEndTrace.trace etr₂))
}
Trace-singleton : ∀ {bss : List BasicStmt} {ρ₁ ρ₂ : Env} →
EndToEndTrace-singleton : ∀ {bss : List BasicStmt} {ρ₁ ρ₂ : Env} →
ρ₁ , bss ⇒ᵇˢ ρ₂ → EndToEndTrace {singleton bss} ρ₁ ρ₂
Trace-singleton ρ₁⇒ρ₂ = record
EndToEndTrace-singleton ρ₁⇒ρ₂ = record
{ idx₁ = zero
; idx₁∈inputs = here refl
; idx₂ = zero
@ -183,23 +221,26 @@ Trace-singleton ρ₁⇒ρ₂ = record
; trace = Trace-single ρ₁⇒ρ₂
}
Trace-singleton[] : ∀ (ρ : Env) → EndToEndTrace {singleton []} ρ ρ
Trace-singleton[] env = Trace-singleton []
EndToEndTrace-singleton[] : ∀ (ρ : Env) → EndToEndTrace {singleton []} ρ ρ
EndToEndTrace-singleton[] env = EndToEndTrace-singleton []
buildCfg-sufficient : ∀ {s : Stmt} {ρ₁ ρ₂ : Env} → ρ₁ , s ⇒ˢ ρ₂ →
EndToEndTrace {buildCfg s} ρ₁ ρ₂
buildCfg-sufficient (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ₂) =
Trace-singleton (ρ₁,bs⇒ρ₂ ∷ [])
EndToEndTrace-singleton (ρ₁,bs⇒ρ₂ ∷ [])
buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ₂ ρ₂,s₂⇒ρ₃) =
buildCfg-sufficient ρ₁,s₁⇒ρ₂ ++ buildCfg-sufficient ρ₂,s₂⇒ρ
buildCfg-sufficient (⇒ˢ-if-true ρ₁ ρ₂ _ _ s₁ s₂ _ _ ρ₁,s₁⇒ρ₂) =
Trace-singleton[] ρ₁ ++
EndToEndTrace-singleton[] ρ₁ ++
(EndToEndTrace-∙ˡ (buildCfg-sufficient ρ₁,s₁⇒ρ₂)) ++
Trace-singleton[] ρ₂
EndToEndTrace-singleton[] ρ₂
buildCfg-sufficient (⇒ˢ-if-false ρ₁ ρ₂ _ s₁ s₂ _ ρ₁,s₂⇒ρ₂) =
Trace-singleton[] ρ₁ ++
EndToEndTrace-singleton[] ρ₁ ++
(EndToEndTrace-∙ʳ {buildCfg s₁} (buildCfg-sufficient ρ₁,s₂⇒ρ₂)) ++
Trace-singleton[] ρ₂
buildCfg-sufficient (⇒ˢ-while-true ρ₁ ρ₂ ρ₃ _ _ s _ _ ρ₁,s⇒ρ₂ ρ₂,ws⇒ρ₃) = {!!}
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-optional-ε {loop (buildCfg s)} {ρ}
EndToEndTrace-loop⁰ {buildCfg s} {ρ}

View File

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

View File

@ -3,7 +3,7 @@ module Utils where
open import Agda.Primitive using () renaming (_⊔_ to _⊔_)
open import Data.Product as Prod using ()
open import Data.Nat using (; suc)
open import Data.List using (List; cartesianProduct; []; _∷_; _++_) renaming (map to mapˡ)
open import Data.List using (List; cartesianProduct; []; _∷_; _++_; foldr) renaming (map to mapˡ)
open import Data.List.Membership.Propositional using (_∈_)
open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
open import Data.List.Relation.Unary.All using (All; []; _∷_; map)
@ -86,3 +86,8 @@ proj₂ (_ , v) = v
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} {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')