Library Coq.Numbers.NatInt.NZDiv
Euclidean Division 
Require Import NZAxioms NZMulOrder.
The first signatures will be common to all divisions over NZ, N and Z 
Module Type DivMod (Import A : Typ).
Parameters Inline div modulo : t -> t -> t.
End DivMod.
Module Type DivModNotation (A : Typ)(Import B : DivMod A).
Infix "/" := div.
Infix "mod" := modulo (at level 40, no associativity).
End DivModNotation.
Module Type DivMod' (A : Typ) := DivMod A <+ DivModNotation A.
Module Type NZDivSpec (Import A : NZOrdAxiomsSig')(Import B : DivMod' A).
Declare Instance div_wd : Proper (eq==>eq==>eq) div.
Declare Instance mod_wd : Proper (eq==>eq==>eq) modulo.
Axiom div_mod : forall a b, b ~= 0 -> a == b*(a/b) + (a mod b).
Axiom mod_bound_pos : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
End NZDivSpec.
The different divisions will only differ in the conditions
    they impose on modulo. For NZ, we have only described the
    behavior on positive numbers.
Module Type NZDiv (A : NZOrdAxiomsSig) := DivMod A <+ NZDivSpec A.
Module Type NZDiv' (A : NZOrdAxiomsSig) := NZDiv A <+ DivModNotation A.
Module Type NZDivProp
(Import A : NZOrdAxiomsSig')
(Import B : NZDiv' A)
(Import C : NZMulOrderProp A).
Uniqueness theorems 
Theorem div_mod_unique :
forall b q1 q2 r1 r2, 0<=r1<b -> 0<=r2<b ->
b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.
Theorem div_unique:
forall a b q r, 0<=a -> 0<=r<b ->
a == b*q + r -> q == a/b.
Theorem mod_unique:
forall a b q r, 0<=a -> 0<=r<b ->
a == b*q + r -> r == a mod b.
Theorem div_unique_exact a b q:
0<=a -> 0<b -> a == b*q -> q == a/b.
A division by itself returns 1 
Lemma div_same : forall a, 0<a -> a/a == 1.
Lemma mod_same : forall a, 0<a -> 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.
Lemma div_0_l: forall a, 0<a -> 0/a == 0.
Lemma mod_0_l: forall a, 0<a -> 0 mod a == 0.
Lemma div_1_r: forall a, 0<=a -> a/1 == a.
Lemma mod_1_r: forall a, 0<=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, 0<=a -> 0<b -> (a*b)/b == a.
Lemma mod_mul : forall a b, 0<=a -> 0<b -> (a*b) mod b == 0.
Theorem mod_le: forall a b, 0<=a -> 0<b -> a mod b <= a.
Lemma 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, 0<=a -> 0<b -> (a/b==0 <-> a<b).
Lemma mod_small_iff : forall a b, 0<=a -> 0<b -> (a mod b == a <-> a<b).
Lemma div_str_pos_iff : forall a b, 0<=a -> 0<b -> (0<a/b <-> b<=a).
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 -> 0<=a<=b -> a/c <= b/c.
The following two properties could be used as specification of div 
Lemma mul_div_le : forall a b, 0<=a -> 0<b -> b*(a/b) <= a.
Lemma mul_succ_div_gt : forall a b, 0<=a -> 0<b -> a < b*(S (a/b)).
The previous inequality is exact iff the modulo is zero. 
Lemma div_exact : forall a b, 0<=a -> 0<b -> (a == b*(a/b) <-> a mod b == 0).
Some additional inequalities about div. 
Theorem div_lt_upper_bound:
forall a b q, 0<=a -> 0<b -> a < b*q -> a/b < q.
Theorem div_le_upper_bound:
forall a b q, 0<=a -> 0<b -> a <= b*q -> a/b <= q.
Theorem div_le_lower_bound:
forall a b q, 0<=a -> 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.
Lemma mod_add : forall a b c, 0<=a -> 0<=a+b*c -> 0<c ->
(a + b * c) mod c == a mod c.
Lemma div_add : forall a b c, 0<=a -> 0<=a+b*c -> 0<c ->
(a + b * c) / c == a / c + b.
Lemma div_add_l: forall a b c, 0<=c -> 0<=a*b+c -> 0<b ->
(a * b + c) / b == a + c / b.
Cancellations. 
Lemma div_mul_cancel_r : forall a b c, 0<=a -> 0<b -> 0<c ->
(a*c)/(b*c) == a/b.
Lemma div_mul_cancel_l : forall a b c, 0<=a -> 0<b -> 0<c ->
(c*a)/(c*b) == a/b.
Lemma mul_mod_distr_l: forall a b c, 0<=a -> 0<b -> 0<c ->
(c*a) mod (c*b) == c * (a mod b).
Lemma mul_mod_distr_r: forall a b c, 0<=a -> 0<b -> 0<c ->
(a*c) mod (b*c) == (a mod b) * c.
Operations modulo. 
Theorem mod_mod: forall a n, 0<=a -> 0<n ->
(a mod n) mod n == a mod n.
Lemma mul_mod_idemp_l : forall a b n, 0<=a -> 0<=b -> 0<n ->
((a mod n)*b) mod n == (a*b) mod n.
Lemma mul_mod_idemp_r : forall a b n, 0<=a -> 0<=b -> 0<n ->
(a*(b mod n)) mod n == (a*b) mod n.
Theorem mul_mod: forall a b n, 0<=a -> 0<=b -> 0<n ->
(a * b) mod n == ((a mod n) * (b mod n)) mod n.
Lemma add_mod_idemp_l : forall a b n, 0<=a -> 0<=b -> 0<n ->
((a mod n)+b) mod n == (a+b) mod n.
Lemma add_mod_idemp_r : forall a b n, 0<=a -> 0<=b -> 0<n ->
(a+(b mod n)) mod n == (a+b) mod n.
Theorem add_mod: forall a b n, 0<=a -> 0<=b -> 0<n ->
(a+b) mod n == (a mod n + b mod n) mod n.
Lemma div_div : forall a b c, 0<=a -> 0<b -> 0<c ->
(a/b)/c == a/(b*c).
Lemma mod_mul_r : forall a b c, 0<=a -> 0<b -> 0<c ->
a mod (b*c) == a mod b + b*((a/b) mod c).
A last inequality: 
Theorem div_mul_le:
forall a b c, 0<=a -> 0<b -> 0<=c -> c*(a/b) <= (c*a)/b.
mod is related to divisibility 
Lemma mod_divides : forall a b, 0<=a -> 0<b ->
(a mod b == 0 <-> exists c, a == b*c).
End NZDivProp.
