Start working on the evaluation operation.

Proving monotonicity is the main hurdle here.

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
Danila Fedorin 2024-03-10 18:13:01 -07:00
parent 0705df708e
commit f21ebdcf46
5 changed files with 63 additions and 5 deletions

View File

@ -1,9 +1,11 @@
module Analysis.Sign where
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
open import Data.Nat using (suc)
open import Data.Product using (_×_; proj₁; _,_)
open import Data.List using (List; _∷_; []; foldr; cartesianProduct; cartesianProductWith)
open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans)
open import Data.List.Membership.Propositional as MemProp using () renaming (_∈_ to _∈ˡ_)
open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; subst)
open import Relation.Nullary using (¬_; Dec; yes; no)
open import Data.Unit using ()
@ -108,6 +110,10 @@ module _ (prog : Program) where
; _≈_ to _≈ᵛ_
; _⊔_ to _⊔ᵛ_
; ≈-dec to ≈ᵛ-dec
; _∈_ to _∈ᵛ_
; _∈k_ to _∈kᵛ_
; _updating_via_ to _updatingᵛ_via_
; locate to locateᵛ
)
open FiniteHeightLattice finiteHeightLatticeᵛ
using ()
@ -129,6 +135,8 @@ module _ (prog : Program) where
( finiteHeightLattice to finiteHeightLatticeᵐ
; FiniteMap to StateVariables
; isLattice to isLatticeᵐ
; _∈k_ to _∈kᵐ_
; locate to locateᵐ
)
open FiniteHeightLattice finiteHeightLatticeᵐ
using ()
@ -159,3 +167,36 @@ module _ (prog : Program) where
( f' to joinAll
; f'-Monotonic to joinAll-Mono
)
-- With 'join' in hand, we need to perform abstract evaluation.
vars-in-Map : (k : String) (vs : VariableSigns)
k ∈ˡ vars k ∈kᵛ vs
vars-in-Map k vs@(m , kvs≡vars) k∈vars rewrite kvs≡vars = k∈vars
states-in-Map : (s : State) (sv : StateVariables) s ∈kᵐ sv
states-in-Map s sv@(m , ksv≡states) rewrite ksv≡states = states-complete s
eval : (e : Expr) ( k k ∈ᵉ e k ∈ˡ vars) VariableSigns SignLattice
eval (e₁ + e₂) k∈e⇒k∈vars vs =
plus (eval e₁ (λ k k∈e₁ k∈e⇒k∈vars k (in⁺₁ k∈e₁)) vs)
(eval e₂ (λ k k∈e₂ k∈e⇒k∈vars k (in⁺₂ k∈e₂)) vs)
eval (e₁ - e₂) k∈e⇒k∈vars vs =
minus (eval e₁ (λ k k∈e₁ k∈e⇒k∈vars k (in⁻₁ k∈e₁)) vs)
(eval e₂ (λ k k∈e₂ k∈e⇒k∈vars k (in⁻₂ k∈e₂)) vs)
eval (` k) k∈e⇒k∈vars vs = proj₁ (locateᵛ {k} {vs} (vars-in-Map k vs (k∈e⇒k∈vars k here)))
eval (# 0) _ _ = [ 0ˢ ]ᵍ
eval (# (suc n')) _ _ = [ + ]ᵍ
updateForState : State StateVariables VariableSigns
updateForState s sv
with code s in p
... | k e =
let
(vs , s,vs∈sv) = locateᵐ {s} {sv} (states-in-Map s sv)
k∈e⇒k∈codes = λ k k∈e subst (λ stmt k ∈ᵗ stmt) (sym p) (in←₂ k∈e)
k∈e⇒k∈vars = λ k k∈e vars-complete s (k∈e⇒k∈codes k k∈e)
in
vs updatingᵛ (k []) via (λ _ eval e k∈e⇒k∈vars vs)
-- module Test = StateVariablesFiniteMap.GeneralizedUpdate states isLatticeᵐ joinAll joinAll-Mono

View File

@ -15,7 +15,7 @@ open import Relation.Nullary using (¬_)
open import Function using (_∘_)
open import Lattice
open import Utils using (Unique; Unique-map; empty; push)
open import Utils using (Unique; Unique-map; empty; push; x∈xs⇒fx∈fxs)
data Expr : Set where
_+_ : Expr Expr Expr
@ -156,6 +156,10 @@ private
, push (z≢mapsfs inds') (Unique-map suc suc-injective unids')
)
indices-complete : (n : ) (f : Fin n) f ∈ˡ (proj₁ (indices n))
indices-complete (suc n') zero = RelAny.here refl
indices-complete (suc n') (suc f') = RelAny.there (x∈xs⇒fx∈fxs suc (indices-complete n' f'))
-- For now, just represent the program and CFG as one type, without branching.
record Program : Set where
@ -179,6 +183,9 @@ record Program : Set where
states : List State
states = proj₁ (indices length)
states-complete : (s : State) s ∈ˡ states
states-complete = indices-complete length
states-Unique : Unique states
states-Unique = proj₂ (indices length)

View File

@ -69,9 +69,15 @@ module WithKeys (ks : List A) where
km₁≡ks
)
_∈_ : A × B FiniteMap Set (a ⊔ℓ b)
_∈_ k,v (m₁ , _) = k,v ∈ˡ (proj₁ m₁)
_∈k_ : A FiniteMap Set a
_∈k_ k (m₁ , _) = k ∈ˡ (keysᵐ m₁)
locate : {k : A} {fm : FiniteMap} k ∈k fm Σ B (λ v (k , v) fm)
locate {k} {fm = (m , _)} k∈kfm = locateᵐ {k} {m} k∈kfm
_updating_via_ : FiniteMap List A (A B) FiniteMap
_updating_via_ (m₁ , ksm₁≡ks) ks f =
( m₁ updatingᵐ ks via f

View File

@ -33,7 +33,6 @@ open import Lattice.Map A B _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A lB
using
( subset-impl
; locate; forget
; _∈_
; Map-functional
; Expr-Provenance
; Expr-Provenance-≡
@ -103,7 +102,7 @@ module IterProdIsomorphism where
_⊔ⁱᵖ_ {ks} = IP._⊔_ (length ks)
_∈ᵐ_ : {ks : List A} A × B FiniteMap ks Set
_∈ᵐ_ {ks} k,v fm = k,v proj₁ fm
_∈ᵐ_ {ks} = _∈_ ks
-- The left inverse is: from (to x) = x
from-to-inverseˡ : {ks : List A} (uks : Unique ks)
@ -156,7 +155,7 @@ module IterProdIsomorphism where
private
first-key-in-map : {k : A} {ks : List A} (fm : FiniteMap (k ks))
Σ B (λ v (k , v) proj₁ fm)
Σ B (λ v (k , v) fm)
first-key-in-map (((k , v) _ , _) , refl) = (v , here refl)
from-first-value : {k : A} {ks : List A} (fm : FiniteMap (k ks))

View File

@ -54,6 +54,11 @@ All-x∈xs : ∀ {a} {A : Set a} (xs : List A) → All (λ x → x ∈ xs) xs
All-x∈xs [] = []
All-x∈xs (x xs') = here refl map there (All-x∈xs xs')
x∈xs⇒fx∈fxs : {a b} {A : Set a} {B : Set b} (f : A B) {x : A} {xs : List A}
x xs (f x) mapˡ f xs
x∈xs⇒fx∈fxs f (here refl) = here refl
x∈xs⇒fx∈fxs f (there x∈xs') = there (x∈xs⇒fx∈fxs f x∈xs')
iterate : {a} {A : Set a} (n : ) (f : A A) A A
iterate 0 _ a = a
iterate (suc n) f a = f (iterate n f a)