agda-spa/Analysis/Sign.agda
Danila Fedorin 2e096bd64e Extract common parts of forward analyses into Forward.agda
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
2024-03-22 17:50:29 -07:00

168 lines
6.4 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module Analysis.Sign where
open import Data.Nat using (suc)
open import Data.Product using (proj₁; _,_)
open import Data.Empty using (⊥; ⊥-elim)
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 (yes; no)
open import Language
open import Lattice
open import Showable using (Showable; show)
import Analysis.Forward
data Sign : Set where
+ : Sign
- : Sign
: Sign
instance
showable : Showable Sign
showable = record
{ show = (λ
{ + "+"
; - "-"
; 0ˢ "0"
})
}
-- g for siGn; s is used for strings and i is not very descriptive.
_≟ᵍ_ : IsDecidable (_≡_ {_} {Sign})
_≟ᵍ_ + + = yes refl
_≟ᵍ_ + - = no (λ ())
_≟ᵍ_ + 0ˢ = no (λ ())
_≟ᵍ_ - + = no (λ ())
_≟ᵍ_ - - = yes refl
_≟ᵍ_ - 0ˢ = no (λ ())
_≟ᵍ_ 0ˢ + = no (λ ())
_≟ᵍ_ 0ˢ - = no (λ ())
_≟ᵍ_ 0ˢ 0ˢ = yes refl
-- embelish 'sign' with a top and bottom element.
open import Lattice.AboveBelow Sign _≡_ (record { ≈-refl = refl; ≈-sym = sym; ≈-trans = trans }) _≟ᵍ_ as AB
using ()
renaming
( AboveBelow to SignLattice
; ≈-dec to ≈ᵍ-dec
; to ⊥ᵍ
; to ⊤ᵍ
; [_] to [_]ᵍ
; _≈_ to _≈ᵍ_
; ≈-⊥-⊥ to ≈ᵍ-⊥ᵍ-⊥ᵍ
; ≈-- to ≈ᵍ-⊤ᵍ-⊤ᵍ
; ≈-lift to ≈ᵍ-lift
; ≈-refl to ≈ᵍ-refl
)
-- 'sign' has no underlying lattice structure, so use the 'plain' above-below lattice.
open AB.Plain 0ˢ using ()
renaming
( isLattice to isLatticeᵍ
; fixedHeight to fixedHeightᵍ
; _≼_ to _≼ᵍ_
; _⊔_ to _⊔ᵍ_
)
open IsLattice isLatticeᵍ using ()
renaming
( ≼-trans to ≼ᵍ-trans
)
plus : SignLattice SignLattice SignLattice
plus ⊥ᵍ _ = ⊥ᵍ
plus _ ⊥ᵍ = ⊥ᵍ
plus ⊤ᵍ _ = ⊤ᵍ
plus _ ⊤ᵍ = ⊤ᵍ
plus [ + ]ᵍ [ + ]ᵍ = [ + ]ᵍ
plus [ + ]ᵍ [ - ]ᵍ = ⊤ᵍ
plus [ + ]ᵍ [ 0ˢ ]ᵍ = [ + ]ᵍ
plus [ - ]ᵍ [ + ]ᵍ = ⊤ᵍ
plus [ - ]ᵍ [ - ]ᵍ = [ - ]ᵍ
plus [ - ]ᵍ [ 0ˢ ]ᵍ = [ - ]ᵍ
plus [ 0ˢ ]ᵍ [ + ]ᵍ = [ + ]ᵍ
plus [ 0ˢ ]ᵍ [ - ]ᵍ = [ - ]ᵍ
plus [ 0ˢ ]ᵍ [ 0ˢ ]ᵍ = [ 0ˢ ]ᵍ
-- this is incredibly tedious: 125 cases per monotonicity proof, and tactics
-- are hard. postulate for now.
postulate plus-Monoˡ : (s₂ : SignLattice) Monotonic _≼ᵍ_ _≼ᵍ_ (λ s₁ plus s₁ s₂)
postulate plus-Monoʳ : (s₁ : SignLattice) Monotonic _≼ᵍ_ _≼ᵍ_ (plus s₁)
minus : SignLattice SignLattice SignLattice
minus ⊥ᵍ _ = ⊥ᵍ
minus _ ⊥ᵍ = ⊥ᵍ
minus ⊤ᵍ _ = ⊤ᵍ
minus _ ⊤ᵍ = ⊤ᵍ
minus [ + ]ᵍ [ + ]ᵍ = ⊤ᵍ
minus [ + ]ᵍ [ - ]ᵍ = [ + ]ᵍ
minus [ + ]ᵍ [ 0ˢ ]ᵍ = [ + ]ᵍ
minus [ - ]ᵍ [ + ]ᵍ = [ - ]ᵍ
minus [ - ]ᵍ [ - ]ᵍ = ⊤ᵍ
minus [ - ]ᵍ [ 0ˢ ]ᵍ = [ - ]ᵍ
minus [ 0ˢ ]ᵍ [ + ]ᵍ = [ - ]ᵍ
minus [ 0ˢ ]ᵍ [ - ]ᵍ = [ + ]ᵍ
minus [ 0ˢ ]ᵍ [ 0ˢ ]ᵍ = [ 0ˢ ]ᵍ
postulate minus-Monoˡ : (s₂ : SignLattice) Monotonic _≼ᵍ_ _≼ᵍ_ (λ s₁ minus s₁ s₂)
postulate minus-Monoʳ : (s₁ : SignLattice) Monotonic _≼ᵍ_ _≼ᵍ_ (minus s₁)
module WithProg (prog : Program) where
open Program prog
module ForwardWithProg = Analysis.Forward.WithProg (record { isLattice = isLatticeᵍ; fixedHeight = fixedHeightᵍ }) ≈ᵍ-dec prog
open ForwardWithProg
eval : (e : Expr) VariableValues SignLattice
eval (e₁ + e₂) vs = plus (eval e₁ vs) (eval e₂ vs)
eval (e₁ - e₂) vs = minus (eval e₁ vs) (eval e₂ vs)
eval (` k) vs
with ∈k-decᵛ k (proj₁ (proj₁ vs))
... | yes k∈vs = proj₁ (locateᵛ {k} {vs} k∈vs)
... | no _ = ⊤ᵍ
eval (# 0) _ = [ 0ˢ ]ᵍ
eval (# (suc n')) _ = [ + ]ᵍ
eval-Mono : (e : Expr) Monotonic _≼ᵛ_ _≼ᵍ_ (eval e)
eval-Mono (e₁ + e₂) {vs₁} {vs₂} vs₁≼vs₂ =
let
-- TODO: can this be done with less boilerplate?
g₁vs₁ = eval e₁ vs₁
g₂vs₁ = eval e₂ vs₁
g₁vs₂ = eval e₁ vs₂
g₂vs₂ = eval e₂ vs₂
in
≼ᵍ-trans
{plus g₁vs₁ g₂vs₁} {plus g₁vs₂ g₂vs₁} {plus g₁vs₂ g₂vs₂}
(plus-Monoˡ g₂vs₁ {g₁vs₁} {g₁vs₂} (eval-Mono e₁ {vs₁} {vs₂} vs₁≼vs₂))
(plus-Monoʳ g₁vs₂ {g₂vs₁} {g₂vs₂} (eval-Mono e₂ {vs₁} {vs₂} vs₁≼vs₂))
eval-Mono (e₁ - e₂) {vs₁} {vs₂} vs₁≼vs₂ =
let
-- TODO: here too -- can this be done with less boilerplate?
g₁vs₁ = eval e₁ vs₁
g₂vs₁ = eval e₂ vs₁
g₁vs₂ = eval e₁ vs₂
g₂vs₂ = eval e₂ vs₂
in
≼ᵍ-trans
{minus g₁vs₁ g₂vs₁} {minus g₁vs₂ g₂vs₁} {minus g₁vs₂ g₂vs₂}
(minus-Monoˡ g₂vs₁ {g₁vs₁} {g₁vs₂} (eval-Mono e₁ {vs₁} {vs₂} vs₁≼vs₂))
(minus-Monoʳ g₁vs₂ {g₂vs₁} {g₂vs₂} (eval-Mono e₂ {vs₁} {vs₂} vs₁≼vs₂))
eval-Mono (` k) {vs₁@((kvs₁ , _) , _)} {vs₂@((kvs₂ , _), _)} vs₁≼vs₂
with ∈k-decᵛ k kvs₁ | ∈k-decᵛ k kvs₂
... | yes k∈kvs₁ | yes k∈kvs₂ =
let
(v₁ , k,v₁∈vs₁) = locateᵛ {k} {vs₁} k∈kvs₁
(v₂ , k,v₂∈vs₂) = locateᵛ {k} {vs₂} k∈kvs₂
in
m₁≼m₂⇒m₁[k]ᵛ≼m₂[k]ᵛ vs₁ vs₂ vs₁≼vs₂ k,v₁∈vs₁ k,v₂∈vs₂
... | yes k∈kvs₁ | no k∉kvs₂ = ⊥-elim (k∉kvs₂ (subst (λ l k ∈ˡ l) (all-equal-keysᵛ vs₁ vs₂) k∈kvs₁))
... | no k∉kvs₁ | yes k∈kvs₂ = ⊥-elim (k∉kvs₁ (subst (λ l k ∈ˡ l) (all-equal-keysᵛ vs₂ vs₁) k∈kvs₂))
... | no k∉kvs₁ | no k∉kvs₂ = IsLattice.≈-refl isLatticeᵍ
eval-Mono (# 0) _ = ≈ᵍ-refl
eval-Mono (# (suc n')) _ = ≈ᵍ-refl
open ForwardWithProg.WithEvaluator eval eval-Mono using (result)
-- For debugging purposes, print out the result.
output = show result