Library Coq.ZArith.Zquot
This file provides results about the Round-Toward-Zero Euclidean
division Z.quotrem, whose projections are Z.quot (noted ÷)
and Z.rem.
This division and Z.div agree only on positive numbers.
Otherwise, Z.div performs Round-Toward-Bottom (a.k.a Floor).
This Z.quot is compatible with the division of usual
programming languages such as Ocaml. In addition, it has nicer
properties with respect to opposite and other usual operations.
The definition of this division is now in file BinIntDef,
while most of the results about here are now in the main module
BinInt.Z, thanks to the generic "Numbers" layer. Remain here:
- some compatibility notation for old names.
- some extra results with less preconditions (in particular exploiting the arbitrary value of division by 0).
Notation Ndiv_Zquot := N2Z.inj_quot (only parsing).
Notation Nmod_Zrem := N2Z.inj_rem (only parsing).
Notation Z_quot_rem_eq := Z.quot_rem' (only parsing).
Notation Zrem_lt := Z.rem_bound_abs (only parsing).
Notation Z_quot_mult := Z.quot_mul (only parsing).
Particular values taken for a÷0 and (Z.rem a 0).
We avise to not rely on these arbitrary values.
The following results are expressed without the b<>0 condition
whenever possible.
Lemma Zrem_0_l a : Z.rem 0 a = 0.
Lemma Zquot_0_l a : 0÷a = 0.
#[global]
Hint Resolve Zrem_0_l Zrem_0_r Zquot_0_l Zquot_0_r Z.quot_1_r Z.rem_1_r
: zarith.
Ltac zero_or_not a :=
destruct (Z.eq_decidable a 0) as [->|?];
[rewrite ?Zquot_0_l, ?Zrem_0_l, ?Zquot_0_r, ?Zrem_0_r;
try lia;
auto with zarith|].
Lemma Z_rem_same a : Z.rem a a = 0.
Lemma Z_rem_mult a b : Z.rem (a*b) b = 0.
Theorem Zquot_opp_l a b : (-a)÷b = -(a÷b).
Theorem Zquot_opp_r a b : a÷(-b) = -(a÷b).
Theorem Zrem_opp_l a b : Z.rem (-a) b = -(Z.rem a b).
Theorem Zrem_opp_r a b : Z.rem a (-b) = Z.rem a b.
Theorem Zquot_opp_opp a b : (-a)÷(-b) = a÷b.
Theorem Zrem_opp_opp a b : Z.rem (-a) (-b) = -(Z.rem a b).
The sign of the remainder is the one of a. Due to the possible
nullity of a, a general result is to be stated in the following form:
This can also be said in a simpler way:
Reformulation of Z.rem_bound_abs in 2 then 4 particular cases.
Theorem Zrem_lt_pos a b : 0<=a -> b<>0 -> 0 <= Z.rem a b < Z.abs b.
Theorem Zrem_lt_neg a b : a<=0 -> b<>0 -> -Z.abs b < Z.rem a b <= 0.
Theorem Zrem_lt_pos_pos a b : 0<=a -> 0<b -> 0 <= Z.rem a b < b.
Theorem Zrem_lt_pos_neg a b : 0<=a -> b<0 -> 0 <= Z.rem a b < -b.
Theorem Zrem_lt_neg_pos a b : a<=0 -> 0<b -> -b < Z.rem a b <= 0.
Theorem Zrem_lt_neg_neg a b : a<=0 -> b<0 -> b < Z.rem a b <= 0.
Definition Remainder a b r :=
(0 <= a /\ 0 <= r < Z.abs b) \/ (a <= 0 /\ -Z.abs b < r <= 0).
Definition Remainder_alt a b r :=
Z.abs r < Z.abs b /\ 0 <= r * a.
Lemma Remainder_equiv : forall a b r,
Remainder a b r <-> Remainder_alt a b r.
Theorem Zquot_mod_unique_full a b q r :
Remainder a b r -> a = b*q + r -> q = a÷b /\ r = Z.rem a b.
Theorem Zquot_unique_full a b q r :
Remainder a b r -> a = b*q + r -> q = a÷b.
Theorem Zrem_unique_full a b q r :
Remainder a b r -> a = b*q + r -> r = Z.rem a b.
As soon as the divisor is greater or equal than 2,
the division is strictly decreasing.
<= is compatible with a positive division.
With our choice of division, rounding of (a÷b) is always done toward 0:
Lemma Z_mult_quot_le a b : 0 <= a -> 0 <= b*(a÷b) <= a.
Lemma Z_mult_quot_ge a b : a <= 0 -> a <= b*(a÷b) <= 0.
The previous inequalities between b*(a÷b) and a are exact
iff the modulo is zero.
A modulo cannot grow beyond its starting point.
Some additional inequalities about Zdiv.
Theorem Zquot_le_upper_bound:
forall a b q, 0 < b -> a <= q*b -> a÷b <= q.
Theorem Zquot_lt_upper_bound:
forall a b q, 0 <= a -> 0 < b -> a < q*b -> a÷b < q.
Theorem Zquot_le_lower_bound:
forall a b q, 0 < b -> q*b <= a -> q <= a÷b.
Theorem Zquot_sgn: forall a b,
0 <= Z.sgn (a÷b) * Z.sgn a * Z.sgn b.
Relations between usual operations and Z.modulo and Z.div
Lemma Z_rem_plus : forall a b c:Z,
0 <= (a+b*c) * a ->
Z.rem (a + b * c) c = Z.rem a c.
Lemma Z_quot_plus : forall a b c:Z,
0 <= (a+b*c) * a -> c<>0 ->
(a + b * c) ÷ c = a ÷ c + b.
Theorem Z_quot_plus_l: forall a b c : Z,
0 <= (a*b+c)*c -> b<>0 ->
b<>0 -> (a * b + c) ÷ b = a + c ÷ b.
Cancellations.
Lemma Zquot_mult_cancel_r : forall a b c:Z,
c<>0 -> (a*c)÷(b*c) = a÷b.
Lemma Zquot_mult_cancel_l : forall a b c:Z,
c<>0 -> (c*a)÷(c*b) = a÷b.
Lemma Zmult_rem_distr_l: forall a b c,
Z.rem (c*a) (c*b) = c * (Z.rem a b).
Lemma Zmult_rem_distr_r: forall a b c,
Z.rem (a*c) (b*c) = (Z.rem a b) * c.
Operations modulo.
Theorem Zrem_rem: forall a n, Z.rem (Z.rem a n) n = Z.rem a n.
Theorem Zmult_rem: forall a b n,
Z.rem (a * b) n = Z.rem (Z.rem a n * Z.rem b n) n.
addition and modulo
Generally speaking, unlike with Zdiv, we don't have
(a+b) rem n = (a rem n + b rem n) rem n
for any a and b.
For instance, take (8 + (-10)) rem 3 = -2 whereas
(8 rem 3 + (-10 rem 3)) rem 3 = 1.
Theorem Zplus_rem: forall a b n,
0 <= a * b ->
Z.rem (a + b) n = Z.rem (Z.rem a n + Z.rem b n) n.
Lemma Zplus_rem_idemp_l: forall a b n,
0 <= a * b ->
Z.rem (Z.rem a n + b) n = Z.rem (a + b) n.
Lemma Zplus_rem_idemp_r: forall a b n,
0 <= a*b ->
Z.rem (b + Z.rem a n) n = Z.rem (b + a) n.
Lemma Zmult_rem_idemp_l: forall a b n, Z.rem (Z.rem a n * b) n = Z.rem (a * b) n.
Lemma Zmult_rem_idemp_r: forall a b n, Z.rem (b * Z.rem a n) n = Z.rem (b * a) n.
Unlike with Zdiv, the following result is true without restrictions.
A last inequality:
Z.rem is related to divisibility (see more in Znumtheory)
Particular case : dividing by 2 is related with parity
Lemma Zquot2_odd_remainder : forall a,
Remainder a 2 (if Z.odd a then Z.sgn a else 0).
Lemma Zrem_odd : forall a, Z.rem a 2 = if Z.odd a then Z.sgn a else 0.
Lemma Zrem_even : forall a, Z.rem a 2 = if Z.even a then 0 else Z.sgn a.
Lemma Zeven_rem : forall a, Z.even a = Z.eqb (Z.rem a 2) 0.
Lemma Zodd_rem : forall a, Z.odd a = negb (Z.eqb (Z.rem a 2) 0).
Theorem Zquotrem_Zdiv_eucl_pos : forall a b:Z, 0 <= a -> 0 < b ->
a÷b = a/b /\ Z.rem a b = a mod b.
Theorem Zquot_Zdiv_pos : forall a b, 0 <= a -> 0 <= b ->
a÷b = a/b.
Theorem Zrem_Zmod_pos : forall a b, 0 <= a -> 0 < b ->
Z.rem a b = a mod b.
Modulos are null at the same places