aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-12-17 16:46:54 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-12-17 16:46:54 +0000
commit35e84a95326c95f9399084c843e244de6ae25753 (patch)
tree46051f847373a93c6cd9c79b583503e79af1f9bb /theories
parent622dbc2eb726949cda5c44639250336586ed8379 (diff)
Division in Numbers : more properties, new filenames based on a paper by R. Boute
Following R. Boute (paper "the Euclidean Definition of the Functions div and mod"): - ZDivFloor.v for Coq historical division (former ZDivCoq.v) - ZDivTrunc.v for Ocaml convention (former ZDivOcaml.v) - ZDivEucl.v for "Mathematical" convention 0<=r (former ZDivMath.v) These property functors are more or less finished (except that sign and abs stuff should be migrated to a separate file). git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12594 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories')
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivEucl.v671
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivFloor.v (renamed from theories/Numbers/Integer/Abstract/ZDivCoq.v)52
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivMath.v396
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivTrunc.v (renamed from theories/Numbers/Integer/Abstract/ZDivOcaml.v)47
-rw-r--r--theories/Numbers/NatInt/NZAxioms.v4
-rw-r--r--theories/Numbers/NatInt/NZDomain.v2
-rw-r--r--theories/Numbers/vo.itarget6
-rw-r--r--theories/ZArith/ZOdiv.v4
8 files changed, 754 insertions, 428 deletions
diff --git a/theories/Numbers/Integer/Abstract/ZDivEucl.v b/theories/Numbers/Integer/Abstract/ZDivEucl.v
new file mode 100644
index 000000000..0f15e9a29
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZDivEucl.v
@@ -0,0 +1,671 @@
+(************************************************************************)
+(* 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, Euclid convention
+
+ We use here the "usual" formulation of the Euclid Theorem
+ [forall a b, b<>0 -> exists b q, a = b*q+r /\ 0 < r < |b| ]
+
+ The outcome of the modulo function is hence always positive.
+ This corresponds to convention "E" 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 [ZDivFloor] for others conventions.
+*)
+
+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.
+
+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].
+
+(*TODO: To migrate later ... *)
+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 abs_opp : forall z, abs (-z) == abs z.
+Proof.
+intros. pos_or_neg z.
+rewrite (abs_pos z), (abs_neg (-z)); try rewrite opp_involutive; auto.
+rewrite (abs_pos (-z)), (abs_neg z); order.
+Qed.
+Lemma abs_nonneg : forall z, 0<=abs z.
+Proof.
+intros. pos_or_neg z.
+rewrite abs_pos; auto.
+rewrite <-abs_opp, abs_pos; order.
+Qed.
+Lemma abs_nz : forall z, z~=0 -> 0<abs z.
+Proof.
+intros. pos_or_neg z.
+rewrite abs_pos; order.
+rewrite <-abs_opp, abs_pos; order.
+Qed.
+Lemma abs_mul : forall a b, abs (a*b) == abs a * abs b.
+Proof.
+intros. pos_or_neg a; pos_or_neg b.
+rewrite 3 abs_pos; auto using mul_nonneg_nonneg.
+rewrite (abs_pos a), 2 abs_neg; try order.
+ rewrite mul_opp_r; auto.
+ rewrite <-mul_opp_r; apply mul_nonneg_nonneg; order.
+rewrite (abs_pos b), 2 abs_neg; try order.
+ rewrite mul_opp_l; auto.
+ rewrite <-mul_opp_l; apply mul_nonneg_nonneg; order.
+rewrite (abs_pos (a*b)), 2 abs_neg; try order.
+ rewrite mul_opp_opp; auto.
+ rewrite <-mul_opp_opp; apply mul_nonneg_nonneg; order.
+Qed.
+(*/TODO *)
+
+(** Uniqueness theorems *)
+
+Theorem div_mod_unique : forall b q1 q2 r1 r2 : t,
+ 0<=r1<abs b -> 0<=r2<abs b ->
+ b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.
+Proof.
+intros b q1 q2 r1 r2 Hr1 Hr2 EQ.
+pos_or_neg b.
+rewrite abs_pos in *; auto.
+apply div_mod_unique with b; auto.
+rewrite abs_neg in *; try order.
+rewrite eq_sym_iff. apply div_mod_unique with (-b); auto.
+rewrite 2 mul_opp_l.
+rewrite add_move_l, sub_opp_r.
+rewrite <-add_assoc.
+symmetry. rewrite add_move_l, sub_opp_r.
+rewrite (add_comm r2), (add_comm r1); auto.
+Qed.
+
+Theorem div_unique:
+ forall a b q r, 0<=r<abs b -> a == b*q + r -> q == a/b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0).
+ pos_or_neg b.
+ rewrite abs_pos in Hr; intuition; order.
+ rewrite <- opp_0, eq_opp_r. rewrite abs_neg in Hr; intuition; order.
+rewrite (div_mod a b Hb) in EQ.
+destruct (div_mod_unique b (a/b) q (a mod b) r); auto.
+apply mod_always_pos.
+Qed.
+
+Theorem mod_unique:
+ forall a b q r, 0<=r<abs b -> a == b*q + r -> r == a mod b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0).
+ pos_or_neg b.
+ rewrite abs_pos in Hr; intuition; order.
+ rewrite <- opp_0, eq_opp_r. rewrite abs_neg in Hr; intuition; order.
+rewrite (div_mod a b Hb) in EQ.
+destruct (div_mod_unique b (a/b) q (a mod b) r); auto.
+apply mod_always_pos.
+Qed.
+
+(** TODO: Provide a [sign] function *)
+Parameter sign : t -> t.
+Parameter sign_pos : forall t, sign t == 1 <-> 0<t.
+Parameter sign_0 : forall t, sign t == 0 <-> t==0.
+Parameter sign_neg : forall t, sign t == -1 <-> t<0.
+Parameter sign_abs : forall t, t * sign t == abs t.
+(** END TODO *)
+
+
+(** Sign rules *)
+
+Lemma div_opp_r : forall a b, b~=0 -> a/(-b) == -(a/b).
+Proof.
+intros. symmetry.
+apply div_unique with (a mod b).
+rewrite abs_opp; apply mod_always_pos.
+rewrite mul_opp_opp; apply div_mod; auto.
+Qed.
+
+Lemma mod_opp_r : forall a b, b~=0 -> a mod (-b) == a mod b.
+Proof.
+intros. symmetry.
+apply mod_unique with (-(a/b)).
+rewrite abs_opp; apply mod_always_pos.
+rewrite mul_opp_opp; apply div_mod; auto.
+Qed.
+
+Lemma div_opp_l_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a)/b == -(a/b).
+Proof.
+intros a b Hb Hab. symmetry.
+apply div_unique with (-(a mod b)).
+rewrite Hab, opp_0. split; [order|].
+pos_or_neg b; [rewrite abs_pos | rewrite abs_neg]; order.
+rewrite mul_opp_r, <-opp_add_distr, <-div_mod; auto.
+Qed.
+
+Lemma div_opp_l_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a)/b == -(a/b)-sign b.
+Proof.
+intros a b Hb Hab. symmetry.
+apply div_unique with (abs b -(a mod b)).
+rewrite lt_sub_lt_add_l.
+rewrite <- le_add_le_sub_l. nzsimpl.
+rewrite <- (add_0_l (abs b)) at 2.
+rewrite <- add_lt_mono_r.
+destruct (mod_always_pos a b); intuition order.
+rewrite <- 2 add_opp_r, mul_add_distr_l, 2 mul_opp_r.
+rewrite sign_abs.
+rewrite add_shuffle2, add_opp_diag_l; nzsimpl.
+rewrite <-opp_add_distr, <-div_mod; order.
+Qed.
+
+Lemma mod_opp_l_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a) mod b == 0.
+Proof.
+intros a b Hb Hab. symmetry.
+apply mod_unique with (-(a/b)).
+split; [order|apply abs_nz; auto].
+rewrite <-opp_0, <-Hab, mul_opp_r, <-opp_add_distr, <-div_mod; auto.
+Qed.
+
+Lemma mod_opp_l_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a) mod b == abs b - (a mod b).
+Proof.
+intros a b Hb Hab. symmetry.
+apply mod_unique with (-(a/b)-sign b).
+rewrite lt_sub_lt_add_l.
+rewrite <- le_add_le_sub_l. nzsimpl.
+rewrite <- (add_0_l (abs b)) at 2.
+rewrite <- add_lt_mono_r.
+destruct (mod_always_pos a b); intuition order.
+rewrite <- 2 add_opp_r, mul_add_distr_l, 2 mul_opp_r.
+rewrite sign_abs.
+rewrite add_shuffle2, add_opp_diag_l; nzsimpl.
+rewrite <-opp_add_distr, <-div_mod; order.
+Qed.
+
+Lemma div_opp_opp_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a)/(-b) == a/b.
+Proof.
+intros. rewrite div_opp_r, div_opp_l_z, opp_involutive; auto.
+Qed.
+
+Lemma div_opp_opp_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a)/(-b) == a/b + sign(b).
+Proof.
+intros. rewrite div_opp_r, div_opp_l_nz; auto.
+rewrite opp_sub_distr, opp_involutive; auto.
+Qed.
+
+Lemma mod_opp_opp_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a) mod (-b) == 0.
+Proof.
+intros. rewrite mod_opp_r, mod_opp_l_z; auto.
+Qed.
+
+Lemma mod_opp_opp_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a) mod (-b) == abs b - a mod b.
+Proof.
+intros. rewrite mod_opp_r, mod_opp_l_nz; auto.
+Qed.
+
+(** A division by itself returns 1 *)
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof.
+intros. symmetry. apply div_unique with 0.
+split; [order|apply abs_nz; auto].
+nzsimpl; 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.
+apply opp_inj. rewrite <- div_opp_r, 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.
+assert (H:=lt_0_1); rewrite abs_pos; intuition; order.
+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.
+apply neq_sym, 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.
+split; [order|apply abs_nz; auto].
+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 -> b~=0 -> a mod b <= a.
+Proof.
+intros. pos_or_neg b. apply mod_le; order.
+rewrite <- mod_opp_r; auto. apply mod_le; order.
+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.
+
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> 0<=a<abs b).
+Proof.
+intros a b Hb.
+split.
+intros EQ.
+rewrite (div_mod a b Hb), EQ; nzsimpl.
+apply mod_always_pos; auto.
+intros. pos_or_neg b.
+apply div_small; auto.
+rewrite <- (abs_pos b); auto.
+apply opp_inj; rewrite opp_0, <- div_opp_r; auto.
+apply div_small; auto.
+rewrite <- (abs_neg b); auto.
+rewrite <-opp_0 in Hb. rewrite eq_opp_r in Hb. order.
+Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> 0<=a<abs b).
+Proof.
+intros.
+rewrite <- div_small_iff, mod_eq; auto.
+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 a b c Hc Hab.
+rewrite lt_eq_cases in Hab. destruct Hab as [LT|EQ];
+ [|rewrite EQ; order].
+rewrite <- lt_succ_r.
+rewrite (mul_lt_mono_pos_l c) by order.
+nzsimpl.
+rewrite (add_lt_mono_r _ _ (a mod c)).
+rewrite <- div_mod by order.
+apply lt_le_trans with b; auto.
+rewrite (div_mod b c) at 1; [| order].
+rewrite <- add_assoc, <- add_le_mono_l.
+apply le_trans with (c+0).
+nzsimpl; destruct (mod_always_pos b c); try order.
+rewrite abs_pos in *; order.
+rewrite <- add_le_mono_l. destruct (mod_always_pos a c); order.
+Qed.
+
+(** In this convention, [div] performs Rounding-Toward-Bottom
+ when divisor is positive, and Rounding-Toward-Top otherwise.
+
+ Since we cannot speak of rational values here, we express this
+ fact by multiplying back by [b], and this leads to a nice
+ unique statement.
+*)
+
+Lemma mul_div_le : forall a b, b~=0 -> b*(a/b) <= a.
+Proof.
+intros.
+rewrite (div_mod a b) at 2; auto.
+rewrite <- (add_0_r (b*(a/b))) at 1.
+rewrite <- add_le_mono_l.
+destruct (mod_always_pos a b); auto.
+Qed.
+
+(** Giving a reversed bound is slightly more complex *)
+
+Lemma mul_succ_div_gt: forall a b, 0<b -> a < b*(S (a/b)).
+Proof.
+intros.
+nzsimpl.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- add_lt_mono_l.
+destruct (mod_always_pos a b).
+rewrite abs_pos in *; order.
+Qed.
+
+Lemma mul_pred_div_gt: forall a b, b<0 -> a < b*(P (a/b)).
+Proof.
+intros a b Hb.
+rewrite mul_pred_r, <- add_opp_r.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- add_lt_mono_l.
+destruct (mod_always_pos a b).
+rewrite <- opp_pos_neg in Hb. rewrite abs_neg in *; order.
+Qed.
+
+(** NB: The three 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).
+Proof.
+intros.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- (add_0_r (b*(a/b))) at 2.
+apply add_cancel_l.
+Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, 0<b -> a < b*q -> a/b < q.
+Proof.
+intros.
+rewrite (mul_lt_mono_pos_l b); auto.
+apply le_lt_trans with a; auto.
+apply mul_div_le; order.
+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 *)
+
+Lemma mod_add : forall a b c, c~=0 ->
+ (a + b * c) mod c == a mod c.
+Proof.
+intros.
+symmetry.
+apply mod_unique with (a/c+b); auto.
+apply mod_always_pos; auto.
+rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+rewrite mul_comm; auto.
+Qed.
+
+Lemma div_add : forall a b c, c~=0 ->
+ (a + b * c) / c == a / c + b.
+Proof.
+intros.
+apply (mul_cancel_l _ _ c); try order.
+apply (add_cancel_r _ _ ((a+b*c) mod c)).
+rewrite <- div_mod, mod_add by order.
+rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+rewrite mul_comm; auto.
+Qed.
+
+Lemma div_add_l: forall a b c, b~=0 ->
+ (a * b + c) / b == a + c / b.
+Proof.
+ intros a b c. rewrite (add_comm _ c), (add_comm a).
+ intros. apply div_add; auto.
+Qed.
+
+(** Cancellations. *)
+
+(** With the current convention, the following isn't always true
+ when [c<0]: [-3*-1 / -2*-1 = 3/2 = 1] while [-3/-2 = 2] *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> 0<c ->
+ (a*c)/(b*c) == a/b.
+Proof.
+intros.
+symmetry.
+apply div_unique with ((a mod b)*c).
+(* ineqs *)
+rewrite abs_mul, (abs_pos c) by order.
+rewrite <-(mul_0_l c), <-mul_lt_mono_pos_r, <-mul_le_mono_pos_r; auto.
+apply mod_always_pos.
+(* equation *)
+rewrite (div_mod a b) at 1; [|order].
+rewrite mul_add_distr_r.
+rewrite add_cancel_r.
+rewrite <- 2 mul_assoc. rewrite (mul_comm c); auto.
+Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> 0<c ->
+ (c*a)/(c*b) == a/b.
+Proof.
+intros. rewrite !(mul_comm c); apply div_mul_cancel_r; auto.
+Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> 0<c ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof.
+intros.
+rewrite <- (add_cancel_l _ _ ((c*b)* ((c*a)/(c*b)))).
+rewrite <- div_mod.
+rewrite div_mul_cancel_l; auto.
+rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+apply div_mod; order.
+rewrite <- neq_mul_0; intuition; order.
+Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> 0<c ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof.
+ intros. rewrite !(mul_comm _ c); rewrite mul_mod_distr_l; auto.
+Qed.
+
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof.
+intros. rewrite mod_small_iff; auto.
+apply mod_always_pos; 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 a b n Hn. symmetry.
+ rewrite (div_mod a n) at 1; [|order].
+ rewrite add_comm, (mul_comm n), (mul_comm _ b).
+ rewrite mul_add_distr_l, mul_assoc.
+ intros. rewrite mod_add; auto.
+ rewrite mul_comm; 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. 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.
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof.
+ intros a b n Hn. symmetry.
+ rewrite (div_mod a n) at 1; [|order].
+ rewrite <- add_assoc, add_comm, mul_comm.
+ intros. rewrite mod_add; 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. rewrite !(add_comm a). apply add_mod_idemp_l; 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. rewrite add_mod_idemp_l, add_mod_idemp_r; auto.
+Qed.
+
+(** With the current convention, the following result isn't always
+ true for negative divisors. For instance
+ [ 3/(-2)/(-2) = 1 <> 0 = 3 / (-2*-2) ]. *)
+
+Lemma div_div : forall a b c, 0<b -> 0<c ->
+ (a/b)/c == a/(b*c).
+Proof.
+ intros a b c Hb Hc.
+ apply div_unique with (b*((a/b) mod c) + a mod b); auto.
+ (* begin 0<= ... <abs(b*c) *)
+ rewrite abs_mul.
+ destruct (mod_always_pos (a/b) c), (mod_always_pos a b); auto using div_pos.
+ split.
+ apply add_nonneg_nonneg; auto.
+ apply mul_nonneg_nonneg; order.
+ apply lt_le_trans with (b*((a/b) mod c) + abs b).
+ rewrite <- add_lt_mono_l; auto.
+ rewrite (abs_pos b) by order.
+ rewrite <- mul_succ_r, <- mul_le_mono_pos_l, le_succ_l; auto.
+ (* end 0<= ... < abs(b*c) *)
+ rewrite (div_mod a b) at 1; [|order].
+ rewrite add_assoc, add_cancel_r.
+ rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+ apply div_mod; 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 a b Hb. split.
+intros Hab. exists (a/b). rewrite (div_mod a b Hb) at 1.
+ rewrite Hab; nzsimpl; auto.
+intros (c,Hc).
+rewrite Hc, mul_comm.
+apply mod_mul; auto.
+Qed.
+
+
+End ZDivPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZDivCoq.v b/theories/Numbers/Integer/Abstract/ZDivFloor.v
index 6fad3ef83..b4a8a4203 100644
--- a/theories/Numbers/Integer/Abstract/ZDivCoq.v
+++ b/theories/Numbers/Integer/Abstract/ZDivFloor.v
@@ -6,11 +6,23 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(** Euclidean Division for integers
+(** * Euclidean Division for integers (Floor convention)
- We use here the historical convention of Coq :
- [a = bq+r /\ 0 < |r| < |b| /\ Sign(r) = Sgn(b)]
- *)
+ 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 [Zdiv] 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.
+*)
Require Import ZAxioms ZProperties NZDiv.
@@ -363,8 +375,14 @@ nzsimpl; destruct (mod_pos_bound b c); order.
rewrite <- add_le_mono_l. destruct (mod_pos_bound a c); order.
Qed.
-(** With this choice of division, rounding of div is done
- toward bottom when the divisor is positive. *)
+(** 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.
Proof.
@@ -375,9 +393,14 @@ rewrite <- add_le_mono_l.
destruct (mod_pos_bound a b); auto.
Qed.
-(** Again with a positive [b], we can give an upper bound for [a].
- Together with the previous inequality, this fact characterizes
- division by a positive number.
+Lemma mul_div_ge : forall a b, b<0 -> a <= b*(a/b).
+Proof.
+intros. rewrite <- div_opp_opp, opp_le_mono, <-mul_opp_l by order.
+apply mul_div_le. rewrite opp_pos_neg; auto.
+Qed.
+
+(** ... 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)).
@@ -389,20 +412,15 @@ rewrite <- add_lt_mono_l.
destruct (mod_pos_bound a b); order.
Qed.
-(** With negative divisor, everything is upside-down *)
-
-Lemma mul_div_ge : forall a b, b<0 -> a <= b*(a/b).
-Proof.
-intros. rewrite <- div_opp_opp, opp_le_mono, <-mul_opp_l by order.
-apply mul_div_le. rewrite opp_pos_neg; auto.
-Qed.
-
Lemma mul_succ_div_lt: forall a b, b<0 -> b*(S (a/b)) < a.
Proof.
intros. rewrite <- div_opp_opp, opp_lt_mono, <-mul_opp_l by order.
apply mul_succ_div_gt. rewrite opp_pos_neg; auto.
Qed.
+(** 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).
diff --git a/theories/Numbers/Integer/Abstract/ZDivMath.v b/theories/Numbers/Integer/Abstract/ZDivMath.v
deleted file mode 100644
index dfc9ee4bc..000000000
--- a/theories/Numbers/Integer/Abstract/ZDivMath.v
+++ /dev/null
@@ -1,396 +0,0 @@
-(************************************************************************)
-(* 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/ZDivTrunc.v
index 73eebd6ae..fe2b262f2 100644
--- a/theories/Numbers/Integer/Abstract/ZDivOcaml.v
+++ b/theories/Numbers/Integer/Abstract/ZDivTrunc.v
@@ -6,12 +6,22 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(** Euclidean Division for integers
+(** * Euclidean Division for integers (Trunc convention)
- We use here the convention of Ocaml and many other system (C, ASM, ...),
- i.e. Round-Toward-Zero :
+ We use here the convention known as Trunc, or Round-Toward-Zero,
+ where [a/b] is the integer with the largest absolute value to
+ be between zero and the exact fraction. It can be summarized by:
- [a = bq+r /\ 0 < |r| < |b| /\ Sign(r) = Sgn(a)]
+ [a = bq+r /\ 0 <= |r| < |b| /\ Sign(r) = Sign(a)]
+
+ This is the convention of Ocaml and many other systems (C, ASM, ...).
+ This convention is named "T" 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 [ZDivFloor] and [ZDivEucl] for others conventions.
*)
Require Import ZAxioms ZProperties NZDiv.
@@ -291,9 +301,34 @@ Qed.
Lemma mul_succ_div_gt: forall a b, 0<=a -> 0<b -> a < b*(S (a/b)).
Proof. exact mul_succ_div_gt. Qed.
-(*TODO: CAS NEGATIF ... *)
+(** Similar results with negative numbers *)
+
+Lemma mul_pred_div_lt: forall a b, a<=0 -> 0<b -> b*(P (a/b)) < a.
+Proof.
+intros.
+rewrite opp_lt_mono, <- mul_opp_r, opp_pred, <- div_opp_l by order.
+apply mul_succ_div_gt; auto.
+rewrite opp_nonneg_nonpos; auto.
+Qed.
+
+Lemma mul_pred_div_gt: forall a b, 0<=a -> b<0 -> a < b*(P (a/b)).
+Proof.
+intros.
+rewrite <- mul_opp_opp, opp_pred, <- div_opp_r by order.
+apply mul_succ_div_gt; auto.
+rewrite opp_pos_neg; auto.
+Qed.
+
+Lemma mul_succ_div_lt: forall a b, a<=0 -> b<0 -> b*(S (a/b)) < a.
+Proof.
+intros.
+rewrite opp_lt_mono, <- mul_opp_l, <- div_opp_opp by order.
+apply mul_succ_div_gt; auto.
+rewrite opp_nonneg_nonpos; auto.
+rewrite opp_pos_neg; auto.
+Qed.
-(** Some previous inequalities are exact iff the modulo is zero. *)
+(** 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).
Proof.
diff --git a/theories/Numbers/NatInt/NZAxioms.v b/theories/Numbers/NatInt/NZAxioms.v
index c0f309329..ff220e793 100644
--- a/theories/Numbers/NatInt/NZAxioms.v
+++ b/theories/Numbers/NatInt/NZAxioms.v
@@ -5,8 +5,8 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* Evgeny Makarov, INRIA, 2007 *)
-(************************************************************************)
+
+(** Initial Author : Evgeny Makarov, INRIA, 2007 *)
(*i $Id$ i*)
diff --git a/theories/Numbers/NatInt/NZDomain.v b/theories/Numbers/NatInt/NZDomain.v
index 60f9287c8..2d627dbc7 100644
--- a/theories/Numbers/NatInt/NZDomain.v
+++ b/theories/Numbers/NatInt/NZDomain.v
@@ -5,8 +5,6 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* Evgeny Makarov, INRIA, 2007 *)
-(************************************************************************)
(*i $Id$ i*)
diff --git a/theories/Numbers/vo.itarget b/theories/Numbers/vo.itarget
index 7c29450ea..b6d2e42aa 100644
--- a/theories/Numbers/vo.itarget
+++ b/theories/Numbers/vo.itarget
@@ -22,9 +22,9 @@ Integer/Abstract/ZLt.vo
Integer/Abstract/ZMulOrder.vo
Integer/Abstract/ZMul.vo
Integer/Abstract/ZProperties.vo
-Integer/Abstract/ZDivCoq.vo
-Integer/Abstract/ZDivOcaml.vo
-Integer/Abstract/ZDivMath.vo
+Integer/Abstract/ZDivFloor.vo
+Integer/Abstract/ZDivTrunc.vo
+Integer/Abstract/ZDivEucl.vo
Integer/BigZ/BigZ.vo
Integer/BigZ/ZMake.vo
Integer/Binary/ZBinary.vo
diff --git a/theories/ZArith/ZOdiv.v b/theories/ZArith/ZOdiv.v
index cc272df5c..957a20ef8 100644
--- a/theories/ZArith/ZOdiv.v
+++ b/theories/ZArith/ZOdiv.v
@@ -8,7 +8,7 @@
Require Import BinPos BinNat Nnat ZArith_base ROmega ZArithRing
- ZBinary ZDivOcaml Morphisms.
+ ZBinary ZDivTrunc Morphisms.
Require Export ZOdiv_def.
Require Zdiv.
@@ -245,7 +245,7 @@ Proof.
Qed.
(** We know enough to prove that [ZOdiv] and [ZOmod] are instances of
- one of the abstract Euclidean divisions of Numbers. *)
+ one of the abstract Euclidean divisions of Numbers. *)
Module ZODiv <: ZDiv ZBinAxiomsMod.
Definition div := ZOdiv.