Correctness proof for common subexpression elimination.
Require Import Coqlib.
Require Import Maps.
Require Import AST.
Require Import Errors.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Memory.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Op.
Require Import Registers.
Require Import RTL.
Require Import RTLtyping.
Require Import Kildall.
Require Import CombineOp.
Require Import CombineOpproof.
Require Import CSE.
Semantic properties of value numberings
Well-formedness of numberings
A numbering is well-formed if all registers mentioned in equations
are less than the ``next'' register number given in the numbering.
This guarantees that the next register is fresh with respect to
the equations.
Definition wf_rhs (
next:
valnum) (
rh:
rhs) :
Prop :=
match rh with
|
Op op vl =>
forall v,
In v vl ->
Plt v next
|
Load chunk addr vl =>
forall v,
In v vl ->
Plt v next
end.
Definition wf_equation (
next:
valnum) (
vr:
valnum) (
rh:
rhs) :
Prop :=
Plt vr next /\
wf_rhs next rh.
Inductive wf_numbering (
n:
numbering) :
Prop :=
wf_numbering_intro
(
EQS:
forall v rh,
In (
v,
rh)
n.(
num_eqs) ->
wf_equation n.(
num_next)
v rh)
(
REG:
forall r v,
PTree.get r n.(
num_reg) =
Some v ->
Plt v n.(
num_next))
(
VAL:
forall r v,
In r (
PMap.get v n.(
num_val)) ->
PTree.get r n.(
num_reg) =
Some v).
Lemma wf_empty_numbering:
wf_numbering empty_numbering.
Proof.
unfold empty_numbering;
split;
simpl;
intros.
contradiction.
rewrite PTree.gempty in H.
congruence.
rewrite PMap.gi in H.
contradiction.
Qed.
Lemma wf_rhs_increasing:
forall next1 next2 rh,
Ple next1 next2 ->
wf_rhs next1 rh ->
wf_rhs next2 rh.
Proof.
intros;
destruct rh;
simpl;
intros;
apply Plt_Ple_trans with next1;
auto.
Qed.
Lemma wf_equation_increasing:
forall next1 next2 vr rh,
Ple next1 next2 ->
wf_equation next1 vr rh ->
wf_equation next2 vr rh.
Proof.
We now show that all operations over numberings
preserve well-formedness.
Lemma wf_valnum_reg:
forall n r n'
v,
wf_numbering n ->
valnum_reg n r = (
n',
v) ->
wf_numbering n' /\
Plt v n'.(
num_next) /\
Ple n.(
num_next)
n'.(
num_next).
Proof.
Lemma wf_valnum_regs:
forall rl n n'
vl,
wf_numbering n ->
valnum_regs n rl = (
n',
vl) ->
wf_numbering n' /\
(
forall v,
In v vl ->
Plt v n'.(
num_next)) /\
Ple n.(
num_next)
n'.(
num_next).
Proof.
induction rl;
intros.
simpl in H0.
inversion H0.
subst n';
subst vl.
simpl.
intuition.
simpl in H0.
caseEq (
valnum_reg n a).
intros n1 v1 EQ1.
caseEq (
valnum_regs n1 rl).
intros ns vs EQS.
rewrite EQ1 in H0;
rewrite EQS in H0;
simpl in H0.
inversion H0.
subst n';
subst vl.
generalize (
wf_valnum_reg _ _ _ _ H EQ1);
intros [
A1 [
B1 C1]].
generalize (
IHrl _ _ _ A1 EQS);
intros [
As [
Bs Cs]].
split.
auto.
split.
simpl;
intros.
elim H1;
intro.
subst v.
eapply Plt_Ple_trans;
eauto.
auto.
eapply Ple_trans;
eauto.
Qed.
Lemma find_valnum_rhs_correct:
forall rh vn eqs,
find_valnum_rhs rh eqs =
Some vn ->
In (
vn,
rh)
eqs.
Proof.
induction eqs;
simpl.
congruence.
case a;
intros v r'.
case (
eq_rhs rh r');
intro.
intro.
left.
congruence.
intro.
right.
auto.
Qed.
Lemma find_valnum_num_correct:
forall rh vn eqs,
find_valnum_num vn eqs =
Some rh ->
In (
vn,
rh)
eqs.
Proof.
induction eqs;
simpl.
congruence.
destruct a as [
v'
r'].
destruct (
peq vn v');
intros.
left.
congruence.
right.
auto.
Qed.
Remark in_remove:
forall (
A:
Type) (
eq:
forall (
x y:
A), {
x=
y}+{
x<>
y})
x y l,
In y (
List.remove eq x l) <->
x <>
y /\
In y l.
Proof.
induction l; simpl.
tauto.
destruct (eq x a).
subst a. rewrite IHl. tauto.
simpl. rewrite IHl. intuition congruence.
Qed.
Lemma wf_forget_reg:
forall n rd r v,
wf_numbering n ->
In r (
PMap.get v (
forget_reg n rd)) ->
r <>
rd /\
PTree.get r n.(
num_reg) =
Some v.
Proof.
unfold forget_reg;
intros.
inversion H.
destruct ((
num_reg n)!
rd)
as [
vd|]
_eqn.
rewrite PMap.gsspec in H0.
destruct (
peq v vd).
subst vd.
rewrite in_remove in H0.
destruct H0.
split.
auto.
eauto.
split;
eauto.
exploit VAL;
eauto.
congruence.
split;
eauto.
exploit VAL;
eauto.
congruence.
Qed.
Lemma wf_update_reg:
forall n rd vd r v,
wf_numbering n ->
In r (
PMap.get v (
update_reg n rd vd)) ->
PTree.get r (
PTree.set rd vd n.(
num_reg)) =
Some v.
Proof.
Lemma wf_add_rhs:
forall n rd rh,
wf_numbering n ->
wf_rhs n.(
num_next)
rh ->
wf_numbering (
add_rhs n rd rh).
Proof.
Lemma wf_add_op:
forall n rd op rs,
wf_numbering n ->
wf_numbering (
add_op n rd op rs).
Proof.
Lemma wf_add_load:
forall n rd chunk addr rs,
wf_numbering n ->
wf_numbering (
add_load n rd chunk addr rs).
Proof.
Lemma wf_add_unknown:
forall n rd,
wf_numbering n ->
wf_numbering (
add_unknown n rd).
Proof.
Remark kill_eqs_in:
forall pred v rhs eqs,
In (
v,
rhs) (
kill_eqs pred eqs) ->
In (
v,
rhs)
eqs /\
pred rhs =
false.
Proof.
induction eqs;
simpl;
intros.
tauto.
destruct (
pred (
snd a))
as []
_eqn.
exploit IHeqs;
eauto.
tauto.
simpl in H;
destruct H.
subst a.
auto.
exploit IHeqs;
eauto.
tauto.
Qed.
Lemma wf_kill_equations:
forall pred n,
wf_numbering n ->
wf_numbering (
kill_equations pred n).
Proof.
intros.
inversion H.
unfold kill_equations;
split;
simpl;
intros.
exploit kill_eqs_in;
eauto.
intros [
A B].
auto.
eauto.
eauto.
Qed.
Lemma wf_add_store:
forall n chunk addr rargs rsrc,
wf_numbering n ->
wf_numbering (
add_store n chunk addr rargs rsrc).
Proof.
Lemma wf_transfer:
forall f pc n,
wf_numbering n ->
wf_numbering (
transfer f pc n).
Proof.
As a consequence, the numberings computed by the static analysis
are well formed.
Theorem wf_analyze:
forall f approx pc,
analyze f =
Some approx ->
wf_numbering approx!!
pc.
Proof.
Properties of satisfiability of numberings
Module ValnumEq.
Definition t :=
valnum.
Definition eq :=
peq.
End ValnumEq.
Module VMap :=
EMap(
ValnumEq).
Section SATISFIABILITY.
Variable ge:
genv.
Variable sp:
val.
Agremment between two mappings from value numbers to values
up to a given value number.
Definition valu_agree (
valu1 valu2:
valnum ->
val) (
upto:
valnum) :
Prop :=
forall v,
Plt v upto ->
valu2 v =
valu1 v.
Lemma valu_agree_refl:
forall valu upto,
valu_agree valu valu upto.
Proof.
intros; red; auto.
Qed.
Lemma valu_agree_trans:
forall valu1 valu2 valu3 upto12 upto23,
valu_agree valu1 valu2 upto12 ->
valu_agree valu2 valu3 upto23 ->
Ple upto12 upto23 ->
valu_agree valu1 valu3 upto12.
Proof.
intros;
red;
intros.
transitivity (
valu2 v).
apply H0.
apply Plt_Ple_trans with upto12;
auto.
apply H;
auto.
Qed.
Lemma valu_agree_list:
forall valu1 valu2 upto vl,
valu_agree valu1 valu2 upto ->
(
forall v,
In v vl ->
Plt v upto) ->
map valu2 vl =
map valu1 vl.
Proof.
The numbering_holds predicate (defined in file CSE) is
extensional with respect to valu_agree.
Lemma numbering_holds_exten:
forall valu1 valu2 n rs m,
valu_agree valu1 valu2 n.(
num_next) ->
wf_numbering n ->
numbering_holds valu1 ge sp rs m n ->
numbering_holds valu2 ge sp rs m n.
Proof.
intros.
inversion H0.
inversion H1.
split;
intros.
exploit EQS;
eauto.
intros [
A B].
red in B.
generalize (
H2 _ _ H4).
unfold equation_holds;
destruct rh.
rewrite (
valu_agree_list valu1 valu2 n.(
num_next)).
rewrite H.
auto.
auto.
auto.
auto.
rewrite (
valu_agree_list valu1 valu2 n.(
num_next)).
rewrite H.
auto.
auto.
auto.
auto.
rewrite H.
auto.
eauto.
Qed.
valnum_reg and valnum_regs preserve the numbering_holds predicate.
Moreover, it is always the case that the returned value number has
the value of the given register in the final assignment of values to
value numbers.
Lemma valnum_reg_holds:
forall valu1 rs n r n'
v m,
wf_numbering n ->
numbering_holds valu1 ge sp rs m n ->
valnum_reg n r = (
n',
v) ->
exists valu2,
numbering_holds valu2 ge sp rs m n' /\
valu2 v =
rs#
r /\
valu_agree valu1 valu2 n.(
num_next).
Proof.
intros until v.
unfold valnum_reg.
caseEq (
n.(
num_reg)!
r).
intros.
inversion H2.
subst n';
subst v0.
inversion H1.
exists valu1.
split.
auto.
split.
symmetry.
auto.
apply valu_agree_refl.
intros.
inversion H2.
subst n'.
subst v.
inversion H1.
set (
valu2 :=
VMap.set n.(
num_next)
rs#
r valu1).
assert (
AG:
valu_agree valu1 valu2 n.(
num_next)).
red;
intros.
unfold valu2.
apply VMap.gso.
auto with coqlib.
destruct (
numbering_holds_exten _ _ _ _ _ AG H0 H1)
as [
A B].
exists valu2.
split.
split;
simpl;
intros.
auto.
unfold valu2,
VMap.set,
ValnumEq.eq.
rewrite PTree.gsspec in H5.
destruct (
peq r0 r).
inv H5.
rewrite peq_true.
auto.
rewrite peq_false.
auto.
assert (
Plt vn (
num_next n)).
inv H0.
eauto.
red;
intros;
subst;
eapply Plt_strict;
eauto.
split.
unfold valu2.
rewrite VMap.gss.
auto.
auto.
Qed.
Lemma valnum_regs_holds:
forall rs rl valu1 n n'
vl m,
wf_numbering n ->
numbering_holds valu1 ge sp rs m n ->
valnum_regs n rl = (
n',
vl) ->
exists valu2,
numbering_holds valu2 ge sp rs m n' /\
List.map valu2 vl =
rs##
rl /\
valu_agree valu1 valu2 n.(
num_next).
Proof.
induction rl;
simpl;
intros.
inversion H1;
subst n';
subst vl.
exists valu1.
split.
auto.
split.
auto.
apply valu_agree_refl.
caseEq (
valnum_reg n a);
intros n1 v1 EQ1.
caseEq (
valnum_regs n1 rl);
intros ns vs EQs.
rewrite EQ1 in H1;
rewrite EQs in H1.
inversion H1.
subst vl;
subst n'.
generalize (
valnum_reg_holds _ _ _ _ _ _ _ H H0 EQ1).
intros [
valu2 [
A [
B C]]].
generalize (
wf_valnum_reg _ _ _ _ H EQ1).
intros [
D [
E F]].
generalize (
IHrl _ _ _ _ _ D A EQs).
intros [
valu3 [
P [
Q R]]].
exists valu3.
split.
auto.
split.
simpl.
rewrite R.
congruence.
auto.
eapply valu_agree_trans;
eauto.
Qed.
A reformulation of the equation_holds predicate in terms
of the value to which a right-hand side of an equation evaluates.
Definition rhs_evals_to
(
valu:
valnum ->
val) (
rh:
rhs) (
m:
mem) (
v:
val) :
Prop :=
match rh with
|
Op op vl =>
eval_operation ge sp op (
List.map valu vl)
m =
Some v
|
Load chunk addr vl =>
exists a,
eval_addressing ge sp addr (
List.map valu vl) =
Some a /\
Mem.loadv chunk m a =
Some v
end.
Lemma equation_evals_to_holds_1:
forall valu rh v vres m,
rhs_evals_to valu rh m v ->
equation_holds valu ge sp m vres rh ->
valu vres =
v.
Proof.
intros until m. unfold rhs_evals_to, equation_holds.
destruct rh. congruence.
intros [a1 [A1 B1]] [a2 [A2 B2]]. congruence.
Qed.
Lemma equation_evals_to_holds_2:
forall valu rh v vres m,
wf_rhs vres rh ->
rhs_evals_to valu rh m v ->
equation_holds (
VMap.set vres v valu)
ge sp m vres rh.
Proof.
The numbering obtained by adding an equation rd = rhs is satisfiable
in a concrete register set where rd is set to the value of rhs.
Lemma add_rhs_satisfiable_gen:
forall n rh valu1 rs rd rs'
m,
wf_numbering n ->
wf_rhs n.(
num_next)
rh ->
numbering_holds valu1 ge sp rs m n ->
rhs_evals_to valu1 rh m (
rs'#
rd) ->
(
forall r,
r <>
rd ->
rs'#
r =
rs#
r) ->
numbering_satisfiable ge sp rs'
m (
add_rhs n rd rh).
Proof.
Lemma add_rhs_satisfiable:
forall n rh valu1 rs rd v m,
wf_numbering n ->
wf_rhs n.(
num_next)
rh ->
numbering_holds valu1 ge sp rs m n ->
rhs_evals_to valu1 rh m v ->
numbering_satisfiable ge sp (
rs#
rd <-
v)
m (
add_rhs n rd rh).
Proof.
add_op returns a numbering that is satisfiable in the register
set after execution of the corresponding Iop instruction.
Lemma add_op_satisfiable:
forall n rs op args dst v m,
wf_numbering n ->
numbering_satisfiable ge sp rs m n ->
eval_operation ge sp op rs##
args m =
Some v ->
numbering_satisfiable ge sp (
rs#
dst <-
v)
m (
add_op n dst op args).
Proof.
add_load returns a numbering that is satisfiable in the register
set after execution of the corresponding Iload instruction.
Lemma add_load_satisfiable:
forall n rs chunk addr args dst a v m,
wf_numbering n ->
numbering_satisfiable ge sp rs m n ->
eval_addressing ge sp addr rs##
args =
Some a ->
Mem.loadv chunk m a =
Some v ->
numbering_satisfiable ge sp (
rs#
dst <-
v)
m (
add_load n dst chunk addr args).
Proof.
intros.
inversion H0.
unfold add_load.
caseEq (
valnum_regs n args).
intros n1 vl VRL.
generalize (
valnum_regs_holds _ _ _ _ _ _ _ H H3 VRL).
intros [
valu2 [
A [
B C]]].
generalize (
wf_valnum_regs _ _ _ _ H VRL).
intros [
D [
E F]].
apply add_rhs_satisfiable with valu2;
auto.
simpl.
exists a;
split;
congruence.
Qed.
add_unknown returns a numbering that is satisfiable in the
register set after setting the target register to any value.
Lemma add_unknown_satisfiable:
forall n rs dst v m,
wf_numbering n ->
numbering_satisfiable ge sp rs m n ->
numbering_satisfiable ge sp (
rs#
dst <-
v)
m (
add_unknown n dst).
Proof.
Satisfiability of kill_equations.
Lemma kill_equations_holds:
forall pred valu n rs m m',
(
forall v r,
equation_holds valu ge sp m v r ->
pred r =
false ->
equation_holds valu ge sp m'
v r) ->
numbering_holds valu ge sp rs m n ->
numbering_holds valu ge sp rs m' (
kill_equations pred n).
Proof.
intros.
destruct H0 as [
A B].
red;
simpl.
split;
intros.
exploit kill_eqs_in;
eauto.
intros [
C D].
eauto.
auto.
Qed.
kill_loads preserves satisfiability. Moreover, the resulting numbering
is satisfiable in any concrete memory state.
Lemma kill_loads_satisfiable:
forall n rs m m',
numbering_satisfiable ge sp rs m n ->
numbering_satisfiable ge sp rs m' (
kill_loads n).
Proof.
add_store returns a numbering that is satisfiable in the memory state
after execution of the corresponding Istore instruction.
Lemma add_store_satisfiable:
forall n rs chunk addr args src a m m',
wf_numbering n ->
numbering_satisfiable ge sp rs m n ->
eval_addressing ge sp addr rs##
args =
Some a ->
Mem.storev chunk m a (
rs#
src) =
Some m' ->
Val.has_type (
rs#
src) (
type_of_chunk chunk) ->
numbering_satisfiable ge sp rs m' (
add_store n chunk addr args src).
Proof.
Correctness of reg_valnum: if it returns a register r,
that register correctly maps back to the given value number.
Lemma reg_valnum_correct:
forall n v r,
wf_numbering n ->
reg_valnum n v =
Some r ->
n.(
num_reg)!
r =
Some v.
Proof.
unfold reg_valnum;
intros.
inv H.
destruct ((
num_val n)#
v)
as [|
r1 rl]
_eqn;
inv H0.
eapply VAL.
rewrite Heql.
auto with coqlib.
Qed.
Correctness of find_rhs: if successful and in a
satisfiable numbering, the returned register does contain the
result value of the operation or memory load.
Lemma find_rhs_correct:
forall valu rs m n rh r,
wf_numbering n ->
numbering_holds valu ge sp rs m n ->
find_rhs n rh =
Some r ->
rhs_evals_to valu rh m rs#
r.
Proof.
Correctness of operator reduction
Section REDUCE.
Variable A:
Type.
Variable f: (
valnum ->
option rhs) ->
A ->
list valnum ->
option (
A *
list valnum).
Variable V:
Type.
Variable rs:
regset.
Variable m:
mem.
Variable sem:
A ->
list val ->
option V.
Hypothesis f_sound:
forall eqs valu op args op'
args',
(
forall v rhs,
eqs v =
Some rhs ->
equation_holds valu ge sp m v rhs) ->
f eqs op args =
Some(
op',
args') ->
sem op' (
map valu args') =
sem op (
map valu args).
Variable n:
numbering.
Variable valu:
valnum ->
val.
Hypothesis n_holds:
numbering_holds valu ge sp rs m n.
Hypothesis n_wf:
wf_numbering n.
Lemma regs_valnums_correct:
forall vl rl,
regs_valnums n vl =
Some rl ->
rs##
rl =
map valu vl.
Proof.
Lemma reduce_rec_sound:
forall niter op args op'
rl'
res,
reduce_rec A f n niter op args =
Some(
op',
rl') ->
sem op (
map valu args) =
Some res ->
sem op' (
rs##
rl') =
Some res.
Proof.
Lemma reduce_sound:
forall op rl vl op'
rl'
res,
reduce A f n op rl vl = (
op',
rl') ->
map valu vl =
rs##
rl ->
sem op rs##
rl =
Some res ->
sem op'
rs##
rl' =
Some res.
Proof.
unfold reduce;
intros.
destruct (
reduce_rec A f n 4%
nat op vl)
as [[
op1 rl1] | ]
_eqn;
inv H.
eapply reduce_rec_sound;
eauto.
congruence.
auto.
Qed.
End REDUCE.
End SATISFIABILITY.
The numberings associated to each instruction by the static analysis
are inductively satisfiable, in the following sense: the numbering
at the function entry point is satisfiable, and for any RTL execution
from pc to pc', satisfiability at pc implies
satisfiability at pc'.
Theorem analysis_correct_1:
forall ge sp rs m f approx pc pc'
i,
analyze f =
Some approx ->
f.(
fn_code)!
pc =
Some i ->
In pc' (
successors_instr i) ->
numbering_satisfiable ge sp rs m (
transfer f pc approx!!
pc) ->
numbering_satisfiable ge sp rs m approx!!
pc'.
Proof.
Theorem analysis_correct_entry:
forall ge sp rs m f approx,
analyze f =
Some approx ->
numbering_satisfiable ge sp rs m approx!!(
f.(
fn_entrypoint)).
Proof.
Semantic preservation
Section PRESERVATION.
Variable prog:
program.
Variable tprog :
program.
Hypothesis TRANSF:
transf_program prog =
OK tprog.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Lemma symbols_preserved:
forall (
s:
ident),
Genv.find_symbol tge s =
Genv.find_symbol ge s.
Proof (
Genv.find_symbol_transf_partial transf_fundef prog TRANSF).
Lemma varinfo_preserved:
forall b,
Genv.find_var_info tge b =
Genv.find_var_info ge b.
Proof (
Genv.find_var_info_transf_partial transf_fundef prog TRANSF).
Lemma functions_translated:
forall (
v:
val) (
f:
RTL.fundef),
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 prog TRANSF).
Lemma funct_ptr_translated:
forall (
b:
block) (
f:
RTL.fundef),
Genv.find_funct_ptr ge b =
Some f ->
exists tf,
Genv.find_funct_ptr tge b =
Some tf /\
transf_fundef f =
OK tf.
Proof (
Genv.find_funct_ptr_transf_partial transf_fundef prog TRANSF).
Lemma sig_preserved:
forall f tf,
transf_fundef f =
OK tf ->
funsig tf =
funsig f.
Proof.
unfold transf_fundef;
intros.
destruct f;
monadInv H;
auto.
unfold transf_function in EQ.
destruct (
type_function f);
try discriminate.
destruct (
analyze f);
try discriminate.
inv EQ;
auto.
Qed.
Lemma find_function_translated:
forall ros rs f,
find_function ge ros rs =
Some f ->
exists tf,
find_function tge ros rs =
Some tf /\
transf_fundef f =
OK tf.
Proof.
Definition transf_function' (
f:
function) (
approxs:
PMap.t numbering) :
function :=
mkfunction
f.(
fn_sig)
f.(
fn_params)
f.(
fn_stacksize)
(
transf_code approxs f.(
fn_code))
f.(
fn_entrypoint).
The proof of semantic preservation is a simulation argument using
diagrams of the following form:
st1 --------------- st2
| |
t| |t
| |
v v
st1'--------------- st2'
Left: RTL execution in the original program. Right: RTL execution in
the optimized program. Precondition (top) and postcondition (bottom):
agreement between the states, including the fact that
the numbering at
pc (returned by the static analysis) is satisfiable.
Inductive match_stackframes:
list stackframe ->
list stackframe ->
typ ->
Prop :=
|
match_stackframes_nil:
match_stackframes nil nil Tint
|
match_stackframes_cons:
forall res sp pc rs f approx tyenv s s'
ty
(
ANALYZE:
analyze f =
Some approx)
(
WTF:
wt_function f tyenv)
(
WTREGS:
wt_regset tyenv rs)
(
TYRES:
tyenv res =
ty)
(
SAT:
forall v m,
numbering_satisfiable ge sp (
rs#
res <-
v)
m approx!!
pc)
(
STACKS:
match_stackframes s s' (
proj_sig_res (
fn_sig f))),
match_stackframes
(
Stackframe res f sp pc rs ::
s)
(
Stackframe res (
transf_function'
f approx)
sp pc rs ::
s')
ty.
Inductive match_states:
state ->
state ->
Prop :=
|
match_states_intro:
forall s sp pc rs m s'
f approx tyenv
(
ANALYZE:
analyze f =
Some approx)
(
WTF:
wt_function f tyenv)
(
WTREGS:
wt_regset tyenv rs)
(
SAT:
numbering_satisfiable ge sp rs m approx!!
pc)
(
STACKS:
match_stackframes s s' (
proj_sig_res (
fn_sig f))),
match_states (
State s f sp pc rs m)
(
State s' (
transf_function'
f approx)
sp pc rs m)
|
match_states_call:
forall s f tf args m s',
match_stackframes s s' (
proj_sig_res (
funsig f)) ->
Val.has_type_list args (
sig_args (
funsig f)) ->
transf_fundef f =
OK tf ->
match_states (
Callstate s f args m)
(
Callstate s'
tf args m)
|
match_states_return:
forall s s'
ty v m,
match_stackframes s s'
ty ->
Val.has_type v ty ->
match_states (
Returnstate s v m)
(
Returnstate s'
v m).
Ltac TransfInstr :=
match goal with
|
H1: (
PTree.get ?
pc ?
c =
Some ?
instr),
f:
function,
approx:
PMap.t numbering |-
_ =>
cut ((
transf_function'
f approx).(
fn_code)!
pc =
Some(
transf_instr approx!!
pc instr));
[
simpl transf_instr
|
unfold transf_function',
transf_code;
simpl;
rewrite PTree.gmap;
unfold option_map;
rewrite H1;
reflexivity ]
end.
The proof of simulation is a case analysis over the transition
in the source code.
Lemma transf_step_correct:
forall s1 t s2,
step ge s1 t s2 ->
forall s1' (
MS:
match_states s1 s1'),
exists s2',
step tge s1'
t s2' /\
match_states s2 s2'.
Proof.
Lemma transf_initial_states:
forall st1,
initial_state prog st1 ->
exists st2,
initial_state tprog st2 /\
match_states st1 st2.
Proof.
Lemma transf_final_states:
forall st1 st2 r,
match_states st1 st2 ->
final_state st1 r ->
final_state st2 r.
Proof.
intros. inv H0. inv H. inv H4. constructor.
Qed.
Theorem transf_program_correct:
forward_simulation (
RTL.semantics prog) (
RTL.semantics tprog).
Proof.
End PRESERVATION.