Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
Require Import AST Linking.
Require Import Values Memory Globalenvs Events Smallstep.
Require Import Registers Op RTL.
Require Import Profiling.
Require Import Lia.
Local Open Scope positive.
Definition match_prog (
p tp:
RTL.program) :=
match_program (
fun ctx f tf =>
tf =
transf_fundef f)
eq p tp.
Lemma transf_program_match:
forall p,
match_prog p (
transf_program p).
Proof.
Section PRESERVATION.
Variables prog tprog:
program.
Hypothesis TRANSL:
match_prog prog tprog.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Lemma functions_translated:
forall v f,
Genv.find_funct ge v =
Some f ->
Genv.find_funct tge v =
Some (
transf_fundef f).
Proof (
Genv.find_funct_transf TRANSL).
Lemma function_ptr_translated:
forall v f,
Genv.find_funct_ptr ge v =
Some f ->
Genv.find_funct_ptr tge v =
Some (
transf_fundef f).
Proof (
Genv.find_funct_ptr_transf TRANSL).
Lemma symbols_preserved:
forall id,
Genv.find_symbol tge id =
Genv.find_symbol ge id.
Proof (
Genv.find_symbol_transf TRANSL).
Lemma senv_preserved:
Senv.equiv ge tge.
Proof (
Genv.senv_transf TRANSL).
Lemma sig_preserved:
forall f,
funsig (
transf_fundef f) =
funsig f.
Proof.
destruct f; reflexivity.
Qed.
Lemma find_function_translated:
forall ros rs fd,
find_function ge ros rs =
Some fd ->
find_function tge ros rs =
Some (
transf_fundef fd).
Proof.
Lemma pair_expand:
forall {
A B :
Type } (
p :
A*
B),
p = ((
fst p), (
snd p)).
Proof.
destruct p; simpl; trivial.
Qed.
Lemma inject_profiling_call_preserves:
forall id body pc extra_pc ifso ifnot pc0,
pc0 <
extra_pc ->
PTree.get pc0 (
snd (
inject_profiling_call id body pc extra_pc ifso ifnot)) =
PTree.get pc0 body.
Proof.
Lemma inject_at_preserves :
forall id body pc extra_pc pc0,
pc0 <
extra_pc ->
pc0 <>
pc ->
PTree.get pc0 (
snd (
inject_at id body pc extra_pc)) =
PTree.get pc0 body.
Proof.
Lemma inject_profiling_call_increases:
forall id body pc extra_pc ifso ifnot,
fst (
inject_profiling_call id body pc extra_pc ifso ifnot) =
extra_pc + 2.
Proof.
Lemma inject_at_increases:
forall id body pc extra_pc,
(
fst (
inject_at id body pc extra_pc)) =
extra_pc + 2.
Proof.
Lemma inject_l_preserves :
forall id injections body extra_pc pc0,
pc0 <
extra_pc ->
List.forallb (
fun injection =>
if peq injection pc0 then false else true)
injections =
true ->
PTree.get pc0 (
snd (
inject_l id body extra_pc injections)) =
PTree.get pc0 body.
Proof.
Fixpoint inject_l_position extra_pc
(
injections :
list node)
(
k :
nat) {
struct injections} :
node :=
match injections with
|
nil =>
extra_pc
|
pc::
l' =>
match k with
|
O =>
extra_pc
|
S k' =>
inject_l_position (
extra_pc + 2)
l' k'
end
end.
Lemma inject_l_position_increases :
forall injections pc k,
pc <=
inject_l_position pc injections k.
Proof.
induction injections; simpl; intros.
lia.
destruct k.
lia.
specialize IHinjections with (pc := pc + 2) (k := k).
lia.
Qed.
Lemma inject_l_injected_pc:
forall f_id injections cond args ifso ifnot expected body injnum pc extra_pc
(
INSTR :
body !
pc =
Some (
Icond cond args ifso ifnot expected))
(
BELOW :
forallb (
fun pc =>
pc <?
extra_pc)
injections =
true)
(
NOREPET :
list_norepet injections)
(
NUMBER :
nth_error injections injnum =
Some pc),
PTree.get pc (
snd (
inject_l f_id body extra_pc injections)) =
Some (
Icond cond args
(
Pos.succ (
inject_l_position extra_pc injections injnum))
(
inject_l_position extra_pc injections injnum)
expected).
Proof.
Lemma inject_l_injected0:
forall f_id cond args ifso ifnot expected injections body injnum pc extra_pc
(
INSTR :
body !
pc =
Some (
Icond cond args ifso ifnot expected))
(
BELOW :
forallb (
fun pc =>
pc <?
extra_pc)
injections =
true)
(
NOREPET :
list_norepet injections)
(
NUMBER :
nth_error injections injnum =
Some pc),
PTree.get (
inject_l_position extra_pc injections injnum)
(
snd (
inject_l f_id body extra_pc injections)) =
Some (
Ibuiltin (
EF_profiling (
branch_id f_id pc) 0%
Z)
nil BR_none ifnot).
Proof.
Lemma inject_l_injected1:
forall f_id cond args ifso ifnot expected injections body injnum pc extra_pc
(
INSTR :
body !
pc =
Some (
Icond cond args ifso ifnot expected))
(
BELOW :
forallb (
fun pc =>
pc <?
extra_pc)
injections =
true)
(
NOREPET :
list_norepet injections)
(
NUMBER :
nth_error injections injnum =
Some pc),
PTree.get (
Pos.succ (
inject_l_position extra_pc injections injnum))
(
snd (
inject_l f_id body extra_pc injections)) =
Some (
Ibuiltin (
EF_profiling (
branch_id f_id pc) 1%
Z)
nil BR_none ifso).
Proof.
Lemma transf_function_at:
forall f pc i
(
CODE :
f.(
fn_code)!
pc =
Some i)
(
INSTR :
match i with
|
Icond _ _ _ _ _ =>
False
| _ =>
True
end),
(
transf_function f).(
fn_code)!
pc =
Some i.
Proof.
Inductive match_frames:
RTL.stackframe ->
RTL.stackframe ->
Prop :=
|
match_frames_intro:
forall res f sp pc rs,
match_frames (
Stackframe res f sp pc rs)
(
Stackframe res (
transf_function f)
sp pc rs).
Inductive match_states:
RTL.state ->
RTL.state ->
Prop :=
|
match_regular_states:
forall stk f sp pc rs m stk'
(
STACKS:
list_forall2 match_frames stk stk'),
match_states (
State stk f sp pc rs m)
(
State stk' (
transf_function f)
sp pc rs m)
|
match_callstates:
forall stk f args m stk'
(
STACKS:
list_forall2 match_frames stk stk'),
match_states (
Callstate stk f args m)
(
Callstate stk' (
transf_fundef f)
args m)
|
match_returnstates:
forall stk v m stk'
(
STACKS:
list_forall2 match_frames stk stk'),
match_states (
Returnstate stk v m)
(
Returnstate stk' v m).
Lemma funsig_preserved:
forall fd,
funsig (
transf_fundef fd) =
funsig fd.
Proof.
destruct fd; simpl; trivial.
Qed.
Lemma stacksize_preserved:
forall f,
fn_stacksize (
transf_function f) =
fn_stacksize f.
Proof.
destruct f; simpl; trivial.
Qed.
Hint Resolve symbols_preserved funsig_preserved external_call_symbols_preserved senv_preserved stacksize_preserved :
profiling.
Lemma step_simulation:
forall s1 t s2 (
STEP :
step ge s1 t s2)
s1' (
MS:
match_states s1 s1'),
exists s2',
plus step tge s1' t s2' /\
match_states s2 s2'.
Proof.
Lemma transf_initial_states:
forall S1,
RTL.initial_state prog S1 ->
exists S2,
RTL.initial_state tprog S2 /\
match_states S1 S2.
Proof.
Lemma transf_final_states:
forall S1 S2 r,
match_states S1 S2 ->
RTL.final_state S1 r ->
RTL.final_state S2 r.
Proof.
intros. inv H0. inv H. inv STACKS. constructor.
Qed.
Theorem transf_program_correct:
forward_simulation (
RTL.semantics prog) (
RTL.semantics tprog).
Proof.
End PRESERVATION.