Correctness proof for code duplication
Require Import AST Linking Errors Globalenvs Smallstep.
Require Import Coqlib Maps Events Values.
Require Import Op RTL Duplicate.
Module DuplicateProof (
D:
DuplicateOracle).
Include Duplicate D.
Local Open Scope positive_scope.
Definition of match_states (independently of the translation)
Inductive match_inst (
dupmap:
PTree.t node):
instruction ->
instruction ->
Prop :=
|
match_inst_nop:
forall n n',
dupmap!
n' = (
Some n) ->
match_inst dupmap (
Inop n) (
Inop n')
|
match_inst_op:
forall n n' op lr r,
dupmap!
n' = (
Some n) ->
match_inst dupmap (
Iop op lr r n) (
Iop op lr r n')
|
match_inst_load:
forall n n' tm m a lr r,
dupmap!
n' = (
Some n) ->
match_inst dupmap (
Iload tm m a lr r n) (
Iload tm m a lr r n')
|
match_inst_store:
forall n n' m a lr r,
dupmap!
n' = (
Some n) ->
match_inst dupmap (
Istore m a lr r n) (
Istore m a lr r n')
|
match_inst_call:
forall n n' s ri lr r,
dupmap!
n' = (
Some n) ->
match_inst dupmap (
Icall s ri lr r n) (
Icall s ri lr r n')
|
match_inst_tailcall:
forall s ri lr,
match_inst dupmap (
Itailcall s ri lr) (
Itailcall s ri lr)
|
match_inst_builtin:
forall n n' ef la br,
dupmap!
n' = (
Some n) ->
match_inst dupmap (
Ibuiltin ef la br n) (
Ibuiltin ef la br n')
|
match_inst_cond:
forall ifso ifso' ifnot ifnot' c lr i i',
dupmap!
ifso' = (
Some ifso) ->
dupmap!
ifnot' = (
Some ifnot) ->
match_inst dupmap (
Icond c lr ifso ifnot i) (
Icond c lr ifso' ifnot' i')
|
match_inst_revcond:
forall ifso ifso' ifnot ifnot' c lr i i',
dupmap!
ifso' = (
Some ifso) ->
dupmap!
ifnot' = (
Some ifnot) ->
match_inst dupmap (
Icond c lr ifso ifnot i) (
Icond (
negate_condition c)
lr ifnot' ifso' i')
|
match_inst_jumptable:
forall ln ln' r,
list_forall2 (
fun n n' => (
dupmap!
n' = (
Some n)))
ln ln' ->
match_inst dupmap (
Ijumptable r ln) (
Ijumptable r ln')
|
match_inst_return:
forall or,
match_inst dupmap (
Ireturn or) (
Ireturn or)
|
match_inst_assert:
forall ifso ifso' c lr,
dupmap!
ifso' = (
Some ifso) ->
match_inst dupmap (
Iassert c lr ifso) (
Iassert c lr ifso').
Record match_function dupmap f f':
Prop := {
dupmap_correct:
forall n n',
dupmap!
n' =
Some n ->
(
forall i, (
fn_code f)!
n =
Some i ->
exists i', (
fn_code f')!
n' =
Some i' /\
match_inst dupmap i i');
dupmap_entrypoint:
dupmap!(
fn_entrypoint f') =
Some (
fn_entrypoint f);
preserv_fnsig:
fn_sig f =
fn_sig f';
preserv_fnparams:
fn_params f =
fn_params f';
preserv_fnstacksize:
fn_stacksize f =
fn_stacksize f'
}.
Inductive match_fundef:
RTL.fundef ->
RTL.fundef ->
Prop :=
|
match_Internal dupmap f f':
match_function dupmap f f' ->
match_fundef (
Internal f) (
Internal f')
|
match_External ef:
match_fundef (
External ef) (
External ef).
Inductive match_stackframes:
stackframe ->
stackframe ->
Prop :=
|
match_stackframe_intro
dupmap res f sp pc rs f' pc'
(
TRANSF:
match_function dupmap f f')
(
DUPLIC:
dupmap!
pc' =
Some pc):
match_stackframes (
Stackframe res f sp pc rs) (
Stackframe res f' sp pc' rs).
Inductive match_states:
state ->
state ->
Prop :=
|
match_states_intro
dupmap st f sp pc rs m st' f' pc'
(
STACKS:
list_forall2 match_stackframes st st')
(
TRANSF:
match_function dupmap f f')
(
DUPLIC:
dupmap!
pc' =
Some pc):
match_states (
State st f sp pc rs m) (
State st' f' sp pc' rs m)
|
match_states_call:
forall st st' f f' args m
(
STACKS:
list_forall2 match_stackframes st st')
(
TRANSF:
match_fundef f f'),
match_states (
Callstate st f args m) (
Callstate st' f' args m)
|
match_states_return:
forall st st' v m
(
STACKS:
list_forall2 match_stackframes st st'),
match_states (
Returnstate st v m) (
Returnstate st' v m).
Auxiliary properties
Theorem transf_function_preserves:
forall f f',
transf_function f =
OK f' ->
fn_sig f =
fn_sig f' /\
fn_params f =
fn_params f' /\
fn_stacksize f =
fn_stacksize f'.
Proof.
intros.
unfold transf_function in H.
destruct (
duplicate_aux _)
as (
tcte &
mp).
destruct tcte as (
tc &
te).
monadInv H.
repeat (
split;
try reflexivity).
Qed.
Lemma verify_mapping_mn_rec_step:
forall dupmap lb b f f',
In b lb ->
verify_mapping_mn_rec dupmap f f' lb =
OK tt ->
verify_mapping_mn dupmap f f' b =
OK tt.
Proof.
induction lb; intros.
- monadInv H0. inversion H.
- inversion H.
+ subst. monadInv H0. destruct x. assumption.
+ monadInv H0. destruct x. eapply IHlb; assumption.
Qed.
Lemma verify_is_copy_correct:
forall dupmap n n',
verify_is_copy dupmap n n' =
OK tt ->
dupmap !
n' =
Some n.
Proof.
intros.
unfold verify_is_copy in H.
destruct (_ !
n')
eqn:
REVM; [|
inversion H].
destruct (
n ?=
p)
eqn:
NP;
try (
inversion H;
fail).
eapply Pos.compare_eq in NP.
subst.
reflexivity.
Qed.
Lemma verify_is_copy_list_correct:
forall dupmap ln ln',
verify_is_copy_list dupmap ln ln' =
OK tt ->
list_forall2 (
fun n n' =>
dupmap !
n' =
Some n)
ln ln'.
Proof.
induction ln.
-
intros.
destruct ln';
monadInv H.
constructor.
-
intros.
destruct ln';
monadInv H.
destruct x.
apply verify_is_copy_correct in EQ.
eapply IHln in EQ0.
constructor;
assumption.
Qed.
Lemma verify_match_inst_correct:
forall dupmap i i',
verify_match_inst dupmap i i' =
OK tt ->
match_inst dupmap i i'.
Proof.
Lemma verify_mapping_mn_correct mp n n' i f f' tc:
mp !
n' =
Some n ->
(
fn_code f) !
n =
Some i ->
(
fn_code f') =
tc ->
verify_mapping_mn mp f f' (
n',
n) =
OK tt ->
exists i',
tc !
n' =
Some i'
/\
match_inst mp i i'.
Proof.
unfold verify_mapping_mn;
intros H H0 H1 H2.
rewrite H0 in H2.
clear H0.
rewrite H1 in H2.
clear H1.
destruct (
tc !
n')
eqn:
TCN; [|
inversion H2].
exists i0.
split;
auto.
eapply verify_match_inst_correct.
assumption.
Qed.
Lemma verify_mapping_mn_rec_correct:
forall mp n n' i f f' tc,
mp !
n' =
Some n ->
(
fn_code f) !
n =
Some i ->
(
fn_code f') =
tc ->
verify_mapping_mn_rec mp f f' (
PTree.elements mp) =
OK tt ->
exists i',
tc !
n' =
Some i'
/\
match_inst mp i i'.
Proof.
Theorem transf_function_correct f f':
transf_function f =
OK f' ->
exists dupmap,
match_function dupmap f f'.
Proof.
Lemma transf_fundef_correct f f':
transf_fundef f =
OK f' ->
match_fundef f f'.
Proof.
Preservation proof
Definition match_prog (
p tp:
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 PRESERVATION.
Variable prog:
program.
Variable tprog:
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:
forall (
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:
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.
Lemma function_sig_translated:
forall f tf,
transf_fundef f =
OK tf ->
funsig tf =
funsig f.
Proof.
intros.
destruct f.
-
simpl in H.
monadInv H.
simpl.
symmetry.
apply transf_function_preserves.
assumption.
-
simpl in H.
monadInv H.
reflexivity.
Qed.
Lemma list_nth_z_dupmap:
forall dupmap ln ln' (
pc pc':
node)
val,
list_nth_z ln val =
Some pc ->
list_forall2 (
fun n n' =>
dupmap!
n' =
Some n)
ln ln' ->
exists pc',
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 p.
split;
auto.
+
inv LFA.
simpl.
rewrite ZEQ.
exploit IHln. 2:
eapply H0.
all:
eauto.
intros (
pc'1 &
LNZ &
REV).
exists pc'1.
split;
auto.
congruence.
Qed.
Theorem transf_initial_states:
forall s1,
initial_state prog s1 ->
exists s2,
initial_state tprog s2 /\
match_states s1 s2.
Proof.
Theorem transf_final_states:
forall s1 s2 r,
match_states s1 s2 ->
final_state s1 r ->
final_state s2 r.
Proof.
intros. inv H0. inv H. inv STACKS. constructor.
Qed.
Theorem step_simulation:
forall s1 t s1',
step ge s1 t s1' ->
forall s2 (
MS:
match_states s1 s2),
exists s2',
step tge s2 t s2'
/\
match_states s1' s2'.
Proof.
Theorem transf_program_correct:
forward_simulation (
semantics prog) (
semantics tprog).
Proof.
End PRESERVATION.
End DuplicateProof.