Correctness proof for the translation from Linear to Mach.
This file proves semantic preservation for the Stacking pass.
Require Import Coqlib.
Require Import Errors.
Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Op.
Require Import Memory.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Locations.
Require Import LTL.
Require Import Linear.
Require Import Lineartyping.
Require Import Mach.
Require Import Bounds.
Require Import Conventions.
Require Import Stacklayout.
Require Import Stacking.
Require Import Annotations.
Properties of frame offsets
Lemma typesize_typesize:
forall ty,
AST.typesize ty = 4 *
Locations.typesize ty.
Proof.
destruct ty; auto.
Qed.
Remark size_type_chunk:
forall ty,
size_chunk (
chunk_of_type ty) =
AST.typesize ty.
Proof.
destruct ty; reflexivity.
Qed.
Section PRESERVATION.
Variable return_address_offset:
Mach.function ->
Mach.code ->
int ->
Prop.
Hypothesis return_address_offset_exists:
forall f sg ros c,
is_tail (
Mcall sg ros ::
c) (
fn_code f) ->
exists ofs,
return_address_offset f c ofs.
Let step :=
Mach.step return_address_offset.
Let step_safe :=
Mach.step_safe return_address_offset.
Variable prog:
Linear.program.
Variable tprog:
Mach.program.
Hypothesis TRANSF:
transf_program prog =
OK tprog.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Section FRAME_PROPERTIES.
Variable f:
Linear.function.
Let b :=
function_bounds f.
Let fe :=
make_env b.
Variable tf:
Mach.function.
Hypothesis TRANSF_F:
transf_function f =
OK tf.
Lemma unfold_transf_function:
tf =
Mach.mkfunction
f.(
Linear.fn_sig)
(
transl_body f fe)
fe.(
fe_stack_data)
fe.(
fe_size)
(
Int.repr fe.(
fe_ofs_link))
(
Int.repr fe.(
fe_ofs_retaddr)).
Proof.
Lemma transf_function_well_typed:
wt_function f =
true.
Proof.
Lemma size_no_overflow:
fe.(
fe_size) <=
Int.max_unsigned.
Proof.
Remark bound_stack_data_stacksize:
f.(
Linear.fn_stacksize) <=
b.(
bound_stack_data).
Proof.
A frame index is valid if it lies within the resource bounds
of the current function.
Definition index_valid (
idx:
frame_index) :=
match idx with
|
FI_link =>
True
|
FI_retaddr =>
True
|
FI_local x ty =>
ty <>
Tlong /\ 0 <=
x /\
x +
typesize ty <=
b.(
bound_local)
|
FI_arg x ty =>
ty <>
Tlong /\ 0 <=
x /\
x +
typesize ty <=
b.(
bound_outgoing)
|
FI_saved_int x => 0 <=
x <
b.(
bound_int_callee_save)
|
FI_saved_float x => 0 <=
x <
b.(
bound_float_callee_save)
end.
Definition type_of_index (
idx:
frame_index) :=
match idx with
|
FI_link =>
Tint
|
FI_retaddr =>
Tint
|
FI_local x ty =>
ty
|
FI_arg x ty =>
ty
|
FI_saved_int x =>
Tany32
|
FI_saved_float x =>
Tany64
end.
Non-overlap between the memory areas corresponding to two
frame indices.
Definition index_diff (
idx1 idx2:
frame_index) :
Prop :=
match idx1,
idx2 with
|
FI_link,
FI_link =>
False
|
FI_retaddr,
FI_retaddr =>
False
|
FI_local x1 ty1,
FI_local x2 ty2 =>
x1 +
typesize ty1 <=
x2 \/
x2 +
typesize ty2 <=
x1
|
FI_arg x1 ty1,
FI_arg x2 ty2 =>
x1 +
typesize ty1 <=
x2 \/
x2 +
typesize ty2 <=
x1
|
FI_saved_int x1,
FI_saved_int x2 =>
x1 <>
x2
|
FI_saved_float x1,
FI_saved_float x2 =>
x1 <>
x2
|
_,
_ =>
True
end.
Lemma index_diff_sym:
forall idx1 idx2,
index_diff idx1 idx2 ->
index_diff idx2 idx1.
Proof.
unfold index_diff;
intros.
destruct idx1;
destruct idx2;
intuition.
Qed.
Ltac AddPosProps :=
generalize (
bound_local_pos b);
intro;
generalize (
bound_int_callee_save_pos b);
intro;
generalize (
bound_float_callee_save_pos b);
intro;
generalize (
bound_outgoing_pos b);
intro;
generalize (
bound_stack_data_pos b);
intro.
Lemma size_pos: 0 <=
fe.(
fe_size).
Proof.
Opaque function_bounds.
Ltac InvIndexValid :=
match goal with
| [
H: ?
ty <>
Tlong /\
_ |-
_ ] =>
destruct H;
generalize (
typesize_pos ty) (
typesize_typesize ty);
intros
end.
Lemma offset_of_index_disj:
forall idx1 idx2,
index_valid idx1 ->
index_valid idx2 ->
index_diff idx1 idx2 ->
offset_of_index fe idx1 +
AST.typesize (
type_of_index idx1) <=
offset_of_index fe idx2 \/
offset_of_index fe idx2 +
AST.typesize (
type_of_index idx2) <=
offset_of_index fe idx1.
Proof.
Lemma offset_of_index_disj_stack_data_1:
forall idx,
index_valid idx ->
offset_of_index fe idx +
AST.typesize (
type_of_index idx) <=
fe.(
fe_stack_data)
\/
fe.(
fe_stack_data) +
b.(
bound_stack_data) <=
offset_of_index fe idx.
Proof.
Lemma offset_of_index_disj_stack_data_2:
forall idx,
index_valid idx ->
offset_of_index fe idx +
AST.typesize (
type_of_index idx) <=
fe.(
fe_stack_data)
\/
fe.(
fe_stack_data) +
f.(
Linear.fn_stacksize) <=
offset_of_index fe idx.
Proof.
Alignment properties
Remark aligned_4_4x:
forall x, (4 | 4 *
x).
Proof.
intro. exists x; ring. Qed.
Remark aligned_4_8x:
forall x, (4 | 8 *
x).
Proof.
intro. exists (x * 2); ring. Qed.
Remark aligned_8_4:
forall x, (8 |
x) -> (4 |
x).
Proof.
Hint Resolve Zdivide_0 Zdivide_refl Zdivide_plus_r
aligned_4_4x aligned_4_8x aligned_8_4:
align_4.
Hint Extern 4 (?
X | ?
Y) => (
exists (
Y/
X);
reflexivity) :
align_4.
Lemma offset_of_index_aligned:
forall idx, (4 |
offset_of_index fe idx).
Proof.
Lemma offset_of_index_aligned_2:
forall idx,
index_valid idx ->
(
align_chunk (
chunk_of_type (
type_of_index idx)) |
offset_of_index fe idx).
Proof.
Lemma fe_stack_data_aligned:
(8 |
fe_stack_data fe).
Proof.
The following lemmas give sufficient conditions for indices
of various kinds to be valid.
Lemma index_local_valid:
forall ofs ty,
slot_within_bounds b Local ofs ty ->
slot_valid f Local ofs ty =
true ->
index_valid (
FI_local ofs ty).
Proof.
Lemma index_arg_valid:
forall ofs ty,
slot_within_bounds b Outgoing ofs ty ->
slot_valid f Outgoing ofs ty =
true ->
index_valid (
FI_arg ofs ty).
Proof.
Lemma index_saved_int_valid:
forall r,
In r int_callee_save_regs ->
index_int_callee_save r <
b.(
bound_int_callee_save) ->
index_valid (
FI_saved_int (
index_int_callee_save r)).
Proof.
Lemma index_saved_float_valid:
forall r,
In r float_callee_save_regs ->
index_float_callee_save r <
b.(
bound_float_callee_save) ->
index_valid (
FI_saved_float (
index_float_callee_save r)).
Proof.
Hint Resolve index_local_valid index_arg_valid
index_saved_int_valid index_saved_float_valid:
stacking.
The offset of a valid index lies within the bounds of the frame.
Lemma offset_of_index_valid:
forall idx,
index_valid idx ->
0 <=
offset_of_index fe idx /\
offset_of_index fe idx +
AST.typesize (
type_of_index idx) <=
fe.(
fe_size).
Proof.
The image of the Linear stack data block lies within the bounds of the frame.
Lemma stack_data_offset_valid:
0 <=
fe.(
fe_stack_data) /\
fe.(
fe_stack_data) +
b.(
bound_stack_data) <=
fe.(
fe_size).
Proof.
Offsets for valid index are representable as signed machine integers
without loss of precision.
Lemma offset_of_index_no_overflow:
forall idx,
index_valid idx ->
Int.unsigned (
Int.repr (
offset_of_index fe idx)) =
offset_of_index fe idx.
Proof.
Likewise, for offsets within the Linear stack slot, after shifting.
Lemma shifted_stack_offset_no_overflow:
forall ofs,
0 <=
Int.unsigned ofs <
Linear.fn_stacksize f ->
Int.unsigned (
Int.add ofs (
Int.repr fe.(
fe_stack_data)))
=
Int.unsigned ofs +
fe.(
fe_stack_data).
Proof.
Contents of frame slots
Inductive index_contains (
m:
mem) (
sp:
block) (
idx:
frame_index) (
v:
val) :
Prop :=
|
index_contains_intro:
index_valid idx ->
Mem.load (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx) =
Some v ->
index_contains m sp idx v.
Lemma index_contains_load_stack:
forall m sp idx v,
index_contains m sp idx v ->
load_stack m (
Vptr sp Int.zero) (
type_of_index idx)
(
Int.repr (
offset_of_index fe idx)) =
Some v.
Proof.
Good variable properties for index_contains
Lemma gss_index_contains_base:
forall idx m m'
sp v,
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v =
Some m' ->
index_valid idx ->
exists v',
index_contains m'
sp idx v'
/\
decode_encode_val v (
chunk_of_type (
type_of_index idx)) (
chunk_of_type (
type_of_index idx))
v'.
Proof.
intros.
exploit Mem.load_store_similar.
eauto.
reflexivity.
omega.
intros [
v' [
A B]].
exists v';
split;
auto.
constructor;
auto.
Qed.
Lemma gss_index_contains:
forall idx m m'
sp v,
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v =
Some m' ->
index_valid idx ->
Val.has_type v (
type_of_index idx) ->
index_contains m'
sp idx v.
Proof.
intros.
exploit gss_index_contains_base;
eauto.
intros [
v' [
A B]].
assert (
v' =
v).
destruct v;
destruct (
type_of_index idx);
simpl in *;
try contradiction;
auto.
subst v'.
auto.
Qed.
Lemma gso_index_contains:
forall idx m m'
sp v idx'
v',
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v =
Some m' ->
index_valid idx ->
index_contains m sp idx'
v' ->
index_diff idx idx' ->
index_contains m'
sp idx'
v'.
Proof.
Lemma store_other_index_contains:
forall chunk m blk ofs v'
m'
sp idx v,
Mem.store chunk m blk ofs v' =
Some m' ->
blk <>
sp \/
(
fe.(
fe_stack_data) <=
ofs /\
ofs +
size_chunk chunk <=
fe.(
fe_stack_data) +
f.(
Linear.fn_stacksize)) ->
index_contains m sp idx v ->
index_contains m'
sp idx v.
Proof.
Definition frame_perm_freeable (
m:
mem) (
sp:
block):
Prop :=
forall ofs,
0 <=
ofs <
fe.(
fe_size) ->
ofs <
fe.(
fe_stack_data) \/
fe.(
fe_stack_data) +
f.(
Linear.fn_stacksize) <=
ofs ->
Mem.perm m sp ofs Cur Freeable.
Lemma offset_of_index_perm:
forall m sp idx,
index_valid idx ->
frame_perm_freeable m sp ->
Mem.range_perm m sp (
offset_of_index fe idx) (
offset_of_index fe idx +
AST.typesize (
type_of_index idx))
Cur Freeable.
Proof.
Lemma store_index_succeeds:
forall m sp idx v,
index_valid idx ->
frame_perm_freeable m sp ->
exists m',
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v =
Some m'.
Proof.
Lemma store_stack_succeeds:
forall m sp idx v m',
index_valid idx ->
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v =
Some m' ->
store_stack m (
Vptr sp Int.zero) (
type_of_index idx) (
Int.repr (
offset_of_index fe idx))
v =
Some m'.
Proof.
A variant of index_contains, up to a memory injection.
Definition index_contains_inj (
j:
meminj) (
m:
mem) (
sp:
block) (
idx:
frame_index) (
v:
val) :
Prop :=
exists v',
index_contains m sp idx v' /\
Val.inject j v v'.
Lemma gss_index_contains_inj:
forall j idx m m'
sp v v',
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v' =
Some m' ->
index_valid idx ->
Val.has_type v (
type_of_index idx) ->
Val.inject j v v' ->
index_contains_inj j m'
sp idx v.
Proof.
intros.
exploit gss_index_contains_base;
eauto.
intros [
v'' [
A B]].
exists v'';
split;
auto.
inv H2;
destruct (
type_of_index idx);
simpl in *;
try contradiction;
subst;
auto.
econstructor;
eauto.
econstructor;
eauto.
econstructor;
eauto.
Qed.
Lemma gss_index_contains_inj':
forall j idx m m'
sp v v',
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v' =
Some m' ->
index_valid idx ->
Val.inject j v v' ->
index_contains_inj j m'
sp idx (
Val.load_result (
chunk_of_type (
type_of_index idx))
v).
Proof.
intros.
exploit gss_index_contains_base;
eauto.
intros [
v'' [
A B]].
exists v'';
split;
auto.
inv H1;
destruct (
type_of_index idx);
simpl in *;
try contradiction;
subst;
auto.
econstructor;
eauto.
econstructor;
eauto.
econstructor;
eauto.
Qed.
Lemma gso_index_contains_inj:
forall j idx m m'
sp v idx'
v',
Mem.store (
chunk_of_type (
type_of_index idx))
m sp (
offset_of_index fe idx)
v =
Some m' ->
index_valid idx ->
index_contains_inj j m sp idx'
v' ->
index_diff idx idx' ->
index_contains_inj j m'
sp idx'
v'.
Proof.
intros.
destruct H1 as [
v'' [
A B]].
exists v'';
split;
auto.
eapply gso_index_contains;
eauto.
Qed.
Lemma store_other_index_contains_inj:
forall j chunk m b ofs v'
m'
sp idx v,
Mem.store chunk m b ofs v' =
Some m' ->
b <>
sp \/
(
fe.(
fe_stack_data) <=
ofs /\
ofs +
size_chunk chunk <=
fe.(
fe_stack_data) +
f.(
Linear.fn_stacksize)) ->
index_contains_inj j m sp idx v ->
index_contains_inj j m'
sp idx v.
Proof.
Lemma index_contains_inj_incr:
forall j m sp idx v j',
index_contains_inj j m sp idx v ->
inject_incr j j' ->
index_contains_inj j'
m sp idx v.
Proof.
intros. destruct H as [v'' [A B]]. exists v''; split; auto. eauto.
Qed.
Lemma index_contains_inj_undef:
forall j m sp idx,
index_valid idx ->
frame_perm_freeable m sp ->
index_contains_inj j m sp idx Vundef.
Proof.
Hint Resolve store_other_index_contains_inj index_contains_inj_incr:
stacking.
Agreement between location sets and Mach states
Agreement with Mach register states
Definition agree_regs (
j:
meminj) (
ls:
locset) (
rs:
regset) :
Prop :=
forall r,
Val.inject j (
ls (
R r)) (
rs r).
Agreement over data stored in memory
Record agree_frame (
j:
meminj) (
ls ls0:
locset)
(
m:
mem) (
sp:
block)
(
m':
mem) (
sp':
block)
(
parent retaddr:
val) :
Prop :=
mk_agree_frame {
Unused registers have the same value as in the caller
agree_unused_reg:
forall r, ~(
mreg_within_bounds b r) ->
ls (
R r) =
ls0 (
R r);
Local and outgoing stack slots (on the Linear side) have
the same values as the one loaded from the current Mach frame
at the corresponding offsets.
agree_locals:
forall ofs ty,
slot_within_bounds b Local ofs ty ->
slot_valid f Local ofs ty =
true ->
index_contains_inj j m'
sp' (
FI_local ofs ty) (
ls (
S Local ofs ty));
agree_outgoing:
forall ofs ty,
slot_within_bounds b Outgoing ofs ty ->
slot_valid f Outgoing ofs ty =
true ->
index_contains_inj j m'
sp' (
FI_arg ofs ty) (
ls (
S Outgoing ofs ty));
Incoming stack slots have the same value as the
corresponding Outgoing stack slots in the caller
agree_incoming:
forall ofs ty,
In (
S Incoming ofs ty) (
loc_parameters f.(
Linear.fn_sig)) ->
ls (
S Incoming ofs ty) =
ls0 (
S Outgoing ofs ty);
The back link and return address slots of the Mach frame contain
the parent and retaddr values, respectively.
agree_link:
index_contains m'
sp'
FI_link parent;
agree_retaddr:
index_contains m'
sp'
FI_retaddr retaddr;
The areas of the frame reserved for saving used callee-save
registers always contain the values that those registers had
in the caller.
agree_saved_int:
forall r,
In r int_callee_save_regs ->
index_int_callee_save r <
b.(
bound_int_callee_save) ->
index_contains_inj j m'
sp' (
FI_saved_int (
index_int_callee_save r)) (
ls0 (
R r));
agree_saved_float:
forall r,
In r float_callee_save_regs ->
index_float_callee_save r <
b.(
bound_float_callee_save) ->
index_contains_inj j m'
sp' (
FI_saved_float (
index_float_callee_save r)) (
ls0 (
R r));
Mapping between the Linear stack pointer and the Mach stack pointer
agree_inj:
j sp =
Some(
sp',
fe.(
fe_stack_data));
agree_inj_unique:
forall b delta,
j b =
Some(
sp',
delta) ->
b =
sp /\
delta =
fe.(
fe_stack_data);
The Linear and Mach stack pointers are valid
agree_valid_linear:
Mem.valid_block m sp;
agree_valid_mach:
Mem.valid_block m'
sp';
Bounds of the Linear stack data block
agree_bounds:
forall ofs p,
Mem.perm m sp ofs Max p -> 0 <=
ofs <
f.(
Linear.fn_stacksize);
Permissions on the frame part of the Mach stack block
agree_perm:
frame_perm_freeable m'
sp'
}.
Hint Resolve agree_unused_reg agree_locals agree_outgoing agree_incoming
agree_link agree_retaddr agree_saved_int agree_saved_float
agree_valid_linear agree_valid_mach agree_perm:
stacking.
Auxiliary predicate used at call points
Definition agree_callee_save (
ls ls0:
locset) :
Prop :=
forall l,
match l with
|
R r => ~
In r destroyed_at_call
|
S _ _ _ =>
True
end ->
ls l =
ls0 l.
Properties of agree_regs.
Values of registers
Lemma agree_reg:
forall j ls rs r,
agree_regs j ls rs ->
Val.inject j (
ls (
R r)) (
rs r).
Proof.
intros. auto.
Qed.
Lemma agree_reglist:
forall j ls rs rl,
agree_regs j ls rs ->
Val.inject_list j (
reglist ls rl) (
rs##
rl).
Proof.
induction rl; simpl; intros.
auto. constructor. eauto with stacking. auto.
Qed.
Hint Resolve agree_reg agree_reglist:
stacking.
Preservation under assignments of machine registers.
Lemma agree_regs_set_reg:
forall j ls rs r v v',
agree_regs j ls rs ->
Val.inject j v v' ->
agree_regs j (
Locmap.set (
R r)
v ls) (
Regmap.set r v'
rs).
Proof.
Lemma agree_regs_set_regs:
forall j rl vl vl'
ls rs,
agree_regs j ls rs ->
Val.inject_list j vl vl' ->
agree_regs j (
Locmap.setlist (
map R rl)
vl ls) (
set_regs rl vl'
rs).
Proof.
induction rl;
simpl;
intros.
auto.
inv H0.
auto.
apply IHrl;
auto.
apply agree_regs_set_reg;
auto.
Qed.
Lemma agree_regs_set_res:
forall j res v v'
ls rs,
agree_regs j ls rs ->
Val.inject j v v' ->
agree_regs j (
Locmap.setres res v ls) (
set_res res v'
rs).
Proof.
Lemma agree_regs_exten:
forall j ls rs ls'
rs',
agree_regs j ls rs ->
(
forall r,
ls' (
R r) =
Vundef \/
ls' (
R r) =
ls (
R r) /\
rs'
r =
rs r) ->
agree_regs j ls'
rs'.
Proof.
intros; red; intros.
destruct (H0 r) as [A | [A B]].
rewrite A. constructor.
rewrite A; rewrite B; auto.
Qed.
Lemma agree_regs_undef_regs:
forall j rl ls rs,
agree_regs j ls rs ->
agree_regs j (
LTL.undef_regs rl ls) (
Mach.undef_regs rl rs).
Proof.
Preservation under assignment of stack slot
Lemma agree_regs_set_slot:
forall j ls rs sl ofs ty v,
agree_regs j ls rs ->
agree_regs j (
Locmap.set (
S sl ofs ty)
v ls)
rs.
Proof.
intros;
red;
intros.
rewrite Locmap.gso;
auto.
red.
auto.
Qed.
Preservation by increasing memory injections
Lemma agree_regs_inject_incr:
forall j ls rs j',
agree_regs j ls rs ->
inject_incr j j' ->
agree_regs j'
ls rs.
Proof.
intros; red; intros; eauto with stacking.
Qed.
Preservation at function entry.
Lemma agree_regs_call_regs:
forall j ls rs,
agree_regs j ls rs ->
agree_regs j (
call_regs ls)
rs.
Proof.
intros.
unfold call_regs;
intros;
red;
intros;
auto.
Qed.
Properties of agree_frame
Preservation under assignment of machine register.
Lemma agree_frame_set_reg:
forall j ls ls0 m sp m'
sp'
parent ra r v,
agree_frame j ls ls0 m sp m'
sp'
parent ra ->
mreg_within_bounds b r ->
agree_frame j (
Locmap.set (
R r)
v ls)
ls0 m sp m'
sp'
parent ra.
Proof.
intros.
inv H;
constructor;
auto;
intros.
rewrite Locmap.gso.
auto.
red.
intuition congruence.
Qed.
Lemma agree_frame_set_regs:
forall j ls0 m sp m'
sp'
parent ra rl vl ls,
agree_frame j ls ls0 m sp m'
sp'
parent ra ->
(
forall r,
In r rl ->
mreg_within_bounds b r) ->
agree_frame j (
Locmap.setlist (
map R rl)
vl ls)
ls0 m sp m'
sp'
parent ra.
Proof.
induction rl;
destruct vl;
simpl;
intros;
intuition.
apply IHrl;
auto.
eapply agree_frame_set_reg;
eauto.
Qed.
Lemma agree_frame_set_res:
forall j ls0 m sp m'
sp'
parent ra res v ls,
agree_frame j ls ls0 m sp m'
sp'
parent ra ->
(
forall r,
In r (
params_of_builtin_res res) ->
mreg_within_bounds b r) ->
agree_frame j (
Locmap.setres res v ls)
ls0 m sp m'
sp'
parent ra.
Proof.
Lemma agree_frame_undef_regs:
forall j ls0 m sp m'
sp'
parent ra regs ls,
agree_frame j ls ls0 m sp m'
sp'
parent ra ->
(
forall r,
In r regs ->
mreg_within_bounds b r) ->
agree_frame j (
LTL.undef_regs regs ls)
ls0 m sp m'
sp'
parent ra.
Proof.
Lemma caller_save_reg_within_bounds:
forall r,
In r destroyed_at_call ->
mreg_within_bounds b r.
Proof.
Lemma agree_frame_undef_locs:
forall j ls0 m sp m'
sp'
parent ra regs ls,
agree_frame j ls ls0 m sp m'
sp'
parent ra ->
incl regs destroyed_at_call ->
agree_frame j (
LTL.undef_regs regs ls)
ls0 m sp m'
sp'
parent ra.
Proof.
Preservation by assignment to local slot
Lemma agree_frame_set_local:
forall j ls ls0 m sp m'
sp'
parent retaddr ofs ty v v'
m'',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
slot_within_bounds b Local ofs ty ->
slot_valid f Local ofs ty =
true ->
Val.inject j v v' ->
Mem.store (
chunk_of_type ty)
m'
sp' (
offset_of_index fe (
FI_local ofs ty))
v' =
Some m'' ->
agree_frame j (
Locmap.set (
S Local ofs ty)
v ls)
ls0 m sp m''
sp'
parent retaddr.
Proof.
Preservation by assignment to outgoing slot
Lemma agree_frame_set_outgoing:
forall j ls ls0 m sp m'
sp'
parent retaddr ofs ty v v'
m'',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
slot_within_bounds b Outgoing ofs ty ->
slot_valid f Outgoing ofs ty =
true ->
Val.inject j v v' ->
Mem.store (
chunk_of_type ty)
m'
sp' (
offset_of_index fe (
FI_arg ofs ty))
v' =
Some m'' ->
agree_frame j (
Locmap.set (
S Outgoing ofs ty)
v ls)
ls0 m sp m''
sp'
parent retaddr.
Proof.
General invariance property with respect to memory changes.
Lemma agree_frame_invariant:
forall j ls ls0 m sp m'
sp'
parent retaddr m1 m1',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
(
Mem.valid_block m sp ->
Mem.valid_block m1 sp) ->
(
forall ofs p,
Mem.perm m1 sp ofs Max p ->
Mem.perm m sp ofs Max p) ->
(
Mem.valid_block m'
sp' ->
Mem.valid_block m1'
sp') ->
(
forall chunk ofs v,
ofs +
size_chunk chunk <=
fe.(
fe_stack_data) \/
fe.(
fe_stack_data) +
f.(
Linear.fn_stacksize) <=
ofs ->
Mem.load chunk m'
sp'
ofs =
Some v ->
Mem.load chunk m1'
sp'
ofs =
Some v) ->
(
forall ofs k p,
ofs <
fe.(
fe_stack_data) \/
fe.(
fe_stack_data) +
f.(
Linear.fn_stacksize) <=
ofs ->
Mem.perm m'
sp'
ofs k p ->
Mem.perm m1'
sp'
ofs k p) ->
agree_frame j ls ls0 m1 sp m1'
sp'
parent retaddr.
Proof.
A variant of the latter, for use with external calls
Lemma agree_frame_extcall_invariant:
forall j ls ls0 m sp m'
sp'
parent retaddr m1 m1',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
(
Mem.valid_block m sp ->
Mem.valid_block m1 sp) ->
(
forall ofs p,
Mem.perm m1 sp ofs Max p ->
Mem.perm m sp ofs Max p) ->
(
Mem.valid_block m'
sp' ->
Mem.valid_block m1'
sp') ->
Mem.unchanged_on (
loc_out_of_reach j m)
m'
m1' ->
agree_frame j ls ls0 m1 sp m1'
sp'
parent retaddr.
Proof.
Preservation by parallel stores in the Linear and Mach codes
Lemma agree_frame_parallel_stores:
forall j ls ls0 m sp m'
sp'
parent retaddr chunk addr addr'
v v'
m1 m1',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
Mem.inject j m m' ->
Val.inject j addr addr' ->
Mem.storev chunk m addr v =
Some m1 ->
Mem.storev chunk m'
addr'
v' =
Some m1' ->
agree_frame j ls ls0 m1 sp m1'
sp'
parent retaddr.
Proof.
Preservation by increasing memory injections (allocations and external calls)
Lemma agree_frame_inject_incr:
forall j ls ls0 m sp m'
sp'
parent retaddr m1 m1'
j',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
inject_incr j j' ->
inject_separated j j'
m1 m1' ->
Mem.valid_block m1'
sp' ->
agree_frame j'
ls ls0 m sp m'
sp'
parent retaddr.
Proof.
intros. inv H. constructor; auto; intros; eauto with stacking.
case_eq (j b0).
intros [b' delta'] EQ. rewrite (H0 _ _ _ EQ) in H. inv H. auto.
intros EQ. exploit H1. eauto. eauto. intros [A B]. contradiction.
Qed.
Remark inject_alloc_separated:
forall j m1 m2 j'
b1 b2 delta,
inject_incr j j' ->
j'
b1 =
Some(
b2,
delta) ->
(
forall b,
b <>
b1 ->
j'
b =
j b) ->
~
Mem.valid_block m1 b1 -> ~
Mem.valid_block m2 b2 ->
inject_separated j j'
m1 m2.
Proof.
intros.
red.
intros.
destruct (
eq_block b0 b1).
subst b0.
rewrite H0 in H5;
inv H5.
tauto.
rewrite H1 in H5.
congruence.
auto.
Qed.
Preservation at return points (when ls is changed but not ls0).
Lemma agree_frame_return:
forall j ls ls0 m sp m'
sp'
parent retaddr ls',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
agree_callee_save ls'
ls ->
agree_frame j ls'
ls0 m sp m'
sp'
parent retaddr.
Proof.
intros.
red in H0.
inv H;
constructor;
auto;
intros.
rewrite H0;
auto.
red;
intros;
elim H.
apply caller_save_reg_within_bounds;
auto.
rewrite H0;
auto.
rewrite H0;
auto.
rewrite H0;
auto.
Qed.
Preservation at tailcalls (when ls0 is changed but not ls).
Lemma agree_frame_tailcall:
forall j ls ls0 m sp m'
sp'
parent retaddr ls0',
agree_frame j ls ls0 m sp m'
sp'
parent retaddr ->
agree_callee_save ls0 ls0' ->
agree_frame j ls ls0'
m sp m'
sp'
parent retaddr.
Proof.
Properties of agree_callee_save.
Lemma agree_callee_save_return_regs:
forall ls1 ls2,
agree_callee_save (
return_regs ls1 ls2)
ls1.
Proof.
Lemma agree_callee_save_set_result:
forall sg vl ls1 ls2,
agree_callee_save ls1 ls2 ->
agree_callee_save (
Locmap.setlist (
map R (
loc_result sg))
vl ls1)
ls2.
Proof.
Properties of destroyed registers.
Lemma check_mreg_list_incl:
forall l1 l2,
forallb (
fun r =>
In_dec mreg_eq r l2)
l1 =
true ->
incl l1 l2.
Proof.
Remark destroyed_by_op_caller_save:
forall op,
incl (
destroyed_by_op op)
destroyed_at_call.
Proof.
Remark destroyed_by_load_caller_save:
forall chunk addr,
incl (
destroyed_by_load chunk addr)
destroyed_at_call.
Proof.
Remark destroyed_by_store_caller_save:
forall chunk addr,
incl (
destroyed_by_store chunk addr)
destroyed_at_call.
Proof.
Remark destroyed_by_cond_caller_save:
forall cond,
incl (
destroyed_by_cond cond)
destroyed_at_call.
Proof.
Remark destroyed_by_jumptable_caller_save:
incl destroyed_by_jumptable destroyed_at_call.
Proof.
Remark destroyed_by_setstack_caller_save:
forall ty,
incl (
destroyed_by_setstack ty)
destroyed_at_call.
Proof.
Remark destroyed_at_function_entry_caller_save:
incl destroyed_at_function_entry destroyed_at_call.
Proof.
Remark temp_for_parent_frame_caller_save:
In temp_for_parent_frame destroyed_at_call.
Proof.
Hint Resolve destroyed_by_op_caller_save destroyed_by_load_caller_save
destroyed_by_store_caller_save
destroyed_by_cond_caller_save destroyed_by_jumptable_caller_save
destroyed_at_function_entry_caller_save:
stacking.
Remark destroyed_by_setstack_function_entry:
forall ty,
incl (
destroyed_by_setstack ty)
destroyed_at_function_entry.
Proof.
Remark transl_destroyed_by_op:
forall op e,
destroyed_by_op (
transl_op e op) =
destroyed_by_op op.
Proof.
intros; destruct op; reflexivity.
Qed.
Remark transl_destroyed_by_load:
forall chunk addr e,
destroyed_by_load chunk (
transl_addr e addr) =
destroyed_by_load chunk addr.
Proof.
intros; destruct chunk; reflexivity.
Qed.
Remark transl_destroyed_by_store:
forall chunk addr e,
destroyed_by_store chunk (
transl_addr e addr) =
destroyed_by_store chunk addr.
Proof.
intros; destruct chunk; reflexivity.
Qed.
Correctness of saving and restoring of callee-save registers
The following lemmas show the correctness of the register saving
code generated by save_callee_save: after this code has executed,
the register save areas of the current frame do contain the
values of the callee-save registers used by the function.
Section SAVE_CALLEE_SAVE.
Variable bound:
frame_env ->
Z.
Variable number:
mreg ->
Z.
Variable mkindex:
Z ->
frame_index.
Variable ty:
typ.
Variable j:
meminj.
Variable cs:
list stackframe.
Variable fb:
block.
Variable sp:
block.
Variable csregs:
list mreg.
Variable ls:
locset.
Inductive stores_in_frame:
mem ->
mem ->
Prop :=
|
stores_in_frame_refl:
forall m,
stores_in_frame m m
|
stores_in_frame_step:
forall m1 chunk ofs v m2 m3,
ofs +
size_chunk chunk <=
fe.(
fe_stack_data)
\/
fe.(
fe_stack_data) +
f.(
Linear.fn_stacksize) <=
ofs ->
Mem.store chunk m1 sp ofs v =
Some m2 ->
stores_in_frame m2 m3 ->
stores_in_frame m1 m3.
Remark stores_in_frame_trans:
forall m1 m2,
stores_in_frame m1 m2 ->
forall m3,
stores_in_frame m2 m3 ->
stores_in_frame m1 m3.
Proof.
induction 1; intros. auto. econstructor; eauto.
Qed.
Hypothesis number_inj:
forall r1 r2,
In r1 csregs ->
In r2 csregs ->
r1 <>
r2 ->
number r1 <>
number r2.
Hypothesis mkindex_valid:
forall r,
In r csregs ->
number r <
bound fe ->
index_valid (
mkindex (
number r)).
Hypothesis mkindex_typ:
forall z,
type_of_index (
mkindex z) =
ty.
Hypothesis mkindex_inj:
forall z1 z2,
z1 <>
z2 ->
mkindex z1 <>
mkindex z2.
Hypothesis mkindex_diff:
forall r idx,
idx <>
mkindex (
number r) ->
index_diff (
mkindex (
number r))
idx.
Hypothesis csregs_typ:
forall r,
In r csregs ->
mreg_type r =
ty.
Hypothesis ls_temp_undef:
forall r,
In r (
destroyed_by_setstack ty) ->
ls (
R r) =
Vundef.
Hypothesis wt_ls:
forall r,
Val.has_type (
ls (
R r)) (
mreg_type r).
Lemma save_callee_save_regs_correct:
forall l k rs m,
incl l csregs ->
list_norepet l ->
frame_perm_freeable m sp ->
agree_regs j ls rs ->
exists rs',
exists m',
star step tge
(
State cs fb (
Vptr sp Int.zero)
(
save_callee_save_regs bound number mkindex ty fe l k)
rs m)
E0 (
State cs fb (
Vptr sp Int.zero)
k rs'
m')
/\ (
forall r,
In r l ->
number r <
bound fe ->
index_contains_inj j m'
sp (
mkindex (
number r)) (
ls (
R r)))
/\ (
forall idx v,
index_valid idx ->
(
forall r,
In r l ->
number r <
bound fe ->
idx <>
mkindex (
number r)) ->
index_contains m sp idx v ->
index_contains m'
sp idx v)
/\
stores_in_frame m m'
/\
frame_perm_freeable m'
sp
/\
agree_regs j ls rs'.
Proof.
Lemma save_callee_save_regs_correct':
forall l k rs m,
incl l csregs ->
list_norepet l ->
frame_perm_freeable m sp ->
agree_regs j ls rs ->
exists rs',
exists m',
star step_safe tge
(
State cs fb (
Vptr sp Int.zero)
(
save_callee_save_regs bound number mkindex ty fe l k)
rs m)
E0 (
State cs fb (
Vptr sp Int.zero)
k rs'
m')
/\ (
forall r,
In r l ->
number r <
bound fe ->
index_contains_inj j m'
sp (
mkindex (
number r)) (
ls (
R r)))
/\ (
forall idx v,
index_valid idx ->
(
forall r,
In r l ->
number r <
bound fe ->
idx <>
mkindex (
number r)) ->
index_contains m sp idx v ->
index_contains m'
sp idx v)
/\
stores_in_frame m m'
/\
frame_perm_freeable m'
sp
/\
agree_regs j ls rs'.
Proof.
End SAVE_CALLEE_SAVE.
Remark LTL_undef_regs_same:
forall r rl ls,
In r rl ->
LTL.undef_regs rl ls (
R r) =
Vundef.
Proof.
induction rl;
simpl;
intros.
contradiction.
unfold Locmap.set.
destruct (
Loc.eq (
R a) (
R r)).
auto.
destruct (
Loc.diff_dec (
R a) (
R r));
auto.
apply IHrl.
intuition congruence.
Qed.
Remark LTL_undef_regs_others:
forall r rl ls, ~
In r rl ->
LTL.undef_regs rl ls (
R r) =
ls (
R r).
Proof.
induction rl;
simpl;
intros.
auto.
rewrite Locmap.gso.
apply IHrl.
intuition.
red.
intuition.
Qed.
Remark LTL_undef_regs_slot:
forall sl ofs ty rl ls,
LTL.undef_regs rl ls (
S sl ofs ty) =
ls (
S sl ofs ty).
Proof.
induction rl;
simpl;
intros.
auto.
rewrite Locmap.gso.
apply IHrl.
red;
auto.
Qed.
Lemma save_callee_save_correct:
forall j ls0 rs sp cs fb k m,
let ls :=
LTL.undef_regs destroyed_at_function_entry ls0 in
agree_regs j ls rs ->
(
forall r,
Val.has_type (
ls (
R r)) (
mreg_type r)) ->
frame_perm_freeable m sp ->
exists rs',
exists m',
star step tge
(
State cs fb (
Vptr sp Int.zero) (
save_callee_save fe k)
rs m)
E0 (
State cs fb (
Vptr sp Int.zero)
k rs'
m')
/\ (
forall r,
In r int_callee_save_regs ->
index_int_callee_save r <
b.(
bound_int_callee_save) ->
index_contains_inj j m'
sp (
FI_saved_int (
index_int_callee_save r)) (
ls (
R r)))
/\ (
forall r,
In r float_callee_save_regs ->
index_float_callee_save r <
b.(
bound_float_callee_save) ->
index_contains_inj j m'
sp (
FI_saved_float (
index_float_callee_save r)) (
ls (
R r)))
/\ (
forall idx v,
index_valid idx ->
match idx with FI_saved_int _ =>
False |
FI_saved_float _ =>
False |
_ =>
True end ->
index_contains m sp idx v ->
index_contains m'
sp idx v)
/\
stores_in_frame sp m m'
/\
frame_perm_freeable m'
sp
/\
agree_regs j ls rs'.
Proof.
Lemma save_callee_save_correct':
forall j ls0 rs sp cs fb k m,
let ls :=
LTL.undef_regs destroyed_at_function_entry ls0 in
agree_regs j ls rs ->
(
forall r,
Val.has_type (
ls (
R r)) (
mreg_type r)) ->
frame_perm_freeable m sp ->
exists rs',
exists m',
star step_safe tge
(
State cs fb (
Vptr sp Int.zero) (
save_callee_save fe k)
rs m)
E0 (
State cs fb (
Vptr sp Int.zero)
k rs'
m')
/\ (
forall r,
In r int_callee_save_regs ->
index_int_callee_save r <
b.(
bound_int_callee_save) ->
index_contains_inj j m'
sp (
FI_saved_int (
index_int_callee_save r)) (
ls (
R r)))
/\ (
forall r,
In r float_callee_save_regs ->
index_float_callee_save r <
b.(
bound_float_callee_save) ->
index_contains_inj j m'
sp (
FI_saved_float (
index_float_callee_save r)) (
ls (
R r)))
/\ (
forall idx v,
index_valid idx ->
match idx with FI_saved_int _ =>
False |
FI_saved_float _ =>
False |
_ =>
True end ->
index_contains m sp idx v ->
index_contains m'
sp idx v)
/\
stores_in_frame sp m m'
/\
frame_perm_freeable m'
sp
/\
agree_regs j ls rs'.
Proof.
Properties of sequences of stores in the frame.
Lemma stores_in_frame_inject:
forall j sp sp'
m,
(
forall b delta,
j b =
Some(
sp',
delta) ->
b =
sp /\
delta =
fe.(
fe_stack_data)) ->
(
forall ofs k p,
Mem.perm m sp ofs k p -> 0 <=
ofs <
f.(
Linear.fn_stacksize)) ->
forall m1 m2,
stores_in_frame sp'
m1 m2 ->
Mem.inject j m m1 ->
Mem.inject j m m2.
Proof.
induction 3;
intros.
auto.
apply IHstores_in_frame.
intros.
eapply Mem.store_outside_inject;
eauto.
intros.
exploit H;
eauto.
intros [
A B];
subst.
exploit H0;
eauto.
omega.
Qed.
Lemma stores_in_frame_valid:
forall b sp m m',
stores_in_frame sp m m' ->
Mem.valid_block m b ->
Mem.valid_block m'
b.
Proof.
induction 1; intros. auto. apply IHstores_in_frame. eauto with mem.
Qed.
Lemma stores_in_frame_perm:
forall b ofs k p sp m m',
stores_in_frame sp m m' ->
Mem.perm m b ofs k p ->
Mem.perm m'
b ofs k p.
Proof.
induction 1; intros. auto. apply IHstores_in_frame. eauto with mem.
Qed.
Lemma stores_in_frame_contents:
forall chunk b ofs sp,
Plt b sp ->
forall m m',
stores_in_frame sp m m' ->
Mem.load chunk m'
b ofs =
Mem.load chunk m b ofs.
Proof.
Lemma undef_regs_type:
forall ty l rl ls,
Val.has_type (
ls l)
ty ->
Val.has_type (
LTL.undef_regs rl ls l)
ty.
Proof.
induction rl;
simpl;
intros.
-
auto.
-
unfold Locmap.set.
destruct (
Loc.eq (
R a)
l).
red;
auto.
destruct (
Loc.diff_dec (
R a)
l);
auto.
red;
auto.
Qed.
As a corollary of the previous lemmas, we obtain the following
correctness theorem for the execution of a function prologue
(allocation of the frame + saving of the link and return address +
saving of the used callee-save registers).
Lemma function_prologue_correct:
forall j ls ls0 ls1 rs rs1 m1 m1'
m2 sp parent ra cs fb k,
agree_regs j ls rs ->
agree_callee_save ls ls0 ->
(
forall r,
Val.has_type (
ls (
R r)) (
mreg_type r)) ->
ls1 =
LTL.undef_regs destroyed_at_function_entry (
LTL.call_regs ls) ->
rs1 =
undef_regs destroyed_at_function_entry rs ->
Mem.inject j m1 m1' ->
Mem.alloc m1 0
f.(
Linear.fn_stacksize) = (
m2,
sp) ->
Val.has_type parent Tint ->
Val.has_type ra Tint ->
exists j',
exists rs',
exists m2',
exists sp',
exists m3',
exists m4',
exists m5',
Mem.alloc m1' 0
tf.(
fn_stacksize) = (
m2',
sp')
/\
store_stack m2' (
Vptr sp'
Int.zero)
Tint tf.(
fn_link_ofs)
parent =
Some m3'
/\
store_stack m3' (
Vptr sp'
Int.zero)
Tint tf.(
fn_retaddr_ofs)
ra =
Some m4'
/\
star step tge
(
State cs fb (
Vptr sp'
Int.zero) (
save_callee_save fe k)
rs1 m4')
E0 (
State cs fb (
Vptr sp'
Int.zero)
k rs'
m5')
/\
agree_regs j'
ls1 rs'
/\
agree_frame j'
ls1 ls0 m2 sp m5'
sp'
parent ra
/\
inject_incr j j'
/\
inject_separated j j'
m1 m1'
/\
Mem.inject j'
m2 m5'
/\
stores_in_frame sp'
m2'
m5'.
Proof.
Lemma function_prologue_correct':
forall j ls ls0 ls1 rs rs1 m1 m1'
m2 sp parent ra cs fb k,
agree_regs j ls rs ->
agree_callee_save ls ls0 ->
(
forall r,
Val.has_type (
ls (
R r)) (
mreg_type r)) ->
ls1 =
LTL.undef_regs destroyed_at_function_entry (
LTL.call_regs ls) ->
rs1 =
undef_regs destroyed_at_function_entry rs ->
Mem.inject j m1 m1' ->
Mem.alloc m1 0
f.(
Linear.fn_stacksize) = (
m2,
sp) ->
Val.has_type parent Tint ->
Val.has_type ra Tint ->
exists j',
exists rs',
exists m2',
exists sp',
exists m3',
exists m4',
exists m5',
Mem.alloc m1' 0
tf.(
fn_stacksize) = (
m2',
sp')
/\
store_stack m2' (
Vptr sp'
Int.zero)
Tint tf.(
fn_link_ofs)
parent =
Some m3'
/\
store_stack m3' (
Vptr sp'
Int.zero)
Tint tf.(
fn_retaddr_ofs)
ra =
Some m4'
/\
star step_safe tge
(
State cs fb (
Vptr sp'
Int.zero) (
save_callee_save fe k)
rs1 m4')
E0 (
State cs fb (
Vptr sp'
Int.zero)
k rs'
m5')
/\
agree_regs j'
ls1 rs'
/\
agree_frame j'
ls1 ls0 m2 sp m5'
sp'
parent ra
/\
inject_incr j j'
/\
inject_separated j j'
m1 m1'
/\
Mem.inject j'
m2 m5'
/\
stores_in_frame sp'
m2'
m5'.
Proof.
The following lemmas show the correctness of the register reloading
code generated by reload_callee_save: after this code has executed,
all callee-save registers contain the same values they had at
function entry.
Section RESTORE_CALLEE_SAVE.
Variable bound:
frame_env ->
Z.
Variable number:
mreg ->
Z.
Variable mkindex:
Z ->
frame_index.
Variable ty:
typ.
Variable csregs:
list mreg.
Variable j:
meminj.
Variable cs:
list stackframe.
Variable fb:
block.
Variable sp:
block.
Variable ls0:
locset.
Variable m:
mem.
Hypothesis mkindex_valid:
forall r,
In r csregs ->
number r <
bound fe ->
index_valid (
mkindex (
number r)).
Hypothesis mkindex_typ:
forall z,
type_of_index (
mkindex z) =
ty.
Hypothesis number_within_bounds:
forall r,
In r csregs ->
(
number r <
bound fe <->
mreg_within_bounds b r).
Hypothesis mkindex_val:
forall r,
In r csregs ->
number r <
bound fe ->
index_contains_inj j m sp (
mkindex (
number r)) (
ls0 (
R r)).
Definition agree_unused (
ls0:
locset) (
rs:
regset) :
Prop :=
forall r, ~(
mreg_within_bounds b r) ->
Val.inject j (
ls0 (
R r)) (
rs r).
Lemma restore_callee_save_regs_correct:
forall l rs k,
incl l csregs ->
list_norepet l ->
agree_unused ls0 rs ->
exists rs',
star step tge
(
State cs fb (
Vptr sp Int.zero)
(
restore_callee_save_regs bound number mkindex ty fe l k)
rs m)
E0 (
State cs fb (
Vptr sp Int.zero)
k rs'
m)
/\ (
forall r,
In r l ->
Val.inject j (
ls0 (
R r)) (
rs'
r))
/\ (
forall r, ~(
In r l) ->
rs'
r =
rs r)
/\
agree_unused ls0 rs'.
Proof.
Lemma restore_callee_save_regs_correct':
forall l rs k,
incl l csregs ->
list_norepet l ->
agree_unused ls0 rs ->
exists rs',
star step_safe tge
(
State cs fb (
Vptr sp Int.zero)
(
restore_callee_save_regs bound number mkindex ty fe l k)
rs m)
E0 (
State cs fb (
Vptr sp Int.zero)
k rs'
m)
/\ (
forall r,
In r l ->
Val.inject j (
ls0 (
R r)) (
rs'
r))
/\ (
forall r, ~(
In r l) ->
rs'
r =
rs r)
/\
agree_unused ls0 rs'.
Proof.
End RESTORE_CALLEE_SAVE.
Lemma restore_callee_save_correct:
forall j ls ls0 m sp m'
sp'
pa ra cs fb rs k,
agree_frame j ls ls0 m sp m'
sp'
pa ra ->
agree_unused j ls0 rs ->
exists rs',
star step tge
(
State cs fb (
Vptr sp'
Int.zero) (
restore_callee_save fe k)
rs m')
E0 (
State cs fb (
Vptr sp'
Int.zero)
k rs'
m')
/\ (
forall r,
In r int_callee_save_regs \/
In r float_callee_save_regs ->
Val.inject j (
ls0 (
R r)) (
rs'
r))
/\ (
forall r,
~(
In r int_callee_save_regs) ->
~(
In r float_callee_save_regs) ->
rs'
r =
rs r).
Proof.
Lemma restore_callee_save_correct':
forall j ls ls0 m sp m'
sp'
pa ra cs fb rs k,
agree_frame j ls ls0 m sp m'
sp'
pa ra ->
agree_unused j ls0 rs ->
exists rs',
star step_safe tge
(
State cs fb (
Vptr sp'
Int.zero) (
restore_callee_save fe k)
rs m')
E0 (
State cs fb (
Vptr sp'
Int.zero)
k rs'
m')
/\ (
forall r,
In r int_callee_save_regs \/
In r float_callee_save_regs ->
Val.inject j (
ls0 (
R r)) (
rs'
r))
/\ (
forall r,
~(
In r int_callee_save_regs) ->
~(
In r float_callee_save_regs) ->
rs'
r =
rs r).
Proof.
As a corollary, we obtain the following correctness result for
the execution of a function epilogue (reloading of used callee-save
registers + reloading of the link and return address + freeing
of the frame).
Lemma function_epilogue_correct:
forall j ls ls0 m sp m'
sp'
pa ra cs fb rs k m1,
agree_regs j ls rs ->
agree_frame j ls ls0 m sp m'
sp'
pa ra ->
Mem.inject j m m' ->
Mem.free m sp 0
f.(
Linear.fn_stacksize) =
Some m1 ->
exists rs1,
exists m1',
load_stack m' (
Vptr sp'
Int.zero)
Tint tf.(
fn_link_ofs) =
Some pa
/\
load_stack m' (
Vptr sp'
Int.zero)
Tint tf.(
fn_retaddr_ofs) =
Some ra
/\
Mem.free m'
sp' 0
tf.(
fn_stacksize) =
Some m1'
/\
star step tge
(
State cs fb (
Vptr sp'
Int.zero) (
restore_callee_save fe k)
rs m')
E0 (
State cs fb (
Vptr sp'
Int.zero)
k rs1 m')
/\
agree_regs j (
return_regs ls0 ls)
rs1
/\
agree_callee_save (
return_regs ls0 ls)
ls0
/\
Mem.inject j m1 m1'.
Proof.
Lemma function_epilogue_correct':
forall j ls ls0 m sp m'
sp'
pa ra cs fb rs k m1,
agree_regs j ls rs ->
agree_frame j ls ls0 m sp m'
sp'
pa ra ->
Mem.inject j m m' ->
Mem.free m sp 0
f.(
Linear.fn_stacksize) =
Some m1 ->
exists rs1,
exists m1',
load_stack m' (
Vptr sp'
Int.zero)
Tint tf.(
fn_link_ofs) =
Some pa
/\
load_stack m' (
Vptr sp'
Int.zero)
Tint tf.(
fn_retaddr_ofs) =
Some ra
/\
Mem.free m'
sp' 0
tf.(
fn_stacksize) =
Some m1'
/\
star step_safe tge
(
State cs fb (
Vptr sp'
Int.zero) (
restore_callee_save fe k)
rs m')
E0 (
State cs fb (
Vptr sp'
Int.zero)
k rs1 m')
/\
agree_regs j (
return_regs ls0 ls)
rs1
/\
agree_callee_save (
return_regs ls0 ls)
ls0
/\
Mem.inject j m1 m1'.
Proof.
End FRAME_PROPERTIES.
Call stack invariant
Inductive match_globalenvs (
j:
meminj) (
bound:
block) :
Prop :=
|
match_globalenvs_intro
(
DOMAIN:
forall b,
Plt b bound ->
j b =
Some(
b, 0))
(
IMAGE:
forall b1 b2 delta,
j b1 =
Some(
b2,
delta) ->
Plt b2 bound ->
b1 =
b2)
(
SYMBOLS:
forall id b,
Genv.find_symbol ge id =
Some b ->
Plt b bound)
(
FUNCTIONS:
forall b fd,
Genv.find_funct_ptr ge b =
Some fd ->
Plt b bound)
(
VARINFOS:
forall b gv,
Genv.find_var_info ge b =
Some gv ->
Plt b bound).
Inductive match_stacks (
j:
meminj) (
m m':
mem):
list Linear.stackframe ->
list stackframe ->
signature ->
block ->
block ->
Prop :=
|
match_stacks_empty:
forall sg hi bound bound',
Ple hi bound ->
Ple hi bound' ->
match_globalenvs j hi ->
tailcall_possible sg ->
match_stacks j m m'
nil nil sg bound bound'
|
match_stacks_cons:
forall f sp ls c cs fb sp'
ra c'
cs'
sg bound bound'
trf
(
TAIL:
is_tail c (
Linear.fn_code f))
(
FINDF:
Genv.find_funct_ptr tge fb =
Some (
Internal trf))
(
TRF:
transf_function f =
OK trf)
(
TRC:
transl_code (
make_env (
function_bounds f))
c =
c')
(
TY_RA:
Val.has_type ra Tint)
(
FRM:
agree_frame f j ls (
parent_locset cs)
m sp m'
sp' (
parent_sp cs') (
parent_ra cs'))
(
ARGS:
forall ofs ty,
In (
S Outgoing ofs ty) (
loc_arguments sg) ->
slot_within_bounds (
function_bounds f)
Outgoing ofs ty)
(
STK:
match_stacks j m m'
cs cs' (
Linear.fn_sig f)
sp sp')
(
BELOW:
Plt sp bound)
(
BELOW':
Plt sp'
bound'),
match_stacks j m m'
(
Linear.Stackframe f (
Vptr sp Int.zero)
ls c ::
cs)
(
Stackframe fb (
Vptr sp'
Int.zero)
ra c' ::
cs')
sg bound bound'.
Invariance with respect to change of bounds.
Lemma match_stacks_change_bounds:
forall j m1 m'
cs cs'
sg bound bound',
match_stacks j m1 m'
cs cs'
sg bound bound' ->
forall xbound xbound',
Ple bound xbound ->
Ple bound'
xbound' ->
match_stacks j m1 m'
cs cs'
sg xbound xbound'.
Proof.
Invariance with respect to change of m.
Lemma match_stacks_change_linear_mem:
forall j m1 m2 m'
cs cs'
sg bound bound',
match_stacks j m1 m'
cs cs'
sg bound bound' ->
(
forall b,
Plt b bound ->
Mem.valid_block m1 b ->
Mem.valid_block m2 b) ->
(
forall b ofs p,
Plt b bound ->
Mem.perm m2 b ofs Max p ->
Mem.perm m1 b ofs Max p) ->
match_stacks j m2 m'
cs cs'
sg bound bound'.
Proof.
induction 1;
intros.
econstructor;
eauto.
econstructor;
eauto.
eapply agree_frame_invariant;
eauto.
apply IHmatch_stacks.
intros.
apply H0;
auto.
apply Plt_trans with sp;
auto.
intros.
apply H1.
apply Plt_trans with sp;
auto.
auto.
Qed.
Invariance with respect to change of m'.
Lemma match_stacks_change_mach_mem:
forall j m m1'
m2'
cs cs'
sg bound bound',
match_stacks j m m1'
cs cs'
sg bound bound' ->
(
forall b,
Plt b bound' ->
Mem.valid_block m1'
b ->
Mem.valid_block m2'
b) ->
(
forall b ofs k p,
Plt b bound' ->
Mem.perm m1'
b ofs k p ->
Mem.perm m2'
b ofs k p) ->
(
forall chunk b ofs v,
Plt b bound' ->
Mem.load chunk m1'
b ofs =
Some v ->
Mem.load chunk m2'
b ofs =
Some v) ->
match_stacks j m m2'
cs cs'
sg bound bound'.
Proof.
induction 1;
intros.
econstructor;
eauto.
econstructor;
eauto.
eapply agree_frame_invariant;
eauto.
apply IHmatch_stacks.
intros;
apply H0;
auto.
apply Plt_trans with sp';
auto.
intros;
apply H1;
auto.
apply Plt_trans with sp';
auto.
intros;
apply H2;
auto.
apply Plt_trans with sp';
auto.
Qed.
A variant of the latter, for use with external calls
Lemma match_stacks_change_mem_extcall:
forall j m1 m2 m1'
m2'
cs cs'
sg bound bound',
match_stacks j m1 m1'
cs cs'
sg bound bound' ->
(
forall b,
Plt b bound ->
Mem.valid_block m1 b ->
Mem.valid_block m2 b) ->
(
forall b ofs p,
Plt b bound ->
Mem.perm m2 b ofs Max p ->
Mem.perm m1 b ofs Max p) ->
(
forall b,
Plt b bound' ->
Mem.valid_block m1'
b ->
Mem.valid_block m2'
b) ->
Mem.unchanged_on (
loc_out_of_reach j m1)
m1'
m2' ->
match_stacks j m2 m2'
cs cs'
sg bound bound'.
Proof.
induction 1;
intros.
econstructor;
eauto.
econstructor;
eauto.
eapply agree_frame_extcall_invariant;
eauto.
apply IHmatch_stacks.
intros;
apply H0;
auto.
apply Plt_trans with sp;
auto.
intros;
apply H1.
apply Plt_trans with sp;
auto.
auto.
intros;
apply H2;
auto.
apply Plt_trans with sp';
auto.
auto.
Qed.
Invariance with respect to change of j.
Lemma match_stacks_change_meminj:
forall j j'
m m'
m1 m1',
inject_incr j j' ->
inject_separated j j'
m1 m1' ->
forall cs cs'
sg bound bound',
match_stacks j m m'
cs cs'
sg bound bound' ->
Ple bound' (
Mem.nextblock m1') ->
match_stacks j'
m m'
cs cs'
sg bound bound'.
Proof.
induction 3;
intros.
apply match_stacks_empty with hi;
auto.
inv H3.
constructor;
auto.
intros.
red in H0.
case_eq (
j b1).
intros [
b'
delta']
EQ.
rewrite (
H _ _ _ EQ)
in H3.
inv H3.
eauto.
intros EQ.
exploit H0;
eauto.
intros [
A B].
elim B.
red.
apply Plt_le_trans with hi.
auto.
apply Ple_trans with bound';
auto.
econstructor;
eauto.
eapply agree_frame_inject_incr;
eauto.
red.
eapply Plt_le_trans;
eauto.
apply IHmatch_stacks.
apply Ple_trans with bound';
auto.
apply Plt_Ple;
auto.
Qed.
Preservation by parallel stores in Linear and Mach.
Lemma match_stacks_parallel_stores:
forall j m m'
chunk addr addr'
v v'
m1 m1',
Mem.inject j m m' ->
Val.inject j addr addr' ->
Mem.storev chunk m addr v =
Some m1 ->
Mem.storev chunk m'
addr'
v' =
Some m1' ->
forall cs cs'
sg bound bound',
match_stacks j m m'
cs cs'
sg bound bound' ->
match_stacks j m1 m1'
cs cs'
sg bound bound'.
Proof.
intros until m1'.
intros MINJ VINJ STORE1 STORE2.
induction 1.
econstructor;
eauto.
econstructor;
eauto.
eapply agree_frame_parallel_stores;
eauto.
Qed.
Invariance by external calls.
Lemma match_stack_change_extcall:
forall ec args m1 res t m2 args'
m1'
res'
t'
m2'
j j',
external_call ec ge args m1 t res m2 ->
external_call ec ge args'
m1'
t'
res'
m2' ->
inject_incr j j' ->
inject_separated j j'
m1 m1' ->
Mem.unchanged_on (
loc_out_of_reach j m1)
m1'
m2' ->
forall cs cs'
sg bound bound',
match_stacks j m1 m1'
cs cs'
sg bound bound' ->
Ple bound (
Mem.nextblock m1) ->
Ple bound' (
Mem.nextblock m1') ->
match_stacks j'
m2 m2'
cs cs'
sg bound bound'.
Proof.
Invariance with respect to change of signature
Lemma match_stacks_change_sig:
forall sg1 j m m'
cs cs'
sg bound bound',
match_stacks j m m'
cs cs'
sg bound bound' ->
tailcall_possible sg1 ->
match_stacks j m m'
cs cs'
sg1 bound bound'.
Proof.
induction 1; intros.
econstructor; eauto.
econstructor; eauto. intros. elim (H0 _ H1).
Qed.
match_stacks implies match_globalenvs, which implies meminj_preserves_globals.
Lemma match_stacks_globalenvs:
forall j m m'
cs cs'
sg bound bound',
match_stacks j m m'
cs cs'
sg bound bound' ->
exists hi,
match_globalenvs j hi.
Proof.
induction 1. exists hi; auto. auto.
Qed.
Lemma match_stacks_preserves_globals:
forall j m m'
cs cs'
sg bound bound',
match_stacks j m m'
cs cs'
sg bound bound' ->
meminj_preserves_globals ge j.
Proof.
intros.
exploit match_stacks_globalenvs;
eauto.
intros [
hi MG].
inv MG.
split.
eauto.
split.
eauto.
intros.
symmetry.
eauto.
Qed.
Typing properties of match_stacks.
Lemma match_stacks_type_sp:
forall j m m'
cs cs'
sg bound bound',
match_stacks j m m'
cs cs'
sg bound bound' ->
Val.has_type (
parent_sp cs')
Tint.
Proof.
induction 1; simpl; auto.
Qed.
Lemma match_stacks_type_retaddr:
forall j m m'
cs cs'
sg bound bound',
match_stacks j m m'
cs cs'
sg bound bound' ->
Val.has_type (
parent_ra cs')
Tint.
Proof.
induction 1; simpl; auto.
Qed.
Syntactic properties of the translation
Preservation of code labels through the translation.
Section LABELS.
Remark find_label_fold_right:
forall (
A:
Type) (
fn:
A ->
Mach.code ->
Mach.code)
lbl,
(
forall x k,
Mach.find_label lbl (
fn x k) =
Mach.find_label lbl k) ->
forall (
args:
list A)
k,
Mach.find_label lbl (
List.fold_right fn k args) =
Mach.find_label lbl k.
Proof.
induction args; simpl. auto.
intros. rewrite H. auto.
Qed.
Remark find_label_save_callee_save:
forall fe lbl k,
Mach.find_label lbl (
save_callee_save fe k) =
Mach.find_label lbl k.
Proof.
Remark find_label_restore_callee_save:
forall fe lbl k,
Mach.find_label lbl (
restore_callee_save fe k) =
Mach.find_label lbl k.
Proof.
Lemma transl_code_eq:
forall fe i c,
transl_code fe (
i ::
c) =
transl_instr fe i (
transl_code fe c).
Proof.
Lemma find_label_transl_code:
forall fe lbl c,
Mach.find_label lbl (
transl_code fe c) =
option_map (
transl_code fe) (
Linear.find_label lbl c).
Proof.
Lemma transl_find_label:
forall f tf lbl c,
transf_function f =
OK tf ->
Linear.find_label lbl f.(
Linear.fn_code) =
Some c ->
Mach.find_label lbl tf.(
Mach.fn_code) =
Some (
transl_code (
make_env (
function_bounds f))
c).
Proof.
End LABELS.
Code tail property for Linear executions.
Lemma find_label_tail:
forall lbl c c',
Linear.find_label lbl c =
Some c' ->
is_tail c'
c.
Proof.
induction c;
simpl.
intros;
discriminate.
intro c'.
case (
Linear.is_label lbl a);
intros.
injection H;
intro;
subst c'.
auto with coqlib.
auto with coqlib.
Qed.
Code tail property for translations
Lemma is_tail_save_callee_save_regs:
forall bound number mkindex ty fe csl k,
is_tail k (
save_callee_save_regs bound number mkindex ty fe csl k).
Proof.
induction csl;
intros;
simpl.
auto with coqlib.
unfold save_callee_save_reg.
destruct (
zlt (
number a) (
bound fe)).
constructor;
auto.
auto.
Qed.
Lemma is_tail_save_callee_save:
forall fe k,
is_tail k (
save_callee_save fe k).
Proof.
Lemma is_tail_restore_callee_save_regs:
forall bound number mkindex ty fe csl k,
is_tail k (
restore_callee_save_regs bound number mkindex ty fe csl k).
Proof.
induction csl;
intros;
simpl.
auto with coqlib.
unfold restore_callee_save_reg.
destruct (
zlt (
number a) (
bound fe)).
constructor;
auto.
auto.
Qed.
Lemma is_tail_restore_callee_save:
forall fe k,
is_tail k (
restore_callee_save fe k).
Proof.
Lemma is_tail_transl_instr:
forall fe i k,
is_tail k (
transl_instr fe i k).
Proof.
Lemma is_tail_transl_code:
forall fe c1 c2,
is_tail c1 c2 ->
is_tail (
transl_code fe c1) (
transl_code fe c2).
Proof.
Lemma is_tail_transf_function:
forall f tf c,
transf_function f =
OK tf ->
is_tail c (
Linear.fn_code f) ->
is_tail (
transl_code (
make_env (
function_bounds f))
c) (
fn_code tf).
Proof.
Semantic preservation
Preservation / translation of global symbols and functions.
Lemma symbols_preserved:
forall id,
Genv.find_symbol tge id =
Genv.find_symbol ge id.
Proof.
Lemma public_preserved:
forall id,
Genv.public_symbol tge id =
Genv.public_symbol ge id.
Proof.
Lemma varinfo_preserved:
forall b,
Genv.find_var_info tge b =
Genv.find_var_info ge b.
Proof.
Lemma functions_translated:
forall v f,
Genv.find_funct ge v =
Some f ->
exists tf,
Genv.find_funct tge v =
Some tf /\
transf_fundef f =
OK tf.
Proof
(
Genv.find_funct_transf_partial transf_fundef _ TRANSF).
Lemma function_ptr_translated:
forall v f,
Genv.find_funct_ptr ge v =
Some f ->
exists tf,
Genv.find_funct_ptr tge v =
Some tf /\
transf_fundef f =
OK tf.
Proof
(
Genv.find_funct_ptr_transf_partial transf_fundef _ TRANSF).
Lemma sig_preserved:
forall f tf,
transf_fundef f =
OK tf ->
Mach.funsig tf =
Linear.funsig f.
Proof.
Lemma find_function_translated:
forall j ls rs m m'
cs cs'
sg bound bound'
ros f,
agree_regs j ls rs ->
match_stacks j m m'
cs cs'
sg bound bound' ->
Linear.find_function ge ros ls =
Some f ->
exists bf,
exists tf,
find_function_ptr tge ros rs =
Some bf
/\
Genv.find_funct_ptr tge bf =
Some tf
/\
transf_fundef f =
OK tf.
Proof.
Preservation of the arguments to an external call.
Section EXTERNAL_ARGUMENTS.
Variable j:
meminj.
Variables m m':
mem.
Variable cs:
list Linear.stackframe.
Variable cs':
list stackframe.
Variable sg:
signature.
Variables bound bound':
block.
Hypothesis MS:
match_stacks j m m'
cs cs'
sg bound bound'.
Variable ls:
locset.
Variable rs:
regset.
Hypothesis AGR:
agree_regs j ls rs.
Hypothesis AGCS:
agree_callee_save ls (
parent_locset cs).
Lemma transl_external_argument:
forall l,
In l (
loc_arguments sg) ->
exists v,
extcall_arg rs m' (
parent_sp cs')
l v /\
Val.inject j (
ls l)
v.
Proof.
Lemma transl_external_arguments_rec:
forall locs,
incl locs (
loc_arguments sg) ->
exists vl,
list_forall2 (
extcall_arg rs m' (
parent_sp cs'))
locs vl /\
Val.inject_list j ls##
locs vl.
Proof.
induction locs;
simpl;
intros.
exists (@
nil val);
split.
constructor.
constructor.
exploit transl_external_argument;
eauto with coqlib.
intros [
v [
A B]].
exploit IHlocs;
eauto with coqlib.
intros [
vl [
C D]].
exists (
v ::
vl);
split;
constructor;
auto.
Qed.
Lemma transl_external_arguments:
exists vl,
extcall_arguments rs m' (
parent_sp cs')
sg vl /\
Val.inject_list j (
ls ## (
loc_arguments sg))
vl.
Proof.
End EXTERNAL_ARGUMENTS.
Preservation of the arguments to a builtin.
Section BUILTIN_ARGUMENTS.
Variable f:
Linear.function.
Let b :=
function_bounds f.
Let fe :=
make_env b.
Variable tf:
Mach.function.
Hypothesis TRANSF_F:
transf_function f =
OK tf.
Variable j:
meminj.
Variables m m':
mem.
Variables ls ls0:
locset.
Variable rs:
regset.
Variables sp sp':
block.
Variables parent retaddr:
val.
Hypothesis AGR:
agree_regs j ls rs.
Hypothesis AGF:
agree_frame f j ls ls0 m sp m'
sp'
parent retaddr.
Hypothesis MINJ:
Mem.inject j m m'.
Hypothesis GINJ:
meminj_preserves_globals ge j.
Lemma transl_builtin_arg_correct:
forall a v,
eval_builtin_arg ge ls (
Vptr sp Int.zero)
m a v ->
(
forall l,
In l (
params_of_builtin_arg a) ->
loc_valid f l =
true) ->
(
forall sl ofs ty,
In (
S sl ofs ty) (
params_of_builtin_arg a) ->
slot_within_bounds b sl ofs ty) ->
exists v',
eval_builtin_arg ge rs (
Vptr sp'
Int.zero)
m' (
transl_builtin_arg fe a)
v'
/\
Val.inject j v v'.
Proof.
Lemma transl_builtin_args_correct:
forall al vl,
eval_builtin_args ge ls (
Vptr sp Int.zero)
m al vl ->
(
forall l,
In l (
params_of_builtin_args al) ->
loc_valid f l =
true) ->
(
forall sl ofs ty,
In (
S sl ofs ty) (
params_of_builtin_args al) ->
slot_within_bounds b sl ofs ty) ->
exists vl',
eval_builtin_args ge rs (
Vptr sp'
Int.zero)
m' (
List.map (
transl_builtin_arg fe)
al)
vl'
/\
Val.inject_list j vl vl'.
Proof.
induction 1;
simpl;
intros VALID BOUNDS.
-
exists (@
nil val);
split;
constructor.
-
exploit transl_builtin_arg_correct;
eauto using in_or_app.
intros (
v1' &
A &
B).
exploit IHlist_forall2;
eauto using in_or_app.
intros (
vl' &
C &
D).
exists (
v1'::
vl');
split;
constructor;
auto.
Qed.
End BUILTIN_ARGUMENTS.
The proof of semantic preservation relies on simulation diagrams
of the following form:
st1 --------------- st2
| |
t| +|t
| |
v v
st1'--------------- st2'
Matching between source and target states is defined by
match_states
below. It implies:
-
Agreement between, on the Linear side, the location sets ls
and parent_locset s of the current function and its caller,
and on the Mach side the register set rs and the contents of
the memory area corresponding to the stack frame.
-
The Linear code c is a suffix of the code of the
function f being executed.
-
Memory injection between the Linear and the Mach memory states.
-
Well-typedness of f.
Inductive match_states:
Linear.state ->
Mach.state ->
Prop :=
|
match_states_intro:
forall cs f sp c ls m cs'
fb sp'
rs m'
j tf
(
MINJ:
Mem.inject j m m')
(
STACKS:
match_stacks j m m'
cs cs'
f.(
Linear.fn_sig)
sp sp')
(
TRANSL:
transf_function f =
OK tf)
(
FIND:
Genv.find_funct_ptr tge fb =
Some (
Internal tf))
(
AGREGS:
agree_regs j ls rs)
(
AGFRAME:
agree_frame f j ls (
parent_locset cs)
m sp m'
sp' (
parent_sp cs') (
parent_ra cs'))
(
TAIL:
is_tail c (
Linear.fn_code f)),
match_states (
Linear.State cs f (
Vptr sp Int.zero)
c ls m)
(
Mach.State cs'
fb (
Vptr sp'
Int.zero) (
transl_code (
make_env (
function_bounds f))
c)
rs m')
|
match_states_call:
forall cs f ls m cs'
fb rs m'
j tf
(
MINJ:
Mem.inject j m m')
(
STACKS:
match_stacks j m m'
cs cs' (
Linear.funsig f) (
Mem.nextblock m) (
Mem.nextblock m'))
(
TRANSL:
transf_fundef f =
OK tf)
(
FIND:
Genv.find_funct_ptr tge fb =
Some tf)
(
AGREGS:
agree_regs j ls rs)
(
AGLOCS:
agree_callee_save ls (
parent_locset cs)),
match_states (
Linear.Callstate cs f ls m)
(
Mach.Callstate cs'
fb rs m')
|
match_states_return:
forall cs ls m cs'
rs m'
j sg
(
MINJ:
Mem.inject j m m')
(
STACKS:
match_stacks j m m'
cs cs'
sg (
Mem.nextblock m) (
Mem.nextblock m'))
(
AGREGS:
agree_regs j ls rs)
(
AGLOCS:
agree_callee_save ls (
parent_locset cs)),
match_states (
Linear.Returnstate cs ls m)
(
Mach.Returnstate cs'
rs m').
Theorem transf_step_correct:
forall s1 t s2,
Linear.step ge s1 t s2 ->
forall (
WTS:
wt_state s1)
s1' (
MS:
match_states s1 s1'),
exists s2',
plus step tge s1'
t s2' /\
match_states s2 s2'.
Proof.
Lemma transf_initial_states:
forall st1,
Linear.initial_state prog st1 ->
exists st2,
Mach.initial_state tprog st2 /\
match_states st1 st2.
Proof.
Lemma transf_final_states:
forall st1 st2 r,
match_states st1 st2 ->
Linear.final_state st1 r ->
Mach.final_state st2 r.
Proof.
intros. inv H0. inv H. inv STACKS.
generalize (AGREGS r0). rewrite H2. intros A; inv A.
econstructor; eauto.
Qed.
Lemma wt_prog:
forall i fd,
In (
i,
Gfun fd)
prog.(
prog_defs) ->
wt_fundef fd.
Proof.
Theorem transf_program_correct:
forward_simulation (
Linear.semantics prog) (
Mach.semantics return_address_offset tprog).
Proof.
Inductive match_sp (
ge:
Genv.t fundef unit) (
j:
meminj) :
val -> (
block *
val) ->
Prop :=
|
match_sp_intro:
forall fb f tf sp tsp,
Genv.find_funct_ptr ge fb =
Some (
Internal tf) ->
j sp =
Some (
tsp,
tf.(
fn_stackdata)) ->
transf_function f =
OK tf ->
match_sp ge j (
Vptr sp Int.zero) (
fb,
Vptr tsp Int.zero).
Lemma list_forall2_in_1:
forall A B P l1 l2 (
a:
A),
list_forall2 P l1 l2 ->
In a l1 ->
exists (
b:
B),
In b l2 /\
P a b.
Proof.
induction l1;
intros.
-
inv H0.
-
destruct H0 as [
H0 |
H0].
+
subst a0.
inv H.
exists b1;
split;
auto;
eapply in_eq.
+
inv H.
exploit IHl1;
eauto.
intros [
b [
HA HB]].
exists b;
split;
auto;
eapply in_cons;
auto.
Qed.
Lemma annot_sem_to_sem_annot:
forall ge j sps tsps alpha a a',
annot_sem (
Genv.find_symbol ge)
sps alpha a ->
list_forall2 (
match_sp ge j)
sps tsps ->
Val.inject j a a' ->
(
forall id b,
Genv.find_symbol ge id =
Some b ->
j b =
Some (
b, 0)) ->
sem_annot ge tsps alpha a'.
Proof.
intros.
destruct H as [
H | [
H |
H]].
-
subst alpha;
left;
reflexivity.
-
right.
left.
destruct H as [
depth [
varname [
base [
bound [
HA [
sp [
ofs [
HB [
HC HD]]]]]]]]].
exists depth,
varname,
base,
bound.
split;
auto.
generalize (
nth_error_in _ _ HB).
intros Hin.
exploit list_forall2_in_1;
eauto.
intros [
fbtsp [
XA XB]].
destruct fbtsp as [
fb tsp].
inv XB.
exists fb,
tf, (
Vptr tsp0 Int.zero),
ofs.
repeat split;
auto.
+
simpl in H1.
simpl.
rewrite Int.add_zero_l in H1.
inv H1.
rewrite H8 in H6;
inv H6.
rewrite Int.add_zero_l.
reflexivity.
+
omega.
+
omega.
-
right.
right.
destruct H as [
id [
base [
bound [
HA [
b [
ofs [
HB [
HC HD]]]]]]]].
generalize (
H2 _ _ HB).
intros HE.
subst a.
inv H1.
rewrite H4 in HE;
inv HE.
rewrite Int.add_zero.
repeat eexists;
eauto;
omega.
Qed.
Definition ok_injective {
A} {
a a':
A} (
H:
OK a =
OK a') :
a =
a' :=
let '
eq_refl :=
H in eq_refl.
Lemma match_stacks_implies_match_sp:
forall j m m'
cs cs'
sig bound bound',
match_stacks j m m'
cs cs'
sig bound bound' ->
list_forall2 (
match_sp tge j) (
map (
fun s =>
match s with |
Linear.Stackframe _ sp _ _ =>
sp end)
cs) (
map (
fun s =>
match s with |
Stackframe fb sp _ _ => (
fb,
sp)
end)
cs').
Proof.
Theorem transf_step_correct':
forall s1 t s2,
Linear.step_safe ge s1 t s2 ->
forall (
WTS:
wt_state s1)
s1' (
MS:
match_states s1 s1'),
exists s2',
plus step_safe tge s1'
t s2' /\
match_states s2 s2'.
Proof.
Theorem transf_program_correct':
forward_simulation (
Linear.semantics_safe prog) (
Mach.semantics_safe return_address_offset tprog).
Proof.
End PRESERVATION.