More on AbstractInterpretation example; need to do a proper abstraction into a new trsys

This commit is contained in:
Adam Chlipala 2016-03-04 16:14:41 -05:00
parent 26023bdcb1
commit d0d6b87a1d
2 changed files with 414 additions and 7 deletions

View file

@ -12,8 +12,8 @@ Record absint := {
Typeof :> Set;
(* This [:>] notation lets us treat any [absint] as its [Typeof],
* automatically. *)
Top : Typeof;
(* The least precise element of the lattice *)
Const : nat -> Typeof;
(* Most accurate representation of a constant *)
Join : Typeof -> Typeof -> Typeof;
(* Least upper bound of two elements *)
Represents : nat -> Typeof -> Prop
@ -21,8 +21,8 @@ Record absint := {
}.
Definition absint_sound (a : absint) :=
(* [Top] really does cover everything. *)
(forall n, a.(Represents) n a.(Top))
(* [Const] gives accurate answers. *)
(forall n, a.(Represents) n (a.(Const) n))
(* [Join] really does return an upper bound. *)
/\ (forall x y n, a.(Represents) n x
@ -33,12 +33,15 @@ Definition absint_sound (a : absint) :=
Definition astate (a : absint) := fmap var a.
Definition astates (a : absint) := fmap cmd (astate a).
Definition compatible1 a (s : astate a) (v : valuation) (c : cmd) : Prop :=
forall x n xa, v $? x = Some n
-> s $? x = Some xa
-> a.(Represents) n xa.
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)
-> compatible1 s v c
-> compatible ss v c.
Definition merge_astate a : astate a -> astate a -> astate a :=
@ -68,6 +71,7 @@ Inductive oneStepClosure a : astates a -> astates a -> Prop :=
oneStepClosure $0 $0
| OscCons : forall ss c s ss' ss'',
(forall v c' v', step (v, c) (v', c')
-> compatible1 s v c
-> compatible ss' v' c')
-> oneStepClosure ss ss''
-> oneStepClosure (ss $+ (c, s)) (merge_astates ss' ss'').
@ -95,6 +99,7 @@ Proof.
econstructor.
eassumption.
unfold compatible1.
simplify.
equality.
@ -123,3 +128,273 @@ Proof.
simplify.
trivial.
Qed.
(** * Example: even-odd analysis *)
Inductive parity := Even | Odd | Either.
Definition isEven (n : nat) := exists k, n = k * 2.
Definition isOdd (n : nat) := exists k, n = k * 2 + 1.
Theorem decide_parity : forall n, isEven n \/ isOdd n.
Proof.
induct n; simplify; propositional.
left; exists 0; linear_arithmetic.
invert H.
right.
exists x; linear_arithmetic.
invert H.
left.
exists (x + 1); linear_arithmetic.
Qed.
Theorem notEven_odd : forall n, ~isEven n -> isOdd n.
Proof.
simplify.
assert (isEven n \/ isOdd n).
apply decide_parity.
propositional.
Qed.
Theorem isEven_0 : isEven 0.
Proof.
exists 0; linear_arithmetic.
Qed.
Theorem isEven_1 : ~isEven 1.
Proof.
propositional; invert H; linear_arithmetic.
Qed.
Theorem isEven_S_Even : forall n, isEven n -> ~isEven (S n).
Proof.
propositional; invert H; invert H0; linear_arithmetic.
Qed.
Theorem isEven_S_Odd : forall n, ~isEven n -> isEven (S n).
Proof.
propositional.
apply notEven_odd in H.
invert H.
exists (x + 1); linear_arithmetic.
Qed.
Hint Resolve isEven_0 isEven_1 isEven_S_Even isEven_S_Odd.
Definition parity_flip (p : parity) :=
match p with
| Even => Odd
| Odd => Even
| Either => Either
end.
Fixpoint parity_const (n : nat) :=
match n with
| O => Even
| S n' => parity_flip (parity_const n')
end.
Definition parity_join (x y : parity) :=
match x, y with
| Even, Even => Even
| Odd, Odd => Odd
| _, _ => Either
end.
Inductive parity_rep : nat -> parity -> Prop :=
| PrEven : forall n,
isEven n
-> parity_rep n Even
| PrOdd : forall n,
~isEven n
-> parity_rep n Odd
| PrEither : forall n,
parity_rep n Either.
Hint Constructors parity_rep.
Definition parity_absint := {|
Const := parity_const;
Join := parity_join;
Represents := parity_rep
|}.
Lemma parity_const_sound : forall n,
parity_rep n (parity_const n).
Proof.
induct n; simplify; eauto.
cases (parity_const n); simplify; eauto.
invert IHn; eauto.
invert IHn; eauto.
Qed.
Hint Resolve parity_const_sound.
Theorem parity_sound : absint_sound parity_absint.
Proof.
unfold absint_sound; propositional.
simplify; eauto.
invert H; cases y; simplify; eauto.
invert H; cases x; simplify; eauto.
Qed.
Definition loopy :=
"n" <- 100;;
"a" <- 0;;
while "n" loop
"a" <- "a" + "n";;
"n" <- "n" - 2
done.
Theorem compatible_skip : forall (s : astate parity_absint) v c c' m,
compatible1 s v c
-> compatible (m $+ (c', s)) v c'.
Proof.
unfold compatible1; simplify.
econstructor.
simplify; equality.
auto.
Qed.
Theorem compatible_skip2 : forall (s : astate parity_absint) v c c' m c'' s',
compatible1 s v c
-> c'' <> c'
-> compatible (m $+ (c', s) $+ (c'', s')) v c'.
Proof.
unfold compatible1; simplify.
econstructor.
simplify; equality.
auto.
Qed.
Theorem compatible_const : forall (s : astate parity_absint) v c c' x n,
compatible1 s v c
-> compatible ($0 $+ (c', s $+ (x, parity_const n))) (v $+ (x, n)) c'.
Proof.
unfold compatible1; simplify.
econstructor.
simplify; equality.
unfold compatible1.
simplify.
cases (x ==v x0); simplify.
invert H1.
invert H0.
eauto.
eapply H.
eassumption.
assumption.
Qed.
Hint Rewrite merge_empty1 merge_empty2 using solve [ eauto 1 ].
Hint Rewrite merge_empty1_alt merge_empty2_alt using congruence.
Lemma merge_astates_fok : forall x : option (astate parity_absint),
match x with Some x' => Some x' | None => None end = x.
Proof.
simplify; cases x; equality.
Qed.
Lemma merge_astates_fok2 : forall x (y : option (astate parity_absint)),
match y with
| Some y' => Some (merge_astate x y')
| None => Some x
end = None -> False.
Proof.
simplify; cases y; equality.
Qed.
Hint Resolve merge_astates_fok merge_astates_fok2.
Hint Rewrite merge_add1 using solve [ eauto | unfold Sets.In; simplify; equality ].
Hint Rewrite merge_add1_alt using solve [ congruence | unfold Sets.In; simplify; equality ].
Ltac inList x xs :=
match xs with
| (x, _) => constr:true
| (_, ?xs') => inList x xs'
| _ => constr:false
end.
Ltac maybe_simplify_map m found kont :=
match m with
| @empty ?A ?B => kont (@empty A B)
| ?m' $+ (?k, ?v) =>
let iL := inList k found in
match iL with
| true => maybe_simplify_map m' found kont
| false =>
maybe_simplify_map m' (k, found) ltac:(fun m' => kont (m' $+ (k, v)))
end
end.
Ltac simplify_map' m found kont :=
match m with
| ?m' $+ (?k, ?v) =>
let iL := inList k found in
match iL with
| true => maybe_simplify_map m' found kont
| false =>
simplify_map' m' (k, found) ltac:(fun m' => kont (m' $+ (k, v)))
end
end.
Ltac simplify_map :=
match goal with
| [ |- context[?m $+ (?k, ?v)] ] =>
simplify_map' (m $+ (k, v)) tt ltac:(fun m' =>
replace (m $+ (k, v)) with m' by maps_equal)
end.
Definition easy :=
"n" <- 10;;
while "n" loop
"n" <- "n" - 2
done.
Theorem easy_even : forall v n,
isEven n
-> invariantFor (trsys_of v easy) (fun p => exists n, fst p $? "n" = Some n /\ isEven n).
Proof.
simplify.
eapply invariant_weaken.
unfold easy.
apply interpret_sound.
Ltac interpret1 :=
eapply InterpretStep; [ (repeat (apply OscNil || eapply OscCons);
invert 1; repeat (match goal with
| [ H : step _ _ |- _ ] => invert H
| [ _ : match ?E with
| Some n => n
| None => 0
end = 0 |- _ ] => cases E; try linear_arithmetic
end; simplify); simplify;
try solve [ eapply compatible_skip; eassumption
| eapply compatible_skip2; eassumption || congruence
| eapply compatible_const; eassumption ]) | ];
repeat match goal with
| [ |- context[@add ?A ?B ?m _ _] ] =>
match m with
| @empty _ _ => fail 1
| _ => unify m (@empty A B)
end
end.
interpret1.
unfold merge_astates, merge_astate; simplify.
interpret1.
unfold merge_astates, merge_astate; simplify.
simplify_map.
interpret1.
unfold merge_astates, merge_astate; simplify.
simplify_map.
interpret1.

132
Map.v
View file

@ -71,6 +71,47 @@ Module Type S.
Axiom lookup_merge : forall A B f (m1 m2 : fmap A B) k,
merge f m1 m2 $? k = f (m1 $? k) (m2 $? k).
Axiom merge_empty1 : forall A B f (m : fmap A B),
(forall x, f None x = x)
-> merge f (@empty _ _) m = m.
Axiom merge_empty2 : forall A B f (m : fmap A B),
(forall x, f x None = x)
-> merge f m (@empty _ _) = m.
Axiom merge_empty1_alt : forall A B f (m : fmap A B),
(forall x, f None x = None)
-> merge f (@empty _ _) m = @empty _ _.
Axiom merge_empty2_alt : forall A B f (m : fmap A B),
(forall x, f x None = None)
-> merge f m (@empty _ _) = @empty _ _.
Axiom merge_add1 : forall A B f (m1 m2 : fmap A B) k v,
(forall x y, f (Some x) y = None -> False)
-> ~k \in dom m1
-> merge f (add m1 k v) m2 = match f (Some v) (lookup m2 k) with
| None => merge f m1 m2
| Some v => add (merge f m1 m2) k v
end.
Axiom merge_add2 : forall A B f (m1 m2 : fmap A B) k v,
(forall x y, f x (Some y) = None -> False)
-> ~k \in dom m2
-> merge f m1 (add m2 k v) = match f (lookup m1 k) (Some v) with
| None => merge f m1 m2
| Some v => add (merge f m1 m2) k v
end.
Axiom merge_add1_alt : forall A B f (m1 m2 : fmap A B) k v,
(forall x y, f (Some x) (Some y) = None -> False)
-> ~k \in dom m1
-> k \in dom m2
-> merge f (add m1 k v) m2 = match f (Some v) (lookup m2 k) with
| None => merge f m1 m2
| Some v => add (merge f m1 m2) k v
end.
Axiom empty_includes : forall A B (m : fmap A B), empty A B $<= m.
Axiom dom_empty : forall A B, dom (empty A B) = {}.
@ -87,6 +128,8 @@ Module Type S.
Hint Rewrite lookup_empty lookup_add_eq lookup_add_ne lookup_remove_eq lookup_remove_ne lookup_merge using congruence.
Hint Rewrite dom_empty dom_add.
Ltac maps_equal :=
apply fmap_ext; intros;
repeat (subst; autorewrite with core; try reflexivity;
@ -237,6 +280,95 @@ Module M : S.
auto.
Qed.
Theorem merge_empty1 : forall A B f (m : fmap A B),
(forall x, f None x = x)
-> merge f (@empty _ _) m = m.
Proof.
intros; apply fmap_ext; unfold lookup, merge; auto.
Qed.
Theorem merge_empty2 : forall A B f (m : fmap A B),
(forall x, f x None = x)
-> merge f m (@empty _ _) = m.
Proof.
intros; apply fmap_ext; unfold lookup, merge; auto.
Qed.
Theorem merge_empty1_alt : forall A B f (m : fmap A B),
(forall x, f None x = None)
-> merge f (@empty _ _) m = @empty _ _.
Proof.
intros; apply fmap_ext; unfold lookup, merge; auto.
Qed.
Theorem merge_empty2_alt : forall A B f (m : fmap A B),
(forall x, f x None = None)
-> merge f m (@empty _ _) = @empty _ _.
Proof.
intros; apply fmap_ext; unfold lookup, merge; auto.
Qed.
Theorem merge_add1 : forall A B f (m1 m2 : fmap A B) k v,
(forall x y, f (Some x) y = None -> False)
-> ~k \in dom m1
-> merge f (add m1 k v) m2 = match f (Some v) (lookup m2 k) with
| None => merge f m1 m2
| Some v => add (merge f m1 m2) k v
end.
Proof.
intros; apply fmap_ext; unfold lookup, merge, add; intros.
destruct (decide (k0 = k)); auto; subst.
case_eq (f (Some v) (m2 k)); intros.
case_eq (decide (k = k)); congruence.
exfalso; eauto.
case_eq (f (Some v) (m2 k)); intros.
destruct (decide (k0 = k)); congruence.
auto.
Qed.
Theorem merge_add2 : forall A B f (m1 m2 : fmap A B) k v,
(forall x y, f x (Some y) = None -> False)
-> ~k \in dom m2
-> merge f m1 (add m2 k v) = match f (lookup m1 k) (Some v) with
| None => merge f m1 m2
| Some v => add (merge f m1 m2) k v
end.
Proof.
intros; apply fmap_ext; unfold lookup, merge, add; intros.
destruct (decide (k0 = k)); auto; subst.
case_eq (f (m1 k) (Some v)); intros.
case_eq (decide (k = k)); congruence.
exfalso; eauto.
case_eq (f (m1 k) (Some v)); intros.
destruct (decide (k0 = k)); congruence.
auto.
Qed.
Theorem merge_add1_alt : forall A B f (m1 m2 : fmap A B) k v,
(forall x y, f (Some x) (Some y) = None -> False)
-> ~k \in dom m1
-> k \in dom m2
-> merge f (add m1 k v) m2 = match f (Some v) (lookup m2 k) with
| None => merge f m1 m2
| Some v => add (merge f m1 m2) k v
end.
Proof.
intros; apply fmap_ext; unfold lookup, merge, add; intros.
destruct (decide (k0 = k)); auto; subst.
case_eq (f (Some v) (m2 k)); intros.
case_eq (decide (k = k)); congruence.
case_eq (m2 k); intros.
rewrite H3 in H2.
exfalso; eauto.
congruence.
case_eq (f (Some v) (m2 k)); intros.
destruct (decide (k0 = k)); congruence.
auto.
Qed.
Theorem empty_includes : forall A B (m : fmap A B), includes (empty (A := A) B) m.
Proof.
unfold includes, empty; intuition congruence.