summaryrefslogtreecommitdiff
path: root/theories/ZArith
diff options
context:
space:
mode:
Diffstat (limited to 'theories/ZArith')
-rw-r--r--theories/ZArith/BinInt.v2011
-rw-r--r--theories/ZArith/BinIntDef.v610
-rw-r--r--theories/ZArith/Int.v12
-rw-r--r--theories/ZArith/Wf_Z.v197
-rw-r--r--theories/ZArith/ZArith.v9
-rw-r--r--theories/ZArith/ZArith_base.v11
-rw-r--r--theories/ZArith/ZArith_dec.v13
-rw-r--r--theories/ZArith/ZOdiv.v947
-rw-r--r--theories/ZArith/ZOdiv_def.v136
-rw-r--r--theories/ZArith/ZOrderedType.v60
-rw-r--r--theories/ZArith/Zabs.v224
-rw-r--r--theories/ZArith/Zbool.v183
-rw-r--r--theories/ZArith/Zcompare.v457
-rw-r--r--theories/ZArith/Zcomplements.v143
-rw-r--r--theories/ZArith/Zdigits.v14
-rw-r--r--theories/ZArith/Zdiv.v897
-rw-r--r--theories/ZArith/Zeuclid.v52
-rw-r--r--theories/ZArith/Zeven.v371
-rw-r--r--theories/ZArith/Zgcd_alt.v8
-rw-r--r--theories/ZArith/Zhints.v441
-rw-r--r--theories/ZArith/Zlogarithm.v48
-rw-r--r--theories/ZArith/Zmax.v109
-rw-r--r--theories/ZArith/Zmin.v89
-rw-r--r--theories/ZArith/Zminmax.v188
-rw-r--r--theories/ZArith/Zmisc.v71
-rw-r--r--theories/ZArith/Znat.v1063
-rw-r--r--theories/ZArith/Znumtheory.v810
-rw-r--r--theories/ZArith/Zorder.v893
-rw-r--r--theories/ZArith/Zpow_alt.v83
-rw-r--r--theories/ZArith/Zpow_def.v42
-rw-r--r--theories/ZArith/Zpow_facts.v510
-rw-r--r--theories/ZArith/Zpower.v427
-rw-r--r--theories/ZArith/Zquot.v536
-rw-r--r--theories/ZArith/Zsqrt_compat.v (renamed from theories/ZArith/Zsqrt.v)24
-rw-r--r--theories/ZArith/Zwf.v4
-rw-r--r--theories/ZArith/auxiliary.v87
-rw-r--r--theories/ZArith/vo.itarget9
37 files changed, 5067 insertions, 6722 deletions
diff --git a/theories/ZArith/BinInt.v b/theories/ZArith/BinInt.v
index e2b89d84..3a5eb885 100644
--- a/theories/ZArith/BinInt.v
+++ b/theories/ZArith/BinInt.v
@@ -1,1158 +1,1585 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: BinInt.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
+Require Export BinNums BinPos Pnat.
+Require Import BinNat Bool Plus Mult Equalities GenericMinMax
+ OrdersFacts ZAxioms ZProperties.
+Require BinIntDef.
(***********************************************************)
-(** Binary Integers (Pierre Crégut, CNET, Lannion, France) *)
+(** * Binary Integers *)
(***********************************************************)
-Require Export BinPos.
-Require Export Pnat.
-Require Import BinNat.
-Require Import Plus.
-Require Import Mult.
+(** Initial author: Pierre Crégut, CNET, Lannion, France *)
-Unset Boxed Definitions.
+(** The type [Z] and its constructors [Z0] and [Zpos] and [Zneg]
+ are now defined in [BinNums.v] *)
-(*****************************)
-(** * Binary integer numbers *)
+Local Open Scope Z_scope.
-Inductive Z : Set :=
- | Z0 : Z
- | Zpos : positive -> Z
- | Zneg : positive -> Z.
+(** Every definitions and early properties about binary integers
+ are placed in a module [Z] for qualification purpose. *)
+Module Z
+ <: ZAxiomsSig
+ <: UsualOrderedTypeFull
+ <: UsualDecidableTypeFull
+ <: TotalOrder.
-(** Automatically open scope positive_scope for the constructors of Z *)
-Delimit Scope Z_scope with Z.
-Bind Scope Z_scope with Z.
-Arguments Scope Zpos [positive_scope].
-Arguments Scope Zneg [positive_scope].
-
-(** ** Subtraction of positive into Z *)
-
-Definition Zdouble_plus_one (x:Z) :=
- match x with
- | Z0 => Zpos 1
- | Zpos p => Zpos p~1
- | Zneg p => Zneg (Pdouble_minus_one p)
- end.
-
-Definition Zdouble_minus_one (x:Z) :=
- match x with
- | Z0 => Zneg 1
- | Zneg p => Zneg p~1
- | Zpos p => Zpos (Pdouble_minus_one p)
- end.
-
-Definition Zdouble (x:Z) :=
- match x with
- | Z0 => Z0
- | Zpos p => Zpos p~0
- | Zneg p => Zneg p~0
- end.
-
-Open Local Scope positive_scope.
-
-Fixpoint ZPminus (x y:positive) {struct y} : Z :=
- match x, y with
- | p~1, q~1 => Zdouble (ZPminus p q)
- | p~1, q~0 => Zdouble_plus_one (ZPminus p q)
- | p~1, 1 => Zpos p~0
- | p~0, q~1 => Zdouble_minus_one (ZPminus p q)
- | p~0, q~0 => Zdouble (ZPminus p q)
- | p~0, 1 => Zpos (Pdouble_minus_one p)
- | 1, q~1 => Zneg q~0
- | 1, q~0 => Zneg (Pdouble_minus_one q)
- | 1, 1 => Z0
- end.
-
-Close Local Scope positive_scope.
-
-(** ** Addition on integers *)
-
-Definition Zplus (x y:Z) :=
- match x, y with
- | Z0, y => y
- | Zpos x', Z0 => Zpos x'
- | Zneg x', Z0 => Zneg x'
- | Zpos x', Zpos y' => Zpos (x' + y')
- | Zpos x', Zneg y' =>
- match (x' ?= y')%positive Eq with
- | Eq => Z0
- | Lt => Zneg (y' - x')
- | Gt => Zpos (x' - y')
- end
- | Zneg x', Zpos y' =>
- match (x' ?= y')%positive Eq with
- | Eq => Z0
- | Lt => Zpos (y' - x')
- | Gt => Zneg (x' - y')
- end
- | Zneg x', Zneg y' => Zneg (x' + y')
- end.
-
-Infix "+" := Zplus : Z_scope.
-
-(** ** Opposite *)
-
-Definition Zopp (x:Z) :=
- match x with
- | Z0 => Z0
- | Zpos x => Zneg x
- | Zneg x => Zpos x
- end.
-
-Notation "- x" := (Zopp x) : Z_scope.
-
-(** ** Successor on integers *)
-
-Definition Zsucc (x:Z) := (x + Zpos 1)%Z.
-
-(** ** Predecessor on integers *)
-
-Definition Zpred (x:Z) := (x + Zneg 1)%Z.
-
-(** ** Subtraction on integers *)
-
-Definition Zminus (m n:Z) := (m + - n)%Z.
-
-Infix "-" := Zminus : Z_scope.
-
-(** ** Multiplication on integers *)
-
-Definition Zmult (x y:Z) :=
- match x, y with
- | Z0, _ => Z0
- | _, Z0 => Z0
- | Zpos x', Zpos y' => Zpos (x' * y')
- | Zpos x', Zneg y' => Zneg (x' * y')
- | Zneg x', Zpos y' => Zneg (x' * y')
- | Zneg x', Zneg y' => Zpos (x' * y')
- end.
-
-Infix "*" := Zmult : Z_scope.
-
-(** ** Comparison of integers *)
-
-Definition Zcompare (x y:Z) :=
- match x, y with
- | Z0, Z0 => Eq
- | Z0, Zpos y' => Lt
- | Z0, Zneg y' => Gt
- | Zpos x', Z0 => Gt
- | Zpos x', Zpos y' => (x' ?= y')%positive Eq
- | Zpos x', Zneg y' => Gt
- | Zneg x', Z0 => Lt
- | Zneg x', Zpos y' => Lt
- | Zneg x', Zneg y' => CompOpp ((x' ?= y')%positive Eq)
- end.
-
-Infix "?=" := Zcompare (at level 70, no associativity) : Z_scope.
+(** * Definitions of operations, now in a separate file *)
-Ltac elim_compare com1 com2 :=
- case (Dcompare (com1 ?= com2)%Z);
- [ idtac | let x := fresh "H" in
- (intro x; case x; clear x) ].
+Include BinIntDef.Z.
+
+(** When including property functors, only inline t eq zero one two *)
-(** ** Sign function *)
+Set Inline Level 30.
-Definition Zsgn (z:Z) : Z :=
- match z with
- | Z0 => Z0
- | Zpos p => Zpos 1
- | Zneg p => Zneg 1
- end.
+(** * Logic Predicates *)
+
+Definition eq := @Logic.eq Z.
+Definition eq_equiv := @eq_equivalence Z.
+
+Definition lt x y := (x ?= y) = Lt.
+Definition gt x y := (x ?= y) = Gt.
+Definition le x y := (x ?= y) <> Gt.
+Definition ge x y := (x ?= y) <> Lt.
+
+Infix "<=" := le : Z_scope.
+Infix "<" := lt : Z_scope.
+Infix ">=" := ge : Z_scope.
+Infix ">" := gt : Z_scope.
+
+Notation "x <= y <= z" := (x <= y /\ y <= z) : Z_scope.
+Notation "x <= y < z" := (x <= y /\ y < z) : Z_scope.
+Notation "x < y < z" := (x < y /\ y < z) : Z_scope.
+Notation "x < y <= z" := (x < y /\ y <= z) : Z_scope.
-(** ** Direct, easier to handle variants of successor and addition *)
+Definition divide x y := exists z, y = z*x.
+Notation "( x | y )" := (divide x y) (at level 0).
-Definition Zsucc' (x:Z) :=
- match x with
- | Z0 => Zpos 1
- | Zpos x' => Zpos (Psucc x')
- | Zneg x' => ZPminus 1 x'
- end.
+Definition Even a := exists b, a = 2*b.
+Definition Odd a := exists b, a = 2*b+1.
-Definition Zpred' (x:Z) :=
- match x with
- | Z0 => Zneg 1
- | Zpos x' => ZPminus x' 1
- | Zneg x' => Zneg (Psucc x')
- end.
+(** * Decidability of equality. *)
-Definition Zplus' (x y:Z) :=
- match x, y with
- | Z0, y => y
- | x, Z0 => x
- | Zpos x', Zpos y' => Zpos (x' + y')
- | Zpos x', Zneg y' => ZPminus x' y'
- | Zneg x', Zpos y' => ZPminus y' x'
- | Zneg x', Zneg y' => Zneg (x' + y')
- end.
+Definition eq_dec (x y : Z) : {x = y} + {x <> y}.
+Proof.
+ decide equality; apply Pos.eq_dec.
+Defined.
-Open Local Scope Z_scope.
+(** * Properties of [pos_sub] *)
-(**********************************************************************)
-(** ** Inductive specification of Z *)
+(** [pos_sub] can be written in term of positive comparison
+ and subtraction (cf. earlier definition of addition of Z) *)
-Theorem Zind :
- forall P:Z -> Prop,
- P Z0 ->
- (forall x:Z, P x -> P (Zsucc' x)) ->
- (forall x:Z, P x -> P (Zpred' x)) -> forall n:Z, P n.
+Lemma pos_sub_spec p q :
+ pos_sub p q =
+ match (p ?= q)%positive with
+ | Eq => 0
+ | Lt => Zneg (q - p)
+ | Gt => Zpos (p - q)
+ end.
Proof.
- intros P H0 Hs Hp z; destruct z.
- assumption.
- apply Pind with (P := fun p => P (Zpos p)).
- change (P (Zsucc' Z0)) in |- *; apply Hs; apply H0.
- intro n; exact (Hs (Zpos n)).
- apply Pind with (P := fun p => P (Zneg p)).
- change (P (Zpred' Z0)) in |- *; apply Hp; apply H0.
- intro n; exact (Hp (Zneg n)).
+ revert q. induction p; destruct q; simpl; trivial;
+ rewrite ?Pos.compare_xI_xI, ?Pos.compare_xO_xI,
+ ?Pos.compare_xI_xO, ?Pos.compare_xO_xO, IHp; simpl;
+ case Pos.compare_spec; intros; simpl; trivial;
+ (now rewrite Pos.sub_xI_xI) || (now rewrite Pos.sub_xO_xO) ||
+ (now rewrite Pos.sub_xO_xI) || (now rewrite Pos.sub_xI_xO) ||
+ subst; unfold Pos.sub; simpl; now rewrite Pos.sub_mask_diag.
Qed.
-(**********************************************************************)
-(** * Misc properties about binary integer operations *)
+(** Particular cases of the previous result *)
+Lemma pos_sub_diag p : pos_sub p p = 0.
+Proof.
+ now rewrite pos_sub_spec, Pos.compare_refl.
+Qed.
-(**********************************************************************)
-(** ** Properties of opposite on binary integer numbers *)
+Lemma pos_sub_lt p q : (p < q)%positive -> pos_sub p q = Zneg (q - p).
+Proof.
+ intros H. now rewrite pos_sub_spec, H.
+Qed.
-Theorem Zopp_0 : Zopp Z0 = Z0.
+Lemma pos_sub_gt p q : (q < p)%positive -> pos_sub p q = Zpos (p - q).
Proof.
- reflexivity.
+ intros H. now rewrite pos_sub_spec, Pos.compare_antisym, H.
Qed.
-Theorem Zopp_neg : forall p:positive, - Zneg p = Zpos p.
+(** The opposite of [pos_sub] is [pos_sub] with reversed arguments *)
+
+Lemma pos_sub_opp p q : - pos_sub p q = pos_sub q p.
Proof.
- reflexivity.
+ revert q; induction p; destruct q; simpl; trivial;
+ rewrite <- IHp; now destruct pos_sub.
Qed.
-(** [opp] is involutive *)
+(** * Results concerning [Zpos] and [Zneg] and the operators *)
-Theorem Zopp_involutive : forall n:Z, - - n = n.
+Lemma opp_Zneg p : - Zneg p = Zpos p.
Proof.
- intro x; destruct x; reflexivity.
+ reflexivity.
Qed.
-(** Injectivity of the opposite *)
+Lemma opp_Zpos p : - Zpos p = Zneg p.
+Proof.
+ reflexivity.
+Qed.
-Theorem Zopp_inj : forall n m:Z, - n = - m -> n = m.
+Lemma succ_Zpos p : succ (Zpos p) = Zpos (Pos.succ p).
Proof.
- intros x y; case x; case y; simpl in |- *; intros;
- [ trivial
- | discriminate H
- | discriminate H
- | discriminate H
- | simplify_eq H; intro E; rewrite E; trivial
- | discriminate H
- | discriminate H
- | discriminate H
- | simplify_eq H; intro E; rewrite E; trivial ].
+ simpl. f_equal. apply Pos.add_1_r.
Qed.
-(**********************************************************************)
-(** ** Other properties of binary integer numbers *)
+Lemma add_Zpos p q : Zpos p + Zpos q = Zpos (p+q).
+Proof.
+ reflexivity.
+Qed.
-Lemma ZL0 : 2%nat = (1 + 1)%nat.
+Lemma add_Zneg p q : Zneg p + Zneg q = Zneg (p+q).
Proof.
- reflexivity.
+ reflexivity.
Qed.
-(**********************************************************************)
-(** * Properties of the addition on integers *)
+Lemma add_Zpos_Zneg p q : Zpos p + Zneg q = pos_sub p q.
+Proof.
+ reflexivity.
+Qed.
-(** ** Zero is left neutral for addition *)
+Lemma add_Zneg_Zpos p q : Zneg p + Zpos q = pos_sub q p.
+Proof.
+ reflexivity.
+Qed.
-Theorem Zplus_0_l : forall n:Z, Z0 + n = n.
+Lemma sub_Zpos n m : (n < m)%positive -> Zpos m - Zpos n = Zpos (m-n).
Proof.
- intro x; destruct x; reflexivity.
+ intros H. simpl. now apply pos_sub_gt.
Qed.
-(** ** Zero is right neutral for addition *)
+Lemma mul_Zpos (p q : positive) : Zpos p * Zpos q = Zpos (p*q).
+Proof.
+ reflexivity.
+Qed.
-Theorem Zplus_0_r : forall n:Z, n + Z0 = n.
+Lemma pow_Zpos p q : (Zpos p)^(Zpos q) = Zpos (p^q).
Proof.
- intro x; destruct x; reflexivity.
+ unfold Pos.pow, pow, pow_pos.
+ symmetry. now apply Pos.iter_swap_gen.
Qed.
-(** ** Addition is commutative *)
+Lemma inj_Zpos p q : Zpos p = Zpos q <-> p = q.
+Proof.
+ split; intros H. now injection H. now f_equal.
+Qed.
-Theorem Zplus_comm : forall n m:Z, n + m = m + n.
+Lemma inj_Zneg p q : Zneg p = Zneg q <-> p = q.
Proof.
- intro x; induction x as [| p| p]; intro y; destruct y as [| q| q];
- simpl in |- *; try reflexivity.
- rewrite Pplus_comm; reflexivity.
- rewrite ZC4; destruct ((q ?= p)%positive Eq); reflexivity.
- rewrite ZC4; destruct ((q ?= p)%positive Eq); reflexivity.
- rewrite Pplus_comm; reflexivity.
+ split; intros H. now injection H. now f_equal.
Qed.
-(** ** Opposite distributes over addition *)
+Lemma pos_xI p : Zpos p~1 = 2 * Zpos p + 1.
+Proof.
+ reflexivity.
+Qed.
-Theorem Zopp_plus_distr : forall n m:Z, - (n + m) = - n + - m.
+Lemma pos_xO p : Zpos p~0 = 2 * Zpos p.
Proof.
- intro x; destruct x as [| p| p]; intro y; destruct y as [| q| q];
- simpl in |- *; reflexivity || destruct ((p ?= q)%positive Eq);
- reflexivity.
+ reflexivity.
Qed.
-Theorem Zopp_succ : forall n:Z, Zopp (Zsucc n) = Zpred (Zopp n).
+Lemma neg_xI p : Zneg p~1 = 2 * Zneg p - 1.
Proof.
-intro; unfold Zsucc; now rewrite Zopp_plus_distr.
+ reflexivity.
Qed.
-(** ** Opposite is inverse for addition *)
+Lemma neg_xO p : Zneg p~0 = 2 * Zneg p.
+Proof.
+ reflexivity.
+Qed.
+
+(** In the following module, we group results that are needed now
+ to prove specifications of operations, but will also be provided
+ later by the generic functor of properties. *)
+
+Module Import Private_BootStrap.
+
+(** * Properties of addition *)
+
+(** ** Zero is neutral for addition *)
-Theorem Zplus_opp_r : forall n:Z, n + - n = Z0.
+Lemma add_0_r n : n + 0 = n.
Proof.
- intro x; destruct x as [| p| p]; simpl in |- *;
- [ reflexivity
- | rewrite (Pcompare_refl p); reflexivity
- | rewrite (Pcompare_refl p); reflexivity ].
+ now destruct n.
Qed.
-Theorem Zplus_opp_l : forall n:Z, - n + n = Z0.
+(** ** Addition is commutative *)
+
+Lemma add_comm n m : n + m = m + n.
+Proof.
+ destruct n, m; simpl; trivial; now rewrite Pos.add_comm.
+Qed.
+
+(** ** Opposite distributes over addition *)
+
+Lemma opp_add_distr n m : - (n + m) = - n + - m.
Proof.
- intro; rewrite Zplus_comm; apply Zplus_opp_r.
+ destruct n, m; simpl; trivial using pos_sub_opp.
Qed.
-Hint Local Resolve Zplus_0_l Zplus_0_r.
+(** ** Opposite is injective *)
+
+Lemma opp_inj n m : -n = -m -> n = m.
+Proof.
+ destruct n, m; simpl; intros H; destr_eq H; now f_equal.
+Qed.
(** ** Addition is associative *)
-Lemma weak_assoc :
- forall (p q:positive) (n:Z), Zpos p + (Zpos q + n) = Zpos p + Zpos q + n.
-Proof.
- intros x y z'; case z';
- [ auto with arith
- | intros z; simpl in |- *; rewrite Pplus_assoc; auto with arith
- | intros z; simpl in |- *; ElimPcompare y z; intros E0; rewrite E0;
- ElimPcompare (x + y)%positive z; intros E1; rewrite E1;
- [ absurd ((x + y ?= z)%positive Eq = Eq);
- [ (* Case 1 *)
- rewrite nat_of_P_gt_Gt_compare_complement_morphism;
- [ discriminate
- | rewrite nat_of_P_plus_morphism; rewrite (Pcompare_Eq_eq y z E0);
- elim (ZL4 x); intros k E2; rewrite E2;
- simpl in |- *; unfold gt, lt in |- *;
- apply le_n_S; apply le_plus_r ]
- | assumption ]
- | absurd ((x + y ?= z)%positive Eq = Lt);
- [ (* Case 2 *)
- rewrite nat_of_P_gt_Gt_compare_complement_morphism;
- [ discriminate
- | rewrite nat_of_P_plus_morphism; rewrite (Pcompare_Eq_eq y z E0);
- elim (ZL4 x); intros k E2; rewrite E2;
- simpl in |- *; unfold gt, lt in |- *;
- apply le_n_S; apply le_plus_r ]
- | assumption ]
- | rewrite (Pcompare_Eq_eq y z E0);
- (* Case 3 *)
- elim (Pminus_mask_Gt (x + z) z);
- [ intros t H; elim H; intros H1 H2; elim H2; intros H3 H4;
- unfold Pminus in |- *; rewrite H1; cut (x = t);
- [ intros E; rewrite E; auto with arith
- | apply Pplus_reg_r with (r := z); rewrite <- H3;
- rewrite Pplus_comm; trivial with arith ]
- | pattern z at 1 in |- *; rewrite <- (Pcompare_Eq_eq y z E0);
- assumption ]
- | elim (Pminus_mask_Gt z y);
- [ (* Case 4 *)
- intros k H; elim H; intros H1 H2; elim H2; intros H3 H4;
- unfold Pminus at 1 in |- *; rewrite H1; cut (x = k);
- [ intros E; rewrite E; rewrite (Pcompare_refl k);
- trivial with arith
- | apply Pplus_reg_r with (r := y); rewrite (Pplus_comm k y);
- rewrite H3; apply Pcompare_Eq_eq; assumption ]
- | apply ZC2; assumption ]
- | elim (Pminus_mask_Gt z y);
- [ (* Case 5 *)
- intros k H; elim H; intros H1 H2; elim H2; intros H3 H4;
- unfold Pminus at 1 3 5 in |- *; rewrite H1;
- cut ((x ?= k)%positive Eq = Lt);
- [ intros E2; rewrite E2; elim (Pminus_mask_Gt k x);
- [ intros i H5; elim H5; intros H6 H7; elim H7; intros H8 H9;
- elim (Pminus_mask_Gt z (x + y));
- [ intros j H10; elim H10; intros H11 H12; elim H12;
- intros H13 H14; unfold Pminus in |- *;
- rewrite H6; rewrite H11; cut (i = j);
- [ intros E; rewrite E; auto with arith
- | apply (Pplus_reg_l (x + y)); rewrite H13;
- rewrite (Pplus_comm x y); rewrite <- Pplus_assoc;
- rewrite H8; assumption ]
- | apply ZC2; assumption ]
- | apply ZC2; assumption ]
- | apply nat_of_P_lt_Lt_compare_complement_morphism;
- apply plus_lt_reg_l with (p := nat_of_P y);
- do 2 rewrite <- nat_of_P_plus_morphism;
- apply nat_of_P_lt_Lt_compare_morphism;
- rewrite H3; rewrite Pplus_comm; assumption ]
- | apply ZC2; assumption ]
- | elim (Pminus_mask_Gt z y);
- [ (* Case 6 *)
- intros k H; elim H; intros H1 H2; elim H2; intros H3 H4;
- elim (Pminus_mask_Gt (x + y) z);
- [ intros i H5; elim H5; intros H6 H7; elim H7; intros H8 H9;
- unfold Pminus in |- *; rewrite H1; rewrite H6;
- cut ((x ?= k)%positive Eq = Gt);
- [ intros H10; elim (Pminus_mask_Gt x k H10); intros j H11;
- elim H11; intros H12 H13; elim H13;
- intros H14 H15; rewrite H10; rewrite H12;
- cut (i = j);
- [ intros H16; rewrite H16; auto with arith
- | apply (Pplus_reg_l (z + k)); rewrite <- (Pplus_assoc z k j);
- rewrite H14; rewrite (Pplus_comm z k);
- rewrite <- Pplus_assoc; rewrite H8;
- rewrite (Pplus_comm x y); rewrite Pplus_assoc;
- rewrite (Pplus_comm k y); rewrite H3;
- trivial with arith ]
- | apply nat_of_P_gt_Gt_compare_complement_morphism;
- unfold lt, gt in |- *;
- apply plus_lt_reg_l with (p := nat_of_P y);
- do 2 rewrite <- nat_of_P_plus_morphism;
- apply nat_of_P_lt_Lt_compare_morphism;
- rewrite H3; rewrite Pplus_comm; apply ZC1;
- assumption ]
- | assumption ]
- | apply ZC2; assumption ]
- | absurd ((x + y ?= z)%positive Eq = Eq);
- [ (* Case 7 *)
- rewrite nat_of_P_gt_Gt_compare_complement_morphism;
- [ discriminate
- | rewrite nat_of_P_plus_morphism; unfold gt in |- *;
- apply lt_le_trans with (m := nat_of_P y);
- [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; assumption
- | apply le_plus_r ] ]
- | assumption ]
- | absurd ((x + y ?= z)%positive Eq = Lt);
- [ (* Case 8 *)
- rewrite nat_of_P_gt_Gt_compare_complement_morphism;
- [ discriminate
- | unfold gt in |- *; apply lt_le_trans with (m := nat_of_P y);
- [ exact (nat_of_P_gt_Gt_compare_morphism y z E0)
- | rewrite nat_of_P_plus_morphism; apply le_plus_r ] ]
- | assumption ]
- | elim Pminus_mask_Gt with (1 := E0); intros k H1;
- (* Case 9 *)
- elim Pminus_mask_Gt with (1 := E1); intros i H2;
- elim H1; intros H3 H4; elim H4; intros H5 H6;
- elim H2; intros H7 H8; elim H8; intros H9 H10;
- unfold Pminus in |- *; rewrite H3; rewrite H7;
- cut ((x + k)%positive = i);
- [ intros E; rewrite E; auto with arith
- | apply (Pplus_reg_l z); rewrite (Pplus_comm x k); rewrite Pplus_assoc;
- rewrite H5; rewrite H9; rewrite Pplus_comm;
- trivial with arith ] ] ].
-Qed.
-
-Hint Local Resolve weak_assoc.
-
-Theorem Zplus_assoc : forall n m p:Z, n + (m + p) = n + m + p.
-Proof.
- intros x y z; case x; case y; case z; auto with arith; intros;
- [ rewrite (Zplus_comm (Zneg p0)); rewrite weak_assoc;
- rewrite (Zplus_comm (Zpos p1 + Zneg p0)); rewrite weak_assoc;
- rewrite (Zplus_comm (Zpos p1)); trivial with arith
- | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg;
- rewrite Zplus_comm; rewrite <- weak_assoc;
- rewrite (Zplus_comm (- Zpos p1));
- rewrite (Zplus_comm (Zpos p0 + - Zpos p1)); rewrite (weak_assoc p);
- rewrite weak_assoc; rewrite (Zplus_comm (Zpos p0));
- trivial with arith
- | rewrite Zplus_comm; rewrite (Zplus_comm (Zpos p0) (Zpos p));
- rewrite <- weak_assoc; rewrite Zplus_comm; rewrite (Zplus_comm (Zpos p0));
- trivial with arith
- | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg;
- rewrite (Zplus_comm (- Zpos p0)); rewrite weak_assoc;
- rewrite (Zplus_comm (Zpos p1 + - Zpos p0)); rewrite weak_assoc;
- rewrite (Zplus_comm (Zpos p)); trivial with arith
- | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg;
- apply weak_assoc
- | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg;
- apply weak_assoc ].
-Qed.
-
-
-Lemma Zplus_assoc_reverse : forall n m p:Z, n + m + p = n + (m + p).
-Proof.
- intros; symmetry in |- *; apply Zplus_assoc.
-Qed.
-
-(** ** Associativity mixed with commutativity *)
-
-Theorem Zplus_permute : forall n m p:Z, n + (m + p) = m + (n + p).
-Proof.
- intros n m p; rewrite Zplus_comm; rewrite <- Zplus_assoc;
- rewrite (Zplus_comm p n); trivial with arith.
-Qed.
-
-(** ** Addition simplifies *)
-
-Theorem Zplus_reg_l : forall n m p:Z, n + m = n + p -> m = p.
- intros n m p H; cut (- n + (n + m) = - n + (n + p));
- [ do 2 rewrite Zplus_assoc; rewrite (Zplus_comm (- n) n);
- rewrite Zplus_opp_r; simpl in |- *; trivial with arith
- | rewrite H; trivial with arith ].
-Qed.
-
-(** ** Addition and successor permutes *)
-
-Lemma Zplus_succ_l : forall n m:Z, Zsucc n + m = Zsucc (n + m).
-Proof.
- intros x y; unfold Zsucc in |- *; rewrite (Zplus_comm (x + y));
- rewrite Zplus_assoc; rewrite (Zplus_comm (Zpos 1));
- trivial with arith.
-Qed.
-
-Lemma Zplus_succ_r_reverse : forall n m:Z, Zsucc (n + m) = n + Zsucc m.
-Proof.
- intros n m; unfold Zsucc in |- *; rewrite Zplus_assoc; trivial with arith.
+Lemma pos_sub_add p q r :
+ pos_sub (p + q) r = Zpos p + pos_sub q r.
+Proof.
+ simpl. rewrite !pos_sub_spec.
+ case (Pos.compare_spec q r); intros E0.
+ - (* q = r *)
+ subst.
+ assert (H := Pos.lt_add_r r p).
+ rewrite Pos.add_comm in H. apply Pos.lt_gt in H.
+ now rewrite H, Pos.add_sub.
+ - (* q < r *)
+ rewrite pos_sub_spec.
+ assert (Hr : (r = (r-q)+q)%positive) by (now rewrite Pos.sub_add).
+ rewrite Hr at 1. rewrite Pos.add_compare_mono_r.
+ case Pos.compare_spec; intros E1; trivial; f_equal.
+ rewrite Pos.add_comm. apply Pos.sub_add_distr.
+ rewrite Hr, Pos.add_comm. now apply Pos.add_lt_mono_r.
+ symmetry. apply Pos.sub_sub_distr; trivial.
+ - (* r < q *)
+ assert (LT : (r < p + q)%positive).
+ { apply Pos.lt_trans with q; trivial.
+ rewrite Pos.add_comm. apply Pos.lt_add_r. }
+ apply Pos.lt_gt in LT. rewrite LT. f_equal.
+ symmetry. now apply Pos.add_sub_assoc.
+Qed.
+
+Lemma add_assoc n m p : n + (m + p) = n + m + p.
+Proof.
+ assert (AUX : forall x y z, Zpos x + (y + z) = Zpos x + y + z).
+ { intros x [|y|y] [|z|z]; rewrite ?add_0_r; trivial.
+ - simpl. now rewrite Pos.add_assoc.
+ - simpl (_ + Zneg _). symmetry. apply pos_sub_add.
+ - simpl (Zneg _ + _); simpl (_ + Zneg _).
+ now rewrite (add_comm _ (Zpos _)), <- 2 pos_sub_add, Pos.add_comm.
+ - apply opp_inj. rewrite !opp_add_distr, opp_Zpos, !opp_Zneg.
+ simpl (Zneg _ + _); simpl (_ + Zneg _).
+ rewrite add_comm, Pos.add_comm. apply pos_sub_add. }
+ destruct n.
+ - trivial.
+ - apply AUX.
+ - apply opp_inj. rewrite !opp_add_distr, opp_Zneg. apply AUX.
+Qed.
+
+(** ** Subtraction and successor *)
+
+Lemma sub_succ_l n m : succ n - m = succ (n - m).
+Proof.
+ unfold sub, succ. now rewrite <- 2 add_assoc, (add_comm 1).
Qed.
-Notation Zplus_succ_r := Zplus_succ_r_reverse (only parsing).
+(** ** Opposite is inverse for addition *)
-Lemma Zplus_succ_comm : forall n m:Z, Zsucc n + m = n + Zsucc m.
+Lemma add_opp_diag_r n : n + - n = 0.
Proof.
- unfold Zsucc in |- *; intros n m; rewrite <- Zplus_assoc;
- rewrite (Zplus_comm (Zpos 1)); trivial with arith.
+ destruct n; simpl; trivial; now rewrite pos_sub_diag.
Qed.
-(** ** Misc properties, usually redundant or non natural *)
-
-Lemma Zplus_0_r_reverse : forall n:Z, n = n + Z0.
+Lemma add_opp_diag_l n : - n + n = 0.
Proof.
- symmetry in |- *; apply Zplus_0_r.
+ rewrite add_comm. apply add_opp_diag_r.
Qed.
-Lemma Zplus_0_simpl_l : forall n m:Z, n + Z0 = m -> n = m.
+(** ** Commutativity of multiplication *)
+
+Lemma mul_comm n m : n * m = m * n.
Proof.
- intros n m; rewrite Zplus_0_r; intro; assumption.
+ destruct n, m; simpl; trivial; f_equal; apply Pos.mul_comm.
Qed.
-Lemma Zplus_0_simpl_l_reverse : forall n m:Z, n = m + Z0 -> n = m.
+(** ** Associativity of multiplication *)
+
+Lemma mul_assoc n m p : n * (m * p) = n * m * p.
Proof.
- intros n m; rewrite Zplus_0_r; intro; assumption.
+ destruct n, m, p; simpl; trivial; f_equal; apply Pos.mul_assoc.
Qed.
-Lemma Zplus_eq_compat : forall n m p q:Z, n = m -> p = q -> n + p = m + q.
+(** Multiplication and constants *)
+
+Lemma mul_1_l n : 1 * n = n.
Proof.
- intros; rewrite H; rewrite H0; reflexivity.
+ now destruct n.
Qed.
-Lemma Zplus_opp_expand : forall n m p:Z, n + - m = n + - p + (p + - m).
+Lemma mul_1_r n : n * 1 = n.
Proof.
- intros x y z.
- rewrite <- (Zplus_assoc x).
- rewrite (Zplus_assoc (- z)).
- rewrite Zplus_opp_l.
- reflexivity.
+ destruct n; simpl; now rewrite ?Pos.mul_1_r.
Qed.
-(************************************************************************)
-(** * Properties of successor and predecessor on binary integer numbers *)
+(** ** Multiplication and Opposite *)
-Theorem Zsucc_discr : forall n:Z, n <> Zsucc n.
+Lemma mul_opp_l n m : - n * m = - (n * m).
Proof.
- intros n; cut (Z0 <> Zpos 1);
- [ unfold not in |- *; intros H1 H2; apply H1; apply (Zplus_reg_l n);
- rewrite Zplus_0_r; exact H2
- | discriminate ].
+ now destruct n, m.
Qed.
-Theorem Zpos_succ_morphism :
- forall p:positive, Zpos (Psucc p) = Zsucc (Zpos p).
+Lemma mul_opp_r n m : n * - m = - (n * m).
Proof.
- intro; rewrite Pplus_one_succ_r; unfold Zsucc in |- *; simpl in |- *;
- trivial with arith.
+ now destruct n, m.
Qed.
-(** ** Successor and predecessor are inverse functions *)
+Lemma mul_opp_opp n m : - n * - m = n * m.
+Proof.
+ now destruct n, m.
+Qed.
-Theorem Zsucc_pred : forall n:Z, n = Zsucc (Zpred n).
+Lemma mul_opp_comm n m : - n * m = n * - m.
Proof.
- intros n; unfold Zsucc, Zpred in |- *; rewrite <- Zplus_assoc; simpl in |- *;
- rewrite Zplus_0_r; trivial with arith.
+ now destruct n, m.
Qed.
-Hint Immediate Zsucc_pred: zarith.
+(** ** Distributivity of multiplication over addition *)
+
+Lemma mul_add_distr_pos (p:positive) n m :
+ Zpos p * (n + m) = Zpos p * n + Zpos p * m.
+Proof.
+ destruct n as [|n|n], m as [|m|m]; simpl; trivial;
+ rewrite ?pos_sub_spec, ?Pos.mul_compare_mono_l; try case Pos.compare_spec;
+ intros; now rewrite ?Pos.mul_add_distr_l, ?Pos.mul_sub_distr_l.
+Qed.
-Theorem Zpred_succ : forall n:Z, n = Zpred (Zsucc n).
+Lemma mul_add_distr_l n m p : n * (m + p) = n * m + n * p.
Proof.
- intros m; unfold Zpred, Zsucc in |- *; rewrite <- Zplus_assoc; simpl in |- *;
- rewrite Zplus_comm; auto with arith.
+ destruct n as [|n|n]. trivial.
+ apply mul_add_distr_pos.
+ rewrite <- opp_Zpos, !mul_opp_l, <- opp_add_distr. f_equal.
+ apply mul_add_distr_pos.
Qed.
-Theorem Zsucc_inj : forall n m:Z, Zsucc n = Zsucc m -> n = m.
+Lemma mul_add_distr_r n m p : (n + m) * p = n * p + m * p.
Proof.
- intros n m H.
- change (Zneg 1 + Zpos 1 + n = Zneg 1 + Zpos 1 + m) in |- *;
- do 2 rewrite <- Zplus_assoc; do 2 rewrite (Zplus_comm (Zpos 1));
- unfold Zsucc in H; rewrite H; trivial with arith.
+ rewrite !(mul_comm _ p). apply mul_add_distr_l.
Qed.
-(*************************************************************************)
-(** ** Properties of the direct definition of successor and predecessor *)
+End Private_BootStrap.
+
+(** * Proofs of specifications *)
+
+(** ** Specification of constants *)
-Theorem Zsucc_succ' : forall n:Z, Zsucc n = Zsucc' n.
+Lemma one_succ : 1 = succ 0.
Proof.
-destruct n as [| p | p]; simpl.
reflexivity.
-now rewrite Pplus_one_succ_r.
-now destruct p as [q | q |].
Qed.
-Theorem Zpred_pred' : forall n:Z, Zpred n = Zpred' n.
+Lemma two_succ : 2 = succ 1.
Proof.
-destruct n as [| p | p]; simpl.
reflexivity.
-now destruct p as [q | q |].
-now rewrite Pplus_one_succ_r.
Qed.
-Theorem Zsucc'_inj : forall n m:Z, Zsucc' n = Zsucc' m -> n = m.
+(** ** Specification of addition *)
+
+Lemma add_0_l n : 0 + n = n.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma add_succ_l n m : succ n + m = succ (n + m).
Proof.
-intros n m; do 2 rewrite <- Zsucc_succ'; now apply Zsucc_inj.
+ unfold succ. now rewrite 2 (add_comm _ 1), add_assoc.
Qed.
-Theorem Zsucc'_pred' : forall n:Z, Zsucc' (Zpred' n) = n.
+(** ** Specification of opposite *)
+
+Lemma opp_0 : -0 = 0.
Proof.
-intro; rewrite <- Zsucc_succ'; rewrite <- Zpred_pred';
-symmetry; apply Zsucc_pred.
+ reflexivity.
Qed.
-Theorem Zpred'_succ' : forall n:Z, Zpred' (Zsucc' n) = n.
+Lemma opp_succ n : -(succ n) = pred (-n).
Proof.
-intro; apply Zsucc'_inj; now rewrite Zsucc'_pred'.
+ unfold succ, pred. apply opp_add_distr.
Qed.
-Theorem Zpred'_inj : forall n m:Z, Zpred' n = Zpred' m -> n = m.
+(** ** Specification of successor and predecessor *)
+
+Lemma succ_pred n : succ (pred n) = n.
Proof.
-intros n m H.
-rewrite <- (Zsucc'_pred' n); rewrite <- (Zsucc'_pred' m); now rewrite H.
+ unfold succ, pred. now rewrite <- add_assoc, add_opp_diag_r, add_0_r.
Qed.
-Theorem Zsucc'_discr : forall n:Z, n <> Zsucc' n.
+Lemma pred_succ n : pred (succ n) = n.
Proof.
- intro x; destruct x; simpl in |- *.
- discriminate.
- injection; apply Psucc_discr.
- destruct p; simpl in |- *.
- discriminate.
- intro H; symmetry in H; injection H; apply double_moins_un_xO_discr.
- discriminate.
+ unfold succ, pred. now rewrite <- add_assoc, add_opp_diag_r, add_0_r.
Qed.
-(** Misc properties, usually redundant or non natural *)
+(** ** Specification of subtraction *)
-Lemma Zsucc_eq_compat : forall n m:Z, n = m -> Zsucc n = Zsucc m.
+Lemma sub_0_r n : n - 0 = n.
Proof.
- intros n m H; rewrite H; reflexivity.
+ apply add_0_r.
Qed.
-Lemma Zsucc_inj_contrapositive : forall n m:Z, n <> m -> Zsucc n <> Zsucc m.
+Lemma sub_succ_r n m : n - succ m = pred (n - m).
Proof.
- unfold not in |- *; intros n m H1 H2; apply H1; apply Zsucc_inj; assumption.
+ unfold sub, succ, pred. now rewrite opp_add_distr, add_assoc.
Qed.
-(**********************************************************************)
-(** * Properties of subtraction on binary integer numbers *)
+(** ** Specification of multiplication *)
-(** ** [minus] and [Z0] *)
+Lemma mul_0_l n : 0 * n = 0.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma mul_succ_l n m : succ n * m = n * m + m.
+Proof.
+ unfold succ. now rewrite mul_add_distr_r, mul_1_l.
+Qed.
+
+(** ** Specification of comparisons and order *)
+
+Lemma eqb_eq n m : (n =? m) = true <-> n = m.
+Proof.
+ destruct n, m; simpl; try (now split).
+ rewrite inj_Zpos. apply Pos.eqb_eq.
+ rewrite inj_Zneg. apply Pos.eqb_eq.
+Qed.
+
+Lemma ltb_lt n m : (n <? m) = true <-> n < m.
+Proof.
+ unfold ltb, lt. destruct compare; easy'.
+Qed.
-Lemma Zminus_0_r : forall n:Z, n - Z0 = n.
+Lemma leb_le n m : (n <=? m) = true <-> n <= m.
Proof.
- intro; unfold Zminus in |- *; simpl in |- *; rewrite Zplus_0_r;
- trivial with arith.
+ unfold leb, le. destruct compare; easy'.
Qed.
-Lemma Zminus_0_l_reverse : forall n:Z, n = n - Z0.
+Lemma compare_eq_iff n m : (n ?= m) = Eq <-> n = m.
Proof.
- intro; symmetry in |- *; apply Zminus_0_r.
+destruct n, m; simpl; rewrite ?CompOpp_iff, ?Pos.compare_eq_iff;
+ split; congruence.
Qed.
-Lemma Zminus_diag : forall n:Z, n - n = Z0.
+Lemma compare_sub n m : (n ?= m) = (n - m ?= 0).
Proof.
- intro; unfold Zminus in |- *; rewrite Zplus_opp_r; trivial with arith.
+ destruct n as [|n|n], m as [|m|m]; simpl; trivial;
+ rewrite <- ? Pos.compare_antisym, ?pos_sub_spec;
+ case Pos.compare_spec; trivial.
Qed.
-Lemma Zminus_diag_reverse : forall n:Z, Z0 = n - n.
+Lemma compare_antisym n m : (m ?= n) = CompOpp (n ?= m).
Proof.
- intro; symmetry in |- *; apply Zminus_diag.
+destruct n, m; simpl; trivial; now rewrite <- ?Pos.compare_antisym.
Qed.
+Lemma compare_lt_iff n m : (n ?= m) = Lt <-> n < m.
+Proof. reflexivity. Qed.
+
+Lemma compare_le_iff n m : (n ?= m) <> Gt <-> n <= m.
+Proof. reflexivity. Qed.
+
+(** Some more advanced properties of comparison and orders,
+ including [compare_spec] and [lt_irrefl] and [lt_eq_cases]. *)
-(** ** Relating [minus] with [plus] and [Zsucc] *)
+Include BoolOrderFacts.
-Lemma Zminus_plus_distr : forall n m p:Z, n - (m + p) = n - m - p.
+(** Remaining specification of [lt] and [le] *)
+
+Lemma lt_succ_r n m : n < succ m <-> n<=m.
Proof.
-intros; unfold Zminus; rewrite Zopp_plus_distr; apply Zplus_assoc.
+ unfold lt, le. rewrite compare_sub, sub_succ_r.
+ rewrite (compare_sub n m).
+ destruct (n-m) as [|[ | | ]|]; easy'.
Qed.
-Lemma Zminus_succ_l : forall n m:Z, Zsucc (n - m) = Zsucc n - m.
+(** ** Specification of minimum and maximum *)
+
+Lemma max_l n m : m<=n -> max n m = n.
Proof.
- intros n m; unfold Zminus, Zsucc in |- *; rewrite (Zplus_comm n (- m));
- rewrite <- Zplus_assoc; apply Zplus_comm.
+ unfold le, max. rewrite (compare_antisym n m).
+ case compare; intuition.
Qed.
-Lemma Zminus_succ_r : forall n m:Z, n - (Zsucc m) = Zpred (n - m).
+Lemma max_r n m : n<=m -> max n m = m.
Proof.
-intros; unfold Zsucc; now rewrite Zminus_plus_distr.
+ unfold le, max. case compare_spec; intuition.
Qed.
-Lemma Zplus_minus_eq : forall n m p:Z, n = m + p -> p = n - m.
+Lemma min_l n m : n<=m -> min n m = n.
Proof.
- intros n m p H; unfold Zminus in |- *; apply (Zplus_reg_l m);
- rewrite (Zplus_comm m (n + - m)); rewrite <- Zplus_assoc;
- rewrite Zplus_opp_l; rewrite Zplus_0_r; rewrite H;
- trivial with arith.
+ unfold le, min. case compare_spec; intuition.
Qed.
-Lemma Zminus_plus : forall n m:Z, n + m - n = m.
+Lemma min_r n m : m<=n -> min n m = m.
Proof.
- intros n m; unfold Zminus in |- *; rewrite (Zplus_comm n m);
- rewrite <- Zplus_assoc; rewrite Zplus_opp_r; apply Zplus_0_r.
+ unfold le, min.
+ rewrite (compare_antisym n m). case compare_spec; intuition.
Qed.
-Lemma Zplus_minus : forall n m:Z, n + (m - n) = m.
+(** ** Specification of absolute value *)
+
+Lemma abs_eq n : 0 <= n -> abs n = n.
Proof.
- unfold Zminus in |- *; intros n m; rewrite Zplus_permute; rewrite Zplus_opp_r;
- apply Zplus_0_r.
+ destruct n; trivial. now destruct 1.
Qed.
-Lemma Zminus_plus_simpl_l : forall n m p:Z, p + n - (p + m) = n - m.
+Lemma abs_neq n : n <= 0 -> abs n = - n.
Proof.
- intros n m p; unfold Zminus in |- *; rewrite Zopp_plus_distr;
- rewrite Zplus_assoc; rewrite (Zplus_comm p); rewrite <- (Zplus_assoc n p);
- rewrite Zplus_opp_r; rewrite Zplus_0_r; trivial with arith.
+ destruct n; trivial. now destruct 1.
Qed.
-Lemma Zminus_plus_simpl_l_reverse : forall n m p:Z, n - m = p + n - (p + m).
+(** ** Specification of sign *)
+
+Lemma sgn_null n : n = 0 -> sgn n = 0.
Proof.
- intros; symmetry in |- *; apply Zminus_plus_simpl_l.
+ intros. now subst.
Qed.
-Lemma Zminus_plus_simpl_r : forall n m p:Z, n + p - (m + p) = n - m.
+Lemma sgn_pos n : 0 < n -> sgn n = 1.
Proof.
- intros x y n.
- unfold Zminus in |- *.
- rewrite Zopp_plus_distr.
- rewrite (Zplus_comm (- y) (- n)).
- rewrite Zplus_assoc.
- rewrite <- (Zplus_assoc x n (- n)).
- rewrite (Zplus_opp_r n).
- rewrite <- Zplus_0_r_reverse.
- reflexivity.
+ now destruct n.
Qed.
-Lemma Zpos_minus_morphism : forall a b:positive, Pcompare a b Eq = Lt ->
- Zpos (b-a) = Zpos b - Zpos a.
+Lemma sgn_neg n : n < 0 -> sgn n = -1.
Proof.
- intros.
- simpl.
- change Eq with (CompOpp Eq).
- rewrite <- Pcompare_antisym.
- rewrite H; simpl; auto.
+ now destruct n.
Qed.
-(** ** Misc redundant properties *)
+(** ** Specification of power *)
-Lemma Zeq_minus : forall n m:Z, n = m -> n - m = Z0.
+Lemma pow_0_r n : n^0 = 1.
Proof.
- intros x y H; rewrite H; symmetry in |- *; apply Zminus_diag_reverse.
+ reflexivity.
Qed.
-Lemma Zminus_eq : forall n m:Z, n - m = Z0 -> n = m.
+Lemma pow_succ_r n m : 0<=m -> n^(succ m) = n * n^m.
Proof.
- intros x y H; rewrite <- (Zplus_minus y x); rewrite H; apply Zplus_0_r.
+ destruct m as [|m|m]; (now destruct 1) || (intros _); simpl; trivial.
+ unfold pow_pos. now rewrite Pos.add_comm, Pos.iter_add.
Qed.
+Lemma pow_neg_r n m : m<0 -> n^m = 0.
+Proof.
+ now destruct m.
+Qed.
-(**********************************************************************)
-(** * Properties of multiplication on binary integer numbers *)
+(** For folding back a [pow_pos] into a [pow] *)
-Theorem Zpos_mult_morphism :
- forall p q:positive, Zpos (p*q) = Zpos p * Zpos q.
+Lemma pow_pos_fold n p : pow_pos n p = n ^ (Zpos p).
Proof.
- auto.
+ reflexivity.
Qed.
-(** ** One is neutral for multiplication *)
+(** ** Specification of square *)
-Theorem Zmult_1_l : forall n:Z, Zpos 1 * n = n.
+Lemma square_spec n : square n = n * n.
Proof.
- intro x; destruct x; reflexivity.
+ destruct n; trivial; simpl; f_equal; apply Pos.square_spec.
Qed.
-Theorem Zmult_1_r : forall n:Z, n * Zpos 1 = n.
+(** ** Specification of square root *)
+
+Lemma sqrtrem_spec n : 0<=n ->
+ let (s,r) := sqrtrem n in n = s*s + r /\ 0 <= r <= 2*s.
Proof.
- intro x; destruct x; simpl in |- *; try rewrite Pmult_1_r; reflexivity.
+ destruct n. now repeat split.
+ generalize (Pos.sqrtrem_spec p). simpl.
+ destruct 1; simpl; subst; now repeat split.
+ now destruct 1.
Qed.
-(** ** Zero property of multiplication *)
+Lemma sqrt_spec n : 0<=n ->
+ let s := sqrt n in s*s <= n < (succ s)*(succ s).
+Proof.
+ destruct n. now repeat split. unfold sqrt.
+ rewrite succ_Zpos. intros _. apply (Pos.sqrt_spec p).
+ now destruct 1.
+Qed.
-Theorem Zmult_0_l : forall n:Z, Z0 * n = Z0.
+Lemma sqrt_neg n : n<0 -> sqrt n = 0.
Proof.
- intro x; destruct x; reflexivity.
+ now destruct n.
Qed.
-Theorem Zmult_0_r : forall n:Z, n * Z0 = Z0.
+Lemma sqrtrem_sqrt n : fst (sqrtrem n) = sqrt n.
Proof.
- intro x; destruct x; reflexivity.
+ destruct n; try reflexivity.
+ unfold sqrtrem, sqrt, Pos.sqrt.
+ destruct (Pos.sqrtrem p) as (s,r). now destruct r.
Qed.
-Hint Local Resolve Zmult_0_l Zmult_0_r.
+(** ** Specification of logarithm *)
-Lemma Zmult_0_r_reverse : forall n:Z, Z0 = n * Z0.
+Lemma log2_spec n : 0 < n -> 2^(log2 n) <= n < 2^(succ (log2 n)).
Proof.
- intro x; destruct x; reflexivity.
+ destruct n as [|[p|p|]|]; intros Hn; split; try easy; unfold log2;
+ rewrite ?succ_Zpos, pow_Zpos.
+ change (2^Pos.size p <= Pos.succ (p~0))%positive.
+ apply Pos.lt_le_incl, Pos.lt_succ_r, Pos.size_le.
+ apply Pos.size_gt.
+ apply Pos.size_le.
+ apply Pos.size_gt.
Qed.
-(** ** Commutativity of multiplication *)
+Lemma log2_nonpos n : n<=0 -> log2 n = 0.
+Proof.
+ destruct n as [|p|p]; trivial; now destruct 1.
+Qed.
+
+(** Specification of parity functions *)
-Theorem Zmult_comm : forall n m:Z, n * m = m * n.
+Lemma even_spec n : even n = true <-> Even n.
Proof.
- intros x y; destruct x as [| p| p]; destruct y as [| q| q]; simpl in |- *;
- try rewrite (Pmult_comm p q); reflexivity.
+ split.
+ exists (div2 n). now destruct n as [|[ | | ]|[ | | ]].
+ intros (m,->). now destruct m.
Qed.
-(** ** Associativity of multiplication *)
+Lemma odd_spec n : odd n = true <-> Odd n.
+Proof.
+ split.
+ exists (div2 n). destruct n as [|[ | | ]|[ | | ]]; simpl; try easy.
+ now rewrite Pos.pred_double_succ.
+ intros (m,->). now destruct m as [|[ | | ]|[ | | ]].
+Qed.
+
+(** ** Multiplication and Doubling *)
+
+Lemma double_spec n : double n = 2*n.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma succ_double_spec n : succ_double n = 2*n + 1.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma pred_double_spec n : pred_double n = 2*n - 1.
+Proof.
+ now destruct n.
+Qed.
+
+(** ** Correctness proofs for Trunc division *)
+
+Lemma pos_div_eucl_eq a b : 0 < b ->
+ let (q, r) := pos_div_eucl a b in Zpos a = q * b + r.
+Proof.
+ intros Hb.
+ induction a; unfold pos_div_eucl; fold pos_div_eucl.
+ - (* ~1 *)
+ destruct pos_div_eucl as (q,r).
+ rewrite pos_xI, IHa, mul_add_distr_l, mul_assoc.
+ destruct ltb.
+ now rewrite add_assoc.
+ rewrite mul_add_distr_r, mul_1_l, <- !add_assoc. f_equal.
+ unfold sub. now rewrite (add_comm _ (-b)), add_assoc, add_opp_diag_r.
+ - (* ~0 *)
+ destruct pos_div_eucl as (q,r).
+ rewrite (pos_xO a), IHa, mul_add_distr_l, mul_assoc.
+ destruct ltb.
+ trivial.
+ rewrite mul_add_distr_r, mul_1_l, <- !add_assoc. f_equal.
+ unfold sub. now rewrite (add_comm _ (-b)), add_assoc, add_opp_diag_r.
+ - (* 1 *)
+ case leb_spec; trivial.
+ intros Hb'.
+ destruct b as [|b|b]; try easy; clear Hb.
+ replace b with 1%positive; trivial.
+ apply Pos.le_antisym. apply Pos.le_1_l. now apply Pos.lt_succ_r.
+Qed.
+
+Lemma div_eucl_eq a b : b<>0 ->
+ let (q, r) := div_eucl a b in a = b * q + r.
+Proof.
+ destruct a as [ |a|a], b as [ |b|b]; unfold div_eucl; trivial;
+ (now destruct 1) || intros _;
+ generalize (pos_div_eucl_eq a (Zpos b) (eq_refl _));
+ destruct pos_div_eucl as (q,r); rewrite <- ?opp_Zpos, mul_comm;
+ intros ->.
+ - (* Zpos Zpos *)
+ trivial.
+ - (* Zpos Zneg *)
+ destruct r as [ |r|r]; rewrite !mul_opp_opp; trivial;
+ rewrite mul_add_distr_l, mul_1_r, <- add_assoc; f_equal;
+ now rewrite add_assoc, add_opp_diag_r.
+ - (* Zneg Zpos *)
+ rewrite (opp_add_distr _ r), <- mul_opp_r.
+ destruct r as [ |r|r]; trivial;
+ rewrite opp_add_distr, mul_add_distr_l, <- add_assoc; f_equal;
+ unfold sub; now rewrite add_assoc, mul_opp_r, mul_1_r, add_opp_diag_l.
+ - (* Zneg Zneg *)
+ now rewrite opp_add_distr, <- mul_opp_l.
+Qed.
+
+Lemma div_mod a b : b<>0 -> a = b*(a/b) + (a mod b).
+Proof.
+ intros Hb. generalize (div_eucl_eq a b Hb).
+ unfold div, modulo. now destruct div_eucl.
+Qed.
+
+Lemma pos_div_eucl_bound a b : 0<b -> 0 <= snd (pos_div_eucl a b) < b.
+Proof.
+ assert (AUX : forall m p, m < Zpos (p~0) -> m - Zpos p < Zpos p).
+ intros m p. unfold lt.
+ rewrite (compare_sub m), (compare_sub _ (Zpos _)). unfold sub.
+ rewrite <- add_assoc. simpl opp; simpl (Zneg _ + _).
+ now rewrite Pos.add_diag.
+ intros Hb.
+ destruct b as [|b|b]; discriminate Hb || clear Hb.
+ induction a; unfold pos_div_eucl; fold pos_div_eucl.
+ (* ~1 *)
+ destruct pos_div_eucl as (q,r).
+ simpl in IHa; destruct IHa as (Hr,Hr').
+ case ltb_spec; intros H; unfold snd. split; trivial. now destruct r.
+ split. unfold le.
+ now rewrite compare_antisym, <- compare_sub, <- compare_antisym.
+ apply AUX. rewrite <- succ_double_spec.
+ destruct r; try easy. unfold lt in *; simpl in *.
+ now rewrite Pos.compare_xI_xO, Hr'.
+ (* ~0 *)
+ destruct pos_div_eucl as (q,r).
+ simpl in IHa; destruct IHa as (Hr,Hr').
+ case ltb_spec; intros H; unfold snd. split; trivial. now destruct r.
+ split. unfold le.
+ now rewrite compare_antisym, <- compare_sub, <- compare_antisym.
+ apply AUX. destruct r; try easy.
+ (* 1 *)
+ case leb_spec; intros H; simpl; split; try easy.
+ red; simpl. now apply Pos.le_succ_l.
+Qed.
-Theorem Zmult_assoc : forall n m p:Z, n * (m * p) = n * m * p.
+Lemma mod_pos_bound a b : 0 < b -> 0 <= a mod b < b.
Proof.
- intros x y z; destruct x; destruct y; destruct z; simpl in |- *;
- try rewrite Pmult_assoc; reflexivity.
+ destruct b as [|b|b]; try easy; intros _.
+ destruct a as [|a|a]; unfold modulo, div_eucl.
+ now split.
+ now apply pos_div_eucl_bound.
+ generalize (pos_div_eucl_bound a (Zpos b) (eq_refl _)).
+ destruct pos_div_eucl as (q,r); unfold snd; intros (Hr,Hr').
+ destruct r as [|r|r]; (now destruct Hr) || clear Hr.
+ now split.
+ split. unfold le.
+ now rewrite compare_antisym, <- compare_sub, <- compare_antisym, Hr'.
+ unfold lt in *; simpl in *. rewrite pos_sub_gt by trivial.
+ simpl. now apply Pos.sub_decr.
Qed.
-Lemma Zmult_assoc_reverse : forall n m p:Z, n * m * p = n * (m * p).
+Definition mod_bound_pos a b (_:0<=a) := mod_pos_bound a b.
+
+Lemma mod_neg_bound a b : b < 0 -> b < a mod b <= 0.
Proof.
- intros n m p; rewrite Zmult_assoc; trivial with arith.
+ destruct b as [|b|b]; try easy; intros _.
+ destruct a as [|a|a]; unfold modulo, div_eucl.
+ now split.
+ generalize (pos_div_eucl_bound a (Zpos b) (eq_refl _)).
+ destruct pos_div_eucl as (q,r); unfold snd; intros (Hr,Hr').
+ destruct r as [|r|r]; (now destruct Hr) || clear Hr.
+ now split.
+ split.
+ unfold lt in *; simpl in *. rewrite pos_sub_lt by trivial.
+ rewrite <- Pos.compare_antisym. now apply Pos.sub_decr.
+ change (Zneg b - Zneg r <= 0). unfold le, lt in *.
+ rewrite <- compare_sub. simpl in *.
+ now rewrite <- Pos.compare_antisym, Hr'.
+ generalize (pos_div_eucl_bound a (Zpos b) (eq_refl _)).
+ destruct pos_div_eucl as (q,r); unfold snd; intros (Hr,Hr').
+ split; destruct r; try easy.
+ red; simpl; now rewrite <- Pos.compare_antisym.
Qed.
-(** ** Associativity mixed with commutativity *)
+(** ** Correctness proofs for Floor division *)
+
+Theorem quotrem_eq a b : let (q,r) := quotrem a b in a = q * b + r.
+Proof.
+ destruct a as [|a|a], b as [|b|b]; simpl; trivial;
+ generalize (N.pos_div_eucl_spec a (Npos b)); case N.pos_div_eucl; trivial;
+ intros q r; rewrite <- ?opp_Zpos;
+ change (Zpos a) with (of_N (Npos a)); intros ->; now destruct q, r.
+Qed.
-Theorem Zmult_permute : forall n m p:Z, n * (m * p) = m * (n * p).
+Lemma quot_rem' a b : a = b*(a÷b) + rem a b.
Proof.
- intros x y z; rewrite (Zmult_assoc y x z); rewrite (Zmult_comm y x).
- apply Zmult_assoc.
+ rewrite mul_comm. generalize (quotrem_eq a b).
+ unfold quot, rem. now destruct quotrem.
Qed.
-(** ** Z is integral *)
+Lemma quot_rem a b : b<>0 -> a = b*(a÷b) + rem a b.
+Proof. intros _. apply quot_rem'. Qed.
-Theorem Zmult_integral_l : forall n m:Z, n <> Z0 -> m * n = Z0 -> m = Z0.
+Lemma rem_bound_pos a b : 0<=a -> 0<b -> 0 <= rem a b < b.
Proof.
- intros x y; destruct x as [| p| p].
- intro H; absurd (Z0 = Z0); trivial.
- intros _ H; destruct y as [| q| q]; reflexivity || discriminate.
- intros _ H; destruct y as [| q| q]; reflexivity || discriminate.
+ intros Ha Hb.
+ destruct b as [|b|b]; (now discriminate Hb) || clear Hb;
+ destruct a as [|a|a]; (now destruct Ha) || clear Ha.
+ compute. now split.
+ unfold rem, quotrem.
+ assert (H := N.pos_div_eucl_remainder a (Npos b)).
+ destruct N.pos_div_eucl as (q,[|r]); simpl; split; try easy.
+ now apply H.
Qed.
+Lemma rem_opp_l' a b : rem (-a) b = - (rem a b).
+Proof.
+ destruct a, b; trivial; unfold rem; simpl;
+ now destruct N.pos_div_eucl as (q,[|r]).
+Qed.
-Theorem Zmult_integral : forall n m:Z, n * m = Z0 -> n = Z0 \/ m = Z0.
+Lemma rem_opp_r' a b : rem a (-b) = rem a b.
Proof.
- intros x y; destruct x; destruct y; auto; simpl in |- *; intro H;
- discriminate H.
+ destruct a, b; trivial; unfold rem; simpl;
+ now destruct N.pos_div_eucl as (q,[|r]).
Qed.
+Lemma rem_opp_l a b : b<>0 -> rem (-a) b = - (rem a b).
+Proof. intros _. apply rem_opp_l'. Qed.
-Lemma Zmult_1_inversion_l :
- forall n m:Z, n * m = Zpos 1 -> n = Zpos 1 \/ n = Zneg 1.
+Lemma rem_opp_r a b : b<>0 -> rem a (-b) = rem a b.
+Proof. intros _. apply rem_opp_r'. Qed.
+
+(** ** Basic properties of divisibility *)
+
+Lemma divide_Zpos p q : (Zpos p|Zpos q) <-> (p|q)%positive.
Proof.
- intros x y; destruct x as [| p| p]; intro; [ discriminate | left | right ];
- (destruct y as [| q| q]; try discriminate; simpl in H; injection H; clear H;
- intro H; rewrite Pmult_1_inversion_l with (1 := H);
- reflexivity).
+ split.
+ intros ([ |r|r],H); simpl in *; destr_eq H. exists r; auto.
+ intros (r,H). exists (Zpos r); simpl; now f_equal.
Qed.
-(** ** Multiplication and Doubling *)
+Lemma divide_Zpos_Zneg_r n p : (n|Zpos p) <-> (n|Zneg p).
+Proof.
+ split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- H.
+Qed.
-Lemma Zdouble_mult : forall z, Zdouble z = (Zpos 2) * z.
+Lemma divide_Zpos_Zneg_l n p : (Zpos p|n) <-> (Zneg p|n).
Proof.
- reflexivity.
+ split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- mul_opp_r.
Qed.
-Lemma Zdouble_plus_one_mult : forall z,
- Zdouble_plus_one z = (Zpos 2) * z + (Zpos 1).
+(** ** Correctness proofs for gcd *)
+
+Lemma ggcd_gcd a b : fst (ggcd a b) = gcd a b.
Proof.
- destruct z; simpl; auto with zarith.
+ destruct a as [ |p|p], b as [ |q|q]; simpl; auto;
+ generalize (Pos.ggcd_gcd p q); destruct Pos.ggcd as (g,(aa,bb));
+ simpl; congruence.
Qed.
-(** ** Multiplication and Opposite *)
+Lemma ggcd_correct_divisors a b :
+ let '(g,(aa,bb)) := ggcd a b in
+ a = g*aa /\ b = g*bb.
+Proof.
+ destruct a as [ |p|p], b as [ |q|q]; simpl; rewrite ?Pos.mul_1_r; auto;
+ generalize (Pos.ggcd_correct_divisors p q);
+ destruct Pos.ggcd as (g,(aa,bb)); simpl; destruct 1; now subst.
+Qed.
-Theorem Zopp_mult_distr_l : forall n m:Z, - (n * m) = - n * m.
+Lemma gcd_divide_l a b : (gcd a b | a).
Proof.
- intros x y; destruct x; destruct y; reflexivity.
+ rewrite <- ggcd_gcd. generalize (ggcd_correct_divisors a b).
+ destruct ggcd as (g,(aa,bb)); simpl. intros (H,_). exists aa.
+ now rewrite mul_comm.
Qed.
-Theorem Zopp_mult_distr_r : forall n m:Z, - (n * m) = n * - m.
+Lemma gcd_divide_r a b : (gcd a b | b).
Proof.
- intros x y; rewrite (Zmult_comm x y); rewrite Zopp_mult_distr_l;
- apply Zmult_comm.
+ rewrite <- ggcd_gcd. generalize (ggcd_correct_divisors a b).
+ destruct ggcd as (g,(aa,bb)); simpl. intros (_,H). exists bb.
+ now rewrite mul_comm.
Qed.
-Lemma Zopp_mult_distr_l_reverse : forall n m:Z, - n * m = - (n * m).
+Lemma gcd_greatest a b c : (c|a) -> (c|b) -> (c | gcd a b).
Proof.
- intros x y; symmetry in |- *; apply Zopp_mult_distr_l.
+ assert (H : forall p q r, (r|Zpos p) -> (r|Zpos q) -> (r|Zpos (Pos.gcd p q))).
+ { intros p q [|r|r] H H'.
+ destruct H; now rewrite mul_comm in *.
+ apply divide_Zpos, Pos.gcd_greatest; now apply divide_Zpos.
+ apply divide_Zpos_Zneg_l, divide_Zpos, Pos.gcd_greatest;
+ now apply divide_Zpos, divide_Zpos_Zneg_l.
+ }
+ destruct a, b; simpl; auto; intros; try apply H; trivial;
+ now apply divide_Zpos_Zneg_r.
Qed.
-Theorem Zmult_opp_comm : forall n m:Z, - n * m = n * - m.
+Lemma gcd_nonneg a b : 0 <= gcd a b.
Proof.
- intros x y; rewrite Zopp_mult_distr_l_reverse; rewrite Zopp_mult_distr_r;
- trivial with arith.
+ now destruct a, b.
Qed.
-Theorem Zmult_opp_opp : forall n m:Z, - n * - m = n * m.
+(** ggcd and opp : an auxiliary result used in QArith *)
+
+Theorem ggcd_opp a b :
+ ggcd (-a) b = (let '(g,(aa,bb)) := ggcd a b in (g,(-aa,bb))).
Proof.
- intros x y; destruct x; destruct y; reflexivity.
+ destruct a as [|a|a], b as [|b|b]; unfold ggcd, opp; auto;
+ destruct (Pos.ggcd a b) as (g,(aa,bb)); auto.
Qed.
-Theorem Zopp_eq_mult_neg_1 : forall n:Z, - n = n * Zneg 1.
+(** ** Conversions between [Z.testbit] and [N.testbit] *)
+
+Lemma testbit_of_N a n :
+ testbit (of_N a) (of_N n) = N.testbit a n.
Proof.
- intro x; induction x; intros; rewrite Zmult_comm; auto with arith.
+ destruct a as [|a], n; simpl; trivial. now destruct a.
Qed.
-(** ** Distributivity of multiplication over addition *)
+Lemma testbit_of_N' a n : 0<=n ->
+ testbit (of_N a) n = N.testbit a (to_N n).
+Proof.
+ intro Hn. rewrite <- testbit_of_N. f_equal.
+ destruct n; trivial; now destruct Hn.
+Qed.
-Lemma weak_Zmult_plus_distr_r :
- forall (p:positive) (n m:Z), Zpos p * (n + m) = Zpos p * n + Zpos p * m.
+Lemma testbit_Zpos a n : 0<=n ->
+ testbit (Zpos a) n = N.testbit (Npos a) (to_N n).
Proof.
- intros x y' z'; case y'; case z'; auto with arith; intros y z;
- (simpl in |- *; rewrite Pmult_plus_distr_l; trivial with arith) ||
- (simpl in |- *; ElimPcompare z y; intros E0; rewrite E0;
- [ rewrite (Pcompare_Eq_eq z y E0); rewrite (Pcompare_refl (x * y));
- trivial with arith
- | cut ((x * z ?= x * y)%positive Eq = Lt);
- [ intros E; rewrite E; rewrite Pmult_minus_distr_l;
- [ trivial with arith | apply ZC2; assumption ]
- | apply nat_of_P_lt_Lt_compare_complement_morphism;
- do 2 rewrite nat_of_P_mult_morphism; elim (ZL4 x);
- intros h H1; rewrite H1; apply mult_S_lt_compat_l;
- exact (nat_of_P_lt_Lt_compare_morphism z y E0) ]
- | cut ((x * z ?= x * y)%positive Eq = Gt);
- [ intros E; rewrite E; rewrite Pmult_minus_distr_l; auto with arith
- | apply nat_of_P_gt_Gt_compare_complement_morphism; unfold gt in |- *;
- do 2 rewrite nat_of_P_mult_morphism; elim (ZL4 x);
- intros h H1; rewrite H1; apply mult_S_lt_compat_l;
- exact (nat_of_P_gt_Gt_compare_morphism z y E0) ] ]).
+ intro Hn. now rewrite <- testbit_of_N'.
Qed.
-Theorem Zmult_plus_distr_r : forall n m p:Z, n * (m + p) = n * m + n * p.
+Lemma testbit_Zneg a n : 0<=n ->
+ testbit (Zneg a) n = negb (N.testbit (Pos.pred_N a) (to_N n)).
Proof.
- intros x y z; case x;
- [ auto with arith
- | intros x'; apply weak_Zmult_plus_distr_r
- | intros p; apply Zopp_inj; rewrite Zopp_plus_distr;
- do 3 rewrite <- Zopp_mult_distr_l_reverse; rewrite Zopp_neg;
- apply weak_Zmult_plus_distr_r ].
+ intro Hn.
+ rewrite <- testbit_of_N' by trivial.
+ destruct n as [ |n|n];
+ [ | simpl; now destruct (Ppred_N a) | now destruct Hn].
+ unfold testbit.
+ now destruct a as [|[ | | ]| ].
Qed.
-Theorem Zmult_plus_distr_l : forall n m p:Z, (n + m) * p = n * p + m * p.
+(** ** Proofs of specifications for bitwise operations *)
+
+Lemma div2_spec a : div2 a = shiftr a 1.
Proof.
- intros n m p; rewrite Zmult_comm; rewrite Zmult_plus_distr_r;
- do 2 rewrite (Zmult_comm p); trivial with arith.
+ reflexivity.
Qed.
-(** ** Distributivity of multiplication over subtraction *)
+Lemma testbit_0_l n : testbit 0 n = false.
+Proof.
+ now destruct n.
+Qed.
-Lemma Zmult_minus_distr_r : forall n m p:Z, (n - m) * p = n * p - m * p.
+Lemma testbit_neg_r a n : n<0 -> testbit a n = false.
Proof.
- intros x y z; unfold Zminus in |- *.
- rewrite <- Zopp_mult_distr_l_reverse.
- apply Zmult_plus_distr_l.
+ now destruct n.
Qed.
+Lemma testbit_odd_0 a : testbit (2*a+1) 0 = true.
+Proof.
+ now destruct a as [|a|[a|a|]].
+Qed.
-Lemma Zmult_minus_distr_l : forall n m p:Z, p * (n - m) = p * n - p * m.
+Lemma testbit_even_0 a : testbit (2*a) 0 = false.
Proof.
- intros x y z; rewrite (Zmult_comm z (x - y)).
- rewrite (Zmult_comm z x).
- rewrite (Zmult_comm z y).
- apply Zmult_minus_distr_r.
+ now destruct a.
Qed.
-(** ** Simplification of multiplication for non-zero integers *)
+Lemma testbit_odd_succ a n : 0<=n ->
+ testbit (2*a+1) (succ n) = testbit a n.
+Proof.
+ destruct n as [|n|n]; (now destruct 1) || intros _.
+ destruct a as [|[a|a|]|[a|a|]]; simpl; trivial. now destruct a.
+ unfold testbit. rewrite succ_Zpos.
+ destruct a as [|a|[a|a|]]; simpl; trivial;
+ rewrite ?Pos.pred_N_succ; now destruct n.
+Qed.
-Lemma Zmult_reg_l : forall n m p:Z, p <> Z0 -> p * n = p * m -> n = m.
+Lemma testbit_even_succ a n : 0<=n ->
+ testbit (2*a) (succ n) = testbit a n.
Proof.
- intros x y z H H0.
- generalize (Zeq_minus _ _ H0).
- intro.
- apply Zminus_eq.
- rewrite <- Zmult_minus_distr_l in H1.
- clear H0; destruct (Zmult_integral _ _ H1).
- contradiction.
- trivial.
+ destruct n as [|n|n]; (now destruct 1) || intros _.
+ destruct a as [|[a|a|]|[a|a|]]; simpl; trivial. now destruct a.
+ unfold testbit. rewrite succ_Zpos.
+ destruct a as [|a|[a|a|]]; simpl; trivial;
+ rewrite ?Pos.pred_N_succ; now destruct n.
Qed.
-Lemma Zmult_reg_r : forall n m p:Z, p <> Z0 -> n * p = m * p -> n = m.
+(** Correctness proofs about [Z.shiftr] and [Z.shiftl] *)
+
+Lemma shiftr_spec_aux a n m : 0<=n -> 0<=m ->
+ testbit (shiftr a n) m = testbit a (m+n).
Proof.
- intros x y z Hz.
- rewrite (Zmult_comm x z).
- rewrite (Zmult_comm y z).
- intro; apply Zmult_reg_l with z; assumption.
+ intros Hn Hm. unfold shiftr.
+ destruct n as [ |n|n]; (now destruct Hn) || clear Hn; simpl.
+ now rewrite add_0_r.
+ assert (forall p, to_N (m + Zpos p) = (to_N m + Npos p)%N).
+ destruct m; trivial; now destruct Hm.
+ assert (forall p, 0 <= m + Zpos p).
+ destruct m; easy || now destruct Hm.
+ destruct a as [ |a|a].
+ (* a = 0 *)
+ replace (Pos.iter n div2 0) with 0
+ by (apply Pos.iter_invariant; intros; subst; trivial).
+ now rewrite 2 testbit_0_l.
+ (* a > 0 *)
+ change (Zpos a) with (of_N (Npos a)) at 1.
+ rewrite <- (Pos.iter_swap_gen _ _ _ Ndiv2) by now intros [|[ | | ]].
+ rewrite testbit_Zpos, testbit_of_N', H; trivial.
+ exact (N.shiftr_spec' (Npos a) (Npos n) (to_N m)).
+ (* a < 0 *)
+ rewrite <- (Pos.iter_swap_gen _ _ _ Pdiv2_up) by trivial.
+ rewrite 2 testbit_Zneg, H; trivial. f_equal.
+ rewrite (Pos.iter_swap_gen _ _ _ _ Ndiv2) by exact N.pred_div2_up.
+ exact (N.shiftr_spec' (Ppred_N a) (Npos n) (to_N m)).
Qed.
-(** ** Addition and multiplication by 2 *)
+Lemma shiftl_spec_low a n m : m<n ->
+ testbit (shiftl a n) m = false.
+Proof.
+ intros H. destruct n as [|n|n], m as [|m|m]; try easy; simpl shiftl.
+ destruct (Pos.succ_pred_or n) as [-> | <-];
+ rewrite ?Pos.iter_succ; apply testbit_even_0.
+ destruct a as [ |a|a].
+ (* a = 0 *)
+ replace (Pos.iter n (mul 2) 0) with 0
+ by (apply Pos.iter_invariant; intros; subst; trivial).
+ apply testbit_0_l.
+ (* a > 0 *)
+ rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
+ rewrite testbit_Zpos by easy.
+ exact (N.shiftl_spec_low (Npos a) (Npos n) (Npos m) H).
+ (* a < 0 *)
+ rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
+ rewrite testbit_Zneg by easy.
+ now rewrite (N.pos_pred_shiftl_low a (Npos n)).
+Qed.
+
+Lemma shiftl_spec_high a n m : 0<=m -> n<=m ->
+ testbit (shiftl a n) m = testbit a (m-n).
+Proof.
+ intros Hm H.
+ destruct n as [ |n|n]. simpl. now rewrite sub_0_r.
+ (* n > 0 *)
+ destruct m as [ |m|m]; try (now destruct H).
+ assert (0 <= Zpos m - Zpos n).
+ red. now rewrite compare_antisym, <- compare_sub, <- compare_antisym.
+ assert (EQ : to_N (Zpos m - Zpos n) = (Npos m - Npos n)%N).
+ red in H. simpl in H. simpl to_N.
+ rewrite pos_sub_spec, Pos.compare_antisym.
+ destruct (Pos.compare_spec n m) as [H'|H'|H']; try (now destruct H).
+ subst. now rewrite N.sub_diag.
+ simpl. destruct (Pos.sub_mask_pos' m n H') as (p & -> & <-).
+ f_equal. now rewrite Pos.add_comm, Pos.add_sub.
+ destruct a; unfold shiftl.
+ (* ... a = 0 *)
+ replace (Pos.iter n (mul 2) 0) with 0
+ by (apply Pos.iter_invariant; intros; subst; trivial).
+ now rewrite 2 testbit_0_l.
+ (* ... a > 0 *)
+ rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
+ rewrite 2 testbit_Zpos, EQ by easy.
+ exact (N.shiftl_spec_high' (Npos p) (Npos n) (Npos m) H).
+ (* ... a < 0 *)
+ rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
+ rewrite 2 testbit_Zneg, EQ by easy. f_equal.
+ simpl to_N.
+ rewrite <- N.shiftl_spec_high by easy.
+ now apply (N.pos_pred_shiftl_high p (Npos n)).
+ (* n < 0 *)
+ unfold sub. simpl.
+ now apply (shiftr_spec_aux a (Zpos n) m).
+Qed.
+
+Lemma shiftr_spec a n m : 0<=m ->
+ testbit (shiftr a n) m = testbit a (m+n).
+Proof.
+ intros Hm.
+ destruct (leb_spec 0 n).
+ now apply shiftr_spec_aux.
+ destruct (leb_spec (-n) m) as [LE|GT].
+ unfold shiftr.
+ rewrite (shiftl_spec_high a (-n) m); trivial. now destruct n.
+ unfold shiftr.
+ rewrite (shiftl_spec_low a (-n) m); trivial.
+ rewrite testbit_neg_r; trivial.
+ red in GT. rewrite compare_sub in GT. now destruct n.
+Qed.
+
+(** Correctness proofs for bitwise operations *)
+
+Lemma lor_spec a b n :
+ testbit (lor a b) n = testbit a n || testbit b n.
+Proof.
+ destruct (leb_spec 0 n) as [Hn|Hn]; [|now rewrite !testbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?testbit_0_l, ?orb_false_r; trivial; unfold lor;
+ rewrite ?testbit_Zpos, ?testbit_Zneg, ?N.pos_pred_succ by trivial.
+ now rewrite <- N.lor_spec.
+ now rewrite N.ldiff_spec, negb_andb, negb_involutive, orb_comm.
+ now rewrite N.ldiff_spec, negb_andb, negb_involutive.
+ now rewrite N.land_spec, negb_andb.
+Qed.
+
+Lemma land_spec a b n :
+ testbit (land a b) n = testbit a n && testbit b n.
+Proof.
+ destruct (leb_spec 0 n) as [Hn|Hn]; [|now rewrite !testbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?testbit_0_l, ?andb_false_r; trivial; unfold land;
+ rewrite ?testbit_Zpos, ?testbit_Zneg, ?testbit_of_N', ?N.pos_pred_succ
+ by trivial.
+ now rewrite <- N.land_spec.
+ now rewrite N.ldiff_spec.
+ now rewrite N.ldiff_spec, andb_comm.
+ now rewrite N.lor_spec, negb_orb.
+Qed.
+
+Lemma ldiff_spec a b n :
+ testbit (ldiff a b) n = testbit a n && negb (testbit b n).
+Proof.
+ destruct (leb_spec 0 n) as [Hn|Hn]; [|now rewrite !testbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?testbit_0_l, ?andb_true_r; trivial; unfold ldiff;
+ rewrite ?testbit_Zpos, ?testbit_Zneg, ?testbit_of_N', ?N.pos_pred_succ
+ by trivial.
+ now rewrite <- N.ldiff_spec.
+ now rewrite N.land_spec, negb_involutive.
+ now rewrite N.lor_spec, negb_orb.
+ now rewrite N.ldiff_spec, negb_involutive, andb_comm.
+Qed.
+
+Lemma lxor_spec a b n :
+ testbit (lxor a b) n = xorb (testbit a n) (testbit b n).
+Proof.
+ destruct (leb_spec 0 n) as [Hn|Hn]; [|now rewrite !testbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?testbit_0_l, ?xorb_false_l, ?xorb_false_r; trivial; unfold lxor;
+ rewrite ?testbit_Zpos, ?testbit_Zneg, ?testbit_of_N', ?N.pos_pred_succ
+ by trivial.
+ now rewrite <- N.lxor_spec.
+ now rewrite N.lxor_spec, negb_xorb_r.
+ now rewrite N.lxor_spec, negb_xorb_l.
+ now rewrite N.lxor_spec, xorb_negb_negb.
+Qed.
+
+(** ** Induction principles based on successor / predecessor *)
+
+Lemma peano_ind (P : Z -> Prop) :
+ P 0 ->
+ (forall x, P x -> P (succ x)) ->
+ (forall x, P x -> P (pred x)) ->
+ forall z, P z.
+Proof.
+ intros H0 Hs Hp z; destruct z.
+ assumption.
+ induction p using Pos.peano_ind.
+ now apply (Hs 0).
+ rewrite <- Pos.add_1_r.
+ now apply (Hs (Zpos p)).
+ induction p using Pos.peano_ind.
+ now apply (Hp 0).
+ rewrite <- Pos.add_1_r.
+ now apply (Hp (Zneg p)).
+Qed.
+
+Lemma bi_induction (P : Z -> Prop) :
+ Proper (eq ==> iff) P ->
+ P 0 ->
+ (forall x, P x <-> P (succ x)) ->
+ forall z, P z.
+Proof.
+ intros _ H0 Hs. induction z using peano_ind.
+ assumption.
+ now apply -> Hs.
+ apply Hs. now rewrite succ_pred.
+Qed.
+
+
+(** * Proofs of morphisms, obvious since eq is Leibniz *)
+
+Local Obligation Tactic := simpl_relation.
+Program Definition succ_wd : Proper (eq==>eq) succ := _.
+Program Definition pred_wd : Proper (eq==>eq) pred := _.
+Program Definition opp_wd : Proper (eq==>eq) opp := _.
+Program Definition add_wd : Proper (eq==>eq==>eq) add := _.
+Program Definition sub_wd : Proper (eq==>eq==>eq) sub := _.
+Program Definition mul_wd : Proper (eq==>eq==>eq) mul := _.
+Program Definition lt_wd : Proper (eq==>eq==>iff) lt := _.
+Program Definition div_wd : Proper (eq==>eq==>eq) div := _.
+Program Definition mod_wd : Proper (eq==>eq==>eq) modulo := _.
+Program Definition quot_wd : Proper (eq==>eq==>eq) quot := _.
+Program Definition rem_wd : Proper (eq==>eq==>eq) rem := _.
+Program Definition pow_wd : Proper (eq==>eq==>eq) pow := _.
+Program Definition testbit_wd : Proper (eq==>eq==>Logic.eq) testbit := _.
+
+Include ZProp
+ <+ UsualMinMaxLogicalProperties <+ UsualMinMaxDecProperties.
+
+(** Otherwise Z stays associated with abstract_scope : (TODO FIX) *)
+Bind Scope Z_scope with Z.
+
+(** In generic statements, the predicates [lt] and [le] have been
+ favored, whereas [gt] and [ge] don't even exist in the abstract
+ layers. The use of [gt] and [ge] is hence not recommended. We provide
+ here the bare minimal results to related them with [lt] and [le]. *)
-Lemma Zplus_diag_eq_mult_2 : forall n:Z, n + n = n * Zpos 2.
+Lemma gt_lt_iff n m : n > m <-> m < n.
Proof.
- intros x; pattern x at 1 2 in |- *; rewrite <- (Zmult_1_r x);
- rewrite <- Zmult_plus_distr_r; reflexivity.
+ unfold lt, gt. now rewrite compare_antisym, CompOpp_iff.
Qed.
-(** ** Multiplication and successor *)
+Lemma gt_lt n m : n > m -> m < n.
+Proof.
+ apply gt_lt_iff.
+Qed.
-Lemma Zmult_succ_r : forall n m:Z, n * Zsucc m = n * m + n.
+Lemma lt_gt n m : n < m -> m > n.
Proof.
- intros n m; unfold Zsucc in |- *; rewrite Zmult_plus_distr_r;
- rewrite (Zmult_comm n (Zpos 1)); rewrite Zmult_1_l;
- trivial with arith.
+ apply gt_lt_iff.
Qed.
-Lemma Zmult_succ_r_reverse : forall n m:Z, n * m + n = n * Zsucc m.
+Lemma ge_le_iff n m : n >= m <-> m <= n.
Proof.
- intros; symmetry in |- *; apply Zmult_succ_r.
+ unfold le, ge. now rewrite compare_antisym, CompOpp_iff.
Qed.
-Lemma Zmult_succ_l : forall n m:Z, Zsucc n * m = n * m + m.
+Lemma ge_le n m : n >= m -> m <= n.
Proof.
- intros n m; unfold Zsucc in |- *; rewrite Zmult_plus_distr_l;
- rewrite Zmult_1_l; trivial with arith.
+ apply ge_le_iff.
Qed.
-Lemma Zmult_succ_l_reverse : forall n m:Z, n * m + m = Zsucc n * m.
+Lemma le_ge n m : n <= m -> m >= n.
Proof.
- intros; symmetry in |- *; apply Zmult_succ_l.
+ apply ge_le_iff.
Qed.
+(** We provide a tactic converting from one style to the other. *)
+Ltac swap_greater := rewrite ?gt_lt_iff in *; rewrite ?ge_le_iff in *.
-(** ** Misc redundant properties *)
+(** Similarly, the boolean comparisons [ltb] and [leb] are favored
+ over their dual [gtb] and [geb]. We prove here the equivalence
+ and a few minimal results. *)
-Lemma Z_eq_mult : forall n m:Z, m = Z0 -> m * n = Z0.
+Lemma gtb_ltb n m : (n >? m) = (m <? n).
Proof.
- intros x y H; rewrite H; auto with arith.
+ unfold gtb, ltb. rewrite compare_antisym. now case compare.
Qed.
+Lemma geb_leb n m : (n >=? m) = (m <=? n).
+Proof.
+ unfold geb, leb. rewrite compare_antisym. now case compare.
+Qed.
+Lemma gtb_lt n m : (n >? m) = true <-> m < n.
+Proof.
+ rewrite gtb_ltb. apply ltb_lt.
+Qed.
-(**********************************************************************)
-(** * Relating binary positive numbers and binary integers *)
+Lemma geb_le n m : (n >=? m) = true <-> m <= n.
+Proof.
+ rewrite geb_leb. apply leb_le.
+Qed.
-Lemma Zpos_eq : forall p q:positive, p = q -> Zpos p = Zpos q.
+Lemma gtb_spec n m : BoolSpec (m<n) (n<=m) (n >? m).
Proof.
- intros; f_equal; auto.
+ rewrite gtb_ltb. apply ltb_spec.
Qed.
-Lemma Zpos_eq_rev : forall p q:positive, Zpos p = Zpos q -> p = q.
+Lemma geb_spec n m : BoolSpec (m<=n) (n<m) (n >=? m).
Proof.
- inversion 1; auto.
+ rewrite geb_leb. apply leb_spec.
Qed.
-Lemma Zpos_eq_iff : forall p q:positive, p = q <-> Zpos p = Zpos q.
+(** TODO : to add in Numbers ? *)
+
+Lemma add_reg_l n m p : n + m = n + p -> m = p.
Proof.
- split; [apply Zpos_eq|apply Zpos_eq_rev].
+ exact (proj1 (add_cancel_l m p n)).
Qed.
-Lemma Zpos_xI : forall p:positive, Zpos p~1 = Zpos 2 * Zpos p + Zpos 1.
+Lemma mul_reg_l n m p : p <> 0 -> p * n = p * m -> n = m.
Proof.
- intro; apply refl_equal.
+ exact (fun Hp => proj1 (mul_cancel_l n m p Hp)).
Qed.
-Lemma Zpos_xO : forall p:positive, Zpos p~0 = Zpos 2 * Zpos p.
+Lemma mul_reg_r n m p : p <> 0 -> n * p = m * p -> n = m.
Proof.
- intro; apply refl_equal.
+ exact (fun Hp => proj1 (mul_cancel_r n m p Hp)).
Qed.
-Lemma Zneg_xI : forall p:positive, Zneg p~1 = Zpos 2 * Zneg p - Zpos 1.
+Lemma opp_eq_mul_m1 n : - n = n * -1.
Proof.
- intro; apply refl_equal.
+ rewrite mul_comm. now destruct n.
Qed.
-Lemma Zneg_xO : forall p:positive, Zneg p~0 = Zpos 2 * Zneg p.
+Lemma add_diag n : n + n = 2 * n.
Proof.
- reflexivity.
+ change 2 with (1+1). now rewrite mul_add_distr_r, !mul_1_l.
Qed.
-Lemma Zpos_plus_distr : forall p q:positive, Zpos (p + q) = Zpos p + Zpos q.
+(** * Comparison and opposite *)
+
+Lemma compare_opp n m : (- n ?= - m) = (m ?= n).
Proof.
- intros p p'; destruct p;
- [ destruct p' as [p0| p0| ]
- | destruct p' as [p0| p0| ]
- | destruct p' as [p| p| ] ]; reflexivity.
+ destruct n, m; simpl; trivial; intros; now rewrite <- Pos.compare_antisym.
Qed.
-Lemma Zneg_plus_distr : forall p q:positive, Zneg (p + q) = Zneg p + Zneg q.
+(** * Comparison and addition *)
+
+Lemma add_compare_mono_l n m p : (n + m ?= n + p) = (m ?= p).
Proof.
- intros p p'; destruct p;
- [ destruct p' as [p0| p0| ]
- | destruct p' as [p0| p0| ]
- | destruct p' as [p| p| ] ]; reflexivity.
+ rewrite (compare_sub m p), compare_sub. f_equal.
+ unfold sub. rewrite opp_add_distr, (add_comm n m), add_assoc.
+ f_equal. now rewrite <- add_assoc, add_opp_diag_r, add_0_r.
Qed.
-(**********************************************************************)
-(** * Order relations *)
+End Z.
+
+(** Export Notations *)
+
+Infix "+" := Z.add : Z_scope.
+Notation "- x" := (Z.opp x) : Z_scope.
+Infix "-" := Z.sub : Z_scope.
+Infix "*" := Z.mul : Z_scope.
+Infix "^" := Z.pow : Z_scope.
+Infix "/" := Z.div : Z_scope.
+Infix "mod" := Z.modulo (at level 40, no associativity) : Z_scope.
+Infix "÷" := Z.quot (at level 40, left associativity) : Z_scope.
+
+(* TODO : transition from Zdivide *)
+Notation "( x | y )" := (Z.divide x y) (at level 0).
-Definition Zlt (x y:Z) := (x ?= y) = Lt.
-Definition Zgt (x y:Z) := (x ?= y) = Gt.
-Definition Zle (x y:Z) := (x ?= y) <> Gt.
-Definition Zge (x y:Z) := (x ?= y) <> Lt.
-Definition Zne (x y:Z) := x <> y.
+Infix "?=" := Z.compare (at level 70, no associativity) : Z_scope.
-Infix "<=" := Zle : Z_scope.
-Infix "<" := Zlt : Z_scope.
-Infix ">=" := Zge : Z_scope.
-Infix ">" := Zgt : Z_scope.
+Infix "<=" := Z.le : Z_scope.
+Infix "<" := Z.lt : Z_scope.
+Infix ">=" := Z.ge : Z_scope.
+Infix ">" := Z.gt : Z_scope.
Notation "x <= y <= z" := (x <= y /\ y <= z) : Z_scope.
Notation "x <= y < z" := (x <= y /\ y < z) : Z_scope.
Notation "x < y < z" := (x < y /\ y < z) : Z_scope.
Notation "x < y <= z" := (x < y /\ y <= z) : Z_scope.
-(**********************************************************************)
-(** * Absolute value on integers *)
-
-Definition Zabs_nat (x:Z) : nat :=
- match x with
- | Z0 => 0%nat
- | Zpos p => nat_of_P p
- | Zneg p => nat_of_P p
- end.
-
-Definition Zabs (z:Z) : Z :=
- match z with
- | Z0 => Z0
- | Zpos p => Zpos p
- | Zneg p => Zpos p
- end.
-
-(**********************************************************************)
-(** * From [nat] to [Z] *)
-
-Definition Z_of_nat (x:nat) :=
- match x with
- | O => Z0
- | S y => Zpos (P_of_succ_nat y)
- end.
-
-Require Import BinNat.
-
-Definition Zabs_N (z:Z) :=
- match z with
- | Z0 => 0%N
- | Zpos p => Npos p
- | Zneg p => Npos p
- end.
-
-Definition Z_of_N (x:N) :=
- match x with
- | N0 => Z0
- | Npos p => Zpos p
- end.
+Infix "=?" := Z.eqb (at level 70, no associativity) : Z_scope.
+Infix "<=?" := Z.leb (at level 70, no associativity) : Z_scope.
+Infix "<?" := Z.ltb (at level 70, no associativity) : Z_scope.
+Infix ">=?" := Z.geb (at level 70, no associativity) : Z_scope.
+Infix ">?" := Z.gtb (at level 70, no associativity) : Z_scope.
+
+(** Compatibility Notations *)
+
+Notation Zdouble_plus_one := Z.succ_double (only parsing).
+Notation Zdouble_minus_one := Z.pred_double (only parsing).
+Notation Zdouble := Z.double (only parsing).
+Notation ZPminus := Z.pos_sub (only parsing).
+Notation Zsucc' := Z.succ (only parsing).
+Notation Zpred' := Z.pred (only parsing).
+Notation Zplus' := Z.add (only parsing).
+Notation Zplus := Z.add (only parsing). (* Slightly incompatible *)
+Notation Zopp := Z.opp (only parsing).
+Notation Zsucc := Z.succ (only parsing).
+Notation Zpred := Z.pred (only parsing).
+Notation Zminus := Z.sub (only parsing).
+Notation Zmult := Z.mul (only parsing).
+Notation Zcompare := Z.compare (only parsing).
+Notation Zsgn := Z.sgn (only parsing).
+Notation Zle := Z.le (only parsing).
+Notation Zge := Z.ge (only parsing).
+Notation Zlt := Z.lt (only parsing).
+Notation Zgt := Z.gt (only parsing).
+Notation Zmax := Z.max (only parsing).
+Notation Zmin := Z.min (only parsing).
+Notation Zabs := Z.abs (only parsing).
+Notation Zabs_nat := Z.abs_nat (only parsing).
+Notation Zabs_N := Z.abs_N (only parsing).
+Notation Z_of_nat := Z.of_nat (only parsing).
+Notation Z_of_N := Z.of_N (only parsing).
+
+Notation Zind := Z.peano_ind (only parsing).
+Notation Zopp_0 := Z.opp_0 (only parsing).
+Notation Zopp_neg := Z.opp_Zneg (only parsing).
+Notation Zopp_involutive := Z.opp_involutive (only parsing).
+Notation Zopp_inj := Z.opp_inj (only parsing).
+Notation Zplus_0_l := Z.add_0_l (only parsing).
+Notation Zplus_0_r := Z.add_0_r (only parsing).
+Notation Zplus_comm := Z.add_comm (only parsing).
+Notation Zopp_plus_distr := Z.opp_add_distr (only parsing).
+Notation Zopp_succ := Z.opp_succ (only parsing).
+Notation Zplus_opp_r := Z.add_opp_diag_r (only parsing).
+Notation Zplus_opp_l := Z.add_opp_diag_l (only parsing).
+Notation Zplus_assoc := Z.add_assoc (only parsing).
+Notation Zplus_permute := Z.add_shuffle3 (only parsing).
+Notation Zplus_reg_l := Z.add_reg_l (only parsing).
+Notation Zplus_succ_l := Z.add_succ_l (only parsing).
+Notation Zplus_succ_comm := Z.add_succ_comm (only parsing).
+Notation Zsucc_discr := Z.neq_succ_diag_r (only parsing).
+Notation Zsucc_inj := Z.succ_inj (only parsing).
+Notation Zsucc'_inj := Z.succ_inj (only parsing).
+Notation Zsucc'_pred' := Z.succ_pred (only parsing).
+Notation Zpred'_succ' := Z.pred_succ (only parsing).
+Notation Zpred'_inj := Z.pred_inj (only parsing).
+Notation Zsucc'_discr := Z.neq_succ_diag_r (only parsing).
+Notation Zminus_0_r := Z.sub_0_r (only parsing).
+Notation Zminus_diag := Z.sub_diag (only parsing).
+Notation Zminus_plus_distr := Z.sub_add_distr (only parsing).
+Notation Zminus_succ_r := Z.sub_succ_r (only parsing).
+Notation Zminus_plus := Z.add_simpl_l (only parsing).
+Notation Zmult_0_l := Z.mul_0_l (only parsing).
+Notation Zmult_0_r := Z.mul_0_r (only parsing).
+Notation Zmult_1_l := Z.mul_1_l (only parsing).
+Notation Zmult_1_r := Z.mul_1_r (only parsing).
+Notation Zmult_comm := Z.mul_comm (only parsing).
+Notation Zmult_assoc := Z.mul_assoc (only parsing).
+Notation Zmult_permute := Z.mul_shuffle3 (only parsing).
+Notation Zmult_1_inversion_l := Z.mul_eq_1 (only parsing).
+Notation Zdouble_mult := Z.double_spec (only parsing).
+Notation Zdouble_plus_one_mult := Z.succ_double_spec (only parsing).
+Notation Zopp_mult_distr_l_reverse := Z.mul_opp_l (only parsing).
+Notation Zmult_opp_opp := Z.mul_opp_opp (only parsing).
+Notation Zmult_opp_comm := Z.mul_opp_comm (only parsing).
+Notation Zopp_eq_mult_neg_1 := Z.opp_eq_mul_m1 (only parsing).
+Notation Zmult_plus_distr_r := Z.mul_add_distr_l (only parsing).
+Notation Zmult_plus_distr_l := Z.mul_add_distr_r (only parsing).
+Notation Zmult_minus_distr_r := Z.mul_sub_distr_r (only parsing).
+Notation Zmult_reg_l := Z.mul_reg_l (only parsing).
+Notation Zmult_reg_r := Z.mul_reg_r (only parsing).
+Notation Zmult_succ_l := Z.mul_succ_l (only parsing).
+Notation Zmult_succ_r := Z.mul_succ_r (only parsing).
+Notation Zpos_xI := Z.pos_xI (only parsing).
+Notation Zpos_xO := Z.pos_xO (only parsing).
+Notation Zneg_xI := Z.neg_xI (only parsing).
+Notation Zneg_xO := Z.neg_xO (only parsing).
+
+Notation Z := Z (only parsing).
+Notation Z_rect := Z_rect (only parsing).
+Notation Z_rec := Z_rec (only parsing).
+Notation Z_ind := Z_ind (only parsing).
+Notation Z0 := Z0 (only parsing).
+Notation Zpos := Zpos (only parsing).
+Notation Zneg := Zneg (only parsing).
+
+(** Compatibility lemmas. These could be notations,
+ but scope information would be lost.
+*)
+
+Notation SYM1 lem := (fun n => eq_sym (lem n)).
+Notation SYM2 lem := (fun n m => eq_sym (lem n m)).
+Notation SYM3 lem := (fun n m p => eq_sym (lem n m p)).
+
+Lemma Zplus_assoc_reverse : forall n m p, n+m+p = n+(m+p).
+Proof (SYM3 Z.add_assoc).
+Lemma Zplus_succ_r_reverse : forall n m, Z.succ (n+m) = n+Z.succ m.
+Proof (SYM2 Z.add_succ_r).
+Notation Zplus_succ_r := Zplus_succ_r_reverse (only parsing).
+Lemma Zplus_0_r_reverse : forall n, n = n + 0.
+Proof (SYM1 Z.add_0_r).
+Lemma Zplus_eq_compat : forall n m p q, n=m -> p=q -> n+p=m+q.
+Proof (f_equal2 Z.add).
+Lemma Zpos_succ_morphism : forall p, Zpos (Psucc p) = Zsucc (Zpos p).
+Proof (SYM1 Z.succ_Zpos).
+Lemma Zsucc_pred : forall n, n = Z.succ (Z.pred n).
+Proof (SYM1 Z.succ_pred).
+Lemma Zpred_succ : forall n, n = Z.pred (Z.succ n).
+Proof (SYM1 Z.pred_succ).
+Lemma Zsucc_eq_compat : forall n m, n = m -> Z.succ n = Z.succ m.
+Proof (f_equal Z.succ).
+Lemma Zminus_0_l_reverse : forall n, n = n - 0.
+Proof (SYM1 Z.sub_0_r).
+Lemma Zminus_diag_reverse : forall n, 0 = n-n.
+Proof (SYM1 Z.sub_diag).
+Lemma Zminus_succ_l : forall n m, Z.succ (n - m) = Z.succ n - m.
+Proof (SYM2 Z.sub_succ_l).
+Lemma Zplus_minus_eq : forall n m p, n = m + p -> p = n - m.
+Proof. intros. now apply Z.add_move_l. Qed.
+Lemma Zplus_minus : forall n m, n + (m - n) = m.
+Proof (fun n m => eq_trans (Z.add_comm n (m-n)) (Z.sub_add n m)).
+Lemma Zminus_plus_simpl_l : forall n m p, p + n - (p + m) = n - m.
+Proof (fun n m p => Z.add_add_simpl_l_l p n m).
+Lemma Zminus_plus_simpl_l_reverse : forall n m p, n - m = p + n - (p + m).
+Proof (SYM3 Zminus_plus_simpl_l).
+Lemma Zminus_plus_simpl_r : forall n m p, n + p - (m + p) = n - m.
+Proof (fun n m p => Z.add_add_simpl_r_r n p m).
+Lemma Zpos_minus_morphism : forall a b,
+ Pcompare a b Eq = Lt -> Zpos (b - a) = Zpos b - Zpos a.
+Proof. intros. now rewrite Z.sub_Zpos. Qed.
+Lemma Zeq_minus : forall n m, n = m -> n - m = 0.
+Proof (fun n m => proj2 (Z.sub_move_0_r n m)).
+Lemma Zminus_eq : forall n m, n - m = 0 -> n = m.
+Proof (fun n m => proj1 (Z.sub_move_0_r n m)).
+Lemma Zpos_mult_morphism : forall p q, Zpos (p * q) = Zpos p * Zpos q.
+Proof (SYM2 Z.mul_Zpos).
+Lemma Zmult_0_r_reverse : forall n, 0 = n * 0.
+Proof (SYM1 Z.mul_0_r).
+Lemma Zmult_assoc_reverse : forall n m p, n * m * p = n * (m * p).
+Proof (SYM3 Z.mul_assoc).
+Lemma Zmult_integral : forall n m, n * m = 0 -> n = 0 \/ m = 0.
+Proof (fun n m => proj1 (Z.mul_eq_0 n m)).
+Lemma Zmult_integral_l : forall n m, n <> 0 -> m * n = 0 -> m = 0.
+Proof (fun n m H H' => Z.mul_eq_0_l m n H' H).
+Lemma Zopp_mult_distr_l : forall n m, - (n * m) = - n * m.
+Proof (SYM2 Z.mul_opp_l).
+Lemma Zopp_mult_distr_r : forall n m, - (n * m) = n * - m.
+Proof (SYM2 Z.mul_opp_r).
+Lemma Zmult_minus_distr_l : forall n m p, p * (n - m) = p * n - p * m.
+Proof (fun n m p => Z.mul_sub_distr_l p n m).
+Lemma Zmult_succ_r_reverse : forall n m, n * m + n = n * Zsucc m.
+Proof (SYM2 Z.mul_succ_r).
+Lemma Zmult_succ_l_reverse : forall n m, n * m + m = Zsucc n * m.
+Proof (SYM2 Z.mul_succ_l).
+Lemma Zpos_eq : forall p q, p = q -> Zpos p = Zpos q.
+Proof (fun p q => proj2 (Z.inj_Zpos p q)).
+Lemma Zpos_eq_rev : forall p q, Zpos p = Zpos q -> p = q.
+Proof (fun p q => proj1 (Z.inj_Zpos p q)).
+Lemma Zpos_eq_iff : forall p q, p = q <-> Zpos p = Zpos q.
+Proof (fun p q => iff_sym (Z.inj_Zpos p q)).
+Lemma Zpos_plus_distr : forall p q, Zpos (p + q) = Zpos p + Zpos q.
+Proof (SYM2 Z.add_Zpos).
+Lemma Zneg_plus_distr : forall p q, Zneg (p + q) = Zneg p + Zneg q.
+Proof (SYM2 Z.add_Zneg).
+
+Hint Immediate Zsucc_pred: zarith.
+
+(* Not kept :
+Zplus_0_simpl_l
+Zplus_0_simpl_l_reverse
+Zplus_opp_expand
+Zsucc_inj_contrapositive
+Zsucc_succ'
+Zpred_pred'
+*)
+
+(* No compat notation for :
+weak_assoc (now Z.add_assoc_pos)
+weak_Zmult_plus_distr_r (now Z.mul_add_distr_pos)
+*)
+
+(** Obsolete stuff *)
+
+Definition Zne (x y:Z) := x <> y. (* TODO : to remove someday ? *)
+
+Ltac elim_compare com1 com2 :=
+ case (Dcompare (com1 ?= com2)%Z);
+ [ idtac | let x := fresh "H" in
+ (intro x; case x; clear x) ].
+
+Lemma ZL0 : 2%nat = (1 + 1)%nat.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma Zplus_diag_eq_mult_2 n : n + n = n * 2.
+Proof.
+ rewrite Z.mul_comm. apply Z.add_diag.
+Qed.
+
+Lemma Z_eq_mult n m : m = 0 -> m * n = 0.
+Proof.
+ intros; now subst.
+Qed.
diff --git a/theories/ZArith/BinIntDef.v b/theories/ZArith/BinIntDef.v
new file mode 100644
index 00000000..d96d20fb
--- /dev/null
+++ b/theories/ZArith/BinIntDef.v
@@ -0,0 +1,610 @@
+(* -*- coding: utf-8 -*- *)
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Export BinNums.
+Require Import BinPos BinNat.
+
+Local Open Scope Z_scope.
+
+(***********************************************************)
+(** * Binary Integers, Definitions of Operations *)
+(***********************************************************)
+
+(** Initial author: Pierre Crégut, CNET, Lannion, France *)
+
+Module Z.
+
+Definition t := Z.
+
+(** ** Constants *)
+
+Definition zero := 0.
+Definition one := 1.
+Definition two := 2.
+
+(** ** Doubling and variants *)
+
+Definition double x :=
+ match x with
+ | 0 => 0
+ | Zpos p => Zpos p~0
+ | Zneg p => Zneg p~0
+ end.
+
+Definition succ_double x :=
+ match x with
+ | 0 => 1
+ | Zpos p => Zpos p~1
+ | Zneg p => Zneg (Pos.pred_double p)
+ end.
+
+Definition pred_double x :=
+ match x with
+ | 0 => -1
+ | Zneg p => Zneg p~1
+ | Zpos p => Zpos (Pos.pred_double p)
+ end.
+
+(** ** Subtraction of positive into Z *)
+
+Fixpoint pos_sub (x y:positive) {struct y} : Z :=
+ match x, y with
+ | p~1, q~1 => double (pos_sub p q)
+ | p~1, q~0 => succ_double (pos_sub p q)
+ | p~1, 1 => Zpos p~0
+ | p~0, q~1 => pred_double (pos_sub p q)
+ | p~0, q~0 => double (pos_sub p q)
+ | p~0, 1 => Zpos (Pos.pred_double p)
+ | 1, q~1 => Zneg q~0
+ | 1, q~0 => Zneg (Pos.pred_double q)
+ | 1, 1 => Z0
+ end%positive.
+
+(** ** Addition *)
+
+Definition add x y :=
+ match x, y with
+ | 0, y => y
+ | x, 0 => x
+ | Zpos x', Zpos y' => Zpos (x' + y')
+ | Zpos x', Zneg y' => pos_sub x' y'
+ | Zneg x', Zpos y' => pos_sub y' x'
+ | Zneg x', Zneg y' => Zneg (x' + y')
+ end.
+
+Infix "+" := add : Z_scope.
+
+(** ** Opposite *)
+
+Definition opp x :=
+ match x with
+ | 0 => 0
+ | Zpos x => Zneg x
+ | Zneg x => Zpos x
+ end.
+
+Notation "- x" := (opp x) : Z_scope.
+
+(** ** Successor *)
+
+Definition succ x := x + 1.
+
+(** ** Predecessor *)
+
+Definition pred x := x + -1.
+
+(** ** Subtraction *)
+
+Definition sub m n := m + -n.
+
+Infix "-" := sub : Z_scope.
+
+(** ** Multiplication *)
+
+Definition mul x y :=
+ match x, y with
+ | 0, _ => 0
+ | _, 0 => 0
+ | Zpos x', Zpos y' => Zpos (x' * y')
+ | Zpos x', Zneg y' => Zneg (x' * y')
+ | Zneg x', Zpos y' => Zneg (x' * y')
+ | Zneg x', Zneg y' => Zpos (x' * y')
+ end.
+
+Infix "*" := mul : Z_scope.
+
+(** ** Power function *)
+
+Definition pow_pos (z:Z) (n:positive) := Pos.iter n (mul z) 1.
+
+Definition pow x y :=
+ match y with
+ | Zpos p => pow_pos x p
+ | 0 => 1
+ | Zneg _ => 0
+ end.
+
+Infix "^" := pow : Z_scope.
+
+(** ** Square *)
+
+Definition square x :=
+ match x with
+ | 0 => 0
+ | Zpos p => Zpos (Pos.square p)
+ | Zneg p => Zpos (Pos.square p)
+ end.
+
+(** ** Comparison *)
+
+Definition compare x y :=
+ match x, y with
+ | 0, 0 => Eq
+ | 0, Zpos y' => Lt
+ | 0, Zneg y' => Gt
+ | Zpos x', 0 => Gt
+ | Zpos x', Zpos y' => (x' ?= y')%positive
+ | Zpos x', Zneg y' => Gt
+ | Zneg x', 0 => Lt
+ | Zneg x', Zpos y' => Lt
+ | Zneg x', Zneg y' => CompOpp ((x' ?= y')%positive)
+ end.
+
+Infix "?=" := compare (at level 70, no associativity) : Z_scope.
+
+(** ** Sign function *)
+
+Definition sgn z :=
+ match z with
+ | 0 => 0
+ | Zpos p => 1
+ | Zneg p => -1
+ end.
+
+(** Boolean equality and comparisons *)
+
+Definition leb x y :=
+ match x ?= y with
+ | Gt => false
+ | _ => true
+ end.
+
+Definition ltb x y :=
+ match x ?= y with
+ | Lt => true
+ | _ => false
+ end.
+
+(** Nota: [geb] and [gtb] are provided for compatibility,
+ but [leb] and [ltb] should rather be used instead, since
+ more results we be available on them. *)
+
+Definition geb x y :=
+ match x ?= y with
+ | Lt => false
+ | _ => true
+ end.
+
+Definition gtb x y :=
+ match x ?= y with
+ | Gt => true
+ | _ => false
+ end.
+
+(** Nota: this [eqb] is not convertible with the generated [Z_beq],
+ since the underlying [Pos.eqb] differs from [positive_beq]
+ (cf BinIntDef). *)
+
+Fixpoint eqb x y :=
+ match x, y with
+ | 0, 0 => true
+ | Zpos p, Zpos q => Pos.eqb p q
+ | Zneg p, Zneg q => Pos.eqb p q
+ | _, _ => false
+ end.
+
+Infix "=?" := eqb (at level 70, no associativity) : Z_scope.
+Infix "<=?" := leb (at level 70, no associativity) : Z_scope.
+Infix "<?" := ltb (at level 70, no associativity) : Z_scope.
+Infix ">=?" := geb (at level 70, no associativity) : Z_scope.
+Infix ">?" := gtb (at level 70, no associativity) : Z_scope.
+
+(** ** Minimum and maximum *)
+
+Definition max n m :=
+ match n ?= m with
+ | Eq | Gt => n
+ | Lt => m
+ end.
+
+Definition min n m :=
+ match n ?= m with
+ | Eq | Lt => n
+ | Gt => m
+ end.
+
+(** ** Absolute value *)
+
+Definition abs z :=
+ match z with
+ | 0 => 0
+ | Zpos p => Zpos p
+ | Zneg p => Zpos p
+ end.
+
+(** ** Conversions *)
+
+(** From [Z] to [nat] via absolute value *)
+
+Definition abs_nat (z:Z) : nat :=
+ match z with
+ | 0 => 0%nat
+ | Zpos p => Pos.to_nat p
+ | Zneg p => Pos.to_nat p
+ end.
+
+(** From [Z] to [N] via absolute value *)
+
+Definition abs_N (z:Z) : N :=
+ match z with
+ | Z0 => 0%N
+ | Zpos p => Npos p
+ | Zneg p => Npos p
+ end.
+
+(** From [Z] to [nat] by rounding negative numbers to 0 *)
+
+Definition to_nat (z:Z) : nat :=
+ match z with
+ | Zpos p => Pos.to_nat p
+ | _ => O
+ end.
+
+(** From [Z] to [N] by rounding negative numbers to 0 *)
+
+Definition to_N (z:Z) : N :=
+ match z with
+ | Zpos p => Npos p
+ | _ => 0%N
+ end.
+
+(** From [nat] to [Z] *)
+
+Definition of_nat (n:nat) : Z :=
+ match n with
+ | O => 0
+ | S n => Zpos (Pos.of_succ_nat n)
+ end.
+
+(** From [N] to [Z] *)
+
+Definition of_N (n:N) : Z :=
+ match n with
+ | N0 => 0
+ | Npos p => Zpos p
+ end.
+
+(** ** Iteration of a function
+
+ By convention, iterating a negative number of times is identity.
+*)
+
+Definition iter (n:Z) {A} (f:A -> A) (x:A) :=
+ match n with
+ | Zpos p => Pos.iter p f x
+ | _ => x
+ end.
+
+(** ** Euclidean divisions for binary integers *)
+
+(** Concerning the many possible variants of integer divisions,
+ see the headers of the generic files [ZDivFloor], [ZDivTrunc],
+ [ZDivEucl], and the article by R. Boute mentioned there.
+ We provide here two flavours, Floor and Trunc, while
+ the Euclid convention can be found in file Zeuclid.v
+ For non-zero b, they all satisfy [a = b*(a/b) + (a mod b)]
+ and [ |a mod b| < |b| ], but the sign of the modulo will differ
+ when [a<0] and/or [b<0].
+*)
+
+(** ** Floor division *)
+
+(** [div_eucl] provides a Truncated-Toward-Bottom (a.k.a Floor)
+ Euclidean division. Its projections are named [div] (noted "/")
+ and [modulo] (noted with an infix "mod").
+ These functions correspond to the `div` and `mod` of Haskell.
+ This is the historical convention of Coq.
+
+ The main properties of this convention are :
+ - we have [sgn (a mod b) = sgn (b)]
+ - [div a b] is the greatest integer smaller or equal to the exact
+ fraction [a/b].
+ - there is no easy sign rule.
+
+ In addition, note that we arbitrary take [a/0 = 0] and [a mod 0 = 0].
+*)
+
+(** First, a division for positive numbers. Even if the second
+ argument is a Z, the answer is arbitrary is it isn't a Zpos. *)
+
+Fixpoint pos_div_eucl (a:positive) (b:Z) : Z * Z :=
+ match a with
+ | xH => if 2 <=? b then (0, 1) else (1, 0)
+ | xO a' =>
+ let (q, r) := pos_div_eucl a' b in
+ let r' := 2 * r in
+ if r' <? b then (2 * q, r') else (2 * q + 1, r' - b)
+ | xI a' =>
+ let (q, r) := pos_div_eucl a' b in
+ let r' := 2 * r + 1 in
+ if r' <? b then (2 * q, r') else (2 * q + 1, r' - b)
+ end.
+
+(** Then the general euclidean division *)
+
+Definition div_eucl (a b:Z) : Z * Z :=
+ match a, b with
+ | 0, _ => (0, 0)
+ | _, 0 => (0, 0)
+ | Zpos a', Zpos _ => pos_div_eucl a' b
+ | Zneg a', Zpos _ =>
+ let (q, r) := pos_div_eucl a' b in
+ match r with
+ | 0 => (- q, 0)
+ | _ => (- (q + 1), b - r)
+ end
+ | Zneg a', Zneg b' =>
+ let (q, r) := pos_div_eucl a' (Zpos b') in (q, - r)
+ | Zpos a', Zneg b' =>
+ let (q, r) := pos_div_eucl a' (Zpos b') in
+ match r with
+ | 0 => (- q, 0)
+ | _ => (- (q + 1), b + r)
+ end
+ end.
+
+Definition div (a b:Z) : Z := let (q, _) := div_eucl a b in q.
+Definition modulo (a b:Z) : Z := let (_, r) := div_eucl a b in r.
+
+Infix "/" := div : Z_scope.
+Infix "mod" := modulo (at level 40, no associativity) : Z_scope.
+
+
+(** ** Trunc Division *)
+
+(** [quotrem] provides a Truncated-Toward-Zero Euclidean division.
+ Its projections are named [quot] (noted "÷") and [rem].
+ These functions correspond to the `quot` and `rem` of Haskell.
+ This division convention is used in most programming languages,
+ e.g. Ocaml.
+
+ With this convention:
+ - we have [sgn(a rem b) = sgn(a)]
+ - sign rule for division: [quot (-a) b = quot a (-b) = -(quot a b)]
+ - and for modulo: [a rem (-b) = a rem b] and [(-a) rem b = -(a rem b)]
+
+ Note that we arbitrary take here [quot a 0 = 0] and [a rem 0 = a].
+*)
+
+Definition quotrem (a b:Z) : Z * Z :=
+ match a, b with
+ | 0, _ => (0, 0)
+ | _, 0 => (0, a)
+ | Zpos a, Zpos b =>
+ let (q, r) := N.pos_div_eucl a (Npos b) in (of_N q, of_N r)
+ | Zneg a, Zpos b =>
+ let (q, r) := N.pos_div_eucl a (Npos b) in (-of_N q, - of_N r)
+ | Zpos a, Zneg b =>
+ let (q, r) := N.pos_div_eucl a (Npos b) in (-of_N q, of_N r)
+ | Zneg a, Zneg b =>
+ let (q, r) := N.pos_div_eucl a (Npos b) in (of_N q, - of_N r)
+ end.
+
+Definition quot a b := fst (quotrem a b).
+Definition rem a b := snd (quotrem a b).
+
+Infix "÷" := quot (at level 40, left associativity) : Z_scope.
+(** No infix notation for rem, otherwise it becomes a keyword *)
+
+
+(** ** Parity functions *)
+
+Definition even z :=
+ match z with
+ | 0 => true
+ | Zpos (xO _) => true
+ | Zneg (xO _) => true
+ | _ => false
+ end.
+
+Definition odd z :=
+ match z with
+ | 0 => false
+ | Zpos (xO _) => false
+ | Zneg (xO _) => false
+ | _ => true
+ end.
+
+
+(** ** Division by two *)
+
+(** [div2] performs rounding toward bottom, it is hence a particular
+ case of [div], and for all relative number [n] we have:
+ [n = 2 * div2 n + if odd n then 1 else 0]. *)
+
+Definition div2 z :=
+ match z with
+ | 0 => 0
+ | Zpos 1 => 0
+ | Zpos p => Zpos (Pos.div2 p)
+ | Zneg p => Zneg (Pos.div2_up p)
+ end.
+
+(** [quot2] performs rounding toward zero, it is hence a particular
+ case of [quot], and for all relative number [n] we have:
+ [n = 2 * quot2 n + if odd n then sgn n else 0]. *)
+
+Definition quot2 (z:Z) :=
+ match z with
+ | 0 => 0
+ | Zpos 1 => 0
+ | Zpos p => Zpos (Pos.div2 p)
+ | Zneg 1 => 0
+ | Zneg p => Zneg (Pos.div2 p)
+ end.
+
+(** NB: [Z.quot2] used to be named [Zdiv2] in Coq <= 8.3 *)
+
+
+(** * Base-2 logarithm *)
+
+Definition log2 z :=
+ match z with
+ | Zpos (p~1) => Zpos (Pos.size p)
+ | Zpos (p~0) => Zpos (Pos.size p)
+ | _ => 0
+ end.
+
+
+(** ** Square root *)
+
+Definition sqrtrem n :=
+ match n with
+ | 0 => (0, 0)
+ | Zpos p =>
+ match Pos.sqrtrem p with
+ | (s, IsPos r) => (Zpos s, Zpos r)
+ | (s, _) => (Zpos s, 0)
+ end
+ | Zneg _ => (0,0)
+ end.
+
+Definition sqrt n :=
+ match n with
+ | Zpos p => Zpos (Pos.sqrt p)
+ | _ => 0
+ end.
+
+
+(** ** Greatest Common Divisor *)
+
+Definition gcd a b :=
+ match a,b with
+ | 0, _ => abs b
+ | _, 0 => abs a
+ | Zpos a, Zpos b => Zpos (Pos.gcd a b)
+ | Zpos a, Zneg b => Zpos (Pos.gcd a b)
+ | Zneg a, Zpos b => Zpos (Pos.gcd a b)
+ | Zneg a, Zneg b => Zpos (Pos.gcd a b)
+ end.
+
+(** A generalized gcd, also computing division of a and b by gcd. *)
+
+Definition ggcd a b : Z*(Z*Z) :=
+ match a,b with
+ | 0, _ => (abs b,(0, sgn b))
+ | _, 0 => (abs a,(sgn a, 0))
+ | Zpos a, Zpos b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zpos aa, Zpos bb))
+ | Zpos a, Zneg b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zpos aa, Zneg bb))
+ | Zneg a, Zpos b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zneg aa, Zpos bb))
+ | Zneg a, Zneg b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zneg aa, Zneg bb))
+ end.
+
+
+(** ** Bitwise functions *)
+
+(** When accessing the bits of negative numbers, all functions
+ below will use the two's complement representation. For instance,
+ [-1] will correspond to an infinite stream of true bits. If this
+ isn't what you're looking for, you can use [abs] first and then
+ access the bits of the absolute value.
+*)
+
+(** [testbit] : accessing the [n]-th bit of a number [a].
+ For negative [n], we arbitrarily answer [false]. *)
+
+Definition testbit a n :=
+ match n with
+ | 0 => odd a
+ | Zpos p =>
+ match a with
+ | 0 => false
+ | Zpos a => Pos.testbit a (Npos p)
+ | Zneg a => negb (N.testbit (Pos.pred_N a) (Npos p))
+ end
+ | Zneg _ => false
+ end.
+
+(** Shifts
+
+ Nota: a shift to the right by [-n] will be a shift to the left
+ by [n], and vice-versa.
+
+ For fulfilling the two's complement convention, shifting to
+ the right a negative number should correspond to a division
+ by 2 with rounding toward bottom, hence the use of [div2]
+ instead of [quot2].
+*)
+
+Definition shiftl a n :=
+ match n with
+ | 0 => a
+ | Zpos p => Pos.iter p (mul 2) a
+ | Zneg p => Pos.iter p div2 a
+ end.
+
+Definition shiftr a n := shiftl a (-n).
+
+(** Bitwise operations [lor] [land] [ldiff] [lxor] *)
+
+Definition lor a b :=
+ match a, b with
+ | 0, _ => b
+ | _, 0 => a
+ | Zpos a, Zpos b => Zpos (Pos.lor a b)
+ | Zneg a, Zpos b => Zneg (N.succ_pos (N.ldiff (Pos.pred_N a) (Npos b)))
+ | Zpos a, Zneg b => Zneg (N.succ_pos (N.ldiff (Pos.pred_N b) (Npos a)))
+ | Zneg a, Zneg b => Zneg (N.succ_pos (N.land (Pos.pred_N a) (Pos.pred_N b)))
+ end.
+
+Definition land a b :=
+ match a, b with
+ | 0, _ => 0
+ | _, 0 => 0
+ | Zpos a, Zpos b => of_N (Pos.land a b)
+ | Zneg a, Zpos b => of_N (N.ldiff (Npos b) (Pos.pred_N a))
+ | Zpos a, Zneg b => of_N (N.ldiff (Npos a) (Pos.pred_N b))
+ | Zneg a, Zneg b => Zneg (N.succ_pos (N.lor (Pos.pred_N a) (Pos.pred_N b)))
+ end.
+
+Definition ldiff a b :=
+ match a, b with
+ | 0, _ => 0
+ | _, 0 => a
+ | Zpos a, Zpos b => of_N (Pos.ldiff a b)
+ | Zneg a, Zpos b => Zneg (N.succ_pos (N.lor (Pos.pred_N a) (Npos b)))
+ | Zpos a, Zneg b => of_N (N.land (Npos a) (Pos.pred_N b))
+ | Zneg a, Zneg b => of_N (N.ldiff (Pos.pred_N b) (Pos.pred_N a))
+ end.
+
+Definition lxor a b :=
+ match a, b with
+ | 0, _ => b
+ | _, 0 => a
+ | Zpos a, Zpos b => of_N (Pos.lxor a b)
+ | Zneg a, Zpos b => Zneg (N.succ_pos (N.lxor (Pos.pred_N a) (Npos b)))
+ | Zpos a, Zneg b => Zneg (N.succ_pos (N.lxor (Npos a) (Pos.pred_N b)))
+ | Zneg a, Zneg b => of_N (N.lxor (Pos.pred_N a) (Pos.pred_N b))
+ end.
+
+End Z. \ No newline at end of file
diff --git a/theories/ZArith/Int.v b/theories/ZArith/Int.v
index c0123ca8..bac50fc4 100644
--- a/theories/ZArith/Int.v
+++ b/theories/ZArith/Int.v
@@ -6,8 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: Int.v 12363 2009-09-28 15:04:07Z letouzey $ *)
-
(** * An light axiomatization of integers (used in FSetAVL). *)
(** We define a signature for an integer datatype based on [Z].
@@ -29,7 +27,7 @@ Module Type Int.
Parameter int : Set.
Parameter i2z : int -> Z.
- Arguments Scope i2z [ Int_scope ].
+ Arguments i2z _%I.
Parameter _0 : int.
Parameter _1 : int.
@@ -222,10 +220,10 @@ Module MoreInt (I:Int).
| (?x \/ ?y) => let ex := p2ep x with ey := p2ep y in constr:(EPor ex ey)
| (~ ?x) => let ex := p2ep x in constr:(EPneg ex)
| (eq (A:=Z) ?x ?y) => let ex := z2ez x with ey := z2ez y in constr:(EPeq ex ey)
- | (?x<?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPlt ex ey)
- | (?x<=?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPle ex ey)
- | (?x>?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPgt ex ey)
- | (?x>=?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPge ex ey)
+ | (?x < ?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPlt ex ey)
+ | (?x <= ?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPle ex ey)
+ | (?x > ?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPgt ex ey)
+ | (?x >= ?y)%Z => let ex := z2ez x with ey := z2ez y in constr:(EPge ex ey)
| ?x => constr:(EPraw x)
end.
diff --git a/theories/ZArith/Wf_Z.v b/theories/ZArith/Wf_Z.v
index 0fe6d623..bcccc126 100644
--- a/theories/ZArith/Wf_Z.v
+++ b/theories/ZArith/Wf_Z.v
@@ -1,123 +1,83 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Wf_Z.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import BinInt.
Require Import Zcompare.
Require Import Zorder.
Require Import Znat.
Require Import Zmisc.
Require Import Wf_nat.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
(** Our purpose is to write an induction shema for {0,1,2,...}
similar to the [nat] schema (Theorem [Natlike_rec]). For that the
following implications will be used :
<<
- (n:nat)(Q n)==(n:nat)(P (inject_nat n)) ===> (x:Z)`x > 0) -> (P x)
+ ∀n:nat, Q n == ∀n:nat, P (Z.of_nat n) ===> ∀x:Z, x <= 0 -> P x
/\
||
||
- (Q O) (n:nat)(Q n)->(Q (S n)) <=== (P 0) (x:Z) (P x) -> (P (Zs x))
+ (Q O) ∧ (∀n:nat, Q n -> Q (S n)) <=== (P 0) ∧ (∀x:Z, P x -> P (Z.succ x))
- <=== (inject_nat (S n))=(Zs (inject_nat n))
+ <=== (Z.of_nat (S n) = Z.succ (Z.of_nat n))
- <=== inject_nat_complete
+ <=== Z_of_nat_complete
>>
Then the diagram will be closed and the theorem proved. *)
-Lemma Z_of_nat_complete :
- forall x:Z, 0 <= x -> exists n : nat, x = Z_of_nat n.
-Proof.
- intro x; destruct x; intros;
- [ exists 0%nat; auto with arith
- | specialize (ZL4 p); intros Hp; elim Hp; intros; exists (S x); intros;
- simpl in |- *; specialize (nat_of_P_o_P_of_succ_nat_eq_succ x);
- intro Hx0; rewrite <- H0 in Hx0; apply f_equal with (f := Zpos);
- apply nat_of_P_inj; auto with arith
- | absurd (0 <= Zneg p);
- [ unfold Zle in |- *; simpl in |- *; do 2 unfold not in |- *;
- auto with arith
- | assumption ] ].
-Qed.
-
-Lemma ZL4_inf : forall y:positive, {h : nat | nat_of_P y = S h}.
+Lemma Z_of_nat_complete (x : Z) :
+ 0 <= x -> exists n : nat, x = Z.of_nat n.
Proof.
- intro y; induction y as [p H| p H1| ];
- [ elim H; intros x H1; exists (S x + S x)%nat; unfold nat_of_P in |- *;
- simpl in |- *; rewrite ZL0; rewrite Pmult_nat_r_plus_morphism;
- unfold nat_of_P in H1; rewrite H1; auto with arith
- | elim H1; intros x H2; exists (x + S x)%nat; unfold nat_of_P in |- *;
- simpl in |- *; rewrite ZL0; rewrite Pmult_nat_r_plus_morphism;
- unfold nat_of_P in H2; rewrite H2; auto with arith
- | exists 0%nat; auto with arith ].
+ intros H. exists (Z.to_nat x). symmetry. now apply Z2Nat.id.
Qed.
-Lemma Z_of_nat_complete_inf :
- forall x:Z, 0 <= x -> {n : nat | x = Z_of_nat n}.
+Lemma Z_of_nat_complete_inf (x : Z) :
+ 0 <= x -> {n : nat | x = Z_of_nat n}.
Proof.
- intro x; destruct x; intros;
- [ exists 0%nat; auto with arith
- | specialize (ZL4_inf p); intros Hp; elim Hp; intros x0 H0; exists (S x0);
- intros; simpl in |- *; specialize (nat_of_P_o_P_of_succ_nat_eq_succ x0);
- intro Hx0; rewrite <- H0 in Hx0; apply f_equal with (f := Zpos);
- apply nat_of_P_inj; auto with arith
- | absurd (0 <= Zneg p);
- [ unfold Zle in |- *; simpl in |- *; do 2 unfold not in |- *;
- auto with arith
- | assumption ] ].
+ intros H. exists (Z.to_nat x). symmetry. now apply Z2Nat.id.
Qed.
Lemma Z_of_nat_prop :
forall P:Z -> Prop,
- (forall n:nat, P (Z_of_nat n)) -> forall x:Z, 0 <= x -> P x.
+ (forall n:nat, P (Z.of_nat n)) -> forall x:Z, 0 <= x -> P x.
Proof.
- intros P H x H0.
- specialize (Z_of_nat_complete x H0).
- intros Hn; elim Hn; intros.
- rewrite H1; apply H.
+ intros P H x Hx. now destruct (Z_of_nat_complete x Hx) as (n,->).
Qed.
Lemma Z_of_nat_set :
forall P:Z -> Set,
(forall n:nat, P (Z_of_nat n)) -> forall x:Z, 0 <= x -> P x.
Proof.
- intros P H x H0.
- specialize (Z_of_nat_complete_inf x H0).
- intros Hn; elim Hn; intros.
- rewrite p; apply H.
+ intros P H x Hx. now destruct (Z_of_nat_complete_inf x Hx) as (n,->).
Qed.
Lemma natlike_ind :
forall P:Z -> Prop,
P 0 ->
- (forall x:Z, 0 <= x -> P x -> P (Zsucc x)) -> forall x:Z, 0 <= x -> P x.
+ (forall x:Z, 0 <= x -> P x -> P (Z.succ x)) ->
+ forall x:Z, 0 <= x -> P x.
Proof.
- intros P H H0 x H1; apply Z_of_nat_prop;
- [ simple induction n;
- [ simpl in |- *; assumption
- | intros; rewrite (inj_S n0); exact (H0 (Z_of_nat n0) (Zle_0_nat n0) H2) ]
- | assumption ].
+ intros P Ho Hrec x Hx; apply Z_of_nat_prop; trivial.
+ induction n. exact Ho.
+ rewrite Nat2Z.inj_succ. apply Hrec; trivial using Nat2Z.is_nonneg.
Qed.
Lemma natlike_rec :
forall P:Z -> Set,
P 0 ->
- (forall x:Z, 0 <= x -> P x -> P (Zsucc x)) -> forall x:Z, 0 <= x -> P x.
+ (forall x:Z, 0 <= x -> P x -> P (Z.succ x)) ->
+ forall x:Z, 0 <= x -> P x.
Proof.
- intros P H H0 x H1; apply Z_of_nat_set;
- [ simple induction n;
- [ simpl in |- *; assumption
- | intros; rewrite (inj_S n0); exact (H0 (Z_of_nat n0) (Zle_0_nat n0) H2) ]
- | assumption ].
+ intros P Ho Hrec x Hx; apply Z_of_nat_set; trivial.
+ induction n. exact Ho.
+ rewrite Nat2Z.inj_succ. apply Hrec; trivial using Nat2Z.is_nonneg.
Qed.
Section Efficient_Rec.
@@ -129,58 +89,44 @@ Section Efficient_Rec.
Let R_wf : well_founded R.
Proof.
- set
- (f :=
- fun z =>
- match z with
- | Zpos p => nat_of_P p
- | Z0 => 0%nat
- | Zneg _ => 0%nat
- end) in *.
- apply well_founded_lt_compat with f.
- unfold R, f in |- *; clear f R.
- intros x y; case x; intros; elim H; clear H.
- case y; intros; apply lt_O_nat_of_P || inversion H0.
- case y; intros; apply nat_of_P_lt_Lt_compare_morphism || inversion H0; auto.
- intros; elim H; auto.
+ apply well_founded_lt_compat with Z.to_nat.
+ intros x y (Hx,H). apply Z2Nat.inj_lt; Z.order.
Qed.
Lemma natlike_rec2 :
forall P:Z -> Type,
P 0 ->
- (forall z:Z, 0 <= z -> P z -> P (Zsucc z)) -> forall z:Z, 0 <= z -> P z.
+ (forall z:Z, 0 <= z -> P z -> P (Z.succ z)) ->
+ forall z:Z, 0 <= z -> P z.
Proof.
- intros P Ho Hrec z; pattern z in |- *;
- apply (well_founded_induction_type R_wf).
- intro x; case x.
- trivial.
- intros.
- assert (0 <= Zpred (Zpos p)).
- apply Zorder.Zlt_0_le_0_pred; unfold Zlt in |- *; simpl in |- *; trivial.
- rewrite Zsucc_pred.
- apply Hrec.
- auto.
- apply X; auto; unfold R in |- *; intuition; apply Zlt_pred.
- intros; elim H; simpl in |- *; trivial.
+ intros P Ho Hrec.
+ induction z as [z IH] using (well_founded_induction_type R_wf).
+ destruct z; intros Hz.
+ - apply Ho.
+ - set (y:=Z.pred (Zpos p)).
+ assert (LE : 0 <= y) by (unfold y; now apply Z.lt_le_pred).
+ assert (EQ : Zpos p = Z.succ y) by (unfold y; now rewrite Z.succ_pred).
+ rewrite EQ. apply Hrec, IH; trivial.
+ split; trivial. unfold y; apply Z.lt_pred_l.
+ - now destruct Hz.
Qed.
- (** A variant of the previous using [Zpred] instead of [Zs]. *)
+ (** A variant of the previous using [Z.pred] instead of [Z.succ]. *)
Lemma natlike_rec3 :
forall P:Z -> Type,
P 0 ->
- (forall z:Z, 0 < z -> P (Zpred z) -> P z) -> forall z:Z, 0 <= z -> P z.
+ (forall z:Z, 0 < z -> P (Z.pred z) -> P z) ->
+ forall z:Z, 0 <= z -> P z.
Proof.
- intros P Ho Hrec z; pattern z in |- *;
- apply (well_founded_induction_type R_wf).
- intro x; case x.
- trivial.
- intros; apply Hrec.
- unfold Zlt in |- *; trivial.
- assert (0 <= Zpred (Zpos p)).
- apply Zorder.Zlt_0_le_0_pred; unfold Zlt in |- *; simpl in |- *; trivial.
- apply X; auto; unfold R in |- *; intuition; apply Zlt_pred.
- intros; elim H; simpl in |- *; trivial.
+ intros P Ho Hrec.
+ induction z as [z IH] using (well_founded_induction_type R_wf).
+ destruct z; intros Hz.
+ - apply Ho.
+ - assert (EQ : 0 <= Z.pred (Zpos p)) by now apply Z.lt_le_pred.
+ apply Hrec. easy. apply IH; trivial. split; trivial.
+ apply Z.lt_pred_l.
+ - now destruct Hz.
Qed.
(** A more general induction principle on non-negative numbers using [Zlt]. *)
@@ -190,15 +136,15 @@ Section Efficient_Rec.
(forall x:Z, (forall y:Z, 0 <= y < x -> P y) -> 0 <= x -> P x) ->
forall x:Z, 0 <= x -> P x.
Proof.
- intros P Hrec z; pattern z in |- *; apply (well_founded_induction_type R_wf).
- intro x; case x; intros.
- apply Hrec; intros.
- assert (H2 : 0 < 0).
- apply Zle_lt_trans with y; intuition.
- inversion H2.
- assumption.
- firstorder.
- unfold Zle, Zcompare in H; elim H; auto.
+ intros P Hrec.
+ induction x as [x IH] using (well_founded_induction_type R_wf).
+ destruct x; intros Hx.
+ - apply Hrec; trivial. intros y (Hy,Hy').
+ assert (0 < 0) by now apply Z.le_lt_trans with y.
+ discriminate.
+ - apply Hrec; trivial. intros y (Hy,Hy').
+ apply IH; trivial. now split.
+ - now destruct Hx.
Defined.
Lemma Zlt_0_ind :
@@ -234,22 +180,15 @@ Section Efficient_Rec.
(forall x:Z, (forall y:Z, z <= y < x -> P y) -> z <= x -> P x) ->
forall x:Z, z <= x -> P x.
Proof.
- intros P z Hrec x.
- assert (Hexpand : forall x, x = x - z + z).
- intro; unfold Zminus; rewrite <- Zplus_assoc; rewrite Zplus_opp_l;
- rewrite Zplus_0_r; trivial.
- intro Hz.
- rewrite (Hexpand x); pattern (x - z) in |- *; apply Zlt_0_rec.
- 2: apply Zplus_le_reg_r with z; rewrite <- Hexpand; assumption.
- intros x0 Hlt_x0 H.
- apply Hrec.
- 2: change z with (0+z); apply Zplus_le_compat_r; assumption.
- intro y; rewrite (Hexpand y); intros.
- destruct H0.
- apply Hlt_x0.
- split.
- apply Zplus_le_reg_r with z; assumption.
- apply Zplus_lt_reg_r with z; assumption.
+ intros P z Hrec x Hx.
+ rewrite <- (Z.sub_simpl_r x z). apply Z.le_0_sub in Hx.
+ pattern (x - z); apply Zlt_0_rec; trivial.
+ clear x Hx. intros x IH Hx.
+ apply Hrec. intros y (Hy,Hy').
+ rewrite <- (Z.sub_simpl_r y z). apply IH; split.
+ now rewrite Z.le_0_sub.
+ now apply Z.lt_sub_lt_add_r.
+ now rewrite <- (Z.add_le_mono_r 0 x z).
Qed.
Lemma Zlt_lower_bound_ind :
diff --git a/theories/ZArith/ZArith.v b/theories/ZArith/ZArith.v
index bc79e373..265e62f0 100644
--- a/theories/ZArith/ZArith.v
+++ b/theories/ZArith/ZArith.v
@@ -1,21 +1,22 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: ZArith.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(** Library for manipulating integers based on binary encoding *)
Require Export ZArith_base.
+(** Extra definitions *)
+
+Require Export Zpow_def.
+
(** Extra modules using [Omega] or [Ring]. *)
Require Export Zcomplements.
-Require Export Zsqrt.
Require Export Zpower.
Require Export Zdiv.
Require Export Zlogarithm.
diff --git a/theories/ZArith/ZArith_base.v b/theories/ZArith/ZArith_base.v
index 8cdae80d..8eeca3b9 100644
--- a/theories/ZArith/ZArith_base.v
+++ b/theories/ZArith/ZArith_base.v
@@ -1,17 +1,16 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: ZArith_base.v 14641 2011-11-06 11:59:10Z herbelin $ *)
-
(** Library for manipulating integers based on binary encoding.
These are the basic modules, required by [Omega] and [Ring] for instance.
The full library is [ZArith]. *)
+Require Export BinNums.
Require Export BinPos.
Require Export BinNat.
Require Export BinInt.
@@ -29,8 +28,8 @@ Require Export Zbool.
Require Export Zmisc.
Require Export Wf_Z.
-Hint Resolve Zle_refl Zplus_comm Zplus_assoc Zmult_comm Zmult_assoc Zplus_0_l
- Zplus_0_r Zmult_1_l Zplus_opp_l Zplus_opp_r Zmult_plus_distr_l
- Zmult_plus_distr_r: zarith.
+Hint Resolve Z.le_refl Z.add_comm Z.add_assoc Z.mul_comm Z.mul_assoc Z.add_0_l
+ Z.add_0_r Z.mul_1_l Z.add_opp_diag_l Z.add_opp_diag_r Z.mul_add_distr_l
+ Z.mul_add_distr_r: zarith.
Require Export Zhints.
diff --git a/theories/ZArith/ZArith_dec.v b/theories/ZArith/ZArith_dec.v
index b6766640..76308e60 100644
--- a/theories/ZArith/ZArith_dec.v
+++ b/theories/ZArith/ZArith_dec.v
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: ZArith_dec.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import Sumbool.
Require Import BinInt.
@@ -38,17 +36,12 @@ Proof.
intro; apply Zcompare_rect.
Defined.
+Notation Z_eq_dec := Z.eq_dec (only parsing).
+
Section decidability.
Variables x y : Z.
- (** * Decidability of equality on binary integers *)
-
- Definition Z_eq_dec : {x = y} + {x <> y}.
- Proof.
- decide equality; apply positive_eq_dec.
- Defined.
-
(** * Decidability of order on binary integers *)
Definition Z_lt_dec : {x < y} + {~ x < y}.
diff --git a/theories/ZArith/ZOdiv.v b/theories/ZArith/ZOdiv.v
deleted file mode 100644
index 70f6866e..00000000
--- a/theories/ZArith/ZOdiv.v
+++ /dev/null
@@ -1,947 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-
-Require Import BinPos BinNat Nnat ZArith_base ROmega ZArithRing.
-Require Export ZOdiv_def.
-Require Zdiv.
-
-Open Scope Z_scope.
-
-(** This file provides results about the Round-Toward-Zero Euclidean
- division [ZOdiv_eucl], whose projections are [ZOdiv] and [ZOmod].
- Definition of this division can be found in file [ZOdiv_def].
-
- This division and the one defined in Zdiv agree only on positive
- numbers. Otherwise, Zdiv performs Round-Toward-Bottom.
-
- The current approach is compatible with the division of usual
- programming languages such as Ocaml. In addition, it has nicer
- properties with respect to opposite and other usual operations.
-*)
-
-(** Since ZOdiv and Zdiv are not meant to be used concurrently,
- we reuse the same notation. *)
-
-Infix "/" := ZOdiv : Z_scope.
-Infix "mod" := ZOmod (at level 40, no associativity) : Z_scope.
-
-Infix "/" := Ndiv : N_scope.
-Infix "mod" := Nmod (at level 40, no associativity) : N_scope.
-
-(** Auxiliary results on the ad-hoc comparison [NPgeb]. *)
-
-Lemma NPgeb_Zge : forall (n:N)(p:positive),
- NPgeb n p = true -> Z_of_N n >= Zpos p.
-Proof.
- destruct n as [|n]; simpl; intros.
- discriminate.
- red; simpl; destruct Pcompare; now auto.
-Qed.
-
-Lemma NPgeb_Zlt : forall (n:N)(p:positive),
- NPgeb n p = false -> Z_of_N n < Zpos p.
-Proof.
- destruct n as [|n]; simpl; intros.
- red; auto.
- red; simpl; destruct Pcompare; now auto.
-Qed.
-
-(** * Relation between division on N and on Z. *)
-
-Lemma Ndiv_Z0div : forall a b:N,
- Z_of_N (a/b) = (Z_of_N a / Z_of_N b).
-Proof.
- intros.
- destruct a; destruct b; simpl; auto.
- unfold Ndiv, ZOdiv; simpl; destruct Pdiv_eucl; auto.
-Qed.
-
-Lemma Nmod_Z0mod : forall a b:N,
- Z_of_N (a mod b) = (Z_of_N a) mod (Z_of_N b).
-Proof.
- intros.
- destruct a; destruct b; simpl; auto.
- unfold Nmod, ZOmod; simpl; destruct Pdiv_eucl; auto.
-Qed.
-
-(** * Characterization of this euclidean division. *)
-
-(** First, the usual equation [a=q*b+r]. Notice that [a mod 0]
- has been chosen to be [a], so this equation holds even for [b=0].
-*)
-
-Theorem N_div_mod_eq : forall a b,
- a = (b * (Ndiv a b) + (Nmod a b))%N.
-Proof.
- intros; generalize (Ndiv_eucl_correct a b).
- unfold Ndiv, Nmod; destruct Ndiv_eucl; simpl.
- intro H; rewrite H; rewrite Nmult_comm; auto.
-Qed.
-
-Theorem ZO_div_mod_eq : forall a b,
- a = b * (ZOdiv a b) + (ZOmod a b).
-Proof.
- intros; generalize (ZOdiv_eucl_correct a b).
- unfold ZOdiv, ZOmod; destruct ZOdiv_eucl; simpl.
- intro H; rewrite H; rewrite Zmult_comm; auto.
-Qed.
-
-(** Then, the inequalities constraining the remainder. *)
-
-Theorem Pdiv_eucl_remainder : forall a b:positive,
- Z_of_N (snd (Pdiv_eucl a b)) < Zpos b.
-Proof.
- induction a; cbv beta iota delta [Pdiv_eucl]; fold Pdiv_eucl; cbv zeta.
- intros b; generalize (IHa b); case Pdiv_eucl.
- intros q1 r1 Hr1; simpl in Hr1.
- case_eq (NPgeb (2*r1+1) b); intros; unfold snd.
- romega with *.
- apply NPgeb_Zlt; auto.
- intros b; generalize (IHa b); case Pdiv_eucl.
- intros q1 r1 Hr1; simpl in Hr1.
- case_eq (NPgeb (2*r1) b); intros; unfold snd.
- romega with *.
- apply NPgeb_Zlt; auto.
- destruct b; simpl; romega with *.
-Qed.
-
-Theorem Nmod_lt : forall (a b:N), b<>0%N ->
- (a mod b < b)%N.
-Proof.
- destruct b as [ |b]; intro H; try solve [elim H;auto].
- destruct a as [ |a]; try solve [compute;auto]; unfold Nmod, Ndiv_eucl.
- generalize (Pdiv_eucl_remainder a b); destruct Pdiv_eucl; simpl.
- romega with *.
-Qed.
-
-(** The remainder is bounded by the divisor, in term of absolute values *)
-
-Theorem ZOmod_lt : forall a b:Z, b<>0 ->
- Zabs (a mod b) < Zabs b.
-Proof.
- destruct b as [ |b|b]; intro H; try solve [elim H;auto];
- destruct a as [ |a|a]; try solve [compute;auto]; unfold ZOmod, ZOdiv_eucl;
- generalize (Pdiv_eucl_remainder a b); destruct Pdiv_eucl; simpl;
- try rewrite Zabs_Zopp; rewrite Zabs_eq; auto; apply Z_of_N_le_0.
-Qed.
-
-(** The sign of the remainder is the one of [a]. Due to the possible
- nullity of [a], a general result is to be stated in the following form:
-*)
-
-Theorem ZOmod_sgn : forall a b:Z,
- 0 <= Zsgn (a mod b) * Zsgn a.
-Proof.
- destruct b as [ |b|b]; destruct a as [ |a|a]; simpl; auto with zarith;
- unfold ZOmod, ZOdiv_eucl; destruct Pdiv_eucl;
- simpl; destruct n0; simpl; auto with zarith.
-Qed.
-
-(** This can also be said in a simplier way: *)
-
-Theorem Zsgn_pos_iff : forall z, 0 <= Zsgn z <-> 0 <= z.
-Proof.
- destruct z; simpl; intuition auto with zarith.
-Qed.
-
-Theorem ZOmod_sgn2 : forall a b:Z,
- 0 <= (a mod b) * a.
-Proof.
- intros; rewrite <-Zsgn_pos_iff, Zsgn_Zmult; apply ZOmod_sgn.
-Qed.
-
-(** Reformulation of [ZOdiv_lt] and [ZOmod_sgn] in 2
- then 4 particular cases. *)
-
-Theorem ZOmod_lt_pos : forall a b:Z, 0<=a -> b<>0 ->
- 0 <= a mod b < Zabs b.
-Proof.
- intros.
- assert (0 <= a mod b).
- generalize (ZOmod_sgn a b).
- destruct (Zle_lt_or_eq 0 a H).
- rewrite <- Zsgn_pos in H1; rewrite H1; romega with *.
- subst a; simpl; auto.
- generalize (ZOmod_lt a b H0); romega with *.
-Qed.
-
-Theorem ZOmod_lt_neg : forall a b:Z, a<=0 -> b<>0 ->
- -Zabs b < a mod b <= 0.
-Proof.
- intros.
- assert (a mod b <= 0).
- generalize (ZOmod_sgn a b).
- destruct (Zle_lt_or_eq a 0 H).
- rewrite <- Zsgn_neg in H1; rewrite H1; romega with *.
- subst a; simpl; auto.
- generalize (ZOmod_lt a b H0); romega with *.
-Qed.
-
-Theorem ZOmod_lt_pos_pos : forall a b:Z, 0<=a -> 0<b -> 0 <= a mod b < b.
-Proof.
- intros; generalize (ZOmod_lt_pos a b); romega with *.
-Qed.
-
-Theorem ZOmod_lt_pos_neg : forall a b:Z, 0<=a -> b<0 -> 0 <= a mod b < -b.
-Proof.
- intros; generalize (ZOmod_lt_pos a b); romega with *.
-Qed.
-
-Theorem ZOmod_lt_neg_pos : forall a b:Z, a<=0 -> 0<b -> -b < a mod b <= 0.
-Proof.
- intros; generalize (ZOmod_lt_neg a b); romega with *.
-Qed.
-
-Theorem ZOmod_lt_neg_neg : forall a b:Z, a<=0 -> b<0 -> b < a mod b <= 0.
-Proof.
- intros; generalize (ZOmod_lt_neg a b); romega with *.
-Qed.
-
-(** * Division and Opposite *)
-
-(* The precise equalities that are invalid with "historic" Zdiv. *)
-
-Theorem ZOdiv_opp_l : forall a b:Z, (-a)/b = -(a/b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold ZOdiv, ZOdiv_eucl; destruct Pdiv_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem ZOdiv_opp_r : forall a b:Z, a/(-b) = -(a/b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold ZOdiv, ZOdiv_eucl; destruct Pdiv_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem ZOmod_opp_l : forall a b:Z, (-a) mod b = -(a mod b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold ZOmod, ZOdiv_eucl; destruct Pdiv_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem ZOmod_opp_r : forall a b:Z, a mod (-b) = a mod b.
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold ZOmod, ZOdiv_eucl; destruct Pdiv_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem ZOdiv_opp_opp : forall a b:Z, (-a)/(-b) = a/b.
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold ZOdiv, ZOdiv_eucl; destruct Pdiv_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem ZOmod_opp_opp : forall a b:Z, (-a) mod (-b) = -(a mod b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold ZOmod, ZOdiv_eucl; destruct Pdiv_eucl; simpl; auto with zarith.
-Qed.
-
-(** * Unicity results *)
-
-Definition Remainder a b r :=
- (0 <= a /\ 0 <= r < Zabs b) \/ (a <= 0 /\ -Zabs b < r <= 0).
-
-Definition Remainder_alt a b r :=
- Zabs r < Zabs b /\ 0 <= r * a.
-
-Lemma Remainder_equiv : forall a b r,
- Remainder a b r <-> Remainder_alt a b r.
-Proof.
- unfold Remainder, Remainder_alt; intuition.
- romega with *.
- romega with *.
- rewrite <-(Zmult_opp_opp).
- apply Zmult_le_0_compat; romega.
- assert (0 <= Zsgn r * Zsgn a) by (rewrite <-Zsgn_Zmult, Zsgn_pos_iff; auto).
- destruct r; simpl Zsgn in *; romega with *.
-Qed.
-
-Theorem ZOdiv_mod_unique_full:
- forall a b q r, Remainder a b r ->
- a = b*q + r -> q = a/b /\ r = a mod b.
-Proof.
- destruct 1 as [(H,H0)|(H,H0)]; intros.
- apply Zdiv.Zdiv_mod_unique with b; auto.
- apply ZOmod_lt_pos; auto.
- romega with *.
- rewrite <- H1; apply ZO_div_mod_eq.
-
- rewrite <- (Zopp_involutive a).
- rewrite ZOdiv_opp_l, ZOmod_opp_l.
- generalize (Zdiv.Zdiv_mod_unique b (-q) (-a/b) (-r) (-a mod b)).
- generalize (ZOmod_lt_pos (-a) b).
- rewrite <-ZO_div_mod_eq, <-Zopp_mult_distr_r, <-Zopp_plus_distr, <-H1.
- romega with *.
-Qed.
-
-Theorem ZOdiv_unique_full:
- forall a b q r, Remainder a b r ->
- a = b*q + r -> q = a/b.
-Proof.
- intros; destruct (ZOdiv_mod_unique_full a b q r); auto.
-Qed.
-
-Theorem ZOdiv_unique:
- forall a b q r, 0 <= a -> 0 <= r < b ->
- a = b*q + r -> q = a/b.
-Proof.
- intros; eapply ZOdiv_unique_full; eauto.
- red; romega with *.
-Qed.
-
-Theorem ZOmod_unique_full:
- forall a b q r, Remainder a b r ->
- a = b*q + r -> r = a mod b.
-Proof.
- intros; destruct (ZOdiv_mod_unique_full a b q r); auto.
-Qed.
-
-Theorem ZOmod_unique:
- forall a b q r, 0 <= a -> 0 <= r < b ->
- a = b*q + r -> r = a mod b.
-Proof.
- intros; eapply ZOmod_unique_full; eauto.
- red; romega with *.
-Qed.
-
-(** * Basic values of divisions and modulo. *)
-
-Lemma ZOmod_0_l: forall a, 0 mod a = 0.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma ZOmod_0_r: forall a, a mod 0 = a.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma ZOdiv_0_l: forall a, 0/a = 0.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma ZOdiv_0_r: forall a, a/0 = 0.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma ZOmod_1_r: forall a, a mod 1 = 0.
-Proof.
- intros; symmetry; apply ZOmod_unique_full with a; auto with zarith.
- rewrite Remainder_equiv; red; simpl; auto with zarith.
-Qed.
-
-Lemma ZOdiv_1_r: forall a, a/1 = a.
-Proof.
- intros; symmetry; apply ZOdiv_unique_full with 0; auto with zarith.
- rewrite Remainder_equiv; red; simpl; auto with zarith.
-Qed.
-
-Hint Resolve ZOmod_0_l ZOmod_0_r ZOdiv_0_l ZOdiv_0_r ZOdiv_1_r ZOmod_1_r
- : zarith.
-
-Lemma ZOdiv_1_l: forall a, 1 < a -> 1/a = 0.
-Proof.
- intros; symmetry; apply ZOdiv_unique with 1; auto with zarith.
-Qed.
-
-Lemma ZOmod_1_l: forall a, 1 < a -> 1 mod a = 1.
-Proof.
- intros; symmetry; apply ZOmod_unique with 0; auto with zarith.
-Qed.
-
-Lemma ZO_div_same : forall a:Z, a<>0 -> a/a = 1.
-Proof.
- intros; symmetry; apply ZOdiv_unique_full with 0; auto with *.
- rewrite Remainder_equiv; red; simpl; romega with *.
-Qed.
-
-Lemma ZO_mod_same : forall a, a mod a = 0.
-Proof.
- destruct a; intros; symmetry.
- compute; auto.
- apply ZOmod_unique with 1; auto with *; romega with *.
- apply ZOmod_unique_full with 1; auto with *; red; romega with *.
-Qed.
-
-Lemma ZO_mod_mult : forall a b, (a*b) mod b = 0.
-Proof.
- intros a b; destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst; simpl; rewrite ZOmod_0_r; auto with zarith.
- symmetry; apply ZOmod_unique_full with a; [ red; romega with * | ring ].
-Qed.
-
-Lemma ZO_div_mult : forall a b:Z, b <> 0 -> (a*b)/b = a.
-Proof.
- intros; symmetry; apply ZOdiv_unique_full with 0; auto with zarith;
- [ red; romega with * | ring].
-Qed.
-
-(** * Order results about ZOmod and ZOdiv *)
-
-(* Division of positive numbers is positive. *)
-
-Lemma ZO_div_pos: forall a b, 0 <= a -> 0 <= b -> 0 <= a/b.
-Proof.
- intros.
- destruct (Zle_lt_or_eq 0 b H0).
- assert (H2:=ZOmod_lt_pos_pos a b H H1).
- rewrite (ZO_div_mod_eq a b) in H.
- destruct (Z_lt_le_dec (a/b) 0); auto.
- assert (b*(a/b) <= -b).
- replace (-b) with (b*-1); [ | ring].
- apply Zmult_le_compat_l; auto with zarith.
- romega.
- subst b; rewrite ZOdiv_0_r; auto.
-Qed.
-
-(** As soon as the divisor is greater or equal than 2,
- the division is strictly decreasing. *)
-
-Lemma ZO_div_lt : forall a b:Z, 0 < a -> 2 <= b -> a/b < a.
-Proof.
- intros.
- assert (Hb : 0 < b) by romega.
- assert (H1 : 0 <= a/b) by (apply ZO_div_pos; auto with zarith).
- assert (H2 : 0 <= a mod b < b) by (apply ZOmod_lt_pos_pos; auto with zarith).
- destruct (Zle_lt_or_eq 0 (a/b) H1) as [H3|H3]; [ | rewrite <- H3; auto].
- pattern a at 2; rewrite (ZO_div_mod_eq a b).
- apply Zlt_le_trans with (2*(a/b)).
- romega.
- apply Zle_trans with (b*(a/b)).
- apply Zmult_le_compat_r; auto.
- romega.
-Qed.
-
-(** A division of a small number by a bigger one yields zero. *)
-
-Theorem ZOdiv_small: forall a b, 0 <= a < b -> a/b = 0.
-Proof.
- intros a b H; apply sym_equal; apply ZOdiv_unique with a; auto with zarith.
-Qed.
-
-(** Same situation, in term of modulo: *)
-
-Theorem ZOmod_small: forall a n, 0 <= a < n -> a mod n = a.
-Proof.
- intros a b H; apply sym_equal; apply ZOmod_unique with 0; auto with zarith.
-Qed.
-
-(** [Zge] is compatible with a positive division. *)
-
-Lemma ZO_div_monotone_pos : forall a b c:Z, 0<=c -> 0<=a<=b -> a/c <= b/c.
-Proof.
- intros.
- destruct H0.
- destruct (Zle_lt_or_eq 0 c H);
- [ clear H | subst c; do 2 rewrite ZOdiv_0_r; auto].
- generalize (ZO_div_mod_eq a c).
- generalize (ZOmod_lt_pos_pos a c H0 H2).
- generalize (ZO_div_mod_eq b c).
- generalize (ZOmod_lt_pos_pos b c (Zle_trans _ _ _ H0 H1) H2).
- intros.
- elim (Z_le_gt_dec (a / c) (b / c)); auto with zarith.
- intro.
- absurd (a - b >= 1).
- omega.
- replace (a-b) with (c * (a/c-b/c) + a mod c - b mod c) by
- (symmetry; pattern a at 1; rewrite H5; pattern b at 1; rewrite H3; ring).
- assert (c * (a / c - b / c) >= c * 1).
- apply Zmult_ge_compat_l.
- omega.
- omega.
- assert (c * 1 = c).
- ring.
- omega.
-Qed.
-
-Lemma ZO_div_monotone : forall a b c, 0<=c -> a<=b -> a/c <= b/c.
-Proof.
- intros.
- destruct (Z_le_gt_dec 0 a).
- apply ZO_div_monotone_pos; auto with zarith.
- destruct (Z_le_gt_dec 0 b).
- apply Zle_trans with 0.
- apply Zle_left_rev.
- simpl.
- rewrite <- ZOdiv_opp_l.
- apply ZO_div_pos; auto with zarith.
- apply ZO_div_pos; auto with zarith.
- rewrite <-(Zopp_involutive a), (ZOdiv_opp_l (-a)).
- rewrite <-(Zopp_involutive b), (ZOdiv_opp_l (-b)).
- generalize (ZO_div_monotone_pos (-b) (-a) c H).
- romega.
-Qed.
-
-(** With our choice of division, rounding of (a/b) is always done toward zero: *)
-
-Lemma ZO_mult_div_le : forall a b:Z, 0 <= a -> 0 <= b*(a/b) <= a.
-Proof.
- intros a b Ha.
- destruct b as [ |b|b].
- simpl; auto with zarith.
- split.
- apply Zmult_le_0_compat; auto with zarith.
- apply ZO_div_pos; auto with zarith.
- generalize (ZO_div_mod_eq a (Zpos b)) (ZOmod_lt_pos_pos a (Zpos b) Ha); romega with *.
- change (Zneg b) with (-Zpos b); rewrite ZOdiv_opp_r, Zmult_opp_opp.
- split.
- apply Zmult_le_0_compat; auto with zarith.
- apply ZO_div_pos; auto with zarith.
- generalize (ZO_div_mod_eq a (Zpos b)) (ZOmod_lt_pos_pos a (Zpos b) Ha); romega with *.
-Qed.
-
-Lemma ZO_mult_div_ge : forall a b:Z, a <= 0 -> a <= b*(a/b) <= 0.
-Proof.
- intros a b Ha.
- destruct b as [ |b|b].
- simpl; auto with zarith.
- split.
- generalize (ZO_div_mod_eq a (Zpos b)) (ZOmod_lt_neg_pos a (Zpos b) Ha); romega with *.
- apply Zle_left_rev; unfold Zplus.
- rewrite Zopp_mult_distr_r, <-ZOdiv_opp_l.
- apply Zmult_le_0_compat; auto with zarith.
- apply ZO_div_pos; auto with zarith.
- change (Zneg b) with (-Zpos b); rewrite ZOdiv_opp_r, Zmult_opp_opp.
- split.
- generalize (ZO_div_mod_eq a (Zpos b)) (ZOmod_lt_neg_pos a (Zpos b) Ha); romega with *.
- apply Zle_left_rev; unfold Zplus.
- rewrite Zopp_mult_distr_r, <-ZOdiv_opp_l.
- apply Zmult_le_0_compat; auto with zarith.
- apply ZO_div_pos; auto with zarith.
-Qed.
-
-(** The previous inequalities between [b*(a/b)] and [a] are exact
- iff the modulo is zero. *)
-
-Lemma ZO_div_exact_full_1 : forall a b:Z, a = b*(a/b) -> a mod b = 0.
-Proof.
- intros; generalize (ZO_div_mod_eq a b); romega.
-Qed.
-
-Lemma ZO_div_exact_full_2 : forall a b:Z, a mod b = 0 -> a = b*(a/b).
-Proof.
- intros; generalize (ZO_div_mod_eq a b); romega.
-Qed.
-
-(** A modulo cannot grow beyond its starting point. *)
-
-Theorem ZOmod_le: forall a b, 0 <= a -> 0 <= b -> a mod b <= a.
-Proof.
- intros a b H1 H2.
- destruct (Zle_lt_or_eq _ _ H2).
- case (Zle_or_lt b a); intros H3.
- case (ZOmod_lt_pos_pos a b); auto with zarith.
- rewrite ZOmod_small; auto with zarith.
- subst; rewrite ZOmod_0_r; auto with zarith.
-Qed.
-
-(** Some additionnal inequalities about Zdiv. *)
-
-Theorem ZOdiv_le_upper_bound:
- forall a b q, 0 < b -> a <= q*b -> a/b <= q.
-Proof.
- intros.
- rewrite <- (ZO_div_mult q b); auto with zarith.
- apply ZO_div_monotone; auto with zarith.
-Qed.
-
-Theorem ZOdiv_lt_upper_bound:
- forall a b q, 0 <= a -> 0 < b -> a < q*b -> a/b < q.
-Proof.
- intros a b q H1 H2 H3.
- apply Zmult_lt_reg_r with b; auto with zarith.
- apply Zle_lt_trans with (2 := H3).
- pattern a at 2; rewrite (ZO_div_mod_eq a b); auto with zarith.
- rewrite (Zmult_comm b); case (ZOmod_lt_pos_pos a b); auto with zarith.
-Qed.
-
-Theorem ZOdiv_le_lower_bound:
- forall a b q, 0 < b -> q*b <= a -> q <= a/b.
-Proof.
- intros.
- rewrite <- (ZO_div_mult q b); auto with zarith.
- apply ZO_div_monotone; auto with zarith.
-Qed.
-
-Theorem ZOdiv_sgn: forall a b,
- 0 <= Zsgn (a/b) * Zsgn a * Zsgn b.
-Proof.
- destruct a as [ |a|a]; destruct b as [ |b|b]; simpl; auto with zarith;
- unfold ZOdiv; simpl; destruct Pdiv_eucl; simpl; destruct n; simpl; auto with zarith.
-Qed.
-
-(** * Relations between usual operations and Zmod and Zdiv *)
-
-(** First, a result that used to be always valid with Zdiv,
- but must be restricted here.
- For instance, now (9+(-5)*2) mod 2 = -1 <> 1 = 9 mod 2 *)
-
-Lemma ZO_mod_plus : forall a b c:Z,
- 0 <= (a+b*c) * a ->
- (a + b * c) mod c = a mod c.
-Proof.
- intros; destruct (Z_eq_dec a 0) as [Ha|Ha].
- subst; simpl; rewrite ZOmod_0_l; apply ZO_mod_mult.
- intros; destruct (Z_eq_dec c 0) as [Hc|Hc].
- subst; do 2 rewrite ZOmod_0_r; romega.
- symmetry; apply ZOmod_unique_full with (a/c+b); auto with zarith.
- rewrite Remainder_equiv; split.
- apply ZOmod_lt; auto.
- apply Zmult_le_0_reg_r with (a*a); eauto.
- destruct a; simpl; auto with zarith.
- replace ((a mod c)*(a+b*c)*(a*a)) with (((a mod c)*a)*((a+b*c)*a)) by ring.
- apply Zmult_le_0_compat; auto.
- apply ZOmod_sgn2.
- rewrite Zmult_plus_distr_r, Zmult_comm.
- generalize (ZO_div_mod_eq a c); romega.
-Qed.
-
-Lemma ZO_div_plus : forall a b c:Z,
- 0 <= (a+b*c) * a -> c<>0 ->
- (a + b * c) / c = a / c + b.
-Proof.
- intros; destruct (Z_eq_dec a 0) as [Ha|Ha].
- subst; simpl; apply ZO_div_mult; auto.
- symmetry.
- apply ZOdiv_unique_full with (a mod c); auto with zarith.
- rewrite Remainder_equiv; split.
- apply ZOmod_lt; auto.
- apply Zmult_le_0_reg_r with (a*a); eauto.
- destruct a; simpl; auto with zarith.
- replace ((a mod c)*(a+b*c)*(a*a)) with (((a mod c)*a)*((a+b*c)*a)) by ring.
- apply Zmult_le_0_compat; auto.
- apply ZOmod_sgn2.
- rewrite Zmult_plus_distr_r, Zmult_comm.
- generalize (ZO_div_mod_eq a c); romega.
-Qed.
-
-Theorem ZO_div_plus_l: forall a b c : Z,
- 0 <= (a*b+c)*c -> b<>0 ->
- b<>0 -> (a * b + c) / b = a + c / b.
-Proof.
- intros a b c; rewrite Zplus_comm; intros; rewrite ZO_div_plus;
- try apply Zplus_comm; auto with zarith.
-Qed.
-
-(** Cancellations. *)
-
-Lemma ZOdiv_mult_cancel_r : forall a b c:Z,
- c<>0 -> (a*c)/(b*c) = a/b.
-Proof.
- intros a b c Hc.
- destruct (Z_eq_dec b 0).
- subst; simpl; do 2 rewrite ZOdiv_0_r; auto.
- symmetry.
- apply ZOdiv_unique_full with ((a mod b)*c); auto with zarith.
- rewrite Remainder_equiv.
- split.
- do 2 rewrite Zabs_Zmult.
- apply Zmult_lt_compat_r.
- romega with *.
- apply ZOmod_lt; auto.
- replace ((a mod b)*c*(a*c)) with (((a mod b)*a)*(c*c)) by ring.
- apply Zmult_le_0_compat.
- apply ZOmod_sgn2.
- destruct c; simpl; auto with zarith.
- pattern a at 1; rewrite (ZO_div_mod_eq a b); ring.
-Qed.
-
-Lemma ZOdiv_mult_cancel_l : forall a b c:Z,
- c<>0 -> (c*a)/(c*b) = a/b.
-Proof.
- intros.
- rewrite (Zmult_comm c a); rewrite (Zmult_comm c b).
- apply ZOdiv_mult_cancel_r; auto.
-Qed.
-
-Lemma ZOmult_mod_distr_l: forall a b c,
- (c*a) mod (c*b) = c * (a mod b).
-Proof.
- intros; destruct (Z_eq_dec c 0) as [Hc|Hc].
- subst; simpl; rewrite ZOmod_0_r; auto.
- destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst; repeat rewrite Zmult_0_r || rewrite ZOmod_0_r; auto.
- assert (c*b <> 0).
- contradict Hc; eapply Zmult_integral_l; eauto.
- rewrite (Zplus_minus_eq _ _ _ (ZO_div_mod_eq (c*a) (c*b))).
- rewrite (Zplus_minus_eq _ _ _ (ZO_div_mod_eq a b)).
- rewrite ZOdiv_mult_cancel_l; auto with zarith.
- ring.
-Qed.
-
-Lemma ZOmult_mod_distr_r: forall a b c,
- (a*c) mod (b*c) = (a mod b) * c.
-Proof.
- intros; repeat rewrite (fun x => (Zmult_comm x c)).
- apply ZOmult_mod_distr_l; auto.
-Qed.
-
-(** Operations modulo. *)
-
-Theorem ZOmod_mod: forall a n, (a mod n) mod n = a mod n.
-Proof.
- intros.
- generalize (ZOmod_sgn2 a n).
- pattern a at 2 4; rewrite (ZO_div_mod_eq a n); auto with zarith.
- rewrite Zplus_comm; rewrite (Zmult_comm n).
- intros.
- apply sym_equal; apply ZO_mod_plus; auto with zarith.
- rewrite Zmult_comm; auto.
-Qed.
-
-Theorem ZOmult_mod: forall a b n,
- (a * b) mod n = ((a mod n) * (b mod n)) mod n.
-Proof.
- intros.
- generalize (Zmult_le_0_compat _ _ (ZOmod_sgn2 a n) (ZOmod_sgn2 b n)).
- pattern a at 2 3; rewrite (ZO_div_mod_eq a n); auto with zarith.
- pattern b at 2 3; rewrite (ZO_div_mod_eq b n); auto with zarith.
- set (A:=a mod n); set (B:=b mod n); set (A':=a/n); set (B':=b/n).
- replace (A*(n*A'+A)*(B*(n*B'+B))) with (((n*A' + A) * (n*B' + B))*(A*B))
- by ring.
- replace ((n*A' + A) * (n*B' + B))
- with (A*B + (A'*B+B'*A+n*A'*B')*n) by ring.
- intros.
- apply ZO_mod_plus; auto with zarith.
-Qed.
-
-(** addition and modulo
-
- Generally speaking, unlike with Zdiv, 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. *)
-
-Theorem ZOplus_mod: forall a b n,
- 0 <= a * b ->
- (a + b) mod n = (a mod n + b mod n) mod n.
-Proof.
- assert (forall a b n, 0<a -> 0<b ->
- (a + b) mod n = (a mod n + b mod n) mod n).
- intros a b n Ha Hb.
- assert (H : 0<=a+b) by (romega with * ); revert H.
- pattern a at 1 2; rewrite (ZO_div_mod_eq a n); auto with zarith.
- pattern b at 1 2; rewrite (ZO_div_mod_eq b n); auto with zarith.
- replace ((n * (a / n) + a mod n) + (n * (b / n) + b mod n))
- with ((a mod n + b mod n) + (a / n + b / n) * n) by ring.
- intros.
- apply ZO_mod_plus; auto with zarith.
- apply Zmult_le_0_compat; auto with zarith.
- apply Zplus_le_0_compat.
- apply Zmult_le_reg_r with a; auto with zarith.
- simpl; apply ZOmod_sgn2; auto.
- apply Zmult_le_reg_r with b; auto with zarith.
- simpl; apply ZOmod_sgn2; auto.
- (* general situation *)
- intros a b n Hab.
- destruct (Z_eq_dec a 0).
- subst; simpl; symmetry; apply ZOmod_mod.
- destruct (Z_eq_dec b 0).
- subst; simpl; do 2 rewrite Zplus_0_r; symmetry; apply ZOmod_mod.
- assert (0<a /\ 0<b \/ a<0 /\ b<0).
- destruct a; destruct b; simpl in *; intuition; romega with *.
- destruct H0.
- apply H; intuition.
- rewrite <-(Zopp_involutive a), <-(Zopp_involutive b).
- rewrite <- Zopp_plus_distr; rewrite ZOmod_opp_l.
- rewrite (ZOmod_opp_l (-a)),(ZOmod_opp_l (-b)).
- match goal with |- _ = (-?x+-?y) mod n =>
- rewrite <-(Zopp_plus_distr x y), ZOmod_opp_l end.
- f_equal; apply H; auto with zarith.
-Qed.
-
-Lemma ZOplus_mod_idemp_l: forall a b n,
- 0 <= a * b ->
- (a mod n + b) mod n = (a + b) mod n.
-Proof.
- intros.
- rewrite ZOplus_mod.
- rewrite ZOmod_mod.
- symmetry.
- apply ZOplus_mod; auto.
- destruct (Z_eq_dec a 0).
- subst; rewrite ZOmod_0_l; auto.
- destruct (Z_eq_dec b 0).
- subst; rewrite Zmult_0_r; auto with zarith.
- apply Zmult_le_reg_r with (a*b).
- assert (a*b <> 0).
- intro Hab.
- rewrite (Zmult_integral_l _ _ n1 Hab) in n0; auto with zarith.
- auto with zarith.
- simpl.
- replace (a mod n * b * (a*b)) with ((a mod n * a)*(b*b)) by ring.
- apply Zmult_le_0_compat.
- apply ZOmod_sgn2.
- destruct b; simpl; auto with zarith.
-Qed.
-
-Lemma ZOplus_mod_idemp_r: forall a b n,
- 0 <= a*b ->
- (b + a mod n) mod n = (b + a) mod n.
-Proof.
- intros.
- rewrite Zplus_comm, (Zplus_comm b a).
- apply ZOplus_mod_idemp_l; auto.
-Qed.
-
-Lemma ZOmult_mod_idemp_l: forall a b n, (a mod n * b) mod n = (a * b) mod n.
-Proof.
- intros; rewrite ZOmult_mod, ZOmod_mod, <- ZOmult_mod; auto.
-Qed.
-
-Lemma ZOmult_mod_idemp_r: forall a b n, (b * (a mod n)) mod n = (b * a) mod n.
-Proof.
- intros; rewrite ZOmult_mod, ZOmod_mod, <- ZOmult_mod; auto.
-Qed.
-
-(** Unlike with Zdiv, the following result is true without restrictions. *)
-
-Lemma ZOdiv_ZOdiv : forall a b c, (a/b)/c = a/(b*c).
-Proof.
- (* particular case: a, b, c positive *)
- assert (forall a b c, a>0 -> b>0 -> c>0 -> (a/b)/c = a/(b*c)).
- intros a b c H H0 H1.
- pattern a at 2;rewrite (ZO_div_mod_eq a b).
- pattern (a/b) at 2;rewrite (ZO_div_mod_eq (a/b) c).
- replace (b * (c * (a / b / c) + (a / b) mod c) + a mod b) with
- ((a / b / c)*(b * c) + (b * ((a / b) mod c) + a mod b)) by ring.
- assert (b*c<>0).
- intro H2;
- assert (H3: c <> 0) by auto with zarith;
- rewrite (Zmult_integral_l _ _ H3 H2) in H0; auto with zarith.
- assert (0<=a/b) by (apply (ZO_div_pos a b); auto with zarith).
- assert (0<=a mod b < b) by (apply ZOmod_lt_pos_pos; auto with zarith).
- assert (0<=(a/b) mod c < c) by
- (apply ZOmod_lt_pos_pos; auto with zarith).
- rewrite ZO_div_plus_l; auto with zarith.
- rewrite (ZOdiv_small (b * ((a / b) mod c) + a mod b)).
- ring.
- split.
- apply Zplus_le_0_compat;auto with zarith.
- apply Zle_lt_trans with (b * ((a / b) mod c) + (b-1)).
- apply Zplus_le_compat;auto with zarith.
- apply Zle_lt_trans with (b * (c-1) + (b - 1)).
- apply Zplus_le_compat;auto with zarith.
- replace (b * (c - 1) + (b - 1)) with (b*c-1);try ring;auto with zarith.
- repeat (apply Zmult_le_0_compat || apply Zplus_le_0_compat); auto with zarith.
- apply (ZO_div_pos (a/b) c); auto with zarith.
- (* b c positive, a general *)
- assert (forall a b c, b>0 -> c>0 -> (a/b)/c = a/(b*c)).
- intros; destruct a as [ |a|a]; try reflexivity.
- apply H; auto with zarith.
- change (Zneg a) with (-Zpos a); repeat rewrite ZOdiv_opp_l.
- f_equal; apply H; auto with zarith.
- (* c positive, a b general *)
- assert (forall a b c, c>0 -> (a/b)/c = a/(b*c)).
- intros; destruct b as [ |b|b].
- repeat rewrite ZOdiv_0_r; reflexivity.
- apply H0; auto with zarith.
- change (Zneg b) with (-Zpos b);
- repeat (rewrite ZOdiv_opp_r || rewrite ZOdiv_opp_l || rewrite <- Zopp_mult_distr_l).
- f_equal; apply H0; auto with zarith.
- (* a b c general *)
- intros; destruct c as [ |c|c].
- rewrite Zmult_0_r; repeat rewrite ZOdiv_0_r; reflexivity.
- apply H1; auto with zarith.
- change (Zneg c) with (-Zpos c);
- rewrite <- Zopp_mult_distr_r; do 2 rewrite ZOdiv_opp_r.
- f_equal; apply H1; auto with zarith.
-Qed.
-
-(** A last inequality: *)
-
-Theorem ZOdiv_mult_le:
- forall a b c, 0<=a -> 0<=b -> 0<=c -> c*(a/b) <= (c*a)/b.
-Proof.
- intros a b c Ha Hb Hc.
- destruct (Zle_lt_or_eq _ _ Ha);
- [ | subst; rewrite ZOdiv_0_l, Zmult_0_r, ZOdiv_0_l; auto].
- destruct (Zle_lt_or_eq _ _ Hb);
- [ | subst; rewrite ZOdiv_0_r, ZOdiv_0_r, Zmult_0_r; auto].
- destruct (Zle_lt_or_eq _ _ Hc);
- [ | subst; rewrite ZOdiv_0_l; auto].
- case (ZOmod_lt_pos_pos a b); auto with zarith; intros Hu1 Hu2.
- case (ZOmod_lt_pos_pos c b); auto with zarith; intros Hv1 Hv2.
- apply Zmult_le_reg_r with b; auto with zarith.
- rewrite <- Zmult_assoc.
- replace (a / b * b) with (a - a mod b).
- replace (c * a / b * b) with (c * a - (c * a) mod b).
- rewrite Zmult_minus_distr_l.
- unfold Zminus; apply Zplus_le_compat_l.
- match goal with |- - ?X <= -?Y => assert (Y <= X); auto with zarith end.
- apply Zle_trans with ((c mod b) * (a mod b)); auto with zarith.
- rewrite ZOmult_mod; auto with zarith.
- apply (ZOmod_le ((c mod b) * (a mod b)) b); auto with zarith.
- apply Zmult_le_compat_r; auto with zarith.
- apply (ZOmod_le c b); auto.
- pattern (c * a) at 1; rewrite (ZO_div_mod_eq (c * a) b); try ring;
- auto with zarith.
- pattern a at 1; rewrite (ZO_div_mod_eq a b); try ring; auto with zarith.
-Qed.
-
-(** ZOmod is related to divisibility (see more in Znumtheory) *)
-
-Lemma ZOmod_divides : forall a b,
- a mod b = 0 <-> exists c, a = b*c.
-Proof.
- split; intros.
- exists (a/b).
- pattern a at 1; rewrite (ZO_div_mod_eq a b).
- rewrite H; auto with zarith.
- destruct H as [c Hc].
- destruct (Z_eq_dec b 0).
- subst b; simpl in *; subst a; auto.
- symmetry.
- apply ZOmod_unique_full with c; auto with zarith.
- red; romega with *.
-Qed.
-
-(** * Interaction with "historic" Zdiv *)
-
-(** They agree at least on positive numbers: *)
-
-Theorem ZOdiv_eucl_Zdiv_eucl_pos : forall a b:Z, 0 <= a -> 0 < b ->
- a/b = Zdiv.Zdiv a b /\ a mod b = Zdiv.Zmod a b.
-Proof.
- intros.
- apply Zdiv.Zdiv_mod_unique with b.
- apply ZOmod_lt_pos; auto with zarith.
- rewrite Zabs_eq; auto with *; apply Zdiv.Z_mod_lt; auto with *.
- rewrite <- Zdiv.Z_div_mod_eq; auto with *.
- symmetry; apply ZO_div_mod_eq; auto with *.
-Qed.
-
-Theorem ZOdiv_Zdiv_pos : forall a b, 0 <= a -> 0 <= b ->
- a/b = Zdiv.Zdiv a b.
-Proof.
- intros a b Ha Hb.
- destruct (Zle_lt_or_eq _ _ Hb).
- generalize (ZOdiv_eucl_Zdiv_eucl_pos a b Ha H); intuition.
- subst; rewrite ZOdiv_0_r, Zdiv.Zdiv_0_r; reflexivity.
-Qed.
-
-Theorem ZOmod_Zmod_pos : forall a b, 0 <= a -> 0 < b ->
- a mod b = Zdiv.Zmod a b.
-Proof.
- intros a b Ha Hb; generalize (ZOdiv_eucl_Zdiv_eucl_pos a b Ha Hb);
- intuition.
-Qed.
-
-(** Modulos are null at the same places *)
-
-Theorem ZOmod_Zmod_zero : forall a b, b<>0 ->
- (a mod b = 0 <-> Zdiv.Zmod a b = 0).
-Proof.
- intros.
- rewrite ZOmod_divides, Zdiv.Zmod_divides; intuition.
-Qed.
diff --git a/theories/ZArith/ZOdiv_def.v b/theories/ZArith/ZOdiv_def.v
deleted file mode 100644
index 71d6cad4..00000000
--- a/theories/ZArith/ZOdiv_def.v
+++ /dev/null
@@ -1,136 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-
-Require Import BinPos BinNat Nnat ZArith_base.
-
-Open Scope Z_scope.
-
-Definition NPgeb (a:N)(b:positive) :=
- match a with
- | N0 => false
- | Npos na => match Pcompare na b Eq with Lt => false | _ => true end
- end.
-
-Fixpoint Pdiv_eucl (a b:positive) : N * N :=
- match a with
- | xH =>
- match b with xH => (1, 0)%N | _ => (0, 1)%N end
- | xO a' =>
- let (q, r) := Pdiv_eucl a' b in
- let r' := (2 * r)%N in
- if (NPgeb r' b) then (2 * q + 1, (Nminus r' (Npos b)))%N
- else (2 * q, r')%N
- | xI a' =>
- let (q, r) := Pdiv_eucl a' b in
- let r' := (2 * r + 1)%N in
- if (NPgeb r' b) then (2 * q + 1, (Nminus r' (Npos b)))%N
- else (2 * q, r')%N
- end.
-
-Definition ZOdiv_eucl (a b:Z) : Z * Z :=
- match a, b with
- | Z0, _ => (Z0, Z0)
- | _, Z0 => (Z0, a)
- | Zpos na, Zpos nb =>
- let (nq, nr) := Pdiv_eucl na nb in
- (Z_of_N nq, Z_of_N nr)
- | Zneg na, Zpos nb =>
- let (nq, nr) := Pdiv_eucl na nb in
- (Zopp (Z_of_N nq), Zopp (Z_of_N nr))
- | Zpos na, Zneg nb =>
- let (nq, nr) := Pdiv_eucl na nb in
- (Zopp (Z_of_N nq), Z_of_N nr)
- | Zneg na, Zneg nb =>
- let (nq, nr) := Pdiv_eucl na nb in
- (Z_of_N nq, Zopp (Z_of_N nr))
- end.
-
-Definition ZOdiv a b := fst (ZOdiv_eucl a b).
-Definition ZOmod a b := snd (ZOdiv_eucl a b).
-
-
-Definition Ndiv_eucl (a b:N) : N * N :=
- match a, b with
- | N0, _ => (N0, N0)
- | _, N0 => (N0, a)
- | Npos na, Npos nb => Pdiv_eucl na nb
- end.
-
-Definition Ndiv a b := fst (Ndiv_eucl a b).
-Definition Nmod a b := snd (Ndiv_eucl a b).
-
-
-(* Proofs of specifications for these euclidean divisions. *)
-
-Theorem NPgeb_correct: forall (a:N)(b:positive),
- if NPgeb a b then a = (Nminus a (Npos b) + Npos b)%N else True.
-Proof.
- destruct a; intros; simpl; auto.
- generalize (Pcompare_Eq_eq p b).
- case_eq (Pcompare p b Eq); intros; auto.
- rewrite H0; auto.
- now rewrite Pminus_mask_diag.
- destruct (Pminus_mask_Gt p b H) as [d [H2 [H3 _]]].
- rewrite H2. rewrite <- H3.
- simpl; f_equal; apply Pplus_comm.
-Qed.
-
-Hint Rewrite Z_of_N_plus Z_of_N_mult Z_of_N_minus Zmult_1_l Zmult_assoc
- Zmult_plus_distr_l Zmult_plus_distr_r : zdiv.
-Hint Rewrite <- Zplus_assoc : zdiv.
-
-Theorem Pdiv_eucl_correct: forall a b,
- let (q,r) := Pdiv_eucl a b in
- Zpos a = Z_of_N q * Zpos b + Z_of_N r.
-Proof.
- induction a; cbv beta iota delta [Pdiv_eucl]; fold Pdiv_eucl; cbv zeta.
- intros b; generalize (IHa b); case Pdiv_eucl.
- intros q1 r1 Hq1.
- generalize (NPgeb_correct (2 * r1 + 1) b); case NPgeb; intros H.
- set (u := Nminus (2 * r1 + 1) (Npos b)) in * |- *.
- assert (HH: Z_of_N u = (Z_of_N (2 * r1 + 1) - Zpos b)%Z).
- rewrite H; autorewrite with zdiv; simpl.
- rewrite Zplus_comm, Zminus_plus; trivial.
- rewrite HH; autorewrite with zdiv; simpl Z_of_N.
- rewrite Zpos_xI, Hq1.
- autorewrite with zdiv; f_equal; rewrite Zplus_minus; trivial.
- rewrite Zpos_xI, Hq1; autorewrite with zdiv; auto.
- intros b; generalize (IHa b); case Pdiv_eucl.
- intros q1 r1 Hq1.
- generalize (NPgeb_correct (2 * r1) b); case NPgeb; intros H.
- set (u := Nminus (2 * r1) (Npos b)) in * |- *.
- assert (HH: Z_of_N u = (Z_of_N (2 * r1) - Zpos b)%Z).
- rewrite H; autorewrite with zdiv; simpl.
- rewrite Zplus_comm, Zminus_plus; trivial.
- rewrite HH; autorewrite with zdiv; simpl Z_of_N.
- rewrite Zpos_xO, Hq1.
- autorewrite with zdiv; f_equal; rewrite Zplus_minus; trivial.
- rewrite Zpos_xO, Hq1; autorewrite with zdiv; auto.
- destruct b; auto.
-Qed.
-
-Theorem ZOdiv_eucl_correct: forall a b,
- let (q,r) := ZOdiv_eucl a b in a = q * b + r.
-Proof.
- destruct a; destruct b; simpl; auto;
- generalize (Pdiv_eucl_correct p p0); case Pdiv_eucl; auto; intros;
- try change (Zneg p) with (Zopp (Zpos p)); rewrite H.
- destruct n; auto.
- repeat (rewrite Zopp_plus_distr || rewrite Zopp_mult_distr_l); trivial.
- repeat (rewrite Zopp_plus_distr || rewrite Zopp_mult_distr_r); trivial.
-Qed.
-
-Theorem Ndiv_eucl_correct: forall a b,
- let (q,r) := Ndiv_eucl a b in a = (q * b + r)%N.
-Proof.
- destruct a; destruct b; simpl; auto;
- generalize (Pdiv_eucl_correct p p0); case Pdiv_eucl; auto; intros;
- destruct n; destruct n0; simpl; simpl in H; try discriminate;
- injection H; intros; subst; trivial.
-Qed.
diff --git a/theories/ZArith/ZOrderedType.v b/theories/ZArith/ZOrderedType.v
deleted file mode 100644
index de4e4e98..00000000
--- a/theories/ZArith/ZOrderedType.v
+++ /dev/null
@@ -1,60 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-Require Import BinInt Zcompare Zorder Zbool ZArith_dec
- Equalities Orders OrdersTac.
-
-Local Open Scope Z_scope.
-
-(** * DecidableType structure for binary integers *)
-
-Module Z_as_UBE <: UsualBoolEq.
- Definition t := Z.
- Definition eq := @eq Z.
- Definition eqb := Zeq_bool.
- Definition eqb_eq x y := iff_sym (Zeq_is_eq_bool x y).
-End Z_as_UBE.
-
-Module Z_as_DT <: UsualDecidableTypeFull := Make_UDTF Z_as_UBE.
-
-(** Note that the last module fulfills by subtyping many other
- interfaces, such as [DecidableType] or [EqualityType]. *)
-
-
-(** * OrderedType structure for binary integers *)
-
-Module Z_as_OT <: OrderedTypeFull.
- Include Z_as_DT.
- Definition lt := Zlt.
- Definition le := Zle.
- Definition compare := Zcompare.
-
- Instance lt_strorder : StrictOrder Zlt.
- Proof. split; [ exact Zlt_irrefl | exact Zlt_trans ]. Qed.
-
- Instance lt_compat : Proper (Logic.eq==>Logic.eq==>iff) Zlt.
- Proof. repeat red; intros; subst; auto. Qed.
-
- Definition le_lteq := Zle_lt_or_eq_iff.
- Definition compare_spec := Zcompare_spec.
-
-End Z_as_OT.
-
-(** Note that [Z_as_OT] can also be seen as a [UsualOrderedType]
- and a [OrderedType] (and also as a [DecidableType]). *)
-
-
-
-(** * An [order] tactic for integers *)
-
-Module ZOrder := OTF_to_OrderTac Z_as_OT.
-Ltac z_order := ZOrder.order.
-
-(** Note that [z_order] is domain-agnostic: it will not prove
- [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x=y]. *)
-
diff --git a/theories/ZArith/Zabs.v b/theories/ZArith/Zabs.v
index 0f6e62b7..23473e93 100644
--- a/theories/ZArith/Zabs.v
+++ b/theories/ZArith/Zabs.v
@@ -1,226 +1,106 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zabs.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-(** Binary Integers (Pierre Crégut (CNET, Lannion, France) *)
+(** Binary Integers : properties of absolute value *)
+(** Initial author : Pierre Crégut (CNET, Lannion, France) *)
+
+(** THIS FILE IS DEPRECATED.
+ It is now almost entirely made of compatibility formulations
+ for results already present in BinInt.Z. *)
Require Import Arith_base.
Require Import BinPos.
Require Import BinInt.
+Require Import Zcompare.
Require Import Zorder.
-Require Import Zmax.
Require Import Znat.
Require Import ZArith_dec.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
(**********************************************************************)
(** * Properties of absolute value *)
-Lemma Zabs_eq : forall n:Z, 0 <= n -> Zabs n = n.
-Proof.
- intro x; destruct x; auto with arith.
- compute in |- *; intros; absurd (Gt = Gt); trivial with arith.
-Qed.
-
-Lemma Zabs_non_eq : forall n:Z, n <= 0 -> Zabs n = - n.
-Proof.
- intro x; destruct x; auto with arith.
- compute in |- *; intros; absurd (Gt = Gt); trivial with arith.
-Qed.
-
-Theorem Zabs_Zopp : forall n:Z, Zabs (- n) = Zabs n.
-Proof.
- intros z; case z; simpl in |- *; auto.
-Qed.
+Notation Zabs_eq := Z.abs_eq (only parsing).
+Notation Zabs_non_eq := Z.abs_neq (only parsing).
+Notation Zabs_Zopp := Z.abs_opp (only parsing).
+Notation Zabs_pos := Z.abs_nonneg (only parsing).
+Notation Zabs_involutive := Z.abs_involutive (only parsing).
+Notation Zabs_eq_case := Z.abs_eq_cases (only parsing).
+Notation Zabs_triangle := Z.abs_triangle (only parsing).
+Notation Zsgn_Zabs := Z.sgn_abs (only parsing).
+Notation Zabs_Zsgn := Z.abs_sgn (only parsing).
+Notation Zabs_Zmult := Z.abs_mul (only parsing).
+Notation Zabs_square := Z.abs_square (only parsing).
(** * Proving a property of the absolute value by cases *)
Lemma Zabs_ind :
forall (P:Z -> Prop) (n:Z),
- (n >= 0 -> P n) -> (n <= 0 -> P (- n)) -> P (Zabs n).
+ (n >= 0 -> P n) -> (n <= 0 -> P (- n)) -> P (Z.abs n).
Proof.
- intros P x H H0; elim (Z_lt_ge_dec x 0); intro.
- assert (x <= 0). apply Zlt_le_weak; assumption.
- rewrite Zabs_non_eq. apply H0. assumption. assumption.
- rewrite Zabs_eq. apply H; assumption. apply Zge_le. assumption.
+ intros. apply Z.abs_case_strong; Z.swap_greater; trivial.
+ intros x y Hx; now subst.
Qed.
-Theorem Zabs_intro : forall P (n:Z), P (- n) -> P n -> P (Zabs n).
+Theorem Zabs_intro : forall P (n:Z), P (- n) -> P n -> P (Z.abs n).
Proof.
- intros P z; case z; simpl in |- *; auto.
+ now destruct n.
Qed.
-Definition Zabs_dec : forall x:Z, {x = Zabs x} + {x = - Zabs x}.
+Definition Zabs_dec : forall x:Z, {x = Z.abs x} + {x = - Z.abs x}.
Proof.
- intro x; destruct x; auto with arith.
+ destruct x; auto.
Defined.
-Lemma Zabs_pos : forall n:Z, 0 <= Zabs n.
- intro x; destruct x; auto with arith; compute in |- *; intros H; inversion H.
-Qed.
-
-Lemma Zabs_involutive : forall x:Z, Zabs (Zabs x) = Zabs x.
-Proof.
- intros; apply Zabs_eq; apply Zabs_pos.
-Qed.
-
-Theorem Zabs_eq_case : forall n m:Z, Zabs n = Zabs m -> n = m \/ n = - m.
-Proof.
- intros z1 z2; case z1; case z2; simpl in |- *; auto;
- try (intros; discriminate); intros p1 p2 H1; injection H1;
- (intros H2; rewrite H2); auto.
-Qed.
-
-Lemma Zabs_spec : forall x:Z,
- 0 <= x /\ Zabs x = x \/
- 0 > x /\ Zabs x = -x.
-Proof.
- intros; unfold Zabs, Zle, Zgt; destruct x; simpl; intuition discriminate.
-Qed.
-
-(** * Triangular inequality *)
-
-Hint Local Resolve Zle_neg_pos: zarith.
-
-Theorem Zabs_triangle : forall n m:Z, Zabs (n + m) <= Zabs n + Zabs m.
-Proof.
- intros z1 z2; case z1; case z2; try (simpl in |- *; auto with zarith; fail).
- intros p1 p2;
- apply Zabs_intro with (P := fun x => x <= Zabs (Zpos p2) + Zabs (Zneg p1));
- try rewrite Zopp_plus_distr; auto with zarith.
- apply Zplus_le_compat; simpl in |- *; auto with zarith.
- apply Zplus_le_compat; simpl in |- *; auto with zarith.
- intros p1 p2;
- apply Zabs_intro with (P := fun x => x <= Zabs (Zpos p2) + Zabs (Zneg p1));
- try rewrite Zopp_plus_distr; auto with zarith.
- apply Zplus_le_compat; simpl in |- *; auto with zarith.
- apply Zplus_le_compat; simpl in |- *; auto with zarith.
-Qed.
-
-(** * Absolute value and multiplication *)
-
-Lemma Zsgn_Zabs : forall n:Z, n * Zsgn n = Zabs n.
-Proof.
- intro x; destruct x; rewrite Zmult_comm; auto with arith.
-Qed.
-
-Lemma Zabs_Zsgn : forall n:Z, Zabs n * Zsgn n = n.
-Proof.
- intro x; destruct x; rewrite Zmult_comm; auto with arith.
-Qed.
-
-Theorem Zabs_Zmult : forall n m:Z, Zabs (n * m) = Zabs n * Zabs m.
-Proof.
- intros z1 z2; case z1; case z2; simpl in |- *; auto.
-Qed.
-
-Theorem Zabs_square : forall a, Zabs a * Zabs a = a * a.
+Lemma Zabs_spec x :
+ 0 <= x /\ Z.abs x = x \/
+ 0 > x /\ Z.abs x = -x.
Proof.
- destruct a; simpl; auto.
-Qed.
-
-(** * Results about absolute value in nat. *)
-
-Theorem inj_Zabs_nat : forall z:Z, Z_of_nat (Zabs_nat z) = Zabs z.
-Proof.
- destruct z; simpl; auto; symmetry; apply Zpos_eq_Z_of_nat_o_nat_of_P.
-Qed.
-
-Theorem Zabs_nat_Z_of_nat: forall n, Zabs_nat (Z_of_nat n) = n.
-Proof.
- destruct n; simpl; auto.
- apply nat_of_P_o_P_of_succ_nat_eq_succ.
-Qed.
-
-Lemma Zabs_nat_mult: forall n m:Z, Zabs_nat (n*m) = (Zabs_nat n * Zabs_nat m)%nat.
-Proof.
- intros; apply inj_eq_rev.
- rewrite inj_mult; repeat rewrite inj_Zabs_nat; apply Zabs_Zmult.
-Qed.
-
-Lemma Zabs_nat_Zsucc:
- forall p, 0 <= p -> Zabs_nat (Zsucc p) = S (Zabs_nat p).
-Proof.
- intros; apply inj_eq_rev.
- rewrite inj_S; repeat rewrite inj_Zabs_nat, Zabs_eq; auto with zarith.
-Qed.
-
-Lemma Zabs_nat_Zplus:
- forall x y, 0<=x -> 0<=y -> Zabs_nat (x+y) = (Zabs_nat x + Zabs_nat y)%nat.
-Proof.
- intros; apply inj_eq_rev.
- rewrite inj_plus; repeat rewrite inj_Zabs_nat, Zabs_eq; auto with zarith.
- apply Zplus_le_0_compat; auto.
-Qed.
-
-Lemma Zabs_nat_Zminus:
- forall x y, 0 <= x <= y -> Zabs_nat (y - x) = (Zabs_nat y - Zabs_nat x)%nat.
-Proof.
- intros x y (H,H').
- assert (0 <= y) by (apply Zle_trans with x; auto).
- assert (0 <= y-x) by (apply Zle_minus_le_0; auto).
- apply inj_eq_rev.
- rewrite inj_minus; repeat rewrite inj_Zabs_nat, Zabs_eq; auto.
- rewrite Zmax_right; auto.
-Qed.
-
-Lemma Zabs_nat_le :
- forall n m:Z, 0 <= n <= m -> (Zabs_nat n <= Zabs_nat m)%nat.
-Proof.
- intros n m (H,H'); apply inj_le_rev.
- repeat rewrite inj_Zabs_nat, Zabs_eq; auto.
- apply Zle_trans with n; auto.
-Qed.
-
-Lemma Zabs_nat_lt :
- forall n m:Z, 0 <= n < m -> (Zabs_nat n < Zabs_nat m)%nat.
-Proof.
- intros n m (H,H'); apply inj_lt_rev.
- repeat rewrite inj_Zabs_nat, Zabs_eq; auto.
- apply Zlt_le_weak; apply Zle_lt_trans with n; auto.
+ Z.swap_greater. apply Z.abs_spec.
Qed.
(** * Some results about the sign function. *)
-Lemma Zsgn_Zmult : forall a b, Zsgn (a*b) = Zsgn a * Zsgn b.
-Proof.
- destruct a; destruct b; simpl; auto.
-Qed.
-
-Lemma Zsgn_Zopp : forall a, Zsgn (-a) = - Zsgn a.
-Proof.
- destruct a; simpl; auto.
-Qed.
+Notation Zsgn_Zmult := Z.sgn_mul (only parsing).
+Notation Zsgn_Zopp := Z.sgn_opp (only parsing).
+Notation Zsgn_pos := Z.sgn_pos_iff (only parsing).
+Notation Zsgn_neg := Z.sgn_neg_iff (only parsing).
+Notation Zsgn_null := Z.sgn_null_iff (only parsing).
(** A characterization of the sign function: *)
-Lemma Zsgn_spec : forall x:Z,
+Lemma Zsgn_spec x :
0 < x /\ Zsgn x = 1 \/
0 = x /\ Zsgn x = 0 \/
0 > x /\ Zsgn x = -1.
Proof.
- intros; unfold Zsgn, Zle, Zgt; destruct x; compute; intuition.
+ intros. Z.swap_greater. apply Z.sgn_spec.
Qed.
-Lemma Zsgn_pos : forall x:Z, Zsgn x = 1 <-> 0 < x.
-Proof.
- destruct x; now intuition.
-Qed.
+(** Compatibility *)
-Lemma Zsgn_neg : forall x:Z, Zsgn x = -1 <-> x < 0.
+Notation inj_Zabs_nat := Zabs2Nat.id_abs (only parsing).
+Notation Zabs_nat_Z_of_nat := Zabs2Nat.id (only parsing).
+Notation Zabs_nat_mult := Zabs2Nat.inj_mul (only parsing).
+Notation Zabs_nat_Zsucc := Zabs2Nat.inj_succ (only parsing).
+Notation Zabs_nat_Zplus := Zabs2Nat.inj_add (only parsing).
+Notation Zabs_nat_Zminus := (fun n m => Zabs2Nat.inj_sub m n) (only parsing).
+Notation Zabs_nat_compare := Zabs2Nat.inj_compare (only parsing).
+
+Lemma Zabs_nat_le n m : 0 <= n <= m -> (Z.abs_nat n <= Z.abs_nat m)%nat.
Proof.
- destruct x; now intuition.
+ intros (H,H'). apply Zabs2Nat.inj_le; trivial. now transitivity n.
Qed.
-Lemma Zsgn_null : forall x:Z, Zsgn x = 0 <-> x = 0.
+Lemma Zabs_nat_lt n m : 0 <= n < m -> (Zabs_nat n < Zabs_nat m)%nat.
Proof.
- destruct x; now intuition.
+ intros (H,H'). apply Zabs2Nat.inj_lt; trivial.
+ transitivity n; trivial. now apply Z.lt_le_incl.
Qed.
-
diff --git a/theories/ZArith/Zbool.v b/theories/ZArith/Zbool.v
index a4eebfb2..d0901282 100644
--- a/theories/ZArith/Zbool.v
+++ b/theories/ZArith/Zbool.v
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: Zbool.v 14641 2011-11-06 11:59:10Z herbelin $ *)
-
Require Import BinInt.
Require Import Zeven.
Require Import Zorder.
@@ -15,7 +13,6 @@ Require Import Zcompare.
Require Import ZArith_dec.
Require Import Sumbool.
-Unset Boxed Definitions.
Open Local Scope Z_scope.
(** * Boolean operations from decidability of order *)
@@ -36,29 +33,13 @@ Definition Zeven_odd_bool (x:Z) := bool_of_sumbool (Zeven_odd_dec x).
(**********************************************************************)
(** * Boolean comparisons of binary integers *)
-Definition Zle_bool (x y:Z) :=
- match x ?= y with
- | Gt => false
- | _ => true
- end.
+Notation Zle_bool := Z.leb (only parsing).
+Notation Zge_bool := Z.geb (only parsing).
+Notation Zlt_bool := Z.ltb (only parsing).
+Notation Zgt_bool := Z.gtb (only parsing).
-Definition Zge_bool (x y:Z) :=
- match x ?= y with
- | Lt => false
- | _ => true
- end.
-
-Definition Zlt_bool (x y:Z) :=
- match x ?= y with
- | Lt => true
- | _ => false
- end.
-
-Definition Zgt_bool (x y:Z) :=
- match x ?= y with
- | Gt => true
- | _ => false
- end.
+(** We now provide a direct [Z.eqb] that doesn't refer to [Z.compare].
+ The old [Zeq_bool] is kept for compatibility. *)
Definition Zeq_bool (x y:Z) :=
match x ?= y with
@@ -74,162 +55,130 @@ Definition Zneq_bool (x y:Z) :=
(** Properties in term of [if ... then ... else ...] *)
-Lemma Zle_cases :
- forall n m:Z, if Zle_bool n m then (n <= m) else (n > m).
+Lemma Zle_cases n m : if n <=? m then n <= m else n > m.
Proof.
- intros x y; unfold Zle_bool, Zle, Zgt in |- *.
- case (x ?= y); auto; discriminate.
+ case Z.leb_spec; now Z.swap_greater.
Qed.
-Lemma Zlt_cases :
- forall n m:Z, if Zlt_bool n m then (n < m) else (n >= m).
+Lemma Zlt_cases n m : if n <? m then n < m else n >= m.
Proof.
- intros x y; unfold Zlt_bool, Zlt, Zge in |- *.
- case (x ?= y); auto; discriminate.
+ case Z.ltb_spec; now Z.swap_greater.
Qed.
-Lemma Zge_cases :
- forall n m:Z, if Zge_bool n m then (n >= m) else (n < m).
+Lemma Zge_cases n m : if n >=? m then n >= m else n < m.
Proof.
- intros x y; unfold Zge_bool, Zge, Zlt in |- *.
- case (x ?= y); auto; discriminate.
+ rewrite Z.geb_leb. case Z.leb_spec; now Z.swap_greater.
Qed.
-Lemma Zgt_cases :
- forall n m:Z, if Zgt_bool n m then (n > m) else (n <= m).
+Lemma Zgt_cases n m : if n >? m then n > m else n <= m.
Proof.
- intros x y; unfold Zgt_bool, Zgt, Zle in |- *.
- case (x ?= y); auto; discriminate.
+ rewrite Z.gtb_ltb. case Z.ltb_spec; now Z.swap_greater.
Qed.
-(** Lemmas on [Zle_bool] used in contrib/graphs *)
+(** Lemmas on [Z.leb] used in contrib/graphs *)
-Lemma Zle_bool_imp_le : forall n m:Z, Zle_bool n m = true -> (n <= m).
+Lemma Zle_bool_imp_le n m : (n <=? m) = true -> (n <= m).
Proof.
- unfold Zle_bool, Zle in |- *. intros x y. unfold not in |- *.
- case (x ?= y); intros; discriminate.
+ apply Z.leb_le.
Qed.
-Lemma Zle_imp_le_bool : forall n m:Z, (n <= m) -> Zle_bool n m = true.
+Lemma Zle_imp_le_bool n m : (n <= m) -> (n <=? m) = true.
Proof.
- unfold Zle, Zle_bool in |- *. intros x y. case (x ?= y); trivial. intro. elim (H (refl_equal _)).
+ apply Z.leb_le.
Qed.
-Lemma Zle_bool_refl : forall n:Z, Zle_bool n n = true.
-Proof.
- intro. apply Zle_imp_le_bool. apply Zeq_le. reflexivity.
-Qed.
+Notation Zle_bool_refl := Z.leb_refl (only parsing).
-Lemma Zle_bool_antisym :
- forall n m:Z, Zle_bool n m = true -> Zle_bool m n = true -> n = m.
+Lemma Zle_bool_antisym n m :
+ (n <=? m) = true -> (m <=? n) = true -> n = m.
Proof.
- intros. apply Zle_antisym. apply Zle_bool_imp_le. assumption.
- apply Zle_bool_imp_le. assumption.
+ rewrite !Z.leb_le. apply Z.le_antisymm.
Qed.
-Lemma Zle_bool_trans :
- forall n m p:Z,
- Zle_bool n m = true -> Zle_bool m p = true -> Zle_bool n p = true.
+Lemma Zle_bool_trans n m p :
+ (n <=? m) = true -> (m <=? p) = true -> (n <=? p) = true.
Proof.
- intros x y z; intros. apply Zle_imp_le_bool. apply Zle_trans with (m := y). apply Zle_bool_imp_le. assumption.
- apply Zle_bool_imp_le. assumption.
+ rewrite !Z.leb_le. apply Z.le_trans.
Qed.
-Definition Zle_bool_total :
- forall x y:Z, {Zle_bool x y = true} + {Zle_bool y x = true}.
+Definition Zle_bool_total x y :
+ { x <=? y = true } + { y <=? x = true }.
Proof.
- intros x y; intros. unfold Zle_bool in |- *. cut ((x ?= y) = Gt <-> (y ?= x) = Lt).
- case (x ?= y). left. reflexivity.
- left. reflexivity.
- right. rewrite (proj1 H (refl_equal _)). reflexivity.
- apply Zcompare_Gt_Lt_antisym.
+ case_eq (x <=? y); intros H.
+ - left; trivial.
+ - right. apply Z.leb_gt in H. now apply Z.leb_le, Z.lt_le_incl.
Defined.
-Lemma Zle_bool_plus_mono :
- forall n m p q:Z,
- Zle_bool n m = true ->
- Zle_bool p q = true -> Zle_bool (n + p) (m + q) = true.
+Lemma Zle_bool_plus_mono n m p q :
+ (n <=? m) = true ->
+ (p <=? q) = true ->
+ (n + p <=? m + q) = true.
Proof.
- intros. apply Zle_imp_le_bool. apply Zplus_le_compat. apply Zle_bool_imp_le. assumption.
- apply Zle_bool_imp_le. assumption.
+ rewrite !Z.leb_le. apply Z.add_le_mono.
Qed.
-Lemma Zone_pos : Zle_bool 1 0 = false.
+Lemma Zone_pos : 1 <=? 0 = false.
Proof.
- reflexivity.
+ reflexivity.
Qed.
-Lemma Zone_min_pos : forall n:Z, Zle_bool n 0 = false -> Zle_bool 1 n = true.
+Lemma Zone_min_pos n : (n <=? 0) = false -> (1 <=? n) = true.
Proof.
- intros x; intros. apply Zle_imp_le_bool. change (Zsucc 0 <= x) in |- *. apply Zgt_le_succ. generalize H.
- unfold Zle_bool, Zgt in |- *. case (x ?= 0). intro H0. discriminate H0.
- intro H0. discriminate H0.
- reflexivity.
+ rewrite Z.leb_le, Z.leb_gt. apply Z.le_succ_l.
Qed.
(** Properties in term of [iff] *)
-Lemma Zle_is_le_bool : forall n m:Z, (n <= m) <-> Zle_bool n m = true.
+Lemma Zle_is_le_bool n m : (n <= m) <-> (n <=? m) = true.
Proof.
- intros. split. intro. apply Zle_imp_le_bool. assumption.
- intro. apply Zle_bool_imp_le. assumption.
+ symmetry. apply Z.leb_le.
Qed.
-Lemma Zge_is_le_bool : forall n m:Z, (n >= m) <-> Zle_bool m n = true.
+Lemma Zge_is_le_bool n m : (n >= m) <-> (m <=? n) = true.
Proof.
- intros. split. intro. apply Zle_imp_le_bool. apply Zge_le. assumption.
- intro. apply Zle_ge. apply Zle_bool_imp_le. assumption.
+ Z.swap_greater. symmetry. apply Z.leb_le.
Qed.
-Lemma Zlt_is_lt_bool : forall n m:Z, (n < m) <-> Zlt_bool n m = true.
+Lemma Zlt_is_lt_bool n m : (n < m) <-> (n <? m) = true.
Proof.
-intros n m; unfold Zlt_bool, Zlt.
-destruct (n ?= m); simpl; split; now intro.
+ symmetry. apply Z.ltb_lt.
Qed.
-Lemma Zgt_is_gt_bool : forall n m:Z, (n > m) <-> Zgt_bool n m = true.
+Lemma Zgt_is_gt_bool n m : (n > m) <-> (n >? m) = true.
Proof.
-intros n m; unfold Zgt_bool, Zgt.
-destruct (n ?= m); simpl; split; now intro.
+ Z.swap_greater. rewrite Z.gtb_ltb. symmetry. apply Z.ltb_lt.
Qed.
-Lemma Zlt_is_le_bool :
- forall n m:Z, (n < m) <-> Zle_bool n (m - 1) = true.
+Lemma Zlt_is_le_bool n m : (n < m) <-> (n <=? m - 1) = true.
Proof.
- intros x y. split. intro. apply Zle_imp_le_bool. apply Zlt_succ_le. rewrite (Zsucc_pred y) in H.
- assumption.
- intro. rewrite (Zsucc_pred y). apply Zle_lt_succ. apply Zle_bool_imp_le. assumption.
+ rewrite Z.leb_le. apply Z.lt_le_pred.
Qed.
-Lemma Zgt_is_le_bool :
- forall n m:Z, (n > m) <-> Zle_bool m (n - 1) = true.
+Lemma Zgt_is_le_bool n m : (n > m) <-> (m <=? n - 1) = true.
Proof.
- intros x y. apply iff_trans with (y < x). split. exact (Zgt_lt x y).
- exact (Zlt_gt y x).
- exact (Zlt_is_le_bool y x).
+ Z.swap_greater. rewrite Z.leb_le. apply Z.lt_le_pred.
Qed.
-Lemma Zeq_is_eq_bool : forall x y, x = y <-> Zeq_bool x y = true.
+(** Properties of the deprecated [Zeq_bool] *)
+
+Lemma Zeq_is_eq_bool x y : x = y <-> Zeq_bool x y = true.
Proof.
- intros; unfold Zeq_bool.
- generalize (Zcompare_Eq_iff_eq x y); destruct Zcompare; intuition;
- try discriminate.
+ unfold Zeq_bool.
+ rewrite <- Z.compare_eq_iff. destruct Z.compare; now split.
Qed.
-Lemma Zeq_bool_eq : forall x y, Zeq_bool x y = true -> x = y.
+Lemma Zeq_bool_eq x y : Zeq_bool x y = true -> x = y.
Proof.
- intros x y H; apply <- Zeq_is_eq_bool; auto.
+ apply Zeq_is_eq_bool.
Qed.
-Lemma Zeq_bool_neq : forall x y, Zeq_bool x y = false -> x <> y.
+Lemma Zeq_bool_neq x y : Zeq_bool x y = false -> x <> y.
Proof.
- unfold Zeq_bool; red ; intros; subst.
- rewrite Zcompare_refl in H.
- discriminate.
+ rewrite Zeq_is_eq_bool; now destruct Zeq_bool.
Qed.
-Lemma Zeq_bool_if : forall x y, if Zeq_bool x y then x=y else x<>y.
+Lemma Zeq_bool_if x y : if Zeq_bool x y then x=y else x<>y.
Proof.
- intros. generalize (Zeq_bool_eq x y)(Zeq_bool_neq x y).
- destruct Zeq_bool; auto.
-Qed. \ No newline at end of file
+ generalize (Zeq_bool_eq x y) (Zeq_bool_neq x y).
+ destruct Zeq_bool; auto.
+Qed.
diff --git a/theories/ZArith/Zcompare.v b/theories/ZArith/Zcompare.v
index ae5302ee..20e1b006 100644
--- a/theories/ZArith/Zcompare.v
+++ b/theories/ZArith/Zcompare.v
@@ -1,387 +1,91 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $$ i*)
+(** Binary Integers : results about Zcompare *)
+(** Initial author: Pierre Crégut (CNET, Lannion, France *)
-(**********************************************************************)
-(** Binary Integers (Pierre Crégut, CNET, Lannion, France) *)
-(**********************************************************************)
+(** THIS FILE IS DEPRECATED.
+ It is now almost entirely made of compatibility formulations
+ for results already present in BinInt.Z. *)
-Require Export BinPos.
-Require Export BinInt.
-Require Import Lt.
-Require Import Gt.
-Require Import Plus.
-Require Import Mult.
+Require Export BinPos BinInt.
+Require Import Lt Gt Plus Mult. (* Useless now, for compatibility only *)
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
(***************************)
(** * Comparison on integers *)
-Lemma Zcompare_refl : forall n:Z, (n ?= n) = Eq.
-Proof.
- intro x; destruct x as [| p| p]; simpl in |- *;
- [ reflexivity | apply Pcompare_refl | rewrite Pcompare_refl; reflexivity ].
-Qed.
-
-Lemma Zcompare_Eq_eq : forall n m:Z, (n ?= m) = Eq -> n = m.
-Proof.
- intros x y; destruct x as [| x'| x']; destruct y as [| y'| y']; simpl in |- *;
- intro H; reflexivity || (try discriminate H);
- [ rewrite (Pcompare_Eq_eq x' y' H); reflexivity
- | rewrite (Pcompare_Eq_eq x' y');
- [ reflexivity
- | destruct ((x' ?= y')%positive Eq); reflexivity || discriminate ] ].
-Qed.
-
-Ltac destr_zcompare :=
- match goal with |- context [Zcompare ?x ?y] =>
- let H := fresh "H" in
- case_eq (Zcompare x y); intro H;
- [generalize (Zcompare_Eq_eq _ _ H); clear H; intro H |
- change (x<y)%Z in H |
- change (x>y)%Z in H ]
- end.
-
-Lemma Zcompare_Eq_iff_eq : forall n m:Z, (n ?= m) = Eq <-> n = m.
-Proof.
- intros x y; split; intro E;
- [ apply Zcompare_Eq_eq; assumption | rewrite E; apply Zcompare_refl ].
-Qed.
-
-Lemma Zcompare_antisym : forall n m:Z, CompOpp (n ?= m) = (m ?= n).
-Proof.
- intros x y; destruct x; destruct y; simpl in |- *;
- reflexivity || discriminate H || rewrite Pcompare_antisym;
- reflexivity.
-Qed.
-
Lemma Zcompare_Gt_Lt_antisym : forall n m:Z, (n ?= m) = Gt <-> (m ?= n) = Lt.
-Proof.
- intros x y.
- rewrite <- Zcompare_antisym. change Gt with (CompOpp Lt).
- split.
- auto using CompOpp_inj.
- intros; f_equal; auto.
-Qed.
-
-Lemma Zcompare_spec : forall n m, CompSpec eq Zlt n m (n ?= m).
-Proof.
- intros.
- destruct (n?=m) as [ ]_eqn:H; constructor; auto.
- apply Zcompare_Eq_eq; auto.
- red; rewrite <- Zcompare_antisym, H; auto.
-Qed.
+Proof Z.gt_lt_iff.
+Lemma Zcompare_antisym n m : CompOpp (n ?= m) = (m ?= n).
+Proof eq_sym (Z.compare_antisym n m).
(** * Transitivity of comparison *)
Lemma Zcompare_Lt_trans :
forall n m p:Z, (n ?= m) = Lt -> (m ?= p) = Lt -> (n ?= p) = Lt.
-Proof.
- intros x y z; case x; case y; case z; simpl;
- try discriminate; auto with arith.
- intros; eapply Plt_trans; eauto.
- intros p q r; rewrite 3 Pcompare_antisym; simpl.
- intros; eapply Plt_trans; eauto.
-Qed.
+Proof Z.lt_trans.
Lemma Zcompare_Gt_trans :
forall n m p:Z, (n ?= m) = Gt -> (m ?= p) = Gt -> (n ?= p) = Gt.
Proof.
- intros n m p Hnm Hmp.
- apply <- Zcompare_Gt_Lt_antisym.
- apply -> Zcompare_Gt_Lt_antisym in Hnm.
- apply -> Zcompare_Gt_Lt_antisym in Hmp.
- eapply Zcompare_Lt_trans; eauto.
+ intros n m p. change (n > m -> m > p -> n > p).
+ Z.swap_greater. intros. now transitivity m.
Qed.
(** * Comparison and opposite *)
-Lemma Zcompare_opp : forall n m:Z, (n ?= m) = (- m ?= - n).
+Lemma Zcompare_opp n m : (n ?= m) = (- m ?= - n).
Proof.
- intros x y; case x; case y; simpl in |- *; auto with arith; intros;
- rewrite <- ZC4; trivial with arith.
+ symmetry. apply Z.compare_opp.
Qed.
-Hint Local Resolve Pcompare_refl.
-
(** * Comparison first-order specification *)
-Lemma Zcompare_Gt_spec :
- forall n m:Z, (n ?= m) = Gt -> exists h : positive, n + - m = Zpos h.
+Lemma Zcompare_Gt_spec n m : (n ?= m) = Gt -> exists h, n + - m = Zpos h.
Proof.
- intros x y; case x; case y;
- [ simpl in |- *; intros H; discriminate H
- | simpl in |- *; intros p H; discriminate H
- | intros p H; exists p; simpl in |- *; auto with arith
- | intros p H; exists p; simpl in |- *; auto with arith
- | intros q p H; exists (p - q)%positive; unfold Zplus, Zopp in |- *;
- unfold Zcompare in H; rewrite H; trivial with arith
- | intros q p H; exists (p + q)%positive; simpl in |- *; trivial with arith
- | simpl in |- *; intros p H; discriminate H
- | simpl in |- *; intros q p H; discriminate H
- | unfold Zcompare in |- *; intros q p; rewrite <- ZC4; intros H;
- exists (q - p)%positive; simpl in |- *; rewrite (ZC1 q p H);
- trivial with arith ].
+ rewrite Z.compare_sub. unfold Z.sub.
+ destruct (n+-m) as [|p|p]; try discriminate. now exists p.
Qed.
(** * Comparison and addition *)
-Lemma weaken_Zcompare_Zplus_compatible :
- (forall (n m:Z) (p:positive), (Zpos p + n ?= Zpos p + m) = (n ?= m)) ->
- forall n m p:Z, (p + n ?= p + m) = (n ?= m).
+Lemma Zcompare_plus_compat n m p : (p + n ?= p + m) = (n ?= m).
Proof.
- intros H x y z; destruct z;
- [ reflexivity
- | apply H
- | rewrite (Zcompare_opp x y); rewrite Zcompare_opp;
- do 2 rewrite Zopp_plus_distr; rewrite Zopp_neg;
- apply H ].
+ apply Z.add_compare_mono_l.
Qed.
-Hint Local Resolve ZC4.
-
-Lemma weak_Zcompare_Zplus_compatible :
- forall (n m:Z) (p:positive), (Zpos p + n ?= Zpos p + m) = (n ?= m).
-Proof.
- intros x y z; case x; case y; simpl in |- *; auto with arith;
- [ intros p; apply nat_of_P_lt_Lt_compare_complement_morphism; apply ZL17
- | intros p; ElimPcompare z p; intros E; rewrite E; auto with arith;
- apply nat_of_P_gt_Gt_compare_complement_morphism;
- rewrite nat_of_P_minus_morphism;
- [ unfold gt in |- *; apply ZL16 | assumption ]
- | intros p; ElimPcompare z p; intros E; auto with arith;
- apply nat_of_P_gt_Gt_compare_complement_morphism;
- unfold gt in |- *; apply ZL17
- | intros p q; ElimPcompare q p; intros E; rewrite E;
- [ rewrite (Pcompare_Eq_eq q p E); apply Pcompare_refl
- | apply nat_of_P_lt_Lt_compare_complement_morphism;
- do 2 rewrite nat_of_P_plus_morphism; apply plus_lt_compat_l;
- apply nat_of_P_lt_Lt_compare_morphism with (1 := E)
- | apply nat_of_P_gt_Gt_compare_complement_morphism; unfold gt in |- *;
- do 2 rewrite nat_of_P_plus_morphism; apply plus_lt_compat_l;
- exact (nat_of_P_gt_Gt_compare_morphism q p E) ]
- | intros p q; ElimPcompare z p; intros E; rewrite E; auto with arith;
- apply nat_of_P_gt_Gt_compare_complement_morphism;
- rewrite nat_of_P_minus_morphism;
- [ unfold gt in |- *; apply lt_trans with (m := nat_of_P z);
- [ apply ZL16 | apply ZL17 ]
- | assumption ]
- | intros p; ElimPcompare z p; intros E; rewrite E; auto with arith;
- simpl in |- *; apply nat_of_P_lt_Lt_compare_complement_morphism;
- rewrite nat_of_P_minus_morphism; [ apply ZL16 | assumption ]
- | intros p q; ElimPcompare z q; intros E; rewrite E; auto with arith;
- simpl in |- *; apply nat_of_P_lt_Lt_compare_complement_morphism;
- rewrite nat_of_P_minus_morphism;
- [ apply lt_trans with (m := nat_of_P z); [ apply ZL16 | apply ZL17 ]
- | assumption ]
- | intros p q; ElimPcompare z q; intros E0; rewrite E0; ElimPcompare z p;
- intros E1; rewrite E1; ElimPcompare q p; intros E2;
- rewrite E2; auto with arith;
- [ absurd ((q ?= p)%positive Eq = Lt);
- [ rewrite <- (Pcompare_Eq_eq z q E0);
- rewrite <- (Pcompare_Eq_eq z p E1); rewrite (Pcompare_refl z);
- discriminate
- | assumption ]
- | absurd ((q ?= p)%positive Eq = Gt);
- [ rewrite <- (Pcompare_Eq_eq z q E0);
- rewrite <- (Pcompare_Eq_eq z p E1); rewrite (Pcompare_refl z);
- discriminate
- | assumption ]
- | absurd ((z ?= p)%positive Eq = Lt);
- [ rewrite (Pcompare_Eq_eq z q E0); rewrite <- (Pcompare_Eq_eq q p E2);
- rewrite (Pcompare_refl q); discriminate
- | assumption ]
- | absurd ((z ?= p)%positive Eq = Lt);
- [ rewrite (Pcompare_Eq_eq z q E0); rewrite E2; discriminate
- | assumption ]
- | absurd ((z ?= p)%positive Eq = Gt);
- [ rewrite (Pcompare_Eq_eq z q E0); rewrite <- (Pcompare_Eq_eq q p E2);
- rewrite (Pcompare_refl q); discriminate
- | assumption ]
- | absurd ((z ?= p)%positive Eq = Gt);
- [ rewrite (Pcompare_Eq_eq z q E0); rewrite E2; discriminate
- | assumption ]
- | absurd ((z ?= q)%positive Eq = Lt);
- [ rewrite (Pcompare_Eq_eq z p E1); rewrite (Pcompare_Eq_eq q p E2);
- rewrite (Pcompare_refl p); discriminate
- | assumption ]
- | absurd ((p ?= q)%positive Eq = Gt);
- [ rewrite <- (Pcompare_Eq_eq z p E1); rewrite E0; discriminate
- | apply ZC2; assumption ]
- | simpl in |- *; rewrite (Pcompare_Eq_eq q p E2);
- rewrite (Pcompare_refl (p - z)); auto with arith
- | simpl in |- *; rewrite <- ZC4;
- apply nat_of_P_gt_Gt_compare_complement_morphism;
- rewrite nat_of_P_minus_morphism;
- [ rewrite nat_of_P_minus_morphism;
- [ unfold gt in |- *; apply plus_lt_reg_l with (p := nat_of_P z);
- rewrite le_plus_minus_r;
- [ rewrite le_plus_minus_r;
- [ apply nat_of_P_lt_Lt_compare_morphism; assumption
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- assumption ]
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- assumption ]
- | apply ZC2; assumption ]
- | apply ZC2; assumption ]
- | simpl in |- *; rewrite <- ZC4;
- apply nat_of_P_lt_Lt_compare_complement_morphism;
- rewrite nat_of_P_minus_morphism;
- [ rewrite nat_of_P_minus_morphism;
- [ apply plus_lt_reg_l with (p := nat_of_P z);
- rewrite le_plus_minus_r;
- [ rewrite le_plus_minus_r;
- [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1;
- assumption
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- assumption ]
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- assumption ]
- | apply ZC2; assumption ]
- | apply ZC2; assumption ]
- | absurd ((z ?= q)%positive Eq = Lt);
- [ rewrite (Pcompare_Eq_eq q p E2); rewrite E1; discriminate
- | assumption ]
- | absurd ((q ?= p)%positive Eq = Lt);
- [ cut ((q ?= p)%positive Eq = Gt);
- [ intros E; rewrite E; discriminate
- | apply nat_of_P_gt_Gt_compare_complement_morphism;
- unfold gt in |- *; apply lt_trans with (m := nat_of_P z);
- [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; assumption
- | apply nat_of_P_lt_Lt_compare_morphism; assumption ] ]
- | assumption ]
- | absurd ((z ?= q)%positive Eq = Gt);
- [ rewrite (Pcompare_Eq_eq z p E1); rewrite (Pcompare_Eq_eq q p E2);
- rewrite (Pcompare_refl p); discriminate
- | assumption ]
- | absurd ((z ?= q)%positive Eq = Gt);
- [ rewrite (Pcompare_Eq_eq z p E1); rewrite ZC1;
- [ discriminate | assumption ]
- | assumption ]
- | absurd ((z ?= q)%positive Eq = Gt);
- [ rewrite (Pcompare_Eq_eq q p E2); rewrite E1; discriminate
- | assumption ]
- | absurd ((q ?= p)%positive Eq = Gt);
- [ rewrite ZC1;
- [ discriminate
- | apply nat_of_P_gt_Gt_compare_complement_morphism;
- unfold gt in |- *; apply lt_trans with (m := nat_of_P z);
- [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; assumption
- | apply nat_of_P_lt_Lt_compare_morphism; assumption ] ]
- | assumption ]
- | simpl in |- *; rewrite (Pcompare_Eq_eq q p E2); apply Pcompare_refl
- | simpl in |- *; apply nat_of_P_gt_Gt_compare_complement_morphism;
- unfold gt in |- *; rewrite nat_of_P_minus_morphism;
- [ rewrite nat_of_P_minus_morphism;
- [ apply plus_lt_reg_l with (p := nat_of_P p);
- rewrite le_plus_minus_r;
- [ rewrite plus_comm; apply plus_lt_reg_l with (p := nat_of_P q);
- rewrite plus_assoc; rewrite le_plus_minus_r;
- [ rewrite (plus_comm (nat_of_P q)); apply plus_lt_compat_l;
- apply nat_of_P_lt_Lt_compare_morphism;
- assumption
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- apply ZC1; assumption ]
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- apply ZC1; assumption ]
- | assumption ]
- | assumption ]
- | simpl in |- *; apply nat_of_P_lt_Lt_compare_complement_morphism;
- rewrite nat_of_P_minus_morphism;
- [ rewrite nat_of_P_minus_morphism;
- [ apply plus_lt_reg_l with (p := nat_of_P q);
- rewrite le_plus_minus_r;
- [ rewrite plus_comm; apply plus_lt_reg_l with (p := nat_of_P p);
- rewrite plus_assoc; rewrite le_plus_minus_r;
- [ rewrite (plus_comm (nat_of_P p)); apply plus_lt_compat_l;
- apply nat_of_P_lt_Lt_compare_morphism;
- apply ZC1; assumption
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- apply ZC1; assumption ]
- | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism;
- apply ZC1; assumption ]
- | assumption ]
- | assumption ] ] ].
-Qed.
-
-Lemma Zcompare_plus_compat : forall n m p:Z, (p + n ?= p + m) = (n ?= m).
+Lemma Zplus_compare_compat (r:comparison) (n m p q:Z) :
+ (n ?= m) = r -> (p ?= q) = r -> (n + p ?= m + q) = r.
Proof.
- exact (weaken_Zcompare_Zplus_compatible weak_Zcompare_Zplus_compatible).
+ rewrite (Z.compare_sub n), (Z.compare_sub p), (Z.compare_sub (n+p)).
+ unfold Z.sub. rewrite Z.opp_add_distr. rewrite Z.add_shuffle1.
+ destruct (n+-m), (p+-q); simpl; intros; now subst.
Qed.
-Lemma Zplus_compare_compat :
- forall (r:comparison) (n m p q:Z),
- (n ?= m) = r -> (p ?= q) = r -> (n + p ?= m + q) = r.
+Lemma Zcompare_succ_Gt n : (Z.succ n ?= n) = Gt.
Proof.
- intros r x y z t; case r;
- [ intros H1 H2; elim (Zcompare_Eq_iff_eq x y); elim (Zcompare_Eq_iff_eq z t);
- intros H3 H4 H5 H6; rewrite H3;
- [ rewrite H5;
- [ elim (Zcompare_Eq_iff_eq (y + t) (y + t)); auto with arith
- | auto with arith ]
- | auto with arith ]
- | intros H1 H2; elim (Zcompare_Gt_Lt_antisym (y + t) (x + z)); intros H3 H4;
- apply H3; apply Zcompare_Gt_trans with (m := y + z);
- [ rewrite Zcompare_plus_compat; elim (Zcompare_Gt_Lt_antisym t z);
- auto with arith
- | do 2 rewrite <- (Zplus_comm z); rewrite Zcompare_plus_compat;
- elim (Zcompare_Gt_Lt_antisym y x); auto with arith ]
- | intros H1 H2; apply Zcompare_Gt_trans with (m := x + t);
- [ rewrite Zcompare_plus_compat; assumption
- | do 2 rewrite <- (Zplus_comm t); rewrite Zcompare_plus_compat;
- assumption ] ].
+ apply Z.lt_gt. apply Z.lt_succ_diag_r.
Qed.
-Lemma Zcompare_succ_Gt : forall n:Z, (Zsucc n ?= n) = Gt.
+Lemma Zcompare_Gt_not_Lt n m : (n ?= m) = Gt <-> (n ?= m+1) <> Lt.
Proof.
- intro x; unfold Zsucc in |- *; pattern x at 2 in |- *;
- rewrite <- (Zplus_0_r x); rewrite Zcompare_plus_compat;
- reflexivity.
-Qed.
-
-Lemma Zcompare_Gt_not_Lt : forall n m:Z, (n ?= m) = Gt <-> (n ?= m + 1) <> Lt.
-Proof.
- intros x y; split;
- [ intro H; elim_compare x (y + 1);
- [ intro H1; rewrite H1; discriminate
- | intros H1; elim Zcompare_Gt_spec with (1 := H); intros h H2;
- absurd ((nat_of_P h > 0)%nat /\ (nat_of_P h < 1)%nat);
- [ unfold not in |- *; intros H3; elim H3; intros H4 H5;
- absurd (nat_of_P h > 0)%nat;
- [ unfold gt in |- *; apply le_not_lt; apply le_S_n; exact H5
- | assumption ]
- | split;
- [ elim (ZL4 h); intros i H3; rewrite H3; apply gt_Sn_O
- | change (nat_of_P h < nat_of_P 1)%nat in |- *;
- apply nat_of_P_lt_Lt_compare_morphism;
- change ((Zpos h ?= 1) = Lt) in |- *; rewrite <- H2;
- rewrite <- (fun m n:Z => Zcompare_plus_compat m n y);
- rewrite (Zplus_comm x); rewrite Zplus_assoc;
- rewrite Zplus_opp_r; simpl in |- *; exact H1 ] ]
- | intros H1; rewrite H1; discriminate ]
- | intros H; elim_compare x (y + 1);
- [ intros H1; elim (Zcompare_Eq_iff_eq x (y + 1)); intros H2 H3;
- rewrite (H2 H1); exact (Zcompare_succ_Gt y)
- | intros H1; absurd ((x ?= y + 1) = Lt); assumption
- | intros H1; apply Zcompare_Gt_trans with (m := Zsucc y);
- [ exact H1 | exact (Zcompare_succ_Gt y) ] ] ].
+ change (n > m <-> n >= m+1). Z.swap_greater. symmetry. apply Z.le_succ_l.
Qed.
(** * Successor and comparison *)
-Lemma Zcompare_succ_compat : forall n m:Z, (Zsucc n ?= Zsucc m) = (n ?= m).
+Lemma Zcompare_succ_compat n m : (Z.succ n ?= Z.succ m) = (n ?= m).
Proof.
- intros n m; unfold Zsucc in |- *; do 2 rewrite (fun t:Z => Zplus_comm t 1);
- rewrite Zcompare_plus_compat; auto with arith.
+ rewrite <- 2 Z.add_1_l. apply Zcompare_plus_compat.
Qed.
(** * Multiplication and comparison *)
@@ -389,28 +93,24 @@ Qed.
Lemma Zcompare_mult_compat :
forall (p:positive) (n m:Z), (Zpos p * n ?= Zpos p * m) = (n ?= m).
Proof.
- intros x; induction x as [p H| p H| ];
- [ intros y z; cut (Zpos (xI p) = Zpos p + Zpos p + 1);
- [ intros E; rewrite E; do 4 rewrite Zmult_plus_distr_l;
- do 2 rewrite Zmult_1_l; apply Zplus_compare_compat;
- [ apply Zplus_compare_compat; apply H | trivial with arith ]
- | simpl in |- *; rewrite (Pplus_diag p); trivial with arith ]
- | intros y z; cut (Zpos (xO p) = Zpos p + Zpos p);
- [ intros E; rewrite E; do 2 rewrite Zmult_plus_distr_l;
- apply Zplus_compare_compat; apply H
- | simpl in |- *; rewrite (Pplus_diag p); trivial with arith ]
- | intros y z; do 2 rewrite Zmult_1_l; trivial with arith ].
+ intros p [|n|n] [|m|m]; simpl; trivial; now rewrite Pos.mul_compare_mono_l.
Qed.
+Lemma Zmult_compare_compat_l n m p:
+ p > 0 -> (n ?= m) = (p * n ?= p * m).
+Proof.
+ intros; destruct p; try discriminate.
+ symmetry. apply Zcompare_mult_compat.
+Qed.
-(** * Reverting [x ?= y] to trichotomy *)
-
-Lemma rename :
- forall (A:Type) (P:A -> Prop) (x:A), (forall y:A, x = y -> P y) -> P x.
+Lemma Zmult_compare_compat_r n m p :
+ p > 0 -> (n ?= m) = (n * p ?= m * p).
Proof.
- auto with arith.
+ intros; rewrite 2 (Zmult_comm _ p); now apply Zmult_compare_compat_l.
Qed.
+(** * Relating [x ?= y] to [=], [<=], [<], [>=] or [>] *)
+
Lemma Zcompare_elim :
forall (c1 c2 c3:Prop) (n m:Z),
(n = m -> c1) ->
@@ -421,11 +121,7 @@ Lemma Zcompare_elim :
| Gt => c3
end.
Proof.
- intros c1 c2 c3 x y; intros.
- apply rename with (x := x ?= y); intro r; elim r;
- [ intro; apply H; apply (Zcompare_Eq_eq x y); assumption
- | unfold Zlt in H0; assumption
- | unfold Zgt in H1; assumption ].
+ intros. case Z.compare_spec; trivial. now Z.swap_greater.
Qed.
Lemma Zcompare_eq_case :
@@ -436,26 +132,9 @@ Lemma Zcompare_eq_case :
| Gt => c3
end.
Proof.
- intros c1 c2 c3 x y; intros.
- rewrite H0; rewrite Zcompare_refl.
- assumption.
+ intros. subst. now rewrite Z.compare_refl.
Qed.
-(** * Decompose an egality between two [?=] relations into 3 implications *)
-
-Lemma Zcompare_egal_dec :
- forall n m p q:Z,
- (n < m -> p < q) ->
- ((n ?= m) = Eq -> (p ?= q) = Eq) ->
- (n > m -> p > q) -> (n ?= m) = (p ?= q).
-Proof.
- intros x1 y1 x2 y2.
- unfold Zgt in |- *; unfold Zlt in |- *; case (x1 ?= y1); case (x2 ?= y2);
- auto with arith; symmetry in |- *; auto with arith.
-Qed.
-
-(** * Relating [x ?= y] to [Zle], [Zlt], [Zge] or [Zgt] *)
-
Lemma Zle_compare :
forall n m:Z,
n <= m -> match n ?= m with
@@ -464,7 +143,7 @@ Lemma Zle_compare :
| Gt => False
end.
Proof.
- intros x y; unfold Zle in |- *; elim (x ?= y); auto with arith.
+ intros. case Z.compare_spec; trivial; Z.order.
Qed.
Lemma Zlt_compare :
@@ -475,8 +154,7 @@ Lemma Zlt_compare :
| Gt => False
end.
Proof.
- intros x y; unfold Zlt in |- *; elim (x ?= y); intros;
- discriminate || trivial with arith.
+ intros x y H; now rewrite H.
Qed.
Lemma Zge_compare :
@@ -487,7 +165,7 @@ Lemma Zge_compare :
| Gt => True
end.
Proof.
- intros x y; unfold Zge in |- *; elim (x ?= y); auto with arith.
+ intros. now case Z.compare_spec.
Qed.
Lemma Zgt_compare :
@@ -498,26 +176,23 @@ Lemma Zgt_compare :
| Gt => True
end.
Proof.
- intros x y; unfold Zgt in |- *; elim (x ?= y); intros;
- discriminate || trivial with arith.
+ intros x y H; now rewrite H.
Qed.
-(*********************)
-(** * Other properties *)
+(** Compatibility notations *)
-Lemma Zmult_compare_compat_l :
- forall n m p:Z, p > 0 -> (n ?= m) = (p * n ?= p * m).
-Proof.
- intros x y z H; destruct z.
- discriminate H.
- rewrite Zcompare_mult_compat; reflexivity.
- discriminate H.
-Qed.
-
-Lemma Zmult_compare_compat_r :
- forall n m p:Z, p > 0 -> (n ?= m) = (n * p ?= m * p).
-Proof.
- intros x y z H; rewrite (Zmult_comm x z); rewrite (Zmult_comm y z);
- apply Zmult_compare_compat_l; assumption.
-Qed.
+Notation Zcompare_refl := Z.compare_refl (only parsing).
+Notation Zcompare_Eq_eq := Z.compare_eq (only parsing).
+Notation Zcompare_Eq_iff_eq := Z.compare_eq_iff (only parsing).
+Notation Zcompare_spec := Z.compare_spec (only parsing).
+Notation Zmin_l := Z.min_l (only parsing).
+Notation Zmin_r := Z.min_r (only parsing).
+Notation Zmax_l := Z.max_l (only parsing).
+Notation Zmax_r := Z.max_r (only parsing).
+Notation Zabs_eq := Z.abs_eq (only parsing).
+Notation Zabs_non_eq := Z.abs_neq (only parsing).
+Notation Zsgn_0 := Z.sgn_null (only parsing).
+Notation Zsgn_1 := Z.sgn_pos (only parsing).
+Notation Zsgn_m1 := Z.sgn_neg (only parsing).
+(** Not kept: Zcompare_egal_dec *)
diff --git a/theories/ZArith/Zcomplements.v b/theories/ZArith/Zcomplements.v
index ca72f8a8..5a2c3cc3 100644
--- a/theories/ZArith/Zcomplements.v
+++ b/theories/ZArith/Zcomplements.v
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zcomplements.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import ZArithRing.
Require Import ZArith_base.
Require Export Omega.
@@ -18,29 +16,7 @@ Open Local Scope Z_scope.
(**********************************************************************)
(** About parity *)
-Lemma two_or_two_plus_one :
- forall n:Z, {y : Z | n = 2 * y} + {y : Z | n = 2 * y + 1}.
-Proof.
- intro x; destruct x.
- left; split with 0; reflexivity.
-
- destruct p.
- right; split with (Zpos p); reflexivity.
-
- left; split with (Zpos p); reflexivity.
-
- right; split with 0; reflexivity.
-
- destruct p.
- right; split with (Zneg (1 + p)).
- rewrite BinInt.Zneg_xI.
- rewrite BinInt.Zneg_plus_distr.
- omega.
-
- left; split with (Zneg p); reflexivity.
-
- right; split with (-1); reflexivity.
-Qed.
+Notation two_or_two_plus_one := Z_modulo_2 (only parsing).
(**********************************************************************)
(** The biggest power of 2 that is stricly less than [a]
@@ -58,31 +34,14 @@ Fixpoint floor_pos (a:positive) : positive :=
Definition floor (a:positive) := Zpos (floor_pos a).
Lemma floor_gt0 : forall p:positive, floor p > 0.
-Proof.
- intro.
- compute in |- *.
- trivial.
-Qed.
+Proof. reflexivity. Qed.
Lemma floor_ok : forall p:positive, floor p <= Zpos p < 2 * floor p.
Proof.
- unfold floor in |- *.
- intro a; induction a as [p| p| ].
-
- simpl in |- *.
- repeat rewrite BinInt.Zpos_xI.
- rewrite (BinInt.Zpos_xO (xO (floor_pos p))).
- rewrite (BinInt.Zpos_xO (floor_pos p)).
- omega.
-
- simpl in |- *.
- repeat rewrite BinInt.Zpos_xI.
- rewrite (BinInt.Zpos_xO (xO (floor_pos p))).
- rewrite (BinInt.Zpos_xO (floor_pos p)).
- rewrite (BinInt.Zpos_xO p).
- omega.
-
- simpl in |- *; omega.
+ unfold floor. induction p; simpl.
+ - rewrite !Z.pos_xI, (Z.pos_xO (xO _)), Z.pos_xO. omega.
+ - rewrite (Z.pos_xO (xO _)), (Z.pos_xO p), Z.pos_xO. omega.
+ - omega.
Qed.
(**********************************************************************)
@@ -90,41 +49,39 @@ Qed.
Theorem Z_lt_abs_rec :
forall P:Z -> Set,
- (forall n:Z, (forall m:Z, Zabs m < Zabs n -> P m) -> P n) ->
+ (forall n:Z, (forall m:Z, Z.abs m < Z.abs n -> P m) -> P n) ->
forall n:Z, P n.
Proof.
intros P HP p.
set (Q := fun z => 0 <= z -> P z * P (- z)) in *.
- cut (Q (Zabs p)); [ intros | apply (Z_lt_rec Q); auto with zarith ].
+ cut (Q (Z.abs p)); [ intros | apply (Z_lt_rec Q); auto with zarith ].
elim (Zabs_dec p); intro eq; rewrite eq; elim H; auto with zarith.
unfold Q in |- *; clear Q; intros.
- apply pair; apply HP.
- rewrite Zabs_eq; auto; intros.
- elim (H (Zabs m)); intros; auto with zarith.
+ split; apply HP.
+ rewrite Z.abs_eq; auto; intros.
+ elim (H (Z.abs m)); intros; auto with zarith.
elim (Zabs_dec m); intro eq; rewrite eq; trivial.
- rewrite Zabs_non_eq; auto with zarith.
- rewrite Zopp_involutive; intros.
- elim (H (Zabs m)); intros; auto with zarith.
+ rewrite Z.abs_neq, Z.opp_involutive; auto with zarith; intros.
+ elim (H (Z.abs m)); intros; auto with zarith.
elim (Zabs_dec m); intro eq; rewrite eq; trivial.
Qed.
Theorem Z_lt_abs_induction :
forall P:Z -> Prop,
- (forall n:Z, (forall m:Z, Zabs m < Zabs n -> P m) -> P n) ->
+ (forall n:Z, (forall m:Z, Z.abs m < Z.abs n -> P m) -> P n) ->
forall n:Z, P n.
Proof.
intros P HP p.
set (Q := fun z => 0 <= z -> P z /\ P (- z)) in *.
- cut (Q (Zabs p)); [ intros | apply (Z_lt_induction Q); auto with zarith ].
+ cut (Q (Z.abs p)); [ intros | apply (Z_lt_induction Q); auto with zarith ].
elim (Zabs_dec p); intro eq; rewrite eq; elim H; auto with zarith.
unfold Q in |- *; clear Q; intros.
split; apply HP.
- rewrite Zabs_eq; auto; intros.
- elim (H (Zabs m)); intros; auto with zarith.
+ rewrite Z.abs_eq; auto; intros.
+ elim (H (Z.abs m)); intros; auto with zarith.
elim (Zabs_dec m); intro eq; rewrite eq; trivial.
- rewrite Zabs_non_eq; auto with zarith.
- rewrite Zopp_involutive; intros.
- elim (H (Zabs m)); intros; auto with zarith.
+ rewrite Z.abs_neq, Z.opp_involutive; auto with zarith; intros.
+ elim (H (Z.abs m)); intros; auto with zarith.
elim (Zabs_dec m); intro eq; rewrite eq; trivial.
Qed.
@@ -134,25 +91,12 @@ Lemma Zcase_sign :
forall (n:Z) (P:Prop), (n = 0 -> P) -> (n > 0 -> P) -> (n < 0 -> P) -> P.
Proof.
intros x P Hzero Hpos Hneg.
- induction x as [| p| p].
- apply Hzero; trivial.
- apply Hpos; apply Zorder.Zgt_pos_0.
- apply Hneg; apply Zorder.Zlt_neg_0.
+ destruct x; [apply Hzero|apply Hpos|apply Hneg]; easy.
Qed.
-Lemma sqr_pos : forall n:Z, n * n >= 0.
+Lemma sqr_pos n : n * n >= 0.
Proof.
- intro x.
- apply (Zcase_sign x (x * x >= 0)).
- intros H; rewrite H; omega.
- intros H; replace 0 with (0 * 0).
- apply Zmult_ge_compat; omega.
- omega.
- intros H; replace 0 with (0 * 0).
- replace (x * x) with (- x * - x).
- apply Zmult_ge_compat; omega.
- ring.
- omega.
+ Z.swap_greater. apply Z.square_nonneg.
Qed.
(**********************************************************************)
@@ -167,7 +111,7 @@ Fixpoint Zlength_aux (acc:Z) (A:Type) (l:list A) : Z :=
end.
Definition Zlength := Zlength_aux 0.
-Implicit Arguments Zlength [A].
+Arguments Zlength [A] l.
Section Zlength_properties.
@@ -175,38 +119,33 @@ Section Zlength_properties.
Implicit Type l : list A.
- Lemma Zlength_correct : forall l, Zlength l = Z_of_nat (length l).
+ Lemma Zlength_correct l : Zlength l = Z.of_nat (length l).
Proof.
- assert (forall l (acc:Z), Zlength_aux acc A l = acc + Z_of_nat (length l)).
- simple induction l.
- simpl in |- *; auto with zarith.
- intros; simpl (length (a :: l0)) in |- *; rewrite Znat.inj_S.
- simpl in |- *; rewrite H; auto with zarith.
- unfold Zlength in |- *; intros; rewrite H; auto.
+ assert (H : forall l acc, Zlength_aux acc A l = acc + Z.of_nat (length l)).
+ clear l. induction l.
+ auto with zarith.
+ intros. simpl length; simpl Zlength_aux.
+ rewrite IHl, Nat2Z.inj_succ; auto with zarith.
+ unfold Zlength. now rewrite H.
Qed.
Lemma Zlength_nil : Zlength (A:=A) nil = 0.
- Proof.
- auto.
- Qed.
+ Proof. reflexivity. Qed.
- Lemma Zlength_cons : forall (x:A) l, Zlength (x :: l) = Zsucc (Zlength l).
+ Lemma Zlength_cons (x:A) l : Zlength (x :: l) = Z.succ (Zlength l).
Proof.
- intros; do 2 rewrite Zlength_correct.
- simpl (length (x :: l)) in |- *; rewrite Znat.inj_S; auto.
+ intros. now rewrite !Zlength_correct, <- Nat2Z.inj_succ.
Qed.
- Lemma Zlength_nil_inv : forall l, Zlength l = 0 -> l = nil.
+ Lemma Zlength_nil_inv l : Zlength l = 0 -> l = nil.
Proof.
- intro l; rewrite Zlength_correct.
- case l; auto.
- intros x l'; simpl (length (x :: l')) in |- *.
- rewrite Znat.inj_S.
- intros; exfalso; generalize (Zle_0_nat (length l')); omega.
+ rewrite Zlength_correct.
+ destruct l as [|x l]; auto.
+ now rewrite <- Nat2Z.inj_0, Nat2Z.inj_iff.
Qed.
End Zlength_properties.
-Implicit Arguments Zlength_correct [A].
-Implicit Arguments Zlength_cons [A].
-Implicit Arguments Zlength_nil_inv [A].
+Arguments Zlength_correct [A] l.
+Arguments Zlength_cons [A] x l.
+Arguments Zlength_nil_inv [A] l _.
diff --git a/theories/ZArith/Zdigits.v b/theories/ZArith/Zdigits.v
index c43b241d..ff1d96df 100644
--- a/theories/ZArith/Zdigits.v
+++ b/theories/ZArith/Zdigits.v
@@ -1,14 +1,12 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zdigits.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(** Bit vectors interpreted as integers.
Contribution by Jean Duprat (ENS Lyon). *)
@@ -47,17 +45,17 @@ Section VALUE_OF_BOOLEAN_VECTORS.
exact 0%Z.
inversion H0.
- exact (bit_value a + 2 * H H2)%Z.
+ exact (bit_value h + 2 * H H2)%Z.
Defined.
Lemma two_compl_value : forall n:nat, Bvector (S n) -> Z.
Proof.
simple induction n; intros.
inversion H.
- exact (- bit_value a)%Z.
+ exact (- bit_value h)%Z.
inversion H0.
- exact (bit_value a + 2 * H H2)%Z.
+ exact (bit_value h + 2 * H H2)%Z.
Defined.
End VALUE_OF_BOOLEAN_VECTORS.
@@ -136,7 +134,7 @@ Section Z_BRIC_A_BRAC.
Lemma binary_value_Sn :
forall (n:nat) (b:bool) (bv:Bvector n),
- binary_value (S n) (Vcons bool b n bv) =
+ binary_value (S n) ( b :: bv) =
(bit_value b + 2 * binary_value n bv)%Z.
Proof.
intros; auto.
@@ -221,7 +219,7 @@ Section Z_BRIC_A_BRAC.
destruct (Zeven.Zeven_odd_dec z); intros.
rewrite <- Zeven.Zeven_div2; auto.
- generalize (Zeven.Zodd_div2 z H z0); omega.
+ generalize (Zeven.Zodd_div2 z z0); omega.
Qed.
Lemma Z_to_two_compl_Sn_z :
diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v
index df22371e..314f696a 100644
--- a/theories/ZArith/Zdiv.v
+++ b/theories/ZArith/Zdiv.v
@@ -1,200 +1,57 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zdiv.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
+(** * Euclidean Division *)
-(* Contribution by Claude Marché and Xavier Urbain *)
-
-(** Euclidean Division
-
- Defines first of function that allows Coq to normalize.
- Then only after proves the main required property.
-*)
+(** Initial Contribution by Claude Marché and Xavier Urbain *)
Require Export ZArith_base.
-Require Import Zbool.
-Require Import Omega.
-Require Import ZArithRing.
-Require Import Zcomplements.
-Require Export Setoid.
-Open Local Scope Z_scope.
-
-(** * Definitions of Euclidian operations *)
-
-(** Euclidean division of a positive by a integer
- (that is supposed to be positive).
-
- Total function than returns an arbitrary value when
- divisor is not positive
-
-*)
-
-Unboxed Fixpoint Zdiv_eucl_POS (a:positive) (b:Z) : Z * Z :=
- match a with
- | xH => if Zge_bool b 2 then (0, 1) else (1, 0)
- | xO a' =>
- let (q, r) := Zdiv_eucl_POS a' b in
- let r' := 2 * r in
- if Zgt_bool b r' then (2 * q, r') else (2 * q + 1, r' - b)
- | xI a' =>
- let (q, r) := Zdiv_eucl_POS a' b in
- let r' := 2 * r + 1 in
- if Zgt_bool b r' then (2 * q, r') else (2 * q + 1, r' - b)
- end.
-
-
-(** Euclidean division of integers.
-
- Total function than returns (0,0) when dividing by 0.
-*)
-
-(**
-
- The pseudo-code is:
-
- if b = 0 : (0,0)
-
- if b <> 0 and a = 0 : (0,0)
-
- if b > 0 and a < 0 : let (q,r) = div_eucl_pos (-a) b in
- if r = 0 then (-q,0) else (-(q+1),b-r)
-
- if b < 0 and a < 0 : let (q,r) = div_eucl (-a) (-b) in (q,-r)
-
- if b < 0 and a > 0 : let (q,r) = div_eucl a (-b) in
- if r = 0 then (-q,0) else (-(q+1),b+r)
-
- In other word, when b is non-zero, q is chosen to be the greatest integer
- smaller or equal to a/b. And sgn(r)=sgn(b) and |r| < |b| (at least when
- r is not null).
-*)
+Require Import Zbool Omega ZArithRing Zcomplements Setoid Morphisms.
+Local Open Scope Z_scope.
-(* Nota: At least two others conventions also exist for euclidean division.
- They all satify the equation a=b*q+r, but differ on the choice of (q,r)
- on negative numbers.
+(** The definition of the division is now in [BinIntDef], the initial
+ specifications and properties are in [BinInt]. *)
- * Ocaml uses Round-Toward-Zero division: (-a)/b = a/(-b) = -(a/b).
- Hence (-a) mod b = - (a mod b)
- a mod (-b) = a mod b
- And: |r| < |b| and sgn(r) = sgn(a) (notice the a here instead of b).
+Notation Zdiv_eucl_POS := Z.pos_div_eucl (only parsing).
+Notation Zdiv_eucl := Z.div_eucl (only parsing).
+Notation Zdiv := Z.div (only parsing).
+Notation Zmod := Z.modulo (only parsing).
- * Another solution is to always pick a non-negative remainder:
- a=b*q+r with 0 <= r < |b|
-*)
-
-Definition Zdiv_eucl (a b:Z) : Z * Z :=
- match a, b with
- | Z0, _ => (0, 0)
- | _, Z0 => (0, 0)
- | Zpos a', Zpos _ => Zdiv_eucl_POS a' b
- | Zneg a', Zpos _ =>
- let (q, r) := Zdiv_eucl_POS a' b in
- match r with
- | Z0 => (- q, 0)
- | _ => (- (q + 1), b - r)
- end
- | Zneg a', Zneg b' => let (q, r) := Zdiv_eucl_POS a' (Zpos b') in (q, - r)
- | Zpos a', Zneg b' =>
- let (q, r) := Zdiv_eucl_POS a' (Zpos b') in
- match r with
- | Z0 => (- q, 0)
- | _ => (- (q + 1), b + r)
- end
- end.
-
-
-(** Division and modulo are projections of [Zdiv_eucl] *)
-
-Definition Zdiv (a b:Z) : Z := let (q, _) := Zdiv_eucl a b in q.
-
-Definition Zmod (a b:Z) : Z := let (_, r) := Zdiv_eucl a b in r.
-
-(** Syntax *)
-
-Infix "/" := Zdiv : Z_scope.
-Infix "mod" := Zmod (at level 40, no associativity) : Z_scope.
-
-(* Tests:
-
-Eval compute in (Zdiv_eucl 7 3).
-
-Eval compute in (Zdiv_eucl (-7) 3).
+Notation Zdiv_eucl_eq := Z.div_eucl_eq (only parsing).
+Notation Z_div_mod_eq_full := Z.div_mod (only parsing).
+Notation Zmod_POS_bound := Z.pos_div_eucl_bound (only parsing).
+Notation Zmod_pos_bound := Z.mod_pos_bound (only parsing).
+Notation Zmod_neg_bound := Z.mod_neg_bound (only parsing).
-Eval compute in (Zdiv_eucl 7 (-3)).
+(** * Main division theorems *)
-Eval compute in (Zdiv_eucl (-7) (-3)).
-
-*)
-
-
-(** * Main division theorem *)
-
-(** First a lemma for two positive arguments *)
+(** NB: many things are stated twice for compatibility reasons *)
Lemma Z_div_mod_POS :
forall b:Z,
b > 0 ->
forall a:positive,
- let (q, r) := Zdiv_eucl_POS a b in Zpos a = b * q + r /\ 0 <= r < b.
+ let (q, r) := Z.pos_div_eucl a b in Zpos a = b * q + r /\ 0 <= r < b.
Proof.
-simple induction a; cbv beta iota delta [Zdiv_eucl_POS] in |- *;
- fold Zdiv_eucl_POS in |- *; cbv zeta.
-
-intro p; case (Zdiv_eucl_POS p b); intros q r [H0 H1].
-generalize (Zgt_cases b (2 * r + 1)).
-case (Zgt_bool b (2 * r + 1));
- (rewrite BinInt.Zpos_xI; rewrite H0; split; [ ring | omega ]).
-
-intros p; case (Zdiv_eucl_POS p b); intros q r [H0 H1].
-generalize (Zgt_cases b (2 * r)).
-case (Zgt_bool b (2 * r)); rewrite BinInt.Zpos_xO;
- change (Zpos (xO p)) with (2 * Zpos p) in |- *; rewrite H0;
- (split; [ ring | omega ]).
-
-generalize (Zge_cases b 2).
-case (Zge_bool b 2); (intros; split; [ try ring | omega ]).
-omega.
+ intros b Hb a. Z.swap_greater.
+ generalize (Z.pos_div_eucl_eq a b Hb) (Z.pos_div_eucl_bound a b Hb).
+ destruct Z.pos_div_eucl. rewrite Z.mul_comm. auto.
Qed.
-(** Then the usual situation of a positive [b] and no restriction on [a] *)
-
-Theorem Z_div_mod :
- forall a b:Z,
- b > 0 -> let (q, r) := Zdiv_eucl a b in a = b * q + r /\ 0 <= r < b.
+Theorem Z_div_mod a b :
+ b > 0 ->
+ let (q, r) := Z.div_eucl a b in a = b * q + r /\ 0 <= r < b.
Proof.
- intros a b; case a; case b; try (simpl in |- *; intros; omega).
- unfold Zdiv_eucl in |- *; intros; apply Z_div_mod_POS; trivial.
-
- intros; discriminate.
-
- intros.
- generalize (Z_div_mod_POS (Zpos p) H p0).
- unfold Zdiv_eucl in |- *.
- case (Zdiv_eucl_POS p0 (Zpos p)).
- intros z z0.
- case z0.
-
- intros [H1 H2].
- split; trivial.
- change (Zneg p0) with (- Zpos p0); rewrite H1; ring.
-
- intros p1 [H1 H2].
- split; trivial.
- change (Zneg p0) with (- Zpos p0); rewrite H1; ring.
- generalize (Zorder.Zgt_pos_0 p1); omega.
-
- intros p1 [H1 H2].
- split; trivial.
- change (Zneg p0) with (- Zpos p0); rewrite H1; ring.
- generalize (Zorder.Zlt_neg_0 p1); omega.
-
- intros; discriminate.
+ Z.swap_greater. intros Hb.
+ assert (Hb' : b<>0) by (now destruct b).
+ generalize (Z.div_eucl_eq a b Hb') (Z.mod_pos_bound a b Hb).
+ unfold Z.modulo. destruct Z.div_eucl. auto.
Qed.
(** For stating the fully general result, let's give a short name
@@ -204,7 +61,7 @@ Definition Remainder r b := 0 <= r < b \/ b < r <= 0.
(** Another equivalent formulation: *)
-Definition Remainder_alt r b := Zabs r < Zabs b /\ Zsgn r <> - Zsgn b.
+Definition Remainder_alt r b := Z.abs r < Z.abs b /\ Z.sgn r <> - Z.sgn b.
(* In the last formulation, [ Zsgn r <> - Zsgn b ] is less nice than saying
[ Zsgn r = Zsgn b ], but at least it works even when [r] is null. *)
@@ -218,90 +75,44 @@ Hint Unfold Remainder.
(** Now comes the fully general result about Euclidean division. *)
-Theorem Z_div_mod_full :
- forall a b:Z,
- b <> 0 -> let (q, r) := Zdiv_eucl a b in a = b * q + r /\ Remainder r b.
+Theorem Z_div_mod_full a b :
+ b <> 0 ->
+ let (q, r) := Z.div_eucl a b in a = b * q + r /\ Remainder r b.
Proof.
- destruct b as [|b|b].
- (* b = 0 *)
- intro H; elim H; auto.
- (* b > 0 *)
- intros _.
- assert (Zpos b > 0) by auto with zarith.
- generalize (Z_div_mod a (Zpos b) H).
- destruct Zdiv_eucl as (q,r); intuition; simpl; auto.
- (* b < 0 *)
- intros _.
- assert (Zpos b > 0) by auto with zarith.
- generalize (Z_div_mod a (Zpos b) H).
- unfold Remainder.
- destruct a as [|a|a].
- (* a = 0 *)
- simpl; intuition.
- (* a > 0 *)
- unfold Zdiv_eucl; destruct Zdiv_eucl_POS as (q,r).
- destruct r as [|r|r]; [ | | omega with *].
- rewrite <- Zmult_opp_comm; simpl Zopp; intuition.
- rewrite <- Zmult_opp_comm; simpl Zopp.
- rewrite Zmult_plus_distr_r; omega with *.
- (* a < 0 *)
- unfold Zdiv_eucl.
- generalize (Z_div_mod_POS (Zpos b) H a).
- destruct Zdiv_eucl_POS as (q,r).
- destruct r as [|r|r]; change (Zneg b) with (-Zpos b).
- rewrite Zmult_opp_comm; omega with *.
- rewrite <- Zmult_opp_comm, Zmult_plus_distr_r;
- repeat rewrite Zmult_opp_comm; omega.
- rewrite Zmult_opp_comm; omega with *.
+ intros Hb.
+ generalize (Z.div_eucl_eq a b Hb)
+ (Z.mod_pos_bound a b) (Z.mod_neg_bound a b).
+ unfold Z.modulo. destruct Z.div_eucl as (q,r).
+ intros EQ POS NEG.
+ split; auto.
+ red; destruct b.
+ now destruct Hb. left; now apply POS. right; now apply NEG.
Qed.
(** The same results as before, stated separately in terms of Zdiv and Zmod *)
-Lemma Z_mod_remainder : forall a b:Z, b<>0 -> Remainder (a mod b) b.
-Proof.
- unfold Zmod; intros a b Hb; generalize (Z_div_mod_full a b Hb); auto.
- destruct Zdiv_eucl; tauto.
-Qed.
-
-Lemma Z_mod_lt : forall a b:Z, b > 0 -> 0 <= a mod b < b.
+Lemma Z_mod_remainder a b : b<>0 -> Remainder (a mod b) b.
Proof.
- unfold Zmod; intros a b Hb; generalize (Z_div_mod a b Hb).
- destruct Zdiv_eucl; tauto.
+ unfold Z.modulo; intros Hb; generalize (Z_div_mod_full a b Hb); auto.
+ destruct Z.div_eucl; tauto.
Qed.
-Lemma Z_mod_neg : forall a b:Z, b < 0 -> b < a mod b <= 0.
-Proof.
- unfold Zmod; intros a b Hb.
- assert (Hb' : b<>0) by (auto with zarith).
- generalize (Z_div_mod_full a b Hb').
- destruct Zdiv_eucl.
- unfold Remainder; intuition.
-Qed.
+Lemma Z_mod_lt a b : b > 0 -> 0 <= a mod b < b.
+Proof (fun Hb => Z.mod_pos_bound a b (Zgt_lt _ _ Hb)).
-Lemma Z_div_mod_eq_full : forall a b:Z, b <> 0 -> a = b*(a/b) + (a mod b).
-Proof.
- unfold Zdiv, Zmod; intros a b Hb; generalize (Z_div_mod_full a b Hb).
- destruct Zdiv_eucl; tauto.
-Qed.
+Lemma Z_mod_neg a b : b < 0 -> b < a mod b <= 0.
+Proof (Z.mod_neg_bound a b).
-Lemma Z_div_mod_eq : forall a b:Z, b > 0 -> a = b*(a/b) + (a mod b).
+Lemma Z_div_mod_eq a b : b > 0 -> a = b*(a/b) + (a mod b).
Proof.
- intros; apply Z_div_mod_eq_full; auto with zarith.
+ intros Hb; apply Z.div_mod; auto with zarith.
Qed.
-Lemma Zmod_eq_full : forall a b:Z, b<>0 -> a mod b = a - (a/b)*b.
-Proof.
- intros.
- rewrite <- Zeq_plus_swap, Zplus_comm, Zmult_comm; symmetry.
- apply Z_div_mod_eq_full; auto.
-Qed.
+Lemma Zmod_eq_full a b : b<>0 -> a mod b = a - (a/b)*b.
+Proof. intros. rewrite Z.mul_comm. now apply Z.mod_eq. Qed.
-Lemma Zmod_eq : forall a b:Z, b>0 -> a mod b = a - (a/b)*b.
-Proof.
- intros.
- rewrite <- Zeq_plus_swap, Zplus_comm, Zmult_comm; symmetry.
- apply Z_div_mod_eq; auto.
-Qed.
+Lemma Zmod_eq a b : b>0 -> a mod b = a - (a/b)*b.
+Proof. intros. apply Zmod_eq_full. now destruct b. Qed.
(** Existence theorem *)
@@ -309,89 +120,51 @@ Theorem Zdiv_eucl_exist : forall (b:Z)(Hb:b>0)(a:Z),
{qr : Z * Z | let (q, r) := qr in a = b * q + r /\ 0 <= r < b}.
Proof.
intros b Hb a.
- exists (Zdiv_eucl a b).
+ exists (Z.div_eucl a b).
exact (Z_div_mod a b Hb).
Qed.
-Implicit Arguments Zdiv_eucl_exist.
+Arguments Zdiv_eucl_exist : default implicits.
(** Uniqueness theorems *)
-Theorem Zdiv_mod_unique :
- forall b q1 q2 r1 r2:Z,
- 0 <= r1 < Zabs b -> 0 <= r2 < Zabs b ->
+Theorem Zdiv_mod_unique b q1 q2 r1 r2 :
+ 0 <= r1 < Z.abs b -> 0 <= r2 < Z.abs b ->
b*q1+r1 = b*q2+r2 -> q1=q2 /\ r1=r2.
Proof.
-intros b q1 q2 r1 r2 Hr1 Hr2 H.
-destruct (Z_eq_dec q1 q2) as [Hq|Hq].
+intros Hr1 Hr2 H. rewrite <- (Z.abs_sgn b), <- !Z.mul_assoc in H.
+destruct (Z.div_mod_unique (Z.abs b) (Z.sgn b * q1) (Z.sgn b * q2) r1 r2); auto.
split; trivial.
-rewrite Hq in H; omega.
-elim (Zlt_not_le (Zabs (r2 - r1)) (Zabs b)).
-omega with *.
-replace (r2-r1) with (b*(q1-q2)) by (rewrite Zmult_minus_distr_l; omega).
-replace (Zabs b) with ((Zabs b)*1) by ring.
-rewrite Zabs_Zmult.
-apply Zmult_le_compat_l; auto with *.
-omega with *.
+apply Z.mul_cancel_l with (Z.sgn b); trivial.
+rewrite Z.sgn_null_iff, <- Z.abs_0_iff. destruct Hr1; Z.order.
Qed.
Theorem Zdiv_mod_unique_2 :
forall b q1 q2 r1 r2:Z,
Remainder r1 b -> Remainder r2 b ->
b*q1+r1 = b*q2+r2 -> q1=q2 /\ r1=r2.
-Proof.
-unfold Remainder.
-intros b q1 q2 r1 r2 Hr1 Hr2 H.
-destruct (Z_eq_dec q1 q2) as [Hq|Hq].
-split; trivial.
-rewrite Hq in H; omega.
-elim (Zlt_not_le (Zabs (r2 - r1)) (Zabs b)).
-omega with *.
-replace (r2-r1) with (b*(q1-q2)) by (rewrite Zmult_minus_distr_l; omega).
-replace (Zabs b) with ((Zabs b)*1) by ring.
-rewrite Zabs_Zmult.
-apply Zmult_le_compat_l; auto with *.
-omega with *.
-Qed.
+Proof Z.div_mod_unique.
Theorem Zdiv_unique_full:
forall a b q r, Remainder r b ->
a = b*q + r -> q = a/b.
-Proof.
- intros.
- assert (b <> 0) by (unfold Remainder in *; omega with *).
- generalize (Z_div_mod_full a b H1).
- unfold Zdiv; destruct Zdiv_eucl as (q',r').
- intros (H2,H3); rewrite H2 in H0.
- destruct (Zdiv_mod_unique_2 b q q' r r'); auto.
-Qed.
+Proof Z.div_unique.
Theorem Zdiv_unique:
forall a b q r, 0 <= r < b ->
a = b*q + r -> q = a/b.
-Proof.
- intros; eapply Zdiv_unique_full; eauto.
-Qed.
+Proof. intros; eapply Zdiv_unique_full; eauto. Qed.
Theorem Zmod_unique_full:
forall a b q r, Remainder r b ->
a = b*q + r -> r = a mod b.
-Proof.
- intros.
- assert (b <> 0) by (unfold Remainder in *; omega with *).
- generalize (Z_div_mod_full a b H1).
- unfold Zmod; destruct Zdiv_eucl as (q',r').
- intros (H2,H3); rewrite H2 in H0.
- destruct (Zdiv_mod_unique_2 b q q' r r'); auto.
-Qed.
+Proof Z.mod_unique.
Theorem Zmod_unique:
forall a b q r, 0 <= r < b ->
a = b*q + r -> r = a mod b.
-Proof.
- intros; eapply Zmod_unique_full; eauto.
-Qed.
+Proof. intros; eapply Zmod_unique_full; eauto. Qed.
(** * Basic values of divisions and modulo. *)
@@ -415,70 +188,44 @@ Proof.
destruct a; simpl; auto.
Qed.
+Ltac zero_or_not a :=
+ destruct (Z.eq_dec a 0);
+ [subst; rewrite ?Zmod_0_l, ?Zdiv_0_l, ?Zmod_0_r, ?Zdiv_0_r;
+ auto with zarith|].
+
Lemma Zmod_1_r: forall a, a mod 1 = 0.
-Proof.
- intros; symmetry; apply Zmod_unique with a; auto with zarith.
-Qed.
+Proof. intros. zero_or_not a. apply Z.mod_1_r. Qed.
Lemma Zdiv_1_r: forall a, a/1 = a.
-Proof.
- intros; symmetry; apply Zdiv_unique with 0; auto with zarith.
-Qed.
+Proof. intros. zero_or_not a. apply Z.div_1_r. Qed.
Hint Resolve Zmod_0_l Zmod_0_r Zdiv_0_l Zdiv_0_r Zdiv_1_r Zmod_1_r
: zarith.
Lemma Zdiv_1_l: forall a, 1 < a -> 1/a = 0.
-Proof.
- intros; symmetry; apply Zdiv_unique with 1; auto with zarith.
-Qed.
+Proof Z.div_1_l.
Lemma Zmod_1_l: forall a, 1 < a -> 1 mod a = 1.
-Proof.
- intros; symmetry; apply Zmod_unique with 0; auto with zarith.
-Qed.
+Proof Z.mod_1_l.
Lemma Z_div_same_full : forall a:Z, a<>0 -> a/a = 1.
-Proof.
- intros; symmetry; apply Zdiv_unique_full with 0; auto with *; red; omega.
-Qed.
+Proof Z.div_same.
Lemma Z_mod_same_full : forall a, a mod a = 0.
-Proof.
- destruct a; intros; symmetry.
- compute; auto.
- apply Zmod_unique with 1; auto with *; omega with *.
- apply Zmod_unique_full with 1; auto with *; red; omega with *.
-Qed.
+Proof. intros. zero_or_not a. apply Z.mod_same; auto. Qed.
Lemma Z_mod_mult : forall a b, (a*b) mod b = 0.
-Proof.
- intros a b; destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst; simpl; rewrite Zmod_0_r; auto.
- symmetry; apply Zmod_unique_full with a; [ red; omega | ring ].
-Qed.
+Proof. intros. zero_or_not b. apply Z.mod_mul. auto. Qed.
Lemma Z_div_mult_full : forall a b:Z, b <> 0 -> (a*b)/b = a.
-Proof.
- intros; symmetry; apply Zdiv_unique_full with 0; auto with zarith;
- [ red; omega | ring].
-Qed.
+Proof Z.div_mul.
(** * Order results about Zmod and Zdiv *)
(* Division of positive numbers is positive. *)
Lemma Z_div_pos: forall a b, b > 0 -> 0 <= a -> 0 <= a/b.
-Proof.
- intros.
- rewrite (Z_div_mod_eq a b H) in H0.
- assert (H1:=Z_mod_lt a b H).
- destruct (Z_lt_le_dec (a/b) 0); auto.
- assert (b*(a/b) <= -b).
- replace (-b) with (b*-1); [ | ring].
- apply Zmult_le_compat_l; auto with zarith.
- omega.
-Qed.
+Proof. intros. apply Z.div_pos; auto with zarith. Qed.
Lemma Z_div_ge0: forall a b, b > 0 -> a >= 0 -> a/b >=0.
Proof.
@@ -489,366 +236,165 @@ Qed.
the division is strictly decreasing. *)
Lemma Z_div_lt : forall a b:Z, b >= 2 -> a > 0 -> a/b < a.
-Proof.
- intros. cut (b > 0); [ intro Hb | omega ].
- generalize (Z_div_mod a b Hb).
- cut (a >= 0); [ intro Ha | omega ].
- generalize (Z_div_ge0 a b Hb Ha).
- unfold Zdiv in |- *; case (Zdiv_eucl a b); intros q r H1 [H2 H3].
- cut (a >= 2 * q -> q < a); [ intro h; apply h; clear h | intros; omega ].
- apply Zge_trans with (b * q).
- omega.
- auto with zarith.
-Qed.
-
+Proof. intros. apply Z.div_lt; auto with zarith. Qed.
(** A division of a small number by a bigger one yields zero. *)
Theorem Zdiv_small: forall a b, 0 <= a < b -> a/b = 0.
-Proof.
- intros a b H; apply sym_equal; apply Zdiv_unique with a; auto with zarith.
-Qed.
+Proof Z.div_small.
(** Same situation, in term of modulo: *)
Theorem Zmod_small: forall a n, 0 <= a < n -> a mod n = a.
-Proof.
- intros a b H; apply sym_equal; apply Zmod_unique with 0; auto with zarith.
-Qed.
+Proof Z.mod_small.
(** [Zge] is compatible with a positive division. *)
Lemma Z_div_ge : forall a b c:Z, c > 0 -> a >= b -> a/c >= b/c.
-Proof.
- intros a b c cPos aGeb.
- generalize (Z_div_mod_eq a c cPos).
- generalize (Z_mod_lt a c cPos).
- generalize (Z_div_mod_eq b c cPos).
- generalize (Z_mod_lt b c cPos).
- intros.
- elim (Z_ge_lt_dec (a / c) (b / c)); trivial.
- intro.
- absurd (b - a >= 1).
- omega.
- replace (b-a) with (c * (b/c-a/c) + b mod c - a mod c) by
- (symmetry; pattern a at 1; rewrite H2; pattern b at 1; rewrite H0; ring).
- assert (c * (b / c - a / c) >= c * 1).
- apply Zmult_ge_compat_l.
- omega.
- omega.
- assert (c * 1 = c).
- ring.
- omega.
-Qed.
+Proof. intros. apply Zle_ge. apply Z.div_le_mono; auto with zarith. Qed.
(** Same, with [Zle]. *)
Lemma Z_div_le : forall a b c:Z, c > 0 -> a <= b -> a/c <= b/c.
-Proof.
- intros a b c H H0.
- apply Zge_le.
- apply Z_div_ge; auto with *.
-Qed.
+Proof. intros. apply Z.div_le_mono; auto with zarith. Qed.
(** With our choice of division, rounding of (a/b) is always done toward bottom: *)
Lemma Z_mult_div_ge : forall a b:Z, b > 0 -> b*(a/b) <= a.
-Proof.
- intros a b H; generalize (Z_div_mod_eq a b H) (Z_mod_lt a b H); omega.
-Qed.
+Proof. intros. apply Z.mul_div_le; auto with zarith. Qed.
Lemma Z_mult_div_ge_neg : forall a b:Z, b < 0 -> b*(a/b) >= a.
-Proof.
- intros a b H.
- generalize (Z_div_mod_eq_full a _ (Zlt_not_eq _ _ H)) (Z_mod_neg a _ H); omega.
-Qed.
+Proof. intros. apply Zle_ge. apply Z.mul_div_ge; auto with zarith. Qed.
(** The previous inequalities are exact iff the modulo is zero. *)
Lemma Z_div_exact_full_1 : forall a b:Z, a = b*(a/b) -> a mod b = 0.
-Proof.
- intros; destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst b; simpl in *; subst; auto.
- generalize (Z_div_mod_eq_full a b Hb); omega.
-Qed.
+Proof. intros a b. zero_or_not b. rewrite Z.div_exact; auto. Qed.
Lemma Z_div_exact_full_2 : forall a b:Z, b <> 0 -> a mod b = 0 -> a = b*(a/b).
-Proof.
- intros; generalize (Z_div_mod_eq_full a b H); omega.
-Qed.
+Proof. intros; rewrite Z.div_exact; auto. Qed.
(** A modulo cannot grow beyond its starting point. *)
Theorem Zmod_le: forall a b, 0 < b -> 0 <= a -> a mod b <= a.
-Proof.
- intros a b H1 H2; case (Zle_or_lt b a); intros H3.
- case (Z_mod_lt a b); auto with zarith.
- rewrite Zmod_small; auto with zarith.
-Qed.
+Proof. intros. apply Z.mod_le; auto. Qed.
(** Some additionnal inequalities about Zdiv. *)
Theorem Zdiv_lt_upper_bound:
forall a b q, 0 < b -> a < q*b -> a/b < q.
-Proof.
- intros a b q H1 H2.
- apply Zmult_lt_reg_r with b; auto with zarith.
- apply Zle_lt_trans with (2 := H2).
- pattern a at 2; rewrite (Z_div_mod_eq a b); auto with zarith.
- rewrite (Zmult_comm b); case (Z_mod_lt a b); auto with zarith.
-Qed.
+Proof. intros a b q; rewrite Z.mul_comm; apply Z.div_lt_upper_bound. Qed.
Theorem Zdiv_le_upper_bound:
forall a b q, 0 < b -> a <= q*b -> a/b <= q.
-Proof.
- intros.
- rewrite <- (Z_div_mult_full q b); auto with zarith.
- apply Z_div_le; auto with zarith.
-Qed.
+Proof. intros a b q; rewrite Z.mul_comm; apply Z.div_le_upper_bound. Qed.
Theorem Zdiv_le_lower_bound:
forall a b q, 0 < b -> q*b <= a -> q <= a/b.
-Proof.
- intros.
- rewrite <- (Z_div_mult_full q b); auto with zarith.
- apply Z_div_le; auto with zarith.
-Qed.
+Proof. intros a b q; rewrite Z.mul_comm; apply Z.div_le_lower_bound. Qed.
(** A division of respect opposite monotonicity for the divisor *)
Lemma Zdiv_le_compat_l: forall p q r, 0 <= p -> 0 < q < r ->
p / r <= p / q.
-Proof.
- intros p q r H H1.
- apply Zdiv_le_lower_bound; auto with zarith.
- rewrite Zmult_comm.
- pattern p at 2; rewrite (Z_div_mod_eq p r); auto with zarith.
- apply Zle_trans with (r * (p / r)); auto with zarith.
- apply Zmult_le_compat_r; auto with zarith.
- apply Zdiv_le_lower_bound; auto with zarith.
- case (Z_mod_lt p r); auto with zarith.
-Qed.
+Proof. intros; apply Z.div_le_compat_l; auto with zarith. Qed.
Theorem Zdiv_sgn: forall a b,
- 0 <= Zsgn (a/b) * Zsgn a * Zsgn b.
+ 0 <= Z.sgn (a/b) * Z.sgn a * Z.sgn b.
Proof.
destruct a as [ |a|a]; destruct b as [ |b|b]; simpl; auto with zarith;
- generalize (Z_div_pos (Zpos a) (Zpos b)); unfold Zdiv, Zdiv_eucl;
- destruct Zdiv_eucl_POS as (q,r); destruct r; omega with *.
+ generalize (Z.div_pos (Zpos a) (Zpos b)); unfold Z.div, Z.div_eucl;
+ destruct Z.pos_div_eucl as (q,r); destruct r; omega with *.
Qed.
(** * Relations between usual operations and Zmod and Zdiv *)
Lemma Z_mod_plus_full : forall a b c:Z, (a + b * c) mod c = a mod c.
-Proof.
- intros; destruct (Z_eq_dec c 0) as [Hc|Hc].
- subst; do 2 rewrite Zmod_0_r; auto.
- symmetry; apply Zmod_unique_full with (a/c+b); auto with zarith.
- red; generalize (Z_mod_lt a c)(Z_mod_neg a c); omega.
- rewrite Zmult_plus_distr_r, Zmult_comm.
- generalize (Z_div_mod_eq_full a c Hc); omega.
-Qed.
+Proof. intros. zero_or_not c. apply Z.mod_add; auto. Qed.
Lemma Z_div_plus_full : forall a b c:Z, c <> 0 -> (a + b * c) / c = a / c + b.
-Proof.
- intro; symmetry.
- apply Zdiv_unique_full with (a mod c); auto with zarith.
- red; generalize (Z_mod_lt a c)(Z_mod_neg a c); omega.
- rewrite Zmult_plus_distr_r, Zmult_comm.
- generalize (Z_div_mod_eq_full a c H); omega.
-Qed.
+Proof Z.div_add.
Theorem Z_div_plus_full_l: forall a b c : Z, b <> 0 -> (a * b + c) / b = a + c / b.
-Proof.
- intros a b c H; rewrite Zplus_comm; rewrite Z_div_plus_full;
- try apply Zplus_comm; auto with zarith.
-Qed.
+Proof Z.div_add_l.
(** [Zopp] and [Zdiv], [Zmod].
Due to the choice of convention for our Euclidean division,
some of the relations about [Zopp] and divisions are rather complex. *)
Lemma Zdiv_opp_opp : forall a b:Z, (-a)/(-b) = a/b.
-Proof.
- intros [|a|a] [|b|b]; try reflexivity; unfold Zdiv; simpl;
- destruct (Zdiv_eucl_POS a (Zpos b)); destruct z0; try reflexivity.
-Qed.
+Proof. intros. zero_or_not b. apply Z.div_opp_opp; auto. Qed.
Lemma Zmod_opp_opp : forall a b:Z, (-a) mod (-b) = - (a mod b).
-Proof.
- intros; destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst; do 2 rewrite Zmod_0_r; auto.
- intros; symmetry.
- apply Zmod_unique_full with ((-a)/(-b)); auto.
- generalize (Z_mod_remainder a b Hb); destruct 1; [right|left]; omega.
- rewrite Zdiv_opp_opp.
- pattern a at 1; rewrite (Z_div_mod_eq_full a b Hb); ring.
-Qed.
+Proof. intros. zero_or_not b. apply Z.mod_opp_opp; auto. Qed.
Lemma Z_mod_zero_opp_full : forall a b:Z, a mod b = 0 -> (-a) mod b = 0.
-Proof.
- intros; destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst; rewrite Zmod_0_r; auto.
- rewrite Z_div_exact_full_2 with a b; auto.
- replace (- (b * (a / b))) with (0 + - (a / b) * b).
- rewrite Z_mod_plus_full; auto.
- ring.
-Qed.
+Proof. intros. zero_or_not b. apply Z.mod_opp_l_z; auto. Qed.
Lemma Z_mod_nz_opp_full : forall a b:Z, a mod b <> 0 ->
(-a) mod b = b - (a mod b).
-Proof.
- intros.
- assert (b<>0) by (contradict H; subst; rewrite Zmod_0_r; auto).
- symmetry; apply Zmod_unique_full with (-1-a/b); auto.
- generalize (Z_mod_remainder a b H0); destruct 1; [left|right]; omega.
- rewrite Zmult_minus_distr_l.
- pattern a at 1; rewrite (Z_div_mod_eq_full a b H0); ring.
-Qed.
+Proof. intros. zero_or_not b. apply Z.mod_opp_l_nz; auto. Qed.
Lemma Z_mod_zero_opp_r : forall a b:Z, a mod b = 0 -> a mod (-b) = 0.
-Proof.
- intros.
- rewrite <- (Zopp_involutive a).
- rewrite Zmod_opp_opp.
- rewrite Z_mod_zero_opp_full; auto.
-Qed.
+Proof. intros. zero_or_not b. apply Z.mod_opp_r_z; auto. Qed.
Lemma Z_mod_nz_opp_r : forall a b:Z, a mod b <> 0 ->
a mod (-b) = (a mod b) - b.
-Proof.
- intros.
- pattern a at 1; rewrite <- (Zopp_involutive a).
- rewrite Zmod_opp_opp.
- rewrite Z_mod_nz_opp_full; auto; omega.
-Qed.
+Proof. intros. zero_or_not b. apply Z.mod_opp_r_nz; auto. Qed.
Lemma Z_div_zero_opp_full : forall a b:Z, a mod b = 0 -> (-a)/b = -(a/b).
-Proof.
- intros; destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst; do 2 rewrite Zdiv_0_r; auto.
- symmetry; apply Zdiv_unique_full with 0; auto.
- red; omega.
- pattern a at 1; rewrite (Z_div_mod_eq_full a b Hb).
- rewrite H; ring.
-Qed.
+Proof. intros. zero_or_not b. apply Z.div_opp_l_z; auto. Qed.
Lemma Z_div_nz_opp_full : forall a b:Z, a mod b <> 0 ->
(-a)/b = -(a/b)-1.
-Proof.
- intros.
- assert (b<>0) by (contradict H; subst; rewrite Zmod_0_r; auto).
- symmetry; apply Zdiv_unique_full with (b-a mod b); auto.
- generalize (Z_mod_remainder a b H0); destruct 1; [left|right]; omega.
- pattern a at 1; rewrite (Z_div_mod_eq_full a b H0); ring.
-Qed.
+Proof. intros a b. zero_or_not b. intros; rewrite Z.div_opp_l_nz; auto. Qed.
Lemma Z_div_zero_opp_r : forall a b:Z, a mod b = 0 -> a/(-b) = -(a/b).
-Proof.
- intros.
- pattern a at 1; rewrite <- (Zopp_involutive a).
- rewrite Zdiv_opp_opp.
- rewrite Z_div_zero_opp_full; auto.
-Qed.
+Proof. intros. zero_or_not b. apply Z.div_opp_r_z; auto. Qed.
Lemma Z_div_nz_opp_r : forall a b:Z, a mod b <> 0 ->
a/(-b) = -(a/b)-1.
-Proof.
- intros.
- pattern a at 1; rewrite <- (Zopp_involutive a).
- rewrite Zdiv_opp_opp.
- rewrite Z_div_nz_opp_full; auto; omega.
-Qed.
+Proof. intros a b. zero_or_not b. intros; rewrite Z.div_opp_r_nz; auto. Qed.
(** Cancellations. *)
-Lemma Zdiv_mult_cancel_r : forall a b c:Z,
+Lemma Zdiv_mult_cancel_r : forall a b c:Z,
c <> 0 -> (a*c)/(b*c) = a/b.
-Proof.
-assert (X: forall a b c, b > 0 -> c > 0 -> (a*c) / (b*c) = a / b).
- intros a b c Hb Hc.
- symmetry.
- apply Zdiv_unique with ((a mod b)*c); auto with zarith.
- destruct (Z_mod_lt a b Hb); split.
- apply Zmult_le_0_compat; auto with zarith.
- apply Zmult_lt_compat_r; auto with zarith.
- pattern a at 1; rewrite (Z_div_mod_eq a b Hb); ring.
-intros a b c Hc.
-destruct (Z_dec b 0) as [Hb|Hb].
-destruct Hb as [Hb|Hb]; destruct (not_Zeq_inf _ _ Hc); auto with *.
-rewrite <- (Zdiv_opp_opp a), <- (Zmult_opp_opp b), <-(Zmult_opp_opp a);
- auto with *.
-rewrite <- (Zdiv_opp_opp a), <- Zdiv_opp_opp, Zopp_mult_distr_l,
- Zopp_mult_distr_l; auto with *.
-rewrite <- Zdiv_opp_opp, Zopp_mult_distr_r, Zopp_mult_distr_r; auto with *.
-rewrite Hb; simpl; do 2 rewrite Zdiv_0_r; auto.
-Qed.
+Proof. intros. zero_or_not b. apply Z.div_mul_cancel_r; auto. Qed.
Lemma Zdiv_mult_cancel_l : forall a b c:Z,
c<>0 -> (c*a)/(c*b) = a/b.
Proof.
- intros.
- rewrite (Zmult_comm c a); rewrite (Zmult_comm c b).
- apply Zdiv_mult_cancel_r; auto.
+ intros. rewrite (Zmult_comm c b); zero_or_not b.
+ rewrite (Zmult_comm b c). apply Z.div_mul_cancel_l; auto.
Qed.
Lemma Zmult_mod_distr_l: forall a b c,
(c*a) mod (c*b) = c * (a mod b).
Proof.
- intros; destruct (Z_eq_dec c 0) as [Hc|Hc].
- subst; simpl; rewrite Zmod_0_r; auto.
- destruct (Z_eq_dec b 0) as [Hb|Hb].
- subst; repeat rewrite Zmult_0_r || rewrite Zmod_0_r; auto.
- assert (c*b <> 0).
- contradict Hc; eapply Zmult_integral_l; eauto.
- rewrite (Zplus_minus_eq _ _ _ (Z_div_mod_eq_full (c*a) (c*b) H)).
- rewrite (Zplus_minus_eq _ _ _ (Z_div_mod_eq_full a b Hb)).
- rewrite Zdiv_mult_cancel_l; auto with zarith.
- ring.
+ intros. zero_or_not c. rewrite (Zmult_comm c b); zero_or_not b.
+ rewrite (Zmult_comm b c). apply Z.mul_mod_distr_l; auto.
Qed.
Lemma Zmult_mod_distr_r: forall a b c,
(a*c) mod (b*c) = (a mod b) * c.
Proof.
- intros; repeat rewrite (fun x => (Zmult_comm x c)).
- apply Zmult_mod_distr_l; auto.
+ intros. zero_or_not b. rewrite (Zmult_comm b c); zero_or_not c.
+ rewrite (Zmult_comm c b). apply Z.mul_mod_distr_r; auto.
Qed.
(** Operations modulo. *)
Theorem Zmod_mod: forall a n, (a mod n) mod n = a mod n.
-Proof.
- intros; destruct (Z_eq_dec n 0) as [Hb|Hb].
- subst; do 2 rewrite Zmod_0_r; auto.
- pattern a at 2; rewrite (Z_div_mod_eq_full a n); auto with zarith.
- rewrite Zplus_comm; rewrite Zmult_comm.
- apply sym_equal; apply Z_mod_plus_full; auto with zarith.
-Qed.
+Proof. intros. zero_or_not n. apply Z.mod_mod; auto. Qed.
Theorem Zmult_mod: forall a b n,
(a * b) mod n = ((a mod n) * (b mod n)) mod n.
-Proof.
- intros; destruct (Z_eq_dec n 0) as [Hb|Hb].
- subst; do 2 rewrite Zmod_0_r; auto.
- pattern a at 1; rewrite (Z_div_mod_eq_full a n); auto with zarith.
- pattern b at 1; rewrite (Z_div_mod_eq_full b n); auto with zarith.
- set (A:=a mod n); set (B:=b mod n); set (A':=a/n); set (B':=b/n).
- replace ((n*A' + A) * (n*B' + B))
- with (A*B + (A'*B+B'*A+n*A'*B')*n) by ring.
- apply Z_mod_plus_full; auto with zarith.
-Qed.
+Proof. intros. zero_or_not n. apply Z.mul_mod; auto. Qed.
Theorem Zplus_mod: forall a b n,
(a + b) mod n = (a mod n + b mod n) mod n.
-Proof.
- intros; destruct (Z_eq_dec n 0) as [Hb|Hb].
- subst; do 2 rewrite Zmod_0_r; auto.
- pattern a at 1; rewrite (Z_div_mod_eq_full a n); auto with zarith.
- pattern b at 1; rewrite (Z_div_mod_eq_full b n); auto with zarith.
- replace ((n * (a / n) + a mod n) + (n * (b / n) + b mod n))
- with ((a mod n + b mod n) + (a / n + b / n) * n) by ring.
- apply Z_mod_plus_full; auto with zarith.
-Qed.
+Proof. intros. zero_or_not n. apply Z.add_mod; auto. Qed.
Theorem Zminus_mod: forall a b n,
(a - b) mod n = (a mod n - b mod n) mod n.
@@ -897,49 +443,48 @@ Qed.
(** For a specific number N, equality modulo N is hence a nice setoid
equivalence, compatible with [+], [-] and [*]. *)
-Definition eqm N a b := (a mod N = b mod N).
+Section EqualityModulo.
+Variable N:Z.
-Lemma eqm_refl N : forall a, (eqm N) a a.
+Definition eqm a b := (a mod N = b mod N).
+Infix "==" := eqm (at level 70).
+
+Lemma eqm_refl : forall a, a == a.
Proof. unfold eqm; auto. Qed.
-Lemma eqm_sym N : forall a b, (eqm N) a b -> (eqm N) b a.
+Lemma eqm_sym : forall a b, a == b -> b == a.
Proof. unfold eqm; auto. Qed.
-Lemma eqm_trans N : forall a b c,
- (eqm N) a b -> (eqm N) b c -> (eqm N) a c.
+Lemma eqm_trans : forall a b c,
+ a == b -> b == c -> a == c.
Proof. unfold eqm; eauto with *. Qed.
-Add Parametric Relation N : Z (eqm N)
- reflexivity proved by (eqm_refl N)
- symmetry proved by (eqm_sym N)
- transitivity proved by (eqm_trans N) as eqm_setoid.
+Instance eqm_setoid : Equivalence eqm.
+Proof.
+ constructor; [exact eqm_refl | exact eqm_sym | exact eqm_trans].
+Qed.
-Add Parametric Morphism N : Zplus
- with signature (eqm N) ==> (eqm N) ==> (eqm N) as Zplus_eqm.
+Instance Zplus_eqm : Proper (eqm ==> eqm ==> eqm) Zplus.
Proof.
- unfold eqm; intros; rewrite Zplus_mod, H, H0, <- Zplus_mod; auto.
+ unfold eqm; repeat red; intros. rewrite Zplus_mod, H, H0, <- Zplus_mod; auto.
Qed.
-Add Parametric Morphism N : Zminus
- with signature (eqm N) ==> (eqm N) ==> (eqm N) as Zminus_eqm.
+Instance Zminus_eqm : Proper (eqm ==> eqm ==> eqm) Zminus.
Proof.
- unfold eqm; intros; rewrite Zminus_mod, H, H0, <- Zminus_mod; auto.
+ unfold eqm; repeat red; intros. rewrite Zminus_mod, H, H0, <- Zminus_mod; auto.
Qed.
-Add Parametric Morphism N : Zmult
- with signature (eqm N) ==> (eqm N) ==> (eqm N) as Zmult_eqm.
+Instance Zmult_eqm : Proper (eqm ==> eqm ==> eqm) Zmult.
Proof.
- unfold eqm; intros; rewrite Zmult_mod, H, H0, <- Zmult_mod; auto.
+ unfold eqm; repeat red; intros. rewrite Zmult_mod, H, H0, <- Zmult_mod; auto.
Qed.
-Add Parametric Morphism N : Zopp
- with signature (eqm N) ==> (eqm N) as Zopp_eqm.
+Instance Zopp_eqm : Proper (eqm ==> eqm) Zopp.
Proof.
- intros; change ((eqm N) (-x) (-y)) with ((eqm N) (0-x) (0-y)).
- rewrite H; red; auto.
+ intros x y H. change ((-x)==(-y)) with ((0-x)==(0-y)). now rewrite H.
Qed.
-Lemma Zmod_eqm N : forall a, (eqm N) (a mod N) a.
+Lemma Zmod_eqm : forall a, (a mod N) == a.
Proof.
intros; exact (Zmod_mod a N).
Qed.
@@ -952,32 +497,12 @@ Qed.
~ (1/3 == 1/1)
*)
+End EqualityModulo.
+
Lemma Zdiv_Zdiv : forall a b c, 0<=b -> 0<=c -> (a/b)/c = a/(b*c).
Proof.
- intros a b c Hb Hc.
- destruct (Zle_lt_or_eq _ _ Hb); [ | subst; rewrite Zdiv_0_r, Zdiv_0_r, Zdiv_0_l; auto].
- destruct (Zle_lt_or_eq _ _ Hc); [ | subst; rewrite Zmult_0_r, Zdiv_0_r, Zdiv_0_r; auto].
- pattern a at 2;rewrite (Z_div_mod_eq_full a b);auto with zarith.
- pattern (a/b) at 2;rewrite (Z_div_mod_eq_full (a/b) c);auto with zarith.
- replace (b * (c * (a / b / c) + (a / b) mod c) + a mod b) with
- ((a / b / c)*(b * c) + (b * ((a / b) mod c) + a mod b)) by ring.
- rewrite Z_div_plus_full_l; auto with zarith.
- rewrite (Zdiv_small (b * ((a / b) mod c) + a mod b)).
- ring.
- split.
- apply Zplus_le_0_compat;auto with zarith.
- apply Zmult_le_0_compat;auto with zarith.
- destruct (Z_mod_lt (a/b) c);auto with zarith.
- destruct (Z_mod_lt a b);auto with zarith.
- apply Zle_lt_trans with (b * ((a / b) mod c) + (b-1)).
- destruct (Z_mod_lt a b);auto with zarith.
- apply Zle_lt_trans with (b * (c-1) + (b - 1)).
- apply Zplus_le_compat;auto with zarith.
- destruct (Z_mod_lt (a/b) c);auto with zarith.
- replace (b * (c - 1) + (b - 1)) with (b*c-1);try ring;auto with zarith.
- intro H1;
- assert (H2: c <> 0) by auto with zarith;
- rewrite (Zmult_integral_l _ _ H2 H1) in H; auto with zarith.
+ intros. zero_or_not b. rewrite Zmult_comm. zero_or_not c.
+ rewrite Z.mul_comm. apply Z.div_div; auto with zarith.
Qed.
(** Unfortunately, the previous result isn't always true on negative numbers.
@@ -988,40 +513,40 @@ Qed.
Theorem Zdiv_mult_le:
forall a b c, 0<=a -> 0<=b -> 0<=c -> c*(a/b) <= (c*a)/b.
Proof.
- intros a b c H1 H2 H3.
- destruct (Zle_lt_or_eq _ _ H2);
- [ | subst; rewrite Zdiv_0_r, Zdiv_0_r, Zmult_0_r; auto].
- case (Z_mod_lt a b); auto with zarith; intros Hu1 Hu2.
- case (Z_mod_lt c b); auto with zarith; intros Hv1 Hv2.
- apply Zmult_le_reg_r with b; auto with zarith.
- rewrite <- Zmult_assoc.
- replace (a / b * b) with (a - a mod b).
- replace (c * a / b * b) with (c * a - (c * a) mod b).
- rewrite Zmult_minus_distr_l.
- unfold Zminus; apply Zplus_le_compat_l.
- match goal with |- - ?X <= -?Y => assert (Y <= X); auto with zarith end.
- apply Zle_trans with ((c mod b) * (a mod b)); auto with zarith.
- rewrite Zmult_mod; auto with zarith.
- apply (Zmod_le ((c mod b) * (a mod b)) b); auto with zarith.
- apply Zmult_le_compat_r; auto with zarith.
- apply (Zmod_le c b); auto.
- pattern (c * a) at 1; rewrite (Z_div_mod_eq (c * a) b); try ring;
- auto with zarith.
- pattern a at 1; rewrite (Z_div_mod_eq a b); try ring; auto with zarith.
-Qed.
+ intros. zero_or_not b. apply Z.div_mul_le; auto with zarith. Qed.
(** Zmod is related to divisibility (see more in Znumtheory) *)
Lemma Zmod_divides : forall a b, b<>0 ->
(a mod b = 0 <-> exists c, a = b*c).
Proof.
- split; intros.
- exists (a/b).
- pattern a at 1; rewrite (Z_div_mod_eq_full a b); auto with zarith.
- destruct H0 as [c Hc].
- symmetry.
- apply Zmod_unique_full with c; auto with zarith.
- red; omega with *.
+ intros. rewrite Z.mod_divide; trivial.
+ split; intros (c,Hc); exists c; subst; auto with zarith.
+Qed.
+
+(** Particular case : dividing by 2 is related with parity *)
+
+Lemma Zdiv2_div : forall a, Z.div2 a = a/2.
+Proof Z.div2_div.
+
+Lemma Zmod_odd : forall a, a mod 2 = if Z.odd a then 1 else 0.
+Proof.
+ intros a. now rewrite <- Z.bit0_odd, <- Z.bit0_mod.
+Qed.
+
+Lemma Zmod_even : forall a, a mod 2 = if Z.even a then 0 else 1.
+Proof.
+ intros a. rewrite Zmod_odd, Zodd_even_bool. now destruct Zeven_bool.
+Qed.
+
+Lemma Zodd_mod : forall a, Z.odd a = Zeq_bool (a mod 2) 1.
+Proof.
+ intros a. rewrite Zmod_odd. now destruct Zodd_bool.
+Qed.
+
+Lemma Zeven_mod : forall a, Z.even a = Zeq_bool (a mod 2) 0.
+Proof.
+ intros a. rewrite Zmod_even. now destruct Zeven_bool.
Qed.
(** * Compatibility *)
@@ -1075,12 +600,12 @@ Fixpoint Zmod_POS (a : positive) (b : Z) : Z :=
| xI a' =>
let r := Zmod_POS a' b in
let r' := (2 * r + 1) in
- if Zgt_bool b r' then r' else (r' - b)
+ if r' <? b then r' else (r' - b)
| xO a' =>
let r := Zmod_POS a' b in
let r' := (2 * r) in
- if Zgt_bool b r' then r' else (r' - b)
- | xH => if Zge_bool b 2 then 1 else 0
+ if r' <? b then r' else (r' - b)
+ | xH => if 2 <=? b then 1 else 0
end.
Definition Zmod' a b :=
@@ -1105,30 +630,28 @@ Definition Zmod' a b :=
end.
-Theorem Zmod_POS_correct: forall a b, Zmod_POS a b = (snd (Zdiv_eucl_POS a b)).
+Theorem Zmod_POS_correct a b : Zmod_POS a b = snd (Z.pos_div_eucl a b).
Proof.
- intros a b; elim a; simpl; auto.
- intros p Rec; rewrite Rec.
- case (Zdiv_eucl_POS p b); intros z1 z2; simpl; auto.
- match goal with |- context [Zgt_bool _ ?X] => case (Zgt_bool b X) end; auto.
- intros p Rec; rewrite Rec.
- case (Zdiv_eucl_POS p b); intros z1 z2; simpl; auto.
- match goal with |- context [Zgt_bool _ ?X] => case (Zgt_bool b X) end; auto.
- case (Zge_bool b 2); auto.
+ induction a as [a IH|a IH| ]; simpl; rewrite ?IH.
+ destruct (Z.pos_div_eucl a b) as (p,q); simpl;
+ case Z.ltb_spec; reflexivity.
+ destruct (Z.pos_div_eucl a b) as (p,q); simpl;
+ case Z.ltb_spec; reflexivity.
+ case Z.leb_spec; trivial.
Qed.
-Theorem Zmod'_correct: forall a b, Zmod' a b = Zmod a b.
+Theorem Zmod'_correct: forall a b, Zmod' a b = a mod b.
Proof.
- intros a b; unfold Zmod; case a; simpl; auto.
+ intros a b; unfold Z.modulo; case a; simpl; auto.
intros p; case b; simpl; auto.
intros p1; refine (Zmod_POS_correct _ _); auto.
intros p1; rewrite Zmod_POS_correct; auto.
- case (Zdiv_eucl_POS p (Zpos p1)); simpl; intros z1 z2; case z2; auto.
+ case (Z.pos_div_eucl p (Zpos p1)); simpl; intros z1 z2; case z2; auto.
intros p; case b; simpl; auto.
intros p1; rewrite Zmod_POS_correct; auto.
- case (Zdiv_eucl_POS p (Zpos p1)); simpl; intros z1 z2; case z2; auto.
+ case (Z.pos_div_eucl p (Zpos p1)); simpl; intros z1 z2; case z2; auto.
intros p1; rewrite Zmod_POS_correct; simpl; auto.
- case (Zdiv_eucl_POS p (Zpos p1)); auto.
+ case (Z.pos_div_eucl p (Zpos p1)); auto.
Qed.
@@ -1140,12 +663,12 @@ Theorem Zdiv_eucl_extended :
forall b:Z,
b <> 0 ->
forall a:Z,
- {qr : Z * Z | let (q, r) := qr in a = b * q + r /\ 0 <= r < Zabs b}.
+ {qr : Z * Z | let (q, r) := qr in a = b * q + r /\ 0 <= r < Z.abs b}.
Proof.
intros b Hb a.
elim (Z_le_gt_dec 0 b); intro Hb'.
cut (b > 0); [ intro Hb'' | omega ].
- rewrite Zabs_eq; [ apply Zdiv_eucl_exist; assumption | assumption ].
+ rewrite Z.abs_eq; [ apply Zdiv_eucl_exist; assumption | assumption ].
cut (- b > 0); [ intro Hb'' | omega ].
elim (Zdiv_eucl_exist Hb'' a); intros qr.
elim qr; intros q r Hqr.
@@ -1153,17 +676,33 @@ Proof.
elim Hqr; intros.
split.
rewrite <- Zmult_opp_comm; assumption.
- rewrite Zabs_non_eq; [ assumption | omega ].
+ rewrite Z.abs_neq; [ assumption | omega ].
Qed.
-Implicit Arguments Zdiv_eucl_extended.
+Arguments Zdiv_eucl_extended : default implicits.
-(** A third convention: Ocaml.
+(** * Division and modulo in Z agree with same in nat: *)
- See files ZOdiv_def.v and ZOdiv.v.
+Require Import NPeano.
- Ocaml uses Round-Toward-Zero division: (-a)/b = a/(-b) = -(a/b).
- Hence (-a) mod b = - (a mod b)
- a mod (-b) = a mod b
- And: |r| < |b| and sgn(r) = sgn(a) (notice the a here instead of b).
-*)
+Lemma div_Zdiv (n m: nat): m <> O ->
+ Z.of_nat (n / m) = Z.of_nat n / Z.of_nat m.
+Proof.
+ intros.
+ apply (Zdiv_unique _ _ _ (Z.of_nat (n mod m))).
+ split. auto with zarith.
+ now apply inj_lt, Nat.mod_upper_bound.
+ rewrite <- inj_mult, <- inj_plus.
+ now apply inj_eq, Nat.div_mod.
+Qed.
+
+Lemma mod_Zmod (n m: nat): m <> O ->
+ Z.of_nat (n mod m) = (Z.of_nat n) mod (Z.of_nat m).
+Proof.
+ intros.
+ apply (Zmod_unique _ _ (Z.of_nat n / Z.of_nat m)).
+ split. auto with zarith.
+ now apply inj_lt, Nat.mod_upper_bound.
+ rewrite <- div_Zdiv, <- inj_mult, <- inj_plus by trivial.
+ now apply inj_eq, Nat.div_mod.
+Qed.
diff --git a/theories/ZArith/Zeuclid.v b/theories/ZArith/Zeuclid.v
new file mode 100644
index 00000000..f1b59749
--- /dev/null
+++ b/theories/ZArith/Zeuclid.v
@@ -0,0 +1,52 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Morphisms BinInt ZDivEucl.
+Local Open Scope Z_scope.
+
+(** * Definitions of division for binary integers, Euclid convention. *)
+
+(** In this convention, the remainder is always positive.
+ For other conventions, see [Z.div] and [Z.quot] in file [BinIntDef].
+ To avoid collision with the other divisions, we place this one
+ under a module.
+*)
+
+Module ZEuclid.
+
+ Definition modulo a b := Z.modulo a (Z.abs b).
+ Definition div a b := (Z.sgn b) * (Z.div a (Z.abs b)).
+
+ Instance mod_wd : Proper (eq==>eq==>eq) modulo.
+ Proof. congruence. Qed.
+ Instance div_wd : Proper (eq==>eq==>eq) div.
+ Proof. congruence. Qed.
+
+ Theorem div_mod a b : b<>0 -> a = b*(div a b) + modulo a b.
+ Proof.
+ intros Hb. unfold div, modulo.
+ rewrite Z.mul_assoc. rewrite Z.sgn_abs. apply Z.div_mod.
+ now destruct b.
+ Qed.
+
+ Lemma mod_always_pos a b : b<>0 -> 0 <= modulo a b < Z.abs b.
+ Proof.
+ intros Hb. unfold modulo.
+ apply Z.mod_pos_bound.
+ destruct b; compute; trivial. now destruct Hb.
+ Qed.
+
+ Lemma mod_bound_pos a b : 0<=a -> 0<b -> 0 <= modulo a b < b.
+ Proof.
+ intros _ Hb. rewrite <- (Z.abs_eq b) at 3 by Z.order.
+ apply mod_always_pos. Z.order.
+ Qed.
+
+ Include ZEuclidProp Z Z Z.
+
+End ZEuclid.
diff --git a/theories/ZArith/Zeven.v b/theories/ZArith/Zeven.v
index 883b7f15..550b66f7 100644
--- a/theories/ZArith/Zeven.v
+++ b/theories/ZArith/Zeven.v
@@ -1,22 +1,25 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zeven.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
+(** Binary Integers : Parity and Division by Two *)
+(** Initial author : Pierre Crégut (CNET, Lannion, France) *)
+
+(** THIS FILE IS DEPRECATED.
+ It is now almost entirely made of compatibility formulations
+ for results already present in BinInt.Z. *)
Require Import BinInt.
Open Scope Z_scope.
-(*******************************************************************)
-(** About parity: even and odd predicates on Z, division by 2 on Z *)
-
-(***************************************************)
-(** * [Zeven], [Zodd] and their related properties *)
+(** Historical formulation of even and odd predicates, based on
+ case analysis. We now rather recommend using [Z.Even] and
+ [Z.Odd], which are based on the exist predicate. *)
Definition Zeven (z:Z) :=
match z with
@@ -35,281 +38,253 @@ Definition Zodd (z:Z) :=
| _ => False
end.
-Definition Zeven_bool (z:Z) :=
- match z with
- | Z0 => true
- | Zpos (xO _) => true
- | Zneg (xO _) => true
- | _ => false
- end.
+Lemma Zeven_equiv z : Zeven z <-> Z.Even z.
+Proof.
+ rewrite <- Z.even_spec.
+ destruct z as [|p|p]; try destruct p; simpl; intuition.
+Qed.
-Definition Zodd_bool (z:Z) :=
- match z with
- | Z0 => false
- | Zpos (xO _) => false
- | Zneg (xO _) => false
- | _ => true
- end.
+Lemma Zodd_equiv z : Zodd z <-> Z.Odd z.
+Proof.
+ rewrite <- Z.odd_spec.
+ destruct z as [|p|p]; try destruct p; simpl; intuition.
+Qed.
+
+Theorem Zeven_ex_iff n : Zeven n <-> exists m, n = 2*m.
+Proof (Zeven_equiv n).
+
+Theorem Zodd_ex_iff n : Zodd n <-> exists m, n = 2*m + 1.
+Proof (Zodd_equiv n).
+
+(** Boolean tests of parity (now in BinInt.Z) *)
+
+Notation Zeven_bool := Z.even (only parsing).
+Notation Zodd_bool := Z.odd (only parsing).
+
+Lemma Zeven_bool_iff n : Z.even n = true <-> Zeven n.
+Proof.
+ now rewrite Z.even_spec, Zeven_equiv.
+Qed.
+
+Lemma Zodd_bool_iff n : Z.odd n = true <-> Zodd n.
+Proof.
+ now rewrite Z.odd_spec, Zodd_equiv.
+Qed.
+
+Ltac boolify_even_odd := rewrite <- ?Zeven_bool_iff, <- ?Zodd_bool_iff.
+
+Lemma Zodd_even_bool n : Z.odd n = negb (Z.even n).
+Proof.
+ symmetry. apply Z.negb_even.
+Qed.
+
+Lemma Zeven_odd_bool n : Z.even n = negb (Z.odd n).
+Proof.
+ symmetry. apply Z.negb_odd.
+Qed.
-Definition Zeven_odd_dec : forall z:Z, {Zeven z} + {Zodd z}.
+Definition Zeven_odd_dec n : {Zeven n} + {Zodd n}.
Proof.
- intro z. case z;
- [ left; compute in |- *; trivial
- | intro p; case p; intros;
- (right; compute in |- *; exact I) || (left; compute in |- *; exact I)
- | intro p; case p; intros;
- (right; compute in |- *; exact I) || (left; compute in |- *; exact I) ].
+ destruct n as [|p|p]; try destruct p; simpl; (now left) || (now right).
Defined.
-Definition Zeven_dec : forall z:Z, {Zeven z} + {~ Zeven z}.
+Definition Zeven_dec n : {Zeven n} + {~ Zeven n}.
Proof.
- intro z. case z;
- [ left; compute in |- *; trivial
- | intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial)
- | intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial) ].
+ destruct n as [|p|p]; try destruct p; simpl; (now left) || (now right).
Defined.
-Definition Zodd_dec : forall z:Z, {Zodd z} + {~ Zodd z}.
+Definition Zodd_dec n : {Zodd n} + {~ Zodd n}.
Proof.
- intro z. case z;
- [ right; compute in |- *; trivial
- | intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial)
- | intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial) ].
+ destruct n as [|p|p]; try destruct p; simpl; (now left) || (now right).
Defined.
-Lemma Zeven_not_Zodd : forall n:Z, Zeven n -> ~ Zodd n.
+Lemma Zeven_not_Zodd n : Zeven n -> ~ Zodd n.
Proof.
- intro z; destruct z; [ idtac | destruct p | destruct p ]; compute in |- *;
- trivial.
+ boolify_even_odd. rewrite <- Z.negb_odd. destruct Z.odd; intuition.
Qed.
-Lemma Zodd_not_Zeven : forall n:Z, Zodd n -> ~ Zeven n.
+Lemma Zodd_not_Zeven n : Zodd n -> ~ Zeven n.
Proof.
- intro z; destruct z; [ idtac | destruct p | destruct p ]; compute in |- *;
- trivial.
+ boolify_even_odd. rewrite <- Z.negb_odd. destruct Z.odd; intuition.
Qed.
-Lemma Zeven_Sn : forall n:Z, Zodd n -> Zeven (Zsucc n).
+Lemma Zeven_Sn n : Zodd n -> Zeven (Z.succ n).
Proof.
- intro z; destruct z; unfold Zsucc in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
- trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ boolify_even_odd. now rewrite Z.even_succ.
Qed.
-Lemma Zodd_Sn : forall n:Z, Zeven n -> Zodd (Zsucc n).
+Lemma Zodd_Sn n : Zeven n -> Zodd (Z.succ n).
Proof.
- intro z; destruct z; unfold Zsucc in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
- trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ boolify_even_odd. now rewrite Z.odd_succ.
Qed.
-Lemma Zeven_pred : forall n:Z, Zodd n -> Zeven (Zpred n).
+Lemma Zeven_pred n : Zodd n -> Zeven (Z.pred n).
Proof.
- intro z; destruct z; unfold Zpred in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
- trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ boolify_even_odd. now rewrite Z.even_pred.
Qed.
-Lemma Zodd_pred : forall n:Z, Zeven n -> Zodd (Zpred n).
+Lemma Zodd_pred n : Zeven n -> Zodd (Z.pred n).
Proof.
- intro z; destruct z; unfold Zpred in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
- trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ boolify_even_odd. now rewrite Z.odd_pred.
Qed.
Hint Unfold Zeven Zodd: zarith.
+Notation Zeven_bool_succ := Z.even_succ (only parsing).
+Notation Zeven_bool_pred := Z.even_pred (only parsing).
+Notation Zodd_bool_succ := Z.odd_succ (only parsing).
+Notation Zodd_bool_pred := Z.odd_pred (only parsing).
(******************************************************************)
-(** * Definition of [Zdiv2] and properties wrt [Zeven] and [Zodd] *)
+(** * Definition of [Zquot2], [Zdiv2] and properties wrt [Zeven]
+ and [Zodd] *)
-(** [Zdiv2] is defined on all [Z], but notice that for odd negative
- integers it is not the euclidean quotient: in that case we have
- [n = 2*(n/2)-1] *)
+Notation Zdiv2 := Z.div2 (only parsing).
+Notation Zquot2 := Z.quot2 (only parsing).
-Definition Zdiv2 (z:Z) :=
- match z with
- | Z0 => 0
- | Zpos xH => 0
- | Zpos p => Zpos (Pdiv2 p)
- | Zneg xH => 0
- | Zneg p => Zneg (Pdiv2 p)
- end.
+(** Properties of [Z.div2] *)
-Lemma Zeven_div2 : forall n:Z, Zeven n -> n = 2 * Zdiv2 n.
+Lemma Zdiv2_odd_eqn n : n = 2*(Z.div2 n) + if Z.odd n then 1 else 0.
+Proof (Z.div2_odd n).
+
+Lemma Zeven_div2 n : Zeven n -> n = 2 * Z.div2 n.
Proof.
- intro x; destruct x.
- auto with arith.
- destruct p; auto with arith.
- intros. absurd (Zeven (Zpos (xI p))); red in |- *; auto with arith.
- intros. absurd (Zeven 1); red in |- *; auto with arith.
- destruct p; auto with arith.
- intros. absurd (Zeven (Zneg (xI p))); red in |- *; auto with arith.
- intros. absurd (Zeven (-1)); red in |- *; auto with arith.
+ boolify_even_odd. rewrite <- Z.negb_odd, Bool.negb_true_iff.
+ intros Hn. rewrite (Zdiv2_odd_eqn n) at 1. now rewrite Hn, Z.add_0_r.
Qed.
-Lemma Zodd_div2 : forall n:Z, n >= 0 -> Zodd n -> n = 2 * Zdiv2 n + 1.
+Lemma Zodd_div2 n : Zodd n -> n = 2 * Z.div2 n + 1.
Proof.
- intro x; destruct x.
- intros. absurd (Zodd 0); red in |- *; auto with arith.
- destruct p; auto with arith.
- intros. absurd (Zodd (Zpos (xO p))); red in |- *; auto with arith.
- intros. absurd (Zneg p >= 0); red in |- *; auto with arith.
+ boolify_even_odd.
+ intros Hn. rewrite (Zdiv2_odd_eqn n) at 1. now rewrite Hn.
Qed.
-Lemma Zodd_div2_neg :
- forall n:Z, n <= 0 -> Zodd n -> n = 2 * Zdiv2 n - 1.
+(** Properties of [Z.quot2] *)
+
+(** TODO: move to Numbers someday *)
+
+Lemma Zquot2_odd_eqn n : n = 2*(Z.quot2 n) + if Z.odd n then Z.sgn n else 0.
Proof.
- intro x; destruct x.
- intros. absurd (Zodd 0); red in |- *; auto with arith.
- intros. absurd (Zneg p >= 0); red in |- *; auto with arith.
- destruct p; auto with arith.
- intros. absurd (Zodd (Zneg (xO p))); red in |- *; auto with arith.
+ now destruct n as [ |[p|p| ]|[p|p| ]].
Qed.
-Lemma Z_modulo_2 :
- forall n:Z, {y : Z | n = 2 * y} + {y : Z | n = 2 * y + 1}.
+Lemma Zeven_quot2 n : Zeven n -> n = 2 * Z.quot2 n.
Proof.
- intros x.
- elim (Zeven_odd_dec x); intro.
- left. split with (Zdiv2 x). exact (Zeven_div2 x a).
- right. generalize b; clear b; case x.
- intro b; inversion b.
- intro p; split with (Zdiv2 (Zpos p)). apply (Zodd_div2 (Zpos p)); trivial.
- unfold Zge, Zcompare in |- *; simpl in |- *; discriminate.
- intro p; split with (Zdiv2 (Zpred (Zneg p))).
- pattern (Zneg p) at 1 in |- *; rewrite (Zsucc_pred (Zneg p)).
- pattern (Zpred (Zneg p)) at 1 in |- *; rewrite (Zeven_div2 (Zpred (Zneg p))).
- reflexivity.
- apply Zeven_pred; assumption.
+ intros Hn. apply Zeven_bool_iff in Hn.
+ rewrite (Zquot2_odd_eqn n) at 1.
+ now rewrite Zodd_even_bool, Hn, Z.add_0_r.
Qed.
-Lemma Zsplit2 :
- forall n:Z,
- {p : Z * Z |
- let (x1, x2) := p in n = x1 + x2 /\ (x1 = x2 \/ x2 = x1 + 1)}.
+Lemma Zodd_quot2 n : n >= 0 -> Zodd n -> n = 2 * Z.quot2 n + 1.
Proof.
- intros x.
- elim (Z_modulo_2 x); intros [y Hy]; rewrite Zmult_comm in Hy;
- rewrite <- Zplus_diag_eq_mult_2 in Hy.
- exists (y, y); split.
- assumption.
- left; reflexivity.
- exists (y, (y + 1)%Z); split.
- rewrite Zplus_assoc; assumption.
- right; reflexivity.
+ intros Hn Hn'. apply Zodd_bool_iff in Hn'.
+ rewrite (Zquot2_odd_eqn n) at 1. rewrite Hn'. f_equal.
+ destruct n; (now destruct Hn) || easy.
Qed.
+Lemma Zodd_quot2_neg n : n <= 0 -> Zodd n -> n = 2 * Z.quot2 n - 1.
+Proof.
+ intros Hn Hn'. apply Zodd_bool_iff in Hn'.
+ rewrite (Zquot2_odd_eqn n) at 1; rewrite Hn'. unfold Z.sub. f_equal.
+ destruct n; (now destruct Hn) || easy.
+Qed.
-Theorem Zeven_ex: forall n, Zeven n -> exists m, n = 2 * m.
+Lemma Zquot2_opp n : Z.quot2 (-n) = - Z.quot2 n.
Proof.
- intro n; exists (Zdiv2 n); apply Zeven_div2; auto.
+ now destruct n as [ |[p|p| ]|[p|p| ]].
Qed.
-Theorem Zodd_ex: forall n, Zodd n -> exists m, n = 2 * m + 1.
+Lemma Zquot2_quot n : Z.quot2 n = n ÷ 2.
Proof.
- destruct n; intros.
- inversion H.
- exists (Zdiv2 (Zpos p)).
- apply Zodd_div2; simpl; auto; compute; inversion 1.
- exists (Zdiv2 (Zneg p) - 1).
- unfold Zminus.
- rewrite Zmult_plus_distr_r.
- rewrite <- Zplus_assoc.
- assert (Zneg p <= 0) by (compute; inversion 1).
- exact (Zodd_div2_neg _ H0 H).
+ assert (AUX : forall m, 0 < m -> Z.quot2 m = m ÷ 2).
+ BeginSubproof.
+ intros m Hm.
+ apply Z.quot_unique with (if Z.odd m then Z.sgn m else 0).
+ now apply Z.lt_le_incl.
+ rewrite Z.sgn_pos by trivial.
+ destruct (Z.odd m); now split.
+ apply Zquot2_odd_eqn.
+ EndSubproof.
+ destruct (Z.lt_trichotomy 0 n) as [POS|[NUL|NEG]].
+ - now apply AUX.
+ - now subst.
+ - apply Z.opp_inj. rewrite <- Z.quot_opp_l, <- Zquot2_opp.
+ apply AUX. now destruct n. easy.
Qed.
-Theorem Zeven_2p: forall p, Zeven (2 * p).
+(** More properties of parity *)
+
+Lemma Z_modulo_2 n : {y | n = 2 * y} + {y | n = 2 * y + 1}.
Proof.
- destruct p; simpl; auto.
+ destruct (Zeven_odd_dec n) as [Hn|Hn].
+ - left. exists (Z.div2 n). exact (Zeven_div2 n Hn).
+ - right. exists (Z.div2 n). exact (Zodd_div2 n Hn).
Qed.
-Theorem Zodd_2p_plus_1: forall p, Zodd (2 * p + 1).
+Lemma Zsplit2 n :
+ {p : Z * Z | let (x1, x2) := p in n = x1 + x2 /\ (x1 = x2 \/ x2 = x1 + 1)}.
Proof.
- destruct p; simpl; auto.
- destruct p; simpl; auto.
+ destruct (Z_modulo_2 n) as [(y,Hy)|(y,Hy)];
+ rewrite Z.mul_comm, <- Zplus_diag_eq_mult_2 in Hy.
+ - exists (y, y). split. assumption. now left.
+ - exists (y, y + 1). split. now rewrite Z.add_assoc. now right.
Qed.
-Theorem Zeven_plus_Zodd: forall a b,
- Zeven a -> Zodd b -> Zodd (a + b).
+Theorem Zeven_ex n : Zeven n -> exists m, n = 2 * m.
Proof.
- intros a b H1 H2; case Zeven_ex with (1 := H1); intros x H3; try rewrite H3; auto.
- case Zodd_ex with (1 := H2); intros y H4; try rewrite H4; auto.
- replace (2 * x + (2 * y + 1)) with (2 * (x + y) + 1); try apply Zodd_2p_plus_1; auto with zarith.
- rewrite Zmult_plus_distr_r, Zplus_assoc; auto.
+ exists (Z.div2 n); apply Zeven_div2; auto.
Qed.
-Theorem Zeven_plus_Zeven: forall a b,
- Zeven a -> Zeven b -> Zeven (a + b).
+Theorem Zodd_ex n : Zodd n -> exists m, n = 2 * m + 1.
Proof.
- intros a b H1 H2; case Zeven_ex with (1 := H1); intros x H3; try rewrite H3; auto.
- case Zeven_ex with (1 := H2); intros y H4; try rewrite H4; auto.
- replace (2 * x + 2 * y) with (2 * (x + y)); try apply Zeven_2p; auto with zarith.
- apply Zmult_plus_distr_r; auto.
+ exists (Z.div2 n); apply Zodd_div2; auto.
Qed.
-Theorem Zodd_plus_Zeven: forall a b,
- Zodd a -> Zeven b -> Zodd (a + b).
+Theorem Zeven_2p p : Zeven (2 * p).
Proof.
- intros a b H1 H2; rewrite Zplus_comm; apply Zeven_plus_Zodd; auto.
+ now destruct p.
Qed.
-Theorem Zodd_plus_Zodd: forall a b,
- Zodd a -> Zodd b -> Zeven (a + b).
+Theorem Zodd_2p_plus_1 p : Zodd (2 * p + 1).
Proof.
- intros a b H1 H2; case Zodd_ex with (1 := H1); intros x H3; try rewrite H3; auto.
- case Zodd_ex with (1 := H2); intros y H4; try rewrite H4; auto.
- replace ((2 * x + 1) + (2 * y + 1)) with (2 * (x + y + 1)); try apply Zeven_2p; auto.
- (* ring part *)
- do 2 rewrite Zmult_plus_distr_r; auto.
- repeat rewrite <- Zplus_assoc; f_equal.
- rewrite (Zplus_comm 1).
- repeat rewrite <- Zplus_assoc; auto.
+ destruct p as [|p|p]; now try destruct p.
Qed.
-Theorem Zeven_mult_Zeven_l: forall a b,
- Zeven a -> Zeven (a * b).
+Theorem Zeven_plus_Zodd a b : Zeven a -> Zodd b -> Zodd (a + b).
Proof.
- intros a b H1; case Zeven_ex with (1 := H1); intros x H3; try rewrite H3; auto.
- replace (2 * x * b) with (2 * (x * b)); try apply Zeven_2p; auto with zarith.
- (* ring part *)
- apply Zmult_assoc.
+ boolify_even_odd. rewrite <- Z.negb_odd, Bool.negb_true_iff.
+ intros Ha Hb. now rewrite Z.odd_add, Ha, Hb.
Qed.
-Theorem Zeven_mult_Zeven_r: forall a b,
- Zeven b -> Zeven (a * b).
+Theorem Zeven_plus_Zeven a b : Zeven a -> Zeven b -> Zeven (a + b).
Proof.
- intros a b H1; case Zeven_ex with (1 := H1); intros x H3; try rewrite H3; auto.
- replace (a * (2 * x)) with (2 * (x * a)); try apply Zeven_2p; auto.
- (* ring part *)
- rewrite (Zmult_comm x a).
- do 2 rewrite Zmult_assoc.
- rewrite (Zmult_comm 2 a); auto.
+ boolify_even_odd. intros Ha Hb. now rewrite Z.even_add, Ha, Hb.
Qed.
-Hint Rewrite Zmult_plus_distr_r Zmult_plus_distr_l
- Zplus_assoc Zmult_1_r Zmult_1_l : Zexpand.
+Theorem Zodd_plus_Zeven a b : Zodd a -> Zeven b -> Zodd (a + b).
+Proof.
+ intros. rewrite Z.add_comm. now apply Zeven_plus_Zodd.
+Qed.
+
+Theorem Zodd_plus_Zodd a b : Zodd a -> Zodd b -> Zeven (a + b).
+Proof.
+ boolify_even_odd. rewrite <- 2 Z.negb_even, 2 Bool.negb_true_iff.
+ intros Ha Hb. now rewrite Z.even_add, Ha, Hb.
+Qed.
+
+Theorem Zeven_mult_Zeven_l a b : Zeven a -> Zeven (a * b).
+Proof.
+ boolify_even_odd. intros Ha. now rewrite Z.even_mul, Ha.
+Qed.
+
+Theorem Zeven_mult_Zeven_r a b : Zeven b -> Zeven (a * b).
+Proof.
+ intros. rewrite Z.mul_comm. now apply Zeven_mult_Zeven_l.
+Qed.
-Theorem Zodd_mult_Zodd: forall a b,
- Zodd a -> Zodd b -> Zodd (a * b).
+Theorem Zodd_mult_Zodd a b : Zodd a -> Zodd b -> Zodd (a * b).
Proof.
- intros a b H1 H2; case Zodd_ex with (1 := H1); intros x H3; try rewrite H3; auto.
- case Zodd_ex with (1 := H2); intros y H4; try rewrite H4; auto.
- replace ((2 * x + 1) * (2 * y + 1)) with (2 * (2 * x * y + x + y) + 1); try apply Zodd_2p_plus_1; auto.
- (* ring part *)
- autorewrite with Zexpand; f_equal.
- repeat rewrite <- Zplus_assoc; f_equal.
- repeat rewrite <- Zmult_assoc; f_equal.
- repeat rewrite Zmult_assoc; f_equal; apply Zmult_comm.
+ boolify_even_odd. intros Ha Hb. now rewrite Z.odd_mul, Ha, Hb.
Qed.
(* for compatibility *)
diff --git a/theories/ZArith/Zgcd_alt.v b/theories/ZArith/Zgcd_alt.v
index 86fe0ef9..ebf3d024 100644
--- a/theories/ZArith/Zgcd_alt.v
+++ b/theories/ZArith/Zgcd_alt.v
@@ -1,14 +1,12 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zgcd_alt.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-(** * Zgcd_alt : an alternate version of Zgcd, based on Euler's algorithm *)
+(** * Zgcd_alt : an alternate version of Zgcd, based on Euclid's algorithm *)
(**
Author: Pierre Letouzey
@@ -17,7 +15,7 @@ Author: Pierre Letouzey
(** The alternate [Zgcd_alt] given here used to be the main [Zgcd]
function (see file [Znumtheory]), but this main [Zgcd] is now
based on a modern binary-efficient algorithm. This earlier
- version, based on Euler's algorithm of iterated modulo, is kept
+ version, based on Euclid's algorithm of iterated modulo, is kept
here due to both its intrinsic interest and its use as reference
point when proving gcd on Int31 numbers *)
diff --git a/theories/ZArith/Zhints.v b/theories/ZArith/Zhints.v
index c2348967..6a14d693 100644
--- a/theories/ZArith/Zhints.v
+++ b/theories/ZArith/Zhints.v
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zhints.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(** This file centralizes the lemmas about [Z], classifying them
according to the way they can be used in automatic search *)
@@ -100,440 +98,3 @@ Hint Resolve
Zplus_le_compat (* :(n,m,p,q:Z)`n <= m`->`p <= q`->`n+p <= m+q` *)
: zarith.
-
-(**********************************************************************)
-(** * Reversible lemmas relating operators *)
-(** Probably to be declared as hints but need to define precedences *)
-
-(** ** Conversion between comparisons/predicates and arithmetic operators *)
-
-(** Lemmas ending by eq *)
-(**
-<<
-Zegal_left: (x,y:Z)`x = y`->`x+(-y) = 0`
-Zabs_eq: (x:Z)`0 <= x`->`|x| = x`
-Zeven_div2: (x:Z)(Zeven x)->`x = 2*(Zdiv2 x)`
-Zodd_div2: (x:Z)`x >= 0`->(Zodd x)->`x = 2*(Zdiv2 x)+1`
->>
-*)
-
-(** Lemmas ending by Zgt *)
-(**
-<<
-Zgt_left_rev: (x,y:Z)`x+(-y) > 0`->`x > y`
-Zgt_left_gt: (x,y:Z)`x > y`->`x+(-y) > 0`
->>
-*)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-Zlt_left_rev: (x,y:Z)`0 < y+(-x)`->`x < y`
-Zlt_left_lt: (x,y:Z)`x < y`->`0 < y+(-x)`
-Zlt_O_minus_lt: (n,m:Z)`0 < n-m`->`m < n`
->>
-*)
-
-(** Lemmas ending by Zle *)
-(**
-<<
-Zle_left: (x,y:Z)`x <= y`->`0 <= y+(-x)`
-Zle_left_rev: (x,y:Z)`0 <= y+(-x)`->`x <= y`
-Zlt_left: (x,y:Z)`x < y`->`0 <= y+(-1)+(-x)`
-Zge_left: (x,y:Z)`x >= y`->`0 <= x+(-y)`
-Zgt_left: (x,y:Z)`x > y`->`0 <= x+(-1)+(-y)`
->>
-*)
-
-(** ** Conversion between nat comparisons and Z comparisons *)
-
-(** Lemmas ending by eq *)
-(**
-<<
-inj_eq: (x,y:nat)x=y->`(inject_nat x) = (inject_nat y)`
->>
-*)
-
-(** Lemmas ending by Zge *)
-(**
-<<
-inj_ge: (x,y:nat)(ge x y)->`(inject_nat x) >= (inject_nat y)`
->>
-*)
-
-(** Lemmas ending by Zgt *)
-(**
-<<
-inj_gt: (x,y:nat)(gt x y)->`(inject_nat x) > (inject_nat y)`
->>
-*)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-inj_lt: (x,y:nat)(lt x y)->`(inject_nat x) < (inject_nat y)`
->>
-*)
-
-(** Lemmas ending by Zle *)
-(**
-<<
-inj_le: (x,y:nat)(le x y)->`(inject_nat x) <= (inject_nat y)`
->>
-*)
-
-(** ** Conversion between comparisons *)
-
-(** Lemmas ending by Zge *)
-(**
-<<
-not_Zlt: (x,y:Z)~`x < y`->`x >= y`
-Zle_ge: (m,n:Z)`m <= n`->`n >= m`
->>
-*)
-
-(** Lemmas ending by Zgt *)
-(**
-<<
-Zle_gt_S: (n,p:Z)`n <= p`->`(Zs p) > n`
-not_Zle: (x,y:Z)~`x <= y`->`x > y`
-Zlt_gt: (m,n:Z)`m < n`->`n > m`
-Zle_S_gt: (n,m:Z)`(Zs n) <= m`->`m > n`
->>
-*)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-not_Zge: (x,y:Z)~`x >= y`->`x < y`
-Zgt_lt: (m,n:Z)`m > n`->`n < m`
-Zle_lt_n_Sm: (n,m:Z)`n <= m`->`n < (Zs m)`
->>
-*)
-
-(** Lemmas ending by Zle *)
-(**
-<<
-Zlt_ZERO_pred_le_ZERO: (x:Z)`0 < x`->`0 <= (Zpred x)`
-not_Zgt: (x,y:Z)~`x > y`->`x <= y`
-Zgt_le_S: (n,p:Z)`p > n`->`(Zs n) <= p`
-Zgt_S_le: (n,p:Z)`(Zs p) > n`->`n <= p`
-Zge_le: (m,n:Z)`m >= n`->`n <= m`
-Zlt_le_S: (n,p:Z)`n < p`->`(Zs n) <= p`
-Zlt_n_Sm_le: (n,m:Z)`n < (Zs m)`->`n <= m`
-Zlt_le_weak: (n,m:Z)`n < m`->`n <= m`
-Zle_refl: (n,m:Z)`n = m`->`n <= m`
->>
-*)
-
-(** ** Irreversible simplification involving several comparaisons *)
-(** useful with clear precedences *)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-Zlt_le_reg :(a,b,c,d:Z)`a < b`->`c <= d`->`a+c < b+d`
-Zle_lt_reg : (a,b,c,d:Z)`a <= b`->`c < d`->`a+c < b+d`
->>
-*)
-
-(** ** What is decreasing here ? *)
-
-(** Lemmas ending by eq *)
-(**
-<<
-Zplus_minus: (n,m,p:Z)`n = m+p`->`p = n-m`
->>
-*)
-
-(** Lemmas ending by Zgt *)
-(**
-<<
-Zgt_pred: (n,p:Z)`p > (Zs n)`->`(Zpred p) > n`
->>
-*)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-Zlt_pred: (n,p:Z)`(Zs n) < p`->`n < (Zpred p)`
->>
-*)
-
-(**********************************************************************)
-(** * Useful Bottom-up lemmas *)
-
-(** ** Bottom-up simplification: should be used *)
-
-(** Lemmas ending by eq *)
-(**
-<<
-Zeq_add_S: (n,m:Z)`(Zs n) = (Zs m)`->`n = m`
-Zsimpl_plus_l: (n,m,p:Z)`n+m = n+p`->`m = p`
-Zplus_unit_left: (n,m:Z)`n+0 = m`->`n = m`
-Zplus_unit_right: (n,m:Z)`n = m+0`->`n = m`
->>
-*)
-
-(** Lemmas ending by Zgt *)
-(**
-<<
-Zsimpl_gt_plus_l: (n,m,p:Z)`p+n > p+m`->`n > m`
-Zsimpl_gt_plus_r: (n,m,p:Z)`n+p > m+p`->`n > m`
-Zgt_S_n: (n,p:Z)`(Zs p) > (Zs n)`->`p > n`
->>
-*)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-Zsimpl_lt_plus_l: (n,m,p:Z)`p+n < p+m`->`n < m`
-Zsimpl_lt_plus_r: (n,m,p:Z)`n+p < m+p`->`n < m`
-Zlt_S_n: (n,m:Z)`(Zs n) < (Zs m)`->`n < m`
->>
-*)
-
-(** Lemmas ending by Zle *)
-(** << Zsimpl_le_plus_l: (p,n,m:Z)`p+n <= p+m`->`n <= m`
-Zsimpl_le_plus_r: (p,n,m:Z)`n+p <= m+p`->`n <= m`
-Zle_S_n: (n,m:Z)`(Zs m) <= (Zs n)`->`m <= n` >> *)
-
-(** ** Bottom-up irreversible (syntactic) simplification *)
-
-(** Lemmas ending by Zle *)
-(**
-<<
-Zle_trans_S: (n,m:Z)`(Zs n) <= m`->`n <= m`
->>
-*)
-
-(** ** Other unclearly simplifying lemmas *)
-
-(** Lemmas ending by Zeq *)
-(**
-<<
-Zmult_eq: (x,y:Z)`x <> 0`->`y*x = 0`->`y = 0`
->>
-*)
-
-(* Lemmas ending by Zgt *)
-(**
-<<
-Zmult_gt: (x,y:Z)`x > 0`->`x*y > 0`->`y > 0`
->>
-*)
-
-(* Lemmas ending by Zlt *)
-(**
-<<
-pZmult_lt: (x,y:Z)`x > 0`->`0 < y*x`->`0 < y`
->>
-*)
-
-(* Lemmas ending by Zle *)
-(**
-<<
-Zmult_le: (x,y:Z)`x > 0`->`0 <= y*x`->`0 <= y`
-OMEGA1: (x,y:Z)`x = y`->`0 <= x`->`0 <= y`
->>
-*)
-
-
-(**********************************************************************)
-(** * Irreversible lemmas with meta-variables *)
-(** To be used by EAuto *)
-
-(* Hints Immediate *)
-(** Lemmas ending by eq *)
-(**
-<<
-Zle_antisym: (n,m:Z)`n <= m`->`m <= n`->`n = m`
->>
-*)
-
-(** Lemmas ending by Zge *)
-(**
-<<
-Zge_trans: (n,m,p:Z)`n >= m`->`m >= p`->`n >= p`
->>
-*)
-
-(** Lemmas ending by Zgt *)
-(**
-<<
-Zgt_trans: (n,m,p:Z)`n > m`->`m > p`->`n > p`
-Zgt_trans_S: (n,m,p:Z)`(Zs n) > m`->`m > p`->`n > p`
-Zle_gt_trans: (n,m,p:Z)`m <= n`->`m > p`->`n > p`
-Zgt_le_trans: (n,m,p:Z)`n > m`->`p <= m`->`n > p`
->>
-*)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-Zlt_trans: (n,m,p:Z)`n < m`->`m < p`->`n < p`
-Zlt_le_trans: (n,m,p:Z)`n < m`->`m <= p`->`n < p`
-Zle_lt_trans: (n,m,p:Z)`n <= m`->`m < p`->`n < p`
->>
-*)
-
-(** Lemmas ending by Zle *)
-(**
-<<
-Zle_trans: (n,m,p:Z)`n <= m`->`m <= p`->`n <= p`
->>
-*)
-
-
-(**********************************************************************)
-(** * Unclear or too specific lemmas *)
-(** Not to be used ? *)
-
-(** ** Irreversible and too specific (not enough regular) *)
-
-(** Lemmas ending by Zle *)
-(**
-<<
-Zle_mult: (x,y:Z)`x > 0`->`0 <= y`->`0 <= y*x`
-Zle_mult_approx: (x,y,z:Z)`x > 0`->`z > 0`->`0 <= y`->`0 <= y*x+z`
-OMEGA6: (x,y,z:Z)`0 <= x`->`y = 0`->`0 <= x+y*z`
-OMEGA7: (x,y,z,t:Z)`z > 0`->`t > 0`->`0 <= x`->`0 <= y`->`0 <= x*z+y*t`
->>
-*)
-
-(** ** Expansion and too specific ? *)
-
-(** Lemmas ending by Zge *)
-(**
-<<
-Zge_mult_simpl: (a,b,c:Z)`c > 0`->`a*c >= b*c`->`a >= b`
->>
-*)
-
-(** Lemmas ending by Zgt *)
-(**
-<<
-Zgt_mult_simpl: (a,b,c:Z)`c > 0`->`a*c > b*c`->`a > b`
-Zgt_square_simpl: (x,y:Z)`x >= 0`->`y >= 0`->`x*x > y*y`->`x > y`
->>
-*)
-
-(** Lemmas ending by Zle *)
-(**
-<<
-Zle_mult_simpl: (a,b,c:Z)`c > 0`->`a*c <= b*c`->`a <= b`
-Zmult_le_approx: (x,y,z:Z)`x > 0`->`x > z`->`0 <= y*x+z`->`0 <= y`
->>
-*)
-
-(** ** Reversible but too specific ? *)
-
-(** Lemmas ending by Zlt *)
-(**
-<<
-Zlt_minus: (n,m:Z)`0 < m`->`n-m < n`
->>
-*)
-
-(**********************************************************************)
-(** * Lemmas to be used as rewrite rules *)
-(** but can also be used as hints *)
-
-(** Left-to-right simplification lemmas (a symbol disappears) *)
-
-(**
-<<
-Zcompare_n_S: (n,m:Z)(Zcompare (Zs n) (Zs m))=(Zcompare n m)
-Zmin_n_n: (n:Z)`(Zmin n n) = n`
-Zmult_1_n: (n:Z)`1*n = n`
-Zmult_n_1: (n:Z)`n*1 = n`
-Zminus_plus: (n,m:Z)`n+m-n = m`
-Zle_plus_minus: (n,m:Z)`n+(m-n) = m`
-Zopp_Zopp: (x:Z)`(-(-x)) = x`
-Zero_left: (x:Z)`0+x = x`
-Zero_right: (x:Z)`x+0 = x`
-Zplus_inverse_r: (x:Z)`x+(-x) = 0`
-Zplus_inverse_l: (x:Z)`(-x)+x = 0`
-Zopp_intro: (x,y:Z)`(-x) = (-y)`->`x = y`
-Zmult_one: (x:Z)`1*x = x`
-Zero_mult_left: (x:Z)`0*x = 0`
-Zero_mult_right: (x:Z)`x*0 = 0`
-Zmult_Zopp_Zopp: (x,y:Z)`(-x)*(-y) = x*y`
->>
-*)
-
-(** Right-to-left simplification lemmas (a symbol disappears) *)
-
-(**
-<<
-Zpred_Sn: (m:Z)`m = (Zpred (Zs m))`
-Zs_pred: (n:Z)`n = (Zs (Zpred n))`
-Zplus_n_O: (n:Z)`n = n+0`
-Zmult_n_O: (n:Z)`0 = n*0`
-Zminus_n_O: (n:Z)`n = n-0`
-Zminus_n_n: (n:Z)`0 = n-n`
-Zred_factor6: (x:Z)`x = x+0`
-Zred_factor0: (x:Z)`x = x*1`
->>
-*)
-
-(** Unclear orientation (no symbol disappears) *)
-
-(**
-<<
-Zplus_n_Sm: (n,m:Z)`(Zs (n+m)) = n+(Zs m)`
-Zmult_n_Sm: (n,m:Z)`n*m+n = n*(Zs m)`
-Zmin_SS: (n,m:Z)`(Zs (Zmin n m)) = (Zmin (Zs n) (Zs m))`
-Zplus_assoc_l: (n,m,p:Z)`n+(m+p) = n+m+p`
-Zplus_assoc_r: (n,m,p:Z)`n+m+p = n+(m+p)`
-Zplus_permute: (n,m,p:Z)`n+(m+p) = m+(n+p)`
-Zplus_Snm_nSm: (n,m:Z)`(Zs n)+m = n+(Zs m)`
-Zminus_plus_simpl: (n,m,p:Z)`n-m = p+n-(p+m)`
-Zminus_Sn_m: (n,m:Z)`(Zs (n-m)) = (Zs n)-m`
-Zmult_plus_distr_l: (n,m,p:Z)`(n+m)*p = n*p+m*p`
-Zmult_minus_distr: (n,m,p:Z)`(n-m)*p = n*p-m*p`
-Zmult_assoc_r: (n,m,p:Z)`n*m*p = n*(m*p)`
-Zmult_assoc_l: (n,m,p:Z)`n*(m*p) = n*m*p`
-Zmult_permute: (n,m,p:Z)`n*(m*p) = m*(n*p)`
-Zmult_Sm_n: (n,m:Z)`n*m+m = (Zs n)*m`
-Zmult_Zplus_distr: (x,y,z:Z)`x*(y+z) = x*y+x*z`
-Zmult_plus_distr: (n,m,p:Z)`(n+m)*p = n*p+m*p`
-Zopp_Zplus: (x,y:Z)`(-(x+y)) = (-x)+(-y)`
-Zplus_sym: (x,y:Z)`x+y = y+x`
-Zplus_assoc: (x,y,z:Z)`x+(y+z) = x+y+z`
-Zmult_sym: (x,y:Z)`x*y = y*x`
-Zmult_assoc: (x,y,z:Z)`x*(y*z) = x*y*z`
-Zopp_Zmult: (x,y:Z)`(-x)*y = (-(x*y))`
-Zplus_S_n: (x,y:Z)`(Zs x)+y = (Zs (x+y))`
-Zopp_one: (x:Z)`(-x) = x*(-1)`
-Zopp_Zmult_r: (x,y:Z)`(-(x*y)) = x*(-y)`
-Zmult_Zopp_left: (x,y:Z)`(-x)*y = x*(-y)`
-Zopp_Zmult_l: (x,y:Z)`(-(x*y)) = (-x)*y`
-Zred_factor1: (x:Z)`x+x = x*2`
-Zred_factor2: (x,y:Z)`x+x*y = x*(1+y)`
-Zred_factor3: (x,y:Z)`x*y+x = x*(1+y)`
-Zred_factor4: (x,y,z:Z)`x*y+x*z = x*(y+z)`
-Zminus_Zplus_compatible: (x,y,n:Z)`x+n-(y+n) = x-y`
-Zmin_plus: (x,y,n:Z)`(Zmin (x+n) (y+n)) = (Zmin x y)+n`
->>
-*)
-
-(** nat <-> Z *)
-(**
-<<
-inj_S: (y:nat)`(inject_nat (S y)) = (Zs (inject_nat y))`
-inj_plus: (x,y:nat)`(inject_nat (plus x y)) = (inject_nat x)+(inject_nat y)`
-inj_mult: (x,y:nat)`(inject_nat (mult x y)) = (inject_nat x)*(inject_nat y)`
-inj_minus1:
- (x,y:nat)(le y x)->`(inject_nat (minus x y)) = (inject_nat x)-(inject_nat y)`
-inj_minus2: (x,y:nat)(gt y x)->`(inject_nat (minus x y)) = 0`
->>
-*)
-
-(** Too specific ? *)
-(**
-<<
-Zred_factor5: (x,y:Z)`x*0+y = y`
->>
-*)
-
diff --git a/theories/ZArith/Zlogarithm.v b/theories/ZArith/Zlogarithm.v
index 59e76830..30948ca7 100644
--- a/theories/ZArith/Zlogarithm.v
+++ b/theories/ZArith/Zlogarithm.v
@@ -1,17 +1,21 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zlogarithm.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(**********************************************************************)
-(** The integer logarithms with base 2.
- There are three logarithms,
+(** The integer logarithms with base 2. *)
+
+(** THIS FILE IS DEPRECATED.
+ Please rather use [Z.log2] (or [Z.log2_up]), which
+ are defined in [BinIntDef], and whose properties can
+ be found in [BinInt.Z]. *)
+
+(* There are three logarithms defined here,
depending on the rounding of the real 2-based logarithm:
- [Log_inf]: [y = (Log_inf x) iff 2^y <= x < 2^(y+1)]
i.e. [Log_inf x] is the biggest integer that is smaller than [Log x]
@@ -20,11 +24,8 @@
- [Log_nearest]: [y= (Log_nearest x) iff 2^(y-1/2) < x <= 2^(y+1/2)]
i.e. [Log_nearest x] is the integer nearest from [Log x] *)
-Require Import ZArith_base.
-Require Import Omega.
-Require Import Zcomplements.
-Require Import Zpower.
-Open Local Scope Z_scope.
+Require Import ZArith_base Omega Zcomplements Zpower.
+Local Open Scope Z_scope.
Section Log_pos. (* Log of positive integers *)
@@ -32,9 +33,9 @@ Section Log_pos. (* Log of positive integers *)
Fixpoint log_inf (p:positive) : Z :=
match p with
- | xH => 0 (* 1 *)
- | xO q => Zsucc (log_inf q) (* 2n *)
- | xI q => Zsucc (log_inf q) (* 2n+1 *)
+ | xH => 0 (* 1 *)
+ | xO q => Zsucc (log_inf q) (* 2n *)
+ | xI q => Zsucc (log_inf q) (* 2n+1 *)
end.
Fixpoint log_sup (p:positive) : Z :=
@@ -46,6 +47,27 @@ Section Log_pos. (* Log of positive integers *)
Hint Unfold log_inf log_sup.
+ Lemma Psize_log_inf : forall p, Zpos (Pos.size p) = Z.succ (log_inf p).
+ Proof.
+ induction p; simpl; now rewrite <- ?Z.succ_Zpos, ?IHp.
+ Qed.
+
+ Lemma Zlog2_log_inf : forall p, Z.log2 (Zpos p) = log_inf p.
+ Proof.
+ unfold Z.log2. destruct p; simpl; trivial; apply Psize_log_inf.
+ Qed.
+
+ Lemma Zlog2_up_log_sup : forall p, Z.log2_up (Zpos p) = log_sup p.
+ Proof.
+ induction p; simpl.
+ - change (Zpos p~1) with (2*(Zpos p)+1).
+ rewrite Z.log2_up_succ_double, Zlog2_log_inf; try easy.
+ unfold Z.succ. now rewrite !(Z.add_comm _ 1), Z.add_assoc.
+ - change (Zpos p~0) with (2*Zpos p).
+ now rewrite Z.log2_up_double, IHp.
+ - reflexivity.
+ Qed.
+
(** Then we give the specifications of [log_inf] and [log_sup]
and prove their validity *)
diff --git a/theories/ZArith/Zmax.v b/theories/ZArith/Zmax.v
index cb2fcf26..999564f0 100644
--- a/theories/ZArith/Zmax.v
+++ b/theories/ZArith/Zmax.v
@@ -1,106 +1,111 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zmax.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-(** THIS FILE IS DEPRECATED. Use [Zminmax] instead. *)
+(** THIS FILE IS DEPRECATED. *)
-Require Export BinInt Zorder Zminmax.
+Require Export BinInt Zcompare Zorder.
-Open Local Scope Z_scope.
-
-(** [Zmax] is now [Zminmax.Zmax]. Code that do things like
- [unfold Zmin.Zmin] will have to be adapted, and neither
- a [Definition] or a [Notation] here can help much. *)
+Local Open Scope Z_scope.
+(** Definition [Zmax] is now [BinInt.Z.max]. *)
(** * Characterization of maximum on binary integer numbers *)
Definition Zmax_case := Z.max_case.
Definition Zmax_case_strong := Z.max_case_strong.
-Lemma Zmax_spec : forall x y,
- x >= y /\ Zmax x y = x \/ x < y /\ Zmax x y = y.
+Lemma Zmax_spec x y :
+ x >= y /\ Z.max x y = x \/ x < y /\ Z.max x y = y.
Proof.
- intros x y. rewrite Zge_iff_le. destruct (Z.max_spec x y); auto.
+ Z.swap_greater. destruct (Z.max_spec x y); auto.
Qed.
-Lemma Zmax_left : forall n m, n>=m -> Zmax n m = n.
-Proof. intros x y. rewrite Zge_iff_le. apply Zmax_l. Qed.
+Lemma Zmax_left n m : n>=m -> Z.max n m = n.
+Proof. Z.swap_greater. apply Zmax_l. Qed.
-Definition Zmax_right : forall n m, n<=m -> Zmax n m = m := Zmax_r.
+Lemma Zmax_right : forall n m, n<=m -> Z.max n m = m. Proof Zmax_r.
(** * Least upper bound properties of max *)
-Definition Zle_max_l : forall n m, n <= Zmax n m := Z.le_max_l.
-Definition Zle_max_r : forall n m, m <= Zmax n m := Z.le_max_r.
+Lemma Zle_max_l : forall n m, n <= Z.max n m. Proof Z.le_max_l.
+Lemma Zle_max_r : forall n m, m <= Z.max n m. Proof Z.le_max_r.
-Definition Zmax_lub : forall n m p, n <= p -> m <= p -> Zmax n m <= p
- := Z.max_lub.
+Lemma Zmax_lub : forall n m p, n <= p -> m <= p -> Z.max n m <= p.
+Proof Z.max_lub.
-Definition Zmax_lub_lt : forall n m p:Z, n < p -> m < p -> Zmax n m < p
- := Z.max_lub_lt.
+Lemma Zmax_lub_lt : forall n m p:Z, n < p -> m < p -> Z.max n m < p.
+Proof Z.max_lub_lt.
(** * Compatibility with order *)
-Definition Zle_max_compat_r : forall n m p, n <= m -> Zmax n p <= Zmax m p
- := Z.max_le_compat_r.
+Lemma Zle_max_compat_r : forall n m p, n <= m -> Z.max n p <= Z.max m p.
+Proof Z.max_le_compat_r.
-Definition Zle_max_compat_l : forall n m p, n <= m -> Zmax p n <= Zmax p m
- := Z.max_le_compat_l.
+Lemma Zle_max_compat_l : forall n m p, n <= m -> Z.max p n <= Z.max p m.
+Proof Z.max_le_compat_l.
(** * Semi-lattice properties of max *)
-Definition Zmax_idempotent : forall n, Zmax n n = n := Z.max_id.
-Definition Zmax_comm : forall n m, Zmax n m = Zmax m n := Z.max_comm.
-Definition Zmax_assoc : forall n m p, Zmax n (Zmax m p) = Zmax (Zmax n m) p
- := Z.max_assoc.
+Lemma Zmax_idempotent : forall n, Z.max n n = n. Proof Z.max_id.
+Lemma Zmax_comm : forall n m, Z.max n m = Z.max m n. Proof Z.max_comm.
+Lemma Zmax_assoc : forall n m p, Z.max n (Z.max m p) = Z.max (Z.max n m) p.
+Proof Z.max_assoc.
(** * Additional properties of max *)
-Lemma Zmax_irreducible_dec : forall n m, {Zmax n m = n} + {Zmax n m = m}.
-Proof. exact Z.max_dec. Qed.
+Lemma Zmax_irreducible_dec : forall n m, {Z.max n m = n} + {Z.max n m = m}.
+Proof Z.max_dec.
-Definition Zmax_le_prime : forall n m p, p <= Zmax n m -> p <= n \/ p <= m
- := Z.max_le.
+Lemma Zmax_le_prime : forall n m p, p <= Z.max n m -> p <= n \/ p <= m.
+Proof Z.max_le.
(** * Operations preserving max *)
-Definition Zsucc_max_distr :
- forall n m:Z, Zsucc (Zmax n m) = Zmax (Zsucc n) (Zsucc m)
- := Z.succ_max_distr.
+Lemma Zsucc_max_distr :
+ forall n m, Z.succ (Z.max n m) = Z.max (Z.succ n) (Z.succ m).
+Proof Z.succ_max_distr.
-Definition Zplus_max_distr_l : forall n m p:Z, Zmax (p + n) (p + m) = p + Zmax n m
- := Z.plus_max_distr_l.
+Lemma Zplus_max_distr_l : forall n m p, Z.max (p + n) (p + m) = p + Z.max n m.
+Proof Z.add_max_distr_l.
-Definition Zplus_max_distr_r : forall n m p:Z, Zmax (n + p) (m + p) = Zmax n m + p
- := Z.plus_max_distr_r.
+Lemma Zplus_max_distr_r : forall n m p, Z.max (n + p) (m + p) = Z.max n m + p.
+Proof Z.add_max_distr_r.
(** * Maximum and Zpos *)
-Definition Zpos_max : forall p q, Zpos (Pmax p q) = Zmax (Zpos p) (Zpos q)
- := Z.pos_max.
+Lemma Zpos_max p q : Zpos (Pos.max p q) = Z.max (Zpos p) (Zpos q).
+Proof.
+ unfold Zmax, Pmax. simpl.
+ case Pos.compare_spec; auto; congruence.
+Qed.
-Definition Zpos_max_1 : forall p, Zmax 1 (Zpos p) = Zpos p
- := Z.pos_max_1.
+Lemma Zpos_max_1 p : Z.max 1 (Zpos p) = Zpos p.
+Proof.
+ now destruct p.
+Qed.
-(** * Characterization of Pminus in term of Zminus and Zmax *)
+(** * Characterization of Pos.sub in term of Z.sub and Z.max *)
-Definition Zpos_minus :
- forall p q, Zpos (Pminus p q) = Zmax 1 (Zpos p - Zpos q)
- := Zpos_minus.
+Lemma Zpos_minus p q : Zpos (p - q) = Z.max 1 (Zpos p - Zpos q).
+Proof.
+ simpl. rewrite Z.pos_sub_spec. case Pos.compare_spec; intros H.
+ subst; now rewrite Pos.sub_diag.
+ now rewrite Pos.sub_lt.
+ symmetry. apply Zpos_max_1.
+Qed.
(* begin hide *)
(* Compatibility *)
-Notation Zmax1 := Zle_max_l (only parsing).
-Notation Zmax2 := Zle_max_r (only parsing).
-Notation Zmax_irreducible_inf := Zmax_irreducible_dec (only parsing).
-Notation Zmax_le_prime_inf := Zmax_le_prime (only parsing).
+Notation Zmax1 := Z.le_max_l (only parsing).
+Notation Zmax2 := Z.le_max_r (only parsing).
+Notation Zmax_irreducible_inf := Z.max_dec (only parsing).
+Notation Zmax_le_prime_inf := Z.max_le (only parsing).
(* end hide *)
diff --git a/theories/ZArith/Zmin.v b/theories/ZArith/Zmin.v
index 7b9ad469..2c5003a6 100644
--- a/theories/ZArith/Zmin.v
+++ b/theories/ZArith/Zmin.v
@@ -1,90 +1,95 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zmin.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-(** THIS FILE IS DEPRECATED. Use [Zminmax] instead. *)
+(** THIS FILE IS DEPRECATED. *)
-Require Import BinInt Zorder Zminmax.
+Require Import BinInt Zcompare Zorder.
-Open Local Scope Z_scope.
-
-(** [Zmin] is now [Zminmax.Zmin]. Code that do things like
- [unfold Zmin.Zmin] will have to be adapted, and neither
- a [Definition] or a [Notation] here can help much. *)
+Local Open Scope Z_scope.
+(** Definition [Zmin] is now [BinInt.Z.min]. *)
(** * Characterization of the minimum on binary integer numbers *)
Definition Zmin_case := Z.min_case.
Definition Zmin_case_strong := Z.min_case_strong.
-Lemma Zmin_spec : forall x y,
- x <= y /\ Zmin x y = x \/ x > y /\ Zmin x y = y.
+Lemma Zmin_spec x y :
+ x <= y /\ Z.min x y = x \/ x > y /\ Z.min x y = y.
Proof.
- intros x y. rewrite Zgt_iff_lt, Z.min_comm. destruct (Z.min_spec y x); auto.
+ Z.swap_greater. rewrite Z.min_comm. destruct (Z.min_spec y x); auto.
Qed.
(** * Greatest lower bound properties of min *)
-Definition Zle_min_l : forall n m, Zmin n m <= n := Z.le_min_l.
-Definition Zle_min_r : forall n m, Zmin n m <= m := Z.le_min_r.
+Lemma Zle_min_l : forall n m, Z.min n m <= n. Proof Z.le_min_l.
+Lemma Zle_min_r : forall n m, Z.min n m <= m. Proof Z.le_min_r.
-Definition Zmin_glb : forall n m p, p <= n -> p <= m -> p <= Zmin n m
- := Z.min_glb.
-Definition Zmin_glb_lt : forall n m p, p < n -> p < m -> p < Zmin n m
- := Z.min_glb_lt.
+Lemma Zmin_glb : forall n m p, p <= n -> p <= m -> p <= Z.min n m.
+Proof Z.min_glb.
+Lemma Zmin_glb_lt : forall n m p, p < n -> p < m -> p < Z.min n m.
+Proof Z.min_glb_lt.
(** * Compatibility with order *)
-Definition Zle_min_compat_r : forall n m p, n <= m -> Zmin n p <= Zmin m p
- := Z.min_le_compat_r.
-Definition Zle_min_compat_l : forall n m p, n <= m -> Zmin p n <= Zmin p m
- := Z.min_le_compat_l.
+Lemma Zle_min_compat_r : forall n m p, n <= m -> Z.min n p <= Z.min m p.
+Proof Z.min_le_compat_r.
+Lemma Zle_min_compat_l : forall n m p, n <= m -> Z.min p n <= Z.min p m.
+Proof Z.min_le_compat_l.
(** * Semi-lattice properties of min *)
-Definition Zmin_idempotent : forall n, Zmin n n = n := Z.min_id.
-Notation Zmin_n_n := Zmin_idempotent (only parsing).
-Definition Zmin_comm : forall n m, Zmin n m = Zmin m n := Z.min_comm.
-Definition Zmin_assoc : forall n m p, Zmin n (Zmin m p) = Zmin (Zmin n m) p
- := Z.min_assoc.
+Lemma Zmin_idempotent : forall n, Z.min n n = n. Proof Z.min_id.
+Notation Zmin_n_n := Z.min_id (only parsing).
+Lemma Zmin_comm : forall n m, Z.min n m = Z.min m n. Proof Z.min_comm.
+Lemma Zmin_assoc : forall n m p, Z.min n (Z.min m p) = Z.min (Z.min n m) p.
+Proof Z.min_assoc.
(** * Additional properties of min *)
-Lemma Zmin_irreducible_inf : forall n m, {Zmin n m = n} + {Zmin n m = m}.
-Proof. exact Z.min_dec. Qed.
+Lemma Zmin_irreducible_inf : forall n m, {Z.min n m = n} + {Z.min n m = m}.
+Proof Z.min_dec.
-Lemma Zmin_irreducible : forall n m, Zmin n m = n \/ Zmin n m = m.
-Proof. intros; destruct (Z.min_dec n m); auto. Qed.
+Lemma Zmin_irreducible n m : Z.min n m = n \/ Z.min n m = m.
+Proof. destruct (Z.min_dec n m); auto. Qed.
Notation Zmin_or := Zmin_irreducible (only parsing).
-Lemma Zmin_le_prime_inf : forall n m p, Zmin n m <= p -> {n <= p} + {m <= p}.
-Proof. intros n m p; apply Zmin_case; auto. Qed.
+Lemma Zmin_le_prime_inf n m p : Z.min n m <= p -> {n <= p} + {m <= p}.
+Proof. apply Zmin_case; auto. Qed.
(** * Operations preserving min *)
-Definition Zsucc_min_distr :
- forall n m, Zsucc (Zmin n m) = Zmin (Zsucc n) (Zsucc m)
- := Z.succ_min_distr.
+Lemma Zsucc_min_distr :
+ forall n m, Z.succ (Z.min n m) = Z.min (Z.succ n) (Z.succ m).
+Proof Z.succ_min_distr.
Notation Zmin_SS := Z.succ_min_distr (only parsing).
-Definition Zplus_min_distr_r :
- forall n m p, Zmin (n + p) (m + p) = Zmin n m + p
- := Z.plus_min_distr_r.
+Lemma Zplus_min_distr_r :
+ forall n m p, Z.min (n + p) (m + p) = Z.min n m + p.
+Proof Z.add_min_distr_r.
-Notation Zmin_plus := Z.plus_min_distr_r (only parsing).
+Notation Zmin_plus := Z.add_min_distr_r (only parsing).
(** * Minimum and Zpos *)
-Definition Zpos_min : forall p q, Zpos (Pmin p q) = Zmin (Zpos p) (Zpos q)
- := Z.pos_min.
+Lemma Zpos_min p q : Zpos (Pos.min p q) = Z.min (Zpos p) (Zpos q).
+Proof.
+ unfold Z.min, Pos.min; simpl. destruct Pos.compare; auto.
+Qed.
+
+Lemma Zpos_min_1 p : Z.min 1 (Zpos p) = 1.
+Proof.
+ now destruct p.
+Qed.
+
+
diff --git a/theories/ZArith/Zminmax.v b/theories/ZArith/Zminmax.v
index 5aebcc55..8908175f 100644
--- a/theories/ZArith/Zminmax.v
+++ b/theories/ZArith/Zminmax.v
@@ -1,194 +1,14 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-Require Import Orders BinInt Zcompare Zorder ZOrderedType
- GenericMinMax.
-
-(** * Maximum and Minimum of two [Z] numbers *)
-
-Local Open Scope Z_scope.
-
-Unboxed Definition Zmax (n m:Z) :=
- match n ?= m with
- | Eq | Gt => n
- | Lt => m
- end.
-
-Unboxed Definition Zmin (n m:Z) :=
- match n ?= m with
- | Eq | Lt => n
- | Gt => m
- end.
-
-(** The functions [Zmax] and [Zmin] implement indeed
- a maximum and a minimum *)
-
-Lemma Zmax_l : forall x y, y<=x -> Zmax x y = x.
-Proof.
- unfold Zle, Zmax. intros x y. rewrite <- (Zcompare_antisym x y).
- destruct (x ?= y); intuition.
-Qed.
-
-Lemma Zmax_r : forall x y, x<=y -> Zmax x y = y.
-Proof.
- unfold Zle, Zmax. intros x y. generalize (Zcompare_Eq_eq x y).
- destruct (x ?= y); intuition.
-Qed.
-
-Lemma Zmin_l : forall x y, x<=y -> Zmin x y = x.
-Proof.
- unfold Zle, Zmin. intros x y. generalize (Zcompare_Eq_eq x y).
- destruct (x ?= y); intuition.
-Qed.
-
-Lemma Zmin_r : forall x y, y<=x -> Zmin x y = y.
-Proof.
- unfold Zle, Zmin. intros x y.
- rewrite <- (Zcompare_antisym x y). generalize (Zcompare_Eq_eq x y).
- destruct (x ?= y); intuition.
-Qed.
-
-Module ZHasMinMax <: HasMinMax Z_as_OT.
- Definition max := Zmax.
- Definition min := Zmin.
- Definition max_l := Zmax_l.
- Definition max_r := Zmax_r.
- Definition min_l := Zmin_l.
- Definition min_r := Zmin_r.
-End ZHasMinMax.
-
-Module Z.
-
-(** We obtain hence all the generic properties of max and min. *)
-
-Include UsualMinMaxProperties Z_as_OT ZHasMinMax.
-
-(** * Properties specific to the [Z] domain *)
-
-(** Compatibilities (consequences of monotonicity) *)
-
-Lemma plus_max_distr_l : forall n m p, Zmax (p + n) (p + m) = p + Zmax n m.
-Proof.
- intros. apply max_monotone.
- intros x y. apply Zplus_le_compat_l.
-Qed.
-
-Lemma plus_max_distr_r : forall n m p, Zmax (n + p) (m + p) = Zmax n m + p.
-Proof.
- intros. rewrite (Zplus_comm n p), (Zplus_comm m p), (Zplus_comm _ p).
- apply plus_max_distr_l.
-Qed.
-
-Lemma plus_min_distr_l : forall n m p, Zmin (p + n) (p + m) = p + Zmin n m.
-Proof.
- intros. apply Z.min_monotone.
- intros x y. apply Zplus_le_compat_l.
-Qed.
-
-Lemma plus_min_distr_r : forall n m p, Zmin (n + p) (m + p) = Zmin n m + p.
-Proof.
- intros. rewrite (Zplus_comm n p), (Zplus_comm m p), (Zplus_comm _ p).
- apply plus_min_distr_l.
-Qed.
-
-Lemma succ_max_distr : forall n m, Zsucc (Zmax n m) = Zmax (Zsucc n) (Zsucc m).
-Proof.
- unfold Zsucc. intros. symmetry. apply plus_max_distr_r.
-Qed.
-
-Lemma succ_min_distr : forall n m, Zsucc (Zmin n m) = Zmin (Zsucc n) (Zsucc m).
-Proof.
- unfold Zsucc. intros. symmetry. apply plus_min_distr_r.
-Qed.
-
-Lemma pred_max_distr : forall n m, Zpred (Zmax n m) = Zmax (Zpred n) (Zpred m).
-Proof.
- unfold Zpred. intros. symmetry. apply plus_max_distr_r.
-Qed.
-
-Lemma pred_min_distr : forall n m, Zsucc (Zmin n m) = Zmin (Zsucc n) (Zsucc m).
-Proof.
- unfold Zpred. intros. symmetry. apply plus_min_distr_r.
-Qed.
-
-(** Anti-monotonicity swaps the role of [min] and [max] *)
-
-Lemma opp_max_distr : forall n m : Z, -(Zmax n m) = Zmin (- n) (- m).
-Proof.
- intros. symmetry. apply min_max_antimonotone.
- intros x x'. red. red. rewrite <- Zcompare_opp; auto.
-Qed.
-
-Lemma opp_min_distr : forall n m : Z, - (Zmin n m) = Zmax (- n) (- m).
-Proof.
- intros. symmetry. apply max_min_antimonotone.
- intros x x'. red. red. rewrite <- Zcompare_opp; auto.
-Qed.
-
-Lemma minus_max_distr_l : forall n m p, Zmax (p - n) (p - m) = p - Zmin n m.
-Proof.
- unfold Zminus. intros. rewrite opp_min_distr. apply plus_max_distr_l.
-Qed.
-
-Lemma minus_max_distr_r : forall n m p, Zmax (n - p) (m - p) = Zmax n m - p.
-Proof.
- unfold Zminus. intros. apply plus_max_distr_r.
-Qed.
-
-Lemma minus_min_distr_l : forall n m p, Zmin (p - n) (p - m) = p - Zmax n m.
-Proof.
- unfold Zminus. intros. rewrite opp_max_distr. apply plus_min_distr_l.
-Qed.
-
-Lemma minus_min_distr_r : forall n m p, Zmin (n - p) (m - p) = Zmin n m - p.
-Proof.
- unfold Zminus. intros. apply plus_min_distr_r.
-Qed.
-
-(** Compatibility with [Zpos] *)
-
-Lemma pos_max : forall p q, Zpos (Pmax p q) = Zmax (Zpos p) (Zpos q).
-Proof.
- intros; unfold Zmax, Pmax; simpl; generalize (Pcompare_Eq_eq p q).
- destruct Pcompare; auto.
- intro H; rewrite H; auto.
-Qed.
-
-Lemma pos_min : forall p q, Zpos (Pmin p q) = Zmin (Zpos p) (Zpos q).
-Proof.
- intros; unfold Zmin, Pmin; simpl; generalize (Pcompare_Eq_eq p q).
- destruct Pcompare; auto.
-Qed.
-
-Lemma pos_max_1 : forall p, Zmax 1 (Zpos p) = Zpos p.
-Proof.
- intros; unfold Zmax; simpl; destruct p; simpl; auto.
-Qed.
-
-Lemma pos_min_1 : forall p, Zmin 1 (Zpos p) = 1.
-Proof.
- intros; unfold Zmax; simpl; destruct p; simpl; auto.
-Qed.
-
-End Z.
-
-
-(** * Characterization of Pminus in term of Zminus and Zmax *)
-
-Lemma Zpos_minus : forall p q, Zpos (Pminus p q) = Zmax 1 (Zpos p - Zpos q).
-Proof.
- intros; simpl. destruct (Pcompare p q Eq) as [ ]_eqn:H.
- rewrite (Pcompare_Eq_eq _ _ H).
- unfold Pminus; rewrite Pminus_mask_diag; reflexivity.
- rewrite Pminus_Lt; auto.
- symmetry. apply Z.pos_max_1.
-Qed.
+Require Import Orders BinInt Zcompare Zorder.
+(** THIS FILE IS DEPRECATED. *)
(*begin hide*)
(* Compatibility with names of the old Zminmax file *)
@@ -199,4 +19,4 @@ Notation Zmin_max_distr_r := Z.min_max_distr (only parsing).
Notation Zmax_min_modular_r := Z.max_min_modular (only parsing).
Notation Zmin_max_modular_r := Z.min_max_modular (only parsing).
Notation max_min_disassoc := Z.max_min_disassoc (only parsing).
-(*end hide*) \ No newline at end of file
+(*end hide*)
diff --git a/theories/ZArith/Zmisc.v b/theories/ZArith/Zmisc.v
index a8872bd5..ff844ec2 100644
--- a/theories/ZArith/Zmisc.v
+++ b/theories/ZArith/Zmisc.v
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zmisc.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import Wf_nat.
Require Import BinInt.
Require Import Zcompare.
@@ -20,72 +18,11 @@ Open Local Scope Z_scope.
(** [n]th iteration of the function [f] *)
-Fixpoint iter_pos (n:positive) (A:Type) (f:A -> A) (x:A) : A :=
- match n with
- | xH => f x
- | xO n' => iter_pos n' A f (iter_pos n' A f x)
- | xI n' => f (iter_pos n' A f (iter_pos n' A f x))
- end.
-
-Definition iter (n:Z) (A:Type) (f:A -> A) (x:A) :=
- match n with
- | Z0 => x
- | Zpos p => iter_pos p A f x
- | Zneg p => x
- end.
-
-Theorem iter_nat_of_P :
- forall (p:positive) (A:Type) (f:A -> A) (x:A),
- iter_pos p A f x = iter_nat (nat_of_P p) A f x.
-Proof.
- intro n; induction n as [p H| p H| ];
- [ intros; simpl in |- *; rewrite (H A f x);
- rewrite (H A f (iter_nat (nat_of_P p) A f x));
- rewrite (ZL6 p); symmetry in |- *; apply f_equal with (f := f);
- apply iter_nat_plus
- | intros; unfold nat_of_P in |- *; simpl in |- *; rewrite (H A f x);
- rewrite (H A f (iter_nat (nat_of_P p) A f x));
- rewrite (ZL6 p); symmetry in |- *; apply iter_nat_plus
- | simpl in |- *; auto with arith ].
-Qed.
+Notation iter := @Z.iter (only parsing).
Lemma iter_nat_of_Z : forall n A f x, 0 <= n ->
- iter n A f x = iter_nat (Zabs_nat n) A f x.
+ iter n A f x = iter_nat (Z.abs_nat n) A f x.
intros n A f x; case n; auto.
-intros p _; unfold iter, Zabs_nat; apply iter_nat_of_P.
+intros p _; unfold Z.iter, Z.abs_nat; apply iter_nat_of_P.
intros p abs; case abs; trivial.
Qed.
-
-Theorem iter_pos_plus :
- forall (p q:positive) (A:Type) (f:A -> A) (x:A),
- iter_pos (p + q) A f x = iter_pos p A f (iter_pos q A f x).
-Proof.
- intros n m; intros.
- rewrite (iter_nat_of_P m A f x).
- rewrite (iter_nat_of_P n A f (iter_nat (nat_of_P m) A f x)).
- rewrite (iter_nat_of_P (n + m) A f x).
- rewrite (nat_of_P_plus_morphism n m).
- apply iter_nat_plus.
-Qed.
-
-(** Preservation of invariants : if [f : A->A] preserves the invariant [Inv],
- then the iterates of [f] also preserve it. *)
-
-Theorem iter_nat_invariant :
- forall (n:nat) (A:Type) (f:A -> A) (Inv:A -> Prop),
- (forall x:A, Inv x -> Inv (f x)) ->
- forall x:A, Inv x -> Inv (iter_nat n A f x).
-Proof.
- simple induction n; intros;
- [ trivial with arith
- | simpl in |- *; apply H0 with (x := iter_nat n0 A f x); apply H;
- trivial with arith ].
-Qed.
-
-Theorem iter_pos_invariant :
- forall (p:positive) (A:Type) (f:A -> A) (Inv:A -> Prop),
- (forall x:A, Inv x -> Inv (f x)) ->
- forall x:A, Inv x -> Inv (iter_pos p A f x).
-Proof.
- intros; rewrite iter_nat_of_P; apply iter_nat_invariant; trivial with arith.
-Qed.
diff --git a/theories/ZArith/Znat.v b/theories/ZArith/Znat.v
index 9585b6f6..e3843990 100644
--- a/theories/ZArith/Znat.v
+++ b/theories/ZArith/Znat.v
@@ -1,286 +1,997 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Znat.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(** Binary Integers (Pierre Crégut, CNET, Lannion, France) *)
Require Export Arith_base.
-Require Import BinPos.
-Require Import BinInt.
-Require Import Zcompare.
-Require Import Zorder.
-Require Import Decidable.
-Require Import Peano_dec.
-Require Import Min Max Zmin Zmax.
-Require Export Compare_dec.
+Require Import BinPos BinInt BinNat Pnat Nnat.
+
+Local Open Scope Z_scope.
+
+(** * Chains of conversions *)
+
+(** When combining successive conversions, we have the following
+ commutative diagram:
+<<
+ ---> Nat ----
+ | ^ |
+ | | v
+ Pos ---------> Z
+ | | ^
+ | v |
+ ----> N -----
+>>
+*)
+
+Lemma nat_N_Z n : Z.of_N (N.of_nat n) = Z.of_nat n.
+Proof.
+ now destruct n.
+Qed.
-Open Local Scope Z_scope.
+Lemma N_nat_Z n : Z.of_nat (N.to_nat n) = Z.of_N n.
+Proof.
+ destruct n; trivial. simpl.
+ destruct (Pos2Nat.is_succ p) as (m,H).
+ rewrite H. simpl. f_equal. now apply SuccNat2Pos.inv.
+Qed.
-Definition neq (x y:nat) := x <> y.
+Lemma positive_nat_Z p : Z.of_nat (Pos.to_nat p) = Zpos p.
+Proof.
+ destruct (Pos2Nat.is_succ p) as (n,H).
+ rewrite H. simpl. f_equal. now apply SuccNat2Pos.inv.
+Qed.
+
+Lemma positive_N_Z p : Z.of_N (Npos p) = Zpos p.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma positive_N_nat p : N.to_nat (Npos p) = Pos.to_nat p.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma positive_nat_N p : N.of_nat (Pos.to_nat p) = Npos p.
+Proof.
+ destruct (Pos2Nat.is_succ p) as (n,H).
+ rewrite H. simpl. f_equal. now apply SuccNat2Pos.inv.
+Qed.
-(************************************************)
-(** Properties of the injection from nat into Z *)
+Lemma Z_N_nat n : N.to_nat (Z.to_N n) = Z.to_nat n.
+Proof.
+ now destruct n.
+Qed.
-(** Injection and successor *)
+Lemma Z_nat_N n : N.of_nat (Z.to_nat n) = Z.to_N n.
+Proof.
+ destruct n; simpl; trivial. apply positive_nat_N.
+Qed.
-Theorem inj_0 : Z_of_nat 0 = 0%Z.
+Lemma Zabs_N_nat n : N.to_nat (Z.abs_N n) = Z.abs_nat n.
Proof.
- reflexivity.
+ now destruct n.
Qed.
-Theorem inj_S : forall n:nat, Z_of_nat (S n) = Zsucc (Z_of_nat n).
+Lemma Zabs_nat_N n : N.of_nat (Z.abs_nat n) = Z.abs_N n.
Proof.
- intro y; induction y as [| n H];
- [ unfold Zsucc in |- *; simpl in |- *; trivial with arith
- | change (Zpos (Psucc (P_of_succ_nat n)) = Zsucc (Z_of_nat (S n))) in |- *;
- rewrite Zpos_succ_morphism; trivial with arith ].
+ destruct n; simpl; trivial; apply positive_nat_N.
Qed.
-(** Injection and equality. *)
-Theorem inj_eq : forall n m:nat, n = m -> Z_of_nat n = Z_of_nat m.
+(** * Conversions between [Z] and [N] *)
+
+Module N2Z.
+
+(** [Z.of_N] is a bijection between [N] and non-negative [Z],
+ with [Z.to_N] (or [Z.abs_N]) as reciprocal.
+ See [Z2N.id] below for the dual equation. *)
+
+Lemma id n : Z.to_N (Z.of_N n) = n.
Proof.
- intros x y H; rewrite H; trivial with arith.
+ now destruct n.
Qed.
-Theorem inj_neq : forall n m:nat, neq n m -> Zne (Z_of_nat n) (Z_of_nat m).
+(** [Z.of_N] is hence injective *)
+
+Lemma inj n m : Z.of_N n = Z.of_N m -> n = m.
Proof.
- unfold neq, Zne, not in |- *; intros x y H1 H2; apply H1; generalize H2;
- case x; case y; intros;
- [ auto with arith
- | discriminate H0
- | discriminate H0
- | simpl in H0; injection H0;
- do 2 rewrite <- nat_of_P_o_P_of_succ_nat_eq_succ;
- intros E; rewrite E; auto with arith ].
+ destruct n, m; simpl; congruence.
Qed.
-Theorem inj_eq_rev : forall n m:nat, Z_of_nat n = Z_of_nat m -> n = m.
+Lemma inj_iff n m : Z.of_N n = Z.of_N m <-> n = m.
Proof.
- intros x y H.
- destruct (eq_nat_dec x y) as [H'|H']; auto.
- exfalso.
- exact (inj_neq _ _ H' H).
+ split. apply inj. intros; now f_equal.
Qed.
-Theorem inj_eq_iff : forall n m:nat, n=m <-> Z_of_nat n = Z_of_nat m.
+(** [Z.of_N] produce non-negative integers *)
+
+Lemma is_nonneg n : 0 <= Z.of_N n.
Proof.
- split; [apply inj_eq | apply inj_eq_rev].
+ now destruct n.
Qed.
+(** [Z.of_N], basic equations *)
+
+Lemma inj_0 : Z.of_N 0 = 0.
+Proof.
+ reflexivity.
+Qed.
-(** Injection and order relations: *)
+Lemma inj_pos p : Z.of_N (Npos p) = Zpos p.
+Proof.
+ reflexivity.
+Qed.
-(** One way ... *)
+(** [Z.of_N] and usual operations. *)
-Theorem inj_le : forall n m:nat, (n <= m)%nat -> Z_of_nat n <= Z_of_nat m.
+Lemma inj_compare n m : (Z.of_N n ?= Z.of_N m) = (n ?= m)%N.
Proof.
- intros x y; intros H; elim H;
- [ unfold Zle in |- *; elim (Zcompare_Eq_iff_eq (Z_of_nat x) (Z_of_nat x));
- intros H1 H2; rewrite H2; [ discriminate | trivial with arith ]
- | intros m H1 H2; apply Zle_trans with (Z_of_nat m);
- [ assumption | rewrite inj_S; apply Zle_succ ] ].
+ now destruct n, m.
Qed.
-Theorem inj_lt : forall n m:nat, (n < m)%nat -> Z_of_nat n < Z_of_nat m.
+Lemma inj_le n m : (n<=m)%N <-> Z.of_N n <= Z.of_N m.
Proof.
- intros x y H; apply Zgt_lt; apply Zle_succ_gt; rewrite <- inj_S; apply inj_le;
- exact H.
+ unfold Z.le. now rewrite inj_compare.
Qed.
-Theorem inj_ge : forall n m:nat, (n >= m)%nat -> Z_of_nat n >= Z_of_nat m.
+Lemma inj_lt n m : (n<m)%N <-> Z.of_N n < Z.of_N m.
Proof.
- intros x y H; apply Zle_ge; apply inj_le; apply H.
+ unfold Z.lt. now rewrite inj_compare.
Qed.
-Theorem inj_gt : forall n m:nat, (n > m)%nat -> Z_of_nat n > Z_of_nat m.
+Lemma inj_ge n m : (n>=m)%N <-> Z.of_N n >= Z.of_N m.
Proof.
- intros x y H; apply Zlt_gt; apply inj_lt; exact H.
+ unfold Z.ge. now rewrite inj_compare.
Qed.
-(** The other way ... *)
+Lemma inj_gt n m : (n>m)%N <-> Z.of_N n > Z.of_N m.
+Proof.
+ unfold Z.gt. now rewrite inj_compare.
+Qed.
-Theorem inj_le_rev : forall n m:nat, Z_of_nat n <= Z_of_nat m -> (n <= m)%nat.
+Lemma inj_abs_N z : Z.of_N (Z.abs_N z) = Z.abs z.
Proof.
- intros x y H.
- destruct (le_lt_dec x y) as [H0|H0]; auto.
- exfalso.
- assert (H1:=inj_lt _ _ H0).
- red in H; red in H1.
- rewrite <- Zcompare_antisym in H; rewrite H1 in H; auto.
+ now destruct z.
Qed.
-Theorem inj_lt_rev : forall n m:nat, Z_of_nat n < Z_of_nat m -> (n < m)%nat.
+Lemma inj_add n m : Z.of_N (n+m) = Z.of_N n + Z.of_N m.
Proof.
- intros x y H.
- destruct (le_lt_dec y x) as [H0|H0]; auto.
- exfalso.
- assert (H1:=inj_le _ _ H0).
- red in H; red in H1.
- rewrite <- Zcompare_antisym in H1; rewrite H in H1; auto.
+ now destruct n, m.
Qed.
-Theorem inj_ge_rev : forall n m:nat, Z_of_nat n >= Z_of_nat m -> (n >= m)%nat.
+Lemma inj_mul n m : Z.of_N (n*m) = Z.of_N n * Z.of_N m.
Proof.
- intros x y H.
- destruct (le_lt_dec y x) as [H0|H0]; auto.
- exfalso.
- assert (H1:=inj_gt _ _ H0).
- red in H; red in H1.
- rewrite <- Zcompare_antisym in H; rewrite H1 in H; auto.
+ now destruct n, m.
Qed.
-Theorem inj_gt_rev : forall n m:nat, Z_of_nat n > Z_of_nat m -> (n > m)%nat.
+Lemma inj_sub_max n m : Z.of_N (n-m) = Z.max 0 (Z.of_N n - Z.of_N m).
Proof.
- intros x y H.
- destruct (le_lt_dec x y) as [H0|H0]; auto.
- exfalso.
- assert (H1:=inj_ge _ _ H0).
- red in H; red in H1.
- rewrite <- Zcompare_antisym in H1; rewrite H in H1; auto.
+ destruct n as [|n], m as [|m]; simpl; trivial.
+ rewrite Z.pos_sub_spec, Pos.compare_sub_mask. unfold Pos.sub.
+ now destruct (Pos.sub_mask n m).
Qed.
-(* Both ways ... *)
+Lemma inj_sub n m : (m<=n)%N -> Z.of_N (n-m) = Z.of_N n - Z.of_N m.
+Proof.
+ intros H. rewrite inj_sub_max.
+ unfold N.le in H.
+ rewrite N.compare_antisym, <- inj_compare, Z.compare_sub in H.
+ destruct (Z.of_N n - Z.of_N m); trivial; now destruct H.
+Qed.
-Theorem inj_le_iff : forall n m:nat, (n<=m)%nat <-> Z_of_nat n <= Z_of_nat m.
+Lemma inj_succ n : Z.of_N (N.succ n) = Z.succ (Z.of_N n).
Proof.
- split; [apply inj_le | apply inj_le_rev].
+ destruct n. trivial. simpl. now rewrite Pos.add_1_r.
Qed.
-Theorem inj_lt_iff : forall n m:nat, (n<m)%nat <-> Z_of_nat n < Z_of_nat m.
+Lemma inj_pred_max n : Z.of_N (N.pred n) = Z.max 0 (Z.pred (Z.of_N n)).
Proof.
- split; [apply inj_lt | apply inj_lt_rev].
+ unfold Z.pred. now rewrite N.pred_sub, inj_sub_max.
Qed.
-Theorem inj_ge_iff : forall n m:nat, (n>=m)%nat <-> Z_of_nat n >= Z_of_nat m.
+Lemma inj_pred n : (0<n)%N -> Z.of_N (N.pred n) = Z.pred (Z.of_N n).
Proof.
- split; [apply inj_ge | apply inj_ge_rev].
+ intros H. unfold Z.pred. rewrite N.pred_sub, inj_sub; trivial.
+ now apply N.le_succ_l in H.
Qed.
-Theorem inj_gt_iff : forall n m:nat, (n>m)%nat <-> Z_of_nat n > Z_of_nat m.
+Lemma inj_min n m : Z.of_N (N.min n m) = Z.min (Z.of_N n) (Z.of_N m).
Proof.
- split; [apply inj_gt | apply inj_gt_rev].
+ unfold Z.min, N.min. rewrite inj_compare. now case N.compare.
Qed.
-(** Injection and usual operations *)
+Lemma inj_max n m : Z.of_N (N.max n m) = Z.max (Z.of_N n) (Z.of_N m).
+Proof.
+ unfold Z.max, N.max. rewrite inj_compare.
+ case N.compare_spec; intros; subst; trivial.
+Qed.
-Theorem inj_plus : forall n m:nat, Z_of_nat (n + m) = Z_of_nat n + Z_of_nat m.
+Lemma inj_div n m : Z.of_N (n/m) = Z.of_N n / Z.of_N m.
Proof.
- intro x; induction x as [| n H]; intro y; destruct y as [| m];
- [ simpl in |- *; trivial with arith
- | simpl in |- *; trivial with arith
- | simpl in |- *; rewrite <- plus_n_O; trivial with arith
- | change (Z_of_nat (S (n + S m)) = Z_of_nat (S n) + Z_of_nat (S m)) in |- *;
- rewrite inj_S; rewrite H; do 2 rewrite inj_S; rewrite Zplus_succ_l;
- trivial with arith ].
+ destruct m as [|m]. now destruct n.
+ apply Z.div_unique_pos with (Z.of_N (n mod (Npos m))).
+ split. apply is_nonneg. apply inj_lt. now apply N.mod_lt.
+ rewrite <- inj_mul, <- inj_add. f_equal. now apply N.div_mod.
Qed.
-Theorem inj_mult : forall n m:nat, Z_of_nat (n * m) = Z_of_nat n * Z_of_nat m.
+Lemma inj_mod n m : (m<>0)%N -> Z.of_N (n mod m) = (Z.of_N n) mod (Z.of_N m).
Proof.
- intro x; induction x as [| n H];
- [ simpl in |- *; trivial with arith
- | intro y; rewrite inj_S; rewrite <- Zmult_succ_l_reverse; rewrite <- H;
- rewrite <- inj_plus; simpl in |- *; rewrite plus_comm;
- trivial with arith ].
+ intros Hm.
+ apply Z.mod_unique_pos with (Z.of_N (n / m)).
+ split. apply is_nonneg. apply inj_lt. now apply N.mod_lt.
+ rewrite <- inj_mul, <- inj_add. f_equal. now apply N.div_mod.
Qed.
-Theorem inj_minus1 :
- forall n m:nat, (m <= n)%nat -> Z_of_nat (n - m) = Z_of_nat n - Z_of_nat m.
+Lemma inj_quot n m : Z.of_N (n/m) = Z.of_N n ÷ Z.of_N m.
Proof.
- intros x y H; apply (Zplus_reg_l (Z_of_nat y)); unfold Zminus in |- *;
- rewrite Zplus_permute; rewrite Zplus_opp_r; rewrite <- inj_plus;
- rewrite <- (le_plus_minus y x H); rewrite Zplus_0_r;
- trivial with arith.
+ destruct m.
+ - now destruct n.
+ - rewrite Z.quot_div_nonneg, inj_div; trivial. apply is_nonneg. easy.
Qed.
-Theorem inj_minus2 : forall n m:nat, (m > n)%nat -> Z_of_nat (n - m) = 0.
+Lemma inj_rem n m : Z.of_N (n mod m) = Z.rem (Z.of_N n) (Z.of_N m).
Proof.
- intros x y H; rewrite not_le_minus_0;
- [ trivial with arith | apply gt_not_le; assumption ].
+ destruct m.
+ - now destruct n.
+ - rewrite Z.rem_mod_nonneg, inj_mod; trivial. easy. apply is_nonneg. easy.
Qed.
-Theorem inj_minus : forall n m:nat,
- Z_of_nat (minus n m) = Zmax 0 (Z_of_nat n - Z_of_nat m).
+Lemma inj_div2 n : Z.of_N (N.div2 n) = Z.div2 (Z.of_N n).
Proof.
- intros.
- rewrite Zmax_comm.
- unfold Zmax.
- destruct (le_lt_dec m n) as [H|H].
+ destruct n as [|p]; trivial. now destruct p.
+Qed.
- rewrite (inj_minus1 _ _ H).
- assert (H':=Zle_minus_le_0 _ _ (inj_le _ _ H)).
- unfold Zle in H'.
- rewrite <- Zcompare_antisym in H'.
- destruct Zcompare; simpl in *; intuition.
+Lemma inj_quot2 n : Z.of_N (N.div2 n) = Z.quot2 (Z.of_N n).
+Proof.
+ destruct n as [|p]; trivial. now destruct p.
+Qed.
- rewrite (inj_minus2 _ _ H).
- assert (H':=Zplus_lt_compat_r _ _ (- Z_of_nat m) (inj_lt _ _ H)).
- rewrite Zplus_opp_r in H'.
- unfold Zminus; rewrite H'; auto.
+Lemma inj_pow n m : Z.of_N (n^m) = (Z.of_N n)^(Z.of_N m).
+Proof.
+ symmetry. destruct n, m; trivial. now apply Z.pow_0_l. apply Z.pow_Zpos.
Qed.
-Theorem inj_min : forall n m:nat,
- Z_of_nat (min n m) = Zmin (Z_of_nat n) (Z_of_nat m).
+End N2Z.
+
+Module Z2N.
+
+(** [Z.to_N] is a bijection between non-negative [Z] and [N],
+ with [Pos.of_N] as reciprocal.
+ See [N2Z.id] above for the dual equation. *)
+
+Lemma id n : 0<=n -> Z.of_N (Z.to_N n) = n.
Proof.
- induction n; destruct m; try (compute; auto; fail).
- simpl min.
- do 3 rewrite inj_S.
- rewrite <- Zsucc_min_distr; f_equal; auto.
+ destruct n; (now destruct 1) || trivial.
Qed.
-Theorem inj_max : forall n m:nat,
- Z_of_nat (max n m) = Zmax (Z_of_nat n) (Z_of_nat m).
+(** [Z.to_N] is hence injective for non-negative integers. *)
+
+Lemma inj n m : 0<=n -> 0<=m -> Z.to_N n = Z.to_N m -> n = m.
Proof.
- induction n; destruct m; try (compute; auto; fail).
- simpl max.
- do 3 rewrite inj_S.
- rewrite <- Zsucc_max_distr; f_equal; auto.
+ intros. rewrite <- (id n), <- (id m) by trivial. now f_equal.
Qed.
-(** Composition of injections **)
+Lemma inj_iff n m : 0<=n -> 0<=m -> (Z.to_N n = Z.to_N m <-> n = m).
+Proof.
+ intros. split. now apply inj. intros; now subst.
+Qed.
+
+(** [Z.to_N], basic equations *)
+
+Lemma inj_0 : Z.to_N 0 = 0%N.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_pos n : Z.to_N (Zpos n) = Npos n.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_neg n : Z.to_N (Zneg n) = 0%N.
+Proof.
+ reflexivity.
+Qed.
+
+(** [Z.to_N] and operations *)
+
+Lemma inj_add n m : 0<=n -> 0<=m -> Z.to_N (n+m) = (Z.to_N n + Z.to_N m)%N.
+Proof.
+ destruct n, m; trivial; (now destruct 1) || (now destruct 2).
+Qed.
+
+Lemma inj_mul n m : 0<=n -> 0<=m -> Z.to_N (n*m) = (Z.to_N n * Z.to_N m)%N.
+Proof.
+ destruct n, m; trivial; (now destruct 1) || (now destruct 2).
+Qed.
+
+Lemma inj_succ n : 0<=n -> Z.to_N (Z.succ n) = N.succ (Z.to_N n).
+Proof.
+ unfold Z.succ. intros. rewrite inj_add by easy. apply N.add_1_r.
+Qed.
+
+Lemma inj_sub n m : 0<=m -> Z.to_N (n - m) = (Z.to_N n - Z.to_N m)%N.
+Proof.
+ destruct n as [|n|n], m as [|m|m]; trivial; try (now destruct 1).
+ intros _. simpl.
+ rewrite Z.pos_sub_spec, Pos.compare_sub_mask. unfold Pos.sub.
+ now destruct (Pos.sub_mask n m).
+Qed.
+
+Lemma inj_pred n : Z.to_N (Z.pred n) = N.pred (Z.to_N n).
+Proof.
+ unfold Z.pred. rewrite <- N.sub_1_r. now apply (inj_sub n 1).
+Qed.
+
+Lemma inj_compare n m : 0<=n -> 0<=m ->
+ (Z.to_N n ?= Z.to_N m)%N = (n ?= m).
+Proof.
+ intros Hn Hm. now rewrite <- N2Z.inj_compare, !id.
+Qed.
+
+Lemma inj_le n m : 0<=n -> 0<=m -> (n<=m <-> (Z.to_N n <= Z.to_N m)%N).
+Proof.
+ intros Hn Hm. unfold Z.le, N.le. now rewrite inj_compare.
+Qed.
+
+Lemma inj_lt n m : 0<=n -> 0<=m -> (n<m <-> (Z.to_N n < Z.to_N m)%N).
+Proof.
+ intros Hn Hm. unfold Z.lt, N.lt. now rewrite inj_compare.
+Qed.
+
+Lemma inj_min n m : Z.to_N (Z.min n m) = N.min (Z.to_N n) (Z.to_N m).
+Proof.
+ destruct n, m; simpl; trivial; unfold Z.min, N.min; simpl;
+ now case Pos.compare.
+Qed.
+
+Lemma inj_max n m : Z.to_N (Z.max n m) = N.max (Z.to_N n) (Z.to_N m).
+Proof.
+ destruct n, m; simpl; trivial; unfold Z.max, N.max; simpl.
+ case Pos.compare_spec; intros; subst; trivial.
+ now case Pos.compare.
+Qed.
+
+Lemma inj_div n m : 0<=n -> 0<=m -> Z.to_N (n/m) = (Z.to_N n / Z.to_N m)%N.
+Proof.
+ destruct n, m; trivial; intros Hn Hm;
+ (now destruct Hn) || (now destruct Hm) || clear.
+ simpl. rewrite <- (N2Z.id (_ / _)). f_equal. now rewrite N2Z.inj_div.
+Qed.
+
+Lemma inj_mod n m : 0<=n -> 0<m ->
+ Z.to_N (n mod m) = ((Z.to_N n) mod (Z.to_N m))%N.
+Proof.
+ destruct n, m; trivial; intros Hn Hm;
+ (now destruct Hn) || (now destruct Hm) || clear.
+ simpl. rewrite <- (N2Z.id (_ mod _)). f_equal. now rewrite N2Z.inj_mod.
+Qed.
-Theorem Zpos_eq_Z_of_nat_o_nat_of_P :
- forall p:positive, Zpos p = Z_of_nat (nat_of_P p).
+Lemma inj_quot n m : 0<=n -> 0<=m -> Z.to_N (n÷m) = (Z.to_N n / Z.to_N m)%N.
Proof.
- intros x; elim x; simpl in |- *; auto.
- intros p H; rewrite ZL6.
- apply f_equal with (f := Zpos).
- apply nat_of_P_inj.
- rewrite nat_of_P_o_P_of_succ_nat_eq_succ; unfold nat_of_P in |- *;
- simpl in |- *.
- rewrite ZL6; auto.
- intros p H; unfold nat_of_P in |- *; simpl in |- *.
- rewrite ZL6; simpl in |- *.
- rewrite inj_plus; repeat rewrite <- H.
- rewrite Zpos_xO; simpl in |- *; rewrite Pplus_diag; reflexivity.
+ destruct m.
+ - now destruct n.
+ - intros. now rewrite Z.quot_div_nonneg, inj_div.
+ - now destruct 2.
Qed.
-(** Misc *)
+Lemma inj_rem n m :0<=n -> 0<=m ->
+ Z.to_N (Z.rem n m) = ((Z.to_N n) mod (Z.to_N m))%N.
+Proof.
+ destruct m.
+ - now destruct n.
+ - intros. now rewrite Z.rem_mod_nonneg, inj_mod.
+ - now destruct 2.
+Qed.
+
+Lemma inj_div2 n : Z.to_N (Z.div2 n) = N.div2 (Z.to_N n).
+Proof.
+ destruct n as [|p|p]; trivial. now destruct p.
+Qed.
+
+Lemma inj_quot2 n : Z.to_N (Z.quot2 n) = N.div2 (Z.to_N n).
+Proof.
+ destruct n as [|p|p]; trivial; now destruct p.
+Qed.
+
+Lemma inj_pow n m : 0<=n -> 0<=m -> Z.to_N (n^m) = ((Z.to_N n)^(Z.to_N m))%N.
+Proof.
+ destruct m.
+ - trivial.
+ - intros. now rewrite <- (N2Z.id (_ ^ _)), N2Z.inj_pow, id.
+ - now destruct 2.
+Qed.
+
+End Z2N.
+
+Module Zabs2N.
+
+(** Results about [Z.abs_N], converting absolute values of [Z] integers
+ to [N]. *)
+
+Lemma abs_N_spec n : Z.abs_N n = Z.to_N (Z.abs n).
+Proof.
+ now destruct n.
+Qed.
+
+Lemma abs_N_nonneg n : 0<=n -> Z.abs_N n = Z.to_N n.
+Proof.
+ destruct n; trivial; now destruct 1.
+Qed.
+
+Lemma id_abs n : Z.of_N (Z.abs_N n) = Z.abs n.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma id n : Z.abs_N (Z.of_N n) = n.
+Proof.
+ now destruct n.
+Qed.
+
+(** [Z.abs_N], basic equations *)
+
+Lemma inj_0 : Z.abs_N 0 = 0%N.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_pos p : Z.abs_N (Zpos p) = Npos p.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_neg p : Z.abs_N (Zneg p) = Npos p.
+Proof.
+ reflexivity.
+Qed.
+
+(** [Z.abs_N] and usual operations, with non-negative integers *)
+
+Lemma inj_opp n : Z.abs_N (-n) = Z.abs_N n.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma inj_succ n : 0<=n -> Z.abs_N (Z.succ n) = N.succ (Z.abs_N n).
+Proof.
+ intros. rewrite !abs_N_nonneg; trivial. now apply Z2N.inj_succ.
+ now apply Z.le_le_succ_r.
+Qed.
+
+Lemma inj_add n m : 0<=n -> 0<=m -> Z.abs_N (n+m) = (Z.abs_N n + Z.abs_N m)%N.
+Proof.
+ intros. rewrite !abs_N_nonneg; trivial. now apply Z2N.inj_add.
+ now apply Z.add_nonneg_nonneg.
+Qed.
+
+Lemma inj_mul n m : Z.abs_N (n*m) = (Z.abs_N n * Z.abs_N m)%N.
+Proof.
+ now destruct n, m.
+Qed.
+
+Lemma inj_sub n m : 0<=m<=n -> Z.abs_N (n-m) = (Z.abs_N n - Z.abs_N m)%N.
+Proof.
+ intros (Hn,H). rewrite !abs_N_nonneg; trivial. now apply Z2N.inj_sub.
+ Z.order.
+ now apply Z.le_0_sub.
+Qed.
+
+Lemma inj_pred n : 0<n -> Z.abs_N (Z.pred n) = N.pred (Z.abs_N n).
+Proof.
+ intros. rewrite !abs_N_nonneg. now apply Z2N.inj_pred.
+ Z.order.
+ apply Z.lt_succ_r. now rewrite Z.succ_pred.
+Qed.
+
+Lemma inj_compare n m : 0<=n -> 0<=m ->
+ (Z.abs_N n ?= Z.abs_N m)%N = (n ?= m).
+Proof.
+ intros. rewrite !abs_N_nonneg by trivial. now apply Z2N.inj_compare.
+Qed.
+
+Lemma inj_le n m : 0<=n -> 0<=m -> (n<=m <-> (Z.abs_N n <= Z.abs_N m)%N).
+Proof.
+ intros Hn Hm. unfold Z.le, N.le. now rewrite inj_compare.
+Qed.
+
+Lemma inj_lt n m : 0<=n -> 0<=m -> (n<m <-> (Z.abs_N n < Z.abs_N m)%N).
+Proof.
+ intros Hn Hm. unfold Z.lt, N.lt. now rewrite inj_compare.
+Qed.
+
+Lemma inj_min n m : 0<=n -> 0<=m ->
+ Z.abs_N (Z.min n m) = N.min (Z.abs_N n) (Z.abs_N m).
+Proof.
+ intros. rewrite !abs_N_nonneg; trivial. now apply Z2N.inj_min.
+ now apply Z.min_glb.
+Qed.
+
+Lemma inj_max n m : 0<=n -> 0<=m ->
+ Z.abs_N (Z.max n m) = N.max (Z.abs_N n) (Z.abs_N m).
+Proof.
+ intros. rewrite !abs_N_nonneg; trivial. now apply Z2N.inj_max.
+ transitivity n; trivial. apply Z.le_max_l.
+Qed.
+
+Lemma inj_quot n m : Z.abs_N (n÷m) = ((Z.abs_N n) / (Z.abs_N m))%N.
+Proof.
+ assert (forall p q, Z.abs_N (Zpos p ÷ Zpos q) = (Npos p / Npos q)%N).
+ intros. rewrite abs_N_nonneg. now apply Z2N.inj_quot. now apply Z.quot_pos.
+ destruct n, m; trivial; simpl.
+ - trivial.
+ - now rewrite <- Z.opp_Zpos, Z.quot_opp_r, inj_opp.
+ - now rewrite <- Z.opp_Zpos, Z.quot_opp_l, inj_opp.
+ - now rewrite <- 2 Z.opp_Zpos, Z.quot_opp_opp.
+Qed.
+
+Lemma inj_rem n m : Z.abs_N (Z.rem n m) = ((Z.abs_N n) mod (Z.abs_N m))%N.
+Proof.
+ assert
+ (forall p q, Z.abs_N (Z.rem (Zpos p) (Zpos q)) = ((Npos p) mod (Npos q))%N).
+ intros. rewrite abs_N_nonneg. now apply Z2N.inj_rem. now apply Z.rem_nonneg.
+ destruct n, m; trivial; simpl.
+ - trivial.
+ - now rewrite <- Z.opp_Zpos, Z.rem_opp_r.
+ - now rewrite <- Z.opp_Zpos, Z.rem_opp_l, inj_opp.
+ - now rewrite <- 2 Z.opp_Zpos, Z.rem_opp_opp, inj_opp.
+Qed.
+
+Lemma inj_pow n m : 0<=m -> Z.abs_N (n^m) = ((Z.abs_N n)^(Z.abs_N m))%N.
+Proof.
+ intros Hm. rewrite abs_N_spec, Z.abs_pow, Z2N.inj_pow, <- abs_N_spec; trivial.
+ f_equal. symmetry; now apply abs_N_nonneg. apply Z.abs_nonneg.
+Qed.
+
+(** [Z.abs_N] and usual operations, statements with [Z.abs] *)
+
+Lemma inj_succ_abs n : Z.abs_N (Z.succ (Z.abs n)) = N.succ (Z.abs_N n).
+Proof.
+ destruct n; simpl; trivial; now rewrite Pos.add_1_r.
+Qed.
+
+Lemma inj_add_abs n m :
+ Z.abs_N (Z.abs n + Z.abs m) = (Z.abs_N n + Z.abs_N m)%N.
+Proof.
+ now destruct n, m.
+Qed.
+
+Lemma inj_mul_abs n m :
+ Z.abs_N (Z.abs n * Z.abs m) = (Z.abs_N n * Z.abs_N m)%N.
+Proof.
+ now destruct n, m.
+Qed.
+
+End Zabs2N.
+
+
+(** * Conversions between [Z] and [nat] *)
+
+Module Nat2Z.
+
+(** [Z.of_nat], basic equations *)
-Theorem intro_Z :
- forall n:nat, exists y : Z, Z_of_nat n = y /\ 0 <= y * 1 + 0.
+Lemma inj_0 : Z.of_nat 0 = 0.
Proof.
- intros x; exists (Z_of_nat x); split;
- [ trivial with arith
- | rewrite Zmult_comm; rewrite Zmult_1_l; rewrite Zplus_0_r;
- unfold Zle in |- *; elim x; intros; simpl in |- *;
- discriminate ].
+ reflexivity.
Qed.
-Lemma Zpos_P_of_succ_nat : forall n:nat,
- Zpos (P_of_succ_nat n) = Zsucc (Z_of_nat n).
+Lemma inj_succ n : Z.of_nat (S n) = Z.succ (Z.of_nat n).
+Proof.
+ destruct n. trivial. simpl. symmetry. apply Z.succ_Zpos.
+Qed.
+
+(** [Z.of_N] produce non-negative integers *)
+
+Lemma is_nonneg n : 0 <= Z.of_nat n.
+Proof.
+ now induction n.
+Qed.
+
+(** [Z.of_nat] is a bijection between [nat] and non-negative [Z],
+ with [Z.to_nat] (or [Z.abs_nat]) as reciprocal.
+ See [Z2Nat.id] below for the dual equation. *)
+
+Lemma id n : Z.to_nat (Z.of_nat n) = n.
+Proof.
+ now rewrite <- nat_N_Z, <- Z_N_nat, N2Z.id, Nat2N.id.
+Qed.
+
+(** [Z.of_nat] is hence injective *)
+
+Lemma inj n m : Z.of_nat n = Z.of_nat m -> n = m.
+Proof.
+ intros H. now rewrite <- (id n), <- (id m), H.
+Qed.
+
+Lemma inj_iff n m : Z.of_nat n = Z.of_nat m <-> n = m.
+Proof.
+ split. apply inj. intros; now f_equal.
+Qed.
+
+(** [Z.of_nat] and usual operations *)
+
+Lemma inj_compare n m : (Z.of_nat n ?= Z.of_nat m) = nat_compare n m.
+Proof.
+ now rewrite <-!nat_N_Z, N2Z.inj_compare, <- Nat2N.inj_compare.
+Qed.
+
+Lemma inj_le n m : (n<=m)%nat <-> Z.of_nat n <= Z.of_nat m.
+Proof.
+ unfold Z.le. now rewrite inj_compare, nat_compare_le.
+Qed.
+
+Lemma inj_lt n m : (n<m)%nat <-> Z.of_nat n < Z.of_nat m.
+Proof.
+ unfold Z.lt. now rewrite inj_compare, nat_compare_lt.
+Qed.
+
+Lemma inj_ge n m : (n>=m)%nat <-> Z.of_nat n >= Z.of_nat m.
+Proof.
+ unfold Z.ge. now rewrite inj_compare, nat_compare_ge.
+Qed.
+
+Lemma inj_gt n m : (n>m)%nat <-> Z.of_nat n > Z.of_nat m.
+Proof.
+ unfold Z.gt. now rewrite inj_compare, nat_compare_gt.
+Qed.
+
+Lemma inj_abs_nat z : Z.of_nat (Z.abs_nat z) = Z.abs z.
+Proof.
+ destruct z; simpl; trivial;
+ destruct (Pos2Nat.is_succ p) as (n,H); rewrite H; simpl; f_equal;
+ now apply SuccNat2Pos.inv.
+Qed.
+
+Lemma inj_add n m : Z.of_nat (n+m) = Z.of_nat n + Z.of_nat m.
+Proof.
+ now rewrite <- !nat_N_Z, Nat2N.inj_add, N2Z.inj_add.
+Qed.
+
+Lemma inj_mul n m : Z.of_nat (n*m) = Z.of_nat n * Z.of_nat m.
+Proof.
+ now rewrite <- !nat_N_Z, Nat2N.inj_mul, N2Z.inj_mul.
+Qed.
+
+Lemma inj_sub_max n m : Z.of_nat (n-m) = Z.max 0 (Z.of_nat n - Z.of_nat m).
+Proof.
+ now rewrite <- !nat_N_Z, Nat2N.inj_sub, N2Z.inj_sub_max.
+Qed.
+
+Lemma inj_sub n m : (m<=n)%nat -> Z.of_nat (n-m) = Z.of_nat n - Z.of_nat m.
+Proof.
+ rewrite nat_compare_le, Nat2N.inj_compare. intros.
+ now rewrite <- !nat_N_Z, Nat2N.inj_sub, N2Z.inj_sub.
+Qed.
+
+Lemma inj_pred_max n : Z.of_nat (pred n) = Z.max 0 (Z.pred (Z.of_nat n)).
+Proof.
+ now rewrite <- !nat_N_Z, Nat2N.inj_pred, N2Z.inj_pred_max.
+Qed.
+
+Lemma inj_pred n : (0<n)%nat -> Z.of_nat (pred n) = Z.pred (Z.of_nat n).
+Proof.
+ rewrite nat_compare_lt, Nat2N.inj_compare. intros.
+ now rewrite <- !nat_N_Z, Nat2N.inj_pred, N2Z.inj_pred.
+Qed.
+
+Lemma inj_min n m : Z.of_nat (min n m) = Z.min (Z.of_nat n) (Z.of_nat m).
+Proof.
+ now rewrite <- !nat_N_Z, Nat2N.inj_min, N2Z.inj_min.
+Qed.
+
+Lemma inj_max n m : Z.of_nat (max n m) = Z.max (Z.of_nat n) (Z.of_nat m).
+Proof.
+ now rewrite <- !nat_N_Z, Nat2N.inj_max, N2Z.inj_max.
+Qed.
+
+End Nat2Z.
+
+Module Z2Nat.
+
+(** [Z.to_nat] is a bijection between non-negative [Z] and [nat],
+ with [Pos.of_nat] as reciprocal.
+ See [nat2Z.id] above for the dual equation. *)
+
+Lemma id n : 0<=n -> Z.of_nat (Z.to_nat n) = n.
+Proof.
+ intros. now rewrite <- Z_N_nat, <- nat_N_Z, N2Nat.id, Z2N.id.
+Qed.
+
+(** [Z.to_nat] is hence injective for non-negative integers. *)
+
+Lemma inj n m : 0<=n -> 0<=m -> Z.to_nat n = Z.to_nat m -> n = m.
+Proof.
+ intros. rewrite <- (id n), <- (id m) by trivial. now f_equal.
+Qed.
+
+Lemma inj_iff n m : 0<=n -> 0<=m -> (Z.to_nat n = Z.to_nat m <-> n = m).
+Proof.
+ intros. split. now apply inj. intros; now subst.
+Qed.
+
+(** [Z.to_nat], basic equations *)
+
+Lemma inj_0 : Z.to_nat 0 = O.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_pos n : Z.to_nat (Zpos n) = Pos.to_nat n.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_neg n : Z.to_nat (Zneg n) = O.
+Proof.
+ reflexivity.
+Qed.
+
+(** [Z.to_nat] and operations *)
+
+Lemma inj_add n m : 0<=n -> 0<=m ->
+ Z.to_nat (n+m) = (Z.to_nat n + Z.to_nat m)%nat.
+Proof.
+ intros. now rewrite <- !Z_N_nat, Z2N.inj_add, N2Nat.inj_add.
+Qed.
+
+Lemma inj_mul n m : 0<=n -> 0<=m ->
+ Z.to_nat (n*m) = (Z.to_nat n * Z.to_nat m)%nat.
+Proof.
+ intros. now rewrite <- !Z_N_nat, Z2N.inj_mul, N2Nat.inj_mul.
+Qed.
+
+Lemma inj_succ n : 0<=n -> Z.to_nat (Z.succ n) = S (Z.to_nat n).
+Proof.
+ intros. now rewrite <- !Z_N_nat, Z2N.inj_succ, N2Nat.inj_succ.
+Qed.
+
+Lemma inj_sub n m : 0<=m -> Z.to_nat (n - m) = (Z.to_nat n - Z.to_nat m)%nat.
+Proof.
+ intros. now rewrite <- !Z_N_nat, Z2N.inj_sub, N2Nat.inj_sub.
+Qed.
+
+Lemma inj_pred n : Z.to_nat (Z.pred n) = pred (Z.to_nat n).
+Proof.
+ now rewrite <- !Z_N_nat, Z2N.inj_pred, N2Nat.inj_pred.
+Qed.
+
+Lemma inj_compare n m : 0<=n -> 0<=m ->
+ nat_compare (Z.to_nat n) (Z.to_nat m) = (n ?= m).
+Proof.
+ intros Hn Hm. now rewrite <- Nat2Z.inj_compare, !id.
+Qed.
+
+Lemma inj_le n m : 0<=n -> 0<=m -> (n<=m <-> (Z.to_nat n <= Z.to_nat m)%nat).
+Proof.
+ intros Hn Hm. unfold Z.le. now rewrite nat_compare_le, inj_compare.
+Qed.
+
+Lemma inj_lt n m : 0<=n -> 0<=m -> (n<m <-> (Z.to_nat n < Z.to_nat m)%nat).
+Proof.
+ intros Hn Hm. unfold Z.lt. now rewrite nat_compare_lt, inj_compare.
+Qed.
+
+Lemma inj_min n m : Z.to_nat (Z.min n m) = min (Z.to_nat n) (Z.to_nat m).
+Proof.
+ now rewrite <- !Z_N_nat, Z2N.inj_min, N2Nat.inj_min.
+Qed.
+
+Lemma inj_max n m : Z.to_nat (Z.max n m) = max (Z.to_nat n) (Z.to_nat m).
+Proof.
+ now rewrite <- !Z_N_nat, Z2N.inj_max, N2Nat.inj_max.
+Qed.
+
+End Z2Nat.
+
+Module Zabs2Nat.
+
+(** Results about [Z.abs_nat], converting absolute values of [Z] integers
+ to [nat]. *)
+
+Lemma abs_nat_spec n : Z.abs_nat n = Z.to_nat (Z.abs n).
+Proof.
+ now destruct n.
+Qed.
+
+Lemma abs_nat_nonneg n : 0<=n -> Z.abs_nat n = Z.to_nat n.
+Proof.
+ destruct n; trivial; now destruct 1.
+Qed.
+
+Lemma id_abs n : Z.of_nat (Z.abs_nat n) = Z.abs n.
+Proof.
+ rewrite <-Zabs_N_nat, N_nat_Z. apply Zabs2N.id_abs.
+Qed.
+
+Lemma id n : Z.abs_nat (Z.of_nat n) = n.
+Proof.
+ now rewrite <-Zabs_N_nat, <-nat_N_Z, Zabs2N.id, Nat2N.id.
+Qed.
+
+(** [Z.abs_nat], basic equations *)
+
+Lemma inj_0 : Z.abs_nat 0 = 0%nat.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_pos p : Z.abs_nat (Zpos p) = Pos.to_nat p.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma inj_neg p : Z.abs_nat (Zneg p) = Pos.to_nat p.
+Proof.
+ reflexivity.
+Qed.
+
+(** [Z.abs_nat] and usual operations, with non-negative integers *)
+
+Lemma inj_succ n : 0<=n -> Z.abs_nat (Z.succ n) = S (Z.abs_nat n).
+Proof.
+ intros. now rewrite <- !Zabs_N_nat, Zabs2N.inj_succ, N2Nat.inj_succ.
+Qed.
+
+Lemma inj_add n m : 0<=n -> 0<=m ->
+ Z.abs_nat (n+m) = (Z.abs_nat n + Z.abs_nat m)%nat.
+Proof.
+ intros. now rewrite <- !Zabs_N_nat, Zabs2N.inj_add, N2Nat.inj_add.
+Qed.
+
+Lemma inj_mul n m : Z.abs_nat (n*m) = (Z.abs_nat n * Z.abs_nat m)%nat.
+Proof.
+ destruct n, m; simpl; trivial using Pos2Nat.inj_mul.
+Qed.
+
+Lemma inj_sub n m : 0<=m<=n ->
+ Z.abs_nat (n-m) = (Z.abs_nat n - Z.abs_nat m)%nat.
+Proof.
+ intros. now rewrite <- !Zabs_N_nat, Zabs2N.inj_sub, N2Nat.inj_sub.
+Qed.
+
+Lemma inj_pred n : 0<n -> Z.abs_nat (Z.pred n) = pred (Z.abs_nat n).
+Proof.
+ intros. now rewrite <- !Zabs_N_nat, Zabs2N.inj_pred, N2Nat.inj_pred.
+Qed.
+
+Lemma inj_compare n m : 0<=n -> 0<=m ->
+ nat_compare (Z.abs_nat n) (Z.abs_nat m) = (n ?= m).
+Proof.
+ intros. now rewrite <- !Zabs_N_nat, <- N2Nat.inj_compare, Zabs2N.inj_compare.
+Qed.
+
+Lemma inj_le n m : 0<=n -> 0<=m -> (n<=m <-> (Z.abs_nat n <= Z.abs_nat m)%nat).
+Proof.
+ intros Hn Hm. unfold Z.le. now rewrite nat_compare_le, inj_compare.
+Qed.
+
+Lemma inj_lt n m : 0<=n -> 0<=m -> (n<m <-> (Z.abs_nat n < Z.abs_nat m)%nat).
+Proof.
+ intros Hn Hm. unfold Z.lt. now rewrite nat_compare_lt, inj_compare.
+Qed.
+
+Lemma inj_min n m : 0<=n -> 0<=m ->
+ Z.abs_nat (Z.min n m) = min (Z.abs_nat n) (Z.abs_nat m).
+Proof.
+ intros. now rewrite <- !Zabs_N_nat, Zabs2N.inj_min, N2Nat.inj_min.
+Qed.
+
+Lemma inj_max n m : 0<=n -> 0<=m ->
+ Z.abs_nat (Z.max n m) = max (Z.abs_nat n) (Z.abs_nat m).
+Proof.
+ intros. now rewrite <- !Zabs_N_nat, Zabs2N.inj_max, N2Nat.inj_max.
+Qed.
+
+(** [Z.abs_nat] and usual operations, statements with [Z.abs] *)
+
+Lemma inj_succ_abs n : Z.abs_nat (Z.succ (Z.abs n)) = S (Z.abs_nat n).
+Proof.
+ now rewrite <- !Zabs_N_nat, Zabs2N.inj_succ_abs, N2Nat.inj_succ.
+Qed.
+
+Lemma inj_add_abs n m :
+ Z.abs_nat (Z.abs n + Z.abs m) = (Z.abs_nat n + Z.abs_nat m)%nat.
+Proof.
+ now rewrite <- !Zabs_N_nat, Zabs2N.inj_add_abs, N2Nat.inj_add.
+Qed.
+
+Lemma inj_mul_abs n m :
+ Z.abs_nat (Z.abs n * Z.abs m) = (Z.abs_nat n * Z.abs_nat m)%nat.
+Proof.
+ now rewrite <- !Zabs_N_nat, Zabs2N.inj_mul_abs, N2Nat.inj_mul.
+Qed.
+
+End Zabs2Nat.
+
+
+(** Compatibility *)
+
+Definition neq (x y:nat) := x <> y.
+
+Lemma inj_neq n m : neq n m -> Zne (Z_of_nat n) (Z_of_nat m).
+Proof. intros H H'. now apply H, Nat2Z.inj. Qed.
+
+Lemma Zpos_P_of_succ_nat n : Zpos (P_of_succ_nat n) = Zsucc (Z_of_nat n).
+Proof (Nat2Z.inj_succ n).
+
+(** For these one, used in omega, a Definition is necessary *)
+
+Definition inj_eq := (f_equal Z.of_nat).
+Definition inj_le n m := proj1 (Nat2Z.inj_le n m).
+Definition inj_lt n m := proj1 (Nat2Z.inj_lt n m).
+Definition inj_ge n m := proj1 (Nat2Z.inj_ge n m).
+Definition inj_gt n m := proj1 (Nat2Z.inj_gt n m).
+
+(** For the others, a Notation is fine *)
+
+Notation inj_0 := Nat2Z.inj_0 (only parsing).
+Notation inj_S := Nat2Z.inj_succ (only parsing).
+Notation inj_compare := Nat2Z.inj_compare (only parsing).
+Notation inj_eq_rev := Nat2Z.inj (only parsing).
+Notation inj_eq_iff := (fun n m => iff_sym (Nat2Z.inj_iff n m)) (only parsing).
+Notation inj_le_iff := Nat2Z.inj_le (only parsing).
+Notation inj_lt_iff := Nat2Z.inj_lt (only parsing).
+Notation inj_ge_iff := Nat2Z.inj_ge (only parsing).
+Notation inj_gt_iff := Nat2Z.inj_gt (only parsing).
+Notation inj_le_rev := (fun n m => proj2 (Nat2Z.inj_le n m)) (only parsing).
+Notation inj_lt_rev := (fun n m => proj2 (Nat2Z.inj_lt n m)) (only parsing).
+Notation inj_ge_rev := (fun n m => proj2 (Nat2Z.inj_ge n m)) (only parsing).
+Notation inj_gt_rev := (fun n m => proj2 (Nat2Z.inj_gt n m)) (only parsing).
+Notation inj_plus := Nat2Z.inj_add (only parsing).
+Notation inj_mult := Nat2Z.inj_mul (only parsing).
+Notation inj_minus1 := Nat2Z.inj_sub (only parsing).
+Notation inj_minus := Nat2Z.inj_sub_max (only parsing).
+Notation inj_min := Nat2Z.inj_min (only parsing).
+Notation inj_max := Nat2Z.inj_max (only parsing).
+
+Notation Z_of_nat_of_P := positive_nat_Z (only parsing).
+Notation Zpos_eq_Z_of_nat_o_nat_of_P :=
+ (fun p => sym_eq (positive_nat_Z p)) (only parsing).
+
+Notation Z_of_nat_of_N := N_nat_Z (only parsing).
+Notation Z_of_N_of_nat := nat_N_Z (only parsing).
+
+Notation Z_of_N_eq := (f_equal Z.of_N) (only parsing).
+Notation Z_of_N_eq_rev := N2Z.inj (only parsing).
+Notation Z_of_N_eq_iff := (fun n m => iff_sym (N2Z.inj_iff n m)) (only parsing).
+Notation Z_of_N_compare := N2Z.inj_compare (only parsing).
+Notation Z_of_N_le_iff := N2Z.inj_le (only parsing).
+Notation Z_of_N_lt_iff := N2Z.inj_lt (only parsing).
+Notation Z_of_N_ge_iff := N2Z.inj_ge (only parsing).
+Notation Z_of_N_gt_iff := N2Z.inj_gt (only parsing).
+Notation Z_of_N_le := (fun n m => proj1 (N2Z.inj_le n m)) (only parsing).
+Notation Z_of_N_lt := (fun n m => proj1 (N2Z.inj_lt n m)) (only parsing).
+Notation Z_of_N_ge := (fun n m => proj1 (N2Z.inj_ge n m)) (only parsing).
+Notation Z_of_N_gt := (fun n m => proj1 (N2Z.inj_gt n m)) (only parsing).
+Notation Z_of_N_le_rev := (fun n m => proj2 (N2Z.inj_le n m)) (only parsing).
+Notation Z_of_N_lt_rev := (fun n m => proj2 (N2Z.inj_lt n m)) (only parsing).
+Notation Z_of_N_ge_rev := (fun n m => proj2 (N2Z.inj_ge n m)) (only parsing).
+Notation Z_of_N_gt_rev := (fun n m => proj2 (N2Z.inj_gt n m)) (only parsing).
+Notation Z_of_N_pos := N2Z.inj_pos (only parsing).
+Notation Z_of_N_abs := N2Z.inj_abs_N (only parsing).
+Notation Z_of_N_le_0 := N2Z.is_nonneg (only parsing).
+Notation Z_of_N_plus := N2Z.inj_add (only parsing).
+Notation Z_of_N_mult := N2Z.inj_mul (only parsing).
+Notation Z_of_N_minus := N2Z.inj_sub_max (only parsing).
+Notation Z_of_N_succ := N2Z.inj_succ (only parsing).
+Notation Z_of_N_min := N2Z.inj_min (only parsing).
+Notation Z_of_N_max := N2Z.inj_max (only parsing).
+Notation Zabs_of_N := Zabs2N.id (only parsing).
+Notation Zabs_N_succ_abs := Zabs2N.inj_succ_abs (only parsing).
+Notation Zabs_N_succ := Zabs2N.inj_succ (only parsing).
+Notation Zabs_N_plus_abs := Zabs2N.inj_add_abs (only parsing).
+Notation Zabs_N_plus := Zabs2N.inj_add (only parsing).
+Notation Zabs_N_mult_abs := Zabs2N.inj_mul_abs (only parsing).
+Notation Zabs_N_mult := Zabs2N.inj_mul (only parsing).
+
+Theorem inj_minus2 : forall n m:nat, (m > n)%nat -> Z_of_nat (n - m) = 0.
Proof.
- intros.
- unfold Z_of_nat.
- destruct n.
- simpl; auto.
- simpl (P_of_succ_nat (S n)).
- apply Zpos_succ_morphism.
+ intros. rewrite not_le_minus_0; auto with arith.
Qed.
diff --git a/theories/ZArith/Znumtheory.v b/theories/ZArith/Znumtheory.v
index 26ff4251..6eb1a709 100644
--- a/theories/ZArith/Znumtheory.v
+++ b/theories/ZArith/Znumtheory.v
@@ -1,19 +1,20 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Znumtheory.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import ZArith_base.
Require Import ZArithRing.
Require Import Zcomplements.
Require Import Zdiv.
Require Import Wf_nat.
-Open Local Scope Z_scope.
+
+(** For compatibility reasons, this Open Scope isn't local as it should *)
+
+Open Scope Z_scope.
(** This file contains some notions of number theory upon Z numbers:
- a divisibility predicate [Zdivide]
@@ -21,308 +22,173 @@ Open Local Scope Z_scope.
- Euclid algorithm [euclid]
- a relatively prime predicate [rel_prime]
- a prime predicate [prime]
- - an efficient [Zgcd] function
+ - properties of the efficient [Z.gcd] function
*)
-(** * Divisibility *)
+Notation Zgcd := Z.gcd (only parsing).
+Notation Zggcd := Z.ggcd (only parsing).
+Notation Zggcd_gcd := Z.ggcd_gcd (only parsing).
+Notation Zggcd_correct_divisors := Z.ggcd_correct_divisors (only parsing).
+Notation Zgcd_divide_l := Z.gcd_divide_l (only parsing).
+Notation Zgcd_divide_r := Z.gcd_divide_r (only parsing).
+Notation Zgcd_greatest := Z.gcd_greatest (only parsing).
+Notation Zgcd_nonneg := Z.gcd_nonneg (only parsing).
+Notation Zggcd_opp := Z.ggcd_opp (only parsing).
-Inductive Zdivide (a b:Z) : Prop :=
- Zdivide_intro : forall q:Z, b = q * a -> Zdivide a b.
+(** The former specialized inductive predicate [Zdivide] is now
+ a generic existential predicate. *)
-(** Syntax for divisibility *)
+Notation Zdivide := Z.divide (only parsing).
-Notation "( a | b )" := (Zdivide a b) (at level 0) : Z_scope.
+(** Its former constructor is now a pseudo-constructor. *)
-(** Results concerning divisibility*)
+Definition Zdivide_intro a b q (H:b=q*a) : Z.divide a b := ex_intro _ q H.
-Lemma Zdivide_refl : forall a:Z, (a | a).
-Proof.
- intros; apply Zdivide_intro with 1; ring.
-Qed.
-
-Lemma Zone_divide : forall a:Z, (1 | a).
-Proof.
- intros; apply Zdivide_intro with a; ring.
-Qed.
-
-Lemma Zdivide_0 : forall a:Z, (a | 0).
-Proof.
- intros; apply Zdivide_intro with 0; ring.
-Qed.
-
-Hint Resolve Zdivide_refl Zone_divide Zdivide_0: zarith.
-
-Lemma Zmult_divide_compat_l : forall a b c:Z, (a | b) -> (c * a | c * b).
-Proof.
- simple induction 1; intros; apply Zdivide_intro with q.
- rewrite H0; ring.
-Qed.
-
-Lemma Zmult_divide_compat_r : forall a b c:Z, (a | b) -> (a * c | b * c).
-Proof.
- intros a b c; rewrite (Zmult_comm a c); rewrite (Zmult_comm b c).
- apply Zmult_divide_compat_l; trivial.
-Qed.
-
-Hint Resolve Zmult_divide_compat_l Zmult_divide_compat_r: zarith.
-
-Lemma Zdivide_plus_r : forall a b c:Z, (a | b) -> (a | c) -> (a | b + c).
-Proof.
- simple induction 1; intros q Hq; simple induction 1; intros q' Hq'.
- apply Zdivide_intro with (q + q').
- rewrite Hq; rewrite Hq'; ring.
-Qed.
-
-Lemma Zdivide_opp_r : forall a b:Z, (a | b) -> (a | - b).
-Proof.
- simple induction 1; intros; apply Zdivide_intro with (- q).
- rewrite H0; ring.
-Qed.
-
-Lemma Zdivide_opp_r_rev : forall a b:Z, (a | - b) -> (a | b).
-Proof.
- intros; replace b with (- - b). apply Zdivide_opp_r; trivial. ring.
-Qed.
+(** Results concerning divisibility*)
-Lemma Zdivide_opp_l : forall a b:Z, (a | b) -> (- a | b).
-Proof.
- simple induction 1; intros; apply Zdivide_intro with (- q).
- rewrite H0; ring.
-Qed.
+Notation Zdivide_refl := Z.divide_refl (only parsing).
+Notation Zone_divide := Z.divide_1_l (only parsing).
+Notation Zdivide_0 := Z.divide_0_r (only parsing).
+Notation Zmult_divide_compat_l := Z.mul_divide_mono_l (only parsing).
+Notation Zmult_divide_compat_r := Z.mul_divide_mono_r (only parsing).
+Notation Zdivide_plus_r := Z.divide_add_r (only parsing).
+Notation Zdivide_minus_l := Z.divide_sub_r (only parsing).
+Notation Zdivide_mult_l := Z.divide_mul_l (only parsing).
+Notation Zdivide_mult_r := Z.divide_mul_r (only parsing).
+Notation Zdivide_factor_r := Z.divide_factor_l (only parsing).
+Notation Zdivide_factor_l := Z.divide_factor_r (only parsing).
-Lemma Zdivide_opp_l_rev : forall a b:Z, (- a | b) -> (a | b).
-Proof.
- intros; replace a with (- - a). apply Zdivide_opp_l; trivial. ring.
-Qed.
+Lemma Zdivide_opp_r a b : (a | b) -> (a | - b).
+Proof. apply Z.divide_opp_r. Qed.
-Lemma Zdivide_minus_l : forall a b c:Z, (a | b) -> (a | c) -> (a | b - c).
-Proof.
- simple induction 1; intros q Hq; simple induction 1; intros q' Hq'.
- apply Zdivide_intro with (q - q').
- rewrite Hq; rewrite Hq'; ring.
-Qed.
+Lemma Zdivide_opp_r_rev a b : (a | - b) -> (a | b).
+Proof. apply Z.divide_opp_r. Qed.
-Lemma Zdivide_mult_l : forall a b c:Z, (a | b) -> (a | b * c).
-Proof.
- simple induction 1; intros q Hq; apply Zdivide_intro with (q * c).
- rewrite Hq; ring.
-Qed.
+Lemma Zdivide_opp_l a b : (a | b) -> (- a | b).
+Proof. apply Z.divide_opp_l. Qed.
-Lemma Zdivide_mult_r : forall a b c:Z, (a | c) -> (a | b * c).
-Proof.
- simple induction 1; intros q Hq; apply Zdivide_intro with (q * b).
- rewrite Hq; ring.
-Qed.
+Lemma Zdivide_opp_l_rev a b : (- a | b) -> (a | b).
+Proof. apply Z.divide_opp_l. Qed.
-Lemma Zdivide_factor_r : forall a b:Z, (a | a * b).
-Proof.
- intros; apply Zdivide_intro with b; ring.
-Qed.
+Theorem Zdivide_Zabs_l a b : (Z.abs a | b) -> (a | b).
+Proof. apply Z.divide_abs_l. Qed.
-Lemma Zdivide_factor_l : forall a b:Z, (a | b * a).
-Proof.
- intros; apply Zdivide_intro with b; ring.
-Qed.
+Theorem Zdivide_Zabs_inv_l a b : (a | b) -> (Z.abs a | b).
+Proof. apply Z.divide_abs_l. Qed.
+Hint Resolve Zdivide_refl Zone_divide Zdivide_0: zarith.
+Hint Resolve Zmult_divide_compat_l Zmult_divide_compat_r: zarith.
Hint Resolve Zdivide_plus_r Zdivide_opp_r Zdivide_opp_r_rev Zdivide_opp_l
Zdivide_opp_l_rev Zdivide_minus_l Zdivide_mult_l Zdivide_mult_r
Zdivide_factor_r Zdivide_factor_l: zarith.
(** Auxiliary result. *)
-Lemma Zmult_one : forall x y:Z, x >= 0 -> x * y = 1 -> x = 1.
+Lemma Zmult_one x y : x >= 0 -> x * y = 1 -> x = 1.
Proof.
- intros x y H H0; destruct (Zmult_1_inversion_l _ _ H0) as [Hpos| Hneg].
- assumption.
- rewrite Hneg in H; simpl in H.
- contradiction (Zle_not_lt 0 (-1)).
- apply Zge_le; assumption.
- apply Zorder.Zlt_neg_0.
+ Z.swap_greater. apply Z.eq_mul_1_nonneg.
Qed.
(** Only [1] and [-1] divide [1]. *)
-Lemma Zdivide_1 : forall x:Z, (x | 1) -> x = 1 \/ x = -1.
-Proof.
- simple induction 1; intros.
- elim (Z_lt_ge_dec 0 x); [ left | right ].
- apply Zmult_one with q; auto with zarith; rewrite H0; ring.
- assert (- x = 1); auto with zarith.
- apply Zmult_one with (- q); auto with zarith; rewrite H0; ring.
-Qed.
+Notation Zdivide_1 := Z.divide_1_r (only parsing).
(** If [a] divides [b] and [b] divides [a] then [a] is [b] or [-b]. *)
-Lemma Zdivide_antisym : forall a b:Z, (a | b) -> (b | a) -> a = b \/ a = - b.
-Proof.
- simple induction 1; intros.
- inversion H1.
- rewrite H0 in H2; clear H H1.
- case (Z_zerop a); intro.
- left; rewrite H0; rewrite e; ring.
- assert (Hqq0 : q0 * q = 1).
- apply Zmult_reg_l with a.
- assumption.
- ring_simplify.
- pattern a at 2 in |- *; rewrite H2; ring.
- assert (q | 1).
- rewrite <- Hqq0; auto with zarith.
- elim (Zdivide_1 q H); intros.
- rewrite H1 in H0; left; omega.
- rewrite H1 in H0; right; omega.
-Qed.
-
-Theorem Zdivide_trans: forall a b c, (a | b) -> (b | c) -> (a | c).
-Proof.
- intros a b c [d H1] [e H2]; exists (d * e); auto with zarith.
- rewrite H2; rewrite H1; ring.
-Qed.
+Notation Zdivide_antisym := Z.divide_antisym (only parsing).
+Notation Zdivide_trans := Z.divide_trans (only parsing).
(** If [a] divides [b] and [b<>0] then [|a| <= |b|]. *)
-Lemma Zdivide_bounds : forall a b:Z, (a | b) -> b <> 0 -> Zabs a <= Zabs b.
+Lemma Zdivide_bounds a b : (a | b) -> b <> 0 -> Z.abs a <= Z.abs b.
Proof.
- simple induction 1; intros.
- assert (Zabs b = Zabs q * Zabs a).
- subst; apply Zabs_Zmult.
- rewrite H2.
- assert (H3 := Zabs_pos q).
- assert (H4 := Zabs_pos a).
- assert (Zabs q * Zabs a >= 1 * Zabs a); auto with zarith.
- apply Zmult_ge_compat; auto with zarith.
- elim (Z_lt_ge_dec (Zabs q) 1); [ intros | auto with zarith ].
- assert (Zabs q = 0).
- omega.
- assert (q = 0).
- rewrite <- (Zabs_Zsgn q).
- rewrite H5; auto with zarith.
- subst q; omega.
+ intros H Hb.
+ rewrite <- Z.divide_abs_l, <- Z.divide_abs_r in H.
+ apply Z.abs_pos in Hb.
+ now apply Z.divide_pos_le.
Qed.
(** [Zdivide] can be expressed using [Zmod]. *)
Lemma Zmod_divide : forall a b, b<>0 -> a mod b = 0 -> (b | a).
Proof.
- intros a b NZ EQ.
- apply Zdivide_intro with (a/b).
- rewrite (Z_div_mod_eq_full a b NZ) at 1.
- rewrite EQ; ring.
+ apply Z.mod_divide.
Qed.
Lemma Zdivide_mod : forall a b, (b | a) -> a mod b = 0.
Proof.
- intros a b (c,->); apply Z_mod_mult.
+ intros a b (c,->); apply Z_mod_mult.
Qed.
(** [Zdivide] is hence decidable *)
-Lemma Zdivide_dec : forall a b:Z, {(a | b)} + {~ (a | b)}.
-Proof.
- intros a b; elim (Ztrichotomy_inf a 0).
- (* a<0 *)
- intros H; elim H; intros.
- case (Z_eq_dec (b mod - a) 0).
- left; apply Zdivide_opp_l_rev; apply Zmod_divide; auto with zarith.
- intro H1; right; intro; elim H1; apply Zdivide_mod; auto with zarith.
- (* a=0 *)
- case (Z_eq_dec b 0); intro.
- left; subst; auto with zarith.
- right; subst; intro H0; inversion H0; omega.
- (* a>0 *)
- intro H; case (Z_eq_dec (b mod a) 0).
- left; apply Zmod_divide; auto with zarith.
- intro H1; right; intro; elim H1; apply Zdivide_mod; auto with zarith.
-Qed.
-
-Theorem Zdivide_Zdiv_eq: forall a b : Z,
- 0 < a -> (a | b) -> b = a * (b / a).
+Lemma Zdivide_dec a b : {(a | b)} + {~ (a | b)}.
Proof.
- intros a b Hb Hc.
- pattern b at 1; rewrite (Z_div_mod_eq b a); auto with zarith.
- rewrite (Zdivide_mod b a); auto with zarith.
-Qed.
-
-Theorem Zdivide_Zdiv_eq_2: forall a b c : Z,
- 0 < a -> (a | b) -> (c * b)/a = c * (b / a).
-Proof.
- intros a b c H1 H2.
- inversion H2 as [z Hz].
- rewrite Hz; rewrite Zmult_assoc.
- repeat rewrite Z_div_mult; auto with zarith.
-Qed.
+ destruct (Z.eq_dec a 0) as [Ha|Ha].
+ destruct (Z.eq_dec b 0) as [Hb|Hb].
+ left; subst; apply Z.divide_0_r.
+ right. subst. contradict Hb. now apply Z.divide_0_l.
+ destruct (Z.eq_dec (b mod a) 0).
+ left. now apply Z.mod_divide.
+ right. now rewrite <- Z.mod_divide.
+Defined.
-Theorem Zdivide_Zabs_l: forall a b, (Zabs a | b) -> (a | b).
+Theorem Zdivide_Zdiv_eq a b : 0 < a -> (a | b) -> b = a * (b / a).
Proof.
- intros a b [x H]; subst b.
- pattern (Zabs a); apply Zabs_intro.
- exists (- x); ring.
- exists x; ring.
+ intros Ha H.
+ rewrite (Z.div_mod b a) at 1; auto with zarith.
+ rewrite Zdivide_mod; auto with zarith.
Qed.
-Theorem Zdivide_Zabs_inv_l: forall a b, (a | b) -> (Zabs a | b).
+Theorem Zdivide_Zdiv_eq_2 a b c :
+ 0 < a -> (a | b) -> (c * b) / a = c * (b / a).
Proof.
- intros a b [x H]; subst b.
- pattern (Zabs a); apply Zabs_intro.
- exists (- x); ring.
- exists x; ring.
+ intros. apply Z.divide_div_mul_exact; auto with zarith.
Qed.
Theorem Zdivide_le: forall a b : Z,
0 <= a -> 0 < b -> (a | b) -> a <= b.
Proof.
- intros a b H1 H2 [q H3]; subst b.
- case (Zle_lt_or_eq 0 a); auto with zarith; intros H3.
- case (Zle_lt_or_eq 0 q); auto with zarith.
- apply (Zmult_le_0_reg_r a); auto with zarith.
- intros H4; apply Zle_trans with (1 * a); auto with zarith.
- intros H4; subst q; omega.
+ intros. now apply Z.divide_pos_le.
Qed.
-Theorem Zdivide_Zdiv_lt_pos: forall a b : Z,
+Theorem Zdivide_Zdiv_lt_pos a b :
1 < a -> 0 < b -> (a | b) -> 0 < b / a < b .
Proof.
- intros a b H1 H2 H3; split.
- apply Zmult_lt_reg_r with a; auto with zarith.
- rewrite (Zmult_comm (Zdiv b a)); rewrite <- Zdivide_Zdiv_eq; auto with zarith.
- apply Zmult_lt_reg_r with a; auto with zarith.
- repeat rewrite (fun x => Zmult_comm x a); auto with zarith.
+ intros H1 H2 H3; split.
+ apply Z.mul_pos_cancel_l with a; auto with zarith.
rewrite <- Zdivide_Zdiv_eq; auto with zarith.
- pattern b at 1; replace b with (1 * b); auto with zarith.
- apply Zmult_lt_compat_r; auto with zarith.
+ now apply Z.div_lt.
Qed.
-Lemma Zmod_div_mod: forall n m a, 0 < n -> 0 < m ->
- (n | m) -> a mod n = (a mod m) mod n.
+Lemma Zmod_div_mod n m a:
+ 0 < n -> 0 < m -> (n | m) -> a mod n = (a mod m) mod n.
Proof.
- intros n m a H1 H2 H3.
- pattern a at 1; rewrite (Z_div_mod_eq a m); auto with zarith.
- case H3; intros q Hq; pattern m at 1; rewrite Hq.
- rewrite (Zmult_comm q).
- rewrite Zplus_mod; auto with zarith.
- rewrite <- Zmult_assoc; rewrite Zmult_mod; auto with zarith.
- rewrite Z_mod_same; try rewrite Zmult_0_l; auto with zarith.
- rewrite (Zmod_small 0); auto with zarith.
- rewrite Zplus_0_l; rewrite Zmod_mod; auto with zarith.
+ intros H1 H2 (p,Hp).
+ rewrite (Z.div_mod a m) at 1; auto with zarith.
+ rewrite Hp at 1.
+ rewrite Z.mul_shuffle0, Z.add_comm, Z.mod_add; auto with zarith.
Qed.
-Lemma Zmod_divide_minus: forall a b c : Z, 0 < b ->
- a mod b = c -> (b | a - c).
+Lemma Zmod_divide_minus a b c:
+ 0 < b -> a mod b = c -> (b | a - c).
Proof.
- intros a b c H H1; apply Zmod_divide; auto with zarith.
+ intros H H1. apply Z.mod_divide; auto with zarith.
rewrite Zminus_mod; auto with zarith.
- rewrite H1; pattern c at 1; rewrite <- (Zmod_small c b); auto with zarith.
- rewrite Zminus_diag; apply Zmod_small; auto with zarith.
- subst; apply Z_mod_lt; auto with zarith.
+ rewrite H1. rewrite <- (Z.mod_small c b) at 1.
+ rewrite Z.sub_diag, Z.mod_0_l; auto with zarith.
+ subst. now apply Z.mod_pos_bound.
Qed.
-Lemma Zdivide_mod_minus: forall a b c : Z, 0 <= c < b ->
- (b | a - c) -> a mod b = c.
+Lemma Zdivide_mod_minus a b c:
+ 0 <= c < b -> (b | a - c) -> a mod b = c.
Proof.
- intros a b c (H1, H2) H3; assert (0 < b); try apply Zle_lt_trans with c; auto.
+ intros (H1, H2) H3.
+ assert (0 < b) by Z.order.
replace a with ((a - c) + c); auto with zarith.
- rewrite Zplus_mod; auto with zarith.
- rewrite (Zdivide_mod (a -c) b); try rewrite Zplus_0_l; auto with zarith.
- rewrite Zmod_mod; try apply Zmod_small; auto with zarith.
+ rewrite Z.add_mod; auto with zarith.
+ rewrite (Zdivide_mod (a-c) b); try rewrite Z.add_0_l; auto with zarith.
+ rewrite Z.mod_mod; try apply Zmod_small; auto with zarith.
Qed.
(** * Greatest common divisor (gcd). *)
@@ -338,12 +204,12 @@ Inductive Zis_gcd (a b d:Z) : Prop :=
(** Trivial properties of [gcd] *)
-Lemma Zis_gcd_sym : forall a b d:Z, Zis_gcd a b d -> Zis_gcd b a d.
+Lemma Zis_gcd_sym : forall a b d, Zis_gcd a b d -> Zis_gcd b a d.
Proof.
- simple induction 1; constructor; intuition.
+ induction 1; constructor; intuition.
Qed.
-Lemma Zis_gcd_0 : forall a:Z, Zis_gcd a 0 a.
+Lemma Zis_gcd_0 : forall a, Zis_gcd a 0 a.
Proof.
constructor; auto with zarith.
Qed.
@@ -358,19 +224,18 @@ Proof.
constructor; auto with zarith.
Qed.
-Lemma Zis_gcd_minus : forall a b d:Z, Zis_gcd a (- b) d -> Zis_gcd b a d.
+Lemma Zis_gcd_minus : forall a b d, Zis_gcd a (- b) d -> Zis_gcd b a d.
Proof.
- simple induction 1; constructor; intuition.
+ induction 1; constructor; intuition.
Qed.
-Lemma Zis_gcd_opp : forall a b d:Z, Zis_gcd a b d -> Zis_gcd b a (- d).
+Lemma Zis_gcd_opp : forall a b d, Zis_gcd a b d -> Zis_gcd b a (- d).
Proof.
- simple induction 1; constructor; intuition.
+ induction 1; constructor; intuition.
Qed.
-Lemma Zis_gcd_0_abs : forall a:Z, Zis_gcd 0 a (Zabs a).
+Lemma Zis_gcd_0_abs a : Zis_gcd 0 a (Z.abs a).
Proof.
- intros a.
apply Zabs_ind.
intros; apply Zis_gcd_sym; apply Zis_gcd_0; auto.
intros; apply Zis_gcd_opp; apply Zis_gcd_0; auto.
@@ -429,7 +294,7 @@ Section extended_euclid_algorithm.
(** The recursive part of Euclid's algorithm uses well-founded
recursion of non-negative integers. It maintains 6 integers
[u1,u2,u3,v1,v2,v3] such that the following invariant holds:
- [u1*a+u2*b=u3] and [v1*a+v2*b=v3] and [gcd(u2,v3)=gcd(a,b)].
+ [u1*a+u2*b=u3] and [v1*a+v2*b=v3] and [gcd(u3,v3)=gcd(a,b)].
*)
Lemma euclid_rec :
@@ -519,14 +384,15 @@ Qed.
Lemma Zis_gcd_mult :
forall a b c d:Z, Zis_gcd a b d -> Zis_gcd (c * a) (c * b) (c * d).
Proof.
- intros a b c d; simple induction 1; constructor; intuition.
- elim (Zis_gcd_bezout a b d H). intros.
- elim H3; intros.
- elim H4; intros.
- apply Zdivide_intro with (u * q + v * q0).
- rewrite <- H5.
+ intros a b c d; simple induction 1. constructor; auto with zarith.
+ intros x Ha Hb.
+ elim (Zis_gcd_bezout a b d H). intros u v Huv.
+ elim Ha; intros a' Ha'.
+ elim Hb; intros b' Hb'.
+ apply Zdivide_intro with (u * a' + v * b').
+ rewrite <- Huv.
replace (c * (u * a + v * b)) with (u * (c * a) + v * (c * b)).
- rewrite H6; rewrite H7; ring.
+ rewrite Ha'; rewrite Hb'; ring.
ring.
Qed.
@@ -625,14 +491,14 @@ Proof.
exists a'; auto with zarith.
exists b'; auto with zarith.
intros x (xa,H5) (xb,H6).
- destruct (H4 (x*g)).
+ destruct (H4 (x*g)) as (x',Hx').
exists xa; rewrite Zmult_assoc; rewrite <- H5; auto.
exists xb; rewrite Zmult_assoc; rewrite <- H6; auto.
- replace g with (1*g) in H7; auto with zarith.
- do 2 rewrite Zmult_assoc in H7.
- generalize (Zmult_reg_r _ _ _ H2 H7); clear H7; intros.
- rewrite Zmult_1_r in H7.
- exists q; auto with zarith.
+ replace g with (1*g) in Hx'; auto with zarith.
+ do 2 rewrite Zmult_assoc in Hx'.
+ apply Zmult_reg_r in Hx'; trivial.
+ rewrite Zmult_1_r in Hx'.
+ exists x'; auto with zarith.
Qed.
Theorem rel_prime_sym: forall a b, rel_prime a b -> rel_prime b a.
@@ -875,249 +741,62 @@ Proof.
contradict Hp; auto with zarith.
Qed.
+(** we now prove that [Z.gcd] is indeed a gcd in
+ the sense of [Zis_gcd]. *)
-(** We could obtain a [Zgcd] function via Euclid algorithm. But we propose
- here a binary version of [Zgcd], faster and executable within Coq.
-
- Algorithm:
-
- gcd 0 b = b
- gcd a 0 = a
- gcd (2a) (2b) = 2(gcd a b)
- gcd (2a+1) (2b) = gcd (2a+1) b
- gcd (2a) (2b+1) = gcd a (2b+1)
- gcd (2a+1) (2b+1) = gcd (b-a) (2*a+1)
- or gcd (a-b) (2*b+1), depending on whether a<b
-*)
+Notation Zgcd_is_pos := Z.gcd_nonneg (only parsing).
-Open Scope positive_scope.
-
-Fixpoint Pgcdn (n: nat) (a b : positive) : positive :=
- match n with
- | O => 1
- | S n =>
- match a,b with
- | xH, _ => 1
- | _, xH => 1
- | xO a, xO b => xO (Pgcdn n a b)
- | a, xO b => Pgcdn n a b
- | xO a, b => Pgcdn n a b
- | xI a', xI b' =>
- match Pcompare a' b' Eq with
- | Eq => a
- | Lt => Pgcdn n (b'-a') a
- | Gt => Pgcdn n (a'-b') b
- end
- end
- end.
-
-Definition Pgcd (a b: positive) := Pgcdn (Psize a + Psize b)%nat a b.
-
-Close Scope positive_scope.
-
-Definition Zgcd (a b : Z) : Z :=
- match a,b with
- | Z0, _ => Zabs b
- | _, Z0 => Zabs a
- | Zpos a, Zpos b => Zpos (Pgcd a b)
- | Zpos a, Zneg b => Zpos (Pgcd a b)
- | Zneg a, Zpos b => Zpos (Pgcd a b)
- | Zneg a, Zneg b => Zpos (Pgcd a b)
- end.
-
-Lemma Zgcd_is_pos : forall a b, 0 <= Zgcd a b.
+Lemma Zgcd_is_gcd : forall a b, Zis_gcd a b (Z.gcd a b).
Proof.
- unfold Zgcd; destruct a; destruct b; auto with zarith.
-Qed.
-
-Lemma Zis_gcd_even_odd : forall a b g, Zis_gcd (Zpos a) (Zpos (xI b)) g ->
- Zis_gcd (Zpos (xO a)) (Zpos (xI b)) g.
-Proof.
- intros.
- destruct H.
- constructor; auto.
- destruct H as (e,H2); exists (2*e); auto with zarith.
- rewrite Zpos_xO; rewrite H2; ring.
- intros.
- apply H1; auto.
- rewrite Zpos_xO in H2.
- rewrite Zpos_xI in H3.
- apply Gauss with 2; auto.
- apply bezout_rel_prime.
- destruct H3 as (bb, H3).
- apply Bezout_intro with bb (-Zpos b).
- omega.
-Qed.
-
-Lemma Pgcdn_correct : forall n a b, (Psize a + Psize b<=n)%nat ->
- Zis_gcd (Zpos a) (Zpos b) (Zpos (Pgcdn n a b)).
-Proof.
- intro n; pattern n; apply lt_wf_ind; clear n; intros.
- destruct n.
- simpl.
- destruct a; simpl in *; try inversion H0.
- destruct a.
- destruct b; simpl.
- case_eq (Pcompare a b Eq); intros.
- (* a = xI, b = xI, compare = Eq *)
- rewrite (Pcompare_Eq_eq _ _ H1); apply Zis_gcd_refl.
- (* a = xI, b = xI, compare = Lt *)
- apply Zis_gcd_sym.
- apply Zis_gcd_for_euclid with 1.
- apply Zis_gcd_sym.
- replace (Zpos (xI b) - 1 * Zpos (xI a)) with (Zpos(xO (b - a))).
- apply Zis_gcd_even_odd.
- apply H; auto.
- simpl in *.
- assert (Psize (b-a) <= Psize b)%nat.
- apply Psize_monotone.
- change (Zpos (b-a) < Zpos b).
- rewrite (Zpos_minus_morphism _ _ H1).
- assert (0 < Zpos a) by (compute; auto).
- omega.
- omega.
- rewrite Zpos_xO; do 2 rewrite Zpos_xI.
- rewrite Zpos_minus_morphism; auto.
- omega.
- (* a = xI, b = xI, compare = Gt *)
- apply Zis_gcd_for_euclid with 1.
- replace (Zpos (xI a) - 1 * Zpos (xI b)) with (Zpos(xO (a - b))).
- apply Zis_gcd_sym.
- apply Zis_gcd_even_odd.
- apply H; auto.
- simpl in *.
- assert (Psize (a-b) <= Psize a)%nat.
- apply Psize_monotone.
- change (Zpos (a-b) < Zpos a).
- rewrite (Zpos_minus_morphism b a).
- assert (0 < Zpos b) by (compute; auto).
- omega.
- rewrite ZC4; rewrite H1; auto.
- omega.
- rewrite Zpos_xO; do 2 rewrite Zpos_xI.
- rewrite Zpos_minus_morphism; auto.
- omega.
- rewrite ZC4; rewrite H1; auto.
- (* a = xI, b = xO *)
- apply Zis_gcd_sym.
- apply Zis_gcd_even_odd.
- apply Zis_gcd_sym.
- apply H; auto.
- simpl in *; omega.
- (* a = xI, b = xH *)
- apply Zis_gcd_1.
- destruct b; simpl.
- (* a = xO, b = xI *)
- apply Zis_gcd_even_odd.
- apply H; auto.
- simpl in *; omega.
- (* a = xO, b = xO *)
- rewrite (Zpos_xO a); rewrite (Zpos_xO b); rewrite (Zpos_xO (Pgcdn n a b)).
- apply Zis_gcd_mult.
- apply H; auto.
- simpl in *; omega.
- (* a = xO, b = xH *)
- apply Zis_gcd_1.
- (* a = xH *)
- simpl; apply Zis_gcd_sym; apply Zis_gcd_1.
-Qed.
-
-Lemma Pgcd_correct : forall a b, Zis_gcd (Zpos a) (Zpos b) (Zpos (Pgcd a b)).
-Proof.
- unfold Pgcd; intros.
- apply Pgcdn_correct; auto.
-Qed.
-
-Lemma Zgcd_is_gcd : forall a b, Zis_gcd a b (Zgcd a b).
-Proof.
- destruct a.
- intros.
- simpl.
- apply Zis_gcd_0_abs.
- destruct b; simpl.
- apply Zis_gcd_0.
- apply Pgcd_correct.
- apply Zis_gcd_sym.
- apply Zis_gcd_minus; simpl.
- apply Pgcd_correct.
- destruct b; simpl.
- apply Zis_gcd_minus; simpl.
- apply Zis_gcd_sym.
- apply Zis_gcd_0.
- apply Zis_gcd_minus; simpl.
- apply Zis_gcd_sym.
- apply Pgcd_correct.
- apply Zis_gcd_sym.
- apply Zis_gcd_minus; simpl.
- apply Zis_gcd_minus; simpl.
- apply Zis_gcd_sym.
- apply Pgcd_correct.
+ constructor.
+ apply Z.gcd_divide_l.
+ apply Z.gcd_divide_r.
+ apply Z.gcd_greatest.
Qed.
Theorem Zgcd_spec : forall x y : Z, {z : Z | Zis_gcd x y z /\ 0 <= z}.
Proof.
- intros x y; exists (Zgcd x y).
- split; [apply Zgcd_is_gcd | apply Zgcd_is_pos].
+ intros x y; exists (Z.gcd x y).
+ split; [apply Zgcd_is_gcd | apply Z.gcd_nonneg].
Qed.
Theorem Zdivide_Zgcd: forall p q r : Z,
- (p | q) -> (p | r) -> (p | Zgcd q r).
+ (p | q) -> (p | r) -> (p | Z.gcd q r).
Proof.
- intros p q r H1 H2.
- assert (H3: (Zis_gcd q r (Zgcd q r))).
- apply Zgcd_is_gcd.
- inversion_clear H3; auto.
+ intros. now apply Z.gcd_greatest.
Qed.
Theorem Zis_gcd_gcd: forall a b c : Z,
- 0 <= c -> Zis_gcd a b c -> Zgcd a b = c.
+ 0 <= c -> Zis_gcd a b c -> Z.gcd a b = c.
Proof.
intros a b c H1 H2.
case (Zis_gcd_uniqueness_apart_sign a b c (Zgcd a b)); auto.
apply Zgcd_is_gcd; auto.
- case Zle_lt_or_eq with (1 := H1); clear H1; intros H1; subst; auto.
- intros H3; subst.
- generalize (Zgcd_is_pos a b); auto with zarith.
- case (Zgcd a b); simpl; auto; intros; discriminate.
-Qed.
-
-Theorem Zgcd_inv_0_l: forall x y, Zgcd x y = 0 -> x = 0.
-Proof.
- intros x y H.
- assert (F1: Zdivide 0 x).
- rewrite <- H.
- generalize (Zgcd_is_gcd x y); intros HH; inversion HH; auto.
- inversion F1 as [z H1].
- rewrite H1; ring.
+ Z.le_elim H1.
+ generalize (Z.gcd_nonneg a b); auto with zarith.
+ subst. now case (Z.gcd a b).
Qed.
-Theorem Zgcd_inv_0_r: forall x y, Zgcd x y = 0 -> y = 0.
-Proof.
- intros x y H.
- assert (F1: Zdivide 0 y).
- rewrite <- H.
- generalize (Zgcd_is_gcd x y); intros HH; inversion HH; auto.
- inversion F1 as [z H1].
- rewrite H1; ring.
-Qed.
+Notation Zgcd_inv_0_l := Z.gcd_eq_0_l (only parsing).
+Notation Zgcd_inv_0_r := Z.gcd_eq_0_r (only parsing).
Theorem Zgcd_div_swap0 : forall a b : Z,
- 0 < Zgcd a b ->
+ 0 < Z.gcd a b ->
0 < b ->
- (a / Zgcd a b) * b = a * (b/Zgcd a b).
+ (a / Z.gcd a b) * b = a * (b/Z.gcd a b).
Proof.
intros a b Hg Hb.
assert (F := Zgcd_is_gcd a b); inversion F as [F1 F2 F3].
- pattern b at 2; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto.
+ pattern b at 2; rewrite (Zdivide_Zdiv_eq (Z.gcd a b) b); auto.
repeat rewrite Zmult_assoc; f_equal.
rewrite Zmult_comm.
rewrite <- Zdivide_Zdiv_eq; auto.
Qed.
Theorem Zgcd_div_swap : forall a b c : Z,
- 0 < Zgcd a b ->
+ 0 < Z.gcd a b ->
0 < b ->
- (c * a) / Zgcd a b * b = c * a * (b/Zgcd a b).
+ (c * a) / Z.gcd a b * b = c * a * (b/Z.gcd a b).
Proof.
intros a b c Hg Hb.
assert (F := Zgcd_is_gcd a b); inversion F as [F1 F2 F3].
@@ -1129,44 +808,21 @@ Proof.
rewrite <- Zdivide_Zdiv_eq; auto.
Qed.
-Lemma Zgcd_comm : forall a b, Zgcd a b = Zgcd b a.
-Proof.
- intros.
- apply Zis_gcd_gcd.
- apply Zgcd_is_pos.
- apply Zis_gcd_sym.
- apply Zgcd_is_gcd.
-Qed.
-
-Lemma Zgcd_ass : forall a b c, Zgcd (Zgcd a b) c = Zgcd a (Zgcd b c).
-Proof.
- intros.
- apply Zis_gcd_gcd.
- apply Zgcd_is_pos.
- destruct (Zgcd_is_gcd a b).
- destruct (Zgcd_is_gcd b c).
- destruct (Zgcd_is_gcd a (Zgcd b c)).
- constructor; eauto using Zdivide_trans.
-Qed.
+Notation Zgcd_comm := Z.gcd_comm (only parsing).
-Lemma Zgcd_Zabs : forall a b, Zgcd (Zabs a) b = Zgcd a b.
+Lemma Zgcd_ass a b c : Zgcd (Zgcd a b) c = Zgcd a (Zgcd b c).
Proof.
- destruct a; simpl; auto.
+ symmetry. apply Z.gcd_assoc.
Qed.
-Lemma Zgcd_0 : forall a, Zgcd a 0 = Zabs a.
-Proof.
- destruct a; simpl; auto.
-Qed.
+Notation Zgcd_Zabs := Z.gcd_abs_l (only parsing).
+Notation Zgcd_0 := Z.gcd_0_r (only parsing).
+Notation Zgcd_1 := Z.gcd_1_r (only parsing).
-Lemma Zgcd_1 : forall a, Zgcd a 1 = 1.
-Proof.
- intros; apply Zis_gcd_gcd; auto with zarith; apply Zis_gcd_1.
-Qed.
Hint Resolve Zgcd_0 Zgcd_1 : zarith.
Theorem Zgcd_1_rel_prime : forall a b,
- Zgcd a b = 1 <-> rel_prime a b.
+ Z.gcd a b = 1 <-> rel_prime a b.
Proof.
unfold rel_prime; split; intro H.
rewrite <- H; apply Zgcd_is_gcd.
@@ -1249,167 +905,3 @@ Proof.
absurd (n = 0); auto with zarith.
case Hr1; auto with zarith.
Qed.
-
-(** A Generalized Gcd that also computes Bezout coefficients.
- The algorithm is the same as for Zgcd. *)
-
-Open Scope positive_scope.
-
-Fixpoint Pggcdn (n: nat) (a b : positive) : (positive*(positive*positive)) :=
- match n with
- | O => (1,(a,b))
- | S n =>
- match a,b with
- | xH, b => (1,(1,b))
- | a, xH => (1,(a,1))
- | xO a, xO b =>
- let (g,p) := Pggcdn n a b in
- (xO g,p)
- | a, xO b =>
- let (g,p) := Pggcdn n a b in
- let (aa,bb) := p in
- (g,(aa, xO bb))
- | xO a, b =>
- let (g,p) := Pggcdn n a b in
- let (aa,bb) := p in
- (g,(xO aa, bb))
- | xI a', xI b' =>
- match Pcompare a' b' Eq with
- | Eq => (a,(1,1))
- | Lt =>
- let (g,p) := Pggcdn n (b'-a') a in
- let (ba,aa) := p in
- (g,(aa, aa + xO ba))
- | Gt =>
- let (g,p) := Pggcdn n (a'-b') b in
- let (ab,bb) := p in
- (g,(bb+xO ab, bb))
- end
- end
- end.
-
-Definition Pggcd (a b: positive) := Pggcdn (Psize a + Psize b)%nat a b.
-
-Open Scope Z_scope.
-
-Definition Zggcd (a b : Z) : Z*(Z*Z) :=
- match a,b with
- | Z0, _ => (Zabs b,(0, Zsgn b))
- | _, Z0 => (Zabs a,(Zsgn a, 0))
- | Zpos a, Zpos b =>
- let (g,p) := Pggcd a b in
- let (aa,bb) := p in
- (Zpos g, (Zpos aa, Zpos bb))
- | Zpos a, Zneg b =>
- let (g,p) := Pggcd a b in
- let (aa,bb) := p in
- (Zpos g, (Zpos aa, Zneg bb))
- | Zneg a, Zpos b =>
- let (g,p) := Pggcd a b in
- let (aa,bb) := p in
- (Zpos g, (Zneg aa, Zpos bb))
- | Zneg a, Zneg b =>
- let (g,p) := Pggcd a b in
- let (aa,bb) := p in
- (Zpos g, (Zneg aa, Zneg bb))
- end.
-
-
-Lemma Pggcdn_gcdn : forall n a b,
- fst (Pggcdn n a b) = Pgcdn n a b.
-Proof.
- induction n.
- simpl; auto.
- destruct a; destruct b; simpl; auto.
- destruct (Pcompare a b Eq); simpl; auto.
- rewrite <- IHn; destruct (Pggcdn n (b-a) (xI a)) as (g,(aa,bb)); simpl; auto.
- rewrite <- IHn; destruct (Pggcdn n (a-b) (xI b)) as (g,(aa,bb)); simpl; auto.
- rewrite <- IHn; destruct (Pggcdn n (xI a) b) as (g,(aa,bb)); simpl; auto.
- rewrite <- IHn; destruct (Pggcdn n a (xI b)) as (g,(aa,bb)); simpl; auto.
- rewrite <- IHn; destruct (Pggcdn n a b) as (g,(aa,bb)); simpl; auto.
-Qed.
-
-Lemma Pggcd_gcd : forall a b, fst (Pggcd a b) = Pgcd a b.
-Proof.
- intros; exact (Pggcdn_gcdn (Psize a+Psize b)%nat a b).
-Qed.
-
-Lemma Zggcd_gcd : forall a b, fst (Zggcd a b) = Zgcd a b.
-Proof.
- destruct a; destruct b; simpl; auto; rewrite <- Pggcd_gcd;
- destruct (Pggcd p p0) as (g,(aa,bb)); simpl; auto.
-Qed.
-
-Open Scope positive_scope.
-
-Lemma Pggcdn_correct_divisors : forall n a b,
- let (g,p) := Pggcdn n a b in
- let (aa,bb):=p in
- (a=g*aa) /\ (b=g*bb).
-Proof.
- induction n.
- simpl; auto.
- destruct a; destruct b; simpl; auto.
- case_eq (Pcompare a b Eq); intros.
- (* Eq *)
- rewrite Pmult_comm; simpl; auto.
- rewrite (Pcompare_Eq_eq _ _ H); auto.
- (* Lt *)
- generalize (IHn (b-a) (xI a)); destruct (Pggcdn n (b-a) (xI a)) as (g,(ba,aa)); simpl.
- intros (H0,H1); split; auto.
- rewrite Pmult_plus_distr_l.
- rewrite Pmult_xO_permute_r.
- rewrite <- H1; rewrite <- H0.
- simpl; f_equal; symmetry.
- apply Pplus_minus; auto.
- rewrite ZC4; rewrite H; auto.
- (* Gt *)
- generalize (IHn (a-b) (xI b)); destruct (Pggcdn n (a-b) (xI b)) as (g,(ab,bb)); simpl.
- intros (H0,H1); split; auto.
- rewrite Pmult_plus_distr_l.
- rewrite Pmult_xO_permute_r.
- rewrite <- H1; rewrite <- H0.
- simpl; f_equal; symmetry.
- apply Pplus_minus; auto.
- (* Then... *)
- generalize (IHn (xI a) b); destruct (Pggcdn n (xI a) b) as (g,(ab,bb)); simpl.
- intros (H0,H1); split; auto.
- rewrite Pmult_xO_permute_r; rewrite H1; auto.
- generalize (IHn a (xI b)); destruct (Pggcdn n a (xI b)) as (g,(ab,bb)); simpl.
- intros (H0,H1); split; auto.
- rewrite Pmult_xO_permute_r; rewrite H0; auto.
- generalize (IHn a b); destruct (Pggcdn n a b) as (g,(ab,bb)); simpl.
- intros (H0,H1); split; subst; auto.
-Qed.
-
-Lemma Pggcd_correct_divisors : forall a b,
- let (g,p) := Pggcd a b in
- let (aa,bb):=p in
- (a=g*aa) /\ (b=g*bb).
-Proof.
- intros a b; exact (Pggcdn_correct_divisors (Psize a + Psize b)%nat a b).
-Qed.
-
-Close Scope positive_scope.
-
-Lemma Zggcd_correct_divisors : forall a b,
- let (g,p) := Zggcd a b in
- let (aa,bb):=p in
- (a=g*aa) /\ (b=g*bb).
-Proof.
- destruct a; destruct b; simpl; auto; try solve [rewrite Pmult_comm; simpl; auto];
- generalize (Pggcd_correct_divisors p p0); destruct (Pggcd p p0) as (g,(aa,bb));
- destruct 1; subst; auto.
-Qed.
-
-Theorem Zggcd_opp: forall x y,
- Zggcd (-x) y = let (p1,p) := Zggcd x y in
- let (p2,p3) := p in
- (p1,(-p2,p3)).
-Proof.
-intros [|x|x] [|y|y]; unfold Zggcd, Zopp; auto.
-case Pggcd; intros p1 (p2, p3); auto.
-case Pggcd; intros p1 (p2, p3); auto.
-case Pggcd; intros p1 (p2, p3); auto.
-case Pggcd; intros p1 (p2, p3); auto.
-Qed.
diff --git a/theories/ZArith/Zorder.v b/theories/ZArith/Zorder.v
index 91c16929..a8cd69bb 100644
--- a/theories/ZArith/Zorder.v
+++ b/theories/ZArith/Zorder.v
@@ -1,337 +1,200 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zorder.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-(** Binary Integers (Pierre Crégut, CNET, Lannion, France) *)
+(** Binary Integers : results about order predicates *)
+(** Initial author : Pierre Crégut (CNET, Lannion, France) *)
-Require Import BinPos.
-Require Import BinInt.
-Require Import Arith_base.
-Require Import Decidable.
-Require Import Zcompare.
+(** THIS FILE IS DEPRECATED.
+ It is now almost entirely made of compatibility formulations
+ for results already present in BinInt.Z. *)
-Open Local Scope Z_scope.
+Require Import BinPos BinInt Decidable Zcompare.
+Require Import Arith_base. (* Useless now, for compatibility only *)
-Implicit Types x y z : Z.
+Local Open Scope Z_scope.
(*********************************************************)
(** Properties of the order relations on binary integers *)
(** * Trichotomy *)
-Theorem Ztrichotomy_inf : forall n m:Z, {n < m} + {n = m} + {n > m}.
+Theorem Ztrichotomy_inf n m : {n < m} + {n = m} + {n > m}.
Proof.
- unfold Zgt, Zlt in |- *; intros m n; assert (H := refl_equal (m ?= n)).
- set (x := m ?= n) in H at 2 |- *.
- destruct x;
- [ left; right; rewrite Zcompare_Eq_eq with (1 := H) | left; left | right ];
- reflexivity.
-Qed.
+ unfold ">", "<". generalize (Z.compare_eq n m).
+ destruct (n ?= m); [ left; right | left; left | right]; auto.
+Defined.
-Theorem Ztrichotomy : forall n m:Z, n < m \/ n = m \/ n > m.
+Theorem Ztrichotomy n m : n < m \/ n = m \/ n > m.
Proof.
- intros m n; destruct (Ztrichotomy_inf m n) as [[Hlt| Heq]| Hgt];
- [ left | right; left | right; right ]; assumption.
+ Z.swap_greater. apply Z.lt_trichotomy.
Qed.
(**********************************************************************)
(** * Decidability of equality and order on Z *)
-Theorem dec_eq : forall n m:Z, decidable (n = m).
-Proof.
- intros x y; unfold decidable in |- *; elim (Zcompare_Eq_iff_eq x y);
- intros H1 H2; elim (Dcompare (x ?= y));
- [ tauto
- | intros H3; right; unfold not in |- *; intros H4; elim H3; rewrite (H2 H4);
- intros H5; discriminate H5 ].
-Qed.
-
-Theorem dec_Zne : forall n m:Z, decidable (Zne n m).
-Proof.
- intros x y; unfold decidable, Zne in |- *; elim (Zcompare_Eq_iff_eq x y).
- intros H1 H2; elim (Dcompare (x ?= y));
- [ right; rewrite H1; auto
- | left; unfold not in |- *; intro; absurd ((x ?= y) = Eq);
- [ elim H; intros HR; rewrite HR; discriminate | auto ] ].
-Qed.
-
-Theorem dec_Zle : forall n m:Z, decidable (n <= m).
-Proof.
- intros x y; unfold decidable, Zle in |- *; elim (x ?= y);
- [ left; discriminate
- | left; discriminate
- | right; unfold not in |- *; intros H; apply H; trivial with arith ].
-Qed.
+Notation dec_eq := Z.eq_decidable (only parsing).
+Notation dec_Zle := Z.le_decidable (only parsing).
+Notation dec_Zlt := Z.lt_decidable (only parsing).
-Theorem dec_Zgt : forall n m:Z, decidable (n > m).
+Theorem dec_Zne n m : decidable (Zne n m).
Proof.
- intros x y; unfold decidable, Zgt in |- *; elim (x ?= y);
- [ right; discriminate | right; discriminate | auto with arith ].
+ destruct (Z.eq_decidable n m); [right|left]; subst; auto.
Qed.
-Theorem dec_Zge : forall n m:Z, decidable (n >= m).
+Theorem dec_Zgt n m : decidable (n > m).
Proof.
- intros x y; unfold decidable, Zge in |- *; elim (x ?= y);
- [ left; discriminate
- | right; unfold not in |- *; intros H; apply H; trivial with arith
- | left; discriminate ].
+ destruct (Z.lt_decidable m n); [left|right]; Z.swap_greater; auto.
Qed.
-Theorem dec_Zlt : forall n m:Z, decidable (n < m).
+Theorem dec_Zge n m : decidable (n >= m).
Proof.
- intros x y; unfold decidable, Zlt in |- *; elim (x ?= y);
- [ right; discriminate | auto with arith | right; discriminate ].
+ destruct (Z.le_decidable m n); [left|right]; Z.swap_greater; auto.
Qed.
-Theorem not_Zeq : forall n m:Z, n <> m -> n < m \/ m < n.
+Theorem not_Zeq n m : n <> m -> n < m \/ m < n.
Proof.
- intros x y; elim (Dcompare (x ?= y));
- [ intros H1 H2; absurd (x = y);
- [ assumption | elim (Zcompare_Eq_iff_eq x y); auto with arith ]
- | unfold Zlt in |- *; intros H; elim H; intros H1;
- [ auto with arith
- | right; elim (Zcompare_Gt_Lt_antisym x y); auto with arith ] ].
+ apply Z.lt_gt_cases.
Qed.
(** * Relating strict and large orders *)
-Lemma Zgt_lt : forall n m:Z, n > m -> m < n.
-Proof.
- unfold Zgt, Zlt in |- *; intros m n H; elim (Zcompare_Gt_Lt_antisym m n);
- auto with arith.
-Qed.
-
-Lemma Zlt_gt : forall n m:Z, n < m -> m > n.
-Proof.
- unfold Zgt, Zlt in |- *; intros m n H; elim (Zcompare_Gt_Lt_antisym n m);
- auto with arith.
-Qed.
+Notation Zgt_lt := Z.gt_lt (only parsing).
+Notation Zlt_gt := Z.lt_gt (only parsing).
+Notation Zge_le := Z.ge_le (only parsing).
+Notation Zle_ge := Z.le_ge (only parsing).
+Notation Zgt_iff_lt := Z.gt_lt_iff (only parsing).
+Notation Zge_iff_le := Z.ge_le_iff (only parsing).
-Lemma Zge_le : forall n m:Z, n >= m -> m <= n.
+Lemma Zle_not_lt n m : n <= m -> ~ m < n.
Proof.
- intros m n; change (~ m < n -> ~ n > m) in |- *; unfold not in |- *;
- intros H1 H2; apply H1; apply Zgt_lt; assumption.
+ apply Z.le_ngt.
Qed.
-Lemma Zle_ge : forall n m:Z, n <= m -> m >= n.
+Lemma Zlt_not_le n m : n < m -> ~ m <= n.
Proof.
- intros m n; change (~ m > n -> ~ n < m) in |- *; unfold not in |- *;
- intros H1 H2; apply H1; apply Zlt_gt; assumption.
+ apply Z.lt_nge.
Qed.
-Lemma Zle_not_gt : forall n m:Z, n <= m -> ~ n > m.
+Lemma Zle_not_gt n m : n <= m -> ~ n > m.
Proof.
trivial.
Qed.
-Lemma Zgt_not_le : forall n m:Z, n > m -> ~ n <= m.
+Lemma Zgt_not_le n m : n > m -> ~ n <= m.
Proof.
- intros n m H1 H2; apply H2; assumption.
+ Z.swap_greater. apply Z.lt_nge.
Qed.
-Lemma Zle_not_lt : forall n m:Z, n <= m -> ~ m < n.
+Lemma Znot_ge_lt n m : ~ n >= m -> n < m.
Proof.
- intros n m H1 H2.
- assert (H3 := Zlt_gt _ _ H2).
- apply Zle_not_gt with n m; assumption.
+ Z.swap_greater. apply Z.nle_gt.
Qed.
-Lemma Zlt_not_le : forall n m:Z, n < m -> ~ m <= n.
-Proof.
- intros n m H1 H2.
- apply Zle_not_lt with m n; assumption.
-Qed.
-
-Lemma Znot_ge_lt : forall n m:Z, ~ n >= m -> n < m.
-Proof.
- unfold Zge, Zlt in |- *; intros x y H; apply dec_not_not;
- [ exact (dec_Zlt x y) | assumption ].
-Qed.
-
-Lemma Znot_lt_ge : forall n m:Z, ~ n < m -> n >= m.
-Proof.
- unfold Zlt, Zge in |- *; auto with arith.
-Qed.
-
-Lemma Znot_gt_le : forall n m:Z, ~ n > m -> n <= m.
+Lemma Znot_lt_ge n m : ~ n < m -> n >= m.
Proof.
trivial.
Qed.
-Lemma Znot_le_gt : forall n m:Z, ~ n <= m -> n > m.
+Lemma Znot_gt_le n m: ~ n > m -> n <= m.
Proof.
- unfold Zle, Zgt in |- *; intros x y H; apply dec_not_not;
- [ exact (dec_Zgt x y) | assumption ].
+ trivial.
Qed.
-Lemma Zge_iff_le : forall n m:Z, n >= m <-> m <= n.
+Lemma Znot_le_gt n m : ~ n <= m -> n > m.
Proof.
- intros x y; intros. split. intro. apply Zge_le. assumption.
- intro. apply Zle_ge. assumption.
+ Z.swap_greater. apply Z.nle_gt.
Qed.
-Lemma Zgt_iff_lt : forall n m:Z, n > m <-> m < n.
+Lemma not_Zne n m : ~ Zne n m -> n = m.
Proof.
- intros x y. split. intro. apply Zgt_lt. assumption.
- intro. apply Zlt_gt. assumption.
+ intros H.
+ destruct (Z.eq_decidable n m); [assumption|now elim H].
Qed.
(** * Equivalence and order properties *)
(** Reflexivity *)
-Lemma Zle_refl : forall n:Z, n <= n.
-Proof.
- intros n; unfold Zle in |- *; rewrite (Zcompare_refl n); discriminate.
-Qed.
+Notation Zle_refl := Z.le_refl (only parsing).
+Notation Zeq_le := Z.eq_le_incl (only parsing).
-Lemma Zeq_le : forall n m:Z, n = m -> n <= m.
-Proof.
- intros; rewrite H; apply Zle_refl.
-Qed.
-
-Hint Resolve Zle_refl: zarith.
+Hint Resolve Z.le_refl: zarith.
(** Antisymmetry *)
-Lemma Zle_antisym : forall n m:Z, n <= m -> m <= n -> n = m.
-Proof.
- intros n m H1 H2; destruct (Ztrichotomy n m) as [Hlt| [Heq| Hgt]].
- absurd (m > n); [ apply Zle_not_gt | apply Zlt_gt ]; assumption.
- assumption.
- absurd (n > m); [ apply Zle_not_gt | idtac ]; assumption.
-Qed.
+Notation Zle_antisym := Z.le_antisymm (only parsing).
(** Asymmetry *)
-Lemma Zgt_asym : forall n m:Z, n > m -> ~ m > n.
-Proof.
- unfold Zgt in |- *; intros n m H; elim (Zcompare_Gt_Lt_antisym n m);
- intros H1 H2; rewrite H1; [ discriminate | assumption ].
-Qed.
+Notation Zlt_asym := Z.lt_asymm (only parsing).
-Lemma Zlt_asym : forall n m:Z, n < m -> ~ m < n.
+Lemma Zgt_asym n m : n > m -> ~ m > n.
Proof.
- intros n m H H1; assert (H2 : m > n). apply Zlt_gt; assumption.
- assert (H3 : n > m). apply Zlt_gt; assumption.
- apply Zgt_asym with m n; assumption.
+ Z.swap_greater. apply Z.lt_asymm.
Qed.
(** Irreflexivity *)
-Lemma Zgt_irrefl : forall n:Z, ~ n > n.
-Proof.
- intros n H; apply (Zgt_asym n n H H).
-Qed.
-
-Lemma Zlt_irrefl : forall n:Z, ~ n < n.
-Proof.
- intros n H; apply (Zlt_asym n n H H).
-Qed.
+Notation Zlt_irrefl := Z.lt_irrefl (only parsing).
+Notation Zlt_not_eq := Z.lt_neq (only parsing).
-Lemma Zlt_not_eq : forall n m:Z, n < m -> n <> m.
+Lemma Zgt_irrefl n : ~ n > n.
Proof.
- unfold not in |- *; intros x y H H0.
- rewrite H0 in H.
- apply (Zlt_irrefl _ H).
+ Z.swap_greater. apply Z.lt_irrefl.
Qed.
(** Large = strict or equal *)
-Lemma Zlt_le_weak : forall n m:Z, n < m -> n <= m.
-Proof.
- intros n m Hlt; apply Znot_gt_le; apply Zgt_asym; apply Zlt_gt; assumption.
-Qed.
-
-Lemma Zle_lt_or_eq : forall n m:Z, n <= m -> n < m \/ n = m.
-Proof.
- intros n m H; destruct (Ztrichotomy n m) as [Hlt| [Heq| Hgt]];
- [ left; assumption
- | right; assumption
- | absurd (n > m); [ apply Zle_not_gt | idtac ]; assumption ].
-Qed.
+Notation Zlt_le_weak := Z.lt_le_incl (only parsing).
+Notation Zle_lt_or_eq_iff := Z.lt_eq_cases (only parsing).
-Lemma Zle_lt_or_eq_iff : forall n m, n <= m <-> n < m \/ n = m.
+Lemma Zle_lt_or_eq n m : n <= m -> n < m \/ n = m.
Proof.
- unfold Zle, Zlt. intros.
- generalize (Zcompare_Eq_iff_eq n m).
- destruct (n ?= m); intuition; discriminate.
+ apply Z.lt_eq_cases.
Qed.
(** Dichotomy *)
-Lemma Zle_or_lt : forall n m:Z, n <= m \/ m < n.
-Proof.
- intros n m; destruct (Ztrichotomy n m) as [Hlt| [Heq| Hgt]];
- [ left; apply Znot_gt_le; intro Hgt; assert (Hgt' := Zlt_gt _ _ Hlt);
- apply Zgt_asym with m n; assumption
- | left; rewrite Heq; apply Zle_refl
- | right; apply Zgt_lt; assumption ].
-Qed.
+Notation Zle_or_lt := Z.le_gt_cases (only parsing).
(** Transitivity of strict orders *)
-Lemma Zgt_trans : forall n m p:Z, n > m -> m > p -> n > p.
-Proof.
- exact Zcompare_Gt_trans.
-Qed.
+Notation Zlt_trans := Z.lt_trans (only parsing).
-Lemma Zlt_trans : forall n m p:Z, n < m -> m < p -> n < p.
-Proof.
- exact Zcompare_Lt_trans.
-Qed.
+Lemma Zgt_trans : forall n m p:Z, n > m -> m > p -> n > p.
+Proof Zcompare_Gt_trans.
(** Mixed transitivity *)
-Lemma Zle_gt_trans : forall n m p:Z, m <= n -> m > p -> n > p.
-Proof.
- intros n m p H1 H2; destruct (Zle_lt_or_eq m n H1) as [Hlt| Heq];
- [ apply Zgt_trans with m; [ apply Zlt_gt; assumption | assumption ]
- | rewrite <- Heq; assumption ].
-Qed.
+Notation Zlt_le_trans := Z.lt_le_trans (only parsing).
+Notation Zle_lt_trans := Z.le_lt_trans (only parsing).
-Lemma Zgt_le_trans : forall n m p:Z, n > m -> p <= m -> n > p.
+Lemma Zle_gt_trans n m p : m <= n -> m > p -> n > p.
Proof.
- intros n m p H1 H2; destruct (Zle_lt_or_eq p m H2) as [Hlt| Heq];
- [ apply Zgt_trans with m; [ assumption | apply Zlt_gt; assumption ]
- | rewrite Heq; assumption ].
+ Z.swap_greater. Z.order.
Qed.
-Lemma Zlt_le_trans : forall n m p:Z, n < m -> m <= p -> n < p.
- intros n m p H1 H2; apply Zgt_lt; apply Zle_gt_trans with (m := m);
- [ assumption | apply Zlt_gt; assumption ].
-Qed.
-
-Lemma Zle_lt_trans : forall n m p:Z, n <= m -> m < p -> n < p.
+Lemma Zgt_le_trans n m p : n > m -> p <= m -> n > p.
Proof.
- intros n m p H1 H2; apply Zgt_lt; apply Zgt_le_trans with (m := m);
- [ apply Zlt_gt; assumption | assumption ].
+ Z.swap_greater. Z.order.
Qed.
(** Transitivity of large orders *)
-Lemma Zle_trans : forall n m p:Z, n <= m -> m <= p -> n <= p.
-Proof.
- intros n m p H1 H2; apply Znot_gt_le.
- intro Hgt; apply Zle_not_gt with n m. assumption.
- exact (Zgt_le_trans n p m Hgt H2).
-Qed.
+Notation Zle_trans := Z.le_trans (only parsing).
-Lemma Zge_trans : forall n m p:Z, n >= m -> m >= p -> n >= p.
+Lemma Zge_trans n m p : n >= m -> m >= p -> n >= p.
Proof.
- intros n m p H1 H2.
- apply Zle_ge.
- apply Zle_trans with m; apply Zge_le; trivial.
+ Z.swap_greater. Z.order.
Qed.
-Hint Resolve Zle_trans: zarith.
-
+Hint Resolve Z.le_trans: zarith.
(** * Compatibility of order and operations on Z *)
@@ -339,700 +202,448 @@ Hint Resolve Zle_trans: zarith.
(** Compatibility of successor wrt to order *)
-Lemma Zsucc_le_compat : forall n m:Z, m <= n -> Zsucc m <= Zsucc n.
+Lemma Zsucc_le_compat n m : m <= n -> Z.succ m <= Z.succ n.
Proof.
- unfold Zle, not in |- *; intros m n H1 H2; apply H1;
- rewrite <- (Zcompare_plus_compat n m 1); do 2 rewrite (Zplus_comm 1);
- exact H2.
+ apply Z.succ_le_mono.
Qed.
-Lemma Zsucc_gt_compat : forall n m:Z, m > n -> Zsucc m > Zsucc n.
+Lemma Zsucc_lt_compat n m : n < m -> Z.succ n < Z.succ m.
Proof.
- unfold Zgt in |- *; intros n m H; rewrite Zcompare_succ_compat;
- auto with arith.
+ apply Z.succ_lt_mono.
Qed.
-Lemma Zsucc_lt_compat : forall n m:Z, n < m -> Zsucc n < Zsucc m.
+Lemma Zsucc_gt_compat n m : m > n -> Z.succ m > Z.succ n.
Proof.
- intros n m H; apply Zgt_lt; apply Zsucc_gt_compat; apply Zlt_gt; assumption.
+ Z.swap_greater. apply Z.succ_lt_mono.
Qed.
Hint Resolve Zsucc_le_compat: zarith.
(** Simplification of successor wrt to order *)
-Lemma Zsucc_gt_reg : forall n m:Z, Zsucc m > Zsucc n -> m > n.
+Lemma Zsucc_gt_reg n m : Z.succ m > Z.succ n -> m > n.
Proof.
- unfold Zsucc, Zgt in |- *; intros n p;
- do 2 rewrite (fun m:Z => Zplus_comm m 1);
- rewrite (Zcompare_plus_compat p n 1); trivial with arith.
+ Z.swap_greater. apply Z.succ_lt_mono.
Qed.
-Lemma Zsucc_le_reg : forall n m:Z, Zsucc m <= Zsucc n -> m <= n.
+Lemma Zsucc_le_reg n m : Z.succ m <= Z.succ n -> m <= n.
Proof.
- unfold Zle, not in |- *; intros m n H1 H2; apply H1; unfold Zsucc in |- *;
- do 2 rewrite <- (Zplus_comm 1); rewrite (Zcompare_plus_compat n m 1);
- assumption.
+ apply Z.succ_le_mono.
Qed.
-Lemma Zsucc_lt_reg : forall n m:Z, Zsucc n < Zsucc m -> n < m.
+Lemma Zsucc_lt_reg n m : Z.succ n < Z.succ m -> n < m.
Proof.
- intros n m H; apply Zgt_lt; apply Zsucc_gt_reg; apply Zlt_gt; assumption.
+ apply Z.succ_lt_mono.
Qed.
(** Special base instances of order *)
-Lemma Zgt_succ : forall n:Z, Zsucc n > n.
-Proof.
- exact Zcompare_succ_Gt.
-Qed.
-
-Lemma Znot_le_succ : forall n:Z, ~ Zsucc n <= n.
-Proof.
- intros n; apply Zgt_not_le; apply Zgt_succ.
-Qed.
+Notation Zlt_succ := Z.lt_succ_diag_r (only parsing).
+Notation Zlt_pred := Z.lt_pred_l (only parsing).
-Lemma Zlt_succ : forall n:Z, n < Zsucc n.
+Lemma Zgt_succ n : Z.succ n > n.
Proof.
- intro n; apply Zgt_lt; apply Zgt_succ.
+ Z.swap_greater. apply Z.lt_succ_diag_r.
Qed.
-Lemma Zlt_pred : forall n:Z, Zpred n < n.
+Lemma Znot_le_succ n : ~ Z.succ n <= n.
Proof.
- intros n; apply Zsucc_lt_reg; rewrite <- Zsucc_pred; apply Zlt_succ.
+ apply Z.lt_nge, Z.lt_succ_diag_r.
Qed.
(** Relating strict and large order using successor or predecessor *)
-Lemma Zgt_le_succ : forall n m:Z, m > n -> Zsucc n <= m.
-Proof.
- unfold Zgt, Zle in |- *; intros n p H; elim (Zcompare_Gt_not_Lt p n);
- intros H1 H2; unfold not in |- *; intros H3; unfold not in H1;
- apply H1;
- [ assumption
- | elim (Zcompare_Gt_Lt_antisym (n + 1) p); intros H4 H5; apply H4; exact H3 ].
-Qed.
+Notation Zlt_succ_r := Z.lt_succ_r (only parsing).
+Notation Zle_succ_l := Z.le_succ_l (only parsing).
-Lemma Zle_gt_succ : forall n m:Z, n <= m -> Zsucc m > n.
+Lemma Zgt_le_succ n m : m > n -> Z.succ n <= m.
Proof.
- intros n p H; apply Zgt_le_trans with p.
- apply Zgt_succ.
- assumption.
+ Z.swap_greater. apply Z.le_succ_l.
Qed.
-Lemma Zle_lt_succ : forall n m:Z, n <= m -> n < Zsucc m.
+Lemma Zle_gt_succ n m : n <= m -> Z.succ m > n.
Proof.
- intros n m H; apply Zgt_lt; apply Zle_gt_succ; assumption.
+ Z.swap_greater. apply Z.lt_succ_r.
Qed.
-Lemma Zlt_le_succ : forall n m:Z, n < m -> Zsucc n <= m.
+Lemma Zle_lt_succ n m : n <= m -> n < Z.succ m.
Proof.
- intros n p H; apply Zgt_le_succ; apply Zlt_gt; assumption.
+ apply Z.lt_succ_r.
Qed.
-Lemma Zgt_succ_le : forall n m:Z, Zsucc m > n -> n <= m.
+Lemma Zlt_le_succ n m : n < m -> Z.succ n <= m.
Proof.
- intros n p H; apply Zsucc_le_reg; apply Zgt_le_succ; assumption.
+ apply Z.le_succ_l.
Qed.
-Lemma Zlt_succ_le : forall n m:Z, n < Zsucc m -> n <= m.
+Lemma Zgt_succ_le n m : Z.succ m > n -> n <= m.
Proof.
- intros n m H; apply Zgt_succ_le; apply Zlt_gt; assumption.
+ Z.swap_greater. apply Z.lt_succ_r.
Qed.
-Lemma Zle_succ_gt : forall n m:Z, Zsucc n <= m -> m > n.
+Lemma Zlt_succ_le n m : n < Z.succ m -> n <= m.
Proof.
- intros n m H; apply Zle_gt_trans with (m := Zsucc n);
- [ assumption | apply Zgt_succ ].
+ apply Z.lt_succ_r.
Qed.
-Lemma Zlt_succ_r : forall n m, n < Zsucc m <-> n <= m.
+Lemma Zle_succ_gt n m : Z.succ n <= m -> m > n.
Proof.
- split; [apply Zlt_succ_le | apply Zle_lt_succ].
+ Z.swap_greater. apply Z.le_succ_l.
Qed.
(** Weakening order *)
-Lemma Zle_succ : forall n:Z, n <= Zsucc n.
-Proof.
- intros n; apply Zgt_succ_le; apply Zgt_trans with (m := Zsucc n);
- apply Zgt_succ.
-Qed.
-
-Hint Resolve Zle_succ: zarith.
-
-Lemma Zle_pred : forall n:Z, Zpred n <= n.
-Proof.
- intros n; pattern n at 2 in |- *; rewrite Zsucc_pred; apply Zle_succ.
-Qed.
-
-Lemma Zlt_lt_succ : forall n m:Z, n < m -> n < Zsucc m.
- intros n m H; apply Zgt_lt; apply Zgt_trans with (m := m);
- [ apply Zgt_succ | apply Zlt_gt; assumption ].
-Qed.
+Notation Zle_succ := Z.le_succ_diag_r (only parsing).
+Notation Zle_pred := Z.le_pred_l (only parsing).
+Notation Zlt_lt_succ := Z.lt_lt_succ_r (only parsing).
+Notation Zle_le_succ := Z.le_le_succ_r (only parsing).
-Lemma Zle_le_succ : forall n m:Z, n <= m -> n <= Zsucc m.
+Lemma Zle_succ_le n m : Z.succ n <= m -> n <= m.
Proof.
- intros x y H.
- apply Zle_trans with y; trivial with zarith.
-Qed.
-
-Lemma Zle_succ_le : forall n m:Z, Zsucc n <= m -> n <= m.
-Proof.
- intros n m H; apply Zle_trans with (m := Zsucc n);
- [ apply Zle_succ | assumption ].
+ intros. now apply Z.lt_le_incl, Z.le_succ_l.
Qed.
+Hint Resolve Z.le_succ_diag_r: zarith.
Hint Resolve Zle_le_succ: zarith.
(** Relating order wrt successor and order wrt predecessor *)
-Lemma Zgt_succ_pred : forall n m:Z, m > Zsucc n -> Zpred m > n.
+Lemma Zgt_succ_pred n m : m > Z.succ n -> Z.pred m > n.
Proof.
- unfold Zgt, Zsucc, Zpred in |- *; intros n p H;
- rewrite <- (fun x y => Zcompare_plus_compat x y 1);
- rewrite (Zplus_comm p); rewrite Zplus_assoc;
- rewrite (fun x => Zplus_comm x n); simpl in |- *;
- assumption.
+ Z.swap_greater. apply Z.lt_succ_lt_pred.
Qed.
-Lemma Zlt_succ_pred : forall n m:Z, Zsucc n < m -> n < Zpred m.
+Lemma Zlt_succ_pred n m : Z.succ n < m -> n < Z.pred m.
Proof.
- intros n p H; apply Zsucc_lt_reg; rewrite <- Zsucc_pred; assumption.
+ apply Z.lt_succ_lt_pred.
Qed.
(** Relating strict order and large order on positive *)
-Lemma Zlt_0_le_0_pred : forall n:Z, 0 < n -> 0 <= Zpred n.
+Lemma Zlt_0_le_0_pred n : 0 < n -> 0 <= Z.pred n.
Proof.
- intros x H.
- rewrite (Zsucc_pred x) in H.
- apply Zgt_succ_le.
- apply Zlt_gt.
- assumption.
+ apply Z.lt_le_pred.
Qed.
-Lemma Zgt_0_le_0_pred : forall n:Z, n > 0 -> 0 <= Zpred n.
+Lemma Zgt_0_le_0_pred n : n > 0 -> 0 <= Z.pred n.
Proof.
- intros; apply Zlt_0_le_0_pred; apply Zgt_lt. assumption.
+ Z.swap_greater. apply Z.lt_le_pred.
Qed.
-
(** Special cases of ordered integers *)
-Lemma Zlt_0_1 : 0 < 1.
-Proof.
- change (0 < Zsucc 0) in |- *. apply Zlt_succ.
-Qed.
-
-Lemma Zle_0_1 : 0 <= 1.
-Proof.
- change (0 <= Zsucc 0) in |- *. apply Zle_succ.
-Qed.
+Notation Zlt_0_1 := Z.lt_0_1 (only parsing).
+Notation Zle_0_1 := Z.le_0_1 (only parsing).
Lemma Zle_neg_pos : forall p q:positive, Zneg p <= Zpos q.
Proof.
- intros p; red in |- *; simpl in |- *; red in |- *; intros H; discriminate.
+ easy.
Qed.
Lemma Zgt_pos_0 : forall p:positive, Zpos p > 0.
Proof.
- unfold Zgt in |- *; trivial.
+ easy.
Qed.
(* weaker but useful (in [Zpower] for instance) *)
Lemma Zle_0_pos : forall p:positive, 0 <= Zpos p.
Proof.
- intro; unfold Zle in |- *; discriminate.
+ easy.
Qed.
Lemma Zlt_neg_0 : forall p:positive, Zneg p < 0.
Proof.
- unfold Zlt in |- *; trivial.
+ easy.
Qed.
-Lemma Zle_0_nat : forall n:nat, 0 <= Z_of_nat n.
+Lemma Zle_0_nat : forall n:nat, 0 <= Z.of_nat n.
Proof.
- simple induction n; simpl in |- *; intros;
- [ apply Zle_refl | unfold Zle in |- *; simpl in |- *; discriminate ].
+ induction n; simpl; intros. apply Z.le_refl. easy.
Qed.
Hint Immediate Zeq_le: zarith.
-(** Transitivity using successor *)
-
-Lemma Zgt_trans_succ : forall n m p:Z, Zsucc n > m -> m > p -> n > p.
-Proof.
- intros n m p H1 H2; apply Zle_gt_trans with (m := m);
- [ apply Zgt_succ_le; assumption | assumption ].
-Qed.
-
(** Derived lemma *)
-Lemma Zgt_succ_gt_or_eq : forall n m:Z, Zsucc n > m -> n > m \/ m = n.
+Lemma Zgt_succ_gt_or_eq n m : Z.succ n > m -> n > m \/ m = n.
Proof.
- intros n m H.
- assert (Hle : m <= n).
- apply Zgt_succ_le; assumption.
- destruct (Zle_lt_or_eq _ _ Hle) as [Hlt| Heq].
- left; apply Zlt_gt; assumption.
- right; assumption.
+ Z.swap_greater. intros. now apply Z.lt_eq_cases, Z.lt_succ_r.
Qed.
(** ** Addition *)
(** Compatibility of addition wrt to order *)
-Lemma Zplus_gt_compat_l : forall n m p:Z, n > m -> p + n > p + m.
-Proof.
- unfold Zgt in |- *; intros n m p H; rewrite (Zcompare_plus_compat n m p);
- assumption.
-Qed.
-
-Lemma Zplus_gt_compat_r : forall n m p:Z, n > m -> n + p > m + p.
-Proof.
- intros n m p H; rewrite (Zplus_comm n p); rewrite (Zplus_comm m p);
- apply Zplus_gt_compat_l; trivial.
-Qed.
-
-Lemma Zplus_le_compat_l : forall n m p:Z, n <= m -> p + n <= p + m.
-Proof.
- intros n m p; unfold Zle, not in |- *; intros H1 H2; apply H1;
- rewrite <- (Zcompare_plus_compat n m p); assumption.
-Qed.
+Notation Zplus_lt_le_compat := Z.add_lt_le_mono (only parsing).
+Notation Zplus_le_lt_compat := Z.add_le_lt_mono (only parsing).
+Notation Zplus_le_compat := Z.add_le_mono (only parsing).
+Notation Zplus_lt_compat := Z.add_lt_mono (only parsing).
-Lemma Zplus_le_compat_r : forall n m p:Z, n <= m -> n + p <= m + p.
+Lemma Zplus_gt_compat_l n m p : n > m -> p + n > p + m.
Proof.
- intros a b c; do 2 rewrite (fun n:Z => Zplus_comm n c);
- exact (Zplus_le_compat_l a b c).
+ Z.swap_greater. apply Z.add_lt_mono_l.
Qed.
-Lemma Zplus_lt_compat_l : forall n m p:Z, n < m -> p + n < p + m.
+Lemma Zplus_gt_compat_r n m p : n > m -> n + p > m + p.
Proof.
- unfold Zlt in |- *; intros n m p; rewrite Zcompare_plus_compat;
- trivial with arith.
+ Z.swap_greater. apply Z.add_lt_mono_r.
Qed.
-Lemma Zplus_lt_compat_r : forall n m p:Z, n < m -> n + p < m + p.
+Lemma Zplus_le_compat_l n m p : n <= m -> p + n <= p + m.
Proof.
- intros n m p H; rewrite (Zplus_comm n p); rewrite (Zplus_comm m p);
- apply Zplus_lt_compat_l; trivial.
+ apply Z.add_le_mono_l.
Qed.
-Lemma Zplus_lt_le_compat : forall n m p q:Z, n < m -> p <= q -> n + p < m + q.
+Lemma Zplus_le_compat_r n m p : n <= m -> n + p <= m + p.
Proof.
- intros a b c d H0 H1.
- apply Zlt_le_trans with (b + c).
- apply Zplus_lt_compat_r; trivial.
- apply Zplus_le_compat_l; trivial.
+ apply Z.add_le_mono_r.
Qed.
-Lemma Zplus_le_lt_compat : forall n m p q:Z, n <= m -> p < q -> n + p < m + q.
+Lemma Zplus_lt_compat_l n m p : n < m -> p + n < p + m.
Proof.
- intros a b c d H0 H1.
- apply Zle_lt_trans with (b + c).
- apply Zplus_le_compat_r; trivial.
- apply Zplus_lt_compat_l; trivial.
+ apply Z.add_lt_mono_l.
Qed.
-Lemma Zplus_le_compat : forall n m p q:Z, n <= m -> p <= q -> n + p <= m + q.
+Lemma Zplus_lt_compat_r n m p : n < m -> n + p < m + p.
Proof.
- intros n m p q; intros H1 H2; apply Zle_trans with (m := n + q);
- [ apply Zplus_le_compat_l; assumption
- | apply Zplus_le_compat_r; assumption ].
-Qed.
-
-
-Lemma Zplus_lt_compat : forall n m p q:Z, n < m -> p < q -> n + p < m + q.
- intros; apply Zplus_le_lt_compat. apply Zlt_le_weak; assumption. assumption.
+ apply Z.add_lt_mono_r.
Qed.
-
(** Compatibility of addition wrt to being positive *)
-Lemma Zplus_le_0_compat : forall n m:Z, 0 <= n -> 0 <= m -> 0 <= n + m.
-Proof.
- intros x y H1 H2; rewrite <- (Zplus_0_l 0); apply Zplus_le_compat; assumption.
-Qed.
+Notation Zplus_le_0_compat := Z.add_nonneg_nonneg (only parsing).
(** Simplification of addition wrt to order *)
-Lemma Zplus_gt_reg_l : forall n m p:Z, p + n > p + m -> n > m.
+Lemma Zplus_le_reg_l n m p : p + n <= p + m -> n <= m.
Proof.
- unfold Zgt in |- *; intros n m p H; rewrite <- (Zcompare_plus_compat n m p);
- assumption.
+ apply Z.add_le_mono_l.
Qed.
-Lemma Zplus_gt_reg_r : forall n m p:Z, n + p > m + p -> n > m.
+Lemma Zplus_le_reg_r n m p : n + p <= m + p -> n <= m.
Proof.
- intros n m p H; apply Zplus_gt_reg_l with p.
- rewrite (Zplus_comm p n); rewrite (Zplus_comm p m); trivial.
+ apply Z.add_le_mono_r.
Qed.
-Lemma Zplus_le_reg_l : forall n m p:Z, p + n <= p + m -> n <= m.
+Lemma Zplus_lt_reg_l n m p : p + n < p + m -> n < m.
Proof.
- intros n m p; unfold Zle, not in |- *; intros H1 H2; apply H1;
- rewrite (Zcompare_plus_compat n m p); assumption.
+ apply Z.add_lt_mono_l.
Qed.
-Lemma Zplus_le_reg_r : forall n m p:Z, n + p <= m + p -> n <= m.
+Lemma Zplus_lt_reg_r n m p : n + p < m + p -> n < m.
Proof.
- intros n m p H; apply Zplus_le_reg_l with p.
- rewrite (Zplus_comm p n); rewrite (Zplus_comm p m); trivial.
+ apply Z.add_lt_mono_r.
Qed.
-Lemma Zplus_lt_reg_l : forall n m p:Z, p + n < p + m -> n < m.
+Lemma Zplus_gt_reg_l n m p : p + n > p + m -> n > m.
Proof.
- unfold Zlt in |- *; intros n m p; rewrite Zcompare_plus_compat;
- trivial with arith.
+ Z.swap_greater. apply Z.add_lt_mono_l.
Qed.
-Lemma Zplus_lt_reg_r : forall n m p:Z, n + p < m + p -> n < m.
+Lemma Zplus_gt_reg_r n m p : n + p > m + p -> n > m.
Proof.
- intros n m p H; apply Zplus_lt_reg_l with p.
- rewrite (Zplus_comm p n); rewrite (Zplus_comm p m); trivial.
+ Z.swap_greater. apply Z.add_lt_mono_r.
Qed.
(** ** Multiplication *)
(** Compatibility of multiplication by a positive wrt to order *)
-Lemma Zmult_le_compat_r : forall n m p:Z, n <= m -> 0 <= p -> n * p <= m * p.
+Lemma Zmult_le_compat_r n m p : n <= m -> 0 <= p -> n * p <= m * p.
Proof.
- intros a b c H H0; destruct c.
- do 2 rewrite Zmult_0_r; assumption.
- rewrite (Zmult_comm a); rewrite (Zmult_comm b).
- unfold Zle in |- *; rewrite Zcompare_mult_compat; assumption.
- unfold Zle in H0; contradiction H0; reflexivity.
+ intros. now apply Z.mul_le_mono_nonneg_r.
Qed.
-Lemma Zmult_le_compat_l : forall n m p:Z, n <= m -> 0 <= p -> p * n <= p * m.
+Lemma Zmult_le_compat_l n m p : n <= m -> 0 <= p -> p * n <= p * m.
Proof.
- intros a b c H1 H2; rewrite (Zmult_comm c a); rewrite (Zmult_comm c b).
- apply Zmult_le_compat_r; trivial.
+ intros. now apply Z.mul_le_mono_nonneg_l.
Qed.
-Lemma Zmult_lt_compat_r : forall n m p:Z, 0 < p -> n < m -> n * p < m * p.
+Lemma Zmult_lt_compat_r n m p : 0 < p -> n < m -> n * p < m * p.
Proof.
- intros x y z H H0; destruct z.
- contradiction (Zlt_irrefl 0).
- rewrite (Zmult_comm x); rewrite (Zmult_comm y).
- unfold Zlt in |- *; rewrite Zcompare_mult_compat; assumption.
- discriminate H.
+ apply Z.mul_lt_mono_pos_r.
Qed.
-Lemma Zmult_gt_compat_r : forall n m p:Z, p > 0 -> n > m -> n * p > m * p.
+Lemma Zmult_gt_compat_r n m p : p > 0 -> n > m -> n * p > m * p.
Proof.
- intros x y z; intros; apply Zlt_gt; apply Zmult_lt_compat_r; apply Zgt_lt;
- assumption.
+ Z.swap_greater. apply Z.mul_lt_mono_pos_r.
Qed.
-Lemma Zmult_gt_0_lt_compat_r :
- forall n m p:Z, p > 0 -> n < m -> n * p < m * p.
+Lemma Zmult_gt_0_lt_compat_r n m p : p > 0 -> n < m -> n * p < m * p.
Proof.
- intros x y z; intros; apply Zmult_lt_compat_r;
- [ apply Zgt_lt; assumption | assumption ].
+ Z.swap_greater. apply Z.mul_lt_mono_pos_r.
Qed.
-Lemma Zmult_gt_0_le_compat_r :
- forall n m p:Z, p > 0 -> n <= m -> n * p <= m * p.
+Lemma Zmult_gt_0_le_compat_r n m p : p > 0 -> n <= m -> n * p <= m * p.
Proof.
- intros x y z Hz Hxy.
- elim (Zle_lt_or_eq x y Hxy).
- intros; apply Zlt_le_weak.
- apply Zmult_gt_0_lt_compat_r; trivial.
- intros; apply Zeq_le.
- rewrite H; trivial.
+ Z.swap_greater. apply Z.mul_le_mono_pos_r.
Qed.
-Lemma Zmult_lt_0_le_compat_r :
- forall n m p:Z, 0 < p -> n <= m -> n * p <= m * p.
+Lemma Zmult_lt_0_le_compat_r n m p : 0 < p -> n <= m -> n * p <= m * p.
Proof.
- intros x y z; intros; apply Zmult_gt_0_le_compat_r; try apply Zlt_gt;
- assumption.
+ apply Z.mul_le_mono_pos_r.
Qed.
-Lemma Zmult_gt_0_lt_compat_l :
- forall n m p:Z, p > 0 -> n < m -> p * n < p * m.
+Lemma Zmult_gt_0_lt_compat_l n m p : p > 0 -> n < m -> p * n < p * m.
Proof.
- intros x y z; intros.
- rewrite (Zmult_comm z x); rewrite (Zmult_comm z y);
- apply Zmult_gt_0_lt_compat_r; assumption.
+ Z.swap_greater. apply Z.mul_lt_mono_pos_l.
Qed.
-Lemma Zmult_lt_compat_l : forall n m p:Z, 0 < p -> n < m -> p * n < p * m.
+Lemma Zmult_lt_compat_l n m p : 0 < p -> n < m -> p * n < p * m.
Proof.
- intros x y z; intros.
- rewrite (Zmult_comm z x); rewrite (Zmult_comm z y);
- apply Zmult_gt_0_lt_compat_r; try apply Zlt_gt; assumption.
+ apply Z.mul_lt_mono_pos_l.
Qed.
-Lemma Zmult_gt_compat_l : forall n m p:Z, p > 0 -> n > m -> p * n > p * m.
+Lemma Zmult_gt_compat_l n m p : p > 0 -> n > m -> p * n > p * m.
Proof.
- intros x y z; intros; rewrite (Zmult_comm z x); rewrite (Zmult_comm z y);
- apply Zmult_gt_compat_r; assumption.
+ Z.swap_greater. apply Z.mul_lt_mono_pos_l.
Qed.
-Lemma Zmult_ge_compat_r : forall n m p:Z, n >= m -> p >= 0 -> n * p >= m * p.
+Lemma Zmult_ge_compat_r n m p : n >= m -> p >= 0 -> n * p >= m * p.
Proof.
- intros a b c H1 H2; apply Zle_ge.
- apply Zmult_le_compat_r; apply Zge_le; trivial.
+ Z.swap_greater. intros. now apply Z.mul_le_mono_nonneg_r.
Qed.
-Lemma Zmult_ge_compat_l : forall n m p:Z, n >= m -> p >= 0 -> p * n >= p * m.
+Lemma Zmult_ge_compat_l n m p : n >= m -> p >= 0 -> p * n >= p * m.
Proof.
- intros a b c H1 H2; apply Zle_ge.
- apply Zmult_le_compat_l; apply Zge_le; trivial.
+ Z.swap_greater. intros. now apply Z.mul_le_mono_nonneg_l.
Qed.
-Lemma Zmult_ge_compat :
- forall n m p q:Z, n >= p -> m >= q -> p >= 0 -> q >= 0 -> n * m >= p * q.
+Lemma Zmult_ge_compat n m p q :
+ n >= p -> m >= q -> p >= 0 -> q >= 0 -> n * m >= p * q.
Proof.
- intros a b c d H0 H1 H2 H3.
- apply Zge_trans with (a * d).
- apply Zmult_ge_compat_l; trivial.
- apply Zge_trans with c; trivial.
- apply Zmult_ge_compat_r; trivial.
+ Z.swap_greater. intros. now apply Z.mul_le_mono_nonneg.
Qed.
-Lemma Zmult_le_compat :
- forall n m p q:Z, n <= p -> m <= q -> 0 <= n -> 0 <= m -> n * m <= p * q.
+Lemma Zmult_le_compat n m p q :
+ n <= p -> m <= q -> 0 <= n -> 0 <= m -> n * m <= p * q.
Proof.
- intros a b c d H0 H1 H2 H3.
- apply Zle_trans with (c * b).
- apply Zmult_le_compat_r; assumption.
- apply Zmult_le_compat_l.
- assumption.
- apply Zle_trans with a; assumption.
+ intros. now apply Z.mul_le_mono_nonneg.
Qed.
(** Simplification of multiplication by a positive wrt to being positive *)
-Lemma Zmult_gt_0_lt_reg_r : forall n m p:Z, p > 0 -> n * p < m * p -> n < m.
+Lemma Zmult_gt_0_lt_reg_r n m p : p > 0 -> n * p < m * p -> n < m.
Proof.
- intros x y z; intros; destruct z.
- contradiction (Zgt_irrefl 0).
- rewrite (Zmult_comm x) in H0; rewrite (Zmult_comm y) in H0.
- unfold Zlt in H0; rewrite Zcompare_mult_compat in H0; assumption.
- discriminate H.
+ Z.swap_greater. apply Z.mul_lt_mono_pos_r.
Qed.
-Lemma Zmult_lt_reg_r : forall n m p:Z, 0 < p -> n * p < m * p -> n < m.
+Lemma Zmult_lt_reg_r n m p : 0 < p -> n * p < m * p -> n < m.
Proof.
- intros a b c H0 H1.
- apply Zmult_gt_0_lt_reg_r with c; try apply Zlt_gt; assumption.
+ apply Z.mul_lt_mono_pos_r.
Qed.
-Lemma Zmult_le_reg_r : forall n m p:Z, p > 0 -> n * p <= m * p -> n <= m.
+Lemma Zmult_le_reg_r n m p : p > 0 -> n * p <= m * p -> n <= m.
Proof.
- intros x y z Hz Hxy.
- elim (Zle_lt_or_eq (x * z) (y * z) Hxy).
- intros; apply Zlt_le_weak.
- apply Zmult_gt_0_lt_reg_r with z; trivial.
- intros; apply Zeq_le.
- apply Zmult_reg_r with z.
- intro. rewrite H0 in Hz. contradiction (Zgt_irrefl 0).
- assumption.
+ Z.swap_greater. apply Z.mul_le_mono_pos_r.
Qed.
-Lemma Zmult_lt_0_le_reg_r : forall n m p:Z, 0 < p -> n * p <= m * p -> n <= m.
+Lemma Zmult_lt_0_le_reg_r n m p : 0 < p -> n * p <= m * p -> n <= m.
Proof.
- intros x y z; intros; apply Zmult_le_reg_r with z.
- try apply Zlt_gt; assumption.
- assumption.
+ apply Z.mul_le_mono_pos_r.
Qed.
-
-Lemma Zmult_ge_reg_r : forall n m p:Z, p > 0 -> n * p >= m * p -> n >= m.
+Lemma Zmult_ge_reg_r n m p : p > 0 -> n * p >= m * p -> n >= m.
Proof.
- intros a b c H1 H2; apply Zle_ge; apply Zmult_le_reg_r with c; trivial.
- apply Zge_le; trivial.
+ Z.swap_greater. apply Z.mul_le_mono_pos_r.
Qed.
-Lemma Zmult_gt_reg_r : forall n m p:Z, p > 0 -> n * p > m * p -> n > m.
+Lemma Zmult_gt_reg_r n m p : p > 0 -> n * p > m * p -> n > m.
Proof.
- intros a b c H1 H2; apply Zlt_gt; apply Zmult_gt_0_lt_reg_r with c; trivial.
- apply Zgt_lt; trivial.
+ Z.swap_greater. apply Z.mul_lt_mono_pos_r.
Qed.
-
-(** Compatibility of multiplication by a positive wrt to being positive *)
-
-Lemma Zmult_le_0_compat : forall n m:Z, 0 <= n -> 0 <= m -> 0 <= n * m.
+Lemma Zmult_lt_compat n m p q :
+ 0 <= n < p -> 0 <= m < q -> n * m < p * q.
Proof.
- intros x y; case x.
- intros; rewrite Zmult_0_l; trivial.
- intros p H1; unfold Zle in |- *.
- pattern 0 at 2 in |- *; rewrite <- (Zmult_0_r (Zpos p)).
- rewrite Zcompare_mult_compat; trivial.
- intros p H1 H2; absurd (0 > Zneg p); trivial.
- unfold Zgt in |- *; simpl in |- *; auto with zarith.
+ intros (Hn,Hnp) (Hm,Hmq). now apply Z.mul_lt_mono_nonneg.
Qed.
-Lemma Zmult_gt_0_compat : forall n m:Z, n > 0 -> m > 0 -> n * m > 0.
+Lemma Zmult_lt_compat2 n m p q :
+ 0 < n <= p -> 0 < m < q -> n * m < p * q.
Proof.
- intros x y; case x.
- intros H; discriminate H.
- intros p H1; unfold Zgt in |- *; pattern 0 at 2 in |- *;
- rewrite <- (Zmult_0_r (Zpos p)).
- rewrite Zcompare_mult_compat; trivial.
- intros p H; discriminate H.
+ intros (Hn, Hnp) (Hm,Hmq).
+ apply Z.le_lt_trans with (p * m).
+ apply Z.mul_le_mono_pos_r; trivial.
+ apply Z.mul_lt_mono_pos_l; Z.order.
Qed.
-Lemma Zmult_lt_0_compat : forall n m:Z, 0 < n -> 0 < m -> 0 < n * m.
+(** Compatibility of multiplication by a positive wrt to being positive *)
+
+Notation Zmult_le_0_compat := Z.mul_nonneg_nonneg (only parsing).
+Notation Zmult_lt_0_compat := Z.mul_pos_pos (only parsing).
+Notation Zmult_lt_O_compat := Z.mul_pos_pos (only parsing).
+
+Lemma Zmult_gt_0_compat n m : n > 0 -> m > 0 -> n * m > 0.
Proof.
- intros a b apos bpos.
- apply Zgt_lt.
- apply Zmult_gt_0_compat; try apply Zlt_gt; assumption.
+ Z.swap_greater. apply Z.mul_pos_pos.
Qed.
-(** For compatibility *)
-Notation Zmult_lt_O_compat := Zmult_lt_0_compat (only parsing).
+(* To remove someday ... *)
-Lemma Zmult_gt_0_le_0_compat : forall n m:Z, n > 0 -> 0 <= m -> 0 <= m * n.
+Lemma Zmult_gt_0_le_0_compat n m : n > 0 -> 0 <= m -> 0 <= m * n.
Proof.
- intros x y H1 H2; apply Zmult_le_0_compat; trivial.
- apply Zlt_le_weak; apply Zgt_lt; trivial.
+ Z.swap_greater. intros. apply Z.mul_nonneg_nonneg. trivial.
+ now apply Z.lt_le_incl.
Qed.
(** Simplification of multiplication by a positive wrt to being positive *)
-Lemma Zmult_le_0_reg_r : forall n m:Z, n > 0 -> 0 <= m * n -> 0 <= m.
+Lemma Zmult_le_0_reg_r n m : n > 0 -> 0 <= m * n -> 0 <= m.
Proof.
- intros x y; case x;
- [ simpl in |- *; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H
- | intros p H1; unfold Zle in |- *; rewrite Zmult_comm;
- pattern 0 at 1 in |- *; rewrite <- (Zmult_0_r (Zpos p));
- rewrite Zcompare_mult_compat; auto with arith
- | intros p; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H ].
+ Z.swap_greater. apply Z.mul_nonneg_cancel_r.
Qed.
-Lemma Zmult_gt_0_lt_0_reg_r : forall n m:Z, n > 0 -> 0 < m * n -> 0 < m.
+Lemma Zmult_lt_0_reg_r n m : 0 < n -> 0 < m * n -> 0 < m.
Proof.
- intros x y; case x;
- [ simpl in |- *; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H
- | intros p H1; unfold Zlt in |- *; rewrite Zmult_comm;
- pattern 0 at 1 in |- *; rewrite <- (Zmult_0_r (Zpos p));
- rewrite Zcompare_mult_compat; auto with arith
- | intros p; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H ].
+ apply Z.mul_pos_cancel_r.
Qed.
-Lemma Zmult_lt_0_reg_r : forall n m:Z, 0 < n -> 0 < m * n -> 0 < m.
+Lemma Zmult_gt_0_lt_0_reg_r n m : n > 0 -> 0 < m * n -> 0 < m.
Proof.
- intros x y; intros; eapply Zmult_gt_0_lt_0_reg_r with x; try apply Zlt_gt;
- assumption.
+ Z.swap_greater. apply Z.mul_pos_cancel_r.
Qed.
-Lemma Zmult_gt_0_reg_l : forall n m:Z, n > 0 -> n * m > 0 -> m > 0.
+Lemma Zmult_gt_0_reg_l n m : n > 0 -> n * m > 0 -> m > 0.
Proof.
- intros x y; case x.
- intros H; discriminate H.
- intros p H1; unfold Zgt in |- *.
- pattern 0 at 1 in |- *; rewrite <- (Zmult_0_r (Zpos p)).
- rewrite Zcompare_mult_compat; trivial.
- intros p H; discriminate H.
+ Z.swap_greater. apply Z.mul_pos_cancel_l.
Qed.
(** ** Square *)
(** Simplification of square wrt order *)
-Lemma Zgt_square_simpl :
- forall n m:Z, n >= 0 -> n * n > m * m -> n > m.
+Lemma Zlt_square_simpl n m : 0 <= n -> m * m < n * n -> m < n.
Proof.
- intros n m H0 H1.
- case (dec_Zlt m n).
- intro; apply Zlt_gt; trivial.
- intros H2; cut (m >= n).
- intros H.
- elim Zgt_not_le with (1 := H1).
- apply Zge_le.
- apply Zmult_ge_compat; auto.
- apply Znot_lt_ge; trivial.
+ apply Z.square_lt_simpl_nonneg.
Qed.
-Lemma Zlt_square_simpl :
- forall n m:Z, 0 <= n -> m * m < n * n -> m < n.
+Lemma Zgt_square_simpl n m : n >= 0 -> n * n > m * m -> n > m.
Proof.
- intros x y H0 H1.
- apply Zgt_lt.
- apply Zgt_square_simpl; try apply Zle_ge; try apply Zlt_gt; assumption.
+ Z.swap_greater. apply Z.square_lt_simpl_nonneg.
Qed.
(** * Equivalence between inequalities *)
-Lemma Zle_plus_swap : forall n m p:Z, n + p <= m <-> n <= m - p.
-Proof.
- intros x y z; intros. split. intro. rewrite <- (Zplus_0_r x). rewrite <- (Zplus_opp_r z).
- rewrite Zplus_assoc. exact (Zplus_le_compat_r _ _ _ H).
- intro. rewrite <- (Zplus_0_r y). rewrite <- (Zplus_opp_l z). rewrite Zplus_assoc.
- apply Zplus_le_compat_r. assumption.
-Qed.
-
-Lemma Zlt_plus_swap : forall n m p:Z, n + p < m <-> n < m - p.
-Proof.
- intros x y z; intros. split. intro. unfold Zminus in |- *. rewrite Zplus_comm. rewrite <- (Zplus_0_l x).
- rewrite <- (Zplus_opp_l z). rewrite Zplus_assoc_reverse. apply Zplus_lt_compat_l. rewrite Zplus_comm.
- assumption.
- intro. rewrite Zplus_comm. rewrite <- (Zplus_0_l y). rewrite <- (Zplus_opp_r z).
- rewrite Zplus_assoc_reverse. apply Zplus_lt_compat_l. rewrite Zplus_comm. assumption.
-Qed.
-
-Lemma Zeq_plus_swap : forall n m p:Z, n + p = m <-> n = m - p.
-Proof.
- intros x y z; intros. split. intro. apply Zplus_minus_eq. symmetry in |- *. rewrite Zplus_comm.
- assumption.
- intro. rewrite H. unfold Zminus in |- *. rewrite Zplus_assoc_reverse.
- rewrite Zplus_opp_l. apply Zplus_0_r.
-Qed.
-
-Lemma Zlt_minus_simpl_swap : forall n m:Z, 0 < m -> n - m < n.
-Proof.
- intros n m H; apply Zplus_lt_reg_l with (p := m); rewrite Zplus_minus;
- pattern n at 1 in |- *; rewrite <- (Zplus_0_r n);
- rewrite (Zplus_comm m n); apply Zplus_lt_compat_l;
- assumption.
-Qed.
-
-Lemma Zlt_0_minus_lt : forall n m:Z, 0 < n - m -> m < n.
-Proof.
- intros n m H; apply Zplus_lt_reg_l with (p := - m); rewrite Zplus_opp_l;
- rewrite Zplus_comm; exact H.
-Qed.
+Notation Zle_plus_swap := Z.le_add_le_sub_r (only parsing).
+Notation Zlt_plus_swap := Z.lt_add_lt_sub_r (only parsing).
+Notation Zlt_minus_simpl_swap := Z.lt_sub_pos (only parsing).
-Lemma Zle_0_minus_le : forall n m:Z, 0 <= n - m -> m <= n.
+Lemma Zeq_plus_swap n m p : n + p = m <-> n = m - p.
Proof.
- intros n m H; apply Zplus_le_reg_l with (p := - m); rewrite Zplus_opp_l;
- rewrite Zplus_comm; exact H.
+ apply Z.add_move_r.
Qed.
-Lemma Zle_minus_le_0 : forall n m:Z, m <= n -> 0 <= n - m.
+Lemma Zlt_0_minus_lt n m : 0 < n - m -> m < n.
Proof.
- intros n m H; unfold Zminus; apply Zplus_le_reg_r with (p := m);
- rewrite <- Zplus_assoc; rewrite Zplus_opp_l; rewrite Zplus_0_r; exact H.
+ apply Z.lt_0_sub.
Qed.
-Lemma Zmult_lt_compat:
- forall n m p q : Z, 0 <= n < p -> 0 <= m < q -> n * m < p * q.
+Lemma Zle_0_minus_le n m : 0 <= n - m -> m <= n.
Proof.
- intros n m p q (H1, H2) (H3,H4).
- assert (0<p) by (apply Zle_lt_trans with n; auto).
- assert (0<q) by (apply Zle_lt_trans with m; auto).
- case Zle_lt_or_eq with (1 := H1); intros H5; auto with zarith.
- case Zle_lt_or_eq with (1 := H3); intros H6; auto with zarith.
- apply Zlt_trans with (n * q).
- apply Zmult_lt_compat_l; auto.
- apply Zmult_lt_compat_r; auto with zarith.
- rewrite <- H6; rewrite Zmult_0_r; apply Zmult_lt_0_compat; auto with zarith.
- rewrite <- H5; simpl; apply Zmult_lt_0_compat; auto with zarith.
+ apply Z.le_0_sub.
Qed.
-Lemma Zmult_lt_compat2:
- forall n m p q : Z, 0 < n <= p -> 0 < m < q -> n * m < p * q.
+Lemma Zle_minus_le_0 n m : m <= n -> 0 <= n - m.
Proof.
- intros n m p q (H1, H2) (H3, H4).
- apply Zle_lt_trans with (p * m).
- apply Zmult_le_compat_r; auto.
- apply Zlt_le_weak; auto.
- apply Zmult_lt_compat_l; auto.
- apply Zlt_le_trans with n; auto.
+ apply Z.le_0_sub.
Qed.
(** For compatibility *)
diff --git a/theories/ZArith/Zpow_alt.v b/theories/ZArith/Zpow_alt.v
new file mode 100644
index 00000000..a90eedb4
--- /dev/null
+++ b/theories/ZArith/Zpow_alt.v
@@ -0,0 +1,83 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import BinInt.
+Local Open Scope Z_scope.
+
+(** An alternative power function for Z *)
+
+(** This [Zpower_alt] is extensionnaly equal to [Z.pow],
+ but not convertible with it. The number of
+ multiplications is logarithmic instead of linear, but
+ these multiplications are bigger. Experimentally, it seems
+ that [Zpower_alt] is slightly quicker than [Z.pow] on average,
+ but can be quite slower on powers of 2.
+*)
+
+Definition Zpower_alt n m :=
+ match m with
+ | Z0 => 1
+ | Zpos p => Pos.iter_op Z.mul p n
+ | Zneg p => 0
+ end.
+
+Infix "^^" := Zpower_alt (at level 30, right associativity) : Z_scope.
+
+Lemma Piter_mul_acc : forall f,
+ (forall x y:Z, (f x)*y = f (x*y)) ->
+ forall p k, Pos.iter p f k = (Pos.iter p f 1)*k.
+Proof.
+ intros f Hf.
+ induction p; simpl; intros.
+ - set (g := Pos.iter p f 1) in *. now rewrite !IHp, Hf, Z.mul_assoc.
+ - set (g := Pos.iter p f 1) in *. now rewrite !IHp, Z.mul_assoc.
+ - now rewrite Hf, Z.mul_1_l.
+Qed.
+
+Lemma Piter_op_square : forall p a,
+ Pos.iter_op Z.mul p (a*a) = (Pos.iter_op Z.mul p a)*(Pos.iter_op Z.mul p a).
+Proof.
+ induction p; simpl; intros; trivial. now rewrite IHp, Z.mul_shuffle1.
+Qed.
+
+Lemma Zpower_equiv a b : a^^b = a^b.
+Proof.
+ destruct b as [|p|p]; trivial.
+ unfold Zpower_alt, Z.pow, Z.pow_pos.
+ revert a.
+ induction p; simpl; intros.
+ - f_equal.
+ rewrite Piter_mul_acc.
+ now rewrite Piter_op_square, IHp.
+ intros. symmetry; apply Z.mul_assoc.
+ - rewrite Piter_mul_acc.
+ now rewrite Piter_op_square, IHp.
+ intros. symmetry; apply Z.mul_assoc.
+ - now Z.nzsimpl.
+Qed.
+
+Lemma Zpower_alt_0_r n : n^^0 = 1.
+Proof. reflexivity. Qed.
+
+Lemma Zpower_alt_succ_r a b : 0<=b -> a^^(Z.succ b) = a * a^^b.
+Proof.
+ destruct b as [|b|b]; intros Hb; simpl.
+ - now Z.nzsimpl.
+ - now rewrite Pos.add_1_r, Pos.iter_op_succ by apply Z.mul_assoc.
+ - now elim Hb.
+Qed.
+
+Lemma Zpower_alt_neg_r a b : b<0 -> a^^b = 0.
+Proof.
+ now destruct b.
+Qed.
+
+Lemma Zpower_alt_Ppow p q : (Zpos p)^^(Zpos q) = Zpos (p^q).
+Proof.
+ now rewrite Zpower_equiv, Z.pow_Zpos.
+Qed.
diff --git a/theories/ZArith/Zpow_def.v b/theories/ZArith/Zpow_def.v
index 620d6324..6f1ebc06 100644
--- a/theories/ZArith/Zpow_def.v
+++ b/theories/ZArith/Zpow_def.v
@@ -1,27 +1,31 @@
-Require Import ZArith_base.
-Require Import Ring_theory.
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
-Open Local Scope Z_scope.
+Require Import BinInt Ring_theory.
+Local Open Scope Z_scope.
-(** [Zpower_pos z n] is the n-th power of [z] when [n] is an binary
- integer (type [positive]) and [z] a signed integer (type [Z]) *)
-Definition Zpower_pos (z:Z) (n:positive) := iter_pos n Z (fun x:Z => z * x) 1.
+(** * Power functions over [Z] *)
-Definition Zpower (x y:Z) :=
- match y with
- | Zpos p => Zpower_pos x p
- | Z0 => 1
- | Zneg p => 0
- end.
+(** Nota : this file is mostly deprecated. The definition of [Z.pow]
+ and its usual properties are now provided by module [BinInt.Z]. *)
-Lemma Zpower_theory : power_theory 1 Zmult (eq (A:=Z)) Z_of_N Zpower.
+Notation Zpower_pos := Z.pow_pos (only parsing).
+Notation Zpower := Z.pow (only parsing).
+Notation Zpower_0_r := Z.pow_0_r (only parsing).
+Notation Zpower_succ_r := Z.pow_succ_r (only parsing).
+Notation Zpower_neg_r := Z.pow_neg_r (only parsing).
+Notation Zpower_Ppow := Z.pow_Zpos (only parsing).
+
+Lemma Zpower_theory : power_theory 1 Z.mul (@eq Z) Z.of_N Z.pow.
Proof.
constructor. intros.
destruct n;simpl;trivial.
- unfold Zpower_pos.
- assert (forall k, iter_pos p Z (fun x : Z => r * x) k = pow_pos Zmult r p*k).
- induction p;simpl;intros;repeat rewrite IHp;trivial;
- repeat rewrite Zmult_assoc;trivial.
- rewrite H;rewrite Zmult_1_r;trivial.
+ unfold Z.pow_pos.
+ rewrite <- (Z.mul_1_r (pow_pos _ _ _)). generalize 1.
+ induction p; simpl; intros; rewrite ?IHp, ?Z.mul_assoc; trivial.
Qed.
-
diff --git a/theories/ZArith/Zpow_facts.v b/theories/ZArith/Zpow_facts.v
index 7879fe42..27e3def4 100644
--- a/theories/ZArith/Zpow_facts.v
+++ b/theories/ZArith/Zpow_facts.v
@@ -1,295 +1,109 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zpow_facts.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-Require Import ZArith_base.
-Require Import ZArithRing.
-Require Import Zcomplements.
+Require Import ZArith_base ZArithRing Zcomplements Zdiv Znumtheory.
Require Export Zpower.
-Require Import Zdiv.
-Require Import Znumtheory.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
-Lemma Zpower_pos_1_r: forall x, Zpower_pos x 1 = x.
-Proof.
- intros x; unfold Zpower_pos; simpl; auto with zarith.
-Qed.
+(** Properties of the power function over [Z] *)
-Lemma Zpower_pos_1_l: forall p, Zpower_pos 1 p = 1.
-Proof.
- induction p.
- (* xI *)
- rewrite xI_succ_xO, <-Pplus_diag, Pplus_one_succ_l.
- repeat rewrite Zpower_pos_is_exp.
- rewrite Zpower_pos_1_r, IHp; auto.
- (* xO *)
- rewrite <- Pplus_diag.
- repeat rewrite Zpower_pos_is_exp.
- rewrite IHp; auto.
- (* xH *)
- rewrite Zpower_pos_1_r; auto.
-Qed.
+(** Nota: the usual properties of [Z.pow] are now already provided
+ by [BinInt.Z]. Only remain here some compatibility elements,
+ as well as more specific results about power and modulo and/or
+ primality. *)
-Lemma Zpower_pos_0_l: forall p, Zpower_pos 0 p = 0.
-Proof.
- induction p.
- change (xI p) with (1 + (xO p))%positive.
- rewrite Zpower_pos_is_exp, Zpower_pos_1_r; auto.
- rewrite <- Pplus_diag.
- rewrite Zpower_pos_is_exp, IHp; auto.
- rewrite Zpower_pos_1_r; auto.
-Qed.
+Lemma Zpower_pos_1_r x : Z.pow_pos x 1 = x.
+Proof (Z.pow_1_r x).
-Lemma Zpower_pos_pos: forall x p,
- 0 < x -> 0 < Zpower_pos x p.
-Proof.
- induction p; intros.
- (* xI *)
- rewrite xI_succ_xO, <-Pplus_diag, Pplus_one_succ_l.
- repeat rewrite Zpower_pos_is_exp.
- rewrite Zpower_pos_1_r.
- repeat apply Zmult_lt_0_compat; auto.
- (* xO *)
- rewrite <- Pplus_diag.
- repeat rewrite Zpower_pos_is_exp.
- repeat apply Zmult_lt_0_compat; auto.
- (* xH *)
- rewrite Zpower_pos_1_r; auto.
-Qed.
+Lemma Zpower_pos_1_l p : Z.pow_pos 1 p = 1.
+Proof. now apply (Z.pow_1_l (Zpos p)). Qed.
+Lemma Zpower_pos_0_l p : Z.pow_pos 0 p = 0.
+Proof. now apply (Z.pow_0_l (Zpos p)). Qed.
-Theorem Zpower_1_r: forall z, z^1 = z.
-Proof.
- exact Zpower_pos_1_r.
-Qed.
-
-Theorem Zpower_1_l: forall z, 0 <= z -> 1^z = 1.
-Proof.
- destruct z; simpl; auto.
- intros; apply Zpower_pos_1_l.
- intros; compute in H; elim H; auto.
-Qed.
+Lemma Zpower_pos_pos x p : 0 < x -> 0 < Z.pow_pos x p.
+Proof. intros. now apply (Z.pow_pos_nonneg x (Zpos p)). Qed.
-Theorem Zpower_0_l: forall z, z<>0 -> 0^z = 0.
-Proof.
- destruct z; simpl; auto with zarith.
- intros; apply Zpower_pos_0_l.
-Qed.
+Notation Zpower_1_r := Z.pow_1_r (only parsing).
+Notation Zpower_1_l := Z.pow_1_l (only parsing).
+Notation Zpower_0_l := Z.pow_0_l' (only parsing).
+Notation Zpower_0_r := Z.pow_0_r (only parsing).
+Notation Zpower_2 := Z.pow_2_r (only parsing).
+Notation Zpower_gt_0 := Z.pow_pos_nonneg (only parsing).
+Notation Zpower_ge_0 := Z.pow_nonneg (only parsing).
+Notation Zpower_Zabs := Z.abs_pow (only parsing).
+Notation Zpower_Zsucc := Z.pow_succ_r (only parsing).
+Notation Zpower_mult := Z.pow_mul_r (only parsing).
+Notation Zpower_le_monotone2 := Z.pow_le_mono_r (only parsing).
-Theorem Zpower_0_r: forall z, z^0 = 1.
-Proof.
- simpl; auto.
-Qed.
-
-Theorem Zpower_2: forall z, z^2 = z * z.
-Proof.
- intros; ring.
-Qed.
-
-Theorem Zpower_gt_0: forall x y,
- 0 < x -> 0 <= y -> 0 < x^y.
-Proof.
- destruct y; simpl; auto with zarith.
- intros; apply Zpower_pos_pos; auto.
- intros; compute in H0; elim H0; auto.
-Qed.
-
-Theorem Zpower_Zabs: forall a b, Zabs (a^b) = (Zabs a)^b.
-Proof.
- intros a b; case (Zle_or_lt 0 b).
- intros Hb; pattern b; apply natlike_ind; auto with zarith.
- intros x Hx Hx1; unfold Zsucc.
- (repeat rewrite Zpower_exp); auto with zarith.
- rewrite Zabs_Zmult; rewrite Hx1.
- f_equal; auto.
- replace (a ^ 1) with a; auto.
- simpl; unfold Zpower_pos; simpl; rewrite Zmult_1_r; auto.
- simpl; unfold Zpower_pos; simpl; rewrite Zmult_1_r; auto.
- case b; simpl; auto with zarith.
- intros p Hp; discriminate.
-Qed.
-
-Theorem Zpower_Zsucc: forall p n, 0 <= n -> p^(Zsucc n) = p * p^n.
-Proof.
- intros p n H.
- unfold Zsucc; rewrite Zpower_exp; auto with zarith.
- rewrite Zpower_1_r; apply Zmult_comm.
-Qed.
-
-Theorem Zpower_mult: forall p q r, 0 <= q -> 0 <= r -> p^(q*r) = (p^q)^r.
-Proof.
- intros p q r H1 H2; generalize H2; pattern r; apply natlike_ind; auto.
- intros H3; rewrite Zmult_0_r; repeat rewrite Zpower_exp_0; auto.
- intros r1 H3 H4 H5.
- unfold Zsucc; rewrite Zpower_exp; auto with zarith.
- rewrite <- H4; try rewrite Zpower_1_r; try rewrite <- Zpower_exp; try f_equal; auto with zarith.
- ring.
- apply Zle_ge; replace 0 with (0 * r1); try apply Zmult_le_compat_r; auto.
-Qed.
-
-Theorem Zpower_le_monotone: forall a b c,
+Theorem Zpower_le_monotone a b c :
0 < a -> 0 <= b <= c -> a^b <= a^c.
-Proof.
- intros a b c H (H1, H2).
- rewrite <- (Zmult_1_r (a ^ b)); replace c with (b + (c - b)); auto with zarith.
- rewrite Zpower_exp; auto with zarith.
- apply Zmult_le_compat_l; auto with zarith.
- assert (0 < a ^ (c - b)); auto with zarith.
- apply Zpower_gt_0; auto with zarith.
- apply Zlt_le_weak; apply Zpower_gt_0; auto with zarith.
-Qed.
+Proof. intros. now apply Z.pow_le_mono_r. Qed.
-Theorem Zpower_lt_monotone: forall a b c,
+Theorem Zpower_lt_monotone a b c :
1 < a -> 0 <= b < c -> a^b < a^c.
-Proof.
- intros a b c H (H1, H2).
- rewrite <- (Zmult_1_r (a ^ b)); replace c with (b + (c - b)); auto with zarith.
- rewrite Zpower_exp; auto with zarith.
- apply Zmult_lt_compat_l; auto with zarith.
- apply Zpower_gt_0; auto with zarith.
- assert (0 < a ^ (c - b)); auto with zarith.
- apply Zpower_gt_0; auto with zarith.
- apply Zlt_le_trans with (a ^1); auto with zarith.
- rewrite Zpower_1_r; auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
-Qed.
-
-Theorem Zpower_gt_1 : forall x y,
- 1 < x -> 0 < y -> 1 < x^y.
-Proof.
- intros x y H1 H2.
- replace 1 with (x ^ 0) by apply Zpower_0_r.
- apply Zpower_lt_monotone; auto with zarith.
-Qed.
+Proof. intros. apply Z.pow_lt_mono_r; auto with zarith. Qed.
-Theorem Zpower_ge_0: forall x y, 0 <= x -> 0 <= x^y.
-Proof.
- intros x y; case y; auto with zarith.
- simpl ; auto with zarith.
- intros p H1; assert (H: 0 <= Zpos p); auto with zarith.
- generalize H; pattern (Zpos p); apply natlike_ind; auto with zarith.
- intros p1 H2 H3 _; unfold Zsucc; rewrite Zpower_exp; simpl; auto with zarith.
- apply Zmult_le_0_compat; auto with zarith.
- generalize H1; case x; compute; intros; auto; try discriminate.
-Qed.
-
-Theorem Zpower_le_monotone2:
- forall a b c, 0 < a -> b <= c -> a^b <= a^c.
-Proof.
- intros a b c H H2.
- destruct (Z_le_gt_dec 0 b).
- apply Zpower_le_monotone; auto.
- replace (a^b) with 0.
- destruct (Z_le_gt_dec 0 c).
- destruct (Zle_lt_or_eq _ _ z0).
- apply Zlt_le_weak;apply Zpower_gt_0;trivial.
- rewrite <- H0;simpl;auto with zarith.
- replace (a^c) with 0. auto with zarith.
- destruct c;trivial;unfold Zgt in z0;discriminate z0.
- destruct b;trivial;unfold Zgt in z;discriminate z.
-Qed.
+Theorem Zpower_gt_1 x y : 1 < x -> 0 < y -> 1 < x^y.
+Proof. apply Z.pow_gt_1. Qed.
-Theorem Zmult_power: forall p q r, 0 <= r ->
- (p*q)^r = p^r * q^r.
-Proof.
- intros p q r H1; generalize H1; pattern r; apply natlike_ind; auto.
- clear r H1; intros r H1 H2 H3.
- unfold Zsucc; rewrite Zpower_exp; auto with zarith.
- rewrite H2; repeat rewrite Zpower_exp; auto with zarith; ring.
-Qed.
+Theorem Zmult_power p q r : 0 <= r -> (p*q)^r = p^r * q^r.
+Proof. intros. apply Z.pow_mul_l. Qed.
-Hint Resolve Zpower_ge_0 Zpower_gt_0: zarith.
+Hint Resolve Z.pow_nonneg Z.pow_pos_nonneg : zarith.
-Theorem Zpower_le_monotone3: forall a b c,
+Theorem Zpower_le_monotone3 a b c :
0 <= c -> 0 <= a <= b -> a^c <= b^c.
-Proof.
- intros a b c H (H1, H2).
- generalize H; pattern c; apply natlike_ind; auto.
- intros x HH HH1 _; unfold Zsucc; repeat rewrite Zpower_exp; auto with zarith.
- repeat rewrite Zpower_1_r.
- apply Zle_trans with (a^x * b); auto with zarith.
-Qed.
+Proof. intros. now apply Z.pow_le_mono_l. Qed.
-Lemma Zpower_le_monotone_inv: forall a b c,
+Lemma Zpower_le_monotone_inv a b c :
1 < a -> 0 < b -> a^b <= a^c -> b <= c.
Proof.
- intros a b c H H0 H1.
- destruct (Z_le_gt_dec b c);trivial.
- assert (2 <= a^b).
- apply Zle_trans with (2^b).
- pattern 2 at 1;replace 2 with (2^1);trivial.
- apply Zpower_le_monotone;auto with zarith.
- apply Zpower_le_monotone3;auto with zarith.
- assert (c > 0).
- destruct (Z_le_gt_dec 0 c);trivial.
- destruct (Zle_lt_or_eq _ _ z0);auto with zarith.
- rewrite <- H3 in H1;simpl in H1; exfalso;omega.
- destruct c;try discriminate z0. simpl in H1. exfalso;omega.
- assert (H4 := Zpower_lt_monotone a c b H). exfalso;omega.
+ intros Ha Hb H. apply (Z.pow_le_mono_r_iff a); trivial.
+ apply Z.lt_le_incl; apply (Z.pow_gt_1 a); trivial.
+ apply Z.lt_le_trans with (a^b); trivial. now apply Z.pow_gt_1.
Qed.
-Theorem Zpower_nat_Zpower: forall p q, 0 <= q ->
- p^q = Zpower_nat p (Zabs_nat q).
-Proof.
- intros p1 q1; case q1; simpl.
- intros _; exact (refl_equal _).
- intros p2 _; apply Zpower_pos_nat.
- intros p2 H1; case H1; auto.
-Qed.
+Notation Zpower_nat_Zpower := Zpower_nat_Zpower (only parsing).
-Theorem Zpower2_lt_lin: forall n, 0 <= n -> n < 2^n.
-Proof.
- intros n; apply (natlike_ind (fun n => n < 2 ^n)); clear n.
- simpl; auto with zarith.
- intros n H1 H2; unfold Zsucc.
- case (Zle_lt_or_eq _ _ H1); clear H1; intros H1.
- apply Zle_lt_trans with (n + n); auto with zarith.
- rewrite Zpower_exp; auto with zarith.
- rewrite Zpower_1_r.
- assert (tmp: forall p, p * 2 = p + p); intros; try ring;
- rewrite tmp; auto with zarith.
- subst n; simpl; unfold Zpower_pos; simpl; auto with zarith.
-Qed.
+Theorem Zpower2_lt_lin n : 0 <= n -> n < 2^n.
+Proof. intros. now apply Z.pow_gt_lin_r. Qed.
-Theorem Zpower2_le_lin: forall n, 0 <= n -> n <= 2^n.
-Proof.
- intros; apply Zlt_le_weak; apply Zpower2_lt_lin; auto.
-Qed.
+Theorem Zpower2_le_lin n : 0 <= n -> n <= 2^n.
+Proof. intros. apply Z.lt_le_incl. now apply Z.pow_gt_lin_r. Qed.
-Lemma Zpower2_Psize :
- forall n p, Zpos p < 2^(Z_of_nat n) <-> (Psize p <= n)%nat.
+Lemma Zpower2_Psize n p :
+ Zpos p < 2^(Z.of_nat n) <-> (Pos.size_nat p <= n)%nat.
Proof.
- induction n.
- destruct p; split; intros H; discriminate H || inversion H.
- destruct p; simpl Psize.
- rewrite inj_S, Zpower_Zsucc; auto with zarith.
- rewrite Zpos_xI; specialize IHn with p; omega.
- rewrite inj_S, Zpower_Zsucc; auto with zarith.
- rewrite Zpos_xO; specialize IHn with p; omega.
- split; auto with arith.
- intros _; apply Zpower_gt_1; auto with zarith.
- rewrite inj_S; generalize (Zle_0_nat n); omega.
+ revert p; induction n.
+ destruct p; now split.
+ assert (Hn := Nat2Z.is_nonneg n).
+ destruct p; simpl Pos.size_nat.
+ - specialize IHn with p.
+ rewrite Z.pos_xI, Nat2Z.inj_succ, Z.pow_succ_r; omega.
+ - specialize IHn with p.
+ rewrite Z.pos_xO, Nat2Z.inj_succ, Z.pow_succ_r; omega.
+ - split; auto with zarith.
+ intros _. apply Z.pow_gt_1. easy.
+ now rewrite Nat2Z.inj_succ, Z.lt_succ_r.
Qed.
(** * Zpower and modulo *)
-Theorem Zpower_mod: forall p q n, 0 < n ->
- (p^q) mod n = ((p mod n)^q) mod n.
+Theorem Zpower_mod p q n :
+ 0 < n -> (p^q) mod n = ((p mod n)^q) mod n.
Proof.
- intros p q n Hn; case (Zle_or_lt 0 q); intros H1.
- generalize H1; pattern q; apply natlike_ind; auto.
- intros q1 Hq1 Rec _; unfold Zsucc; repeat rewrite Zpower_exp; repeat rewrite Zpower_1_r; auto with zarith.
- rewrite (fun x => (Zmult_mod x p)); try rewrite Rec; auto with zarith.
- rewrite (fun x y => (Zmult_mod (x ^y))); try f_equal; auto with zarith.
- f_equal; auto; apply sym_equal; apply Zmod_mod; auto with zarith.
- generalize H1; case q; simpl; auto.
- intros; discriminate.
+ intros Hn; destruct (Z.le_gt_cases 0 q) as [H1|H1].
+ - pattern q; apply natlike_ind; trivial.
+ clear q H1. intros q Hq Rec. rewrite !Z.pow_succ_r; trivial.
+ rewrite Z.mul_mod_idemp_l; auto with zarith.
+ rewrite Z.mul_mod, Rec, <- Z.mul_mod; auto with zarith.
+ - rewrite !Z.pow_neg_r; auto with zarith.
Qed.
(** A direct way to compute Zpower modulo **)
@@ -313,153 +127,113 @@ Fixpoint Zpow_mod_pos (a: Z)(m: positive)(n : Z) : Z :=
Definition Zpow_mod a m n :=
match m with
- | 0 => 1
+ | 0 => 1 mod n
| Zpos p => Zpow_mod_pos a p n
| Zneg p => 0
end.
-Theorem Zpow_mod_pos_correct: forall a m n, 0 < n ->
- Zpow_mod_pos a m n = (Zpower_pos a m) mod n.
+Theorem Zpow_mod_pos_correct a m n :
+ n <> 0 -> Zpow_mod_pos a m n = (Z.pow_pos a m) mod n.
Proof.
- intros a m; elim m; simpl; auto.
- intros p Rec n H1; rewrite xI_succ_xO, Pplus_one_succ_r, <-Pplus_diag; auto.
- repeat rewrite Zpower_pos_is_exp; auto.
- repeat rewrite Rec; auto.
- rewrite Zpower_pos_1_r.
- repeat rewrite (fun x => (Zmult_mod x a)); auto with zarith.
- rewrite (Zmult_mod (Zpower_pos a p)); auto with zarith.
- case (Zpower_pos a p mod n); auto.
- intros p Rec n H1; rewrite <- Pplus_diag; auto.
- repeat rewrite Zpower_pos_is_exp; auto.
- repeat rewrite Rec; auto.
- rewrite (Zmult_mod (Zpower_pos a p)); auto with zarith.
- case (Zpower_pos a p mod n); auto.
- unfold Zpower_pos; simpl; rewrite Zmult_1_r; auto with zarith.
+ intros Hn. induction m.
+ - rewrite Pos.xI_succ_xO at 2. rewrite <- Pos.add_1_r, <- Pos.add_diag.
+ rewrite 2 Zpower_pos_is_exp, Zpower_pos_1_r.
+ rewrite Z.mul_mod, (Z.mul_mod (Z.pow_pos a m)) by trivial.
+ rewrite <- IHm, <- Z.mul_mod by trivial.
+ simpl. now destruct (Zpow_mod_pos a m n).
+ - rewrite <- Pos.add_diag at 2.
+ rewrite Zpower_pos_is_exp.
+ rewrite Z.mul_mod by trivial.
+ rewrite <- IHm.
+ simpl. now destruct (Zpow_mod_pos a m n).
+ - now rewrite Zpower_pos_1_r.
Qed.
-Theorem Zpow_mod_correct: forall a m n, 1 < n -> 0 <= m ->
- Zpow_mod a m n = (a ^ m) mod n.
+Theorem Zpow_mod_correct a m n :
+ n <> 0 -> Zpow_mod a m n = (a ^ m) mod n.
Proof.
- intros a m n; case m; simpl.
- intros; apply sym_equal; apply Zmod_small; auto with zarith.
- intros; apply Zpow_mod_pos_correct; auto with zarith.
- intros p H H1; case H1; auto.
+ intros Hn. destruct m; simpl.
+ - trivial.
+ - apply Zpow_mod_pos_correct; auto with zarith.
+ - rewrite Z.mod_0_l; auto with zarith.
Qed.
(* Complements about power and number theory. *)
-Lemma Zpower_divide: forall p q, 0 < q -> (p | p ^ q).
+Lemma Zpower_divide p q : 0 < q -> (p | p ^ q).
Proof.
- intros p q H; exists (p ^(q - 1)).
- pattern p at 3; rewrite <- (Zpower_1_r p); rewrite <- Zpower_exp; try f_equal; auto with zarith.
+ exists (p^(q - 1)).
+ rewrite Z.mul_comm, <- Z.pow_succ_r; f_equal; auto with zarith.
Qed.
-Theorem rel_prime_Zpower_r: forall i p q, 0 < i ->
- rel_prime p q -> rel_prime p (q^i).
+Theorem rel_prime_Zpower_r i p q :
+ 0 <= i -> rel_prime p q -> rel_prime p (q^i).
Proof.
- intros i p q Hi Hpq; generalize Hi; pattern i; apply natlike_ind; auto with zarith; clear i Hi.
- intros H; contradict H; auto with zarith.
- intros i Hi Rec _; rewrite Zpower_Zsucc; auto.
+ intros Hi Hpq; pattern i; apply natlike_ind; auto with zarith.
+ simpl. apply rel_prime_sym, rel_prime_1.
+ clear i Hi. intros i Hi Rec; rewrite Z.pow_succ_r; auto.
apply rel_prime_mult; auto.
- case Zle_lt_or_eq with (1 := Hi); intros Hi1; subst; auto.
- rewrite Zpower_0_r; apply rel_prime_sym; apply rel_prime_1.
Qed.
-Theorem rel_prime_Zpower: forall i j p q, 0 <= i -> 0 <= j ->
- rel_prime p q -> rel_prime (p^i) (q^j).
+Theorem rel_prime_Zpower i j p q :
+ 0 <= i -> 0 <= j -> rel_prime p q -> rel_prime (p^i) (q^j).
Proof.
- intros i j p q Hi; generalize Hi j p q; pattern i; apply natlike_ind; auto with zarith; clear i Hi j p q.
- intros _ j p q H H1; rewrite Zpower_0_r; apply rel_prime_1.
- intros n Hn Rec _ j p q Hj Hpq.
- rewrite Zpower_Zsucc; auto.
- case Zle_lt_or_eq with (1 := Hj); intros Hj1; subst.
- apply rel_prime_sym; apply rel_prime_mult; auto.
- apply rel_prime_sym; apply rel_prime_Zpower_r; auto with arith.
- apply rel_prime_sym; apply Rec; auto.
- rewrite Zpower_0_r; apply rel_prime_sym; apply rel_prime_1.
+ intros Hi Hj H. apply rel_prime_Zpower_r; trivial.
+ apply rel_prime_sym. apply rel_prime_Zpower_r; trivial.
+ now apply rel_prime_sym.
Qed.
-Theorem prime_power_prime: forall p q n, 0 <= n ->
- prime p -> prime q -> (p | q^n) -> p = q.
+Theorem prime_power_prime p q n :
+ 0 <= n -> prime p -> prime q -> (p | q^n) -> p = q.
Proof.
- intros p q n Hn Hp Hq; pattern n; apply natlike_ind; auto; clear n Hn.
- rewrite Zpower_0_r; intros.
- assert (2<=p) by (apply prime_ge_2; auto).
- assert (p<=1) by (apply Zdivide_le; auto with zarith).
- omega.
- intros n1 H H1.
- unfold Zsucc; rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith.
- assert (2<=p) by (apply prime_ge_2; auto).
- assert (2<=q) by (apply prime_ge_2; auto).
- intros H3; case prime_mult with (2 := H3); auto.
- intros; apply prime_div_prime; auto.
+ intros Hn Hp Hq; pattern n; apply natlike_ind; auto; clear n Hn.
+ - simpl; intros.
+ assert (2<=p) by (apply prime_ge_2; auto).
+ assert (p<=1) by (apply Z.divide_pos_le; auto with zarith).
+ omega.
+ - intros n Hn Rec.
+ rewrite Z.pow_succ_r by trivial. intros.
+ assert (2<=p) by (apply prime_ge_2; auto).
+ assert (2<=q) by (apply prime_ge_2; auto).
+ destruct prime_mult with (2 := H); auto.
+ apply prime_div_prime; auto.
Qed.
-Theorem Zdivide_power_2: forall x p n, 0 <= n -> 0 <= x -> prime p ->
- (x | p^n) -> exists m, x = p^m.
+Theorem Zdivide_power_2 x p n :
+ 0 <= n -> 0 <= x -> prime p -> (x | p^n) -> exists m, x = p^m.
Proof.
- intros x p n Hn Hx; revert p n Hn; generalize Hx.
+ intros Hn Hx; revert p n Hn. generalize Hx.
pattern x; apply Z_lt_induction; auto.
clear x Hx; intros x IH Hx p n Hn Hp H.
- case Zle_lt_or_eq with (1 := Hx); auto; clear Hx; intros Hx; subst.
- case (Zle_lt_or_eq 1 x); auto with zarith; clear Hx; intros Hx; subst.
+ Z.le_elim Hx; subst.
+ apply Z.le_succ_l in Hx; simpl in Hx.
+ Z.le_elim Hx; subst.
(* x > 1 *)
- case (prime_dec x); intros H2.
- exists 1; rewrite Zpower_1_r; apply prime_power_prime with n; auto.
- case not_prime_divide with (2 := H2); auto.
- intros p1 ((H3, H4), (q1, Hq1)); subst.
- case (IH p1) with p n; auto with zarith.
- apply Zdivide_trans with (2 := H); exists q1; auto with zarith.
- intros r1 Hr1.
- case (IH q1) with p n; auto with zarith.
- case (Zle_lt_or_eq 0 q1).
- apply Zmult_le_0_reg_r with p1; auto with zarith.
+ case (prime_dec x); intros Hpr.
+ exists 1; rewrite Z.pow_1_r; apply prime_power_prime with n; auto.
+ case not_prime_divide with (2 := Hpr); auto.
+ intros p1 ((Hp1, Hpq1),(q1,->)).
+ assert (Hq1 : 0 < q1) by (apply Z.mul_lt_mono_pos_r with p1; auto with zarith).
+ destruct (IH p1) with p n as (r1,Hr1); auto with zarith.
+ transitivity (q1 * p1); trivial. exists q1; auto with zarith.
+ destruct (IH q1) with p n as (r2,Hr2); auto with zarith.
split; auto with zarith.
- pattern q1 at 1; replace q1 with (q1 * 1); auto with zarith.
- apply Zmult_lt_compat_l; auto with zarith.
- intros H5; subst; contradict Hx; auto with zarith.
- apply Zmult_le_0_reg_r with p1; auto with zarith.
- apply Zdivide_trans with (2 := H); exists p1; auto with zarith.
- intros r2 Hr2; exists (r2 + r1); subst.
- apply sym_equal; apply Zpower_exp.
- generalize Hx; case r2; simpl; auto with zarith.
- intros; red; simpl; intros; discriminate.
- generalize H3; case r1; simpl; auto with zarith.
- intros; red; simpl; intros; discriminate.
+ rewrite <- (Z.mul_1_r q1) at 1.
+ apply Z.mul_lt_mono_pos_l; auto with zarith.
+ transitivity (q1 * p1); trivial. exists p1; auto with zarith.
+ exists (r2 + r1); subst.
+ symmetry. apply Z.pow_add_r.
+ generalize Hq1; case r2; now auto with zarith.
+ generalize Hp1; case r1; now auto with zarith.
(* x = 1 *)
- exists 0; rewrite Zpower_0_r; auto.
+ exists 0; rewrite Z.pow_0_r; auto.
(* x = 0 *)
- exists n; destruct H; rewrite Zmult_0_r in H; auto.
+ exists n; destruct H; rewrite Z.mul_0_r in H; auto.
Qed.
(** * Zsquare: a direct definition of [z^2] *)
-Fixpoint Psquare (p: positive): positive :=
- match p with
- | xH => xH
- | xO p => xO (xO (Psquare p))
- | xI p => xI (xO (Pplus (Psquare p) p))
- end.
-
-Definition Zsquare p :=
- match p with
- | Z0 => Z0
- | Zpos p => Zpos (Psquare p)
- | Zneg p => Zpos (Psquare p)
- end.
-
-Theorem Psquare_correct: forall p, Psquare p = (p * p)%positive.
-Proof.
- induction p; simpl; auto; f_equal; rewrite IHp.
- apply trans_equal with (xO p + xO (p*p))%positive; auto.
- rewrite (Pplus_comm (xO p)); auto.
- rewrite Pmult_xI_permute_r; rewrite Pplus_assoc.
- f_equal; auto.
- symmetry; apply Pplus_diag.
- symmetry; apply Pmult_xO_permute_r.
-Qed.
-
-Theorem Zsquare_correct: forall p, Zsquare p = p * p.
-Proof.
- intro p; case p; simpl; auto; intros p1; rewrite Psquare_correct; auto.
-Qed.
+Notation Psquare := Pos.square (only parsing).
+Notation Zsquare := Z.square (only parsing).
+Notation Psquare_correct := Pos.square_spec (only parsing).
+Notation Zsquare_correct := Z.square_spec (only parsing).
diff --git a/theories/ZArith/Zpower.v b/theories/ZArith/Zpower.v
index 038748b5..5052d01a 100644
--- a/theories/ZArith/Zpower.v
+++ b/theories/ZArith/Zpower.v
@@ -1,79 +1,89 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Zpower.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-Require Import Wf_nat.
-Require Import ZArith_base.
+Require Import Wf_nat ZArith_base Omega Zcomplements.
Require Export Zpow_def.
-Require Import Omega.
-Require Import Zcomplements.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
+
+(** * Power functions over [Z] *)
-Infix "^" := Zpower : Z_scope.
+(** Nota : this file is mostly deprecated. The definition of [Z.pow]
+ and its usual properties are now provided by module [BinInt.Z].
+ Powers of 2 are also available there (see [Z.shiftl] and [Z.shiftr]).
+ Only remain here:
+ - [Zpower_nat] : a power function with a [nat] exponent
+ - old-style powers of two, such as [two_p]
+ - [Zdiv_rest] : a division + modulo when the divisor is a power of 2
+*)
-(** * Definition of powers over [Z]*)
(** [Zpower_nat z n] is the n-th power of [z] when [n] is an unary
integer (type [nat]) and [z] a signed integer (type [Z]) *)
-Definition Zpower_nat (z:Z) (n:nat) := iter_nat n Z (fun x:Z => z * x) 1.
+Definition Zpower_nat (z:Z) (n:nat) := nat_iter n (Z.mul z) 1.
+
+Lemma Zpower_nat_0_r z : Zpower_nat z 0 = 1.
+Proof. reflexivity. Qed.
+
+Lemma Zpower_nat_succ_r n z : Zpower_nat z (S n) = z * (Zpower_nat z n).
+Proof. reflexivity. Qed.
(** [Zpower_nat_is_exp] says [Zpower_nat] is a morphism for
- [plus : nat->nat] and [Zmult : Z->Z] *)
+ [plus : nat->nat->nat] and [Z.mul : Z->Z->Z] *)
Lemma Zpower_nat_is_exp :
forall (n m:nat) (z:Z),
Zpower_nat z (n + m) = Zpower_nat z n * Zpower_nat z m.
Proof.
- intros; elim n;
- [ simpl in |- *; elim (Zpower_nat z m); auto with zarith
- | unfold Zpower_nat in |- *; intros; simpl in |- *; rewrite H;
- apply Zmult_assoc ].
+ induction n.
+ - intros. now rewrite Zpower_nat_0_r, Z.mul_1_l.
+ - intros. simpl. now rewrite 2 Zpower_nat_succ_r, IHn, Z.mul_assoc.
+Qed.
+
+(** Conversions between powers of unary and binary integers *)
+
+Lemma Zpower_pos_nat (z : Z) (p : positive) :
+ Z.pow_pos z p = Zpower_nat z (Pos.to_nat p).
+Proof.
+ apply Pos2Nat.inj_iter.
Qed.
-(** This theorem shows that powers of unary and binary integers
- are the same thing, modulo the function convert : [positive -> nat] *)
+Lemma Zpower_nat_Z (z : Z) (n : nat) :
+ Zpower_nat z n = z ^ (Z.of_nat n).
+Proof.
+ induction n. trivial.
+ rewrite Zpower_nat_succ_r, Nat2Z.inj_succ, Z.pow_succ_r.
+ now f_equal.
+ apply Nat2Z.is_nonneg.
+Qed.
-Lemma Zpower_pos_nat :
- forall (z:Z) (p:positive), Zpower_pos z p = Zpower_nat z (nat_of_P p).
+Theorem Zpower_nat_Zpower z n : 0 <= n ->
+ z^n = Zpower_nat z (Z.abs_nat n).
Proof.
- intros; unfold Zpower_pos in |- *; unfold Zpower_nat in |- *;
- apply iter_nat_of_P.
+ intros. now rewrite Zpower_nat_Z, Zabs2Nat.id_abs, Z.abs_eq.
Qed.
-(** Using the theorem [Zpower_pos_nat] and the lemma [Zpower_nat_is_exp] we
- deduce that the function [[n:positive](Zpower_pos z n)] is a morphism
- for [add : positive->positive] and [Zmult : Z->Z] *)
+(** The function [(Z.pow_pos z)] is a morphism
+ for [Pos.add : positive->positive->positive] and [Z.mul : Z->Z->Z] *)
-Lemma Zpower_pos_is_exp :
- forall (n m:positive) (z:Z),
- Zpower_pos z (n + m) = Zpower_pos z n * Zpower_pos z m.
+Lemma Zpower_pos_is_exp (n m : positive)(z:Z) :
+ Z.pow_pos z (n + m) = Z.pow_pos z n * Z.pow_pos z m.
Proof.
- intros.
- rewrite (Zpower_pos_nat z n).
- rewrite (Zpower_pos_nat z m).
- rewrite (Zpower_pos_nat z (n + m)).
- rewrite (nat_of_P_plus_morphism n m).
- apply Zpower_nat_is_exp.
+ now apply (Z.pow_add_r z (Zpos n) (Zpos m)).
Qed.
Hint Immediate Zpower_nat_is_exp Zpower_pos_is_exp : zarith.
Hint Unfold Zpower_pos Zpower_nat: zarith.
-Theorem Zpower_exp :
- forall x n m:Z, n >= 0 -> m >= 0 -> x ^ (n + m) = x ^ n * x ^ m.
+Theorem Zpower_exp x n m :
+ n >= 0 -> m >= 0 -> x ^ (n + m) = x ^ n * x ^ m.
Proof.
- destruct n; destruct m; auto with zarith.
- simpl; intros; apply Zred_factor0.
- simpl; auto with zarith.
- intros; compute in H0; elim H0; auto.
- intros; compute in H; elim H; auto.
+ Z.swap_greater. apply Z.pow_add_r.
Qed.
Section Powers_of_2.
@@ -81,178 +91,137 @@ Section Powers_of_2.
(** * Powers of 2 *)
(** For the powers of two, that will be widely used, a more direct
- calculus is possible. We will also prove some properties such
- as [(x:positive) x < 2^x] that are true for all integers bigger
- than 2 but more difficult to prove and useless. *)
-
- (** [shift n m] computes [2^n * m], or [m] shifted by [n] positions *)
+ calculus is possible. [shift n m] computes [2^n * m], i.e.
+ [m] shifted by [n] positions *)
- Definition shift_nat (n:nat) (z:positive) := iter_nat n positive xO z.
- Definition shift_pos (n z:positive) := iter_pos n positive xO z.
+ Definition shift_nat (n:nat) (z:positive) := nat_iter n xO z.
+ Definition shift_pos (n z:positive) := Pos.iter n xO z.
Definition shift (n:Z) (z:positive) :=
match n with
| Z0 => z
- | Zpos p => iter_pos p positive xO z
+ | Zpos p => Pos.iter p xO z
| Zneg p => z
end.
Definition two_power_nat (n:nat) := Zpos (shift_nat n 1).
Definition two_power_pos (x:positive) := Zpos (shift_pos x 1).
- Lemma two_power_nat_S :
- forall n:nat, two_power_nat (S n) = 2 * two_power_nat n.
+ Definition two_p (x:Z) :=
+ match x with
+ | Z0 => 1
+ | Zpos y => two_power_pos y
+ | Zneg y => 0
+ end.
+
+ (** Equivalence with notions defined in BinInt *)
+
+ Lemma shift_nat_equiv n p : shift_nat n p = Pos.shiftl_nat p n.
+ Proof. reflexivity. Qed.
+
+ Lemma shift_pos_equiv n p : shift_pos n p = Pos.shiftl p (Npos n).
+ Proof. reflexivity. Qed.
+
+ Lemma shift_equiv n p : 0<=n -> Zpos (shift n p) = Z.shiftl (Zpos p) n.
Proof.
- intro; simpl in |- *; apply refl_equal.
+ destruct n.
+ - trivial.
+ - simpl; intros. now apply Pos.iter_swap_gen.
+ - now destruct 1.
Qed.
- Lemma shift_nat_plus :
- forall (n m:nat) (x:positive),
- shift_nat (n + m) x = shift_nat n (shift_nat m x).
+ Lemma two_power_nat_equiv n : two_power_nat n = 2 ^ (Z.of_nat n).
Proof.
- intros; unfold shift_nat in |- *; apply iter_nat_plus.
+ induction n.
+ - trivial.
+ - now rewrite Nat2Z.inj_succ, Z.pow_succ_r, <- IHn by apply Nat2Z.is_nonneg.
Qed.
- Theorem shift_nat_correct :
- forall (n:nat) (x:positive), Zpos (shift_nat n x) = Zpower_nat 2 n * Zpos x.
+ Lemma two_power_pos_equiv p : two_power_pos p = 2 ^ Zpos p.
Proof.
- unfold shift_nat in |- *; simple induction n;
- [ simpl in |- *; trivial with zarith
- | intros; replace (Zpower_nat 2 (S n0)) with (2 * Zpower_nat 2 n0);
- [ rewrite <- Zmult_assoc; rewrite <- (H x); simpl in |- *; reflexivity
- | auto with zarith ] ].
+ now apply Pos.iter_swap_gen.
Qed.
- Theorem two_power_nat_correct :
- forall n:nat, two_power_nat n = Zpower_nat 2 n.
+ Lemma two_p_equiv x : two_p x = 2 ^ x.
Proof.
- intro n.
- unfold two_power_nat in |- *.
- rewrite (shift_nat_correct n).
- omega.
+ destruct x; trivial. apply two_power_pos_equiv.
Qed.
- (** Second we show that [two_power_pos] and [two_power_nat] are the same *)
- Lemma shift_pos_nat :
- forall p x:positive, shift_pos p x = shift_nat (nat_of_P p) x.
+ (** Properties of these old versions of powers of two *)
+
+ Lemma two_power_nat_S n : two_power_nat (S n) = 2 * two_power_nat n.
+ Proof. reflexivity. Qed.
+
+ Lemma shift_nat_plus n m x :
+ shift_nat (n + m) x = shift_nat n (shift_nat m x).
Proof.
- unfold shift_pos in |- *.
- unfold shift_nat in |- *.
- intros; apply iter_nat_of_P.
+ apply iter_nat_plus.
Qed.
- Lemma two_power_pos_nat :
- forall p:positive, two_power_pos p = two_power_nat (nat_of_P p).
+ Theorem shift_nat_correct n x :
+ Zpos (shift_nat n x) = Zpower_nat 2 n * Zpos x.
Proof.
- intro; unfold two_power_pos in |- *; unfold two_power_nat in |- *.
- apply f_equal with (f := Zpos).
- apply shift_pos_nat.
+ induction n.
+ - trivial.
+ - now rewrite Zpower_nat_succ_r, <- Z.mul_assoc, <- IHn.
Qed.
- (** Then we deduce that [two_power_pos] is also correct *)
+ Theorem two_power_nat_correct n : two_power_nat n = Zpower_nat 2 n.
+ Proof.
+ now rewrite two_power_nat_equiv, Zpower_nat_Z.
+ Qed.
- Theorem shift_pos_correct :
- forall p x:positive, Zpos (shift_pos p x) = Zpower_pos 2 p * Zpos x.
+ Lemma shift_pos_nat p x : shift_pos p x = shift_nat (Pos.to_nat p) x.
Proof.
- intros.
- rewrite (shift_pos_nat p x).
- rewrite (Zpower_pos_nat 2 p).
- apply shift_nat_correct.
+ apply Pos2Nat.inj_iter.
Qed.
- Theorem two_power_pos_correct :
- forall x:positive, two_power_pos x = Zpower_pos 2 x.
+ Lemma two_power_pos_nat p : two_power_pos p = two_power_nat (Pos.to_nat p).
Proof.
- intro.
- rewrite two_power_pos_nat.
- rewrite Zpower_pos_nat.
- apply two_power_nat_correct.
+ unfold two_power_pos. now rewrite shift_pos_nat.
Qed.
- (** Some consequences *)
+ Theorem shift_pos_correct p x :
+ Zpos (shift_pos p x) = Zpower_pos 2 p * Zpos x.
+ Proof.
+ now rewrite shift_pos_nat, Zpower_pos_nat, shift_nat_correct.
+ Qed.
- Theorem two_power_pos_is_exp :
- forall x y:positive,
- two_power_pos (x + y) = two_power_pos x * two_power_pos y.
+ Theorem two_power_pos_correct x : two_power_pos x = Z.pow_pos 2 x.
Proof.
- intros.
- rewrite (two_power_pos_correct (x + y)).
- rewrite (two_power_pos_correct x).
- rewrite (two_power_pos_correct y).
- apply Zpower_pos_is_exp.
+ apply two_power_pos_equiv.
Qed.
- (** The exponentiation [z -> 2^z] for [z] a signed integer.
- For convenience, we assume that [2^z = 0] for all [z < 0]
- We could also define a inductive type [Log_result] with
- 3 contructors [ Zero | Pos positive -> | minus_infty]
- but it's more complexe and not so useful. *)
+ Theorem two_power_pos_is_exp x y :
+ two_power_pos (x + y) = two_power_pos x * two_power_pos y.
+ Proof.
+ rewrite 3 two_power_pos_equiv. now apply (Z.pow_add_r 2 (Zpos x) (Zpos y)).
+ Qed.
- Definition two_p (x:Z) :=
- match x with
- | Z0 => 1
- | Zpos y => two_power_pos y
- | Zneg y => 0
- end.
+ Lemma two_p_correct x : two_p x = 2^x.
+ Proof (two_p_equiv x).
- Theorem two_p_is_exp :
- forall x y:Z, 0 <= x -> 0 <= y -> two_p (x + y) = two_p x * two_p y.
+ Theorem two_p_is_exp x y :
+ 0 <= x -> 0 <= y -> two_p (x + y) = two_p x * two_p y.
Proof.
- simple induction x;
- [ simple induction y; simpl in |- *; auto with zarith
- | simple induction y;
- [ unfold two_p in |- *; rewrite (Zmult_comm (two_power_pos p) 1);
- rewrite (Zmult_1_l (two_power_pos p)); auto with zarith
- | unfold Zplus in |- *; unfold two_p in |- *; intros;
- apply two_power_pos_is_exp
- | intros; unfold Zle in H0; unfold Zcompare in H0;
- absurd (Datatypes.Gt = Datatypes.Gt); trivial with zarith ]
- | simple induction y;
- [ simpl in |- *; auto with zarith
- | intros; unfold Zle in H; unfold Zcompare in H;
- absurd (Datatypes.Gt = Datatypes.Gt); trivial with zarith
- | intros; unfold Zle in H; unfold Zcompare in H;
- absurd (Datatypes.Gt = Datatypes.Gt); trivial with zarith ] ].
+ rewrite !two_p_equiv. apply Z.pow_add_r.
Qed.
- Lemma two_p_gt_ZERO : forall x:Z, 0 <= x -> two_p x > 0.
+ Lemma two_p_gt_ZERO x : 0 <= x -> two_p x > 0.
Proof.
- simple induction x; intros;
- [ simpl in |- *; omega
- | simpl in |- *; unfold two_power_pos in |- *; apply Zorder.Zgt_pos_0
- | absurd (0 <= Zneg p);
- [ simpl in |- *; unfold Zle in |- *; unfold Zcompare in |- *;
- do 2 unfold not in |- *; auto with zarith
- | assumption ] ].
+ Z.swap_greater. rewrite two_p_equiv. now apply Z.pow_pos_nonneg.
Qed.
- Lemma two_p_S : forall x:Z, 0 <= x -> two_p (Zsucc x) = 2 * two_p x.
+ Lemma two_p_S x : 0 <= x -> two_p (Z.succ x) = 2 * two_p x.
Proof.
- intros; unfold Zsucc in |- *.
- rewrite (two_p_is_exp x 1 H (Zorder.Zle_0_pos 1)).
- apply Zmult_comm.
+ rewrite !two_p_equiv. now apply Z.pow_succ_r.
Qed.
- Lemma two_p_pred : forall x:Z, 0 <= x -> two_p (Zpred x) < two_p x.
+ Lemma two_p_pred x : 0 <= x -> two_p (Z.pred x) < two_p x.
Proof.
- intros; apply natlike_ind with (P := fun x:Z => two_p (Zpred x) < two_p x);
- [ simpl in |- *; unfold Zlt in |- *; auto with zarith
- | intros; elim (Zle_lt_or_eq 0 x0 H0);
- [ intros;
- replace (two_p (Zpred (Zsucc x0))) with (two_p (Zsucc (Zpred x0)));
- [ rewrite (two_p_S (Zpred x0));
- [ rewrite (two_p_S x0); [ omega | assumption ]
- | apply Zorder.Zlt_0_le_0_pred; assumption ]
- | rewrite <- (Zsucc_pred x0); rewrite <- (Zpred_succ x0);
- trivial with zarith ]
- | intro Hx0; rewrite <- Hx0; simpl in |- *; unfold Zlt in |- *;
- auto with zarith ]
- | assumption ].
+ rewrite !two_p_equiv. intros. apply Z.pow_lt_mono_r; auto with zarith.
Qed.
- Lemma Zlt_lt_double : forall x y:Z, 0 <= x < y -> x < 2 * y.
- intros; omega. Qed.
-
- End Powers_of_2.
+End Powers_of_2.
Hint Resolve two_p_gt_ZERO: zarith.
Hint Immediate two_p_pred two_p_S: zarith.
@@ -261,100 +230,88 @@ Section power_div_with_rest.
(** * Division by a power of two. *)
- (** To [n:Z] and [p:positive], [q],[r] are associated such that
- [n = 2^p.q + r] and [0 <= r < 2^p] *)
+ (** To [x:Z] and [p:positive], [q],[r] are associated such that
+ [x = 2^p.q + r] and [0 <= r < 2^p] *)
- (** Invariant: [d*q + r = d'*q + r /\ d' = 2*d /\ 0<= r < d /\ 0 <= r' < d'] *)
+ (** Invariant: [d*q + r = d'*q + r /\ d' = 2*d /\ 0<=r<d /\ 0<=r'<d'] *)
Definition Zdiv_rest_aux (qrd:Z * Z * Z) :=
- let (qr, d) := qrd in
- let (q, r) := qr in
- (match q with
- | Z0 => (0, r)
- | Zpos xH => (0, d + r)
- | Zpos (xI n) => (Zpos n, d + r)
- | Zpos (xO n) => (Zpos n, r)
- | Zneg xH => (-1, d + r)
- | Zneg (xI n) => (Zneg n - 1, d + r)
- | Zneg (xO n) => (Zneg n, r)
- end, 2 * d).
+ let '(q,r,d) := qrd in
+ (match q with
+ | Z0 => (0, r)
+ | Zpos xH => (0, d + r)
+ | Zpos (xI n) => (Zpos n, d + r)
+ | Zpos (xO n) => (Zpos n, r)
+ | Zneg xH => (-1, d + r)
+ | Zneg (xI n) => (Zneg n - 1, d + r)
+ | Zneg (xO n) => (Zneg n, r)
+ end, 2 * d).
Definition Zdiv_rest (x:Z) (p:positive) :=
- let (qr, d) := iter_pos p _ Zdiv_rest_aux (x, 0, 1) in qr.
+ let (qr, d) := Pos.iter p Zdiv_rest_aux (x, 0, 1) in qr.
- Lemma Zdiv_rest_correct1 :
- forall (x:Z) (p:positive),
- let (qr, d) := iter_pos p _ Zdiv_rest_aux (x, 0, 1) in d = two_power_pos p.
+ Lemma Zdiv_rest_correct1 (x:Z) (p:positive) :
+ let (_, d) := Pos.iter p Zdiv_rest_aux (x, 0, 1) in
+ d = two_power_pos p.
Proof.
- intros x p; rewrite (iter_nat_of_P p _ Zdiv_rest_aux (x, 0, 1));
- rewrite (two_power_pos_nat p); elim (nat_of_P p);
- simpl in |- *;
- [ trivial with zarith
- | intro n; rewrite (two_power_nat_S n); unfold Zdiv_rest_aux at 2 in |- *;
- elim (iter_nat n (Z * Z * Z) Zdiv_rest_aux (x, 0, 1));
- destruct a; intros; apply f_equal with (f := fun z:Z => 2 * z);
- assumption ].
+ rewrite Pos2Nat.inj_iter, two_power_pos_nat.
+ induction (Pos.to_nat p); simpl; trivial.
+ destruct (nat_iter n Zdiv_rest_aux (x,0,1)) as ((q,r),d).
+ unfold Zdiv_rest_aux. rewrite two_power_nat_S; now f_equal.
Qed.
- Lemma Zdiv_rest_correct2 :
- forall (x:Z) (p:positive),
- let (qr, d) := iter_pos p _ Zdiv_rest_aux (x, 0, 1) in
- let (q, r) := qr in x = q * d + r /\ 0 <= r < d.
+ Lemma Zdiv_rest_correct2 (x:Z) (p:positive) :
+ let '(q,r,d) := Pos.iter p Zdiv_rest_aux (x, 0, 1) in
+ x = q * d + r /\ 0 <= r < d.
Proof.
- intros;
- apply iter_pos_invariant with
- (f := Zdiv_rest_aux)
- (Inv := fun qrd:Z * Z * Z =>
- let (qr, d) := qrd in
- let (q, r) := qr in x = q * d + r /\ 0 <= r < d);
- [ intro x0; elim x0; intro y0; elim y0; intros q r d;
- unfold Zdiv_rest_aux in |- *; elim q;
- [ omega
- | destruct p0;
- [ rewrite BinInt.Zpos_xI; intro; elim H; intros; split;
- [ rewrite H0; rewrite Zplus_assoc; rewrite Zmult_plus_distr_l;
- rewrite Zmult_1_l; rewrite Zmult_assoc;
- rewrite (Zmult_comm (Zpos p0) 2); apply refl_equal
- | omega ]
- | rewrite BinInt.Zpos_xO; intro; elim H; intros; split;
- [ rewrite H0; rewrite Zmult_assoc; rewrite (Zmult_comm (Zpos p0) 2);
- apply refl_equal
- | omega ]
- | omega ]
- | destruct p0;
- [ rewrite BinInt.Zneg_xI; unfold Zminus in |- *; intro; elim H; intros;
- split;
- [ rewrite H0; rewrite Zplus_assoc;
- apply f_equal with (f := fun z:Z => z + r);
- do 2 rewrite Zmult_plus_distr_l; rewrite Zmult_assoc;
- rewrite (Zmult_comm (Zneg p0) 2); rewrite <- Zplus_assoc;
- apply f_equal with (f := fun z:Z => 2 * Zneg p0 * d + z);
- omega
- | omega ]
- | rewrite BinInt.Zneg_xO; unfold Zminus in |- *; intro; elim H; intros;
- split;
- [ rewrite H0; rewrite Zmult_assoc; rewrite (Zmult_comm (Zneg p0) 2);
- apply refl_equal
- | omega ]
- | omega ] ]
- | omega ].
+ apply Pos.iter_invariant; [|omega].
+ intros ((q,r),d) (H,H'). unfold Zdiv_rest_aux.
+ destruct q as [ |[q|q| ]|[q|q| ]]; try omega.
+ - rewrite Z.pos_xI, Z.mul_add_distr_r in H.
+ rewrite Z.mul_shuffle3, Z.mul_assoc. omega.
+ - rewrite Z.pos_xO in H.
+ rewrite Z.mul_shuffle3, Z.mul_assoc. omega.
+ - rewrite Z.neg_xI, Z.mul_sub_distr_r in H.
+ rewrite Z.mul_sub_distr_r, Z.mul_shuffle3, Z.mul_assoc. omega.
+ - rewrite Z.neg_xO in H.
+ rewrite Z.mul_shuffle3, Z.mul_assoc. omega.
Qed.
+ (** Old-style rich specification by proof of existence *)
+
Inductive Zdiv_rest_proofs (x:Z) (p:positive) : Set :=
Zdiv_rest_proof :
forall q r:Z,
x = q * two_power_pos p + r ->
0 <= r -> r < two_power_pos p -> Zdiv_rest_proofs x p.
- Lemma Zdiv_rest_correct : forall (x:Z) (p:positive), Zdiv_rest_proofs x p.
+ Lemma Zdiv_rest_correct (x:Z) (p:positive) : Zdiv_rest_proofs x p.
Proof.
- intros x p.
generalize (Zdiv_rest_correct1 x p); generalize (Zdiv_rest_correct2 x p).
- elim (iter_pos p (Z * Z * Z) Zdiv_rest_aux (x, 0, 1)).
- simple induction a.
- intros.
- elim H; intros H1 H2; clear H.
- rewrite H0 in H1; rewrite H0 in H2; elim H2; intros;
- apply Zdiv_rest_proof with (q := a0) (r := b); assumption.
+ destruct (Pos.iter p Zdiv_rest_aux (x, 0, 1)) as ((q,r),d).
+ intros (H1,(H2,H3)) ->. now exists q r.
+ Qed.
+
+ (** Direct correctness of [Zdiv_rest] *)
+
+ Lemma Zdiv_rest_ok x p :
+ let (q,r) := Zdiv_rest x p in
+ x = q * 2^(Zpos p) + r /\ 0 <= r < 2^(Zpos p).
+ Proof.
+ unfold Zdiv_rest.
+ generalize (Zdiv_rest_correct1 x p); generalize (Zdiv_rest_correct2 x p).
+ destruct (Pos.iter p Zdiv_rest_aux (x, 0, 1)) as ((q,r),d).
+ intros H ->. now rewrite two_power_pos_equiv in H.
+ Qed.
+
+ (** Equivalence with [Z.shiftr] *)
+
+ Lemma Zdiv_rest_shiftr x p :
+ fst (Zdiv_rest x p) = Z.shiftr x (Zpos p).
+ Proof.
+ generalize (Zdiv_rest_ok x p). destruct (Zdiv_rest x p) as (q,r).
+ intros (H,H'). simpl.
+ rewrite Z.shiftr_div_pow2 by easy.
+ apply Z.div_unique_pos with r; trivial. now rewrite Z.mul_comm.
Qed.
End power_div_with_rest.
diff --git a/theories/ZArith/Zquot.v b/theories/ZArith/Zquot.v
new file mode 100644
index 00000000..9a95669f
--- /dev/null
+++ b/theories/ZArith/Zquot.v
@@ -0,0 +1,536 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Nnat ZArith_base ROmega ZArithRing Zdiv Morphisms.
+
+Local Open Scope Z_scope.
+
+(** This file provides results about the Round-Toward-Zero Euclidean
+ division [Zquotrem], whose projections are [Zquot] and [Zrem].
+ Definition of this division can be found in file [BinIntDef].
+
+ This division and the one defined in Zdiv agree only on positive
+ numbers. Otherwise, Zdiv performs Round-Toward-Bottom (a.k.a Floor).
+
+ The current approach is compatible with the division of usual
+ programming languages such as Ocaml. In addition, it has nicer
+ properties with respect to opposite and other usual operations.
+*)
+
+(** * Relation between division on N and on Z. *)
+
+Lemma Ndiv_Zquot : forall a b:N,
+ Z_of_N (a/b) = (Z_of_N a ÷ Z_of_N b).
+Proof.
+ intros.
+ destruct a; destruct b; simpl; auto.
+ unfold N.div, Z.quot; simpl. destruct N.pos_div_eucl; auto.
+Qed.
+
+Lemma Nmod_Zrem : forall a b:N,
+ Z.of_N (a mod b) = Z.rem (Z.of_N a) (Z.of_N b).
+Proof.
+ intros.
+ destruct a; destruct b; simpl; auto.
+ unfold N.modulo, Z.rem; simpl; destruct N.pos_div_eucl; auto.
+Qed.
+
+(** * Characterization of this euclidean division. *)
+
+(** First, the usual equation [a=q*b+r]. Notice that [a mod 0]
+ has been chosen to be [a], so this equation holds even for [b=0].
+*)
+
+Notation Z_quot_rem_eq := Z.quot_rem' (only parsing).
+
+(** Then, the inequalities constraining the remainder:
+ The remainder is bounded by the divisor, in term of absolute values *)
+
+Theorem Zrem_lt : forall a b:Z, b<>0 ->
+ Z.abs (Z.rem a b) < Z.abs b.
+Proof.
+ apply Z.rem_bound_abs.
+Qed.
+
+(** The sign of the remainder is the one of [a]. Due to the possible
+ nullity of [a], a general result is to be stated in the following form:
+*)
+
+Theorem Zrem_sgn a b : 0 <= Z.sgn (Z.rem a b) * Z.sgn a.
+Proof.
+ destruct b as [ |b|b]; destruct a as [ |a|a]; simpl; auto with zarith;
+ unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl;
+ simpl; destruct n0; simpl; auto with zarith.
+Qed.
+
+(** This can also be said in a simplier way: *)
+
+Theorem Zrem_sgn2 a b : 0 <= (Z.rem a b) * a.
+Proof.
+ rewrite <-Z.sgn_nonneg, Z.sgn_mul; apply Zrem_sgn.
+Qed.
+
+(** Reformulation of [Zquot_lt] and [Zrem_sgn] in 2
+ then 4 particular cases. *)
+
+Theorem Zrem_lt_pos a b : 0<=a -> b<>0 -> 0 <= Z.rem a b < Z.abs b.
+Proof.
+ intros.
+ assert (0 <= Z.rem a b).
+ generalize (Zrem_sgn a b).
+ destruct (Zle_lt_or_eq 0 a H).
+ rewrite <- Zsgn_pos in H1; rewrite H1. romega with *.
+ subst a; simpl; auto.
+ generalize (Zrem_lt a b H0); romega with *.
+Qed.
+
+Theorem Zrem_lt_neg a b : a<=0 -> b<>0 -> -Z.abs b < Z.rem a b <= 0.
+Proof.
+ intros.
+ assert (Z.rem a b <= 0).
+ generalize (Zrem_sgn a b).
+ destruct (Zle_lt_or_eq a 0 H).
+ rewrite <- Zsgn_neg in H1; rewrite H1; romega with *.
+ subst a; simpl; auto.
+ generalize (Zrem_lt a b H0); romega with *.
+Qed.
+
+Theorem Zrem_lt_pos_pos a b : 0<=a -> 0<b -> 0 <= Z.rem a b < b.
+Proof.
+ intros; generalize (Zrem_lt_pos a b); romega with *.
+Qed.
+
+Theorem Zrem_lt_pos_neg a b : 0<=a -> b<0 -> 0 <= Z.rem a b < -b.
+Proof.
+ intros; generalize (Zrem_lt_pos a b); romega with *.
+Qed.
+
+Theorem Zrem_lt_neg_pos a b : a<=0 -> 0<b -> -b < Z.rem a b <= 0.
+Proof.
+ intros; generalize (Zrem_lt_neg a b); romega with *.
+Qed.
+
+Theorem Zrem_lt_neg_neg a b : a<=0 -> b<0 -> b < Z.rem a b <= 0.
+Proof.
+ intros; generalize (Zrem_lt_neg a b); romega with *.
+Qed.
+
+(** * Division and Opposite *)
+
+(* The precise equalities that are invalid with "historic" Zdiv. *)
+
+Theorem Zquot_opp_l a b : (-a)÷b = -(a÷b).
+Proof.
+ destruct a; destruct b; simpl; auto;
+ unfold Z.quot, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
+Qed.
+
+Theorem Zquot_opp_r a b : a÷(-b) = -(a÷b).
+Proof.
+ destruct a; destruct b; simpl; auto;
+ unfold Z.quot, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
+Qed.
+
+Theorem Zrem_opp_l a b : Z.rem (-a) b = -(Z.rem a b).
+Proof.
+ destruct a; destruct b; simpl; auto;
+ unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
+Qed.
+
+Theorem Zrem_opp_r a b : Z.rem a (-b) = Z.rem a b.
+Proof.
+ destruct a; destruct b; simpl; auto;
+ unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
+Qed.
+
+Theorem Zquot_opp_opp a b : (-a)÷(-b) = a÷b.
+Proof.
+ destruct a; destruct b; simpl; auto;
+ unfold Z.quot, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
+Qed.
+
+Theorem Zrem_opp_opp a b : Z.rem (-a) (-b) = -(Z.rem a b).
+Proof.
+ destruct a; destruct b; simpl; auto;
+ unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
+Qed.
+
+(** * Unicity results *)
+
+Definition Remainder a b r :=
+ (0 <= a /\ 0 <= r < Z.abs b) \/ (a <= 0 /\ -Z.abs b < r <= 0).
+
+Definition Remainder_alt a b r :=
+ Z.abs r < Z.abs b /\ 0 <= r * a.
+
+Lemma Remainder_equiv : forall a b r,
+ Remainder a b r <-> Remainder_alt a b r.
+Proof.
+ unfold Remainder, Remainder_alt; intuition.
+ romega with *.
+ romega with *.
+ rewrite <-(Zmult_opp_opp).
+ apply Zmult_le_0_compat; romega.
+ assert (0 <= Z.sgn r * Z.sgn a) by (rewrite <-Z.sgn_mul, Z.sgn_nonneg; auto).
+ destruct r; simpl Z.sgn in *; romega with *.
+Qed.
+
+Theorem Zquot_mod_unique_full:
+ forall a b q r, Remainder a b r ->
+ a = b*q + r -> q = a÷b /\ r = Z.rem a b.
+Proof.
+ destruct 1 as [(H,H0)|(H,H0)]; intros.
+ apply Zdiv_mod_unique with b; auto.
+ apply Zrem_lt_pos; auto.
+ romega with *.
+ rewrite <- H1; apply Z_quot_rem_eq.
+
+ rewrite <- (Zopp_involutive a).
+ rewrite Zquot_opp_l, Zrem_opp_l.
+ generalize (Zdiv_mod_unique b (-q) (-a÷b) (-r) (Z.rem (-a) b)).
+ generalize (Zrem_lt_pos (-a) b).
+ rewrite <-Z_quot_rem_eq, <-Zopp_mult_distr_r, <-Zopp_plus_distr, <-H1.
+ romega with *.
+Qed.
+
+Theorem Zquot_unique_full:
+ forall a b q r, Remainder a b r ->
+ a = b*q + r -> q = a÷b.
+Proof.
+ intros; destruct (Zquot_mod_unique_full a b q r); auto.
+Qed.
+
+Theorem Zquot_unique:
+ forall a b q r, 0 <= a -> 0 <= r < b ->
+ a = b*q + r -> q = a÷b.
+Proof. exact Z.quot_unique. Qed.
+
+Theorem Zrem_unique_full:
+ forall a b q r, Remainder a b r ->
+ a = b*q + r -> r = Z.rem a b.
+Proof.
+ intros; destruct (Zquot_mod_unique_full a b q r); auto.
+Qed.
+
+Theorem Zrem_unique:
+ forall a b q r, 0 <= a -> 0 <= r < b ->
+ a = b*q + r -> r = Z.rem a b.
+Proof. exact Z.rem_unique. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma Zrem_0_l: forall a, Z.rem 0 a = 0.
+Proof.
+ destruct a; simpl; auto.
+Qed.
+
+Lemma Zrem_0_r: forall a, Z.rem a 0 = a.
+Proof.
+ destruct a; simpl; auto.
+Qed.
+
+Lemma Zquot_0_l: forall a, 0÷a = 0.
+Proof.
+ destruct a; simpl; auto.
+Qed.
+
+Lemma Zquot_0_r: forall a, a÷0 = 0.
+Proof.
+ destruct a; simpl; auto.
+Qed.
+
+Lemma Zrem_1_r: forall a, Z.rem a 1 = 0.
+Proof. exact Z.rem_1_r. Qed.
+
+Lemma Zquot_1_r: forall a, a÷1 = a.
+Proof. exact Z.quot_1_r. Qed.
+
+Hint Resolve Zrem_0_l Zrem_0_r Zquot_0_l Zquot_0_r Zquot_1_r Zrem_1_r
+ : zarith.
+
+Lemma Zquot_1_l: forall a, 1 < a -> 1÷a = 0.
+Proof. exact Z.quot_1_l. Qed.
+
+Lemma Zrem_1_l: forall a, 1 < a -> Z.rem 1 a = 1.
+Proof. exact Z.rem_1_l. Qed.
+
+Lemma Z_quot_same : forall a:Z, a<>0 -> a÷a = 1.
+Proof. exact Z.quot_same. Qed.
+
+Ltac zero_or_not a :=
+ destruct (Z.eq_dec a 0);
+ [subst; rewrite ?Zrem_0_l, ?Zquot_0_l, ?Zrem_0_r, ?Zquot_0_r;
+ auto with zarith|].
+
+Lemma Z_rem_same : forall a, Z.rem a a = 0.
+Proof. intros. zero_or_not a. apply Z.rem_same; auto. Qed.
+
+Lemma Z_rem_mult : forall a b, Z.rem (a*b) b = 0.
+Proof. intros. zero_or_not b. apply Z.rem_mul; auto. Qed.
+
+Lemma Z_quot_mult : forall a b:Z, b <> 0 -> (a*b)÷b = a.
+Proof. exact Z.quot_mul. Qed.
+
+(** * Order results about Zrem and Zquot *)
+
+(* Division of positive numbers is positive. *)
+
+Lemma Z_quot_pos: forall a b, 0 <= a -> 0 <= b -> 0 <= a÷b.
+Proof. intros. zero_or_not b. apply Z.quot_pos; auto with zarith. Qed.
+
+(** As soon as the divisor is greater or equal than 2,
+ the division is strictly decreasing. *)
+
+Lemma Z_quot_lt : forall a b:Z, 0 < a -> 2 <= b -> a÷b < a.
+Proof. intros. apply Z.quot_lt; auto with zarith. Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem Zquot_small: forall a b, 0 <= a < b -> a÷b = 0.
+Proof. exact Z.quot_small. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem Zrem_small: forall a n, 0 <= a < n -> Z.rem a n = a.
+Proof. exact Z.rem_small. Qed.
+
+(** [Zge] is compatible with a positive division. *)
+
+Lemma Z_quot_monotone : forall a b c, 0<=c -> a<=b -> a÷c <= b÷c.
+Proof. intros. zero_or_not c. apply Z.quot_le_mono; auto with zarith. Qed.
+
+(** With our choice of division, rounding of (a÷b) is always done toward zero: *)
+
+Lemma Z_mult_quot_le : forall a b:Z, 0 <= a -> 0 <= b*(a÷b) <= a.
+Proof. intros. zero_or_not b. apply Z.mul_quot_le; auto with zarith. Qed.
+
+Lemma Z_mult_quot_ge : forall a b:Z, a <= 0 -> a <= b*(a÷b) <= 0.
+Proof. intros. zero_or_not b. apply Z.mul_quot_ge; auto with zarith. Qed.
+
+(** The previous inequalities between [b*(a÷b)] and [a] are exact
+ iff the modulo is zero. *)
+
+Lemma Z_quot_exact_full : forall a b:Z, a = b*(a÷b) <-> Z.rem a b = 0.
+Proof. intros. zero_or_not b. intuition. apply Z.quot_exact; auto. Qed.
+
+(** A modulo cannot grow beyond its starting point. *)
+
+Theorem Zrem_le: forall a b, 0 <= a -> 0 <= b -> Z.rem a b <= a.
+Proof. intros. zero_or_not b. apply Z.rem_le; auto with zarith. Qed.
+
+(** Some additionnal inequalities about Zdiv. *)
+
+Theorem Zquot_le_upper_bound:
+ forall a b q, 0 < b -> a <= q*b -> a÷b <= q.
+Proof. intros a b q; rewrite Zmult_comm; apply Z.quot_le_upper_bound. Qed.
+
+Theorem Zquot_lt_upper_bound:
+ forall a b q, 0 <= a -> 0 < b -> a < q*b -> a÷b < q.
+Proof. intros a b q; rewrite Zmult_comm; apply Z.quot_lt_upper_bound. Qed.
+
+Theorem Zquot_le_lower_bound:
+ forall a b q, 0 < b -> q*b <= a -> q <= a÷b.
+Proof. intros a b q; rewrite Zmult_comm; apply Z.quot_le_lower_bound. Qed.
+
+Theorem Zquot_sgn: forall a b,
+ 0 <= Z.sgn (a÷b) * Z.sgn a * Z.sgn b.
+Proof.
+ destruct a as [ |a|a]; destruct b as [ |b|b]; simpl; auto with zarith;
+ unfold Z.quot; simpl; destruct N.pos_div_eucl; simpl; destruct n; simpl; auto with zarith.
+Qed.
+
+(** * Relations between usual operations and Zmod and Zdiv *)
+
+(** First, a result that used to be always valid with Zdiv,
+ but must be restricted here.
+ For instance, now (9+(-5)*2) rem 2 = -1 <> 1 = 9 rem 2 *)
+
+Lemma Z_rem_plus : forall a b c:Z,
+ 0 <= (a+b*c) * a ->
+ Z.rem (a + b * c) c = Z.rem a c.
+Proof. intros. zero_or_not c. apply Z.rem_add; auto with zarith. Qed.
+
+Lemma Z_quot_plus : forall a b c:Z,
+ 0 <= (a+b*c) * a -> c<>0 ->
+ (a + b * c) ÷ c = a ÷ c + b.
+Proof. intros. apply Z.quot_add; auto with zarith. Qed.
+
+Theorem Z_quot_plus_l: forall a b c : Z,
+ 0 <= (a*b+c)*c -> b<>0 ->
+ b<>0 -> (a * b + c) ÷ b = a + c ÷ b.
+Proof. intros. apply Z.quot_add_l; auto with zarith. Qed.
+
+(** Cancellations. *)
+
+Lemma Zquot_mult_cancel_r : forall a b c:Z,
+ c<>0 -> (a*c)÷(b*c) = a÷b.
+Proof. intros. zero_or_not b. apply Z.quot_mul_cancel_r; auto. Qed.
+
+Lemma Zquot_mult_cancel_l : forall a b c:Z,
+ c<>0 -> (c*a)÷(c*b) = a÷b.
+Proof.
+ intros. rewrite (Zmult_comm c b). zero_or_not b.
+ rewrite (Zmult_comm b c). apply Z.quot_mul_cancel_l; auto.
+Qed.
+
+Lemma Zmult_rem_distr_l: forall a b c,
+ Z.rem (c*a) (c*b) = c * (Z.rem a b).
+Proof.
+ intros. zero_or_not c. rewrite (Zmult_comm c b). zero_or_not b.
+ rewrite (Zmult_comm b c). apply Z.mul_rem_distr_l; auto.
+Qed.
+
+Lemma Zmult_rem_distr_r: forall a b c,
+ Z.rem (a*c) (b*c) = (Z.rem a b) * c.
+Proof.
+ intros. zero_or_not b. rewrite (Zmult_comm b c). zero_or_not c.
+ rewrite (Zmult_comm c b). apply Z.mul_rem_distr_r; auto.
+Qed.
+
+(** Operations modulo. *)
+
+Theorem Zrem_rem: forall a n, Z.rem (Z.rem a n) n = Z.rem a n.
+Proof. intros. zero_or_not n. apply Z.rem_rem; auto. Qed.
+
+Theorem Zmult_rem: forall a b n,
+ Z.rem (a * b) n = Z.rem (Z.rem a n * Z.rem b n) n.
+Proof. intros. zero_or_not n. apply Z.mul_rem; auto. Qed.
+
+(** addition and modulo
+
+ Generally speaking, unlike with Zdiv, we don't have
+ (a+b) rem n = (a rem n + b rem n) rem n
+ for any a and b.
+ For instance, take (8 + (-10)) rem 3 = -2 whereas
+ (8 rem 3 + (-10 rem 3)) rem 3 = 1. *)
+
+Theorem Zplus_rem: forall a b n,
+ 0 <= a * b ->
+ Z.rem (a + b) n = Z.rem (Z.rem a n + Z.rem b n) n.
+Proof. intros. zero_or_not n. apply Z.add_rem; auto. Qed.
+
+Lemma Zplus_rem_idemp_l: forall a b n,
+ 0 <= a * b ->
+ Z.rem (Z.rem a n + b) n = Z.rem (a + b) n.
+Proof. intros. zero_or_not n. apply Z.add_rem_idemp_l; auto. Qed.
+
+Lemma Zplus_rem_idemp_r: forall a b n,
+ 0 <= a*b ->
+ Z.rem (b + Z.rem a n) n = Z.rem (b + a) n.
+Proof.
+ intros. zero_or_not n. apply Z.add_rem_idemp_r; auto.
+ rewrite Zmult_comm; auto.
+Qed.
+
+Lemma Zmult_rem_idemp_l: forall a b n, Z.rem (Z.rem a n * b) n = Z.rem (a * b) n.
+Proof. intros. zero_or_not n. apply Z.mul_rem_idemp_l; auto. Qed.
+
+Lemma Zmult_rem_idemp_r: forall a b n, Z.rem (b * Z.rem a n) n = Z.rem (b * a) n.
+Proof. intros. zero_or_not n. apply Z.mul_rem_idemp_r; auto. Qed.
+
+(** Unlike with Zdiv, the following result is true without restrictions. *)
+
+Lemma Zquot_Zquot : forall a b c, (a÷b)÷c = a÷(b*c).
+Proof.
+ intros. zero_or_not b. rewrite Zmult_comm. zero_or_not c.
+ rewrite Zmult_comm. apply Z.quot_quot; auto.
+Qed.
+
+(** A last inequality: *)
+
+Theorem Zquot_mult_le:
+ forall a b c, 0<=a -> 0<=b -> 0<=c -> c*(a÷b) <= (c*a)÷b.
+Proof. intros. zero_or_not b. apply Z.quot_mul_le; auto with zarith. Qed.
+
+(** Z.rem is related to divisibility (see more in Znumtheory) *)
+
+Lemma Zrem_divides : forall a b,
+ Z.rem a b = 0 <-> exists c, a = b*c.
+Proof.
+ intros. zero_or_not b. firstorder.
+ rewrite Z.rem_divide; trivial.
+ split; intros (c,Hc); exists c; subst; auto with zarith.
+Qed.
+
+(** Particular case : dividing by 2 is related with parity *)
+
+Lemma Zquot2_odd_remainder : forall a,
+ Remainder a 2 (if Z.odd a then Z.sgn a else 0).
+Proof.
+ intros [ |p|p]. simpl.
+ left. simpl. auto with zarith.
+ left. destruct p; simpl; auto with zarith.
+ right. destruct p; simpl; split; now auto with zarith.
+Qed.
+
+Notation Zquot2_quot := Zquot2_quot (only parsing).
+
+Lemma Zrem_odd : forall a, Z.rem a 2 = if Z.odd a then Z.sgn a else 0.
+Proof.
+ intros. symmetry.
+ apply Zrem_unique_full with (Zquot2 a).
+ apply Zquot2_odd_remainder.
+ apply Zquot2_odd_eqn.
+Qed.
+
+Lemma Zrem_even : forall a, Z.rem a 2 = if Z.even a then 0 else Z.sgn a.
+Proof.
+ intros a. rewrite Zrem_odd, Zodd_even_bool. now destruct Zeven_bool.
+Qed.
+
+Lemma Zeven_rem : forall a, Z.even a = Zeq_bool (Z.rem a 2) 0.
+Proof.
+ intros a. rewrite Zrem_even.
+ destruct a as [ |p|p]; trivial; now destruct p.
+Qed.
+
+Lemma Zodd_rem : forall a, Z.odd a = negb (Zeq_bool (Z.rem a 2) 0).
+Proof.
+ intros a. rewrite Zrem_odd.
+ destruct a as [ |p|p]; trivial; now destruct p.
+Qed.
+
+(** * Interaction with "historic" Zdiv *)
+
+(** They agree at least on positive numbers: *)
+
+Theorem Zquotrem_Zdiv_eucl_pos : forall a b:Z, 0 <= a -> 0 < b ->
+ a÷b = a/b /\ Z.rem a b = a mod b.
+Proof.
+ intros.
+ apply Zdiv_mod_unique with b.
+ apply Zrem_lt_pos; auto with zarith.
+ rewrite Zabs_eq; auto with *; apply Z_mod_lt; auto with *.
+ rewrite <- Z_div_mod_eq; auto with *.
+ symmetry; apply Z_quot_rem_eq; auto with *.
+Qed.
+
+Theorem Zquot_Zdiv_pos : forall a b, 0 <= a -> 0 <= b ->
+ a÷b = a/b.
+Proof.
+ intros a b Ha Hb.
+ destruct (Zle_lt_or_eq _ _ Hb).
+ generalize (Zquotrem_Zdiv_eucl_pos a b Ha H); intuition.
+ subst; rewrite Zquot_0_r, Zdiv_0_r; reflexivity.
+Qed.
+
+Theorem Zrem_Zmod_pos : forall a b, 0 <= a -> 0 < b ->
+ Z.rem a b = a mod b.
+Proof.
+ intros a b Ha Hb; generalize (Zquotrem_Zdiv_eucl_pos a b Ha Hb);
+ intuition.
+Qed.
+
+(** Modulos are null at the same places *)
+
+Theorem Zrem_Zmod_zero : forall a b, b<>0 ->
+ (Z.rem a b = 0 <-> a mod b = 0).
+Proof.
+ intros.
+ rewrite Zrem_divides, Zmod_divides; intuition.
+Qed.
diff --git a/theories/ZArith/Zsqrt.v b/theories/ZArith/Zsqrt_compat.v
index 1a67bbb2..4584c3f8 100644
--- a/theories/ZArith/Zsqrt.v
+++ b/theories/ZArith/Zsqrt_compat.v
@@ -1,18 +1,27 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: Zsqrt.v 14641 2011-11-06 11:59:10Z herbelin $ *)
-
Require Import ZArithRing.
Require Import Omega.
Require Export ZArith_base.
Open Local Scope Z_scope.
+(** THIS FILE IS DEPRECATED
+
+ Instead of the various [Zsqrt] defined here, please use rather
+ [Z.sqrt] (or [Z.sqrtrem]). The latter are pure functions without
+ proof parts, and more results are available about them.
+ Some equivalence proofs between the old and the new versions
+ can be found below. Importing ZArith will provides by default
+ the new versions.
+
+*)
+
(**********************************************************************)
(** Definition and properties of square root on Z *)
@@ -213,3 +222,12 @@ Proof.
Qed.
+(** Equivalence between Zsqrt_plain and [Z.sqrt] *)
+
+Lemma Zsqrt_equiv : forall n, Zsqrt_plain n = Z.sqrt n.
+Proof.
+ intros. destruct (Z_le_gt_dec 0 n).
+ symmetry. apply Z.sqrt_unique; trivial.
+ now apply Zsqrt_interval.
+ now destruct n.
+Qed. \ No newline at end of file
diff --git a/theories/ZArith/Zwf.v b/theories/ZArith/Zwf.v
index 53f167e8..30802f82 100644
--- a/theories/ZArith/Zwf.v
+++ b/theories/ZArith/Zwf.v
@@ -1,13 +1,11 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: Zwf.v 14641 2011-11-06 11:59:10Z herbelin $ *)
-
Require Import ZArith_base.
Require Export Wf_nat.
Require Import Omega.
diff --git a/theories/ZArith/auxiliary.v b/theories/ZArith/auxiliary.v
index ade35bef..742f4bde 100644
--- a/theories/ZArith/auxiliary.v
+++ b/theories/ZArith/auxiliary.v
@@ -1,14 +1,12 @@
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: auxiliary.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(** Binary Integers (Pierre Crégut, CNET, Lannion, France) *)
Require Export Arith_base.
@@ -18,96 +16,79 @@ Require Import Decidable.
Require Import Peano_dec.
Require Export Compare_dec.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
(***************************************************************)
(** * Moving terms from one side to the other of an inequality *)
-Theorem Zne_left : forall n m:Z, Zne n m -> Zne (n + - m) 0.
+Theorem Zne_left n m : Zne n m -> Zne (n + - m) 0.
Proof.
- intros x y; unfold Zne in |- *; unfold not in |- *; intros H1 H2; apply H1;
- apply Zplus_reg_l with (- y); rewrite Zplus_opp_l;
- rewrite Zplus_comm; trivial with arith.
+ unfold Zne. now rewrite <- Z.sub_move_0_r.
Qed.
-Theorem Zegal_left : forall n m:Z, n = m -> n + - m = 0.
+Theorem Zegal_left n m : n = m -> n + - m = 0.
Proof.
- intros x y H; apply (Zplus_reg_l y); rewrite Zplus_permute;
- rewrite Zplus_opp_r; do 2 rewrite Zplus_0_r; assumption.
+ apply Z.sub_move_0_r.
Qed.
-Theorem Zle_left : forall n m:Z, n <= m -> 0 <= m + - n.
+Theorem Zle_left n m : n <= m -> 0 <= m + - n.
Proof.
- intros x y H; replace 0 with (x + - x).
- apply Zplus_le_compat_r; trivial.
- apply Zplus_opp_r.
+ apply Z.le_0_sub.
Qed.
-Theorem Zle_left_rev : forall n m:Z, 0 <= m + - n -> n <= m.
+Theorem Zle_left_rev n m : 0 <= m + - n -> n <= m.
Proof.
- intros x y H; apply Zplus_le_reg_r with (- x).
- rewrite Zplus_opp_r; trivial.
+ apply Z.le_0_sub.
Qed.
-Theorem Zlt_left_rev : forall n m:Z, 0 < m + - n -> n < m.
+Theorem Zlt_left_rev n m : 0 < m + - n -> n < m.
Proof.
- intros x y H; apply Zplus_lt_reg_r with (- x).
- rewrite Zplus_opp_r; trivial.
+ apply Z.lt_0_sub.
Qed.
-Theorem Zlt_left : forall n m:Z, n < m -> 0 <= m + -1 + - n.
+Theorem Zlt_left_lt n m : n < m -> 0 < m + - n.
Proof.
- intros x y H; apply Zle_left; apply Zsucc_le_reg;
- change (Zsucc x <= Zsucc (Zpred y)) in |- *; rewrite <- Zsucc_pred;
- apply Zlt_le_succ; assumption.
+ apply Z.lt_0_sub.
Qed.
-Theorem Zlt_left_lt : forall n m:Z, n < m -> 0 < m + - n.
+Theorem Zlt_left n m : n < m -> 0 <= m + -1 + - n.
Proof.
- intros x y H; replace 0 with (x + - x).
- apply Zplus_lt_compat_r; trivial.
- apply Zplus_opp_r.
+ intros. rewrite Z.add_shuffle0. change (-1) with (- Z.succ 0).
+ now apply Z.le_0_sub, Z.le_succ_l, Z.lt_0_sub.
Qed.
-Theorem Zge_left : forall n m:Z, n >= m -> 0 <= n + - m.
+Theorem Zge_left n m : n >= m -> 0 <= n + - m.
Proof.
- intros x y H; apply Zle_left; apply Zge_le; assumption.
+ Z.swap_greater. apply Z.le_0_sub.
Qed.
-Theorem Zgt_left : forall n m:Z, n > m -> 0 <= n + -1 + - m.
+Theorem Zgt_left n m : n > m -> 0 <= n + -1 + - m.
Proof.
- intros x y H; apply Zlt_left; apply Zgt_lt; assumption.
+ Z.swap_greater. apply Zlt_left.
Qed.
-Theorem Zgt_left_gt : forall n m:Z, n > m -> n + - m > 0.
+Theorem Zgt_left_gt n m : n > m -> n + - m > 0.
Proof.
- intros x y H; replace 0 with (y + - y).
- apply Zplus_gt_compat_r; trivial.
- apply Zplus_opp_r.
+ Z.swap_greater. apply Z.lt_0_sub.
Qed.
-Theorem Zgt_left_rev : forall n m:Z, n + - m > 0 -> n > m.
+Theorem Zgt_left_rev n m : n + - m > 0 -> n > m.
Proof.
- intros x y H; apply Zplus_gt_reg_r with (- y).
- rewrite Zplus_opp_r; trivial.
+ Z.swap_greater. apply Z.lt_0_sub.
Qed.
-Theorem Zle_mult_approx :
- forall n m p:Z, n > 0 -> p > 0 -> 0 <= m -> 0 <= m * n + p.
+Theorem Zle_mult_approx n m p :
+ n > 0 -> p > 0 -> 0 <= m -> 0 <= m * n + p.
Proof.
- intros x y z H1 H2 H3; apply Zle_trans with (m := y * x);
- [ apply Zmult_gt_0_le_0_compat; assumption
- | pattern (y * x) at 1 in |- *; rewrite <- Zplus_0_r;
- apply Zplus_le_compat_l; apply Zlt_le_weak; apply Zgt_lt;
- assumption ].
+ Z.swap_greater. intros. Z.order_pos.
Qed.
-Theorem Zmult_le_approx :
- forall n m p:Z, n > 0 -> n > p -> 0 <= m * n + p -> 0 <= m.
+Theorem Zmult_le_approx n m p :
+ n > 0 -> n > p -> 0 <= m * n + p -> 0 <= m.
Proof.
- intros x y z H1 H2 H3; apply Zlt_succ_le; apply Zmult_gt_0_lt_0_reg_r with x;
- [ assumption
- | apply Zle_lt_trans with (1 := H3); rewrite <- Zmult_succ_l_reverse;
- apply Zplus_lt_compat_l; apply Zgt_lt; assumption ].
+ Z.swap_greater. intros. apply Z.lt_succ_r.
+ apply Z.mul_pos_cancel_r with n; trivial. Z.nzsimpl.
+ apply Z.le_lt_trans with (m*n+p); trivial.
+ now apply Z.add_lt_mono_l.
Qed.
diff --git a/theories/ZArith/vo.itarget b/theories/ZArith/vo.itarget
index 3efa7055..178111cd 100644
--- a/theories/ZArith/vo.itarget
+++ b/theories/ZArith/vo.itarget
@@ -1,4 +1,5 @@
auxiliary.vo
+BinIntDef.vo
BinInt.vo
Int.vo
Wf_Z.vo
@@ -13,6 +14,7 @@ Zcomplements.vo
Zdiv.vo
Zeven.vo
Zgcd_alt.vo
+Zpow_alt.vo
Zhints.vo
Zlogarithm.vo
Zmax.vo
@@ -21,12 +23,11 @@ Zmin.vo
Zmisc.vo
Znat.vo
Znumtheory.vo
-ZOdiv_def.vo
-ZOdiv.vo
+Zquot.vo
Zorder.vo
Zpow_def.vo
Zpower.vo
Zpow_facts.vo
-Zsqrt.vo
+Zsqrt_compat.vo
Zwf.vo
-ZOrderedType.vo
+Zeuclid.vo