Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
Require Import AST Linking.
Require Import Memory Registers Op RTL Maps.
Require Import Globalenvs Values.
Require Import Linking Values Memory Globalenvs Events Smallstep.
Require Import Registers Op RTL.
Require Import CSE3analysis CSE2deps CSE2depsproof HashedSet.
Require Import RTLtyping.
Require Import Lia.
Lemma rel_leb_correct:
forall x x',
rel_leb x x' =
true <->
RELATION.ge x' x.
Proof.
#[
global]
Hint Resolve rel_leb_correct :
cse3.
Lemma relb_leb_correct:
forall x x',
relb_leb x x' =
true <->
RB.ge x' x.
Proof.
unfold relb_leb,
RB.ge.
destruct x;
destruct x';
split;
trivial;
try contradiction;
discriminate.
Qed.
#[
global]
Hint Resolve relb_leb_correct :
cse3.
Theorem loadv_storev_really_same:
forall chunk:
memory_chunk,
forall m1:
mem,
forall addr v:
val,
forall m2:
mem,
forall ty :
typ,
forall TYPE:
Val.has_type v ty,
forall STORE:
Mem.storev chunk m1 addr v =
Some m2,
forall COMPATIBLE:
loadv_storev_compatible_type chunk ty =
true,
Mem.loadv chunk m2 addr =
Some v.
Proof.
Lemma subst_args_notin :
forall (
rs :
regset)
dst v args,
~
In dst args ->
(
rs #
dst <-
v) ##
args =
rs ##
args.
Proof.
induction args;
simpl;
trivial.
intro NOTIN.
destruct (
peq a dst).
{
subst a.
intuition congruence.
}
rewrite Regmap.gso by congruence.
f_equal.
apply IHargs.
intuition congruence.
Qed.
Lemma add_i_j_adds :
forall i j m,
PSet.contains (
Regmap.get i (
add_i_j i j m))
j =
true.
Proof.
#[
global]
Hint Resolve add_i_j_adds:
cse3.
Lemma add_i_j_monotone :
forall i j i' j' m,
PSet.contains (
Regmap.get i' m)
j' =
true ->
PSet.contains (
Regmap.get i' (
add_i_j i j m))
j' =
true.
Proof.
#[
global]
Hint Resolve add_i_j_monotone:
cse3.
Lemma add_ilist_j_monotone :
forall ilist j i' j' m,
PSet.contains (
Regmap.get i' m)
j' =
true ->
PSet.contains (
Regmap.get i' (
add_ilist_j ilist j m))
j' =
true.
Proof.
induction ilist; simpl; intros until m; intro CONTAINS; auto with cse3.
Qed.
#[
global]
Hint Resolve add_ilist_j_monotone:
cse3.
Lemma add_ilist_j_adds :
forall ilist j m,
forall i,
In i ilist ->
PSet.contains (
Regmap.get i (
add_ilist_j ilist j m))
j =
true.
Proof.
induction ilist; simpl; intros until i; intro IN.
contradiction.
destruct IN as [HEAD | TAIL]; subst; auto with cse3.
Qed.
#[
global]
Hint Resolve add_ilist_j_adds:
cse3.
Definition xlget_kills (
eqs :
list (
eq_id *
equation_or_condition))
(
m :
Regmap.t PSet.t) :
Regmap.t PSet.t :=
List.fold_left (
fun already (
item :
eq_id *
equation_or_condition) =>
match snd item with
|
Equ lhs sop args =>
add_i_j lhs (
fst item)
(
add_ilist_j args (
fst item)
already)
|
Cond cond args =>
add_ilist_j args (
fst item)
already
end)
eqs m.
Definition xlget_mem_kills (
eqs :
list (
positive *
equation_or_condition))
(
m :
PSet.t) :
PSet.t :=
(
fold_left
(
fun (
a :
PSet.t) (
item :
positive *
equation_or_condition) =>
if eq_cond_depends_on_mem (
snd item)
then PSet.add (
fst item)
a
else a
)
eqs m).
Definition xlget_store_kills (
eqs :
list (
positive *
equation_or_condition))
(
m :
PSet.t) :
PSet.t :=
(
fold_left
(
fun (
a :
PSet.t) (
item :
positive *
equation_or_condition) =>
if eq_cond_depends_on_store (
snd item)
then PSet.add (
fst item)
a
else a
)
eqs m).
Lemma xlget_kills_monotone :
forall eqs m i j,
PSet.contains (
Regmap.get i m)
j =
true ->
PSet.contains (
Regmap.get i (
xlget_kills eqs m))
j =
true.
Proof.
induction eqs; simpl; trivial.
intros.
destruct a as [id eq_cond]; cbn.
destruct eq_cond as [eq_lhs eq_sop eq_args | eq_cond eq_args]; auto with cse3.
Qed.
#[
global]
Hint Resolve xlget_kills_monotone :
cse3.
Lemma xlget_mem_kills_monotone :
forall eqs m j,
PSet.contains m j =
true ->
PSet.contains (
xlget_mem_kills eqs m)
j =
true.
Proof.
induction eqs;
simpl;
trivial.
intros.
destruct a as [
id eq_cond];
cbn.
destruct eq_cond_depends_on_mem.
-
apply IHeqs.
destruct (
peq id j).
+
subst j.
apply PSet.gadds.
+
rewrite PSet.gaddo by congruence.
trivial.
-
auto.
Qed.
#[
global]
Hint Resolve xlget_mem_kills_monotone :
cse3.
Lemma xlget_store_kills_monotone :
forall eqs m j,
PSet.contains m j =
true ->
PSet.contains (
xlget_store_kills eqs m)
j =
true.
Proof.
#[
global]
Hint Resolve xlget_store_kills_monotone :
cse3.
Lemma xlget_kills_has_lhs :
forall eqs m lhs sop args j,
In (
j, (
Equ lhs sop args))
eqs ->
PSet.contains (
Regmap.get lhs (
xlget_kills eqs m))
j =
true.
Proof.
induction eqs; simpl.
contradiction.
intros until j.
intro HEAD_TAIL.
destruct HEAD_TAIL as [HEAD | TAIL]; subst; simpl.
- auto with cse3.
- eapply IHeqs. eassumption.
Qed.
#[
global]
Hint Resolve xlget_kills_has_lhs :
cse3.
Lemma xlget_kills_has_arg :
forall eqs m lhs sop arg args j,
In (
j, (
Equ lhs sop args))
eqs ->
In arg args ->
PSet.contains (
Regmap.get arg (
xlget_kills eqs m))
j =
true.
Proof.
induction eqs; simpl.
contradiction.
intros until j.
intros HEAD_TAIL ARG.
destruct HEAD_TAIL as [HEAD | TAIL]; subst; simpl.
- auto with cse3.
- eapply IHeqs; eassumption.
Qed.
#[
global]
Hint Resolve xlget_kills_has_arg :
cse3.
Lemma xlget_cond_kills_has_arg :
forall eqs m cond arg args j,
In (
j, (
Cond cond args))
eqs ->
In arg args ->
PSet.contains (
Regmap.get arg (
xlget_kills eqs m))
j =
true.
Proof.
induction eqs; simpl.
contradiction.
intros until j.
intros HEAD_TAIL ARG.
destruct HEAD_TAIL as [HEAD | TAIL]; subst; simpl.
- auto with cse3.
- eapply IHeqs; eassumption.
Qed.
#[
global]
Hint Resolve xlget_cond_kills_has_arg :
cse3.
Lemma get_kills_has_lhs :
forall eqs lhs sop args j,
PTree.get j eqs =
Some (
Equ lhs sop args) ->
PSet.contains (
Regmap.get lhs (
get_reg_kills eqs))
j =
true.
Proof.
#[
global]
Hint Resolve get_kills_has_lhs :
cse3.
Lemma context_from_hints_get_kills_has_lhs :
forall hints lhs sop args j,
PTree.get j (
hint_eq_catalog hints) =
Some (
Equ lhs sop args) ->
PSet.contains (
eq_kill_reg (
context_from_hints hints)
lhs)
j =
true.
Proof.
#[
global]
Hint Resolve context_from_hints_get_kills_has_lhs :
cse3.
Lemma get_kills_has_arg :
forall eqs lhs sop arg args j,
PTree.get j eqs =
Some (
Equ lhs sop args) ->
In arg args ->
PSet.contains (
Regmap.get arg (
get_reg_kills eqs))
j =
true.
Proof.
#[
global]
Hint Resolve get_kills_has_arg :
cse3.
Lemma context_from_hints_get_kills_has_arg :
forall hints lhs sop arg args j,
PTree.get j (
hint_eq_catalog hints) =
Some (
Equ lhs sop args) ->
In arg args ->
PSet.contains (
eq_kill_reg (
context_from_hints hints)
arg)
j =
true.
Proof.
#[
global]
Hint Resolve context_from_hints_get_kills_has_arg :
cse3.
Lemma get_cond_kills_has_arg :
forall eqs cond arg args j,
PTree.get j eqs =
Some (
Cond cond args) ->
In arg args ->
PSet.contains (
Regmap.get arg (
get_reg_kills eqs))
j =
true.
Proof.
#[
global]
Hint Resolve get_cond_kills_has_arg :
cse3.
Lemma context_from_hints_get_cond_kills_has_arg :
forall hints cond arg args j,
PTree.get j (
hint_eq_catalog hints) =
Some (
Cond cond args) ->
In arg args ->
PSet.contains (
eq_kill_reg (
context_from_hints hints)
arg)
j =
true.
Proof.
#[
global]
Hint Resolve context_from_hints_get_cond_kills_has_arg :
cse3.
Lemma xlget_kills_has_eq_depends_on_mem :
forall eqs eq j m,
In (
j,
eq)
eqs ->
eq_cond_depends_on_mem eq =
true ->
PSet.contains (
xlget_mem_kills eqs m)
j =
true.
Proof.
#[
global]
Hint Resolve xlget_kills_has_eq_depends_on_mem :
cse3.
Lemma get_kills_has_eq_depends_on_mem :
forall eqs eq j,
PTree.get j eqs =
Some eq ->
eq_cond_depends_on_mem eq =
true ->
PSet.contains (
get_mem_kills eqs)
j =
true.
Proof.
Lemma context_from_hints_get_kills_has_eq_depends_on_mem :
forall hints eq j,
PTree.get j (
hint_eq_catalog hints) =
Some eq ->
eq_cond_depends_on_mem eq =
true ->
PSet.contains (
eq_kill_mem (
context_from_hints hints)
tt)
j =
true.
Proof.
#[
global]
Hint Resolve context_from_hints_get_kills_has_eq_depends_on_mem :
cse3.
Lemma xlget_kills_has_eq_depends_on_store :
forall eqs eq j m,
In (
j,
eq)
eqs ->
eq_cond_depends_on_store eq =
true ->
PSet.contains (
xlget_store_kills eqs m)
j =
true.
Proof.
#[
global]
Hint Resolve xlget_kills_has_eq_depends_on_store :
cse3.
Lemma get_kills_has_eq_depends_on_store :
forall eqs eq j,
PTree.get j eqs =
Some eq ->
eq_cond_depends_on_store eq =
true ->
PSet.contains (
get_store_kills eqs)
j =
true.
Proof.
Lemma context_from_hints_get_kills_has_eq_depends_on_store :
forall hints eq j,
PTree.get j (
hint_eq_catalog hints) =
Some eq ->
eq_cond_depends_on_store eq =
true ->
PSet.contains (
eq_kill_store (
context_from_hints hints)
tt)
j =
true.
Proof.
#[
global]
Hint Resolve context_from_hints_get_kills_has_eq_depends_on_store :
cse3.
Definition eq_involves (
eq :
equation_or_condition) (
i :
reg) :=
match eq with
|
Equ lhs sop args =>
i =
lhs \/
In i args
|
Cond cond args =>
In i args
end.
Section SOUNDNESS.
Context {
F V :
Type}.
Context {
genv:
Genv.t F V}.
Context {
sp :
val}.
Context {
ctx :
eq_context}.
Definition sem_rhs (
sop :
sym_op) (
args :
list reg)
(
rs :
regset) (
m :
mem) (
v' :
val) :=
match sop with
|
SOp op =>
match eval_operation genv sp op (
rs ##
args)
m with
|
Some v =>
v' =
v
|
None =>
False
end
|
SLoad chunk addr =>
match
match eval_addressing genv sp addr (
rs ##
args)
with
|
Some a =>
Mem.loadv chunk m a
|
None =>
None
end
with
|
Some dat =>
v' =
dat
|
None =>
v' =
Vundef
end
end.
Definition sem_eq (
eq :
equation_or_condition) (
rs :
regset) (
m :
mem) :=
match eq with
|
Equ lhs sop args =>
sem_rhs sop args rs m (
rs #
lhs)
|
Cond cond args =>
eval_condition cond (
rs ##
args)
m =
Some true
end.
Definition sem_rel (
rel :
RELATION.t) (
rs :
regset) (
m :
mem) :=
forall i eq,
PSet.contains rel i =
true ->
eq_catalog ctx i =
Some eq ->
sem_eq eq rs m.
Set Apply With Renaming.
Lemma sem_rel_glb:
forall rel1 rel2 rs m,
(
sem_rel (
RELATION.glb rel1 rel2)
rs m) <->
((
sem_rel rel1 rs m) /\
(
sem_rel rel2 rs m)).
Proof.
intros.
unfold sem_rel,
RELATION.glb.
split.
-
intro IMPLIES.
split;
intros i eq CONTAINS;
(
specialize IMPLIES with (
i:=
i) (
eq0:=
eq) ||
specialize IMPLIES with (
i:=
i) (
eq:=
eq));
rewrite PSet.gunion in IMPLIES;
rewrite orb_true_iff in IMPLIES;
intuition.
-
intros (
IMPLIES1 &
IMPLIES2)
i eq.
rewrite PSet.gunion.
rewrite orb_true_iff.
specialize IMPLIES1 with (
i:=
i) (
eq0:=
eq) ||
specialize IMPLIES1 with (
i:=
i) (
eq:=
eq) .
specialize IMPLIES2 with (
i:=
i) (
eq0:=
eq) ||
specialize IMPLIES2 with (
i:=
i) (
eq:=
eq) .
intuition.
Qed.
Hypothesis ctx_kill_reg_has_lhs :
forall lhs sop args j,
eq_catalog ctx j =
Some (
Equ lhs sop args) ->
PSet.contains (
eq_kill_reg ctx lhs)
j =
true.
Hypothesis ctx_kill_reg_has_arg :
forall lhs sop args j,
eq_catalog ctx j =
Some (
Equ lhs sop args) ->
forall arg,
In arg args ->
PSet.contains (
eq_kill_reg ctx arg)
j =
true.
Hypothesis ctx_cond_kill_reg_has_arg :
forall cond args j,
eq_catalog ctx j =
Some (
Cond cond args) ->
forall arg,
In arg args ->
PSet.contains (
eq_kill_reg ctx arg)
j =
true.
Hypothesis ctx_kill_mem_has_depends_on_mem :
forall eq j,
eq_catalog ctx j =
Some eq ->
eq_cond_depends_on_mem eq =
true ->
PSet.contains (
eq_kill_mem ctx tt)
j =
true.
Hypothesis ctx_kill_store_has_depends_on_store :
forall eq j,
eq_catalog ctx j =
Some eq ->
eq_cond_depends_on_store eq =
true ->
PSet.contains (
eq_kill_store ctx tt)
j =
true.
Theorem kill_reg_sound :
forall rel rs m dst v,
(
sem_rel rel rs m) ->
(
sem_rel (
kill_reg (
ctx:=
ctx)
dst rel) (
rs#
dst <-
v)
m).
Proof.
Hint Resolve kill_reg_sound :
cse3.
Theorem kill_reg_sound2 :
forall rel rs m dst,
(
sem_rel rel rs m) ->
(
sem_rel (
kill_reg (
ctx:=
ctx)
dst rel)
rs m).
Proof.
Lemma pick_source_sound :
forall (
l :
list reg),
match pick_source l with
|
Some x =>
In x l
|
None =>
True
end.
Proof.
unfold pick_source.
destruct l;
simpl;
trivial.
left;
trivial.
Qed.
Hint Resolve pick_source_sound :
cse3.
Theorem forward_move_sound :
forall rel rs m x,
(
sem_rel rel rs m) ->
rs # (
forward_move (
ctx :=
ctx)
rel x) =
rs #
x.
Proof.
Hint Resolve forward_move_sound :
cse3.
Theorem forward_move_l_sound :
forall rel rs m l,
(
sem_rel rel rs m) ->
rs ## (
forward_move_l (
ctx :=
ctx)
rel l) =
rs ##
l.
Proof.
induction l;
simpl;
intros;
trivial.
erewrite forward_move_sound by eassumption.
intuition congruence.
Qed.
Hint Resolve forward_move_l_sound :
cse3.
Theorem kill_mem_sound :
forall rel rs m m',
(
sem_rel rel rs m) ->
(
sem_rel (
kill_mem (
ctx:=
ctx)
rel)
rs m').
Proof.
Hint Resolve kill_mem_sound :
cse3.
Lemma store_preserves_validity:
forall m m' wchunk a v
(
STORE :
Mem.storev wchunk m a v =
Some m')
(
b :
block) (
z :
Z),
Mem.valid_pointer m' b z =
Mem.valid_pointer m b z.
Proof.
Hint Resolve store_preserves_validity :
cse3.
Theorem kill_store_sound :
forall rel rs m m' wchunk a v,
(
sem_rel rel rs m) ->
(
Mem.storev wchunk m a v =
Some m') ->
(
sem_rel (
kill_store (
ctx:=
ctx)
rel)
rs m').
Proof.
Hint Resolve kill_store_sound :
cse3.
Theorem eq_find_sound:
forall no eq id,
eq_find (
ctx :=
ctx)
no eq =
Some id ->
eq_catalog ctx id =
Some eq.
Proof.
Hint Resolve eq_find_sound :
cse3.
Theorem is_condition_present_sound :
forall node rel cond args rs m
(
REL :
sem_rel rel rs m)
(
COND : (
is_condition_present (
ctx :=
ctx)
node rel cond args) =
true),
(
eval_condition cond (
rs ##
args)
m) =
Some true.
Proof.
Hint Resolve is_condition_present_sound :
cse3.
Theorem rhs_find_sound:
forall no sop args rel src rs m,
sem_rel rel rs m ->
rhs_find (
ctx :=
ctx)
no sop args rel =
Some src ->
sem_rhs sop args rs m (
rs #
src).
Proof.
Hint Resolve rhs_find_sound :
cse3.
Theorem forward_move_rhs_sound :
forall sop args rel rs m v,
(
sem_rel rel rs m) ->
(
sem_rhs sop args rs m v) ->
(
sem_rhs sop (
forward_move_l (
ctx :=
ctx)
rel args)
rs m v).
Proof.
intros until v.
intros REL RHS.
destruct sop;
simpl in *.
all:
erewrite forward_move_l_sound by eassumption;
assumption.
Qed.
Hint Resolve forward_move_rhs_sound :
cse3.
Lemma arg_not_replaced:
forall (
rs :
regset)
dst v args,
~
In dst args ->
(
rs #
dst <-
v) ##
args =
rs ##
args.
Proof.
induction args;
simpl;
trivial.
intuition.
f_equal;
trivial.
apply Regmap.gso;
congruence.
Qed.
Lemma sem_rhs_depends_on_args_only:
forall sop args rs dst m v,
sem_rhs sop args rs m v ->
~
In dst args ->
sem_rhs sop args (
rs #
dst <-
v)
m v.
Proof.
Lemma replace_sound:
forall no eqno dst sop args rel rs m v,
sem_rel rel rs m ->
sem_rhs sop args rs m v ->
~
In dst args ->
eq_find (
ctx :=
ctx)
no (
Equ dst sop args) =
Some eqno ->
sem_rel (
PSet.add eqno (
kill_reg (
ctx :=
ctx)
dst rel)) (
rs #
dst <-
v)
m.
Proof.
Lemma sem_rhs_det:
forall {
sop} {
args} {
rs} {
m} {
v} {
v'},
sem_rhs sop args rs m v ->
sem_rhs sop args rs m v' ->
v =
v'.
Proof.
Lemma arglist_idem_write:
forall {
A :
Type}
args (
rs :
Regmap.t A)
dst,
(
rs #
dst <- (
rs #
dst)) ##
args =
rs ##
args.
Proof.
induction args;
trivial.
intros.
cbn.
f_equal;
trivial.
apply Regmap.gsident.
Qed.
Lemma sem_rhs_idem_write:
forall sop args rs dst m v,
sem_rhs sop args rs m v ->
sem_rhs sop args (
rs #
dst <- (
rs #
dst))
m v.
Proof.
Theorem oper2_sound:
forall no dst sop args rel rs m v,
sem_rel rel rs m ->
not (
In dst args) ->
sem_rhs sop args rs m v ->
sem_rel (
oper2 (
ctx :=
ctx)
no dst sop args rel) (
rs #
dst <-
v)
m.
Proof.
unfold oper2.
intros until v.
intros REL NOTIN RHS.
pose proof (
eq_find_sound no (
Equ dst sop args))
as EQ_FIND_SOUND.
destruct eq_find.
2:
auto with cse3;
fail.
specialize EQ_FIND_SOUND with (
id :=
e).
intuition.
intros i eq CONTAINS.
destruct (
peq i e).
{
subst i.
rewrite H.
clear H.
intro Z.
inv Z.
unfold sem_eq.
simpl.
rewrite Regmap.gss.
apply sem_rhs_depends_on_args_only;
auto.
}
intros INi.
destruct (
PSet.contains rel e)
eqn:
CONTAINSe.
{
pose proof (
REL e (
Equ dst sop args)
CONTAINSe H)
as RELe.
pose proof (
REL i eq CONTAINS INi)
as RELi.
destruct eq as [
eq_lhs eq_sop eq_args |
eq_cond eq_args];
cbn in *.
-
replace v with (
rs #
dst)
by (
eapply sem_rhs_det;
eassumption).
rewrite Regmap.gsident.
apply sem_rhs_idem_write.
assumption.
-
replace v with (
rs #
dst)
by (
eapply sem_rhs_det;
eassumption).
rewrite arglist_idem_write.
assumption.
}
rewrite PSet.gaddo in CONTAINS by congruence.
apply (
kill_reg_sound rel rs m dst v REL i eq);
auto.
Qed.
Hint Resolve oper2_sound :
cse3.
Theorem oper1_sound:
forall no dst sop args rel rs m v,
sem_rel rel rs m ->
sem_rhs sop args rs m v ->
sem_rel (
oper1 (
ctx :=
ctx)
no dst sop args rel) (
rs #
dst <-
v)
m.
Proof.
intros.
unfold oper1.
destruct in_dec;
auto with cse3.
Qed.
Hint Resolve oper1_sound :
cse3.
Lemma rel_idem_replace:
forall rel rs r m,
sem_rel rel rs m ->
sem_rel rel rs #
r <- (
rs #
r)
m.
Proof.
Lemma move_sound :
forall no :
node,
forall rel :
RELATION.t,
forall src dst :
reg,
forall rs m,
sem_rel rel rs m ->
sem_rel (
move (
ctx:=
ctx)
no src dst rel) (
rs #
dst <- (
rs #
src))
m.
Proof.
Hint Resolve move_sound :
cse3.
Theorem oper_sound:
forall no dst sop args rel rs m v,
sem_rel rel rs m ->
sem_rhs sop args rs m v ->
sem_rel (
oper (
ctx :=
ctx)
no dst sop args rel) (
rs #
dst <-
v)
m.
Proof.
Hint Resolve oper_sound :
cse3.
Theorem clever_kill_store_sound:
forall chunk addr args a src rel rs m m',
sem_rel rel rs m ->
eval_addressing genv sp addr (
rs ##
args) =
Some a ->
Mem.storev chunk m a (
rs #
src) =
Some m' ->
sem_rel (
clever_kill_store (
ctx:=
ctx)
chunk addr args src rel)
rs m'.
Proof.
unfold clever_kill_store.
intros until m'.
intros REL ADDR STORE i eq CONTAINS CATALOG.
autorewrite with pset in CONTAINS.
destruct (
PSet.contains rel i)
eqn:
RELi;
simpl in CONTAINS.
2:
discriminate.
rewrite CATALOG in CONTAINS.
unfold sem_rel in REL.
specialize REL with (
i :=
i) (
eq0 :=
eq) ||
specialize REL with (
i :=
i) (
eq :=
eq).
destruct eq as [
eq_lhs eq_sop eq_args |
eq_cond eq_args];
simpl in *.
*
unfold sem_eq in *.
simpl in *.
destruct eq_sop as [
op' |
chunk' addr'];
simpl.
-
rewrite op_valid_pointer_eq with (
m2 :=
m).
+
cbn in *.
apply REL;
auto.
+
eapply store_preserves_validity;
eauto.
-
simpl in REL.
erewrite ctx_kill_store_has_depends_on_store in CONTAINS by eauto.
simpl in CONTAINS.
rewrite negb_true_iff in CONTAINS.
destruct (
eval_addressing genv sp addr' rs ##
eq_args)
as [
a'|]
eqn:
ADDR'.
+
erewrite may_overlap_sound with (
chunk:=
chunk) (
addr:=
addr) (
args:=
args) (
chunk':=
chunk') (
addr':=
addr') (
args':=
eq_args);
try eassumption.
apply REL;
auto.
+
apply REL;
auto.
*
rewrite cond_valid_pointer_eq with (
m2 :=
m).
auto.
eapply store_preserves_validity;
eauto.
Qed.
Hint Resolve clever_kill_store_sound :
cse3.
Theorem store2_sound:
forall chunk addr args a src rel rs m m',
sem_rel rel rs m ->
eval_addressing genv sp addr (
rs ##
args) =
Some a ->
Mem.storev chunk m a (
rs #
src) =
Some m' ->
sem_rel (
store2 (
ctx:=
ctx)
chunk addr args src rel)
rs m'.
Proof.
Hint Resolve store2_sound :
cse3.
Theorem store1_sound:
forall no chunk addr args a src rel tenv rs m m',
sem_rel rel rs m ->
wt_regset tenv rs ->
eval_addressing genv sp addr (
rs ##
args) =
Some a ->
Mem.storev chunk m a (
rs#
src) =
Some m' ->
sem_rel (
store1 (
ctx:=
ctx)
no chunk addr args src (
tenv src)
rel)
rs m'.
Proof.
unfold store1.
intros until m'.
intros REL WT ADDR STORE.
assert (
sem_rel (
store2 (
ctx:=
ctx)
chunk addr args src rel)
rs m')
as REL' by eauto with cse3.
destruct loadv_storev_compatible_type eqn:
COMPATIBLE.
2:
auto;
fail.
destruct eq_find as [
eq_id | ]
eqn:
FIND.
2:
auto;
fail.
intros i eq CONTAINS CATALOG.
destruct (
peq i eq_id).
{
subst i.
(
rewrite eq_find_sound with (
no:=
no) (
eq0:=
Equ src (
SLoad chunk addr)
args)
in CATALOG ||
rewrite eq_find_sound with (
no:=
no) (
eq:=
Equ src (
SLoad chunk addr)
args)
in CATALOG);
trivial.
inv CATALOG.
unfold sem_eq.
simpl.
rewrite ADDR.
rewrite loadv_storev_really_same with (
m1:=
m) (
v:=
rs#
src) (
ty:=(
tenv src));
trivial.
}
unfold sem_rel in REL'.
rewrite PSet.gaddo in CONTAINS by congruence.
eauto.
Qed.
Hint Resolve store1_sound :
cse3.
Theorem store_sound:
forall no chunk addr args a src rel tenv rs m m',
sem_rel rel rs m ->
wt_regset tenv rs ->
eval_addressing genv sp addr (
rs ##
args) =
Some a ->
Mem.storev chunk m a (
rs#
src) =
Some m' ->
sem_rel (
store (
ctx:=
ctx)
no tenv chunk addr args src rel)
rs m'.
Proof.
Hint Resolve store_sound :
cse3.
Lemma kill_builtin_res_sound:
forall res (
m :
mem) (
rs :
regset)
vres (
rel :
RELATION.t)
(
REL :
sem_rel rel rs m),
(
sem_rel (
kill_builtin_res (
ctx:=
ctx)
res rel)
(
regmap_setres res vres rs)
m).
Proof.
destruct res;
simpl;
intros;
trivial.
apply kill_reg_sound;
trivial.
Qed.
Hint Resolve kill_builtin_res_sound :
cse3.
Lemma top_sound:
forall rs m, (
sem_rel RELATION.top rs m).
Proof.
Hint Resolve top_sound :
cse3.
Lemma external_call_sound:
forall ge ef (
rel :
RELATION.t) (
m m' :
mem) (
rs :
regset)
vargs t vres
(
REL :
sem_rel rel rs m)
(
CALL :
external_call ef ge vargs m t vres m'),
sem_rel (
apply_external_call (
ctx:=
ctx)
ef rel)
rs m'.
Proof.
Hint Resolve external_call_sound :
cse3.
Definition sem_rel_b (
rel :
RB.t) (
rs :
regset) (
m :
mem) :=
match rel with
|
None =>
False
|
Some rel =>
sem_rel rel rs m
end.
Lemma apply_cond1_sound :
forall pc cond args rel rs m
(
COND : (
eval_condition cond (
rs ##
args)
m) =
Some true)
(
REL : (
sem_rel rel rs m)),
(
sem_rel_b (
apply_cond1 (
ctx:=
ctx)
pc cond args rel)
rs m).
Proof.
Lemma apply_cond0_sound :
forall pc cond args rel rs m
(
COND : (
eval_condition cond (
rs ##
args)
m) =
Some true)
(
REL : (
sem_rel rel rs m)),
(
sem_rel (
apply_cond0 (
ctx:=
ctx)
pc cond args rel)
rs m).
Proof.
intros.
unfold apply_cond0.
destruct eq_find as [
eq_id | ]
eqn:
FIND;
cbn.
2:
assumption.
pose proof (
eq_find_sound pc (
Cond cond args)
eq_id FIND)
as FIND_SOUND.
intros eq_id' eq' CONTAINS CATALOG.
destruct (
peq eq_id eq_id').
{
subst eq_id'.
unfold sem_eq.
rewrite FIND_SOUND in CATALOG.
inv CATALOG.
assumption.
}
rewrite PSet.gaddo in CONTAINS by assumption.
unfold sem_rel in REL.
eapply REL;
eassumption.
Qed.
Lemma apply_cond_sound :
forall pc cond args rel rs m
(
COND : (
eval_condition cond (
rs ##
args)
m) =
Some true)
(
REL : (
sem_rel rel rs m)),
(
sem_rel_b (
apply_cond (
ctx:=
ctx)
pc cond args rel)
rs m).
Proof.
End SOUNDNESS.