Correctness of instruction selection for integer division
Require Import Zquot Coqlib.
Require Import AST Integers Floats Values Memory Globalenvs Events.
Require Import Cminor Op CminorSel.
Require Import SelectOp SelectOpproof SplitLong SplitLongproof SelectLong SelectLongproof SelectDiv.
Local Open Scope cminorsel_scope.
Main approximation theorems
Section Z_DIV_MUL.
Variable N:
Z.
(* number of relevant bits *)
Hypothesis N_pos:
N >= 0.
Variable d:
Z.
(* divisor *)
Hypothesis d_pos:
d > 0.
This is theorem 4.2 from Granlund and Montgomery, PLDI 1994.
Lemma Zdiv_mul_pos:
forall m l,
l >= 0 ->
two_p (
N+
l) <=
m *
d <=
two_p (
N+
l) +
two_p l ->
forall n,
0 <=
n <
two_p N ->
Zdiv n d =
Zdiv (
m *
n) (
two_p (
N +
l)).
Proof.
Lemma Zdiv_unique_2:
forall x y q,
y > 0 -> 0 <
y *
q -
x <=
y ->
Zdiv x y =
q - 1.
Proof.
intros.
apply Zdiv_unique with (
x - (
q - 1) *
y).
ring.
replace ((
q - 1) *
y)
with (
y *
q -
y)
by ring.
omega.
Qed.
Lemma Zdiv_mul_opp:
forall m l,
l >= 0 ->
two_p (
N+
l) <
m *
d <=
two_p (
N+
l) +
two_p l ->
forall n,
0 <
n <=
two_p N ->
Zdiv n d = -
Zdiv (
m * (-
n)) (
two_p (
N +
l)) - 1.
Proof.
This is theorem 5.1 from Granlund and Montgomery, PLDI 1994.
Lemma Zquot_mul:
forall m l,
l >= 0 ->
two_p (
N+
l) <
m *
d <=
two_p (
N+
l) +
two_p l ->
forall n,
-
two_p N <=
n <
two_p N ->
Z.quot n d =
Zdiv (
m *
n) (
two_p (
N +
l)) + (
if zlt n 0
then 1
else 0).
Proof.
End Z_DIV_MUL.
Correctness of the division parameters
Lemma divs_mul_params_sound:
forall d m p,
divs_mul_params d =
Some(
p,
m) ->
0 <=
m <
Int.modulus /\ 0 <=
p < 32 /\
forall n,
Int.min_signed <=
n <=
Int.max_signed ->
Z.quot n d =
Zdiv (
m *
n) (
two_p (32 +
p)) + (
if zlt n 0
then 1
else 0).
Proof with
Lemma divu_mul_params_sound:
forall d m p,
divu_mul_params d =
Some(
p,
m) ->
0 <=
m <
Int.modulus /\ 0 <=
p < 32 /\
forall n,
0 <=
n <
Int.modulus ->
Zdiv n d =
Zdiv (
m *
n) (
two_p (32 +
p)).
Proof with
Lemma divs_mul_shift_gen:
forall x y m p,
divs_mul_params (
Int.signed y) =
Some(
p,
m) ->
0 <=
m <
Int.modulus /\ 0 <=
p < 32 /\
Int.divs x y =
Int.add (
Int.shr (
Int.repr ((
Int.signed x *
m) /
Int.modulus)) (
Int.repr p))
(
Int.shru x (
Int.repr 31)).
Proof.
Theorem divs_mul_shift_1:
forall x y m p,
divs_mul_params (
Int.signed y) =
Some(
p,
m) ->
m <
Int.half_modulus ->
0 <=
p < 32 /\
Int.divs x y =
Int.add (
Int.shr (
Int.mulhs x (
Int.repr m)) (
Int.repr p))
(
Int.shru x (
Int.repr 31)).
Proof.
Theorem divs_mul_shift_2:
forall x y m p,
divs_mul_params (
Int.signed y) =
Some(
p,
m) ->
m >=
Int.half_modulus ->
0 <=
p < 32 /\
Int.divs x y =
Int.add (
Int.shr (
Int.add (
Int.mulhs x (
Int.repr m))
x) (
Int.repr p))
(
Int.shru x (
Int.repr 31)).
Proof.
Theorem divu_mul_shift:
forall x y m p,
divu_mul_params (
Int.unsigned y) =
Some(
p,
m) ->
0 <=
p < 32 /\
Int.divu x y =
Int.shru (
Int.mulhu x (
Int.repr m)) (
Int.repr p).
Proof.
Same, for 64-bit integers
Lemma divls_mul_params_sound:
forall d m p,
divls_mul_params d =
Some(
p,
m) ->
0 <=
m <
Int64.modulus /\ 0 <=
p < 64 /\
forall n,
Int64.min_signed <=
n <=
Int64.max_signed ->
Z.quot n d =
Zdiv (
m *
n) (
two_p (64 +
p)) + (
if zlt n 0
then 1
else 0).
Proof with
Lemma divlu_mul_params_sound:
forall d m p,
divlu_mul_params d =
Some(
p,
m) ->
0 <=
m <
Int64.modulus /\ 0 <=
p < 64 /\
forall n,
0 <=
n <
Int64.modulus ->
Zdiv n d =
Zdiv (
m *
n) (
two_p (64 +
p)).
Proof with
Remark int64_shr'
_div_two_p:
forall x y,
Int64.shr'
x y =
Int64.repr (
Int64.signed x /
two_p (
Int.unsigned y)).
Proof.
Lemma divls_mul_shift_gen:
forall x y m p,
divls_mul_params (
Int64.signed y) =
Some(
p,
m) ->
0 <=
m <
Int64.modulus /\ 0 <=
p < 64 /\
Int64.divs x y =
Int64.add (
Int64.shr' (
Int64.repr ((
Int64.signed x *
m) /
Int64.modulus)) (
Int.repr p))
(
Int64.shru x (
Int64.repr 63)).
Proof.
Theorem divls_mul_shift_1:
forall x y m p,
divls_mul_params (
Int64.signed y) =
Some(
p,
m) ->
m <
Int64.half_modulus ->
0 <=
p < 64 /\
Int64.divs x y =
Int64.add (
Int64.shr' (
Int64.mulhs x (
Int64.repr m)) (
Int.repr p))
(
Int64.shru'
x (
Int.repr 63)).
Proof.
Theorem divls_mul_shift_2:
forall x y m p,
divls_mul_params (
Int64.signed y) =
Some(
p,
m) ->
m >=
Int64.half_modulus ->
0 <=
p < 64 /\
Int64.divs x y =
Int64.add (
Int64.shr' (
Int64.add (
Int64.mulhs x (
Int64.repr m))
x) (
Int.repr p))
(
Int64.shru'
x (
Int.repr 63)).
Proof.
Remark int64_shru'
_div_two_p:
forall x y,
Int64.shru'
x y =
Int64.repr (
Int64.unsigned x /
two_p (
Int.unsigned y)).
Proof.
Theorem divlu_mul_shift:
forall x y m p,
divlu_mul_params (
Int64.unsigned y) =
Some(
p,
m) ->
0 <=
p < 64 /\
Int64.divu x y =
Int64.shru' (
Int64.mulhu x (
Int64.repr m)) (
Int.repr p).
Proof.
Correctness of the smart constructors for division and modulus
Section CMCONSTRS.
Context mem `{
external_calls:
ExternalCalls mem} `{!
I64HelpersCorrect mem}.
Variable prog:
program.
Variable hf:
helper_functions.
Hypothesis HELPERS:
helper_functions_declared prog hf.
Let ge :=
Genv.globalenv prog.
Variable sp:
val.
Variable e:
env.
Variable m:
mem.
Lemma eval_divu_mul:
forall le x y p M,
divu_mul_params (
Int.unsigned y) =
Some(
p,
M) ->
nth_error le O =
Some (
Vint x) ->
eval_expr ge sp e m le (
divu_mul p M) (
Vint (
Int.divu x y)).
Proof.
Theorem eval_divuimm:
forall le e1 x n2 z,
eval_expr ge sp e m le e1 x ->
Val.divu x (
Vint n2) =
Some z ->
exists v,
eval_expr ge sp e m le (
divuimm e1 n2)
v /\
Val.lessdef z v.
Proof.
Theorem eval_divu:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.divu x y =
Some z ->
exists v,
eval_expr ge sp e m le (
divu a b)
v /\
Val.lessdef z v.
Proof.
Lemma eval_mod_from_div:
forall le a n x y,
eval_expr ge sp e m le a (
Vint y) ->
nth_error le O =
Some (
Vint x) ->
eval_expr ge sp e m le (
mod_from_div a n) (
Vint (
Int.sub x (
Int.mul y n))).
Proof.
unfold mod_from_div;
intros.
exploit eval_mulimm;
eauto.
instantiate (1 :=
n).
intros [
v [
A B]].
simpl in B.
inv B.
EvalOp.
Qed.
Theorem eval_moduimm:
forall le e1 x n2 z,
eval_expr ge sp e m le e1 x ->
Val.modu x (
Vint n2) =
Some z ->
exists v,
eval_expr ge sp e m le (
moduimm e1 n2)
v /\
Val.lessdef z v.
Proof.
Theorem eval_modu:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.modu x y =
Some z ->
exists v,
eval_expr ge sp e m le (
modu a b)
v /\
Val.lessdef z v.
Proof.
Lemma eval_divs_mul:
forall le x y p M,
divs_mul_params (
Int.signed y) =
Some(
p,
M) ->
nth_error le O =
Some (
Vint x) ->
eval_expr ge sp e m le (
divs_mul p M) (
Vint (
Int.divs x y)).
Proof.
Theorem eval_divsimm:
forall le e1 x n2 z,
eval_expr ge sp e m le e1 x ->
Val.divs x (
Vint n2) =
Some z ->
exists v,
eval_expr ge sp e m le (
divsimm e1 n2)
v /\
Val.lessdef z v.
Proof.
Theorem eval_divs:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.divs x y =
Some z ->
exists v,
eval_expr ge sp e m le (
divs a b)
v /\
Val.lessdef z v.
Proof.
Theorem eval_modsimm:
forall le e1 x n2 z,
eval_expr ge sp e m le e1 x ->
Val.mods x (
Vint n2) =
Some z ->
exists v,
eval_expr ge sp e m le (
modsimm e1 n2)
v /\
Val.lessdef z v.
Proof.
Theorem eval_mods:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.mods x y =
Some z ->
exists v,
eval_expr ge sp e m le (
mods a b)
v /\
Val.lessdef z v.
Proof.
Lemma eval_modl_from_divl:
forall le a n x y,
eval_expr ge sp e m le a (
Vlong y) ->
nth_error le O =
Some (
Vlong x) ->
eval_expr ge sp e m le (
modl_from_divl a n) (
Vlong (
Int64.sub x (
Int64.mul y n))).
Proof.
unfold modl_from_divl;
intros.
exploit eval_mullimm;
eauto.
eauto.
instantiate (1 :=
n).
intros (
v1 &
A1 &
B1).
assert (
A0:
eval_expr ge sp e m le (
Eletvar O) (
Vlong x))
by (
constructor;
auto).
exploit eval_subl;
auto.
eexact A0.
eexact A1.
intros (
v2 &
A2 &
B2).
simpl in B1;
inv B1.
simpl in B2;
inv B2.
exact A2.
Qed.
Lemma eval_divlu_mull:
forall le x y p M,
divlu_mul_params (
Int64.unsigned y) =
Some(
p,
M) ->
nth_error le O =
Some (
Vlong x) ->
eval_expr ge sp e m le (
divlu_mull p M) (
Vlong (
Int64.divu x y)).
Proof.
Theorem eval_divlu:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.divlu x y =
Some z ->
exists v,
eval_expr ge sp e m le (
divlu a b)
v /\
Val.lessdef z v.
Proof.
Theorem eval_modlu:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.modlu x y =
Some z ->
exists v,
eval_expr ge sp e m le (
modlu a b)
v /\
Val.lessdef z v.
Proof.
Lemma eval_divls_mull:
forall le x y p M,
divls_mul_params (
Int64.signed y) =
Some(
p,
M) ->
nth_error le O =
Some (
Vlong x) ->
eval_expr ge sp e m le (
divls_mull p M) (
Vlong (
Int64.divs x y)).
Proof.
Theorem eval_divls:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.divls x y =
Some z ->
exists v,
eval_expr ge sp e m le (
divls a b)
v /\
Val.lessdef z v.
Proof.
Theorem eval_modls:
forall le a b x y z,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
Val.modls x y =
Some z ->
exists v,
eval_expr ge sp e m le (
modls a b)
v /\
Val.lessdef z v.
Proof.
Floating-point division
Theorem eval_divf:
forall le a b x y,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
exists v,
eval_expr ge sp e m le (
divf a b)
v /\
Val.lessdef (
Val.divf x y)
v.
Proof.
intros until y.
unfold divf.
destruct (
divf_match b);
intros.
-
unfold divfimm.
destruct (
Float.exact_inverse n2)
as [
n2' | ]
eqn:
EINV.
+
inv H0.
inv H4.
simpl in H6.
inv H6.
econstructor;
split.
EvalOp.
constructor.
eauto.
constructor.
EvalOp.
simpl;
eauto.
constructor.
simpl;
eauto.
destruct x;
simpl;
auto.
erewrite Float.div_mul_inverse;
eauto.
+
TrivialExists.
-
TrivialExists.
Qed.
Theorem eval_divfs:
forall le a b x y,
eval_expr ge sp e m le a x ->
eval_expr ge sp e m le b y ->
exists v,
eval_expr ge sp e m le (
divfs a b)
v /\
Val.lessdef (
Val.divfs x y)
v.
Proof.
intros until y.
unfold divfs.
destruct (
divfs_match b);
intros.
-
unfold divfsimm.
destruct (
Float32.exact_inverse n2)
as [
n2' | ]
eqn:
EINV.
+
inv H0.
inv H4.
simpl in H6.
inv H6.
econstructor;
split.
EvalOp.
constructor.
eauto.
constructor.
EvalOp.
simpl;
eauto.
constructor.
simpl;
eauto.
destruct x;
simpl;
auto.
erewrite Float32.div_mul_inverse;
eauto.
+
TrivialExists.
-
TrivialExists.
Qed.
End CMCONSTRS.