Correctness proof for operator strength reduction.
Require Import Coqlib Compopts.
Require Import Integers Floats Values Memory Globalenvs Events.
Require Import Op Registers RTL ValueDomain ValueAOp ValueAnalysis.
Require Import ConstpropOp.
Section STRENGTH_REDUCTION.
Variable bc:
block_classification.
Variable ge:
genv.
Hypothesis GENV:
genv_match bc ge.
Variable sp:
block.
Hypothesis STACK:
bc sp =
BCstack.
Variable ae:
AE.t.
Variable e:
regset.
Variable m:
mem.
Hypothesis MATCH:
ematch bc e ae.
Lemma match_G:
forall r id ofs,
AE.get r ae =
Ptr(
Gl id ofs) ->
Val.lessdef e#
r (
Genv.symbol_address ge id ofs).
Proof.
Lemma match_S:
forall r ofs,
AE.get r ae =
Ptr(
Stk ofs) ->
Val.lessdef e#
r (
Vptr sp ofs).
Proof.
Ltac InvApproxRegs :=
match goal with
| [
H:
_ ::
_ =
_ ::
_ |-
_ ] =>
injection H;
clear H;
intros;
InvApproxRegs
| [
H: ?
v =
AE.get ?
r ae |-
_ ] =>
generalize (
MATCH r);
rewrite <-
H;
clear H;
intro;
InvApproxRegs
|
_ =>
idtac
end.
Ltac SimplVM :=
match goal with
| [
H:
vmatch _ ?
v (
I ?
n) |-
_ ] =>
let E :=
fresh in
assert (
E:
v =
Vint n)
by (
inversion H;
auto);
rewrite E in *;
clear H;
SimplVM
| [
H:
vmatch _ ?
v (
L ?
n) |-
_ ] =>
let E :=
fresh in
assert (
E:
v =
Vlong n)
by (
inversion H;
auto);
rewrite E in *;
clear H;
SimplVM
| [
H:
vmatch _ ?
v (
F ?
n) |-
_ ] =>
let E :=
fresh in
assert (
E:
v =
Vfloat n)
by (
inversion H;
auto);
rewrite E in *;
clear H;
SimplVM
| [
H:
vmatch _ ?
v (
FS ?
n) |-
_ ] =>
let E :=
fresh in
assert (
E:
v =
Vsingle n)
by (
inversion H;
auto);
rewrite E in *;
clear H;
SimplVM
| [
H:
vmatch _ ?
v (
Ptr(
Gl ?
id ?
ofs)) |-
_ ] =>
let E :=
fresh in
assert (
E:
Val.lessdef v (
Genv.symbol_address ge id ofs))
by (
eapply vmatch_ptr_gl;
eauto);
clear H;
SimplVM
| [
H:
vmatch _ ?
v (
Ptr(
Stk ?
ofs)) |-
_ ] =>
let E :=
fresh in
assert (
E:
Val.lessdef v (
Vptr sp ofs))
by (
eapply vmatch_ptr_stk;
eauto);
clear H;
SimplVM
|
_ =>
idtac
end.
Lemma eval_Olea_ptr:
forall a el,
eval_operation ge (
Vptr sp Ptrofs.zero) (
Olea_ptr a)
el m =
eval_addressing ge (
Vptr sp Ptrofs.zero)
a el.
Proof.
Lemma const_for_result_correct:
forall a op v,
const_for_result a =
Some op ->
vmatch bc v a ->
exists v',
eval_operation ge (
Vptr sp Ptrofs.zero)
op nil m =
Some v' /\
Val.lessdef v v'.
Proof.
Lemma cond_strength_reduction_correct:
forall cond args vl,
vl =
map (
fun r =>
AE.get r ae)
args ->
let (
cond',
args') :=
cond_strength_reduction cond args vl in
eval_condition cond'
e##
args'
m =
eval_condition cond e##
args m.
Proof.
Lemma addr_strength_reduction_32_generic_correct:
forall addr args vl res,
vl =
map (
fun r =>
AE.get r ae)
args ->
eval_addressing32 ge (
Vptr sp Ptrofs.zero)
addr e##
args =
Some res ->
let (
addr',
args') :=
addr_strength_reduction_32_generic addr args vl in
exists res',
eval_addressing32 ge (
Vptr sp Ptrofs.zero)
addr'
e##
args' =
Some res' /\
Val.lessdef res res'.
Proof.
Lemma addr_strength_reduction_32_correct:
forall addr args vl res,
vl =
map (
fun r =>
AE.get r ae)
args ->
eval_addressing32 ge (
Vptr sp Ptrofs.zero)
addr e##
args =
Some res ->
let (
addr',
args') :=
addr_strength_reduction_32 addr args vl in
exists res',
eval_addressing32 ge (
Vptr sp Ptrofs.zero)
addr'
e##
args' =
Some res' /\
Val.lessdef res res'.
Proof.
Lemma addr_strength_reduction_64_generic_correct:
forall addr args vl res,
vl =
map (
fun r =>
AE.get r ae)
args ->
eval_addressing64 ge (
Vptr sp Ptrofs.zero)
addr e##
args =
Some res ->
let (
addr',
args') :=
addr_strength_reduction_64_generic addr args vl in
exists res',
eval_addressing64 ge (
Vptr sp Ptrofs.zero)
addr'
e##
args' =
Some res' /\
Val.lessdef res res'.
Proof.
Lemma addr_strength_reduction_64_correct:
forall addr args vl res,
vl =
map (
fun r =>
AE.get r ae)
args ->
eval_addressing64 ge (
Vptr sp Ptrofs.zero)
addr e##
args =
Some res ->
let (
addr',
args') :=
addr_strength_reduction_64 addr args vl in
exists res',
eval_addressing64 ge (
Vptr sp Ptrofs.zero)
addr'
e##
args' =
Some res' /\
Val.lessdef res res'.
Proof.
Lemma addr_strength_reduction_correct:
forall addr args vl res,
vl =
map (
fun r =>
AE.get r ae)
args ->
eval_addressing ge (
Vptr sp Ptrofs.zero)
addr e##
args =
Some res ->
let (
addr',
args') :=
addr_strength_reduction addr args vl in
exists res',
eval_addressing ge (
Vptr sp Ptrofs.zero)
addr'
e##
args' =
Some res' /\
Val.lessdef res res'.
Proof.
Lemma make_cmp_base_correct:
forall c args vl,
vl =
map (
fun r =>
AE.get r ae)
args ->
let (
op',
args') :=
make_cmp_base c args vl in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op'
e##
args'
m =
Some v
/\
Val.lessdef (
Val.of_optbool (
eval_condition c e##
args m))
v.
Proof.
Lemma make_cmp_correct:
forall c args vl,
vl =
map (
fun r =>
AE.get r ae)
args ->
let (
op',
args') :=
make_cmp c args vl in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op'
e##
args'
m =
Some v
/\
Val.lessdef (
Val.of_optbool (
eval_condition c e##
args m))
v.
Proof.
Lemma make_select_correct:
forall c ty r1 r2 args vl,
vl =
map (
fun r =>
AE.get r ae)
args ->
let (
op',
args') :=
make_select c ty r1 r2 args vl in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op'
e##
args'
m =
Some v
/\
Val.lessdef (
Val.select (
eval_condition c e##
args m)
e#
r1 e#
r2 ty)
v.
Proof.
Lemma make_addimm_correct:
forall n r,
let (
op,
args) :=
make_addimm n r in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.add e#
r (
Vint n))
v.
Proof.
Lemma make_shlimm_correct:
forall n r1 r2,
e#
r2 =
Vint n ->
let (
op,
args) :=
make_shlimm n r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.shl e#
r1 (
Vint n))
v.
Proof.
Lemma make_shrimm_correct:
forall n r1 r2,
e#
r2 =
Vint n ->
let (
op,
args) :=
make_shrimm n r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.shr e#
r1 (
Vint n))
v.
Proof.
Lemma make_shruimm_correct:
forall n r1 r2,
e#
r2 =
Vint n ->
let (
op,
args) :=
make_shruimm n r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.shru e#
r1 (
Vint n))
v.
Proof.
Lemma make_mulimm_correct:
forall n r1,
let (
op,
args) :=
make_mulimm n r1 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.mul e#
r1 (
Vint n))
v.
Proof.
Lemma make_divimm_correct:
forall n r1 r2 v,
Val.divs e#
r1 e#
r2 =
Some v ->
e#
r2 =
Vint n ->
let (
op,
args) :=
make_divimm n r1 r2 in
exists w,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some w /\
Val.lessdef v w.
Proof.
Lemma make_divuimm_correct:
forall n r1 r2 v,
Val.divu e#
r1 e#
r2 =
Some v ->
e#
r2 =
Vint n ->
let (
op,
args) :=
make_divuimm n r1 r2 in
exists w,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some w /\
Val.lessdef v w.
Proof.
Lemma make_moduimm_correct:
forall n r1 r2 v,
Val.modu e#
r1 e#
r2 =
Some v ->
e#
r2 =
Vint n ->
let (
op,
args) :=
make_moduimm n r1 r2 in
exists w,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some w /\
Val.lessdef v w.
Proof.
Lemma make_andimm_correct:
forall n r x,
vmatch bc e#
r x ->
let (
op,
args) :=
make_andimm n r x in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.and e#
r (
Vint n))
v.
Proof.
Lemma make_orimm_correct:
forall n r,
let (
op,
args) :=
make_orimm n r in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.or e#
r (
Vint n))
v.
Proof.
Lemma make_xorimm_correct:
forall n r,
let (
op,
args) :=
make_xorimm n r in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.xor e#
r (
Vint n))
v.
Proof.
Lemma make_addlimm_correct:
forall n r,
let (
op,
args) :=
make_addlimm n r in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.addl e#
r (
Vlong n))
v.
Proof.
Lemma make_shllimm_correct:
forall n r1 r2,
e#
r2 =
Vint n ->
let (
op,
args) :=
make_shllimm n r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.shll e#
r1 (
Vint n))
v.
Proof.
Lemma make_shrlimm_correct:
forall n r1 r2,
e#
r2 =
Vint n ->
let (
op,
args) :=
make_shrlimm n r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.shrl e#
r1 (
Vint n))
v.
Proof.
Lemma make_shrluimm_correct:
forall n r1 r2,
e#
r2 =
Vint n ->
let (
op,
args) :=
make_shrluimm n r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.shrlu e#
r1 (
Vint n))
v.
Proof.
Lemma make_mullimm_correct:
forall n r1,
let (
op,
args) :=
make_mullimm n r1 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.mull e#
r1 (
Vlong n))
v.
Proof.
Lemma make_divlimm_correct:
forall n r1 r2 v,
Val.divls e#
r1 e#
r2 =
Some v ->
e#
r2 =
Vlong n ->
let (
op,
args) :=
make_divlimm n r1 r2 in
exists w,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some w /\
Val.lessdef v w.
Proof.
intros;
unfold make_divlimm.
destruct (
Int64.is_power2'
n)
eqn:?.
destruct (
Int.ltu i (
Int.repr 63))
eqn:?.
rewrite H0 in H.
econstructor;
split.
simpl;
eauto.
eapply Val.divls_pow2;
eauto.
auto.
exists v;
auto.
exists v;
auto.
Qed.
Lemma make_divluimm_correct:
forall n r1 r2 v,
Val.divlu e#
r1 e#
r2 =
Some v ->
e#
r2 =
Vlong n ->
let (
op,
args) :=
make_divluimm n r1 r2 in
exists w,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some w /\
Val.lessdef v w.
Proof.
intros;
unfold make_divluimm.
destruct (
Int64.is_power2'
n)
eqn:?.
econstructor;
split.
simpl;
eauto.
rewrite H0 in H.
destruct (
e#
r1);
inv H.
destruct (
Int64.eq n Int64.zero);
inv H2.
simpl.
erewrite Int64.is_power2'
_range by eauto.
erewrite Int64.divu_pow2'
by eauto.
auto.
exists v;
auto.
Qed.
Lemma make_modluimm_correct:
forall n r1 r2 v,
Val.modlu e#
r1 e#
r2 =
Some v ->
e#
r2 =
Vlong n ->
let (
op,
args) :=
make_modluimm n r1 r2 in
exists w,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some w /\
Val.lessdef v w.
Proof.
Lemma make_andlimm_correct:
forall n r x,
let (
op,
args) :=
make_andlimm n r x in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.andl e#
r (
Vlong n))
v.
Proof.
Lemma make_orlimm_correct:
forall n r,
let (
op,
args) :=
make_orlimm n r in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.orl e#
r (
Vlong n))
v.
Proof.
Lemma make_xorlimm_correct:
forall n r,
let (
op,
args) :=
make_xorlimm n r in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.xorl e#
r (
Vlong n))
v.
Proof.
Lemma make_mulfimm_correct:
forall n r1 r2,
e#
r2 =
Vfloat n ->
let (
op,
args) :=
make_mulfimm n r1 r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.mulf e#
r1 e#
r2)
v.
Proof.
Lemma make_mulfimm_correct_2:
forall n r1 r2,
e#
r1 =
Vfloat n ->
let (
op,
args) :=
make_mulfimm n r2 r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.mulf e#
r1 e#
r2)
v.
Proof.
Lemma make_mulfsimm_correct:
forall n r1 r2,
e#
r2 =
Vsingle n ->
let (
op,
args) :=
make_mulfsimm n r1 r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.mulfs e#
r1 e#
r2)
v.
Proof.
Lemma make_mulfsimm_correct_2:
forall n r1 r2,
e#
r1 =
Vsingle n ->
let (
op,
args) :=
make_mulfsimm n r2 r1 r2 in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.mulfs e#
r1 e#
r2)
v.
Proof.
Lemma make_cast8signed_correct:
forall r x,
vmatch bc e#
r x ->
let (
op,
args) :=
make_cast8signed r x in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.sign_ext 8
e#
r)
v.
Proof.
Lemma make_cast8unsigned_correct:
forall r x,
vmatch bc e#
r x ->
let (
op,
args) :=
make_cast8unsigned r x in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.zero_ext 8
e#
r)
v.
Proof.
Lemma make_cast16signed_correct:
forall r x,
vmatch bc e#
r x ->
let (
op,
args) :=
make_cast16signed r x in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.sign_ext 16
e#
r)
v.
Proof.
Lemma make_cast16unsigned_correct:
forall r x,
vmatch bc e#
r x ->
let (
op,
args) :=
make_cast16unsigned r x in
exists v,
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v /\
Val.lessdef (
Val.zero_ext 16
e#
r)
v.
Proof.
Lemma op_strength_reduction_correct:
forall op args vl v,
vl =
map (
fun r =>
AE.get r ae)
args ->
eval_operation ge (
Vptr sp Ptrofs.zero)
op e##
args m =
Some v ->
let (
op',
args') :=
op_strength_reduction op args vl in
exists w,
eval_operation ge (
Vptr sp Ptrofs.zero)
op'
e##
args'
m =
Some w /\
Val.lessdef v w.
Proof.
End STRENGTH_REDUCTION.