Switch maps (and consequently most of the code) to using instances

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
2025-01-04 21:16:22 -08:00
parent d90b544436
commit d96eb97b69
16 changed files with 257 additions and 240 deletions

View File

@@ -2,10 +2,10 @@ open import Language hiding (_[_])
open import Lattice
module Analysis.Forward
{L : Set} {h}
(L : Set) {h}
{_≈ˡ_ : L L Set} {_⊔ˡ_ : L L L} {_⊓ˡ_ : L L L}
(isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_)
(≈ˡ-dec : IsDecidable _≈ˡ_) where
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
{{≈ˡ-dec : IsDecidable _≈ˡ_}} where
open import Data.Empty using (⊥-elim)
open import Data.String using (String)
@@ -20,8 +20,8 @@ open IsFiniteHeightLattice isFiniteHeightLatticeˡ
using () renaming (isLattice to isLatticeˡ)
module WithProg (prog : Program) where
open import Analysis.Forward.Lattices isFiniteHeightLatticeˡ ≈ˡ-dec prog
open import Analysis.Forward.Evaluation isFiniteHeightLatticeˡ ≈ˡ-dec prog
open import Analysis.Forward.Lattices L prog hiding (≈ᵛ-Decidable) -- to disambiguate instance resolution
open import Analysis.Forward.Evaluation L prog
open Program prog
private module WithStmtEvaluator {{evaluator : StmtEvaluator}} where
@@ -43,7 +43,7 @@ module WithProg (prog : Program) where
(flip (eval s)) (eval-Monoʳ s)
vs₁≼vs₂
open StateVariablesFiniteMap.GeneralizedUpdate isLatticeᵐ (λ x x) (λ a₁≼a₂ a₁≼a₂) updateVariablesForState updateVariablesForState-Monoʳ states
open StateVariablesFiniteMap.GeneralizedUpdate {{isLatticeᵐ}} (λ x x) (λ a₁≼a₂ a₁≼a₂) updateVariablesForState updateVariablesForState-Monoʳ states
using ()
renaming
( f' to updateAll

View File

@@ -2,14 +2,14 @@ open import Language hiding (_[_])
open import Lattice
module Analysis.Forward.Adapters
{L : Set} {h}
(L : Set) {h}
{_≈ˡ_ : L L Set} {_⊔ˡ_ : L L L} {_⊓ˡ_ : L L L}
(isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_)
(≈ˡ-dec : IsDecidable _≈ˡ_)
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
{{≈ˡ-dec : IsDecidable _≈ˡ_}}
(prog : Program) where
open import Analysis.Forward.Lattices isFiniteHeightLatticeˡ ≈ˡ-dec prog
open import Analysis.Forward.Evaluation isFiniteHeightLatticeˡ ≈ˡ-dec prog
open import Analysis.Forward.Lattices L prog
open import Analysis.Forward.Evaluation L prog
open import Data.Empty using (⊥-elim)
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
@@ -41,7 +41,7 @@ module ExprToStmtAdapter {{ exprEvaluator : ExprEvaluator }} where
-- for an assignment, and update the corresponding key. Use Exercise 4.26's
-- generalized update to set the single key's value.
private module _ (k : String) (e : Expr) where
open VariableValuesFiniteMap.GeneralizedUpdate isLatticeᵛ (λ x x) (λ a₁≼a₂ a₁≼a₂) (λ _ evalᵉ e) (λ _ {vs₁} {vs₂} vs₁≼vs₂ evalᵉ-Monoʳ e {vs₁} {vs₂} vs₁≼vs₂) (k [])
open VariableValuesFiniteMap.GeneralizedUpdate {{isLatticeᵛ}} (λ x x) (λ a₁≼a₂ a₁≼a₂) (λ _ evalᵉ e) (λ _ {vs₁} {vs₂} vs₁≼vs₂ evalᵉ-Monoʳ e {vs₁} {vs₂} vs₁≼vs₂) (k [])
using ()
renaming
( f' to updateVariablesFromExpression

View File

@@ -2,13 +2,13 @@ open import Language hiding (_[_])
open import Lattice
module Analysis.Forward.Evaluation
{L : Set} {h}
(L : Set) {h}
{_≈ˡ_ : L L Set} {_⊔ˡ_ : L L L} {_⊓ˡ_ : L L L}
(isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_)
(≈ˡ-dec : IsDecidable _≈ˡ_)
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
{{≈ˡ-dec : IsDecidable _≈ˡ_}}
(prog : Program) where
open import Analysis.Forward.Lattices isFiniteHeightLatticeˡ ≈ˡ-dec prog
open import Analysis.Forward.Lattices L prog
open import Data.Product using (_,_)
open IsFiniteHeightLattice isFiniteHeightLatticeˡ

View File

@@ -2,20 +2,21 @@ open import Language hiding (_[_])
open import Lattice
module Analysis.Forward.Lattices
{L : Set} {h}
(L : Set) {h}
{_≈ˡ_ : L L Set} {_⊔ˡ_ : L L L} {_⊓ˡ_ : L L L}
(isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_)
(≈ˡ-Decidable : IsDecidable _≈ˡ_)
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
{{≈ˡ-Decidable : IsDecidable _≈ˡ_}}
(prog : Program) where
open import Data.String using () renaming (_≟_ to _≟ˢ_)
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
open import Data.Product using (proj₁; proj₂; _,_)
open import Data.Unit using (tt)
open import Data.Sum using (inj₁; inj₂)
open import Data.List using (List; _∷_; []; foldr)
open import Data.List.Membership.Propositional using () renaming (_∈_ to _∈ˡ_)
open import Data.List.Relation.Unary.Any as Any using ()
open import Relation.Binary.PropositionalEquality using (_≡_; refl)
open import Utils using (Pairwise; _⇒_; __)
open import Utils using (Pairwise; _⇒_; __; it)
open IsFiniteHeightLattice isFiniteHeightLatticeˡ
using ()
@@ -29,14 +30,15 @@ open Program prog
import Lattice.FiniteMap
import Chain
≡-Decidable-String = record { R-dec = _≟ˢ_ }
≡-Decidable-State = record { R-dec = _≟_ }
instance
≡-Decidable-String = record { R-dec = _≟ˢ_ }
≡-Decidable-State = record { R-dec = _≟_ }
-- The variable -> abstract value (e.g. sign) map is a finite value-map
-- with keys strings. Use a bundle to avoid explicitly specifying operators.
-- It's helpful to export these via 'public' since consumers tend to
-- use various variable lattice operations.
module VariableValuesFiniteMap = Lattice.FiniteMap.WithKeys ≡-Decidable-String isLatticeˡ vars
module VariableValuesFiniteMap = Lattice.FiniteMap.WithKeys String L tt vars
open VariableValuesFiniteMap
using ()
renaming
@@ -45,7 +47,7 @@ open VariableValuesFiniteMap
; _≈_ to _≈ᵛ_
; _⊔_ to _⊔ᵛ_
; _≼_ to _≼ᵛ_
; ₂-Decidable⇒≈-Decidable to ˡ-Decidable⇒≈ᵛ-Decidable
; ≈-Decidable to ≈ᵛ-Decidable
; _∈_ to _∈ᵛ_
; _∈k_ to _∈kᵛ_
; _updating_via_ to _updatingᵛ_via_
@@ -63,27 +65,25 @@ open IsLattice isLatticeᵛ
; ⊔-idemp to ⊔ᵛ-idemp
)
public
open Lattice.FiniteMap.IterProdIsomorphism ≡-Decidable-String isLatticeˡ
open Lattice.FiniteMap.IterProdIsomorphism String L _
using ()
renaming
( Provenance-union to Provenance-unionᵐ
)
public
open Lattice.FiniteMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight ≡-Decidable-String isLatticeˡ vars-Unique ≈ˡ-Decidable _ fixedHeightˡ
open Lattice.FiniteMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight String L _ vars-Unique
using ()
renaming
( isFiniteHeightLattice to isFiniteHeightLatticeᵛ
; fixedHeight to fixedHeightᵛ
; ⊥-contains-bottoms to ⊥ᵛ-contains-bottoms
)
public
≈ᵛ-Decidable = ≈ˡ-Decidable⇒≈ᵛ-Decidable ≈ˡ-Decidable
joinSemilatticeᵛ = IsFiniteHeightLattice.joinSemilattice isFiniteHeightLatticeᵛ
fixedHeightᵛ = IsFiniteHeightLattice.fixedHeight isFiniteHeightLatticeᵛ
⊥ᵛ = Chain.Height.⊥ fixedHeightᵛ
-- Finally, the map we care about is (state -> (variables -> value)). Bring that in.
module StateVariablesFiniteMap = Lattice.FiniteMap.WithKeys ≡-Decidable-State isLatticeᵛ states
module StateVariablesFiniteMap = Lattice.FiniteMap.WithKeys State VariableValues tt states
open StateVariablesFiniteMap
using (_[_]; []-∈; m₁≼m₂⇒m₁[ks]≼m₂[ks]; m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂)
renaming
@@ -94,11 +94,11 @@ open StateVariablesFiniteMap
; _∈k_ to _∈kᵐ_
; locate to locateᵐ
; _≼_ to _≼ᵐ_
; ₂-Decidable⇒≈-Decidable to ᵛ-Decidable⇒≈ᵐ-Decidable
; ≈-Decidable to ≈ᵐ-Decidable
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ
)
public
open Lattice.FiniteMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight ≡-Decidable-State isLatticeᵛ states-Unique ≈ᵛ-Decidable _ fixedHeightᵛ
open Lattice.FiniteMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight State VariableValues _ states-Unique
using ()
renaming
( isFiniteHeightLattice to isFiniteHeightLatticeᵐ
@@ -111,9 +111,6 @@ open IsFiniteHeightLattice isFiniteHeightLatticeᵐ
)
public
≈ᵐ-Decidable = ≈ᵛ-Decidable⇒≈ᵐ-Decidable ≈ᵛ-Decidable
fixedHeightᵐ = IsFiniteHeightLattice.fixedHeight isFiniteHeightLatticeᵐ
-- We now have our (state -> (variables -> value)) map.
-- Define a couple of helpers to retrieve values from it. Specifically,
-- since the State type is as specific as possible, it's always possible to
@@ -147,12 +144,12 @@ joinForKey k states = foldr _⊔ᵛ_ ⊥ᵛ (states [ incoming k ])
joinForKey-Mono : (k : State) Monotonic _≼ᵐ_ _≼ᵛ_ (joinForKey k)
joinForKey-Mono k {fm₁} {fm₂} fm₁≼fm₂ =
foldr-Mono joinSemilatticeᵛ joinSemilatticeᵛ (fm₁ [ incoming k ]) (fm₂ [ incoming k ]) _⊔ᵛ_ ⊥ᵛ ⊥ᵛ
foldr-Mono it it (fm₁ [ incoming k ]) (fm₂ [ incoming k ]) _⊔ᵛ_ ⊥ᵛ ⊥ᵛ
(m₁≼m₂⇒m₁[ks]≼m₂[ks] fm₁ fm₂ (incoming k) fm₁≼fm₂)
(⊔ᵛ-idemp ⊥ᵛ) ⊔ᵛ-Monotonicʳ ⊔ᵛ-Monotonicˡ
-- The name f' comes from the formulation of Exercise 4.26.
open StateVariablesFiniteMap.GeneralizedUpdate isLatticeᵐ (λ x x) (λ a₁≼a₂ a₁≼a₂) joinForKey joinForKey-Mono states
open StateVariablesFiniteMap.GeneralizedUpdate {{isLatticeᵐ}} (λ x x) (λ a₁≼a₂ a₁≼a₂) joinForKey joinForKey-Mono states
using ()
renaming
( f' to joinAll

View File

@@ -52,7 +52,6 @@ open import Lattice.AboveBelow Sign _≡_ (record { ≈-refl = refl; ≈-sym = s
using ()
renaming
( AboveBelow to SignLattice
; ≈-Decidable to ≈ᵍ-Decidable
; to ⊥ᵍ
; to ⊤ᵍ
; [_] to [_]ᵍ
@@ -72,10 +71,7 @@ open AB.Plain 0ˢ using ()
; _⊓_ to _⊓ᵍ_
)
open IsLattice isLatticeᵍ using ()
renaming
( ≼-trans to ≼ᵍ-trans
)
open IsLattice isLatticeᵍ using () renaming (≼-trans to ≼ᵍ-trans)
plus : SignLattice SignLattice SignLattice
plus ⊥ᵍ _ = ⊥ᵍ
@@ -175,9 +171,9 @@ instance
module WithProg (prog : Program) where
open Program prog
open import Analysis.Forward.Lattices isFiniteHeightLatticeᵍ ≈ᵍ-Decidable prog
open import Analysis.Forward.Evaluation isFiniteHeightLatticeᵍ ≈ᵍ-Decidable prog
open import Analysis.Forward.Adapters isFiniteHeightLatticeᵍ ≈ᵍ-Decidable prog
open import Analysis.Forward.Lattices SignLattice prog
open import Analysis.Forward.Evaluation SignLattice prog
open import Analysis.Forward.Adapters SignLattice prog
eval : (e : Expr) VariableValues SignLattice
eval (e₁ + e₂) vs = plus (eval e₁ vs) (eval e₂ vs)
@@ -233,7 +229,7 @@ module WithProg (prog : Program) where
SignEval = record { eval = eval; eval-Monoʳ = eval-Monoʳ }
-- For debugging purposes, print out the result.
output = show (Analysis.Forward.WithProg.result isFiniteHeightLatticeᵍ ≈ᵍ-Decidable prog)
output = show (Analysis.Forward.WithProg.result SignLattice prog)
-- This should have fewer cases -- the same number as the actual 'plus' above.
-- But agda only simplifies on first argument, apparently, so we are stuck