aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Numbers/Integer/Abstract
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-12-15 18:20:03 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-12-15 18:20:03 +0000
commit4bf2fe115c9ea22d9e2b4d3bb392de2d4cf23adc (patch)
treef2c40e96395bc921bbcf4f7b29f2fdf92c63a266 /theories/Numbers/Integer/Abstract
parent5976fd4370daefbe8eb597af64968f499ad94d46 (diff)
A generic euclidean division in Numbers (Still Work-In-Progress)
- For Z, we propose 3 conventions for the sign of the remainder... - Instanciation for nat in NPeano. - Beginning of instanciation in ZOdiv. Still many proofs to finish, etc, etc, but soon we will have a decent properties database for all divisions of all instances of Numbers (e.g. BigZ). git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12590 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories/Numbers/Integer/Abstract')
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivCoq.v390
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivMath.v396
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivOcaml.v541
-rw-r--r--theories/Numbers/Integer/Abstract/ZMulOrder.v6
4 files changed, 1327 insertions, 6 deletions
diff --git a/theories/Numbers/Integer/Abstract/ZDivCoq.v b/theories/Numbers/Integer/Abstract/ZDivCoq.v
new file mode 100644
index 000000000..402d520d5
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZDivCoq.v
@@ -0,0 +1,390 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Euclidean Division for integers
+
+ We use here the historical convention of Coq :
+ [a = bq+r /\ 0 < |r| < |b| /\ Sign(r) = Sgn(b)]
+ *)
+
+Require Import ZAxioms ZProperties NZDiv.
+
+Open Scope NumScope.
+
+Module Type ZDiv (Import Z : ZAxiomsSig).
+
+ Parameter Inline div : t -> t -> t.
+ Parameter Inline modulo : t -> t -> t.
+
+ Infix "/" := div : NumScope.
+ Infix "mod" := modulo (at level 40, no associativity) : NumScope.
+
+ Instance div_wd : Proper (eq==>eq==>eq) div.
+ 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_pos_bound : forall a b, 0 < b -> 0 <= a mod b < b.
+ Axiom mod_neg_bound : forall a b, b < 0 -> b < a mod b <= 0.
+
+End ZDiv.
+
+Module Type ZDivSig := ZAxiomsSig <+ ZDiv.
+
+Module ZDivPropFunct (Import Z : ZDivSig).
+ (* TODO: en faire un arg du foncteur + comprendre le bug de SearchAbout *)
+ Module Import ZP := ZPropFunct Z.
+
+(** We benefit from what already exists for NZ *)
+
+ Module Z' <: NZDivSig.
+ Include Z.
+ Lemma mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+ Proof. intros; apply mod_pos_bound; auto. Qed.
+ End Z'.
+ Module Import NZDivP := NZDivPropFunct Z'.
+
+(** Another formulation of the main equation *)
+
+Lemma mod_eq :
+ forall a b, b~=0 -> a mod b == a - b*(a/b).
+Proof.
+intros.
+rewrite <- add_move_l.
+symmetry. apply div_mod; auto.
+Qed.
+
+(** A few sign rules (simple ones) *)
+
+Lemma div_mod_opp_opp : forall a b, b~=0 ->
+ (-a/-b) == a/b /\ (-a) mod (-b) == -(a mod b).
+Proof.
+intros a b Hb.
+assert (-b ~= 0).
+ contradict Hb. rewrite eq_opp_l, opp_0 in Hb; auto.
+assert (EQ := opp_involutive a).
+rewrite (div_mod a b) in EQ at 2; auto.
+rewrite (div_mod (-a) (-b)) in EQ; auto.
+
+destruct (lt_ge_cases 0 b).
+rewrite opp_add_distr in EQ.
+rewrite <- mul_opp_l, opp_involutive in EQ.
+destruct (div_mod_unique b (-a/-b) (a/b) (-(-a mod -b)) (a mod b)); auto.
+rewrite <- (opp_involutive b) at 3.
+rewrite <- opp_lt_mono.
+rewrite opp_nonneg_nonpos.
+destruct (mod_neg_bound (-a) (-b)); auto.
+rewrite opp_neg_pos; auto.
+apply mod_pos_bound; auto.
+split; auto.
+rewrite eq_opp_r; auto.
+
+rewrite eq_opp_l in EQ.
+rewrite opp_add_distr in EQ.
+rewrite <- mul_opp_l in EQ.
+destruct (div_mod_unique (-b) (-a/-b) (a/b) (-a mod -b) (-(a mod b))); auto.
+apply mod_pos_bound; auto.
+rewrite opp_pos_neg; order.
+rewrite <- opp_lt_mono.
+rewrite opp_nonneg_nonpos.
+destruct (mod_neg_bound a b); intuition; order.
+Qed.
+
+Lemma div_opp_opp : forall a b, b~=0 -> -a/-b == a/b.
+Proof.
+intros; destruct (div_mod_opp_opp a b); auto.
+Qed.
+
+Lemma mod_opp_opp : forall a b, b~=0 -> (-a) mod (-b) == - (a mod b).
+Proof.
+intros; destruct (div_mod_opp_opp a b); auto.
+Qed.
+
+
+(** 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.
+Proof.
+intros b q1 q2 r1 r2 Hr1 Hr2 EQ.
+destruct Hr1; destruct Hr2; try (intuition; order).
+apply div_mod_unique with b; auto.
+rewrite <- opp_inj_wd in EQ.
+rewrite 2 opp_add_distr in EQ. rewrite <- 2 mul_opp_l in EQ.
+rewrite <- (opp_inj_wd r1 r2).
+apply div_mod_unique with (-b); auto.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; intuition.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; intuition.
+Qed.
+
+Theorem div_unique:
+ forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> q == a/b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0) by (destruct Hr; intuition; order).
+rewrite (div_mod a b Hb) in EQ.
+destruct (div_mod_unique b (a/b) q (a mod b) r); auto.
+destruct Hr; [left; apply mod_pos_bound|right; apply mod_neg_bound];
+ intuition order.
+Qed.
+
+Theorem div_unique_pos:
+ forall a b q r, 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto. Qed.
+
+Theorem div_unique_neg:
+ forall a b q r, 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto. Qed.
+
+Theorem mod_unique:
+ forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> r == a mod b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0) by (destruct Hr; intuition; order).
+rewrite (div_mod a b Hb) in EQ.
+destruct (div_mod_unique b (a/b) q (a mod b) r); auto.
+destruct Hr; [left; apply mod_pos_bound|right; apply mod_neg_bound];
+ intuition order.
+Qed.
+
+Theorem mod_unique_pos:
+ forall a b q r, 0<=r<b -> a == b*q + r -> r == a mod b.
+Proof. intros; apply mod_unique with q; auto. Qed.
+
+Theorem mod_unique_neg:
+ forall a b q r, b<r<=0 -> a == b*q + r -> r == a mod b.
+Proof. intros; apply mod_unique with q; auto. Qed.
+
+
+(** A division by itself returns 1 *)
+
+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].
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof.
+intros. pos_or_neg a. apply div_same; order.
+rewrite <- div_opp_opp; auto. apply div_same; auto.
+Qed.
+
+Lemma mod_same : forall a, a~=0 -> a mod a == 0.
+Proof.
+intros. rewrite mod_eq, div_same; auto. nzsimpl. apply sub_diag.
+Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, 0<=a<b -> a/b == 0.
+Proof. exact div_small. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.
+Proof. exact mod_small. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, a~=0 -> 0/a == 0.
+Proof.
+intros. pos_or_neg a. apply div_0_l; order.
+rewrite <- div_opp_opp, opp_0; auto. apply div_0_l; auto.
+Qed.
+
+Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.
+Proof.
+intros; rewrite mod_eq, div_0_l; nzsimpl; auto.
+Qed.
+
+Lemma div_1_r: forall a, a/1 == a.
+Proof.
+intros. symmetry. apply div_unique with 0. left. split; order || apply lt_0_1.
+nzsimpl; auto.
+Qed.
+
+Lemma mod_1_r: forall a, a mod 1 == 0.
+Proof.
+intros. rewrite mod_eq, div_1_r; nzsimpl; auto using sub_diag.
+intro EQ; symmetry in EQ; revert EQ; apply lt_neq; apply lt_0_1.
+Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof. exact div_1_l. Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof. exact mod_1_l. Qed.
+
+Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.
+Proof.
+intros. symmetry. apply div_unique with 0.
+destruct (lt_ge_cases 0 b); [left|right]; split; order.
+nzsimpl; apply mul_comm.
+Qed.
+
+Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.
+Proof.
+intros. rewrite mod_eq, div_mul; auto. rewrite mul_comm; apply sub_diag.
+Qed.
+
+(** * 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.
+Proof. exact mod_le. Qed.
+
+Theorem div_pos : forall a b, 0<=a -> 0<b -> 0<= a/b.
+Proof. exact div_pos. Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof. exact div_str_pos. Qed.
+
+(* A REVOIR APRES LA REGLE DES SIGNES
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> 0<=a<b \/ b<a<=0).
+intros. apply div_small_iff; auto'. Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> a<b).
+Proof. intros. apply mod_small_iff; auto'. Qed.
+
+Lemma div_str_pos_iff : forall a b, b~=0 -> (0<a/b <-> b<=a).
+Proof. intros. apply div_str_pos_iff; auto'. Qed.
+*)
+
+(** 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.
+Proof. exact div_lt. Qed.
+
+(* STILL TODO !!
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono : forall a b c, 0<c -> a<=b -> a/c <= b/c.
+Proof.
+intros. destruct (le_gt_cases 0 a).
+apply div_le; auto.
+destruct (lt_ge_cases 0 b).
+apply le_trans with 0.
+ admit. (* !!! *)
+apply div_pos; order.
+Admitted. (* !!! *)
+
+Lemma mul_div_le : forall a b, b~=0 -> b*(a/b) <= a.
+Proof. intros. apply mul_div_le; auto'. Qed.
+
+Lemma mul_succ_div_gt: forall a b, b~=0 -> a < b*(S (a/b)).
+Proof. intros; apply mul_succ_div_gt; auto'. Qed.
+
+(** The previous inequality is exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).
+Proof. intros. apply div_exact; auto'. Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, b~=0 -> a < b*q -> a/b < q.
+Proof. intros. apply div_lt_upper_bound; auto'. Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, b~=0 -> a <= b*q -> a/b <= q.
+Proof. intros; apply div_le_upper_bound; auto'. Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, b~=0 -> b*q <= a -> q <= a/b.
+Proof. intros; apply div_le_lower_bound; auto'. Qed.
+
+(** A division respects opposite monotonicity for the divisor *)
+
+Lemma div_le_compat_l: forall p q r, 0<q<r -> p/r <= p/q.
+Proof. intros. apply div_le_compat_l. auto'. auto. Qed.
+
+(** * 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.
+Proof. intros. apply mod_add; auto'. Qed.
+
+Lemma div_add : forall a b c, c~=0 ->
+ (a + b * c) / c == a / c + b.
+Proof. intros. apply div_add; auto'. Qed.
+
+Lemma div_add_l: forall a b c, b~=0 ->
+ (a * b + c) / b == a + c / b.
+Proof. intros. apply div_add_l; auto'. Qed.
+
+(** Cancellations. *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> c~=0 ->
+ (a*c)/(b*c) == a/b.
+Proof. intros. apply div_mul_cancel_r; auto'. Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> c~=0 ->
+ (c*a)/(c*b) == a/b.
+Proof. intros. apply div_mul_cancel_l; auto'. Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> c~=0 ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof. intros. apply mul_mod_distr_l; auto'. Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> c~=0 ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof. intros. apply mul_mod_distr_r; auto'. Qed.
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof. intros. apply mod_mod; auto'. Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof. intros. apply mul_mod_idemp_l; auto'. Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof. intros. apply mul_mod_idemp_r; auto'. Qed.
+
+Theorem mul_mod: forall a b n, n~=0 ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof. intros. apply mul_mod; auto'. Qed.
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof. intros. apply add_mod_idemp_l; auto'. Qed.
+
+Lemma add_mod_idemp_r : forall a b n, n~=0 ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof. intros. apply add_mod_idemp_r; auto'. Qed.
+
+Theorem add_mod: forall a b n, n~=0 ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof. intros. apply add_mod; auto'. Qed.
+
+Lemma div_div : forall a b c, b~=0 -> c~=0 ->
+ (a/b)/c == a/(b*c).
+Proof. intros. apply div_div; auto'. Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, b~=0 -> c*(a/b) <= (c*a)/b.
+Proof. intros. apply div_mul_le; auto'. Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, b~=0 ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof. intros. apply mod_divides; auto'. Qed.
+*)
+
+End ZDivPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZDivMath.v b/theories/Numbers/Integer/Abstract/ZDivMath.v
new file mode 100644
index 000000000..dfc9ee4bc
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZDivMath.v
@@ -0,0 +1,396 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Euclidean Division for integers
+
+ We use here the "mathematical" convention, i.e. Round-Toward-Bottom :
+ [a = bq+r /\ 0 < r < |b| ]
+ *)
+
+Require Import ZAxioms ZProperties NZDiv.
+
+Open Scope NumScope.
+
+Module Type ZDiv (Import Z : ZAxiomsSig).
+
+ Parameter Inline div : t -> t -> t.
+ Parameter Inline modulo : t -> t -> t.
+
+ Infix "/" := div : NumScope.
+ Infix "mod" := modulo (at level 40, no associativity) : NumScope.
+
+ Instance div_wd : Proper (eq==>eq==>eq) div.
+ Instance mod_wd : Proper (eq==>eq==>eq) modulo.
+
+ Definition abs z := max z (-z).
+
+ Axiom div_mod : forall a b, b ~= 0 -> a == b*(a/b) + (a mod b).
+ Axiom mod_always_pos : forall a b, 0 <= a mod b < abs b.
+
+End ZDiv.
+
+Module Type ZDivSig := ZAxiomsSig <+ ZDiv.
+
+Module ZDivPropFunct (Import Z : ZDivSig).
+ (* TODO: en faire un arg du foncteur + comprendre le bug de SearchAbout *)
+ Module Import ZP := ZPropFunct Z.
+
+(** We benefit from what already exists for NZ *)
+
+ Module Z' <: NZDivSig.
+ Include Z.
+ Lemma mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+ Proof.
+ intros. rewrite <- (max_l b (-b)) at 3.
+ apply mod_always_pos.
+ apply le_trans with 0; [ rewrite opp_nonpos_nonneg |]; order.
+ Qed.
+ End Z'.
+ Module Import NZDivP := NZDivPropFunct Z'.
+
+(** Another formulation of the main equation *)
+
+Lemma mod_eq :
+ forall a b, b~=0 -> a mod b == a - b*(a/b).
+Proof.
+intros.
+rewrite <- add_move_l.
+symmetry. apply div_mod; auto.
+Qed.
+
+(* STILL TODO ...
+
+(** A few sign rules (simple ones) *)
+
+Lemma div_mod_opp_opp : forall a b, b~=0 ->
+ (-a/-b) == a/b /\ (-a) mod (-b) == -(a mod b).
+Proof.
+intros a b Hb.
+assert (-b ~= 0).
+ contradict Hb. rewrite eq_opp_l, opp_0 in Hb; auto.
+assert (EQ := opp_involutive a).
+rewrite (div_mod a b) in EQ at 2; auto.
+rewrite (div_mod (-a) (-b)) in EQ; auto.
+
+destruct (lt_ge_cases 0 b).
+rewrite opp_add_distr in EQ.
+rewrite <- mul_opp_l, opp_involutive in EQ.
+destruct (div_mod_unique b (-a/-b) (a/b) (-(-a mod -b)) (a mod b)); auto.
+rewrite <- (opp_involutive b) at 3.
+rewrite <- opp_lt_mono.
+rewrite opp_nonneg_nonpos.
+destruct (mod_neg_bound (-a) (-b)); auto.
+rewrite opp_neg_pos; auto.
+apply mod_pos_bound; auto.
+split; auto.
+rewrite eq_opp_r; auto.
+
+rewrite eq_opp_l in EQ.
+rewrite opp_add_distr in EQ.
+rewrite <- mul_opp_l in EQ.
+destruct (div_mod_unique (-b) (-a/-b) (a/b) (-a mod -b) (-(a mod b))); auto.
+apply mod_pos_bound; auto.
+rewrite opp_pos_neg; order.
+rewrite <- opp_lt_mono.
+rewrite opp_nonneg_nonpos.
+destruct (mod_neg_bound a b); intuition; order.
+Qed.
+
+Lemma div_opp_opp : forall a b, b~=0 -> -a/-b == a/b.
+Proof.
+intros; destruct (div_mod_opp_opp a b); auto.
+Qed.
+
+Lemma mod_opp_opp : forall a b, b~=0 -> (-a) mod (-b) == - (a mod b).
+Proof.
+intros; destruct (div_mod_opp_opp a b); auto.
+Qed.
+
+
+(** 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.
+Proof.
+intros b q1 q2 r1 r2 Hr1 Hr2 EQ.
+destruct Hr1; destruct Hr2; try (intuition; order).
+apply div_mod_unique with b; auto.
+rewrite <- opp_inj_wd in EQ.
+rewrite 2 opp_add_distr in EQ. rewrite <- 2 mul_opp_l in EQ.
+rewrite <- (opp_inj_wd r1 r2).
+apply div_mod_unique with (-b); auto.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; intuition.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; intuition.
+Qed.
+
+Theorem div_unique:
+ forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> q == a/b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0) by (destruct Hr; intuition; order).
+rewrite (div_mod a b Hb) in EQ.
+destruct (div_mod_unique b (a/b) q (a mod b) r); auto.
+destruct Hr; [left; apply mod_pos_bound|right; apply mod_neg_bound];
+ intuition order.
+Qed.
+
+Theorem div_unique_pos:
+ forall a b q r, 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto. Qed.
+
+Theorem div_unique_neg:
+ forall a b q r, 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto. Qed.
+
+Theorem mod_unique:
+ forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> r == a mod b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0) by (destruct Hr; intuition; order).
+rewrite (div_mod a b Hb) in EQ.
+destruct (div_mod_unique b (a/b) q (a mod b) r); auto.
+destruct Hr; [left; apply mod_pos_bound|right; apply mod_neg_bound];
+ intuition order.
+Qed.
+
+Theorem mod_unique_pos:
+ forall a b q r, 0<=r<b -> a == b*q + r -> r == a mod b.
+Proof. intros; apply mod_unique with q; auto. Qed.
+
+Theorem mod_unique_neg:
+ forall a b q r, b<r<=0 -> a == b*q + r -> r == a mod b.
+Proof. intros; apply mod_unique with q; auto. Qed.
+
+
+(** A division by itself returns 1 *)
+
+Ltac pos_or_neg a :=
+ destruct (le_gt_cases 0 a) as [LE|LT]; [|rewrite <- opp_pos_neg in LT].
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof.
+intros. pos_or_neg a. apply div_same; order.
+rewrite <- div_opp_opp; auto. apply div_same; auto.
+Qed.
+
+Lemma mod_same : forall a, a~=0 -> a mod a == 0.
+Proof.
+intros. rewrite mod_eq, div_same; auto. nzsimpl. apply sub_diag.
+Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, 0<=a<b -> a/b == 0.
+Proof. exact div_small. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.
+Proof. exact mod_small. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, a~=0 -> 0/a == 0.
+Proof.
+intros. pos_or_neg a. apply div_0_l; order.
+rewrite <- div_opp_opp, opp_0; auto. apply div_0_l; auto.
+Qed.
+
+Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.
+Proof.
+intros; rewrite mod_eq, div_0_l; nzsimpl; auto.
+Qed.
+
+Lemma div_1_r: forall a, a/1 == a.
+Proof.
+intros. symmetry. apply div_unique with 0. left. split; order || apply lt_0_1.
+nzsimpl; auto.
+Qed.
+
+Lemma mod_1_r: forall a, a mod 1 == 0.
+Proof.
+intros. rewrite mod_eq, div_1_r; nzsimpl; auto using sub_diag.
+intro EQ; symmetry in EQ; revert EQ; apply lt_neq; apply lt_0_1.
+Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof. exact div_1_l. Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof. exact mod_1_l. Qed.
+
+Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.
+Proof.
+intros. symmetry. apply div_unique with 0.
+destruct (lt_ge_cases 0 b); [left|right]; split; order.
+nzsimpl; apply mul_comm.
+Qed.
+
+Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.
+Proof.
+intros. rewrite mod_eq, div_mul; auto. rewrite mul_comm; apply sub_diag.
+Qed.
+
+(** * 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.
+Proof. exact mod_le. Qed.
+
+Theorem div_pos : forall a b, 0<=a -> 0<b -> 0<= a/b.
+Proof. exact div_pos. Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof. exact div_str_pos. Qed.
+
+(* A REVOIR APRES LA REGLE DES SIGNES
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> 0<=a<b \/ b<a<=0).
+intros. apply div_small_iff; auto'. Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> a<b).
+Proof. intros. apply mod_small_iff; auto'. Qed.
+
+Lemma div_str_pos_iff : forall a b, b~=0 -> (0<a/b <-> b<=a).
+Proof. intros. apply div_str_pos_iff; auto'. Qed.
+*)
+
+(** 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.
+Proof. exact div_lt. Qed.
+
+(* STILL TODO !!
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono: forall a b c, 0<c -> a<=b -> a/c <= b/c.
+Proof.
+intros. destruct (le_gt_cases 0 a).
+apply div_le_mono; auto.
+destruct (lt_ge_cases 0 b).
+apply le_trans with 0.
+ admit. (* !!! *)
+apply div_pos; order.
+Admitted. (* !!! *)
+
+Lemma mul_div_le : forall a b, b~=0 -> b*(a/b) <= a.
+Proof. intros. apply mul_div_le; auto'. Qed.
+
+Lemma mul_succ_div_gt: forall a b, b~=0 -> a < b*(S (a/b)).
+Proof. intros; apply mul_succ_div_gt; auto'. Qed.
+
+(** The previous inequality is exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).
+Proof. intros. apply div_exact; auto'. Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, b~=0 -> a < b*q -> a/b < q.
+Proof. intros. apply div_lt_upper_bound; auto'. Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, b~=0 -> a <= b*q -> a/b <= q.
+Proof. intros; apply div_le_upper_bound; auto'. Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, b~=0 -> b*q <= a -> q <= a/b.
+Proof. intros; apply div_le_lower_bound; auto'. Qed.
+
+(** A division respects opposite monotonicity for the divisor *)
+
+Lemma div_le_compat_l: forall p q r, 0<q<r -> p/r <= p/q.
+Proof. intros. apply div_le_compat_l. auto'. auto. Qed.
+
+(** * 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.
+Proof. intros. apply mod_add; auto'. Qed.
+
+Lemma div_add : forall a b c, c~=0 ->
+ (a + b * c) / c == a / c + b.
+Proof. intros. apply div_add; auto'. Qed.
+
+Lemma div_add_l: forall a b c, b~=0 ->
+ (a * b + c) / b == a + c / b.
+Proof. intros. apply div_add_l; auto'. Qed.
+
+(** Cancellations. *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> c~=0 ->
+ (a*c)/(b*c) == a/b.
+Proof. intros. apply div_mul_cancel_r; auto'. Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> c~=0 ->
+ (c*a)/(c*b) == a/b.
+Proof. intros. apply div_mul_cancel_l; auto'. Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> c~=0 ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof. intros. apply mul_mod_distr_l; auto'. Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> c~=0 ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof. intros. apply mul_mod_distr_r; auto'. Qed.
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof. intros. apply mod_mod; auto'. Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof. intros. apply mul_mod_idemp_l; auto'. Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof. intros. apply mul_mod_idemp_r; auto'. Qed.
+
+Theorem mul_mod: forall a b n, n~=0 ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof. intros. apply mul_mod; auto'. Qed.
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof. intros. apply add_mod_idemp_l; auto'. Qed.
+
+Lemma add_mod_idemp_r : forall a b n, n~=0 ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof. intros. apply add_mod_idemp_r; auto'. Qed.
+
+Theorem add_mod: forall a b n, n~=0 ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof. intros. apply add_mod; auto'. Qed.
+
+Lemma div_div : forall a b c, b~=0 -> c~=0 ->
+ (a/b)/c == a/(b*c).
+Proof. intros. apply div_div; auto'. Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, b~=0 -> c*(a/b) <= (c*a)/b.
+Proof. intros. apply div_mul_le; auto'. Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, b~=0 ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof. intros. apply mod_divides; auto'. Qed.
+*)
+*)
+
+End ZDivPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZDivOcaml.v b/theories/Numbers/Integer/Abstract/ZDivOcaml.v
new file mode 100644
index 000000000..2f68da933
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZDivOcaml.v
@@ -0,0 +1,541 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Euclidean Division for integers
+
+ We use here the convention of Ocaml and many other system (C, ASM, ...),
+ i.e. Round-Toward-Zero :
+
+ [a = bq+r /\ 0 < |r| < |b| /\ Sign(r) = Sgn(a)]
+*)
+
+Require Import ZAxioms ZProperties NZDiv.
+
+Open Scope NumScope.
+
+Module Type ZDiv (Import Z : ZAxiomsSig).
+
+ Parameter Inline div : t -> t -> t.
+ Parameter Inline modulo : t -> t -> t.
+
+ Infix "/" := div : NumScope.
+ Infix "mod" := modulo (at level 40, no associativity) : NumScope.
+
+ Instance div_wd : Proper (eq==>eq==>eq) div.
+ 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 : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+ Axiom mod_opp_l : forall a b, b ~= 0 -> (-a) mod b == - (a mod b).
+ Axiom mod_opp_r : forall a b, b ~= 0 -> a mod (-b) == a mod b.
+
+End ZDiv.
+
+Module Type ZDivSig := ZAxiomsSig <+ ZDiv.
+
+Module ZDivPropFunct (Import Z : ZDivSig).
+ (* TODO: en faire un arg du foncteur + comprendre le bug de SearchAbout *)
+ Module Import ZP := ZPropFunct Z.
+
+(** We benefit from what already exists for NZ *)
+
+ Module Import NZDivP := NZDivPropFunct Z.
+
+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].
+
+(** Another formulation of the main equation *)
+
+Lemma mod_eq :
+ forall a b, b~=0 -> a mod b == a - b*(a/b).
+Proof.
+intros.
+rewrite <- add_move_l.
+symmetry. apply div_mod; auto.
+Qed.
+
+(** A few sign rules (simple ones) *)
+
+Lemma mod_opp_opp : forall a b, b ~= 0 -> (-a) mod (-b) == - (a mod b).
+Proof. intros. rewrite mod_opp_r, mod_opp_l; auto. Qed.
+
+Lemma div_opp_l : forall a b, b ~= 0 -> (-a)/b == -(a/b).
+Proof.
+intros.
+rewrite <- (mul_cancel_l _ _ b); auto.
+rewrite <- (add_cancel_r _ _ ((-a) mod b)); auto.
+rewrite <- div_mod; auto.
+rewrite mod_opp_l, mul_opp_r, <- opp_add_distr, <- div_mod; auto.
+Qed.
+
+Lemma div_opp_r : forall a b, b ~= 0 -> a/(-b) == -(a/b).
+Proof.
+intros.
+assert (-b ~= 0) by (rewrite eq_opp_l, opp_0; auto).
+rewrite <- (mul_cancel_l _ _ (-b)); auto.
+rewrite <- (add_cancel_r _ _ (a mod (-b))); auto.
+rewrite <- div_mod; auto.
+rewrite mod_opp_r, mul_opp_opp, <- div_mod; auto.
+Qed.
+
+Lemma div_opp_opp : forall a b, b ~= 0 -> (-a)/(-b) == a/b.
+Proof. intros. rewrite div_opp_r, div_opp_l, opp_involutive; auto. Qed.
+
+(** The sign of [a mod b] is the one of [a] *)
+
+(* TODO: a proper sgn function and theory *)
+
+Lemma mod_sign : forall a b, b~=0 -> 0 <= (a mod b) * a.
+Proof.
+assert (Aux : forall a b, 0<b -> 0 <= (a mod b) * a).
+ intros. pos_or_neg a.
+ apply mul_nonneg_nonneg; auto. destruct (mod_bound a b); auto.
+ rewrite <- mul_opp_opp, <- mod_opp_l by order.
+ apply mul_nonneg_nonneg; try order. destruct (mod_bound (-a) b); try order.
+intros. pos_or_neg b. apply Aux; order.
+rewrite <- mod_opp_r by order. apply Aux; order.
+Qed.
+
+
+(** 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.
+Proof.
+intros b q1 q2 r1 r2 Hr1 Hr2 EQ.
+destruct Hr1; destruct Hr2; try (intuition; order).
+apply div_mod_unique with b; auto.
+rewrite <- opp_inj_wd in EQ.
+rewrite 2 opp_add_distr in EQ. rewrite <- 2 mul_opp_l in EQ.
+rewrite <- (opp_inj_wd r1 r2).
+apply div_mod_unique with (-b); auto.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; intuition.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; intuition.
+Qed.
+
+Theorem div_unique:
+ forall a b q r, 0<=a -> 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto. Qed.
+
+Theorem mod_unique:
+ forall a b q r, 0<=a -> 0<=r<b -> a == b*q + r -> r == a mod b.
+Proof. intros; apply mod_unique with q; auto. Qed.
+
+(** A division by itself returns 1 *)
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof.
+intros. pos_or_neg a. apply div_same; order.
+rewrite <- div_opp_opp; auto. apply div_same; auto.
+Qed.
+
+Lemma mod_same : forall a, a~=0 -> a mod a == 0.
+Proof.
+intros. rewrite mod_eq, div_same; auto. nzsimpl. apply sub_diag.
+Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, 0<=a<b -> a/b == 0.
+Proof. exact div_small. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.
+Proof. exact mod_small. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, a~=0 -> 0/a == 0.
+Proof.
+intros. pos_or_neg a. apply div_0_l; order.
+rewrite <- div_opp_opp, opp_0; auto. apply div_0_l; auto.
+Qed.
+
+Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.
+Proof.
+intros; rewrite mod_eq, div_0_l; nzsimpl; auto.
+Qed.
+
+Lemma div_1_r: forall a, a/1 == a.
+Proof.
+intros. pos_or_neg a. apply div_1_r; auto.
+apply opp_inj. rewrite <- div_opp_l. apply div_1_r; order.
+intro EQ; symmetry in EQ; revert EQ; apply lt_neq, lt_0_1.
+Qed.
+
+Lemma mod_1_r: forall a, a mod 1 == 0.
+Proof.
+intros. rewrite mod_eq, div_1_r; nzsimpl; auto using sub_diag.
+intro EQ; symmetry in EQ; revert EQ; apply lt_neq; apply lt_0_1.
+Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof. exact div_1_l. Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof. exact mod_1_l. Qed.
+
+Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.
+Proof.
+intros. pos_or_neg a; pos_or_neg b. apply div_mul; order.
+rewrite <- div_opp_opp, <- mul_opp_r by order. apply div_mul; order.
+rewrite <- opp_inj_wd, <- div_opp_l, <- mul_opp_l by order. apply div_mul; order.
+rewrite <- opp_inj_wd, <- div_opp_r, <- mul_opp_opp by order. apply div_mul; order.
+Qed.
+
+Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.
+Proof.
+intros. rewrite mod_eq, div_mul; auto. rewrite mul_comm; apply sub_diag.
+Qed.
+
+(** * 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.
+Proof. exact mod_le. Qed.
+
+Theorem div_pos : forall a b, 0<=a -> 0<b -> 0<= a/b.
+Proof. exact div_pos. Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof. exact div_str_pos. Qed.
+
+(** TODO: TO MIGRATE LATER *)
+Definition abs z := max z (-z).
+Lemma abs_pos : forall z, 0<=z -> abs z == z.
+Proof.
+intros; apply max_l. apply le_trans with 0; auto.
+rewrite opp_nonpos_nonneg; auto.
+Qed.
+Lemma abs_neg : forall z, 0<=-z -> abs z == -z.
+Proof.
+intros; apply max_r. apply le_trans with 0; auto.
+rewrite <- opp_nonneg_nonpos; auto.
+Qed.
+
+Lemma eq_sym_iff : forall x y, x==y <-> y==x.
+Proof.
+intros; split; symmetry; auto.
+Qed.
+
+(** END TODO *)
+
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> abs a < abs b).
+Proof.
+intros. pos_or_neg a; pos_or_neg b.
+rewrite div_small_iff; try order. rewrite 2 abs_pos; intuition; order.
+rewrite <- opp_inj_wd, opp_0, <- div_opp_r, div_small_iff by order.
+ rewrite (abs_pos a), (abs_neg b); intuition; order.
+rewrite <- opp_inj_wd, opp_0, <- div_opp_l, div_small_iff by order.
+ rewrite (abs_neg a), (abs_pos b); intuition; order.
+rewrite <- div_opp_opp, div_small_iff by order.
+ rewrite (abs_neg a), (abs_neg b); intuition; order.
+Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> abs a < abs b).
+Proof.
+intros. rewrite mod_eq, <- div_small_iff by order.
+rewrite sub_move_r, <- (add_0_r a) at 1. rewrite add_cancel_l.
+rewrite eq_sym_iff, eq_mul_0. intuition.
+Qed.
+
+(** 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.
+Proof. exact div_lt. Qed.
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono : forall a b c, 0<c -> a<=b -> a/c <= b/c.
+Proof.
+intros. pos_or_neg a. apply div_le_mono; auto.
+pos_or_neg b. apply le_trans with 0.
+ rewrite <- opp_nonneg_nonpos, <- div_opp_l by order.
+ apply div_pos; order.
+ apply div_pos; order.
+rewrite opp_le_mono in *. rewrite <- 2 div_opp_l by order.
+ apply div_le_mono; intuition; order.
+Qed.
+
+(** With this choice of division,
+ rounding of div is always done toward zero: *)
+
+Lemma mul_div_le : forall a b, 0<=a -> b~=0 -> 0 <= b*(a/b) <= a.
+Proof.
+intros. pos_or_neg b.
+split.
+apply mul_nonneg_nonneg; [|apply div_pos]; order.
+apply mul_div_le; order.
+rewrite <- mul_opp_opp, <- div_opp_r by order.
+split.
+apply mul_nonneg_nonneg; [|apply div_pos]; order.
+apply mul_div_le; order.
+Qed.
+
+Lemma mul_div_ge : forall a b, a<=0 -> b~=0 -> a <= b*(a/b) <= 0.
+Proof.
+intros.
+rewrite <- opp_nonneg_nonpos, opp_le_mono, <-mul_opp_r, <-div_opp_l by order.
+destruct (mul_div_le (-a) b); auto.
+rewrite opp_nonneg_nonpos; auto.
+Qed.
+
+(** For positive numbers, considering [S (a/b)] leads to an upper bound for [a] *)
+
+Lemma mul_succ_div_gt: forall a b, 0<=a -> 0<b -> a < b*(S (a/b)).
+Proof. exact mul_succ_div_gt. Qed.
+
+(** Some previous inequalities are exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).
+Proof.
+intros. rewrite mod_eq by order. rewrite sub_move_r; nzsimpl; intuition.
+Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, 0<=a -> 0<b -> a < b*q -> a/b < q.
+Proof. exact div_lt_upper_bound. Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, 0<b -> a <= b*q -> a/b <= q.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; auto. rewrite mul_comm; auto.
+Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, 0<b -> b*q <= a -> q <= a/b.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; auto. rewrite mul_comm; auto.
+Qed.
+
+(** 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.
+Proof. exact div_le_compat_l. Qed.
+
+(** * Relations between usual operations and mod and div *)
+
+(** Unlike with other division conventions, some results here aren't
+ always valid, and need to be restricted. For instance
+ [(a+b*c) mod c <> a mod c] for [a=9,b=-5,c=2] *)
+
+Lemma mod_add : forall a b c, c~=0 -> 0 <= (a+b*c)*a ->
+ (a + b * c) mod c == a mod c.
+Proof.
+assert (forall a b c, c~=0 -> 0<=a -> 0<=a+b*c -> (a+b*c) mod c == a mod c).
+ intros. pos_or_neg c. apply mod_add; order.
+ rewrite <- (mod_opp_r a), <- (mod_opp_r (a+b*c)) by order.
+ rewrite <- mul_opp_opp in *.
+ apply mod_add; order.
+intros a b c Hc Habc.
+destruct (le_0_mul _ _ Habc) as [(Habc',Ha)|(Habc',Ha)]; auto.
+apply opp_inj. revert Ha Habc'.
+rewrite <- 2 opp_nonneg_nonpos.
+rewrite <- 2 mod_opp_l, opp_add_distr, <- mul_opp_l by order; auto.
+Qed.
+
+Lemma div_add : forall a b c, c~=0 -> 0 <= (a+b*c)*a ->
+ (a + b * c) / c == a / c + b.
+Proof.
+intros.
+rewrite <- (mul_cancel_l _ _ c) by auto.
+rewrite <- (add_cancel_r _ _ ((a+b*c) mod c)).
+rewrite <- div_mod; auto.
+rewrite mod_add; auto.
+rewrite mul_add_distr_l, add_shuffle0, <-div_mod, mul_comm; auto.
+Qed.
+
+Lemma div_add_l: forall a b c, b~=0 -> 0 <= (a*b+c)*c ->
+ (a * b + c) / b == a + c / b.
+Proof.
+ intros a b c. rewrite add_comm, (add_comm a). intros; apply div_add; auto.
+Qed.
+
+(** Cancellations. *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> c~=0 ->
+ (a*c)/(b*c) == a/b.
+Proof.
+assert (Aux1 : forall a b c, 0<=a -> 0<b -> c~=0 -> (a*c)/(b*c) == a/b).
+ intros. pos_or_neg c. apply div_mul_cancel_r; order.
+ rewrite <- div_opp_opp, <- 2 mul_opp_r. apply div_mul_cancel_r; order.
+ rewrite <- neq_mul_0; intuition order.
+assert (Aux2 : forall a b c, 0<=a -> b~=0 -> c~=0 -> (a*c)/(b*c) == a/b).
+ intros. pos_or_neg b. apply Aux1; order.
+ apply opp_inj. rewrite <- 2 div_opp_r, <- mul_opp_l; try order. apply Aux1; order.
+ rewrite <- neq_mul_0; intuition order.
+intros. pos_or_neg a. apply Aux2; order.
+apply opp_inj. rewrite <- 2 div_opp_l, <- mul_opp_l; try order. apply Aux2; order.
+rewrite <- neq_mul_0; intuition order.
+Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> c~=0 ->
+ (c*a)/(c*b) == a/b.
+Proof.
+intros. rewrite !(mul_comm c); apply div_mul_cancel_r; auto.
+Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> c~=0 ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof.
+intros.
+assert (b*c ~= 0) by (rewrite <- neq_mul_0; intuition).
+rewrite ! mod_eq by auto.
+rewrite div_mul_cancel_r by order.
+rewrite mul_sub_distr_r, <- !mul_assoc, (mul_comm (a/b) c); auto.
+Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> c~=0 ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof.
+intros; rewrite !(mul_comm c); apply mul_mod_distr_r; auto.
+Qed.
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof.
+intros. pos_or_neg a; pos_or_neg n. apply mod_mod; order.
+rewrite <- ! (mod_opp_r _ n) by auto. apply mod_mod; order.
+apply opp_inj. rewrite <- !mod_opp_l by order. apply mod_mod; order.
+apply opp_inj. rewrite <- !mod_opp_opp by order. apply mod_mod; order.
+Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof.
+assert (Aux1 : forall a b n, 0<=a -> 0<=b -> n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n).
+ intros. pos_or_neg n. apply mul_mod_idemp_l; order.
+ rewrite <- ! (mod_opp_r _ n) by order. apply mul_mod_idemp_l; order.
+assert (Aux2 : forall a b n, 0<=a -> n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n).
+ intros. pos_or_neg b. apply Aux1; auto.
+ apply opp_inj. rewrite <-2 mod_opp_l, <-2 mul_opp_r by order.
+ apply Aux1; order.
+intros a b n Hn. pos_or_neg a. apply Aux2; auto.
+apply opp_inj. rewrite <-2 mod_opp_l, <-2 mul_opp_l, <-mod_opp_l by order.
+apply Aux2; order.
+Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof.
+intros. rewrite !(mul_comm a). apply mul_mod_idemp_l; auto.
+Qed.
+
+Theorem mul_mod: forall a b n, n~=0 ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof.
+intros. rewrite mul_mod_idemp_l, mul_mod_idemp_r; auto.
+Qed.
+
+(** addition and modulo
+
+ Generally speaking, unlike with other conventions, we don't have
+ [(a+b) mod n = (a mod n + b mod n) mod n]
+ for any a and b.
+ For instance, take (8 + (-10)) mod 3 = -2 whereas
+ (8 mod 3 + (-10 mod 3)) mod 3 = 1.
+*)
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 -> 0 <= a*b ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof.
+assert (Aux : forall a b n, 0<=a -> 0<=b -> n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n).
+ intros. pos_or_neg n. apply add_mod_idemp_l; order.
+ rewrite <- ! (mod_opp_r _ n) by order. apply add_mod_idemp_l; order.
+intros a b n Hn Hab. destruct (le_0_mul _ _ Hab) as [(Ha,Hb)|(Ha,Hb)].
+apply Aux; auto.
+apply opp_inj. rewrite <-2 mod_opp_l, 2 opp_add_distr, <-mod_opp_l by order.
+apply Aux; auto.
+rewrite opp_nonneg_nonpos; auto.
+rewrite opp_nonneg_nonpos; auto.
+Qed.
+
+Lemma add_mod_idemp_r : forall a b n, n~=0 -> 0 <= a*b ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof.
+intros. rewrite !(add_comm a). apply add_mod_idemp_l; auto.
+rewrite mul_comm; auto.
+Qed.
+
+Theorem add_mod: forall a b n, n~=0 -> 0 <= a*b ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof.
+intros a b n Hn Hab. rewrite add_mod_idemp_l, add_mod_idemp_r; auto.
+destruct (le_0_mul _ _ Hab) as [(Ha,Hb)|(Ha,Hb)];
+ destruct (le_0_mul _ _ (mod_sign b n Hn)) as [(Hb',Hm)|(Hb',Hm)];
+ auto using mul_nonneg_nonneg, mul_nonpos_nonpos.
+ setoid_replace b with 0 by order. rewrite mod_0_l by order. nzsimpl; order.
+ setoid_replace b with 0 by order. rewrite mod_0_l by order. nzsimpl; order.
+Qed.
+
+
+(** Conversely, the following result needs less restrictions here. *)
+
+Lemma div_div : forall a b c, b~=0 -> c~=0 ->
+ (a/b)/c == a/(b*c).
+Proof.
+assert (Aux1 : forall a b c, 0<=a -> 0<b -> c~=0 -> (a/b)/c == a/(b*c)).
+ intros. pos_or_neg c. apply div_div; order.
+ apply opp_inj. rewrite <- 2 div_opp_r, <- mul_opp_r; auto. apply div_div; order.
+ rewrite <- neq_mul_0; intuition order.
+assert (Aux2 : forall a b c, 0<=a -> b~=0 -> c~=0 -> (a/b)/c == a/(b*c)).
+ intros. pos_or_neg b. apply Aux1; order.
+ apply opp_inj. rewrite <- div_opp_l, <- 2 div_opp_r, <- mul_opp_l; auto.
+ rewrite <- neq_mul_0; intuition order.
+intros. pos_or_neg a. apply Aux2; order.
+apply opp_inj. rewrite <- 3 div_opp_l; try order. apply Aux2; order.
+rewrite <- neq_mul_0; intuition order.
+Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, 0<=a -> 0<b -> 0<=c -> c*(a/b) <= (c*a)/b.
+Proof. exact div_mul_le. Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, b~=0 ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof.
+ intros.
+ pos_or_neg a; pos_or_neg b.
+ apply mod_divides; order.
+ rewrite <- mod_opp_r, mod_divides by order.
+ split; intros (c,Hc); exists (-c).
+ rewrite mul_opp_r, <- mul_opp_l; auto.
+ rewrite mul_opp_opp; auto.
+ rewrite <- opp_inj_wd, opp_0, <- mod_opp_l, mod_divides by order.
+ split; intros (c,Hc); exists (-c).
+ rewrite mul_opp_r, eq_opp_r; auto.
+ rewrite mul_opp_r, opp_inj_wd; auto.
+ rewrite <- opp_inj_wd, opp_0, <- mod_opp_opp, mod_divides by order.
+ split; intros (c,Hc); exists c.
+ rewrite <-opp_inj_wd, <- mul_opp_l; auto.
+ rewrite mul_opp_l, opp_inj_wd; auto.
+Qed.
+
+End ZDivPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZMulOrder.v b/theories/Numbers/Integer/Abstract/ZMulOrder.v
index 4f11dcc5c..a9cf0dc4d 100644
--- a/theories/Numbers/Integer/Abstract/ZMulOrder.v
+++ b/theories/Numbers/Integer/Abstract/ZMulOrder.v
@@ -34,12 +34,6 @@ now apply mul_le_mono_nonpos_l.
apply mul_le_mono_nonpos_r; [now apply le_trans with q | assumption].
Qed.
-Theorem mul_nonneg_nonneg : forall n m, 0 <= n -> 0 <= m -> 0 <= n * m.
-Proof.
-intros n m H1 H2.
-rewrite <- (mul_0_l m). now apply mul_le_mono_nonneg_r.
-Qed.
-
Theorem mul_nonpos_nonpos : forall n m, n <= 0 -> m <= 0 -> 0 <= n * m.
Proof.
intros n m H1 H2.