mirror of
https://github.com/achlipala/frap.git
synced 2025-01-07 16:44:14 +00:00
Start of AbstractInterpretation: interpret_sound
This commit is contained in:
parent
e06af75c78
commit
26023bdcb1
2 changed files with 138 additions and 1 deletions
125
AbstractInterpretation.v
Normal file
125
AbstractInterpretation.v
Normal file
|
@ -0,0 +1,125 @@
|
||||||
|
(** Formal Reasoning About Programs <http://adam.chlipala.net/frap/>
|
||||||
|
* Chapter 7: Abstract Interpretation and Dataflow Analysis
|
||||||
|
* Author: Adam Chlipala
|
||||||
|
* License: https://creativecommons.org/licenses/by-nc-nd/4.0/ *)
|
||||||
|
|
||||||
|
Require Import Frap Imp.
|
||||||
|
|
||||||
|
Set Implicit Arguments.
|
||||||
|
|
||||||
|
|
||||||
|
Record absint := {
|
||||||
|
Typeof :> Set;
|
||||||
|
(* This [:>] notation lets us treat any [absint] as its [Typeof],
|
||||||
|
* automatically. *)
|
||||||
|
Top : Typeof;
|
||||||
|
(* The least precise element of the lattice *)
|
||||||
|
Join : Typeof -> Typeof -> Typeof;
|
||||||
|
(* Least upper bound of two elements *)
|
||||||
|
Represents : nat -> Typeof -> Prop
|
||||||
|
(* Which lattice elements represent which numbers? *)
|
||||||
|
}.
|
||||||
|
|
||||||
|
Definition absint_sound (a : absint) :=
|
||||||
|
(* [Top] really does cover everything. *)
|
||||||
|
(forall n, a.(Represents) n a.(Top))
|
||||||
|
|
||||||
|
(* [Join] really does return an upper bound. *)
|
||||||
|
/\ (forall x y n, a.(Represents) n x
|
||||||
|
-> a.(Represents) n (a.(Join) x y))
|
||||||
|
/\ (forall x y n, a.(Represents) n y
|
||||||
|
-> a.(Represents) n (a.(Join) x y)).
|
||||||
|
|
||||||
|
Definition astate (a : absint) := fmap var a.
|
||||||
|
Definition astates (a : absint) := fmap cmd (astate a).
|
||||||
|
|
||||||
|
Inductive compatible a (ss : astates a) (v : valuation) (c : cmd) : Prop :=
|
||||||
|
| Compatible : forall s,
|
||||||
|
ss $? c = Some s
|
||||||
|
-> (forall x n xa, v $? x = Some n
|
||||||
|
-> s $? x = Some xa
|
||||||
|
-> a.(Represents) n xa)
|
||||||
|
-> compatible ss v c.
|
||||||
|
|
||||||
|
Definition merge_astate a : astate a -> astate a -> astate a :=
|
||||||
|
merge (fun x y =>
|
||||||
|
match x with
|
||||||
|
| None => None
|
||||||
|
| Some x' =>
|
||||||
|
match y with
|
||||||
|
| None => None
|
||||||
|
| Some y' => Some (a.(Join) x' y')
|
||||||
|
end
|
||||||
|
end).
|
||||||
|
|
||||||
|
Definition merge_astates a : astates a -> astates a -> astates a :=
|
||||||
|
merge (fun x y =>
|
||||||
|
match x with
|
||||||
|
| None => y
|
||||||
|
| Some x' =>
|
||||||
|
match y with
|
||||||
|
| None => Some x'
|
||||||
|
| Some y' => Some (merge_astate x' y')
|
||||||
|
end
|
||||||
|
end).
|
||||||
|
|
||||||
|
Inductive oneStepClosure a : astates a -> astates a -> Prop :=
|
||||||
|
| OscNil :
|
||||||
|
oneStepClosure $0 $0
|
||||||
|
| OscCons : forall ss c s ss' ss'',
|
||||||
|
(forall v c' v', step (v, c) (v', c')
|
||||||
|
-> compatible ss' v' c')
|
||||||
|
-> oneStepClosure ss ss''
|
||||||
|
-> oneStepClosure (ss $+ (c, s)) (merge_astates ss' ss'').
|
||||||
|
|
||||||
|
Inductive interpret a : astates a -> astates a -> Prop :=
|
||||||
|
| InterpretDone : forall ss,
|
||||||
|
(forall v c, compatible ss v c
|
||||||
|
-> forall v' c', step (v, c) (v', c')
|
||||||
|
-> compatible ss v' c')
|
||||||
|
-> interpret ss ss
|
||||||
|
|
||||||
|
| InterpretStep : forall ss ss' ss'',
|
||||||
|
oneStepClosure ss ss'
|
||||||
|
-> interpret (merge_astates ss ss') ss''
|
||||||
|
-> interpret ss ss''.
|
||||||
|
|
||||||
|
Lemma interpret_sound' : forall v c a (ss ss' : astates a),
|
||||||
|
interpret ss ss'
|
||||||
|
-> ss $? c = Some $0
|
||||||
|
-> invariantFor (trsys_of v c) (fun p => compatible ss' (fst p) (snd p)).
|
||||||
|
Proof.
|
||||||
|
induct 1; simplify.
|
||||||
|
|
||||||
|
apply invariant_induction; simplify; propositional; subst; simplify.
|
||||||
|
|
||||||
|
econstructor.
|
||||||
|
eassumption.
|
||||||
|
simplify.
|
||||||
|
equality.
|
||||||
|
|
||||||
|
cases s; cases s'; simplify.
|
||||||
|
eapply H.
|
||||||
|
eassumption.
|
||||||
|
assumption.
|
||||||
|
|
||||||
|
apply IHinterpret.
|
||||||
|
unfold merge_astates; simplify.
|
||||||
|
rewrite H1.
|
||||||
|
cases (ss' $? c); trivial.
|
||||||
|
f_equal.
|
||||||
|
maps_equal.
|
||||||
|
unfold merge_astate; simplify.
|
||||||
|
trivial.
|
||||||
|
Qed.
|
||||||
|
|
||||||
|
Theorem interpret_sound : forall v c a (ss : astates a),
|
||||||
|
interpret ($0 $+ (c, $0)) ss
|
||||||
|
-> invariantFor (trsys_of v c) (fun p => compatible ss (fst p) (snd p)).
|
||||||
|
Proof.
|
||||||
|
simplify.
|
||||||
|
eapply interpret_sound'.
|
||||||
|
eassumption.
|
||||||
|
simplify.
|
||||||
|
trivial.
|
||||||
|
Qed.
|
14
Map.v
14
Map.v
|
@ -9,6 +9,7 @@ Module Type S.
|
||||||
Parameter add : forall A B, fmap A B -> A -> B -> fmap A B.
|
Parameter add : forall A B, fmap A B -> A -> B -> fmap A B.
|
||||||
Parameter remove : forall A B, fmap A B -> A -> fmap A B.
|
Parameter remove : forall A B, fmap A B -> A -> fmap A B.
|
||||||
Parameter join : forall A B, fmap A B -> fmap A B -> fmap A B.
|
Parameter join : forall A B, fmap A B -> fmap A B -> fmap A B.
|
||||||
|
Parameter merge : forall A B, (option B -> option B -> option B) -> fmap A B -> fmap A B -> fmap A B.
|
||||||
Parameter lookup : forall A B, fmap A B -> A -> option B.
|
Parameter lookup : forall A B, fmap A B -> A -> option B.
|
||||||
Parameter includes : forall A B, fmap A B -> fmap A B -> Prop.
|
Parameter includes : forall A B, fmap A B -> fmap A B -> Prop.
|
||||||
|
|
||||||
|
@ -67,6 +68,9 @@ Module Type S.
|
||||||
Axiom join_assoc : forall A B (m1 m2 m3 : fmap A B),
|
Axiom join_assoc : forall A B (m1 m2 m3 : fmap A B),
|
||||||
(m1 $++ m2) $++ m3 = m1 $++ (m2 $++ m3).
|
(m1 $++ m2) $++ m3 = m1 $++ (m2 $++ m3).
|
||||||
|
|
||||||
|
Axiom lookup_merge : forall A B f (m1 m2 : fmap A B) k,
|
||||||
|
merge f m1 m2 $? k = f (m1 $? k) (m2 $? k).
|
||||||
|
|
||||||
Axiom empty_includes : forall A B (m : fmap A B), empty A B $<= m.
|
Axiom empty_includes : forall A B (m : fmap A B), empty A B $<= m.
|
||||||
|
|
||||||
Axiom dom_empty : forall A B, dom (empty A B) = {}.
|
Axiom dom_empty : forall A B, dom (empty A B) = {}.
|
||||||
|
@ -81,7 +85,7 @@ Module Type S.
|
||||||
|
|
||||||
Hint Resolve includes_lookup includes_add empty_includes.
|
Hint Resolve includes_lookup includes_add empty_includes.
|
||||||
|
|
||||||
Hint Rewrite lookup_empty lookup_add_eq lookup_add_ne lookup_remove_eq lookup_remove_ne using congruence.
|
Hint Rewrite lookup_empty lookup_add_eq lookup_add_ne lookup_remove_eq lookup_remove_ne lookup_merge using congruence.
|
||||||
|
|
||||||
Ltac maps_equal :=
|
Ltac maps_equal :=
|
||||||
apply fmap_ext; intros;
|
apply fmap_ext; intros;
|
||||||
|
@ -121,6 +125,8 @@ Module M : S.
|
||||||
| None => m2 k
|
| None => m2 k
|
||||||
| x => x
|
| x => x
|
||||||
end.
|
end.
|
||||||
|
Definition merge A B f (m1 m2 : fmap A B) : fmap A B :=
|
||||||
|
fun k => f (m1 k) (m2 k).
|
||||||
Definition lookup A B (m : fmap A B) (k : A) := m k.
|
Definition lookup A B (m : fmap A B) (k : A) := m k.
|
||||||
Definition includes A B (m1 m2 : fmap A B) :=
|
Definition includes A B (m1 m2 : fmap A B) :=
|
||||||
forall k v, m1 k = Some v -> m2 k = Some v.
|
forall k v, m1 k = Some v -> m2 k = Some v.
|
||||||
|
@ -225,6 +231,12 @@ Module M : S.
|
||||||
destruct (m1 k); auto.
|
destruct (m1 k); auto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
Theorem lookup_merge : forall A B f (m1 m2 : fmap A B) k,
|
||||||
|
lookup (merge f m1 m2) k = f (m1 k) (m2 k).
|
||||||
|
Proof.
|
||||||
|
auto.
|
||||||
|
Qed.
|
||||||
|
|
||||||
Theorem empty_includes : forall A B (m : fmap A B), includes (empty (A := A) B) m.
|
Theorem empty_includes : forall A B (m : fmap A B), includes (empty (A := A) B) m.
|
||||||
Proof.
|
Proof.
|
||||||
unfold includes, empty; intuition congruence.
|
unfold includes, empty; intuition congruence.
|
||||||
|
|
Loading…
Reference in a new issue