Formalizations of machine integers modulo 2N.
Require Import Axioms.
Require Import Coqlib.
Comparisons
Inductive comparison :
Type :=
|
Ceq :
comparison (* same *)
|
Cne :
comparison (* different *)
|
Clt :
comparison (* less than *)
|
Cle :
comparison (* less than or equal *)
|
Cgt :
comparison (* greater than *)
|
Cge :
comparison.
(* greater than or equal *)
Definition negate_comparison (
c:
comparison):
comparison :=
match c with
|
Ceq =>
Cne
|
Cne =>
Ceq
|
Clt =>
Cge
|
Cle =>
Cgt
|
Cgt =>
Cle
|
Cge =>
Clt
end.
Definition swap_comparison (
c:
comparison):
comparison :=
match c with
|
Ceq =>
Ceq
|
Cne =>
Cne
|
Clt =>
Cgt
|
Cle =>
Cge
|
Cgt =>
Clt
|
Cge =>
Cle
end.
Parameterization by the word size, in bits.
Module Type WORDSIZE.
Variable wordsize:
nat.
Axiom wordsize_not_zero:
wordsize <> 0%
nat.
End WORDSIZE.
Module Make(
WS:
WORDSIZE).
Definition wordsize:
nat :=
WS.wordsize.
Definition modulus :
Z :=
two_power_nat wordsize.
Definition half_modulus :
Z :=
modulus / 2.
Definition max_unsigned :
Z :=
modulus - 1.
Definition max_signed :
Z :=
half_modulus - 1.
Definition min_signed :
Z := -
half_modulus.
Remark wordsize_pos:
Z_of_nat wordsize > 0.
Proof.
Remark modulus_power:
modulus =
two_p (
Z_of_nat wordsize).
Proof.
Remark modulus_pos:
modulus > 0.
Proof.
Representation of machine integers
A machine integer (type int) is represented as a Coq arbitrary-precision
integer (type Z) plus a proof that it is in the range 0 (included) to
modulus (excluded.
Record int:
Type :=
mkint {
intval:
Z;
intrange: 0 <=
intval <
modulus }.
The unsigned and signed functions return the Coq integer corresponding
to the given machine integer, interpreted as unsigned or signed
respectively.
Definition unsigned (
n:
int) :
Z :=
intval n.
Definition signed (
n:
int) :
Z :=
if zlt (
unsigned n)
half_modulus
then unsigned n
else unsigned n -
modulus.
Conversely, repr takes a Coq integer and returns the corresponding
machine integer. The argument is treated modulo modulus.
Definition repr (
x:
Z) :
int :=
mkint (
Zmod x modulus) (
Z_mod_lt x modulus modulus_pos).
Definition zero :=
repr 0.
Definition one :=
repr 1.
Definition mone :=
repr (-1).
Definition iwordsize :=
repr (
Z_of_nat wordsize).
Lemma mkint_eq:
forall x y Px Py,
x =
y ->
mkint x Px =
mkint y Py.
Proof.
intros.
subst y.
generalize (
proof_irr Px Py);
intro.
subst Py.
reflexivity.
Qed.
Lemma eq_dec:
forall (
x y:
int), {
x =
y} + {
x <>
y}.
Proof.
intros.
destruct x;
destruct y.
case (
zeq intval0 intval1);
intro.
left.
apply mkint_eq.
auto.
right.
red;
intro.
injection H.
exact n.
Qed.
Arithmetic and logical operations over machine integers
Definition eq (
x y:
int) :
bool :=
if zeq (
unsigned x) (
unsigned y)
then true else false.
Definition lt (
x y:
int) :
bool :=
if zlt (
signed x) (
signed y)
then true else false.
Definition ltu (
x y:
int) :
bool :=
if zlt (
unsigned x) (
unsigned y)
then true else false.
Definition neg (
x:
int) :
int :=
repr (-
unsigned x).
Definition add (
x y:
int) :
int :=
repr (
unsigned x +
unsigned y).
Definition sub (
x y:
int) :
int :=
repr (
unsigned x -
unsigned y).
Definition mul (
x y:
int) :
int :=
repr (
unsigned x *
unsigned y).
Zdiv_round and Zmod_round implement signed division and modulus
with round-towards-zero behavior.
Definition Zdiv_round (
x y:
Z) :
Z :=
if zlt x 0
then
if zlt y 0
then (-
x) / (-
y)
else - ((-
x) /
y)
else
if zlt y 0
then -(
x / (-
y))
else x /
y.
Definition Zmod_round (
x y:
Z) :
Z :=
x - (
Zdiv_round x y) *
y.
Definition divs (
x y:
int) :
int :=
repr (
Zdiv_round (
signed x) (
signed y)).
Definition mods (
x y:
int) :
int :=
repr (
Zmod_round (
signed x) (
signed y)).
Definition divu (
x y:
int) :
int :=
repr (
unsigned x /
unsigned y).
Definition modu (
x y:
int) :
int :=
repr (
Zmod (
unsigned x) (
unsigned y)).
Definition add_carry (
x y cin:
int):
int :=
if zlt (
unsigned x +
unsigned y +
unsigned cin)
modulus
then zero
else one.
For bitwise operations, we need to convert between Coq integers Z
and their bit-level representations. Bit-level representations are
represented as characteristic functions, that is, functions f
of type nat -> bool such that f i is the value of the i-th bit
of the number. For i less than 0 or greater or equal to wordsize,
the characteristic function is false.
Definition Z_shift_add (
b:
bool) (
x:
Z) :=
if b then 2 *
x + 1
else 2 *
x.
Definition Z_bin_decomp (
x:
Z) :
bool *
Z :=
match x with
|
Z0 => (
false, 0)
|
Zpos p =>
match p with
|
xI q => (
true,
Zpos q)
|
xO q => (
false,
Zpos q)
|
xH => (
true, 0)
end
|
Zneg p =>
match p with
|
xI q => (
true,
Zneg q - 1)
|
xO q => (
false,
Zneg q)
|
xH => (
true, -1)
end
end.
Fixpoint bits_of_Z (
n:
nat) (
x:
Z) {
struct n}:
Z ->
bool :=
match n with
|
O =>
(
fun i:
Z =>
false)
|
S m =>
let (
b,
y) :=
Z_bin_decomp x in
let f :=
bits_of_Z m y in
(
fun i:
Z =>
if zeq i 0
then b else f (
i - 1))
end.
Fixpoint Z_of_bits (
n:
nat) (
f:
Z ->
bool) (
i:
Z) {
struct n}:
Z :=
match n with
|
O => 0
|
S m =>
Z_shift_add (
f i) (
Z_of_bits m f (
Zsucc i))
end.
Bitwise logical "and", "or" and "xor" operations.
Definition bitwise_binop (
f:
bool ->
bool ->
bool) (
x y:
int) :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
let fy :=
bits_of_Z wordsize (
unsigned y)
in
repr(
Z_of_bits wordsize (
fun i =>
f (
fx i) (
fy i)) 0).
Definition and (
x y:
int):
int :=
bitwise_binop andb x y.
Definition or (
x y:
int):
int :=
bitwise_binop orb x y.
Definition xor (
x y:
int) :
int :=
bitwise_binop xorb x y.
Definition not (
x:
int) :
int :=
xor x mone.
Shifts and rotates.
Definition shl (
x y:
int):
int :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
repr (
Z_of_bits wordsize fx (-
unsigned y)).
Definition shru (
x y:
int):
int :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
repr (
Z_of_bits wordsize fx (
unsigned y)).
Definition shr (
x y:
int):
int :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
let sx :=
fun i =>
fx (
if zlt i (
Z_of_nat wordsize)
then i else Z_of_nat wordsize - 1)
in
repr (
Z_of_bits wordsize sx (
unsigned y)).
Definition rol (
x y:
int) :
int :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
let rx :=
fun i =>
fx (
Zmod i (
Z_of_nat wordsize))
in
repr (
Z_of_bits wordsize rx (-
unsigned y)).
Definition ror (
x y:
int) :
int :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
let rx :=
fun i =>
fx (
Zmod i (
Z_of_nat wordsize))
in
repr (
Z_of_bits wordsize rx (
unsigned y)).
Definition rolm (
x a m:
int):
int :=
and (
rol x a)
m.
Viewed as signed divisions by powers of two, shrx rounds towards
zero, while shr rounds towards minus infinity.
Definition shrx (
x y:
int):
int :=
divs x (
shl one y).
shr_carry x y is 1 if x is negative and at least one 1 bit is shifted away.
Definition shr_carry (
x y:
int) :=
if lt x zero &&
negb (
eq (
and x (
sub (
shl one y)
one))
zero)
then one else zero.
Zero and sign extensions
Definition zero_ext (
n:
Z) (
x:
int) :
int :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
repr (
Z_of_bits wordsize (
fun i =>
if zlt i n then fx i else false) 0).
Definition sign_ext (
n:
Z) (
x:
int) :
int :=
let fx :=
bits_of_Z wordsize (
unsigned x)
in
repr (
Z_of_bits wordsize (
fun i =>
if zlt i n then fx i else fx (
n - 1)) 0).
Decomposition of a number as a sum of powers of two.
Fixpoint Z_one_bits (
n:
nat) (
x:
Z) (
i:
Z) {
struct n}:
list Z :=
match n with
|
O =>
nil
|
S m =>
let (
b,
y) :=
Z_bin_decomp x in
if b then i ::
Z_one_bits m y (
i+1)
else Z_one_bits m y (
i+1)
end.
Definition one_bits (
x:
int) :
list int :=
List.map repr (
Z_one_bits wordsize (
unsigned x) 0).
Recognition of powers of two.
Definition is_power2 (
x:
int) :
option int :=
match Z_one_bits wordsize (
unsigned x) 0
with
|
i ::
nil =>
Some (
repr i)
|
_ =>
None
end.
Comparisons.
Definition cmp (
c:
comparison) (
x y:
int) :
bool :=
match c with
|
Ceq =>
eq x y
|
Cne =>
negb (
eq x y)
|
Clt =>
lt x y
|
Cle =>
negb (
lt y x)
|
Cgt =>
lt y x
|
Cge =>
negb (
lt x y)
end.
Definition cmpu (
c:
comparison) (
x y:
int) :
bool :=
match c with
|
Ceq =>
eq x y
|
Cne =>
negb (
eq x y)
|
Clt =>
ltu x y
|
Cle =>
negb (
ltu y x)
|
Cgt =>
ltu y x
|
Cge =>
negb (
ltu x y)
end.
Definition is_false (
x:
int) :
Prop :=
x =
zero.
Definition is_true (
x:
int) :
Prop :=
x <>
zero.
Definition notbool (
x:
int) :
int :=
if eq x zero then one else zero.
Properties of integers and integer arithmetic
Properties of modulus, max_unsigned, etc.
Remark half_modulus_power:
half_modulus =
two_p (
Z_of_nat wordsize - 1).
Proof.
Remark half_modulus_modulus:
modulus = 2 *
half_modulus.
Proof.
Relative positions, from greatest to smallest:
max_unsigned
max_signed
2*wordsize-1
wordsize
0
min_signed
Remark half_modulus_pos:
half_modulus > 0.
Proof.
Remark min_signed_neg:
min_signed < 0.
Proof.
Remark max_signed_pos:
max_signed >= 0.
Proof.
Remark wordsize_max_unsigned:
Z_of_nat wordsize <=
max_unsigned.
Proof.
Remark two_wordsize_max_unsigned: 2 *
Z_of_nat wordsize - 1 <=
max_unsigned.
Proof.
Remark max_signed_unsigned:
max_signed <
max_unsigned.
Proof.
Modulo arithmetic
We define and state properties of equality and arithmetic modulo a
positive integer.
Section EQ_MODULO.
Variable modul:
Z.
Hypothesis modul_pos:
modul > 0.
Definition eqmod (
x y:
Z) :
Prop :=
exists k,
x =
k *
modul +
y.
Lemma eqmod_refl:
forall x,
eqmod x x.
Proof.
intros; red. exists 0. omega.
Qed.
Lemma eqmod_refl2:
forall x y,
x =
y ->
eqmod x y.
Proof.
Lemma eqmod_sym:
forall x y,
eqmod x y ->
eqmod y x.
Proof.
intros x y [k EQ]; red. exists (-k). subst x. ring.
Qed.
Lemma eqmod_trans:
forall x y z,
eqmod x y ->
eqmod y z ->
eqmod x z.
Proof.
intros x y z [k1 EQ1] [k2 EQ2]; red.
exists (k1 + k2). subst x; subst y. ring.
Qed.
Lemma eqmod_small_eq:
forall x y,
eqmod x y -> 0 <=
x <
modul -> 0 <=
y <
modul ->
x =
y.
Proof.
intros x y [
k EQ]
I1 I2.
generalize (
Zdiv_unique _ _ _ _ EQ I2).
intro.
rewrite (
Zdiv_small x modul I1)
in H.
subst k.
omega.
Qed.
Lemma eqmod_mod_eq:
forall x y,
eqmod x y ->
x mod modul =
y mod modul.
Proof.
Lemma eqmod_mod:
forall x,
eqmod x (
x mod modul).
Proof.
Lemma eqmod_add:
forall a b c d,
eqmod a b ->
eqmod c d ->
eqmod (
a +
c) (
b +
d).
Proof.
intros a b c d [k1 EQ1] [k2 EQ2]; red.
subst a; subst c. exists (k1 + k2). ring.
Qed.
Lemma eqmod_neg:
forall x y,
eqmod x y ->
eqmod (-
x) (-
y).
Proof.
intros x y [k EQ]; red. exists (-k). rewrite EQ. ring.
Qed.
Lemma eqmod_sub:
forall a b c d,
eqmod a b ->
eqmod c d ->
eqmod (
a -
c) (
b -
d).
Proof.
intros a b c d [k1 EQ1] [k2 EQ2]; red.
subst a; subst c. exists (k1 - k2). ring.
Qed.
Lemma eqmod_mult:
forall a b c d,
eqmod a c ->
eqmod b d ->
eqmod (
a *
b) (
c *
d).
Proof.
intros a b c d [
k1 EQ1] [
k2 EQ2];
red.
subst a;
subst b.
exists (
k1 *
k2 *
modul +
c *
k2 +
k1 *
d).
ring.
Qed.
End EQ_MODULO.
Lemma eqmod_divides:
forall n m x y,
eqmod n x y ->
Zdivide m n ->
eqmod m x y.
Proof.
intros.
destruct H as [
k1 EQ1].
destruct H0 as [
k2 EQ2].
exists (
k1*
k2).
rewrite <-
Zmult_assoc.
rewrite <-
EQ2.
auto.
Qed.
We then specialize these definitions to equality modulo
2wordsize.
Hint Resolve modulus_pos:
ints.
Definition eqm :=
eqmod modulus.
Lemma eqm_refl:
forall x,
eqm x x.
Proof (
eqmod_refl modulus).
Hint Resolve eqm_refl:
ints.
Lemma eqm_refl2:
forall x y,
x =
y ->
eqm x y.
Proof (
eqmod_refl2 modulus).
Hint Resolve eqm_refl2:
ints.
Lemma eqm_sym:
forall x y,
eqm x y ->
eqm y x.
Proof (
eqmod_sym modulus).
Hint Resolve eqm_sym:
ints.
Lemma eqm_trans:
forall x y z,
eqm x y ->
eqm y z ->
eqm x z.
Proof (
eqmod_trans modulus).
Hint Resolve eqm_trans:
ints.
Lemma eqm_small_eq:
forall x y,
eqm x y -> 0 <=
x <
modulus -> 0 <=
y <
modulus ->
x =
y.
Proof (
eqmod_small_eq modulus).
Hint Resolve eqm_small_eq:
ints.
Lemma eqm_add:
forall a b c d,
eqm a b ->
eqm c d ->
eqm (
a +
c) (
b +
d).
Proof (
eqmod_add modulus).
Hint Resolve eqm_add:
ints.
Lemma eqm_neg:
forall x y,
eqm x y ->
eqm (-
x) (-
y).
Proof (
eqmod_neg modulus).
Hint Resolve eqm_neg:
ints.
Lemma eqm_sub:
forall a b c d,
eqm a b ->
eqm c d ->
eqm (
a -
c) (
b -
d).
Proof (
eqmod_sub modulus).
Hint Resolve eqm_sub:
ints.
Lemma eqm_mult:
forall a b c d,
eqm a c ->
eqm b d ->
eqm (
a *
b) (
c *
d).
Proof (
eqmod_mult modulus).
Hint Resolve eqm_mult:
ints.
Properties of the coercions between Z and int
Lemma eqm_samerepr:
forall x y,
eqm x y ->
repr x =
repr y.
Proof.
Lemma eqm_unsigned_repr:
forall z,
eqm z (
unsigned (
repr z)).
Proof.
unfold eqm,
repr,
unsigned;
intros;
simpl.
apply eqmod_mod.
auto with ints.
Qed.
Hint Resolve eqm_unsigned_repr:
ints.
Lemma eqm_unsigned_repr_l:
forall a b,
eqm a b ->
eqm (
unsigned (
repr a))
b.
Proof.
Hint Resolve eqm_unsigned_repr_l:
ints.
Lemma eqm_unsigned_repr_r:
forall a b,
eqm a b ->
eqm a (
unsigned (
repr b)).
Proof.
Hint Resolve eqm_unsigned_repr_r:
ints.
Lemma eqm_signed_unsigned:
forall x,
eqm (
signed x) (
unsigned x).
Proof.
Theorem unsigned_range:
forall i, 0 <=
unsigned i <
modulus.
Proof.
destruct i; auto.
Qed.
Hint Resolve unsigned_range:
ints.
Theorem unsigned_range_2:
forall i, 0 <=
unsigned i <=
max_unsigned.
Proof.
intro;
unfold max_unsigned.
generalize (
unsigned_range i).
omega.
Qed.
Hint Resolve unsigned_range_2:
ints.
Theorem signed_range:
forall i,
min_signed <=
signed i <=
max_signed.
Proof.
Theorem repr_unsigned:
forall i,
repr (
unsigned i) =
i.
Proof.
Hint Resolve repr_unsigned:
ints.
Lemma repr_signed:
forall i,
repr (
signed i) =
i.
Proof.
Hint Resolve repr_signed:
ints.
Lemma eqm_repr_eq:
forall x y,
eqm x (
unsigned y) ->
repr x =
y.
Proof.
Theorem unsigned_repr:
forall z, 0 <=
z <=
max_unsigned ->
unsigned (
repr z) =
z.
Proof.
intros.
unfold repr,
unsigned;
simpl.
apply Zmod_small.
unfold max_unsigned in H.
fold modulus.
omega.
Qed.
Hint Resolve unsigned_repr:
ints.
Theorem signed_repr:
forall z,
min_signed <=
z <=
max_signed ->
signed (
repr z) =
z.
Proof.
Theorem signed_eq_unsigned:
forall x,
unsigned x <=
max_signed ->
signed x =
unsigned x.
Proof.
intros.
unfold signed.
destruct (
zlt (
unsigned x)
half_modulus).
auto.
unfold max_signed in H.
omegaContradiction.
Qed.
Theorem signed_positive:
forall x,
signed x >= 0 <->
unsigned x <=
max_signed.
Proof.
Properties of zero, one, minus one
Theorem unsigned_zero:
unsigned zero = 0.
Proof.
Theorem unsigned_one:
unsigned one = 1.
Proof.
Theorem unsigned_mone:
unsigned mone =
modulus - 1.
Proof.
Theorem signed_zero:
signed zero = 0.
Proof.
Theorem signed_mone:
signed mone = -1.
Proof.
Theorem one_not_zero:
one <>
zero.
Proof.
Theorem unsigned_repr_wordsize:
unsigned iwordsize =
Z_of_nat wordsize.
Proof.
Properties of equality
Theorem eq_sym:
forall x y,
eq x y =
eq y x.
Proof.
Theorem eq_spec:
forall (
x y:
int),
if eq x y then x =
y else x <>
y.
Proof.
intros;
unfold eq.
case (
eq_dec x y);
intro.
subst y.
rewrite zeq_true.
auto.
rewrite zeq_false.
auto.
destruct x;
destruct y.
simpl.
red;
intro.
elim n.
apply mkint_eq.
auto.
Qed.
Theorem eq_true:
forall x,
eq x x =
true.
Proof.
intros.
generalize (
eq_spec x x);
case (
eq x x);
intros;
congruence.
Qed.
Theorem eq_false:
forall x y,
x <>
y ->
eq x y =
false.
Proof.
intros.
generalize (
eq_spec x y);
case (
eq x y);
intros;
congruence.
Qed.
Properties of addition
Theorem add_unsigned:
forall x y,
add x y =
repr (
unsigned x +
unsigned y).
Proof.
intros; reflexivity.
Qed.
Theorem add_signed:
forall x y,
add x y =
repr (
signed x +
signed y).
Proof.
Theorem add_commut:
forall x y,
add x y =
add y x.
Proof.
intros; unfold add. decEq. omega. Qed.
Theorem add_zero:
forall x,
add x zero =
x.
Proof.
Theorem add_zero_l:
forall x,
add zero x =
x.
Proof.
Theorem add_assoc:
forall x y z,
add (
add x y)
z =
add x (
add y z).
Proof.
Theorem add_permut:
forall x y z,
add x (
add y z) =
add y (
add x z).
Proof.
Theorem add_neg_zero:
forall x,
add x (
neg x) =
zero.
Proof.
Theorem unsigned_add_carry:
forall x y,
unsigned (
add x y) =
unsigned x +
unsigned y -
unsigned (
add_carry x y zero) *
modulus.
Proof.
Corollary unsigned_add_either:
forall x y,
unsigned (
add x y) =
unsigned x +
unsigned y
\/
unsigned (
add x y) =
unsigned x +
unsigned y -
modulus.
Proof.
Properties of negation
Theorem neg_repr:
forall z,
neg (
repr z) =
repr (-
z).
Proof.
intros;
unfold neg.
apply eqm_samerepr.
auto with ints.
Qed.
Theorem neg_zero:
neg zero =
zero.
Proof.
Theorem neg_involutive:
forall x,
neg (
neg x) =
x.
Proof.
Theorem neg_add_distr:
forall x y,
neg(
add x y) =
add (
neg x) (
neg y).
Proof.
Properties of subtraction
Theorem sub_zero_l:
forall x,
sub x zero =
x.
Proof.
Theorem sub_zero_r:
forall x,
sub zero x =
neg x.
Proof.
Theorem sub_add_opp:
forall x y,
sub x y =
add x (
neg y).
Proof.
Theorem sub_idem:
forall x,
sub x x =
zero.
Proof.
intros; unfold sub. unfold zero. decEq. omega.
Qed.
Theorem sub_add_l:
forall x y z,
sub (
add x y)
z =
add (
sub x z)
y.
Proof.
Theorem sub_add_r:
forall x y z,
sub x (
add y z) =
add (
sub x z) (
neg y).
Proof.
Theorem sub_shifted:
forall x y z,
sub (
add x z) (
add y z) =
sub x y.
Proof.
Theorem sub_signed:
forall x y,
sub x y =
repr (
signed x -
signed y).
Proof.
Properties of multiplication
Theorem mul_commut:
forall x y,
mul x y =
mul y x.
Proof.
intros; unfold mul. decEq. ring.
Qed.
Theorem mul_zero:
forall x,
mul x zero =
zero.
Proof.
intros;
unfold mul.
rewrite unsigned_zero.
unfold zero.
decEq.
ring.
Qed.
Theorem mul_one:
forall x,
mul x one =
x.
Proof.
Theorem mul_mone:
forall x,
mul x mone =
neg x.
Proof.
Theorem mul_assoc:
forall x y z,
mul (
mul x y)
z =
mul x (
mul y z).
Proof.
Theorem mul_add_distr_l:
forall x y z,
mul (
add x y)
z =
add (
mul x z) (
mul y z).
Proof.
intros;
unfold mul,
add.
apply eqm_samerepr.
set (
x' :=
unsigned x).
set (
y' :=
unsigned y).
set (
z' :=
unsigned z).
apply eqm_trans with ((
x' +
y') *
z').
auto with ints.
replace ((
x' +
y') *
z')
with (
x' *
z' +
y' *
z').
auto with ints.
ring.
Qed.
Theorem mul_add_distr_r:
forall x y z,
mul x (
add y z) =
add (
mul x y) (
mul x z).
Proof.
Theorem neg_mul_distr_l:
forall x y,
neg(
mul x y) =
mul (
neg x)
y.
Proof.
intros.
unfold mul,
neg.
set (
x' :=
unsigned x).
set (
y' :=
unsigned y).
apply eqm_samerepr.
apply eqm_trans with (- (
x' *
y')).
auto with ints.
replace (- (
x' *
y'))
with ((-
x') *
y')
by ring.
auto with ints.
Qed.
Theorem neg_mul_distr_r:
forall x y,
neg(
mul x y) =
mul x (
neg y).
Proof.
Theorem mul_signed:
forall x y,
mul x y =
repr (
signed x *
signed y).
Proof.
Properties of division and modulus
Lemma modu_divu_Euclid:
forall x y,
y <>
zero ->
x =
add (
mul (
divu x y)
y) (
modu x y).
Proof.
Theorem modu_divu:
forall x y,
y <>
zero ->
modu x y =
sub x (
mul (
divu x y)
y).
Proof.
Theorem mods_divs:
forall x y,
mods x y =
sub x (
mul (
divs x y)
y).
Proof.
Theorem divu_one:
forall x,
divu x one =
x.
Proof.
Theorem modu_one:
forall x,
modu x one =
zero.
Proof.
Theorem divs_mone:
forall x,
divs x mone =
neg x.
Proof.
Theorem mods_mone:
forall x,
mods x mone =
zero.
Proof.
Properties of binary decompositions
Lemma Z_shift_add_bin_decomp:
forall x,
Z_shift_add (
fst (
Z_bin_decomp x)) (
snd (
Z_bin_decomp x)) =
x.
Proof.
destruct x;
simpl.
auto.
destruct p;
reflexivity.
destruct p;
try reflexivity.
simpl.
assert (
forall z, 2 * (
z + 1) - 1 = 2 *
z + 1).
intro;
omega.
generalize (
H (
Zpos p));
simpl.
congruence.
Qed.
Lemma Z_bin_decomp_shift_add:
forall b x,
Z_bin_decomp (
Z_shift_add b x) = (
b,
x).
Proof.
Lemma Z_shift_add_inj:
forall b1 x1 b2 x2,
Z_shift_add b1 x1 =
Z_shift_add b2 x2 ->
b1 =
b2 /\
x1 =
x2.
Proof.
intros.
assert ((
b1,
x1) = (
b2,
x2)).
repeat rewrite <-
Z_bin_decomp_shift_add.
rewrite H;
auto.
split;
congruence.
Qed.
Lemma Z_of_bits_exten:
forall f1 f2 n i1 i2,
(
forall i, 0 <=
i <
Z_of_nat n ->
f1 (
i +
i1) =
f2 (
i +
i2)) ->
Z_of_bits n f1 i1 =
Z_of_bits n f2 i2.
Proof.
induction n;
intros;
simpl.
auto.
rewrite inj_S in H.
decEq.
apply (
H 0).
omega.
apply IHn.
intros.
replace (
i +
Zsucc i1)
with (
Zsucc i +
i1)
by omega.
replace (
i +
Zsucc i2)
with (
Zsucc i +
i2)
by omega.
apply H.
omega.
Qed.
Lemma Z_of_bits_of_Z:
forall x,
eqm (
Z_of_bits wordsize (
bits_of_Z wordsize x) 0)
x.
Proof.
Lemma bits_of_Z_zero:
forall n x,
bits_of_Z n 0
x =
false.
Proof.
induction n;
simpl;
intros.
auto.
case (
zeq x 0);
intro.
auto.
auto.
Qed.
Remark Z_bin_decomp_2xm1:
forall x,
Z_bin_decomp (2 *
x - 1) = (
true,
x - 1).
Proof.
Lemma bits_of_Z_two_p:
forall n x i,
x >= 0 -> 0 <=
i <
Z_of_nat n ->
bits_of_Z n (
two_p x - 1)
i =
zlt i x.
Proof.
induction n;
intros.
simpl in H0.
omegaContradiction.
destruct (
zeq x 0).
subst x.
change (
two_p 0 - 1)
with 0.
rewrite bits_of_Z_zero.
unfold proj_sumbool;
rewrite zlt_false.
auto.
omega.
simpl.
replace (
two_p x)
with (2 *
two_p (
x - 1)).
rewrite Z_bin_decomp_2xm1.
destruct (
zeq i 0).
subst.
unfold proj_sumbool.
rewrite zlt_true.
auto.
omega.
rewrite inj_S in H0.
rewrite IHn.
unfold proj_sumbool.
destruct (
zlt i x).
apply zlt_true.
omega.
apply zlt_false.
omega.
omega.
omega.
rewrite <-
two_p_S.
decEq.
omega.
omega.
Qed.
Lemma bits_of_Z_mone:
forall x,
0 <=
x <
Z_of_nat wordsize ->
bits_of_Z wordsize (
modulus - 1)
x =
true.
Proof.
Lemma Z_of_bits_range:
forall f n i, 0 <=
Z_of_bits n f i <
two_power_nat n.
Proof.
Lemma Z_of_bits_range_1:
forall f i, 0 <=
Z_of_bits wordsize f i <
modulus.
Proof.
Hint Resolve Z_of_bits_range_1:
ints.
Lemma Z_of_bits_range_2:
forall f i, 0 <=
Z_of_bits wordsize f i <=
max_unsigned.
Proof.
Hint Resolve Z_of_bits_range_2:
ints.
Lemma bits_of_Z_of_bits_gen:
forall n f i j,
0 <=
i <
Z_of_nat n ->
bits_of_Z n (
Z_of_bits n f j)
i =
f (
i +
j).
Proof.
induction n;
intros;
simpl.
simpl in H.
omegaContradiction.
rewrite Z_bin_decomp_shift_add.
destruct (
zeq i 0).
f_equal.
omega.
rewrite IHn.
f_equal.
omega.
rewrite inj_S in H.
omega.
Qed.
Lemma bits_of_Z_of_bits:
forall n f i,
0 <=
i <
Z_of_nat n ->
bits_of_Z n (
Z_of_bits n f 0)
i =
f i.
Proof.
Lemma bits_of_Z_below:
forall n x i,
i < 0 ->
bits_of_Z n x i =
false.
Proof.
induction n;
intros;
simpl.
auto.
destruct (
Z_bin_decomp x).
rewrite zeq_false.
apply IHn.
omega.
omega.
Qed.
Lemma bits_of_Z_above:
forall n x i,
i >=
Z_of_nat n ->
bits_of_Z n x i =
false.
Proof.
induction n;
intros;
simpl.
auto.
caseEq (
Z_bin_decomp x);
intros b x1 EQ.
rewrite zeq_false.
rewrite IHn.
destruct x;
simpl in EQ.
inv EQ.
auto.
destruct p;
inv EQ;
auto.
destruct p;
inv EQ;
auto.
rewrite inj_S in H.
omega.
rewrite inj_S in H.
omega.
Qed.
Lemma bits_of_Z_greater:
forall n x i,
0 <=
x <
two_p i ->
bits_of_Z n x i =
false.
Proof.
induction n;
intros.
auto.
destruct (
zlt i 0).
apply bits_of_Z_below.
auto.
simpl.
destruct (
Z_bin_decomp x)
as [
b x1]
_eqn.
destruct (
zeq i 0).
subst i.
simpl in H.
assert (
x = 0)
by omega.
subst x.
simpl in Heqp.
congruence.
apply IHn.
rewrite <- (
Z_shift_add_bin_decomp x)
in H.
rewrite Heqp in H.
simpl in H.
replace i with (
Zsucc (
i-1))
in H by omega.
rewrite two_p_S in H.
unfold Z_shift_add in H.
destruct b;
omega.
omega.
Qed.
Lemma bits_of_Z_of_bits_gen':
forall n f i j,
bits_of_Z n (
Z_of_bits n f j)
i =
if zlt i 0
then false
else if zle (
Z_of_nat n)
i then false
else f (
i +
j).
Proof.
Lemma Z_of_bits_excl:
forall n f g h j,
(
forall i,
j <=
i <
j +
Z_of_nat n ->
f i &&
g i =
false) ->
(
forall i,
j <=
i <
j +
Z_of_nat n ->
f i ||
g i =
h i) ->
Z_of_bits n f j +
Z_of_bits n g j =
Z_of_bits n h j.
Proof.
induction n.
intros;
reflexivity.
intros.
simpl.
rewrite inj_S in H.
rewrite inj_S in H0.
rewrite <- (
IHn f g h (
Zsucc j)).
assert (
j <=
j <
j +
Zsucc(
Z_of_nat n)).
omega.
unfold Z_shift_add.
rewrite <-
H0;
auto.
caseEq (
f j);
intros;
caseEq (
g j);
intros;
simpl orb.
generalize (
H j H1).
rewrite H2;
rewrite H3.
simpl andb;
congruence.
omega.
omega.
omega.
intros;
apply H.
omega.
intros;
apply H0.
omega.
Qed.
Lemma Z_of_bits_compose:
forall f m n i,
Z_of_bits (
m +
n)
f i =
Z_of_bits n f (
i +
Z_of_nat m) *
two_power_nat m
+
Z_of_bits m f i.
Proof.
Lemma Z_of_bits_truncate:
forall f n i,
eqm (
Z_of_bits (
wordsize +
n)
f i) (
Z_of_bits wordsize f i).
Proof.
Lemma Z_of_bits_false:
forall f n i,
(
forall j,
i <=
j <
i +
Z_of_nat n ->
f j =
false) ->
Z_of_bits n f i = 0.
Proof.
induction n;
intros;
simpl.
auto.
rewrite inj_S in H.
rewrite H.
rewrite IHn.
auto.
intros;
apply H;
omega.
omega.
Qed.
Lemma Z_of_bits_complement:
forall f n i,
Z_of_bits n (
fun j =>
negb (
f j))
i =
two_power_nat n - 1 -
Z_of_bits n f i.
Proof.
Lemma Z_of_bits_true:
forall f n i,
(
forall j,
i <=
j <
i +
Z_of_nat n ->
f j =
true) ->
Z_of_bits n f i =
two_power_nat n - 1.
Proof.
Properties of bitwise and, or, xor
Lemma bitwise_binop_commut:
forall f,
(
forall a b,
f a b =
f b a) ->
forall x y,
bitwise_binop f x y =
bitwise_binop f y x.
Proof.
unfold bitwise_binop;
intros.
decEq;
apply Z_of_bits_exten;
intros.
auto.
Qed.
Lemma bitwise_binop_assoc:
forall f,
(
forall a b c,
f a (
f b c) =
f (
f a b)
c) ->
forall x y z,
bitwise_binop f (
bitwise_binop f x y)
z =
bitwise_binop f x (
bitwise_binop f y z).
Proof.
Lemma bitwise_binop_idem:
forall f,
(
forall a,
f a a =
a) ->
forall x,
bitwise_binop f x x =
x.
Proof.
Theorem and_commut:
forall x y,
and x y =
and y x.
Proof (
bitwise_binop_commut andb andb_comm).
Theorem and_assoc:
forall x y z,
and (
and x y)
z =
and x (
and y z).
Proof (
bitwise_binop_assoc andb andb_assoc).
Theorem and_zero:
forall x,
and x zero =
zero.
Proof.
Theorem and_mone:
forall x,
and x mone =
x.
Proof.
Theorem and_idem:
forall x,
and x x =
x.
Proof.
Theorem or_commut:
forall x y,
or x y =
or y x.
Proof (
bitwise_binop_commut orb orb_comm).
Theorem or_assoc:
forall x y z,
or (
or x y)
z =
or x (
or y z).
Proof (
bitwise_binop_assoc orb orb_assoc).
Theorem or_zero:
forall x,
or x zero =
x.
Proof.
Theorem or_mone:
forall x,
or x mone =
mone.
Proof.
Theorem or_idem:
forall x,
or x x =
x.
Proof.
Theorem and_or_distrib:
forall x y z,
and x (
or y z) =
or (
and x y) (
and x z).
Proof.
Theorem xor_commut:
forall x y,
xor x y =
xor y x.
Proof (
bitwise_binop_commut xorb xorb_comm).
Theorem xor_assoc:
forall x y z,
xor (
xor x y)
z =
xor x (
xor y z).
Proof.
Theorem xor_zero:
forall x,
xor x zero =
x.
Proof.
Theorem xor_idem:
forall x,
xor x x =
zero.
Proof.
Theorem xor_zero_one:
xor zero one =
one.
Proof.
Theorem xor_one_one:
xor one one =
zero.
Proof.
Theorem and_xor_distrib:
forall x y z,
and x (
xor y z) =
xor (
and x y) (
and x z).
Proof.
intros;
unfold and,
xor,
bitwise_binop.
repeat rewrite unsigned_repr;
auto with ints.
decEq;
apply Z_of_bits_exten;
intros.
repeat rewrite bits_of_Z_of_bits;
repeat rewrite Zplus_0_r;
auto.
assert (
forall a b c,
a && (
xorb b c) =
xorb (
a &&
b) (
a &&
c)).
destruct a;
destruct b;
destruct c;
reflexivity.
auto.
Qed.
Properties of bitwise complement.
Theorem not_involutive:
forall (
x:
int),
not (
not x) =
x.
Proof.
Theorem not_zero:
not zero =
mone.
Proof.
Theorem not_mone:
not mone =
zero.
Proof.
Theorem not_or_and_not:
forall x y,
not (
or x y) =
and (
not x) (
not y).
Proof.
Theorem not_and_or_not:
forall x y,
not (
and x y) =
or (
not x) (
not y).
Proof.
Theorem and_not_self:
forall x,
and x (
not x) =
zero.
Proof.
Theorem or_not_self:
forall x,
or x (
not x) =
mone.
Proof.
Theorem xor_not_self:
forall x,
xor x (
not x) =
mone.
Proof.
Theorem not_neg:
forall x,
not x =
add (
neg x)
mone.
Proof.
Theorem neg_not:
forall x,
neg x =
add (
not x)
one.
Proof.
Connections between add and bitwise logical operations.
Theorem add_is_or:
forall x y,
and x y =
zero ->
add x y =
or x y.
Proof.
Theorem xor_is_or:
forall x y,
and x y =
zero ->
xor x y =
or x y.
Proof.
Theorem add_is_xor:
forall x y,
and x y =
zero ->
add x y =
xor x y.
Proof.
Theorem add_and:
forall x y z,
and y z =
zero ->
add (
and x y) (
and x z) =
and x (
or y z).
Proof.
Properties of shifts
Theorem shl_zero:
forall x,
shl x zero =
x.
Proof.
Lemma bitwise_binop_shl:
forall f x y n,
f false false =
false ->
bitwise_binop f (
shl x n) (
shl y n) =
shl (
bitwise_binop f x y)
n.
Proof.
Theorem and_shl:
forall x y n,
and (
shl x n) (
shl y n) =
shl (
and x y)
n.
Proof.
Theorem or_shl:
forall x y n,
or (
shl x n) (
shl y n) =
shl (
or x y)
n.
Proof.
Theorem xor_shl:
forall x y n,
xor (
shl x n) (
shl y n) =
shl (
xor x y)
n.
Proof.
Lemma ltu_inv:
forall x y,
ltu x y =
true -> 0 <=
unsigned x <
unsigned y.
Proof.
Theorem shl_shl:
forall x y z,
ltu y iwordsize =
true ->
ltu z iwordsize =
true ->
ltu (
add y z)
iwordsize =
true ->
shl (
shl x y)
z =
shl x (
add y z).
Proof.
Theorem shru_zero:
forall x,
shru x zero =
x.
Proof.
Lemma bitwise_binop_shru:
forall f x y n,
f false false =
false ->
bitwise_binop f (
shru x n) (
shru y n) =
shru (
bitwise_binop f x y)
n.
Proof.
Theorem and_shru:
forall x y n,
and (
shru x n) (
shru y n) =
shru (
and x y)
n.
Proof.
Theorem or_shru:
forall x y n,
or (
shru x n) (
shru y n) =
shru (
or x y)
n.
Proof.
Theorem xor_shru:
forall x y n,
xor (
shru x n) (
shru y n) =
shru (
xor x y)
n.
Proof.
Theorem shru_shru:
forall x y z,
ltu y iwordsize =
true ->
ltu z iwordsize =
true ->
ltu (
add y z)
iwordsize =
true ->
shru (
shru x y)
z =
shru x (
add y z).
Proof.
Theorem shr_zero:
forall x,
shr x zero =
x.
Proof.
Lemma bitwise_binop_shr:
forall f x y n,
bitwise_binop f (
shr x n) (
shr y n) =
shr (
bitwise_binop f x y)
n.
Proof.
Theorem and_shr:
forall x y n,
and (
shr x n) (
shr y n) =
shr (
and x y)
n.
Proof.
Theorem or_shr:
forall x y n,
or (
shr x n) (
shr y n) =
shr (
or x y)
n.
Proof.
Theorem xor_shr:
forall x y n,
xor (
shr x n) (
shr y n) =
shr (
xor x y)
n.
Proof.
Theorem shr_shr:
forall x y z,
ltu y iwordsize =
true ->
ltu z iwordsize =
true ->
ltu (
add y z)
iwordsize =
true ->
shr (
shr x y)
z =
shr x (
add y z).
Proof.
Remark two_p_m1_range:
forall n,
0 <=
n <=
Z_of_nat wordsize ->
0 <=
two_p n - 1 <=
max_unsigned.
Proof.
Theorem shru_shl_and:
forall x y,
ltu y iwordsize =
true ->
shru (
shl x y)
y =
and x (
repr (
two_p (
Z_of_nat wordsize -
unsigned y) - 1)).
Proof.
Theorem and_shr_shru:
forall x y z,
and (
shr x z) (
shru y z) =
shru (
and x y)
z.
Proof.
Theorem shr_and_shru_and:
forall x y z,
shru (
shl z y)
y =
z ->
and (
shr x y)
z =
and (
shru x y)
z.
Proof.
Properties of rotations
Theorem shl_rolm:
forall x n,
ltu n iwordsize =
true ->
shl x n =
rolm x n (
shl mone n).
Proof.
Theorem shru_rolm:
forall x n,
ltu n iwordsize =
true ->
shru x n =
rolm x (
sub iwordsize n) (
shru mone n).
Proof.
Theorem rol_zero:
forall x,
rol x zero =
x.
Proof.
Lemma bitwise_binop_rol:
forall f x y n,
bitwise_binop f (
rol x n) (
rol y n) =
rol (
bitwise_binop f x y)
n.
Proof.
Theorem rol_and:
forall x y n,
rol (
and x y)
n =
and (
rol x n) (
rol y n).
Proof.
Theorem rol_or:
forall x y n,
rol (
or x y)
n =
or (
rol x n) (
rol y n).
Proof.
Theorem rol_xor:
forall x y n,
rol (
xor x y)
n =
xor (
rol x n) (
rol y n).
Proof.
Theorem rol_rol:
forall x n m,
Zdivide (
Z_of_nat wordsize)
modulus ->
rol (
rol x n)
m =
rol x (
modu (
add n m)
iwordsize).
Proof.
Theorem rolm_zero:
forall x m,
rolm x zero m =
and x m.
Proof.
intros.
unfold rolm.
rewrite rol_zero.
auto.
Qed.
Theorem rolm_rolm:
forall x n1 m1 n2 m2,
Zdivide (
Z_of_nat wordsize)
modulus ->
rolm (
rolm x n1 m1)
n2 m2 =
rolm x (
modu (
add n1 n2)
iwordsize)
(
and (
rol m1 n2)
m2).
Proof.
Theorem or_rolm:
forall x n m1 m2,
or (
rolm x n m1) (
rolm x n m2) =
rolm x n (
or m1 m2).
Proof.
Theorem ror_rol:
forall x y,
ltu y iwordsize =
true ->
ror x y =
rol x (
sub iwordsize y).
Proof.
Theorem or_ror:
forall x y z,
ltu y iwordsize =
true ->
ltu z iwordsize =
true ->
add y z =
iwordsize ->
ror x z =
or (
shl x y) (
shru x z).
Proof.
Properties of Z_one_bits and is_power2.
Fixpoint powerserie (
l:
list Z):
Z :=
match l with
|
nil => 0
|
x ::
xs =>
two_p x +
powerserie xs
end.
Lemma Z_bin_decomp_range:
forall x n,
0 <=
x < 2 *
n -> 0 <=
snd (
Z_bin_decomp x) <
n.
Proof.
Lemma Z_one_bits_powerserie:
forall x, 0 <=
x <
modulus ->
x =
powerserie (
Z_one_bits wordsize x 0).
Proof.
Lemma Z_one_bits_range:
forall x i,
In i (
Z_one_bits wordsize x 0) -> 0 <=
i <
Z_of_nat wordsize.
Proof.
assert (
forall n x i j,
In j (
Z_one_bits n x i) ->
i <=
j <
i +
Z_of_nat n).
induction n;
simpl In.
intros;
elim H.
intros x i j.
destruct (
Z_bin_decomp x).
case b.
rewrite inj_S.
simpl.
intros [
A|
B].
subst j.
omega.
generalize (
IHn _ _ _ B).
omega.
intros B.
rewrite inj_S.
generalize (
IHn _ _ _ B).
omega.
intros.
generalize (
H wordsize x 0
i H0).
omega.
Qed.
Lemma is_power2_rng:
forall n logn,
is_power2 n =
Some logn ->
0 <=
unsigned logn <
Z_of_nat wordsize.
Proof.
Theorem is_power2_range:
forall n logn,
is_power2 n =
Some logn ->
ltu logn iwordsize =
true.
Proof.
Lemma is_power2_correct:
forall n logn,
is_power2 n =
Some logn ->
unsigned n =
two_p (
unsigned logn).
Proof.
Remark two_p_range:
forall n,
0 <=
n <
Z_of_nat wordsize ->
0 <=
two_p n <=
max_unsigned.
Proof.
Remark Z_one_bits_zero:
forall n i,
Z_one_bits n 0
i =
nil.
Proof.
induction n; intros; simpl; auto.
Qed.
Remark Z_one_bits_two_p:
forall n x i,
0 <=
x <
Z_of_nat n ->
Z_one_bits n (
two_p x)
i = (
i +
x) ::
nil.
Proof.
induction n;
intros;
simpl.
simpl in H.
omegaContradiction.
rewrite inj_S in H.
assert (
x = 0 \/ 0 <
x)
by omega.
destruct H0.
subst x;
simpl.
decEq.
omega.
apply Z_one_bits_zero.
replace (
two_p x)
with (
Z_shift_add false (
two_p (
x-1))).
rewrite Z_bin_decomp_shift_add.
replace (
i +
x)
with ((
i + 1) + (
x - 1))
by omega.
apply IHn.
omega.
unfold Z_shift_add.
rewrite <-
two_p_S.
decEq;
omega.
omega.
Qed.
Lemma is_power2_two_p:
forall n, 0 <=
n <
Z_of_nat wordsize ->
is_power2 (
repr (
two_p n)) =
Some (
repr n).
Proof.
Relation between bitwise operations and multiplications / divisions by powers of 2
Left shifts and multiplications by powers of 2.
Lemma Z_of_bits_shift_left:
forall f m,
m >= 0 ->
(
forall i,
i < 0 ->
f i =
false) ->
eqm (
Z_of_bits wordsize f (-
m)) (
two_p m *
Z_of_bits wordsize f 0).
Proof.
Lemma shl_mul_two_p:
forall x y,
shl x y =
mul x (
repr (
two_p (
unsigned y))).
Proof.
Theorem shl_mul:
forall x y,
shl x y =
mul x (
shl one y).
Proof.
Theorem mul_pow2:
forall x n logn,
is_power2 n =
Some logn ->
mul x n =
shl x logn.
Proof.
Theorem shifted_or_is_add:
forall x y n,
0 <=
n <
Z_of_nat wordsize ->
unsigned y <
two_p n ->
or (
shl x (
repr n))
y =
repr(
unsigned x *
two_p n +
unsigned y).
Proof.
Unsigned right shifts and unsigned divisions by powers of 2.
Lemma Z_of_bits_shift_right:
forall m f,
m >= 0 ->
(
forall i,
i >=
Z_of_nat wordsize ->
f i =
false) ->
exists k,
Z_of_bits wordsize f 0 =
k +
two_p m *
Z_of_bits wordsize f m /\ 0 <=
k <
two_p m.
Proof.
Lemma shru_div_two_p:
forall x y,
shru x y =
repr (
unsigned x /
two_p (
unsigned y)).
Proof.
Theorem divu_pow2:
forall x n logn,
is_power2 n =
Some logn ->
divu x n =
shru x logn.
Proof.
Signed right shifts and signed divisions by powers of 2.
Lemma Z_of_bits_shift_right_neg:
forall m f,
m >= 0 ->
(
forall i,
i >=
Z_of_nat wordsize ->
f i =
true) ->
exists k,
Z_of_bits wordsize f 0 -
modulus =
k +
two_p m * (
Z_of_bits wordsize f m -
modulus)
/\ 0 <=
k <
two_p m.
Proof.
Lemma sign_bit_of_Z_rec:
forall n x,
0 <=
x <
two_power_nat (
S n) ->
bits_of_Z (
S n)
x (
Z_of_nat n) =
if zlt x (
two_power_nat n)
then false else true.
Proof.
Lemma sign_bit_of_Z:
forall x,
bits_of_Z wordsize (
unsigned x) (
Z_of_nat wordsize - 1) =
if zlt (
unsigned x)
half_modulus then false else true.
Proof.
Lemma shr_div_two_p:
forall x y,
shr x y =
repr (
signed x /
two_p (
unsigned y)).
Proof.
Theorem divs_pow2:
forall x n logn,
is_power2 n =
Some logn ->
divs x n =
shrx x logn.
Proof.
Unsigned modulus over 2^n is masking with 2^n-1.
Lemma Z_of_bits_mod_mask:
forall f n,
0 <=
n <=
Z_of_nat wordsize ->
Z_of_bits wordsize (
fun i =>
if zlt i n then f i else false) 0 =
(
Z_of_bits wordsize f 0)
mod (
two_p n).
Proof.
Theorem modu_and:
forall x n logn,
is_power2 n =
Some logn ->
modu x n =
and x (
sub n one).
Proof.
Properties of shrx (signed division by a power of 2)
Lemma Zdiv_round_Zdiv:
forall x y,
y > 0 ->
Zdiv_round x y =
if zlt x 0
then (
x +
y - 1) /
y else x /
y.
Proof.
intros.
unfold Zdiv_round.
destruct (
zlt x 0).
rewrite zlt_false;
try omega.
generalize (
Z_div_mod_eq (-
x)
y H).
generalize (
Z_mod_lt (-
x)
y H).
set (
q := (-
x) /
y).
set (
r := (-
x)
mod y).
intros.
symmetry.
apply Zdiv_unique with (
y -
r - 1).
replace x with (- (
y *
q) -
r)
by omega.
replace (-(
y *
q))
with ((-
q) *
y)
by ring.
omega.
omega.
apply zlt_false.
omega.
Qed.
Theorem shrx_shr:
forall x y,
ltu y (
repr (
Z_of_nat wordsize - 1)) =
true ->
shrx x y =
shr (
if lt x zero then add x (
sub (
shl one y)
one)
else x)
y.
Proof.
Lemma Zdiv_shift:
forall x y,
y > 0 ->
(
x + (
y - 1)) /
y =
x /
y +
if zeq (
Zmod x y) 0
then 0
else 1.
Proof.
intros.
generalize (
Z_div_mod_eq x y H).
generalize (
Z_mod_lt x y H).
set (
q :=
x /
y).
set (
r :=
x mod y).
intros.
destruct (
zeq r 0).
apply Zdiv_unique with (
y - 1).
rewrite H1.
rewrite e.
ring.
omega.
apply Zdiv_unique with (
r - 1).
rewrite H1.
ring.
omega.
Qed.
Theorem shrx_carry:
forall x y,
ltu y (
repr (
Z_of_nat wordsize - 1)) =
true ->
shrx x y =
add (
shr x y) (
shr_carry x y).
Proof.
Connections between shr and shru.
Lemma shr_shru_positive:
forall x y,
signed x >= 0 ->
shr x y =
shru x y.
Proof.
Lemma and_positive:
forall x y,
signed y >= 0 ->
signed (
and x y) >= 0.
Proof.
Theorem shr_and_is_shru_and:
forall x y z,
lt y zero =
false ->
shr (
and x y)
z =
shru (
and x y)
z.
Proof.
Properties of integer zero extension and sign extension.
Section EXTENSIONS.
Variable n:
Z.
Hypothesis RANGE: 0 <
n <
Z_of_nat wordsize.
Remark two_p_n_pos:
two_p n > 0.
Proof.
Remark two_p_n_range:
0 <=
two_p n <=
max_unsigned.
Proof.
Remark two_p_n_range':
two_p n <=
max_signed + 1.
Proof.
Remark unsigned_repr_two_p:
unsigned (
repr (
two_p n)) =
two_p n.
Proof.
Remark eqm_eqmod_two_p:
forall a b,
eqm a b ->
eqmod (
two_p n)
a b.
Proof.
Theorem zero_ext_and:
forall x,
zero_ext n x =
and x (
repr (
two_p n - 1)).
Proof.
Theorem zero_ext_mod:
forall x,
unsigned (
zero_ext n x) =
Zmod (
unsigned x) (
two_p n).
Proof.
Theorem zero_ext_idem:
forall x,
zero_ext n (
zero_ext n x) =
zero_ext n x.
Proof.
Theorem sign_ext_idem:
forall x,
sign_ext n (
sign_ext n x) =
sign_ext n x.
Proof.
Theorem sign_ext_zero_ext:
forall x,
sign_ext n (
zero_ext n x) =
sign_ext n x.
Proof.
Theorem zero_ext_sign_ext:
forall x,
zero_ext n (
sign_ext n x) =
zero_ext n x.
Proof.
Theorem sign_ext_equal_if_zero_equal:
forall x y,
zero_ext n x =
zero_ext n y ->
sign_ext n x =
sign_ext n y.
Proof.
Theorem zero_ext_shru_shl:
forall x,
let y :=
repr (
Z_of_nat wordsize -
n)
in
zero_ext n x =
shru (
shl x y)
y.
Proof.
Theorem sign_ext_shr_shl:
forall x,
let y :=
repr (
Z_of_nat wordsize -
n)
in
sign_ext n x =
shr (
shl x y)
y.
Proof.
zero_ext n x is the unique integer congruent to x modulo 2^n
in the range 0...2^n-1.
Lemma zero_ext_range:
forall x, 0 <=
unsigned (
zero_ext n x) <
two_p n.
Proof.
Lemma eqmod_zero_ext:
forall x,
eqmod (
two_p n) (
unsigned (
zero_ext n x)) (
unsigned x).
Proof.
sign_ext n x is the unique integer congruent to x modulo 2^n
in the range -2^(n-1)...2^(n-1) - 1.
Lemma sign_ext_div:
forall x,
signed (
sign_ext n x) =
signed (
repr (
unsigned x *
two_p (
Z_of_nat wordsize -
n))) /
two_p (
Z_of_nat wordsize -
n).
Proof.
Lemma sign_ext_range:
forall x, -
two_p (
n-1) <=
signed (
sign_ext n x) <
two_p (
n-1).
Proof.
Lemma eqmod_sign_ext:
forall x,
eqmod (
two_p n) (
signed (
sign_ext n x)) (
unsigned x).
Proof.
Lemma eqmod_sign_ext':
forall x,
eqmod (
two_p n) (
unsigned (
sign_ext n x)) (
unsigned x).
Proof.
End EXTENSIONS.
Theorem zero_ext_widen:
forall x n n',
0 <
n <
Z_of_nat wordsize ->
n <=
n' <
Z_of_nat wordsize ->
zero_ext n' (
zero_ext n x) =
zero_ext n x.
Proof.
Theorem sign_ext_widen:
forall x n n',
0 <
n <
Z_of_nat wordsize ->
n <=
n' <
Z_of_nat wordsize ->
sign_ext n' (
sign_ext n x) =
sign_ext n x.
Proof.
Theorem sign_zero_ext_widen:
forall x n n',
0 <
n <
Z_of_nat wordsize ->
n <
n' <
Z_of_nat wordsize ->
sign_ext n' (
zero_ext n x) =
zero_ext n x.
Proof.
Theorem zero_ext_narrow:
forall x n n',
0 <
n <
Z_of_nat wordsize ->
n <=
n' <
Z_of_nat wordsize ->
zero_ext n (
zero_ext n'
x) =
zero_ext n x.
Proof.
Theorem zero_sign_ext_narrow:
forall x n n',
0 <
n <
Z_of_nat wordsize ->
n <=
n' <
Z_of_nat wordsize ->
zero_ext n (
sign_ext n'
x) =
zero_ext n x.
Proof.
Properties of one_bits (decomposition in sum of powers of two)
Theorem one_bits_range:
forall x i,
In i (
one_bits x) ->
ltu i iwordsize =
true.
Proof.
Fixpoint int_of_one_bits (
l:
list int) :
int :=
match l with
|
nil =>
zero
|
a ::
b =>
add (
shl one a) (
int_of_one_bits b)
end.
Theorem one_bits_decomp:
forall x,
x =
int_of_one_bits (
one_bits x).
Proof.
Properties of comparisons
Theorem negate_cmp:
forall c x y,
cmp (
negate_comparison c)
x y =
negb (
cmp c x y).
Proof.
intros.
destruct c;
simpl;
try rewrite negb_elim;
auto.
Qed.
Theorem negate_cmpu:
forall c x y,
cmpu (
negate_comparison c)
x y =
negb (
cmpu c x y).
Proof.
intros.
destruct c;
simpl;
try rewrite negb_elim;
auto.
Qed.
Theorem swap_cmp:
forall c x y,
cmp (
swap_comparison c)
x y =
cmp c y x.
Proof.
intros.
destruct c;
simpl;
auto.
apply eq_sym.
decEq.
apply eq_sym.
Qed.
Theorem swap_cmpu:
forall c x y,
cmpu (
swap_comparison c)
x y =
cmpu c y x.
Proof.
intros.
destruct c;
simpl;
auto.
apply eq_sym.
decEq.
apply eq_sym.
Qed.
Lemma translate_eq:
forall x y d,
eq (
add x d) (
add y d) =
eq x y.
Proof.
Lemma translate_ltu:
forall x y d,
0 <=
unsigned x +
unsigned d <=
max_unsigned ->
0 <=
unsigned y +
unsigned d <=
max_unsigned ->
ltu (
add x d) (
add y d) =
ltu x y.
Proof.
Theorem translate_cmpu:
forall c x y d,
0 <=
unsigned x +
unsigned d <=
max_unsigned ->
0 <=
unsigned y +
unsigned d <=
max_unsigned ->
cmpu c (
add x d) (
add y d) =
cmpu c x y.
Proof.
Lemma translate_lt:
forall x y d,
min_signed <=
signed x +
signed d <=
max_signed ->
min_signed <=
signed y +
signed d <=
max_signed ->
lt (
add x d) (
add y d) =
lt x y.
Proof.
Theorem translate_cmp:
forall c x y d,
min_signed <=
signed x +
signed d <=
max_signed ->
min_signed <=
signed y +
signed d <=
max_signed ->
cmp c (
add x d) (
add y d) =
cmp c x y.
Proof.
Theorem notbool_isfalse_istrue:
forall x,
is_false x ->
is_true (
notbool x).
Proof.
Theorem notbool_istrue_isfalse:
forall x,
is_true x ->
is_false (
notbool x).
Proof.
unfold is_false,
is_true,
notbool;
intros.
generalize (
eq_spec x zero).
case (
eq x zero);
intro.
contradiction.
auto.
Qed.
Theorem shru_lt_zero:
forall x,
shru x (
repr (
Z_of_nat wordsize - 1)) =
if lt x zero then one else zero.
Proof.
Theorem ltu_range_test:
forall x y,
ltu x y =
true ->
unsigned y <=
max_signed ->
0 <=
signed x <
unsigned y.
Proof.
Non-overlapping test
Definition no_overlap (
ofs1:
int) (
sz1:
Z) (
ofs2:
int) (
sz2:
Z) :
bool :=
let x1 :=
unsigned ofs1 in let x2 :=
unsigned ofs2 in
zlt (
x1 +
sz1)
modulus &&
zlt (
x2 +
sz2)
modulus
&& (
zle (
x1 +
sz1)
x2 ||
zle (
x2 +
sz2)
x1).
Lemma no_overlap_sound:
forall ofs1 sz1 ofs2 sz2 base,
sz1 > 0 ->
sz2 > 0 ->
no_overlap ofs1 sz1 ofs2 sz2 =
true ->
unsigned (
add base ofs1) +
sz1 <=
unsigned (
add base ofs2)
\/
unsigned (
add base ofs2) +
sz2 <=
unsigned (
add base ofs1).
Proof.
End Make.
Specialization to integers of size 8, 32, and 64 bits
Module Wordsize_32.
Definition wordsize := 32%
nat.
Remark wordsize_not_zero:
wordsize <> 0%
nat.
Proof.
unfold wordsize; congruence. Qed.
End Wordsize_32.
Module Int :=
Make(
Wordsize_32).
Notation int :=
Int.int.
Remark int_wordsize_divides_modulus:
Zdivide (
Z_of_nat Int.wordsize)
Int.modulus.
Proof.
exists (
two_p (32-5));
reflexivity.
Qed.
Module Wordsize_8.
Definition wordsize := 8%
nat.
Remark wordsize_not_zero:
wordsize <> 0%
nat.
Proof.
unfold wordsize; congruence. Qed.
End Wordsize_8.
Module Byte :=
Make(
Wordsize_8).
Notation byte :=
Byte.int.
Module Wordsize_64.
Definition wordsize := 64%
nat.
Remark wordsize_not_zero:
wordsize <> 0%
nat.
Proof.
unfold wordsize; congruence. Qed.
End Wordsize_64.
Module Int64 :=
Make(
Wordsize_64).
Notation int64 :=
Int64.int.