Compare commits
148 Commits
f0da9a9020
...
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 | |||
| cf824dc744 | |||
| 70847d51db | |||
| d96eb97b69 | |||
| d90b544436 | |||
| b0488c9cc6 | |||
| 8abf6f8670 | |||
| 4da9b6d3cd | |||
| c2c04e3ecd | |||
| f01df5af4b | |||
| b28994e1d2 | |||
| 10332351ea | |||
| 9131214880 | |||
| 4fba1fe79a | |||
| 828b652d3b | |||
| 12971450e3 | |||
| 7d2928ed81 | |||
| 5f946de5e8 | |||
| 04bafb2d55 | |||
| e0248397b7 | |||
| 41ada43047 | |||
| a081edb881 | |||
| 3d2a507f2f | |||
| 82027ecd04 | |||
| 734e82ff6d | |||
| 69d1ecebae | |||
| b78cb91f2a | |||
| 16fa4cd1d8 | |||
| 95669b2c65 | |||
| 6857f60465 | |||
| f4392b32c0 | |||
| 794c04eee9 | |||
| 80069e76e6 | |||
| a22c0c9252 | |||
| 20dc99ba1f | |||
| b3a62da1fb | |||
| f0b0d51b48 | |||
| 8ff88f9f9e | |||
| c1581075d3 | |||
| 838aaf9c58 | |||
| 4ac9dffa9b | |||
| 3f5551d70c | |||
| 5837fdf19b | |||
| 4350919871 | |||
| 7fe46b014c | |||
| 66d229c493 | |||
| 1b8bea8957 | |||
| dd8cdcd10c | |||
| ad4592d4d2 | |||
| 8d0d87d2d9 | |||
| cfa3375de5 | |||
| 6b116ed960 | |||
| 3859826293 | |||
| be50c76cb1 | |||
| 112a5087ef | |||
| ccb7abc501 | |||
| 91b5d108f6 | |||
| c574ca9c56 | |||
| bbfba34e05 | |||
| aec15573fc | |||
| b4d395767d | |||
| 07550bc214 | |||
| 9366ec4a97 | |||
| 0fb884eb07 | |||
| 6b44ac1feb | |||
| 69a4e8eb5c | |||
| 4fee16413a | |||
| 316e56f2bc | |||
| ab40a27e78 | |||
| f555947184 | |||
| 660f6594fd | |||
| fb32315f58 | |||
| 037358308f | |||
| f2b8084a9c | |||
| c00c8e3e85 | |||
| b134c143ca | |||
| e218d1b7a3 | |||
| 6e3f06ca5d | |||
| 54b11d21b0 | |||
| f3e0d5f2e3 | |||
| 855bf3f56c | |||
| 2f91ca113e | |||
| 7571cb7451 | |||
| fc27b045d3 | |||
| de956cdc6a | |||
| 7ed7f20227 | |||
| 163108b9b3 | |||
| 8dc5c40eae | |||
| 44f04e4020 | |||
| 4fe0d147fa | |||
| ba1c9b3ec8 | |||
| b6e357787f | |||
| ce3fa182fe | |||
| 71cb97ad8c | |||
| 57606636a7 | |||
| da2f7f51d7 | |||
| 2db11dcfc7 | |||
| 3e2719d45f | |||
| 78252b6c9e | |||
| 85fdf544b9 | |||
| 4f14a7b765 | |||
| bc5b4b7d9e | |||
| 520b2b514c | |||
| f7ac22257e | |||
| b72ad070ba | |||
| 195537fe15 | |||
| d4b0796715 | |||
| b505063771 | |||
| 844c99336a | |||
| 5d56a7ce2d | |||
| 2e096bd64e | |||
| 1a7b2a1736 |
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
|
||||||
144
Analysis/Forward.agda
Normal file
144
Analysis/Forward.agda
Normal file
@@ -0,0 +1,144 @@
|
|||||||
|
open import Language hiding (_[_])
|
||||||
|
open import Lattice
|
||||||
|
|
||||||
|
module Analysis.Forward
|
||||||
|
(L : Set) {h}
|
||||||
|
{_≈ˡ_ : L → L → Set} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
||||||
|
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
|
||||||
|
{{≈ˡ-dec : IsDecidable _≈ˡ_}} where
|
||||||
|
|
||||||
|
open import Data.Empty using (⊥-elim)
|
||||||
|
open import Data.Unit using (⊤)
|
||||||
|
open import Data.String using (String)
|
||||||
|
open import Data.Product using (_,_)
|
||||||
|
open import Data.List using (_∷_; []; foldr; foldl)
|
||||||
|
open import Data.List.Relation.Unary.Any as Any using ()
|
||||||
|
open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; sym; subst)
|
||||||
|
open import Relation.Nullary using (yes; no)
|
||||||
|
open import Function using (_∘_; flip)
|
||||||
|
|
||||||
|
open IsFiniteHeightLattice isFiniteHeightLatticeˡ
|
||||||
|
using () renaming (isLattice to isLatticeˡ)
|
||||||
|
|
||||||
|
module WithProg (prog : Program) where
|
||||||
|
open import Analysis.Forward.Lattices L prog hiding (≈ᵛ-Decidable; ≈ᵐ-Decidable) -- to disambiguate instance resolution
|
||||||
|
open import Analysis.Forward.Evaluation L prog
|
||||||
|
open Program prog
|
||||||
|
|
||||||
|
private module WithStmtEvaluator {{evaluator : StmtEvaluator}} where
|
||||||
|
open StmtEvaluator evaluator
|
||||||
|
|
||||||
|
updateVariablesForState : State → StateVariables → VariableValues
|
||||||
|
updateVariablesForState s sv =
|
||||||
|
foldl (flip (eval s)) (variablesAt s sv) (code s)
|
||||||
|
|
||||||
|
updateVariablesForState-Monoʳ : ∀ (s : State) → Monotonic _≼ᵐ_ _≼ᵛ_ (updateVariablesForState s)
|
||||||
|
updateVariablesForState-Monoʳ s {sv₁} {sv₂} sv₁≼sv₂ =
|
||||||
|
let
|
||||||
|
bss = code s
|
||||||
|
(vs₁ , s,vs₁∈sv₁) = locateᵐ {s} {sv₁} (states-in-Map s sv₁)
|
||||||
|
(vs₂ , s,vs₂∈sv₂) = locateᵐ {s} {sv₂} (states-in-Map s sv₂)
|
||||||
|
vs₁≼vs₂ = m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ sv₁ sv₂ sv₁≼sv₂ s,vs₁∈sv₁ s,vs₂∈sv₂
|
||||||
|
in
|
||||||
|
foldl-Mono' (IsLattice.joinSemilattice isLatticeᵛ) bss
|
||||||
|
(flip (eval s)) (eval-Monoʳ s)
|
||||||
|
vs₁≼vs₂
|
||||||
|
|
||||||
|
open StateVariablesFiniteMap.GeneralizedUpdate {{isLatticeᵐ}} (λ x → x) (λ a₁≼a₂ → a₁≼a₂) updateVariablesForState updateVariablesForState-Monoʳ states
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( f' to updateAll
|
||||||
|
; f'-Monotonic to updateAll-Mono
|
||||||
|
; f'-k∈ks-≡ to updateAll-k∈ks-≡
|
||||||
|
)
|
||||||
|
public
|
||||||
|
|
||||||
|
-- Finally, the whole analysis consists of getting the 'join'
|
||||||
|
-- of all incoming states, then applying the per-state evaluation
|
||||||
|
-- function. This is just a composition, and is trivially monotonic.
|
||||||
|
|
||||||
|
analyze : StateVariables → StateVariables
|
||||||
|
analyze = updateAll ∘ joinAll
|
||||||
|
|
||||||
|
analyze-Mono : Monotonic _≼ᵐ_ _≼ᵐ_ analyze
|
||||||
|
analyze-Mono {sv₁} {sv₂} sv₁≼sv₂ =
|
||||||
|
updateAll-Mono {joinAll sv₁} {joinAll sv₂}
|
||||||
|
(joinAll-Mono {sv₁} {sv₂} sv₁≼sv₂)
|
||||||
|
|
||||||
|
-- The fixed point of the 'analyze' function is our final goal.
|
||||||
|
open import Fixedpoint analyze (λ {m₁} {m₂} m₁≼m₂ → analyze-Mono {m₁} {m₂} m₁≼m₂)
|
||||||
|
using ()
|
||||||
|
renaming (aᶠ to result; aᶠ≈faᶠ to result≈analyze-result)
|
||||||
|
public
|
||||||
|
|
||||||
|
variablesAt-updateAll : ∀ (s : State) (sv : StateVariables) →
|
||||||
|
variablesAt s (updateAll sv) ≡ updateVariablesForState s sv
|
||||||
|
variablesAt-updateAll s sv
|
||||||
|
with (vs , s,vs∈usv) ← locateᵐ {s} {updateAll sv} (states-in-Map s (updateAll sv)) =
|
||||||
|
updateAll-k∈ks-≡ {l = sv} (states-complete s) s,vs∈usv
|
||||||
|
|
||||||
|
module WithValidInterpretation {{latticeInterpretationˡ : LatticeInterpretation isLatticeˡ}}
|
||||||
|
{{validEvaluator : ValidStmtEvaluator evaluator latticeInterpretationˡ}} (dummy : ⊤) where
|
||||||
|
open ValidStmtEvaluator validEvaluator
|
||||||
|
|
||||||
|
eval-fold-valid : ∀ {s bss vs ρ₁ ρ₂} → ρ₁ , bss ⇒ᵇˢ ρ₂ → ⟦ vs ⟧ᵛ ρ₁ → ⟦ foldl (flip (eval s)) vs bss ⟧ᵛ ρ₂
|
||||||
|
eval-fold-valid {_} [] ⟦vs⟧ρ = ⟦vs⟧ρ
|
||||||
|
eval-fold-valid {s} {bs ∷ bss'} {vs} {ρ₁} {ρ₂} (ρ₁,bs⇒ρ ∷ ρ,bss'⇒ρ₂) ⟦vs⟧ρ₁ =
|
||||||
|
eval-fold-valid
|
||||||
|
{bss = bss'} {eval s bs vs} ρ,bss'⇒ρ₂
|
||||||
|
(valid ρ₁,bs⇒ρ ⟦vs⟧ρ₁)
|
||||||
|
|
||||||
|
updateVariablesForState-matches : ∀ {s sv ρ₁ ρ₂} → ρ₁ , (code s) ⇒ᵇˢ ρ₂ → ⟦ variablesAt s sv ⟧ᵛ ρ₁ → ⟦ updateVariablesForState s sv ⟧ᵛ ρ₂
|
||||||
|
updateVariablesForState-matches = eval-fold-valid
|
||||||
|
|
||||||
|
updateAll-matches : ∀ {s sv ρ₁ ρ₂} → ρ₁ , (code s) ⇒ᵇˢ ρ₂ → ⟦ variablesAt s sv ⟧ᵛ ρ₁ → ⟦ variablesAt s (updateAll sv) ⟧ᵛ ρ₂
|
||||||
|
updateAll-matches {s} {sv} ρ₁,bss⇒ρ₂ ⟦vs⟧ρ₁
|
||||||
|
rewrite variablesAt-updateAll s sv =
|
||||||
|
updateVariablesForState-matches {s} {sv} ρ₁,bss⇒ρ₂ ⟦vs⟧ρ₁
|
||||||
|
|
||||||
|
stepTrace : ∀ {s₁ ρ₁ ρ₂} → ⟦ joinForKey s₁ result ⟧ᵛ ρ₁ → ρ₁ , (code s₁) ⇒ᵇˢ ρ₂ → ⟦ variablesAt s₁ result ⟧ᵛ ρ₂
|
||||||
|
stepTrace {s₁} {ρ₁} {ρ₂} ⟦joinForKey-s₁⟧ρ₁ ρ₁,bss⇒ρ₂ =
|
||||||
|
let
|
||||||
|
-- I'd use rewrite, but Agda gets a memory overflow (?!).
|
||||||
|
⟦joinAll-result⟧ρ₁ =
|
||||||
|
subst (λ vs → ⟦ vs ⟧ᵛ ρ₁)
|
||||||
|
(sym (variablesAt-joinAll s₁ result))
|
||||||
|
⟦joinForKey-s₁⟧ρ₁
|
||||||
|
⟦analyze-result⟧ρ₂ =
|
||||||
|
updateAll-matches {sv = joinAll result}
|
||||||
|
ρ₁,bss⇒ρ₂ ⟦joinAll-result⟧ρ₁
|
||||||
|
analyze-result≈result =
|
||||||
|
≈ᵐ-sym {result} {updateAll (joinAll result)}
|
||||||
|
result≈analyze-result
|
||||||
|
analyze-s₁≈s₁ =
|
||||||
|
variablesAt-≈ s₁ (updateAll (joinAll result))
|
||||||
|
result (analyze-result≈result)
|
||||||
|
in
|
||||||
|
⟦⟧ᵛ-respects-≈ᵛ {variablesAt s₁ (updateAll (joinAll result))} {variablesAt s₁ result} (analyze-s₁≈s₁) ρ₂ ⟦analyze-result⟧ρ₂
|
||||||
|
|
||||||
|
walkTrace : ∀ {s₁ s₂ ρ₁ ρ₂} → ⟦ joinForKey s₁ result ⟧ᵛ ρ₁ → Trace {graph} s₁ s₂ ρ₁ ρ₂ → ⟦ variablesAt s₂ result ⟧ᵛ ρ₂
|
||||||
|
walkTrace {s₁} {s₁} {ρ₁} {ρ₂} ⟦joinForKey-s₁⟧ρ₁ (Trace-single ρ₁,bss⇒ρ₂) =
|
||||||
|
stepTrace {s₁} {ρ₁} {ρ₂} ⟦joinForKey-s₁⟧ρ₁ ρ₁,bss⇒ρ₂
|
||||||
|
walkTrace {s₁} {s₂} {ρ₁} {ρ₂} ⟦joinForKey-s₁⟧ρ₁ (Trace-edge {ρ₂ = ρ} {idx₂ = s} ρ₁,bss⇒ρ s₁→s₂ tr) =
|
||||||
|
let
|
||||||
|
⟦result-s₁⟧ρ =
|
||||||
|
stepTrace {s₁} {ρ₁} {ρ} ⟦joinForKey-s₁⟧ρ₁ ρ₁,bss⇒ρ
|
||||||
|
s₁∈incomingStates =
|
||||||
|
[]-∈ result (edge⇒incoming s₁→s₂)
|
||||||
|
(variablesAt-∈ s₁ result)
|
||||||
|
⟦joinForKey-s⟧ρ =
|
||||||
|
⟦⟧ᵛ-foldr ⟦result-s₁⟧ρ s₁∈incomingStates
|
||||||
|
in
|
||||||
|
walkTrace ⟦joinForKey-s⟧ρ tr
|
||||||
|
|
||||||
|
joinForKey-initialState-⊥ᵛ : joinForKey initialState result ≡ ⊥ᵛ
|
||||||
|
joinForKey-initialState-⊥ᵛ = cong (λ ins → foldr _⊔ᵛ_ ⊥ᵛ (result [ ins ])) initialState-pred-∅
|
||||||
|
|
||||||
|
⟦joinAll-initialState⟧ᵛ∅ : ⟦ joinForKey initialState result ⟧ᵛ []
|
||||||
|
⟦joinAll-initialState⟧ᵛ∅ = subst (λ vs → ⟦ vs ⟧ᵛ []) (sym joinForKey-initialState-⊥ᵛ) ⟦⊥ᵛ⟧ᵛ∅
|
||||||
|
|
||||||
|
analyze-correct : ∀ {ρ : Env} → [] , rootStmt ⇒ˢ ρ → ⟦ variablesAt finalState result ⟧ᵛ ρ
|
||||||
|
analyze-correct {ρ} ∅,s⇒ρ = walkTrace {initialState} {finalState} {[]} {ρ} ⟦joinAll-initialState⟧ᵛ∅ (trace ∅,s⇒ρ)
|
||||||
|
|
||||||
|
open WithStmtEvaluator using (result; analyze; result≈analyze-result) public
|
||||||
|
open WithStmtEvaluator.WithValidInterpretation using (analyze-correct) public
|
||||||
100
Analysis/Forward/Adapters.agda
Normal file
100
Analysis/Forward/Adapters.agda
Normal file
@@ -0,0 +1,100 @@
|
|||||||
|
open import Language hiding (_[_])
|
||||||
|
open import Lattice
|
||||||
|
|
||||||
|
module Analysis.Forward.Adapters
|
||||||
|
(L : Set) {h}
|
||||||
|
{_≈ˡ_ : L → L → Set} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
||||||
|
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
|
||||||
|
{{≈ˡ-dec : IsDecidable _≈ˡ_}}
|
||||||
|
(prog : Program) where
|
||||||
|
|
||||||
|
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 _≟ˢ_)
|
||||||
|
open import Data.Product using (_,_)
|
||||||
|
open import Data.List using (_∷_; []; foldr; foldl)
|
||||||
|
open import Data.List.Relation.Unary.Any as Any using ()
|
||||||
|
open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; sym; subst)
|
||||||
|
open import Relation.Nullary using (yes; no)
|
||||||
|
open import Function using (_∘_; flip)
|
||||||
|
|
||||||
|
open IsFiniteHeightLattice isFiniteHeightLatticeˡ
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( isLattice to isLatticeˡ
|
||||||
|
; _≼_ to _≼ˡ_
|
||||||
|
)
|
||||||
|
open Program prog
|
||||||
|
|
||||||
|
-- Now, allow StmtEvaluators to be auto-constructed from ExprEvaluators.
|
||||||
|
module ExprToStmtAdapter {{ exprEvaluator : ExprEvaluator }} where
|
||||||
|
open ExprEvaluator exprEvaluator
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( eval to evalᵉ
|
||||||
|
; eval-Monoʳ to evalᵉ-Monoʳ
|
||||||
|
)
|
||||||
|
|
||||||
|
-- For a particular evaluation function, we need to perform an evaluation
|
||||||
|
-- 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 ∷ [])
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( f' to updateVariablesFromExpression
|
||||||
|
; f'-Monotonic to updateVariablesFromExpression-Mono
|
||||||
|
; f'-k∈ks-≡ to updateVariablesFromExpression-k∈ks-≡
|
||||||
|
; f'-k∉ks-backward to updateVariablesFromExpression-k∉ks-backward
|
||||||
|
)
|
||||||
|
public
|
||||||
|
|
||||||
|
-- The per-state update function makes use of the single-key setter,
|
||||||
|
-- updateVariablesFromExpression, for the case where the statement
|
||||||
|
-- is an assignment.
|
||||||
|
--
|
||||||
|
-- This per-state function adjusts the variables in that state,
|
||||||
|
-- also monotonically; we derive the for-each-state update from
|
||||||
|
-- the Exercise 4.26 again.
|
||||||
|
|
||||||
|
evalᵇ : State → BasicStmt → VariableValues → VariableValues
|
||||||
|
evalᵇ _ (k ← e) vs = updateVariablesFromExpression k e vs
|
||||||
|
evalᵇ _ noop vs = vs
|
||||||
|
|
||||||
|
evalᵇ-Monoʳ : ∀ (s : State) (bs : BasicStmt) → Monotonic _≼ᵛ_ _≼ᵛ_ (evalᵇ s bs)
|
||||||
|
evalᵇ-Monoʳ _ (k ← e) {vs₁} {vs₂} vs₁≼vs₂ = updateVariablesFromExpression-Mono k e {vs₁} {vs₂} vs₁≼vs₂
|
||||||
|
evalᵇ-Monoʳ _ noop vs₁≼vs₂ = vs₁≼vs₂
|
||||||
|
|
||||||
|
instance
|
||||||
|
stmtEvaluator : StmtEvaluator
|
||||||
|
stmtEvaluator = record { eval = evalᵇ ; eval-Monoʳ = evalᵇ-Monoʳ }
|
||||||
|
|
||||||
|
-- Moreover, correct StmtEvaluators can be constructed from correct
|
||||||
|
-- ExprEvaluators.
|
||||||
|
module _ {{latticeInterpretationˡ : LatticeInterpretation isLatticeˡ}}
|
||||||
|
{{isValid : ValidExprEvaluator exprEvaluator latticeInterpretationˡ}} where
|
||||||
|
open ValidExprEvaluator isValid using () renaming (valid to validᵉ)
|
||||||
|
|
||||||
|
evalᵇ-valid : ∀ {s vs ρ₁ ρ₂ bs} → ρ₁ , bs ⇒ᵇ ρ₂ → ⟦ vs ⟧ᵛ ρ₁ → ⟦ evalᵇ s bs vs ⟧ᵛ ρ₂
|
||||||
|
evalᵇ-valid {_} {vs} {ρ₁} {ρ₁} {_} (⇒ᵇ-noop ρ₁) ⟦vs⟧ρ₁ = ⟦vs⟧ρ₁
|
||||||
|
evalᵇ-valid {_} {vs} {ρ₁} {_} {_} (⇒ᵇ-← ρ₁ k e v ρ,e⇒v) ⟦vs⟧ρ₁ {k'} {l} k',l∈vs' {v'} k',v'∈ρ₂
|
||||||
|
with k ≟ˢ k' | k',v'∈ρ₂
|
||||||
|
... | yes refl | here _ v _
|
||||||
|
rewrite updateVariablesFromExpression-k∈ks-≡ k e {l = vs} (Any.here refl) k',l∈vs' =
|
||||||
|
validᵉ ρ,e⇒v ⟦vs⟧ρ₁
|
||||||
|
... | yes k≡k' | there _ _ _ _ _ k'≢k _ = ⊥-elim (k'≢k (sym k≡k'))
|
||||||
|
... | no k≢k' | here _ _ _ = ⊥-elim (k≢k' refl)
|
||||||
|
... | no k≢k' | there _ _ _ _ _ _ k',v'∈ρ₁ =
|
||||||
|
let
|
||||||
|
k'∉[k] = (λ { (Any.here refl) → k≢k' refl })
|
||||||
|
k',l∈vs = updateVariablesFromExpression-k∉ks-backward k e {l = vs} k'∉[k] k',l∈vs'
|
||||||
|
in
|
||||||
|
⟦vs⟧ρ₁ k',l∈vs k',v'∈ρ₁
|
||||||
|
|
||||||
|
instance
|
||||||
|
validStmtEvaluator : ValidStmtEvaluator stmtEvaluator latticeInterpretationˡ
|
||||||
|
validStmtEvaluator = record
|
||||||
|
{ valid = λ {a} {b} {c} {d} → evalᵇ-valid {a} {b} {c} {d}
|
||||||
|
}
|
||||||
66
Analysis/Forward/Evaluation.agda
Normal file
66
Analysis/Forward/Evaluation.agda
Normal file
@@ -0,0 +1,66 @@
|
|||||||
|
open import Language hiding (_[_])
|
||||||
|
open import Lattice
|
||||||
|
|
||||||
|
module Analysis.Forward.Evaluation
|
||||||
|
(L : Set) {h}
|
||||||
|
{_≈ˡ_ : L → L → Set} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
||||||
|
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
|
||||||
|
{{≈ˡ-dec : IsDecidable _≈ˡ_}}
|
||||||
|
(prog : Program) where
|
||||||
|
|
||||||
|
open import Analysis.Forward.Lattices L prog
|
||||||
|
open import Data.Product using (_,_)
|
||||||
|
|
||||||
|
open IsFiniteHeightLattice isFiniteHeightLatticeˡ
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( isLattice to isLatticeˡ
|
||||||
|
; _≼_ to _≼ˡ_
|
||||||
|
)
|
||||||
|
open Program prog
|
||||||
|
|
||||||
|
-- The "full" version of the analysis ought to define a function
|
||||||
|
-- that analyzes each basic statement. For some analyses, the state ID
|
||||||
|
-- is used as part of the lattice, so include it here.
|
||||||
|
record StmtEvaluator : Set where
|
||||||
|
field
|
||||||
|
eval : State → BasicStmt → VariableValues → VariableValues
|
||||||
|
eval-Monoʳ : ∀ (s : State) (bs : BasicStmt) → Monotonic _≼ᵛ_ _≼ᵛ_ (eval s bs)
|
||||||
|
|
||||||
|
-- For some "simpler" analyes, all we need to do is analyze the expressions.
|
||||||
|
-- For that purpose, provide a simpler evaluator type.
|
||||||
|
record ExprEvaluator : Set where
|
||||||
|
field
|
||||||
|
eval : Expr → VariableValues → L
|
||||||
|
eval-Monoʳ : ∀ (e : Expr) → Monotonic _≼ᵛ_ _≼ˡ_ (eval e)
|
||||||
|
|
||||||
|
-- Evaluators have a notion of being "valid", in which the (symbolic)
|
||||||
|
-- manipulations on lattice elements they perform match up with
|
||||||
|
-- the semantics. Define what it means to be valid for statement and
|
||||||
|
-- expression-based evaluators. Define "IsValidExprEvaluator"
|
||||||
|
-- and "IsValidStmtEvaluator" standalone so that users can use them
|
||||||
|
-- in their type expressions.
|
||||||
|
|
||||||
|
module _ {{evaluator : ExprEvaluator}} {{interpretation : LatticeInterpretation isLatticeˡ}} where
|
||||||
|
open ExprEvaluator evaluator
|
||||||
|
open LatticeInterpretation interpretation
|
||||||
|
|
||||||
|
IsValidExprEvaluator : Set
|
||||||
|
IsValidExprEvaluator = ∀ {vs ρ e v} → ρ , e ⇒ᵉ v → ⟦ vs ⟧ᵛ ρ → ⟦ eval e vs ⟧ˡ v
|
||||||
|
|
||||||
|
record ValidExprEvaluator (evaluator : ExprEvaluator)
|
||||||
|
(interpretation : LatticeInterpretation isLatticeˡ) : Set₁ where
|
||||||
|
field
|
||||||
|
valid : IsValidExprEvaluator {{evaluator}} {{interpretation}}
|
||||||
|
|
||||||
|
module _ {{evaluator : StmtEvaluator}} {{interpretation : LatticeInterpretation isLatticeˡ}} where
|
||||||
|
open StmtEvaluator evaluator
|
||||||
|
open LatticeInterpretation interpretation
|
||||||
|
|
||||||
|
IsValidStmtEvaluator : Set
|
||||||
|
IsValidStmtEvaluator = ∀ {s vs ρ₁ ρ₂ bs} → ρ₁ , bs ⇒ᵇ ρ₂ → ⟦ vs ⟧ᵛ ρ₁ → ⟦ eval s bs vs ⟧ᵛ ρ₂
|
||||||
|
|
||||||
|
record ValidStmtEvaluator (evaluator : StmtEvaluator)
|
||||||
|
(interpretation : LatticeInterpretation isLatticeˡ) : Set₁ where
|
||||||
|
field
|
||||||
|
valid : IsValidStmtEvaluator {{evaluator}} {{interpretation}}
|
||||||
195
Analysis/Forward/Lattices.agda
Normal file
195
Analysis/Forward/Lattices.agda
Normal file
@@ -0,0 +1,195 @@
|
|||||||
|
open import Language hiding (_[_])
|
||||||
|
open import Lattice
|
||||||
|
|
||||||
|
module Analysis.Forward.Lattices
|
||||||
|
(L : Set) {h}
|
||||||
|
{_≈ˡ_ : L → L → Set} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
||||||
|
{{isFiniteHeightLatticeˡ : IsFiniteHeightLattice L h _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
|
||||||
|
{{≈ˡ-Decidable : IsDecidable _≈ˡ_}}
|
||||||
|
(prog : Program) where
|
||||||
|
|
||||||
|
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
|
||||||
|
open import Data.Product using (proj₁; proj₂; _,_)
|
||||||
|
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; _⇒_; _∨_; it)
|
||||||
|
|
||||||
|
open IsFiniteHeightLattice isFiniteHeightLatticeˡ
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( isLattice to isLatticeˡ
|
||||||
|
; fixedHeight to fixedHeightˡ
|
||||||
|
; ≈-sym to ≈ˡ-sym
|
||||||
|
)
|
||||||
|
open Program prog
|
||||||
|
|
||||||
|
import Lattice.FiniteMap
|
||||||
|
import Chain
|
||||||
|
|
||||||
|
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 String L vars
|
||||||
|
open VariableValuesFiniteMap
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( FiniteMap to VariableValues
|
||||||
|
; isLattice to isLatticeᵛ
|
||||||
|
; _≈_ to _≈ᵛ_
|
||||||
|
; _⊔_ to _⊔ᵛ_
|
||||||
|
; _≼_ to _≼ᵛ_
|
||||||
|
; ≈-Decidable to ≈ᵛ-Decidable
|
||||||
|
; _∈_ to _∈ᵛ_
|
||||||
|
; _∈k_ to _∈kᵛ_
|
||||||
|
; _updating_via_ to _updatingᵛ_via_
|
||||||
|
; locate to locateᵛ
|
||||||
|
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵛ≼m₂[k]ᵛ
|
||||||
|
; ∈k-dec to ∈k-decᵛ
|
||||||
|
; all-equal-keys to all-equal-keysᵛ
|
||||||
|
; Provenance-union to Provenance-unionᵛ
|
||||||
|
; ⊔-Monotonicˡ to ⊔ᵛ-Monotonicˡ
|
||||||
|
; ⊔-Monotonicʳ to ⊔ᵛ-Monotonicʳ
|
||||||
|
; ⊔-idemp to ⊔ᵛ-idemp
|
||||||
|
)
|
||||||
|
public
|
||||||
|
open VariableValuesFiniteMap.FixedHeight vars-Unique
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( isFiniteHeightLattice to isFiniteHeightLatticeᵛ
|
||||||
|
; fixedHeight to fixedHeightᵛ
|
||||||
|
; ⊥-contains-bottoms to ⊥ᵛ-contains-bottoms
|
||||||
|
)
|
||||||
|
public
|
||||||
|
|
||||||
|
⊥ᵛ = Chain.Height.⊥ fixedHeightᵛ
|
||||||
|
|
||||||
|
-- Finally, the map we care about is (state -> (variables -> value)). Bring that in.
|
||||||
|
module StateVariablesFiniteMap = Lattice.FiniteMap State VariableValues states
|
||||||
|
open StateVariablesFiniteMap
|
||||||
|
using (_[_]; []-∈; m₁≼m₂⇒m₁[ks]≼m₂[ks]; m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂)
|
||||||
|
renaming
|
||||||
|
( FiniteMap to StateVariables
|
||||||
|
; isLattice to isLatticeᵐ
|
||||||
|
; _≈_ to _≈ᵐ_
|
||||||
|
; _∈_ to _∈ᵐ_
|
||||||
|
; _∈k_ to _∈kᵐ_
|
||||||
|
; locate to locateᵐ
|
||||||
|
; _≼_ to _≼ᵐ_
|
||||||
|
; ≈-Decidable to ≈ᵐ-Decidable
|
||||||
|
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ
|
||||||
|
; ≈-sym to ≈ᵐ-sym
|
||||||
|
)
|
||||||
|
public
|
||||||
|
open StateVariablesFiniteMap.FixedHeight states-Unique
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( isFiniteHeightLattice to isFiniteHeightLatticeᵐ
|
||||||
|
)
|
||||||
|
public
|
||||||
|
|
||||||
|
-- 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
|
||||||
|
-- retrieve the variable values at each state.
|
||||||
|
|
||||||
|
states-in-Map : ∀ (s : State) (sv : StateVariables) → s ∈kᵐ sv
|
||||||
|
states-in-Map s sv@(m , ksv≡states) rewrite ksv≡states = states-complete s
|
||||||
|
|
||||||
|
variablesAt : State → StateVariables → VariableValues
|
||||||
|
variablesAt s sv = proj₁ (locateᵐ {s} {sv} (states-in-Map s sv))
|
||||||
|
|
||||||
|
variablesAt-∈ : ∀ (s : State) (sv : StateVariables) → (s , variablesAt s sv) ∈ᵐ sv
|
||||||
|
variablesAt-∈ s sv = proj₂ (locateᵐ {s} {sv} (states-in-Map s sv))
|
||||||
|
|
||||||
|
variablesAt-≈ : ∀ s sv₁ sv₂ → sv₁ ≈ᵐ sv₂ → variablesAt s sv₁ ≈ᵛ variablesAt s sv₂
|
||||||
|
variablesAt-≈ s sv₁ sv₂ sv₁≈sv₂ =
|
||||||
|
m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ sv₁ sv₂ sv₁≈sv₂
|
||||||
|
(states-in-Map s sv₁) (states-in-Map s sv₂)
|
||||||
|
|
||||||
|
-- build up the 'join' function, which follows from Exercise 4.26's
|
||||||
|
--
|
||||||
|
-- L₁ → (A → L₂)
|
||||||
|
--
|
||||||
|
-- Construction, with L₁ = (A → L₂), and f = id
|
||||||
|
|
||||||
|
joinForKey : State → StateVariables → VariableValues
|
||||||
|
joinForKey k states = foldr _⊔ᵛ_ ⊥ᵛ (states [ incoming k ])
|
||||||
|
|
||||||
|
-- The per-key join is made up of map key accesses (which are monotonic)
|
||||||
|
-- and folds using the join operation (also monotonic)
|
||||||
|
|
||||||
|
joinForKey-Mono : ∀ (k : State) → Monotonic _≼ᵐ_ _≼ᵛ_ (joinForKey k)
|
||||||
|
joinForKey-Mono k {fm₁} {fm₂} fm₁≼fm₂ =
|
||||||
|
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
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( f' to joinAll
|
||||||
|
; f'-Monotonic to joinAll-Mono
|
||||||
|
; f'-k∈ks-≡ to joinAll-k∈ks-≡
|
||||||
|
)
|
||||||
|
public
|
||||||
|
|
||||||
|
variablesAt-joinAll : ∀ (s : State) (sv : StateVariables) →
|
||||||
|
variablesAt s (joinAll sv) ≡ joinForKey s sv
|
||||||
|
variablesAt-joinAll s sv
|
||||||
|
with (vs , s,vs∈usv) ← locateᵐ {s} {joinAll sv} (states-in-Map s (joinAll sv)) =
|
||||||
|
joinAll-k∈ks-≡ {l = sv} (states-complete s) s,vs∈usv
|
||||||
|
|
||||||
|
-- Elements of the lattice type L describe individual variables. What
|
||||||
|
-- exactly each lattice element says about the variable is defined
|
||||||
|
-- by a LatticeInterpretation element. We've now constructed the
|
||||||
|
-- (Variable → L) lattice, which describes states, and we need to lift
|
||||||
|
-- the "meaning" of the element lattice to a descriptions of states.
|
||||||
|
module _ {{latticeInterpretationˡ : LatticeInterpretation isLatticeˡ}} where
|
||||||
|
open LatticeInterpretation latticeInterpretationˡ
|
||||||
|
using ()
|
||||||
|
renaming
|
||||||
|
( ⟦_⟧ to ⟦_⟧ˡ
|
||||||
|
; ⟦⟧-respects-≈ to ⟦⟧ˡ-respects-≈ˡ
|
||||||
|
; ⟦⟧-⊔-∨ to ⟦⟧ˡ-⊔ˡ-∨
|
||||||
|
)
|
||||||
|
public
|
||||||
|
|
||||||
|
⟦_⟧ᵛ : VariableValues → Env → Set
|
||||||
|
⟦_⟧ᵛ vs ρ = ∀ {k l} → (k , l) ∈ᵛ vs → ∀ {v} → (k , v) Language.∈ ρ → ⟦ l ⟧ˡ v
|
||||||
|
|
||||||
|
⟦⊥ᵛ⟧ᵛ∅ : ⟦ ⊥ᵛ ⟧ᵛ []
|
||||||
|
⟦⊥ᵛ⟧ᵛ∅ _ ()
|
||||||
|
|
||||||
|
⟦⟧ᵛ-respects-≈ᵛ : ∀ {vs₁ vs₂ : VariableValues} → vs₁ ≈ᵛ vs₂ → ⟦ vs₁ ⟧ᵛ ⇒ ⟦ vs₂ ⟧ᵛ
|
||||||
|
⟦⟧ᵛ-respects-≈ᵛ {m₁ , _} {m₂ , _}
|
||||||
|
(m₁⊆m₂ , m₂⊆m₁) ρ ⟦vs₁⟧ρ {k} {l} k,l∈m₂ {v} k,v∈ρ =
|
||||||
|
let
|
||||||
|
(l' , (l≈l' , k,l'∈m₁)) = m₂⊆m₁ _ _ k,l∈m₂
|
||||||
|
⟦l'⟧v = ⟦vs₁⟧ρ k,l'∈m₁ k,v∈ρ
|
||||||
|
in
|
||||||
|
⟦⟧ˡ-respects-≈ˡ (≈ˡ-sym l≈l') v ⟦l'⟧v
|
||||||
|
|
||||||
|
⟦⟧ᵛ-⊔ᵛ-∨ : ∀ {vs₁ vs₂ : VariableValues} → (⟦ vs₁ ⟧ᵛ ∨ ⟦ vs₂ ⟧ᵛ) ⇒ ⟦ vs₁ ⊔ᵛ vs₂ ⟧ᵛ
|
||||||
|
⟦⟧ᵛ-⊔ᵛ-∨ {vs₁} {vs₂} ρ ⟦vs₁⟧ρ∨⟦vs₂⟧ρ {k} {l} k,l∈vs₁₂ {v} k,v∈ρ
|
||||||
|
with ((l₁ , l₂) , (refl , (k,l₁∈vs₁ , k,l₂∈vs₂)))
|
||||||
|
← Provenance-unionᵛ vs₁ vs₂ k,l∈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∈ρ))
|
||||||
|
|
||||||
|
⟦⟧ᵛ-foldr : ∀ {vs : VariableValues} {vss : List VariableValues} {ρ : Env} →
|
||||||
|
⟦ vs ⟧ᵛ ρ → vs ∈ˡ vss → ⟦ foldr _⊔ᵛ_ ⊥ᵛ vss ⟧ᵛ ρ
|
||||||
|
⟦⟧ᵛ-foldr {vs} {vs ∷ vss'} {ρ = ρ} ⟦vs⟧ρ (Any.here refl) =
|
||||||
|
⟦⟧ᵛ-⊔ᵛ-∨ {vs₁ = vs} {vs₂ = foldr _⊔ᵛ_ ⊥ᵛ vss'} ρ (inj₁ ⟦vs⟧ρ)
|
||||||
|
⟦⟧ᵛ-foldr {vs} {vs' ∷ vss'} {ρ = ρ} ⟦vs⟧ρ (Any.there vs∈vss') =
|
||||||
|
⟦⟧ᵛ-⊔ᵛ-∨ {vs₁ = vs'} {vs₂ = foldr _⊔ᵛ_ ⊥ᵛ vss'} ρ
|
||||||
|
(inj₂ (⟦⟧ᵛ-foldr ⟦vs⟧ρ vs∈vss'))
|
||||||
@@ -1,20 +1,23 @@
|
|||||||
module Analysis.Sign where
|
module Analysis.Sign where
|
||||||
|
|
||||||
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
|
open import Data.Integer as Int using (ℤ; +_; -[1+_])
|
||||||
open import Data.Nat using (suc)
|
open import Data.Nat as Nat using (ℕ; suc; zero)
|
||||||
open import Data.Product using (_×_; proj₁; _,_)
|
open import Data.Product using (Σ; proj₁; proj₂; _,_)
|
||||||
open import Data.List using (List; _∷_; []; foldr; cartesianProduct; cartesianProductWith)
|
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 Data.List.Membership.Propositional as MemProp using () renaming (_∈_ to _∈ˡ_)
|
||||||
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; subst)
|
open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; subst)
|
||||||
open import Relation.Nullary using (¬_; Dec; yes; no)
|
open import Relation.Nullary using (¬_; yes; no)
|
||||||
open import Data.Unit using (⊤)
|
|
||||||
open import Function using (_∘_)
|
|
||||||
|
|
||||||
open import Language
|
open import Language
|
||||||
open import Lattice
|
open import Lattice
|
||||||
open import Utils using (Pairwise)
|
open import Equivalence
|
||||||
open import Showable using (Showable; show)
|
open import Showable using (Showable; show)
|
||||||
import Lattice.FiniteValueMap
|
open import Utils using (_⇒_; _∧_; _∨_)
|
||||||
|
open import Analysis.Utils using (eval-combine₂)
|
||||||
|
import Analysis.Forward
|
||||||
|
|
||||||
data Sign : Set where
|
data Sign : Set where
|
||||||
+ : Sign
|
+ : Sign
|
||||||
@@ -32,7 +35,7 @@ instance
|
|||||||
}
|
}
|
||||||
|
|
||||||
-- g for siGn; s is used for strings and i is not very descriptive.
|
-- g for siGn; s is used for strings and i is not very descriptive.
|
||||||
_≟ᵍ_ : IsDecidable (_≡_ {_} {Sign})
|
_≟ᵍ_ : Decidable (_≡_ {_} {Sign})
|
||||||
_≟ᵍ_ + + = yes refl
|
_≟ᵍ_ + + = yes refl
|
||||||
_≟ᵍ_ + - = no (λ ())
|
_≟ᵍ_ + - = no (λ ())
|
||||||
_≟ᵍ_ + 0ˢ = no (λ ())
|
_≟ᵍ_ + 0ˢ = no (λ ())
|
||||||
@@ -43,12 +46,22 @@ _≟ᵍ_ 0ˢ + = no (λ ())
|
|||||||
_≟ᵍ_ 0ˢ - = no (λ ())
|
_≟ᵍ_ 0ˢ - = no (λ ())
|
||||||
_≟ᵍ_ 0ˢ 0ˢ = yes refl
|
_≟ᵍ_ 0ˢ 0ˢ = yes refl
|
||||||
|
|
||||||
|
instance
|
||||||
|
≡-equiv : IsEquivalence Sign _≡_
|
||||||
|
≡-equiv = record
|
||||||
|
{ ≈-refl = refl
|
||||||
|
; ≈-sym = sym
|
||||||
|
; ≈-trans = trans
|
||||||
|
}
|
||||||
|
|
||||||
|
≡-Decidable-Sign : IsDecidable {_} {Sign} _≡_
|
||||||
|
≡-Decidable-Sign = record { R-dec = _≟ᵍ_ }
|
||||||
|
|
||||||
-- embelish 'sign' with a top and bottom element.
|
-- embelish 'sign' with a top and bottom element.
|
||||||
open import Lattice.AboveBelow Sign _≡_ (record { ≈-refl = refl; ≈-sym = sym; ≈-trans = trans }) _≟ᵍ_ as AB
|
open import Lattice.AboveBelow Sign _ as AB
|
||||||
using ()
|
using ()
|
||||||
renaming
|
renaming
|
||||||
( AboveBelow to SignLattice
|
( AboveBelow to SignLattice
|
||||||
; ≈-dec to ≈ᵍ-dec
|
|
||||||
; ⊥ to ⊥ᵍ
|
; ⊥ to ⊥ᵍ
|
||||||
; ⊤ to ⊤ᵍ
|
; ⊤ to ⊤ᵍ
|
||||||
; [_] to [_]ᵍ
|
; [_] to [_]ᵍ
|
||||||
@@ -61,16 +74,12 @@ open import Lattice.AboveBelow Sign _≡_ (record { ≈-refl = refl; ≈-sym = s
|
|||||||
-- 'sign' has no underlying lattice structure, so use the 'plain' above-below lattice.
|
-- 'sign' has no underlying lattice structure, so use the 'plain' above-below lattice.
|
||||||
open AB.Plain 0ˢ using ()
|
open AB.Plain 0ˢ using ()
|
||||||
renaming
|
renaming
|
||||||
( finiteHeightLattice to finiteHeightLatticeᵍ
|
( isLattice to isLatticeᵍ
|
||||||
; isLattice to isLatticeᵍ
|
; isFiniteHeightLattice to isFiniteHeightLatticeᵍ
|
||||||
; fixedHeight to fixedHeightᵍ
|
|
||||||
; _≼_ to _≼ᵍ_
|
; _≼_ to _≼ᵍ_
|
||||||
; _⊔_ to _⊔ᵍ_
|
; _⊔_ to _⊔ᵍ_
|
||||||
)
|
; _⊓_ to _⊓ᵍ_
|
||||||
|
; ≼-trans to ≼ᵍ-trans
|
||||||
open IsLattice isLatticeᵍ using ()
|
|
||||||
renaming
|
|
||||||
( ≼-trans to ≼ᵍ-trans
|
|
||||||
)
|
)
|
||||||
|
|
||||||
plus : SignLattice → SignLattice → SignLattice
|
plus : SignLattice → SignLattice → SignLattice
|
||||||
@@ -93,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 _ ⊥ᵍ = ⊥ᵍ
|
||||||
@@ -111,199 +123,177 @@ 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
|
||||||
|
⟦_⟧ᵍ ⊥ᵍ _ = ⊥
|
||||||
|
⟦_⟧ᵍ ⊤ᵍ _ = ⊤
|
||||||
|
⟦_⟧ᵍ [ + ]ᵍ v = Σ ℕ (λ n → v ≡ ↑ᶻ (+_ (suc n)))
|
||||||
|
⟦_⟧ᵍ [ 0ˢ ]ᵍ v = v ≡ ↑ᶻ (+_ zero)
|
||||||
|
⟦_⟧ᵍ [ - ]ᵍ v = Σ ℕ (λ n → v ≡ ↑ᶻ -[1+ n ])
|
||||||
|
|
||||||
|
⟦⟧ᵍ-respects-≈ᵍ : ∀ {s₁ s₂ : SignLattice} → s₁ ≈ᵍ s₂ → ⟦ s₁ ⟧ᵍ ⇒ ⟦ s₂ ⟧ᵍ
|
||||||
|
⟦⟧ᵍ-respects-≈ᵍ ≈ᵍ-⊥ᵍ-⊥ᵍ v bot = bot
|
||||||
|
⟦⟧ᵍ-respects-≈ᵍ ≈ᵍ-⊤ᵍ-⊤ᵍ v top = top
|
||||||
|
⟦⟧ᵍ-respects-≈ᵍ (≈ᵍ-lift { + } { + } refl) v proof = proof
|
||||||
|
⟦⟧ᵍ-respects-≈ᵍ (≈ᵍ-lift { - } { - } refl) v proof = proof
|
||||||
|
⟦⟧ᵍ-respects-≈ᵍ (≈ᵍ-lift { 0ˢ } { 0ˢ } refl) v proof = proof
|
||||||
|
|
||||||
|
⟦⟧ᵍ-⊔ᵍ-∨ : ∀ {s₁ s₂ : SignLattice} → (⟦ 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₂ : ∀ {s₁ s₂ : Sign} → ¬ s₁ ≡ s₂ → ∀ {v} → ¬ ((⟦ [ s₁ ]ᵍ ⟧ᵍ ∧ ⟦ [ s₂ ]ᵍ ⟧ᵍ) v)
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { + } { + } +≢+ _ = ⊥-elim (+≢+ refl)
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { + } { - } _ ((n , refl) , (m , ()))
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { + } { 0ˢ } _ ((n , refl) , ())
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { 0ˢ } { + } _ (refl , (m , ()))
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { 0ˢ } { 0ˢ } +≢+ _ = ⊥-elim (+≢+ refl)
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { 0ˢ } { - } _ (refl , (m , ()))
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { - } { + } _ ((n , refl) , (m , ()))
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { - } { 0ˢ } _ ((n , refl) , ())
|
||||||
|
s₁≢s₂⇒¬s₁∧s₂ { - } { - } +≢+ _ = ⊥-elim (+≢+ refl)
|
||||||
|
|
||||||
|
⟦⟧ᵍ-⊓ᵍ-∧ : ∀ {s₁ s₂ : SignLattice} → (⟦ 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
|
module WithProg (prog : Program) where
|
||||||
open Program prog
|
open Program prog
|
||||||
|
|
||||||
-- The variable -> sign map is a finite value-map with keys strings. Use a bundle to avoid explicitly specifying operators.
|
open import Analysis.Forward.Lattices SignLattice prog
|
||||||
module VariableSignsFiniteMap = Lattice.FiniteValueMap.WithKeys _≟ˢ_ isLatticeᵍ vars
|
open import Analysis.Forward.Evaluation SignLattice prog
|
||||||
open VariableSignsFiniteMap
|
open import Analysis.Forward.Adapters SignLattice prog
|
||||||
using ()
|
|
||||||
renaming
|
|
||||||
( FiniteMap to VariableSigns
|
|
||||||
; isLattice to isLatticeᵛ
|
|
||||||
; _≈_ to _≈ᵛ_
|
|
||||||
; _⊔_ to _⊔ᵛ_
|
|
||||||
; _≼_ to _≼ᵛ_
|
|
||||||
; ≈₂-dec⇒≈-dec to ≈ᵍ-dec⇒≈ᵛ-dec
|
|
||||||
; _∈_ to _∈ᵛ_
|
|
||||||
; _∈k_ to _∈kᵛ_
|
|
||||||
; _updating_via_ to _updatingᵛ_via_
|
|
||||||
; locate to locateᵛ
|
|
||||||
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵛ≼m₂[k]ᵛ
|
|
||||||
)
|
|
||||||
open IsLattice isLatticeᵛ
|
|
||||||
using ()
|
|
||||||
renaming
|
|
||||||
( ⊔-Monotonicˡ to ⊔ᵛ-Monotonicˡ
|
|
||||||
; ⊔-Monotonicʳ to ⊔ᵛ-Monotonicʳ
|
|
||||||
; ⊔-idemp to ⊔ᵛ-idemp
|
|
||||||
)
|
|
||||||
open Lattice.FiniteValueMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight _≟ˢ_ isLatticeᵍ vars-Unique ≈ᵍ-dec _ fixedHeightᵍ
|
|
||||||
using ()
|
|
||||||
renaming
|
|
||||||
( isFiniteHeightLattice to isFiniteHeightLatticeᵛ
|
|
||||||
)
|
|
||||||
|
|
||||||
≈ᵛ-dec = ≈ᵍ-dec⇒≈ᵛ-dec ≈ᵍ-dec
|
eval : ∀ (e : Expr) → VariableValues → SignLattice
|
||||||
joinSemilatticeᵛ = IsFiniteHeightLattice.joinSemilattice isFiniteHeightLatticeᵛ
|
eval (e₁ + e₂) vs = plus (eval e₁ vs) (eval e₂ vs)
|
||||||
fixedHeightᵛ = IsFiniteHeightLattice.fixedHeight isFiniteHeightLatticeᵛ
|
eval (e₁ - e₂) vs = minus (eval e₁ vs) (eval e₂ vs)
|
||||||
⊥ᵛ = proj₁ (proj₁ (proj₁ fixedHeightᵛ))
|
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')) _ = [ + ]ᵍ
|
||||||
|
|
||||||
-- Finally, the map we care about is (state -> (variables -> sign)). Bring that in.
|
eval-Monoʳ : ∀ (e : Expr) → Monotonic _≼ᵛ_ _≼ᵍ_ (eval e)
|
||||||
module StateVariablesFiniteMap = Lattice.FiniteValueMap.WithKeys _≟_ isLatticeᵛ states
|
eval-Monoʳ (e₁ + e₂) {vs₁} {vs₂} vs₁≼vs₂ =
|
||||||
open StateVariablesFiniteMap
|
eval-combine₂ (λ {x} {y} {z} → ≼ᵍ-trans {x} {y} {z})
|
||||||
using (_[_]; m₁≼m₂⇒m₁[ks]≼m₂[ks])
|
plus plus-Mono₂ {o₁ = eval e₁ vs₁}
|
||||||
renaming
|
(eval-Monoʳ e₁ vs₁≼vs₂) (eval-Monoʳ e₂ vs₁≼vs₂)
|
||||||
( FiniteMap to StateVariables
|
eval-Monoʳ (e₁ - e₂) {vs₁} {vs₂} vs₁≼vs₂ =
|
||||||
; isLattice to isLatticeᵐ
|
eval-combine₂ (λ {x} {y} {z} → ≼ᵍ-trans {x} {y} {z})
|
||||||
; _∈k_ to _∈kᵐ_
|
minus minus-Mono₂ {o₁ = eval e₁ vs₁}
|
||||||
; locate to locateᵐ
|
(eval-Monoʳ e₁ vs₁≼vs₂) (eval-Monoʳ e₂ vs₁≼vs₂)
|
||||||
; _≼_ to _≼ᵐ_
|
eval-Monoʳ (` k) {vs₁@((kvs₁ , _) , _)} {vs₂@((kvs₂ , _), _)} vs₁≼vs₂
|
||||||
; ≈₂-dec⇒≈-dec to ≈ᵛ-dec⇒≈ᵐ-dec
|
with ∈k-decᵛ k kvs₁ | ∈k-decᵛ k kvs₂
|
||||||
; m₁≼m₂⇒m₁[k]≼m₂[k] to m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ
|
... | yes k∈kvs₁ | yes k∈kvs₂ =
|
||||||
)
|
|
||||||
open Lattice.FiniteValueMap.IterProdIsomorphism.WithUniqueKeysAndFixedHeight _≟_ isLatticeᵛ states-Unique ≈ᵛ-dec _ fixedHeightᵛ
|
|
||||||
using ()
|
|
||||||
renaming
|
|
||||||
( isFiniteHeightLattice to isFiniteHeightLatticeᵐ
|
|
||||||
)
|
|
||||||
|
|
||||||
≈ᵐ-dec = ≈ᵛ-dec⇒≈ᵐ-dec ≈ᵛ-dec
|
|
||||||
fixedHeightᵐ = IsFiniteHeightLattice.fixedHeight isFiniteHeightLatticeᵐ
|
|
||||||
|
|
||||||
-- build up the 'join' function, which follows from Exercise 4.26's
|
|
||||||
--
|
|
||||||
-- L₁ → (A → L₂)
|
|
||||||
--
|
|
||||||
-- Construction, with L₁ = (A → L₂), and f = id
|
|
||||||
|
|
||||||
joinForKey : State → StateVariables → VariableSigns
|
|
||||||
joinForKey k states = foldr _⊔ᵛ_ ⊥ᵛ (states [ incoming k ])
|
|
||||||
|
|
||||||
-- The per-key join is made up of map key accesses (which are monotonic)
|
|
||||||
-- and folds using the join operation (also monotonic)
|
|
||||||
|
|
||||||
joinForKey-Mono : ∀ (k : State) → Monotonic _≼ᵐ_ _≼ᵛ_ (joinForKey k)
|
|
||||||
joinForKey-Mono k {fm₁} {fm₂} fm₁≼fm₂ =
|
|
||||||
foldr-Mono joinSemilatticeᵛ joinSemilatticeᵛ (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 states isLatticeᵐ (λ x → x) (λ a₁≼a₂ → a₁≼a₂) joinForKey joinForKey-Mono states
|
|
||||||
renaming
|
|
||||||
( 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')) _ _ = [ + ]ᵍ
|
|
||||||
|
|
||||||
eval-Mono : ∀ (e : Expr) (k∈e⇒k∈vars : ∀ k → k ∈ᵉ e → k ∈ˡ vars) → Monotonic _≼ᵛ_ _≼ᵍ_ (eval e k∈e⇒k∈vars)
|
|
||||||
eval-Mono (e₁ + e₂) k∈e⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂ =
|
|
||||||
let
|
|
||||||
-- TODO: can this be done with less boilerplate?
|
|
||||||
k∈e₁⇒k∈vars = λ k k∈e₁ → k∈e⇒k∈vars k (in⁺₁ k∈e₁)
|
|
||||||
k∈e₂⇒k∈vars = λ k k∈e₂ → k∈e⇒k∈vars k (in⁺₂ k∈e₂)
|
|
||||||
g₁vs₁ = eval e₁ k∈e₁⇒k∈vars vs₁
|
|
||||||
g₂vs₁ = eval e₂ k∈e₂⇒k∈vars vs₁
|
|
||||||
g₁vs₂ = eval e₁ k∈e₁⇒k∈vars vs₂
|
|
||||||
g₂vs₂ = eval e₂ k∈e₂⇒k∈vars 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₁ k∈e₁⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂))
|
|
||||||
(plus-Monoʳ g₁vs₂ {g₂vs₁} {g₂vs₂} (eval-Mono e₂ k∈e₂⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂))
|
|
||||||
eval-Mono (e₁ - e₂) k∈e⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂ =
|
|
||||||
let
|
|
||||||
-- TODO: here too -- can this be done with less boilerplate?
|
|
||||||
k∈e₁⇒k∈vars = λ k k∈e₁ → k∈e⇒k∈vars k (in⁻₁ k∈e₁)
|
|
||||||
k∈e₂⇒k∈vars = λ k k∈e₂ → k∈e⇒k∈vars k (in⁻₂ k∈e₂)
|
|
||||||
g₁vs₁ = eval e₁ k∈e₁⇒k∈vars vs₁
|
|
||||||
g₂vs₁ = eval e₂ k∈e₂⇒k∈vars vs₁
|
|
||||||
g₁vs₂ = eval e₁ k∈e₁⇒k∈vars vs₂
|
|
||||||
g₂vs₂ = eval e₂ k∈e₂⇒k∈vars 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₁ k∈e₁⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂))
|
|
||||||
(minus-Monoʳ g₁vs₂ {g₂vs₁} {g₂vs₂} (eval-Mono e₂ k∈e₂⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂))
|
|
||||||
eval-Mono (` k) k∈e⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂ =
|
|
||||||
let
|
|
||||||
(v₁ , k,v₁∈vs₁) = locateᵛ {k} {vs₁} (vars-in-Map k vs₁ (k∈e⇒k∈vars k here))
|
|
||||||
(v₂ , k,v₂∈vs₂) = locateᵛ {k} {vs₂} (vars-in-Map k vs₂ (k∈e⇒k∈vars k here))
|
|
||||||
in
|
|
||||||
m₁≼m₂⇒m₁[k]ᵛ≼m₂[k]ᵛ vs₁ vs₂ vs₁≼vs₂ k,v₁∈vs₁ k,v₂∈vs₂
|
|
||||||
eval-Mono (# 0) _ _ = ≈ᵍ-refl
|
|
||||||
eval-Mono (# (suc n')) _ _ = ≈ᵍ-refl
|
|
||||||
|
|
||||||
private module _ (k : String) (e : Expr) (k∈e⇒k∈vars : ∀ k → k ∈ᵉ e → k ∈ˡ vars) where
|
|
||||||
open VariableSignsFiniteMap.GeneralizedUpdate vars isLatticeᵛ (λ x → x) (λ a₁≼a₂ → a₁≼a₂) (λ _ → eval e k∈e⇒k∈vars) (λ _ {vs₁} {vs₂} vs₁≼vs₂ → eval-Mono e k∈e⇒k∈vars {vs₁} {vs₂} vs₁≼vs₂) (k ∷ [])
|
|
||||||
renaming
|
|
||||||
( f' to updateVariablesFromExpression
|
|
||||||
; f'-Monotonic to updateVariablesFromExpression-Mono
|
|
||||||
)
|
|
||||||
public
|
|
||||||
|
|
||||||
updateVariablesForState : State → StateVariables → VariableSigns
|
|
||||||
updateVariablesForState s sv
|
|
||||||
-- More weirdness here. Apparently, capturing the with-equality proof
|
|
||||||
-- using 'in p' makes code that reasons about this function (below)
|
|
||||||
-- throw ill-typed with-abstraction errors. Instead, make use of the
|
|
||||||
-- fact that later with-clauses are generalized over earlier ones to
|
|
||||||
-- construct a specialization of vars-complete for (code s).
|
|
||||||
with code s | (λ k → vars-complete {k} s)
|
|
||||||
... | k ← e | k∈codes⇒k∈vars =
|
|
||||||
let
|
let
|
||||||
(vs , s,vs∈sv) = locateᵐ {s} {sv} (states-in-Map s sv)
|
(v₁ , k,v₁∈vs₁) = locateᵛ {k} {vs₁} k∈kvs₁
|
||||||
|
(v₂ , k,v₂∈vs₂) = locateᵛ {k} {vs₂} k∈kvs₂
|
||||||
in
|
in
|
||||||
updateVariablesFromExpression k e (λ k k∈e → k∈codes⇒k∈vars k (in←₂ k∈e)) vs
|
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
|
||||||
|
|
||||||
updateVariablesForState-Monoʳ : ∀ (s : State) → Monotonic _≼ᵐ_ _≼ᵛ_ (updateVariablesForState s)
|
instance
|
||||||
updateVariablesForState-Monoʳ s {sv₁} {sv₂} sv₁≼sv₂
|
SignEval : ExprEvaluator
|
||||||
with code s | (λ k → vars-complete {k} s)
|
SignEval = record { eval = eval; eval-Monoʳ = eval-Monoʳ }
|
||||||
... | k ← e | k∈codes⇒k∈vars =
|
|
||||||
let
|
|
||||||
(vs₁ , s,vs₁∈sv₁) = locateᵐ {s} {sv₁} (states-in-Map s sv₁)
|
|
||||||
(vs₂ , s,vs₂∈sv₂) = locateᵐ {s} {sv₂} (states-in-Map s sv₂)
|
|
||||||
vs₁≼vs₂ = m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ sv₁ sv₂ sv₁≼sv₂ s,vs₁∈sv₁ s,vs₂∈sv₂
|
|
||||||
in
|
|
||||||
updateVariablesFromExpression-Mono k e (λ k k∈e → k∈codes⇒k∈vars k (in←₂ k∈e)) {vs₁} {vs₂} vs₁≼vs₂
|
|
||||||
|
|
||||||
open StateVariablesFiniteMap.GeneralizedUpdate states isLatticeᵐ (λ x → x) (λ a₁≼a₂ → a₁≼a₂) updateVariablesForState updateVariablesForState-Monoʳ states
|
-- For debugging purposes, print out the result.
|
||||||
renaming
|
output = show (Analysis.Forward.WithProg.result SignLattice prog)
|
||||||
( f' to updateAll
|
|
||||||
; f'-Monotonic to updateAll-Mono
|
|
||||||
)
|
|
||||||
|
|
||||||
analyze : StateVariables → StateVariables
|
-- This should have fewer cases -- the same number as the actual 'plus' above.
|
||||||
analyze = updateAll ∘ joinAll
|
-- 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 {[ + ]ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
plus-valid {[ - ]ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
plus-valid {[ 0ˢ ]ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
plus-valid {⊤ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
plus-valid {⊤ᵍ} {[ + ]ᵍ} _ _ = tt
|
||||||
|
plus-valid {⊤ᵍ} {[ - ]ᵍ} _ _ = tt
|
||||||
|
plus-valid {⊤ᵍ} {[ 0ˢ ]ᵍ} _ _ = tt
|
||||||
|
plus-valid {⊤ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
plus-valid {[ + ]ᵍ} {[ + ]ᵍ} (n₁ , refl) (n₂ , refl) = (_ , refl)
|
||||||
|
plus-valid {[ + ]ᵍ} {[ - ]ᵍ} _ _ = tt
|
||||||
|
plus-valid {[ + ]ᵍ} {[ 0ˢ ]ᵍ} (n₁ , refl) refl = (_ , refl)
|
||||||
|
plus-valid {[ + ]ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
plus-valid {[ - ]ᵍ} {[ + ]ᵍ} _ _ = tt
|
||||||
|
plus-valid {[ - ]ᵍ} {[ - ]ᵍ} (n₁ , refl) (n₂ , refl) = (_ , refl)
|
||||||
|
plus-valid {[ - ]ᵍ} {[ 0ˢ ]ᵍ} (n₁ , refl) refl = (_ , refl)
|
||||||
|
plus-valid {[ - ]ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
plus-valid {[ 0ˢ ]ᵍ} {[ + ]ᵍ} refl (n₂ , refl) = (_ , refl)
|
||||||
|
plus-valid {[ 0ˢ ]ᵍ} {[ - ]ᵍ} refl (n₂ , refl) = (_ , refl)
|
||||||
|
plus-valid {[ 0ˢ ]ᵍ} {[ 0ˢ ]ᵍ} refl refl = refl
|
||||||
|
plus-valid {[ 0ˢ ]ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
|
||||||
analyze-Mono : Monotonic _≼ᵐ_ _≼ᵐ_ analyze
|
-- Same for this one. It should be easier, but Agda won't simplify.
|
||||||
analyze-Mono {sv₁} {sv₂} sv₁≼sv₂ = updateAll-Mono {joinAll sv₁} {joinAll sv₂} (joinAll-Mono {sv₁} {sv₂} sv₁≼sv₂)
|
minus-valid : ∀ {g₁ g₂} {z₁ z₂} → ⟦ g₁ ⟧ᵍ (↑ᶻ z₁) → ⟦ g₂ ⟧ᵍ (↑ᶻ z₂) → ⟦ minus g₁ g₂ ⟧ᵍ (↑ᶻ (z₁ Int.- z₂))
|
||||||
|
minus-valid {⊥ᵍ} {_} ⊥ _ = ⊥
|
||||||
|
minus-valid {[ + ]ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
minus-valid {[ - ]ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
minus-valid {[ 0ˢ ]ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
minus-valid {⊤ᵍ} {⊥ᵍ} _ ⊥ = ⊥
|
||||||
|
minus-valid {⊤ᵍ} {[ + ]ᵍ} _ _ = tt
|
||||||
|
minus-valid {⊤ᵍ} {[ - ]ᵍ} _ _ = tt
|
||||||
|
minus-valid {⊤ᵍ} {[ 0ˢ ]ᵍ} _ _ = tt
|
||||||
|
minus-valid {⊤ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
minus-valid {[ + ]ᵍ} {[ + ]ᵍ} _ _ = tt
|
||||||
|
minus-valid {[ + ]ᵍ} {[ - ]ᵍ} (n₁ , refl) (n₂ , refl) = (_ , refl)
|
||||||
|
minus-valid {[ + ]ᵍ} {[ 0ˢ ]ᵍ} (n₁ , refl) refl = (_ , refl)
|
||||||
|
minus-valid {[ + ]ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
minus-valid {[ - ]ᵍ} {[ + ]ᵍ} (n₁ , refl) (n₂ , refl) = (_ , refl)
|
||||||
|
minus-valid {[ - ]ᵍ} {[ - ]ᵍ} _ _ = tt
|
||||||
|
minus-valid {[ - ]ᵍ} {[ 0ˢ ]ᵍ} (n₁ , refl) refl = (_ , refl)
|
||||||
|
minus-valid {[ - ]ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
minus-valid {[ 0ˢ ]ᵍ} {[ + ]ᵍ} refl (n₂ , refl) = (_ , refl)
|
||||||
|
minus-valid {[ 0ˢ ]ᵍ} {[ - ]ᵍ} refl (n₂ , refl) = (_ , refl)
|
||||||
|
minus-valid {[ 0ˢ ]ᵍ} {[ 0ˢ ]ᵍ} refl refl = refl
|
||||||
|
minus-valid {[ 0ˢ ]ᵍ} {⊤ᵍ} _ _ = tt
|
||||||
|
|
||||||
open import Fixedpoint ≈ᵐ-dec isFiniteHeightLatticeᵐ analyze (λ {m₁} {m₂} m₁≼m₂ → analyze-Mono {m₁} {m₂} m₁≼m₂)
|
eval-valid : IsValidExprEvaluator
|
||||||
using ()
|
eval-valid (⇒ᵉ-+ ρ e₁ e₂ z₁ z₂ ρ,e₁⇒z₁ ρ,e₂⇒z₂) ⟦vs⟧ρ =
|
||||||
renaming (aᶠ to signs)
|
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 (⇒ᵉ-ℕ ρ 0) _ = refl
|
||||||
|
eval-valid (⇒ᵉ-ℕ ρ (suc n')) _ = (n' , refl)
|
||||||
|
|
||||||
|
instance
|
||||||
|
SignEvalValid : ValidExprEvaluator SignEval latticeInterpretationᵍ
|
||||||
|
SignEvalValid = record { valid = eval-valid }
|
||||||
|
|
||||||
-- Debugging code: print the resulting map.
|
analyze-correct = Analysis.Forward.WithProg.analyze-correct SignLattice prog tt
|
||||||
output = show signs
|
|
||||||
|
|||||||
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₄)
|
||||||
13
Chain.agda
13
Chain.agda
@@ -10,7 +10,7 @@ open import Data.Nat using (ℕ; suc; _+_; _≤_)
|
|||||||
open import Data.Nat.Properties using (+-comm; m+1+n≰m)
|
open import Data.Nat.Properties using (+-comm; m+1+n≰m)
|
||||||
open import Data.Product using (_×_; Σ; _,_)
|
open import Data.Product using (_×_; Σ; _,_)
|
||||||
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl)
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl)
|
||||||
open import Data.Empty using (⊥)
|
open import Data.Empty as Empty using ()
|
||||||
|
|
||||||
open IsEquivalence ≈-equiv
|
open IsEquivalence ≈-equiv
|
||||||
|
|
||||||
@@ -38,11 +38,16 @@ module _ where
|
|||||||
Bounded : ℕ → Set a
|
Bounded : ℕ → Set a
|
||||||
Bounded bound = ∀ {a₁ a₂ : A} {n : ℕ} → Chain a₁ a₂ n → n ≤ bound
|
Bounded bound = ∀ {a₁ a₂ : A} {n : ℕ} → Chain a₁ a₂ n → n ≤ bound
|
||||||
|
|
||||||
Bounded-suc-n : ∀ {a₁ a₂ : A} {n : ℕ} → Bounded n → Chain a₁ a₂ (suc n) → ⊥
|
Bounded-suc-n : ∀ {a₁ a₂ : A} {n : ℕ} → Bounded n → Chain a₁ a₂ (suc n) → Empty.⊥
|
||||||
Bounded-suc-n {a₁} {a₂} {n} bounded c = (m+1+n≰m n n+1≤n)
|
Bounded-suc-n {a₁} {a₂} {n} bounded c = (m+1+n≰m n n+1≤n)
|
||||||
where
|
where
|
||||||
n+1≤n : n + 1 ≤ n
|
n+1≤n : n + 1 ≤ n
|
||||||
n+1≤n rewrite (+-comm n 1) = bounded c
|
n+1≤n rewrite (+-comm n 1) = bounded c
|
||||||
|
|
||||||
Height : ℕ → Set a
|
record Height (height : ℕ) : Set a where
|
||||||
Height height = (Σ (A × A) (λ (a₁ , a₂) → Chain a₁ a₂ height) × Bounded height)
|
field
|
||||||
|
⊥ : A
|
||||||
|
⊤ : A
|
||||||
|
|
||||||
|
longestChain : Chain ⊥ ⊤ height
|
||||||
|
bounded : Bounded height
|
||||||
|
|||||||
@@ -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}
|
||||||
(≈-dec : 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
|
||||||
@@ -17,27 +17,19 @@ open import Data.Empty using (⊥-elim)
|
|||||||
open import Relation.Binary.PropositionalEquality using (_≡_; sym)
|
open import Relation.Binary.PropositionalEquality using (_≡_; sym)
|
||||||
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
||||||
|
|
||||||
|
open IsDecidable ≈-Decidable using () renaming (R-dec to ≈-dec)
|
||||||
open IsFiniteHeightLattice flA
|
open IsFiniteHeightLattice flA
|
||||||
|
|
||||||
import Chain
|
import Chain
|
||||||
module ChainA = Chain _≈_ ≈-equiv _≺_ ≺-cong
|
module ChainA = Chain _≈_ ≈-equiv _≺_ ≺-cong
|
||||||
|
|
||||||
private
|
private
|
||||||
⊥ᴬ : A
|
open ChainA.Height fixedHeight
|
||||||
⊥ᴬ = proj₁ (proj₁ (proj₁ fixedHeight))
|
using ()
|
||||||
|
renaming
|
||||||
⊥ᴬ≼ : ∀ (a : A) → ⊥ᴬ ≼ a
|
( ⊥ to ⊥ᴬ
|
||||||
⊥ᴬ≼ a with ≈-dec a ⊥ᴬ
|
; bounded to boundedᴬ
|
||||||
... | 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 (proj₂ fixedHeight) (ChainA.step x≺⊥ᴬ ≈-refl (proj₂ (proj₁ fixedHeight))))
|
|
||||||
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
|
||||||
@@ -45,7 +37,7 @@ private
|
|||||||
-- out, we have exceeded h steps, which shouldn't be possible.
|
-- out, we have exceeded h steps, which shouldn't be possible.
|
||||||
|
|
||||||
doStep : ∀ (g hᶜ : ℕ) (a₁ a₂ : A) (c : ChainA.Chain a₁ a₂ hᶜ) (g+hᶜ≡h : g + hᶜ ≡ suc h) (a₂≼fa₂ : a₂ ≼ f a₂) → Σ A (λ a → a ≈ f a)
|
doStep : ∀ (g hᶜ : ℕ) (a₁ a₂ : A) (c : ChainA.Chain a₁ a₂ hᶜ) (g+hᶜ≡h : g + hᶜ ≡ suc h) (a₂≼fa₂ : a₂ ≼ f a₂) → Σ A (λ a → a ≈ f a)
|
||||||
doStep 0 hᶜ a₁ a₂ c g+hᶜ≡sh a₂≼fa₂ rewrite g+hᶜ≡sh = ⊥-elim (ChainA.Bounded-suc-n (proj₂ fixedHeight) c)
|
doStep 0 hᶜ a₁ a₂ c g+hᶜ≡sh a₂≼fa₂ rewrite g+hᶜ≡sh = ⊥-elim (ChainA.Bounded-suc-n boundedᴬ c)
|
||||||
doStep (suc g') hᶜ a₁ a₂ c g+hᶜ≡sh a₂≼fa₂ rewrite sym (+-suc g' hᶜ)
|
doStep (suc g') hᶜ a₁ a₂ c g+hᶜ≡sh a₂≼fa₂ rewrite sym (+-suc g' hᶜ)
|
||||||
with ≈-dec a₂ (f a₂)
|
with ≈-dec a₂ (f a₂)
|
||||||
... | yes a₂≈fa₂ = (a₂ , a₂≈fa₂)
|
... | yes a₂≈fa₂ = (a₂ , a₂≈fa₂)
|
||||||
@@ -58,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
|
||||||
@@ -67,15 +59,15 @@ aᶠ≈faᶠ : aᶠ ≈ f aᶠ
|
|||||||
aᶠ≈faᶠ = proj₂ fix
|
aᶠ≈faᶠ = proj₂ fix
|
||||||
|
|
||||||
private
|
private
|
||||||
stepPreservesLess : ∀ (g hᶜ : ℕ) (a₁ a₂ a : A) (a≈fa : a ≈ f a) (a₂≼a : a₂ ≼ a)
|
stepPreservesLess : ∀ (g hᶜ : ℕ) (a₁ a₂ b : A) (b≈fb : b ≈ f b) (a₂≼a : a₂ ≼ b)
|
||||||
(c : ChainA.Chain a₁ a₂ hᶜ) (g+hᶜ≡h : g + hᶜ ≡ suc h)
|
(c : ChainA.Chain a₁ a₂ hᶜ) (g+hᶜ≡h : g + hᶜ ≡ suc h)
|
||||||
(a₂≼fa₂ : a₂ ≼ f a₂) →
|
(a₂≼fa₂ : a₂ ≼ f a₂) →
|
||||||
proj₁ (doStep g hᶜ a₁ a₂ c g+hᶜ≡h a₂≼fa₂) ≼ a
|
proj₁ (doStep g hᶜ a₁ a₂ c g+hᶜ≡h a₂≼fa₂) ≼ b
|
||||||
stepPreservesLess 0 _ _ _ _ _ _ c g+hᶜ≡sh _ rewrite g+hᶜ≡sh = ⊥-elim (ChainA.Bounded-suc-n (proj₂ fixedHeight) c)
|
stepPreservesLess 0 _ _ _ _ _ _ c g+hᶜ≡sh _ rewrite g+hᶜ≡sh = ⊥-elim (ChainA.Bounded-suc-n boundedᴬ c)
|
||||||
stepPreservesLess (suc g') hᶜ a₁ a₂ a a≈fa a₂≼a c g+hᶜ≡sh a₂≼fa₂ rewrite sym (+-suc g' hᶜ)
|
stepPreservesLess (suc g') hᶜ a₁ a₂ b b≈fb a₂≼b c g+hᶜ≡sh a₂≼fa₂ rewrite sym (+-suc g' hᶜ)
|
||||||
with ≈-dec a₂ (f a₂)
|
with ≈-dec a₂ (f a₂)
|
||||||
... | yes _ = a₂≼a
|
... | yes _ = a₂≼b
|
||||||
... | no _ = stepPreservesLess g' _ _ _ a a≈fa (≼-cong ≈-refl (≈-sym a≈fa) (Monotonicᶠ a₂≼a)) _ _ _
|
... | 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 ⊥ᴬ))
|
||||||
|
|||||||
@@ -38,8 +38,9 @@ module TransportFiniteHeight
|
|||||||
open IsEquivalence ≈₁-equiv using () renaming (≈-sym to ≈₁-sym; ≈-trans to ≈₁-trans)
|
open IsEquivalence ≈₁-equiv using () renaming (≈-sym to ≈₁-sym; ≈-trans to ≈₁-trans)
|
||||||
open IsEquivalence ≈₂-equiv using () renaming (≈-sym to ≈₂-sym; ≈-trans to ≈₂-trans)
|
open IsEquivalence ≈₂-equiv using () renaming (≈-sym to ≈₂-sym; ≈-trans to ≈₂-trans)
|
||||||
|
|
||||||
open import Chain _≈₁_ ≈₁-equiv _≺₁_ ≺₁-cong using () renaming (Chain to Chain₁; done to done₁; step to step₁)
|
import Chain
|
||||||
open import Chain _≈₂_ ≈₂-equiv _≺₂_ ≺₂-cong using () renaming (Chain to Chain₂; done to done₂; step to step₂)
|
open Chain _≈₁_ ≈₁-equiv _≺₁_ ≺₁-cong using () renaming (Chain to Chain₁; done to done₁; step to step₁)
|
||||||
|
open Chain _≈₂_ ≈₂-equiv _≺₂_ ≺₂-cong using () renaming (Chain to Chain₂; done to done₂; step to step₂)
|
||||||
|
|
||||||
private
|
private
|
||||||
f-Injective : Injective _≈₁_ _≈₂_ f
|
f-Injective : Injective _≈₁_ _≈₂_ f
|
||||||
@@ -62,20 +63,30 @@ module TransportFiniteHeight
|
|||||||
portChain₂ (done₂ a₂≈a₁) = done₁ (g-preserves-≈₂ a₂≈a₁)
|
portChain₂ (done₂ a₂≈a₁) = done₁ (g-preserves-≈₂ a₂≈a₁)
|
||||||
portChain₂ (step₂ {b₁} {b₂} (b₁≼b₂ , b₁̷≈b₂) b₂≈b₂' c) = step₁ (≈₁-trans (≈₁-sym (g-⊔-distr b₁ b₂)) (g-preserves-≈₂ b₁≼b₂) , g-preserves-̷≈ b₁̷≈b₂) (g-preserves-≈₂ b₂≈b₂') (portChain₂ c)
|
portChain₂ (step₂ {b₁} {b₂} (b₁≼b₂ , b₁̷≈b₂) b₂≈b₂' c) = step₁ (≈₁-trans (≈₁-sym (g-⊔-distr b₁ b₂)) (g-preserves-≈₂ b₁≼b₂) , g-preserves-̷≈ b₁̷≈b₂) (g-preserves-≈₂ b₂≈b₂') (portChain₂ c)
|
||||||
|
|
||||||
isFiniteHeightLattice : IsFiniteHeightLattice B height _≈₂_ _⊔₂_ _⊓₂_
|
open Chain.Height (IsFiniteHeightLattice.fixedHeight fhlA)
|
||||||
isFiniteHeightLattice =
|
using ()
|
||||||
let
|
renaming (⊥ to ⊥₁; ⊤ to ⊤₁; bounded to bounded₁; longestChain to c)
|
||||||
(((a₁ , a₂) , c) , bounded₁) = IsFiniteHeightLattice.fixedHeight fhlA
|
|
||||||
in record
|
instance
|
||||||
{ isLattice = lB
|
fixedHeight : IsLattice.FixedHeight lB height
|
||||||
; fixedHeight = (((f a₁ , f a₂), portChain₁ c) , λ c' → bounded₁ (portChain₂ c'))
|
fixedHeight = record
|
||||||
|
{ ⊥ = f ⊥₁
|
||||||
|
; ⊤ = f ⊤₁
|
||||||
|
; longestChain = portChain₁ c
|
||||||
|
; bounded = λ c' → bounded₁ (portChain₂ c')
|
||||||
}
|
}
|
||||||
|
|
||||||
finiteHeightLattice : FiniteHeightLattice B
|
isFiniteHeightLattice : IsFiniteHeightLattice B height _≈₂_ _⊔₂_ _⊓₂_
|
||||||
finiteHeightLattice = record
|
isFiniteHeightLattice = record
|
||||||
{ height = height
|
{ isLattice = lB
|
||||||
; _≈_ = _≈₂_
|
; fixedHeight = fixedHeight
|
||||||
; _⊔_ = _⊔₂_
|
}
|
||||||
; _⊓_ = _⊓₂_
|
|
||||||
; isFiniteHeightLattice = isFiniteHeightLattice
|
finiteHeightLattice : FiniteHeightLattice B
|
||||||
}
|
finiteHeightLattice = record
|
||||||
|
{ height = height
|
||||||
|
; _≈_ = _≈₂_
|
||||||
|
; _⊔_ = _⊔₂_
|
||||||
|
; _⊓_ = _⊓₂_
|
||||||
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
||||||
|
}
|
||||||
|
|||||||
239
Language.agda
239
Language.agda
@@ -1,175 +1,57 @@
|
|||||||
module Language where
|
module Language where
|
||||||
|
|
||||||
open import Data.Nat using (ℕ; suc; pred)
|
open import Language.Base public
|
||||||
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
|
open import Language.Semantics public
|
||||||
open import Data.Product using (Σ; _,_; proj₁; proj₂)
|
open import Language.Traces public
|
||||||
open import Data.Vec using (Vec; foldr; lookup; _∷_)
|
open import Language.Graphs public
|
||||||
open import Data.List using ([]; _∷_; List) renaming (foldr to foldrˡ; map to mapˡ)
|
open import Language.Properties public
|
||||||
open import Data.List.Membership.Propositional as MemProp using () renaming (_∈_ to _∈ˡ_)
|
|
||||||
open import Data.List.Relation.Unary.All using (All; []; _∷_)
|
open import Data.Fin using (Fin; suc; zero)
|
||||||
|
open import Data.Fin.Properties as FinProp using (suc-injective)
|
||||||
|
open import Data.List as List using (List; []; _∷_)
|
||||||
|
open import Data.List.Membership.Propositional as ListMem using ()
|
||||||
|
open import Data.List.Membership.Propositional.Properties as ListMemProp using (∈-filter⁺)
|
||||||
open import Data.List.Relation.Unary.Any as RelAny using ()
|
open import Data.List.Relation.Unary.Any as RelAny using ()
|
||||||
open import Data.Fin using (Fin; suc; zero; fromℕ; inject₁) renaming (_≟_ to _≟ᶠ_)
|
open import Data.Nat using (ℕ; suc)
|
||||||
open import Data.Fin.Properties using (suc-injective)
|
open import Data.Product using (_,_; Σ; proj₁; proj₂)
|
||||||
open import Relation.Binary.PropositionalEquality using (cong; _≡_; refl)
|
open import Data.Product.Properties as ProdProp using ()
|
||||||
|
open import Data.String using (String) renaming (_≟_ to _≟ˢ_)
|
||||||
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
|
open import Relation.Binary.PropositionalEquality using (_≡_; refl)
|
||||||
open import Relation.Nullary using (¬_)
|
open import Relation.Nullary using (¬_)
|
||||||
open import Function using (_∘_)
|
|
||||||
|
|
||||||
open import Lattice
|
open import Lattice
|
||||||
open import Utils using (Unique; Unique-map; empty; push; x∈xs⇒fx∈fxs)
|
open import Utils using (Unique; push; Unique-map; x∈xs⇒fx∈fxs)
|
||||||
|
open import Lattice.MapSet String {{record { R-dec = _≟ˢ_ }}} _ using ()
|
||||||
data Expr : Set where
|
|
||||||
_+_ : Expr → Expr → Expr
|
|
||||||
_-_ : Expr → Expr → Expr
|
|
||||||
`_ : String → Expr
|
|
||||||
#_ : ℕ → Expr
|
|
||||||
|
|
||||||
data Stmt : Set where
|
|
||||||
_←_ : String → Expr → Stmt
|
|
||||||
|
|
||||||
open import Lattice.MapSet _≟ˢ_
|
|
||||||
renaming
|
renaming
|
||||||
( MapSet to StringSet
|
( MapSet to StringSet
|
||||||
; insert to insertˢ
|
|
||||||
; to-List to to-Listˢ
|
; to-List to to-Listˢ
|
||||||
; empty to emptyˢ
|
|
||||||
; singleton to singletonˢ
|
|
||||||
; _⊔_ to _⊔ˢ_
|
|
||||||
; `_ to `ˢ_
|
|
||||||
; _∈_ to _∈ˢ_
|
|
||||||
; ⊔-preserves-∈k₁ to ⊔ˢ-preserves-∈k₁
|
|
||||||
; ⊔-preserves-∈k₂ to ⊔ˢ-preserves-∈k₂
|
|
||||||
)
|
)
|
||||||
|
|
||||||
data _∈ᵉ_ : String → Expr → Set where
|
|
||||||
in⁺₁ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₁ → k ∈ᵉ (e₁ + e₂)
|
|
||||||
in⁺₂ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₂ → k ∈ᵉ (e₁ + e₂)
|
|
||||||
in⁻₁ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₁ → k ∈ᵉ (e₁ - e₂)
|
|
||||||
in⁻₂ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₂ → k ∈ᵉ (e₁ - e₂)
|
|
||||||
here : ∀ {k : String} → k ∈ᵉ (` k)
|
|
||||||
|
|
||||||
data _∈ᵗ_ : String → Stmt → Set where
|
|
||||||
in←₁ : ∀ {k : String} {e : Expr} → k ∈ᵗ (k ← e)
|
|
||||||
in←₂ : ∀ {k k' : String} {e : Expr} → k ∈ᵉ e → k ∈ᵗ (k' ← e)
|
|
||||||
|
|
||||||
private
|
|
||||||
Expr-vars : Expr → StringSet
|
|
||||||
Expr-vars (l + r) = Expr-vars l ⊔ˢ Expr-vars r
|
|
||||||
Expr-vars (l - r) = Expr-vars l ⊔ˢ Expr-vars r
|
|
||||||
Expr-vars (` s) = singletonˢ s
|
|
||||||
Expr-vars (# _) = emptyˢ
|
|
||||||
|
|
||||||
∈-Expr-vars⇒∈ : ∀ {k : String} (e : Expr) → k ∈ˢ (Expr-vars e) → k ∈ᵉ e
|
|
||||||
∈-Expr-vars⇒∈ {k} (e₁ + e₂) k∈vs
|
|
||||||
with Expr-Provenance k ((`ˢ (Expr-vars e₁)) ∪ (`ˢ (Expr-vars e₂))) k∈vs
|
|
||||||
... | in₁ (single k,tt∈vs₁) _ = (in⁺₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
|
||||||
... | in₂ _ (single k,tt∈vs₂) = (in⁺₂ (∈-Expr-vars⇒∈ e₂ (forget k,tt∈vs₂)))
|
|
||||||
... | bothᵘ (single k,tt∈vs₁) _ = (in⁺₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
|
||||||
∈-Expr-vars⇒∈ {k} (e₁ - e₂) k∈vs
|
|
||||||
with Expr-Provenance k ((`ˢ (Expr-vars e₁)) ∪ (`ˢ (Expr-vars e₂))) k∈vs
|
|
||||||
... | in₁ (single k,tt∈vs₁) _ = (in⁻₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
|
||||||
... | in₂ _ (single k,tt∈vs₂) = (in⁻₂ (∈-Expr-vars⇒∈ e₂ (forget k,tt∈vs₂)))
|
|
||||||
... | bothᵘ (single k,tt∈vs₁) _ = (in⁻₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
|
||||||
∈-Expr-vars⇒∈ {k} (` k) (RelAny.here refl) = here
|
|
||||||
|
|
||||||
∈⇒∈-Expr-vars : ∀ {k : String} {e : Expr} → k ∈ᵉ e → k ∈ˢ (Expr-vars e)
|
|
||||||
∈⇒∈-Expr-vars {k} {e₁ + e₂} (in⁺₁ k∈e₁) =
|
|
||||||
⊔ˢ-preserves-∈k₁ {m₁ = Expr-vars e₁}
|
|
||||||
{m₂ = Expr-vars e₂}
|
|
||||||
(∈⇒∈-Expr-vars k∈e₁)
|
|
||||||
∈⇒∈-Expr-vars {k} {e₁ + e₂} (in⁺₂ k∈e₂) =
|
|
||||||
⊔ˢ-preserves-∈k₂ {m₁ = Expr-vars e₁}
|
|
||||||
{m₂ = Expr-vars e₂}
|
|
||||||
(∈⇒∈-Expr-vars k∈e₂)
|
|
||||||
∈⇒∈-Expr-vars {k} {e₁ - e₂} (in⁻₁ k∈e₁) =
|
|
||||||
⊔ˢ-preserves-∈k₁ {m₁ = Expr-vars e₁}
|
|
||||||
{m₂ = Expr-vars e₂}
|
|
||||||
(∈⇒∈-Expr-vars k∈e₁)
|
|
||||||
∈⇒∈-Expr-vars {k} {e₁ - e₂} (in⁻₂ k∈e₂) =
|
|
||||||
⊔ˢ-preserves-∈k₂ {m₁ = Expr-vars e₁}
|
|
||||||
{m₂ = Expr-vars e₂}
|
|
||||||
(∈⇒∈-Expr-vars k∈e₂)
|
|
||||||
∈⇒∈-Expr-vars here = RelAny.here refl
|
|
||||||
|
|
||||||
Stmt-vars : Stmt → StringSet
|
|
||||||
Stmt-vars (x ← e) = (singletonˢ x) ⊔ˢ (Expr-vars e)
|
|
||||||
|
|
||||||
∈-Stmt-vars⇒∈ : ∀ {k : String} (s : Stmt) → k ∈ˢ (Stmt-vars s) → k ∈ᵗ s
|
|
||||||
∈-Stmt-vars⇒∈ {k} (k' ← e) k∈vs
|
|
||||||
with Expr-Provenance k ((`ˢ (singletonˢ k')) ∪ (`ˢ (Expr-vars e))) k∈vs
|
|
||||||
... | in₁ (single (RelAny.here refl)) _ = in←₁
|
|
||||||
... | in₂ _ (single k,tt∈vs') = in←₂ (∈-Expr-vars⇒∈ e (forget k,tt∈vs'))
|
|
||||||
... | bothᵘ (single (RelAny.here refl)) _ = in←₁
|
|
||||||
|
|
||||||
∈⇒∈-Stmt-vars : ∀ {k : String} {s : Stmt} → k ∈ᵗ s → k ∈ˢ (Stmt-vars s)
|
|
||||||
∈⇒∈-Stmt-vars {k} {k ← e} in←₁ =
|
|
||||||
⊔ˢ-preserves-∈k₁ {m₁ = singletonˢ k}
|
|
||||||
{m₂ = Expr-vars e}
|
|
||||||
(RelAny.here refl)
|
|
||||||
∈⇒∈-Stmt-vars {k} {k' ← e} (in←₂ k∈e) =
|
|
||||||
⊔ˢ-preserves-∈k₂ {m₁ = singletonˢ k'}
|
|
||||||
{m₂ = Expr-vars e}
|
|
||||||
(∈⇒∈-Expr-vars k∈e)
|
|
||||||
|
|
||||||
Stmts-vars : ∀ {n : ℕ} → Vec Stmt n → StringSet
|
|
||||||
Stmts-vars = foldr (λ n → StringSet)
|
|
||||||
(λ {k} stmt acc → (Stmt-vars stmt) ⊔ˢ acc) emptyˢ
|
|
||||||
|
|
||||||
∈-Stmts-vars⇒∈ : ∀ {n : ℕ} {k : String} (ss : Vec Stmt n) →
|
|
||||||
k ∈ˢ (Stmts-vars ss) → Σ (Fin n) (λ f → k ∈ᵗ lookup ss f)
|
|
||||||
∈-Stmts-vars⇒∈ {suc n'} {k} (s ∷ ss') k∈vss
|
|
||||||
with Expr-Provenance k ((`ˢ (Stmt-vars s)) ∪ (`ˢ (Stmts-vars ss'))) k∈vss
|
|
||||||
... | in₁ (single k,tt∈vs) _ = (zero , ∈-Stmt-vars⇒∈ s (forget k,tt∈vs))
|
|
||||||
... | in₂ _ (single k,tt∈vss') =
|
|
||||||
let
|
|
||||||
(f' , k∈s') = ∈-Stmts-vars⇒∈ ss' (forget k,tt∈vss')
|
|
||||||
in
|
|
||||||
(suc f' , k∈s')
|
|
||||||
... | bothᵘ (single k,tt∈vs) _ = (zero , ∈-Stmt-vars⇒∈ s (forget k,tt∈vs))
|
|
||||||
|
|
||||||
∈⇒∈-Stmts-vars : ∀ {n : ℕ} {k : String} {ss : Vec Stmt n} {f : Fin n} →
|
|
||||||
k ∈ᵗ lookup ss f → k ∈ˢ (Stmts-vars ss)
|
|
||||||
∈⇒∈-Stmts-vars {suc n} {k} {s ∷ ss'} {zero} k∈s =
|
|
||||||
⊔ˢ-preserves-∈k₁ {m₁ = Stmt-vars s}
|
|
||||||
{m₂ = Stmts-vars ss'}
|
|
||||||
(∈⇒∈-Stmt-vars k∈s)
|
|
||||||
∈⇒∈-Stmts-vars {suc n} {k} {s ∷ ss'} {suc f'} k∈ss' =
|
|
||||||
⊔ˢ-preserves-∈k₂ {m₁ = Stmt-vars s}
|
|
||||||
{m₂ = Stmts-vars ss'}
|
|
||||||
(∈⇒∈-Stmts-vars {n} {k} {ss'} {f'} k∈ss')
|
|
||||||
|
|
||||||
-- Creating a new number from a natural number can never create one
|
|
||||||
-- equal to one you get from weakening the bounds on another number.
|
|
||||||
z≢sf : ∀ {n : ℕ} (f : Fin n) → ¬ (zero ≡ suc f)
|
|
||||||
z≢sf f ()
|
|
||||||
|
|
||||||
z≢mapsfs : ∀ {n : ℕ} (fs : List (Fin n)) → All (λ sf → ¬ zero ≡ sf) (mapˡ suc fs)
|
|
||||||
z≢mapsfs [] = []
|
|
||||||
z≢mapsfs (f ∷ fs') = z≢sf f ∷ z≢mapsfs fs'
|
|
||||||
|
|
||||||
indices : ∀ (n : ℕ) → Σ (List (Fin n)) Unique
|
|
||||||
indices 0 = ([] , empty)
|
|
||||||
indices (suc n') =
|
|
||||||
let
|
|
||||||
(inds' , unids') = indices n'
|
|
||||||
in
|
|
||||||
( zero ∷ mapˡ suc inds'
|
|
||||||
, 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
|
record Program : Set where
|
||||||
field
|
field
|
||||||
length : ℕ
|
rootStmt : Stmt
|
||||||
stmts : Vec Stmt length
|
|
||||||
|
graph : Graph
|
||||||
|
graph = wrap (buildCfg rootStmt)
|
||||||
|
|
||||||
|
State : Set
|
||||||
|
State = Graph.Index graph
|
||||||
|
|
||||||
|
initialState : State
|
||||||
|
initialState = proj₁ (wrap-input (buildCfg rootStmt))
|
||||||
|
|
||||||
|
finalState : State
|
||||||
|
finalState = proj₁ (wrap-output (buildCfg rootStmt))
|
||||||
|
|
||||||
|
trace : ∀ {ρ : Env} → [] , rootStmt ⇒ˢ ρ → Trace {graph} initialState finalState [] ρ
|
||||||
|
trace {ρ} ∅,s⇒ρ
|
||||||
|
with MkEndToEndTrace idx₁ (RelAny.here refl) idx₂ (RelAny.here refl) tr
|
||||||
|
← EndToEndTrace-wrap (buildCfg-sufficient ∅,s⇒ρ) = tr
|
||||||
|
|
||||||
private
|
private
|
||||||
vars-Set : StringSet
|
vars-Set : StringSet
|
||||||
vars-Set = Stmts-vars stmts
|
vars-Set = Stmt-vars rootStmt
|
||||||
|
|
||||||
vars : List String
|
vars : List String
|
||||||
vars = to-Listˢ vars-Set
|
vars = to-Listˢ vars-Set
|
||||||
@@ -177,35 +59,36 @@ record Program : Set where
|
|||||||
vars-Unique : Unique vars
|
vars-Unique : Unique vars
|
||||||
vars-Unique = proj₂ vars-Set
|
vars-Unique = proj₂ vars-Set
|
||||||
|
|
||||||
State : Set
|
|
||||||
State = Fin length
|
|
||||||
|
|
||||||
states : List State
|
states : List State
|
||||||
states = proj₁ (indices length)
|
states = indices graph
|
||||||
|
|
||||||
states-complete : ∀ (s : State) → s ∈ˡ states
|
states-complete : ∀ (s : State) → s ListMem.∈ states
|
||||||
states-complete = indices-complete length
|
states-complete = indices-complete graph
|
||||||
|
|
||||||
states-Unique : Unique states
|
states-Unique : Unique states
|
||||||
states-Unique = proj₂ (indices length)
|
states-Unique = indices-Unique graph
|
||||||
|
|
||||||
code : State → Stmt
|
code : State → List BasicStmt
|
||||||
code = lookup stmts
|
code st = graph [ st ]
|
||||||
|
|
||||||
vars-complete : ∀ {k : String} (s : State) → k ∈ᵗ (code s) → k ∈ˡ vars
|
-- vars-complete : ∀ {k : String} (s : State) → k ∈ᵇ (code s) → k ListMem.∈ vars
|
||||||
vars-complete {k} s = ∈⇒∈-Stmts-vars {length} {k} {stmts} {s}
|
-- vars-complete {k} s = ∈⇒∈-Stmts-vars {length} {k} {stmts} {s}
|
||||||
|
|
||||||
_≟_ : IsDecidable (_≡_ {_} {State})
|
_≟_ : Decidable (_≡_ {_} {State})
|
||||||
_≟_ = _≟ᶠ_
|
_≟_ = FinProp._≟_
|
||||||
|
|
||||||
-- Computations for incoming and outgoing edges will have to change too
|
_≟ᵉ_ : Decidable (_≡_ {_} {Graph.Edge graph})
|
||||||
-- when we support branching etc.
|
_≟ᵉ_ = ProdProp.≡-dec _≟_ _≟_
|
||||||
|
|
||||||
|
open import Data.List.Membership.DecPropositional _≟ᵉ_ using (_∈?_)
|
||||||
|
|
||||||
incoming : State → List State
|
incoming : State → List State
|
||||||
incoming
|
incoming = predecessors graph
|
||||||
with length
|
|
||||||
... | 0 = (λ ())
|
initialState-pred-∅ : incoming initialState ≡ []
|
||||||
... | suc n' = (λ
|
initialState-pred-∅ =
|
||||||
{ zero → []
|
wrap-preds-∅ (buildCfg rootStmt) initialState (RelAny.here refl)
|
||||||
; (suc f') → (inject₁ f') ∷ []
|
|
||||||
})
|
edge⇒incoming : ∀ {s₁ s₂ : State} → (s₁ , s₂) ListMem.∈ (Graph.edges graph) →
|
||||||
|
s₁ ListMem.∈ (incoming s₂)
|
||||||
|
edge⇒incoming = edge⇒predecessor graph
|
||||||
|
|||||||
145
Language/Base.agda
Normal file
145
Language/Base.agda
Normal file
@@ -0,0 +1,145 @@
|
|||||||
|
module Language.Base where
|
||||||
|
|
||||||
|
open import Data.List as List using (List)
|
||||||
|
open import Data.Nat using (ℕ; suc)
|
||||||
|
open import Data.Product using (Σ; _,_; proj₁)
|
||||||
|
open import Data.String as String using (String)
|
||||||
|
open import Data.Vec using (Vec; foldr; lookup)
|
||||||
|
open import Relation.Binary.PropositionalEquality using (_≡_; refl)
|
||||||
|
|
||||||
|
open import Lattice
|
||||||
|
|
||||||
|
data Expr : Set where
|
||||||
|
_+_ : Expr → Expr → Expr
|
||||||
|
_-_ : Expr → Expr → Expr
|
||||||
|
`_ : String → Expr
|
||||||
|
#_ : ℕ → Expr
|
||||||
|
|
||||||
|
data BasicStmt : Set where
|
||||||
|
_←_ : String → Expr → BasicStmt
|
||||||
|
noop : BasicStmt
|
||||||
|
|
||||||
|
infixr 2 _then_
|
||||||
|
infix 3 if_then_else_
|
||||||
|
infix 3 while_repeat_
|
||||||
|
data Stmt : Set where
|
||||||
|
⟨_⟩ : BasicStmt → Stmt
|
||||||
|
_then_ : Stmt → Stmt → Stmt
|
||||||
|
if_then_else_ : Expr → Stmt → Stmt → Stmt
|
||||||
|
while_repeat_ : Expr → Stmt → Stmt
|
||||||
|
|
||||||
|
data _∈ᵉ_ : String → Expr → Set where
|
||||||
|
in⁺₁ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₁ → k ∈ᵉ (e₁ + e₂)
|
||||||
|
in⁺₂ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₂ → k ∈ᵉ (e₁ + e₂)
|
||||||
|
in⁻₁ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₁ → k ∈ᵉ (e₁ - e₂)
|
||||||
|
in⁻₂ : ∀ {e₁ e₂ : Expr} {k : String} → k ∈ᵉ e₂ → k ∈ᵉ (e₁ - e₂)
|
||||||
|
here : ∀ {k : String} → k ∈ᵉ (` k)
|
||||||
|
|
||||||
|
data _∈ᵇ_ : String → BasicStmt → Set where
|
||||||
|
in←₁ : ∀ {k : String} {e : Expr} → k ∈ᵇ (k ← e)
|
||||||
|
in←₂ : ∀ {k k' : String} {e : Expr} → k ∈ᵉ e → k ∈ᵇ (k' ← e)
|
||||||
|
|
||||||
|
open import Lattice.MapSet String {{record { R-dec = String._≟_ }}} _
|
||||||
|
renaming
|
||||||
|
( MapSet to StringSet
|
||||||
|
; insert to insertˢ
|
||||||
|
; empty to emptyˢ
|
||||||
|
; singleton to singletonˢ
|
||||||
|
; _⊔_ to _⊔ˢ_
|
||||||
|
; `_ to `ˢ_
|
||||||
|
; _∈_ to _∈ˢ_
|
||||||
|
; ⊔-preserves-∈k₁ to ⊔ˢ-preserves-∈k₁
|
||||||
|
; ⊔-preserves-∈k₂ to ⊔ˢ-preserves-∈k₂
|
||||||
|
)
|
||||||
|
|
||||||
|
Expr-vars : Expr → StringSet
|
||||||
|
Expr-vars (l + r) = Expr-vars l ⊔ˢ Expr-vars r
|
||||||
|
Expr-vars (l - r) = Expr-vars l ⊔ˢ Expr-vars r
|
||||||
|
Expr-vars (` s) = singletonˢ s
|
||||||
|
Expr-vars (# _) = emptyˢ
|
||||||
|
|
||||||
|
-- ∈-Expr-vars⇒∈ : ∀ {k : String} (e : Expr) → k ∈ˢ (Expr-vars e) → k ∈ᵉ e
|
||||||
|
-- ∈-Expr-vars⇒∈ {k} (e₁ + e₂) k∈vs
|
||||||
|
-- with Expr-Provenance k ((`ˢ (Expr-vars e₁)) ∪ (`ˢ (Expr-vars e₂))) k∈vs
|
||||||
|
-- ... | in₁ (single k,tt∈vs₁) _ = (in⁺₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
||||||
|
-- ... | in₂ _ (single k,tt∈vs₂) = (in⁺₂ (∈-Expr-vars⇒∈ e₂ (forget k,tt∈vs₂)))
|
||||||
|
-- ... | bothᵘ (single k,tt∈vs₁) _ = (in⁺₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
||||||
|
-- ∈-Expr-vars⇒∈ {k} (e₁ - e₂) k∈vs
|
||||||
|
-- with Expr-Provenance k ((`ˢ (Expr-vars e₁)) ∪ (`ˢ (Expr-vars e₂))) k∈vs
|
||||||
|
-- ... | in₁ (single k,tt∈vs₁) _ = (in⁻₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
||||||
|
-- ... | in₂ _ (single k,tt∈vs₂) = (in⁻₂ (∈-Expr-vars⇒∈ e₂ (forget k,tt∈vs₂)))
|
||||||
|
-- ... | bothᵘ (single k,tt∈vs₁) _ = (in⁻₁ (∈-Expr-vars⇒∈ e₁ (forget k,tt∈vs₁)))
|
||||||
|
-- ∈-Expr-vars⇒∈ {k} (` k) (RelAny.here refl) = here
|
||||||
|
|
||||||
|
-- ∈⇒∈-Expr-vars : ∀ {k : String} {e : Expr} → k ∈ᵉ e → k ∈ˢ (Expr-vars e)
|
||||||
|
-- ∈⇒∈-Expr-vars {k} {e₁ + e₂} (in⁺₁ k∈e₁) =
|
||||||
|
-- ⊔ˢ-preserves-∈k₁ {m₁ = Expr-vars e₁}
|
||||||
|
-- {m₂ = Expr-vars e₂}
|
||||||
|
-- (∈⇒∈-Expr-vars k∈e₁)
|
||||||
|
-- ∈⇒∈-Expr-vars {k} {e₁ + e₂} (in⁺₂ k∈e₂) =
|
||||||
|
-- ⊔ˢ-preserves-∈k₂ {m₁ = Expr-vars e₁}
|
||||||
|
-- {m₂ = Expr-vars e₂}
|
||||||
|
-- (∈⇒∈-Expr-vars k∈e₂)
|
||||||
|
-- ∈⇒∈-Expr-vars {k} {e₁ - e₂} (in⁻₁ k∈e₁) =
|
||||||
|
-- ⊔ˢ-preserves-∈k₁ {m₁ = Expr-vars e₁}
|
||||||
|
-- {m₂ = Expr-vars e₂}
|
||||||
|
-- (∈⇒∈-Expr-vars k∈e₁)
|
||||||
|
-- ∈⇒∈-Expr-vars {k} {e₁ - e₂} (in⁻₂ k∈e₂) =
|
||||||
|
-- ⊔ˢ-preserves-∈k₂ {m₁ = Expr-vars e₁}
|
||||||
|
-- {m₂ = Expr-vars e₂}
|
||||||
|
-- (∈⇒∈-Expr-vars k∈e₂)
|
||||||
|
-- ∈⇒∈-Expr-vars here = RelAny.here refl
|
||||||
|
|
||||||
|
BasicStmt-vars : BasicStmt → StringSet
|
||||||
|
BasicStmt-vars (x ← e) = (singletonˢ x) ⊔ˢ (Expr-vars e)
|
||||||
|
BasicStmt-vars noop = emptyˢ
|
||||||
|
|
||||||
|
Stmt-vars : Stmt → StringSet
|
||||||
|
Stmt-vars ⟨ bs ⟩ = BasicStmt-vars bs
|
||||||
|
Stmt-vars (s₁ then s₂) = (Stmt-vars s₁) ⊔ˢ (Stmt-vars s₂)
|
||||||
|
Stmt-vars (if e then s₁ else s₂) = ((Expr-vars e) ⊔ˢ (Stmt-vars s₁)) ⊔ˢ (Stmt-vars s₂)
|
||||||
|
Stmt-vars (while e repeat s) = (Expr-vars e) ⊔ˢ (Stmt-vars s)
|
||||||
|
|
||||||
|
-- ∈-Stmt-vars⇒∈ : ∀ {k : String} (s : Stmt) → k ∈ˢ (Stmt-vars s) → k ∈ᵇ s
|
||||||
|
-- ∈-Stmt-vars⇒∈ {k} (k' ← e) k∈vs
|
||||||
|
-- with Expr-Provenance k ((`ˢ (singletonˢ k')) ∪ (`ˢ (Expr-vars e))) k∈vs
|
||||||
|
-- ... | in₁ (single (RelAny.here refl)) _ = in←₁
|
||||||
|
-- ... | in₂ _ (single k,tt∈vs') = in←₂ (∈-Expr-vars⇒∈ e (forget k,tt∈vs'))
|
||||||
|
-- ... | bothᵘ (single (RelAny.here refl)) _ = in←₁
|
||||||
|
|
||||||
|
-- ∈⇒∈-Stmt-vars : ∀ {k : String} {s : Stmt} → k ∈ᵇ s → k ∈ˢ (Stmt-vars s)
|
||||||
|
-- ∈⇒∈-Stmt-vars {k} {k ← e} in←₁ =
|
||||||
|
-- ⊔ˢ-preserves-∈k₁ {m₁ = singletonˢ k}
|
||||||
|
-- {m₂ = Expr-vars e}
|
||||||
|
-- (RelAny.here refl)
|
||||||
|
-- ∈⇒∈-Stmt-vars {k} {k' ← e} (in←₂ k∈e) =
|
||||||
|
-- ⊔ˢ-preserves-∈k₂ {m₁ = singletonˢ k'}
|
||||||
|
-- {m₂ = Expr-vars e}
|
||||||
|
-- (∈⇒∈-Expr-vars k∈e)
|
||||||
|
|
||||||
|
Stmts-vars : ∀ {n : ℕ} → Vec Stmt n → StringSet
|
||||||
|
Stmts-vars = foldr (λ n → StringSet)
|
||||||
|
(λ {k} stmt acc → (Stmt-vars stmt) ⊔ˢ acc) emptyˢ
|
||||||
|
|
||||||
|
-- ∈-Stmts-vars⇒∈ : ∀ {n : ℕ} {k : String} (ss : Vec Stmt n) →
|
||||||
|
-- k ∈ˢ (Stmts-vars ss) → Σ (Fin n) (λ f → k ∈ᵇ lookup ss f)
|
||||||
|
-- ∈-Stmts-vars⇒∈ {suc n'} {k} (s ∷ ss') k∈vss
|
||||||
|
-- with Expr-Provenance k ((`ˢ (Stmt-vars s)) ∪ (`ˢ (Stmts-vars ss'))) k∈vss
|
||||||
|
-- ... | in₁ (single k,tt∈vs) _ = (zero , ∈-Stmt-vars⇒∈ s (forget k,tt∈vs))
|
||||||
|
-- ... | in₂ _ (single k,tt∈vss') =
|
||||||
|
-- let
|
||||||
|
-- (f' , k∈s') = ∈-Stmts-vars⇒∈ ss' (forget k,tt∈vss')
|
||||||
|
-- in
|
||||||
|
-- (suc f' , k∈s')
|
||||||
|
-- ... | bothᵘ (single k,tt∈vs) _ = (zero , ∈-Stmt-vars⇒∈ s (forget k,tt∈vs))
|
||||||
|
|
||||||
|
-- ∈⇒∈-Stmts-vars : ∀ {n : ℕ} {k : String} {ss : Vec Stmt n} {f : Fin n} →
|
||||||
|
-- k ∈ᵇ lookup ss f → k ∈ˢ (Stmts-vars ss)
|
||||||
|
-- ∈⇒∈-Stmts-vars {suc n} {k} {s ∷ ss'} {zero} k∈s =
|
||||||
|
-- ⊔ˢ-preserves-∈k₁ {m₁ = Stmt-vars s}
|
||||||
|
-- {m₂ = Stmts-vars ss'}
|
||||||
|
-- (∈⇒∈-Stmt-vars k∈s)
|
||||||
|
-- ∈⇒∈-Stmts-vars {suc n} {k} {s ∷ ss'} {suc f'} k∈ss' =
|
||||||
|
-- ⊔ˢ-preserves-∈k₂ {m₁ = Stmt-vars s}
|
||||||
|
-- {m₂ = Stmts-vars ss'}
|
||||||
|
-- (∈⇒∈-Stmts-vars {n} {k} {ss'} {f'} k∈ss')
|
||||||
177
Language/Graphs.agda
Normal file
177
Language/Graphs.agda
Normal file
@@ -0,0 +1,177 @@
|
|||||||
|
module Language.Graphs where
|
||||||
|
|
||||||
|
open import Language.Base using (Expr; Stmt; BasicStmt; ⟨_⟩; _then_; if_then_else_; while_repeat_)
|
||||||
|
|
||||||
|
open import Data.Fin as Fin using (Fin; suc; zero)
|
||||||
|
open import Data.Fin.Properties as FinProp using (suc-injective)
|
||||||
|
open import Data.List as List using (List; []; _∷_)
|
||||||
|
open import Data.List.Membership.Propositional as ListMem using ()
|
||||||
|
open import Data.List.Membership.Propositional.Properties as ListMemProp using (∈-filter⁺; ∈-filter⁻)
|
||||||
|
open import Data.List.Relation.Unary.All using (All; []; _∷_)
|
||||||
|
open import Data.List.Relation.Unary.Any as RelAny using ()
|
||||||
|
open import Data.Nat as Nat using (ℕ; suc)
|
||||||
|
open import Data.Nat.Properties using (+-assoc; +-comm)
|
||||||
|
open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂)
|
||||||
|
open import Data.Vec using (Vec; []; _∷_; lookup; cast; _++_)
|
||||||
|
open import Data.Vec.Properties using (cast-is-id; ++-assoc; lookup-++ˡ; cast-sym; ++-identityʳ; lookup-++ʳ)
|
||||||
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; sym; refl; subst; trans)
|
||||||
|
open import Relation.Nullary using (¬_)
|
||||||
|
|
||||||
|
open import Lattice
|
||||||
|
open import Utils using (Unique; push; Unique-map; x∈xs⇒fx∈fxs; ∈-cartesianProduct)
|
||||||
|
|
||||||
|
record Graph : Set where
|
||||||
|
constructor MkGraph
|
||||||
|
field
|
||||||
|
size : ℕ
|
||||||
|
|
||||||
|
Index : Set
|
||||||
|
Index = Fin size
|
||||||
|
|
||||||
|
Edge : Set
|
||||||
|
Edge = Index × Index
|
||||||
|
|
||||||
|
field
|
||||||
|
nodes : Vec (List BasicStmt) size
|
||||||
|
edges : List Edge
|
||||||
|
inputs : List Index
|
||||||
|
outputs : List Index
|
||||||
|
|
||||||
|
_↑ˡ_ : ∀ {n} → (Fin n × Fin n) → ∀ m → (Fin (n Nat.+ m) × Fin (n Nat.+ m))
|
||||||
|
_↑ˡ_ (idx₁ , idx₂) m = (idx₁ Fin.↑ˡ m , idx₂ Fin.↑ˡ m)
|
||||||
|
|
||||||
|
_↑ʳ_ : ∀ {m} n → (Fin m × Fin m) → Fin (n Nat.+ m) × Fin (n Nat.+ m)
|
||||||
|
_↑ʳ_ n (idx₁ , idx₂) = (n Fin.↑ʳ idx₁ , n Fin.↑ʳ idx₂)
|
||||||
|
|
||||||
|
_↑ˡⁱ_ : ∀ {n} → List (Fin n) → ∀ m → List (Fin (n Nat.+ m))
|
||||||
|
_↑ˡⁱ_ l m = List.map (Fin._↑ˡ m) l
|
||||||
|
|
||||||
|
_↑ʳⁱ_ : ∀ {m} n → List (Fin m) → List (Fin (n Nat.+ m))
|
||||||
|
_↑ʳⁱ_ n l = List.map (n Fin.↑ʳ_) l
|
||||||
|
|
||||||
|
_↑ˡᵉ_ : ∀ {n} → List (Fin n × Fin n) → ∀ m → List (Fin (n Nat.+ m) × Fin (n Nat.+ m))
|
||||||
|
_↑ˡᵉ_ l m = List.map (_↑ˡ m) l
|
||||||
|
|
||||||
|
_↑ʳᵉ_ : ∀ {m} n → List (Fin m × Fin m) → List (Fin (n Nat.+ m) × Fin (n Nat.+ m))
|
||||||
|
_↑ʳᵉ_ n l = List.map (n ↑ʳ_) l
|
||||||
|
|
||||||
|
infixr 5 _∙_
|
||||||
|
_∙_ : Graph → Graph → Graph
|
||||||
|
_∙_ g₁ g₂ = record
|
||||||
|
{ size = Graph.size g₁ Nat.+ Graph.size g₂
|
||||||
|
; nodes = Graph.nodes g₁ ++ Graph.nodes g₂
|
||||||
|
; edges = (Graph.edges g₁ ↑ˡᵉ Graph.size g₂) List.++
|
||||||
|
(Graph.size g₁ ↑ʳᵉ Graph.edges g₂)
|
||||||
|
; inputs = (Graph.inputs g₁ ↑ˡⁱ Graph.size g₂) List.++
|
||||||
|
(Graph.size g₁ ↑ʳⁱ Graph.inputs g₂)
|
||||||
|
; outputs = (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂) List.++
|
||||||
|
(Graph.size g₁ ↑ʳⁱ Graph.outputs g₂)
|
||||||
|
}
|
||||||
|
|
||||||
|
infixr 5 _↦_
|
||||||
|
_↦_ : Graph → Graph → Graph
|
||||||
|
_↦_ g₁ g₂ = record
|
||||||
|
{ size = Graph.size g₁ Nat.+ Graph.size g₂
|
||||||
|
; nodes = Graph.nodes g₁ ++ Graph.nodes g₂
|
||||||
|
; edges = (Graph.edges g₁ ↑ˡᵉ Graph.size g₂) List.++
|
||||||
|
(Graph.size g₁ ↑ʳᵉ Graph.edges g₂) List.++
|
||||||
|
(List.cartesianProduct (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂)
|
||||||
|
(Graph.size g₁ ↑ʳⁱ Graph.inputs g₂))
|
||||||
|
; inputs = Graph.inputs g₁ ↑ˡⁱ Graph.size g₂
|
||||||
|
; outputs = Graph.size g₁ ↑ʳⁱ Graph.outputs g₂
|
||||||
|
}
|
||||||
|
|
||||||
|
loop : Graph → Graph
|
||||||
|
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.++
|
||||||
|
((suc zero , zero) ∷ (zero , suc zero) ∷ [])
|
||||||
|
; inputs = zero ∷ []
|
||||||
|
; outputs = (suc zero) ∷ []
|
||||||
|
}
|
||||||
|
|
||||||
|
infixr 5 _skipto_
|
||||||
|
_skipto_ : Graph → Graph → Graph
|
||||||
|
_skipto_ g₁ g₂ = record (g₁ ∙ g₂)
|
||||||
|
{ edges = Graph.edges (g₁ ∙ g₂) List.++
|
||||||
|
(List.cartesianProduct (Graph.inputs g₁ ↑ˡⁱ Graph.size g₂)
|
||||||
|
(Graph.size g₁ ↑ʳⁱ Graph.inputs g₂))
|
||||||
|
; inputs = Graph.inputs g₁ ↑ˡⁱ Graph.size g₂
|
||||||
|
; outputs = Graph.size g₁ ↑ʳⁱ Graph.inputs g₂
|
||||||
|
}
|
||||||
|
|
||||||
|
_[_] : ∀ (g : Graph) → Graph.Index g → List BasicStmt
|
||||||
|
_[_] g idx = lookup (Graph.nodes g) idx
|
||||||
|
|
||||||
|
singleton : List BasicStmt → Graph
|
||||||
|
singleton bss = record
|
||||||
|
{ size = 1
|
||||||
|
; nodes = bss ∷ []
|
||||||
|
; edges = []
|
||||||
|
; inputs = zero ∷ []
|
||||||
|
; outputs = zero ∷ []
|
||||||
|
}
|
||||||
|
|
||||||
|
wrap : Graph → Graph
|
||||||
|
wrap g = singleton [] ↦ g ↦ singleton []
|
||||||
|
|
||||||
|
buildCfg : Stmt → Graph
|
||||||
|
buildCfg ⟨ bs₁ ⟩ = singleton (bs₁ ∷ [])
|
||||||
|
buildCfg (s₁ then s₂) = buildCfg s₁ ↦ buildCfg s₂
|
||||||
|
buildCfg (if _ then s₁ else s₂) = buildCfg s₁ ∙ buildCfg s₂
|
||||||
|
buildCfg (while _ repeat s) = loop (buildCfg s)
|
||||||
|
|
||||||
|
private
|
||||||
|
z≢sf : ∀ {n : ℕ} (f : Fin n) → ¬ (zero ≡ suc f)
|
||||||
|
z≢sf f ()
|
||||||
|
|
||||||
|
z≢mapsfs : ∀ {n : ℕ} (fs : List (Fin n)) → All (λ sf → ¬ zero ≡ sf) (List.map suc fs)
|
||||||
|
z≢mapsfs [] = []
|
||||||
|
z≢mapsfs (f ∷ fs') = z≢sf f ∷ z≢mapsfs fs'
|
||||||
|
|
||||||
|
finValues : ∀ (n : ℕ) → Σ (List (Fin n)) Unique
|
||||||
|
finValues 0 = ([] , Utils.empty)
|
||||||
|
finValues (suc n') =
|
||||||
|
let
|
||||||
|
(inds' , unids') = finValues n'
|
||||||
|
in
|
||||||
|
( zero ∷ List.map suc inds'
|
||||||
|
, push (z≢mapsfs inds') (Unique-map suc suc-injective unids')
|
||||||
|
)
|
||||||
|
|
||||||
|
finValues-complete : ∀ (n : ℕ) (f : Fin n) → f ListMem.∈ (proj₁ (finValues n))
|
||||||
|
finValues-complete (suc n') zero = RelAny.here refl
|
||||||
|
finValues-complete (suc n') (suc f') = RelAny.there (x∈xs⇒fx∈fxs suc (finValues-complete n' f'))
|
||||||
|
|
||||||
|
module _ (g : Graph) where
|
||||||
|
open import Data.Product.Properties as ProdProp using ()
|
||||||
|
private _≟_ = ProdProp.≡-dec (FinProp._≟_ {Graph.size g})
|
||||||
|
(FinProp._≟_ {Graph.size g})
|
||||||
|
|
||||||
|
open import Data.List.Membership.DecPropositional (_≟_) using (_∈?_)
|
||||||
|
|
||||||
|
indices : List (Graph.Index g)
|
||||||
|
indices = proj₁ (finValues (Graph.size g))
|
||||||
|
|
||||||
|
indices-complete : ∀ (idx : (Graph.Index g)) → idx ListMem.∈ indices
|
||||||
|
indices-complete = finValues-complete (Graph.size g)
|
||||||
|
|
||||||
|
indices-Unique : Unique indices
|
||||||
|
indices-Unique = proj₂ (finValues (Graph.size g))
|
||||||
|
|
||||||
|
predecessors : (Graph.Index g) → List (Graph.Index g)
|
||||||
|
predecessors idx = List.filter (λ idx' → (idx' , idx) ∈? (Graph.edges g)) indices
|
||||||
|
|
||||||
|
edge⇒predecessor : ∀ {idx₁ idx₂ : Graph.Index g} → (idx₁ , idx₂) ListMem.∈ (Graph.edges g) →
|
||||||
|
idx₁ ListMem.∈ (predecessors idx₂)
|
||||||
|
edge⇒predecessor {idx₁} {idx₂} idx₁,idx₂∈es =
|
||||||
|
∈-filter⁺ (λ idx' → (idx' , idx₂) ∈? (Graph.edges g))
|
||||||
|
(indices-complete idx₁) idx₁,idx₂∈es
|
||||||
|
|
||||||
|
predecessor⇒edge : ∀ {idx₁ idx₂ : Graph.Index g} → idx₁ ListMem.∈ (predecessors idx₂) →
|
||||||
|
(idx₁ , idx₂) ListMem.∈ (Graph.edges g)
|
||||||
|
predecessor⇒edge {idx₁} {idx₂} idx₁∈pred =
|
||||||
|
proj₂ (∈-filter⁻ (λ idx' → (idx' , idx₂) ∈? (Graph.edges g)) {v = idx₁} {xs = indices} idx₁∈pred )
|
||||||
296
Language/Properties.agda
Normal file
296
Language/Properties.agda
Normal file
@@ -0,0 +1,296 @@
|
|||||||
|
module Language.Properties where
|
||||||
|
|
||||||
|
open import Language.Base
|
||||||
|
open import Language.Semantics
|
||||||
|
open import Language.Graphs
|
||||||
|
open import Language.Traces
|
||||||
|
|
||||||
|
open import Data.Fin as Fin using (suc; zero)
|
||||||
|
open import Data.Fin.Properties as FinProp using (suc-injective)
|
||||||
|
open import Data.List as List using (List; _∷_; [])
|
||||||
|
open import Data.List.Properties using (filter-none)
|
||||||
|
open import Data.List.Relation.Unary.Any using (here; there)
|
||||||
|
open import Data.List.Relation.Unary.All using (All; []; _∷_; map; tabulate)
|
||||||
|
open import Data.List.Membership.Propositional as ListMem using ()
|
||||||
|
open import Data.List.Membership.Propositional.Properties as ListMemProp using ()
|
||||||
|
open import Data.Nat as Nat using ()
|
||||||
|
open import Data.Product using (Σ; _,_; _×_; proj₂)
|
||||||
|
open import Data.Product.Properties as ProdProp using ()
|
||||||
|
open import Data.Sum using (inj₁; inj₂)
|
||||||
|
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; cong)
|
||||||
|
open import Relation.Nullary using (¬_)
|
||||||
|
|
||||||
|
open import Utils using (x∈xs⇒fx∈fxs; ∈-cartesianProduct; concat-∈)
|
||||||
|
|
||||||
|
-- All of the below helpers are to reason about what edges aren't included
|
||||||
|
-- when combinings graphs. The currenty most important use for this is proving
|
||||||
|
-- that the entry node has no incoming edges.
|
||||||
|
--
|
||||||
|
-- -------------- Begin ugly code to make this work ----------------
|
||||||
|
↑-≢ : ∀ {n m} (f₁ : Fin.Fin n) (f₂ : Fin.Fin m) → ¬ (f₁ Fin.↑ˡ m) ≡ (n Fin.↑ʳ f₂)
|
||||||
|
↑-≢ zero f₂ ()
|
||||||
|
↑-≢ (suc f₁') f₂ f₁≡f₂ = ↑-≢ f₁' f₂ (suc-injective f₁≡f₂)
|
||||||
|
|
||||||
|
idx→f∉↑ʳᵉ : ∀ {n m} (idx : Fin.Fin (n Nat.+ m)) (f : Fin.Fin n) (es₂ : List (Fin.Fin m × Fin.Fin m)) → ¬ (idx , f Fin.↑ˡ m) ListMem.∈ (n ↑ʳᵉ es₂)
|
||||||
|
idx→f∉↑ʳᵉ idx f ((idx₁ , idx₂) ∷ es') (here idx,f≡idx₁,idx₂) = ↑-≢ f idx₂ (cong proj₂ idx,f≡idx₁,idx₂)
|
||||||
|
idx→f∉↑ʳᵉ idx f (_ ∷ es₂') (there idx→f∈es₂') = idx→f∉↑ʳᵉ idx f es₂' idx→f∈es₂'
|
||||||
|
|
||||||
|
idx→f∉pair : ∀ {n m} (idx idx' : Fin.Fin (n Nat.+ m)) (f : Fin.Fin n) (inputs₂ : List (Fin.Fin m)) → ¬ (idx , f Fin.↑ˡ m) ListMem.∈ (List.map (idx' ,_) (n ↑ʳⁱ inputs₂))
|
||||||
|
idx→f∉pair idx idx' f [] ()
|
||||||
|
idx→f∉pair idx idx' f (input ∷ inputs') (here idx,f≡idx',input) = ↑-≢ f input (cong proj₂ idx,f≡idx',input)
|
||||||
|
idx→f∉pair idx idx' f (_ ∷ inputs₂') (there idx,f∈inputs₂') = idx→f∉pair idx idx' f inputs₂' idx,f∈inputs₂'
|
||||||
|
|
||||||
|
idx→f∉cart : ∀ {n m} (idx : Fin.Fin (n Nat.+ m)) (f : Fin.Fin n) (outputs₁ : List (Fin.Fin n)) (inputs₂ : List (Fin.Fin m)) → ¬ (idx , f Fin.↑ˡ m) ListMem.∈ (List.cartesianProduct (outputs₁ ↑ˡⁱ m) (n ↑ʳⁱ inputs₂))
|
||||||
|
idx→f∉cart idx f [] inputs₂ ()
|
||||||
|
idx→f∉cart {n} {m} idx f (output ∷ outputs₁') inputs₂ idx,f∈pair++cart
|
||||||
|
with ListMemProp.∈-++⁻ (List.map (output Fin.↑ˡ m ,_) (n ↑ʳⁱ inputs₂)) idx,f∈pair++cart
|
||||||
|
... | inj₁ idx,f∈pair = idx→f∉pair idx (output Fin.↑ˡ m) f inputs₂ idx,f∈pair
|
||||||
|
... | inj₂ idx,f∈cart = idx→f∉cart idx f outputs₁' inputs₂ idx,f∈cart
|
||||||
|
|
||||||
|
help : let g₁ = singleton [] in
|
||||||
|
∀ (g₂ : Graph) (idx₁ : Graph.Index g₁) (idx : Graph.Index (g₁ ↦ g₂)) →
|
||||||
|
¬ (idx , idx₁ Fin.↑ˡ Graph.size g₂) ListMem.∈ ((Graph.size g₁ ↑ʳᵉ Graph.edges g₂) List.++
|
||||||
|
(List.cartesianProduct (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂)
|
||||||
|
(Graph.size g₁ ↑ʳⁱ Graph.inputs g₂)))
|
||||||
|
help g₂ idx₁ idx idx,idx₁∈g
|
||||||
|
with ListMemProp.∈-++⁻ (Graph.size (singleton []) ↑ʳᵉ Graph.edges g₂) idx,idx₁∈g
|
||||||
|
... | inj₁ idx,idx₁∈edges₂ = idx→f∉↑ʳᵉ idx idx₁ (Graph.edges g₂) idx,idx₁∈edges₂
|
||||||
|
... | inj₂ idx,idx₁∈cart = idx→f∉cart idx idx₁ (Graph.outputs (singleton [])) (Graph.inputs g₂) idx,idx₁∈cart
|
||||||
|
|
||||||
|
helpAll : let g₁ = singleton [] in
|
||||||
|
∀ (g₂ : Graph) (idx₁ : Graph.Index g₁) →
|
||||||
|
All (λ idx → ¬ (idx , idx₁ Fin.↑ˡ Graph.size g₂) ListMem.∈ ((Graph.size g₁ ↑ʳᵉ Graph.edges g₂) List.++
|
||||||
|
(List.cartesianProduct (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂)
|
||||||
|
(Graph.size g₁ ↑ʳⁱ Graph.inputs g₂)))) (indices (g₁ ↦ g₂))
|
||||||
|
helpAll g₂ idx₁ = tabulate (λ {idx} _ → help g₂ idx₁ idx)
|
||||||
|
|
||||||
|
module _ (g : Graph) where
|
||||||
|
wrap-preds-∅ : (idx : Graph.Index (wrap g)) →
|
||||||
|
idx ListMem.∈ Graph.inputs (wrap g) → predecessors (wrap g) idx ≡ []
|
||||||
|
wrap-preds-∅ zero (here refl) =
|
||||||
|
filter-none (λ idx' → (idx' , zero) ∈?
|
||||||
|
(Graph.edges (wrap g)))
|
||||||
|
(helpAll (g ↦ singleton []) zero)
|
||||||
|
where open import Data.List.Membership.DecPropositional (ProdProp.≡-dec (FinProp._≟_ {Graph.size (wrap g)}) (FinProp._≟_ {Graph.size (wrap g)})) using (_∈?_)
|
||||||
|
-- -------------- End ugly code to make this work ----------------
|
||||||
|
|
||||||
|
|
||||||
|
module _ (g : Graph) where
|
||||||
|
wrap-input : Σ (Graph.Index (wrap g)) (λ idx → Graph.inputs (wrap g) ≡ idx ∷ [])
|
||||||
|
wrap-input = (_ , refl)
|
||||||
|
|
||||||
|
wrap-output : Σ (Graph.Index (wrap g)) (λ idx → Graph.outputs (wrap g) ≡ idx ∷ [])
|
||||||
|
wrap-output = (_ , refl)
|
||||||
|
|
||||||
|
Trace-∙ˡ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} →
|
||||||
|
Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ →
|
||||||
|
Trace {g₁ ∙ g₂} (idx₁ Fin.↑ˡ Graph.size g₂) (idx₂ Fin.↑ˡ Graph.size g₂) ρ₁ ρ₂
|
||||||
|
Trace-∙ˡ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂)
|
||||||
|
rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-single ρ₁⇒ρ₂
|
||||||
|
Trace-∙ˡ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
|
||||||
|
rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (_↑ˡ Graph.size g₂) idx₁→idx))
|
||||||
|
(Trace-∙ˡ tr')
|
||||||
|
|
||||||
|
Trace-∙ʳ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₂} {ρ₁ ρ₂ : Env} →
|
||||||
|
Trace {g₂} idx₁ idx₂ ρ₁ ρ₂ →
|
||||||
|
Trace {g₁ ∙ g₂} (Graph.size g₁ Fin.↑ʳ idx₁) (Graph.size g₁ Fin.↑ʳ idx₂) ρ₁ ρ₂
|
||||||
|
Trace-∙ʳ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂)
|
||||||
|
rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-single ρ₁⇒ρ₂
|
||||||
|
Trace-∙ʳ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
|
||||||
|
rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ʳ _ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx))
|
||||||
|
(Trace-∙ʳ tr')
|
||||||
|
|
||||||
|
EndToEndTrace-∙ˡ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ : Env} →
|
||||||
|
EndToEndTrace {g₁} ρ₁ ρ₂ →
|
||||||
|
EndToEndTrace {g₁ ∙ g₂} ρ₁ ρ₂
|
||||||
|
EndToEndTrace-∙ˡ {g₁} {g₂} etr = record
|
||||||
|
{ idx₁ = EndToEndTrace.idx₁ etr Fin.↑ˡ Graph.size g₂
|
||||||
|
; idx₁∈inputs = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂)
|
||||||
|
(EndToEndTrace.idx₁∈inputs etr))
|
||||||
|
; idx₂ = EndToEndTrace.idx₂ etr Fin.↑ˡ Graph.size g₂
|
||||||
|
; idx₂∈outputs = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂)
|
||||||
|
(EndToEndTrace.idx₂∈outputs etr))
|
||||||
|
; trace = Trace-∙ˡ (EndToEndTrace.trace etr)
|
||||||
|
}
|
||||||
|
|
||||||
|
EndToEndTrace-∙ʳ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ : Env} →
|
||||||
|
EndToEndTrace {g₂} ρ₁ ρ₂ →
|
||||||
|
EndToEndTrace {g₁ ∙ g₂} ρ₁ ρ₂
|
||||||
|
EndToEndTrace-∙ʳ {g₁} {g₂} etr = record
|
||||||
|
{ idx₁ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₁ etr
|
||||||
|
; idx₁∈inputs = ListMemProp.∈-++⁺ʳ (Graph.inputs g₁ ↑ˡⁱ Graph.size g₂)
|
||||||
|
((x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_)
|
||||||
|
(EndToEndTrace.idx₁∈inputs etr)))
|
||||||
|
; idx₂ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₂ etr
|
||||||
|
; idx₂∈outputs = ListMemProp.∈-++⁺ʳ (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂)
|
||||||
|
((x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_)
|
||||||
|
(EndToEndTrace.idx₂∈outputs etr)))
|
||||||
|
|
||||||
|
; trace = Trace-∙ʳ (EndToEndTrace.trace etr)
|
||||||
|
}
|
||||||
|
|
||||||
|
Trace-↦ˡ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} →
|
||||||
|
Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ →
|
||||||
|
Trace {g₁ ↦ g₂} (idx₁ Fin.↑ˡ Graph.size g₂) (idx₂ Fin.↑ˡ Graph.size g₂) ρ₁ ρ₂
|
||||||
|
Trace-↦ˡ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂)
|
||||||
|
rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-single ρ₁⇒ρ₂
|
||||||
|
Trace-↦ˡ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
|
||||||
|
rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (_↑ˡ Graph.size g₂) idx₁→idx))
|
||||||
|
(Trace-↦ˡ tr')
|
||||||
|
|
||||||
|
Trace-↦ʳ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₂} {ρ₁ ρ₂ : Env} →
|
||||||
|
Trace {g₂} idx₁ idx₂ ρ₁ ρ₂ →
|
||||||
|
Trace {g₁ ↦ g₂} (Graph.size g₁ Fin.↑ʳ idx₁) (Graph.size g₁ Fin.↑ʳ idx₂) ρ₁ ρ₂
|
||||||
|
Trace-↦ʳ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂)
|
||||||
|
rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-single ρ₁⇒ρ₂
|
||||||
|
Trace-↦ʳ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr')
|
||||||
|
rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) =
|
||||||
|
Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ʳ (Graph.edges g₁ ↑ˡᵉ Graph.size g₂)
|
||||||
|
(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} (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 {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} (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₂
|
||||||
|
}
|
||||||
|
|
||||||
|
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} →
|
||||||
|
EndToEndTrace {g₁} ρ₁ ρ₂ → EndToEndTrace {g₂} ρ₂ ρ₃ →
|
||||||
|
EndToEndTrace {g₁ ↦ g₂} ρ₁ ρ₃
|
||||||
|
_++_ {g₁} {g₂} etr₁ etr₂
|
||||||
|
= record
|
||||||
|
{ idx₁ = EndToEndTrace.idx₁ etr₁ Fin.↑ˡ Graph.size g₂
|
||||||
|
; idx₁∈inputs = x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂) (EndToEndTrace.idx₁∈inputs etr₁)
|
||||||
|
; idx₂ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₂ etr₂
|
||||||
|
; idx₂∈outputs = x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_) (EndToEndTrace.idx₂∈outputs etr₂)
|
||||||
|
; trace =
|
||||||
|
let
|
||||||
|
o∈tr₁ = x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂) (EndToEndTrace.idx₂∈outputs etr₁)
|
||||||
|
i∈tr₂ = x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_) (EndToEndTrace.idx₁∈inputs etr₂)
|
||||||
|
oi∈es = ListMemProp.∈-++⁺ʳ (Graph.edges g₁ ↑ˡᵉ Graph.size g₂)
|
||||||
|
(ListMemProp.∈-++⁺ʳ (Graph.size g₁ ↑ʳᵉ Graph.edges g₂)
|
||||||
|
(∈-cartesianProduct o∈tr₁ i∈tr₂))
|
||||||
|
in
|
||||||
|
(Trace-↦ˡ {g₁} {g₂} (EndToEndTrace.trace etr₁)) ++⟨ oi∈es ⟩
|
||||||
|
(Trace-↦ʳ {g₁} {g₂} (EndToEndTrace.trace etr₂))
|
||||||
|
}
|
||||||
|
|
||||||
|
EndToEndTrace-singleton : ∀ {bss : List BasicStmt} {ρ₁ ρ₂ : Env} →
|
||||||
|
ρ₁ , bss ⇒ᵇˢ ρ₂ → EndToEndTrace {singleton bss} ρ₁ ρ₂
|
||||||
|
EndToEndTrace-singleton ρ₁⇒ρ₂ = record
|
||||||
|
{ idx₁ = zero
|
||||||
|
; idx₁∈inputs = here refl
|
||||||
|
; idx₂ = zero
|
||||||
|
; idx₂∈outputs = here refl
|
||||||
|
; trace = Trace-single ρ₁⇒ρ₂
|
||||||
|
}
|
||||||
|
|
||||||
|
EndToEndTrace-singleton[] : ∀ (ρ : Env) → EndToEndTrace {singleton []} ρ ρ
|
||||||
|
EndToEndTrace-singleton[] env = EndToEndTrace-singleton []
|
||||||
|
|
||||||
|
EndToEndTrace-wrap : ∀ {g : Graph} {ρ₁ ρ₂ : Env} →
|
||||||
|
EndToEndTrace {g} ρ₁ ρ₂ → EndToEndTrace {wrap g} ρ₁ ρ₂
|
||||||
|
EndToEndTrace-wrap {g} {ρ₁} {ρ₂} etr = EndToEndTrace-singleton[] ρ₁ ++ etr ++ EndToEndTrace-singleton[] ρ₂
|
||||||
|
|
||||||
|
buildCfg-sufficient : ∀ {s : Stmt} {ρ₁ ρ₂ : Env} → ρ₁ , s ⇒ˢ ρ₂ →
|
||||||
|
EndToEndTrace {buildCfg s} ρ₁ ρ₂
|
||||||
|
buildCfg-sufficient (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ₂) =
|
||||||
|
EndToEndTrace-singleton (ρ₁,bs⇒ρ₂ ∷ [])
|
||||||
|
buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ₂ ρ₂,s₂⇒ρ₃) =
|
||||||
|
buildCfg-sufficient ρ₁,s₁⇒ρ₂ ++ buildCfg-sufficient ρ₂,s₂⇒ρ₃
|
||||||
|
buildCfg-sufficient (⇒ˢ-if-true ρ₁ ρ₂ _ _ s₁ s₂ _ _ ρ₁,s₁⇒ρ₂) =
|
||||||
|
EndToEndTrace-∙ˡ (buildCfg-sufficient ρ₁,s₁⇒ρ₂)
|
||||||
|
buildCfg-sufficient (⇒ˢ-if-false ρ₁ ρ₂ _ s₁ s₂ _ ρ₁,s₂⇒ρ₂) =
|
||||||
|
EndToEndTrace-∙ʳ {buildCfg s₁} (buildCfg-sufficient ρ₁,s₂⇒ρ₂)
|
||||||
|
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-loop⁰ {buildCfg s} {ρ}
|
||||||
73
Language/Semantics.agda
Normal file
73
Language/Semantics.agda
Normal file
@@ -0,0 +1,73 @@
|
|||||||
|
module Language.Semantics where
|
||||||
|
|
||||||
|
open import Language.Base
|
||||||
|
|
||||||
|
open import Agda.Primitive using (lsuc)
|
||||||
|
open import Data.Integer using (ℤ; +_) renaming (_+_ to _+ᶻ_; _-_ to _-ᶻ_)
|
||||||
|
open import Data.Product using (_×_; _,_)
|
||||||
|
open import Data.String using (String)
|
||||||
|
open import Data.List as List using (List)
|
||||||
|
open import Data.Nat using (ℕ)
|
||||||
|
open import Relation.Nullary using (¬_)
|
||||||
|
open import Relation.Binary.PropositionalEquality using (_≡_)
|
||||||
|
|
||||||
|
open import Lattice
|
||||||
|
open import Utils using (_⇒_; _∧_; _∨_)
|
||||||
|
|
||||||
|
data Value : Set where
|
||||||
|
↑ᶻ : ℤ → Value
|
||||||
|
|
||||||
|
Env : Set
|
||||||
|
Env = List (String × Value)
|
||||||
|
|
||||||
|
data _∈_ : (String × Value) → Env → Set where
|
||||||
|
here : ∀ (s : String) (v : Value) (ρ : Env) → (s , v) ∈ ((s , v) List.∷ ρ)
|
||||||
|
there : ∀ (s s' : String) (v v' : Value) (ρ : Env) → ¬ (s ≡ s') → (s , v) ∈ ρ → (s , v) ∈ ((s' , v') List.∷ ρ)
|
||||||
|
|
||||||
|
data _,_⇒ᵉ_ : Env → Expr → Value → Set where
|
||||||
|
⇒ᵉ-ℕ : ∀ (ρ : Env) (n : ℕ) → ρ , (# n) ⇒ᵉ (↑ᶻ (+ n))
|
||||||
|
⇒ᵉ-Var : ∀ (ρ : Env) (x : String) (v : Value) → (x , v) ∈ ρ → ρ , (` x) ⇒ᵉ v
|
||||||
|
⇒ᵉ-+ : ∀ (ρ : Env) (e₁ e₂ : Expr) (z₁ z₂ : ℤ) →
|
||||||
|
ρ , e₁ ⇒ᵉ (↑ᶻ z₁) → ρ , e₂ ⇒ᵉ (↑ᶻ z₂) →
|
||||||
|
ρ , (e₁ + e₂) ⇒ᵉ (↑ᶻ (z₁ +ᶻ z₂))
|
||||||
|
⇒ᵉ-- : ∀ (ρ : Env) (e₁ e₂ : Expr) (z₁ z₂ : ℤ) →
|
||||||
|
ρ , e₁ ⇒ᵉ (↑ᶻ z₁) → ρ , e₂ ⇒ᵉ (↑ᶻ z₂) →
|
||||||
|
ρ , (e₁ - e₂) ⇒ᵉ (↑ᶻ (z₁ -ᶻ z₂))
|
||||||
|
|
||||||
|
data _,_⇒ᵇ_ : Env → BasicStmt → Env → Set where
|
||||||
|
⇒ᵇ-noop : ∀ (ρ : Env) → ρ , noop ⇒ᵇ ρ
|
||||||
|
⇒ᵇ-← : ∀ (ρ : Env) (x : String) (e : Expr) (v : Value) →
|
||||||
|
ρ , e ⇒ᵉ v → ρ , (x ← e) ⇒ᵇ ((x , v) List.∷ ρ)
|
||||||
|
|
||||||
|
data _,_⇒ᵇˢ_ : Env → List BasicStmt → Env → Set where
|
||||||
|
[] : ∀ {ρ : Env} → ρ , List.[] ⇒ᵇˢ ρ
|
||||||
|
_∷_ : ∀ {ρ₁ ρ₂ ρ₃ : Env} {bs : BasicStmt} {bss : List BasicStmt} →
|
||||||
|
ρ₁ , bs ⇒ᵇ ρ₂ → ρ₂ , bss ⇒ᵇˢ ρ₃ → ρ₁ , (bs List.∷ bss) ⇒ᵇˢ ρ₃
|
||||||
|
|
||||||
|
data _,_⇒ˢ_ : Env → Stmt → Env → Set where
|
||||||
|
⇒ˢ-⟨⟩ : ∀ (ρ₁ ρ₂ : Env) (bs : BasicStmt) →
|
||||||
|
ρ₁ , bs ⇒ᵇ ρ₂ → ρ₁ , ⟨ bs ⟩ ⇒ˢ ρ₂
|
||||||
|
⇒ˢ-then : ∀ (ρ₁ ρ₂ ρ₃ : Env) (s₁ s₂ : Stmt) →
|
||||||
|
ρ₁ , s₁ ⇒ˢ ρ₂ → ρ₂ , s₂ ⇒ˢ ρ₃ →
|
||||||
|
ρ₁ , (s₁ then s₂) ⇒ˢ ρ₃
|
||||||
|
⇒ˢ-if-true : ∀ (ρ₁ ρ₂ : Env) (e : Expr) (z : ℤ) (s₁ s₂ : Stmt) →
|
||||||
|
ρ₁ , e ⇒ᵉ (↑ᶻ z) → ¬ z ≡ (+ 0) → ρ₁ , s₁ ⇒ˢ ρ₂ →
|
||||||
|
ρ₁ , (if e then s₁ else s₂) ⇒ˢ ρ₂
|
||||||
|
⇒ˢ-if-false : ∀ (ρ₁ ρ₂ : Env) (e : Expr) (s₁ s₂ : Stmt) →
|
||||||
|
ρ₁ , e ⇒ᵉ (↑ᶻ (+ 0)) → ρ₁ , s₂ ⇒ˢ ρ₂ →
|
||||||
|
ρ₁ , (if e then s₁ else s₂) ⇒ˢ ρ₂
|
||||||
|
⇒ˢ-while-true : ∀ (ρ₁ ρ₂ ρ₃ : Env) (e : Expr) (z : ℤ) (s : Stmt) →
|
||||||
|
ρ₁ , e ⇒ᵉ (↑ᶻ z) → ¬ z ≡ (+ 0) → ρ₁ , s ⇒ˢ ρ₂ → ρ₂ , (while e repeat s) ⇒ˢ ρ₃ →
|
||||||
|
ρ₁ , (while e repeat s) ⇒ˢ ρ₃
|
||||||
|
⇒ˢ-while-false : ∀ (ρ : Env) (e : Expr) (s : Stmt) →
|
||||||
|
ρ , e ⇒ᵉ (↑ᶻ (+ 0)) →
|
||||||
|
ρ , (while e repeat s) ⇒ˢ ρ
|
||||||
|
|
||||||
|
record LatticeInterpretation {l} {L : Set l} {_≈_ : L → L → Set l}
|
||||||
|
{_⊔_ : L → L → L} {_⊓_ : L → L → L}
|
||||||
|
(isLattice : IsLattice L _≈_ _⊔_ _⊓_) : Set (lsuc l) where
|
||||||
|
field
|
||||||
|
⟦_⟧ : L → Value → Set
|
||||||
|
⟦⟧-respects-≈ : ∀ {l₁ l₂ : L} → l₁ ≈ l₂ → ⟦ l₁ ⟧ ⇒ ⟦ l₂ ⟧
|
||||||
|
⟦⟧-⊔-∨ : ∀ {l₁ l₂ : L} → (⟦ l₁ ⟧ ∨ ⟦ l₂ ⟧) ⇒ ⟦ l₁ ⊔ l₂ ⟧
|
||||||
|
⟦⟧-⊓-∧ : ∀ {l₁ l₂ : L} → (⟦ l₁ ⟧ ∧ ⟦ l₂ ⟧) ⇒ ⟦ l₁ ⊓ l₂ ⟧
|
||||||
36
Language/Traces.agda
Normal file
36
Language/Traces.agda
Normal file
@@ -0,0 +1,36 @@
|
|||||||
|
module Language.Traces where
|
||||||
|
|
||||||
|
open import Language.Base
|
||||||
|
open import Language.Semantics using (Env; _,_⇒ᵇˢ_)
|
||||||
|
open import Language.Graphs
|
||||||
|
|
||||||
|
open import Data.Product using (_,_)
|
||||||
|
open import Data.List.Membership.Propositional using (_∈_)
|
||||||
|
|
||||||
|
module _ {g : Graph} where
|
||||||
|
open Graph g using (Index; edges; inputs; outputs)
|
||||||
|
|
||||||
|
data Trace : Index → Index → Env → Env → Set where
|
||||||
|
Trace-single : ∀ {ρ₁ ρ₂ : Env} {idx : Index} →
|
||||||
|
ρ₁ , (g [ idx ]) ⇒ᵇˢ ρ₂ → Trace idx idx ρ₁ ρ₂
|
||||||
|
Trace-edge : ∀ {ρ₁ ρ₂ ρ₃ : Env} {idx₁ idx₂ idx₃ : Index} →
|
||||||
|
ρ₁ , (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₄ ρ₁ ρ₃
|
||||||
|
_++⟨_⟩_ (Trace-single ρ₁⇒ρ₂) idx₂→idx₃ tr = Trace-edge ρ₁⇒ρ₂ idx₂→idx₃ tr
|
||||||
|
_++⟨_⟩_ (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
|
||||||
|
|
||||||
|
idx₂ : Index
|
||||||
|
idx₂∈outputs : idx₂ ∈ outputs
|
||||||
|
|
||||||
|
trace : Trace idx₁ idx₂ ρ₁ ρ₂
|
||||||
88
Lattice.agda
88
Lattice.agda
@@ -4,15 +4,17 @@ 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₂)
|
||||||
open import Agda.Primitive using (lsuc; Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (lsuc; Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
open import Function.Definitions using (Injective)
|
open import Function.Definitions using (Injective)
|
||||||
|
|
||||||
IsDecidable : ∀ {a} {A : Set a} (R : A → A → Set a) → Set a
|
record IsDecidable {a} {A : Set a} (R : A → A → Set a) : Set a where
|
||||||
IsDecidable {a} {A} R = ∀ (a₁ a₂ : A) → Dec (R a₁ a₂)
|
field
|
||||||
|
R-dec : ∀ (a₁ a₂ : A) → Dec (R a₁ a₂)
|
||||||
|
|
||||||
module _ {a b} {A : Set a} {B : Set b}
|
module _ {a b} {A : Set a} {B : Set b}
|
||||||
(_≼₁_ : A → A → Set a) (_≼₂_ : B → B → Set b) where
|
(_≼₁_ : A → A → Set a) (_≼₂_ : B → B → Set b) where
|
||||||
@@ -20,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
|
||||||
@@ -137,7 +151,7 @@ module _ {a b} {A : Set a} {B : Set b}
|
|||||||
const-Mono : ∀ (x : B) → Monotonic _≼₁_ _≼₂_ (λ _ → x)
|
const-Mono : ∀ (x : B) → Monotonic _≼₁_ _≼₂_ (λ _ → x)
|
||||||
const-Mono x _ = ⊔₂-idemp x
|
const-Mono x _ = ⊔₂-idemp x
|
||||||
|
|
||||||
open import Data.List as List using (List; foldr; _∷_)
|
open import Data.List as List using (List; foldr; foldl; _∷_)
|
||||||
open import Utils using (Pairwise; _∷_)
|
open import Utils using (Pairwise; _∷_)
|
||||||
|
|
||||||
foldr-Mono : ∀ (l₁ l₂ : List A) (f : A → B → B) (b₁ b₂ : B) →
|
foldr-Mono : ∀ (l₁ l₂ : List A) (f : A → B → B) (b₁ b₂ : B) →
|
||||||
@@ -150,14 +164,44 @@ module _ {a b} {A : Set a} {B : Set b}
|
|||||||
≼₂-trans (f-Mono₁ (foldr f b₁ xs) x≼y)
|
≼₂-trans (f-Mono₁ (foldr f b₁ xs) x≼y)
|
||||||
(f-Mono₂ y (foldr-Mono xs ys f b₁ b₂ xs≼ys b₁≼b₂ f-Mono₁ f-Mono₂))
|
(f-Mono₂ y (foldr-Mono xs ys f b₁ b₂ xs≼ys b₁≼b₂ f-Mono₁ f-Mono₂))
|
||||||
|
|
||||||
|
foldl-Mono : ∀ (l₁ l₂ : List A) (f : B → A → B) (b₁ b₂ : B) →
|
||||||
|
Pairwise _≼₁_ l₁ l₂ → b₁ ≼₂ b₂ →
|
||||||
|
(∀ a → Monotonic _≼₂_ _≼₂_ (λ b → f b a)) →
|
||||||
|
(∀ b → Monotonic _≼₁_ _≼₂_ (f b)) →
|
||||||
|
foldl f b₁ l₁ ≼₂ foldl f b₂ l₂
|
||||||
|
foldl-Mono List.[] List.[] f b₁ b₂ _ b₁≼b₂ _ _ = b₁≼b₂
|
||||||
|
foldl-Mono (x ∷ xs) (y ∷ ys) f b₁ b₂ (x≼y ∷ xs≼ys) b₁≼b₂ f-Mono₁ f-Mono₂ =
|
||||||
|
foldl-Mono xs ys f (f b₁ x) (f b₂ y) xs≼ys (≼₂-trans (f-Mono₁ x b₁≼b₂) (f-Mono₂ b₂ x≼y)) f-Mono₁ f-Mono₂
|
||||||
|
|
||||||
|
module _ {a b} {A : Set a} {B : Set b}
|
||||||
|
{_≈₂_ : B → B → Set b} {_⊔₂_ : B → B → B}
|
||||||
|
(lB : IsSemilattice B _≈₂_ _⊔₂_) where
|
||||||
|
|
||||||
|
open IsSemilattice lB using () renaming (_≼_ to _≼₂_; ⊔-idemp to ⊔₂-idemp; ≼-trans to ≼₂-trans)
|
||||||
|
|
||||||
|
open import Data.List as List using (List; foldr; foldl; _∷_)
|
||||||
|
open import Utils using (Pairwise; _∷_)
|
||||||
|
|
||||||
|
foldr-Mono' : ∀ (l : List A) (f : A → B → B) →
|
||||||
|
(∀ a → Monotonic _≼₂_ _≼₂_ (f a)) →
|
||||||
|
Monotonic _≼₂_ _≼₂_ (λ b → foldr f b l)
|
||||||
|
foldr-Mono' List.[] f _ b₁≼b₂ = b₁≼b₂
|
||||||
|
foldr-Mono' (x ∷ xs) f f-Mono₂ b₁≼b₂ = f-Mono₂ x (foldr-Mono' xs f f-Mono₂ b₁≼b₂)
|
||||||
|
|
||||||
|
foldl-Mono' : ∀ (l : List A) (f : B → A → B) →
|
||||||
|
(∀ b → Monotonic _≼₂_ _≼₂_ (λ a → f a b)) →
|
||||||
|
Monotonic _≼₂_ _≼₂_ (λ b → foldl f b l)
|
||||||
|
foldl-Mono' List.[] f _ b₁≼b₂ = b₁≼b₂
|
||||||
|
foldl-Mono' (x ∷ xs) f f-Mono₁ b₁≼b₂ = foldl-Mono' xs f f-Mono₁ (f-Mono₁ x b₁≼b₂)
|
||||||
|
|
||||||
record IsLattice {a} (A : Set a)
|
record IsLattice {a} (A : Set a)
|
||||||
(_≈_ : A → A → Set a)
|
(_≈_ : A → A → Set a)
|
||||||
(_⊔_ : A → A → A)
|
(_⊔_ : A → A → A)
|
||||||
(_⊓_ : A → A → A) : Set a where
|
(_⊓_ : A → A → A) : Set a where
|
||||||
|
|
||||||
field
|
field
|
||||||
joinSemilattice : IsSemilattice A _≈_ _⊔_
|
{{joinSemilattice}} : IsSemilattice A _≈_ _⊔_
|
||||||
meetSemilattice : IsSemilattice A _≈_ _⊓_
|
{{meetSemilattice}} : IsSemilattice A _≈_ _⊓_
|
||||||
|
|
||||||
absorb-⊔-⊓ : (x y : A) → (x ⊔ (x ⊓ y)) ≈ x
|
absorb-⊔-⊓ : (x y : A) → (x ⊔ (x ⊓ y)) ≈ x
|
||||||
absorb-⊓-⊔ : (x y : A) → (x ⊓ (x ⊔ y)) ≈ x
|
absorb-⊓-⊔ : (x y : A) → (x ⊓ (x ⊔ y)) ≈ x
|
||||||
@@ -186,12 +230,34 @@ record IsFiniteHeightLattice {a} (A : Set a)
|
|||||||
(_⊓_ : A → A → A) : Set (lsuc a) where
|
(_⊓_ : A → A → A) : Set (lsuc a) where
|
||||||
|
|
||||||
field
|
field
|
||||||
isLattice : IsLattice A _≈_ _⊔_ _⊓_
|
{{isLattice}} : IsLattice A _≈_ _⊔_ _⊓_
|
||||||
|
|
||||||
open IsLattice isLattice public
|
open IsLattice isLattice public
|
||||||
|
|
||||||
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}
|
||||||
@@ -221,7 +287,7 @@ record Semilattice {a} (A : Set a) : Set (lsuc a) where
|
|||||||
_≈_ : A → A → Set a
|
_≈_ : A → A → Set a
|
||||||
_⊔_ : A → A → A
|
_⊔_ : A → A → A
|
||||||
|
|
||||||
isSemilattice : IsSemilattice A _≈_ _⊔_
|
{{isSemilattice}} : IsSemilattice A _≈_ _⊔_
|
||||||
|
|
||||||
open IsSemilattice isSemilattice public
|
open IsSemilattice isSemilattice public
|
||||||
|
|
||||||
@@ -232,7 +298,7 @@ record Lattice {a} (A : Set a) : Set (lsuc a) where
|
|||||||
_⊔_ : A → A → A
|
_⊔_ : A → A → A
|
||||||
_⊓_ : A → A → A
|
_⊓_ : A → A → A
|
||||||
|
|
||||||
isLattice : IsLattice A _≈_ _⊔_ _⊓_
|
{{isLattice}} : IsLattice A _≈_ _⊔_ _⊓_
|
||||||
|
|
||||||
open IsLattice isLattice public
|
open IsLattice isLattice public
|
||||||
|
|
||||||
@@ -243,6 +309,6 @@ record FiniteHeightLattice {a} (A : Set a) : Set (lsuc a) where
|
|||||||
_⊔_ : A → A → A
|
_⊔_ : A → A → A
|
||||||
_⊓_ : A → A → A
|
_⊓_ : A → A → A
|
||||||
|
|
||||||
isFiniteHeightLattice : IsFiniteHeightLattice A height _≈_ _⊔_ _⊓_
|
{{isFiniteHeightLattice}} : IsFiniteHeightLattice A height _≈_ _⊔_ _⊓_
|
||||||
|
|
||||||
open IsFiniteHeightLattice isFiniteHeightLattice public
|
open IsFiniteHeightLattice isFiniteHeightLattice public
|
||||||
|
|||||||
@@ -1,17 +1,19 @@
|
|||||||
open import Lattice
|
open import Lattice
|
||||||
open import Equivalence
|
open import Equivalence
|
||||||
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
||||||
|
open import Data.Unit using () renaming (⊤ to ⊤ᵘ)
|
||||||
|
|
||||||
module Lattice.AboveBelow {a} (A : Set a)
|
module Lattice.AboveBelow {a} (A : Set a)
|
||||||
(_≈₁_ : A → A → Set a)
|
{_≈₁_ : A → A → Set a}
|
||||||
(≈₁-equiv : IsEquivalence A _≈₁_)
|
{{≈₁-equiv : IsEquivalence A _≈₁_}}
|
||||||
(≈₁-dec : IsDecidable _≈₁_) where
|
{{≈₁-Decidable : IsDecidable _≈₁_}} (dummy : ⊤ᵘ) where
|
||||||
|
|
||||||
open import Data.Empty using (⊥-elim)
|
open import Data.Empty using (⊥-elim)
|
||||||
open import Data.Product using (_,_)
|
open import Data.Product using (_,_)
|
||||||
open import Data.Nat using (_≤_; ℕ; z≤n; s≤s; suc)
|
open import Data.Nat using (_≤_; ℕ; z≤n; s≤s; suc)
|
||||||
open import Function using (_∘_)
|
open import Function using (_∘_)
|
||||||
open import Showable using (Showable; show)
|
open import Showable using (Showable; show)
|
||||||
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
open import Relation.Binary.PropositionalEquality as Eq
|
open import Relation.Binary.PropositionalEquality as Eq
|
||||||
using (_≡_; sym; subst; refl)
|
using (_≡_; sym; subst; refl)
|
||||||
|
|
||||||
@@ -20,6 +22,8 @@ import Chain
|
|||||||
open IsEquivalence ≈₁-equiv using ()
|
open IsEquivalence ≈₁-equiv using ()
|
||||||
renaming (≈-refl to ≈₁-refl; ≈-sym to ≈₁-sym; ≈-trans to ≈₁-trans)
|
renaming (≈-refl to ≈₁-refl; ≈-sym to ≈₁-sym; ≈-trans to ≈₁-trans)
|
||||||
|
|
||||||
|
open IsDecidable ≈₁-Decidable using () renaming (R-dec to ≈₁-dec)
|
||||||
|
|
||||||
data AboveBelow : Set a where
|
data AboveBelow : Set a where
|
||||||
⊥ : AboveBelow
|
⊥ : AboveBelow
|
||||||
⊤ : AboveBelow
|
⊤ : AboveBelow
|
||||||
@@ -62,7 +66,7 @@ data _≈_ : AboveBelow → AboveBelow → Set a where
|
|||||||
; ≈-trans = ≈-trans
|
; ≈-trans = ≈-trans
|
||||||
}
|
}
|
||||||
|
|
||||||
≈-dec : IsDecidable _≈_
|
≈-dec : Decidable _≈_
|
||||||
≈-dec ⊥ ⊥ = yes ≈-⊥-⊥
|
≈-dec ⊥ ⊥ = yes ≈-⊥-⊥
|
||||||
≈-dec ⊤ ⊤ = yes ≈-⊤-⊤
|
≈-dec ⊤ ⊤ = yes ≈-⊤-⊤
|
||||||
≈-dec [ x ] [ y ]
|
≈-dec [ x ] [ y ]
|
||||||
@@ -76,6 +80,10 @@ data _≈_ : AboveBelow → AboveBelow → Set a where
|
|||||||
≈-dec [ x ] ⊥ = no λ ()
|
≈-dec [ x ] ⊥ = no λ ()
|
||||||
≈-dec [ x ] ⊤ = no λ ()
|
≈-dec [ x ] ⊤ = no λ ()
|
||||||
|
|
||||||
|
instance
|
||||||
|
≈-Decidable : IsDecidable _≈_
|
||||||
|
≈-Decidable = record { R-dec = ≈-dec }
|
||||||
|
|
||||||
-- Any object can be wrapped in an 'above below' to make it a lattice,
|
-- Any object can be wrapped in an 'above below' to make it a lattice,
|
||||||
-- since ⊤ and ⊥ are the largest and least elements, and the rest are left
|
-- since ⊤ and ⊥ are the largest and least elements, and the rest are left
|
||||||
-- unordered. That's what this module does.
|
-- unordered. That's what this module does.
|
||||||
@@ -169,14 +177,15 @@ module Plain (x : A) where
|
|||||||
⊔-idemp ⊥ = ≈-⊥-⊥
|
⊔-idemp ⊥ = ≈-⊥-⊥
|
||||||
⊔-idemp [ x ] rewrite x≈y⇒[x]⊔[y]≡[x] (≈₁-refl {x}) = ≈-refl
|
⊔-idemp [ x ] rewrite x≈y⇒[x]⊔[y]≡[x] (≈₁-refl {x}) = ≈-refl
|
||||||
|
|
||||||
isJoinSemilattice : IsSemilattice AboveBelow _≈_ _⊔_
|
instance
|
||||||
isJoinSemilattice = record
|
isJoinSemilattice : IsSemilattice AboveBelow _≈_ _⊔_
|
||||||
{ ≈-equiv = ≈-equiv
|
isJoinSemilattice = record
|
||||||
; ≈-⊔-cong = ≈-⊔-cong
|
{ ≈-equiv = ≈-equiv
|
||||||
; ⊔-assoc = ⊔-assoc
|
; ≈-⊔-cong = ≈-⊔-cong
|
||||||
; ⊔-comm = ⊔-comm
|
; ⊔-assoc = ⊔-assoc
|
||||||
; ⊔-idemp = ⊔-idemp
|
; ⊔-comm = ⊔-comm
|
||||||
}
|
; ⊔-idemp = ⊔-idemp
|
||||||
|
}
|
||||||
|
|
||||||
_⊓_ : AboveBelow → AboveBelow → AboveBelow
|
_⊓_ : AboveBelow → AboveBelow → AboveBelow
|
||||||
⊥ ⊓ x = ⊥
|
⊥ ⊓ x = ⊥
|
||||||
@@ -262,14 +271,15 @@ module Plain (x : A) where
|
|||||||
⊓-idemp ⊤ = ≈-⊤-⊤
|
⊓-idemp ⊤ = ≈-⊤-⊤
|
||||||
⊓-idemp [ x ] rewrite x≈y⇒[x]⊓[y]≡[x] (≈₁-refl {x}) = ≈-refl
|
⊓-idemp [ x ] rewrite x≈y⇒[x]⊓[y]≡[x] (≈₁-refl {x}) = ≈-refl
|
||||||
|
|
||||||
isMeetSemilattice : IsSemilattice AboveBelow _≈_ _⊓_
|
instance
|
||||||
isMeetSemilattice = record
|
isMeetSemilattice : IsSemilattice AboveBelow _≈_ _⊓_
|
||||||
{ ≈-equiv = ≈-equiv
|
isMeetSemilattice = record
|
||||||
; ≈-⊔-cong = ≈-⊓-cong
|
{ ≈-equiv = ≈-equiv
|
||||||
; ⊔-assoc = ⊓-assoc
|
; ≈-⊔-cong = ≈-⊓-cong
|
||||||
; ⊔-comm = ⊓-comm
|
; ⊔-assoc = ⊓-assoc
|
||||||
; ⊔-idemp = ⊓-idemp
|
; ⊔-comm = ⊓-comm
|
||||||
}
|
; ⊔-idemp = ⊓-idemp
|
||||||
|
}
|
||||||
|
|
||||||
absorb-⊔-⊓ : ∀ (ab₁ ab₂ : AboveBelow) → (ab₁ ⊔ (ab₁ ⊓ ab₂)) ≈ ab₁
|
absorb-⊔-⊓ : ∀ (ab₁ ab₂ : AboveBelow) → (ab₁ ⊔ (ab₁ ⊓ ab₂)) ≈ ab₁
|
||||||
absorb-⊔-⊓ ⊥ ab₂ rewrite ⊥⊓x≡⊥ ab₂ = ≈-⊥-⊥
|
absorb-⊔-⊓ ⊥ ab₂ rewrite ⊥⊓x≡⊥ ab₂ = ≈-⊥-⊥
|
||||||
@@ -294,23 +304,24 @@ module Plain (x : A) where
|
|||||||
... | no x̷≈y rewrite x̷≈y⇒[x]⊔[y]≡⊤ x̷≈y rewrite x⊓⊤≡x [ x ] = ≈-refl
|
... | no x̷≈y rewrite x̷≈y⇒[x]⊔[y]≡⊤ x̷≈y rewrite x⊓⊤≡x [ x ] = ≈-refl
|
||||||
|
|
||||||
|
|
||||||
isLattice : IsLattice AboveBelow _≈_ _⊔_ _⊓_
|
instance
|
||||||
isLattice = record
|
isLattice : IsLattice AboveBelow _≈_ _⊔_ _⊓_
|
||||||
{ joinSemilattice = isJoinSemilattice
|
isLattice = record
|
||||||
; meetSemilattice = isMeetSemilattice
|
{ joinSemilattice = isJoinSemilattice
|
||||||
; absorb-⊔-⊓ = absorb-⊔-⊓
|
; meetSemilattice = isMeetSemilattice
|
||||||
; absorb-⊓-⊔ = absorb-⊓-⊔
|
; absorb-⊔-⊓ = absorb-⊔-⊓
|
||||||
}
|
; absorb-⊓-⊔ = absorb-⊓-⊔
|
||||||
|
}
|
||||||
|
|
||||||
lattice : Lattice AboveBelow
|
lattice : Lattice AboveBelow
|
||||||
lattice = record
|
lattice = record
|
||||||
{ _≈_ = _≈_
|
{ _≈_ = _≈_
|
||||||
; _⊔_ = _⊔_
|
; _⊔_ = _⊔_
|
||||||
; _⊓_ = _⊓_
|
; _⊓_ = _⊓_
|
||||||
; 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 , λ ())
|
||||||
@@ -354,20 +365,26 @@ module Plain (x : A) where
|
|||||||
isLongest {⊥} (step {_} {[ x ]} _ (≈-lift _) (step [x]≺y y≈z c@(step _ _ _)))
|
isLongest {⊥} (step {_} {[ x ]} _ (≈-lift _) (step [x]≺y y≈z c@(step _ _ _)))
|
||||||
rewrite [x]≺y⇒y≡⊤ _ _ [x]≺y with ≈-⊤-⊤ ← y≈z = ⊥-elim (¬-Chain-⊤ c)
|
rewrite [x]≺y⇒y≡⊤ _ _ [x]≺y with ≈-⊤-⊤ ← y≈z = ⊥-elim (¬-Chain-⊤ c)
|
||||||
|
|
||||||
fixedHeight : IsLattice.FixedHeight isLattice 2
|
instance
|
||||||
fixedHeight = (((⊥ , ⊤) , longestChain) , isLongest)
|
fixedHeight : IsLattice.FixedHeight isLattice 2
|
||||||
|
fixedHeight = record
|
||||||
|
{ ⊥ = ⊥
|
||||||
|
; ⊤ = ⊤
|
||||||
|
; longestChain = longestChain
|
||||||
|
; bounded = isLongest
|
||||||
|
}
|
||||||
|
|
||||||
isFiniteHeightLattice : IsFiniteHeightLattice AboveBelow 2 _≈_ _⊔_ _⊓_
|
isFiniteHeightLattice : IsFiniteHeightLattice AboveBelow 2 _≈_ _⊔_ _⊓_
|
||||||
isFiniteHeightLattice = record
|
isFiniteHeightLattice = record
|
||||||
{ isLattice = isLattice
|
{ isLattice = isLattice
|
||||||
; fixedHeight = fixedHeight
|
; fixedHeight = fixedHeight
|
||||||
}
|
}
|
||||||
|
|
||||||
finiteHeightLattice : FiniteHeightLattice AboveBelow
|
finiteHeightLattice : FiniteHeightLattice AboveBelow
|
||||||
finiteHeightLattice = record
|
finiteHeightLattice = record
|
||||||
{ height = 2
|
{ height = 2
|
||||||
; _≈_ = _≈_
|
; _≈_ = _≈_
|
||||||
; _⊔_ = _⊔_
|
; _⊔_ = _⊔_
|
||||||
; _⊓_ = _⊓_
|
; _⊓_ = _⊓_
|
||||||
; isFiniteHeightLattice = isFiniteHeightLattice
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
||||||
}
|
}
|
||||||
|
|||||||
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 }
|
||||||
@@ -3,16 +3,28 @@ open import Relation.Binary.PropositionalEquality as Eq
|
|||||||
using (_≡_;refl; sym; trans; cong; subst)
|
using (_≡_;refl; sym; trans; cong; subst)
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
open import Data.List using (List; _∷_; [])
|
open import Data.List using (List; _∷_; [])
|
||||||
|
open import Data.Unit using (⊤)
|
||||||
|
|
||||||
module Lattice.FiniteMap {a b : Level} {A : Set a} {B : Set b}
|
module Lattice.FiniteMap (A : Set) (B : Set)
|
||||||
{_≈₂_ : B → B → Set b}
|
{_≈₂_ : B → B → Set}
|
||||||
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
||||||
(≡-dec-A : IsDecidable (_≡_ {a} {A}))
|
{{≡-Decidable-A : IsDecidable {_} {A} _≡_}}
|
||||||
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
{{lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_}} (ks : List A) where
|
||||||
|
|
||||||
open IsLattice lB using () renaming (_≼_ to _≼₂_)
|
open IsLattice lB using () renaming (_≼_ to _≼₂_)
|
||||||
open import Lattice.Map ≡-dec-A lB as Map
|
open import Lattice.Map A B _ as Map
|
||||||
using (Map; ⊔-equal-keys; ⊓-equal-keys; ∈k-dec)
|
using
|
||||||
|
( Map
|
||||||
|
; ⊔-equal-keys
|
||||||
|
; ⊓-equal-keys
|
||||||
|
; subset-impl
|
||||||
|
; Map-functional
|
||||||
|
; Expr-Provenance
|
||||||
|
; Expr-Provenance-≡
|
||||||
|
; `_; _∪_; _∩_
|
||||||
|
; in₁; in₂; bothᵘ; single
|
||||||
|
; ⊔-combines
|
||||||
|
)
|
||||||
renaming
|
renaming
|
||||||
( _≈_ to _≈ᵐ_
|
( _≈_ to _≈ᵐ_
|
||||||
; _⊔_ to _⊔ᵐ_
|
; _⊔_ to _⊔ᵐ_
|
||||||
@@ -28,27 +40,42 @@ open import Lattice.Map ≡-dec-A lB as Map
|
|||||||
; ⊓-idemp to ⊓ᵐ-idemp
|
; ⊓-idemp to ⊓ᵐ-idemp
|
||||||
; absorb-⊔-⊓ to absorb-⊔ᵐ-⊓ᵐ
|
; absorb-⊔-⊓ to absorb-⊔ᵐ-⊓ᵐ
|
||||||
; absorb-⊓-⊔ to absorb-⊓ᵐ-⊔ᵐ
|
; absorb-⊓-⊔ to absorb-⊓ᵐ-⊔ᵐ
|
||||||
; ≈-dec to ≈ᵐ-dec
|
; ≈-Decidable to ≈ᵐ-Decidable
|
||||||
; _[_] to _[_]ᵐ
|
; _[_] to _[_]ᵐ
|
||||||
|
; []-∈ to []ᵐ-∈
|
||||||
; 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]ᵐ
|
||||||
|
; m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ to m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ᵐ
|
||||||
; locate to locateᵐ
|
; locate to locateᵐ
|
||||||
; keys to keysᵐ
|
; keys to keysᵐ
|
||||||
; _updating_via_ to _updatingᵐ_via_
|
; _updating_via_ to _updatingᵐ_via_
|
||||||
; updating-via-keys-≡ to updatingᵐ-via-keys-≡
|
; updating-via-keys-≡ to updatingᵐ-via-keys-≡
|
||||||
|
; updating-via-k∈ks to updatingᵐ-via-k∈ks
|
||||||
|
; updating-via-k∈ks-≡ to updatingᵐ-via-k∈ks-≡
|
||||||
|
; updating-via-∈k-forward to updatingᵐ-via-∈k-forward
|
||||||
|
; updating-via-k∉ks-forward to updatingᵐ-via-k∉ks-forward
|
||||||
|
; updating-via-k∉ks-backward to updatingᵐ-via-k∉ks-backward
|
||||||
; f'-Monotonic to f'-Monotonicᵐ
|
; f'-Monotonic to f'-Monotonicᵐ
|
||||||
; _≼_ to _≼ᵐ_
|
; _≼_ to _≼ᵐ_
|
||||||
|
; ∈k-dec to ∈k-decᵐ
|
||||||
)
|
)
|
||||||
|
open import Data.Empty using (⊥-elim)
|
||||||
|
open import Data.List using (List; length; []; _∷_; map)
|
||||||
open import Data.List.Membership.Propositional using () renaming (_∈_ to _∈ˡ_)
|
open import Data.List.Membership.Propositional using () renaming (_∈_ to _∈ˡ_)
|
||||||
open import Data.Product using (_×_; _,_; Σ; proj₁ ; proj₂)
|
open import Data.List.Properties using (∷-injectiveʳ)
|
||||||
|
open import Data.List.Relation.Unary.All using (All)
|
||||||
|
open import Data.List.Relation.Unary.Any using (Any; here; there)
|
||||||
|
open import Data.Nat using (ℕ)
|
||||||
|
open import Data.Product using (_×_; _,_; Σ; proj₁; proj₂)
|
||||||
open import Equivalence
|
open import Equivalence
|
||||||
open import Function using (_∘_)
|
open import Function using (_∘_)
|
||||||
open import Relation.Nullary using (¬_; Dec; yes; no)
|
open import Relation.Nullary using (¬_; Dec; yes; no)
|
||||||
open import Utils using (Pairwise; _∷_; [])
|
open import Utils using (Pairwise; _∷_; []; Unique; push; empty; All¬-¬Any)
|
||||||
open import Data.Empty using (⊥-elim)
|
|
||||||
open import Showable using (Showable; show)
|
open import Showable using (Showable; show)
|
||||||
|
open import Isomorphism using (IsInverseˡ; IsInverseʳ)
|
||||||
|
open import Chain using (Height)
|
||||||
|
|
||||||
module WithKeys (ks : List A) where
|
private module WithKeys (ks : List A) where
|
||||||
FiniteMap : Set (a ⊔ℓ b)
|
FiniteMap : Set
|
||||||
FiniteMap = Σ Map (λ m → Map.keys m ≡ ks)
|
FiniteMap = Σ Map (λ m → Map.keys m ≡ ks)
|
||||||
|
|
||||||
instance
|
instance
|
||||||
@@ -56,11 +83,15 @@ module WithKeys (ks : List A) where
|
|||||||
Showable FiniteMap
|
Showable FiniteMap
|
||||||
showable = record { show = λ (m₁ , _) → show m₁ }
|
showable = record { show = λ (m₁ , _) → show m₁ }
|
||||||
|
|
||||||
_≈_ : FiniteMap → FiniteMap → Set (a ⊔ℓ b)
|
_≈_ : FiniteMap → FiniteMap → Set
|
||||||
_≈_ (m₁ , _) (m₂ , _) = m₁ ≈ᵐ m₂
|
_≈_ (m₁ , _) (m₂ , _) = m₁ ≈ᵐ m₂
|
||||||
|
|
||||||
≈₂-dec⇒≈-dec : IsDecidable _≈₂_ → IsDecidable _≈_
|
instance
|
||||||
≈₂-dec⇒≈-dec ≈₂-dec fm₁ fm₂ = ≈ᵐ-dec ≈₂-dec (proj₁ fm₁) (proj₁ fm₂)
|
≈-Decidable : {{ IsDecidable _≈₂_ }} → IsDecidable _≈_
|
||||||
|
≈-Decidable {{≈₂-Decidable}} = record
|
||||||
|
{ R-dec = λ fm₁ fm₂ → IsDecidable.R-dec (≈ᵐ-Decidable {{≈₂-Decidable}})
|
||||||
|
(proj₁ fm₁) (proj₁ fm₂)
|
||||||
|
}
|
||||||
|
|
||||||
_⊔_ : FiniteMap → FiniteMap → FiniteMap
|
_⊔_ : FiniteMap → FiniteMap → FiniteMap
|
||||||
_⊔_ (m₁ , km₁≡ks) (m₂ , km₂≡ks) =
|
_⊔_ (m₁ , km₁≡ks) (m₂ , km₂≡ks) =
|
||||||
@@ -76,12 +107,16 @@ module WithKeys (ks : List A) where
|
|||||||
km₁≡ks
|
km₁≡ks
|
||||||
)
|
)
|
||||||
|
|
||||||
_∈_ : A × B → FiniteMap → Set (a ⊔ℓ b)
|
_∈_ : A × B → FiniteMap → Set
|
||||||
_∈_ k,v (m₁ , _) = k,v ∈ˡ (proj₁ m₁)
|
_∈_ k,v (m₁ , _) = k,v ∈ˡ (proj₁ m₁)
|
||||||
|
|
||||||
_∈k_ : A → FiniteMap → Set a
|
_∈k_ : A → FiniteMap → Set
|
||||||
_∈k_ k (m₁ , _) = k ∈ˡ (keysᵐ m₁)
|
_∈k_ k (m₁ , _) = k ∈ˡ (keysᵐ m₁)
|
||||||
|
|
||||||
|
open Map using (forget) public
|
||||||
|
|
||||||
|
∈k-dec = ∈k-decᵐ
|
||||||
|
|
||||||
locate : ∀ {k : A} {fm : FiniteMap} → k ∈k fm → Σ B (λ v → (k , v) ∈ fm)
|
locate : ∀ {k : A} {fm : FiniteMap} → k ∈k fm → Σ B (λ v → (k , v) ∈ fm)
|
||||||
locate {k} {fm = (m , _)} k∈kfm = locateᵐ {k} {m} k∈kfm
|
locate {k} {fm = (m , _)} k∈kfm = locateᵐ {k} {m} k∈kfm
|
||||||
|
|
||||||
@@ -94,6 +129,10 @@ module WithKeys (ks : List A) where
|
|||||||
_[_] : FiniteMap → List A → List B
|
_[_] : FiniteMap → List A → List B
|
||||||
_[_] (m₁ , _) ks = m₁ [ ks ]ᵐ
|
_[_] (m₁ , _) ks = m₁ [ ks ]ᵐ
|
||||||
|
|
||||||
|
[]-∈ : ∀ {k : A} {v : B} {ks' : List A} (fm : FiniteMap) →
|
||||||
|
k ∈ˡ ks' → (k , v) ∈ fm → v ∈ˡ (fm [ ks' ])
|
||||||
|
[]-∈ {k} {v} {ks'} (m , _) k∈ks' k,v∈fm = []ᵐ-∈ m k,v∈fm k∈ks'
|
||||||
|
|
||||||
≈-equiv : IsEquivalence FiniteMap _≈_
|
≈-equiv : IsEquivalence FiniteMap _≈_
|
||||||
≈-equiv = record
|
≈-equiv = record
|
||||||
{ ≈-refl =
|
{ ≈-refl =
|
||||||
@@ -104,55 +143,62 @@ 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
|
||||||
|
|
||||||
isUnionSemilattice : IsSemilattice FiniteMap _≈_ _⊔_
|
instance
|
||||||
isUnionSemilattice = record
|
isUnionSemilattice : IsSemilattice FiniteMap _≈_ _⊔_
|
||||||
{ ≈-equiv = ≈-equiv
|
isUnionSemilattice = record
|
||||||
; ≈-⊔-cong =
|
{ ≈-equiv = ≈-equiv
|
||||||
λ {(m₁ , _)} {(m₂ , _)} {(m₃ , _)} {(m₄ , _)} m₁≈m₂ m₃≈m₄ →
|
; ≈-⊔-cong =
|
||||||
≈ᵐ-⊔ᵐ-cong {m₁} {m₂} {m₃} {m₄} m₁≈m₂ m₃≈m₄
|
λ {(m₁ , _)} {(m₂ , _)} {(m₃ , _)} {(m₄ , _)} m₁≈m₂ m₃≈m₄ →
|
||||||
; ⊔-assoc = λ (m₁ , _) (m₂ , _) (m₃ , _) → ⊔ᵐ-assoc m₁ m₂ m₃
|
≈ᵐ-⊔ᵐ-cong {m₁} {m₂} {m₃} {m₄} m₁≈m₂ m₃≈m₄
|
||||||
; ⊔-comm = λ (m₁ , _) (m₂ , _) → ⊔ᵐ-comm m₁ m₂
|
; ⊔-assoc = λ (m₁ , _) (m₂ , _) (m₃ , _) → ⊔ᵐ-assoc m₁ m₂ m₃
|
||||||
; ⊔-idemp = λ (m , _) → ⊔ᵐ-idemp m
|
; ⊔-comm = λ (m₁ , _) (m₂ , _) → ⊔ᵐ-comm m₁ m₂
|
||||||
}
|
; ⊔-idemp = λ (m , _) → ⊔ᵐ-idemp m
|
||||||
|
}
|
||||||
|
|
||||||
isIntersectSemilattice : IsSemilattice FiniteMap _≈_ _⊓_
|
isIntersectSemilattice : IsSemilattice FiniteMap _≈_ _⊓_
|
||||||
isIntersectSemilattice = record
|
isIntersectSemilattice = record
|
||||||
{ ≈-equiv = ≈-equiv
|
{ ≈-equiv = ≈-equiv
|
||||||
; ≈-⊔-cong =
|
; ≈-⊔-cong =
|
||||||
λ {(m₁ , _)} {(m₂ , _)} {(m₃ , _)} {(m₄ , _)} m₁≈m₂ m₃≈m₄ →
|
λ {(m₁ , _)} {(m₂ , _)} {(m₃ , _)} {(m₄ , _)} m₁≈m₂ m₃≈m₄ →
|
||||||
≈ᵐ-⊓ᵐ-cong {m₁} {m₂} {m₃} {m₄} m₁≈m₂ m₃≈m₄
|
≈ᵐ-⊓ᵐ-cong {m₁} {m₂} {m₃} {m₄} m₁≈m₂ m₃≈m₄
|
||||||
; ⊔-assoc = λ (m₁ , _) (m₂ , _) (m₃ , _) → ⊓ᵐ-assoc m₁ m₂ m₃
|
; ⊔-assoc = λ (m₁ , _) (m₂ , _) (m₃ , _) → ⊓ᵐ-assoc m₁ m₂ m₃
|
||||||
; ⊔-comm = λ (m₁ , _) (m₂ , _) → ⊓ᵐ-comm m₁ m₂
|
; ⊔-comm = λ (m₁ , _) (m₂ , _) → ⊓ᵐ-comm m₁ m₂
|
||||||
; ⊔-idemp = λ (m , _) → ⊓ᵐ-idemp m
|
; ⊔-idemp = λ (m , _) → ⊓ᵐ-idemp m
|
||||||
}
|
}
|
||||||
|
|
||||||
isLattice : IsLattice FiniteMap _≈_ _⊔_ _⊓_
|
isLattice : IsLattice FiniteMap _≈_ _⊔_ _⊓_
|
||||||
isLattice = record
|
isLattice = record
|
||||||
{ joinSemilattice = isUnionSemilattice
|
{ joinSemilattice = isUnionSemilattice
|
||||||
; meetSemilattice = isIntersectSemilattice
|
; meetSemilattice = isIntersectSemilattice
|
||||||
; absorb-⊔-⊓ = λ (m₁ , _) (m₂ , _) → absorb-⊔ᵐ-⊓ᵐ m₁ m₂
|
; absorb-⊔-⊓ = λ (m₁ , _) (m₂ , _) → absorb-⊔ᵐ-⊓ᵐ m₁ m₂
|
||||||
; absorb-⊓-⊔ = λ (m₁ , _) (m₂ , _) → absorb-⊓ᵐ-⊔ᵐ m₁ m₂
|
; absorb-⊓-⊔ = λ (m₁ , _) (m₂ , _) → absorb-⊓ᵐ-⊔ᵐ m₁ m₂
|
||||||
}
|
}
|
||||||
|
|
||||||
open IsLattice isLattice using (_≼_; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
|
lattice : Lattice FiniteMap
|
||||||
|
lattice = record
|
||||||
|
{ _≈_ = _≈_
|
||||||
|
; _⊔_ = _⊔_
|
||||||
|
; _⊓_ = _⊓_
|
||||||
|
; isLattice = isLattice
|
||||||
|
}
|
||||||
|
|
||||||
lattice : Lattice FiniteMap
|
open IsLattice isLattice using (_≼_; ⊔-idemp; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
|
||||||
lattice = record
|
|
||||||
{ _≈_ = _≈_
|
|
||||||
; _⊔_ = _⊔_
|
|
||||||
; _⊓_ = _⊓_
|
|
||||||
; isLattice = isLattice
|
|
||||||
}
|
|
||||||
|
|
||||||
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₂
|
||||||
m₁≼m₂⇒m₁[k]≼m₂[k] (m₁ , _) (m₂ , _) m₁≼m₂ k,v₁∈m₁ k,v₂∈m₂ = m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ m₁ m₂ m₁≼m₂ k,v₁∈m₁ k,v₂∈m₂
|
m₁≼m₂⇒m₁[k]≼m₂[k] (m₁ , _) (m₂ , _) m₁≼m₂ k,v₁∈m₁ k,v₂∈m₂ = m₁≼m₂⇒m₁[k]ᵐ≼m₂[k]ᵐ m₁ m₂ m₁≼m₂ k,v₁∈m₁ k,v₂∈m₂
|
||||||
|
|
||||||
|
m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ : ∀ (fm₁ fm₂ : FiniteMap) {k : A} →
|
||||||
|
fm₁ ≈ fm₂ → ∀ (k∈kfm₁ : k ∈k fm₁) (k∈kfm₂ : k ∈k fm₂) →
|
||||||
|
proj₁ (locate {fm = fm₁} k∈kfm₁) ≈₂ proj₁ (locate {fm = fm₂} k∈kfm₂)
|
||||||
|
m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ (m₁ , _) (m₂ , _) = m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ᵐ m₁ m₂
|
||||||
|
|
||||||
module GeneralizedUpdate
|
module GeneralizedUpdate
|
||||||
{l} {L : Set l}
|
{l} {L : Set l}
|
||||||
{_≈ˡ_ : L → L → Set l} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
{_≈ˡ_ : L → L → Set l} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
||||||
(lL : IsLattice L _≈ˡ_ _⊔ˡ_ _⊓ˡ_)
|
{{lL : IsLattice L _≈ˡ_ _⊔ˡ_ _⊓ˡ_}}
|
||||||
(f : L → FiniteMap) (f-Monotonic : Monotonic (IsLattice._≼_ lL) _≼_ f)
|
(f : L → FiniteMap) (f-Monotonic : Monotonic (IsLattice._≼_ lL) _≼_ f)
|
||||||
(g : A → L → B) (g-Monotonicʳ : ∀ k → Monotonic (IsLattice._≼_ lL) _≼₂_ (g k))
|
(g : A → L → B) (g-Monotonicʳ : ∀ k → Monotonic (IsLattice._≼_ lL) _≼₂_ (g k))
|
||||||
(ks : List A) where
|
(ks : List A) where
|
||||||
@@ -166,7 +212,22 @@ module WithKeys (ks : List A) where
|
|||||||
f' l = (f l) updating ks via (updater l)
|
f' l = (f l) updating ks via (updater l)
|
||||||
|
|
||||||
f'-Monotonic : Monotonic _≼ˡ_ _≼_ f'
|
f'-Monotonic : Monotonic _≼ˡ_ _≼_ f'
|
||||||
f'-Monotonic {l₁} {l₂} l₁≼l₂ = f'-Monotonicᵐ lL (proj₁ ∘ f) f-Monotonic g g-Monotonicʳ ks l₁≼l₂
|
f'-Monotonic {l₁} {l₂} l₁≼l₂ = f'-Monotonicᵐ (proj₁ ∘ f) f-Monotonic g g-Monotonicʳ ks l₁≼l₂
|
||||||
|
|
||||||
|
f'-∈k-forward : ∀ {k l} → k ∈k (f l) → k ∈k (f' l)
|
||||||
|
f'-∈k-forward {k} {l} = updatingᵐ-via-∈k-forward (proj₁ (f l)) ks (updater l)
|
||||||
|
|
||||||
|
f'-k∈ks : ∀ {k l} → k ∈ˡ ks → k ∈k (f' l) → (k , updater l k) ∈ (f' l)
|
||||||
|
f'-k∈ks {k} {l} = updatingᵐ-via-k∈ks (proj₁ (f l)) (updater l)
|
||||||
|
|
||||||
|
f'-k∈ks-≡ : ∀ {k v l} → k ∈ˡ ks → (k , v) ∈ (f' l) → v ≡ updater l k
|
||||||
|
f'-k∈ks-≡ {k} {v} {l} = updatingᵐ-via-k∈ks-≡ (proj₁ (f l)) (updater l)
|
||||||
|
|
||||||
|
f'-k∉ks-forward : ∀ {k v l} → ¬ k ∈ˡ ks → (k , v) ∈ (f l) → (k , v) ∈ (f' l)
|
||||||
|
f'-k∉ks-forward {k} {v} {l} = updatingᵐ-via-k∉ks-forward (proj₁ (f l)) (updater l)
|
||||||
|
|
||||||
|
f'-k∉ks-backward : ∀ {k v l} → ¬ k ∈ˡ ks → (k , v) ∈ (f' l) → (k , v) ∈ (f l)
|
||||||
|
f'-k∉ks-backward {k} {v} {l} = updatingᵐ-via-k∉ks-backward (proj₁ (f l)) (updater l)
|
||||||
|
|
||||||
all-equal-keys : ∀ (fm₁ fm₂ : FiniteMap) → (Map.keys (proj₁ fm₁) ≡ Map.keys (proj₁ fm₂))
|
all-equal-keys : ∀ (fm₁ fm₂ : FiniteMap) → (Map.keys (proj₁ fm₁) ≡ Map.keys (proj₁ fm₂))
|
||||||
all-equal-keys (fm₁ , km₁≡ks) (fm₂ , km₂≡ks) = trans km₁≡ks (sym km₂≡ks)
|
all-equal-keys (fm₁ , km₁≡ks) (fm₂ , km₂≡ks) = trans km₁≡ks (sym km₂≡ks)
|
||||||
@@ -182,7 +243,7 @@ module WithKeys (ks : List A) where
|
|||||||
fm₁ ≼ fm₂ → Pairwise _≼₂_ (fm₁ [ ks' ]) (fm₂ [ ks' ])
|
fm₁ ≼ fm₂ → Pairwise _≼₂_ (fm₁ [ ks' ]) (fm₂ [ ks' ])
|
||||||
m₁≼m₂⇒m₁[ks]≼m₂[ks] _ _ [] _ = []
|
m₁≼m₂⇒m₁[ks]≼m₂[ks] _ _ [] _ = []
|
||||||
m₁≼m₂⇒m₁[ks]≼m₂[ks] fm₁@(m₁ , km₁≡ks) fm₂@(m₂ , km₂≡ks) (k ∷ ks'') m₁≼m₂
|
m₁≼m₂⇒m₁[ks]≼m₂[ks] fm₁@(m₁ , km₁≡ks) fm₂@(m₂ , km₂≡ks) (k ∷ ks'') m₁≼m₂
|
||||||
with ∈k-dec k (proj₁ m₁) | ∈k-dec k (proj₁ m₂)
|
with ∈k-decᵐ k (proj₁ m₁) | ∈k-decᵐ k (proj₁ m₂)
|
||||||
... | yes k∈km₁ | yes k∈km₂ =
|
... | yes k∈km₁ | yes k∈km₂ =
|
||||||
let
|
let
|
||||||
(v₁ , k,v₁∈m₁) = locateᵐ {m = m₁} k∈km₁
|
(v₁ , k,v₁∈m₁) = locateᵐ {m = m₁} k∈km₁
|
||||||
@@ -193,4 +254,382 @@ 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₂))
|
||||||
|
|
||||||
|
_∈ᵐ_ : ∀ {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 Lattice.Unit using ()
|
||||||
|
renaming
|
||||||
|
( _≈_ to _≈ᵘ_
|
||||||
|
; _⊔_ to _⊔ᵘ_
|
||||||
|
; _⊓_ to _⊓ᵘ_
|
||||||
|
; ≈-Decidable to ≈ᵘ-Decidable
|
||||||
|
; isLattice to isLatticeᵘ
|
||||||
|
; ≈-equiv to ≈ᵘ-equiv
|
||||||
|
; fixedHeight to fixedHeightᵘ
|
||||||
|
)
|
||||||
|
open import Lattice.IterProd B ⊤ _
|
||||||
|
as IP
|
||||||
|
using (IterProd)
|
||||||
|
open IsLattice lB using ()
|
||||||
|
renaming
|
||||||
|
( ≈-trans to ≈₂-trans
|
||||||
|
; ≈-sym to ≈₂-sym
|
||||||
|
; FixedHeight to FixedHeight₂
|
||||||
|
)
|
||||||
|
|
||||||
|
from : ∀ {ks : List A} → FiniteMap ks → IterProd (length ks)
|
||||||
|
from {[]} (([] , _) , _) = tt
|
||||||
|
from {k ∷ ks'} (((k' , v) ∷ fm' , push _ uks') , refl) =
|
||||||
|
(v , from ((fm' , uks'), refl))
|
||||||
|
|
||||||
|
to : ∀ {ks : List A} → Unique ks → IterProd (length ks) → FiniteMap ks
|
||||||
|
to {[]} _ ⊤ = (([] , empty) , refl)
|
||||||
|
to {k ∷ ks'} (push k≢ks' uks') (v , rest) =
|
||||||
|
let
|
||||||
|
((fm' , ufm') , fm'≡ks') = to uks' rest
|
||||||
|
|
||||||
|
-- This would be easier if we pattern matched on the equiality proof
|
||||||
|
-- to get refl, but that makes it harder to reason about 'to' when
|
||||||
|
-- the arguments are not known to be refl.
|
||||||
|
k≢fm' = subst (λ ks → All (λ k' → ¬ k ≡ k') ks) (sym fm'≡ks') k≢ks'
|
||||||
|
kvs≡ks = cong (k ∷_) fm'≡ks'
|
||||||
|
in
|
||||||
|
(((k , v) ∷ fm' , push k≢fm' ufm') , kvs≡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}
|
||||||
|
|
||||||
|
to-build : ∀ {b : B} {ks : List A} (uks : Unique ks) →
|
||||||
|
let fm = to uks (IP.build b tt (length ks))
|
||||||
|
in ∀ (k : A) (v : B) → (k , v) ∈ᵐ fm → v ≡ b
|
||||||
|
to-build {b} {k ∷ ks'} (push _ uks') k v (here refl) = refl
|
||||||
|
to-build {b} {k ∷ ks'} (push _ uks') k' v (there k',v∈m') =
|
||||||
|
to-build {ks = ks'} uks' k' v k',v∈m'
|
||||||
|
|
||||||
|
|
||||||
|
-- The left inverse is: from (to x) = x
|
||||||
|
from-to-inverseˡ : ∀ {ks : List A} (uks : Unique ks) →
|
||||||
|
IsInverseˡ (_≈_ ks) (_≈ⁱᵖ_ {length ks})
|
||||||
|
(from {ks}) (to {ks} uks)
|
||||||
|
from-to-inverseˡ {[]} _ _ = IsEquivalence.≈-refl (IP.≈-equiv {0})
|
||||||
|
from-to-inverseˡ {k ∷ ks'} (push k≢ks' uks') (v , rest)
|
||||||
|
with ((fm' , ufm') , refl) ← to uks' rest in p rewrite sym p =
|
||||||
|
(IsLattice.≈-refl lB , from-to-inverseˡ {ks'} uks' rest)
|
||||||
|
-- the rewrite here is needed because the IH is in terms of `to uks' rest`,
|
||||||
|
-- but we end up with the 'unpacked' form (fm', ...). So, put it back
|
||||||
|
-- in the 'packed' form after we've performed enough inspection
|
||||||
|
-- to know we take the cons branch of `to`.
|
||||||
|
|
||||||
|
-- The map has its own uniqueness proof, but the call to 'to' needs a standalone
|
||||||
|
-- uniqueness proof too. Work with both proofs as needed to thread things through.
|
||||||
|
--
|
||||||
|
-- The right inverse is: to (from x) = x
|
||||||
|
from-to-inverseʳ : ∀ {ks : List A} (uks : Unique ks) →
|
||||||
|
IsInverseʳ (_≈_ ks) (_≈ⁱᵖ_ {length ks})
|
||||||
|
(from {ks}) (to {ks} uks)
|
||||||
|
from-to-inverseʳ {[]} _ (([] , empty) , kvs≡ks) rewrite kvs≡ks =
|
||||||
|
( (λ k v ())
|
||||||
|
, (λ k v ())
|
||||||
|
)
|
||||||
|
from-to-inverseʳ {k ∷ ks'} uks@(push _ uks'₁) fm₁@(((k , v) ∷ fm'₁ , push _ uks'₂) , refl)
|
||||||
|
with to uks'₁ (from ((fm'₁ , uks'₂) , refl))
|
||||||
|
| from-to-inverseʳ {ks'} uks'₁ ((fm'₁ , uks'₂) , refl)
|
||||||
|
... | ((fm'₂ , ufm'₂) , _)
|
||||||
|
| (fm'₂⊆fm'₁ , fm'₁⊆fm'₂) = (m₂⊆m₁ , m₁⊆m₂)
|
||||||
|
where
|
||||||
|
kvs₁ = (k , v) ∷ fm'₁
|
||||||
|
kvs₂ = (k , v) ∷ fm'₂
|
||||||
|
|
||||||
|
m₁⊆m₂ : subset-impl kvs₁ kvs₂
|
||||||
|
m₁⊆m₂ k' v' (here refl) =
|
||||||
|
(v' , (IsLattice.≈-refl lB , here refl))
|
||||||
|
m₁⊆m₂ k' v' (there k',v'∈fm'₁) =
|
||||||
|
let (v'' , (v'≈v'' , k',v''∈fm'₂)) =
|
||||||
|
fm'₁⊆fm'₂ k' v' k',v'∈fm'₁
|
||||||
|
in (v'' , (v'≈v'' , there k',v''∈fm'₂))
|
||||||
|
|
||||||
|
m₂⊆m₁ : subset-impl kvs₂ kvs₁
|
||||||
|
m₂⊆m₁ k' v' (here refl) =
|
||||||
|
(v' , (IsLattice.≈-refl lB , here refl))
|
||||||
|
m₂⊆m₁ k' v' (there k',v'∈fm'₂) =
|
||||||
|
let (v'' , (v'≈v'' , k',v''∈fm'₁)) =
|
||||||
|
fm'₂⊆fm'₁ k' v' k',v'∈fm'₂
|
||||||
|
in (v'' , (v'≈v'' , there k',v''∈fm'₁))
|
||||||
|
|
||||||
|
private
|
||||||
|
first-key-in-map : ∀ {k : A} {ks : List A} (fm : FiniteMap (k ∷ ks)) →
|
||||||
|
Σ 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)) →
|
||||||
|
proj₁ (from fm) ≡ proj₁ (first-key-in-map fm)
|
||||||
|
from-first-value {k} {ks} (((k , v) ∷ _ , push _ _) , refl) = refl
|
||||||
|
|
||||||
|
-- We need pop because reasoning about two distinct 'refl' pattern
|
||||||
|
-- matches is giving us unification errors. So, stash the 'refl' pattern
|
||||||
|
-- matching into a helper functions, and write solutions in terms
|
||||||
|
-- of that.
|
||||||
|
pop : ∀ {k : A} {ks : List A} → FiniteMap (k ∷ ks) → FiniteMap ks
|
||||||
|
pop (((_ ∷ fm') , push _ ufm') , refl) = ((fm' , ufm') , refl)
|
||||||
|
|
||||||
|
pop-≈ : ∀ {k : A} {ks : List A} (fm₁ fm₂ : FiniteMap (k ∷ ks)) →
|
||||||
|
_≈_ _ fm₁ fm₂ → _≈_ _ (pop fm₁) (pop fm₂)
|
||||||
|
pop-≈ {k} {ks} fm₁ fm₂ (fm₁⊆fm₂ , fm₂⊆fm₁) =
|
||||||
|
(narrow fm₁⊆fm₂ , narrow fm₂⊆fm₁)
|
||||||
|
where
|
||||||
|
narrow₁ : ∀ {fm₁ fm₂ : FiniteMap (k ∷ ks)} →
|
||||||
|
fm₁ ⊆ᵐ fm₂ → pop fm₁ ⊆ᵐ fm₂
|
||||||
|
narrow₁ {(_ ∷ _ , push _ _) , refl} kvs₁⊆kvs₂ k' v' k',v'∈fm'₁ =
|
||||||
|
kvs₁⊆kvs₂ k' v' (there k',v'∈fm'₁)
|
||||||
|
|
||||||
|
narrow₂ : ∀ {fm₁ : FiniteMap ks} {fm₂ : FiniteMap (k ∷ ks)} →
|
||||||
|
fm₁ ⊆ᵐ fm₂ → fm₁ ⊆ᵐ pop fm₂
|
||||||
|
narrow₂ {fm₁} {fm₂ = (_ ∷ fm'₂ , push k≢ks _) , kvs≡ks@refl} kvs₁⊆kvs₂ k' v' k',v'∈fm'₁
|
||||||
|
with kvs₁⊆kvs₂ k' v' k',v'∈fm'₁
|
||||||
|
... | (v'' , (v'≈v'' , here refl)) rewrite sym (proj₂ fm₁) =
|
||||||
|
⊥-elim (All¬-¬Any k≢ks (forget k',v'∈fm'₁))
|
||||||
|
... | (v'' , (v'≈v'' , there k',v'∈fm'₂)) =
|
||||||
|
(v'' , (v'≈v'' , k',v'∈fm'₂))
|
||||||
|
|
||||||
|
narrow : ∀ {fm₁ fm₂ : FiniteMap (k ∷ ks)} →
|
||||||
|
fm₁ ⊆ᵐ fm₂ → pop fm₁ ⊆ᵐ pop fm₂
|
||||||
|
narrow {fm₁} {fm₂} x = narrow₂ {pop fm₁} (narrow₁ {fm₂ = fm₂} x)
|
||||||
|
|
||||||
|
k,v∈pop⇒k,v∈ : ∀ {k : A} {ks : List A} {k' : A} {v : B} (fm : FiniteMap (k ∷ ks)) →
|
||||||
|
(k' , v) ∈ᵐ pop fm → (¬ k ≡ k' × ((k' , v) ∈ᵐ fm))
|
||||||
|
k,v∈pop⇒k,v∈ {k} {ks} {k'} {v} (m@((k , _) ∷ fm' , push k≢ks uks') , refl) k',v∈fm =
|
||||||
|
( (λ { refl → All¬-¬Any k≢ks (forget k',v∈fm) })
|
||||||
|
, there k',v∈fm
|
||||||
|
)
|
||||||
|
|
||||||
|
k,v∈⇒k,v∈pop : ∀ {k : A} {ks : List A} {k' : A} {v : B} (fm : FiniteMap (k ∷ ks)) →
|
||||||
|
¬ k ≡ k' → (k' , v) ∈ᵐ fm → (k' , v) ∈ᵐ pop fm
|
||||||
|
k,v∈⇒k,v∈pop (m@(_ ∷ _ , push k≢ks _) , refl) k≢k' (here refl) = ⊥-elim (k≢k' refl)
|
||||||
|
k,v∈⇒k,v∈pop (m@(_ ∷ _ , push k≢ks _) , refl) k≢k' (there k,v'∈fm') = k,v'∈fm'
|
||||||
|
|
||||||
|
pop-⊔-distr : ∀ {k : A} {ks : List A} (fm₁ fm₂ : FiniteMap (k ∷ ks)) →
|
||||||
|
_≈_ _ (pop (_⊔_ _ fm₁ fm₂)) ((_⊔_ _ (pop fm₁) (pop fm₂)))
|
||||||
|
pop-⊔-distr {k} {ks} fm₁@(m₁ , _) fm₂@(m₂ , _) =
|
||||||
|
(pfm₁fm₂⊆pfm₁pfm₂ , pfm₁pfm₂⊆pfm₁fm₂)
|
||||||
|
where
|
||||||
|
-- pfm₁fm₂⊆pfm₁pfm₂ = {!!}
|
||||||
|
pfm₁fm₂⊆pfm₁pfm₂ : pop (_⊔_ _ fm₁ fm₂) ⊆ᵐ (_⊔_ _ (pop fm₁) (pop fm₂))
|
||||||
|
pfm₁fm₂⊆pfm₁pfm₂ k' v' k',v'∈pfm₁fm₂
|
||||||
|
with (k≢k' , k',v'∈fm₁fm₂) ← k,v∈pop⇒k,v∈ (_⊔_ _ fm₁ fm₂) k',v'∈pfm₁fm₂
|
||||||
|
with ((v₁ , v₂) , (refl , (k,v₁∈fm₁ , k,v₂∈fm₂)))
|
||||||
|
← Provenance-union fm₁ fm₂ k',v'∈fm₁fm₂
|
||||||
|
with k',v₁∈pfm₁ ← k,v∈⇒k,v∈pop fm₁ k≢k' k,v₁∈fm₁
|
||||||
|
with k',v₂∈pfm₂ ← k,v∈⇒k,v∈pop fm₂ k≢k' k,v₂∈fm₂
|
||||||
|
=
|
||||||
|
( v₁ ⊔₂ v₂
|
||||||
|
, (IsLattice.≈-refl lB
|
||||||
|
, ⊔-combines {m₁ = proj₁ (pop fm₁)}
|
||||||
|
{m₂ = proj₁ (pop fm₂)}
|
||||||
|
k',v₁∈pfm₁ k',v₂∈pfm₂
|
||||||
|
)
|
||||||
|
)
|
||||||
|
|
||||||
|
pfm₁pfm₂⊆pfm₁fm₂ : (_⊔_ _ (pop fm₁) (pop fm₂)) ⊆ᵐ pop (_⊔_ _ fm₁ fm₂)
|
||||||
|
pfm₁pfm₂⊆pfm₁fm₂ k' v' k',v'∈pfm₁pfm₂
|
||||||
|
with ((v₁ , v₂) , (refl , (k,v₁∈pfm₁ , k,v₂∈pfm₂)))
|
||||||
|
← Provenance-union (pop fm₁) (pop fm₂) k',v'∈pfm₁pfm₂
|
||||||
|
with (k≢k' , k',v₁∈fm₁) ← k,v∈pop⇒k,v∈ fm₁ k,v₁∈pfm₁
|
||||||
|
with (_ , k',v₂∈fm₂) ← k,v∈pop⇒k,v∈ fm₂ k,v₂∈pfm₂
|
||||||
|
=
|
||||||
|
( v₁ ⊔₂ v₂
|
||||||
|
, ( IsLattice.≈-refl lB
|
||||||
|
, k,v∈⇒k,v∈pop (_⊔_ _ fm₁ fm₂) k≢k'
|
||||||
|
(⊔-combines {m₁ = m₁} {m₂ = m₂}
|
||||||
|
k',v₁∈fm₁ k',v₂∈fm₂)
|
||||||
|
)
|
||||||
|
)
|
||||||
|
|
||||||
|
from-rest : ∀ {k : A} {ks : List A} (fm : FiniteMap (k ∷ ks)) →
|
||||||
|
proj₂ (from fm) ≡ from (pop fm)
|
||||||
|
from-rest (((_ ∷ fm') , push _ ufm') , refl) = refl
|
||||||
|
|
||||||
|
from-preserves-≈ : ∀ {ks : List A} → {fm₁ fm₂ : FiniteMap ks} →
|
||||||
|
_≈_ _ fm₁ fm₂ → (_≈ⁱᵖ_ {length ks}) (from fm₁) (from fm₂)
|
||||||
|
from-preserves-≈ {[]} {_} {_} _ = IsEquivalence.≈-refl ≈ᵘ-equiv
|
||||||
|
from-preserves-≈ {k ∷ ks'} {fm₁@(m₁ , _)} {fm₂@(m₂ , _)} fm₁≈fm₂@(kvs₁⊆kvs₂ , kvs₂⊆kvs₁)
|
||||||
|
with first-key-in-map fm₁
|
||||||
|
| first-key-in-map fm₂
|
||||||
|
| from-first-value fm₁
|
||||||
|
| from-first-value fm₂
|
||||||
|
... | (v₁ , k,v₁∈fm₁) | (v₂ , k,v₂∈fm₂) | refl | refl
|
||||||
|
with kvs₁⊆kvs₂ _ _ k,v₁∈fm₁
|
||||||
|
... | (v₁' , (v₁≈v₁' , k,v₁'∈fm₂))
|
||||||
|
rewrite Map-functional {m = m₂} k,v₂∈fm₂ k,v₁'∈fm₂
|
||||||
|
rewrite from-rest fm₁ rewrite from-rest fm₂
|
||||||
|
=
|
||||||
|
( v₁≈v₁'
|
||||||
|
, from-preserves-≈ {ks'} {pop fm₁} {pop fm₂}
|
||||||
|
(pop-≈ fm₁ fm₂ fm₁≈fm₂)
|
||||||
|
)
|
||||||
|
|
||||||
|
to-preserves-≈ : ∀ {ks : List A} (uks : Unique ks) {ip₁ ip₂ : IterProd (length ks)} →
|
||||||
|
_≈ⁱᵖ_ {length ks} ip₁ ip₂ → _≈_ _ (to uks ip₁) (to uks ip₂)
|
||||||
|
to-preserves-≈ {[]} empty {tt} {tt} _ = ((λ k v ()), (λ k v ()))
|
||||||
|
to-preserves-≈ {k ∷ ks'} uks@(push k≢ks' uks') {ip₁@(v₁ , rest₁)} {ip₂@(v₂ , rest₂)} (v₁≈v₂ , rest₁≈rest₂) = (fm₁⊆fm₂ , fm₂⊆fm₁)
|
||||||
|
where
|
||||||
|
inductive-step : ∀ {v₁ v₂ : B} {rest₁ rest₂ : IterProd (length ks')} →
|
||||||
|
v₁ ≈₂ v₂ → _≈ⁱᵖ_ {length ks'} rest₁ rest₂ →
|
||||||
|
to uks (v₁ , rest₁) ⊆ᵐ to uks (v₂ , rest₂)
|
||||||
|
inductive-step {v₁} {v₂} {rest₁} {rest₂} v₁≈v₂ rest₁≈rest₂ k v k,v∈kvs₁
|
||||||
|
with ((fm'₁ , ufm'₁) , fm'₁≡ks') ← to uks' rest₁ in p₁
|
||||||
|
with ((fm'₂ , ufm'₂) , fm'₂≡ks') ← to uks' rest₂ in p₂
|
||||||
|
with k,v∈kvs₁
|
||||||
|
... | here refl = (v₂ , (v₁≈v₂ , here refl))
|
||||||
|
... | there k,v∈fm'₁ with refl ← p₁ with refl ← p₂ =
|
||||||
|
let
|
||||||
|
(fm'₁⊆fm'₂ , _) = to-preserves-≈ uks' {rest₁} {rest₂}
|
||||||
|
rest₁≈rest₂
|
||||||
|
(v' , (v≈v' , k,v'∈kvs₁)) = fm'₁⊆fm'₂ k v k,v∈fm'₁
|
||||||
|
in
|
||||||
|
(v' , (v≈v' , there k,v'∈kvs₁))
|
||||||
|
|
||||||
|
fm₁⊆fm₂ : to uks ip₁ ⊆ᵐ to uks ip₂
|
||||||
|
fm₁⊆fm₂ = inductive-step v₁≈v₂ rest₁≈rest₂
|
||||||
|
|
||||||
|
fm₂⊆fm₁ : to uks ip₂ ⊆ᵐ to uks ip₁
|
||||||
|
fm₂⊆fm₁ = inductive-step (≈₂-sym v₁≈v₂)
|
||||||
|
(IP.≈-sym {length ks'} rest₁≈rest₂)
|
||||||
|
|
||||||
|
from-⊔-distr : ∀ {ks : List A} → (fm₁ fm₂ : FiniteMap ks) →
|
||||||
|
_≈ⁱᵖ_ {length ks} (from (_⊔_ _ fm₁ fm₂))
|
||||||
|
(_⊔ⁱᵖ_ {ks} (from fm₁) (from fm₂))
|
||||||
|
from-⊔-distr {[]} fm₁ fm₂ = IsEquivalence.≈-refl ≈ᵘ-equiv
|
||||||
|
from-⊔-distr {k ∷ ks} fm₁@(m₁ , _) fm₂@(m₂ , _)
|
||||||
|
with first-key-in-map (_⊔_ _ fm₁ fm₂)
|
||||||
|
| first-key-in-map fm₁
|
||||||
|
| first-key-in-map fm₂
|
||||||
|
| from-first-value (_⊔_ _ fm₁ fm₂)
|
||||||
|
| from-first-value fm₁ | from-first-value fm₂
|
||||||
|
... | (v , k,v∈fm₁fm₂) | (v₁ , k,v₁∈fm₁) | (v₂ , k,v₂∈fm₂) | refl | refl | refl
|
||||||
|
with Expr-Provenance k ((` m₁) ∪ (` m₂)) (forget k,v∈fm₁fm₂)
|
||||||
|
... | (_ , (in₁ _ k∉km₂ , _)) = ⊥-elim (k∉km₂ (forget k,v₂∈fm₂))
|
||||||
|
... | (_ , (in₂ k∉km₁ _ , _)) = ⊥-elim (k∉km₁ (forget k,v₁∈fm₁))
|
||||||
|
... | (v₁⊔v₂ , (bothᵘ {v₁'} {v₂'} (single k,v₁'∈m₁) (single k,v₂'∈m₂) , k,v₁⊔v₂∈m₁m₂))
|
||||||
|
rewrite Map-functional {m = m₁} k,v₁∈fm₁ k,v₁'∈m₁
|
||||||
|
rewrite Map-functional {m = m₂} k,v₂∈fm₂ k,v₂'∈m₂
|
||||||
|
rewrite Map-functional {m = proj₁ (_⊔_ _ fm₁ fm₂)} k,v∈fm₁fm₂ k,v₁⊔v₂∈m₁m₂
|
||||||
|
rewrite from-rest (_⊔_ _ fm₁ fm₂) rewrite from-rest fm₁ rewrite from-rest fm₂
|
||||||
|
= ( IsLattice.≈-refl lB
|
||||||
|
, IsEquivalence.≈-trans
|
||||||
|
(IP.≈-equiv {length ks})
|
||||||
|
(from-preserves-≈ {_} {pop (_⊔_ _ fm₁ fm₂)}
|
||||||
|
{_⊔_ _ (pop fm₁) (pop fm₂)}
|
||||||
|
(pop-⊔-distr fm₁ fm₂))
|
||||||
|
((from-⊔-distr (pop fm₁) (pop fm₂)))
|
||||||
|
)
|
||||||
|
|
||||||
|
|
||||||
|
to-⊔-distr : ∀ {ks : List A} (uks : Unique ks) → (ip₁ ip₂ : IterProd (length ks)) →
|
||||||
|
_≈_ _ (to uks (_⊔ⁱᵖ_ {ks} ip₁ ip₂)) ((_⊔_ _ (to uks ip₁) (to uks ip₂)))
|
||||||
|
to-⊔-distr {[]} empty tt tt = ((λ k v ()), (λ k v ()))
|
||||||
|
to-⊔-distr {ks@(k ∷ ks')} uks@(push k≢ks' uks') ip₁@(v₁ , rest₁) ip₂@(v₂ , rest₂) = (fm⊆fm₁fm₂ , fm₁fm₂⊆fm)
|
||||||
|
where
|
||||||
|
fm₁ = to uks ip₁
|
||||||
|
fm₁' = to uks' rest₁
|
||||||
|
fm₂ = to uks ip₂
|
||||||
|
fm₂' = to uks' rest₂
|
||||||
|
fm = to uks (_⊔ⁱᵖ_ {k ∷ ks'} ip₁ ip₂)
|
||||||
|
|
||||||
|
fm⊆fm₁fm₂ : fm ⊆ᵐ (_⊔_ _ fm₁ fm₂)
|
||||||
|
fm⊆fm₁fm₂ k v (here refl) =
|
||||||
|
(v₁ ⊔₂ v₂
|
||||||
|
, (IsLattice.≈-refl lB
|
||||||
|
, ⊔-combines {k} {v₁} {v₂} {proj₁ fm₁} {proj₁ fm₂}
|
||||||
|
(here refl) (here refl)
|
||||||
|
)
|
||||||
|
)
|
||||||
|
fm⊆fm₁fm₂ k' v (there k',v∈fm')
|
||||||
|
with (fm'⊆fm'₁fm'₂ , _) ← to-⊔-distr uks' rest₁ rest₂
|
||||||
|
with (v' , (v₁⊔v₂≈v' , k',v'∈fm'₁fm'₂))
|
||||||
|
← fm'⊆fm'₁fm'₂ k' v k',v∈fm'
|
||||||
|
with (_ , (refl , (v₁∈fm'₁ , v₂∈fm'₂)))
|
||||||
|
← Provenance-union fm₁' fm₂' k',v'∈fm'₁fm'₂ =
|
||||||
|
( v'
|
||||||
|
, ( v₁⊔v₂≈v'
|
||||||
|
, ⊔-combines {m₁ = proj₁ fm₁} {m₂ = proj₁ fm₂}
|
||||||
|
(there v₁∈fm'₁) (there v₂∈fm'₂)
|
||||||
|
)
|
||||||
|
)
|
||||||
|
|
||||||
|
fm₁fm₂⊆fm : (_⊔_ _ fm₁ fm₂) ⊆ᵐ fm
|
||||||
|
fm₁fm₂⊆fm k' v k',v∈fm₁fm₂
|
||||||
|
with (_ , fm'₁fm'₂⊆fm')
|
||||||
|
← to-⊔-distr uks' rest₁ rest₂
|
||||||
|
with (_ , (refl , (v₁∈fm₁ , v₂∈fm₂)))
|
||||||
|
← Provenance-union fm₁ fm₂ k',v∈fm₁fm₂
|
||||||
|
with v₁∈fm₁ | v₂∈fm₂
|
||||||
|
... | here refl | here refl =
|
||||||
|
(v , (IsLattice.≈-refl lB , here refl))
|
||||||
|
... | here refl | there k',v₂∈fm₂' =
|
||||||
|
⊥-elim (All¬-¬Any k≢ks' (subst (k' ∈ˡ_) (proj₂ fm₂')
|
||||||
|
(forget k',v₂∈fm₂')))
|
||||||
|
... | there k',v₁∈fm₁' | here refl =
|
||||||
|
⊥-elim (All¬-¬Any k≢ks' (subst (k' ∈ˡ_) (proj₂ fm₁')
|
||||||
|
(forget k',v₁∈fm₁')))
|
||||||
|
... | there k',v₁∈fm₁' | there k',v₂∈fm₂' =
|
||||||
|
let
|
||||||
|
k',v₁v₂∈fm₁'fm₂' =
|
||||||
|
⊔-combines {m₁ = proj₁ fm₁'} {m₂ = proj₁ fm₂'}
|
||||||
|
k',v₁∈fm₁' k',v₂∈fm₂'
|
||||||
|
(v' , (v₁⊔v₂≈v' , v'∈fm')) =
|
||||||
|
fm'₁fm'₂⊆fm' _ _ k',v₁v₂∈fm₁'fm₂'
|
||||||
|
in
|
||||||
|
(v' , (v₁⊔v₂≈v' , there v'∈fm'))
|
||||||
|
|
||||||
|
module FixedHeight {ks : List A} {{≈₂-Decidable : IsDecidable _≈₂_}} {h₂ : ℕ} {{fhB : FixedHeight₂ h₂}} (uks : Unique ks) where
|
||||||
|
import Isomorphism
|
||||||
|
open Isomorphism.TransportFiniteHeight
|
||||||
|
(IP.isFiniteHeightLattice {k = length ks} {{fhB = fixedHeightᵘ}}) (isLattice ks)
|
||||||
|
{f = to uks} {g = from {ks}}
|
||||||
|
(to-preserves-≈ uks) (from-preserves-≈ {ks})
|
||||||
|
(to-⊔-distr uks) (from-⊔-distr {ks})
|
||||||
|
(from-to-inverseʳ uks) (from-to-inverseˡ uks)
|
||||||
|
using (isFiniteHeightLattice; finiteHeightLattice; fixedHeight) public
|
||||||
|
|
||||||
|
-- Helpful lemma: all entries of the 'bottom' map are assigned to bottom.
|
||||||
|
|
||||||
|
open Height (IsFiniteHeightLattice.fixedHeight isFiniteHeightLattice) using (⊥)
|
||||||
|
|
||||||
|
⊥-contains-bottoms : ∀ {k : A} {v : B} → (k , v) ∈ᵐ ⊥ → v ≡ (Height.⊥ fhB)
|
||||||
|
⊥-contains-bottoms {k} {v} k,v∈⊥
|
||||||
|
rewrite IP.⊥-built {length ks} {{fhB = fixedHeightᵘ}} =
|
||||||
|
to-build uks k v k,v∈⊥
|
||||||
|
|
||||||
|
open WithKeys ks public
|
||||||
|
module FixedHeight = IterProdIsomorphism.FixedHeight
|
||||||
|
|||||||
@@ -1,409 +0,0 @@
|
|||||||
-- Because iterated products currently require both A and B to be of the same
|
|
||||||
-- universe, and the FiniteMap is written in a universe-polymorphic way,
|
|
||||||
-- specialize the FiniteMap module with Set-typed types only.
|
|
||||||
|
|
||||||
open import Lattice
|
|
||||||
open import Equivalence
|
|
||||||
open import Relation.Binary.PropositionalEquality as Eq
|
|
||||||
using (_≡_; refl; sym; trans; cong; subst)
|
|
||||||
open import Relation.Binary.Definitions using (Decidable)
|
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
|
||||||
open import Function.Definitions using (Inverseˡ; Inverseʳ)
|
|
||||||
|
|
||||||
module Lattice.FiniteValueMap {A : Set} {B : Set}
|
|
||||||
{_≈₂_ : B → B → Set}
|
|
||||||
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
|
||||||
(≡-dec-A : Decidable (_≡_ {_} {A}))
|
|
||||||
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
|
||||||
|
|
||||||
open import Data.List using (List; length; []; _∷_; map)
|
|
||||||
open import Data.List.Membership.Propositional using () renaming (_∈_ to _∈ˡ_)
|
|
||||||
open import Data.Nat using (ℕ)
|
|
||||||
open import Data.Product using (Σ; proj₁; proj₂; _×_)
|
|
||||||
open import Data.Empty using (⊥-elim)
|
|
||||||
open import Utils using (Unique; push; empty; All¬-¬Any)
|
|
||||||
open import Data.Product using (_,_)
|
|
||||||
open import Data.List.Properties using (∷-injectiveʳ)
|
|
||||||
open import Data.List.Relation.Unary.All using (All)
|
|
||||||
open import Data.List.Relation.Unary.Any using (Any; here; there)
|
|
||||||
open import Relation.Nullary using (¬_)
|
|
||||||
open import Isomorphism using (IsInverseˡ; IsInverseʳ)
|
|
||||||
|
|
||||||
open import Lattice.Map ≡-dec-A lB
|
|
||||||
using
|
|
||||||
( subset-impl
|
|
||||||
; locate; forget
|
|
||||||
; Map-functional
|
|
||||||
; Expr-Provenance
|
|
||||||
; Expr-Provenance-≡
|
|
||||||
; _∩_; _∪_; `_
|
|
||||||
; in₁; in₂; bothᵘ; single
|
|
||||||
; ⊔-combines
|
|
||||||
)
|
|
||||||
open import Lattice.FiniteMap ≡-dec-A lB public
|
|
||||||
|
|
||||||
module IterProdIsomorphism where
|
|
||||||
open import Data.Unit using (⊤; tt)
|
|
||||||
open import Lattice.Unit using ()
|
|
||||||
renaming
|
|
||||||
( _≈_ to _≈ᵘ_
|
|
||||||
; _⊔_ to _⊔ᵘ_
|
|
||||||
; _⊓_ to _⊓ᵘ_
|
|
||||||
; ≈-dec to ≈ᵘ-dec
|
|
||||||
; isLattice to isLatticeᵘ
|
|
||||||
; ≈-equiv to ≈ᵘ-equiv
|
|
||||||
; fixedHeight to fixedHeightᵘ
|
|
||||||
)
|
|
||||||
open import Lattice.IterProd _≈₂_ _≈ᵘ_ _⊔₂_ _⊔ᵘ_ _⊓₂_ _⊓ᵘ_ lB isLatticeᵘ
|
|
||||||
as IP
|
|
||||||
using (IterProd)
|
|
||||||
open IsLattice lB using ()
|
|
||||||
renaming
|
|
||||||
( ≈-trans to ≈₂-trans
|
|
||||||
; ≈-sym to ≈₂-sym
|
|
||||||
; FixedHeight to FixedHeight₂
|
|
||||||
)
|
|
||||||
|
|
||||||
from : ∀ {ks : List A} → FiniteMap ks → IterProd (length ks)
|
|
||||||
from {[]} (([] , _) , _) = tt
|
|
||||||
from {k ∷ ks'} (((k' , v) ∷ fm' , push _ uks') , refl) =
|
|
||||||
(v , from ((fm' , uks'), refl))
|
|
||||||
|
|
||||||
to : ∀ {ks : List A} → Unique ks → IterProd (length ks) → FiniteMap ks
|
|
||||||
to {[]} _ ⊤ = (([] , empty) , refl)
|
|
||||||
to {k ∷ ks'} (push k≢ks' uks') (v , rest) =
|
|
||||||
let
|
|
||||||
((fm' , ufm') , fm'≡ks') = to uks' rest
|
|
||||||
|
|
||||||
-- This would be easier if we pattern matched on the equiality proof
|
|
||||||
-- to get refl, but that makes it harder to reason about 'to' when
|
|
||||||
-- the arguments are not known to be refl.
|
|
||||||
k≢fm' = subst (λ ks → All (λ k' → ¬ k ≡ k') ks) (sym fm'≡ks') k≢ks'
|
|
||||||
kvs≡ks = cong (k ∷_) fm'≡ks'
|
|
||||||
in
|
|
||||||
(((k , v) ∷ fm' , push k≢fm' ufm') , kvs≡ks)
|
|
||||||
|
|
||||||
|
|
||||||
private
|
|
||||||
_≈ᵐ_ : ∀ {ks : List A} → FiniteMap ks → FiniteMap ks → Set
|
|
||||||
_≈ᵐ_ {ks} = _≈_ ks
|
|
||||||
|
|
||||||
_⊔ᵐ_ : ∀ {ks : List A} → FiniteMap ks → FiniteMap ks → FiniteMap ks
|
|
||||||
_⊔ᵐ_ {ks} = _⊔_ ks
|
|
||||||
|
|
||||||
_⊆ᵐ_ : ∀ {ks₁ ks₂ : List A} → FiniteMap ks₁ → FiniteMap ks₂ → Set
|
|
||||||
_⊆ᵐ_ fm₁ fm₂ = subset-impl (proj₁ (proj₁ fm₁)) (proj₁ (proj₁ fm₂))
|
|
||||||
|
|
||||||
_≈ⁱᵖ_ : ∀ {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
|
|
||||||
|
|
||||||
-- The left inverse is: from (to x) = x
|
|
||||||
from-to-inverseˡ : ∀ {ks : List A} (uks : Unique ks) →
|
|
||||||
IsInverseˡ (_≈ᵐ_ {ks}) (_≈ⁱᵖ_ {length ks})
|
|
||||||
(from {ks}) (to {ks} uks)
|
|
||||||
from-to-inverseˡ {[]} _ _ = IsEquivalence.≈-refl (IP.≈-equiv 0)
|
|
||||||
from-to-inverseˡ {k ∷ ks'} (push k≢ks' uks') (v , rest)
|
|
||||||
with ((fm' , ufm') , refl) ← to uks' rest in p rewrite sym p =
|
|
||||||
(IsLattice.≈-refl lB , from-to-inverseˡ {ks'} uks' rest)
|
|
||||||
-- the rewrite here is needed because the IH is in terms of `to uks' rest`,
|
|
||||||
-- but we end up with the 'unpacked' form (fm', ...). So, put it back
|
|
||||||
-- in the 'packed' form after we've performed enough inspection
|
|
||||||
-- to know we take the cons branch of `to`.
|
|
||||||
|
|
||||||
-- The map has its own uniqueness proof, but the call to 'to' needs a standalone
|
|
||||||
-- uniqueness proof too. Work with both proofs as needed to thread things through.
|
|
||||||
--
|
|
||||||
-- The right inverse is: to (from x) = x
|
|
||||||
from-to-inverseʳ : ∀ {ks : List A} (uks : Unique ks) →
|
|
||||||
IsInverseʳ (_≈ᵐ_ {ks}) (_≈ⁱᵖ_ {length ks})
|
|
||||||
(from {ks}) (to {ks} uks)
|
|
||||||
from-to-inverseʳ {[]} _ (([] , empty) , kvs≡ks) rewrite kvs≡ks =
|
|
||||||
( (λ k v ())
|
|
||||||
, (λ k v ())
|
|
||||||
)
|
|
||||||
from-to-inverseʳ {k ∷ ks'} uks@(push _ uks'₁) fm₁@(((k , v) ∷ fm'₁ , push _ uks'₂) , refl)
|
|
||||||
with to uks'₁ (from ((fm'₁ , uks'₂) , refl))
|
|
||||||
| from-to-inverseʳ {ks'} uks'₁ ((fm'₁ , uks'₂) , refl)
|
|
||||||
... | ((fm'₂ , ufm'₂) , _)
|
|
||||||
| (fm'₂⊆fm'₁ , fm'₁⊆fm'₂) = (m₂⊆m₁ , m₁⊆m₂)
|
|
||||||
where
|
|
||||||
kvs₁ = (k , v) ∷ fm'₁
|
|
||||||
kvs₂ = (k , v) ∷ fm'₂
|
|
||||||
|
|
||||||
m₁⊆m₂ : subset-impl kvs₁ kvs₂
|
|
||||||
m₁⊆m₂ k' v' (here refl) =
|
|
||||||
(v' , (IsLattice.≈-refl lB , here refl))
|
|
||||||
m₁⊆m₂ k' v' (there k',v'∈fm'₁) =
|
|
||||||
let (v'' , (v'≈v'' , k',v''∈fm'₂)) =
|
|
||||||
fm'₁⊆fm'₂ k' v' k',v'∈fm'₁
|
|
||||||
in (v'' , (v'≈v'' , there k',v''∈fm'₂))
|
|
||||||
|
|
||||||
m₂⊆m₁ : subset-impl kvs₂ kvs₁
|
|
||||||
m₂⊆m₁ k' v' (here refl) =
|
|
||||||
(v' , (IsLattice.≈-refl lB , here refl))
|
|
||||||
m₂⊆m₁ k' v' (there k',v'∈fm'₂) =
|
|
||||||
let (v'' , (v'≈v'' , k',v''∈fm'₁)) =
|
|
||||||
fm'₂⊆fm'₁ k' v' k',v'∈fm'₂
|
|
||||||
in (v'' , (v'≈v'' , there k',v''∈fm'₁))
|
|
||||||
|
|
||||||
private
|
|
||||||
first-key-in-map : ∀ {k : A} {ks : List A} (fm : FiniteMap (k ∷ ks)) →
|
|
||||||
Σ 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)) →
|
|
||||||
proj₁ (from fm) ≡ proj₁ (first-key-in-map fm)
|
|
||||||
from-first-value {k} {ks} (((k , v) ∷ _ , push _ _) , refl) = refl
|
|
||||||
|
|
||||||
-- We need pop because reasoning about two distinct 'refl' pattern
|
|
||||||
-- matches is giving us unification errors. So, stash the 'refl' pattern
|
|
||||||
-- matching into a helper functions, and write solutions in terms
|
|
||||||
-- of that.
|
|
||||||
pop : ∀ {k : A} {ks : List A} → FiniteMap (k ∷ ks) → FiniteMap ks
|
|
||||||
pop (((_ ∷ fm') , push _ ufm') , refl) = ((fm' , ufm') , refl)
|
|
||||||
|
|
||||||
pop-≈ : ∀ {k : A} {ks : List A} (fm₁ fm₂ : FiniteMap (k ∷ ks)) →
|
|
||||||
fm₁ ≈ᵐ fm₂ → pop fm₁ ≈ᵐ pop fm₂
|
|
||||||
pop-≈ {k} {ks} fm₁ fm₂ (fm₁⊆fm₂ , fm₂⊆fm₁) =
|
|
||||||
(narrow fm₁⊆fm₂ , narrow fm₂⊆fm₁)
|
|
||||||
where
|
|
||||||
narrow₁ : ∀ {fm₁ fm₂ : FiniteMap (k ∷ ks)} →
|
|
||||||
fm₁ ⊆ᵐ fm₂ → pop fm₁ ⊆ᵐ fm₂
|
|
||||||
narrow₁ {(_ ∷ _ , push _ _) , refl} kvs₁⊆kvs₂ k' v' k',v'∈fm'₁ =
|
|
||||||
kvs₁⊆kvs₂ k' v' (there k',v'∈fm'₁)
|
|
||||||
|
|
||||||
narrow₂ : ∀ {fm₁ : FiniteMap ks} {fm₂ : FiniteMap (k ∷ ks)} →
|
|
||||||
fm₁ ⊆ᵐ fm₂ → fm₁ ⊆ᵐ pop fm₂
|
|
||||||
narrow₂ {fm₁} {fm₂ = (_ ∷ fm'₂ , push k≢ks _) , kvs≡ks@refl} kvs₁⊆kvs₂ k' v' k',v'∈fm'₁
|
|
||||||
with kvs₁⊆kvs₂ k' v' k',v'∈fm'₁
|
|
||||||
... | (v'' , (v'≈v'' , here refl)) rewrite sym (proj₂ fm₁) =
|
|
||||||
⊥-elim (All¬-¬Any k≢ks (forget k',v'∈fm'₁))
|
|
||||||
... | (v'' , (v'≈v'' , there k',v'∈fm'₂)) =
|
|
||||||
(v'' , (v'≈v'' , k',v'∈fm'₂))
|
|
||||||
|
|
||||||
narrow : ∀ {fm₁ fm₂ : FiniteMap (k ∷ ks)} →
|
|
||||||
fm₁ ⊆ᵐ fm₂ → pop fm₁ ⊆ᵐ pop fm₂
|
|
||||||
narrow {fm₁} {fm₂} x = narrow₂ {pop fm₁} (narrow₁ {fm₂ = fm₂} x)
|
|
||||||
|
|
||||||
k,v∈pop⇒k,v∈ : ∀ {k : A} {ks : List A} {k' : A} {v : B} (fm : FiniteMap (k ∷ ks)) →
|
|
||||||
(k' , v) ∈ᵐ pop fm → (¬ k ≡ k' × ((k' , v) ∈ᵐ fm))
|
|
||||||
k,v∈pop⇒k,v∈ {k} {ks} {k'} {v} (m@((k , _) ∷ fm' , push k≢ks uks') , refl) k',v∈fm =
|
|
||||||
( (λ { refl → All¬-¬Any k≢ks (forget k',v∈fm) })
|
|
||||||
, there k',v∈fm
|
|
||||||
)
|
|
||||||
|
|
||||||
k,v∈⇒k,v∈pop : ∀ {k : A} {ks : List A} {k' : A} {v : B} (fm : FiniteMap (k ∷ ks)) →
|
|
||||||
¬ k ≡ k' → (k' , v) ∈ᵐ fm → (k' , v) ∈ᵐ pop fm
|
|
||||||
k,v∈⇒k,v∈pop (m@(_ ∷ _ , push k≢ks _) , refl) k≢k' (here refl) = ⊥-elim (k≢k' refl)
|
|
||||||
k,v∈⇒k,v∈pop (m@(_ ∷ _ , push k≢ks _) , refl) k≢k' (there k,v'∈fm') = 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) ∈ᵐ (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₂)))
|
|
||||||
|
|
||||||
pop-⊔-distr : ∀ {k : A} {ks : List A} (fm₁ fm₂ : FiniteMap (k ∷ ks)) →
|
|
||||||
pop (fm₁ ⊔ᵐ fm₂) ≈ᵐ (pop fm₁ ⊔ᵐ pop fm₂)
|
|
||||||
pop-⊔-distr {k} {ks} fm₁@(m₁ , _) fm₂@(m₂ , _) =
|
|
||||||
(pfm₁fm₂⊆pfm₁pfm₂ , pfm₁pfm₂⊆pfm₁fm₂)
|
|
||||||
where
|
|
||||||
-- pfm₁fm₂⊆pfm₁pfm₂ = {!!}
|
|
||||||
pfm₁fm₂⊆pfm₁pfm₂ : pop (fm₁ ⊔ᵐ fm₂) ⊆ᵐ (pop fm₁ ⊔ᵐ pop fm₂)
|
|
||||||
pfm₁fm₂⊆pfm₁pfm₂ k' v' k',v'∈pfm₁fm₂
|
|
||||||
with (k≢k' , k',v'∈fm₁fm₂) ← k,v∈pop⇒k,v∈ (fm₁ ⊔ᵐ fm₂) k',v'∈pfm₁fm₂
|
|
||||||
with ((v₁ , v₂) , (refl , (k,v₁∈fm₁ , k,v₂∈fm₂)))
|
|
||||||
← Provenance-union fm₁ fm₂ k',v'∈fm₁fm₂
|
|
||||||
with k',v₁∈pfm₁ ← k,v∈⇒k,v∈pop fm₁ k≢k' k,v₁∈fm₁
|
|
||||||
with k',v₂∈pfm₂ ← k,v∈⇒k,v∈pop fm₂ k≢k' k,v₂∈fm₂
|
|
||||||
=
|
|
||||||
( v₁ ⊔₂ v₂
|
|
||||||
, (IsLattice.≈-refl lB
|
|
||||||
, ⊔-combines {m₁ = proj₁ (pop fm₁)}
|
|
||||||
{m₂ = proj₁ (pop fm₂)}
|
|
||||||
k',v₁∈pfm₁ k',v₂∈pfm₂
|
|
||||||
)
|
|
||||||
)
|
|
||||||
|
|
||||||
pfm₁pfm₂⊆pfm₁fm₂ : (pop fm₁ ⊔ᵐ pop fm₂) ⊆ᵐ pop (fm₁ ⊔ᵐ fm₂)
|
|
||||||
pfm₁pfm₂⊆pfm₁fm₂ k' v' k',v'∈pfm₁pfm₂
|
|
||||||
with ((v₁ , v₂) , (refl , (k,v₁∈pfm₁ , k,v₂∈pfm₂)))
|
|
||||||
← Provenance-union (pop fm₁) (pop fm₂) k',v'∈pfm₁pfm₂
|
|
||||||
with (k≢k' , k',v₁∈fm₁) ← k,v∈pop⇒k,v∈ fm₁ k,v₁∈pfm₁
|
|
||||||
with (_ , k',v₂∈fm₂) ← k,v∈pop⇒k,v∈ fm₂ k,v₂∈pfm₂
|
|
||||||
=
|
|
||||||
( v₁ ⊔₂ v₂
|
|
||||||
, ( IsLattice.≈-refl lB
|
|
||||||
, k,v∈⇒k,v∈pop (fm₁ ⊔ᵐ fm₂) k≢k'
|
|
||||||
(⊔-combines {m₁ = m₁} {m₂ = m₂}
|
|
||||||
k',v₁∈fm₁ k',v₂∈fm₂)
|
|
||||||
)
|
|
||||||
)
|
|
||||||
|
|
||||||
from-rest : ∀ {k : A} {ks : List A} (fm : FiniteMap (k ∷ ks)) →
|
|
||||||
proj₂ (from fm) ≡ from (pop fm)
|
|
||||||
from-rest (((_ ∷ fm') , push _ ufm') , refl) = refl
|
|
||||||
|
|
||||||
from-preserves-≈ : ∀ {ks : List A} → {fm₁ fm₂ : FiniteMap ks} →
|
|
||||||
fm₁ ≈ᵐ fm₂ → (_≈ⁱᵖ_ {length ks}) (from fm₁) (from fm₂)
|
|
||||||
from-preserves-≈ {[]} {_} {_} _ = IsEquivalence.≈-refl ≈ᵘ-equiv
|
|
||||||
from-preserves-≈ {k ∷ ks'} {fm₁@(m₁ , _)} {fm₂@(m₂ , _)} fm₁≈fm₂@(kvs₁⊆kvs₂ , kvs₂⊆kvs₁)
|
|
||||||
with first-key-in-map fm₁
|
|
||||||
| first-key-in-map fm₂
|
|
||||||
| from-first-value fm₁
|
|
||||||
| from-first-value fm₂
|
|
||||||
... | (v₁ , k,v₁∈fm₁) | (v₂ , k,v₂∈fm₂) | refl | refl
|
|
||||||
with kvs₁⊆kvs₂ _ _ k,v₁∈fm₁
|
|
||||||
... | (v₁' , (v₁≈v₁' , k,v₁'∈fm₂))
|
|
||||||
rewrite Map-functional {m = m₂} k,v₂∈fm₂ k,v₁'∈fm₂
|
|
||||||
rewrite from-rest fm₁ rewrite from-rest fm₂
|
|
||||||
=
|
|
||||||
( v₁≈v₁'
|
|
||||||
, from-preserves-≈ {ks'} {pop fm₁} {pop fm₂}
|
|
||||||
(pop-≈ fm₁ fm₂ fm₁≈fm₂)
|
|
||||||
)
|
|
||||||
|
|
||||||
to-preserves-≈ : ∀ {ks : List A} (uks : Unique ks) {ip₁ ip₂ : IterProd (length ks)} →
|
|
||||||
_≈ⁱᵖ_ {length ks} ip₁ ip₂ → to uks ip₁ ≈ᵐ to uks ip₂
|
|
||||||
to-preserves-≈ {[]} empty {tt} {tt} _ = ((λ k v ()), (λ k v ()))
|
|
||||||
to-preserves-≈ {k ∷ ks'} uks@(push k≢ks' uks') {ip₁@(v₁ , rest₁)} {ip₂@(v₂ , rest₂)} (v₁≈v₂ , rest₁≈rest₂) = (fm₁⊆fm₂ , fm₂⊆fm₁)
|
|
||||||
where
|
|
||||||
inductive-step : ∀ {v₁ v₂ : B} {rest₁ rest₂ : IterProd (length ks')} →
|
|
||||||
v₁ ≈₂ v₂ → _≈ⁱᵖ_ {length ks'} rest₁ rest₂ →
|
|
||||||
to uks (v₁ , rest₁) ⊆ᵐ to uks (v₂ , rest₂)
|
|
||||||
inductive-step {v₁} {v₂} {rest₁} {rest₂} v₁≈v₂ rest₁≈rest₂ k v k,v∈kvs₁
|
|
||||||
with ((fm'₁ , ufm'₁) , fm'₁≡ks') ← to uks' rest₁ in p₁
|
|
||||||
with ((fm'₂ , ufm'₂) , fm'₂≡ks') ← to uks' rest₂ in p₂
|
|
||||||
with k,v∈kvs₁
|
|
||||||
... | here refl = (v₂ , (v₁≈v₂ , here refl))
|
|
||||||
... | there k,v∈fm'₁ with refl ← p₁ with refl ← p₂ =
|
|
||||||
let
|
|
||||||
(fm'₁⊆fm'₂ , _) = to-preserves-≈ uks' {rest₁} {rest₂}
|
|
||||||
rest₁≈rest₂
|
|
||||||
(v' , (v≈v' , k,v'∈kvs₁)) = fm'₁⊆fm'₂ k v k,v∈fm'₁
|
|
||||||
in
|
|
||||||
(v' , (v≈v' , there k,v'∈kvs₁))
|
|
||||||
|
|
||||||
fm₁⊆fm₂ : to uks ip₁ ⊆ᵐ to uks ip₂
|
|
||||||
fm₁⊆fm₂ = inductive-step v₁≈v₂ rest₁≈rest₂
|
|
||||||
|
|
||||||
fm₂⊆fm₁ : to uks ip₂ ⊆ᵐ to uks ip₁
|
|
||||||
fm₂⊆fm₁ = inductive-step (≈₂-sym v₁≈v₂)
|
|
||||||
(IP.≈-sym (length ks') rest₁≈rest₂)
|
|
||||||
|
|
||||||
from-⊔-distr : ∀ {ks : List A} → (fm₁ fm₂ : FiniteMap ks) →
|
|
||||||
_≈ⁱᵖ_ {length ks} (from (fm₁ ⊔ᵐ fm₂))
|
|
||||||
(_⊔ⁱᵖ_ {ks} (from fm₁) (from fm₂))
|
|
||||||
from-⊔-distr {[]} fm₁ fm₂ = IsEquivalence.≈-refl ≈ᵘ-equiv
|
|
||||||
from-⊔-distr {k ∷ ks} fm₁@(m₁ , _) fm₂@(m₂ , _)
|
|
||||||
with first-key-in-map (fm₁ ⊔ᵐ fm₂)
|
|
||||||
| first-key-in-map fm₁
|
|
||||||
| first-key-in-map fm₂
|
|
||||||
| from-first-value (fm₁ ⊔ᵐ fm₂)
|
|
||||||
| from-first-value fm₁ | from-first-value fm₂
|
|
||||||
... | (v , k,v∈fm₁fm₂) | (v₁ , k,v₁∈fm₁) | (v₂ , k,v₂∈fm₂) | refl | refl | refl
|
|
||||||
with Expr-Provenance k ((` m₁) ∪ (` m₂)) (forget k,v∈fm₁fm₂)
|
|
||||||
... | (_ , (in₁ _ k∉km₂ , _)) = ⊥-elim (k∉km₂ (forget k,v₂∈fm₂))
|
|
||||||
... | (_ , (in₂ k∉km₁ _ , _)) = ⊥-elim (k∉km₁ (forget k,v₁∈fm₁))
|
|
||||||
... | (v₁⊔v₂ , (bothᵘ {v₁'} {v₂'} (single k,v₁'∈m₁) (single k,v₂'∈m₂) , k,v₁⊔v₂∈m₁m₂))
|
|
||||||
rewrite Map-functional {m = m₁} k,v₁∈fm₁ k,v₁'∈m₁
|
|
||||||
rewrite Map-functional {m = m₂} k,v₂∈fm₂ k,v₂'∈m₂
|
|
||||||
rewrite Map-functional {m = proj₁ (fm₁ ⊔ᵐ fm₂)} k,v∈fm₁fm₂ k,v₁⊔v₂∈m₁m₂
|
|
||||||
rewrite from-rest (fm₁ ⊔ᵐ fm₂) rewrite from-rest fm₁ rewrite from-rest fm₂
|
|
||||||
= ( IsLattice.≈-refl lB
|
|
||||||
, IsEquivalence.≈-trans
|
|
||||||
(IP.≈-equiv (length ks))
|
|
||||||
(from-preserves-≈ {_} {pop (fm₁ ⊔ᵐ fm₂)}
|
|
||||||
{pop fm₁ ⊔ᵐ pop fm₂}
|
|
||||||
(pop-⊔-distr fm₁ fm₂))
|
|
||||||
((from-⊔-distr (pop fm₁) (pop fm₂)))
|
|
||||||
)
|
|
||||||
|
|
||||||
|
|
||||||
to-⊔-distr : ∀ {ks : List A} (uks : Unique ks) → (ip₁ ip₂ : IterProd (length ks)) →
|
|
||||||
to uks (_⊔ⁱᵖ_ {ks} ip₁ ip₂) ≈ᵐ (to uks ip₁ ⊔ᵐ to uks ip₂)
|
|
||||||
to-⊔-distr {[]} empty tt tt = ((λ k v ()), (λ k v ()))
|
|
||||||
to-⊔-distr {ks@(k ∷ ks')} uks@(push k≢ks' uks') ip₁@(v₁ , rest₁) ip₂@(v₂ , rest₂) = (fm⊆fm₁fm₂ , fm₁fm₂⊆fm)
|
|
||||||
where
|
|
||||||
fm₁ = to uks ip₁
|
|
||||||
fm₁' = to uks' rest₁
|
|
||||||
fm₂ = to uks ip₂
|
|
||||||
fm₂' = to uks' rest₂
|
|
||||||
fm = to uks (_⊔ⁱᵖ_ {k ∷ ks'} ip₁ ip₂)
|
|
||||||
|
|
||||||
fm⊆fm₁fm₂ : fm ⊆ᵐ (fm₁ ⊔ᵐ fm₂)
|
|
||||||
fm⊆fm₁fm₂ k v (here refl) =
|
|
||||||
(v₁ ⊔₂ v₂
|
|
||||||
, (IsLattice.≈-refl lB
|
|
||||||
, ⊔-combines {k} {v₁} {v₂} {proj₁ fm₁} {proj₁ fm₂}
|
|
||||||
(here refl) (here refl)
|
|
||||||
)
|
|
||||||
)
|
|
||||||
fm⊆fm₁fm₂ k' v (there k',v∈fm')
|
|
||||||
with (fm'⊆fm'₁fm'₂ , _) ← to-⊔-distr uks' rest₁ rest₂
|
|
||||||
with (v' , (v₁⊔v₂≈v' , k',v'∈fm'₁fm'₂))
|
|
||||||
← fm'⊆fm'₁fm'₂ k' v k',v∈fm'
|
|
||||||
with (_ , (refl , (v₁∈fm'₁ , v₂∈fm'₂)))
|
|
||||||
← Provenance-union fm₁' fm₂' k',v'∈fm'₁fm'₂ =
|
|
||||||
( v'
|
|
||||||
, ( v₁⊔v₂≈v'
|
|
||||||
, ⊔-combines {m₁ = proj₁ fm₁} {m₂ = proj₁ fm₂}
|
|
||||||
(there v₁∈fm'₁) (there v₂∈fm'₂)
|
|
||||||
)
|
|
||||||
)
|
|
||||||
|
|
||||||
fm₁fm₂⊆fm : (fm₁ ⊔ᵐ fm₂) ⊆ᵐ fm
|
|
||||||
fm₁fm₂⊆fm k' v k',v∈fm₁fm₂
|
|
||||||
with (_ , fm'₁fm'₂⊆fm')
|
|
||||||
← to-⊔-distr uks' rest₁ rest₂
|
|
||||||
with (_ , (refl , (v₁∈fm₁ , v₂∈fm₂)))
|
|
||||||
← Provenance-union fm₁ fm₂ k',v∈fm₁fm₂
|
|
||||||
with v₁∈fm₁ | v₂∈fm₂
|
|
||||||
... | here refl | here refl =
|
|
||||||
(v , (IsLattice.≈-refl lB , here refl))
|
|
||||||
... | here refl | there k',v₂∈fm₂' =
|
|
||||||
⊥-elim (All¬-¬Any k≢ks' (subst (k' ∈ˡ_) (proj₂ fm₂')
|
|
||||||
(forget k',v₂∈fm₂')))
|
|
||||||
... | there k',v₁∈fm₁' | here refl =
|
|
||||||
⊥-elim (All¬-¬Any k≢ks' (subst (k' ∈ˡ_) (proj₂ fm₁')
|
|
||||||
(forget k',v₁∈fm₁')))
|
|
||||||
... | there k',v₁∈fm₁' | there k',v₂∈fm₂' =
|
|
||||||
let
|
|
||||||
k',v₁v₂∈fm₁'fm₂' =
|
|
||||||
⊔-combines {m₁ = proj₁ fm₁'} {m₂ = proj₁ fm₂'}
|
|
||||||
k',v₁∈fm₁' k',v₂∈fm₂'
|
|
||||||
(v' , (v₁⊔v₂≈v' , v'∈fm')) =
|
|
||||||
fm'₁fm'₂⊆fm' _ _ k',v₁v₂∈fm₁'fm₂'
|
|
||||||
in
|
|
||||||
(v' , (v₁⊔v₂≈v' , there v'∈fm'))
|
|
||||||
|
|
||||||
module WithUniqueKeysAndFixedHeight {ks : List A} (uks : Unique ks) (≈₂-dec : Decidable _≈₂_) (h₂ : ℕ) (fhB : FixedHeight₂ h₂) where
|
|
||||||
import Isomorphism
|
|
||||||
open Isomorphism.TransportFiniteHeight
|
|
||||||
(IP.isFiniteHeightLattice (length ks) ≈₂-dec ≈ᵘ-dec h₂ 0 fhB fixedHeightᵘ) (isLattice ks)
|
|
||||||
{f = to uks} {g = from {ks}}
|
|
||||||
(to-preserves-≈ uks) (from-preserves-≈ {ks})
|
|
||||||
(to-⊔-distr uks) (from-⊔-distr {ks})
|
|
||||||
(from-to-inverseʳ uks) (from-to-inverseˡ uks)
|
|
||||||
using (isFiniteHeightLattice; finiteHeightLattice) public
|
|
||||||
@@ -1,19 +1,22 @@
|
|||||||
open import Lattice
|
open import Lattice
|
||||||
|
open import Data.Unit using (⊤)
|
||||||
|
|
||||||
-- Due to universe levels, it becomes relatively annoying to handle the case
|
-- Due to universe levels, it becomes relatively annoying to handle the case
|
||||||
-- where the levels of A and B are not the same. For the time being, constrain
|
-- where the levels of A and B are not the same. For the time being, constrain
|
||||||
-- them to be the same.
|
-- them to be the same.
|
||||||
|
|
||||||
module Lattice.IterProd {a} {A B : Set a}
|
module Lattice.IterProd {a} (A B : Set a)
|
||||||
(_≈₁_ : A → A → Set a) (_≈₂_ : B → B → Set a)
|
{_≈₁_ : A → A → Set a} {_≈₂_ : B → B → Set a}
|
||||||
(_⊔₁_ : A → A → A) (_⊔₂_ : B → B → B)
|
{_⊔₁_ : A → A → A} {_⊔₂_ : B → B → B}
|
||||||
(_⊓₁_ : A → A → A) (_⊓₂_ : B → B → B)
|
{_⊓₁_ : A → A → A} {_⊓₂_ : B → B → B}
|
||||||
(lA : IsLattice A _≈₁_ _⊔₁_ _⊓₁_) (lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
{{lA : IsLattice A _≈₁_ _⊔₁_ _⊓₁_}} {{lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_}} (dummy : ⊤) where
|
||||||
|
|
||||||
open import Agda.Primitive using (lsuc)
|
open import Agda.Primitive using (lsuc)
|
||||||
open import Data.Nat using (ℕ; suc; _+_)
|
open import Data.Nat using (ℕ; zero; suc; _+_)
|
||||||
open import Data.Product using (_×_)
|
open import Data.Product using (_×_; _,_; proj₁; proj₂)
|
||||||
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; cong)
|
||||||
open import Utils using (iterate)
|
open import Utils using (iterate)
|
||||||
|
open import Chain using (Height)
|
||||||
|
|
||||||
open IsLattice lA renaming (FixedHeight to FixedHeight₁)
|
open IsLattice lA renaming (FixedHeight to FixedHeight₁)
|
||||||
open IsLattice lB renaming (FixedHeight to FixedHeight₂)
|
open IsLattice lB renaming (FixedHeight to FixedHeight₂)
|
||||||
@@ -30,31 +33,50 @@ IterProd k = iterate k (λ t → A × t) B
|
|||||||
-- that are built up by the two iterations. So, do everything in one iteration.
|
-- that are built up by the two iterations. So, do everything in one iteration.
|
||||||
-- This requires some odd code.
|
-- This requires some odd code.
|
||||||
|
|
||||||
|
build : A → B → (k : ℕ) → IterProd k
|
||||||
|
build _ b zero = b
|
||||||
|
build a b (suc s) = (a , build a b s)
|
||||||
|
|
||||||
private
|
private
|
||||||
record RequiredForFixedHeight : Set (lsuc a) where
|
record RequiredForFixedHeight : Set (lsuc a) where
|
||||||
field
|
field
|
||||||
≈₁-dec : IsDecidable _≈₁_
|
{{≈₁-Decidable}} : IsDecidable _≈₁_
|
||||||
≈₂-dec : IsDecidable _≈₂_
|
{{≈₂-Decidable}} : IsDecidable _≈₂_
|
||||||
h₁ h₂ : ℕ
|
h₁ h₂ : ℕ
|
||||||
fhA : FixedHeight₁ h₁
|
{{fhA}} : FixedHeight₁ h₁
|
||||||
fhB : FixedHeight₂ h₂
|
{{fhB}} : FixedHeight₂ h₂
|
||||||
|
|
||||||
record IsFiniteHeightAndDecEq {A : Set a} {_≈_ : A → A → Set a} {_⊔_ : A → A → A} {_⊓_ : A → A → A} (isLattice : IsLattice A _≈_ _⊔_ _⊓_) : Set (lsuc a) where
|
⊥₁ : A
|
||||||
|
⊥₁ = Height.⊥ fhA
|
||||||
|
|
||||||
|
⊥₂ : B
|
||||||
|
⊥₂ = Height.⊥ fhB
|
||||||
|
|
||||||
|
⊥k : ∀ (k : ℕ) → IterProd k
|
||||||
|
⊥k = build ⊥₁ ⊥₂
|
||||||
|
|
||||||
|
record IsFiniteHeightWithBotAndDecEq {A : Set a} {_≈_ : A → A → Set a} {_⊔_ : A → A → A} {_⊓_ : A → A → A} (isLattice : IsLattice A _≈_ _⊔_ _⊓_) (⊥ : A) : Set (lsuc a) where
|
||||||
field
|
field
|
||||||
height : ℕ
|
height : ℕ
|
||||||
fixedHeight : IsLattice.FixedHeight isLattice height
|
fixedHeight : IsLattice.FixedHeight isLattice height
|
||||||
≈-dec : IsDecidable _≈_
|
≈-Decidable : IsDecidable _≈_
|
||||||
|
|
||||||
|
⊥-correct : Height.⊥ fixedHeight ≡ ⊥
|
||||||
|
|
||||||
|
record Everything (k : ℕ) : Set (lsuc a) where
|
||||||
|
T = IterProd k
|
||||||
|
|
||||||
record Everything (A : Set a) : Set (lsuc a) where
|
|
||||||
field
|
field
|
||||||
_≈_ : A → A → Set a
|
_≈_ : T → T → Set a
|
||||||
_⊔_ : A → A → A
|
_⊔_ : T → T → T
|
||||||
_⊓_ : A → A → A
|
_⊓_ : T → T → T
|
||||||
|
|
||||||
isLattice : IsLattice A _≈_ _⊔_ _⊓_
|
isLattice : IsLattice T _≈_ _⊔_ _⊓_
|
||||||
isFiniteHeightIfSupported : RequiredForFixedHeight → IsFiniteHeightAndDecEq isLattice
|
isFiniteHeightIfSupported :
|
||||||
|
∀ (req : RequiredForFixedHeight) →
|
||||||
|
IsFiniteHeightWithBotAndDecEq isLattice (RequiredForFixedHeight.⊥k req k)
|
||||||
|
|
||||||
everything : ∀ (k : ℕ) → Everything (IterProd k)
|
everything : ∀ (k : ℕ) → Everything k
|
||||||
everything 0 = record
|
everything 0 = record
|
||||||
{ _≈_ = _≈₂_
|
{ _≈_ = _≈₂_
|
||||||
; _⊔_ = _⊔₂_
|
; _⊔_ = _⊔₂_
|
||||||
@@ -63,7 +85,8 @@ private
|
|||||||
; isFiniteHeightIfSupported = λ req → record
|
; isFiniteHeightIfSupported = λ req → record
|
||||||
{ height = RequiredForFixedHeight.h₂ req
|
{ height = RequiredForFixedHeight.h₂ req
|
||||||
; fixedHeight = RequiredForFixedHeight.fhB req
|
; fixedHeight = RequiredForFixedHeight.fhB req
|
||||||
; ≈-dec = RequiredForFixedHeight.≈₂-dec req
|
; ≈-Decidable = RequiredForFixedHeight.≈₂-Decidable req
|
||||||
|
; ⊥-correct = refl
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
everything (suc k') = record
|
everything (suc k') = record
|
||||||
@@ -76,61 +99,70 @@ private
|
|||||||
fhlRest = Everything.isFiniteHeightIfSupported everythingRest req
|
fhlRest = Everything.isFiniteHeightIfSupported everythingRest req
|
||||||
in
|
in
|
||||||
record
|
record
|
||||||
{ height = (RequiredForFixedHeight.h₁ req) + IsFiniteHeightAndDecEq.height fhlRest
|
{ height = (RequiredForFixedHeight.h₁ req) + IsFiniteHeightWithBotAndDecEq.height fhlRest
|
||||||
; fixedHeight =
|
; fixedHeight =
|
||||||
P.fixedHeight
|
P.fixedHeight
|
||||||
(RequiredForFixedHeight.≈₁-dec req) (IsFiniteHeightAndDecEq.≈-dec fhlRest)
|
{{≈₂-Decidable = IsFiniteHeightWithBotAndDecEq.≈-Decidable fhlRest}}
|
||||||
(RequiredForFixedHeight.h₁ req) (IsFiniteHeightAndDecEq.height fhlRest)
|
{{fhB = IsFiniteHeightWithBotAndDecEq.fixedHeight fhlRest}}
|
||||||
(RequiredForFixedHeight.fhA req) (IsFiniteHeightAndDecEq.fixedHeight fhlRest)
|
; ≈-Decidable = P.≈-Decidable {{≈₂-Decidable = IsFiniteHeightWithBotAndDecEq.≈-Decidable fhlRest}}
|
||||||
; ≈-dec = P.≈-dec (RequiredForFixedHeight.≈₁-dec req) (IsFiniteHeightAndDecEq.≈-dec fhlRest)
|
; ⊥-correct =
|
||||||
|
cong ((Height.⊥ (RequiredForFixedHeight.fhA req)) ,_)
|
||||||
|
(IsFiniteHeightWithBotAndDecEq.⊥-correct fhlRest)
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
where
|
where
|
||||||
everythingRest = everything k'
|
everythingRest = everything k'
|
||||||
|
import Lattice.Prod A (IterProd k') {{lB = Everything.isLattice everythingRest}} as P
|
||||||
|
|
||||||
import Lattice.Prod
|
module _ {k : ℕ} where
|
||||||
_≈₁_ (Everything._≈_ everythingRest)
|
open Everything (everything k) using (_≈_; _⊔_; _⊓_) public
|
||||||
_⊔₁_ (Everything._⊔_ everythingRest)
|
open Lattice.IsLattice (Everything.isLattice (everything k)) public
|
||||||
_⊓₁_ (Everything._⊓_ everythingRest)
|
|
||||||
lA (Everything.isLattice everythingRest) as P
|
|
||||||
|
|
||||||
module _ (k : ℕ) where
|
instance
|
||||||
open Everything (everything k) using (_≈_; _⊔_; _⊓_; isLattice) public
|
isLattice = Everything.isLattice (everything k)
|
||||||
open Lattice.IsLattice isLattice public
|
|
||||||
|
|
||||||
lattice : Lattice (IterProd k)
|
lattice : Lattice (IterProd k)
|
||||||
lattice = record
|
lattice = record
|
||||||
{ _≈_ = _≈_
|
{ _≈_ = _≈_
|
||||||
; _⊔_ = _⊔_
|
|
||||||
; _⊓_ = _⊓_
|
|
||||||
; isLattice = isLattice
|
|
||||||
}
|
|
||||||
|
|
||||||
module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_)
|
|
||||||
(h₁ h₂ : ℕ)
|
|
||||||
(fhA : FixedHeight₁ h₁) (fhB : FixedHeight₂ h₂) where
|
|
||||||
|
|
||||||
private
|
|
||||||
required : RequiredForFixedHeight
|
|
||||||
required = record
|
|
||||||
{ ≈₁-dec = ≈₁-dec
|
|
||||||
; ≈₂-dec = ≈₂-dec
|
|
||||||
; h₁ = h₁
|
|
||||||
; h₂ = h₂
|
|
||||||
; fhA = fhA
|
|
||||||
; fhB = fhB
|
|
||||||
}
|
|
||||||
|
|
||||||
isFiniteHeightLattice = record
|
|
||||||
{ isLattice = isLattice
|
|
||||||
; fixedHeight = IsFiniteHeightAndDecEq.fixedHeight (Everything.isFiniteHeightIfSupported (everything k) required)
|
|
||||||
}
|
|
||||||
|
|
||||||
finiteHeightLattice : FiniteHeightLattice (IterProd k)
|
|
||||||
finiteHeightLattice = record
|
|
||||||
{ height = IsFiniteHeightAndDecEq.height (Everything.isFiniteHeightIfSupported (everything k) required)
|
|
||||||
; _≈_ = _≈_
|
|
||||||
; _⊔_ = _⊔_
|
; _⊔_ = _⊔_
|
||||||
; _⊓_ = _⊓_
|
; _⊓_ = _⊓_
|
||||||
; isFiniteHeightLattice = isFiniteHeightLattice
|
; isLattice = isLattice
|
||||||
}
|
}
|
||||||
|
|
||||||
|
module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDecidable _≈₂_}}
|
||||||
|
{h₁ h₂ : ℕ}
|
||||||
|
{{fhA : FixedHeight₁ h₁}} {{fhB : FixedHeight₂ h₂}} where
|
||||||
|
|
||||||
|
private
|
||||||
|
isFiniteHeightWithBotAndDecEq =
|
||||||
|
Everything.isFiniteHeightIfSupported (everything k)
|
||||||
|
record
|
||||||
|
{ ≈₁-Decidable = ≈₁-Decidable
|
||||||
|
; ≈₂-Decidable = ≈₂-Decidable
|
||||||
|
; h₁ = h₁
|
||||||
|
; h₂ = h₂
|
||||||
|
; fhA = fhA
|
||||||
|
; fhB = fhB
|
||||||
|
}
|
||||||
|
open IsFiniteHeightWithBotAndDecEq isFiniteHeightWithBotAndDecEq using (height; ⊥-correct)
|
||||||
|
|
||||||
|
instance
|
||||||
|
fixedHeight = IsFiniteHeightWithBotAndDecEq.fixedHeight isFiniteHeightWithBotAndDecEq
|
||||||
|
|
||||||
|
isFiniteHeightLattice = record
|
||||||
|
{ isLattice = isLattice
|
||||||
|
; fixedHeight = fixedHeight
|
||||||
|
}
|
||||||
|
|
||||||
|
finiteHeightLattice : FiniteHeightLattice (IterProd k)
|
||||||
|
finiteHeightLattice = record
|
||||||
|
{ height = height
|
||||||
|
; _≈_ = _≈_
|
||||||
|
; _⊔_ = _⊔_
|
||||||
|
; _⊓_ = _⊓_
|
||||||
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
||||||
|
}
|
||||||
|
|
||||||
|
⊥-built : Height.⊥ fixedHeight ≡ (build (Height.⊥ fhA) (Height.⊥ fhB) k)
|
||||||
|
⊥-built = ⊥-correct
|
||||||
|
|
||||||
|
|||||||
@@ -1,13 +1,14 @@
|
|||||||
open import Lattice
|
open import Lattice
|
||||||
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym; trans; cong; subst)
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym; trans; cong; subst)
|
||||||
open import Relation.Binary.Definitions using (Decidable)
|
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
|
open import Data.Unit using (⊤)
|
||||||
|
|
||||||
module Lattice.Map {a b : Level} {A : Set a} {B : Set b}
|
module Lattice.Map {a b : Level} (A : Set a) (B : Set b)
|
||||||
{_≈₂_ : B → B → Set b}
|
{_≈₂_ : B → B → Set b}
|
||||||
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
||||||
(≡-dec-A : Decidable (_≡_ {a} {A}))
|
{{≡-Decidable-A : IsDecidable {a} {A} _≡_}}
|
||||||
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
{{lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_}}
|
||||||
|
(dummy : ⊤) where
|
||||||
|
|
||||||
open import Data.List.Membership.Propositional as MemProp using () renaming (_∈_ to _∈ˡ_)
|
open import Data.List.Membership.Propositional as MemProp using () renaming (_∈_ to _∈ˡ_)
|
||||||
|
|
||||||
@@ -23,6 +24,8 @@ open import Utils using (Unique; push; Unique-append; All¬-¬Any; All-x∈xs)
|
|||||||
open import Data.String using () renaming (_++_ to _++ˢ_)
|
open import Data.String using () renaming (_++_ to _++ˢ_)
|
||||||
open import Showable using (Showable; show)
|
open import Showable using (Showable; show)
|
||||||
|
|
||||||
|
open IsDecidable ≡-Decidable-A using () renaming (R-dec to _≟ᴬ_)
|
||||||
|
|
||||||
open IsLattice lB using () renaming
|
open IsLattice lB using () renaming
|
||||||
( ≈-refl to ≈₂-refl; ≈-sym to ≈₂-sym; ≈-trans to ≈₂-trans
|
( ≈-refl to ≈₂-refl; ≈-sym to ≈₂-sym; ≈-trans to ≈₂-trans
|
||||||
; ≈-⊔-cong to ≈₂-⊔₂-cong; ≈-⊓-cong to ≈₂-⊓₂-cong
|
; ≈-⊔-cong to ≈₂-⊔₂-cong; ≈-⊓-cong to ≈₂-⊓₂-cong
|
||||||
@@ -41,7 +44,7 @@ private module ImplKeys where
|
|||||||
∈k-dec : ∀ (k : A) (l : List (A × B)) → Dec (k ∈ˡ (ImplKeys.keys l))
|
∈k-dec : ∀ (k : A) (l : List (A × B)) → Dec (k ∈ˡ (ImplKeys.keys l))
|
||||||
∈k-dec k [] = no (λ ())
|
∈k-dec k [] = no (λ ())
|
||||||
∈k-dec k ((k' , v) ∷ xs)
|
∈k-dec k ((k' , v) ∷ xs)
|
||||||
with (≡-dec-A k k')
|
with (k ≟ᴬ k')
|
||||||
... | yes k≡k' = yes (here k≡k')
|
... | yes k≡k' = yes (here k≡k')
|
||||||
... | no k≢k' with (∈k-dec k xs)
|
... | no k≢k' with (∈k-dec k xs)
|
||||||
... | yes k∈kxs = yes (there k∈kxs)
|
... | yes k∈kxs = yes (there k∈kxs)
|
||||||
@@ -76,7 +79,7 @@ private module _ where
|
|||||||
k∈-dec : ∀ (k : A) (l : List A) → Dec (k ∈ l)
|
k∈-dec : ∀ (k : A) (l : List A) → Dec (k ∈ l)
|
||||||
k∈-dec k [] = no (λ ())
|
k∈-dec k [] = no (λ ())
|
||||||
k∈-dec k (x ∷ xs)
|
k∈-dec k (x ∷ xs)
|
||||||
with (≡-dec-A k x)
|
with (k ≟ᴬ x)
|
||||||
... | yes refl = yes (here refl)
|
... | yes refl = yes (here refl)
|
||||||
... | no k≢x with (k∈-dec k xs)
|
... | no k≢x with (k∈-dec k xs)
|
||||||
... | yes k∈xs = yes (there k∈xs)
|
... | yes k∈xs = yes (there k∈xs)
|
||||||
@@ -113,7 +116,7 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
|
|
||||||
insert : A → B → List (A × B) → List (A × B)
|
insert : A → B → List (A × B) → List (A × B)
|
||||||
insert k v [] = (k , v) ∷ []
|
insert k v [] = (k , v) ∷ []
|
||||||
insert k v (x@(k' , v') ∷ xs) with ≡-dec-A k k'
|
insert k v (x@(k' , v') ∷ xs) with k ≟ᴬ k'
|
||||||
... | yes _ = (k' , f v v') ∷ xs
|
... | yes _ = (k' , f v v') ∷ xs
|
||||||
... | no _ = x ∷ insert k v xs
|
... | no _ = x ∷ insert k v xs
|
||||||
|
|
||||||
@@ -123,11 +126,11 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
insert-keys-∈ : ∀ {k : A} {v : B} {l : List (A × B)} →
|
insert-keys-∈ : ∀ {k : A} {v : B} {l : List (A × B)} →
|
||||||
k ∈k l → keys l ≡ keys (insert k v l)
|
k ∈k l → keys l ≡ keys (insert k v l)
|
||||||
insert-keys-∈ {k} {v} {(k' , v') ∷ xs} (here k≡k')
|
insert-keys-∈ {k} {v} {(k' , v') ∷ xs} (here k≡k')
|
||||||
with (≡-dec-A k k')
|
with (k ≟ᴬ k')
|
||||||
... | yes _ = refl
|
... | yes _ = refl
|
||||||
... | no k≢k' = ⊥-elim (k≢k' k≡k')
|
... | no k≢k' = ⊥-elim (k≢k' k≡k')
|
||||||
insert-keys-∈ {k} {v} {(k' , _) ∷ xs} (there k∈kxs)
|
insert-keys-∈ {k} {v} {(k' , _) ∷ xs} (there k∈kxs)
|
||||||
with (≡-dec-A k k')
|
with (k ≟ᴬ k')
|
||||||
... | yes _ = refl
|
... | yes _ = refl
|
||||||
... | no _ = cong (λ xs' → k' ∷ xs') (insert-keys-∈ k∈kxs)
|
... | no _ = cong (λ xs' → k' ∷ xs') (insert-keys-∈ k∈kxs)
|
||||||
|
|
||||||
@@ -135,7 +138,7 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
¬ (k ∈k l) → (keys l ++ (k ∷ [])) ≡ keys (insert k v l)
|
¬ (k ∈k l) → (keys l ++ (k ∷ [])) ≡ keys (insert k v l)
|
||||||
insert-keys-∉ {k} {v} {[]} _ = refl
|
insert-keys-∉ {k} {v} {[]} _ = refl
|
||||||
insert-keys-∉ {k} {v} {(k' , v') ∷ xs} k∉kl
|
insert-keys-∉ {k} {v} {(k' , v') ∷ xs} k∉kl
|
||||||
with (≡-dec-A k k')
|
with (k ≟ᴬ k')
|
||||||
... | yes k≡k' = ⊥-elim (k∉kl (here k≡k'))
|
... | yes k≡k' = ⊥-elim (k∉kl (here k≡k'))
|
||||||
... | no _ = cong (λ xs' → k' ∷ xs')
|
... | no _ = cong (λ xs' → k' ∷ xs')
|
||||||
(insert-keys-∉ (λ k∈kxs → k∉kl (there k∈kxs)))
|
(insert-keys-∉ (λ k∈kxs → k∉kl (there k∈kxs)))
|
||||||
@@ -171,7 +174,7 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
¬ k ∈k l → (k , v) ∈ insert k v l
|
¬ k ∈k l → (k , v) ∈ insert k v l
|
||||||
insert-fresh {l = []} k∉kl = here refl
|
insert-fresh {l = []} k∉kl = here refl
|
||||||
insert-fresh {k} {l = (k' , v') ∷ xs} k∉kl
|
insert-fresh {k} {l = (k' , v') ∷ xs} k∉kl
|
||||||
with ≡-dec-A k k'
|
with k ≟ᴬ k'
|
||||||
... | yes k≡k' = ⊥-elim (k∉kl (here k≡k'))
|
... | yes k≡k' = ⊥-elim (k∉kl (here k≡k'))
|
||||||
... | no _ = there (insert-fresh (λ k∈kxs → k∉kl (there k∈kxs)))
|
... | no _ = there (insert-fresh (λ k∈kxs → k∉kl (there k∈kxs)))
|
||||||
|
|
||||||
@@ -180,9 +183,9 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
insert-preserves-∉k {l = []} k≢k' k∉kl (here k≡k') = k≢k' k≡k'
|
insert-preserves-∉k {l = []} k≢k' k∉kl (here k≡k') = k≢k' k≡k'
|
||||||
insert-preserves-∉k {l = []} k≢k' k∉kl (there ())
|
insert-preserves-∉k {l = []} k≢k' k∉kl (there ())
|
||||||
insert-preserves-∉k {k} {k'} {v'} {(k'' , v'') ∷ xs} k≢k' k∉kl k∈kil
|
insert-preserves-∉k {k} {k'} {v'} {(k'' , v'') ∷ xs} k≢k' k∉kl k∈kil
|
||||||
with ≡-dec-A k k''
|
with k ≟ᴬ k''
|
||||||
... | yes k≡k'' = k∉kl (here k≡k'')
|
... | yes k≡k'' = k∉kl (here k≡k'')
|
||||||
... | no k≢k'' with ≡-dec-A k' k'' | k∈kil
|
... | no k≢k'' with k' ≟ᴬ k'' | k∈kil
|
||||||
... | yes k'≡k'' | here k≡k'' = k≢k'' k≡k''
|
... | yes k'≡k'' | here k≡k'' = k≢k'' k≡k''
|
||||||
... | yes k'≡k'' | there k∈kxs = k∉kl (there k∈kxs)
|
... | yes k'≡k'' | there k∈kxs = k∉kl (there k∈kxs)
|
||||||
... | no k'≢k'' | here k≡k'' = k∉kl (here k≡k'')
|
... | no k'≢k'' | here k≡k'' = k∉kl (here k≡k'')
|
||||||
@@ -193,18 +196,18 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
¬ k ∈k l₁ → ¬ k ∈k l₂ → ¬ k ∈k union l₁ l₂
|
¬ k ∈k l₁ → ¬ k ∈k l₂ → ¬ k ∈k union l₁ l₂
|
||||||
union-preserves-∉ {l₁ = []} _ k∉kl₂ = k∉kl₂
|
union-preserves-∉ {l₁ = []} _ k∉kl₂ = k∉kl₂
|
||||||
union-preserves-∉ {k} {(k' , v') ∷ xs₁} k∉kl₁ k∉kl₂
|
union-preserves-∉ {k} {(k' , v') ∷ xs₁} k∉kl₁ k∉kl₂
|
||||||
with ≡-dec-A k k'
|
with k ≟ᴬ k'
|
||||||
... | yes k≡k' = ⊥-elim (k∉kl₁ (here k≡k'))
|
... | yes k≡k' = ⊥-elim (k∉kl₁ (here k≡k'))
|
||||||
... | no k≢k' = insert-preserves-∉k k≢k' (union-preserves-∉ (λ k∈kxs₁ → k∉kl₁ (there k∈kxs₁)) k∉kl₂)
|
... | no k≢k' = insert-preserves-∉k k≢k' (union-preserves-∉ (λ k∈kxs₁ → k∉kl₁ (there k∈kxs₁)) k∉kl₂)
|
||||||
|
|
||||||
insert-preserves-∈k : ∀ {k k' : A} {v' : B} {l : List (A × B)} →
|
insert-preserves-∈k : ∀ {k k' : A} {v' : B} {l : List (A × B)} →
|
||||||
k ∈k l → k ∈k insert k' v' l
|
k ∈k l → k ∈k insert k' v' l
|
||||||
insert-preserves-∈k {k} {k'} {v'} {(k'' , v'') ∷ xs} (here k≡k'')
|
insert-preserves-∈k {k} {k'} {v'} {(k'' , v'') ∷ xs} (here k≡k'')
|
||||||
with (≡-dec-A k' k'')
|
with k' ≟ᴬ k''
|
||||||
... | yes _ = here k≡k''
|
... | yes _ = here k≡k''
|
||||||
... | no _ = here k≡k''
|
... | no _ = here k≡k''
|
||||||
insert-preserves-∈k {k} {k'} {v'} {(k'' , v'') ∷ xs} (there k∈kxs)
|
insert-preserves-∈k {k} {k'} {v'} {(k'' , v'') ∷ xs} (there k∈kxs)
|
||||||
with (≡-dec-A k' k'')
|
with k' ≟ᴬ k''
|
||||||
... | yes _ = there k∈kxs
|
... | yes _ = there k∈kxs
|
||||||
... | no _ = there (insert-preserves-∈k k∈kxs)
|
... | no _ = there (insert-preserves-∈k k∈kxs)
|
||||||
|
|
||||||
@@ -236,11 +239,11 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
insert-preserves-∈ : ∀ {k k' : A} {v v' : B} {l : List (A × B)} →
|
insert-preserves-∈ : ∀ {k k' : A} {v v' : B} {l : List (A × B)} →
|
||||||
¬ k ≡ k' → (k , v) ∈ l → (k , v) ∈ insert k' v' l
|
¬ k ≡ k' → (k , v) ∈ l → (k , v) ∈ insert k' v' l
|
||||||
insert-preserves-∈ {k} {k'} {l = x ∷ xs} k≢k' (here k,v=x)
|
insert-preserves-∈ {k} {k'} {l = x ∷ xs} k≢k' (here k,v=x)
|
||||||
rewrite sym k,v=x with ≡-dec-A k' k
|
rewrite sym k,v=x with k' ≟ᴬ k
|
||||||
... | yes k'≡k = ⊥-elim (k≢k' (sym k'≡k))
|
... | yes k'≡k = ⊥-elim (k≢k' (sym k'≡k))
|
||||||
... | no _ = here refl
|
... | no _ = here refl
|
||||||
insert-preserves-∈ {k} {k'} {l = (k'' , _) ∷ xs} k≢k' (there k,v∈xs)
|
insert-preserves-∈ {k} {k'} {l = (k'' , _) ∷ xs} k≢k' (there k,v∈xs)
|
||||||
with ≡-dec-A k' k''
|
with k' ≟ᴬ k''
|
||||||
... | yes _ = there k,v∈xs
|
... | yes _ = there k,v∈xs
|
||||||
... | no _ = there (insert-preserves-∈ k≢k' k,v∈xs)
|
... | no _ = there (insert-preserves-∈ k≢k' k,v∈xs)
|
||||||
|
|
||||||
@@ -259,7 +262,7 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
k,v∈mxs₁l = union-preserves-∈₁ uxs₁ k,v∈xs₁ k∉kl₂
|
k,v∈mxs₁l = union-preserves-∈₁ uxs₁ k,v∈xs₁ k∉kl₂
|
||||||
|
|
||||||
k≢k' : ¬ k ≡ k'
|
k≢k' : ¬ k ≡ k'
|
||||||
k≢k' with ≡-dec-A k k'
|
k≢k' with k ≟ᴬ k'
|
||||||
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs₁ (∈-cong proj₁ k,v∈xs₁))
|
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs₁ (∈-cong proj₁ k,v∈xs₁))
|
||||||
... | no k≢k' = k≢k'
|
... | no k≢k' = k≢k'
|
||||||
union-preserves-∈₁ {l₁ = (k' , v') ∷ xs₁} (push k'≢xs₁ uxs₁) (here k,v≡k',v') k∉kl₂
|
union-preserves-∈₁ {l₁ = (k' , v') ∷ xs₁} (push k'≢xs₁ uxs₁) (here k,v≡k',v') k∉kl₂
|
||||||
@@ -270,11 +273,11 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
Unique (keys l) → (k , v') ∈ l → (k , f v v') ∈ (insert k v l)
|
Unique (keys l) → (k , v') ∈ l → (k , f v v') ∈ (insert k v l)
|
||||||
insert-combines {l = (k' , v'') ∷ xs} _ (here k,v'≡k',v'')
|
insert-combines {l = (k' , v'') ∷ xs} _ (here k,v'≡k',v'')
|
||||||
rewrite cong proj₁ k,v'≡k',v'' rewrite cong proj₂ k,v'≡k',v''
|
rewrite cong proj₁ k,v'≡k',v'' rewrite cong proj₂ k,v'≡k',v''
|
||||||
with ≡-dec-A k' k'
|
with k' ≟ᴬ k'
|
||||||
... | yes _ = here refl
|
... | yes _ = here refl
|
||||||
... | no k≢k' = ⊥-elim (k≢k' refl)
|
... | no k≢k' = ⊥-elim (k≢k' refl)
|
||||||
insert-combines {k} {l = (k' , v'') ∷ xs} (push k'≢xs uxs) (there k,v'∈xs)
|
insert-combines {k} {l = (k' , v'') ∷ xs} (push k'≢xs uxs) (there k,v'∈xs)
|
||||||
with ≡-dec-A k k'
|
with k ≟ᴬ k'
|
||||||
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs (∈-cong proj₁ k,v'∈xs))
|
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs (∈-cong proj₁ k,v'∈xs))
|
||||||
... | no k≢k' = there (insert-combines uxs k,v'∈xs)
|
... | no k≢k' = there (insert-combines uxs k,v'∈xs)
|
||||||
|
|
||||||
@@ -288,13 +291,13 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
insert-preserves-∈ k≢k' (union-combines uxs₁ ul₂ k,v₁∈xs₁ k,v₂∈l₂)
|
insert-preserves-∈ k≢k' (union-combines uxs₁ ul₂ k,v₁∈xs₁ k,v₂∈l₂)
|
||||||
where
|
where
|
||||||
k≢k' : ¬ k ≡ k'
|
k≢k' : ¬ k ≡ k'
|
||||||
k≢k' with ≡-dec-A k k'
|
k≢k' with k ≟ᴬ k'
|
||||||
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs₁ (∈-cong proj₁ k,v₁∈xs₁))
|
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs₁ (∈-cong proj₁ k,v₁∈xs₁))
|
||||||
... | no k≢k' = k≢k'
|
... | no k≢k' = k≢k'
|
||||||
|
|
||||||
update : A → B → List (A × B) → List (A × B)
|
update : A → B → List (A × B) → List (A × B)
|
||||||
update k v [] = []
|
update k v [] = []
|
||||||
update k v ((k' , v') ∷ xs) with ≡-dec-A k k'
|
update k v ((k' , v') ∷ xs) with k ≟ᴬ k'
|
||||||
... | yes _ = (k' , f v v') ∷ xs
|
... | yes _ = (k' , f v v') ∷ xs
|
||||||
... | no _ = (k' , v') ∷ update k v xs
|
... | no _ = (k' , v') ∷ update k v xs
|
||||||
|
|
||||||
@@ -314,7 +317,7 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
keys l ≡ keys (update k v l)
|
keys l ≡ keys (update k v l)
|
||||||
update-keys {l = []} = refl
|
update-keys {l = []} = refl
|
||||||
update-keys {k} {v} {l = (k' , v') ∷ xs}
|
update-keys {k} {v} {l = (k' , v') ∷ xs}
|
||||||
with ≡-dec-A k k'
|
with k ≟ᴬ k'
|
||||||
... | yes _ = refl
|
... | yes _ = refl
|
||||||
... | no _ rewrite update-keys {k} {v} {xs} = refl
|
... | no _ rewrite update-keys {k} {v} {xs} = refl
|
||||||
|
|
||||||
@@ -431,11 +434,11 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
¬ k ≡ k' → (k , v) ∈ l → (k , v) ∈ update k' v' l
|
¬ k ≡ k' → (k , v) ∈ l → (k , v) ∈ update k' v' l
|
||||||
update-preserves-∈ {k} {k'} {v} {v'} {(k'' , v'') ∷ xs} k≢k' (here k,v≡k'',v'')
|
update-preserves-∈ {k} {k'} {v} {v'} {(k'' , v'') ∷ xs} k≢k' (here k,v≡k'',v'')
|
||||||
rewrite cong proj₁ k,v≡k'',v'' rewrite cong proj₂ k,v≡k'',v''
|
rewrite cong proj₁ k,v≡k'',v'' rewrite cong proj₂ k,v≡k'',v''
|
||||||
with ≡-dec-A k' k''
|
with k' ≟ᴬ k''
|
||||||
... | yes k'≡k'' = ⊥-elim (k≢k' (sym k'≡k''))
|
... | yes k'≡k'' = ⊥-elim (k≢k' (sym k'≡k''))
|
||||||
... | no _ = here refl
|
... | no _ = here refl
|
||||||
update-preserves-∈ {k} {k'} {v} {v'} {(k'' , v'') ∷ xs} k≢k' (there k,v∈xs)
|
update-preserves-∈ {k} {k'} {v} {v'} {(k'' , v'') ∷ xs} k≢k' (there k,v∈xs)
|
||||||
with ≡-dec-A k' k''
|
with k' ≟ᴬ k''
|
||||||
... | yes _ = there k,v∈xs
|
... | yes _ = there k,v∈xs
|
||||||
... | no _ = there (update-preserves-∈ k≢k' k,v∈xs)
|
... | no _ = there (update-preserves-∈ k≢k' k,v∈xs)
|
||||||
|
|
||||||
@@ -449,11 +452,11 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
Unique (keys l) → (k , v) ∈ l → (k , f v' v) ∈ update k v' l
|
Unique (keys l) → (k , v) ∈ l → (k , f v' v) ∈ update k v' l
|
||||||
update-combines {k} {v} {v'} {(k' , v'') ∷ xs} _ (here k,v=k',v'')
|
update-combines {k} {v} {v'} {(k' , v'') ∷ xs} _ (here k,v=k',v'')
|
||||||
rewrite cong proj₁ k,v=k',v'' rewrite cong proj₂ k,v=k',v''
|
rewrite cong proj₁ k,v=k',v'' rewrite cong proj₂ k,v=k',v''
|
||||||
with ≡-dec-A k' k'
|
with k' ≟ᴬ k'
|
||||||
... | yes _ = here refl
|
... | yes _ = here refl
|
||||||
... | no k'≢k' = ⊥-elim (k'≢k' refl)
|
... | no k'≢k' = ⊥-elim (k'≢k' refl)
|
||||||
update-combines {k} {v} {v'} {(k' , v'') ∷ xs} (push k'≢xs uxs) (there k,v∈xs)
|
update-combines {k} {v} {v'} {(k' , v'') ∷ xs} (push k'≢xs uxs) (there k,v∈xs)
|
||||||
with ≡-dec-A k k'
|
with k ≟ᴬ k'
|
||||||
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs (∈-cong proj₁ k,v∈xs))
|
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs (∈-cong proj₁ k,v∈xs))
|
||||||
... | no _ = there (update-combines uxs k,v∈xs)
|
... | no _ = there (update-combines uxs k,v∈xs)
|
||||||
|
|
||||||
@@ -467,7 +470,7 @@ private module ImplInsert (f : B → B → B) where
|
|||||||
update-preserves-∈ k≢k' (updates-combine uxs₁ u₂ k,v₁∈xs k,v₂∈l₂)
|
update-preserves-∈ k≢k' (updates-combine uxs₁ u₂ k,v₁∈xs k,v₂∈l₂)
|
||||||
where
|
where
|
||||||
k≢k' : ¬ k ≡ k'
|
k≢k' : ¬ k ≡ k'
|
||||||
k≢k' with ≡-dec-A k k'
|
k≢k' with k ≟ᴬ k'
|
||||||
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs (∈-cong proj₁ k,v₁∈xs))
|
... | yes k≡k' rewrite k≡k' = ⊥-elim (All¬-¬Any k'≢xs (∈-cong proj₁ k,v₁∈xs))
|
||||||
... | no k≢k' = k≢k'
|
... | no k≢k' = k≢k'
|
||||||
|
|
||||||
@@ -625,7 +628,8 @@ Expr-Provenance-≡ {k} {v} e k,v∈e
|
|||||||
with (v' , (p , k,v'∈e)) ← Expr-Provenance k e (forget k,v∈e)
|
with (v' , (p , k,v'∈e)) ← Expr-Provenance k e (forget k,v∈e)
|
||||||
rewrite Map-functional {m = ⟦ e ⟧} k,v∈e k,v'∈e = p
|
rewrite Map-functional {m = ⟦ e ⟧} k,v∈e k,v'∈e = p
|
||||||
|
|
||||||
module _ (≈₂-dec : IsDecidable _≈₂_) where
|
module _ {{≈₂-Decidable : IsDecidable _≈₂_}} where
|
||||||
|
open IsDecidable ≈₂-Decidable using () renaming (R-dec to ≈₂-dec)
|
||||||
private module _ where
|
private module _ where
|
||||||
data SubsetInfo (m₁ m₂ : Map) : Set (a ⊔ℓ b) where
|
data SubsetInfo (m₁ m₂ : Map) : Set (a ⊔ℓ b) where
|
||||||
extra : (k : A) → k ∈k m₁ → ¬ k ∈k m₂ → SubsetInfo m₁ m₂
|
extra : (k : A) → k ∈k m₁ → ¬ k ∈k m₂ → SubsetInfo m₁ m₂
|
||||||
@@ -676,6 +680,9 @@ module _ (≈₂-dec : IsDecidable _≈₂_) where
|
|||||||
... | _ | no m₂̷⊆m₁ = no (λ (_ , m₂⊆m₁) → m₂̷⊆m₁ m₂⊆m₁)
|
... | _ | no m₂̷⊆m₁ = no (λ (_ , m₂⊆m₁) → m₂̷⊆m₁ m₂⊆m₁)
|
||||||
... | no m₁̷⊆m₂ | _ = no (λ (m₁⊆m₂ , _) → m₁̷⊆m₂ m₁⊆m₂)
|
... | no m₁̷⊆m₂ | _ = no (λ (m₁⊆m₂ , _) → m₁̷⊆m₂ m₁⊆m₂)
|
||||||
|
|
||||||
|
≈-Decidable : IsDecidable _≈_
|
||||||
|
≈-Decidable = record { R-dec = ≈-dec }
|
||||||
|
|
||||||
private module I⊔ = ImplInsert _⊔₂_
|
private module I⊔ = ImplInsert _⊔₂_
|
||||||
private module I⊓ = ImplInsert _⊓₂_
|
private module I⊓ = ImplInsert _⊓₂_
|
||||||
|
|
||||||
@@ -1026,7 +1033,7 @@ updating-via-k∉ks-backward m = transform-k∉ks-backward
|
|||||||
|
|
||||||
module _ {l} {L : Set l}
|
module _ {l} {L : Set l}
|
||||||
{_≈ˡ_ : L → L → Set l} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
{_≈ˡ_ : L → L → Set l} {_⊔ˡ_ : L → L → L} {_⊓ˡ_ : L → L → L}
|
||||||
(lL : IsLattice L _≈ˡ_ _⊔ˡ_ _⊓ˡ_) where
|
{{lL : IsLattice L _≈ˡ_ _⊔ˡ_ _⊓ˡ_}} where
|
||||||
open IsLattice lL using () renaming (_≼_ to _≼ˡ_)
|
open IsLattice lL using () renaming (_≼_ to _≼ˡ_)
|
||||||
|
|
||||||
module _ (f : L → Map) (f-Monotonic : Monotonic _≼ˡ_ _≼_ f)
|
module _ (f : L → Map) (f-Monotonic : Monotonic _≼ˡ_ _≼_ f)
|
||||||
@@ -1112,6 +1119,19 @@ _[_] m (k ∷ ks)
|
|||||||
... | yes k∈km = proj₁ (locate {m = m} k∈km) ∷ (m [ ks ])
|
... | yes k∈km = proj₁ (locate {m = m} k∈km) ∷ (m [ ks ])
|
||||||
... | no _ = m [ ks ]
|
... | no _ = m [ ks ]
|
||||||
|
|
||||||
|
[]-∈ : ∀ {k : A} {v : B} {ks : List A} (m : Map) →
|
||||||
|
(k , v) ∈ m → k ∈ˡ ks → v ∈ˡ (m [ ks ])
|
||||||
|
[]-∈ {k} {v} {ks} m k,v∈m (here refl)
|
||||||
|
with ∈k-dec k (proj₁ m)
|
||||||
|
... | no k∉km = ⊥-elim (k∉km (forget k,v∈m))
|
||||||
|
... | yes k∈km
|
||||||
|
with (v' , k,v'∈m) ← locate {m = m} k∈km
|
||||||
|
rewrite Map-functional {m = m} k,v'∈m k,v∈m = here refl
|
||||||
|
[]-∈ {k} {v} {k' ∷ ks'} m k,v∈m (there k∈ks')
|
||||||
|
with ∈k-dec k' (proj₁ m)
|
||||||
|
... | no _ = []-∈ m k,v∈m k∈ks'
|
||||||
|
... | yes _ = there ([]-∈ m k,v∈m k∈ks')
|
||||||
|
|
||||||
m₁≼m₂⇒m₁[k]≼m₂[k] : ∀ (m₁ m₂ : Map) {k : A} {v₁ v₂ : B} →
|
m₁≼m₂⇒m₁[k]≼m₂[k] : ∀ (m₁ m₂ : Map) {k : A} {v₁ v₂ : B} →
|
||||||
m₁ ≼ m₂ → (k , v₁) ∈ m₁ → (k , v₂) ∈ m₂ → v₁ ≼₂ v₂
|
m₁ ≼ m₂ → (k , v₁) ∈ m₁ → (k , v₂) ∈ m₂ → v₁ ≼₂ v₂
|
||||||
m₁≼m₂⇒m₁[k]≼m₂[k] m₁ m₂ m₁≼m₂ k,v₁∈m₁ k,v₂∈m₂
|
m₁≼m₂⇒m₁[k]≼m₂[k] m₁ m₂ m₁≼m₂ k,v₁∈m₁ k,v₂∈m₂
|
||||||
@@ -1129,3 +1149,12 @@ m₁≼m₂⇒k∈km₁⇒k∈km₂ m₁ m₂ m₁≼m₂ k∈km₁ =
|
|||||||
(v' , (v≈v' , k,v'∈m₂)) = (proj₁ m₁≼m₂) _ _ k,v∈m₁m₂
|
(v' , (v≈v' , k,v'∈m₂)) = (proj₁ m₁≼m₂) _ _ k,v∈m₁m₂
|
||||||
in
|
in
|
||||||
forget k,v'∈m₂
|
forget k,v'∈m₂
|
||||||
|
|
||||||
|
m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ : ∀ (m₁ m₂ : Map) {k : A} →
|
||||||
|
m₁ ≈ m₂ → ∀ (k∈km₁ : k ∈k m₁) (k∈km₂ : k ∈k m₂) →
|
||||||
|
proj₁ (locate {m = m₁} k∈km₁) ≈₂ proj₁ (locate {m = m₂} k∈km₂)
|
||||||
|
m₁≈m₂⇒k∈m₁⇒k∈km₂⇒v₁≈v₂ m₁ m₂ {k} (m₁⊆m₂ , m₂⊆m₁) k∈km₁ k∈km₂
|
||||||
|
with (v₁ , k,v₁∈m₁) ← locate {m = m₁} k∈km₁
|
||||||
|
with (v₂ , k,v₂∈m₂) ← locate {m = m₂} k∈km₂
|
||||||
|
with (v₂' , (v₁≈v₂' , k,v₂'∈m₂)) ← m₁⊆m₂ k v₁ k,v₁∈m₁
|
||||||
|
rewrite Map-functional {m = m₂} k,v₂∈m₂ k,v₂'∈m₂ = v₁≈v₂'
|
||||||
|
|||||||
@@ -1,9 +1,9 @@
|
|||||||
open import Lattice
|
open import Lattice
|
||||||
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym; trans; cong; subst)
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym; trans; cong; subst)
|
||||||
open import Relation.Binary.Definitions using (Decidable)
|
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
|
open import Data.Unit using (⊤)
|
||||||
|
|
||||||
module Lattice.MapSet {a : Level} {A : Set a} (≡-dec-A : Decidable (_≡_ {a} {A})) where
|
module Lattice.MapSet {a : Level} (A : Set a) {{≡-Decidable-A : IsDecidable (_≡_ {a} {A})}} (dummy : ⊤) where
|
||||||
|
|
||||||
open import Data.List using (List; map)
|
open import Data.List using (List; map)
|
||||||
open import Data.Product using (_,_; proj₁)
|
open import Data.Product using (_,_; proj₁)
|
||||||
@@ -12,7 +12,7 @@ open import Function using (_∘_)
|
|||||||
open import Lattice.Unit using (⊤; tt) renaming (_≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_; isLattice to ⊤-isLattice)
|
open import Lattice.Unit using (⊤; tt) renaming (_≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_; isLattice to ⊤-isLattice)
|
||||||
import Lattice.Map
|
import Lattice.Map
|
||||||
|
|
||||||
private module UnitMap = Lattice.Map ≡-dec-A ⊤-isLattice
|
private module UnitMap = Lattice.Map A ⊤ dummy
|
||||||
open UnitMap
|
open UnitMap
|
||||||
using (Map; Expr; ⟦_⟧)
|
using (Map; Expr; ⟦_⟧)
|
||||||
renaming
|
renaming
|
||||||
|
|||||||
@@ -18,31 +18,32 @@ private
|
|||||||
≡-⊓-cong : ∀ {a₁ a₂ a₃ a₄} → a₁ ≡ a₂ → a₃ ≡ a₄ → (a₁ ⊓ a₃) ≡ (a₂ ⊓ a₄)
|
≡-⊓-cong : ∀ {a₁ a₂ a₃ a₄} → a₁ ≡ a₂ → a₃ ≡ a₄ → (a₁ ⊓ a₃) ≡ (a₂ ⊓ a₄)
|
||||||
≡-⊓-cong a₁≡a₂ a₃≡a₄ rewrite a₁≡a₂ rewrite a₃≡a₄ = refl
|
≡-⊓-cong a₁≡a₂ a₃≡a₄ rewrite a₁≡a₂ rewrite a₃≡a₄ = refl
|
||||||
|
|
||||||
isMaxSemilattice : IsSemilattice ℕ _≡_ _⊔_
|
instance
|
||||||
isMaxSemilattice = record
|
isMaxSemilattice : IsSemilattice ℕ _≡_ _⊔_
|
||||||
{ ≈-equiv = record
|
isMaxSemilattice = record
|
||||||
{ ≈-refl = refl
|
{ ≈-equiv = record
|
||||||
; ≈-sym = sym
|
{ ≈-refl = refl
|
||||||
; ≈-trans = trans
|
; ≈-sym = sym
|
||||||
|
; ≈-trans = trans
|
||||||
|
}
|
||||||
|
; ≈-⊔-cong = ≡-⊔-cong
|
||||||
|
; ⊔-assoc = ⊔-assoc
|
||||||
|
; ⊔-comm = ⊔-comm
|
||||||
|
; ⊔-idemp = ⊔-idem
|
||||||
}
|
}
|
||||||
; ≈-⊔-cong = ≡-⊔-cong
|
|
||||||
; ⊔-assoc = ⊔-assoc
|
|
||||||
; ⊔-comm = ⊔-comm
|
|
||||||
; ⊔-idemp = ⊔-idem
|
|
||||||
}
|
|
||||||
|
|
||||||
isMinSemilattice : IsSemilattice ℕ _≡_ _⊓_
|
isMinSemilattice : IsSemilattice ℕ _≡_ _⊓_
|
||||||
isMinSemilattice = record
|
isMinSemilattice = record
|
||||||
{ ≈-equiv = record
|
{ ≈-equiv = record
|
||||||
{ ≈-refl = refl
|
{ ≈-refl = refl
|
||||||
; ≈-sym = sym
|
; ≈-sym = sym
|
||||||
; ≈-trans = trans
|
; ≈-trans = trans
|
||||||
|
}
|
||||||
|
; ≈-⊔-cong = ≡-⊓-cong
|
||||||
|
; ⊔-assoc = ⊓-assoc
|
||||||
|
; ⊔-comm = ⊓-comm
|
||||||
|
; ⊔-idemp = ⊓-idem
|
||||||
}
|
}
|
||||||
; ≈-⊔-cong = ≡-⊓-cong
|
|
||||||
; ⊔-assoc = ⊓-assoc
|
|
||||||
; ⊔-comm = ⊓-comm
|
|
||||||
; ⊔-idemp = ⊓-idem
|
|
||||||
}
|
|
||||||
|
|
||||||
private
|
private
|
||||||
max-bound₁ : {x y z : ℕ} → x ⊔ y ≡ z → x ≤ z
|
max-bound₁ : {x y z : ℕ} → x ⊔ y ≡ z → x ≤ z
|
||||||
@@ -74,18 +75,19 @@ private
|
|||||||
helper : x ⊔ (x ⊓ y) ≤ x ⊔ x → x ⊔ x ≡ x → x ⊔ (x ⊓ y) ≤ x
|
helper : x ⊔ (x ⊓ y) ≤ x ⊔ x → x ⊔ x ≡ x → x ⊔ (x ⊓ y) ≤ x
|
||||||
helper x⊔x⊓y≤x⊔x x⊔x≡x rewrite x⊔x≡x = x⊔x⊓y≤x⊔x
|
helper x⊔x⊓y≤x⊔x x⊔x≡x rewrite x⊔x≡x = x⊔x⊓y≤x⊔x
|
||||||
|
|
||||||
isLattice : IsLattice ℕ _≡_ _⊔_ _⊓_
|
instance
|
||||||
isLattice = record
|
isLattice : IsLattice ℕ _≡_ _⊔_ _⊓_
|
||||||
{ joinSemilattice = isMaxSemilattice
|
isLattice = record
|
||||||
; meetSemilattice = isMinSemilattice
|
{ joinSemilattice = isMaxSemilattice
|
||||||
; absorb-⊔-⊓ = λ x y → maxmin-absorb {x} {y}
|
; meetSemilattice = isMinSemilattice
|
||||||
; absorb-⊓-⊔ = λ x y → minmax-absorb {x} {y}
|
; absorb-⊔-⊓ = λ x y → maxmin-absorb {x} {y}
|
||||||
}
|
; absorb-⊓-⊔ = λ x y → minmax-absorb {x} {y}
|
||||||
|
}
|
||||||
|
|
||||||
lattice : Lattice ℕ
|
lattice : Lattice ℕ
|
||||||
lattice = record
|
lattice = record
|
||||||
{ _≈_ = _≡_
|
{ _≈_ = _≡_
|
||||||
; _⊔_ = _⊔_
|
; _⊔_ = _⊔_
|
||||||
; _⊓_ = _⊓_
|
; _⊓_ = _⊓_
|
||||||
; isLattice = isLattice
|
; isLattice = isLattice
|
||||||
}
|
}
|
||||||
|
|||||||
@@ -1,10 +1,10 @@
|
|||||||
open import Lattice
|
open import Lattice
|
||||||
|
|
||||||
module Lattice.Prod {a b} {A : Set a} {B : Set b}
|
module Lattice.Prod {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}
|
||||||
(_⊓₁_ : A → A → A) (_⊓₂_ : B → B → B)
|
{_⊓₁_ : A → A → A} {_⊓₂_ : B → B → B}
|
||||||
(lA : IsLattice A _≈₁_ _⊔₁_ _⊓₁_) (lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
{{lA : IsLattice A _≈₁_ _⊔₁_ _⊓₁_}} {{lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_}} where
|
||||||
|
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
open import Data.Nat using (ℕ; _≤_; _+_; suc)
|
open import Data.Nat using (ℕ; _≤_; _+_; suc)
|
||||||
@@ -12,6 +12,7 @@ open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂)
|
|||||||
open import Data.Empty using (⊥-elim)
|
open import Data.Empty using (⊥-elim)
|
||||||
open import Relation.Binary.Core using (_Preserves_⟶_ )
|
open import Relation.Binary.Core using (_Preserves_⟶_ )
|
||||||
open import Relation.Binary.PropositionalEquality using (sym; subst)
|
open import Relation.Binary.PropositionalEquality using (sym; subst)
|
||||||
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
open import Relation.Nullary using (¬_; yes; no)
|
open import Relation.Nullary using (¬_; yes; no)
|
||||||
open import Equivalence
|
open import Equivalence
|
||||||
import Chain
|
import Chain
|
||||||
@@ -39,13 +40,14 @@ open IsLattice lB using () renaming
|
|||||||
_≈_ : A × B → A × B → Set (a ⊔ℓ b)
|
_≈_ : A × B → A × B → Set (a ⊔ℓ b)
|
||||||
(a₁ , b₁) ≈ (a₂ , b₂) = (a₁ ≈₁ a₂) × (b₁ ≈₂ b₂)
|
(a₁ , b₁) ≈ (a₂ , b₂) = (a₁ ≈₁ a₂) × (b₁ ≈₂ b₂)
|
||||||
|
|
||||||
≈-equiv : IsEquivalence (A × B) _≈_
|
instance
|
||||||
≈-equiv = record
|
≈-equiv : IsEquivalence (A × B) _≈_
|
||||||
{ ≈-refl = λ {p} → (≈₁-refl , ≈₂-refl)
|
≈-equiv = record
|
||||||
; ≈-sym = λ {p₁} {p₂} (a₁≈a₂ , b₁≈b₂) → (≈₁-sym a₁≈a₂ , ≈₂-sym b₁≈b₂)
|
{ ≈-refl = λ {p} → (≈₁-refl , ≈₂-refl)
|
||||||
; ≈-trans = λ {p₁} {p₂} {p₃} (a₁≈a₂ , b₁≈b₂) (a₂≈a₃ , b₂≈b₃) →
|
; ≈-sym = λ {p₁} {p₂} (a₁≈a₂ , b₁≈b₂) → (≈₁-sym a₁≈a₂ , ≈₂-sym b₁≈b₂)
|
||||||
( ≈₁-trans a₁≈a₂ a₂≈a₃ , ≈₂-trans b₁≈b₂ b₂≈b₃ )
|
; ≈-trans = λ {p₁} {p₂} {p₃} (a₁≈a₂ , b₁≈b₂) (a₂≈a₃ , b₂≈b₃) →
|
||||||
}
|
( ≈₁-trans a₁≈a₂ a₂≈a₃ , ≈₂-trans b₁≈b₂ b₂≈b₃ )
|
||||||
|
}
|
||||||
|
|
||||||
_⊔_ : A × B → A × B → A × B
|
_⊔_ : A × B → A × B → A × B
|
||||||
(a₁ , b₁) ⊔ (a₂ , b₂) = (a₁ ⊔₁ a₂ , b₁ ⊔₂ b₂)
|
(a₁ , b₁) ⊔ (a₂ , b₂) = (a₁ ⊔₁ a₂ , b₁ ⊔₂ b₂)
|
||||||
@@ -75,124 +77,124 @@ private module ProdIsSemilattice (f₁ : A → A → A) (f₂ : B → B → B) (
|
|||||||
)
|
)
|
||||||
}
|
}
|
||||||
|
|
||||||
isJoinSemilattice : IsSemilattice (A × B) _≈_ _⊔_
|
instance
|
||||||
isJoinSemilattice = ProdIsSemilattice.isSemilattice _⊔₁_ _⊔₂_ joinSemilattice₁ joinSemilattice₂
|
isJoinSemilattice : IsSemilattice (A × B) _≈_ _⊔_
|
||||||
|
isJoinSemilattice = ProdIsSemilattice.isSemilattice _⊔₁_ _⊔₂_ joinSemilattice₁ joinSemilattice₂
|
||||||
|
|
||||||
isMeetSemilattice : IsSemilattice (A × B) _≈_ _⊓_
|
isMeetSemilattice : IsSemilattice (A × B) _≈_ _⊓_
|
||||||
isMeetSemilattice = ProdIsSemilattice.isSemilattice _⊓₁_ _⊓₂_ meetSemilattice₁ meetSemilattice₂
|
isMeetSemilattice = ProdIsSemilattice.isSemilattice _⊓₁_ _⊓₂_ meetSemilattice₁ meetSemilattice₂
|
||||||
|
|
||||||
isLattice : IsLattice (A × B) _≈_ _⊔_ _⊓_
|
isLattice : IsLattice (A × B) _≈_ _⊔_ _⊓_
|
||||||
isLattice = record
|
isLattice = record
|
||||||
{ joinSemilattice = isJoinSemilattice
|
{ joinSemilattice = isJoinSemilattice
|
||||||
; meetSemilattice = isMeetSemilattice
|
; meetSemilattice = isMeetSemilattice
|
||||||
; absorb-⊔-⊓ = λ (a₁ , b₁) (a₂ , b₂) →
|
; absorb-⊔-⊓ = λ (a₁ , b₁) (a₂ , b₂) →
|
||||||
( IsLattice.absorb-⊔-⊓ lA a₁ a₂
|
( IsLattice.absorb-⊔-⊓ lA a₁ a₂
|
||||||
, IsLattice.absorb-⊔-⊓ lB b₁ b₂
|
, IsLattice.absorb-⊔-⊓ lB b₁ b₂
|
||||||
)
|
)
|
||||||
; absorb-⊓-⊔ = λ (a₁ , b₁) (a₂ , b₂) →
|
; absorb-⊓-⊔ = λ (a₁ , b₁) (a₂ , b₂) →
|
||||||
( IsLattice.absorb-⊓-⊔ lA a₁ a₂
|
( IsLattice.absorb-⊓-⊔ lA a₁ a₂
|
||||||
, IsLattice.absorb-⊓-⊔ lB b₁ b₂
|
, IsLattice.absorb-⊓-⊔ lB b₁ b₂
|
||||||
)
|
)
|
||||||
}
|
}
|
||||||
|
|
||||||
lattice : Lattice (A × B)
|
lattice : Lattice (A × B)
|
||||||
lattice = record
|
lattice = record
|
||||||
{ _≈_ = _≈_
|
{ _≈_ = _≈_
|
||||||
; _⊔_ = _⊔_
|
; _⊔_ = _⊔_
|
||||||
; _⊓_ = _⊓_
|
; _⊓_ = _⊓_
|
||||||
; isLattice = isLattice
|
; isLattice = isLattice
|
||||||
}
|
}
|
||||||
|
|
||||||
module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_) where
|
open IsLattice isLattice using (_≼_; _≺_; ≺-cong) public
|
||||||
≈-dec : IsDecidable _≈_
|
|
||||||
|
module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDecidable _≈₂_}} where
|
||||||
|
open IsDecidable ≈₁-Decidable using () renaming (R-dec to ≈₁-dec)
|
||||||
|
open IsDecidable ≈₂-Decidable using () renaming (R-dec to ≈₂-dec)
|
||||||
|
|
||||||
|
≈-dec : Decidable _≈_
|
||||||
≈-dec (a₁ , b₁) (a₂ , b₂)
|
≈-dec (a₁ , b₁) (a₂ , b₂)
|
||||||
with ≈₁-dec a₁ a₂ | ≈₂-dec b₁ b₂
|
with ≈₁-dec a₁ a₂ | ≈₂-dec b₁ b₂
|
||||||
... | yes a₁≈a₂ | yes b₁≈b₂ = yes (a₁≈a₂ , b₁≈b₂)
|
... | yes a₁≈a₂ | yes b₁≈b₂ = yes (a₁≈a₂ , b₁≈b₂)
|
||||||
... | no a₁̷≈a₂ | _ = no (λ (a₁≈a₂ , _) → a₁̷≈a₂ a₁≈a₂)
|
... | no a₁̷≈a₂ | _ = no (λ (a₁≈a₂ , _) → a₁̷≈a₂ a₁≈a₂)
|
||||||
... | _ | no b₁̷≈b₂ = no (λ (_ , b₁≈b₂) → b₁̷≈b₂ b₁≈b₂)
|
... | _ | no b₁̷≈b₂ = no (λ (_ , b₁≈b₂) → b₁̷≈b₂ b₁≈b₂)
|
||||||
|
|
||||||
|
instance
|
||||||
|
≈-Decidable : IsDecidable _≈_
|
||||||
|
≈-Decidable = record { R-dec = ≈-dec }
|
||||||
|
|
||||||
module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_)
|
module _ {h₁ h₂ : ℕ}
|
||||||
(h₁ h₂ : ℕ)
|
{{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
|
||||||
|
|
||||||
module ChainA = Chain _≈₁_ ≈₁-equiv _≺₁_ ≺₁-cong
|
module ChainA = Chain _≈₁_ ≈₁-equiv _≺₁_ ≺₁-cong
|
||||||
module ChainB = Chain _≈₂_ ≈₂-equiv _≺₂_ ≺₂-cong
|
module ChainB = Chain _≈₂_ ≈₂-equiv _≺₂_ ≺₂-cong
|
||||||
module ProdChain = Chain _≈_ ≈-equiv _≺_ ≺-cong
|
module ProdChain = Chain _≈_ ≈-equiv _≺_ ≺-cong
|
||||||
|
|
||||||
open ChainA using () renaming (Chain to Chain₁; done to done₁; step to step₁; Chain-≈-cong₁ to Chain₁-≈-cong₁)
|
open ChainA using () renaming (Chain to Chain₁; done to done₁; step to step₁; Chain-≈-cong₁ to Chain₁-≈-cong₁)
|
||||||
open ChainB using () renaming (Chain to Chain₂; done to done₂; step to step₂; Chain-≈-cong₁ to Chain₂-≈-cong₁)
|
open ChainB using () renaming (Chain to Chain₂; done to done₂; step to step₂; Chain-≈-cong₁ to Chain₂-≈-cong₁)
|
||||||
open ProdChain using (Chain; concat; done; step)
|
open ProdChain using (Chain; concat; done; step)
|
||||||
|
|
||||||
private
|
private
|
||||||
a,∙-Monotonic : ∀ (a : A) → Monotonic _≼₂_ _≼_ (λ b → (a , b))
|
a,∙-Monotonic : ∀ (a : A) → Monotonic _≼₂_ _≼_ (λ b → (a , b))
|
||||||
a,∙-Monotonic a {b₁} {b₂} b₁⊔b₂≈b₂ = (⊔₁-idemp a , b₁⊔b₂≈b₂)
|
a,∙-Monotonic a {b₁} {b₂} b₁⊔b₂≈b₂ = (⊔₁-idemp a , b₁⊔b₂≈b₂)
|
||||||
|
|
||||||
a,∙-Preserves-≈₂ : ∀ (a : A) → (λ b → (a , b)) Preserves _≈₂_ ⟶ _≈_
|
a,∙-Preserves-≈₂ : ∀ (a : A) → (λ b → (a , b)) Preserves _≈₂_ ⟶ _≈_
|
||||||
a,∙-Preserves-≈₂ a {b₁} {b₂} b₁≈b₂ = (≈₁-refl , b₁≈b₂)
|
a,∙-Preserves-≈₂ a {b₁} {b₂} b₁≈b₂ = (≈₁-refl , b₁≈b₂)
|
||||||
|
|
||||||
∙,b-Monotonic : ∀ (b : B) → Monotonic _≼₁_ _≼_ (λ a → (a , b))
|
∙,b-Monotonic : ∀ (b : B) → Monotonic _≼₁_ _≼_ (λ a → (a , b))
|
||||||
∙,b-Monotonic b {a₁} {a₂} a₁⊔a₂≈a₂ = (a₁⊔a₂≈a₂ , ⊔₂-idemp b)
|
∙,b-Monotonic b {a₁} {a₂} a₁⊔a₂≈a₂ = (a₁⊔a₂≈a₂ , ⊔₂-idemp b)
|
||||||
|
|
||||||
∙,b-Preserves-≈₁ : ∀ (b : B) → (λ a → (a , b)) Preserves _≈₁_ ⟶ _≈_
|
∙,b-Preserves-≈₁ : ∀ (b : B) → (λ a → (a , b)) Preserves _≈₁_ ⟶ _≈_
|
||||||
∙,b-Preserves-≈₁ b {a₁} {a₂} a₁≈a₂ = (a₁≈a₂ , ≈₂-refl)
|
∙,b-Preserves-≈₁ b {a₁} {a₂} a₁≈a₂ = (a₁≈a₂ , ≈₂-refl)
|
||||||
|
|
||||||
amin : A
|
open ChainA.Height fhA using () renaming (⊥ to ⊥₁; ⊤ to ⊤₁; longestChain to longestChain₁; bounded to bounded₁)
|
||||||
amin = proj₁ (proj₁ (proj₁ fhA))
|
open ChainB.Height fhB using () renaming (⊥ to ⊥₂; ⊤ to ⊤₂; longestChain to longestChain₂; bounded to bounded₂)
|
||||||
|
|
||||||
amax : A
|
unzip : ∀ {a₁ a₂ : A} {b₁ b₂ : B} {n : ℕ} → Chain (a₁ , b₁) (a₂ , b₂) n → Σ (ℕ × ℕ) (λ (n₁ , n₂) → ((Chain₁ a₁ a₂ n₁ × Chain₂ b₁ b₂ n₂) × (n ≤ n₁ + n₂)))
|
||||||
amax = proj₂ (proj₁ (proj₁ fhA))
|
unzip (done (a₁≈a₂ , b₁≈b₂)) = ((0 , 0) , ((done₁ a₁≈a₂ , done₂ b₁≈b₂) , ≤-refl))
|
||||||
|
unzip {a₁} {a₂} {b₁} {b₂} {n} (step {(a₁ , b₁)} {(a , b)} ((a₁≼a , b₁≼b) , a₁b₁̷≈ab) (a≈a' , b≈b') a'b'a₂b₂)
|
||||||
bmin : B
|
with ≈₁-dec a₁ a | ≈₂-dec b₁ b | unzip a'b'a₂b₂
|
||||||
bmin = proj₁ (proj₁ (proj₁ fhB))
|
... | yes a₁≈a | yes b₁≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) = ⊥-elim (a₁b₁̷≈ab (a₁≈a , b₁≈b))
|
||||||
|
... | no a₁̷≈a | yes b₁≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) =
|
||||||
bmax : B
|
((suc n₁ , n₂) , ((step₁ (a₁≼a , a₁̷≈a) a≈a' c₁ , Chain₂-≈-cong₁ (≈₂-sym (≈₂-trans b₁≈b b≈b')) c₂), +-monoʳ-≤ 1 (n≤n₁+n₂)))
|
||||||
bmax = proj₂ (proj₁ (proj₁ fhB))
|
... | yes a₁≈a | no b₁̷≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) =
|
||||||
|
((n₁ , suc n₂) , ( (Chain₁-≈-cong₁ (≈₁-sym (≈₁-trans a₁≈a a≈a')) c₁ , step₂ (b₁≼b , b₁̷≈b) b≈b' c₂)
|
||||||
unzip : ∀ {a₁ a₂ : A} {b₁ b₂ : B} {n : ℕ} → Chain (a₁ , b₁) (a₂ , b₂) n → Σ (ℕ × ℕ) (λ (n₁ , n₂) → ((Chain₁ a₁ a₂ n₁ × Chain₂ b₁ b₂ n₂) × (n ≤ n₁ + n₂)))
|
, subst (n ≤_) (sym (+-suc n₁ n₂)) (+-monoʳ-≤ 1 n≤n₁+n₂)
|
||||||
unzip (done (a₁≈a₂ , b₁≈b₂)) = ((0 , 0) , ((done₁ a₁≈a₂ , done₂ b₁≈b₂) , ≤-refl))
|
|
||||||
unzip {a₁} {a₂} {b₁} {b₂} {n} (step {(a₁ , b₁)} {(a , b)} ((a₁≼a , b₁≼b) , a₁b₁̷≈ab) (a≈a' , b≈b') a'b'a₂b₂)
|
|
||||||
with ≈₁-dec a₁ a | ≈₂-dec b₁ b | unzip a'b'a₂b₂
|
|
||||||
... | yes a₁≈a | yes b₁≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) = ⊥-elim (a₁b₁̷≈ab (a₁≈a , b₁≈b))
|
|
||||||
... | no a₁̷≈a | yes b₁≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) =
|
|
||||||
((suc n₁ , n₂) , ((step₁ (a₁≼a , a₁̷≈a) a≈a' c₁ , Chain₂-≈-cong₁ (≈₂-sym (≈₂-trans b₁≈b b≈b')) c₂), +-monoʳ-≤ 1 (n≤n₁+n₂)))
|
|
||||||
... | yes a₁≈a | no b₁̷≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) =
|
|
||||||
((n₁ , suc n₂) , ( (Chain₁-≈-cong₁ (≈₁-sym (≈₁-trans a₁≈a a≈a')) c₁ , step₂ (b₁≼b , b₁̷≈b) b≈b' c₂)
|
|
||||||
, subst (n ≤_) (sym (+-suc n₁ n₂)) (+-monoʳ-≤ 1 n≤n₁+n₂)
|
|
||||||
))
|
|
||||||
... | no a₁̷≈a | no b₁̷≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) =
|
|
||||||
((suc n₁ , suc n₂) , ( (step₁ (a₁≼a , a₁̷≈a) a≈a' c₁ , step₂ (b₁≼b , b₁̷≈b) b≈b' c₂)
|
|
||||||
, m≤n⇒m≤o+n 1 (subst (n ≤_) (sym (+-suc n₁ n₂)) (+-monoʳ-≤ 1 n≤n₁+n₂))
|
|
||||||
))
|
))
|
||||||
|
... | no a₁̷≈a | no b₁̷≈b | ((n₁ , n₂) , ((c₁ , c₂) , n≤n₁+n₂)) =
|
||||||
|
((suc n₁ , suc n₂) , ( (step₁ (a₁≼a , a₁̷≈a) a≈a' c₁ , step₂ (b₁≼b , b₁̷≈b) b≈b' c₂)
|
||||||
|
, m≤n⇒m≤o+n 1 (subst (n ≤_) (sym (+-suc n₁ n₂)) (+-monoʳ-≤ 1 n≤n₁+n₂))
|
||||||
|
))
|
||||||
|
|
||||||
fixedHeight : IsLattice.FixedHeight isLattice (h₁ + h₂)
|
instance
|
||||||
fixedHeight =
|
fixedHeight : IsLattice.FixedHeight isLattice (h₁ + h₂)
|
||||||
( ( ((amin , bmin) , (amax , bmax))
|
fixedHeight = record
|
||||||
, concat
|
{ ⊥ = (⊥₁ , ⊥₂)
|
||||||
(ChainMapping₁.Chain-map (λ a → (a , bmin)) (∙,b-Monotonic _) proj₁ (∙,b-Preserves-≈₁ _) (proj₂ (proj₁ fhA)))
|
; ⊤ = (⊤₁ , ⊤₂)
|
||||||
(ChainMapping₂.Chain-map (λ b → (amax , b)) (a,∙-Monotonic _) proj₂ (a,∙-Preserves-≈₂ _) (proj₂ (proj₁ fhB)))
|
; longestChain = concat
|
||||||
)
|
(ChainMapping₁.Chain-map (λ a → (a , ⊥₂)) (∙,b-Monotonic _) proj₁ (∙,b-Preserves-≈₁ _) longestChain₁)
|
||||||
, λ a₁b₁a₂b₂ → let ((n₁ , n₂) , ((a₁a₂ , b₁b₂) , n≤n₁+n₂)) = unzip a₁b₁a₂b₂
|
(ChainMapping₂.Chain-map (λ b → (⊤₁ , b)) (a,∙-Monotonic _) proj₂ (a,∙-Preserves-≈₂ _) longestChain₂)
|
||||||
in ≤-trans n≤n₁+n₂ (+-mono-≤ (proj₂ fhA a₁a₂) (proj₂ fhB b₁b₂))
|
; bounded = λ a₁b₁a₂b₂ →
|
||||||
)
|
let ((n₁ , n₂) , ((a₁a₂ , b₁b₂) , n≤n₁+n₂)) = unzip a₁b₁a₂b₂
|
||||||
|
in ≤-trans n≤n₁+n₂ (+-mono-≤ (bounded₁ a₁a₂) (bounded₂ b₁b₂))
|
||||||
|
}
|
||||||
|
|
||||||
isFiniteHeightLattice : IsFiniteHeightLattice (A × B) (h₁ + h₂) _≈_ _⊔_ _⊓_
|
isFiniteHeightLattice : IsFiniteHeightLattice (A × B) (h₁ + h₂) _≈_ _⊔_ _⊓_
|
||||||
isFiniteHeightLattice = record
|
isFiniteHeightLattice = record
|
||||||
{ isLattice = isLattice
|
{ isLattice = isLattice
|
||||||
; fixedHeight = fixedHeight
|
; fixedHeight = fixedHeight
|
||||||
}
|
}
|
||||||
|
|
||||||
finiteHeightLattice : FiniteHeightLattice (A × B)
|
finiteHeightLattice : FiniteHeightLattice (A × B)
|
||||||
finiteHeightLattice = record
|
finiteHeightLattice = record
|
||||||
{ height = h₁ + h₂
|
{ height = h₁ + h₂
|
||||||
; _≈_ = _≈_
|
; _≈_ = _≈_
|
||||||
; _⊔_ = _⊔_
|
; _⊔_ = _⊔_
|
||||||
; _⊓_ = _⊓_
|
; _⊓_ = _⊓_
|
||||||
; isFiniteHeightLattice = isFiniteHeightLattice
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
||||||
}
|
}
|
||||||
|
|||||||
@@ -7,6 +7,7 @@ open import Data.Unit using (⊤; tt) public
|
|||||||
open import Data.Unit.Properties using (_≟_; ≡-setoid)
|
open import Data.Unit.Properties using (_≟_; ≡-setoid)
|
||||||
open import Relation.Binary using (Setoid)
|
open import Relation.Binary using (Setoid)
|
||||||
open import Relation.Binary.PropositionalEquality as Eq using (_≡_)
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_)
|
||||||
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
||||||
open import Equivalence
|
open import Equivalence
|
||||||
open import Lattice
|
open import Lattice
|
||||||
@@ -24,9 +25,13 @@ _≈_ = _≡_
|
|||||||
; ≈-trans = trans
|
; ≈-trans = trans
|
||||||
}
|
}
|
||||||
|
|
||||||
≈-dec : IsDecidable _≈_
|
≈-dec : Decidable _≈_
|
||||||
≈-dec = _≟_
|
≈-dec = _≟_
|
||||||
|
|
||||||
|
instance
|
||||||
|
≈-Decidable : IsDecidable _≈_
|
||||||
|
≈-Decidable = record { R-dec = ≈-dec }
|
||||||
|
|
||||||
_⊔_ : ⊤ → ⊤ → ⊤
|
_⊔_ : ⊤ → ⊤ → ⊤
|
||||||
tt ⊔ tt = tt
|
tt ⊔ tt = tt
|
||||||
|
|
||||||
@@ -45,14 +50,15 @@ tt ⊓ tt = tt
|
|||||||
⊔-idemp : (x : ⊤) → (x ⊔ x) ≈ x
|
⊔-idemp : (x : ⊤) → (x ⊔ x) ≈ x
|
||||||
⊔-idemp tt = Eq.refl
|
⊔-idemp tt = Eq.refl
|
||||||
|
|
||||||
isJoinSemilattice : IsSemilattice ⊤ _≈_ _⊔_
|
instance
|
||||||
isJoinSemilattice = record
|
isJoinSemilattice : IsSemilattice ⊤ _≈_ _⊔_
|
||||||
{ ≈-equiv = ≈-equiv
|
isJoinSemilattice = record
|
||||||
; ≈-⊔-cong = ≈-⊔-cong
|
{ ≈-equiv = ≈-equiv
|
||||||
; ⊔-assoc = ⊔-assoc
|
; ≈-⊔-cong = ≈-⊔-cong
|
||||||
; ⊔-comm = ⊔-comm
|
; ⊔-assoc = ⊔-assoc
|
||||||
; ⊔-idemp = ⊔-idemp
|
; ⊔-comm = ⊔-comm
|
||||||
}
|
; ⊔-idemp = ⊔-idemp
|
||||||
|
}
|
||||||
|
|
||||||
≈-⊓-cong : ∀ {ab₁ ab₂ ab₃ ab₄} → ab₁ ≈ ab₂ → ab₃ ≈ ab₄ → (ab₁ ⊓ ab₃) ≈ (ab₂ ⊓ ab₄)
|
≈-⊓-cong : ∀ {ab₁ ab₂ ab₃ ab₄} → ab₁ ≈ ab₂ → ab₃ ≈ ab₄ → (ab₁ ⊓ ab₃) ≈ (ab₂ ⊓ ab₄)
|
||||||
≈-⊓-cong {tt} {tt} {tt} {tt} _ _ = Eq.refl
|
≈-⊓-cong {tt} {tt} {tt} {tt} _ _ = Eq.refl
|
||||||
@@ -66,36 +72,32 @@ isJoinSemilattice = record
|
|||||||
⊓-idemp : (x : ⊤) → (x ⊓ x) ≈ x
|
⊓-idemp : (x : ⊤) → (x ⊓ x) ≈ x
|
||||||
⊓-idemp tt = Eq.refl
|
⊓-idemp tt = Eq.refl
|
||||||
|
|
||||||
isMeetSemilattice : IsSemilattice ⊤ _≈_ _⊓_
|
instance
|
||||||
isMeetSemilattice = record
|
isMeetSemilattice : IsSemilattice ⊤ _≈_ _⊓_
|
||||||
{ ≈-equiv = ≈-equiv
|
isMeetSemilattice = record
|
||||||
; ≈-⊔-cong = ≈-⊓-cong
|
{ ≈-equiv = ≈-equiv
|
||||||
; ⊔-assoc = ⊓-assoc
|
; ≈-⊔-cong = ≈-⊓-cong
|
||||||
; ⊔-comm = ⊓-comm
|
; ⊔-assoc = ⊓-assoc
|
||||||
; ⊔-idemp = ⊓-idemp
|
; ⊔-comm = ⊓-comm
|
||||||
}
|
; ⊔-idemp = ⊓-idemp
|
||||||
|
}
|
||||||
|
|
||||||
absorb-⊔-⊓ : (x y : ⊤) → (x ⊔ (x ⊓ y)) ≈ x
|
instance
|
||||||
absorb-⊔-⊓ tt tt = Eq.refl
|
isLattice : IsLattice ⊤ _≈_ _⊔_ _⊓_
|
||||||
|
isLattice = record
|
||||||
|
{ joinSemilattice = isJoinSemilattice
|
||||||
|
; meetSemilattice = isMeetSemilattice
|
||||||
|
; absorb-⊔-⊓ = λ { tt tt → Eq.refl }
|
||||||
|
; absorb-⊓-⊔ = λ { tt tt → Eq.refl }
|
||||||
|
}
|
||||||
|
|
||||||
absorb-⊓-⊔ : (x y : ⊤) → (x ⊓ (x ⊔ y)) ≈ x
|
lattice : Lattice ⊤
|
||||||
absorb-⊓-⊔ tt tt = Eq.refl
|
lattice = record
|
||||||
|
{ _≈_ = _≈_
|
||||||
isLattice : IsLattice ⊤ _≈_ _⊔_ _⊓_
|
; _⊔_ = _⊔_
|
||||||
isLattice = record
|
; _⊓_ = _⊓_
|
||||||
{ joinSemilattice = isJoinSemilattice
|
; isLattice = isLattice
|
||||||
; meetSemilattice = isMeetSemilattice
|
}
|
||||||
; absorb-⊔-⊓ = absorb-⊔-⊓
|
|
||||||
; absorb-⊓-⊔ = absorb-⊓-⊔
|
|
||||||
}
|
|
||||||
|
|
||||||
lattice : Lattice ⊤
|
|
||||||
lattice = record
|
|
||||||
{ _≈_ = _≈_
|
|
||||||
; _⊔_ = _⊔_
|
|
||||||
; _⊓_ = _⊓_
|
|
||||||
; isLattice = isLattice
|
|
||||||
}
|
|
||||||
|
|
||||||
open Chain _≈_ ≈-equiv (IsLattice._≺_ isLattice) (IsLattice.≺-cong isLattice)
|
open Chain _≈_ ≈-equiv (IsLattice._≺_ isLattice) (IsLattice.≺-cong isLattice)
|
||||||
|
|
||||||
@@ -107,20 +109,26 @@ private
|
|||||||
isLongest {tt} {tt} (step (tt⊔tt≈tt , tt̷≈tt) _ _) = ⊥-elim (tt̷≈tt refl)
|
isLongest {tt} {tt} (step (tt⊔tt≈tt , tt̷≈tt) _ _) = ⊥-elim (tt̷≈tt refl)
|
||||||
isLongest (done _) = z≤n
|
isLongest (done _) = z≤n
|
||||||
|
|
||||||
fixedHeight : IsLattice.FixedHeight isLattice 0
|
instance
|
||||||
fixedHeight = (((tt , tt) , longestChain) , isLongest)
|
fixedHeight : IsLattice.FixedHeight isLattice 0
|
||||||
|
fixedHeight = record
|
||||||
|
{ ⊥ = tt
|
||||||
|
; ⊤ = tt
|
||||||
|
; longestChain = longestChain
|
||||||
|
; bounded = isLongest
|
||||||
|
}
|
||||||
|
|
||||||
isFiniteHeightLattice : IsFiniteHeightLattice ⊤ 0 _≈_ _⊔_ _⊓_
|
isFiniteHeightLattice : IsFiniteHeightLattice ⊤ 0 _≈_ _⊔_ _⊓_
|
||||||
isFiniteHeightLattice = record
|
isFiniteHeightLattice = record
|
||||||
{ isLattice = isLattice
|
{ isLattice = isLattice
|
||||||
; fixedHeight = fixedHeight
|
; fixedHeight = fixedHeight
|
||||||
}
|
}
|
||||||
|
|
||||||
finiteHeightLattice : FiniteHeightLattice ⊤
|
finiteHeightLattice : FiniteHeightLattice ⊤
|
||||||
finiteHeightLattice = record
|
finiteHeightLattice = record
|
||||||
{ height = 0
|
{ height = 0
|
||||||
; _≈_ = _≈_
|
; _≈_ = _≈_
|
||||||
; _⊔_ = _⊔_
|
; _⊔_ = _⊔_
|
||||||
; _⊓_ = _⊓_
|
; _⊓_ = _⊓_
|
||||||
; isFiniteHeightLattice = isFiniteHeightLattice
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
||||||
}
|
}
|
||||||
|
|||||||
45
Main.agda
45
Main.agda
@@ -1,25 +1,46 @@
|
|||||||
|
{-# 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 : Vec Stmt _
|
testCode : Stmt
|
||||||
testCode =
|
testCode =
|
||||||
("zero" ← (# 0)) ∷
|
⟨ "zero" ← (# 0) ⟩ then
|
||||||
("pos" ← ((` "zero") Expr.+ (# 1))) ∷
|
⟨ "pos" ← ((` "zero") Expr.+ (# 1)) ⟩ then
|
||||||
("neg" ← ((` "zero") Expr.- (# 1))) ∷
|
⟨ "neg" ← ((` "zero") Expr.- (# 1)) ⟩ then
|
||||||
("unknown" ← ((` "pos") Expr.+ (` "neg"))) ∷
|
⟨ "unknown" ← ((` "pos") Expr.+ (` "neg")) ⟩
|
||||||
[]
|
|
||||||
|
testCodeCond₁ : Stmt
|
||||||
|
testCodeCond₁ =
|
||||||
|
⟨ "var" ← (# 1) ⟩ then
|
||||||
|
if (` "var") then (
|
||||||
|
⟨ "var" ← ((` "var") Expr.+ (# 1)) ⟩
|
||||||
|
) else (
|
||||||
|
⟨ "var" ← ((` "var") Expr.- (# 1)) ⟩ then
|
||||||
|
⟨ "var" ← (# 1) ⟩
|
||||||
|
)
|
||||||
|
|
||||||
|
testCodeCond₂ : Stmt
|
||||||
|
testCodeCond₂ =
|
||||||
|
⟨ "var" ← (# 1) ⟩ then
|
||||||
|
if (` "var") then (
|
||||||
|
⟨ "x" ← (# 1) ⟩
|
||||||
|
) else (
|
||||||
|
⟨ noop ⟩
|
||||||
|
)
|
||||||
|
|
||||||
testProgram : Program
|
testProgram : Program
|
||||||
testProgram = record
|
testProgram = record
|
||||||
{ length = _
|
{ rootStmt = testCode
|
||||||
; stmts = testCode
|
|
||||||
}
|
}
|
||||||
|
|
||||||
open WithProg testProgram using (output)
|
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))
|
||||||
|
|||||||
44
Utils.agda
44
Utils.agda
@@ -1,14 +1,18 @@
|
|||||||
module Utils where
|
module Utils where
|
||||||
|
|
||||||
open import Agda.Primitive using () renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using () renaming (_⊔_ to _⊔ℓ_)
|
||||||
|
open import Data.Product as Prod using (_×_)
|
||||||
open import Data.Nat using (ℕ; suc)
|
open import Data.Nat using (ℕ; suc)
|
||||||
open import Data.List using (List; []; _∷_; _++_) renaming (map to mapˡ)
|
open import Data.List using (List; cartesianProduct; []; _∷_; _++_; foldr; filter) renaming (map to mapˡ)
|
||||||
open import Data.List.Membership.Propositional using (_∈_)
|
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)
|
open import Data.List.Relation.Unary.All using (All; []; _∷_; map)
|
||||||
open import Data.List.Relation.Unary.Any using (Any; here; there) -- TODO: re-export these with nicer names from map
|
open import Data.List.Relation.Unary.Any using (Any; here; there) -- TODO: re-export these with nicer names from map
|
||||||
|
open import Data.Sum using (_⊎_)
|
||||||
open import Function.Definitions using (Injective)
|
open import Function.Definitions using (Injective)
|
||||||
open import Relation.Binary.PropositionalEquality using (_≡_; sym; refl)
|
open import Relation.Binary.PropositionalEquality using (_≡_; sym; refl; cong)
|
||||||
open import Relation.Nullary using (¬_)
|
open import Relation.Nullary using (¬_; yes; no)
|
||||||
|
open import Relation.Unary using (Decidable)
|
||||||
|
|
||||||
data Unique {c} {C : Set c} : List C → Set c where
|
data Unique {c} {C : Set c} : List C → Set c where
|
||||||
empty : Unique []
|
empty : Unique []
|
||||||
@@ -68,3 +72,37 @@ data Pairwise {a} {b} {c} {A : Set a} {B : Set b} (P : A → B → Set c) : List
|
|||||||
_∷_ : ∀ {x : A} {y : B} {xs : List A} {ys : List B} →
|
_∷_ : ∀ {x : A} {y : B} {xs : List A} {ys : List B} →
|
||||||
P x y → Pairwise P xs ys →
|
P x y → Pairwise P xs ys →
|
||||||
Pairwise P (x ∷ xs) (y ∷ ys)
|
Pairwise P (x ∷ xs) (y ∷ ys)
|
||||||
|
|
||||||
|
∈-cartesianProduct : ∀ {a b} {A : Set a} {B : Set b}
|
||||||
|
{x : A} {xs : List A} {y : B} {ys : List B} →
|
||||||
|
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')
|
||||||
|
|
||||||
|
filter-++ : ∀ {a p} {A : Set a} (l₁ l₂ : List A) {P : A → Set p} (P? : Decidable P) →
|
||||||
|
filter P? (l₁ ++ l₂) ≡ filter P? l₁ ++ filter P? l₂
|
||||||
|
filter-++ [] l₂ P? = refl
|
||||||
|
filter-++ (x ∷ xs) l₂ P?
|
||||||
|
with P? x
|
||||||
|
... | yes _ = cong (x ∷_) (filter-++ xs l₂ P?)
|
||||||
|
... | no _ = (filter-++ xs l₂ P?)
|
||||||
|
|
||||||
|
_⇒_ : ∀ {a p₁ p₂} {A : Set a} (P : A → Set p₁) (Q : A → Set p₂) →
|
||||||
|
Set (a ⊔ℓ p₁ ⊔ℓ p₂)
|
||||||
|
_⇒_ P Q = ∀ a → P a → Q a
|
||||||
|
|
||||||
|
_∨_ : ∀ {a p₁ p₂} {A : Set a} (P : A → Set p₁) (Q : A → Set p₂) →
|
||||||
|
A → Set (p₁ ⊔ℓ p₂)
|
||||||
|
_∨_ P Q a = P a ⊎ Q a
|
||||||
|
|
||||||
|
_∧_ : ∀ {a p₁ p₂} {A : Set a} (P : A → Set p₁) (Q : A → Set p₂) →
|
||||||
|
A → Set (p₁ ⊔ℓ p₂)
|
||||||
|
_∧_ P Q a = P a × Q a
|
||||||
|
|
||||||
|
it : ∀ {a} {A : Set a} {{_ : A}} → A
|
||||||
|
it {{x}} = x
|
||||||
|
|||||||
Reference in New Issue
Block a user