Correctness of instruction selection
Require Import FunInd.
Require Import Coqlib Maps.
Require Import AST Linking Errors Integers.
Require Import Values Memory Builtins Events Globalenvs Smallstep.
Require Import Switch Cminor Op CminorSel Cminortyping.
Require Import SelectOp SelectDiv SplitLong SelectLong Selection.
Require Import SelectOpproof SelectDivproof SplitLongproof SelectLongproof.
Local Open Scope cminorsel_scope.
Local Open Scope error_monad_scope.
Relational specification of instruction selection
Definition match_fundef (
cunit:
Cminor.program) (
f:
Cminor.fundef) (
tf:
CminorSel.fundef) :
Prop :=
exists hf,
helper_functions_declared cunit hf /\
sel_fundef (
prog_defmap cunit)
hf f =
OK tf.
Definition match_prog (
p:
Cminor.program) (
tp:
CminorSel.program) :=
match_program match_fundef eq p tp.
Processing of helper functions
Lemma record_globdefs_sound:
forall dm id gd, (
record_globdefs dm)!
id =
Some gd ->
dm!
id =
Some gd.
Proof.
Lemma lookup_helper_correct_1:
forall globs name sg id,
lookup_helper globs name sg =
OK id ->
globs!
id =
Some (
Gfun (
External (
EF_runtime name sg))).
Proof.
Lemma lookup_helper_correct:
forall p name sg id,
lookup_helper (
record_globdefs (
prog_defmap p))
name sg =
OK id ->
helper_declared p id name sg.
Proof.
Lemma get_helpers_correct:
forall p hf,
get_helpers (
prog_defmap p) =
OK hf ->
helper_functions_declared p hf.
Proof.
Theorem transf_program_match:
forall p tp,
sel_program p =
OK tp ->
match_prog p tp.
Proof.
Lemma helper_functions_declared_linkorder:
forall (
p p':
Cminor.program)
hf,
helper_functions_declared p hf ->
linkorder p p' ->
helper_functions_declared p'
hf.
Proof.
Correctness of the instruction selection functions for expressions
Section PRESERVATION.
Variable prog:
Cminor.program.
Variable tprog:
CminorSel.program.
Let ge :=
Genv.globalenv prog.
Let tge :=
Genv.globalenv tprog.
Hypothesis TRANSF:
match_prog prog tprog.
Lemma wt_prog :
wt_program prog.
Proof.
red;
intros.
destruct TRANSF as [
A _].
exploit list_forall2_in_left;
eauto.
intros ((
i' &
gd') &
B & (
C &
D)).
simpl in *.
inv D.
destruct H2 as (
hf &
P &
Q).
destruct f;
monadInv Q.
-
monadInv EQ.
econstructor;
apply type_function_sound;
eauto.
-
constructor.
Qed.
Lemma symbols_preserved:
forall (
s:
ident),
Genv.find_symbol tge s =
Genv.find_symbol ge s.
Proof (
Genv.find_symbol_match TRANSF).
Lemma senv_preserved:
Senv.equiv (
Genv.to_senv ge) (
Genv.to_senv tge).
Proof (
Genv.senv_match TRANSF).
Lemma function_ptr_translated:
forall (
b:
block) (
f:
Cminor.fundef),
Genv.find_funct_ptr ge b =
Some f ->
exists cu tf,
Genv.find_funct_ptr tge b =
Some tf /\
match_fundef cu f tf /\
linkorder cu prog.
Proof (
Genv.find_funct_ptr_match TRANSF).
Lemma functions_translated:
forall (
v v':
val) (
f:
Cminor.fundef),
Genv.find_funct ge v =
Some f ->
Val.lessdef v v' ->
exists cu tf,
Genv.find_funct tge v' =
Some tf /\
match_fundef cu f tf /\
linkorder cu prog.
Proof.
Lemma sig_function_translated:
forall cu f tf,
match_fundef cu f tf ->
funsig tf =
Cminor.funsig f.
Proof.
intros. destruct H as (hf & P & Q). destruct f; monadInv Q; auto. monadInv EQ; auto.
Qed.
Lemma stackspace_function_translated:
forall dm hf f tf,
sel_function dm hf f =
OK tf ->
fn_stackspace tf =
Cminor.fn_stackspace f.
Proof.
intros. monadInv H. auto.
Qed.
Lemma helper_functions_preserved:
forall hf,
helper_functions_declared prog hf ->
helper_functions_declared tprog hf.
Proof.
Section CMCONSTR.
Variable cunit:
Cminor.program.
Variable hf:
helper_functions.
Hypothesis LINK:
linkorder cunit prog.
Hypothesis HF:
helper_functions_declared cunit hf.
Let HF':
helper_functions_declared tprog hf.
Proof.
Variable sp:
val.
Variable e:
env.
Variable m:
mem.
Lemma eval_condexpr_of_expr:
forall a le v b,
eval_expr tge sp e m le a v ->
Val.bool_of_val v b ->
eval_condexpr tge sp e m le (
condexpr_of_expr a)
b.
Proof.
intros until a.
functional induction (
condexpr_of_expr a);
intros.
inv H.
econstructor;
eauto.
simpl in H6.
inv H6.
apply Val.bool_of_val_of_optbool.
auto.
inv H.
econstructor;
eauto.
destruct va;
eauto.
inv H.
econstructor;
eauto.
econstructor.
constructor.
eauto.
constructor.
simpl.
inv H0.
auto.
Qed.
Lemma eval_condition_of_expr:
forall a le v b,
eval_expr tge sp e m le a v ->
Val.bool_of_val v b ->
exists vl,
eval_exprlist tge sp e m le (
snd (
condition_of_expr a))
vl
/\
eval_condition (
fst (
condition_of_expr a))
vl m =
Some b.
Proof.
intros a;
functional induction (
condition_of_expr a);
intros;
simpl.
-
inv H.
exists vl;
split;
auto.
simpl in H6.
inv H6.
apply Val.bool_of_val_of_optbool in H0.
auto.
-
exists (
v ::
nil);
split.
constructor;
auto;
constructor.
inv H0;
simpl;
auto.
Qed.
Lemma eval_load:
forall le a v chunk v',
eval_expr tge sp e m le a v ->
Mem.loadv chunk m v =
Some v' ->
eval_expr tge sp e m le (
load chunk a)
v'.
Proof.
intros.
generalize H0;
destruct v;
simpl;
intro;
try discriminate.
unfold load.
generalize (
eval_addressing _ _ _ _ _ chunk _ _ _ _ H (
eq_refl _)).
destruct (
addressing chunk a).
intros [
vl [
EV EQ]].
eapply eval_Eload;
eauto.
Qed.
Lemma eval_store:
forall chunk a1 a2 v1 v2 f k m',
eval_expr tge sp e m nil a1 v1 ->
eval_expr tge sp e m nil a2 v2 ->
Mem.storev chunk m v1 v2 =
Some m' ->
step tge (
State f (
store chunk a1 a2)
k sp e m)
E0 (
State f Sskip k sp e m').
Proof.
Correctness of instruction selection for operators
Lemma eval_sel_unop:
forall le op a1 v1 v,
eval_expr tge sp e m le a1 v1 ->
eval_unop op v1 =
Some v ->
exists v',
eval_expr tge sp e m le (
sel_unop op a1)
v' /\
Val.lessdef v v'.
Proof.
Lemma eval_sel_binop:
forall le op a1 a2 v1 v2 v,
eval_expr tge sp e m le a1 v1 ->
eval_expr tge sp e m le a2 v2 ->
eval_binop op v1 v2 m =
Some v ->
exists v',
eval_expr tge sp e m le (
sel_binop op a1 a2)
v' /\
Val.lessdef v v'.
Proof.
Lemma eval_sel_select:
forall le a1 a2 a3 v1 v2 v3 b ty,
eval_expr tge sp e m le a1 v1 ->
eval_expr tge sp e m le a2 v2 ->
eval_expr tge sp e m le a3 v3 ->
Val.bool_of_val v1 b ->
exists v,
eval_expr tge sp e m le (
sel_select ty a1 a2 a3)
v
/\
Val.lessdef (
Val.select (
Some b)
v2 v3 ty)
v.
Proof.
Known built-in functions
Lemma eval_sel_known_builtin:
forall bf args a vl v le,
sel_known_builtin bf args =
Some a ->
eval_exprlist tge sp e m le args vl ->
builtin_function_sem bf vl =
Some v ->
exists v',
eval_expr tge sp e m le a v' /\
Val.lessdef v v'.
Proof.
intros until le;
intros SEL ARGS SEM.
destruct bf as [
bf|
bf];
simpl in SEL.
-
destruct bf;
try discriminate.
+
inv ARGS;
try discriminate.
inv H0;
try discriminate.
inv H2;
try discriminate.
inv H3;
try discriminate.
inv SEL.
simpl in SEM.
destruct v1;
inv SEM.
replace (
Val.normalize (
if Int.eq i Int.zero then v2 else v0)
t)
with (
Val.select (
Some (
negb (
Int.eq i Int.zero)))
v0 v2 t)
by (
destruct (
Int.eq i Int.zero);
reflexivity).
eapply eval_sel_select;
eauto.
constructor.
+
inv ARGS;
try discriminate.
inv H0;
try discriminate.
inv SEL.
simpl in SEM;
inv SEM.
apply eval_absf;
auto.
+
inv ARGS;
try discriminate.
inv H0;
try discriminate.
inv SEL.
simpl in SEM;
inv SEM.
apply eval_absfs;
auto.
-
eapply eval_platform_builtin;
eauto.
Qed.
End CMCONSTR.
Recognition of calls to built-in functions
Lemma expr_is_addrof_ident_correct:
forall e id,
expr_is_addrof_ident e =
Some id ->
e =
Cminor.Econst (
Cminor.Oaddrsymbol id Ptrofs.zero).
Proof.
Lemma classify_call_correct:
forall unit sp e m a v fd,
linkorder unit prog ->
Cminor.eval_expr ge sp e m a v ->
Genv.find_funct ge v =
Some fd ->
match classify_call (
prog_defmap unit)
a with
|
Call_default =>
True
|
Call_imm id =>
exists b,
Genv.find_symbol ge id =
Some b /\
v =
Vptr b Ptrofs.zero
|
Call_builtin ef =>
fd =
External ef
end.
Proof.
Translation of switch statements
Inductive Rint:
Z ->
val ->
Prop :=
|
Rint_intro:
forall n,
Rint (
Int.unsigned n) (
Vint n).
Inductive Rlong:
Z ->
val ->
Prop :=
|
Rlong_intro:
forall n,
Rlong (
Int64.unsigned n) (
Vlong n).
Section SEL_SWITCH.
Variable make_cmp_eq:
expr ->
Z ->
expr.
Variable make_cmp_ltu:
expr ->
Z ->
expr.
Variable make_sub:
expr ->
Z ->
expr.
Variable make_to_int:
expr ->
expr.
Variable modulus:
Z.
Variable R:
Z ->
val ->
Prop.
Hypothesis eval_make_cmp_eq:
forall sp e m le a v i n,
eval_expr tge sp e m le a v ->
R i v -> 0 <=
n <
modulus ->
eval_expr tge sp e m le (
make_cmp_eq a n) (
Val.of_bool (
zeq i n)).
Hypothesis eval_make_cmp_ltu:
forall sp e m le a v i n,
eval_expr tge sp e m le a v ->
R i v -> 0 <=
n <
modulus ->
eval_expr tge sp e m le (
make_cmp_ltu a n) (
Val.of_bool (
zlt i n)).
Hypothesis eval_make_sub:
forall sp e m le a v i n,
eval_expr tge sp e m le a v ->
R i v -> 0 <=
n <
modulus ->
exists v',
eval_expr tge sp e m le (
make_sub a n)
v' /\
R ((
i -
n)
mod modulus)
v'.
Hypothesis eval_make_to_int:
forall sp e m le a v i,
eval_expr tge sp e m le a v ->
R i v ->
exists v',
eval_expr tge sp e m le (
make_to_int a)
v' /\
Rint (
i mod Int.modulus)
v'.
Lemma sel_switch_correct_rec:
forall sp e m varg i x,
R i varg ->
forall t arg le,
wf_comptree modulus t ->
nth_error le arg =
Some varg ->
comptree_match modulus i t =
Some x ->
eval_exitexpr tge sp e m le (
sel_switch make_cmp_eq make_cmp_ltu make_sub make_to_int arg t)
x.
Proof.
Lemma sel_switch_correct:
forall dfl cases arg sp e m varg i t le,
validate_switch modulus dfl cases t =
true ->
eval_expr tge sp e m le arg varg ->
R i varg ->
0 <=
i <
modulus ->
eval_exitexpr tge sp e m le
(
XElet arg (
sel_switch make_cmp_eq make_cmp_ltu make_sub make_to_int O t))
(
switch_target i dfl cases).
Proof.
End SEL_SWITCH.
Section SEL_SWITCH_INT.
Variable cunit:
Cminor.program.
Variable hf:
helper_functions.
Hypothesis LINK:
linkorder cunit prog.
Hypothesis HF:
helper_functions_declared cunit hf.
Let HF':
helper_functions_declared tprog hf.
Proof.
Lemma sel_switch_int_correct:
forall dfl cases arg sp e m i t le,
validate_switch Int.modulus dfl cases t =
true ->
eval_expr tge sp e m le arg (
Vint i) ->
eval_exitexpr tge sp e m le (
XElet arg (
sel_switch_int O t)) (
switch_target (
Int.unsigned i)
dfl cases).
Proof.
Lemma sel_switch_long_correct:
forall dfl cases arg sp e m i t le,
validate_switch Int64.modulus dfl cases t =
true ->
eval_expr tge sp e m le arg (
Vlong i) ->
eval_exitexpr tge sp e m le (
XElet arg (
sel_switch_long O t)) (
switch_target (
Int64.unsigned i)
dfl cases).
Proof.
End SEL_SWITCH_INT.
Compatibility of evaluation functions with the "less defined than" relation.
Ltac TrivialExists :=
match goal with
| [ |-
exists v,
Some ?
x =
Some v /\
_ ] =>
exists x;
split;
auto
|
_ =>
idtac
end.
Lemma eval_unop_lessdef:
forall op v1 v1'
v,
eval_unop op v1 =
Some v ->
Val.lessdef v1 v1' ->
exists v',
eval_unop op v1' =
Some v' /\
Val.lessdef v v'.
Proof.
intros until v; intros EV LD. inv LD.
exists v; auto.
destruct op; simpl in *; inv EV; TrivialExists.
Qed.
Lemma eval_binop_lessdef:
forall op v1 v1'
v2 v2'
v m m',
eval_binop op v1 v2 m =
Some v ->
Val.lessdef v1 v1' ->
Val.lessdef v2 v2' ->
Mem.extends m m' ->
exists v',
eval_binop op v1'
v2'
m' =
Some v' /\
Val.lessdef v v'.
Proof.
Semantic preservation for instruction selection.
Relationship between the local environments.
Definition env_lessdef (
e1 e2:
env) :
Prop :=
forall id v1,
e1!
id =
Some v1 ->
exists v2,
e2!
id =
Some v2 /\
Val.lessdef v1 v2.
Lemma set_var_lessdef:
forall e1 e2 id v1 v2,
env_lessdef e1 e2 ->
Val.lessdef v1 v2 ->
env_lessdef (
PTree.set id v1 e1) (
PTree.set id v2 e2).
Proof.
intros;
red;
intros.
rewrite PTree.gsspec in *.
destruct (
peq id0 id).
exists v2;
split;
congruence.
auto.
Qed.
Lemma set_optvar_lessdef:
forall e1 e2 optid v1 v2,
env_lessdef e1 e2 ->
Val.lessdef v1 v2 ->
env_lessdef (
set_optvar optid v1 e1) (
set_optvar optid v2 e2).
Proof.
Lemma set_params_lessdef:
forall il vl1 vl2,
Val.lessdef_list vl1 vl2 ->
env_lessdef (
set_params vl1 il) (
set_params vl2 il).
Proof.
Lemma set_locals_lessdef:
forall e1 e2,
env_lessdef e1 e2 ->
forall il,
env_lessdef (
set_locals il e1) (
set_locals il e2).
Proof.
Semantic preservation for expressions.
Section EXPRESSIONS.
Variable cunit:
Cminor.program.
Variable hf:
helper_functions.
Hypothesis LINK:
linkorder cunit prog.
Hypothesis HF:
helper_functions_declared cunit hf.
Lemma sel_expr_correct:
forall sp e m a v,
Cminor.eval_expr ge sp e m a v ->
forall e'
le m',
env_lessdef e e' ->
Mem.extends m m' ->
exists v',
eval_expr tge sp e'
m'
le (
sel_expr a)
v' /\
Val.lessdef v v'.
Proof.
Lemma sel_exprlist_correct:
forall sp e m a v,
Cminor.eval_exprlist ge sp e m a v ->
forall e'
le m',
env_lessdef e e' ->
Mem.extends m m' ->
exists v',
eval_exprlist tge sp e'
m'
le (
sel_exprlist a)
v' /\
Val.lessdef_list v v'.
Proof.
induction 1;
intros;
simpl.
exists (@
nil val);
split;
auto.
constructor.
exploit sel_expr_correct;
eauto.
intros [
v1' [
A B]].
exploit IHeval_exprlist;
eauto.
intros [
vl' [
C D]].
exists (
v1' ::
vl');
split;
auto.
constructor;
eauto.
Qed.
Lemma sel_select_opt_correct:
forall ty cond a1 a2 a sp e m vcond v1 v2 b e'
m'
le,
sel_select_opt ty cond a1 a2 =
Some a ->
Cminor.eval_expr ge sp e m cond vcond ->
Cminor.eval_expr ge sp e m a1 v1 ->
Cminor.eval_expr ge sp e m a2 v2 ->
Val.bool_of_val vcond b ->
env_lessdef e e' ->
Mem.extends m m' ->
exists v',
eval_expr tge sp e'
m'
le a v' /\
Val.lessdef (
Val.select (
Some b)
v1 v2 ty)
v'.
Proof.
Lemma sel_builtin_arg_correct:
forall sp e e'
m m'
a v c,
env_lessdef e e' ->
Mem.extends m m' ->
Cminor.eval_expr ge sp e m a v ->
exists v',
CminorSel.eval_builtin_arg tge sp e'
m' (
sel_builtin_arg a c)
v'
/\
Val.lessdef v v'.
Proof.
Lemma sel_builtin_args_correct:
forall sp e e'
m m',
env_lessdef e e' ->
Mem.extends m m' ->
forall al vl,
Cminor.eval_exprlist ge sp e m al vl ->
forall cl,
exists vl',
list_forall2 (
CminorSel.eval_builtin_arg tge sp e'
m')
(
sel_builtin_args al cl)
vl'
/\
Val.lessdef_list vl vl'.
Proof.
induction 3;
intros;
simpl.
-
exists (@
nil val);
split;
constructor.
-
exploit sel_builtin_arg_correct;
eauto.
intros (
v1' &
A &
B).
edestruct IHeval_exprlist as (
vl' &
C &
D).
exists (
v1' ::
vl');
split;
auto.
constructor;
eauto.
Qed.
Lemma sel_builtin_res_correct:
forall oid v e v'
e',
env_lessdef e e' ->
Val.lessdef v v' ->
env_lessdef (
set_optvar oid v e) (
set_builtin_res (
sel_builtin_res oid)
v'
e').
Proof.
Lemma sel_builtin_default_correct:
forall optid ef al sp e1 m1 vl t v m2 e1'
m1'
f k,
Cminor.eval_exprlist ge sp e1 m1 al vl ->
external_call ef ge vl m1 t v m2 ->
env_lessdef e1 e1' ->
Mem.extends m1 m1' ->
exists e2'
m2',
plus step tge (
State f (
sel_builtin_default optid ef al)
k sp e1'
m1')
t (
State f Sskip k sp e2'
m2')
/\
env_lessdef (
set_optvar optid v e1)
e2'
/\
Mem.extends m2 m2'.
Proof.
Lemma sel_builtin_correct:
forall optid ef al sp e1 m1 vl t v m2 e1'
m1'
f k,
Cminor.eval_exprlist ge sp e1 m1 al vl ->
external_call ef ge vl m1 t v m2 ->
env_lessdef e1 e1' ->
Mem.extends m1 m1' ->
exists e2'
m2',
plus step tge (
State f (
sel_builtin optid ef al)
k sp e1'
m1')
t (
State f Sskip k sp e2'
m2')
/\
env_lessdef (
set_optvar optid v e1)
e2'
/\
Mem.extends m2 m2'.
Proof.
If-conversion
Lemma classify_stmt_sound_1:
forall f sp e m s k,
classify_stmt s =
SCskip ->
star Cminor.step ge (
Cminor.State f s k sp e m)
E0 (
Cminor.State f Cminor.Sskip k sp e m).
Proof.
Lemma classify_stmt_sound_2:
forall f sp e m a id v,
Cminor.eval_expr ge sp e m a v ->
forall s k,
classify_stmt s =
SCassign id a ->
star Cminor.step ge (
Cminor.State f s k sp e m)
E0 (
Cminor.State f Cminor.Sskip k sp (
PTree.set id v e)
m).
Proof.
intros until s;
functional induction (
classify_stmt s);
intros;
try discriminate.
-
inv H0.
apply star_one.
constructor;
auto.
-
eapply star_trans;
eauto.
eapply star_two.
constructor.
constructor.
traceEq.
traceEq.
-
eapply star_left.
constructor.
eapply star_right.
eauto.
constructor.
traceEq.
traceEq.
Qed.
Lemma classify_stmt_wt:
forall env tyret id a s,
classify_stmt s =
SCassign id a ->
wt_stmt env tyret s ->
wt_expr env a (
env id).
Proof.
intros until s;
functional induction (
classify_stmt s);
intros CL WT;
try discriminate.
-
inv CL;
inv WT;
auto.
-
inv WT;
eauto.
-
inv WT;
eauto.
Qed.
Lemma eval_select_safe_exprs:
forall a1 a2 f env ty e e'
m m'
sp cond vb b id s,
safe_expr (
known_id f)
a1 =
true ->
safe_expr (
known_id f)
a2 =
true ->
option_map (
fun sel =>
Sassign id sel) (
sel_select_opt ty cond a1 a2) =
Some s ->
Cminor.eval_expr ge sp e m cond vb ->
Val.bool_of_val vb b ->
wt_expr env a1 ty ->
wt_expr env a2 ty ->
def_env f e ->
wt_env env e ->
Cminor.eval_expr ge sp e m cond vb ->
Val.bool_of_val vb b ->
env_lessdef e e' ->
Mem.extends m m' ->
exists a'
v1 v2 v',
s =
Sassign id a'
/\
Cminor.eval_expr ge sp e m a1 v1
/\
Cminor.eval_expr ge sp e m a2 v2
/\
eval_expr tge sp e'
m'
nil a'
v'
/\
Val.lessdef (
if b then v1 else v2)
v'.
Proof.
Lemma if_conversion_correct:
forall f env tyret cond ifso ifnot s vb b k f'
k'
sp e m e'
m',
if_conversion (
known_id f)
env cond ifso ifnot =
Some s ->
def_env f e ->
wt_env env e ->
wt_stmt env tyret ifso ->
wt_stmt env tyret ifnot ->
Cminor.eval_expr ge sp e m cond vb ->
Val.bool_of_val vb b ->
env_lessdef e e' ->
Mem.extends m m' ->
let s0 :=
if b then ifso else ifnot in
exists e1 e1',
step tge (
State f'
s k'
sp e'
m')
E0 (
State f'
Sskip k'
sp e1'
m')
/\
star Cminor.step ge (
Cminor.State f s0 k sp e m)
E0 (
Cminor.State f Cminor.Sskip k sp e1 m)
/\
env_lessdef e1 e1'.
Proof.
End EXPRESSIONS.
Semantic preservation for functions and statements.
Inductive match_cont:
Cminor.program ->
helper_functions ->
known_idents ->
typenv ->
Cminor.cont ->
CminorSel.cont ->
Prop :=
|
match_cont_seq:
forall cunit hf ki env s s'
k k',
sel_stmt (
prog_defmap cunit)
ki env s =
OK s' ->
match_cont cunit hf ki env k k' ->
match_cont cunit hf ki env (
Cminor.Kseq s k) (
Kseq s'
k')
|
match_cont_block:
forall cunit hf ki env k k',
match_cont cunit hf ki env k k' ->
match_cont cunit hf ki env (
Cminor.Kblock k) (
Kblock k')
|
match_cont_other:
forall cunit hf ki env k k',
match_call_cont k k' ->
match_cont cunit hf ki env k k'
with match_call_cont:
Cminor.cont ->
CminorSel.cont ->
Prop :=
|
match_cont_stop:
match_call_cont Cminor.Kstop Kstop
|
match_cont_call:
forall cunit hf env id f sp e k f'
e'
k',
linkorder cunit prog ->
helper_functions_declared cunit hf ->
sel_function (
prog_defmap cunit)
hf f =
OK f' ->
type_function f =
OK env ->
match_cont cunit hf (
known_id f)
env k k' ->
env_lessdef e e' ->
match_call_cont (
Cminor.Kcall id f sp e k) (
Kcall id f'
sp e'
k').
Inductive match_states:
Cminor.state ->
CminorSel.state ->
Prop :=
|
match_state:
forall cunit hf f f'
s k s'
k'
sp e m e'
m'
env
(
LINK:
linkorder cunit prog)
(
HF:
helper_functions_declared cunit hf)
(
TF:
sel_function (
prog_defmap cunit)
hf f =
OK f')
(
TYF:
type_function f =
OK env)
(
TS:
sel_stmt (
prog_defmap cunit) (
known_id f)
env s =
OK s')
(
MC:
match_cont cunit hf (
known_id f)
env k k')
(
LD:
env_lessdef e e')
(
ME:
Mem.extends m m'),
match_states
(
Cminor.State f s k sp e m)
(
State f'
s'
k'
sp e'
m')
|
match_callstate:
forall cunit f f'
args args'
k k'
m m'
(
LINK:
linkorder cunit prog)
(
TF:
match_fundef cunit f f')
(
MC:
match_call_cont k k')
(
LD:
Val.lessdef_list args args')
(
ME:
Mem.extends m m'),
match_states
(
Cminor.Callstate f args k m)
(
Callstate f'
args'
k'
m')
|
match_returnstate:
forall v v'
k k'
m m'
(
MC:
match_call_cont k k')
(
LD:
Val.lessdef v v')
(
ME:
Mem.extends m m'),
match_states
(
Cminor.Returnstate v k m)
(
Returnstate v'
k'
m')
|
match_builtin_1:
forall cunit hf ef args optid f sp e k m al f'
e'
k'
m'
env
(
LINK:
linkorder cunit prog)
(
HF:
helper_functions_declared cunit hf)
(
TF:
sel_function (
prog_defmap cunit)
hf f =
OK f')
(
TYF:
type_function f =
OK env)
(
MC:
match_cont cunit hf (
known_id f)
env k k')
(
EA:
Cminor.eval_exprlist ge sp e m al args)
(
LDE:
env_lessdef e e')
(
ME:
Mem.extends m m'),
match_states
(
Cminor.Callstate (
External ef)
args (
Cminor.Kcall optid f sp e k)
m)
(
State f' (
sel_builtin optid ef al)
k'
sp e'
m')
|
match_builtin_2:
forall cunit hf v v'
optid f sp e k m f'
e'
m'
k'
env
(
LINK:
linkorder cunit prog)
(
HF:
helper_functions_declared cunit hf)
(
TF:
sel_function (
prog_defmap cunit)
hf f =
OK f')
(
TYF:
type_function f =
OK env)
(
MC:
match_cont cunit hf (
known_id f)
env k k')
(
LDV:
Val.lessdef v v')
(
LDE:
env_lessdef (
set_optvar optid v e)
e')
(
ME:
Mem.extends m m'),
match_states
(
Cminor.Returnstate v (
Cminor.Kcall optid f sp e k)
m)
(
State f'
Sskip k'
sp e'
m').
Remark call_cont_commut:
forall cunit hf ki env k k',
match_cont cunit hf ki env k k' ->
match_call_cont (
Cminor.call_cont k) (
call_cont k').
Proof.
induction 1; simpl; auto. inversion H; subst; auto.
Qed.
Remark match_is_call_cont:
forall cunit hf ki env k k',
match_cont cunit ki env hf k k' ->
Cminor.is_call_cont k ->
match_call_cont k k' /\
is_call_cont k'.
Proof.
destruct 1; intros; try contradiction. split; auto. inv H; auto.
Qed.
Definition nolabel (
s:
Cminor.stmt) :
Prop :=
forall lbl k,
Cminor.find_label lbl s k =
None.
Definition nolabel' (
s:
stmt) :
Prop :=
forall lbl k,
find_label lbl s k =
None.
Lemma classify_stmt_nolabel:
forall s,
classify_stmt s <>
SCother ->
nolabel s.
Proof.
intros s.
functional induction (
classify_stmt s);
intros.
-
red;
auto.
-
red;
auto.
-
apply IHs0 in H.
red;
intros;
simpl.
apply H.
-
apply IHs0 in H.
red;
intros;
simpl.
rewrite H;
auto.
-
congruence.
Qed.
Lemma if_conversion_base_nolabel:
forall (
hf:
helper_functions)
ki env a id a1 a2 s,
if_conversion_base ki env a id a1 a2 =
Some s ->
nolabel'
s.
Proof.
Lemma if_conversion_nolabel:
forall (
hf:
helper_functions)
ki env a s1 s2 s,
if_conversion ki env a s1 s2 =
Some s ->
nolabel s1 /\
nolabel s2 /\
nolabel'
s.
Proof.
Remark sel_builtin_nolabel:
forall (
hf:
helper_functions)
optid ef args,
nolabel' (
sel_builtin optid ef args).
Proof.
Remark find_label_commut:
forall cunit hf ki env lbl s k s'
k',
match_cont cunit hf ki env k k' ->
sel_stmt (
prog_defmap cunit)
ki env s =
OK s' ->
match Cminor.find_label lbl s k,
find_label lbl s'
k'
with
|
None,
None =>
True
|
Some(
s1,
k1),
Some(
s1',
k1') =>
sel_stmt (
prog_defmap cunit)
ki env s1 =
OK s1' /\
match_cont cunit hf ki env k1 k1'
|
_,
_ =>
False
end.
Proof.
Definition measure (
s:
Cminor.state) :
nat :=
match s with
|
Cminor.Callstate _ _ _ _ => 0%
nat
|
Cminor.State _ _ _ _ _ _ => 1%
nat
|
Cminor.Returnstate _ _ _ => 2%
nat
end.
Lemma sel_step_correct:
forall S1 t S2,
Cminor.step ge S1 t S2 ->
forall T1,
match_states S1 T1 ->
wt_state S1 ->
(
exists T2,
plus step tge T1 t T2 /\
match_states S2 T2)
\/ (
measure S2 <
measure S1 /\
t =
E0 /\
match_states S2 T1)%
nat
\/ (
exists S3 T2,
star Cminor.step ge S2 E0 S3 /\
step tge T1 t T2 /\
match_states S3 T2).
Proof.
Lemma sel_initial_states:
forall S,
Cminor.initial_state prog S ->
exists R,
initial_state tprog R /\
match_states S R.
Proof.
Lemma sel_final_states:
forall S R r,
match_states S R ->
Cminor.final_state S r ->
final_state R r.
Proof.
intros. inv H0. inv H. inv MC. inv LD. constructor.
Qed.
Theorem transf_program_correct:
forward_simulation (
Cminor.semantics prog) (
CminorSel.semantics tprog).
Proof.
End PRESERVATION.
Commutation with linking
Instance TransfSelectionLink :
TransfLink match_prog.
Proof.
red;
intros.
destruct (
link_linkorder _ _ _ H)
as [
LO1 LO2].
eapply link_match_program;
eauto.
intros.
elim H3;
intros hf1 [
A1 B1].
elim H4;
intros hf2 [
A2 B2].
Local Transparent Linker_fundef.
simpl in *.
destruct f1,
f2;
simpl in *;
monadInv B1;
monadInv B2;
simpl.
-
discriminate.
-
destruct e;
inv H2.
econstructor;
eauto.
-
destruct e;
inv H2.
econstructor;
eauto.
-
destruct (
external_function_eq e e0);
inv H2.
econstructor;
eauto.
Qed.