Use instance search to avoid multiply-nested modules
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
@@ -159,19 +159,20 @@ s₁≢s₂⇒¬s₁∧s₂ { - } { - } +≢+ _ = ⊥-elim (+≢+ refl)
|
||||
⟦⟧ᵍ-⊓ᵍ-∧ {[ g₁ ]ᵍ} {⊥ᵍ} x (_ , bot) = bot
|
||||
⟦⟧ᵍ-⊓ᵍ-∧ {[ g₁ ]ᵍ} {⊤ᵍ} x (px₁ , _) = px₁
|
||||
|
||||
latticeInterpretationᵍ : LatticeInterpretation isLatticeᵍ
|
||||
latticeInterpretationᵍ = record
|
||||
{ ⟦_⟧ = ⟦_⟧ᵍ
|
||||
; ⟦⟧-respects-≈ = ⟦⟧ᵍ-respects-≈ᵍ
|
||||
; ⟦⟧-⊔-∨ = ⟦⟧ᵍ-⊔ᵍ-∨
|
||||
; ⟦⟧-⊓-∧ = ⟦⟧ᵍ-⊓ᵍ-∧
|
||||
}
|
||||
instance
|
||||
latticeInterpretationᵍ : LatticeInterpretation isLatticeᵍ
|
||||
latticeInterpretationᵍ = record
|
||||
{ ⟦_⟧ = ⟦_⟧ᵍ
|
||||
; ⟦⟧-respects-≈ = ⟦⟧ᵍ-respects-≈ᵍ
|
||||
; ⟦⟧-⊔-∨ = ⟦⟧ᵍ-⊔ᵍ-∨
|
||||
; ⟦⟧-⊓-∧ = ⟦⟧ᵍ-⊓ᵍ-∧
|
||||
}
|
||||
|
||||
module WithProg (prog : Program) where
|
||||
open Program prog
|
||||
|
||||
module ForwardWithProg = Analysis.Forward.WithProg (record { isLattice = isLatticeᵍ; fixedHeight = fixedHeightᵍ }) ≈ᵍ-dec prog
|
||||
open ForwardWithProg
|
||||
open ForwardWithProg hiding (analyze-correct)
|
||||
|
||||
eval : ∀ (e : Expr) → VariableValues → SignLattice
|
||||
eval (e₁ + e₂) vs = plus (eval e₁ vs) (eval e₂ vs)
|
||||
@@ -222,15 +223,13 @@ module WithProg (prog : Program) where
|
||||
eval-Mono (# 0) _ = ≈ᵍ-refl
|
||||
eval-Mono (# (suc n')) _ = ≈ᵍ-refl
|
||||
|
||||
module ForwardWithEval = ForwardWithProg.WithEvaluator eval eval-Mono
|
||||
open ForwardWithEval using (result)
|
||||
instance
|
||||
SignEval : Evaluator
|
||||
SignEval = record { eval = eval; eval-Mono = eval-Mono }
|
||||
|
||||
-- For debugging purposes, print out the result.
|
||||
output = show result
|
||||
|
||||
module ForwardWithInterp = ForwardWithEval.WithInterpretation latticeInterpretationᵍ
|
||||
open ForwardWithInterp using (⟦_⟧ᵛ; InterpretationValid)
|
||||
|
||||
-- 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.
|
||||
@@ -281,16 +280,16 @@ module WithProg (prog : Program) where
|
||||
minus-valid {[ 0ˢ ]ᵍ} {[ 0ˢ ]ᵍ} refl refl = refl
|
||||
minus-valid {[ 0ˢ ]ᵍ} {⊤ᵍ} _ _ = tt
|
||||
|
||||
eval-Valid : InterpretationValid
|
||||
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⟧ρ
|
||||
eval-valid : IsValid
|
||||
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 (⇒ᵉ-ℕ ρ 0) _ = refl
|
||||
eval-Valid (⇒ᵉ-ℕ ρ (suc n')) _ = (n' , refl)
|
||||
eval-valid (⇒ᵉ-ℕ ρ 0) _ = refl
|
||||
eval-valid (⇒ᵉ-ℕ ρ (suc n')) _ = (n' , refl)
|
||||
|
||||
open ForwardWithInterp.WithValidity eval-Valid using (analyze-correct) public
|
||||
analyze-correct = ForwardWithProg.analyze-correct
|
||||
|
||||
Reference in New Issue
Block a user