Require Import Axioms.
Require Import FunInd.
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 KillUselessMoves.
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 transf_function_at:
forall f pc i,
f.(
fn_code)!
pc =
Some i ->
(
transf_function f).(
fn_code)!
pc =
Some(
transf_instr pc i).
Proof.
Ltac TR_AT :=
match goal with
| [
A: (
fn_code _)!_ =
Some _ |- _ ] =>
generalize (
transf_function_at _ _ _
A);
intros
end.
Section SAME_RS.
Context {
A :
Type}.
Definition same_rs (
rs rs' :
Regmap.t A) :=
forall x,
rs #
x =
rs' #
x.
Lemma same_rs_refl :
forall rs,
same_rs rs rs.
Proof.
Lemma same_rs_comm :
forall rs rs', (
same_rs rs rs') -> (
same_rs rs' rs).
Proof.
Lemma same_rs_trans :
forall rs1 rs2 rs3,
(
same_rs rs1 rs2) -> (
same_rs rs2 rs3) -> (
same_rs rs1 rs3).
Proof.
Lemma same_rs_idem_write :
forall rs r,
(
same_rs rs (
rs #
r <- (
rs #
r))).
Proof.
Lemma same_rs_read:
forall rs rs' r, (
same_rs rs rs') ->
rs #
r =
rs' #
r.
Proof.
Lemma same_rs_subst:
forall rs rs' l, (
same_rs rs rs') ->
rs ##
l =
rs' ##
l.
Proof.
induction l; cbn; intuition congruence.
Qed.
Lemma same_rs_write:
forall rs rs' r x,
(
same_rs rs rs') -> (
same_rs (
rs #
r <-
x) (
rs' #
r <-
x)).
Proof.
Lemma same_rs_setres:
forall rs rs' (
SAME:
same_rs rs rs')
res vres,
same_rs (
regmap_setres res vres rs) (
regmap_setres res vres rs').
Proof.
End SAME_RS.
Lemma same_find_function:
forall tge rs rs' (
SAME:
same_rs rs rs')
ros,
find_function tge ros rs =
find_function tge ros rs'.
Proof.
destruct ros;
cbn.
{
rewrite (
same_rs_read rs rs' r SAME).
reflexivity. }
reflexivity.
Qed.
Inductive match_frames:
RTL.stackframe ->
RTL.stackframe ->
Prop :=
|
match_frames_intro:
forall res f sp pc rs rs' (
SAME :
same_rs rs 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 rs' m stk'
(
SAME:
same_rs rs rs')
(
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 step_simulation:
forall S1 t S2,
RTL.step ge S1 t S2 ->
forall S1',
match_states S1 S1' ->
exists S2',
RTL.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.