Abstract domain for neededness analysis
Require Import Coqlib.
Require Import Maps.
Require Import IntvSets.
Require Import AST.
Require Import Zbits.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Memory.
Require Import Globalenvs.
Require Import Events.
Require Import Lattice.
Require Import Registers.
Require Import ValueDomain.
Require Import Op.
Require Import RTL.
Neededness for values
Inductive nval :
Type :=
|
Nothing (* value is entirely unused *)
|
I (
m:
int)
(* only need the bits that are 1 in m (32-bit ints) *)
|
L (
m:
int64)
(* only need the bits that are 1 in m (64 bit ints) *)
|
All.
(* every bit of the value is used *)
Definition eq_nval (
x y:
nval) : {
x=
y} + {
x<>
y}.
Proof.
Agreement between two values relative to a need.
Definition iagree (
p q mask:
int) :
Prop :=
forall i, 0 <=
i <
Int.zwordsize ->
Int.testbit mask i =
true ->
Int.testbit p i =
Int.testbit q i.
Definition lagree (
p q mask:
int64) :
Prop :=
forall i, 0 <=
i <
Int64.zwordsize ->
Int64.testbit mask i =
true ->
Int64.testbit p i =
Int64.testbit q i.
Definition vagree (
v w:
val) (
x:
nval) :
Prop :=
match x with
|
Nothing =>
True
|
I m =>
match v,
w with
|
Vint p,
Vint q =>
iagree p q m
|
Vint p, _ =>
False
| _, _ =>
True
end
|
L m =>
match v,
w with
|
Vlong p,
Vlong q =>
lagree p q m
|
Vlong p, _ =>
False
| _, _ =>
True
end
|
All =>
Val.lessdef v w
end.
Lemma vagree_same:
forall v x,
vagree v v x.
Proof.
intros. destruct x; simpl; auto; destruct v; auto; red; auto.
Qed.
Lemma vagree_lessdef:
forall v w x,
Val.lessdef v w ->
vagree v w x.
Proof.
intros.
inv H.
apply vagree_same.
destruct x;
simpl;
auto.
Qed.
Lemma lessdef_vagree:
forall v w,
vagree v w All ->
Val.lessdef v w.
Proof.
intros. simpl in H. auto.
Qed.
Global Hint Resolve vagree_same vagree_lessdef lessdef_vagree:
na.
Inductive vagree_list:
list val ->
list val ->
list nval ->
Prop :=
|
vagree_list_nil:
forall nvl,
vagree_list nil nil nvl
|
vagree_list_default:
forall v1 vl1 v2 vl2,
vagree v1 v2 All ->
vagree_list vl1 vl2 nil ->
vagree_list (
v1 ::
vl1) (
v2 ::
vl2)
nil
|
vagree_list_cons:
forall v1 vl1 v2 vl2 nv1 nvl,
vagree v1 v2 nv1 ->
vagree_list vl1 vl2 nvl ->
vagree_list (
v1 ::
vl1) (
v2 ::
vl2) (
nv1 ::
nvl).
Lemma lessdef_vagree_list:
forall vl1 vl2,
vagree_list vl1 vl2 nil ->
Val.lessdef_list vl1 vl2.
Proof.
induction vl1; intros; inv H; constructor; auto with na.
Qed.
Lemma vagree_lessdef_list:
forall vl1 vl2,
Val.lessdef_list vl1 vl2 ->
forall nvl,
vagree_list vl1 vl2 nvl.
Proof.
induction 1; intros.
constructor.
destruct nvl; constructor; auto with na.
Qed.
Global Hint Resolve lessdef_vagree_list vagree_lessdef_list:
na.
Ordering and least upper bound between value needs
Inductive nge:
nval ->
nval ->
Prop :=
|
nge_nothing:
forall x,
nge All x
|
nge_all:
forall x,
nge x Nothing
|
nge_int:
forall m1 m2,
(
forall i, 0 <=
i <
Int.zwordsize ->
Int.testbit m2 i =
true ->
Int.testbit m1 i =
true) ->
nge (
I m1) (
I m2)
|
nge_long:
forall m1 m2,
(
forall i, 0 <=
i <
Int64.zwordsize ->
Int64.testbit m2 i =
true ->
Int64.testbit m1 i =
true) ->
nge (
L m1) (
L m2).
Lemma nge_refl:
forall x,
nge x x.
Proof.
destruct x; constructor; auto.
Qed.
Global Hint Constructors nge:
na.
Global Hint Resolve nge_refl:
na.
Lemma nge_trans:
forall x y,
nge x y ->
forall z,
nge y z ->
nge x z.
Proof.
induction 1; intros w VG; inv VG; eauto with na.
Qed.
Lemma nge_agree:
forall v w x y,
nge x y ->
vagree v w x ->
vagree v w y.
Proof.
induction 1; simpl; auto.
- destruct v; auto with na.
- destruct v, w; intuition. red; auto.
- destruct v, w; intuition. red; auto.
Qed.
Definition nlub (
x y:
nval) :
nval :=
match x,
y with
|
Nothing, _ =>
y
| _,
Nothing =>
x
|
I m1,
I m2 =>
I (
Int.or m1 m2)
|
L m1,
L m2 =>
L (
Int64.or m1 m2)
| _, _ =>
All
end.
Lemma nge_lub_l:
forall x y,
nge (
nlub x y)
x.
Proof.
unfold nlub;
destruct x,
y;
auto with na.
all:
constructor;
intros;
autorewrite with ints;
auto.
all:
rewrite H0;
auto.
Qed.
Lemma nge_lub_r:
forall x y,
nge (
nlub x y)
y.
Proof.
unfold nlub;
destruct x,
y;
auto with na.
all:
constructor;
intros;
autorewrite with ints;
auto.
all:
rewrite H0;
apply orb_true_r.
Qed.
Properties of agreement between 32-bit integers
Lemma iagree_refl:
forall p m,
iagree p p m.
Proof.
intros; red; auto.
Qed.
Remark eq_same_bits:
forall i x y,
x =
y ->
Int.testbit x i =
Int.testbit y i.
Proof.
intros; congruence.
Qed.
Lemma iagree_and_eq:
forall x y mask,
iagree x y mask <->
Int.and x mask =
Int.and y mask.
Proof.
Lemma iagree_mone:
forall p q,
iagree p q Int.mone ->
p =
q.
Proof.
Lemma iagree_zero:
forall p q,
iagree p q Int.zero.
Proof.
Lemma iagree_and:
forall x y n m,
iagree x y (
Int.and m n) ->
iagree (
Int.and x n) (
Int.and y n)
m.
Proof.
Lemma iagree_not:
forall x y m,
iagree x y m ->
iagree (
Int.not x) (
Int.not y)
m.
Proof.
intros; red; intros; autorewrite with ints; auto. f_equal; auto.
Qed.
Lemma iagree_not':
forall x y m,
iagree (
Int.not x) (
Int.not y)
m ->
iagree x y m.
Proof.
Lemma iagree_or:
forall x y n m,
iagree x y (
Int.and m (
Int.not n)) ->
iagree (
Int.or x n) (
Int.or y n)
m.
Proof.
Lemma iagree_bitwise_binop:
forall sem f,
(
forall x y i, 0 <=
i <
Int.zwordsize ->
Int.testbit (
f x y)
i =
sem (
Int.testbit x i) (
Int.testbit y i)) ->
forall x1 x2 y1 y2 m,
iagree x1 y1 m ->
iagree x2 y2 m ->
iagree (
f x1 x2) (
f y1 y2)
m.
Proof.
intros; red; intros. rewrite ! H by auto. f_equal; auto.
Qed.
Lemma iagree_shl:
forall x y m n,
iagree x y (
Int.shru m n) ->
iagree (
Int.shl x n) (
Int.shl y n)
m.
Proof.
Lemma iagree_shru:
forall x y m n,
iagree x y (
Int.shl m n) ->
iagree (
Int.shru x n) (
Int.shru y n)
m.
Proof.
Lemma iagree_shr_1:
forall x y m n,
Int.shru (
Int.shl m n)
n =
m ->
iagree x y (
Int.shl m n) ->
iagree (
Int.shr x n) (
Int.shr y n)
m.
Proof.
Lemma iagree_shr:
forall x y m n,
iagree x y (
Int.or (
Int.shl m n) (
Int.repr Int.min_signed)) ->
iagree (
Int.shr x n) (
Int.shr y n)
m.
Proof.
Lemma iagree_rol:
forall p q m amount,
iagree p q (
Int.ror m amount) ->
iagree (
Int.rol p amount) (
Int.rol q amount)
m.
Proof.
Lemma iagree_ror:
forall p q m amount,
iagree p q (
Int.rol m amount) ->
iagree (
Int.ror p amount) (
Int.ror q amount)
m.
Proof.
Lemma eqmod_iagree:
forall m x y,
eqmod (
two_p (
Int.size m))
x y ->
iagree (
Int.repr x) (
Int.repr y)
m.
Proof.
Definition complete_mask (
m:
int) :=
Int.zero_ext (
Int.size m)
Int.mone.
Lemma iagree_eqmod:
forall x y m,
iagree x y (
complete_mask m) ->
eqmod (
two_p (
Int.size m)) (
Int.unsigned x) (
Int.unsigned y).
Proof.
Lemma complete_mask_idem:
forall m,
complete_mask (
complete_mask m) =
complete_mask m.
Proof.
Properties of agreement between 64-bit integers
Lemma lagree_refl:
forall p m,
lagree p p m.
Proof.
intros; red; auto.
Qed.
Remark eq_same_bits64:
forall i x y,
x =
y ->
Int64.testbit x i =
Int64.testbit y i.
Proof.
intros; congruence.
Qed.
Lemma lagree_and_eq:
forall x y mask,
lagree x y mask <->
Int64.and x mask =
Int64.and y mask.
Proof.
Lemma lagree_mone:
forall p q,
lagree p q Int64.mone ->
p =
q.
Proof.
Lemma lagree_zero:
forall p q,
lagree p q Int64.zero.
Proof.
Lemma lagree_and:
forall x y n m,
lagree x y (
Int64.and m n) ->
lagree (
Int64.and x n) (
Int64.and y n)
m.
Proof.
Lemma lagree_not:
forall x y m,
lagree x y m ->
lagree (
Int64.not x) (
Int64.not y)
m.
Proof.
intros; red; intros; autorewrite with ints; auto. f_equal; auto.
Qed.
Lemma lagree_not':
forall x y m,
lagree (
Int64.not x) (
Int64.not y)
m ->
lagree x y m.
Proof.
Lemma lagree_or:
forall x y n m,
lagree x y (
Int64.and m (
Int64.not n)) ->
lagree (
Int64.or x n) (
Int64.or y n)
m.
Proof.
Lemma lagree_bitwise_binop:
forall sem f,
(
forall x y i, 0 <=
i <
Int64.zwordsize ->
Int64.testbit (
f x y)
i =
sem (
Int64.testbit x i) (
Int64.testbit y i)) ->
forall x1 x2 y1 y2 m,
lagree x1 y1 m ->
lagree x2 y2 m ->
lagree (
f x1 x2) (
f y1 y2)
m.
Proof.
intros; red; intros. rewrite ! H by auto. f_equal; auto.
Qed.
Lemma lagree_shl:
forall x y m n,
lagree x y (
Int64.shru' m n) ->
lagree (
Int64.shl' x n) (
Int64.shl' y n)
m.
Proof.
Lemma lagree_shru:
forall x y m n,
lagree x y (
Int64.shl' m n) ->
lagree (
Int64.shru' x n) (
Int64.shru' y n)
m.
Proof.
Lemma lagree_shr_1:
forall x y m n,
Int64.shru' (
Int64.shl' m n)
n =
m ->
lagree x y (
Int64.shl' m n) ->
lagree (
Int64.shr' x n) (
Int64.shr' y n)
m.
Proof.
intros;
red;
intros.
rewrite <-
H in H2.
rewrite Int64.bits_shru' in H2 by auto.
rewrite !
Int64.bits_shr' by auto.
destruct (
zlt (
i +
Int.unsigned n)
Int64.zwordsize).
-
apply H0;
auto.
generalize (
Int.unsigned_range n);
lia.
-
discriminate.
Qed.
Lemma lagree_shr:
forall x y m n,
lagree x y (
Int64.or (
Int64.shl' m n) (
Int64.repr Int64.min_signed)) ->
lagree (
Int64.shr' x n) (
Int64.shr' y n)
m.
Proof.
Lemma lagree_rol':
forall p q m amount,
lagree p q (
Int64.ror m (
Int64.repr (
Int.unsigned amount))) ->
lagree (
Int64.rol' p amount) (
Int64.rol' q amount)
m.
Proof.
Lemma lagree_rol:
forall p q m amount,
lagree p q (
Int64.ror m amount) ->
lagree (
Int64.rol p amount) (
Int64.rol q amount)
m.
Proof.
Lemma lagree_ror:
forall p q m amount,
lagree p q (
Int64.rol' m amount) ->
lagree (
Int64.ror' p amount) (
Int64.ror' q amount)
m.
Proof.
Lemma eqmod_lagree:
forall m x y,
eqmod (
two_p (
Int64.size m))
x y ->
lagree (
Int64.repr x) (
Int64.repr y)
m.
Proof.
Definition complete_mask64 (
m:
int64) :=
Int64.zero_ext (
Int64.size m)
Int64.mone.
Lemma lagree_eqmod:
forall x y m,
lagree x y (
complete_mask64 m) ->
eqmod (
two_p (
Int64.size m)) (
Int64.unsigned x) (
Int64.unsigned y).
Proof.
Lemma complete_mask64_idem:
forall m,
complete_mask64 (
complete_mask64 m) =
complete_mask64 m.
Proof.
Abstract operations over value needs.
Ltac InvAgree :=
simpl vagree in *;
repeat (
auto ||
exact Logic.I ||
match goal with
| [
H:
False |- _ ] =>
contradiction
| [
H:
match ?
v with Vundef => _ |
Vint _ => _ |
Vlong _ => _ |
Vfloat _ => _ |
Vsingle _ => _ |
Vptr _ _ => _
end |- _ ] =>
destruct v
| [ |-
context [
if Archi.ptr64 then _
else _] ] =>
destruct Archi.ptr64 eqn:?
end).
And immediate, or immediate
Definition andimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.and m n)
| _ =>
I n
end.
Lemma andimm_sound:
forall v w x n,
vagree v w (
andimm x n) ->
vagree (
Val.and v (
Vint n)) (
Val.and w (
Vint n))
x.
Proof.
unfold andimm;
intros;
destruct x;
simpl in *;
unfold Val.and.
-
auto.
-
InvAgree.
apply iagree_and;
auto.
-
InvAgree.
-
InvAgree.
rewrite iagree_and_eq in H.
rewrite H;
auto.
Qed.
Definition orimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.and m (
Int.not n))
| _ =>
I (
Int.not n)
end.
Lemma orimm_sound:
forall v w x n,
vagree v w (
orimm x n) ->
vagree (
Val.or v (
Vint n)) (
Val.or w (
Vint n))
x.
Proof.
64 bit and immediate, or immediate
Definition andlimm (
x:
nval) (
n:
int64) :=
match x with
|
Nothing =>
Nothing
|
L m =>
L (
Int64.and m n)
| _ =>
L n
end.
Lemma andlimm_sound:
forall v w x n,
vagree v w (
andlimm x n) ->
vagree (
Val.andl v (
Vlong n)) (
Val.andl w (
Vlong n))
x.
Proof.
Definition orlimm (
x:
nval) (
n:
int64) :=
match x with
|
Nothing =>
Nothing
|
L m =>
L (
Int64.and m (
Int64.not n))
| _ =>
L (
Int64.not n)
end.
Lemma orlimm_sound:
forall v w x n,
vagree v w (
orlimm x n) ->
vagree (
Val.orl v (
Vlong n)) (
Val.orl w (
Vlong n))
x.
Proof.
Bitwise operations: and, or, xor, not
Definition bitwise (
x:
nval) :=
x.
Remark bitwise_idem:
forall nv,
bitwise (
bitwise nv) =
bitwise nv.
Proof.
auto. Qed.
Lemma vagree_bitwise_binop:
forall f,
(
forall p1 p2 q1 q2 m,
iagree p1 q1 m ->
iagree p2 q2 m ->
iagree (
f p1 p2) (
f q1 q2)
m) ->
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
match v1,
v2 with Vint i1,
Vint i2 =>
Vint(
f i1 i2) | _, _ =>
Vundef end)
(
match w1,
w2 with Vint i1,
Vint i2 =>
Vint(
f i1 i2) | _, _ =>
Vundef end)
x.
Proof.
unfold bitwise;
intros.
destruct x;
simpl in *.
-
auto.
-
InvAgree.
-
InvAgree.
-
inv H0;
auto.
inv H1;
auto.
destruct w1;
auto.
Qed.
Lemma and_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
Val.and v1 v2) (
Val.and w1 w2)
x.
Proof (
vagree_bitwise_binop Int.and (
iagree_bitwise_binop andb Int.and Int.bits_and)).
Lemma or_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
Val.or v1 v2) (
Val.or w1 w2)
x.
Proof (
vagree_bitwise_binop Int.or (
iagree_bitwise_binop orb Int.or Int.bits_or)).
Lemma xor_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
Val.xor v1 v2) (
Val.xor w1 w2)
x.
Proof (
vagree_bitwise_binop Int.xor (
iagree_bitwise_binop xorb Int.xor Int.bits_xor)).
Lemma notint_sound:
forall v w x,
vagree v w (
bitwise x) ->
vagree (
Val.notint v) (
Val.notint w)
x.
Proof.
Bitwise operations: andl, orl, xorl, notl
Lemma vagree_bitwise64_binop:
forall f,
(
forall p1 p2 q1 q2 m,
lagree p1 q1 m ->
lagree p2 q2 m ->
lagree (
f p1 p2) (
f q1 q2)
m) ->
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
match v1,
v2 with Vlong i1,
Vlong i2 =>
Vlong(
f i1 i2) | _, _ =>
Vundef end)
(
match w1,
w2 with Vlong i1,
Vlong i2 =>
Vlong(
f i1 i2) | _, _ =>
Vundef end)
x.
Proof.
unfold bitwise;
intros.
destruct x;
simpl in *.
-
auto.
-
InvAgree.
-
InvAgree.
-
inv H0;
auto.
inv H1;
auto.
destruct w1;
auto.
Qed.
Lemma andl_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
Val.andl v1 v2) (
Val.andl w1 w2)
x.
Proof (
vagree_bitwise64_binop Int64.and (
lagree_bitwise_binop andb Int64.and Int64.bits_and)).
Lemma orl_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
Val.orl v1 v2) (
Val.orl w1 w2)
x.
Proof (
vagree_bitwise64_binop Int64.or (
lagree_bitwise_binop orb Int64.or Int64.bits_or)).
Lemma xorl_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
bitwise x) ->
vagree v2 w2 (
bitwise x) ->
vagree (
Val.xorl v1 v2) (
Val.xorl w1 w2)
x.
Proof (
vagree_bitwise64_binop Int64.xor (
lagree_bitwise_binop xorb Int64.xor Int64.bits_xor)).
Lemma notl_sound:
forall v w x,
vagree v w (
bitwise x) ->
vagree (
Val.notl v) (
Val.notl w)
x.
Proof.
Shifts and rotates
Definition shlimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.shru m n)
| _ =>
I (
Int.shru Int.mone n)
end.
Lemma shlimm_sound:
forall v w x n,
vagree v w (
shlimm x n) ->
vagree (
Val.shl v (
Vint n)) (
Val.shl w (
Vint n))
x.
Proof.
Definition shruimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.shl m n)
| _ =>
I (
Int.shl Int.mone n)
end.
Lemma shruimm_sound:
forall v w x n,
vagree v w (
shruimm x n) ->
vagree (
Val.shru v (
Vint n)) (
Val.shru w (
Vint n))
x.
Proof.
Definition shrimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
let m' :=
Int.shl m n in
if Int.eq_dec (
Int.shru m' n)
m
then m'
else Int.or m' (
Int.repr Int.min_signed))
| _ =>
I (
Int.or (
Int.shl Int.mone n) (
Int.repr Int.min_signed))
end.
Lemma shrimm_sound:
forall v w x n,
vagree v w (
shrimm x n) ->
vagree (
Val.shr v (
Vint n)) (
Val.shr w (
Vint n))
x.
Proof.
Definition rol (
x:
nval) (
amount:
int) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.ror m amount)
| _ =>
All
end.
Lemma rol_sound:
forall v w x n,
vagree v w (
rol x n) ->
vagree (
Val.rol v (
Vint n)) (
Val.rol w (
Vint n))
x.
Proof.
unfold rol,
Val.rol;
intros.
destruct x;
simpl in *.
-
auto.
-
InvAgree.
apply iagree_rol;
auto.
-
destruct v;
auto.
-
inv H;
auto.
Qed.
Definition ror (
x:
nval) (
amount:
int) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.rol m amount)
| _ =>
All
end.
Lemma ror_sound:
forall v w x n,
vagree v w (
ror x n) ->
vagree (
Val.ror v (
Vint n)) (
Val.ror w (
Vint n))
x.
Proof.
unfold ror,
Val.ror;
intros.
destruct x;
simpl in *.
-
auto.
-
InvAgree.
apply iagree_ror;
auto.
-
destruct v;
auto.
-
inv H;
auto.
Qed.
Definition rolm (
x:
nval) (
amount mask:
int) :=
rol (
andimm x mask)
amount.
Lemma rolm_sound:
forall v w x amount mask,
vagree v w (
rolm x amount mask) ->
vagree (
Val.rolm v amount mask) (
Val.rolm w amount mask)
x.
Proof.
64 bit shifts and rotates
Definition shllimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
L m =>
L (
Int64.shru' m n)
| _ =>
L (
Int64.shru' Int64.mone n)
end.
Lemma shllimm_sound:
forall v w x n,
vagree v w (
shllimm x n) ->
vagree (
Val.shll v (
Vint n)) (
Val.shll w (
Vint n))
x.
Proof.
Definition shrluimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
L m =>
L (
Int64.shl' m n)
| _ =>
L (
Int64.shl' Int64.mone n)
end.
Lemma shrluimm_sound:
forall v w x n,
vagree v w (
shrluimm x n) ->
vagree (
Val.shrlu v (
Vint n)) (
Val.shrlu w (
Vint n))
x.
Proof.
Definition shrlimm (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
Nothing
|
L m =>
L (
let m' :=
Int64.shl' m n in
if Int64.eq_dec (
Int64.shru' m' n)
m
then m'
else Int64.or m' (
Int64.repr Int64.min_signed))
| _ =>
L (
Int64.or (
Int64.shl' Int64.mone n) (
Int64.repr Int64.min_signed))
end.
Lemma shrlimm_sound:
forall v w x n,
vagree v w (
shrlimm x n) ->
vagree (
Val.shrl v (
Vint n)) (
Val.shrl w (
Vint n))
x.
Proof.
Definition roll (
x:
nval) (
amount:
int) :=
match x with
|
Nothing =>
Nothing
|
L m =>
L (
Int64.ror m (
Int64.repr (
Int.unsigned amount)))
| _ =>
All
end.
Lemma roll_sound:
forall v w x n,
vagree v w (
roll x n) ->
vagree (
Val.roll v (
Vint n)) (
Val.roll w (
Vint n))
x.
Proof.
unfold roll,
Val.roll;
intros.
destruct x;
simpl in *.
-
auto.
-
destruct v;
auto.
-
InvAgree.
apply lagree_rol';
auto.
-
inv H;
auto.
Qed.
Definition rorl (
x:
nval) (
amount:
int) :=
match x with
|
Nothing =>
Nothing
|
L m =>
L (
Int64.rol' m amount)
| _ =>
All
end.
Lemma rorl_sound:
forall v w x n,
vagree v w (
rorl x n) ->
vagree (
Val.rorl v (
Vint n)) (
Val.rorl w (
Vint n))
x.
Proof.
unfold ror,
Val.ror;
intros.
destruct x;
simpl in *.
-
auto.
-
destruct v;
simpl;
auto.
-
InvAgree.
apply lagree_ror;
auto.
-
inv H;
auto.
Qed.
Definition rolml (
x:
nval) (
amount :
int) (
mask:
int64) :=
roll (
andlimm x mask)
amount.
Lemma rolml_sound:
forall v w x amount mask,
vagree v w (
rolml x amount mask) ->
vagree (
Val.rolml v amount mask) (
Val.rolml w amount mask)
x.
Proof.
Modular arithmetic operations: add, mul, opposite.
Also subtraction, but only on 64-bit targets, otherwise
the pointer - pointer case does not fit.
Definition modarith (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
complete_mask m)
|
L m =>
L (
complete_mask64 m)
|
All =>
All
end.
Lemma add_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
modarith x) ->
vagree v2 w2 (
modarith x) ->
vagree (
Val.add v1 v2) (
Val.add w1 w2)
x.
Proof.
Lemma sub_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
modarith x) ->
vagree v2 w2 (
modarith x) ->
Archi.ptr64 =
true ->
vagree (
Val.sub v1 v2) (
Val.sub w1 w2)
x.
Proof.
Remark modarith_idem:
forall nv,
modarith (
modarith nv) =
modarith nv.
Proof.
Lemma mul_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
modarith x) ->
vagree v2 w2 (
modarith x) ->
vagree (
Val.mul v1 v2) (
Val.mul w1 w2)
x.
Proof.
Lemma neg_sound:
forall v w x,
vagree v w (
modarith x) ->
vagree (
Val.neg v) (
Val.neg w)
x.
Proof.
Lemma addl_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
modarith x) ->
vagree v2 w2 (
modarith x) ->
vagree (
Val.addl v1 v2) (
Val.addl w1 w2)
x.
Proof.
Lemma subl_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
modarith x) ->
vagree v2 w2 (
modarith x) ->
Archi.ptr64 =
false ->
vagree (
Val.subl v1 v2) (
Val.subl w1 w2)
x.
Proof.
unfold modarith;
intros.
destruct x;
simpl in *.
-
auto.
-
InvAgree;
simpl;
rewrite H1;
auto.
-
unfold Val.subl;
rewrite H1;
InvAgree.
apply eqmod_lagree.
apply eqmod_sub;
apply lagree_eqmod;
auto.
-
inv H;
auto.
inv H0;
auto.
destruct w1;
auto.
Qed.
Lemma mull_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
modarith x) ->
vagree v2 w2 (
modarith x) ->
vagree (
Val.mull v1 v2) (
Val.mull w1 w2)
x.
Proof.
Lemma negl_sound:
forall v w x,
vagree v w (
modarith x) ->
vagree (
Val.negl v) (
Val.negl w)
x.
Proof.
Conversions: zero extension, sign extension, 32/64 bit conversions
Definition zero_ext (
n:
Z) (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.zero_ext n m)
| _ =>
I (
Int.zero_ext n Int.mone)
end.
Lemma zero_ext_sound:
forall v w x n,
vagree v w (
zero_ext n x) ->
0 <=
n ->
vagree (
Val.zero_ext n v) (
Val.zero_ext n w)
x.
Proof.
unfold zero_ext;
intros.
destruct x;
simpl in *.
-
auto.
-
unfold Val.zero_ext;
InvAgree.
red;
intros.
autorewrite with ints;
try lia.
destruct (
zlt i1 n);
auto.
apply H;
auto.
autorewrite with ints;
try lia.
rewrite zlt_true;
auto.
-
InvAgree.
-
unfold Val.zero_ext;
InvAgree;
auto.
apply Val.lessdef_same.
f_equal.
Int.bit_solve;
try lia.
destruct (
zlt i1 n);
auto.
apply H;
auto.
autorewrite with ints;
try lia.
apply zlt_true;
auto.
Qed.
Definition sign_ext (
n:
Z) (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
I m =>
I (
Int.or (
Int.zero_ext n m) (
Int.shl Int.one (
Int.repr (
n - 1))))
| _ =>
I (
Int.zero_ext n Int.mone)
end.
Lemma sign_ext_sound:
forall v w x n,
vagree v w (
sign_ext n x) ->
0 <
n ->
vagree (
Val.sign_ext n v) (
Val.sign_ext n w)
x.
Proof.
Definition loword (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
I m =>
L (
Int64.ofwords Int.zero m)
| _ =>
L (
Int64.ofwords Int.zero Int.mone)
end.
Lemma loword_sound:
forall v w x,
vagree v w (
loword x) ->
vagree (
Val.loword v) (
Val.loword w)
x.
Proof.
Definition hiword (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
I m =>
L (
Int64.ofwords m Int.zero)
| _ =>
L (
Int64.ofwords Int.mone Int.zero)
end.
Lemma hiword_sound:
forall v w x,
vagree v w (
hiword x) ->
vagree (
Val.hiword v) (
Val.hiword w)
x.
Proof.
Definition makelong_lo (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
L m =>
I (
Int64.loword m)
| _ =>
All
end.
Definition makelong_hi (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
L m =>
I (
Int64.hiword m)
| _ =>
All
end.
Lemma makelong_sound:
forall v1 w1 v2 w2 x,
vagree v1 w1 (
makelong_hi x) ->
vagree v2 w2 (
makelong_lo x) ->
vagree (
Val.longofwords v1 v2) (
Val.longofwords w1 w2)
x.
Proof.
Definition longofintu (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
L m =>
I (
Int64.loword m)
| _ =>
All
end.
Lemma longofintu_sound:
forall v w x,
vagree v w (
longofintu x) ->
vagree (
Val.longofintu v) (
Val.longofintu w)
x.
Proof.
Definition longofint (
x:
nval) :=
match x with
|
Nothing =>
Nothing
|
L m =>
I (
Int.or (
Int64.loword m) (
Int.shl Int.one (
Int.repr 31)))
| _ =>
All
end.
Lemma longofint_sound:
forall v w x,
vagree v w (
longofint x) ->
vagree (
Val.longofint v) (
Val.longofint w)
x.
Proof.
The needs of a memory store concerning the value being stored.
Definition store_argument (
chunk:
memory_chunk) :=
match chunk with
|
Mbool |
Mint8signed |
Mint8unsigned =>
I (
Int.repr 255)
|
Mint16signed |
Mint16unsigned =>
I (
Int.repr 65535)
| _ =>
All
end.
Lemma store_argument_sound:
forall chunk v w,
vagree v w (
store_argument chunk) ->
list_forall2 memval_lessdef (
encode_val chunk v) (
encode_val chunk w).
Proof.
Lemma store_argument_load_result:
forall chunk v w,
vagree v w (
store_argument chunk) ->
Val.lessdef (
Val.load_result chunk v) (
Val.load_result chunk w).
Proof.
The needs of a comparison
Definition maskzero (
n:
int) :=
I n.
Lemma maskzero_sound:
forall v w n b,
vagree v w (
maskzero n) ->
Val.maskzero_bool v n =
Some b ->
Val.maskzero_bool w n =
Some b.
Proof.
The needs of a select
Lemma normalize_sound:
forall v w x ty,
vagree v w x ->
vagree (
Val.normalize v ty) (
Val.normalize w ty)
x.
Proof.
intros.
destruct x;
simpl in *.
-
auto.
-
unfold Val.normalize.
destruct v.
auto.
destruct w;
try contradiction.
destruct ty;
auto.
destruct ty;
auto.
destruct ty;
auto.
destruct ty;
auto.
destruct ty;
destruct Archi.ptr64;
auto.
-
unfold Val.normalize.
destruct v;
auto.
destruct ty;
auto.
destruct w;
try contradiction.
destruct ty;
auto.
destruct ty;
auto.
destruct ty;
auto.
destruct ty;
destruct Archi.ptr64;
auto.
-
apply Val.normalize_lessdef;
auto.
Qed.
Lemma select_sound:
forall ob v1 v2 w1 w2 ty x,
vagree v1 w1 x ->
vagree v2 w2 x ->
vagree (
Val.select ob v1 v2 ty) (
Val.select ob w1 w2 ty)
x.
Proof.
The default abstraction: if the result is unused, the arguments are
unused; otherwise, the arguments are needed in full.
Definition default (
x:
nval) :=
match x with
|
Nothing =>
Nothing
| _ =>
All
end.
Section DEFAULT.
Variable ge:
genv.
Variable sp:
block.
Variables m1 m2:
mem.
Hypothesis PERM:
forall b ofs k p,
Mem.perm m1 b ofs k p ->
Mem.perm m2 b ofs k p.
Let valid_pointer_inj:
forall b1 ofs b2 delta,
inject_id b1 =
Some(
b2,
delta) ->
Mem.valid_pointer m1 b1 (
Ptrofs.unsigned ofs) =
true ->
Mem.valid_pointer m2 b2 (
Ptrofs.unsigned (
Ptrofs.add ofs (
Ptrofs.repr delta))) =
true.
Proof.
Let weak_valid_pointer_inj:
forall b1 ofs b2 delta,
inject_id b1 =
Some(
b2,
delta) ->
Mem.weak_valid_pointer m1 b1 (
Ptrofs.unsigned ofs) =
true ->
Mem.weak_valid_pointer m2 b2 (
Ptrofs.unsigned (
Ptrofs.add ofs (
Ptrofs.repr delta))) =
true.
Proof.
Let weak_valid_pointer_no_overflow:
forall b1 ofs b2 delta,
inject_id b1 =
Some(
b2,
delta) ->
Mem.weak_valid_pointer m1 b1 (
Ptrofs.unsigned ofs) =
true ->
0 <=
Ptrofs.unsigned ofs +
Ptrofs.unsigned (
Ptrofs.repr delta) <=
Ptrofs.max_unsigned.
Proof.
Let valid_different_pointers_inj:
forall b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
b1 <>
b2 ->
Mem.valid_pointer m1 b1 (
Ptrofs.unsigned ofs1) =
true ->
Mem.valid_pointer m1 b2 (
Ptrofs.unsigned ofs2) =
true ->
inject_id b1 =
Some (
b1',
delta1) ->
inject_id b2 =
Some (
b2',
delta2) ->
b1' <>
b2' \/
Ptrofs.unsigned (
Ptrofs.add ofs1 (
Ptrofs.repr delta1)) <>
Ptrofs.unsigned (
Ptrofs.add ofs2 (
Ptrofs.repr delta2)).
Proof.
unfold inject_id;
intros.
left;
congruence.
Defined.
Lemma default_needs_of_condition_sound:
forall cond args1 b args2,
eval_condition cond args1 m1 =
Some b ->
vagree_list args1 args2 nil ->
eval_condition cond args2 m2 =
Some b.
Proof.
Lemma default_needs_of_operation_sound:
forall op args1 v1 args2 nv,
eval_operation ge (
Vptr sp Ptrofs.zero)
op args1 m1 =
Some v1 ->
vagree_list args1 args2 nil
\/
vagree_list args1 args2 (
default nv ::
nil)
\/
vagree_list args1 args2 (
default nv ::
default nv ::
nil)
\/
vagree_list args1 args2 (
default nv ::
default nv ::
default nv ::
nil) ->
nv <>
Nothing ->
exists v2,
eval_operation ge (
Vptr sp Ptrofs.zero)
op args2 m2 =
Some v2
/\
vagree v1 v2 nv.
Proof.
End DEFAULT.
Detecting operations that are redundant and can be turned into a move
Definition andimm_redundant (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
true
|
I m =>
Int.eq_dec (
Int.and m (
Int.not n))
Int.zero
| _ =>
false
end.
Lemma andimm_redundant_sound:
forall v w x n,
andimm_redundant x n =
true ->
vagree v w (
andimm x n) ->
vagree (
Val.and v (
Vint n))
w x.
Proof.
unfold andimm_redundant;
intros.
destruct x;
try discriminate.
-
simpl;
auto.
-
InvBooleans.
simpl in *;
unfold Val.and;
InvAgree.
red;
intros.
exploit (
eq_same_bits i1);
eauto.
autorewrite with ints;
auto.
rewrite H2;
simpl;
intros.
destruct (
Int.testbit n i1)
eqn:
N;
try discriminate.
rewrite andb_true_r.
apply H0;
auto.
autorewrite with ints;
auto.
rewrite H2,
N;
auto.
Qed.
Definition orimm_redundant (
x:
nval) (
n:
int) :=
match x with
|
Nothing =>
true
|
I m =>
Int.eq_dec (
Int.and m n)
Int.zero
| _ =>
false
end.
Lemma orimm_redundant_sound:
forall v w x n,
orimm_redundant x n =
true ->
vagree v w (
orimm x n) ->
vagree (
Val.or v (
Vint n))
w x.
Proof.
Definition rolm_redundant (
x:
nval) (
amount mask:
int) :=
Int.eq_dec amount Int.zero &&
andimm_redundant x mask.
Lemma rolm_redundant_sound:
forall v w x amount mask,
rolm_redundant x amount mask =
true ->
vagree v w (
rolm x amount mask) ->
vagree (
Val.rolm v amount mask)
w x.
Proof.
Definition zero_ext_redundant (
n:
Z) (
x:
nval) :=
match x with
|
Nothing =>
true
|
I m =>
Int.eq_dec (
Int.zero_ext n m)
m
| _ =>
false
end.
Lemma zero_ext_redundant_sound:
forall v w x n,
zero_ext_redundant n x =
true ->
vagree v w (
zero_ext n x) ->
0 <=
n ->
vagree (
Val.zero_ext n v)
w x.
Proof.
unfold zero_ext_redundant;
intros.
destruct x;
try discriminate.
-
auto.
-
simpl in *;
InvAgree.
simpl.
InvBooleans.
rewrite <-
H.
red;
intros;
autorewrite with ints;
try lia.
destruct (
zlt i1 n).
apply H0;
auto.
rewrite Int.bits_zero_ext in H3 by lia.
rewrite zlt_false in H3 by auto.
discriminate.
Qed.
Definition sign_ext_redundant (
n:
Z) (
x:
nval) :=
match x with
|
Nothing =>
true
|
I m =>
Int.eq_dec (
Int.zero_ext n m)
m
| _ =>
false
end.
Lemma sign_ext_redundant_sound:
forall v w x n,
sign_ext_redundant n x =
true ->
vagree v w (
sign_ext n x) ->
0 <
n ->
vagree (
Val.sign_ext n v)
w x.
Proof.
unfold sign_ext_redundant;
intros.
destruct x;
try discriminate.
-
auto.
-
simpl in *;
InvAgree.
simpl.
InvBooleans.
rewrite <-
H.
red;
intros;
autorewrite with ints;
try lia.
destruct (
zlt i1 n).
apply H0;
auto.
rewrite Int.bits_or;
auto.
rewrite H3;
auto.
rewrite Int.bits_zero_ext in H3 by lia.
rewrite zlt_false in H3 by auto.
discriminate.
Qed.
Definition andlimm_redundant (
x:
nval) (
n:
int64) :=
match x with
|
Nothing =>
true
|
L m =>
Int64.eq_dec (
Int64.and m (
Int64.not n))
Int64.zero
| _ =>
false
end.
Lemma andlimm_redundant_sound:
forall v w x n,
andlimm_redundant x n =
true ->
vagree v w (
andlimm x n) ->
vagree (
Val.andl v (
Vlong n))
w x.
Proof.
unfold andlimm_redundant;
intros.
destruct x;
try discriminate.
-
simpl;
auto.
-
InvBooleans.
simpl in *;
unfold Val.andl;
InvAgree.
red;
intros.
exploit (
eq_same_bits64 i1);
eauto.
autorewrite with ints;
auto.
rewrite H2;
simpl;
intros.
destruct (
Int64.testbit n i1)
eqn:
N;
try discriminate.
rewrite andb_true_r.
apply H0;
auto.
autorewrite with ints;
auto.
rewrite H2,
N;
auto.
Qed.
Definition orlimm_redundant (
x:
nval) (
n:
int64) :=
match x with
|
Nothing =>
true
|
L m =>
Int64.eq_dec (
Int64.and m n)
Int64.zero
| _ =>
false
end.
Lemma orlimm_redundant_sound:
forall v w x n,
orlimm_redundant x n =
true ->
vagree v w (
orlimm x n) ->
vagree (
Val.orl v (
Vlong n))
w x.
Proof.
Definition rolml_redundant (
x:
nval) (
amount :
int) (
mask:
int64) :=
Int.eq_dec amount Int.zero &&
andlimm_redundant x mask.
Lemma rolml_redundant_sound:
forall v w x amount mask,
rolml_redundant x amount mask =
true ->
vagree v w (
rolml x amount mask) ->
vagree (
Val.rolml v amount mask)
w x.
Proof.
Neededness for register environments
Module NVal <:
SEMILATTICE.
Definition t :=
nval.
Definition eq (
x y:
t) := (
x =
y).
Definition eq_refl:
forall x,
eq x x := (@
eq_refl t).
Definition eq_sym:
forall x y,
eq x y ->
eq y x := (@
eq_sym t).
Definition eq_trans:
forall x y z,
eq x y ->
eq y z ->
eq x z := (@
eq_trans t).
Definition beq (
x y:
t) :
bool :=
proj_sumbool (
eq_nval x y).
Lemma beq_correct:
forall x y,
beq x y =
true ->
eq x y.
Proof.
unfold beq;
intros.
InvBooleans.
auto. Qed.
Definition ge :=
nge.
Lemma ge_refl:
forall x y,
eq x y ->
ge x y.
Proof.
Lemma ge_trans:
forall x y z,
ge x y ->
ge y z ->
ge x z.
Proof.
Definition bot :
t :=
Nothing.
Lemma ge_bot:
forall x,
ge x bot.
Proof.
intros. constructor. Qed.
Definition lub :=
nlub.
Lemma ge_lub_left:
forall x y,
ge (
lub x y)
x.
Proof nge_lub_l.
Lemma ge_lub_right:
forall x y,
ge (
lub x y)
y.
Proof nge_lub_r.
End NVal.
Module NE :=
LPMap1(
NVal).
Definition nenv :=
NE.t.
Definition nreg (
ne:
nenv) (
r:
reg) :=
NE.get r ne.
Definition eagree (
e1 e2:
regset) (
ne:
nenv) :
Prop :=
forall r,
vagree e1#
r e2#
r (
NE.get r ne).
Lemma nreg_agree:
forall rs1 rs2 ne r,
eagree rs1 rs2 ne ->
vagree rs1#
r rs2#
r (
nreg ne r).
Proof.
intros. apply H.
Qed.
Global Hint Resolve nreg_agree:
na.
Lemma eagree_ge:
forall e1 e2 ne ne',
eagree e1 e2 ne ->
NE.ge ne ne' ->
eagree e1 e2 ne'.
Proof.
intros;
red;
intros.
apply nge_agree with (
NE.get r ne);
auto.
apply H0.
Qed.
Lemma eagree_bot:
forall e1 e2,
eagree e1 e2 NE.bot.
Proof.
Lemma eagree_same:
forall e ne,
eagree e e ne.
Proof.
Lemma eagree_update_1:
forall e1 e2 ne v1 v2 nv r,
eagree e1 e2 ne ->
vagree v1 v2 nv ->
eagree (
e1#
r <-
v1) (
e2#
r <-
v2) (
NE.set r nv ne).
Proof.
Lemma eagree_update:
forall e1 e2 ne v1 v2 r,
vagree v1 v2 (
nreg ne r) ->
eagree e1 e2 (
NE.set r Nothing ne) ->
eagree (
e1#
r <-
v1) (
e2#
r <-
v2)
ne.
Proof.
intros;
red;
intros.
specialize (
H0 r0).
rewrite NE.gsspec in H0.
rewrite !
PMap.gsspec.
destruct (
peq r0 r).
subst r0.
auto.
auto.
Qed.
Lemma eagree_update_dead:
forall e1 e2 ne v1 r,
nreg ne r =
Nothing ->
eagree e1 e2 ne ->
eagree (
e1#
r <-
v1)
e2 ne.
Proof.
intros;
red;
intros.
rewrite PMap.gsspec.
destruct (
peq r0 r);
auto.
subst.
unfold nreg in H.
rewrite H.
red;
auto.
Qed.
Neededness for memory locations
Inductive nmem :
Type :=
|
NMemDead
|
NMem (
stk:
ISet.t) (
gl:
PTree.t ISet.t).
Interpretation of
nmem:
-
NMemDead: all memory locations are unused (dead). Acts as bottom.
-
NMem stk gl: all memory locations are used, except:
- the stack locations whose offset is in the interval stk
- the global locations whose offset is in the corresponding entry of gl.
Section LOCATIONS.
Variable ge:
genv.
Variable sp:
block.
Inductive nlive:
nmem ->
block ->
Z ->
Prop :=
|
nlive_intro:
forall stk gl b ofs
(
STK:
b =
sp -> ~
ISet.In ofs stk)
(
GL:
forall id iv,
Genv.find_symbol ge id =
Some b ->
gl!
id =
Some iv ->
~
ISet.In ofs iv),
nlive (
NMem stk gl)
b ofs.
All locations are live
Definition nmem_all :=
NMem ISet.empty (
PTree.empty _).
Lemma nlive_all:
forall b ofs,
nlive nmem_all b ofs.
Proof.
Add a range of live locations to nm. The range starts at
the abstract pointer p and has length sz.
Definition nmem_add (
nm:
nmem) (
p:
aptr) (
sz:
Z) :
nmem :=
match nm with
|
NMemDead =>
nmem_all (* very conservative, should never happen *)
|
NMem stk gl =>
match p with
|
Gl id ofs =>
match gl!
id with
|
Some iv =>
NMem stk (
PTree.set id (
ISet.remove (
Ptrofs.unsigned ofs) (
Ptrofs.unsigned ofs +
sz)
iv)
gl)
|
None =>
nm
end
|
Glo id =>
NMem stk (
PTree.remove id gl)
|
Stk ofs =>
NMem (
ISet.remove (
Ptrofs.unsigned ofs) (
Ptrofs.unsigned ofs +
sz)
stk)
gl
|
Stack =>
NMem ISet.empty gl
| _ =>
nmem_all
end
end.
Lemma nlive_add:
forall bc b ofs p nm sz i,
genv_match bc ge ->
bc sp =
BCstack ->
pmatch bc b ofs p ->
Ptrofs.unsigned ofs <=
i <
Ptrofs.unsigned ofs +
sz ->
nlive (
nmem_add nm p sz)
b i.
Proof.
Lemma incl_nmem_add:
forall nm b i p sz,
nlive nm b i ->
nlive (
nmem_add nm p sz)
b i.
Proof.
intros.
inversion H;
subst.
unfold nmem_add;
destruct p;
try (
apply nlive_all).
-
destruct gl!
id as [
iv|]
eqn:
NG.
+
split;
simpl;
intros.
auto.
rewrite PTree.gsspec in H1.
destruct (
peq id0 id);
eauto.
inv H1.
rewrite ISet.In_remove.
intros [
P Q].
eelim GL;
eauto.
+
auto.
-
split;
simpl;
intros.
auto.
rewrite PTree.grspec in H1.
destruct (
PTree.elt_eq id0 id).
congruence.
eauto.
-
split;
simpl;
intros.
rewrite ISet.In_remove.
intros [
P Q].
eelim STK;
eauto.
eauto.
-
split;
simpl;
intros.
apply ISet.In_empty.
eauto.
Qed.
Remove a range of locations from nm, marking these locations as dead.
The range starts at the abstract pointer p and has length sz.
Definition nmem_remove (
nm:
nmem) (
p:
aptr) (
sz:
Z) :
nmem :=
match nm with
|
NMemDead =>
NMemDead
|
NMem stk gl =>
match p with
|
Gl id ofs =>
let iv' :=
match gl!
id with
|
Some iv =>
ISet.add (
Ptrofs.unsigned ofs) (
Ptrofs.unsigned ofs +
sz)
iv
|
None =>
ISet.interval (
Ptrofs.unsigned ofs) (
Ptrofs.unsigned ofs +
sz)
end in
NMem stk (
PTree.set id iv' gl)
|
Stk ofs =>
NMem (
ISet.add (
Ptrofs.unsigned ofs) (
Ptrofs.unsigned ofs +
sz)
stk)
gl
| _ =>
nm
end
end.
Lemma nlive_remove:
forall bc b ofs p nm sz b' i,
genv_match bc ge ->
bc sp =
BCstack ->
pmatch bc b ofs p ->
nlive nm b' i ->
b' <>
b \/
i <
Ptrofs.unsigned ofs \/
Ptrofs.unsigned ofs +
sz <=
i ->
nlive (
nmem_remove nm p sz)
b' i.
Proof.
Test (conservatively) whether some locations in the range delimited
by p and sz can be live in nm.
Definition nmem_contains (
nm:
nmem) (
p:
aptr) (
sz:
Z) :=
match nm with
|
NMemDead =>
false
|
NMem stk gl =>
match p with
|
Gl id ofs =>
match gl!
id with
|
Some iv =>
negb (
ISet.contains (
Ptrofs.unsigned ofs) (
Ptrofs.unsigned ofs +
sz)
iv)
|
None =>
true
end
|
Stk ofs =>
negb (
ISet.contains (
Ptrofs.unsigned ofs) (
Ptrofs.unsigned ofs +
sz)
stk)
| _ =>
true (* conservative answer *)
end
end.
Lemma nlive_contains:
forall bc b ofs p nm sz i,
genv_match bc ge ->
bc sp =
BCstack ->
pmatch bc b ofs p ->
nmem_contains nm p sz =
false ->
Ptrofs.unsigned ofs <=
i <
Ptrofs.unsigned ofs +
sz ->
~(
nlive nm b i).
Proof.
Kill all stack locations between 0 and sz, and mark everything else
as live. This reflects the effect of freeing the stack block at
a Ireturn or Itailcall instruction.
Definition nmem_dead_stack (
sz:
Z) :=
NMem (
ISet.interval 0
sz) (
PTree.empty _).
Lemma nlive_dead_stack:
forall sz b' i,
b' <>
sp \/ ~(0 <=
i <
sz) ->
nlive (
nmem_dead_stack sz)
b' i.
Proof.
Least upper bound
Definition nmem_lub (
nm1 nm2:
nmem) :
nmem :=
match nm1,
nm2 with
|
NMemDead, _ =>
nm2
| _,
NMemDead =>
nm1
|
NMem stk1 gl1,
NMem stk2 gl2 =>
NMem (
ISet.inter stk1 stk2)
(
PTree.combine
(
fun o1 o2 =>
match o1,
o2 with
|
Some iv1,
Some iv2 =>
Some(
ISet.inter iv1 iv2)
| _, _ =>
None
end)
gl1 gl2)
end.
Lemma nlive_lub_l:
forall nm1 nm2 b i,
nlive nm1 b i ->
nlive (
nmem_lub nm1 nm2)
b i.
Proof.
intros.
inversion H;
subst.
destruct nm2;
simpl.
auto.
constructor;
simpl;
intros.
-
rewrite ISet.In_inter.
intros [
P Q].
eelim STK;
eauto.
-
rewrite PTree.gcombine in H1 by auto.
destruct gl!
id as [
iv1|]
eqn:
NG1;
try discriminate;
destruct gl0!
id as [
iv2|]
eqn:
NG2;
inv H1.
rewrite ISet.In_inter.
intros [
P Q].
eelim GL;
eauto.
Qed.
Lemma nlive_lub_r:
forall nm1 nm2 b i,
nlive nm2 b i ->
nlive (
nmem_lub nm1 nm2)
b i.
Proof.
intros.
inversion H;
subst.
destruct nm1;
simpl.
auto.
constructor;
simpl;
intros.
-
rewrite ISet.In_inter.
intros [
P Q].
eelim STK;
eauto.
-
rewrite PTree.gcombine in H1 by auto.
destruct gl0!
id as [
iv1|]
eqn:
NG1;
try discriminate;
destruct gl!
id as [
iv2|]
eqn:
NG2;
inv H1.
rewrite ISet.In_inter.
intros [
P Q].
eelim GL;
eauto.
Qed.
Boolean-valued equality test
Definition nmem_beq (
nm1 nm2:
nmem) :
bool :=
match nm1,
nm2 with
|
NMemDead,
NMemDead =>
true
|
NMem stk1 gl1,
NMem stk2 gl2 =>
ISet.beq stk1 stk2 &&
PTree.beq ISet.beq gl1 gl2
| _, _ =>
false
end.
Lemma nmem_beq_sound:
forall nm1 nm2 b ofs,
nmem_beq nm1 nm2 =
true ->
(
nlive nm1 b ofs <->
nlive nm2 b ofs).
Proof.
unfold nmem_beq;
intros.
destruct nm1 as [ |
stk1 gl1];
destruct nm2 as [ |
stk2 gl2];
try discriminate.
-
split;
intros L;
inv L.
-
InvBooleans.
rewrite ISet.beq_spec in H0.
rewrite PTree.beq_correct in H1.
split;
intros L;
inv L;
constructor;
intros.
+
rewrite <-
H0.
eauto.
+
specialize (
H1 id).
rewrite H2 in H1.
destruct gl1!
id as [
iv1|]
eqn:
NG;
try contradiction.
rewrite ISet.beq_spec in H1.
rewrite <-
H1.
eauto.
+
rewrite H0.
eauto.
+
specialize (
H1 id).
rewrite H2 in H1.
destruct gl2!
id as [
iv2|]
eqn:
NG;
try contradiction.
rewrite ISet.beq_spec in H1.
rewrite H1.
eauto.
Qed.
End LOCATIONS.
The lattice for dataflow analysis
Module NA <:
SEMILATTICE.
Definition t := (
nenv *
nmem)%
type.
Definition eq (
x y:
t) :=
NE.eq (
fst x) (
fst y) /\
(
forall ge sp b ofs,
nlive ge sp (
snd x)
b ofs <->
nlive ge sp (
snd y)
b ofs).
Lemma eq_refl:
forall x,
eq x x.
Proof.
unfold eq;
destruct x;
simpl;
split.
apply NE.eq_refl.
tauto.
Qed.
Lemma eq_sym:
forall x y,
eq x y ->
eq y x.
Proof.
unfold eq;
destruct x,
y;
simpl.
intros [
A B].
split.
apply NE.eq_sym;
auto.
intros.
rewrite B.
tauto.
Qed.
Lemma eq_trans:
forall x y z,
eq x y ->
eq y z ->
eq x z.
Proof.
unfold eq;
destruct x,
y,
z;
simpl.
intros [
A B] [
C D];
split.
eapply NE.eq_trans;
eauto.
intros.
rewrite B;
auto.
Qed.
Definition beq (
x y:
t) :
bool :=
NE.beq (
fst x) (
fst y) &&
nmem_beq (
snd x) (
snd y).
Lemma beq_correct:
forall x y,
beq x y =
true ->
eq x y.
Proof.
Definition ge (
x y:
t) :
Prop :=
NE.ge (
fst x) (
fst y) /\
(
forall ge sp b ofs,
nlive ge sp (
snd y)
b ofs ->
nlive ge sp (
snd x)
b ofs).
Lemma ge_refl:
forall x y,
eq x y ->
ge x y.
Proof.
unfold eq,
ge;
destruct x,
y;
simpl.
intros [
A B];
split.
apply NE.ge_refl;
auto.
intros.
apply B;
auto.
Qed.
Lemma ge_trans:
forall x y z,
ge x y ->
ge y z ->
ge x z.
Proof.
unfold ge;
destruct x,
y,
z;
simpl.
intros [
A B] [
C D];
split.
eapply NE.ge_trans;
eauto.
auto.
Qed.
Definition bot :
t := (
NE.bot,
NMemDead).
Lemma ge_bot:
forall x,
ge x bot.
Proof.
unfold ge,
bot;
destruct x;
simpl.
split.
apply NE.ge_bot.
intros.
inv H.
Qed.
Definition lub (
x y:
t) :
t :=
(
NE.lub (
fst x) (
fst y),
nmem_lub (
snd x) (
snd y)).
Lemma ge_lub_left:
forall x y,
ge (
lub x y)
x.
Proof.
Lemma ge_lub_right:
forall x y,
ge (
lub x y)
y.
Proof.
End NA.