Additional operations and proofs about binary integers,
on top of the ZArith standard library.
Require Import Psatz Zquot.
Require Import Coqlib.
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. lia.
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.
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 ->
Z.divide m n ->
eqmod m x y.
Proof.
intros.
destruct H as [
k1 EQ1].
destruct H0 as [
k2 EQ2].
exists (
k1*
k2).
rewrite <-
Z.mul_assoc.
rewrite <-
EQ2.
auto.
Qed.
Fast normalization modulo 2^n
Fixpoint P_mod_two_p (
p:
positive) (
n:
nat) {
struct n} :
Z :=
match n with
|
O => 0
|
S m =>
match p with
|
xH => 1
|
xO q =>
Z.double (
P_mod_two_p q m)
|
xI q =>
Z.succ_double (
P_mod_two_p q m)
end
end.
Definition Z_mod_two_p (
x:
Z) (
n:
nat) :
Z :=
match x with
|
Z0 => 0
|
Zpos p =>
P_mod_two_p p n
|
Zneg p =>
let r :=
P_mod_two_p p n in if zeq r 0
then 0
else two_power_nat n -
r
end.
Lemma P_mod_two_p_range:
forall n p, 0 <=
P_mod_two_p p n <
two_power_nat n.
Proof.
Lemma P_mod_two_p_eq:
forall n p,
P_mod_two_p p n = (
Zpos p)
mod (
two_power_nat n).
Proof.
Lemma Z_mod_two_p_range:
forall n x, 0 <=
Z_mod_two_p x n <
two_power_nat n.
Proof.
Lemma Z_mod_two_p_eq:
forall n x,
Z_mod_two_p x n =
x mod (
two_power_nat n).
Proof.
Bit-level operations and properties
Shift x left by one and insert b as the low bit of the result.
Definition Zshiftin (
b:
bool) (
x:
Z) :
Z :=
if b then Z.succ_double x else Z.double x.
Remark Ztestbit_0:
forall n,
Z.testbit 0
n =
false.
Proof Z.testbit_0_l.
Remark Ztestbit_1:
forall n,
Z.testbit 1
n =
zeq n 0.
Proof.
intros. destruct n; simpl; auto.
Qed.
Remark Ztestbit_m1:
forall n, 0 <=
n ->
Z.testbit (-1)
n =
true.
Proof.
intros. destruct n; simpl; auto.
Qed.
Remark Zshiftin_spec:
forall b x,
Zshiftin b x = 2 *
x + (
if b then 1
else 0).
Proof.
Remark Zshiftin_inj:
forall b1 x1 b2 x2,
Zshiftin b1 x1 =
Zshiftin b2 x2 ->
b1 =
b2 /\
x1 =
x2.
Proof.
intros.
rewrite !
Zshiftin_spec in H.
destruct b1;
destruct b2.
split; [
auto|
lia].
extlia.
extlia.
split; [
auto|
lia].
Qed.
Remark Zdecomp:
forall x,
x =
Zshiftin (
Z.odd x) (
Z.div2 x).
Proof.
intros.
destruct x;
simpl.
-
auto.
-
destruct p;
auto.
-
destruct p;
auto.
simpl.
rewrite Pos.pred_double_succ.
auto.
Qed.
Remark Ztestbit_shiftin:
forall b x n,
0 <=
n ->
Z.testbit (
Zshiftin b x)
n =
if zeq n 0
then b else Z.testbit x (
Z.pred n).
Proof.
Remark Ztestbit_shiftin_base:
forall b x,
Z.testbit (
Zshiftin b x) 0 =
b.
Proof.
Remark Ztestbit_shiftin_succ:
forall b x n, 0 <=
n ->
Z.testbit (
Zshiftin b x) (
Z.succ n) =
Z.testbit x n.
Proof.
Lemma Zshiftin_ind:
forall (
P:
Z ->
Prop),
P 0 ->
(
forall b x, 0 <=
x ->
P x ->
P (
Zshiftin b x)) ->
forall x, 0 <=
x ->
P x.
Proof.
Lemma Zshiftin_pos_ind:
forall (
P:
Z ->
Prop),
P 1 ->
(
forall b x, 0 <
x ->
P x ->
P (
Zshiftin b x)) ->
forall x, 0 <
x ->
P x.
Proof.
Bit-wise decomposition (Z.testbit)
Remark Ztestbit_eq:
forall n x, 0 <=
n ->
Z.testbit x n =
if zeq n 0
then Z.odd x else Z.testbit (
Z.div2 x) (
Z.pred n).
Proof.
Remark Ztestbit_base:
forall x,
Z.testbit x 0 =
Z.odd x.
Proof.
Remark Ztestbit_succ:
forall n x, 0 <=
n ->
Z.testbit x (
Z.succ n) =
Z.testbit (
Z.div2 x)
n.
Proof.
Lemma eqmod_same_bits:
forall n x y,
(
forall i, 0 <=
i <
Z.of_nat n ->
Z.testbit x i =
Z.testbit y i) ->
eqmod (
two_power_nat n)
x y.
Proof.
Lemma same_bits_eqmod:
forall n x y i,
eqmod (
two_power_nat n)
x y -> 0 <=
i <
Z.of_nat n ->
Z.testbit x i =
Z.testbit y i.
Proof.
Lemma equal_same_bits:
forall x y,
(
forall i, 0 <=
i ->
Z.testbit x i =
Z.testbit y i) ->
x =
y.
Proof Z.bits_inj'.
Lemma Z_one_complement:
forall i, 0 <=
i ->
forall x,
Z.testbit (-
x-1)
i =
negb (
Z.testbit x i).
Proof.
Lemma Ztestbit_above:
forall n x i,
0 <=
x <
two_power_nat n ->
i >=
Z.of_nat n ->
Z.testbit x i =
false.
Proof.
Lemma Ztestbit_above_neg:
forall n x i,
-
two_power_nat n <=
x < 0 ->
i >=
Z.of_nat n ->
Z.testbit x i =
true.
Proof.
Lemma Zsign_bit:
forall n x,
0 <=
x <
two_power_nat (
S n) ->
Z.testbit x (
Z.of_nat n) =
if zlt x (
two_power_nat n)
then false else true.
Proof.
Lemma Ztestbit_le:
forall x y,
0 <=
y ->
(
forall i, 0 <=
i ->
Z.testbit x i =
true ->
Z.testbit y i =
true) ->
x <=
y.
Proof.
Lemma Ztestbit_mod_two_p:
forall n x i,
0 <=
n -> 0 <=
i ->
Z.testbit (
x mod (
two_p n))
i =
if zlt i n then Z.testbit x i else false.
Proof.
Corollary Ztestbit_two_p_m1:
forall n i, 0 <=
n -> 0 <=
i ->
Z.testbit (
two_p n - 1)
i =
if zlt i n then true else false.
Proof.
Corollary Ztestbit_neg_two_p:
forall n i, 0 <=
n -> 0 <=
i ->
Z.testbit (- (
two_p n))
i =
if zlt i n then false else true.
Proof.
Lemma Z_add_is_or:
forall i, 0 <=
i ->
forall x y,
(
forall j, 0 <=
j <=
i ->
Z.testbit x j &&
Z.testbit y j =
false) ->
Z.testbit (
x +
y)
i =
Z.testbit x i ||
Z.testbit y i.
Proof.
Zero and sign extensions
In pseudo-code:
Fixpoint Zzero_ext (n: Z) (x: Z) : Z :=
if zle n 0 then
0
else
Zshiftin (Z.odd x) (Zzero_ext (Z.pred n) (Z.div2 x)).
Fixpoint Zsign_ext (n: Z) (x: Z) : Z :=
if zle n 1 then
if Z.odd x then -1 else 0
else
Zshiftin (Z.odd x) (Zzero_ext (Z.pred n) (Z.div2 x)).
We encode this
nat-like recursion using the
Z.iter iteration
function, in order to make the
Zzero_ext and
Zsign_ext
functions efficiently executable within Coq.
Definition Zzero_ext (
n:
Z) (
x:
Z) :
Z :=
Z.iter n
(
fun rec x =>
Zshiftin (
Z.odd x) (
rec (
Z.div2 x)))
(
fun x => 0)
x.
Definition Zsign_ext (
n:
Z) (
x:
Z) :
Z :=
Z.iter (
Z.pred n)
(
fun rec x =>
Zshiftin (
Z.odd x) (
rec (
Z.div2 x)))
(
fun x =>
if Z.odd x &&
zlt 0
n then -1
else 0)
x.
Lemma Ziter_base:
forall (
A:
Type)
n (
f:
A ->
A)
x,
n <= 0 ->
Z.iter n f x =
x.
Proof.
intros.
unfold Z.iter.
destruct n;
auto.
compute in H.
elim H;
auto.
Qed.
Lemma Ziter_succ:
forall (
A:
Type)
n (
f:
A ->
A)
x,
0 <=
n ->
Z.iter (
Z.succ n)
f x =
f (
Z.iter n f x).
Proof.
Lemma Znatlike_ind:
forall (
P:
Z ->
Prop),
(
forall n,
n <= 0 ->
P n) ->
(
forall n, 0 <=
n ->
P n ->
P (
Z.succ n)) ->
forall n,
P n.
Proof.
intros.
destruct (
zle 0
n).
apply natlike_ind;
auto.
apply H;
lia.
apply H.
lia.
Qed.
Lemma Zzero_ext_spec:
forall n x i, 0 <=
i ->
Z.testbit (
Zzero_ext n x)
i =
if zlt i n then Z.testbit x i else false.
Proof.
Lemma Zsign_ext_spec:
forall n x i, 0 <=
i ->
Z.testbit (
Zsign_ext n x)
i =
Z.testbit x (
if zlt i n then i else n - 1).
Proof.
Zzero_ext n x is x modulo 2^n
Lemma Zzero_ext_mod:
forall n x, 0 <=
n ->
Zzero_ext n x =
x mod (
two_p n).
Proof.
Zzero_ext n x is the unique integer congruent to x modulo 2^n in the range 0...2^n-1.
Lemma Zzero_ext_range:
forall n x, 0 <=
n -> 0 <=
Zzero_ext n x <
two_p n.
Proof.
Lemma eqmod_Zzero_ext:
forall n x, 0 <=
n ->
eqmod (
two_p n) (
Zzero_ext n x)
x.
Proof.
Relation between Zsign_ext n x and (Zzero_ext n x]
Lemma Zsign_ext_zero_ext:
forall n, 0 <=
n ->
forall x,
Zsign_ext n x =
Zzero_ext n x - (
if Z.testbit x (
n - 1)
then two_p n else 0).
Proof.
Zsign_ext n x is the unique integer congruent to x modulo 2^n
in the range -2^(n-1)...2^(n-1) - 1.
Lemma Zsign_ext_range:
forall n x, 0 <
n -> -
two_p (
n-1) <=
Zsign_ext n x <
two_p (
n-1).
Proof.
Lemma eqmod_Zsign_ext:
forall n x, 0 <=
n ->
eqmod (
two_p n) (
Zsign_ext n x)
x.
Proof.
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 =>
if Z.odd x
then i ::
Z_one_bits m (
Z.div2 x) (
i+1)
else Z_one_bits m (
Z.div2 x) (
i+1)
end.
Fixpoint powerserie (
l:
list Z):
Z :=
match l with
|
nil => 0
|
x ::
xs =>
two_p x +
powerserie xs
end.
Lemma Z_one_bits_powerserie:
forall n x, 0 <=
x <
two_power_nat n ->
x =
powerserie (
Z_one_bits n x 0).
Proof.
Lemma Z_one_bits_range:
forall n x i,
In i (
Z_one_bits n x 0) -> 0 <=
i <
Z.of_nat n.
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.
tauto.
intros x i j.
rewrite Nat2Z.inj_succ.
assert (
In j (
Z_one_bits n (
Z.div2 x) (
i + 1)) ->
i <=
j <
i +
Z.succ (
Z.of_nat n)).
intros.
exploit IHn;
eauto.
lia.
destruct (
Z.odd x);
simpl.
intros [
A|
B].
subst j.
lia.
auto.
auto.
}
intros.
generalize (
H n x 0
i H0).
lia.
Qed.
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.
Recognition of powers of two
Fixpoint P_is_power2 (
p:
positive) :
bool :=
match p with
|
xH =>
true
|
xO q =>
P_is_power2 q
|
xI q =>
false
end.
Definition Z_is_power2 (
x:
Z) :
option Z :=
match x with
|
Z0 =>
None
|
Zpos p =>
if P_is_power2 p then Some (
Z.log2 x)
else None
|
Zneg _ =>
None
end.
Remark P_is_power2_sound:
forall p,
P_is_power2 p =
true ->
Z.pos p =
two_p (
Z.log2 (
Z.pos p)).
Proof.
Lemma Z_is_power2_nonneg:
forall x i,
Z_is_power2 x =
Some i -> 0 <=
i.
Proof.
Lemma Z_is_power2_sound:
forall x i,
Z_is_power2 x =
Some i ->
x =
two_p i /\
i =
Z.log2 x.
Proof.
Corollary Z_is_power2_range:
forall n x i,
0 <=
n -> 0 <=
x <
two_p n ->
Z_is_power2 x =
Some i -> 0 <=
i <
n.
Proof.
Lemma Z_is_power2_complete:
forall i, 0 <=
i ->
Z_is_power2 (
two_p i) =
Some i.
Proof.
Definition Z_is_power2m1 (
x:
Z) :
option Z :=
Z_is_power2 (
Z.succ x).
Lemma Z_is_power2m1_nonneg:
forall x i,
Z_is_power2m1 x =
Some i -> 0 <=
i.
Proof.
Lemma Z_is_power2m1_sound:
forall x i,
Z_is_power2m1 x =
Some i ->
x =
two_p i - 1.
Proof.
Lemma Z_is_power2m1_complete:
forall i, 0 <=
i ->
Z_is_power2m1 (
two_p i - 1) =
Some i.
Proof.
Lemma Z_is_power2m1_range:
forall n x i,
0 <=
n -> 0 <=
x <
two_p n ->
Z_is_power2m1 x =
Some i -> 0 <=
i <=
n.
Proof.
Relation between bitwise operations and multiplications / divisions by powers of 2
Left shifts and multiplications by powers of 2.
Lemma Zshiftl_mul_two_p:
forall x n, 0 <=
n ->
Z.shiftl x n =
x *
two_p n.
Proof.
Right shifts and divisions by powers of 2.
Lemma Zshiftr_div_two_p:
forall x n, 0 <=
n ->
Z.shiftr x n =
x /
two_p n.
Proof.
Properties of shrx (signed division by a power of 2)
Lemma Zquot_Zdiv:
forall x y,
y > 0 ->
Z.quot x y =
if zlt x 0
then (
x +
y - 1) /
y else x /
y.
Proof.
intros.
destruct (
zlt x 0).
-
symmetry.
apply Zquot_unique_full with ((
x +
y - 1)
mod y - (
y - 1)).
+
red.
right;
split.
lia.
exploit (
Z_mod_lt (
x +
y - 1)
y);
auto.
rewrite Z.abs_eq.
lia.
lia.
+
transitivity ((
y * ((
x +
y - 1) /
y) + (
x +
y - 1)
mod y) - (
y-1)).
rewrite <-
Z_div_mod_eq.
ring.
auto.
ring.
-
apply Zquot_Zdiv_pos;
lia.
Qed.
Lemma Zdiv_shift:
forall x y,
y > 0 ->
(
x + (
y - 1)) /
y =
x /
y +
if zeq (
Z.modulo 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.
lia.
apply Zdiv_unique with (
r - 1).
rewrite H1.
ring.
lia.
Qed.
Size of integers, in bits.
Definition Zsize (
x:
Z) :
Z :=
match x with
|
Zpos p =>
Zpos (
Pos.size p)
|
_ => 0
end.
Remark Zsize_pos:
forall x, 0 <=
Zsize x.
Proof.
destruct x; simpl. lia. compute; intuition congruence. lia.
Qed.
Remark Zsize_pos':
forall x, 0 <
x -> 0 <
Zsize x.
Proof.
destruct x; simpl; intros; try discriminate. compute; auto.
Qed.
Lemma Zsize_shiftin:
forall b x, 0 <
x ->
Zsize (
Zshiftin b x) =
Z.succ (
Zsize x).
Proof.
Lemma Ztestbit_size_1:
forall x, 0 <
x ->
Z.testbit x (
Z.pred (
Zsize x)) =
true.
Proof.
Lemma Ztestbit_size_2:
forall x, 0 <=
x ->
forall i,
i >=
Zsize x ->
Z.testbit x i =
false.
Proof.
Lemma Zsize_interval_1:
forall x, 0 <=
x -> 0 <=
x <
two_p (
Zsize x).
Proof.
Lemma Zsize_interval_2:
forall x n, 0 <=
n -> 0 <=
x <
two_p n ->
n >=
Zsize x.
Proof.
Lemma Zsize_monotone:
forall x y, 0 <=
x <=
y ->
Zsize x <=
Zsize y.
Proof.
Bit insertion, bit extraction
Extract and optionally sign-extend bits from...from+len-1 of x
Definition Zextract_u (
x:
Z) (
from:
Z) (
len:
Z) :
Z :=
Zzero_ext len (
Z.shiftr x from).
Definition Zextract_s (
x:
Z) (
from:
Z) (
len:
Z) :
Z :=
Zsign_ext len (
Z.shiftr x from).
Lemma Zextract_u_spec:
forall x from len i,
0 <=
from -> 0 <=
len -> 0 <=
i ->
Z.testbit (
Zextract_u x from len)
i =
if zlt i len then Z.testbit x (
from +
i)
else false.
Proof.
Lemma Zextract_s_spec:
forall x from len i,
0 <=
from -> 0 <
len -> 0 <=
i ->
Z.testbit (
Zextract_s x from len)
i =
Z.testbit x (
from + (
if zlt i len then i else len - 1)).
Proof.
Insert bits 0...len-1 of y into bits to...to+len-1 of x
Definition Zinsert (
x y:
Z) (
to:
Z) (
len:
Z) :
Z :=
let mask :=
Z.shiftl (
two_p len - 1)
to in
Z.lor (
Z.land (
Z.shiftl y to)
mask) (
Z.ldiff x mask).
Lemma Zinsert_spec:
forall x y to len i,
0 <=
to -> 0 <=
len -> 0 <=
i ->
Z.testbit (
Zinsert x y to len)
i =
if zle to i &&
zlt i (
to +
len)
then Z.testbit y (
i -
to)
else Z.testbit x i.
Proof.