Make 'IsDecidable' into a record to aid instance search

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
2025-01-04 18:58:56 -08:00
parent 8abf6f8670
commit b0488c9cc6
14 changed files with 146 additions and 115 deletions

View File

@@ -7,6 +7,7 @@ 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.Definitions using (Decidable)
open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; trans; subst)
open import Relation.Nullary using (¬_; yes; no)
@@ -32,7 +33,7 @@ instance
}
-- g for siGn; s is used for strings and i is not very descriptive.
_≟ᵍ_ : IsDecidable (_≡_ {_} {Sign})
_≟ᵍ_ : Decidable (_≡_ {_} {Sign})
_≟ᵍ_ + + = yes refl
_≟ᵍ_ + - = no (λ ())
_≟ᵍ_ + 0ˢ = no (λ ())
@@ -43,12 +44,15 @@ _≟ᵍ_ 0ˢ + = no (λ ())
_≟ᵍ_ 0ˢ - = no (λ ())
_≟ᵍ_ 0ˢ 0ˢ = yes refl
≡-Decidable-Sign : IsDecidable {_} {Sign} _≡_
≡-Decidable-Sign = record { R-dec = _≟ᵍ_ }
-- 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 _≡_ (record { ≈-refl = refl; ≈-sym = sym; ≈-trans = trans }) ≡-Decidable-Sign as AB
using ()
renaming
( AboveBelow to SignLattice
; ≈-dec to ≈ᵍ-dec
; ≈-Decidable to ≈ᵍ-Decidable
; to ⊥ᵍ
; to ⊤ᵍ
; [_] to [_]ᵍ
@@ -171,9 +175,9 @@ instance
module WithProg (prog : Program) where
open Program prog
open import Analysis.Forward.Lattices isFiniteHeightLatticeᵍ ≈ᵍ-dec prog
open import Analysis.Forward.Evaluation isFiniteHeightLatticeᵍ ≈ᵍ-dec prog
open import Analysis.Forward.Adapters isFiniteHeightLatticeᵍ ≈ᵍ-dec prog
open import Analysis.Forward.Lattices isFiniteHeightLatticeᵍ ≈ᵍ-Decidable prog
open import Analysis.Forward.Evaluation isFiniteHeightLatticeᵍ ≈ᵍ-Decidable prog
open import Analysis.Forward.Adapters isFiniteHeightLatticeᵍ ≈ᵍ-Decidable prog
eval : (e : Expr) VariableValues SignLattice
eval (e₁ + e₂) vs = plus (eval e₁ vs) (eval e₂ vs)
@@ -229,7 +233,7 @@ module WithProg (prog : Program) where
SignEval = record { eval = eval; eval-Monoʳ = eval-Monoʳ }
-- For debugging purposes, print out the result.
output = show (Analysis.Forward.WithProg.result isFiniteHeightLatticeᵍ ≈ᵍ-dec prog)
output = show (Analysis.Forward.WithProg.result isFiniteHeightLatticeᵍ ≈ᵍ-Decidable 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