Correctness proof for Cminor generation.
Require Import Coq.Program.Equality.
Require Import FSets.
Require Import Coqlib.
Require Intv.
Require Import Errors.
Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Memdata.
Require Import Memory.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Switch.
Require Import Csharpminor.
Require Import Cminor.
Require Import Cminorgen.
Open Local Scope error_monad_scope.
Section TRANSLATION.
Variable prog:
Csharpminor.program.
Variable tprog:
program.
Hypothesis TRANSL:
transl_program prog =
OK tprog.
Let ge :
Csharpminor.genv :=
Genv.globalenv prog.
Let gce :
compilenv :=
build_global_compilenv prog.
Let tge:
genv :=
Genv.globalenv tprog.
Lemma symbols_preserved:
forall (
s:
ident),
Genv.find_symbol tge s =
Genv.find_symbol ge s.
Proof (
Genv.find_symbol_transf_partial2 (
transl_fundef gce)
transl_globvar _ TRANSL).
Lemma function_ptr_translated:
forall (
b:
block) (
f:
Csharpminor.fundef),
Genv.find_funct_ptr ge b =
Some f ->
exists tf,
Genv.find_funct_ptr tge b =
Some tf /\
transl_fundef gce f =
OK tf.
Proof (
Genv.find_funct_ptr_transf_partial2 (
transl_fundef gce)
transl_globvar _ TRANSL).
Lemma functions_translated:
forall (
v:
val) (
f:
Csharpminor.fundef),
Genv.find_funct ge v =
Some f ->
exists tf,
Genv.find_funct tge v =
Some tf /\
transl_fundef gce f =
OK tf.
Proof (
Genv.find_funct_transf_partial2 (
transl_fundef gce)
transl_globvar _ TRANSL).
Lemma var_info_translated:
forall b v,
Genv.find_var_info ge b =
Some v ->
exists tv,
Genv.find_var_info tge b =
Some tv /\
transf_globvar transl_globvar v =
OK tv.
Proof (
Genv.find_var_info_transf_partial2 (
transl_fundef gce)
transl_globvar _ TRANSL).
Lemma var_info_rev_translated:
forall b tv,
Genv.find_var_info tge b =
Some tv ->
exists v,
Genv.find_var_info ge b =
Some v /\
transf_globvar transl_globvar v =
OK tv.
Proof (
Genv.find_var_info_rev_transf_partial2 (
transl_fundef gce)
transl_globvar _ TRANSL).
Lemma sig_preserved_body:
forall f tf cenv size,
transl_funbody cenv size f =
OK tf ->
tf.(
fn_sig) =
Csharpminor.fn_sig f.
Proof.
intros. monadInv H. reflexivity.
Qed.
Lemma sig_preserved:
forall f tf,
transl_fundef gce f =
OK tf ->
Cminor.funsig tf =
Csharpminor.funsig f.
Proof.
Definition global_compilenv_match (
ce:
compilenv) (
ge:
Csharpminor.genv) :
Prop :=
forall id,
match ce!!
id with
|
Var_global_scalar chunk =>
forall b gv,
Genv.find_symbol ge id =
Some b ->
Genv.find_var_info ge b =
Some gv ->
gv.(
gvar_info) =
Vscalar chunk
|
Var_global_array =>
True
|
_ =>
False
end.
Lemma global_compilenv_charact:
global_compilenv_match gce ge.
Proof.
Derived properties of memory operations
Lemma load_freelist:
forall fbl chunk m b ofs m',
(
forall b'
lo hi,
In (
b',
lo,
hi)
fbl ->
b' <>
b) ->
Mem.free_list m fbl =
Some m' ->
Mem.load chunk m'
b ofs =
Mem.load chunk m b ofs.
Proof.
induction fbl;
intros.
simpl in H0.
congruence.
destruct a as [[
b'
lo]
hi].
generalize H0.
simpl.
case_eq (
Mem.free m b'
lo hi);
try congruence.
intros m1 FR1 FRL.
transitivity (
Mem.load chunk m1 b ofs).
eapply IHfbl;
eauto.
intros.
eapply H.
eauto with coqlib.
eapply Mem.load_free;
eauto.
left.
apply sym_not_equal.
eapply H.
auto with coqlib.
Qed.
Lemma perm_freelist:
forall fbl m m'
b ofs k p,
Mem.free_list m fbl =
Some m' ->
Mem.perm m'
b ofs k p ->
Mem.perm m b ofs k p.
Proof.
induction fbl;
simpl;
intros until p.
congruence.
destruct a as [[
b'
lo]
hi].
case_eq (
Mem.free m b'
lo hi);
try congruence.
intros.
eauto with mem.
Qed.
Lemma nextblock_freelist:
forall fbl m m',
Mem.free_list m fbl =
Some m' ->
Mem.nextblock m' =
Mem.nextblock m.
Proof.
induction fbl;
intros until m';
simpl.
congruence.
destruct a as [[
b lo]
hi].
case_eq (
Mem.free m b lo hi);
intros;
try congruence.
transitivity (
Mem.nextblock m0).
eauto.
eapply Mem.nextblock_free;
eauto.
Qed.
Lemma free_list_freeable:
forall l m m',
Mem.free_list m l =
Some m' ->
forall b lo hi,
In (
b,
lo,
hi)
l ->
Mem.range_perm m b lo hi Cur Freeable.
Proof.
induction l;
simpl;
intros.
contradiction.
revert H.
destruct a as [[
b'
lo']
hi'].
caseEq (
Mem.free m b'
lo'
hi');
try congruence.
intros m1 FREE1 FREE2.
destruct H0.
inv H.
eauto with mem.
red;
intros.
eapply Mem.perm_free_3;
eauto.
exploit IHl;
eauto.
Qed.
Lemma nextblock_storev:
forall chunk m addr v m',
Mem.storev chunk m addr v =
Some m' ->
Mem.nextblock m' =
Mem.nextblock m.
Proof.
unfold Mem.storev;
intros.
destruct addr;
try discriminate.
eapply Mem.nextblock_store;
eauto.
Qed.
Correspondence between Csharpminor's and Cminor's environments and memory states
In Csharpminor, every variable is stored in a separate memory block.
In the corresponding Cminor code, some of these variables reside in
the local variable environment; others are sub-blocks of the stack data
block. We capture these changes in memory via a memory injection
f:
-
f b = None means that the Csharpminor block b no longer exist
in the execution of the generated Cminor code. This corresponds to
a Csharpminor local variable translated to a Cminor local variable.
-
f b = Some(b', ofs) means that Csharpminor block b corresponds
to a sub-block of Cminor block b at offset ofs.
A memory injection
f defines a relation
val_inject f between
values and a relation
Mem.inject f between memory states.
These relations will be used intensively
in our proof of simulation between Csharpminor and Cminor executions.
In this section, we define the relation between
Csharpminor and Cminor environments.
Matching for a Csharpminor variable
id.
-
If this variable is mapped to a Cminor local variable, the
corresponding Csharpminor memory block b must no longer exist in
Cminor (f b = None). Moreover, the content of block b must
match the value of id found in the Cminor local environment te.
-
If this variable is mapped to a sub-block of the Cminor stack data
at offset ofs, the address of this variable in Csharpminor Vptr b
Int.zero must match the address of the sub-block Vptr sp (Int.repr
ofs).
Inductive match_var (
f:
meminj) (
id:
ident)
(
e:
Csharpminor.env) (
m:
mem) (
te:
env) (
sp:
block) :
var_info ->
Prop :=
|
match_var_local:
forall chunk b v v',
PTree.get id e =
Some (
b,
Vscalar chunk) ->
Mem.load chunk m b 0 =
Some v ->
f b =
None ->
PTree.get (
for_var id)
te =
Some v' ->
val_inject f v v' ->
match_var f id e m te sp (
Var_local chunk)
|
match_var_stack_scalar:
forall chunk ofs b,
PTree.get id e =
Some (
b,
Vscalar chunk) ->
val_inject f (
Vptr b Int.zero) (
Vptr sp (
Int.repr ofs)) ->
match_var f id e m te sp (
Var_stack_scalar chunk ofs)
|
match_var_stack_array:
forall ofs sz al b,
PTree.get id e =
Some (
b,
Varray sz al) ->
val_inject f (
Vptr b Int.zero) (
Vptr sp (
Int.repr ofs)) ->
match_var f id e m te sp (
Var_stack_array ofs sz al)
|
match_var_global_scalar:
forall chunk,
PTree.get id e =
None ->
(
forall b gv,
Genv.find_symbol ge id =
Some b ->
Genv.find_var_info ge b =
Some gv ->
gvar_info gv =
Vscalar chunk) ->
match_var f id e m te sp (
Var_global_scalar chunk)
|
match_var_global_array:
PTree.get id e =
None ->
match_var f id e m te sp (
Var_global_array).
Matching between a Csharpminor environment e and a Cminor
environment te. The lo and hi parameters delimit the range
of addresses for the blocks referenced from te.
Record match_env (
f:
meminj) (
cenv:
compilenv)
(
e:
Csharpminor.env) (
le:
Csharpminor.temp_env) (
m:
mem)
(
te:
env) (
sp:
block)
(
lo hi:
Z) :
Prop :=
mk_match_env {
Each variable mentioned in the compilation environment must match
as defined above.
me_vars:
forall id,
match_var f id e m te sp (
PMap.get id cenv);
Temporaries match
me_temps:
forall id v,
le!
id =
Some v ->
exists v',
te!(
for_temp id) =
Some v' /\
val_inject f v v';
lo, hi is a proper interval.
me_low_high:
lo <=
hi;
Every block appearing in the Csharpminor environment e must be
in the range lo, hi.
me_bounded:
forall id b lv,
PTree.get id e =
Some(
b,
lv) ->
lo <=
b <
hi;
Distinct Csharpminor local variables must be mapped to distinct blocks.
me_inj:
forall id1 b1 lv1 id2 b2 lv2,
PTree.get id1 e =
Some(
b1,
lv1) ->
PTree.get id2 e =
Some(
b2,
lv2) ->
id1 <>
id2 ->
b1 <>
b2;
All blocks mapped to sub-blocks of the Cminor stack data must be
images of variables from the Csharpminor environment e
me_inv:
forall b delta,
f b =
Some(
sp,
delta) ->
exists id,
exists lv,
PTree.get id e =
Some(
b,
lv);
All Csharpminor blocks below lo (i.e. allocated before the blocks
referenced from e) must map to blocks that are below sp
(i.e. allocated before the stack data for the current Cminor function).
me_incr:
forall b tb delta,
f b =
Some(
tb,
delta) ->
b <
lo ->
tb <
sp;
The sizes of blocks appearing in e agree with their types
me_bounds:
forall id b lv ofs p,
PTree.get id e =
Some(
b,
lv) ->
Mem.perm m b ofs Max p -> 0 <=
ofs <
sizeof lv
}.
Hint Resolve me_low_high.
The remainder of this section is devoted to showing preservation
of the match_en invariant under various assignments and memory
stores. First: preservation by memory stores to ``mapped'' blocks
(block that have a counterpart in the Cminor execution).
Ltac geninv x :=
let H :=
fresh in (
generalize x;
intro H;
inv H).
Lemma match_env_store_mapped:
forall f cenv e le m1 m2 te sp lo hi chunk b ofs v,
f b <>
None ->
Mem.store chunk m1 b ofs v =
Some m2 ->
match_env f cenv e le m1 te sp lo hi ->
match_env f cenv e le m2 te sp lo hi.
Proof.
intros;
inv H1;
constructor;
auto.
intros.
geninv (
me_vars0 id);
econstructor;
eauto.
rewrite <-
H4.
eapply Mem.load_store_other;
eauto.
left.
congruence.
intros.
eauto with mem.
Qed.
Preservation by assignment to a Csharpminor variable that is
translated to a Cminor local variable. The value being assigned
must be normalized with respect to the memory chunk of the variable.
Remark val_normalized_has_type:
forall v chunk,
val_normalized v chunk ->
Val.has_type v (
type_of_chunk chunk).
Proof.
intros.
red in H.
rewrite <-
H.
destruct chunk;
destruct v;
exact I.
Qed.
Lemma match_env_store_local:
forall f cenv e le m1 m2 te sp lo hi id b chunk v tv,
e!
id =
Some(
b,
Vscalar chunk) ->
val_normalized v chunk ->
val_inject f v tv ->
Mem.store chunk m1 b 0
v =
Some m2 ->
match_env f cenv e le m1 te sp lo hi ->
match_env f cenv e le m2 (
PTree.set (
for_var id)
tv te)
sp lo hi.
Proof.
intros.
inv H3.
constructor;
auto.
intros.
geninv (
me_vars0 id0).
case (
peq id id0);
intro.
subst id0.
assert (
b0 =
b)
by congruence.
subst.
assert (
chunk0 =
chunk)
by congruence.
subst.
econstructor.
eauto.
eapply Mem.load_store_same;
eauto.
auto.
rewrite PTree.gss.
reflexivity.
red in H0.
rewrite H0.
auto.
econstructor;
eauto.
rewrite <-
H6.
eapply Mem.load_store_other;
eauto.
rewrite PTree.gso;
auto.
unfold for_var;
congruence.
econstructor;
eauto.
econstructor;
eauto.
econstructor;
eauto.
econstructor;
eauto.
intros.
rewrite PTree.gso.
auto.
unfold for_temp,
for_var;
congruence.
intros.
eauto with mem.
Qed.
Preservation by assignment to a Csharpminor temporary and the
corresponding Cminor local variable.
Lemma match_env_set_temp:
forall f cenv e le m te sp lo hi id v tv,
val_inject f v tv ->
match_env f cenv e le m te sp lo hi ->
match_env f cenv e (
PTree.set id v le)
m (
PTree.set (
for_temp id)
tv te)
sp lo hi.
Proof.
intros.
inv H0.
constructor;
auto.
intros.
geninv (
me_vars0 id0).
econstructor;
eauto.
rewrite PTree.gso.
auto.
unfold for_var,
for_temp;
congruence.
econstructor;
eauto.
econstructor;
eauto.
econstructor;
eauto.
econstructor;
eauto.
intros.
rewrite PTree.gsspec in H0.
destruct (
peq id0 id).
inv H0.
exists tv;
split;
auto.
apply PTree.gss.
rewrite PTree.gso.
eauto.
unfold for_temp;
congruence.
Qed.
The match_env relation is preserved by any memory operation
that preserves sizes and loads from blocks in the lo, hi range.
Lemma match_env_invariant:
forall f cenv e le m1 m2 te sp lo hi,
(
forall b ofs chunk v,
lo <=
b <
hi ->
Mem.load chunk m1 b ofs =
Some v ->
Mem.load chunk m2 b ofs =
Some v) ->
(
forall b ofs p,
lo <=
b <
hi ->
Mem.perm m2 b ofs Max p ->
Mem.perm m1 b ofs Max p) ->
match_env f cenv e le m1 te sp lo hi ->
match_env f cenv e le m2 te sp lo hi.
Proof.
intros. inv H1. constructor; eauto.
intros. geninv (me_vars0 id); econstructor; eauto.
Qed.
match_env is insensitive to the Cminor values of stack-allocated data.
Lemma match_env_extensional:
forall f cenv e le m te1 sp lo hi te2,
match_env f cenv e le m te1 sp lo hi ->
(
forall id chunk,
cenv!!
id =
Var_local chunk ->
te2!(
for_var id) =
te1!(
for_var id)) ->
(
forall id v,
le!
id =
Some v ->
te2!(
for_temp id) =
te1!(
for_temp id)) ->
match_env f cenv e le m te2 sp lo hi.
Proof.
intros. inv H; econstructor; eauto.
intros. geninv (me_vars0 id); econstructor; eauto. rewrite <- H6. eauto.
intros. rewrite (H1 _ _ H). auto.
Qed.
match_env and allocations
Inductive alloc_condition:
var_info ->
var_kind ->
block ->
option (
block *
Z) ->
Prop :=
|
alloc_cond_local:
forall chunk sp,
alloc_condition (
Var_local chunk) (
Vscalar chunk)
sp None
|
alloc_cond_stack_scalar:
forall chunk pos sp,
alloc_condition (
Var_stack_scalar chunk pos) (
Vscalar chunk)
sp (
Some(
sp,
pos))
|
alloc_cond_stack_array:
forall pos sz al sp,
alloc_condition (
Var_stack_array pos sz al) (
Varray sz al)
sp (
Some(
sp,
pos)).
Lemma match_env_alloc_same:
forall f1 cenv e le m1 te sp lo lv m2 b f2 id info tv,
match_env f1 cenv e le m1 te sp lo (
Mem.nextblock m1) ->
Mem.alloc m1 0 (
sizeof lv) = (
m2,
b) ->
inject_incr f1 f2 ->
alloc_condition info lv sp (
f2 b) ->
(
forall b',
b' <>
b ->
f2 b' =
f1 b') ->
te!(
for_var id) =
Some tv ->
e!
id =
None ->
match_env f2 (
PMap.set id info cenv) (
PTree.set id (
b,
lv)
e)
le m2 te sp lo (
Mem.nextblock m2).
Proof.
Lemma match_env_alloc_other:
forall f1 cenv e le m1 te sp lo hi sz m2 b f2,
match_env f1 cenv e le m1 te sp lo hi ->
Mem.alloc m1 0
sz = (
m2,
b) ->
inject_incr f1 f2 ->
(
forall b',
b' <>
b ->
f2 b' =
f1 b') ->
hi <=
b ->
match f2 b with None =>
True |
Some(
b',
ofs) =>
sp <
b'
end ->
match_env f2 cenv e le m2 te sp lo hi.
Proof.
intros until f2;
intros ME ALLOC INCR OTHER BOUND TBOUND.
inv ME.
assert (
BELOW:
forall id b'
lv,
e!
id =
Some(
b',
lv) ->
b' <>
b).
intros.
exploit me_bounded0;
eauto.
exploit Mem.alloc_result;
eauto.
unfold block in *;
omega.
econstructor;
eauto.
intros.
geninv (
me_vars0 id);
econstructor.
eauto.
eapply Mem.load_alloc_other;
eauto.
rewrite OTHER;
eauto.
eauto.
eapply val_inject_incr;
eauto.
eauto.
eapply val_inject_incr;
eauto.
eauto.
eapply val_inject_incr;
eauto.
auto.
auto.
auto.
intros.
exploit me_temps0;
eauto.
intros [
v' [
A B]].
exists v';
split;
auto.
eapply val_inject_incr;
eauto.
intros.
rewrite OTHER in H.
eauto.
red;
intro;
subst b0.
rewrite H in TBOUND.
omegaContradiction.
intros.
eapply me_incr0;
eauto.
rewrite <-
OTHER;
eauto.
exploit Mem.alloc_result;
eauto.
unfold block in *;
omega.
intros.
exploit Mem.perm_alloc_inv.
eexact ALLOC.
eauto.
rewrite zeq_false.
eauto.
exploit me_bounded0;
eauto.
Qed.
match_env and external calls
Remark inject_incr_separated_same:
forall f1 f2 m1 m1',
inject_incr f1 f2 ->
inject_separated f1 f2 m1 m1' ->
forall b,
Mem.valid_block m1 b ->
f2 b =
f1 b.
Proof.
intros. case_eq (f1 b).
intros [b' delta] EQ. apply H; auto.
intros EQ. case_eq (f2 b).
intros [b'1 delta1] EQ1. exploit H0; eauto. intros [C D]. contradiction.
auto.
Qed.
Remark inject_incr_separated_same':
forall f1 f2 m1 m1',
inject_incr f1 f2 ->
inject_separated f1 f2 m1 m1' ->
forall b b'
delta,
f2 b =
Some(
b',
delta) ->
Mem.valid_block m1'
b' ->
f1 b =
Some(
b',
delta).
Proof.
intros. case_eq (f1 b).
intros [b'1 delta1] EQ. exploit H; eauto. congruence.
intros. exploit H0; eauto. intros [C D]. contradiction.
Qed.
Lemma match_env_external_call:
forall f1 cenv e le m1 te sp lo hi m2 f2 m1',
match_env f1 cenv e le m1 te sp lo hi ->
mem_unchanged_on (
loc_unmapped f1)
m1 m2 ->
inject_incr f1 f2 ->
inject_separated f1 f2 m1 m1' ->
(
forall b ofs p,
Mem.valid_block m1 b ->
Mem.perm m2 b ofs Max p ->
Mem.perm m1 b ofs Max p) ->
hi <=
Mem.nextblock m1 ->
sp <
Mem.nextblock m1' ->
match_env f2 cenv e le m2 te sp lo hi.
Proof.
intros until m1'.
intros ME UNCHANGED INCR SEPARATED BOUNDS VALID VALID'.
destruct UNCHANGED as [
UNCHANGED1 UNCHANGED2].
inversion ME.
constructor;
auto.
intros.
geninv (
me_vars0 id);
try (
econstructor;
eauto;
fail).
econstructor.
eauto.
apply UNCHANGED2;
eauto.
rewrite <-
H3.
eapply inject_incr_separated_same;
eauto.
red.
exploit me_bounded0;
eauto.
omega.
eauto.
eauto.
intros.
exploit me_temps0;
eauto.
intros [
v' [
A B]].
exists v';
split;
auto.
eapply val_inject_incr;
eauto.
intros.
apply me_inv0 with delta.
eapply inject_incr_separated_same';
eauto.
intros.
exploit inject_incr_separated_same;
eauto.
instantiate (1 :=
b).
red;
omega.
intros.
apply me_incr0 with b delta.
congruence.
auto.
intros.
eapply me_bounds0;
eauto.
eapply BOUNDS;
eauto.
red.
exploit me_bounded0;
eauto.
omega.
Qed.
Invariant on abstract call stacks
Call stacks represent abstractly the execution state of the current
Csharpminor and Cminor functions, as well as the states of the
calling functions. A call stack is a list of frames, each frame
collecting information on the current execution state of a Csharpminor
function and its Cminor translation.
Inductive frame :
Type :=
Frame(
cenv:
compilenv)
(
tf:
Cminor.function)
(
e:
Csharpminor.env)
(
le:
Csharpminor.temp_env)
(
te:
Cminor.env)
(
sp:
block)
(
lo hi:
Z).
Definition callstack :
Type :=
list frame.
Global environments match if the memory injection f leaves unchanged
the references to global symbols and functions.
Inductive match_globalenvs (
f:
meminj) (
bound:
Z):
Prop :=
|
mk_match_globalenvs
(
POS:
bound > 0)
(
DOMAIN:
forall b,
b <
bound ->
f b =
Some(
b, 0))
(
IMAGE:
forall b1 b2 delta,
f b1 =
Some(
b2,
delta) ->
b2 <
bound ->
b1 =
b2)
(
SYMBOLS:
forall id b,
Genv.find_symbol ge id =
Some b ->
b <
bound)
(
INFOS:
forall b gv,
Genv.find_var_info ge b =
Some gv ->
b <
bound).
Matching of call stacks imply:
-
matching of environments for each of the frames
-
matching of the global environments
-
separation conditions over the memory blocks allocated for C#minor local variables;
-
separation conditions over the memory blocks allocated for Cminor stack data;
-
freeable permissions on the parts of the Cminor stack data blocks
that are not images of C#minor local variable blocks.
Definition padding_freeable (
f:
meminj) (
e:
Csharpminor.env) (
tm:
mem) (
sp:
block) (
sz:
Z) :
Prop :=
forall ofs,
0 <=
ofs <
sz ->
Mem.perm tm sp ofs Cur Freeable
\/
exists id,
exists b,
exists lv,
exists delta,
e!
id =
Some(
b,
lv) /\
f b =
Some(
sp,
delta) /\
delta <=
ofs <
delta +
sizeof lv.
Inductive match_callstack (
f:
meminj) (
m:
mem) (
tm:
mem):
callstack ->
Z ->
Z ->
Prop :=
|
mcs_nil:
forall hi bound tbound,
match_globalenvs f hi ->
hi <=
bound ->
hi <=
tbound ->
match_callstack f m tm nil bound tbound
|
mcs_cons:
forall cenv tf e le te sp lo hi cs bound tbound
(
BOUND:
hi <=
bound)
(
TBOUND:
sp <
tbound)
(
MENV:
match_env f cenv e le m te sp lo hi)
(
PERM:
padding_freeable f e tm sp tf.(
fn_stackspace))
(
MCS:
match_callstack f m tm cs lo sp),
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs)
bound tbound.
match_callstack implies match_globalenvs.
Lemma match_callstack_match_globalenvs:
forall f m tm cs bound tbound,
match_callstack f m tm cs bound tbound ->
exists hi,
match_globalenvs f hi.
Proof.
induction 1; eauto.
Qed.
We now show invariance properties for match_callstack that
generalize those for match_env.
Lemma padding_freeable_invariant:
forall f1 m1 tm1 sp sz cenv e le te lo hi f2 tm2,
padding_freeable f1 e tm1 sp sz ->
match_env f1 cenv e le m1 te sp lo hi ->
(
forall ofs,
Mem.perm tm1 sp ofs Cur Freeable ->
Mem.perm tm2 sp ofs Cur Freeable) ->
(
forall b,
b <
hi ->
f2 b =
f1 b) ->
padding_freeable f2 e tm2 sp sz.
Proof.
intros;
red;
intros.
exploit H;
eauto.
intros [
A | [
id [
b [
lv [
delta [
A [
B C]]]]]]].
left;
auto.
exploit me_bounded;
eauto.
intros [
D E].
right;
exists id;
exists b;
exists lv;
exists delta;
split.
auto.
rewrite H2;
auto.
Qed.
Lemma match_callstack_store_mapped:
forall f m tm chunk b b'
delta ofs ofs'
v tv m'
tm',
f b =
Some(
b',
delta) ->
Mem.store chunk m b ofs v =
Some m' ->
Mem.store chunk tm b'
ofs'
tv =
Some tm' ->
forall cs lo hi,
match_callstack f m tm cs lo hi ->
match_callstack f m'
tm'
cs lo hi.
Proof.
Lemma match_callstack_storev_mapped:
forall f m tm chunk a ta v tv m'
tm',
val_inject f a ta ->
Mem.storev chunk m a v =
Some m' ->
Mem.storev chunk tm ta tv =
Some tm' ->
forall cs lo hi,
match_callstack f m tm cs lo hi ->
match_callstack f m'
tm'
cs lo hi.
Proof.
Lemma match_callstack_invariant:
forall f m tm cs bound tbound,
match_callstack f m tm cs bound tbound ->
forall m'
tm',
(
forall cenv e le te sp lo hi,
hi <=
bound ->
match_env f cenv e le m te sp lo hi ->
match_env f cenv e le m'
te sp lo hi) ->
(
forall b ofs k p,
b <
tbound ->
Mem.perm tm b ofs k p ->
Mem.perm tm'
b ofs k p) ->
match_callstack f m'
tm'
cs bound tbound.
Proof.
induction 1;
intros.
econstructor;
eauto.
constructor;
auto.
eapply padding_freeable_invariant;
eauto.
eapply IHmatch_callstack;
eauto.
intros.
eapply H0;
eauto.
inv MENV;
omega.
intros.
apply H1;
auto.
inv MENV;
omega.
Qed.
Lemma match_callstack_store_local:
forall f cenv e le te sp lo hi cs bound tbound m1 m2 tm tf id b chunk v tv,
e!
id =
Some(
b,
Vscalar chunk) ->
val_normalized v chunk ->
val_inject f v tv ->
Mem.store chunk m1 b 0
v =
Some m2 ->
match_callstack f m1 tm (
Frame cenv tf e le te sp lo hi ::
cs)
bound tbound ->
match_callstack f m2 tm (
Frame cenv tf e le (
PTree.set (
for_var id)
tv te)
sp lo hi ::
cs)
bound tbound.
Proof.
A variant of match_callstack_store_local where the Cminor environment
te already associates to id a value that matches the assigned value.
In this case, match_callstack is preserved even if no assignment
takes place on the Cminor side.
Lemma match_callstack_store_local_unchanged:
forall f cenv e le te sp lo hi cs bound tbound m1 m2 id b chunk v tv tf tm,
e!
id =
Some(
b,
Vscalar chunk) ->
val_normalized v chunk ->
val_inject f v tv ->
Mem.store chunk m1 b 0
v =
Some m2 ->
te!(
for_var id) =
Some tv ->
match_callstack f m1 tm (
Frame cenv tf e le te sp lo hi ::
cs)
bound tbound ->
match_callstack f m2 tm (
Frame cenv tf e le te sp lo hi ::
cs)
bound tbound.
Proof.
Lemma match_callstack_set_temp:
forall f cenv e le te sp lo hi cs bound tbound m tm tf id v tv,
val_inject f v tv ->
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs)
bound tbound ->
match_callstack f m tm (
Frame cenv tf e (
PTree.set id v le) (
PTree.set (
for_temp id)
tv te)
sp lo hi ::
cs)
bound tbound.
Proof.
Lemma match_callstack_incr_bound:
forall f m tm cs bound tbound bound'
tbound',
match_callstack f m tm cs bound tbound ->
bound <=
bound' ->
tbound <=
tbound' ->
match_callstack f m tm cs bound'
tbound'.
Proof.
intros. inv H.
econstructor; eauto. omega. omega.
constructor; auto. omega. omega.
Qed.
Preservation of match_callstack by freeing all blocks allocated
for local variables at function entry (on the Csharpminor side)
and simultaneously freeing the Cminor stack data block.
Lemma in_blocks_of_env:
forall e id b lv,
e!
id =
Some(
b,
lv) ->
In (
b, 0,
sizeof lv) (
blocks_of_env e).
Proof.
Lemma in_blocks_of_env_inv:
forall b lo hi e,
In (
b,
lo,
hi) (
blocks_of_env e) ->
exists id,
exists lv,
e!
id =
Some(
b,
lv) /\
lo = 0 /\
hi =
sizeof lv.
Proof.
unfold blocks_of_env;
intros.
exploit list_in_map_inv;
eauto.
intros [[
id [
b'
lv]] [
A B]].
unfold block_of_binding in A.
inv A.
exists id;
exists lv;
intuition.
apply PTree.elements_complete.
auto.
Qed.
Lemma match_callstack_freelist:
forall f cenv tf e le te sp lo hi cs m m'
tm,
Mem.inject f m tm ->
Mem.free_list m (
blocks_of_env e) =
Some m' ->
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m) (
Mem.nextblock tm) ->
exists tm',
Mem.free tm sp 0
tf.(
fn_stackspace) =
Some tm'
/\
match_callstack f m'
tm'
cs (
Mem.nextblock m') (
Mem.nextblock tm')
/\
Mem.inject f m'
tm'.
Proof.
Preservation of match_callstack by allocations.
Lemma match_callstack_alloc_below:
forall f1 m1 tm sz m2 b f2,
Mem.alloc m1 0
sz = (
m2,
b) ->
inject_incr f1 f2 ->
(
forall b',
b' <>
b ->
f2 b' =
f1 b') ->
forall cs bound tbound,
match_callstack f1 m1 tm cs bound tbound ->
bound <=
b ->
match f2 b with None =>
True |
Some(
b',
ofs) =>
tbound <=
b'
end ->
match_callstack f2 m2 tm cs bound tbound.
Proof.
induction 4;
intros.
apply mcs_nil with hi;
auto.
inv H2.
constructor;
auto.
intros.
destruct (
eq_block b1 b).
subst.
rewrite H2 in H6.
omegaContradiction.
rewrite H1 in H2;
eauto.
constructor;
auto.
eapply match_env_alloc_other;
eauto.
omega.
destruct (
f2 b);
auto.
destruct p;
omega.
eapply padding_freeable_invariant;
eauto.
intros.
apply H1.
unfold block;
omega.
apply IHmatch_callstack.
inv MENV;
omega.
destruct (
f2 b);
auto.
destruct p;
omega.
Qed.
Lemma match_callstack_alloc_left:
forall f1 m1 tm cenv tf e le te sp lo cs lv m2 b f2 info id tv,
match_callstack f1 m1 tm
(
Frame cenv tf e le te sp lo (
Mem.nextblock m1) ::
cs)
(
Mem.nextblock m1) (
Mem.nextblock tm) ->
Mem.alloc m1 0 (
sizeof lv) = (
m2,
b) ->
inject_incr f1 f2 ->
alloc_condition info lv sp (
f2 b) ->
(
forall b',
b' <>
b ->
f2 b' =
f1 b') ->
te!(
for_var id) =
Some tv ->
e!
id =
None ->
match_callstack f2 m2 tm
(
Frame (
PMap.set id info cenv)
tf (
PTree.set id (
b,
lv)
e)
le te sp lo (
Mem.nextblock m2) ::
cs)
(
Mem.nextblock m2) (
Mem.nextblock tm).
Proof.
intros until tv;
intros MCS ALLOC INCR ACOND OTHER TE E.
inv MCS.
exploit Mem.alloc_result;
eauto.
intro RESULT.
exploit Mem.nextblock_alloc;
eauto.
intro NEXT.
constructor.
omega.
auto.
eapply match_env_alloc_same;
eauto.
red;
intros.
exploit PERM;
eauto.
intros [
A | [
id' [
b' [
lv' [
delta' [
A [
B C]]]]]]].
left;
auto.
right;
exists id';
exists b';
exists lv';
exists delta'.
split.
rewrite PTree.gso;
auto.
congruence.
split.
apply INCR;
auto.
auto.
eapply match_callstack_alloc_below;
eauto.
inv MENV.
unfold block in *;
omega.
inv ACOND.
auto.
omega.
omega.
Qed.
Lemma match_callstack_alloc_right:
forall f m tm cs tf sp tm'
te,
match_callstack f m tm cs (
Mem.nextblock m) (
Mem.nextblock tm) ->
Mem.alloc tm 0
tf.(
fn_stackspace) = (
tm',
sp) ->
Mem.inject f m tm ->
match_callstack f m tm'
(
Frame gce tf empty_env empty_temp_env te sp (
Mem.nextblock m) (
Mem.nextblock m) ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm').
Proof.
Decidability of the predicate "this is not a padding location"
Definition is_reachable (
f:
meminj) (
e:
Csharpminor.env) (
sp:
block) (
ofs:
Z) :
Prop :=
exists id,
exists b,
exists lv,
exists delta,
e!
id =
Some(
b,
lv) /\
f b =
Some(
sp,
delta) /\
delta <=
ofs <
delta +
sizeof lv.
Lemma is_reachable_dec:
forall f e sp ofs,
is_reachable f e sp ofs \/ ~
is_reachable f e sp ofs.
Proof.
Preservation of match_callstack by external calls.
Lemma match_callstack_external_call:
forall f1 f2 m1 m2 m1'
m2',
mem_unchanged_on (
loc_unmapped f1)
m1 m2 ->
mem_unchanged_on (
loc_out_of_reach f1 m1)
m1'
m2' ->
inject_incr f1 f2 ->
inject_separated f1 f2 m1 m1' ->
(
forall b ofs p,
Mem.valid_block m1 b ->
Mem.perm m2 b ofs Max p ->
Mem.perm m1 b ofs Max p) ->
forall cs bound tbound,
match_callstack f1 m1 m1'
cs bound tbound ->
bound <=
Mem.nextblock m1 ->
tbound <=
Mem.nextblock m1' ->
match_callstack f2 m2 m2'
cs bound tbound.
Proof.
intros until m2'.
intros UNMAPPED OUTOFREACH INCR SEPARATED MAXPERMS.
destruct OUTOFREACH as [
OUTOFREACH1 OUTOFREACH2].
induction 1;
intros.
apply mcs_nil with hi;
auto.
inv H.
constructor;
auto.
intros.
case_eq (
f1 b1).
intros [
b2'
delta']
EQ.
rewrite (
INCR _ _ _ EQ)
in H.
inv H.
eauto.
intro EQ.
exploit SEPARATED;
eauto.
intros [
A B].
elim B.
red.
omega.
constructor.
auto.
auto.
eapply match_env_external_call;
eauto.
omega.
omega.
red;
intros.
destruct (
is_reachable_dec f1 e sp ofs).
destruct H3 as [
id [
b [
lv [
delta [
A [
B C]]]]]].
right;
exists id;
exists b;
exists lv;
exists delta.
split.
auto.
split.
apply INCR;
auto.
auto.
exploit PERM;
eauto.
intros [
A|
A];
try contradiction.
left.
apply OUTOFREACH1;
auto.
red;
intros.
red;
intros;
elim H3.
exploit me_inv;
eauto.
intros [
id [
lv B]].
exploit me_bounds;
eauto.
intros C.
red.
exists id;
exists b0;
exists lv;
exists delta.
intuition omega.
eapply IHmatch_callstack;
eauto.
inv MENV;
omega.
omega.
Qed.
Remark external_call_nextblock_incr:
forall ef vargs m1 t vres m2,
external_call ef ge vargs m1 t vres m2 ->
Mem.nextblock m1 <=
Mem.nextblock m2.
Proof.
Remark inj_preserves_globals:
forall f hi,
match_globalenvs f hi ->
meminj_preserves_globals ge f.
Proof.
intros. inv H.
split. intros. apply DOMAIN. eapply SYMBOLS. eauto.
split. intros. apply DOMAIN. eapply INFOS. eauto.
intros. symmetry. eapply IMAGE; eauto.
Qed.
Properties of compile-time approximations of values
Definition val_match_approx (
a:
approx) (
v:
val) :
Prop :=
match a with
|
Int1 =>
v =
Val.zero_ext 1
v
|
Int7 =>
v =
Val.zero_ext 8
v /\
v =
Val.sign_ext 8
v
|
Int8u =>
v =
Val.zero_ext 8
v
|
Int8s =>
v =
Val.sign_ext 8
v
|
Int15 =>
v =
Val.zero_ext 16
v /\
v =
Val.sign_ext 16
v
|
Int16u =>
v =
Val.zero_ext 16
v
|
Int16s =>
v =
Val.sign_ext 16
v
|
Float32 =>
v =
Val.singleoffloat v
|
Any =>
True
end.
Remark undef_match_approx:
forall a,
val_match_approx a Vundef.
Proof.
destruct a; simpl; auto.
Qed.
Lemma val_match_approx_increasing:
forall a1 a2 v,
Approx.bge a1 a2 =
true ->
val_match_approx a2 v ->
val_match_approx a1 v.
Proof.
assert (
A:
forall v,
v =
Val.zero_ext 8
v ->
v =
Val.zero_ext 16
v).
intros.
rewrite H.
destruct v;
simpl;
auto.
decEq.
symmetry.
apply Int.zero_ext_widen.
compute;
auto.
split.
omega.
compute;
auto.
assert (
B:
forall v,
v =
Val.sign_ext 8
v ->
v =
Val.sign_ext 16
v).
intros.
rewrite H.
destruct v;
simpl;
auto.
decEq.
symmetry.
apply Int.sign_ext_widen.
compute;
auto.
split.
omega.
compute;
auto.
assert (
C:
forall v,
v =
Val.zero_ext 8
v ->
v =
Val.sign_ext 16
v).
intros.
rewrite H.
destruct v;
simpl;
auto.
decEq.
symmetry.
apply Int.sign_zero_ext_widen.
compute;
auto.
split.
omega.
compute;
auto.
assert (
D:
forall v,
v =
Val.zero_ext 1
v ->
v =
Val.zero_ext 8
v).
intros.
rewrite H.
destruct v;
simpl;
auto.
decEq.
symmetry.
apply Int.zero_ext_widen.
compute;
auto.
split.
omega.
compute;
auto.
assert (
E:
forall v,
v =
Val.zero_ext 1
v ->
v =
Val.sign_ext 8
v).
intros.
rewrite H.
destruct v;
simpl;
auto.
decEq.
symmetry.
apply Int.sign_zero_ext_widen.
compute;
auto.
split.
omega.
compute;
auto.
intros.
unfold Approx.bge in H;
destruct a1;
try discriminate;
destruct a2;
simpl in *;
try discriminate;
intuition.
Qed.
Lemma approx_lub_ge_left:
forall x y,
Approx.bge (
Approx.lub x y)
x =
true.
Proof.
destruct x; destruct y; auto.
Qed.
Lemma approx_lub_ge_right:
forall x y,
Approx.bge (
Approx.lub x y)
y =
true.
Proof.
destruct x; destruct y; auto.
Qed.
Lemma approx_of_int_sound:
forall n,
val_match_approx (
Approx.of_int n) (
Vint n).
Proof.
Lemma approx_of_float_sound:
forall f,
val_match_approx (
Approx.of_float f) (
Vfloat f).
Proof.
Lemma approx_of_chunk_sound:
forall chunk m b ofs v,
Mem.load chunk m b ofs =
Some v ->
val_match_approx (
Approx.of_chunk chunk)
v.
Proof.
intros.
exploit Mem.load_cast;
eauto.
destruct chunk;
intros;
simpl;
auto.
Qed.
Lemma approx_of_unop_sound:
forall op v1 v a1,
eval_unop op v1 =
Some v ->
val_match_approx a1 v1 ->
val_match_approx (
Approx.unop op a1)
v.
Proof.
Lemma approx_bitwise_and_sound:
forall a1 v1 a2 v2,
val_match_approx a1 v1 ->
val_match_approx a2 v2 ->
val_match_approx (
Approx.bitwise_and a1 a2) (
Val.and v1 v2).
Proof.
Lemma approx_bitwise_or_sound:
forall (
sem_op:
val ->
val ->
val)
a1 v1 a2 v2,
(
forall a b c,
sem_op (
Val.and a (
Vint c)) (
Val.and b (
Vint c)) =
Val.and (
sem_op a b) (
Vint c)) ->
val_match_approx a1 v1 ->
val_match_approx a2 v2 ->
val_match_approx (
Approx.bitwise_or a1 a2) (
sem_op v1 v2).
Proof.
Lemma approx_of_binop_sound:
forall op v1 a1 v2 a2 m v,
eval_binop op v1 v2 m =
Some v ->
val_match_approx a1 v1 ->
val_match_approx a2 v2 ->
val_match_approx (
Approx.binop op a1 a2)
v.
Proof.
Lemma approx_unop_is_redundant_sound:
forall op a v,
Approx.unop_is_redundant op a =
true ->
val_match_approx a v ->
eval_unop op v =
Some v.
Proof.
Compatibility of evaluation functions with respect to memory injections.
Remark val_inject_val_of_bool:
forall f b,
val_inject f (
Val.of_bool b) (
Val.of_bool b).
Proof.
intros; destruct b; constructor.
Qed.
Remark val_inject_val_of_optbool:
forall f ob,
val_inject f (
Val.of_optbool ob) (
Val.of_optbool ob).
Proof.
intros; destruct ob; simpl. destruct b; constructor. constructor.
Qed.
Ltac TrivialExists :=
match goal with
| [ |-
exists y,
Some ?
x =
Some y /\
val_inject _ _ _ ] =>
exists x;
split; [
auto |
try(
econstructor;
eauto)]
| [ |-
exists y,
_ /\
val_inject _ (
Vint ?
x)
_ ] =>
exists (
Vint x);
split; [
eauto with evalexpr |
constructor]
| [ |-
exists y,
_ /\
val_inject _ (
Vfloat ?
x)
_ ] =>
exists (
Vfloat x);
split; [
eauto with evalexpr |
constructor]
|
_ =>
idtac
end.
Compatibility of eval_unop with respect to val_inject.
Lemma eval_unop_compat:
forall f op v1 tv1 v,
eval_unop op v1 =
Some v ->
val_inject f v1 tv1 ->
exists tv,
eval_unop op tv1 =
Some tv
/\
val_inject f v tv.
Proof.
destruct op;
simpl;
intros.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
apply val_inject_val_of_bool.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
apply val_inject_val_of_bool.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
inv H;
inv H0;
simpl;
TrivialExists.
inv H0;
simpl in H;
inv H.
simpl.
destruct (
Float.intoffloat f0);
simpl in *;
inv H1.
TrivialExists.
inv H0;
simpl in H;
inv H.
simpl.
destruct (
Float.intuoffloat f0);
simpl in *;
inv H1.
TrivialExists.
inv H0;
simpl in H;
inv H.
simpl.
TrivialExists.
inv H0;
simpl in H;
inv H.
simpl.
TrivialExists.
Qed.
Compatibility of eval_binop with respect to val_inject.
Lemma eval_binop_compat:
forall f op v1 tv1 v2 tv2 v m tm,
eval_binop op v1 v2 m =
Some v ->
val_inject f v1 tv1 ->
val_inject f v2 tv2 ->
Mem.inject f m tm ->
exists tv,
eval_binop op tv1 tv2 tm =
Some tv
/\
val_inject f v tv.
Proof.
Correctness of Cminor construction functions
Lemma make_stackaddr_correct:
forall sp te tm ofs,
eval_expr tge (
Vptr sp Int.zero)
te tm
(
make_stackaddr ofs) (
Vptr sp (
Int.repr ofs)).
Proof.
Lemma make_globaladdr_correct:
forall sp te tm id b,
Genv.find_symbol tge id =
Some b ->
eval_expr tge (
Vptr sp Int.zero)
te tm
(
make_globaladdr id) (
Vptr b Int.zero).
Proof.
intros;
unfold make_globaladdr.
eapply eval_Econst.
simpl.
rewrite H.
auto.
Qed.
Correctness of make_store.
Inductive val_lessdef_upto (
m:
int):
val ->
val ->
Prop :=
|
val_lessdef_upto_base:
forall v1 v2,
Val.lessdef v1 v2 ->
val_lessdef_upto m v1 v2
|
val_lessdef_upto_int:
forall n1 n2,
Int.and n1 m =
Int.and n2 m ->
val_lessdef_upto m (
Vint n1) (
Vint n2).
Hint Resolve val_lessdef_upto_base.
Remark val_lessdef_upto_and:
forall m v1 v2 p,
val_lessdef_upto m v1 v2 ->
Int.and p m =
m ->
val_lessdef_upto m (
Val.and v1 (
Vint p))
v2.
Proof.
Remark val_lessdef_upto_zero_ext:
forall m v1 v2 p,
val_lessdef_upto m v1 v2 ->
Int.and (
Int.repr (
two_p p - 1))
m =
m -> 0 <
p < 32 ->
val_lessdef_upto m (
Val.zero_ext p v1)
v2.
Proof.
Remark val_lessdef_upto_sign_ext:
forall m v1 v2 p,
val_lessdef_upto m v1 v2 ->
Int.and (
Int.repr (
two_p p - 1))
m =
m -> 0 <
p < 32 ->
val_lessdef_upto m (
Val.sign_ext p v1)
v2.
Proof.
Remark val_lessdef_upto_shru:
forall m v1 v2 p,
val_lessdef_upto (
Int.shl m p)
v1 v2 ->
Int.shru (
Int.shl m p)
p =
m ->
val_lessdef_upto m (
Val.shru v1 (
Vint p)) (
Val.shru v2 (
Vint p)).
Proof.
Remark val_lessdef_upto_shr:
forall m v1 v2 p,
val_lessdef_upto (
Int.shl m p)
v1 v2 ->
Int.shru (
Int.shl m p)
p =
m ->
val_lessdef_upto m (
Val.shr v1 (
Vint p)) (
Val.shr v2 (
Vint p)).
Proof.
Lemma eval_uncast_int:
forall m sp te tm a x,
eval_expr tge sp te tm a x ->
exists v,
eval_expr tge sp te tm (
uncast_int m a)
v /\
val_lessdef_upto m x v.
Proof.
assert (
EQ:
forall p q,
Int.eq p q =
true ->
p =
q).
intros.
generalize (
Int.eq_spec p q).
rewrite H;
auto.
intros until a.
functional induction (
uncast_int m a);
intros.
inv H.
simpl in H4;
inv H4.
exploit IHe;
eauto.
intros [
v [
A B]].
exists v;
split;
auto.
apply val_lessdef_upto_zero_ext;
auto.
compute;
auto.
exists x;
auto.
inv H.
simpl in H4;
inv H4.
exploit IHe;
eauto.
intros [
v [
A B]].
exists v;
split;
auto.
apply val_lessdef_upto_sign_ext;
auto.
compute;
auto.
exists x;
auto.
inv H.
simpl in H4;
inv H4.
exploit IHe;
eauto.
intros [
v [
A B]].
exists v;
split;
auto.
apply val_lessdef_upto_zero_ext;
auto.
compute;
auto.
exists x;
auto.
inv H.
simpl in H4;
inv H4.
exploit IHe;
eauto.
intros [
v [
A B]].
exists v;
split;
auto.
apply val_lessdef_upto_sign_ext;
auto.
compute;
auto.
exists x;
auto.
inv H.
simpl in H6;
inv H6.
inv H5.
simpl in H0.
inv H0.
exploit IHe;
eauto.
intros [
v [
A B]].
exists v;
split;
auto.
apply val_lessdef_upto_and;
auto.
exists x;
auto.
inv H.
simpl in H6;
inv H6.
inv H5.
simpl in H0.
inv H0.
exploit IHe;
eauto.
intros [
v [
A B]].
exists (
Val.shru v (
Vint n));
split.
econstructor.
eauto.
econstructor.
simpl;
reflexivity.
auto.
apply val_lessdef_upto_shru;
auto.
exists x;
auto.
inv H.
simpl in H6;
inv H6.
inv H5.
simpl in H0.
inv H0.
exploit IHe;
eauto.
intros [
v [
A B]].
exists (
Val.shr v (
Vint n));
split.
econstructor.
eauto.
econstructor.
simpl;
reflexivity.
auto.
apply val_lessdef_upto_shr;
auto.
exists x;
auto.
exists x;
split;
auto.
Qed.
Inductive val_lessdef_upto_single:
val ->
val ->
Prop :=
|
val_lessdef_upto_single_base:
forall v1 v2,
Val.lessdef v1 v2 ->
val_lessdef_upto_single v1 v2
|
val_lessdef_upto_single_float:
forall n1 n2,
Float.singleoffloat n1 =
Float.singleoffloat n2 ->
val_lessdef_upto_single (
Vfloat n1) (
Vfloat n2).
Hint Resolve val_lessdef_upto_single_base.
Lemma eval_uncast_float32:
forall sp te tm a x,
eval_expr tge sp te tm a x ->
exists v,
eval_expr tge sp te tm (
uncast_float32 a)
v /\
val_lessdef_upto_single x v.
Proof.
Inductive val_content_inject (
f:
meminj):
memory_chunk ->
val ->
val ->
Prop :=
|
val_content_inject_8_signed:
forall n1 n2,
Int.sign_ext 8
n1 =
Int.sign_ext 8
n2 ->
val_content_inject f Mint8signed (
Vint n1) (
Vint n2)
|
val_content_inject_8_unsigned:
forall n1 n2,
Int.zero_ext 8
n1 =
Int.zero_ext 8
n2 ->
val_content_inject f Mint8unsigned (
Vint n1) (
Vint n2)
|
val_content_inject_16_signed:
forall n1 n2,
Int.sign_ext 16
n1 =
Int.sign_ext 16
n2 ->
val_content_inject f Mint16signed (
Vint n1) (
Vint n2)
|
val_content_inject_16_unsigned:
forall n1 n2,
Int.zero_ext 16
n1 =
Int.zero_ext 16
n2 ->
val_content_inject f Mint16unsigned (
Vint n1) (
Vint n2)
|
val_content_inject_single:
forall n1 n2,
Float.singleoffloat n1 =
Float.singleoffloat n2 ->
val_content_inject f Mfloat32 (
Vfloat n1) (
Vfloat n2)
|
val_content_inject_base:
forall chunk v1 v2,
val_inject f v1 v2 ->
val_content_inject f chunk v1 v2.
Hint Resolve val_content_inject_base.
Lemma eval_store_arg:
forall f sp te tm a v va chunk,
eval_expr tge sp te tm a va ->
val_inject f v va ->
exists vb,
eval_expr tge sp te tm (
store_arg chunk a)
vb
/\
val_content_inject f chunk v vb.
Proof.
Lemma storev_mapped_content_inject:
forall f chunk m1 a1 v1 n1 m2 a2 v2,
Mem.inject f m1 m2 ->
Mem.storev chunk m1 a1 v1 =
Some n1 ->
val_inject f a1 a2 ->
val_content_inject f chunk v1 v2 ->
exists n2,
Mem.storev chunk m2 a2 v2 =
Some n2 /\
Mem.inject f n1 n2.
Proof.
Lemma make_store_correct:
forall f sp te tm addr tvaddr rhs tvrhs chunk m vaddr vrhs m'
fn k,
eval_expr tge sp te tm addr tvaddr ->
eval_expr tge sp te tm rhs tvrhs ->
Mem.storev chunk m vaddr vrhs =
Some m' ->
Mem.inject f m tm ->
val_inject f vaddr tvaddr ->
val_inject f vrhs tvrhs ->
exists tm',
exists tvrhs',
step tge (
State fn (
make_store chunk addr rhs)
k sp te tm)
E0 (
State fn Sskip k sp te tm')
/\
Mem.storev chunk tm tvaddr tvrhs' =
Some tm'
/\
Mem.inject f m'
tm'.
Proof.
Correctness of make_unop.
Lemma eval_make_unop:
forall sp te tm a v op v',
eval_expr tge sp te tm a v ->
eval_unop op v =
Some v' ->
exists v'',
eval_expr tge sp te tm (
make_unop op a)
v'' /\
Val.lessdef v'
v''.
Proof.
Lemma make_unop_correct:
forall f sp te tm a v op v'
tv,
eval_expr tge sp te tm a tv ->
eval_unop op v =
Some v' ->
val_inject f v tv ->
exists tv',
eval_expr tge sp te tm (
make_unop op a)
tv' /\
val_inject f v'
tv'.
Proof.
intros.
exploit eval_unop_compat;
eauto.
intros [
tv' [
A B]].
exploit eval_make_unop;
eauto.
intros [
tv'' [
C D]].
exists tv'';
split;
auto.
inv D.
auto.
inv B.
auto.
Qed.
Correctness of the variable accessors var_get, var_addr,
and var_set.
Lemma var_get_correct:
forall cenv id a app f tf e le te sp lo hi m cs tm b chunk v,
var_get cenv id =
OK (
a,
app) ->
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m) (
Mem.nextblock tm) ->
Mem.inject f m tm ->
eval_var_ref ge e id b chunk ->
Mem.load chunk m b 0 =
Some v ->
exists tv,
eval_expr tge (
Vptr sp Int.zero)
te tm a tv
/\
val_inject f v tv
/\
val_match_approx app v.
Proof.
Lemma var_addr_correct:
forall cenv id a app f tf e le te sp lo hi m cs tm b,
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m) (
Mem.nextblock tm) ->
var_addr cenv id =
OK (
a,
app) ->
eval_var_addr ge e id b ->
exists tv,
eval_expr tge (
Vptr sp Int.zero)
te tm a tv
/\
val_inject f (
Vptr b Int.zero)
tv
/\
val_match_approx app (
Vptr b Int.zero).
Proof.
Lemma var_set_correct:
forall cenv id rhs a f tf e le te sp lo hi m cs tm tv v m'
fn k,
var_set cenv id rhs =
OK a ->
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m) (
Mem.nextblock tm) ->
eval_expr tge (
Vptr sp Int.zero)
te tm rhs tv ->
val_inject f v tv ->
Mem.inject f m tm ->
exec_assign ge e m id v m' ->
exists te',
exists tm',
step tge (
State fn a k (
Vptr sp Int.zero)
te tm)
E0 (
State fn Sskip k (
Vptr sp Int.zero)
te'
tm') /\
Mem.inject f m'
tm' /\
match_callstack f m'
tm' (
Frame cenv tf e le te'
sp lo hi ::
cs) (
Mem.nextblock m') (
Mem.nextblock tm') /\
(
forall id',
id' <>
for_var id ->
te'!
id' =
te!
id').
Proof.
Lemma match_callstack_extensional:
forall f cenv tf e le te1 te2 sp lo hi cs bound tbound m tm,
(
forall id chunk,
cenv!!
id =
Var_local chunk ->
te2!(
for_var id) =
te1!(
for_var id)) ->
(
forall id v,
le!
id =
Some v ->
te2!(
for_temp id) =
te1!(
for_temp id)) ->
match_callstack f m tm (
Frame cenv tf e le te1 sp lo hi ::
cs)
bound tbound ->
match_callstack f m tm (
Frame cenv tf e le te2 sp lo hi ::
cs)
bound tbound.
Proof.
Lemma var_set_self_correct_scalar:
forall cenv id s a f tf e le te sp lo hi m cs tm tv v m'
fn k,
var_set_self cenv id s =
OK a ->
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m) (
Mem.nextblock tm) ->
val_inject f v tv ->
Mem.inject f m tm ->
exec_assign ge e m id v m' ->
te!(
for_var id) =
Some tv ->
exists tm',
star step tge (
State fn a k (
Vptr sp Int.zero)
te tm)
E0 (
State fn s k (
Vptr sp Int.zero)
te tm') /\
Mem.inject f m'
tm' /\
match_callstack f m'
tm' (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m') (
Mem.nextblock tm').
Proof.
Lemma var_set_self_correct_array:
forall cenv id s a f tf e le te sp lo hi m cs tm tv b v sz al m'
fn k,
var_set_self cenv id s =
OK a ->
match_callstack f m tm (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m) (
Mem.nextblock tm) ->
val_inject f v tv ->
Mem.inject f m tm ->
PTree.get id e =
Some(
b,
Varray sz al) ->
extcall_memcpy_sem sz (
Zmin al 4)
ge
(
Vptr b Int.zero ::
v ::
nil)
m E0 Vundef m' ->
te!(
for_var id) =
Some tv ->
exists f',
exists tm',
star step tge (
State fn a k (
Vptr sp Int.zero)
te tm)
E0 (
State fn s k (
Vptr sp Int.zero)
te tm') /\
Mem.inject f'
m'
tm' /\
match_callstack f'
m'
tm' (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m') (
Mem.nextblock tm') /\
inject_incr f f'.
Proof.
Correctness of stack allocation of local variables
This section shows the correctness of the translation of Csharpminor
local variables, either as Cminor local variables or as sub-blocks
of the Cminor stack data. This is the most difficult part of the proof.
Remark array_alignment_pos:
forall sz,
array_alignment sz > 0.
Proof.
unfold array_alignment;
intros.
destruct (
zlt sz 2).
omega.
destruct (
zlt sz 4).
omega.
destruct (
zlt sz 8);
omega.
Qed.
Remark assign_variable_incr:
forall atk id lv cenv sz cenv'
sz',
assign_variable atk (
id,
lv) (
cenv,
sz) = (
cenv',
sz') ->
sz <=
sz'.
Proof.
Remark assign_variables_incr:
forall atk vars cenv sz cenv'
sz',
assign_variables atk vars (
cenv,
sz) = (
cenv',
sz') ->
sz <=
sz'.
Proof.
Remark inj_offset_aligned_array:
forall stacksize sz,
Mem.inj_offset_aligned (
align stacksize (
array_alignment sz))
sz.
Proof.
intros;
red;
intros.
apply Zdivides_trans with (
array_alignment sz).
unfold align_chunk.
unfold array_alignment.
generalize Zone_divide;
intro.
generalize Zdivide_refl;
intro.
assert (2 | 4).
exists 2;
auto.
assert (2 | 8).
exists 4;
auto.
assert (4 | 8).
exists 2;
auto.
destruct (
zlt sz 2).
destruct chunk;
simpl in *;
auto;
omegaContradiction.
destruct (
zlt sz 4).
destruct chunk;
simpl in *;
auto;
omegaContradiction.
destruct (
zlt sz 8).
destruct chunk;
simpl in *;
auto;
omegaContradiction.
destruct chunk;
simpl;
auto.
apply align_divides.
apply array_alignment_pos.
Qed.
Remark inj_offset_aligned_array':
forall stacksize sz,
Mem.inj_offset_aligned (
align stacksize (
array_alignment sz)) (
Zmax 0
sz).
Proof.
Remark inj_offset_aligned_var:
forall stacksize chunk,
Mem.inj_offset_aligned (
align stacksize (
size_chunk chunk)) (
size_chunk chunk).
Proof.
Lemma match_callstack_alloc_variable:
forall atk id lv cenv sz cenv'
sz'
tm sp e tf m m'
b te le lo cs f tv,
assign_variable atk (
id,
lv) (
cenv,
sz) = (
cenv',
sz') ->
Mem.valid_block tm sp ->
(
forall ofs k p,
Mem.perm tm sp ofs k p -> 0 <=
ofs <
tf.(
fn_stackspace)) ->
Mem.range_perm tm sp 0
tf.(
fn_stackspace)
Cur Freeable ->
tf.(
fn_stackspace) <=
Int.max_unsigned ->
Mem.alloc m 0 (
sizeof lv) = (
m',
b) ->
match_callstack f m tm
(
Frame cenv tf e le te sp lo (
Mem.nextblock m) ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm) ->
Mem.inject f m tm ->
0 <=
sz ->
sz' <=
tf.(
fn_stackspace) ->
(
forall b delta ofs k p,
f b =
Some(
sp,
delta) ->
Mem.perm m b ofs k p ->
ofs +
delta <
sz) ->
e!
id =
None ->
te!(
for_var id) =
Some tv ->
exists f',
inject_incr f f'
/\
Mem.inject f'
m'
tm
/\
match_callstack f'
m'
tm
(
Frame cenv'
tf (
PTree.set id (
b,
lv)
e)
le te sp lo (
Mem.nextblock m') ::
cs)
(
Mem.nextblock m') (
Mem.nextblock tm)
/\ (
forall b delta ofs k p,
f'
b =
Some(
sp,
delta) ->
Mem.perm m'
b ofs k p ->
ofs +
delta <
sz').
Proof.
intros until tv.
intros ASV VALID BOUNDS PERMS NOOV ALLOC MCS INJ LO HI RANGE E TE.
generalize ASV.
unfold assign_variable.
caseEq lv.
intros chunk LV.
case (
Identset.mem id atk).
set (
ofs :=
align sz (
size_chunk chunk)).
intro EQ;
injection EQ;
intros;
clear EQ.
rewrite <-
H0.
generalize (
size_chunk_pos chunk);
intro SIZEPOS.
generalize (
align_le sz (
size_chunk chunk)
SIZEPOS).
fold ofs.
intro SZOFS.
exploit Mem.alloc_left_mapped_inject.
eauto.
eauto.
eauto.
instantiate (1 :=
ofs).
omega.
intros.
exploit BOUNDS;
eauto.
omega.
intros.
apply Mem.perm_implies with Freeable;
auto with mem.
apply Mem.perm_cur.
apply PERMS.
rewrite LV in H1.
simpl in H1.
omega.
rewrite LV;
simpl.
rewrite Zminus_0_r.
unfold ofs.
apply inj_offset_aligned_var.
intros.
generalize (
RANGE _ _ _ _ _ H1 H2).
omega.
intros [
f1 [
MINJ1 [
INCR1 [
SAME OTHER]]]].
exists f1;
split.
auto.
split.
auto.
split.
eapply match_callstack_alloc_left;
eauto.
rewrite <-
LV;
auto.
rewrite SAME;
constructor.
intros.
exploit Mem.perm_alloc_inv;
eauto.
destruct (
zeq b0 b).
subst b0.
assert (
delta =
ofs)
by congruence.
subst delta.
rewrite LV.
simpl.
omega.
intro.
rewrite OTHER in H1;
eauto.
generalize (
RANGE _ _ _ _ _ H1 H3).
omega.
intro EQ;
injection EQ;
intros;
clear EQ.
subst sz'.
rewrite <-
H0.
exploit Mem.alloc_left_unmapped_inject;
eauto.
intros [
f1 [
MINJ1 [
INCR1 [
SAME OTHER]]]].
exists f1;
split.
auto.
split.
auto.
split.
eapply match_callstack_alloc_left;
eauto.
rewrite <-
LV;
auto.
rewrite SAME;
constructor.
intros.
exploit Mem.perm_alloc_inv;
eauto.
destruct (
zeq b0 b).
subst b0.
congruence.
rewrite OTHER in H;
eauto.
intros dim al LV EQ.
injection EQ;
clear EQ;
intros.
rewrite <-
H.
assert (0 <=
Zmax 0
dim).
apply Zmax1.
generalize (
align_le sz (
array_alignment dim) (
array_alignment_pos dim)).
intro.
set (
ofs :=
align sz (
array_alignment dim))
in *.
exploit Mem.alloc_left_mapped_inject.
eauto.
eauto.
eauto.
instantiate (1 :=
ofs).
generalize Int.min_signed_neg.
omega.
intros.
exploit BOUNDS;
eauto.
generalize Int.min_signed_neg.
omega.
intros.
apply Mem.perm_implies with Freeable;
auto with mem.
apply Mem.perm_cur.
apply PERMS.
rewrite LV in H3.
simpl in H3.
omega.
rewrite LV;
simpl.
rewrite Zminus_0_r.
unfold ofs.
apply inj_offset_aligned_array'.
intros.
generalize (
RANGE _ _ _ _ _ H3 H4).
omega.
intros [
f1 [
MINJ1 [
INCR1 [
SAME OTHER]]]].
exists f1;
split.
auto.
split.
auto.
split.
subst cenv'.
eapply match_callstack_alloc_left;
eauto.
rewrite <-
LV;
auto.
rewrite SAME;
constructor.
intros.
exploit Mem.perm_alloc_inv;
eauto.
destruct (
zeq b0 b).
subst b0.
assert (
delta =
ofs)
by congruence.
subst delta.
rewrite LV.
simpl.
omega.
intro.
rewrite OTHER in H3;
eauto.
generalize (
RANGE _ _ _ _ _ H3 H5).
omega.
Qed.
Lemma match_callstack_alloc_variables_rec:
forall tm sp cenv'
tf le te lo cs atk,
Mem.valid_block tm sp ->
(
forall ofs k p,
Mem.perm tm sp ofs k p -> 0 <=
ofs <
tf.(
fn_stackspace)) ->
Mem.range_perm tm sp 0
tf.(
fn_stackspace)
Cur Freeable ->
tf.(
fn_stackspace) <=
Int.max_unsigned ->
forall e m vars e'
m',
alloc_variables e m vars e'
m' ->
forall f cenv sz,
assign_variables atk vars (
cenv,
sz) = (
cenv',
tf.(
fn_stackspace)) ->
match_callstack f m tm
(
Frame cenv tf e le te sp lo (
Mem.nextblock m) ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm) ->
Mem.inject f m tm ->
0 <=
sz ->
(
forall b delta ofs k p,
f b =
Some(
sp,
delta) ->
Mem.perm m b ofs k p ->
ofs +
delta <
sz) ->
(
forall id lv,
In (
id,
lv)
vars ->
te!(
for_var id) <>
None) ->
list_norepet (
List.map (@
fst ident var_kind)
vars) ->
(
forall id lv,
In (
id,
lv)
vars ->
e!
id =
None) ->
exists f',
inject_incr f f'
/\
Mem.inject f'
m'
tm
/\
match_callstack f'
m'
tm
(
Frame cenv'
tf e'
le te sp lo (
Mem.nextblock m') ::
cs)
(
Mem.nextblock m') (
Mem.nextblock tm).
Proof.
intros until atk.
intros VALID BOUNDS PERM NOOV.
induction 1.
intros.
simpl in H.
inversion H;
subst cenv sz.
exists f.
split.
apply inject_incr_refl.
split.
auto.
auto.
intros until sz.
change (
assign_variables atk ((
id,
lv) ::
vars) (
cenv,
sz))
with (
assign_variables atk vars (
assign_variable atk (
id,
lv) (
cenv,
sz))).
caseEq (
assign_variable atk (
id,
lv) (
cenv,
sz)).
intros cenv1 sz1 ASV1 ASVS MATCH MINJ SZPOS BOUND DEFINED NOREPET UNDEFINED.
assert (
DEFINED1:
forall id0 lv0,
In (
id0,
lv0)
vars ->
te!(
for_var id0) <>
None).
intros.
eapply DEFINED.
simpl.
right.
eauto.
assert (
exists tv,
te!(
for_var id) =
Some tv).
assert (
te!(
for_var id) <>
None).
eapply DEFINED.
simpl;
left;
auto.
destruct (
te!(
for_var id)).
exists v;
auto.
congruence.
destruct H1 as [
tv TEID].
assert (
sz1 <=
fn_stackspace tf).
eapply assign_variables_incr;
eauto.
exploit match_callstack_alloc_variable;
eauto with coqlib.
intros [
f1 [
INCR1 [
INJ1 [
MCS1 BOUND1]]]].
exploit IHalloc_variables;
eauto.
apply Zle_trans with sz;
auto.
eapply assign_variable_incr;
eauto.
inv NOREPET;
auto.
intros.
rewrite PTree.gso.
eapply UNDEFINED;
eauto with coqlib.
simpl in NOREPET.
inversion NOREPET.
red;
intro;
subst id0.
elim H5.
change id with (
fst (
id,
lv0)).
apply List.in_map.
auto.
intros [
f2 [
INCR2 [
INJ2 MCS2]]].
exists f2;
intuition.
eapply inject_incr_trans;
eauto.
Qed.
Lemma set_params_defined:
forall params args id,
In id params -> (
set_params args params)!
id <>
None.
Proof.
induction params;
simpl;
intros.
elim H.
destruct args.
rewrite PTree.gsspec.
case (
peq id a);
intro.
congruence.
eapply IHparams.
elim H;
intro.
congruence.
auto.
rewrite PTree.gsspec.
case (
peq id a);
intro.
congruence.
eapply IHparams.
elim H;
intro.
congruence.
auto.
Qed.
Lemma set_locals_defined:
forall e vars id,
In id vars \/
e!
id <>
None -> (
set_locals vars e)!
id <>
None.
Proof.
induction vars;
simpl;
intros.
tauto.
rewrite PTree.gsspec.
case (
peq id a);
intro.
congruence.
apply IHvars.
assert (
a <>
id).
congruence.
tauto.
Qed.
Lemma set_locals_params_defined:
forall args params vars id,
In id (
params ++
vars) ->
(
set_locals vars (
set_params args params))!
id <>
None.
Proof.
Preservation of match_callstack by simultaneous allocation
of Csharpminor local variables and of the Cminor stack data block.
Lemma match_callstack_alloc_variables:
forall fn cenv tf m e m'
tm tm'
sp f cs targs,
build_compilenv gce fn = (
cenv,
tf.(
fn_stackspace)) ->
tf.(
fn_stackspace) <=
Int.max_unsigned ->
list_norepet (
fn_params_names fn ++
fn_vars_names fn) ->
alloc_variables Csharpminor.empty_env m (
fn_variables fn)
e m' ->
Mem.alloc tm 0
tf.(
fn_stackspace) = (
tm',
sp) ->
match_callstack f m tm cs (
Mem.nextblock m) (
Mem.nextblock tm) ->
Mem.inject f m tm ->
let tparams :=
List.map for_var (
fn_params_names fn)
in
let tvars :=
List.map for_var (
fn_vars_names fn)
in
let ttemps :=
List.map for_temp (
Csharpminor.fn_temps fn)
in
let te :=
set_locals (
tvars ++
ttemps) (
set_params targs tparams)
in
exists f',
inject_incr f f'
/\
Mem.inject f'
m'
tm'
/\
match_callstack f'
m'
tm'
(
Frame cenv tf e empty_temp_env te sp (
Mem.nextblock m) (
Mem.nextblock m') ::
cs)
(
Mem.nextblock m') (
Mem.nextblock tm').
Proof.
Correctness of the code generated by store_parameters
to store in memory the values of parameters that are stack-allocated.
Inductive vars_vals_match (
f:
meminj):
list (
ident *
var_kind) ->
list val ->
env ->
Prop :=
|
vars_vals_nil:
forall te,
vars_vals_match f nil nil te
|
vars_vals_scalar:
forall te id chunk vars v vals tv,
te!(
for_var id) =
Some tv ->
val_inject f v tv ->
val_normalized v chunk ->
vars_vals_match f vars vals te ->
vars_vals_match f ((
id,
Vscalar chunk) ::
vars) (
v ::
vals)
te
|
vars_vals_array:
forall te id sz al vars v vals tv,
te!(
for_var id) =
Some tv ->
val_inject f v tv ->
vars_vals_match f vars vals te ->
vars_vals_match f ((
id,
Varray sz al) ::
vars) (
v ::
vals)
te.
Lemma vars_vals_match_extensional:
forall f vars vals te,
vars_vals_match f vars vals te ->
forall te',
(
forall id lv,
In (
id,
lv)
vars ->
te'!(
for_var id) =
te!(
for_var id)) ->
vars_vals_match f vars vals te'.
Proof.
induction 1; intros.
constructor.
econstructor; eauto.
rewrite <- H. eauto with coqlib.
apply IHvars_vals_match. intros. eapply H3; eauto with coqlib.
econstructor; eauto.
rewrite <- H. eauto with coqlib.
apply IHvars_vals_match. intros. eapply H2; eauto with coqlib.
Qed.
Lemma vars_vals_match_incr:
forall f f',
inject_incr f f' ->
forall vars vals te,
vars_vals_match f vars vals te ->
vars_vals_match f'
vars vals te.
Proof.
induction 2; intros; econstructor; eauto.
Qed.
Lemma store_parameters_correct:
forall e le te m1 params vl m2,
bind_parameters ge e m1 params vl m2 ->
forall s f cenv tf sp lo hi cs tm1 fn k,
vars_vals_match f params vl te ->
list_norepet (
List.map variable_name params) ->
Mem.inject f m1 tm1 ->
match_callstack f m1 tm1 (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m1) (
Mem.nextblock tm1) ->
store_parameters cenv params =
OK s ->
exists f',
exists tm2,
star step tge (
State fn s k (
Vptr sp Int.zero)
te tm1)
E0 (
State fn Sskip k (
Vptr sp Int.zero)
te tm2)
/\
Mem.inject f'
m2 tm2
/\
match_callstack f'
m2 tm2 (
Frame cenv tf e le te sp lo hi ::
cs) (
Mem.nextblock m2) (
Mem.nextblock tm2)
/\
inject_incr f f'.
Proof.
induction 1.
intros;
simpl.
monadInv H3.
exists f;
exists tm1.
split.
constructor.
auto.
intros until k.
intros VVM NOREPET MINJ MATCH STOREP.
monadInv STOREP.
inv VVM.
inv NOREPET.
exploit var_set_self_correct_scalar;
eauto.
econstructor;
eauto.
econstructor;
eauto.
intros [
tm2 [
EXEC1 [
MINJ1 MATCH1]]].
exploit IHbind_parameters;
eauto.
intros [
f' [
tm3 [
EXEC2 [
MINJ2 [
MATCH2 INCR2]]]]].
exists f';
exists tm3.
split.
eapply star_trans;
eauto.
auto.
intros until k.
intros VVM NOREPET MINJ MATCH STOREP.
monadInv STOREP.
inv VVM.
inv NOREPET.
exploit var_set_self_correct_array;
eauto.
intros [
f2 [
tm2 [
EXEC1 [
MINJ1 [
MATCH1 INCR1]]]]].
exploit IHbind_parameters.
eapply vars_vals_match_incr;
eauto.
auto.
eauto.
eauto.
eauto.
intros [
f3 [
tm3 [
EXEC2 [
MINJ2 [
MATCH2 INCR2]]]]].
exists f3;
exists tm3.
split.
eapply star_trans;
eauto.
split.
auto.
split.
auto.
eapply inject_incr_trans;
eauto.
Qed.
Definition val_normalized' (
v:
val) (
vk:
var_kind) :
Prop :=
match vk with
|
Vscalar chunk =>
val_normalized v chunk
|
Varray _ _ =>
True
end.
Lemma vars_vals_match_holds_1:
forall f params args targs,
list_norepet (
List.map variable_name params) ->
val_list_inject f args targs ->
list_forall2 val_normalized'
args (
List.map variable_kind params) ->
vars_vals_match f params args
(
set_params targs (
List.map for_var (
List.map variable_name params))).
Proof.
Lemma vars_vals_match_holds_2:
forall f params args e,
vars_vals_match f params args e ->
forall vl,
(
forall id1 id2,
In id1 (
List.map variable_name params) ->
In id2 vl ->
for_var id1 <>
id2) ->
vars_vals_match f params args (
set_locals vl e).
Proof.
Lemma vars_vals_match_holds:
forall f params args targs vars temps,
list_norepet (
List.map variable_name params ++
vars) ->
val_list_inject f args targs ->
list_forall2 val_normalized'
args (
List.map variable_kind params) ->
vars_vals_match f params args
(
set_locals (
List.map for_var vars ++
List.map for_temp temps)
(
set_params targs (
List.map for_var (
List.map variable_name params)))).
Proof.
Remark bind_parameters_normalized:
forall e m params args m',
bind_parameters ge e m params args m' ->
list_forall2 val_normalized'
args (
List.map variable_kind params).
Proof.
induction 1; simpl.
constructor.
constructor; auto.
constructor; auto. red; auto.
Qed.
The main result in this section: the behaviour of function entry
in the generated Cminor code (allocate stack data block and store
parameters whose address is taken) simulates what happens at function
entry in the original Csharpminor (allocate one block per local variable
and initialize the blocks corresponding to function parameters).
Lemma function_entry_ok:
forall fn m e m1 vargs m2 f cs tm cenv tf tm1 sp tvargs s fn'
k,
list_norepet (
fn_params_names fn ++
fn_vars_names fn) ->
alloc_variables empty_env m (
fn_variables fn)
e m1 ->
bind_parameters ge e m1 fn.(
Csharpminor.fn_params)
vargs m2 ->
match_callstack f m tm cs (
Mem.nextblock m) (
Mem.nextblock tm) ->
build_compilenv gce fn = (
cenv,
tf.(
fn_stackspace)) ->
tf.(
fn_stackspace) <=
Int.max_unsigned ->
Mem.alloc tm 0
tf.(
fn_stackspace) = (
tm1,
sp) ->
let tparams :=
List.map for_var (
fn_params_names fn)
in
let tvars :=
List.map for_var (
fn_vars_names fn)
in
let ttemps :=
List.map for_temp (
Csharpminor.fn_temps fn)
in
let te :=
set_locals (
tvars ++
ttemps) (
set_params tvargs tparams)
in
val_list_inject f vargs tvargs ->
Mem.inject f m tm ->
store_parameters cenv fn.(
Csharpminor.fn_params) =
OK s ->
exists f2,
exists tm2,
star step tge (
State fn'
s k (
Vptr sp Int.zero)
te tm1)
E0 (
State fn'
Sskip k (
Vptr sp Int.zero)
te tm2)
/\
Mem.inject f2 m2 tm2
/\
inject_incr f f2
/\
match_callstack f2 m2 tm2
(
Frame cenv tf e empty_temp_env te sp (
Mem.nextblock m) (
Mem.nextblock m1) ::
cs)
(
Mem.nextblock m2) (
Mem.nextblock tm2).
Proof.
Semantic preservation for the translation
The proof of semantic preservation uses simulation diagrams of the
following form:
e, m1, s ----------------- sp, te1, tm1, ts
| |
t| |t
v v
e, m2, out --------------- sp, te2, tm2, tout
where
ts is the Cminor statement obtained by translating the
Csharpminor statement
s. The left vertical arrow is an execution
of a Csharpminor statement. The right vertical arrow is an execution
of a Cminor statement. The precondition (top vertical bar)
includes a
mem_inject relation between the memory states
m1 and
tm1,
and a
match_callstack relation for any callstack having
e,
te1,
sp as top frame. The postcondition (bottom vertical bar)
is the existence of a memory injection
f2 that extends the injection
f1 we started with, preserves the
match_callstack relation for
the transformed callstack at the final state, and validates a
outcome_inject relation between the outcomes
out and
tout.
Semantic preservation for expressions
Remark bool_of_val_inject:
forall f v tv b,
Val.bool_of_val v b ->
val_inject f v tv ->
Val.bool_of_val tv b.
Proof.
intros. inv H0; inv H; constructor; auto.
Qed.
Lemma transl_constant_correct:
forall f sp cst v,
Csharpminor.eval_constant cst =
Some v ->
let (
tcst,
a) :=
transl_constant cst in
exists tv,
eval_constant tge sp tcst =
Some tv
/\
val_inject f v tv
/\
val_match_approx a v.
Proof.
Lemma transl_expr_correct:
forall f m tm cenv tf e le te sp lo hi cs
(
MINJ:
Mem.inject f m tm)
(
MATCH:
match_callstack f m tm
(
Frame cenv tf e le te sp lo hi ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm)),
forall a v,
Csharpminor.eval_expr ge e le m a v ->
forall ta app
(
TR:
transl_expr cenv a =
OK (
ta,
app)),
exists tv,
eval_expr tge (
Vptr sp Int.zero)
te tm ta tv
/\
val_inject f v tv
/\
val_match_approx app v.
Proof.
induction 3;
intros;
simpl in TR;
try (
monadInv TR).
eapply var_get_correct;
eauto.
inv MATCH.
inv MENV.
exploit me_temps0;
eauto.
intros [
tv [
A B]].
exists tv;
split.
constructor;
auto.
split.
auto.
exact I.
eapply var_addr_correct;
eauto.
exploit transl_constant_correct;
eauto.
destruct (
transl_constant cst)
as [
tcst a];
inv TR.
intros [
tv [
A [
B C]]].
exists tv;
split.
constructor;
eauto.
eauto.
exploit IHeval_expr;
eauto.
intros [
tv1 [
EVAL1 [
INJ1 APP1]]].
unfold Csharpminor.eval_unop in H0.
destruct (
Approx.unop_is_redundant op x0)
as []
_eqn;
inv EQ0.
exploit approx_unop_is_redundant_sound;
eauto.
intros.
replace v with v1 by congruence.
exists tv1;
auto.
exploit make_unop_correct;
eauto.
intros [
tv [
A B]].
exists tv;
split.
auto.
split.
auto.
eapply approx_of_unop_sound;
eauto.
exploit IHeval_expr1;
eauto.
intros [
tv1 [
EVAL1 [
INJ1 APP1]]].
exploit IHeval_expr2;
eauto.
intros [
tv2 [
EVAL2 [
INJ2 APP2]]].
exploit eval_binop_compat;
eauto.
intros [
tv [
EVAL INJ]].
exists tv;
split.
econstructor;
eauto.
split.
auto.
eapply approx_of_binop_sound;
eauto.
exploit IHeval_expr;
eauto.
intros [
tv1 [
EVAL1 [
INJ1 APP1]]].
exploit Mem.loadv_inject;
eauto.
intros [
tv [
LOAD INJ]].
exists tv;
split.
econstructor;
eauto.
split.
auto.
destruct v1;
simpl in H0;
try discriminate.
eapply approx_of_chunk_sound;
eauto.
exploit IHeval_expr1;
eauto.
intros [
tv1 [
EVAL1 [
INJ1 APP1]]].
assert (
transl_expr cenv (
if vb1 then b else c) =
OK ((
if vb1 then x1 else x3), (
if vb1 then x2 else x4))).
destruct vb1;
auto.
exploit IHeval_expr2;
eauto.
intros [
tv2 [
EVAL2 [
INJ2 APP2]]].
exists tv2;
split.
eapply eval_Econdition;
eauto.
eapply bool_of_val_inject;
eauto.
split.
auto.
apply val_match_approx_increasing with (
if vb1 then x2 else x4);
auto.
destruct vb1.
apply approx_lub_ge_left.
apply approx_lub_ge_right.
Qed.
Lemma transl_exprlist_correct:
forall f m tm cenv tf e le te sp lo hi cs
(
MINJ:
Mem.inject f m tm)
(
MATCH:
match_callstack f m tm
(
Frame cenv tf e le te sp lo hi ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm)),
forall a v,
Csharpminor.eval_exprlist ge e le m a v ->
forall ta
(
TR:
transl_exprlist cenv a =
OK ta),
exists tv,
eval_exprlist tge (
Vptr sp Int.zero)
te tm ta tv
/\
val_list_inject f v tv.
Proof.
induction 3;
intros;
monadInv TR.
exists (@
nil val);
split.
constructor.
constructor.
exploit transl_expr_correct;
eauto.
intros [
tv1 [
EVAL1 [
VINJ1 APP1]]].
exploit IHeval_exprlist;
eauto.
intros [
tv2 [
EVAL2 VINJ2]].
exists (
tv1 ::
tv2);
split.
constructor;
auto.
constructor;
auto.
Qed.
Semantic preservation for statements and functions
Inductive match_cont:
Csharpminor.cont ->
Cminor.cont ->
option typ ->
compilenv ->
exit_env ->
callstack ->
Prop :=
|
match_Kstop:
forall ty cenv xenv,
match_cont Csharpminor.Kstop Kstop ty cenv xenv nil
|
match_Kseq:
forall s k ts tk ty cenv xenv cs,
transl_stmt ty cenv xenv s =
OK ts ->
match_cont k tk ty cenv xenv cs ->
match_cont (
Csharpminor.Kseq s k) (
Kseq ts tk)
ty cenv xenv cs
|
match_Kseq2:
forall s1 s2 k ts1 tk ty cenv xenv cs,
transl_stmt ty cenv xenv s1 =
OK ts1 ->
match_cont (
Csharpminor.Kseq s2 k)
tk ty cenv xenv cs ->
match_cont (
Csharpminor.Kseq (
Csharpminor.Sseq s1 s2)
k)
(
Kseq ts1 tk)
ty cenv xenv cs
|
match_Kblock:
forall k tk ty cenv xenv cs,
match_cont k tk ty cenv xenv cs ->
match_cont (
Csharpminor.Kblock k) (
Kblock tk)
ty cenv (
true ::
xenv)
cs
|
match_Kblock2:
forall k tk ty cenv xenv cs,
match_cont k tk ty cenv xenv cs ->
match_cont k (
Kblock tk)
ty cenv (
false ::
xenv)
cs
|
match_Kcall:
forall optid fn e le k tfn sp te tk ty cenv xenv lo hi cs sz cenv',
transl_funbody cenv sz fn =
OK tfn ->
match_cont k tk fn.(
fn_return)
cenv xenv cs ->
match_cont (
Csharpminor.Kcall optid fn e le k)
(
Kcall (
option_map for_temp optid)
tfn (
Vptr sp Int.zero)
te tk)
ty cenv'
nil
(
Frame cenv tfn e le te sp lo hi ::
cs).
Inductive match_states:
Csharpminor.state ->
Cminor.state ->
Prop :=
|
match_state:
forall fn s k e le m tfn ts tk sp te tm cenv xenv f lo hi cs sz
(
TRF:
transl_funbody cenv sz fn =
OK tfn)
(
TR:
transl_stmt fn.(
fn_return)
cenv xenv s =
OK ts)
(
MINJ:
Mem.inject f m tm)
(
MCS:
match_callstack f m tm
(
Frame cenv tfn e le te sp lo hi ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm))
(
MK:
match_cont k tk fn.(
fn_return)
cenv xenv cs),
match_states (
Csharpminor.State fn s k e le m)
(
State tfn ts tk (
Vptr sp Int.zero)
te tm)
|
match_state_seq:
forall fn s1 s2 k e le m tfn ts1 tk sp te tm cenv xenv f lo hi cs sz
(
TRF:
transl_funbody cenv sz fn =
OK tfn)
(
TR:
transl_stmt fn.(
fn_return)
cenv xenv s1 =
OK ts1)
(
MINJ:
Mem.inject f m tm)
(
MCS:
match_callstack f m tm
(
Frame cenv tfn e le te sp lo hi ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm))
(
MK:
match_cont (
Csharpminor.Kseq s2 k)
tk fn.(
fn_return)
cenv xenv cs),
match_states (
Csharpminor.State fn (
Csharpminor.Sseq s1 s2)
k e le m)
(
State tfn ts1 tk (
Vptr sp Int.zero)
te tm)
|
match_callstate:
forall fd args k m tfd targs tk tm f cs cenv
(
TR:
transl_fundef gce fd =
OK tfd)
(
MINJ:
Mem.inject f m tm)
(
MCS:
match_callstack f m tm cs (
Mem.nextblock m) (
Mem.nextblock tm))
(
MK:
match_cont k tk (
Csharpminor.funsig fd).(
sig_res)
cenv nil cs)
(
ISCC:
Csharpminor.is_call_cont k)
(
ARGSINJ:
val_list_inject f args targs),
match_states (
Csharpminor.Callstate fd args k m)
(
Callstate tfd targs tk tm)
|
match_returnstate:
forall v k m tv tk tm f cs ty cenv
(
MINJ:
Mem.inject f m tm)
(
MCS:
match_callstack f m tm cs (
Mem.nextblock m) (
Mem.nextblock tm))
(
MK:
match_cont k tk ty cenv nil cs)
(
RESINJ:
val_inject f v tv),
match_states (
Csharpminor.Returnstate v k m)
(
Returnstate tv tk tm).
Remark val_inject_function_pointer:
forall bound v fd f tv,
Genv.find_funct tge v =
Some fd ->
match_globalenvs f bound ->
val_inject f v tv ->
tv =
v.
Proof.
Lemma match_call_cont:
forall k tk ty cenv xenv cs,
match_cont k tk ty cenv xenv cs ->
match_cont (
Csharpminor.call_cont k) (
call_cont tk)
ty cenv nil cs.
Proof.
induction 1; simpl; auto; econstructor; eauto.
Qed.
Lemma match_is_call_cont:
forall tfn te sp tm k tk ty cenv xenv cs,
match_cont k tk ty cenv xenv cs ->
Csharpminor.is_call_cont k ->
exists tk',
star step tge (
State tfn Sskip tk sp te tm)
E0 (
State tfn Sskip tk'
sp te tm)
/\
is_call_cont tk'
/\
match_cont k tk'
ty cenv nil cs.
Proof.
induction 1;
simpl;
intros;
try contradiction.
econstructor;
split.
apply star_refl.
split.
exact I.
econstructor;
eauto.
exploit IHmatch_cont;
eauto.
intros [
tk' [
A B]].
exists tk';
split.
eapply star_left;
eauto.
constructor.
traceEq.
auto.
econstructor;
split.
apply star_refl.
split.
exact I.
econstructor;
eauto.
Qed.
Properties of switch compilation
Remark switch_table_shift:
forall n sl base dfl,
switch_target n (
S dfl) (
switch_table sl (
S base)) =
S (
switch_target n dfl (
switch_table sl base)).
Proof.
induction sl;
intros;
simpl.
auto.
destruct (
Int.eq n i);
auto.
Qed.
Remark length_switch_table:
forall sl base1 base2,
length (
switch_table sl base1) =
length (
switch_table sl base2).
Proof.
induction sl; intros; simpl. auto. decEq; auto.
Qed.
Inductive transl_lblstmt_cont (
ty:
option typ) (
cenv:
compilenv) (
xenv:
exit_env):
lbl_stmt ->
cont ->
cont ->
Prop :=
|
tlsc_default:
forall s k ts,
transl_stmt ty cenv (
switch_env (
LSdefault s)
xenv)
s =
OK ts ->
transl_lblstmt_cont ty cenv xenv (
LSdefault s)
k (
Kblock (
Kseq ts k))
|
tlsc_case:
forall i s ls k ts k',
transl_stmt ty cenv (
switch_env (
LScase i s ls)
xenv)
s =
OK ts ->
transl_lblstmt_cont ty cenv xenv ls k k' ->
transl_lblstmt_cont ty cenv xenv (
LScase i s ls)
k (
Kblock (
Kseq ts k')).
Lemma switch_descent:
forall ty cenv xenv k ls body s,
transl_lblstmt ty cenv (
switch_env ls xenv)
ls body =
OK s ->
exists k',
transl_lblstmt_cont ty cenv xenv ls k k'
/\ (
forall f sp e m,
plus step tge (
State f s k sp e m)
E0 (
State f body k'
sp e m)).
Proof.
induction ls;
intros.
monadInv H.
econstructor;
split.
econstructor;
eauto.
intros.
eapply plus_left.
constructor.
apply star_one.
constructor.
traceEq.
monadInv H.
exploit IHls;
eauto.
intros [
k' [
A B]].
econstructor;
split.
econstructor;
eauto.
intros.
eapply plus_star_trans.
eauto.
eapply star_left.
constructor.
apply star_one.
constructor.
reflexivity.
traceEq.
Qed.
Lemma switch_ascent:
forall f n sp e m ty cenv xenv k ls k1,
let tbl :=
switch_table ls O in
let ls' :=
select_switch n ls in
transl_lblstmt_cont ty cenv xenv ls k k1 ->
exists k2,
star step tge (
State f (
Sexit (
switch_target n (
length tbl)
tbl))
k1 sp e m)
E0 (
State f (
Sexit O)
k2 sp e m)
/\
transl_lblstmt_cont ty cenv xenv ls'
k k2.
Proof.
induction ls;
intros;
unfold tbl,
ls';
simpl.
inv H.
econstructor;
split.
apply star_refl.
econstructor;
eauto.
simpl in H.
inv H.
rewrite Int.eq_sym.
destruct (
Int.eq i n).
econstructor;
split.
apply star_refl.
econstructor;
eauto.
exploit IHls;
eauto.
intros [
k2 [
A B]].
rewrite (
length_switch_table ls 1%
nat 0%
nat).
rewrite switch_table_shift.
econstructor;
split.
eapply star_left.
constructor.
eapply star_left.
constructor.
eexact A.
reflexivity.
traceEq.
exact B.
Qed.
Lemma switch_match_cont:
forall ty cenv xenv k cs tk ls tk',
match_cont k tk ty cenv xenv cs ->
transl_lblstmt_cont ty cenv xenv ls tk tk' ->
match_cont (
Csharpminor.Kseq (
seq_of_lbl_stmt ls)
k)
tk'
ty cenv (
false ::
switch_env ls xenv)
cs.
Proof.
Lemma transl_lblstmt_suffix:
forall n ty cenv xenv ls body ts,
transl_lblstmt ty cenv (
switch_env ls xenv)
ls body =
OK ts ->
let ls' :=
select_switch n ls in
exists body',
exists ts',
transl_lblstmt ty cenv (
switch_env ls'
xenv)
ls'
body' =
OK ts'.
Proof.
induction ls;
simpl;
intros.
monadInv H.
exists body;
econstructor.
rewrite EQ;
eauto.
simpl.
reflexivity.
monadInv H.
destruct (
Int.eq i n).
exists body;
econstructor.
simpl.
rewrite EQ;
simpl.
rewrite EQ0;
simpl.
reflexivity.
eauto.
Qed.
Lemma switch_match_states:
forall fn k e le m tfn ts tk sp te tm cenv xenv f lo hi cs sz ls body tk'
(
TRF:
transl_funbody cenv sz fn =
OK tfn)
(
TR:
transl_lblstmt (
fn_return fn)
cenv (
switch_env ls xenv)
ls body =
OK ts)
(
MINJ:
Mem.inject f m tm)
(
MCS:
match_callstack f m tm
(
Frame cenv tfn e le te sp lo hi ::
cs)
(
Mem.nextblock m) (
Mem.nextblock tm))
(
MK:
match_cont k tk (
fn_return fn)
cenv xenv cs)
(
TK:
transl_lblstmt_cont (
fn_return fn)
cenv xenv ls tk tk'),
exists S,
plus step tge (
State tfn (
Sexit O)
tk' (
Vptr sp Int.zero)
te tm)
E0 S
/\
match_states (
Csharpminor.State fn (
seq_of_lbl_stmt ls)
k e le m)
S.
Proof.
Commutation between find_label and compilation
Section FIND_LABEL.
Variable lbl:
label.
Variable ty:
option typ.
Variable cenv:
compilenv.
Variable cs:
callstack.
Remark find_label_var_set:
forall id e s k,
var_set cenv id e =
OK s ->
find_label lbl s k =
None.
Proof.
intros.
unfold var_set in H.
destruct (
cenv!!
id);
try (
monadInv H;
reflexivity).
Qed.
Remark find_label_var_set_self:
forall id s0 s k,
var_set_self cenv id s0 =
OK s ->
find_label lbl s k =
find_label lbl s0 k.
Proof.
intros.
unfold var_set_self in H.
destruct (
cenv!!
id);
try (
monadInv H;
reflexivity).
Qed.
Lemma transl_lblstmt_find_label_context:
forall xenv ls body ts tk1 tk2 ts'
tk',
transl_lblstmt ty cenv (
switch_env ls xenv)
ls body =
OK ts ->
transl_lblstmt_cont ty cenv xenv ls tk1 tk2 ->
find_label lbl body tk2 =
Some (
ts',
tk') ->
find_label lbl ts tk1 =
Some (
ts',
tk').
Proof.
induction ls; intros.
monadInv H. inv H0. simpl. simpl in H2. replace x with ts by congruence. rewrite H1. auto.
monadInv H. inv H0.
eapply IHls. eauto. eauto. simpl in H6. replace x with ts0 by congruence. simpl.
rewrite H1. auto.
Qed.
Lemma transl_find_label:
forall s k xenv ts tk,
transl_stmt ty cenv xenv s =
OK ts ->
match_cont k tk ty cenv xenv cs ->
match Csharpminor.find_label lbl s k with
|
None =>
find_label lbl ts tk =
None
|
Some(
s',
k') =>
exists ts',
exists tk',
exists xenv',
find_label lbl ts tk =
Some(
ts',
tk')
/\
transl_stmt ty cenv xenv'
s' =
OK ts'
/\
match_cont k'
tk'
ty cenv xenv'
cs
end
with transl_lblstmt_find_label:
forall ls xenv body k ts tk tk1,
transl_lblstmt ty cenv (
switch_env ls xenv)
ls body =
OK ts ->
match_cont k tk ty cenv xenv cs ->
transl_lblstmt_cont ty cenv xenv ls tk tk1 ->
find_label lbl body tk1 =
None ->
match Csharpminor.find_label_ls lbl ls k with
|
None =>
find_label lbl ts tk =
None
|
Some(
s',
k') =>
exists ts',
exists tk',
exists xenv',
find_label lbl ts tk =
Some(
ts',
tk')
/\
transl_stmt ty cenv xenv'
s' =
OK ts'
/\
match_cont k'
tk'
ty cenv xenv'
cs
end.
Proof.
intros.
destruct s;
try (
monadInv H);
simpl;
auto.
eapply find_label_var_set;
eauto.
exploit (
transl_find_label s1).
eauto.
eapply match_Kseq.
eexact EQ1.
eauto.
destruct (
Csharpminor.find_label lbl s1 (
Csharpminor.Kseq s2 k))
as [[
s'
k'] | ].
intros [
ts' [
tk' [
xenv' [
A [
B C]]]]].
exists ts';
exists tk';
exists xenv'.
intuition.
rewrite A;
auto.
intro.
rewrite H.
apply transl_find_label with xenv;
auto.
exploit (
transl_find_label s1).
eauto.
eauto.
destruct (
Csharpminor.find_label lbl s1 k)
as [[
s'
k'] | ].
intros [
ts' [
tk' [
xenv' [
A [
B C]]]]].
exists ts';
exists tk';
exists xenv'.
intuition.
rewrite A;
auto.
intro.
rewrite H.
apply transl_find_label with xenv;
auto.
apply transl_find_label with xenv.
auto.
econstructor;
eauto.
simpl.
rewrite EQ;
auto.
apply transl_find_label with (
true ::
xenv).
auto.
constructor;
auto.
exploit switch_descent;
eauto.
intros [
k' [
A B]].
eapply transl_lblstmt_find_label.
eauto.
eauto.
eauto.
reflexivity.
destruct o;
monadInv H;
auto.
destruct (
ident_eq lbl l).
exists x;
exists tk;
exists xenv;
auto.
apply transl_find_label with xenv;
auto.
intros.
destruct ls;
monadInv H;
simpl.
inv H1.
simpl in H3.
replace x with ts by congruence.
rewrite H2.
eapply transl_find_label;
eauto.
inv H1.
simpl in H7.
exploit (
transl_find_label s).
eauto.
eapply switch_match_cont;
eauto.
destruct (
Csharpminor.find_label lbl s (
Csharpminor.Kseq (
seq_of_lbl_stmt ls)
k))
as [[
s'
k''] | ].
intros [
ts' [
tk' [
xenv' [
A [
B C]]]]].
exists ts';
exists tk';
exists xenv';
intuition.
eapply transl_lblstmt_find_label_context;
eauto.
simpl.
replace x with ts0 by congruence.
rewrite H2.
auto.
intro.
eapply transl_lblstmt_find_label.
eauto.
auto.
eauto.
simpl.
replace x with ts0 by congruence.
rewrite H2.
auto.
Qed.
Remark find_label_store_parameters:
forall vars s k,
store_parameters cenv vars =
OK s ->
find_label lbl s k =
None.
Proof.
End FIND_LABEL.
Lemma transl_find_label_body:
forall cenv xenv size f tf k tk cs lbl s'
k',
transl_funbody cenv size f =
OK tf ->
match_cont k tk (
fn_return f)
cenv xenv cs ->
Csharpminor.find_label lbl f.(
Csharpminor.fn_body) (
Csharpminor.call_cont k) =
Some (
s',
k') ->
exists ts',
exists tk',
exists xenv',
find_label lbl tf.(
fn_body) (
call_cont tk) =
Some(
ts',
tk')
/\
transl_stmt (
fn_return f)
cenv xenv'
s' =
OK ts'
/\
match_cont k'
tk' (
fn_return f)
cenv xenv'
cs.
Proof.
The simulation diagram.
Fixpoint seq_left_depth (
s:
Csharpminor.stmt) :
nat :=
match s with
|
Csharpminor.Sseq s1 s2 =>
S (
seq_left_depth s1)
|
_ =>
O
end.
Definition measure (
S:
Csharpminor.state) :
nat :=
match S with
|
Csharpminor.State fn s k e le m =>
seq_left_depth s
|
_ =>
O
end.
Lemma transl_step_correct:
forall S1 t S2,
Csharpminor.step ge S1 t S2 ->
forall T1,
match_states S1 T1 ->
(
exists T2,
plus step tge T1 t T2 /\
match_states S2 T2)
\/ (
measure S2 <
measure S1 /\
t =
E0 /\
match_states S2 T1)%
nat.
Proof.
Lemma match_globalenvs_init:
forall m,
Genv.init_mem prog =
Some m ->
match_globalenvs (
Mem.flat_inj (
Mem.nextblock m)) (
Mem.nextblock m).
Proof.
Lemma transl_initial_states:
forall S,
Csharpminor.initial_state prog S ->
exists R,
Cminor.initial_state tprog R /\
match_states S R.
Proof.
Lemma transl_final_states:
forall S R r,
match_states S R ->
Csharpminor.final_state S r ->
Cminor.final_state R r.
Proof.
intros. inv H0. inv H. inv MK. inv RESINJ. constructor.
Qed.
Theorem transl_program_correct:
forward_simulation (
Csharpminor.semantics prog) (
Cminor.semantics tprog).
Proof.
End TRANSLATION.