Require Import Coqlib Errors.
Require Import Integers Floats AST Linking.
Require Import Values Memory Events Globalenvs Smallstep.
Require Import Op Locations Machblock Conventions Asmblock.
Require Import Asmblockgenproof0 Asmblockprops.
Require Import PostpassScheduling.
Require Import Asmblockgenproof.
Require Import Axioms.
Require Import Lia.
Local Open Scope error_monad_scope.
Definition match_prog (
p tp:
Asmvliw.program) :=
match_program (
fun _
f tf =>
transf_fundef f =
OK tf)
eq p tp.
Lemma transf_program_match:
forall p tp,
transf_program p =
OK tp ->
match_prog p tp.
Proof.
Lemma regset_double_set_id:
forall r (
rs:
regset)
v1 v2,
(
rs #
r <-
v1 #
r <-
v2) = (
rs #
r <-
v2).
Proof.
Lemma exec_body_pc_var:
forall l ge rs m rs' m' v,
exec_body ge l rs m =
Next rs' m' ->
exec_body ge l (
rs #
PC <-
v)
m =
Next (
rs' #
PC <-
v)
m'.
Proof.
Lemma pc_set_add:
forall rs v r x y,
0 <=
x <=
Ptrofs.max_unsigned ->
0 <=
y <=
Ptrofs.max_unsigned ->
rs #
r <- (
Val.offset_ptr v (
Ptrofs.repr (
x +
y))) =
rs #
r <- (
Val.offset_ptr (
rs #
r <- (
Val.offset_ptr v (
Ptrofs.repr x))
r) (
Ptrofs.repr y)).
Proof.
Lemma concat2_straight:
forall a b bb rs m rs'' m'' f ge,
concat2 a b =
OK bb ->
exec_bblock ge f bb rs m =
Next rs'' m'' ->
exists rs' m',
exec_bblock ge f a rs m =
Next rs' m'
/\
rs' PC =
Val.offset_ptr (
rs PC) (
Ptrofs.repr (
size a))
/\
exec_bblock ge f b rs' m' =
Next rs'' m''.
Proof.
Lemma concat_all_exec_bblock (
ge:
Genv.t fundef unit) (
f:
function) :
forall a bb rs m lbb rs'' m'',
lbb <>
nil ->
concat_all (
a ::
lbb) =
OK bb ->
exec_bblock ge f bb rs m =
Next rs'' m'' ->
exists bb' rs' m',
concat_all lbb =
OK bb'
/\
exec_bblock ge f a rs m =
Next rs' m'
/\
rs' PC =
Val.offset_ptr (
rs PC) (
Ptrofs.repr (
size a))
/\
exec_bblock ge f bb' rs' m' =
Next rs'' m''.
Proof.
intros until m''.
intros Hnonil CONC EXEB.
simpl in CONC.
destruct lbb as [|
b lbb];
try contradiction.
clear Hnonil.
monadInv CONC.
exploit concat2_straight;
eauto.
intros (
rs' &
m' &
EXEB1 &
PCeq &
EXEB2).
exists x.
repeat econstructor.
all:
eauto.
Qed.
Lemma ptrofs_add_repr :
forall a b,
Ptrofs.unsigned (
Ptrofs.add (
Ptrofs.repr a) (
Ptrofs.repr b)) =
Ptrofs.unsigned (
Ptrofs.repr (
a +
b)).
Proof.
Section PRESERVATION_ASMBLOCK.
Variables prog tprog:
program.
Hypothesis TRANSL:
match_prog prog tprog.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Lemma transf_function_no_overflow:
forall f tf,
transf_function f =
OK tf ->
size_blocks tf.(
fn_blocks) <=
Ptrofs.max_unsigned.
Proof.
Lemma symbols_preserved:
forall id,
Genv.find_symbol tge id =
Genv.find_symbol ge id.
Proof (
Genv.find_symbol_match TRANSL).
Lemma senv_preserved:
Senv.equiv ge tge.
Proof (
Genv.senv_match TRANSL).
Lemma functions_translated:
forall v f,
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 TRANSL).
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 (
Genv.find_funct_ptr_transf_partial TRANSL).
Lemma functions_transl:
forall fb f tf,
Genv.find_funct_ptr ge fb =
Some (
Internal f) ->
transf_function f =
OK tf ->
Genv.find_funct_ptr tge fb =
Some (
Internal tf).
Proof.
intros.
exploit function_ptr_translated;
eauto.
intros (
tf' &
A &
B).
monadInv B.
rewrite H0 in EQ.
inv EQ.
auto.
Qed.
Inductive match_states:
state ->
state ->
Prop :=
|
match_states_intro:
forall s1 s2,
s1 =
s2 ->
match_states s1 s2.
Lemma prog_main_preserved:
prog_main tprog =
prog_main prog.
Proof (
match_program_main TRANSL).
Lemma prog_main_address_preserved:
(
Genv.symbol_address (
Genv.globalenv prog) (
prog_main prog)
Ptrofs.zero) =
(
Genv.symbol_address (
Genv.globalenv tprog) (
prog_main tprog)
Ptrofs.zero).
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. econstructor; eauto.
Qed.
Lemma tail_find_bblock:
forall lbb pos bb,
find_bblock pos lbb =
Some bb ->
exists c,
code_tail pos lbb (
bb::
c).
Proof.
induction lbb.
-
intros.
simpl in H.
inv H.
-
intros.
simpl in H.
destruct (
zlt pos 0);
try (
inv H;
fail).
destruct (
zeq pos 0).
+
inv H.
exists lbb.
constructor;
auto.
+
apply IHlbb in H.
destruct H as (
c &
TAIL).
exists c.
enough (
pos =
pos -
size a +
size a)
as ->.
apply code_tail_S;
auto.
lia.
Qed.
Lemma code_tail_head_app:
forall l pos c1 c2,
code_tail pos c1 c2 ->
code_tail (
pos +
size_blocks l) (
l++
c1)
c2.
Proof.
Lemma transf_blocks_verified:
forall c tc pos bb c',
transf_blocks c =
OK tc ->
code_tail pos c (
bb::
c') ->
exists lbb,
verified_schedule bb =
OK lbb
/\
exists tc',
code_tail pos tc (
lbb ++
tc').
Proof.
induction c;
intros.
-
simpl in H.
inv H.
inv H0.
-
inv H0.
+
monadInv H.
exists x0.
split;
simpl;
auto.
eexists;
eauto.
econstructor;
eauto.
+
unfold transf_blocks in H.
fold transf_blocks in H.
monadInv H.
exploit IHc;
eauto.
intros (
lbb &
TRANS &
tc' &
TAIL).
repeat eexists;
eauto.
erewrite verified_schedule_size;
eauto.
apply code_tail_head_app.
eauto.
Qed.
Lemma transf_find_bblock:
forall ofs f bb tf,
find_bblock (
Ptrofs.unsigned ofs) (
fn_blocks f) =
Some bb ->
transf_function f =
OK tf ->
exists lbb,
verified_schedule bb =
OK lbb
/\
exists c,
code_tail (
Ptrofs.unsigned ofs) (
fn_blocks tf) (
lbb ++
c).
Proof.
Lemma symbol_address_preserved:
forall l ofs,
Genv.symbol_address ge l ofs =
Genv.symbol_address tge l ofs.
Proof.
Lemma head_tail {
A:
Type}:
forall (
l:
list A)
hd,
hd::
l =
hd :: (
tail (
hd::
l)).
Proof.
intros. simpl. auto.
Qed.
Lemma verified_schedule_not_empty:
forall bb lbb,
verified_schedule bb =
OK lbb ->
lbb <>
nil.
Proof.
Lemma header_nil_label_pos_none:
forall lbb l p,
Forall (
fun b =>
header b =
nil)
lbb ->
label_pos l p lbb =
None.
Proof.
induction lbb.
-
intros.
simpl.
auto.
-
intros.
inv H.
simpl.
unfold is_label.
rewrite H2.
destruct (
in_dec l nil). {
inv i. }
auto.
Qed.
Lemma verified_schedule_label:
forall bb tbb lbb l,
verified_schedule bb =
OK (
tbb ::
lbb) ->
is_label l bb =
is_label l tbb
/\
label_pos l 0
lbb =
None.
Proof.
Lemma label_pos_app_none:
forall c c' l p p',
label_pos l p c =
None ->
label_pos l (
p' +
size_blocks c)
c' =
label_pos l p' (
c ++
c').
Proof.
induction c.
-
intros.
simpl in *.
rewrite Z.add_0_r.
reflexivity.
-
intros.
simpl in *.
destruct (
is_label _ _)
eqn:
ISLABEL.
+
discriminate.
+
eapply IHc in H.
rewrite Z.add_assoc.
eauto.
Qed.
Remark label_pos_pvar_none_add:
forall tc l p p' k,
label_pos l (
p+
k)
tc =
None ->
label_pos l (
p'+
k)
tc =
None.
Proof.
induction tc.
-
intros.
simpl.
auto.
-
intros.
simpl in *.
destruct (
is_label _ _)
eqn:
ISLBL.
+
discriminate.
+
pose (
IHtc l p p' (
k +
size a)).
repeat (
rewrite Z.add_assoc in e).
auto.
Qed.
Lemma label_pos_pvar_none:
forall tc l p p',
label_pos l p tc =
None ->
label_pos l p' tc =
None.
Proof.
Remark label_pos_pvar_some_add_add:
forall tc l p p' k k',
label_pos l (
p+
k')
tc =
Some (
p+
k) ->
label_pos l (
p'+
k')
tc =
Some (
p'+
k).
Proof.
induction tc.
-
intros.
simpl in H.
discriminate.
-
intros.
simpl in *.
destruct (
is_label _ _)
eqn:
ISLBL.
+
inv H.
assert (
k =
k')
by lia.
subst.
reflexivity.
+
pose (
IHtc l p p' k (
k' +
size a)).
repeat (
rewrite Z.add_assoc in e).
auto.
Qed.
Lemma label_pos_pvar_some_add:
forall tc l p p' k,
label_pos l p tc =
Some (
p+
k) ->
label_pos l p' tc =
Some (
p'+
k).
Proof.
Remark label_pos_pvar_add:
forall c tc l p p' k,
label_pos l (
p+
k)
c =
label_pos l p tc ->
label_pos l (
p'+
k)
c =
label_pos l p' tc.
Proof.
Lemma label_pos_pvar:
forall c tc l p p',
label_pos l p c =
label_pos l p tc ->
label_pos l p' c =
label_pos l p' tc.
Proof.
Lemma label_pos_head_app:
forall c bb lbb l tc p,
verified_schedule bb =
OK lbb ->
label_pos l p c =
label_pos l p tc ->
label_pos l p (
bb ::
c) =
label_pos l p (
lbb ++
tc).
Proof.
Lemma label_pos_preserved:
forall c tc l,
transf_blocks c =
OK tc ->
label_pos l 0
c =
label_pos l 0
tc.
Proof.
Lemma label_pos_preserved_blocks:
forall l f tf,
transf_function f =
OK tf ->
label_pos l 0 (
fn_blocks f) =
label_pos l 0 (
fn_blocks tf).
Proof.
Lemma transf_exec_control:
forall f tf ex rs m,
transf_function f =
OK tf ->
exec_control ge f ex rs m =
exec_control tge tf ex rs m.
Proof.
Lemma transf_exec_basic_instr:
forall i rs m,
exec_basic_instr ge i rs m =
exec_basic_instr tge i rs m.
Proof.
Lemma transf_exec_body:
forall bdy rs m,
exec_body ge bdy rs m =
exec_body tge bdy rs m.
Proof.
Lemma transf_exec_bblock:
forall f tf bb rs m,
transf_function f =
OK tf ->
exec_bblock ge f bb rs m =
exec_bblock tge tf bb rs m.
Proof.
Lemma transf_step_simu:
forall tf b lbb ofs c tbb rs m rs' m',
Genv.find_funct_ptr tge b =
Some (
Internal tf) ->
size_blocks (
fn_blocks tf) <=
Ptrofs.max_unsigned ->
rs PC =
Vptr b ofs ->
code_tail (
Ptrofs.unsigned ofs) (
fn_blocks tf) (
lbb ++
c) ->
concat_all lbb =
OK tbb ->
exec_bblock tge tf tbb rs m =
Next rs' m' ->
plus step tge (
State rs m)
E0 (
State rs' m').
Proof.
induction lbb.
-
intros until m'.
simpl.
intros.
discriminate.
-
intros until m'.
intros GFIND SIZE PCeq TAIL CONC EXEB.
destruct lbb.
+
simpl in *.
clear IHlbb.
inv CONC.
eapply plus_one.
econstructor;
eauto.
eapply find_bblock_tail;
eauto.
+
exploit concat_all_exec_bblock;
eauto;
try discriminate.
intros (
tbb0 &
rs0 &
m0 &
CONC0 &
EXEB0 &
PCeq' &
EXEB1).
eapply plus_left.
econstructor.
3:
eapply find_bblock_tail.
rewrite <-
app_comm_cons in TAIL. 3:
eauto.
all:
eauto.
eapply plus_star.
eapply IHlbb;
eauto.
rewrite PCeq in PCeq'.
simpl in PCeq'.
all:
eauto.
eapply code_tail_next_int;
eauto.
Qed.
Theorem transf_step_correct:
forall s1 t s2,
step ge s1 t s2 ->
forall s1' (
MS:
match_states s1 s1'),
(
exists s2',
plus step tge s1' t s2' /\
match_states s2 s2').
Proof.
Theorem transf_program_correct_Asmblock:
forward_simulation (
Asmblock.semantics prog) (
Asmblock.semantics tprog).
Proof.
End PRESERVATION_ASMBLOCK.
Require Import Asmvliw.
Lemma verified_par_checks_alls_bundles lb x:
forall bundle,
verify_par lb =
OK x ->
List.In bundle lb ->
verify_par_bblock bundle =
OK tt.
Proof.
induction lb; simpl; try tauto.
intros bundle H; monadInv H.
destruct 1; subst; eauto.
destruct x0; auto.
Qed.
Lemma verified_schedule_nob_checks_alls_bundles bb lb bundle:
verified_schedule_nob bb =
OK lb ->
List.In bundle lb ->
verify_par_bblock bundle =
OK tt.
Proof.
Lemma verify_par_bblock_PExpand bb i:
exit bb =
Some (
PExpand i) ->
verify_par_bblock bb =
OK tt.
Proof.
destruct bb as [
h bdy ext H];
simpl.
intros;
subst.
destruct i.
generalize H.
rewrite <-
wf_bblock_refl in H.
destruct H as [
H H0].
unfold builtin_alone in H0.
erewrite H0;
eauto.
Qed.
Local Hint Resolve verified_schedule_nob_checks_alls_bundles:
core.
Lemma verified_schedule_checks_alls_bundles bb lb bundle:
verified_schedule bb =
OK lb ->
List.In bundle lb ->
verify_par_bblock bundle =
OK tt.
Proof.
Lemma transf_blocks_checks_all_bundles lbb:
forall lb bundle,
transf_blocks lbb =
OK lb ->
List.In bundle lb ->
verify_par_bblock bundle =
OK tt.
Proof.
Lemma find_bblock_Some_in lb:
forall ofs b,
find_bblock ofs lb =
Some b ->
List.In b lb.
Proof.
induction lb;
simpl;
try congruence.
intros ofs b.
destruct (
zlt ofs 0);
try congruence.
destruct (
zeq ofs 0);
eauto.
intros X;
inversion X;
eauto.
Qed.
Section PRESERVATION_ASMVLIW.
Variables prog tprog:
program.
Hypothesis TRANSL:
match_prog prog tprog.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Lemma all_bundles_are_checked b ofs f bundle:
Genv.find_funct_ptr (
globalenv (
Asmblock.semantics tprog))
b =
Some (
Internal f) ->
find_bblock ofs (
fn_blocks f) =
Some bundle ->
verify_par_bblock bundle =
OK tt.
Proof.
Lemma checked_bundles_are_parexec_equiv f bundle rs rs' m m':
exec_bblock (
globalenv (
Asmblock.semantics tprog))
f bundle rs m =
Next rs' m' ->
verify_par_bblock bundle =
OK tt ->
det_parexec (
globalenv (
semantics tprog))
f bundle rs m rs' m'.
Proof.
Lemma seqexec_parexec_equiv b ofs f bundle rs rs' m m':
Genv.find_funct_ptr (
globalenv (
Asmblock.semantics tprog))
b =
Some (
Internal f) ->
find_bblock (
Ptrofs.unsigned ofs) (
fn_blocks f) =
Some bundle ->
exec_bblock (
globalenv (
Asmblock.semantics tprog))
f bundle rs m =
Next rs' m' ->
det_parexec (
globalenv (
semantics tprog))
f bundle rs m rs' m'.
Proof.
Theorem transf_program_correct_Asmvliw:
forward_simulation (
Asmblock.semantics tprog) (
Asmvliw.semantics tprog).
Proof.
End PRESERVATION_ASMVLIW.
Section PRESERVATION.
Variables prog tprog:
program.
Hypothesis TRANSL:
match_prog prog tprog.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Theorem transf_program_correct:
forward_simulation (
Asmblock.semantics prog) (
Asmvliw.semantics tprog).
Proof.
End PRESERVATION.