Compare commits
28 Commits
36ae125e1e
...
dag-lattic
| Author | SHA1 | Date | |
|---|---|---|---|
| 299938d97e | |||
| 927030c337 | |||
| ef3c351bb0 | |||
| 84c4ea6936 | |||
| a277c8f969 | |||
| d1700f23fa | |||
| eb2d64f3b5 | |||
| 14214ab5e7 | |||
| baece236d3 | |||
| 6f642d85e0 | |||
| 25fa0140f0 | |||
| 27621992ad | |||
| e409cceae5 | |||
| 8cb082e3c5 | |||
| c199e9616f | |||
| f5457d8841 | |||
| d99d4a2893 | |||
| fbb98de40f | |||
| 706b593d1d | |||
| 45606679f5 | |||
| 7e099a2561 | |||
| 2808759338 | |||
| 42bb8f8792 | |||
| 05e693594d | |||
| 90e0046707 | |||
| 13eee93255 | |||
| 6243326665 | |||
| 7b2114cd0f |
@@ -21,7 +21,7 @@ open IsFiniteHeightLattice isFiniteHeightLatticeˡ
|
||||
using () renaming (isLattice to isLatticeˡ)
|
||||
|
||||
module WithProg (prog : Program) where
|
||||
open import Analysis.Forward.Lattices L prog hiding (≈ᵛ-Decidable) -- to disambiguate instance resolution
|
||||
open import Analysis.Forward.Lattices L prog hiding (≈ᵛ-Decidable; ≈ᵐ-Decidable) -- to disambiguate instance resolution
|
||||
open import Analysis.Forward.Evaluation L prog
|
||||
open Program prog
|
||||
|
||||
@@ -66,7 +66,7 @@ module WithProg (prog : Program) where
|
||||
(joinAll-Mono {sv₁} {sv₂} sv₁≼sv₂)
|
||||
|
||||
-- The fixed point of the 'analyze' function is our final goal.
|
||||
open import Fixedpoint ≈ᵐ-Decidable isFiniteHeightLatticeᵐ analyze (λ {m₁} {m₂} m₁≼m₂ → analyze-Mono {m₁} {m₂} m₁≼m₂)
|
||||
open import Fixedpoint analyze (λ {m₁} {m₂} m₁≼m₂ → analyze-Mono {m₁} {m₂} m₁≼m₂)
|
||||
using ()
|
||||
renaming (aᶠ to result; aᶠ≈faᶠ to result≈analyze-result)
|
||||
public
|
||||
|
||||
@@ -5,8 +5,8 @@ module Fixedpoint {a} {A : Set a}
|
||||
{h : ℕ}
|
||||
{_≈_ : A → A → Set a}
|
||||
{_⊔_ : A → A → A} {_⊓_ : A → A → A}
|
||||
(≈-Decidable : IsDecidable _≈_)
|
||||
(flA : IsFiniteHeightLattice A h _≈_ _⊔_ _⊓_)
|
||||
{{ ≈-Decidable : IsDecidable _≈_ }}
|
||||
{{flA : IsFiniteHeightLattice A h _≈_ _⊔_ _⊓_}}
|
||||
(f : A → A)
|
||||
(Monotonicᶠ : Monotonic (IsFiniteHeightLattice._≼_ flA)
|
||||
(IsFiniteHeightLattice._≼_ flA) f) where
|
||||
@@ -28,24 +28,9 @@ private
|
||||
using ()
|
||||
renaming
|
||||
( ⊥ to ⊥ᴬ
|
||||
; longestChain to longestChainᴬ
|
||||
; bounded to boundedᴬ
|
||||
)
|
||||
|
||||
|
||||
⊥ᴬ≼ : ∀ (a : A) → ⊥ᴬ ≼ a
|
||||
⊥ᴬ≼ a with ≈-dec a ⊥ᴬ
|
||||
... | yes a≈⊥ᴬ = ≼-cong a≈⊥ᴬ ≈-refl (≼-refl a)
|
||||
... | no a̷≈⊥ᴬ with ≈-dec ⊥ᴬ (a ⊓ ⊥ᴬ)
|
||||
... | yes ⊥ᴬ≈a⊓⊥ᴬ = ≈-trans (⊔-comm ⊥ᴬ a) (≈-trans (≈-⊔-cong (≈-refl {a}) ⊥ᴬ≈a⊓⊥ᴬ) (absorb-⊔-⊓ a ⊥ᴬ))
|
||||
... | no ⊥ᴬ̷≈a⊓⊥ᴬ = ⊥-elim (ChainA.Bounded-suc-n boundedᴬ (ChainA.step x≺⊥ᴬ ≈-refl longestChainᴬ))
|
||||
where
|
||||
⊥ᴬ⊓a̷≈⊥ᴬ : ¬ (⊥ᴬ ⊓ a) ≈ ⊥ᴬ
|
||||
⊥ᴬ⊓a̷≈⊥ᴬ = λ ⊥ᴬ⊓a≈⊥ᴬ → ⊥ᴬ̷≈a⊓⊥ᴬ (≈-trans (≈-sym ⊥ᴬ⊓a≈⊥ᴬ) (⊓-comm _ _))
|
||||
|
||||
x≺⊥ᴬ : (⊥ᴬ ⊓ a) ≺ ⊥ᴬ
|
||||
x≺⊥ᴬ = (≈-trans (⊔-comm _ _) (≈-trans (≈-refl {⊥ᴬ ⊔ (⊥ᴬ ⊓ a)}) (absorb-⊔-⊓ ⊥ᴬ a)) , ⊥ᴬ⊓a̷≈⊥ᴬ)
|
||||
|
||||
-- using 'g', for gas, here helps make sure the function terminates.
|
||||
-- since A forms a fixed-height lattice, we must find a solution after
|
||||
-- 'h' steps at most. Gas is set up such that as soon as it runs
|
||||
@@ -65,7 +50,7 @@ private
|
||||
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 = 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ᶠ = proj₁ fix
|
||||
@@ -85,4 +70,4 @@ private
|
||||
... | 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≈fa = stepPreservesLess (suc h) 0 ⊥ᴬ ⊥ᴬ a a≈fa (⊥ᴬ≼ a) (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥ᴬ≼ (f ⊥ᴬ))
|
||||
aᶠ≼ a a≈fa = stepPreservesLess (suc h) 0 ⊥ᴬ ⊥ᴬ a a≈fa (⊥≼ a) (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥≼ (f ⊥ᴬ))
|
||||
|
||||
@@ -68,6 +68,7 @@ module TransportFiniteHeight
|
||||
renaming (⊥ to ⊥₁; ⊤ to ⊤₁; bounded to bounded₁; longestChain to c)
|
||||
|
||||
instance
|
||||
fixedHeight : IsLattice.FixedHeight lB height
|
||||
fixedHeight = record
|
||||
{ ⊥ = f ⊥₁
|
||||
; ⊤ = f ⊤₁
|
||||
|
||||
@@ -124,28 +124,6 @@ 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})
|
||||
@@ -154,13 +132,13 @@ module _ (g : Graph) where
|
||||
open import Data.List.Membership.DecPropositional (_≟_) using (_∈?_)
|
||||
|
||||
indices : List (Graph.Index g)
|
||||
indices = proj₁ (finValues (Graph.size g))
|
||||
indices = proj₁ (fins (Graph.size g))
|
||||
|
||||
indices-complete : ∀ (idx : (Graph.Index g)) → idx ListMem.∈ indices
|
||||
indices-complete = finValues-complete (Graph.size g)
|
||||
indices-complete = fins-complete (Graph.size g)
|
||||
|
||||
indices-Unique : Unique indices
|
||||
indices-Unique = proj₂ (finValues (Graph.size g))
|
||||
indices-Unique = proj₂ (fins (Graph.size g))
|
||||
|
||||
predecessors : (Graph.Index g) → List (Graph.Index g)
|
||||
predecessors idx = List.filter (λ idx' → (idx' , idx) ∈? (Graph.edges g)) indices
|
||||
|
||||
25
Lattice.agda
25
Lattice.agda
@@ -4,7 +4,8 @@ open import Equivalence
|
||||
import Chain
|
||||
|
||||
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.Product using (_×_; Σ; _,_)
|
||||
open import Data.Sum using (_⊎_; inj₁; inj₂)
|
||||
@@ -236,6 +237,28 @@ record IsFiniteHeightLattice {a} (A : Set a)
|
||||
field
|
||||
{{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}
|
||||
{_≈₁_ : A → A → Set a} {_≈₂_ : B → B → Set b}
|
||||
{_⊔₁_ : A → A → A} {_⊔₂_ : B → B → B}
|
||||
|
||||
1149
Lattice/Builder.agda
1149
Lattice/Builder.agda
File diff suppressed because it is too large
Load Diff
76
Utils.agda
76
Utils.agda
@@ -1,19 +1,31 @@
|
||||
module Utils where
|
||||
|
||||
open import Agda.Primitive using () renaming (_⊔_ to _⊔ℓ_)
|
||||
open import Data.Product as Prod using (_×_)
|
||||
open import Data.Product as Prod using (Σ; _×_; _,_; proj₁; proj₂)
|
||||
open import Data.Nat using (ℕ; suc)
|
||||
open import Data.Fin as Fin using (Fin; suc; zero)
|
||||
open import Data.Fin.Properties using (suc-injective)
|
||||
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 (_∈_; lose)
|
||||
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.Any using (Any; here; there) -- TODO: re-export these with nicer names from map
|
||||
open import Data.List.Relation.Unary.All using (All; []; _∷_; map; all?; lookup)
|
||||
open import Data.List.Relation.Unary.All.Properties using (++⁻ˡ; ++⁻ʳ)
|
||||
open import Data.List.Relation.Unary.Any as Any using (Any; here; there; any?) -- TODO: re-export these with nicer names from map
|
||||
open import Data.Sum using (_⊎_)
|
||||
open import Function.Definitions using (Injective)
|
||||
open import Relation.Binary using (Antisymmetric) renaming (Decidable to Decidable²)
|
||||
open import Relation.Binary.PropositionalEquality using (_≡_; sym; refl; cong)
|
||||
open import Relation.Nullary using (¬_; yes; no)
|
||||
open import Relation.Nullary using (¬_; yes; no; Dec)
|
||||
open import Relation.Nullary.Decidable using (¬?)
|
||||
open import Relation.Unary using (Decidable)
|
||||
|
||||
All¬-¬Any : ∀ {p c} {C : Set c} {P : C → Set p} {l : List C} → All (λ x → ¬ P x) l → ¬ Any P l
|
||||
All¬-¬Any {l = x ∷ xs} (¬Px ∷ _) (here Px) = ¬Px Px
|
||||
All¬-¬Any {l = x ∷ xs} (_ ∷ ¬Pxs) (there Pxs) = All¬-¬Any ¬Pxs Pxs
|
||||
|
||||
Decidable-¬ : ∀ {p c} {C : Set c} {P : C → Set p} → Decidable P → Decidable (λ x → ¬ P x)
|
||||
Decidable-¬ Decidable-P x = ¬? (Decidable-P x)
|
||||
|
||||
data Unique {c} {C : Set c} : List C → Set c where
|
||||
empty : Unique []
|
||||
push : ∀ {x : C} {xs : List C}
|
||||
@@ -34,6 +46,24 @@ Unique-append {c} {C} {x} {x' ∷ xs'} x∉xs (push x'≢ uxs') =
|
||||
help {[]} _ = x'≢x ∷ []
|
||||
help {e ∷ es} (x'≢e ∷ x'≢es) = x'≢e ∷ help x'≢es
|
||||
|
||||
Unique-++⁻ˡ : ∀ {c} {C : Set c} (xs : List C) {ys : List C} → Unique (xs ++ ys) → Unique xs
|
||||
Unique-++⁻ˡ [] Unique-ys = empty
|
||||
Unique-++⁻ˡ (x ∷ xs) {ys} (push x≢xs++ys Unique-xs++ys) = push (++⁻ˡ xs {ys = ys} x≢xs++ys) (Unique-++⁻ˡ xs Unique-xs++ys)
|
||||
|
||||
Unique-++⁻ʳ : ∀ {c} {C : Set c} (xs : List C) {ys : List C} → Unique (xs ++ ys) → Unique ys
|
||||
Unique-++⁻ʳ [] Unique-ys = Unique-ys
|
||||
Unique-++⁻ʳ (x ∷ xs) {ys} (push x≢xs++ys Unique-xs++ys) = Unique-++⁻ʳ xs Unique-xs++ys
|
||||
|
||||
Unique-∈-++ˡ : ∀ {c} {C : Set c} {x : C} (xs : List C) {ys : List C} → Unique (xs ++ ys) → x ∈ xs → ¬ x ∈ ys
|
||||
Unique-∈-++ˡ [] _ ()
|
||||
Unique-∈-++ˡ {x = x} (x' ∷ xs) (push x≢xs++ys _) (here refl) = All¬-¬Any (++⁻ʳ xs x≢xs++ys)
|
||||
Unique-∈-++ˡ {x = x} (x' ∷ xs) (push _ Unique-xs++ys) (there x̷∈xs) = Unique-∈-++ˡ xs Unique-xs++ys x̷∈xs
|
||||
|
||||
Unique-narrow : ∀ {c} {C : Set c} {x : C} (xs : List C) {ys : List C} → Unique (xs ++ ys) → x ∈ xs → Unique (x ∷ ys)
|
||||
Unique-narrow [] _ ()
|
||||
Unique-narrow {x = x} (x' ∷ xs) (push x≢xs++ys Unique-xs++ys) (here refl) = push (++⁻ʳ xs x≢xs++ys) (Unique-++⁻ʳ xs Unique-xs++ys)
|
||||
Unique-narrow {x = x} (x' ∷ xs) (push _ Unique-xs++ys) (there x̷∈xs) = Unique-narrow xs Unique-xs++ys x̷∈xs
|
||||
|
||||
All-≢-map : ∀ {c d} {C : Set c} {D : Set d} (x : C) {xs : List C} (f : C → D) →
|
||||
Injective (_≡_ {_} {C}) (_≡_ {_} {D}) f →
|
||||
All (λ x' → ¬ x ≡ x') xs → All (λ y' → ¬ (f x) ≡ y') (mapˡ f xs)
|
||||
@@ -46,9 +76,8 @@ Unique-map : ∀ {c d} {C : Set c} {D : Set d} {l : List C} (f : C → D) →
|
||||
Unique-map {l = []} _ _ _ = empty
|
||||
Unique-map {l = x ∷ xs} f f-Injecitve (push x≢xs uxs) = push (All-≢-map x f f-Injecitve x≢xs) (Unique-map f f-Injecitve uxs)
|
||||
|
||||
All¬-¬Any : ∀ {p c} {C : Set c} {P : C → Set p} {l : List C} → All (λ x → ¬ P x) l → ¬ Any P l
|
||||
All¬-¬Any {l = x ∷ xs} (¬Px ∷ _) (here Px) = ¬Px Px
|
||||
All¬-¬Any {l = x ∷ xs} (_ ∷ ¬Pxs) (there Pxs) = All¬-¬Any ¬Pxs Pxs
|
||||
¬Any-map : ∀ {p₁ p₂ c} {C : Set c} {P₁ : C → Set p₁} {P₂ : C → Set p₂} {l : List C} → (∀ {x} → P₁ x → P₂ x) → ¬ Any P₂ l → ¬ Any P₁ l
|
||||
¬Any-map f ¬Any-P₂ Any-P₁ = ¬Any-P₂ (Any.map f Any-P₁)
|
||||
|
||||
All-single : ∀ {p c} {C : Set c} {P : C → Set p} {c : C} {l : List C} → All P l → c ∈ l → P c
|
||||
All-single {c = c} {l = x ∷ xs} (p ∷ ps) (here refl) = p
|
||||
@@ -106,3 +135,34 @@ _∧_ P Q a = P a × Q a
|
||||
|
||||
it : ∀ {a} {A : Set a} {{_ : A}} → A
|
||||
it {{x}} = x
|
||||
|
||||
z≢sf : ∀ {n : ℕ} (f : Fin n) → ¬ (Fin.zero ≡ Fin.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'
|
||||
|
||||
fins : ∀ (n : ℕ) → Σ (List (Fin n)) Unique
|
||||
fins 0 = ([] , empty)
|
||||
fins (suc n') =
|
||||
let
|
||||
(inds' , unids') = fins n'
|
||||
in
|
||||
( zero ∷ mapˡ suc inds'
|
||||
, push (z≢mapsfs inds') (Unique-map suc suc-injective unids')
|
||||
)
|
||||
|
||||
fins-complete : ∀ (n : ℕ) (f : Fin n) → f ∈ (proj₁ (fins n))
|
||||
fins-complete (suc n') zero = here refl
|
||||
fins-complete (suc n') (suc f') = there (x∈xs⇒fx∈fxs suc (fins-complete n' f'))
|
||||
|
||||
findUniversal : ∀ {p c} {C : Set c} {R : C → C → Set p} (l : List C) → Decidable² R →
|
||||
Dec (Any (λ x → All (R x) l) l)
|
||||
findUniversal l Rdec = any? (λ x → all? (Rdec x) l) l
|
||||
|
||||
findUniversal-unique : ∀ {p c} {C : Set c} (R : C → C → Set p) (l : List C) →
|
||||
Antisymmetric _≡_ R →
|
||||
∀ x₁ x₂ → x₁ ∈ l → x₂ ∈ l → All (R x₁) l → All (R x₂) l →
|
||||
x₁ ≡ x₂
|
||||
findUniversal-unique R l Rantisym x₁ x₂ x₁∈l x₂∈l Allx₁ Allx₂ = Rantisym (lookup Allx₁ x₂∈l) (lookup Allx₂ x₁∈l)
|
||||
|
||||
Reference in New Issue
Block a user