Library Coq.Numbers.Integer.Abstract.ZDivFloor


Require Import ZAxioms ZMulOrder ZSgnAbs NZDiv.

Euclidean Division for integers (Floor convention)

We use here the convention known as Floor, or Round-Toward-Bottom, where a/b is the closest integer below the exact fraction. It can be summarized by:
a = bq+r /\ 0 <= |r| < |b| /\ Sign(r) = Sign(b)
This is the convention followed historically by Z.div in Coq, and corresponds to convention "F" in the following paper:
R. Boute, "The Euclidean definition of the functions div and mod", ACM Transactions on Programming Languages and Systems, Vol. 14, No.2, pp. 127-144, April 1992.
See files ZDivTrunc and ZDivEucl for others conventions.

Module Type ZDivProp
 (Import A : ZAxiomsSig')
 (Import B : ZMulOrderProp A)
 (Import C : ZSgnAbsProp A B).

We benefit from what already exists for NZ
Module Import Private_NZDiv := Nop <+ NZDivProp A A B.

Another formulation of the main equation

Lemma mod_eq :
 forall a b, b~=0 -> a mod b == a - b*(a/b).

We have a general bound for absolute values

Lemma mod_bound_abs :
 forall a b, b~=0 -> abs (a mod b) < abs b.

Uniqueness theorems

Theorem div_mod_unique : forall b q1 q2 r1 r2 : t,
  (0<=r1<b \/ b<r1<=0) -> (0<=r2<b \/ b<r2<=0) ->
  b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.

Theorem div_unique:
 forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> q == a/b.

Theorem div_unique_pos:
 forall a b q r, 0<=r<b -> a == b*q + r -> q == a/b.

Theorem div_unique_neg:
 forall a b q r, b<r<=0 -> a == b*q + r -> q == a/b.

Theorem mod_unique:
 forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> r == a mod b.

Theorem mod_unique_pos:
 forall a b q r, 0<=r<b -> a == b*q + r -> r == a mod b.

Theorem mod_unique_neg:
 forall a b q r, b<r<=0 -> a == b*q + r -> r == a mod b.

Sign rules

Ltac pos_or_neg a :=
 let LT := fresh "LT" in
 let LE := fresh "LE" in
 destruct (le_gt_cases 0 a) as [LE|LT]; [|rewrite <- opp_pos_neg in LT].

Fact mod_bound_or : forall a b, b~=0 -> 0<=a mod b<b \/ b<a mod b<=0.

Fact opp_mod_bound_or : forall a b, b~=0 ->
 0 <= -(a mod b) < -b \/ -b < -(a mod b) <= 0.

Lemma div_opp_opp : forall a b, b~=0 -> -a/-b == a/b.

Lemma mod_opp_opp : forall a b, b~=0 -> (-a) mod (-b) == - (a mod b).

With the current conventions, the other sign rules are rather complex.

Lemma div_opp_l_z :
 forall a b, b~=0 -> a mod b == 0 -> (-a)/b == -(a/b).

Lemma div_opp_l_nz :
 forall a b, b~=0 -> a mod b ~= 0 -> (-a)/b == -(a/b)-1.

Lemma mod_opp_l_z :
 forall a b, b~=0 -> a mod b == 0 -> (-a) mod b == 0.

Lemma mod_opp_l_nz :
 forall a b, b~=0 -> a mod b ~= 0 -> (-a) mod b == b - a mod b.

Lemma div_opp_r_z :
 forall a b, b~=0 -> a mod b == 0 -> a/(-b) == -(a/b).

Lemma div_opp_r_nz :
 forall a b, b~=0 -> a mod b ~= 0 -> a/(-b) == -(a/b)-1.

Lemma mod_opp_r_z :
 forall a b, b~=0 -> a mod b == 0 -> a mod (-b) == 0.

Lemma mod_opp_r_nz :
 forall a b, b~=0 -> a mod b ~= 0 -> a mod (-b) == (a mod b) - b.

The sign of a mod b is the one of b (when it isn't null)

Lemma mod_sign_nz : forall a b, b~=0 -> a mod b ~= 0 ->
 sgn (a mod b) == sgn b.

Lemma mod_sign : forall a b, b~=0 -> sgn (a mod b) ~= -sgn b.

Lemma mod_sign_mul : forall a b, b~=0 -> 0 <= (a mod b) * b.

A division by itself returns 1

Lemma div_same : forall a, a~=0 -> a/a == 1.

Lemma mod_same : forall a, a~=0 -> a mod a == 0.

A division of a small number by a bigger one yields zero.

Theorem div_small: forall a b, 0<=a<b -> a/b == 0.

Same situation, in term of modulo:

Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.

Basic values of divisions and modulo.


Lemma div_0_l: forall a, a~=0 -> 0/a == 0.

Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.

Lemma div_1_r: forall a, a/1 == a.

Lemma mod_1_r: forall a, a mod 1 == 0.

Lemma div_1_l: forall a, 1<a -> 1/a == 0.

Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.

Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.

Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.

Theorem div_unique_exact a b q: b~=0 -> a == b*q -> q == a/b.

Order results about mod and div

A modulo cannot grow beyond its starting point.

Theorem mod_le: forall a b, 0<=a -> 0<b -> a mod b <= a.

Theorem div_pos : forall a b, 0<=a -> 0<b -> 0<= a/b.

Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.

Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> 0<=a<b \/ b<a<=0).

Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> 0<=a<b \/ b<a<=0).

As soon as the divisor is strictly greater than 1, the division is strictly decreasing.

Lemma div_lt : forall a b, 0<a -> 1<b -> a/b < a.

le is compatible with a positive division.

Lemma div_le_mono : forall a b c, 0<c -> a<=b -> a/c <= b/c.

In this convention, div performs Rounding-Toward-Bottom.
Since we cannot speak of rational values here, we express this fact by multiplying back by b, and this leads to separates statements according to the sign of b.
First, a/b is below the exact fraction ...

Lemma mul_div_le : forall a b, 0<b -> b*(a/b) <= a.

Lemma mul_div_ge : forall a b, b<0 -> a <= b*(a/b).

... and moreover it is the larger such integer, since S(a/b) is strictly above the exact fraction.

Lemma mul_succ_div_gt: forall a b, 0<b -> a < b*(S (a/b)).

Lemma mul_succ_div_lt: forall a b, b<0 -> b*(S (a/b)) < a.

NB: The four previous properties could be used as specifications for div.
Inequality mul_div_le is exact iff the modulo is zero.

Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).

Some additional inequalities about div.

Theorem div_lt_upper_bound:
  forall a b q, 0<b -> a < b*q -> a/b < q.

Theorem div_le_upper_bound:
  forall a b q, 0<b -> a <= b*q -> a/b <= q.

Theorem div_le_lower_bound:
  forall a b q, 0<b -> b*q <= a -> q <= a/b.

A division respects opposite monotonicity for the divisor

Lemma div_le_compat_l: forall p q r, 0<=p -> 0<q<=r -> p/r <= p/q.

Relations between usual operations and mod and div


Lemma mod_add : forall a b c, c~=0 ->
 (a + b * c) mod c == a mod c.

Lemma div_add : forall a b c, c~=0 ->
 (a + b * c) / c == a / c + b.

Lemma div_add_l: forall a b c, b~=0 ->
 (a * b + c) / b == a + c / b.

Cancellations.

Lemma div_mul_cancel_r : forall a b c, b~=0 -> c~=0 ->
 (a*c)/(b*c) == a/b.

Lemma div_mul_cancel_l : forall a b c, b~=0 -> c~=0 ->
 (c*a)/(c*b) == a/b.

Lemma mul_mod_distr_l: forall a b c, b~=0 -> c~=0 ->
  (c*a) mod (c*b) == c * (a mod b).

Lemma mul_mod_distr_r: forall a b c, b~=0 -> c~=0 ->
  (a*c) mod (b*c) == (a mod b) * c.

Operations modulo.

Theorem mod_mod: forall a n, n~=0 ->
 (a mod n) mod n == a mod n.

Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
 ((a mod n)*b) mod n == (a*b) mod n.

Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
 (a*(b mod n)) mod n == (a*b) mod n.

Theorem mul_mod: forall a b n, n~=0 ->
 (a * b) mod n == ((a mod n) * (b mod n)) mod n.

Lemma add_mod_idemp_l : forall a b n, n~=0 ->
 ((a mod n)+b) mod n == (a+b) mod n.

Lemma add_mod_idemp_r : forall a b n, n~=0 ->
 (a+(b mod n)) mod n == (a+b) mod n.

Theorem add_mod: forall a b n, n~=0 ->
 (a+b) mod n == (a mod n + b mod n) mod n.

With the current convention, the following result isn't always true with a negative last divisor. For instance 3/(-2)/(-2) = 1 <> 0 = 3 / (-2*-2) , or 5/2/(-2) = -1 <> -2 = 5 / (2*-2) .

Lemma div_div : forall a b c, b~=0 -> 0<c ->
 (a/b)/c == a/(b*c).

Similarly, the following result doesn't always hold when c<0. For instance 3 mod (-2*-2)) = 3 while 3 mod (-2) + (-2)*((3/-2) mod -2) = -1.

Lemma rem_mul_r : forall a b c, b~=0 -> 0<c ->
 a mod (b*c) == a mod b + b*((a/b) mod c).

Lemma mod_div: forall a b, b~=0 ->
 a mod b / b == 0.

Lemma add_mul_mod_distr_l : forall a b c d, 0<=a -> 0<b -> 0<=d<c ->
 (c*a+d) mod (c*b) == c*(a mod b)+d.

Lemma add_mul_mod_distr_r: forall a b c d, 0<=a -> 0<b -> 0<=d<c ->
 (a*c+d) mod (b*c) == (a mod b)*c+d.

A last inequality:

Theorem div_mul_le:
 forall a b c, 0<=a -> 0<b -> 0<=c -> c*(a/b) <= (c*a)/b.

End ZDivProp.