Compare commits
37 Commits
cf824dc744
...
main
| Author | SHA1 | Date | |
|---|---|---|---|
| f5457d8841 | |||
| d99d4a2893 | |||
| fbb98de40f | |||
| 706b593d1d | |||
| 45606679f5 | |||
| 7e099a2561 | |||
| 2808759338 | |||
| 42bb8f8792 | |||
| 05e693594d | |||
| 90e0046707 | |||
| 13eee93255 | |||
| 6243326665 | |||
| 7b2114cd0f | |||
| 36ae125e1e | |||
| 6055a79e6a | |||
| 01f7f678d3 | |||
| 14f1494fc3 | |||
| d3bac2fe60 | |||
| 5705f256fd | |||
| d59ae90cef | |||
| c1c34c69a5 | |||
| d2faada90a | |||
| 7fdbf0397d | |||
| fdef8c0a60 | |||
| c48bd0272e | |||
| d251915772 | |||
| da6e82d04b | |||
| dd101c6e9b | |||
| a611dd0f31 | |||
| 33cc0f9fe9 | |||
| 9f2790c500 | |||
| 105321971f | |||
| 236c92a5ef | |||
| ca375976b7 | |||
| c0238fea25 | |||
| 1432dfa669 | |||
| ffe9d193d9 |
223
Analysis/Constant.agda
Normal file
223
Analysis/Constant.agda
Normal file
@@ -0,0 +1,223 @@
|
|||||||
|
|
||||||
|
module Analysis.Constant where
|
||||||
|
|
||||||
|
open import Data.Integer as Int using (ℤ; +_; -[1+_]; _≟_)
|
||||||
|
open import Data.Integer.Show as IntShow using ()
|
||||||
|
open import Data.Nat as Nat using (ℕ; suc; zero)
|
||||||
|
open import Data.Product using (Σ; proj₁; proj₂; _,_)
|
||||||
|
open import Data.Sum using (inj₁; inj₂)
|
||||||
|
open import Data.Empty using (⊥; ⊥-elim)
|
||||||
|
open import Data.Unit using (⊤; tt)
|
||||||
|
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 Equivalence
|
||||||
|
open import Language
|
||||||
|
open import Lattice
|
||||||
|
open import Showable using (Showable; show)
|
||||||
|
open import Utils using (_⇒_; _∧_; _∨_)
|
||||||
|
open import Analysis.Utils using (eval-combine₂)
|
||||||
|
import Analysis.Forward
|
||||||
|
|
||||||
|
instance
|
||||||
|
showable : Showable ℤ
|
||||||
|
showable = record
|
||||||
|
{ show = IntShow.show
|
||||||
|
}
|
||||||
|
|
||||||
|
instance
|
||||||
|
≡-equiv : IsEquivalence ℤ _≡_
|
||||||
|
≡-equiv = record
|
||||||
|
{ ≈-refl = refl
|
||||||
|
; ≈-sym = sym
|
||||||
|
; ≈-trans = trans
|
||||||
|
}
|
||||||
|
|
||||||
|
≡-Decidable-ℤ : IsDecidable {_} {ℤ} _≡_
|
||||||
|
≡-Decidable-ℤ = record { R-dec = _≟_ }
|
||||||
|
|
||||||
|
-- embelish 'ℕ' with a top and bottom element.
|
||||||
|
open import Lattice.AboveBelow ℤ _ as AB
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( AboveBelow to ConstLattice
|
||||||
|
; ≈-dec to ≈ᶜ-dec
|
||||||
|
; ⊥ to ⊥ᶜ
|
||||||
|
; ⊤ to ⊤ᶜ
|
||||||
|
; [_] to [_]ᶜ
|
||||||
|
; _≈_ to _≈ᶜ_
|
||||||
|
; ≈-⊥-⊥ to ≈ᶜ-⊥ᶜ-⊥ᶜ
|
||||||
|
; ≈-⊤-⊤ to ≈ᶜ-⊤ᶜ-⊤ᶜ
|
||||||
|
; ≈-lift to ≈ᶜ-lift
|
||||||
|
; ≈-refl to ≈ᶜ-refl
|
||||||
|
)
|
||||||
|
-- 'ℕ''s structure is not finite, so just use a 'plain' above-below Lattice.
|
||||||
|
open AB.Plain (+ 0) using ()
|
||||||
|
renaming
|
||||||
|
( isLattice to isLatticeᶜ
|
||||||
|
; isFiniteHeightLattice to isFiniteHeightLatticeᵍ
|
||||||
|
; _≼_ to _≼ᶜ_
|
||||||
|
; _⊔_ to _⊔ᶜ_
|
||||||
|
; _⊓_ to _⊓ᶜ_
|
||||||
|
; ≼-trans to ≼ᶜ-trans
|
||||||
|
; ≼-refl to ≼ᶜ-refl
|
||||||
|
)
|
||||||
|
|
||||||
|
plus : ConstLattice → ConstLattice → ConstLattice
|
||||||
|
plus ⊥ᶜ _ = ⊥ᶜ
|
||||||
|
plus _ ⊥ᶜ = ⊥ᶜ
|
||||||
|
plus ⊤ᶜ _ = ⊤ᶜ
|
||||||
|
plus _ ⊤ᶜ = ⊤ᶜ
|
||||||
|
plus [ z₁ ]ᶜ [ z₂ ]ᶜ = [ z₁ Int.+ z₂ ]ᶜ
|
||||||
|
|
||||||
|
-- this is incredibly tedious: 125 cases per monotonicity proof, and tactics
|
||||||
|
-- are hard. postulate for now.
|
||||||
|
postulate plus-Monoˡ : ∀ (s₂ : ConstLattice) → Monotonic _≼ᶜ_ _≼ᶜ_ (λ s₁ → plus s₁ s₂)
|
||||||
|
postulate plus-Monoʳ : ∀ (s₁ : ConstLattice) → Monotonic _≼ᶜ_ _≼ᶜ_ (plus s₁)
|
||||||
|
|
||||||
|
plus-Mono₂ : Monotonic₂ _≼ᶜ_ _≼ᶜ_ _≼ᶜ_ plus
|
||||||
|
plus-Mono₂ = (plus-Monoˡ , plus-Monoʳ)
|
||||||
|
|
||||||
|
minus : ConstLattice → ConstLattice → ConstLattice
|
||||||
|
minus ⊥ᶜ _ = ⊥ᶜ
|
||||||
|
minus _ ⊥ᶜ = ⊥ᶜ
|
||||||
|
minus ⊤ᶜ _ = ⊤ᶜ
|
||||||
|
minus _ ⊤ᶜ = ⊤ᶜ
|
||||||
|
minus [ z₁ ]ᶜ [ z₂ ]ᶜ = [ z₁ Int.- z₂ ]ᶜ
|
||||||
|
|
||||||
|
postulate minus-Monoˡ : ∀ (s₂ : ConstLattice) → Monotonic _≼ᶜ_ _≼ᶜ_ (λ s₁ → minus s₁ s₂)
|
||||||
|
postulate minus-Monoʳ : ∀ (s₁ : ConstLattice) → Monotonic _≼ᶜ_ _≼ᶜ_ (minus s₁)
|
||||||
|
|
||||||
|
minus-Mono₂ : Monotonic₂ _≼ᶜ_ _≼ᶜ_ _≼ᶜ_ minus
|
||||||
|
minus-Mono₂ = (minus-Monoˡ , minus-Monoʳ)
|
||||||
|
|
||||||
|
⟦_⟧ᶜ : ConstLattice → Value → Set
|
||||||
|
⟦_⟧ᶜ ⊥ᶜ _ = ⊥
|
||||||
|
⟦_⟧ᶜ ⊤ᶜ _ = ⊤
|
||||||
|
⟦_⟧ᶜ [ z ]ᶜ v = v ≡ ↑ᶻ z
|
||||||
|
|
||||||
|
⟦⟧ᶜ-respects-≈ᶜ : ∀ {s₁ s₂ : ConstLattice} → s₁ ≈ᶜ s₂ → ⟦ s₁ ⟧ᶜ ⇒ ⟦ s₂ ⟧ᶜ
|
||||||
|
⟦⟧ᶜ-respects-≈ᶜ ≈ᶜ-⊥ᶜ-⊥ᶜ v bot = bot
|
||||||
|
⟦⟧ᶜ-respects-≈ᶜ ≈ᶜ-⊤ᶜ-⊤ᶜ v top = top
|
||||||
|
⟦⟧ᶜ-respects-≈ᶜ (≈ᶜ-lift { z } { z } refl) v proof = proof
|
||||||
|
|
||||||
|
⟦⟧ᶜ-⊔ᶜ-∨ : ∀ {s₁ s₂ : ConstLattice} → (⟦ s₁ ⟧ᶜ ∨ ⟦ s₂ ⟧ᶜ) ⇒ ⟦ s₁ ⊔ᶜ s₂ ⟧ᶜ
|
||||||
|
⟦⟧ᶜ-⊔ᶜ-∨ {⊥ᶜ} x (inj₂ px₂) = px₂
|
||||||
|
⟦⟧ᶜ-⊔ᶜ-∨ {⊤ᶜ} x _ = tt
|
||||||
|
⟦⟧ᶜ-⊔ᶜ-∨ {[ s₁ ]ᶜ} {[ s₂ ]ᶜ} x px
|
||||||
|
with s₁ ≟ s₂
|
||||||
|
... | no _ = tt
|
||||||
|
... | yes refl
|
||||||
|
with px
|
||||||
|
... | inj₁ px₁ = px₁
|
||||||
|
... | inj₂ px₂ = px₂
|
||||||
|
⟦⟧ᶜ-⊔ᶜ-∨ {[ s₁ ]ᶜ} {⊥ᶜ} x (inj₁ px₁) = px₁
|
||||||
|
⟦⟧ᶜ-⊔ᶜ-∨ {[ s₁ ]ᶜ} {⊤ᶜ} x _ = tt
|
||||||
|
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ : ∀ {z₁ z₂ : ℤ} → ¬ z₁ ≡ z₂ → ∀ {v} → ¬ ((⟦ [ z₁ ]ᶜ ⟧ᶜ ∧ ⟦ [ z₂ ]ᶜ ⟧ᶜ) v)
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { z₁ } { z₂ } z₁≢z₂ {v} (v≡z₁ , v≡z₂)
|
||||||
|
with refl ← trans (sym v≡z₁) v≡z₂ = z₁≢z₂ refl
|
||||||
|
|
||||||
|
⟦⟧ᶜ-⊓ᶜ-∧ : ∀ {s₁ s₂ : ConstLattice} → (⟦ s₁ ⟧ᶜ ∧ ⟦ s₂ ⟧ᶜ) ⇒ ⟦ s₁ ⊓ᶜ s₂ ⟧ᶜ
|
||||||
|
⟦⟧ᶜ-⊓ᶜ-∧ {⊥ᶜ} x (bot , _) = bot
|
||||||
|
⟦⟧ᶜ-⊓ᶜ-∧ {⊤ᶜ} x (_ , px₂) = px₂
|
||||||
|
⟦⟧ᶜ-⊓ᶜ-∧ {[ s₁ ]ᶜ} {[ s₂ ]ᶜ} x (px₁ , px₂)
|
||||||
|
with s₁ ≟ s₂
|
||||||
|
... | no s₁≢s₂ = s₁≢s₂⇒¬s₁∧s₂ s₁≢s₂ (px₁ , px₂)
|
||||||
|
... | yes refl = px₁
|
||||||
|
⟦⟧ᶜ-⊓ᶜ-∧ {[ g₁ ]ᶜ} {⊥ᶜ} x (_ , bot) = bot
|
||||||
|
⟦⟧ᶜ-⊓ᶜ-∧ {[ g₁ ]ᶜ} {⊤ᶜ} x (px₁ , _) = px₁
|
||||||
|
|
||||||
|
instance
|
||||||
|
latticeInterpretationᶜ : LatticeInterpretation isLatticeᶜ
|
||||||
|
latticeInterpretationᶜ = record
|
||||||
|
{ ⟦_⟧ = ⟦_⟧ᶜ
|
||||||
|
; ⟦⟧-respects-≈ = ⟦⟧ᶜ-respects-≈ᶜ
|
||||||
|
; ⟦⟧-⊔-∨ = ⟦⟧ᶜ-⊔ᶜ-∨
|
||||||
|
; ⟦⟧-⊓-∧ = ⟦⟧ᶜ-⊓ᶜ-∧
|
||||||
|
}
|
||||||
|
|
||||||
|
module WithProg (prog : Program) where
|
||||||
|
open Program prog
|
||||||
|
|
||||||
|
open import Analysis.Forward.Lattices ConstLattice prog
|
||||||
|
open import Analysis.Forward.Evaluation ConstLattice prog
|
||||||
|
open import Analysis.Forward.Adapters ConstLattice prog
|
||||||
|
|
||||||
|
eval : ∀ (e : Expr) → VariableValues → ConstLattice
|
||||||
|
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 (# n) _ = [ + n ]ᶜ
|
||||||
|
|
||||||
|
eval-Monoʳ : ∀ (e : Expr) → Monotonic _≼ᵛ_ _≼ᶜ_ (eval e)
|
||||||
|
eval-Monoʳ (e₁ + e₂) {vs₁} {vs₂} vs₁≼vs₂ =
|
||||||
|
eval-combine₂ (λ {x} {y} {z} → ≼ᶜ-trans {x} {y} {z})
|
||||||
|
plus plus-Mono₂ {o₁ = eval e₁ vs₁}
|
||||||
|
(eval-Monoʳ e₁ vs₁≼vs₂) (eval-Monoʳ e₂ vs₁≼vs₂)
|
||||||
|
eval-Monoʳ (e₁ - e₂) {vs₁} {vs₂} vs₁≼vs₂ =
|
||||||
|
eval-combine₂ (λ {x} {y} {z} → ≼ᶜ-trans {x} {y} {z})
|
||||||
|
minus minus-Mono₂ {o₁ = eval e₁ vs₁}
|
||||||
|
(eval-Monoʳ e₁ vs₁≼vs₂) (eval-Monoʳ e₂ 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ʳ (# n) _ = ≼ᶜ-refl ([ + n ]ᶜ)
|
||||||
|
|
||||||
|
instance
|
||||||
|
ConstEval : ExprEvaluator
|
||||||
|
ConstEval = record { eval = eval; eval-Monoʳ = eval-Monoʳ }
|
||||||
|
|
||||||
|
-- For debugging purposes, print out the result.
|
||||||
|
output = show (Analysis.Forward.WithProg.result ConstLattice 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
|
||||||
|
-- listing them all.
|
||||||
|
plus-valid : ∀ {g₁ g₂} {z₁ z₂} → ⟦ g₁ ⟧ᶜ (↑ᶻ z₁) → ⟦ g₂ ⟧ᶜ (↑ᶻ z₂) → ⟦ plus g₁ g₂ ⟧ᶜ (↑ᶻ (z₁ Int.+ z₂))
|
||||||
|
plus-valid {⊥ᶜ} {_} ⊥ _ = ⊥
|
||||||
|
plus-valid {[ z ]ᶜ} {⊥ᶜ} _ ⊥ = ⊥
|
||||||
|
plus-valid {⊤ᶜ} {⊥ᶜ} _ ⊥ = ⊥
|
||||||
|
plus-valid {⊤ᶜ} {[ z ]ᶜ} _ _ = tt
|
||||||
|
plus-valid {⊤ᶜ} {⊤ᶜ} _ _ = tt
|
||||||
|
plus-valid {[ z₁ ]ᶜ} {[ z₂ ]ᶜ} refl refl = refl
|
||||||
|
plus-valid {[ z ]ᶜ} {⊤ᶜ} _ _ = tt
|
||||||
|
--
|
||||||
|
-- Same for this one. It should be easier, but Agda won't simplify.
|
||||||
|
minus-valid : ∀ {g₁ g₂} {z₁ z₂} → ⟦ g₁ ⟧ᶜ (↑ᶻ z₁) → ⟦ g₂ ⟧ᶜ (↑ᶻ z₂) → ⟦ minus g₁ g₂ ⟧ᶜ (↑ᶻ (z₁ Int.- z₂))
|
||||||
|
minus-valid {⊥ᶜ} {_} ⊥ _ = ⊥
|
||||||
|
minus-valid {[ z ]ᶜ} {⊥ᶜ} _ ⊥ = ⊥
|
||||||
|
minus-valid {⊤ᶜ} {⊥ᶜ} _ ⊥ = ⊥
|
||||||
|
minus-valid {⊤ᶜ} {[ z ]ᶜ} _ _ = tt
|
||||||
|
minus-valid {⊤ᶜ} {⊤ᶜ} _ _ = tt
|
||||||
|
minus-valid {[ z₁ ]ᶜ} {[ z₂ ]ᶜ} refl refl = refl
|
||||||
|
minus-valid {[ z ]ᶜ} {⊤ᶜ} _ _ = tt
|
||||||
|
|
||||||
|
eval-valid : IsValidExprEvaluator
|
||||||
|
eval-valid (⇒ᵉ-+ ρ e₁ e₂ z₁ z₂ ρ,e₁⇒z₁ ρ,e₂⇒z₂) ⟦vs⟧ρ =
|
||||||
|
plus-valid (eval-valid ρ,e₁⇒z₁ ⟦vs⟧ρ) (eval-valid ρ,e₂⇒z₂ ⟦vs⟧ρ)
|
||||||
|
eval-valid (⇒ᵉ-- ρ e₁ e₂ z₁ z₂ ρ,e₁⇒z₁ ρ,e₂⇒z₂) ⟦vs⟧ρ =
|
||||||
|
minus-valid (eval-valid ρ,e₁⇒z₁ ⟦vs⟧ρ) (eval-valid ρ,e₂⇒z₂ ⟦vs⟧ρ)
|
||||||
|
eval-valid {vs} (⇒ᵉ-Var ρ x v x,v∈ρ) ⟦vs⟧ρ
|
||||||
|
with ∈k-decᵛ x (proj₁ (proj₁ vs))
|
||||||
|
... | yes x∈kvs = ⟦vs⟧ρ (proj₂ (locateᵛ {x} {vs} x∈kvs)) x,v∈ρ
|
||||||
|
... | no x∉kvs = tt
|
||||||
|
eval-valid (⇒ᵉ-ℕ ρ n) _ = refl
|
||||||
|
|
||||||
|
instance
|
||||||
|
ConstEvalValid : ValidExprEvaluator ConstEval latticeInterpretationᶜ
|
||||||
|
ConstEvalValid = record { valid = eval-valid }
|
||||||
|
|
||||||
|
analyze-correct = Analysis.Forward.WithProg.analyze-correct ConstLattice prog tt
|
||||||
@@ -8,6 +8,7 @@ module Analysis.Forward
|
|||||||
{{≈ˡ-dec : IsDecidable _≈ˡ_}} where
|
{{≈ˡ-dec : IsDecidable _≈ˡ_}} where
|
||||||
|
|
||||||
open import Data.Empty using (⊥-elim)
|
open import Data.Empty using (⊥-elim)
|
||||||
|
open import Data.Unit using (⊤)
|
||||||
open import Data.String using (String)
|
open import Data.String using (String)
|
||||||
open import Data.Product using (_,_)
|
open import Data.Product using (_,_)
|
||||||
open import Data.List using (_∷_; []; foldr; foldl)
|
open import Data.List using (_∷_; []; foldr; foldl)
|
||||||
@@ -20,7 +21,7 @@ open IsFiniteHeightLattice isFiniteHeightLatticeˡ
|
|||||||
using () renaming (isLattice to isLatticeˡ)
|
using () renaming (isLattice to isLatticeˡ)
|
||||||
|
|
||||||
module WithProg (prog : Program) where
|
module WithProg (prog : Program) where
|
||||||
open import Analysis.Forward.Lattices L prog hiding (≈ᵛ-Decidable) -- to disambiguate instance resolution
|
open import Analysis.Forward.Lattices L prog hiding (≈ᵛ-Decidable; ≈ᵐ-Decidable) -- to disambiguate instance resolution
|
||||||
open import Analysis.Forward.Evaluation L prog
|
open import Analysis.Forward.Evaluation L prog
|
||||||
open Program prog
|
open Program prog
|
||||||
|
|
||||||
@@ -65,7 +66,7 @@ module WithProg (prog : Program) where
|
|||||||
(joinAll-Mono {sv₁} {sv₂} sv₁≼sv₂)
|
(joinAll-Mono {sv₁} {sv₂} sv₁≼sv₂)
|
||||||
|
|
||||||
-- The fixed point of the 'analyze' function is our final goal.
|
-- The fixed point of the 'analyze' function is our final goal.
|
||||||
open import Fixedpoint ≈ᵐ-Decidable isFiniteHeightLatticeᵐ analyze (λ {m₁} {m₂} m₁≼m₂ → analyze-Mono {m₁} {m₂} m₁≼m₂)
|
open import Fixedpoint analyze (λ {m₁} {m₂} m₁≼m₂ → analyze-Mono {m₁} {m₂} m₁≼m₂)
|
||||||
using ()
|
using ()
|
||||||
renaming (aᶠ to result; aᶠ≈faᶠ to result≈analyze-result)
|
renaming (aᶠ to result; aᶠ≈faᶠ to result≈analyze-result)
|
||||||
public
|
public
|
||||||
@@ -77,7 +78,7 @@ module WithProg (prog : Program) where
|
|||||||
updateAll-k∈ks-≡ {l = sv} (states-complete s) s,vs∈usv
|
updateAll-k∈ks-≡ {l = sv} (states-complete s) s,vs∈usv
|
||||||
|
|
||||||
module WithValidInterpretation {{latticeInterpretationˡ : LatticeInterpretation isLatticeˡ}}
|
module WithValidInterpretation {{latticeInterpretationˡ : LatticeInterpretation isLatticeˡ}}
|
||||||
{{validEvaluator : ValidStmtEvaluator evaluator latticeInterpretationˡ}} where
|
{{validEvaluator : ValidStmtEvaluator evaluator latticeInterpretationˡ}} (dummy : ⊤) where
|
||||||
open ValidStmtEvaluator validEvaluator
|
open ValidStmtEvaluator validEvaluator
|
||||||
|
|
||||||
eval-fold-valid : ∀ {s bss vs ρ₁ ρ₂} → ρ₁ , bss ⇒ᵇˢ ρ₂ → ⟦ vs ⟧ᵛ ρ₁ → ⟦ foldl (flip (eval s)) vs bss ⟧ᵛ ρ₂
|
eval-fold-valid : ∀ {s bss vs ρ₁ ρ₂} → ρ₁ , bss ⇒ᵇˢ ρ₂ → ⟦ vs ⟧ᵛ ρ₁ → ⟦ foldl (flip (eval s)) vs bss ⟧ᵛ ρ₂
|
||||||
|
|||||||
@@ -10,7 +10,6 @@ module Analysis.Forward.Lattices
|
|||||||
|
|
||||||
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
|
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
|
||||||
open import Data.Product using (proj₁; proj₂; _,_)
|
open import Data.Product using (proj₁; proj₂; _,_)
|
||||||
open import Data.Unit using (tt)
|
|
||||||
open import Data.Sum using (inj₁; inj₂)
|
open import Data.Sum using (inj₁; inj₂)
|
||||||
open import Data.List using (List; _∷_; []; foldr)
|
open import Data.List using (List; _∷_; []; foldr)
|
||||||
open import Data.List.Membership.Propositional using () renaming (_∈_ to _∈ˡ_)
|
open import Data.List.Membership.Propositional using () renaming (_∈_ to _∈ˡ_)
|
||||||
@@ -38,7 +37,7 @@ instance
|
|||||||
-- with keys strings. Use a bundle to avoid explicitly specifying operators.
|
-- with keys strings. Use a bundle to avoid explicitly specifying operators.
|
||||||
-- It's helpful to export these via 'public' since consumers tend to
|
-- It's helpful to export these via 'public' since consumers tend to
|
||||||
-- use various variable lattice operations.
|
-- use various variable lattice operations.
|
||||||
module VariableValuesFiniteMap = Lattice.FiniteMap.WithKeys String L tt vars
|
module VariableValuesFiniteMap = Lattice.FiniteMap String L vars
|
||||||
open VariableValuesFiniteMap
|
open VariableValuesFiniteMap
|
||||||
using ()
|
using ()
|
||||||
renaming
|
renaming
|
||||||
@@ -55,23 +54,13 @@ open VariableValuesFiniteMap
|
|||||||
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵛ≼m₂[k]ᵛ
|
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵛ≼m₂[k]ᵛ
|
||||||
; ∈k-dec to ∈k-decᵛ
|
; ∈k-dec to ∈k-decᵛ
|
||||||
; all-equal-keys to all-equal-keysᵛ
|
; all-equal-keys to all-equal-keysᵛ
|
||||||
)
|
; Provenance-union to Provenance-unionᵛ
|
||||||
public
|
; ⊔-Monotonicˡ to ⊔ᵛ-Monotonicˡ
|
||||||
open IsLattice isLatticeᵛ
|
|
||||||
using ()
|
|
||||||
renaming
|
|
||||||
( ⊔-Monotonicˡ to ⊔ᵛ-Monotonicˡ
|
|
||||||
; ⊔-Monotonicʳ to ⊔ᵛ-Monotonicʳ
|
; ⊔-Monotonicʳ to ⊔ᵛ-Monotonicʳ
|
||||||
; ⊔-idemp to ⊔ᵛ-idemp
|
; ⊔-idemp to ⊔ᵛ-idemp
|
||||||
)
|
)
|
||||||
public
|
public
|
||||||
open Lattice.FiniteMap.IterProdIsomorphism String L _
|
open VariableValuesFiniteMap.FixedHeight vars-Unique
|
||||||
using ()
|
|
||||||
renaming
|
|
||||||
( Provenance-union to Provenance-unionᵐ
|
|
||||||
)
|
|
||||||
public
|
|
||||||
open Lattice.FiniteMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight String L _ vars-Unique
|
|
||||||
using ()
|
using ()
|
||||||
renaming
|
renaming
|
||||||
( isFiniteHeightLattice to isFiniteHeightLatticeᵛ
|
( isFiniteHeightLattice to isFiniteHeightLatticeᵛ
|
||||||
@@ -83,7 +72,7 @@ open Lattice.FiniteMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight String L
|
|||||||
⊥ᵛ = Chain.Height.⊥ fixedHeightᵛ
|
⊥ᵛ = Chain.Height.⊥ fixedHeightᵛ
|
||||||
|
|
||||||
-- Finally, the map we care about is (state -> (variables -> value)). Bring that in.
|
-- Finally, the map we care about is (state -> (variables -> value)). Bring that in.
|
||||||
module StateVariablesFiniteMap = Lattice.FiniteMap.WithKeys State VariableValues tt states
|
module StateVariablesFiniteMap = Lattice.FiniteMap State VariableValues states
|
||||||
open StateVariablesFiniteMap
|
open StateVariablesFiniteMap
|
||||||
using (_[_]; []-∈; m₁≼m₂⇒m₁[ks]≼m₂[ks]; m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂)
|
using (_[_]; []-∈; m₁≼m₂⇒m₁[ks]≼m₂[ks]; m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂)
|
||||||
renaming
|
renaming
|
||||||
@@ -96,20 +85,15 @@ open StateVariablesFiniteMap
|
|||||||
; _≼_ to _≼ᵐ_
|
; _≼_ to _≼ᵐ_
|
||||||
; ≈-Decidable to ≈ᵐ-Decidable
|
; ≈-Decidable to ≈ᵐ-Decidable
|
||||||
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ
|
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ
|
||||||
|
; ≈-sym to ≈ᵐ-sym
|
||||||
)
|
)
|
||||||
public
|
public
|
||||||
open Lattice.FiniteMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight State VariableValues _ states-Unique
|
open StateVariablesFiniteMap.FixedHeight states-Unique
|
||||||
using ()
|
using ()
|
||||||
renaming
|
renaming
|
||||||
( isFiniteHeightLattice to isFiniteHeightLatticeᵐ
|
( isFiniteHeightLattice to isFiniteHeightLatticeᵐ
|
||||||
)
|
)
|
||||||
public
|
public
|
||||||
open IsFiniteHeightLattice isFiniteHeightLatticeᵐ
|
|
||||||
using ()
|
|
||||||
renaming
|
|
||||||
( ≈-sym to ≈ᵐ-sym
|
|
||||||
)
|
|
||||||
public
|
|
||||||
|
|
||||||
-- We now have our (state -> (variables -> value)) map.
|
-- We now have our (state -> (variables -> value)) map.
|
||||||
-- Define a couple of helpers to retrieve values from it. Specifically,
|
-- Define a couple of helpers to retrieve values from it. Specifically,
|
||||||
@@ -197,7 +181,7 @@ module _ {{latticeInterpretationˡ : LatticeInterpretation isLatticeˡ}} where
|
|||||||
⟦⟧ᵛ-⊔ᵛ-∨ : ∀ {vs₁ vs₂ : VariableValues} → (⟦ vs₁ ⟧ᵛ ∨ ⟦ vs₂ ⟧ᵛ) ⇒ ⟦ vs₁ ⊔ᵛ vs₂ ⟧ᵛ
|
⟦⟧ᵛ-⊔ᵛ-∨ : ∀ {vs₁ vs₂ : VariableValues} → (⟦ vs₁ ⟧ᵛ ∨ ⟦ vs₂ ⟧ᵛ) ⇒ ⟦ vs₁ ⊔ᵛ vs₂ ⟧ᵛ
|
||||||
⟦⟧ᵛ-⊔ᵛ-∨ {vs₁} {vs₂} ρ ⟦vs₁⟧ρ∨⟦vs₂⟧ρ {k} {l} k,l∈vs₁₂ {v} k,v∈ρ
|
⟦⟧ᵛ-⊔ᵛ-∨ {vs₁} {vs₂} ρ ⟦vs₁⟧ρ∨⟦vs₂⟧ρ {k} {l} k,l∈vs₁₂ {v} k,v∈ρ
|
||||||
with ((l₁ , l₂) , (refl , (k,l₁∈vs₁ , k,l₂∈vs₂)))
|
with ((l₁ , l₂) , (refl , (k,l₁∈vs₁ , k,l₂∈vs₂)))
|
||||||
← Provenance-unionᵐ vs₁ vs₂ k,l∈vs₁₂
|
← Provenance-unionᵛ vs₁ vs₂ k,l∈vs₁₂
|
||||||
with ⟦vs₁⟧ρ∨⟦vs₂⟧ρ
|
with ⟦vs₁⟧ρ∨⟦vs₂⟧ρ
|
||||||
... | inj₁ ⟦vs₁⟧ρ = ⟦⟧ˡ-⊔ˡ-∨ {l₁} {l₂} v (inj₁ (⟦vs₁⟧ρ k,l₁∈vs₁ k,v∈ρ))
|
... | inj₁ ⟦vs₁⟧ρ = ⟦⟧ˡ-⊔ˡ-∨ {l₁} {l₂} v (inj₁ (⟦vs₁⟧ρ k,l₁∈vs₁ k,v∈ρ))
|
||||||
... | inj₂ ⟦vs₂⟧ρ = ⟦⟧ˡ-⊔ˡ-∨ {l₁} {l₂} v (inj₂ (⟦vs₂⟧ρ k,l₂∈vs₂ k,v∈ρ))
|
... | inj₂ ⟦vs₂⟧ρ = ⟦⟧ˡ-⊔ˡ-∨ {l₁} {l₂} v (inj₂ (⟦vs₂⟧ρ k,l₂∈vs₂ k,v∈ρ))
|
||||||
|
|||||||
@@ -16,6 +16,7 @@ open import Lattice
|
|||||||
open import Equivalence
|
open import Equivalence
|
||||||
open import Showable using (Showable; show)
|
open import Showable using (Showable; show)
|
||||||
open import Utils using (_⇒_; _∧_; _∨_)
|
open import Utils using (_⇒_; _∧_; _∨_)
|
||||||
|
open import Analysis.Utils using (eval-combine₂)
|
||||||
import Analysis.Forward
|
import Analysis.Forward
|
||||||
|
|
||||||
data Sign : Set where
|
data Sign : Set where
|
||||||
@@ -78,10 +79,9 @@ open AB.Plain 0ˢ using ()
|
|||||||
; _≼_ to _≼ᵍ_
|
; _≼_ to _≼ᵍ_
|
||||||
; _⊔_ to _⊔ᵍ_
|
; _⊔_ to _⊔ᵍ_
|
||||||
; _⊓_ to _⊓ᵍ_
|
; _⊓_ to _⊓ᵍ_
|
||||||
|
; ≼-trans to ≼ᵍ-trans
|
||||||
)
|
)
|
||||||
|
|
||||||
open IsLattice isLatticeᵍ using () renaming (≼-trans to ≼ᵍ-trans)
|
|
||||||
|
|
||||||
plus : SignLattice → SignLattice → SignLattice
|
plus : SignLattice → SignLattice → SignLattice
|
||||||
plus ⊥ᵍ _ = ⊥ᵍ
|
plus ⊥ᵍ _ = ⊥ᵍ
|
||||||
plus _ ⊥ᵍ = ⊥ᵍ
|
plus _ ⊥ᵍ = ⊥ᵍ
|
||||||
@@ -102,6 +102,9 @@ plus [ 0ˢ ]ᵍ [ 0ˢ ]ᵍ = [ 0ˢ ]ᵍ
|
|||||||
postulate plus-Monoˡ : ∀ (s₂ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (λ s₁ → plus s₁ s₂)
|
postulate plus-Monoˡ : ∀ (s₂ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (λ s₁ → plus s₁ s₂)
|
||||||
postulate plus-Monoʳ : ∀ (s₁ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (plus s₁)
|
postulate plus-Monoʳ : ∀ (s₁ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (plus s₁)
|
||||||
|
|
||||||
|
plus-Mono₂ : Monotonic₂ _≼ᵍ_ _≼ᵍ_ _≼ᵍ_ plus
|
||||||
|
plus-Mono₂ = (plus-Monoˡ , plus-Monoʳ)
|
||||||
|
|
||||||
minus : SignLattice → SignLattice → SignLattice
|
minus : SignLattice → SignLattice → SignLattice
|
||||||
minus ⊥ᵍ _ = ⊥ᵍ
|
minus ⊥ᵍ _ = ⊥ᵍ
|
||||||
minus _ ⊥ᵍ = ⊥ᵍ
|
minus _ ⊥ᵍ = ⊥ᵍ
|
||||||
@@ -120,6 +123,9 @@ minus [ 0ˢ ]ᵍ [ 0ˢ ]ᵍ = [ 0ˢ ]ᵍ
|
|||||||
postulate minus-Monoˡ : ∀ (s₂ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (λ s₁ → minus s₁ s₂)
|
postulate minus-Monoˡ : ∀ (s₂ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (λ s₁ → minus s₁ s₂)
|
||||||
postulate minus-Monoʳ : ∀ (s₁ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (minus s₁)
|
postulate minus-Monoʳ : ∀ (s₁ : SignLattice) → Monotonic _≼ᵍ_ _≼ᵍ_ (minus s₁)
|
||||||
|
|
||||||
|
minus-Mono₂ : Monotonic₂ _≼ᵍ_ _≼ᵍ_ _≼ᵍ_ minus
|
||||||
|
minus-Mono₂ = (minus-Monoˡ , minus-Monoʳ)
|
||||||
|
|
||||||
⟦_⟧ᵍ : SignLattice → Value → Set
|
⟦_⟧ᵍ : SignLattice → Value → Set
|
||||||
⟦_⟧ᵍ ⊥ᵍ _ = ⊥
|
⟦_⟧ᵍ ⊥ᵍ _ = ⊥
|
||||||
⟦_⟧ᵍ ⊤ᵍ _ = ⊤
|
⟦_⟧ᵍ ⊤ᵍ _ = ⊤
|
||||||
@@ -196,29 +202,13 @@ module WithProg (prog : Program) where
|
|||||||
|
|
||||||
eval-Monoʳ : ∀ (e : Expr) → Monotonic _≼ᵛ_ _≼ᵍ_ (eval e)
|
eval-Monoʳ : ∀ (e : Expr) → Monotonic _≼ᵛ_ _≼ᵍ_ (eval e)
|
||||||
eval-Monoʳ (e₁ + e₂) {vs₁} {vs₂} vs₁≼vs₂ =
|
eval-Monoʳ (e₁ + e₂) {vs₁} {vs₂} vs₁≼vs₂ =
|
||||||
let
|
eval-combine₂ (λ {x} {y} {z} → ≼ᵍ-trans {x} {y} {z})
|
||||||
-- TODO: can this be done with less boilerplate?
|
plus plus-Mono₂ {o₁ = eval e₁ vs₁}
|
||||||
g₁vs₁ = eval e₁ vs₁
|
(eval-Monoʳ e₁ vs₁≼vs₂) (eval-Monoʳ e₂ vs₁≼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₂ =
|
eval-Monoʳ (e₁ - e₂) {vs₁} {vs₂} vs₁≼vs₂ =
|
||||||
let
|
eval-combine₂ (λ {x} {y} {z} → ≼ᵍ-trans {x} {y} {z})
|
||||||
-- TODO: here too -- can this be done with less boilerplate?
|
minus minus-Mono₂ {o₁ = eval e₁ vs₁}
|
||||||
g₁vs₁ = eval e₁ vs₁
|
(eval-Monoʳ e₁ vs₁≼vs₂) (eval-Monoʳ e₂ vs₁≼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₂
|
eval-Monoʳ (` k) {vs₁@((kvs₁ , _) , _)} {vs₂@((kvs₂ , _), _)} vs₁≼vs₂
|
||||||
with ∈k-decᵛ k kvs₁ | ∈k-decᵛ k kvs₂
|
with ∈k-decᵛ k kvs₁ | ∈k-decᵛ k kvs₂
|
||||||
... | yes k∈kvs₁ | yes k∈kvs₂ =
|
... | yes k∈kvs₁ | yes k∈kvs₂ =
|
||||||
@@ -302,4 +292,8 @@ module WithProg (prog : Program) where
|
|||||||
eval-valid (⇒ᵉ-ℕ ρ 0) _ = refl
|
eval-valid (⇒ᵉ-ℕ ρ 0) _ = refl
|
||||||
eval-valid (⇒ᵉ-ℕ ρ (suc n')) _ = (n' , refl)
|
eval-valid (⇒ᵉ-ℕ ρ (suc n')) _ = (n' , refl)
|
||||||
|
|
||||||
analyze-correct = Analysis.Forward.WithProg.analyze-correct
|
instance
|
||||||
|
SignEvalValid : ValidExprEvaluator SignEval latticeInterpretationᵍ
|
||||||
|
SignEvalValid = record { valid = eval-valid }
|
||||||
|
|
||||||
|
analyze-correct = Analysis.Forward.WithProg.analyze-correct SignLattice prog tt
|
||||||
|
|||||||
15
Analysis/Utils.agda
Normal file
15
Analysis/Utils.agda
Normal file
@@ -0,0 +1,15 @@
|
|||||||
|
module Analysis.Utils where
|
||||||
|
|
||||||
|
open import Data.Product using (_,_)
|
||||||
|
open import Lattice
|
||||||
|
|
||||||
|
module _ {o} {O : Set o} {_≼ᴼ_ : O → O → Set o}
|
||||||
|
(≼ᴼ-trans : ∀ {o₁ o₂ o₃} → o₁ ≼ᴼ o₂ → o₂ ≼ᴼ o₃ → o₁ ≼ᴼ o₃)
|
||||||
|
(combine : O → O → O) (combine-Mono₂ : Monotonic₂ _≼ᴼ_ _≼ᴼ_ _≼ᴼ_ combine) where
|
||||||
|
|
||||||
|
eval-combine₂ : {o₁ o₂ o₃ o₄ : O} → o₁ ≼ᴼ o₃ → o₂ ≼ᴼ o₄ →
|
||||||
|
combine o₁ o₂ ≼ᴼ combine o₃ o₄
|
||||||
|
eval-combine₂ {o₁} {o₂} {o₃} {o₄} o₁≼o₃ o₂≼o₄ =
|
||||||
|
let (combine-Monoˡ , combine-Monoʳ) = combine-Mono₂
|
||||||
|
in ≼ᴼ-trans (combine-Monoˡ o₂ o₁≼o₃)
|
||||||
|
(combine-Monoʳ o₃ o₂≼o₄)
|
||||||
@@ -5,8 +5,8 @@ module Fixedpoint {a} {A : Set a}
|
|||||||
{h : ℕ}
|
{h : ℕ}
|
||||||
{_≈_ : A → A → Set a}
|
{_≈_ : A → A → Set a}
|
||||||
{_⊔_ : A → A → A} {_⊓_ : A → A → A}
|
{_⊔_ : A → A → A} {_⊓_ : A → A → A}
|
||||||
(≈-Decidable : IsDecidable _≈_)
|
{{ ≈-Decidable : IsDecidable _≈_ }}
|
||||||
(flA : IsFiniteHeightLattice A h _≈_ _⊔_ _⊓_)
|
{{flA : IsFiniteHeightLattice A h _≈_ _⊔_ _⊓_}}
|
||||||
(f : A → A)
|
(f : A → A)
|
||||||
(Monotonicᶠ : Monotonic (IsFiniteHeightLattice._≼_ flA)
|
(Monotonicᶠ : Monotonic (IsFiniteHeightLattice._≼_ flA)
|
||||||
(IsFiniteHeightLattice._≼_ flA) f) where
|
(IsFiniteHeightLattice._≼_ flA) f) where
|
||||||
@@ -28,24 +28,9 @@ private
|
|||||||
using ()
|
using ()
|
||||||
renaming
|
renaming
|
||||||
( ⊥ to ⊥ᴬ
|
( ⊥ to ⊥ᴬ
|
||||||
; longestChain to longestChainᴬ
|
|
||||||
; bounded to boundedᴬ
|
; bounded to boundedᴬ
|
||||||
)
|
)
|
||||||
|
|
||||||
|
|
||||||
⊥ᴬ≼ : ∀ (a : A) → ⊥ᴬ ≼ a
|
|
||||||
⊥ᴬ≼ a with ≈-dec a ⊥ᴬ
|
|
||||||
... | yes a≈⊥ᴬ = ≼-cong a≈⊥ᴬ ≈-refl (≼-refl a)
|
|
||||||
... | no a̷≈⊥ᴬ with ≈-dec ⊥ᴬ (a ⊓ ⊥ᴬ)
|
|
||||||
... | yes ⊥ᴬ≈a⊓⊥ᴬ = ≈-trans (⊔-comm ⊥ᴬ a) (≈-trans (≈-⊔-cong (≈-refl {a}) ⊥ᴬ≈a⊓⊥ᴬ) (absorb-⊔-⊓ a ⊥ᴬ))
|
|
||||||
... | no ⊥ᴬ̷≈a⊓⊥ᴬ = ⊥-elim (ChainA.Bounded-suc-n boundedᴬ (ChainA.step x≺⊥ᴬ ≈-refl longestChainᴬ))
|
|
||||||
where
|
|
||||||
⊥ᴬ⊓a̷≈⊥ᴬ : ¬ (⊥ᴬ ⊓ a) ≈ ⊥ᴬ
|
|
||||||
⊥ᴬ⊓a̷≈⊥ᴬ = λ ⊥ᴬ⊓a≈⊥ᴬ → ⊥ᴬ̷≈a⊓⊥ᴬ (≈-trans (≈-sym ⊥ᴬ⊓a≈⊥ᴬ) (⊓-comm _ _))
|
|
||||||
|
|
||||||
x≺⊥ᴬ : (⊥ᴬ ⊓ a) ≺ ⊥ᴬ
|
|
||||||
x≺⊥ᴬ = (≈-trans (⊔-comm _ _) (≈-trans (≈-refl {⊥ᴬ ⊔ (⊥ᴬ ⊓ a)}) (absorb-⊔-⊓ ⊥ᴬ a)) , ⊥ᴬ⊓a̷≈⊥ᴬ)
|
|
||||||
|
|
||||||
-- using 'g', for gas, here helps make sure the function terminates.
|
-- using 'g', for gas, here helps make sure the function terminates.
|
||||||
-- since A forms a fixed-height lattice, we must find a solution after
|
-- since A forms a fixed-height lattice, we must find a solution after
|
||||||
-- 'h' steps at most. Gas is set up such that as soon as it runs
|
-- 'h' steps at most. Gas is set up such that as soon as it runs
|
||||||
@@ -65,7 +50,7 @@ private
|
|||||||
c' rewrite +-comm 1 hᶜ = ChainA.concat c (ChainA.step a₂≺fa₂ ≈-refl (ChainA.done (≈-refl {f a₂})))
|
c' rewrite +-comm 1 hᶜ = ChainA.concat c (ChainA.step a₂≺fa₂ ≈-refl (ChainA.done (≈-refl {f a₂})))
|
||||||
|
|
||||||
fix : Σ A (λ a → a ≈ f a)
|
fix : Σ A (λ a → a ≈ f a)
|
||||||
fix = doStep (suc h) 0 ⊥ᴬ ⊥ᴬ (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥ᴬ≼ (f ⊥ᴬ))
|
fix = doStep (suc h) 0 ⊥ᴬ ⊥ᴬ (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥≼ (f ⊥ᴬ))
|
||||||
|
|
||||||
aᶠ : A
|
aᶠ : A
|
||||||
aᶠ = proj₁ fix
|
aᶠ = proj₁ fix
|
||||||
@@ -85,4 +70,4 @@ private
|
|||||||
... | no _ = stepPreservesLess g' _ _ _ b b≈fb (≼-cong ≈-refl (≈-sym b≈fb) (Monotonicᶠ a₂≼b)) _ _ _
|
... | no _ = stepPreservesLess g' _ _ _ b b≈fb (≼-cong ≈-refl (≈-sym b≈fb) (Monotonicᶠ a₂≼b)) _ _ _
|
||||||
|
|
||||||
aᶠ≼ : ∀ (a : A) → a ≈ f a → aᶠ ≼ a
|
aᶠ≼ : ∀ (a : A) → a ≈ f a → aᶠ ≼ a
|
||||||
aᶠ≼ a a≈fa = stepPreservesLess (suc h) 0 ⊥ᴬ ⊥ᴬ a a≈fa (⊥ᴬ≼ a) (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥ᴬ≼ (f ⊥ᴬ))
|
aᶠ≼ a a≈fa = stepPreservesLess (suc h) 0 ⊥ᴬ ⊥ᴬ a a≈fa (⊥≼ a) (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥≼ (f ⊥ᴬ))
|
||||||
|
|||||||
@@ -68,6 +68,7 @@ module TransportFiniteHeight
|
|||||||
renaming (⊥ to ⊥₁; ⊤ to ⊤₁; bounded to bounded₁; longestChain to c)
|
renaming (⊥ to ⊥₁; ⊤ to ⊤₁; bounded to bounded₁; longestChain to c)
|
||||||
|
|
||||||
instance
|
instance
|
||||||
|
fixedHeight : IsLattice.FixedHeight lB height
|
||||||
fixedHeight = record
|
fixedHeight = record
|
||||||
{ ⊥ = f ⊥₁
|
{ ⊥ = f ⊥₁
|
||||||
; ⊤ = f ⊤₁
|
; ⊤ = f ⊤₁
|
||||||
|
|||||||
37
Lattice.agda
37
Lattice.agda
@@ -4,7 +4,8 @@ open import Equivalence
|
|||||||
import Chain
|
import Chain
|
||||||
|
|
||||||
open import Relation.Binary.Core using (_Preserves_⟶_ )
|
open import Relation.Binary.Core using (_Preserves_⟶_ )
|
||||||
open import Relation.Nullary using (Dec; ¬_)
|
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
||||||
|
open import Data.Empty using (⊥-elim)
|
||||||
open import Data.Nat as Nat using (ℕ)
|
open import Data.Nat as Nat using (ℕ)
|
||||||
open import Data.Product using (_×_; Σ; _,_)
|
open import Data.Product using (_×_; Σ; _,_)
|
||||||
open import Data.Sum using (_⊎_; inj₁; inj₂)
|
open import Data.Sum using (_⊎_; inj₁; inj₂)
|
||||||
@@ -21,6 +22,18 @@ module _ {a b} {A : Set a} {B : Set b}
|
|||||||
Monotonic : (A → B) → Set (a ⊔ℓ b)
|
Monotonic : (A → B) → Set (a ⊔ℓ b)
|
||||||
Monotonic f = ∀ {a₁ a₂ : A} → a₁ ≼₁ a₂ → f a₁ ≼₂ f a₂
|
Monotonic f = ∀ {a₁ a₂ : A} → a₁ ≼₁ a₂ → f a₁ ≼₂ f a₂
|
||||||
|
|
||||||
|
Monotonicˡ : ∀ {c} {C : Set c} → (A → C → B) → Set (a ⊔ℓ b ⊔ℓ c)
|
||||||
|
Monotonicˡ f = ∀ c → Monotonic (λ a → f a c)
|
||||||
|
|
||||||
|
Monotonicʳ : ∀ {c} {C : Set c} → (C → A → B) → Set (a ⊔ℓ b ⊔ℓ c)
|
||||||
|
Monotonicʳ f = ∀ a → Monotonic (f a)
|
||||||
|
|
||||||
|
module _ {a b c} {A : Set a} {B : Set b} {C : Set c}
|
||||||
|
(_≼₁_ : A → A → Set a) (_≼₂_ : B → B → Set b) (_≼₃_ : C → C → Set c) where
|
||||||
|
|
||||||
|
Monotonic₂ : (A → B → C) → Set (a ⊔ℓ b ⊔ℓ c)
|
||||||
|
Monotonic₂ f = Monotonicˡ _≼₁_ _≼₃_ f × Monotonicʳ _≼₂_ _≼₃_ f
|
||||||
|
|
||||||
record IsSemilattice {a} (A : Set a)
|
record IsSemilattice {a} (A : Set a)
|
||||||
(_≈_ : A → A → Set a)
|
(_≈_ : A → A → Set a)
|
||||||
(_⊔_ : A → A → A) : Set a where
|
(_⊔_ : A → A → A) : Set a where
|
||||||
@@ -224,6 +237,28 @@ record IsFiniteHeightLattice {a} (A : Set a)
|
|||||||
field
|
field
|
||||||
{{fixedHeight}} : FixedHeight h
|
{{fixedHeight}} : FixedHeight h
|
||||||
|
|
||||||
|
-- If the equality is decidable, we can prove that the top and bottom
|
||||||
|
-- elements of the chain are least and greatest elements of the lattice
|
||||||
|
module _ {{≈-Decidable : IsDecidable _≈_}} where
|
||||||
|
open IsDecidable ≈-Decidable using () renaming (R-dec to ≈-dec)
|
||||||
|
|
||||||
|
module MyChain = Chain _≈_ ≈-equiv _≺_ ≺-cong
|
||||||
|
open MyChain.Height fixedHeight using (⊥; ⊤) public
|
||||||
|
open MyChain.Height fixedHeight using (bounded; longestChain)
|
||||||
|
|
||||||
|
⊥≼ : ∀ (a : A) → ⊥ ≼ a
|
||||||
|
⊥≼ a with ≈-dec a ⊥
|
||||||
|
... | yes a≈⊥ = ≼-cong a≈⊥ ≈-refl (≼-refl a)
|
||||||
|
... | no a̷≈⊥ with ≈-dec ⊥ (a ⊓ ⊥)
|
||||||
|
... | yes ⊥≈a⊓⊥ = ≈-trans (⊔-comm ⊥ a) (≈-trans (≈-⊔-cong (≈-refl {a}) ⊥≈a⊓⊥) (absorb-⊔-⊓ a ⊥))
|
||||||
|
... | no ⊥ᴬ̷≈a⊓⊥ = ⊥-elim (MyChain.Bounded-suc-n bounded (MyChain.step x≺⊥ ≈-refl longestChain))
|
||||||
|
where
|
||||||
|
⊥⊓a̷≈⊥ : ¬ (⊥ ⊓ a) ≈ ⊥
|
||||||
|
⊥⊓a̷≈⊥ = λ ⊥⊓a≈⊥ → ⊥ᴬ̷≈a⊓⊥ (≈-trans (≈-sym ⊥⊓a≈⊥) (⊓-comm _ _))
|
||||||
|
|
||||||
|
x≺⊥ : (⊥ ⊓ a) ≺ ⊥
|
||||||
|
x≺⊥ = (≈-trans (⊔-comm _ _) (≈-trans (≈-refl {⊥ ⊔ (⊥ ⊓ a)}) (absorb-⊔-⊓ ⊥ a)) , ⊥⊓a̷≈⊥)
|
||||||
|
|
||||||
module ChainMapping {a b} {A : Set a} {B : Set b}
|
module ChainMapping {a b} {A : Set a} {B : Set b}
|
||||||
{_≈₁_ : A → A → Set a} {_≈₂_ : B → B → Set b}
|
{_≈₁_ : A → A → Set a} {_≈₂_ : B → B → Set b}
|
||||||
{_⊔₁_ : A → A → A} {_⊔₂_ : B → B → B}
|
{_⊔₁_ : A → A → A} {_⊔₂_ : B → B → B}
|
||||||
|
|||||||
@@ -321,7 +321,7 @@ module Plain (x : A) where
|
|||||||
; isLattice = isLattice
|
; isLattice = isLattice
|
||||||
}
|
}
|
||||||
|
|
||||||
open IsLattice isLattice using (_≼_; _≺_; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
|
open IsLattice isLattice using (_≼_; _≺_; ≼-trans; ≼-refl; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
|
||||||
|
|
||||||
⊥≺[x] : ∀ (x : A) → ⊥ ≺ [ x ]
|
⊥≺[x] : ∀ (x : A) → ⊥ ≺ [ x ]
|
||||||
⊥≺[x] x = (≈-refl , λ ())
|
⊥≺[x] x = (≈-refl , λ ())
|
||||||
|
|||||||
1200
Lattice/Builder.agda
Normal file
1200
Lattice/Builder.agda
Normal file
File diff suppressed because it is too large
Load Diff
169
Lattice/ExtendBelow.agda
Normal file
169
Lattice/ExtendBelow.agda
Normal file
@@ -0,0 +1,169 @@
|
|||||||
|
open import Lattice
|
||||||
|
|
||||||
|
module Lattice.ExtendBelow {a} (A : Set a)
|
||||||
|
{_≈₁_ : A → A → Set a} {_⊔₁_ : A → A → A} {_⊓₁_ : A → A → A}
|
||||||
|
{{lA : IsLattice A _≈₁_ _⊔₁_ _⊓₁_}} where
|
||||||
|
|
||||||
|
open import Equivalence
|
||||||
|
open import Showable using (Showable; show)
|
||||||
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
|
open import Relation.Binary.PropositionalEquality using (refl)
|
||||||
|
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
||||||
|
|
||||||
|
open IsLattice lA using ()
|
||||||
|
renaming ( ≈-equiv to ≈₁-equiv
|
||||||
|
; ≈-⊔-cong to ≈₁-⊔₁-cong
|
||||||
|
; ⊔-assoc to ⊔₁-assoc; ⊔-comm to ⊔₁-comm; ⊔-idemp to ⊔₁-idemp
|
||||||
|
; ≈-⊓-cong to ≈₁-⊓₁-cong
|
||||||
|
; ⊓-assoc to ⊓₁-assoc; ⊓-comm to ⊓₁-comm; ⊓-idemp to ⊓₁-idemp
|
||||||
|
; absorb-⊔-⊓ to absorb-⊔₁-⊓₁; absorb-⊓-⊔ to absorb-⊓₁-⊔₁
|
||||||
|
)
|
||||||
|
|
||||||
|
open IsEquivalence ≈₁-equiv using ()
|
||||||
|
renaming (≈-refl to ≈₁-refl; ≈-sym to ≈₁-sym; ≈-trans to ≈₁-trans)
|
||||||
|
|
||||||
|
data ExtendBelow : Set a where
|
||||||
|
[_] : A → ExtendBelow
|
||||||
|
⊥ : ExtendBelow
|
||||||
|
|
||||||
|
instance
|
||||||
|
showable : {{ showableA : Showable A }} → Showable ExtendBelow
|
||||||
|
showable = record
|
||||||
|
{ show = (λ
|
||||||
|
{ ⊥ → "⊥"
|
||||||
|
; [ a ] → show a
|
||||||
|
})
|
||||||
|
}
|
||||||
|
|
||||||
|
data _≈_ : ExtendBelow → ExtendBelow → Set a where
|
||||||
|
≈-⊥-⊥ : ⊥ ≈ ⊥
|
||||||
|
≈-lift : ∀ {x y : A} → x ≈₁ y → [ x ] ≈ [ y ]
|
||||||
|
|
||||||
|
≈-refl : ∀ {ab : ExtendBelow} → ab ≈ ab
|
||||||
|
≈-refl {⊥} = ≈-⊥-⊥
|
||||||
|
≈-refl {[ x ]} = ≈-lift ≈₁-refl
|
||||||
|
|
||||||
|
≈-sym : ∀ {ab₁ ab₂ : ExtendBelow} → ab₁ ≈ ab₂ → ab₂ ≈ ab₁
|
||||||
|
≈-sym ≈-⊥-⊥ = ≈-⊥-⊥
|
||||||
|
≈-sym (≈-lift x≈₁y) = ≈-lift (≈₁-sym x≈₁y)
|
||||||
|
|
||||||
|
≈-trans : ∀ {ab₁ ab₂ ab₃ : ExtendBelow} → ab₁ ≈ ab₂ → ab₂ ≈ ab₃ → ab₁ ≈ ab₃
|
||||||
|
≈-trans ≈-⊥-⊥ ≈-⊥-⊥ = ≈-⊥-⊥
|
||||||
|
≈-trans (≈-lift a₁≈a₂) (≈-lift a₂≈a₃) = ≈-lift (≈₁-trans a₁≈a₂ a₂≈a₃)
|
||||||
|
|
||||||
|
instance
|
||||||
|
≈-equiv : IsEquivalence ExtendBelow _≈_
|
||||||
|
≈-equiv = record
|
||||||
|
{ ≈-refl = ≈-refl
|
||||||
|
; ≈-sym = ≈-sym
|
||||||
|
; ≈-trans = ≈-trans
|
||||||
|
}
|
||||||
|
|
||||||
|
_⊔_ : ExtendBelow → ExtendBelow → ExtendBelow
|
||||||
|
_⊔_ ⊥ x = x
|
||||||
|
_⊔_ [ a₁ ] ⊥ = [ a₁ ]
|
||||||
|
_⊔_ [ a₁ ] [ a₂ ] = [ a₁ ⊔₁ a₂ ]
|
||||||
|
|
||||||
|
_⊓_ : ExtendBelow → ExtendBelow → ExtendBelow
|
||||||
|
_⊓_ ⊥ x = ⊥
|
||||||
|
_⊓_ [ _ ] ⊥ = ⊥
|
||||||
|
_⊓_ [ a₁ ] [ a₂ ] = [ a₁ ⊓₁ a₂ ]
|
||||||
|
|
||||||
|
≈-⊔-cong : ∀ {x₁ x₂ x₃ x₄} → x₁ ≈ x₂ → x₃ ≈ x₄ →
|
||||||
|
(x₁ ⊔ x₃) ≈ (x₂ ⊔ x₄)
|
||||||
|
≈-⊔-cong .{⊥} .{⊥} {x₃} {x₄} ≈-⊥-⊥ [a₃]≈[a₄] = [a₃]≈[a₄]
|
||||||
|
≈-⊔-cong {x₁ = [ a₁ ]} {x₂ = [ a₂ ]} .{⊥} .{⊥} [a₁]≈[a₂] ≈-⊥-⊥ = [a₁]≈[a₂]
|
||||||
|
≈-⊔-cong {x₁ = [ a₁ ]} {x₂ = [ a₂ ]} {x₃ = [ a₃ ]} {x₄ = [ a₄ ]} (≈-lift a₁≈a₂) (≈-lift a₃≈a₄) = ≈-lift (≈₁-⊔₁-cong a₁≈a₂ a₃≈a₄)
|
||||||
|
|
||||||
|
⊔-assoc : ∀ (x₁ x₂ x₃ : ExtendBelow) → ((x₁ ⊔ x₂) ⊔ x₃) ≈ (x₁ ⊔ (x₂ ⊔ x₃))
|
||||||
|
⊔-assoc ⊥ x₁ x₂ = ≈-refl
|
||||||
|
⊔-assoc [ a₁ ] ⊥ x₂ = ≈-refl
|
||||||
|
⊔-assoc [ a₁ ] [ a₂ ] ⊥ = ≈-refl
|
||||||
|
⊔-assoc [ a₁ ] [ a₂ ] [ a₃ ] = ≈-lift (⊔₁-assoc a₁ a₂ a₃)
|
||||||
|
|
||||||
|
⊔-comm : ∀ (x₁ x₂ : ExtendBelow) → (x₁ ⊔ x₂) ≈ (x₂ ⊔ x₁)
|
||||||
|
⊔-comm ⊥ ⊥ = ≈-refl
|
||||||
|
⊔-comm ⊥ [ a₂ ] = ≈-refl
|
||||||
|
⊔-comm [ a₁ ] ⊥ = ≈-refl
|
||||||
|
⊔-comm [ a₁ ] [ a₂ ] = ≈-lift (⊔₁-comm a₁ a₂)
|
||||||
|
|
||||||
|
⊔-idemp : ∀ (x : ExtendBelow) → (x ⊔ x) ≈ x
|
||||||
|
⊔-idemp ⊥ = ≈-refl
|
||||||
|
⊔-idemp [ a ] = ≈-lift (⊔₁-idemp a)
|
||||||
|
|
||||||
|
≈-⊓-cong : ∀ {x₁ x₂ x₃ x₄} → x₁ ≈ x₂ → x₃ ≈ x₄ →
|
||||||
|
(x₁ ⊓ x₃) ≈ (x₂ ⊓ x₄)
|
||||||
|
≈-⊓-cong .{⊥} .{⊥} {x₃} {x₄} ≈-⊥-⊥ [a₃]≈[a₄] = ≈-⊥-⊥
|
||||||
|
≈-⊓-cong {x₁ = [ a₁ ]} {x₂ = [ a₂ ]} .{⊥} .{⊥} [a₁]≈[a₂] ≈-⊥-⊥ = ≈-⊥-⊥
|
||||||
|
≈-⊓-cong {x₁ = [ a₁ ]} {x₂ = [ a₂ ]} {x₃ = [ a₃ ]} {x₄ = [ a₄ ]} (≈-lift a₁≈a₂) (≈-lift a₃≈a₄) = ≈-lift (≈₁-⊓₁-cong a₁≈a₂ a₃≈a₄)
|
||||||
|
|
||||||
|
⊓-assoc : ∀ (x₁ x₂ x₃ : ExtendBelow) → ((x₁ ⊓ x₂) ⊓ x₃) ≈ (x₁ ⊓ (x₂ ⊓ x₃))
|
||||||
|
⊓-assoc ⊥ x₁ x₂ = ≈-refl
|
||||||
|
⊓-assoc [ a₁ ] ⊥ x₂ = ≈-refl
|
||||||
|
⊓-assoc [ a₁ ] [ a₂ ] ⊥ = ≈-refl
|
||||||
|
⊓-assoc [ a₁ ] [ a₂ ] [ a₃ ] = ≈-lift (⊓₁-assoc a₁ a₂ a₃)
|
||||||
|
|
||||||
|
⊓-comm : ∀ (x₁ x₂ : ExtendBelow) → (x₁ ⊓ x₂) ≈ (x₂ ⊓ x₁)
|
||||||
|
⊓-comm ⊥ ⊥ = ≈-refl
|
||||||
|
⊓-comm ⊥ [ a₂ ] = ≈-refl
|
||||||
|
⊓-comm [ a₁ ] ⊥ = ≈-refl
|
||||||
|
⊓-comm [ a₁ ] [ a₂ ] = ≈-lift (⊓₁-comm a₁ a₂)
|
||||||
|
|
||||||
|
⊓-idemp : ∀ (x : ExtendBelow) → (x ⊓ x) ≈ x
|
||||||
|
⊓-idemp ⊥ = ≈-refl
|
||||||
|
⊓-idemp [ a ] = ≈-lift (⊓₁-idemp a)
|
||||||
|
|
||||||
|
absorb-⊔-⊓ : (x₁ x₂ : ExtendBelow) → (x₁ ⊔ (x₁ ⊓ x₂)) ≈ x₁
|
||||||
|
absorb-⊔-⊓ ⊥ ⊥ = ≈-refl
|
||||||
|
absorb-⊔-⊓ ⊥ [ a₂ ] = ≈-refl
|
||||||
|
absorb-⊔-⊓ [ a₁ ] ⊥ = ≈-refl
|
||||||
|
absorb-⊔-⊓ [ a₁ ] [ a₂ ] = ≈-lift (absorb-⊔₁-⊓₁ a₁ a₂)
|
||||||
|
|
||||||
|
absorb-⊓-⊔ : (x₁ x₂ : ExtendBelow) → (x₁ ⊓ (x₁ ⊔ x₂)) ≈ x₁
|
||||||
|
absorb-⊓-⊔ ⊥ ⊥ = ≈-refl
|
||||||
|
absorb-⊓-⊔ ⊥ [ a₂ ] = ≈-refl
|
||||||
|
absorb-⊓-⊔ [ a₁ ] ⊥ = ⊓-idemp [ a₁ ]
|
||||||
|
absorb-⊓-⊔ [ a₁ ] [ a₂ ] = ≈-lift (absorb-⊓₁-⊔₁ a₁ a₂)
|
||||||
|
|
||||||
|
instance
|
||||||
|
isJoinSemilattice : IsSemilattice ExtendBelow _≈_ _⊔_
|
||||||
|
isJoinSemilattice = record
|
||||||
|
{ ≈-equiv = ≈-equiv
|
||||||
|
; ≈-⊔-cong = ≈-⊔-cong
|
||||||
|
; ⊔-assoc = ⊔-assoc
|
||||||
|
; ⊔-comm = ⊔-comm
|
||||||
|
; ⊔-idemp = ⊔-idemp
|
||||||
|
}
|
||||||
|
|
||||||
|
isMeetSemilattice : IsSemilattice ExtendBelow _≈_ _⊓_
|
||||||
|
isMeetSemilattice = record
|
||||||
|
{ ≈-equiv = ≈-equiv
|
||||||
|
; ≈-⊔-cong = ≈-⊓-cong
|
||||||
|
; ⊔-assoc = ⊓-assoc
|
||||||
|
; ⊔-comm = ⊓-comm
|
||||||
|
; ⊔-idemp = ⊓-idemp
|
||||||
|
}
|
||||||
|
|
||||||
|
isLattice : IsLattice ExtendBelow _≈_ _⊔_ _⊓_
|
||||||
|
isLattice = record
|
||||||
|
{ joinSemilattice = isJoinSemilattice
|
||||||
|
; meetSemilattice = isMeetSemilattice
|
||||||
|
; absorb-⊔-⊓ = absorb-⊔-⊓
|
||||||
|
; absorb-⊓-⊔ = absorb-⊓-⊔
|
||||||
|
}
|
||||||
|
|
||||||
|
module _ {{≈₁-Decidable : IsDecidable _≈₁_}} where
|
||||||
|
open IsDecidable ≈₁-Decidable using () renaming (R-dec to ≈₁-dec)
|
||||||
|
|
||||||
|
≈-dec : Decidable _≈_
|
||||||
|
≈-dec ⊥ ⊥ = yes ≈-⊥-⊥
|
||||||
|
≈-dec [ a₁ ] [ a₂ ]
|
||||||
|
with ≈₁-dec a₁ a₂
|
||||||
|
... | yes a₁≈a₂ = yes (≈-lift a₁≈a₂)
|
||||||
|
... | no a₁̷≈a₂ = no (λ { (≈-lift a₁≈a₂) → a₁̷≈a₂ a₁≈a₂ })
|
||||||
|
≈-dec ⊥ [ _ ] = no (λ ())
|
||||||
|
≈-dec [ _ ] ⊥ = no (λ ())
|
||||||
|
|
||||||
|
instance
|
||||||
|
≈-Decidable : IsDecidable _≈_
|
||||||
|
≈-Decidable = record { R-dec = ≈-dec }
|
||||||
@@ -9,10 +9,10 @@ module Lattice.FiniteMap (A : Set) (B : Set)
|
|||||||
{_≈₂_ : B → B → Set}
|
{_≈₂_ : B → B → Set}
|
||||||
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
||||||
{{≡-Decidable-A : IsDecidable {_} {A} _≡_}}
|
{{≡-Decidable-A : IsDecidable {_} {A} _≡_}}
|
||||||
{{lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_}} (dummy : ⊤) where
|
{{lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_}} (ks : List A) where
|
||||||
|
|
||||||
open IsLattice lB using () renaming (_≼_ to _≼₂_)
|
open IsLattice lB using () renaming (_≼_ to _≼₂_)
|
||||||
open import Lattice.Map A B dummy as Map
|
open import Lattice.Map A B _ as Map
|
||||||
using
|
using
|
||||||
( Map
|
( Map
|
||||||
; ⊔-equal-keys
|
; ⊔-equal-keys
|
||||||
@@ -74,7 +74,7 @@ open import Showable using (Showable; show)
|
|||||||
open import Isomorphism using (IsInverseˡ; IsInverseʳ)
|
open import Isomorphism using (IsInverseˡ; IsInverseʳ)
|
||||||
open import Chain using (Height)
|
open import Chain using (Height)
|
||||||
|
|
||||||
module WithKeys (ks : List A) where
|
private module WithKeys (ks : List A) where
|
||||||
FiniteMap : Set
|
FiniteMap : Set
|
||||||
FiniteMap = Σ Map (λ m → Map.keys m ≡ ks)
|
FiniteMap = Σ Map (λ m → Map.keys m ≡ ks)
|
||||||
|
|
||||||
@@ -143,6 +143,7 @@ module WithKeys (ks : List A) where
|
|||||||
λ {(m₁ , _)} {(m₂ , _)} {(m₃ , _)} →
|
λ {(m₁ , _)} {(m₂ , _)} {(m₃ , _)} →
|
||||||
IsEquivalence.≈-trans ≈ᵐ-equiv {m₁} {m₂} {m₃}
|
IsEquivalence.≈-trans ≈ᵐ-equiv {m₁} {m₂} {m₃}
|
||||||
}
|
}
|
||||||
|
open IsEquivalence ≈-equiv public
|
||||||
|
|
||||||
instance
|
instance
|
||||||
isUnionSemilattice : IsSemilattice FiniteMap _≈_ _⊔_
|
isUnionSemilattice : IsSemilattice FiniteMap _≈_ _⊔_
|
||||||
@@ -183,7 +184,7 @@ module WithKeys (ks : List A) where
|
|||||||
; isLattice = isLattice
|
; isLattice = isLattice
|
||||||
}
|
}
|
||||||
|
|
||||||
open IsLattice isLattice using (_≼_; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
|
open IsLattice isLattice using (_≼_; ⊔-idemp; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
|
||||||
|
|
||||||
m₁≼m₂⇒m₁[k]≼m₂[k] : ∀ (fm₁ fm₂ : FiniteMap) {k : A} {v₁ v₂ : B} →
|
m₁≼m₂⇒m₁[k]≼m₂[k] : ∀ (fm₁ fm₂ : FiniteMap) {k : A} {v₁ v₂ : B} →
|
||||||
fm₁ ≼ fm₂ → (k , v₁) ∈ fm₁ → (k , v₂) ∈ fm₂ → v₁ ≼₂ v₂
|
fm₁ ≼ fm₂ → (k , v₁) ∈ fm₁ → (k , v₂) ∈ fm₂ → v₁ ≼₂ v₂
|
||||||
@@ -253,9 +254,35 @@ module WithKeys (ks : List A) where
|
|||||||
... | yes k∈km₁ | no k∉km₂ = ⊥-elim (∈k-exclusive fm₁ fm₂ (k∈km₁ , k∉km₂))
|
... | yes k∈km₁ | no k∉km₂ = ⊥-elim (∈k-exclusive fm₁ fm₂ (k∈km₁ , k∉km₂))
|
||||||
... | no k∉km₁ | yes k∈km₂ = ⊥-elim (∈k-exclusive fm₂ fm₁ (k∈km₂ , k∉km₁))
|
... | no k∉km₁ | yes k∈km₂ = ⊥-elim (∈k-exclusive fm₂ fm₁ (k∈km₂ , k∉km₁))
|
||||||
|
|
||||||
open WithKeys public
|
private
|
||||||
|
_⊆ᵐ_ : ∀ {ks₁ ks₂ : List A} → WithKeys.FiniteMap ks₁ → WithKeys.FiniteMap ks₂ → Set
|
||||||
|
_⊆ᵐ_ fm₁ fm₂ = subset-impl (proj₁ (proj₁ fm₁)) (proj₁ (proj₁ fm₂))
|
||||||
|
|
||||||
module IterProdIsomorphism where
|
_∈ᵐ_ : ∀ {ks : List A} → A × B → WithKeys.FiniteMap ks → Set
|
||||||
|
_∈ᵐ_ {ks} = WithKeys._∈_ ks
|
||||||
|
|
||||||
|
FromBothMaps : ∀ (k : A) (v : B) {ks : List A} (fm₁ fm₂ : WithKeys.FiniteMap ks) → Set
|
||||||
|
FromBothMaps k v fm₁ fm₂ =
|
||||||
|
Σ (B × B)
|
||||||
|
(λ (v₁ , v₂) → ( (v ≡ v₁ ⊔₂ v₂) × ((k , v₁) ∈ᵐ fm₁ × (k , v₂) ∈ᵐ fm₂)))
|
||||||
|
|
||||||
|
Provenance-union : ∀ {ks : List A} (fm₁ fm₂ : WithKeys.FiniteMap ks) {k : A} {v : B} →
|
||||||
|
(k , v) ∈ᵐ (WithKeys._⊔_ ks fm₁ fm₂) → FromBothMaps k v fm₁ fm₂
|
||||||
|
Provenance-union fm₁@(m₁ , ks₁≡ks) fm₂@(m₂ , ks₂≡ks) {k} {v} k,v∈fm₁fm₂
|
||||||
|
with Expr-Provenance-≡ ((` m₁) ∪ (` m₂)) k,v∈fm₁fm₂
|
||||||
|
... | in₁ (single k,v∈m₁) k∉km₂
|
||||||
|
with k∈km₁ ← (WithKeys.forget k,v∈m₁)
|
||||||
|
rewrite trans ks₁≡ks (sym ks₂≡ks) =
|
||||||
|
⊥-elim (k∉km₂ k∈km₁)
|
||||||
|
... | in₂ k∉km₁ (single k,v∈m₂)
|
||||||
|
with k∈km₂ ← (WithKeys.forget k,v∈m₂)
|
||||||
|
rewrite trans ks₁≡ks (sym ks₂≡ks) =
|
||||||
|
⊥-elim (k∉km₁ k∈km₂)
|
||||||
|
... | bothᵘ {v₁} {v₂} (single k,v₁∈m₁) (single k,v₂∈m₂) =
|
||||||
|
((v₁ , v₂) , (refl , (k,v₁∈m₁ , k,v₂∈m₂)))
|
||||||
|
|
||||||
|
private module IterProdIsomorphism where
|
||||||
|
open WithKeys
|
||||||
open import Data.Unit using (tt)
|
open import Data.Unit using (tt)
|
||||||
open import Lattice.Unit using ()
|
open import Lattice.Unit using ()
|
||||||
renaming
|
renaming
|
||||||
@@ -267,7 +294,7 @@ module IterProdIsomorphism where
|
|||||||
; ≈-equiv to ≈ᵘ-equiv
|
; ≈-equiv to ≈ᵘ-equiv
|
||||||
; fixedHeight to fixedHeightᵘ
|
; fixedHeight to fixedHeightᵘ
|
||||||
)
|
)
|
||||||
open import Lattice.IterProd B ⊤ dummy
|
open import Lattice.IterProd B ⊤ _
|
||||||
as IP
|
as IP
|
||||||
using (IterProd)
|
using (IterProd)
|
||||||
open IsLattice lB using ()
|
open IsLattice lB using ()
|
||||||
@@ -296,20 +323,12 @@ module IterProdIsomorphism where
|
|||||||
in
|
in
|
||||||
(((k , v) ∷ fm' , push k≢fm' ufm') , kvs≡ks)
|
(((k , v) ∷ fm' , push k≢fm' ufm') , kvs≡ks)
|
||||||
|
|
||||||
|
_≈ⁱᵖ_ : ∀ {n : ℕ} → IterProd n → IterProd n → Set
|
||||||
|
_≈ⁱᵖ_ {n} = IP._≈_ {n}
|
||||||
|
|
||||||
private
|
_⊔ⁱᵖ_ : ∀ {ks : List A} →
|
||||||
_⊆ᵐ_ : ∀ {ks₁ ks₂ : List A} → FiniteMap ks₁ → FiniteMap ks₂ → Set
|
IterProd (length ks) → IterProd (length ks) → IterProd (length ks)
|
||||||
_⊆ᵐ_ fm₁ fm₂ = subset-impl (proj₁ (proj₁ fm₁)) (proj₁ (proj₁ fm₂))
|
_⊔ⁱᵖ_ {ks} = IP._⊔_ {length ks}
|
||||||
|
|
||||||
_≈ⁱᵖ_ : ∀ {n : ℕ} → IterProd n → IterProd n → Set
|
|
||||||
_≈ⁱᵖ_ {n} = IP._≈_ {n}
|
|
||||||
|
|
||||||
_⊔ⁱᵖ_ : ∀ {ks : List A} →
|
|
||||||
IterProd (length ks) → IterProd (length ks) → IterProd (length ks)
|
|
||||||
_⊔ⁱᵖ_ {ks} = IP._⊔_ {length ks}
|
|
||||||
|
|
||||||
_∈ᵐ_ : ∀ {ks : List A} → A × B → FiniteMap ks → Set
|
|
||||||
_∈ᵐ_ {ks} = _∈_ ks
|
|
||||||
|
|
||||||
to-build : ∀ {b : B} {ks : List A} (uks : Unique ks) →
|
to-build : ∀ {b : B} {ks : List A} (uks : Unique ks) →
|
||||||
let fm = to uks (IP.build b tt (length ks))
|
let fm = to uks (IP.build b tt (length ks))
|
||||||
@@ -368,26 +387,6 @@ module IterProdIsomorphism where
|
|||||||
fm'₂⊆fm'₁ k' v' k',v'∈fm'₂
|
fm'₂⊆fm'₁ k' v' k',v'∈fm'₂
|
||||||
in (v'' , (v'≈v'' , there k',v''∈fm'₁))
|
in (v'' , (v'≈v'' , there k',v''∈fm'₁))
|
||||||
|
|
||||||
FromBothMaps : ∀ (k : A) (v : B) {ks : List A} (fm₁ fm₂ : FiniteMap ks) → Set
|
|
||||||
FromBothMaps k v fm₁ fm₂ =
|
|
||||||
Σ (B × B)
|
|
||||||
(λ (v₁ , v₂) → ( (v ≡ v₁ ⊔₂ v₂) × ((k , v₁) ∈ᵐ fm₁ × (k , v₂) ∈ᵐ fm₂)))
|
|
||||||
|
|
||||||
Provenance-union : ∀ {ks : List A} (fm₁ fm₂ : FiniteMap ks) {k : A} {v : B} →
|
|
||||||
(k , v) ∈ᵐ (_⊔_ ks fm₁ fm₂) → FromBothMaps k v fm₁ fm₂
|
|
||||||
Provenance-union fm₁@(m₁ , ks₁≡ks) fm₂@(m₂ , ks₂≡ks) {k} {v} k,v∈fm₁fm₂
|
|
||||||
with Expr-Provenance-≡ ((` m₁) ∪ (` m₂)) k,v∈fm₁fm₂
|
|
||||||
... | in₁ (single k,v∈m₁) k∉km₂
|
|
||||||
with k∈km₁ ← (forget k,v∈m₁)
|
|
||||||
rewrite trans ks₁≡ks (sym ks₂≡ks) =
|
|
||||||
⊥-elim (k∉km₂ k∈km₁)
|
|
||||||
... | in₂ k∉km₁ (single k,v∈m₂)
|
|
||||||
with k∈km₂ ← (forget k,v∈m₂)
|
|
||||||
rewrite trans ks₁≡ks (sym ks₂≡ks) =
|
|
||||||
⊥-elim (k∉km₁ k∈km₂)
|
|
||||||
... | bothᵘ {v₁} {v₂} (single k,v₁∈m₁) (single k,v₂∈m₂) =
|
|
||||||
((v₁ , v₂) , (refl , (k,v₁∈m₁ , k,v₂∈m₂)))
|
|
||||||
|
|
||||||
private
|
private
|
||||||
first-key-in-map : ∀ {k : A} {ks : List A} (fm : FiniteMap (k ∷ ks)) →
|
first-key-in-map : ∀ {k : A} {ks : List A} (fm : FiniteMap (k ∷ ks)) →
|
||||||
Σ B (λ v → (k , v) ∈ᵐ fm)
|
Σ B (λ v → (k , v) ∈ᵐ fm)
|
||||||
@@ -613,7 +612,7 @@ module IterProdIsomorphism where
|
|||||||
in
|
in
|
||||||
(v' , (v₁⊔v₂≈v' , there v'∈fm'))
|
(v' , (v₁⊔v₂≈v' , there v'∈fm'))
|
||||||
|
|
||||||
module WithUniqueKeysAndFixedHeight {ks : List A} (uks : Unique ks) {{≈₂-Decidable : IsDecidable _≈₂_}} {h₂ : ℕ} {{fhB : FixedHeight₂ h₂}} where
|
module FixedHeight {ks : List A} {{≈₂-Decidable : IsDecidable _≈₂_}} {h₂ : ℕ} {{fhB : FixedHeight₂ h₂}} (uks : Unique ks) where
|
||||||
import Isomorphism
|
import Isomorphism
|
||||||
open Isomorphism.TransportFiniteHeight
|
open Isomorphism.TransportFiniteHeight
|
||||||
(IP.isFiniteHeightLattice {k = length ks} {{fhB = fixedHeightᵘ}}) (isLattice ks)
|
(IP.isFiniteHeightLattice {k = length ks} {{fhB = fixedHeightᵘ}}) (isLattice ks)
|
||||||
@@ -631,3 +630,6 @@ module IterProdIsomorphism where
|
|||||||
⊥-contains-bottoms {k} {v} k,v∈⊥
|
⊥-contains-bottoms {k} {v} k,v∈⊥
|
||||||
rewrite IP.⊥-built {length ks} {{fhB = fixedHeightᵘ}} =
|
rewrite IP.⊥-built {length ks} {{fhB = fixedHeightᵘ}} =
|
||||||
to-build uks k v k,v∈⊥
|
to-build uks k v k,v∈⊥
|
||||||
|
|
||||||
|
open WithKeys ks public
|
||||||
|
module FixedHeight = IterProdIsomorphism.FixedHeight
|
||||||
|
|||||||
@@ -106,6 +106,8 @@ instance
|
|||||||
; isLattice = isLattice
|
; isLattice = isLattice
|
||||||
}
|
}
|
||||||
|
|
||||||
|
open IsLattice isLattice using (_≼_; _≺_; ≺-cong) public
|
||||||
|
|
||||||
module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDecidable _≈₂_}} where
|
module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDecidable _≈₂_}} where
|
||||||
open IsDecidable ≈₁-Decidable using () renaming (R-dec to ≈₁-dec)
|
open IsDecidable ≈₁-Decidable using () renaming (R-dec to ≈₁-dec)
|
||||||
open IsDecidable ≈₂-Decidable using () renaming (R-dec to ≈₂-dec)
|
open IsDecidable ≈₂-Decidable using () renaming (R-dec to ≈₂-dec)
|
||||||
@@ -125,7 +127,6 @@ module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDeci
|
|||||||
{{fhA : FixedHeight₁ h₁}} {{fhB : FixedHeight₂ h₂}} where
|
{{fhA : FixedHeight₁ h₁}} {{fhB : FixedHeight₂ h₂}} where
|
||||||
|
|
||||||
open import Data.Nat.Properties
|
open import Data.Nat.Properties
|
||||||
open IsLattice isLattice using (_≼_; _≺_; ≺-cong)
|
|
||||||
|
|
||||||
module ChainMapping₁ = ChainMapping joinSemilattice₁ isJoinSemilattice
|
module ChainMapping₁ = ChainMapping joinSemilattice₁ isJoinSemilattice
|
||||||
module ChainMapping₂ = ChainMapping joinSemilattice₂ isJoinSemilattice
|
module ChainMapping₂ = ChainMapping joinSemilattice₂ isJoinSemilattice
|
||||||
|
|||||||
11
Main.agda
11
Main.agda
@@ -1,11 +1,13 @@
|
|||||||
{-# OPTIONS --guardedness #-}
|
{-# OPTIONS --guardedness #-}
|
||||||
module Main where
|
module Main where
|
||||||
|
|
||||||
open import Language
|
open import Language hiding (_++_)
|
||||||
open import Analysis.Sign
|
|
||||||
open import Data.Vec using (Vec; _∷_; [])
|
open import Data.Vec using (Vec; _∷_; [])
|
||||||
open import IO
|
open import IO
|
||||||
open import Level using (0ℓ)
|
open import Level using (0ℓ)
|
||||||
|
open import Data.String using (_++_)
|
||||||
|
import Analysis.Constant as ConstantAnalysis
|
||||||
|
import Analysis.Sign as SignAnalysis
|
||||||
|
|
||||||
testCode : Stmt
|
testCode : Stmt
|
||||||
testCode =
|
testCode =
|
||||||
@@ -38,6 +40,7 @@ testProgram = record
|
|||||||
{ rootStmt = testCode
|
{ rootStmt = testCode
|
||||||
}
|
}
|
||||||
|
|
||||||
open WithProg testProgram using (output; analyze-correct)
|
open SignAnalysis.WithProg testProgram using (analyze-correct) renaming (output to output-Sign)
|
||||||
|
open ConstantAnalysis.WithProg testProgram using (analyze-correct) renaming (output to output-Const)
|
||||||
|
|
||||||
main = run {0ℓ} (putStrLn output)
|
main = run {0ℓ} (putStrLn (output-Const ++ "\n" ++ output-Sign))
|
||||||
|
|||||||
Reference in New Issue
Block a user