(************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* 0 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 ZDivSpecific. Module Type ZDiv (Z:ZAxiomsSig) := DivMod Z <+ NZDivCommon Z <+ ZDivSpecific Z. Module Type ZDivSig := ZAxiomsExtSig <+ ZDiv. Module Type ZDivSig' := ZAxiomsExtSig' <+ ZDiv <+ DivModNotation. Module ZDivPropFunct (Import Z : ZDivSig')(Import ZP : ZPropSig Z). (** We benefit from what already exists for NZ *) Module Import NZDivP := NZDivPropFunct Z ZP 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. now apply div_mod. Qed. (** A few sign rules (simple ones) *) Lemma mod_opp_opp : forall a b, b ~= 0 -> (-a) mod (-b) == - (a mod b). Proof. intros. now rewrite mod_opp_r, mod_opp_l. Qed. Lemma div_opp_l : forall a b, b ~= 0 -> (-a)/b == -(a/b). Proof. intros. rewrite <- (mul_cancel_l _ _ b) by trivial. rewrite <- (add_cancel_r _ _ ((-a) mod b)). now rewrite <- div_mod, mod_opp_l, mul_opp_r, <- opp_add_distr, <- div_mod. Qed. Lemma div_opp_r : forall a b, b ~= 0 -> a/(-b) == -(a/b). Proof. intros. assert (-b ~= 0) by (now rewrite eq_opp_l, opp_0). rewrite <- (mul_cancel_l _ _ (-b)) by trivial. rewrite <- (add_cancel_r _ _ (a mod (-b))). now rewrite <- div_mod, mod_opp_r, mul_opp_opp, <- div_mod. Qed. Lemma div_opp_opp : forall a b, b ~= 0 -> (-a)/(-b) == a/b. Proof. intros. now rewrite div_opp_r, div_opp_l, opp_involutive. 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 0 <= (a mod b) * a). intros. pos_or_neg a. apply mul_nonneg_nonneg; trivial. now destruct (mod_bound a b). rewrite <- mul_opp_opp, <- mod_opp_l by order. apply mul_nonneg_nonneg; try order. destruct (mod_bound (-a) b); 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 (0<=r2 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; trivial. rewrite <- (opp_inj_wd r1 r2). apply div_mod_unique with (-b); trivial. rewrite <- opp_lt_mono, opp_nonneg_nonpos; tauto. rewrite <- opp_lt_mono, opp_nonneg_nonpos; tauto. now rewrite 2 mul_opp_l, <- 2 opp_add_distr, opp_inj_wd. Qed. Theorem div_unique: forall a b q r, 0<=a -> 0<=r a == b*q + r -> q == a/b. Proof. intros; now apply div_unique with r. Qed. Theorem mod_unique: forall a b q r, 0<=a -> 0<=r a == b*q + r -> r == a mod b. Proof. intros; now apply mod_unique with q. 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 by trivial. now apply div_same. Qed. Lemma mod_same : forall a, a~=0 -> a mod a == 0. Proof. intros. rewrite mod_eq, div_same by trivial. 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 a/b == 0. Proof. exact div_small. Qed. (** Same situation, in term of modulo: *) Theorem mod_small: forall a b, 0<=a 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 by trivial. now apply div_0_l. Qed. Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0. Proof. intros; rewrite mod_eq, div_0_l; now nzsimpl. Qed. Lemma div_1_r: forall a, a/1 == a. Proof. intros. pos_or_neg a. now apply div_1_r. 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 1/a == 0. Proof. exact div_1_l. Qed. Lemma mod_1_l: forall a, 1 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 by trivial. 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 a mod b <= a. Proof. exact mod_le. Qed. Theorem div_pos : forall a b, 0<=a -> 0 0<= a/b. Proof. exact div_pos. Qed. Lemma div_str_pos : forall a b, 0 0 < a/b. Proof. exact div_str_pos. Qed. 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_eq; intuition; order. rewrite <- opp_inj_wd, opp_0, <- div_opp_r, div_small_iff by order. rewrite (abs_eq a), (abs_neq' b); intuition; order. rewrite <- opp_inj_wd, opp_0, <- div_opp_l, div_small_iff by order. rewrite (abs_neq' a), (abs_eq b); intuition; order. rewrite <- div_opp_opp, div_small_iff by order. rewrite (abs_neq' a), (abs_neq' 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. tauto. Qed. (** As soon as the divisor is strictly greater than 1, the division is strictly decreasing. *) Lemma div_lt : forall a b, 0 1 a/b < a. Proof. exact div_lt. Qed. (** [le] is compatible with a positive division. *) Lemma div_le_mono : forall a b c, 0 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. rewrite <- opp_nonneg_nonpos in *. destruct (mul_div_le (-a) b); tauto. 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 a < b*(S (a/b)). Proof. exact mul_succ_div_gt. Qed. (** Similar results with negative numbers *) Lemma mul_pred_div_lt: forall a b, a<=0 -> 0 b*(P (a/b)) < a. Proof. intros. rewrite opp_lt_mono, <- mul_opp_r, opp_pred, <- div_opp_l by order. rewrite <- opp_nonneg_nonpos in *. now apply mul_succ_div_gt. 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. rewrite <- opp_pos_neg in *. now apply mul_succ_div_gt. 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. rewrite <- opp_nonneg_nonpos, <- opp_pos_neg in *. now apply mul_succ_div_gt. Qed. (** 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 mod_eq by order. rewrite sub_move_r; nzsimpl; tauto. Qed. (** Some additionnal inequalities about div. *) Theorem div_lt_upper_bound: forall a b q, 0<=a -> 0 a < b*q -> a/b < q. Proof. exact div_lt_upper_bound. Qed. Theorem div_le_upper_bound: forall a b q, 0 a <= b*q -> a/b <= q. Proof. intros. rewrite <- (div_mul q b) by order. apply div_le_mono; trivial. now rewrite mul_comm. Qed. Theorem div_le_lower_bound: forall a b q, 0 b*q <= a -> q <= a/b. Proof. intros. rewrite <- (div_mul q b) by order. apply div_le_mono; trivial. now rewrite mul_comm. Qed. (** A division respects opposite monotonicity for the divisor *) Lemma div_le_compat_l: forall p q r, 0<=p -> 0 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 trivial. rewrite <- (add_cancel_r _ _ ((a+b*c) mod c)). rewrite <- div_mod, mod_add by trivial. now rewrite mul_add_distr_l, add_shuffle0, <-div_mod, mul_comm. 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). now apply div_add. 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 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); now apply div_mul_cancel_r. 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; tauto). rewrite ! mod_eq by trivial. rewrite div_mul_cancel_r by order. now rewrite mul_sub_distr_r, <- !mul_assoc, (mul_comm (a/b) c). 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); now apply mul_mod_distr_r. 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 trivial. 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. now apply Aux1. 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. now apply Aux2. 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). now apply mul_mod_idemp_l. Qed. Theorem mul_mod: forall a b n, n~=0 -> (a * b) mod n == ((a mod n) * (b mod n)) mod n. Proof. intros. now rewrite mul_mod_idemp_l, mul_mod_idemp_r. 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)]. now apply Aux. apply opp_inj. rewrite <-2 mod_opp_l, 2 opp_add_distr, <-mod_opp_l by order. rewrite <- opp_nonneg_nonpos in *. now apply Aux. 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; trivial. now rewrite mul_comm. 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; trivial. reflexivity. 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 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; trivial. 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; trivial. apply Aux1; trivial. 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. tauto. Qed. (** A last inequality: *) Theorem div_mul_le: forall a b c, 0<=a -> 0 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; now nzsimpl. intros (c,Hc). rewrite Hc, mul_comm. now apply mod_mul. Qed. End ZDivPropFunct.