Require Import Coqlib Maps.
Require Import AST Values Events Globalenvs Smallstep.
Require Import RTL Op OptionMonad.
Require Import Errors Linking IPass BTLtoRTL.
Module BTLtoRTL_Translationproof (
B:
BTLtoRTL_TranslationOracle).
Include BTLtoRTL_Translation B.
Inductive match_fundef:
BTL.fundef ->
RTL.fundef ->
Prop :=
|
match_Internal dupmap f f':
match_function dupmap f f' false ->
match_fundef (
Internal f) (
Internal f')
|
match_External ef:
match_fundef (
External ef) (
External ef).
Inductive match_stackframes:
BTL.stackframe ->
RTL.stackframe ->
Prop :=
|
match_stackframe_intro
dupmap res f sp pc rs f' pc'
(
TRANSF:
match_function dupmap f f' false)
(
DUPLIC:
dupmap!
pc =
Some pc')
:
match_stackframes (
BTL.Stackframe res f sp pc rs) (
RTL.Stackframe res f' sp pc' rs).
Inductive match_states:
BTL.state ->
RTL.state ->
Prop :=
|
match_states_intro
dupmap st f sp pc rs m bc st' f' pc'
(
STACKS:
list_forall2 match_stackframes st st')
(
TRANSF:
match_function dupmap f f' false)
(
DUPLIC:
dupmap!
pc =
Some pc')
:
match_states (
State st f sp pc rs m bc) (
RTL.State st' f' sp pc' rs m)
|
match_states_call
st st' f f' args m bc
(
STACKS:
list_forall2 match_stackframes st st')
(
TRANSF:
match_fundef f f')
:
match_states (
Callstate st f args m bc) (
RTL.Callstate st' f' args m)
|
match_states_return
st st' v m bc
(
STACKS:
list_forall2 match_stackframes st st')
:
match_states (
Returnstate st v m bc) (
RTL.Returnstate st' v m)
.
Lemma transf_function_correct f f':
transf_function f =
OK f' ->
exists dupmap,
match_function dupmap f f' false.
Proof.
Lemma transf_fundef_correct f f':
transf_fundef f =
OK f' ->
match_fundef f f'.
Proof.
Definition match_prog (
p:
program) (
tp:
RTL.program) :=
match_program (
fun _
f tf =>
transf_fundef f =
OK tf)
eq p tp.
Lemma transf_program_match:
forall prog tprog,
transf_program prog =
OK tprog ->
match_prog prog tprog.
Proof.
Section RTL_SIMULATES_BTL.
Variable prog:
program.
Variable tprog:
RTL.program.
Hypothesis TRANSL:
match_prog prog tprog.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Lemma symbols_preserved s:
Genv.find_symbol tge s =
Genv.find_symbol ge s.
Proof.
Lemma senv_preserved:
Senv.equiv ge tge.
Proof.
Lemma functions_translated (
v:
val) (
f:
fundef):
Genv.find_funct ge v =
Some f ->
exists tf cunit,
transf_fundef f =
OK tf /\
Genv.find_funct tge v =
Some tf /\
linkorder cunit prog.
Proof.
Lemma function_ptr_translated 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.
Lemma function_sig_translated f tf:
transf_fundef f =
OK tf ->
RTL.funsig tf =
funsig f.
Proof.
Lemma transf_initial_states s1:
initial_state prog s1 ->
exists s2,
RTL.initial_state tprog s2 /\
match_states s1 s2.
Proof.
Lemma transf_final_states s1 s2 r:
match_states s1 s2 ->
final_state s1 r ->
RTL.final_state s2 r.
Proof.
intros. inv H0. inv H. inv STACKS. constructor.
Qed.
Lemma find_function_preserved ri rs0 fd
(
FIND :
find_function ge ri rs0 =
Some fd)
:
exists fd',
RTL.find_function tge ri rs0 =
Some fd'
/\
transf_fundef fd =
OK fd'.
Proof.
Lemma list_nth_z_dupmap:
forall dupmap btlcode ln ln' (
pc pc':
node)
val,
list_nth_z ln val =
Some pc ->
list_forall2 (
fun n n' =>
match_nodes dupmap btlcode false n n')
ln ln' ->
exists (
pc':
node),
list_nth_z ln' val =
Some pc'
/\
dupmap!
pc =
Some pc'.
Proof.
induction ln;
intros until val;
intros LNZ LFA.
-
inv LNZ.
-
inv LNZ.
destruct (
zeq val 0)
eqn:
ZEQ.
+
inv H0.
destruct ln';
inv LFA.
simpl.
exists n.
destruct H2 as (
osn_pc &
MSN).
inv MSN;
split;
auto.
+
inv LFA.
simpl.
rewrite ZEQ.
exploit IHln. 2:
eapply H0.
all:
eauto.
Qed.
Inductive cond_star_step (
P:
Prop):
RTL.state ->
trace ->
RTL.state ->
Prop :=
|
css_refl s:
P ->
cond_star_step P s E0 s
|
css_plus s1 t s2:
plus RTL.step tge s1 t s2 ->
cond_star_step P s1 t s2
.
Lemma css_plus_trans P Q s0 s1 s2 t:
plus RTL.step tge s0 E0 s1 ->
cond_star_step P s1 t s2 ->
cond_star_step Q s0 t s2.
Proof.
Lemma css_E0_trans isfst isfst' s0 s1 s2:
cond_star_step (
isfst=
false)
s0 E0 s1 ->
cond_star_step (
false=
isfst')
s1 E0 s2 ->
cond_star_step (
isfst=
isfst')
s0 E0 s2.
Proof.
intros S1 S2.
inversion S1;
subst;
eauto.
inversion S2;
subst;
eauto.
eapply css_plus_trans;
eauto.
Qed.
Lemma css_star P s0 s1 t:
cond_star_step P s0 t s1 ->
star RTL.step tge s0 t s1.
Proof.
Local Hint Constructors RTL.step match_states:
core.
Local Hint Resolve css_refl plus_one transf_fundef_correct:
core.
Lemma iblock_istep_simulation sp bc dupmap stack' f' rs0 m0 ib rs1 m1 ofin btlcode
(
IBIS:
iblock_istep AnnotNone ge sp bc rs0 m0 ib rs1 m1 ofin):
forall pc0 opc isfst
(
MIB:
match_iblock dupmap btlcode (
RTL.fn_code f')
false isfst pc0 ib opc),
match ofin with
|
None =>
exists pc1,(
opc =
Some pc1) /\
cond_star_step (
isfst =
false) (
RTL.State stack' f' sp pc0 rs0 m0)
E0 (
RTL.State stack' f' sp pc1 rs1 m1)
|
Some fin =>
exists isfst' pc1 iinfo,
match_iblock dupmap btlcode (
RTL.fn_code f')
false isfst' pc1 (
BF fin iinfo)
None
/\
cond_star_step (
isfst =
isfst') (
RTL.State stack' f' sp pc0 rs0 m0)
E0 (
RTL.State stack' f' sp pc1 rs1 m1)
end.
Proof.
induction IBIS;
simpl;
intros.
-
assert (
X:
opc =
None). {
inv MIB;
auto. }
subst.
repeat eexists;
eauto.
-
inv MIB;
eexists;
split;
econstructor;
eauto.
-
inv MIB.
exists pc';
split;
auto;
constructor.
apply plus_one.
eapply exec_Iop;
eauto.
erewrite <-
eval_operation_preserved;
eauto.
intros;
rewrite symbols_preserved;
trivial.
-
inv MIB.
exists pc';
split;
auto;
constructor.
apply plus_one.
eapply exec_Iload;
eauto.
apply may_undef_prop_lax in UDP as ->.
erewrite has_loaded_preserved by (
apply symbols_preserved).
exact LOAD.
-
inv MIB.
exists pc';
split;
auto;
constructor.
apply plus_one.
eapply exec_Istore;
eauto.
all:
erewrite <-
eval_addressing_preserved;
eauto;
intros;
rewrite symbols_preserved;
trivial.
-
inv MIB;
eauto.
-
inv MIB.
exploit IHIBIS1;
eauto.
intros (
pc1 &
EQpc1 &
STEP1);
inv EQpc1.
exploit IHIBIS2;
eauto.
destruct ofin;
simpl.
+
intros (
ifst2 &
pc2 &
iinfo &
M2 &
STEP2).
repeat eexists;
eauto.
eapply css_E0_trans;
eauto.
+
intros (
pc2 &
EQpc2 &
STEP2);
inv EQpc2.
eexists;
split;
auto.
eapply css_E0_trans;
eauto.
-
inv MIB.
rename H10 into JOIN.
destruct b;
exploit IHIBIS;
eauto.
+
destruct ofin;
simpl.
*
intros (
isfst' &
pc1 &
iinfo' &
MI &
STAR).
repeat eexists;
eauto.
eapply css_plus_trans;
eauto.
*
intros (
pc1 &
OPC &
PLUS).
inv OPC.
inv JOIN;
eexists;
split;
eauto.
all:
eapply css_plus_trans;
eauto.
+
destruct ofin;
simpl.
*
intros (
isfst' &
pc1 &
iinfo' &
MI &
STAR).
repeat eexists;
eauto.
eapply css_plus_trans;
eauto.
*
intros (
pc1 &
OPC &
PLUS).
subst.
inv JOIN;
eexists;
split;
eauto.
all:
eapply css_plus_trans;
eauto.
Qed.
Lemma final_simu_except_goto sp dupmap stack stack' f f' rs1 m1 bc pc1 fin t s btlcode
(
STACKS :
list_forall2 match_stackframes stack stack')
(
TRANSF :
match_function dupmap f f' false)
(
FS :
final_step ge stack f sp rs1 m1 bc fin t s)
(
i :
instruction)
(
ATpc1 : (
RTL.fn_code f') !
pc1 =
Some i)
(
MF :
match_final_inst dupmap btlcode false fin i)
:
exists s',
RTL.step tge (
RTL.State stack' f' sp pc1 rs1 m1)
t s' /\
match_states s s'.
Proof.
Lemma iblock_step_simulation sp dupmap stack stack' f f' ib rs0 m0 rs1 m1 bc pc0 fin t s btlcode
(
STACKS:
list_forall2 match_stackframes stack stack')
(
TRANSF:
match_function dupmap f f' false)
(
IBIS:
iblock_istep AnnotNone ge sp bc rs0 m0 ib rs1 m1 (
Some fin))
(
MIB :
match_iblock dupmap btlcode (
RTL.fn_code f')
false true pc0 ib None)
(
FS :
final_step ge stack f sp rs1 m1 bc fin t s)
:
exists s',
plus RTL.step tge (
RTL.State stack' f' sp pc0 rs0 m0)
t s' /\
match_states s s'.
Proof.
intros;
exploit iblock_istep_simulation;
eauto.
simpl.
intros (
isfst' &
pc1 &
iinfo &
MI &
STAR).
clear IBIS MIB.
inv MI.
-
exploit final_simu_except_goto;
eauto.
intros (
s' &
STEP &
MS).
eexists;
split.
+
eapply plus_right.
eapply css_star;
eauto.
eapply STEP.
econstructor.
+
eapply MS.
-
inv FS.
inv STAR;
try congruence.
eexists;
split.
eauto.
destruct H2 as (
osn_pc &
MSN);
inv MSN.
econstructor;
eauto.
Qed.
Theorem plus_simulation s1 t s1':
step AnnotNone ge s1 t s1' ->
forall s2 (
MS:
match_states s1 s2),
exists s2',
plus RTL.step tge s2 t s2'
/\
match_states s1' s2'.
Proof.
Theorem transf_program_correct:
forward_simulation (
BTL.sem AnnotNone prog) (
RTL.semantics tprog).
Proof.
End RTL_SIMULATES_BTL.
Program Definition ipass :
IPass (
BTL.sem AnnotNone)
RTL.semantics :=
{|
ipass_spec :=
mkpass match_prog;
ipass_impl :=
transf_program;
|}.
Next Obligation.
Next Obligation.
End BTLtoRTL_Translationproof.