aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--Makefile.common10
-rw-r--r--theories/Init/Tactics.v38
-rw-r--r--theories/Ints/Basic_type.v3
-rw-r--r--theories/Ints/List/Iterator.v180
-rw-r--r--theories/Ints/List/LPermutation.v509
-rw-r--r--theories/Ints/List/ListAux.v272
-rw-r--r--theories/Ints/List/UList.v286
-rw-r--r--theories/Ints/List/ZProgression.v105
-rw-r--r--theories/Ints/Tactic.v76
-rw-r--r--theories/Ints/Z/IntsZmisc.v185
-rw-r--r--theories/Ints/Z/Pmod.v565
-rw-r--r--theories/Ints/Z/Ppow.v39
-rw-r--r--theories/Ints/Z/ZAux.v1367
-rw-r--r--theories/Ints/Z/ZDivModAux.v479
-rw-r--r--theories/Ints/Z/ZPowerAux.v203
-rw-r--r--theories/Ints/Z/ZSum.v321
-rw-r--r--theories/Ints/num/GenAdd.v9
-rw-r--r--theories/Ints/num/GenBase.v22
-rw-r--r--theories/Ints/num/GenDiv.v58
-rw-r--r--theories/Ints/num/GenDivn1.v31
-rw-r--r--theories/Ints/num/GenLift.v16
-rw-r--r--theories/Ints/num/GenMul.v49
-rw-r--r--theories/Ints/num/GenSqrt.v119
-rw-r--r--theories/Ints/num/GenSub.v11
-rw-r--r--theories/Ints/num/Nbasic.v7
-rw-r--r--theories/Ints/num/Q0Make.v1
-rw-r--r--theories/Ints/num/QbiMake.v1
-rw-r--r--theories/Ints/num/QifMake.v1
-rw-r--r--theories/Ints/num/QpMake.v1
-rw-r--r--theories/Ints/num/QvMake.v1
-rw-r--r--theories/Ints/num/ZMake.v1
-rw-r--r--theories/Ints/num/Zn2Z.v7
-rw-r--r--theories/Ints/num/ZnZ.v1
-rw-r--r--theories/Numbers/Integer/TreeMod/ZTreeMod.v25
-rw-r--r--theories/QArith/Qpower.v2
-rw-r--r--theories/QArith/Qround.v2
-rw-r--r--theories/ZArith/Zabs.v112
-rw-r--r--theories/ZArith/Zdiv.v1107
-rw-r--r--theories/ZArith/Zeven.v128
-rw-r--r--theories/ZArith/Zmax.v14
-rw-r--r--theories/ZArith/Znat.v5
-rw-r--r--theories/ZArith/Znumtheory.v325
-rw-r--r--theories/ZArith/Zorder.v28
-rw-r--r--theories/ZArith/Zpow_facts.v451
-rw-r--r--theories/ZArith/Zpower.v66
-rw-r--r--theories/ZArith/Zsqrt.v51
46 files changed, 2307 insertions, 4983 deletions
diff --git a/Makefile.common b/Makefile.common
index bf466da7a..47102d69f 100644
--- a/Makefile.common
+++ b/Makefile.common
@@ -509,12 +509,10 @@ ZARITHVO:=\
theories/ZArith/Zwf.vo theories/ZArith/ZArith_base.vo \
theories/ZArith/Zbool.vo theories/ZArith/Zbinary.vo \
theories/ZArith/Znumtheory.vo theories/ZArith/Int.vo \
- theories/ZArith/Zpow_def.vo
+ theories/ZArith/Zpow_def.vo theories/ZArith/Zpow_facts.vo
INTSVO:=\
- theories/Ints/Z/IntsZmisc.vo theories/Ints/Z/Pmod.vo \
- theories/Ints/Tactic.vo theories/Ints/Z/ZAux.vo \
- theories/Ints/Z/ZPowerAux.vo theories/Ints/Z/ZDivModAux.vo \
+ theories/Ints/Z/ZAux.vo \
theories/Ints/Basic_type.vo theories/Ints/Int31.vo \
theories/Ints/num/GenBase.vo theories/Ints/num/ZnZ.vo \
theories/Ints/num/GenAdd.vo theories/Ints/num/GenSub.vo \
@@ -524,10 +522,6 @@ INTSVO:=\
theories/Ints/num/Nbasic.vo theories/Ints/num/NMake.vo \
theories/Ints/BigN.vo theories/Ints/num/ZMake.vo \
theories/Ints/BigZ.vo theories/Ints/num/BigQ.vo
-# theories/Ints/List/ListAux.vo
-# theories/Ints/List/LPermutation.vo theories/Ints/List/Iterator.vo \
-# theories/Ints/List/ZProgression.vo
-# theories/Ints/Z/ZSum.vo theories/Ints/Z/Ppow.vo \
# spiwack : should use the genN.ml to create NMake eventually
# arnaud : see above
diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v
index 61a6c8e77..29ec5f888 100644
--- a/theories/Init/Tactics.v
+++ b/theories/Init/Tactics.v
@@ -17,9 +17,43 @@ Require Import Logic.
Tactic Notation "revert" ne_hyp_list(l) := generalize l; clear l.
+(**************************************
+ A tactic for proof by contradiction
+ with contradict H
+ H: ~A |- B gives |- A
+ H: ~A |- ~ B gives H: B |- A
+ H: A |- B gives |- ~ A
+ H: A |- B gives |- ~ A
+ H: A |- ~ B gives H: A |- ~ A
+**************************************)
+
+Ltac contradict name :=
+ let term := type of name in (
+ match term with
+ (~_) =>
+ match goal with
+ |- ~ _ => let x := fresh in
+ (intros x; case name;
+ generalize x; clear x name;
+ intro name)
+ | |- _ => case name; clear name
+ end
+ | _ =>
+ match goal with
+ |- ~ _ => let x := fresh in
+ (intros x; absurd term;
+ [idtac | exact name]; generalize x; clear x name;
+ intros name)
+ | |- _ => generalize name; absurd term;
+ [idtac | exact name]; clear name
+ end
+ end).
+
(* to contradict an hypothesis without copying its type. *)
+
Ltac absurd_hyp h :=
+ (* idtac "absurd_hyp is OBSOLETE: use contradict instead."; *)
let T := type of h in
absurd T.
@@ -29,7 +63,9 @@ absurd_hyp H; [apply t | assumption].
(* Transforming a negative goal [ H:~A |- ~B ] into a positive one [ B |- A ]*)
-Ltac swap H := intro; apply H; clear H.
+Ltac swap H :=
+ (* idtac "swap is OBSOLETE: use contradict instead."; *)
+ intro; apply H; clear H.
(* A case with no loss of information. *)
diff --git a/theories/Ints/Basic_type.v b/theories/Ints/Basic_type.v
index f481f3942..72a027cb6 100644
--- a/theories/Ints/Basic_type.v
+++ b/theories/Ints/Basic_type.v
@@ -9,9 +9,6 @@
Set Implicit Arguments.
Require Import ZArith.
-Require Import ZDivModAux.
-Require Import ZPowerAux.
-
Open Local Scope Z_scope.
Section Carry.
diff --git a/theories/Ints/List/Iterator.v b/theories/Ints/List/Iterator.v
deleted file mode 100644
index 327a1454b..000000000
--- a/theories/Ints/List/Iterator.v
+++ /dev/null
@@ -1,180 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-Require Export List.
-Require Export LPermutation.
-Require Import Arith.
-
-Section Iterator.
-Variables A B : Set.
-Variable zero : B.
-Variable f : A -> B.
-Variable g : B -> B -> B.
-Hypothesis g_zero : forall a, g a zero = a.
-Hypothesis g_trans : forall a b c, g a (g b c) = g (g a b) c.
-Hypothesis g_sym : forall a b, g a b = g b a.
-
-Definition iter := fold_right (fun a r => g (f a) r) zero.
-Hint Unfold iter .
-
-Theorem iter_app: forall l1 l2, iter (app l1 l2) = g (iter l1) (iter l2).
-intros l1; elim l1; simpl; auto.
-intros l2; rewrite g_sym; auto.
-intros a l H l2; rewrite H.
-rewrite g_trans; auto.
-Qed.
-
-Theorem iter_permutation: forall l1 l2, permutation l1 l2 -> iter l1 = iter l2.
-intros l1 l2 H; elim H; simpl; auto; clear H l1 l2.
-intros a l1 l2 H1 H2; apply f_equal2 with ( f := g ); auto.
-intros a b l; (repeat rewrite g_trans).
-apply f_equal2 with ( f := g ); auto.
-intros l1 l2 l3 H H0 H1 H2; apply trans_equal with ( 1 := H0 ); auto.
-Qed.
-
-Lemma iter_inv:
- forall P l,
- P zero ->
- (forall a b, P a -> P b -> P (g a b)) ->
- (forall x, In x l -> P (f x)) -> P (iter l).
-intros P l H H0; (elim l; simpl; auto).
-Qed.
-Variable next : A -> A.
-
-Fixpoint progression (m : A) (n : nat) {struct n} : list A :=
- match n with 0 => nil
- | S n1 => cons m (progression (next m) n1) end.
-
-Fixpoint next_n (c : A) (n : nat) {struct n} : A :=
- match n with 0 => c | S n1 => next_n (next c) n1 end.
-
-Theorem progression_app:
- forall a b n m,
- le m n ->
- b = next_n a m ->
- progression a n = app (progression a m) (progression b (n - m)).
-intros a b n m; generalize a b n; clear a b n; elim m; clear m; simpl.
-intros a b n H H0; apply f_equal2 with ( f := progression ); auto with arith.
-intros m H a b n; case n; simpl; clear n.
-intros H1; absurd (0 < 1 + m); auto with arith.
-intros n H0 H1; apply f_equal2 with ( f := @cons A ); auto with arith.
-Qed.
-
-Let iter_progression := fun m n => iter (progression m n).
-
-Theorem iter_progression_app:
- forall a b n m,
- le m n ->
- b = next_n a m ->
- iter (progression a n) =
- g (iter (progression a m)) (iter (progression b (n - m))).
-intros a b n m H H0; unfold iter_progression; rewrite (progression_app a b n m);
- (try apply iter_app); auto.
-Qed.
-
-Theorem length_progression: forall z n, length (progression z n) = n.
-intros z n; generalize z; elim n; simpl; auto.
-Qed.
-
-End Iterator.
-Implicit Arguments iter [A B].
-Implicit Arguments progression [A].
-Implicit Arguments next_n [A].
-Hint Unfold iter .
-Hint Unfold progression .
-Hint Unfold next_n .
-
-Theorem iter_ext:
- forall (A B : Set) zero (f1 : A -> B) f2 g l,
- (forall a, In a l -> f1 a = f2 a) -> iter zero f1 g l = iter zero f2 g l.
-intros A B zero f1 f2 g l; elim l; simpl; auto.
-intros a l0 H H0; apply f_equal2 with ( f := g ); auto.
-Qed.
-
-Theorem iter_map:
- forall (A B C : Set) zero (f : B -> C) g (k : A -> B) l,
- iter zero f g (map k l) = iter zero (fun x => f (k x)) g l.
-intros A B C zero f g k l; elim l; simpl; auto.
-intros; apply f_equal2 with ( f := g ); auto with arith.
-Qed.
-
-Theorem iter_comp:
- forall (A B : Set) zero (f1 f2 : A -> B) g l,
- (forall a, g a zero = a) ->
- (forall a b c, g a (g b c) = g (g a b) c) ->
- (forall a b, g a b = g b a) ->
- g (iter zero f1 g l) (iter zero f2 g l) =
- iter zero (fun x => g (f1 x) (f2 x)) g l.
-intros A B zero f1 f2 g l g_zero g_trans g_sym; elim l; simpl; auto.
-intros a l0 H; rewrite <- H; (repeat rewrite <- g_trans).
-apply f_equal2 with ( f := g ); auto.
-(repeat rewrite g_trans); apply f_equal2 with ( f := g ); auto.
-Qed.
-
-Theorem iter_com:
- forall (A B : Set) zero (f : A -> A -> B) g l1 l2,
- (forall a, g a zero = a) ->
- (forall a b c, g a (g b c) = g (g a b) c) ->
- (forall a b, g a b = g b a) ->
- iter zero (fun x => iter zero (fun y => f x y) g l1) g l2 =
- iter zero (fun y => iter zero (fun x => f x y) g l2) g l1.
-intros A B zero f g l1 l2 H H0 H1; generalize l2; elim l1; simpl; auto;
- clear l1 l2.
-intros l2; elim l2; simpl; auto with arith.
-intros; rewrite H1; rewrite H; auto with arith.
-intros a l1 H2 l2; case l2; clear l2; simpl; auto.
-elim l1; simpl; auto with arith.
-intros; rewrite H1; rewrite H; auto with arith.
-intros b l2.
-rewrite <- (iter_comp
- _ _ zero (fun x => f x a)
- (fun x => iter zero (fun (y : A) => f x y) g l1)); auto with arith.
-rewrite <- (iter_comp
- _ _ zero (fun y => f b y)
- (fun y => iter zero (fun (x : A) => f x y) g l2)); auto with arith.
-(repeat rewrite H0); auto.
-apply f_equal2 with ( f := g ); auto.
-(repeat rewrite <- H0); auto.
-apply f_equal2 with ( f := g ); auto.
-Qed.
-
-Theorem iter_comp_const:
- forall (A B : Set) zero (f : A -> B) g k l,
- k zero = zero ->
- (forall a b, k (g a b) = g (k a) (k b)) ->
- k (iter zero f g l) = iter zero (fun x => k (f x)) g l.
-intros A B zero f g k l H H0; elim l; simpl; auto.
-intros a l0 H1; rewrite H0; apply f_equal2 with ( f := g ); auto.
-Qed.
-
-Lemma next_n_S: forall n m, next_n S n m = plus n m.
-intros n m; generalize n; elim m; clear n m; simpl; auto with arith.
-intros m H n; case n; simpl; auto with arith.
-rewrite H; auto with arith.
-intros n1; rewrite H; simpl; auto with arith.
-Qed.
-
-Theorem progression_S_le_init:
- forall n m p, In p (progression S n m) -> le n p.
-intros n m; generalize n; elim m; clear n m; simpl; auto.
-intros; contradiction.
-intros m H n p [H1|H1]; auto with arith.
-subst n; auto.
-apply le_S_n; auto with arith.
-Qed.
-
-Theorem progression_S_le_end:
- forall n m p, In p (progression S n m) -> lt p (n + m).
-intros n m; generalize n; elim m; clear n m; simpl; auto.
-intros; contradiction.
-intros m H n p [H1|H1]; auto with arith.
-subst n; auto with arith.
-rewrite <- plus_n_Sm; auto with arith.
-rewrite <- plus_n_Sm; auto with arith.
-generalize (H (S n) p); auto with arith.
-Qed.
diff --git a/theories/Ints/List/LPermutation.v b/theories/Ints/List/LPermutation.v
deleted file mode 100644
index 9270ded43..000000000
--- a/theories/Ints/List/LPermutation.v
+++ /dev/null
@@ -1,509 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-(**********************************************************************
- Permutation.v
-
- Defintion and properties of permutations
- **********************************************************************)
-Require Export List.
-Require Export ListAux.
-
-Section permutation.
-Variable A : Set.
-
-(**************************************
- Definition of permutations as sequences of adjacent transpositions
- **************************************)
-
-Inductive permutation : list A -> list A -> Prop :=
- | permutation_nil : permutation nil nil
- | permutation_skip :
- forall (a : A) (l1 l2 : list A),
- permutation l2 l1 -> permutation (a :: l2) (a :: l1)
- | permutation_swap :
- forall (a b : A) (l : list A), permutation (a :: b :: l) (b :: a :: l)
- | permutation_trans :
- forall l1 l2 l3 : list A,
- permutation l1 l2 -> permutation l2 l3 -> permutation l1 l3.
-Hint Constructors permutation.
-
-(**************************************
- Reflexivity
- **************************************)
-
-Theorem permutation_refl : forall l : list A, permutation l l.
-simple induction l.
-apply permutation_nil.
-intros a l1 H.
-apply permutation_skip with (1 := H).
-Qed.
-Hint Resolve permutation_refl.
-
-(**************************************
- Symmetry
- **************************************)
-
-Theorem permutation_sym :
- forall l m : list A, permutation l m -> permutation m l.
-intros l1 l2 H'; elim H'.
-apply permutation_nil.
-intros a l1' l2' H1 H2.
-apply permutation_skip with (1 := H2).
-intros a b l1'.
-apply permutation_swap.
-intros l1' l2' l3' H1 H2 H3 H4.
-apply permutation_trans with (1 := H4) (2 := H2).
-Qed.
-
-(**************************************
- Compatibility with list length
- **************************************)
-
-Theorem permutation_length :
- forall l m : list A, permutation l m -> length l = length m.
-intros l m H'; elim H'; simpl in |- *; auto.
-intros l1 l2 l3 H'0 H'1 H'2 H'3.
-rewrite <- H'3; auto.
-Qed.
-
-(**************************************
- A permutation of the nil list is the nil list
- **************************************)
-
-Theorem permutation_nil_inv : forall l : list A, permutation l nil -> l = nil.
-intros l H; generalize (permutation_length _ _ H); case l; simpl in |- *;
- auto.
-intros; discriminate.
-Qed.
-
-(**************************************
- A permutation of the singleton list is the singleton list
- **************************************)
-
-Let permutation_one_inv_aux :
- forall l1 l2 : list A,
- permutation l1 l2 -> forall a : A, l1 = a :: nil -> l2 = a :: nil.
-intros l1 l2 H; elim H; clear H l1 l2; auto.
-intros a l3 l4 H0 H1 b H2.
-eq_tac.
-injection H2; auto.
-apply permutation_nil_inv; auto.
-injection H2; intros H3 H4; rewrite <- H3; auto.
-apply permutation_sym; auto.
-intros; discriminate.
-Qed.
-
-Theorem permutation_one_inv :
- forall (a : A) (l : list A), permutation (a :: nil) l -> l = a :: nil.
-intros a l H; apply permutation_one_inv_aux with (l1 := a :: nil); auto.
-Qed.
-
-(**************************************
- Compatibility with the belonging
- **************************************)
-
-Theorem permutation_in :
- forall (a : A) (l m : list A), permutation l m -> In a l -> In a m.
-intros a l m H; elim H; simpl in |- *; auto; intuition.
-Qed.
-
-(**************************************
- Compatibility with the append function
- **************************************)
-
-Theorem permutation_app_comp :
- forall l1 l2 l3 l4,
- permutation l1 l2 -> permutation l3 l4 -> permutation (l1 ++ l3) (l2 ++ l4).
-intros l1 l2 l3 l4 H1; generalize l3 l4; elim H1; clear H1 l1 l2 l3 l4;
- simpl in |- *; auto.
-intros a b l l3 l4 H.
-cut (permutation (l ++ l3) (l ++ l4)); auto.
-intros; apply permutation_trans with (a :: b :: l ++ l4); auto.
-elim l; simpl in |- *; auto.
-intros l1 l2 l3 H H0 H1 H2 l4 l5 H3.
-apply permutation_trans with (l2 ++ l4); auto.
-Qed.
-Hint Resolve permutation_app_comp.
-
-(**************************************
- Swap two sublists
- **************************************)
-
-Theorem permutation_app_swap :
- forall l1 l2, permutation (l1 ++ l2) (l2 ++ l1).
-intros l1; elim l1; auto.
-intros; rewrite <- app_nil_end; auto.
-intros a l H l2.
-replace (l2 ++ a :: l) with ((l2 ++ a :: nil) ++ l).
-apply permutation_trans with (l ++ l2 ++ a :: nil); auto.
-apply permutation_trans with (((a :: nil) ++ l2) ++ l); auto.
-simpl in |- *; auto.
-apply permutation_trans with (l ++ (a :: nil) ++ l2); auto.
-apply permutation_sym; auto.
-replace (l2 ++ a :: l) with ((l2 ++ a :: nil) ++ l).
-apply permutation_app_comp; auto.
-elim l2; simpl in |- *; auto.
-intros a0 l0 H0.
-apply permutation_trans with (a0 :: a :: l0); auto.
-apply (app_ass l2 (a :: nil) l).
-apply (app_ass l2 (a :: nil) l).
-Qed.
-
-(**************************************
- A transposition is a permutation
- **************************************)
-
-Theorem permutation_transposition :
- forall a b l1 l2 l3,
- permutation (l1 ++ a :: l2 ++ b :: l3) (l1 ++ b :: l2 ++ a :: l3).
-intros a b l1 l2 l3.
-apply permutation_app_comp; auto.
-change
- (permutation ((a :: nil) ++ l2 ++ (b :: nil) ++ l3)
- ((b :: nil) ++ l2 ++ (a :: nil) ++ l3)) in |- *.
-repeat rewrite <- app_ass.
-apply permutation_app_comp; auto.
-apply permutation_trans with ((b :: nil) ++ (a :: nil) ++ l2); auto.
-apply permutation_app_swap; auto.
-repeat rewrite app_ass.
-apply permutation_app_comp; auto.
-apply permutation_app_swap; auto.
-Qed.
-
-(**************************************
- An element of a list can be put on top of the list to get a permutation
- **************************************)
-
-Theorem in_permutation_ex :
- forall a l, In a l -> exists l1 : list A, permutation (a :: l1) l.
-intros a l; elim l; simpl in |- *; auto.
-intros H; case H; auto.
-intros a0 l0 H [H0| H0].
-exists l0; rewrite H0; auto.
-case H; auto; intros l1 Hl1; exists (a0 :: l1).
-apply permutation_trans with (a0 :: a :: l1); auto.
-Qed.
-
-(**************************************
- A permutation of a cons can be inverted
- **************************************)
-
-Let permutation_cons_ex_aux :
- forall (a : A) (l1 l2 : list A),
- permutation l1 l2 ->
- forall l11 l12 : list A,
- l1 = l11 ++ a :: l12 ->
- exists l3 : list A,
- (exists l4 : list A,
- l2 = l3 ++ a :: l4 /\ permutation (l11 ++ l12) (l3 ++ l4)).
-intros a l1 l2 H; elim H; clear H l1 l2.
-intros l11 l12; case l11; simpl in |- *; intros; discriminate.
-intros a0 l1 l2 H H0 l11 l12; case l11; simpl in |- *.
-exists (nil (A:=A)); exists l1; simpl in |- *; split; auto.
-eq_tac; injection H1; auto.
-injection H1; intros H2 H3; rewrite <- H2; auto.
-intros a1 l111 H1.
-case (H0 l111 l12); auto.
-injection H1; auto.
-intros l3 (l4, (Hl1, Hl2)).
-exists (a0 :: l3); exists l4; split; simpl in |- *; auto.
-eq_tac; injection H1; auto.
-injection H1; intros H2 H3; rewrite H3; auto.
-intros a0 b l l11 l12; case l11; simpl in |- *.
-case l12; try (intros; discriminate).
-intros a1 l0 H; exists (b :: nil); exists l0; simpl in |- *; split; auto.
-repeat eq_tac; injection H; auto.
-injection H; intros H1 H2 H3; rewrite H2; auto.
-intros a1 l111; case l111; simpl in |- *.
-intros H; exists (nil (A:=A)); exists (a0 :: l12); simpl in |- *; split; auto.
-repeat eq_tac; injection H; auto.
-injection H; intros H1 H2 H3; rewrite H3; auto.
-intros a2 H1111 H; exists (a2 :: a1 :: H1111); exists l12; simpl in |- *;
- split; auto.
-repeat eq_tac; injection H; auto.
-intros l1 l2 l3 H H0 H1 H2 l11 l12 H3.
-case H0 with (1 := H3).
-intros l4 (l5, (Hl1, Hl2)).
-case H2 with (1 := Hl1).
-intros l6 (l7, (Hl3, Hl4)).
-exists l6; exists l7; split; auto.
-apply permutation_trans with (1 := Hl2); auto.
-Qed.
-
-Theorem permutation_cons_ex :
- forall (a : A) (l1 l2 : list A),
- permutation (a :: l1) l2 ->
- exists l3 : list A,
- (exists l4 : list A, l2 = l3 ++ a :: l4 /\ permutation l1 (l3 ++ l4)).
-intros a l1 l2 H.
-apply (permutation_cons_ex_aux a (a :: l1) l2 H nil l1); simpl in |- *; auto.
-Qed.
-
-(**************************************
- A permutation can be simply inverted if the two list starts with a cons
- **************************************)
-
-Theorem permutation_inv :
- forall (a : A) (l1 l2 : list A),
- permutation (a :: l1) (a :: l2) -> permutation l1 l2.
-intros a l1 l2 H; case permutation_cons_ex with (1 := H).
-intros l3 (l4, (Hl1, Hl2)).
-apply permutation_trans with (1 := Hl2).
-generalize Hl1; case l3; simpl in |- *; auto.
-intros H1; injection H1; intros H2; rewrite H2; auto.
-intros a0 l5 H1; injection H1; intros H2 H3; rewrite H2; rewrite H3; auto.
-apply permutation_trans with (a0 :: l4 ++ l5); auto.
-apply permutation_skip; apply permutation_app_swap.
-apply (permutation_app_swap (a0 :: l4) l5).
-Qed.
-
-(**************************************
- Take a list and return tle list of all pairs of an element of the
- list and the remaining list
- **************************************)
-
-Fixpoint split_one (l : list A) : list (A * list A) :=
- match l with
- | nil => nil (A:=A * list A)
- | a :: l1 =>
- (a, l1)
- :: map (fun p : A * list A => (fst p, a :: snd p)) (split_one l1)
- end.
-
-(**************************************
- The pairs of the list are a permutation
- **************************************)
-
-Theorem split_one_permutation :
- forall (a : A) (l1 l2 : list A),
- In (a, l1) (split_one l2) -> permutation (a :: l1) l2.
-intros a l1 l2; generalize a l1; elim l2; clear a l1 l2; simpl in |- *; auto.
-intros a l1 H1; case H1.
-intros a l H a0 l1 [H0| H0].
-injection H0; intros H1 H2; rewrite H2; rewrite H1; auto.
-generalize H H0; elim (split_one l); simpl in |- *; auto.
-intros H1 H2; case H2.
-intros a1 l0 H1 H2 [H3| H3]; auto.
-injection H3; intros H4 H5; (rewrite <- H4; rewrite <- H5).
-apply permutation_trans with (a :: fst a1 :: snd a1); auto.
-apply permutation_skip.
-apply H2; auto.
-case a1; simpl in |- *; auto.
-Qed.
-
-(**************************************
- All elements of the list are there
- **************************************)
-
-Theorem split_one_in_ex :
- forall (a : A) (l1 : list A),
- In a l1 -> exists l2 : list A, In (a, l2) (split_one l1).
-intros a l1; elim l1; simpl in |- *; auto.
-intros H; case H.
-intros a0 l H [H0| H0]; auto.
-exists l; left; eq_tac; auto.
-case H; auto.
-intros x H1; exists (a0 :: x); right; auto.
-apply
- (in_map (fun p : A * list A => (fst p, a0 :: snd p)) (split_one l) (a, x));
- auto.
-Qed.
-
-(**************************************
- An auxillary function to generate all permutations
- **************************************)
-
-Fixpoint all_permutations_aux (l : list A) (n : nat) {struct n} :
- list (list A) :=
- match n with
- | O => nil :: nil
- | S n1 =>
- flat_map
- (fun p : A * list A =>
- map (cons (fst p)) (all_permutations_aux (snd p) n1)) (
- split_one l)
- end.
-(**************************************
- Generate all the permutations
- **************************************)
-
-Definition all_permutations (l : list A) := all_permutations_aux l (length l).
-
-(**************************************
- All the elements of the list are permutations
- **************************************)
-
-Let all_permutations_aux_permutation :
- forall (n : nat) (l1 l2 : list A),
- n = length l2 -> In l1 (all_permutations_aux l2 n) -> permutation l1 l2.
-intros n; elim n; simpl in |- *; auto.
-intros l1 l2; case l2.
-simpl in |- *; intros H0 [H1| H1].
-rewrite <- H1; auto.
-case H1.
-simpl in |- *; intros; discriminate.
-intros n0 H l1 l2 H0 H1.
-case in_flat_map_ex with (1 := H1).
-clear H1; intros x; case x; clear x; intros a1 l3 (H1, H2).
-case in_map_inv with (1 := H2).
-simpl in |- *; intros y (H3, H4).
-rewrite H4; auto.
-apply permutation_trans with (a1 :: l3); auto.
-apply permutation_skip; auto.
-apply H with (2 := H3).
-apply eq_add_S.
-apply trans_equal with (1 := H0).
-change (length l2 = length (a1 :: l3)) in |- *.
-apply permutation_length; auto.
-apply permutation_sym; apply split_one_permutation; auto.
-apply split_one_permutation; auto.
-Qed.
-
-Theorem all_permutations_permutation :
- forall l1 l2 : list A, In l1 (all_permutations l2) -> permutation l1 l2.
-intros l1 l2 H; apply all_permutations_aux_permutation with (n := length l2);
- auto.
-Qed.
-
-(**************************************
- A permutation is in the list
- **************************************)
-
-Let permutation_all_permutations_aux :
- forall (n : nat) (l1 l2 : list A),
- n = length l2 -> permutation l1 l2 -> In l1 (all_permutations_aux l2 n).
-intros n; elim n; simpl in |- *; auto.
-intros l1 l2; case l2.
-intros H H0; rewrite permutation_nil_inv with (1 := H0); auto with datatypes.
-simpl in |- *; intros; discriminate.
-intros n0 H l1; case l1.
-intros l2 H0 H1;
- rewrite permutation_nil_inv with (1 := permutation_sym _ _ H1) in H0;
- discriminate.
-clear l1; intros a1 l1 l2 H1 H2.
-case (split_one_in_ex a1 l2); auto.
-apply permutation_in with (1 := H2); auto with datatypes.
-intros x H0.
-apply in_flat_map with (b := (a1, x)); auto.
-apply in_map; simpl in |- *.
-apply H; auto.
-apply eq_add_S.
-apply trans_equal with (1 := H1).
-change (length l2 = length (a1 :: x)) in |- *.
-apply permutation_length; auto.
-apply permutation_sym; apply split_one_permutation; auto.
-apply permutation_inv with (a := a1).
-apply permutation_trans with (1 := H2).
-apply permutation_sym; apply split_one_permutation; auto.
-Qed.
-
-Theorem permutation_all_permutations :
- forall l1 l2 : list A, permutation l1 l2 -> In l1 (all_permutations l2).
-intros l1 l2 H; unfold all_permutations in |- *;
- apply permutation_all_permutations_aux; auto.
-Qed.
-
-(**************************************
- Permutation is decidable
- **************************************)
-
-Definition permutation_dec :
- (forall a b : A, {a = b} + {a <> b}) ->
- forall l1 l2 : list A, {permutation l1 l2} + {~ permutation l1 l2}.
-intros H l1 l2.
-case (In_dec (list_eq_dec H) l1 (all_permutations l2)).
-intros i; left; apply all_permutations_permutation; auto.
-intros i; right; contradict i; apply permutation_all_permutations; auto.
-Defined.
-
-End permutation.
-
-(**************************************
- Hints
- **************************************)
-
-Hint Constructors permutation.
-Hint Resolve permutation_refl.
-Hint Resolve permutation_app_comp.
-Hint Resolve permutation_app_swap.
-
-(**************************************
- Implicits
- **************************************)
-
-Implicit Arguments permutation [A].
-Implicit Arguments split_one [A].
-Implicit Arguments all_permutations [A].
-Implicit Arguments permutation_dec [A].
-
-(**************************************
- Permutation is compatible with map
- **************************************)
-
-Theorem permutation_map :
- forall (A B : Set) (f : A -> B) l1 l2,
- permutation l1 l2 -> permutation (map f l1) (map f l2).
-intros A B f l1 l2 H; elim H; simpl in |- *; auto.
-intros l0 l3 l4 H0 H1 H2 H3; apply permutation_trans with (2 := H3); auto.
-Qed.
-Hint Resolve permutation_map.
-
-(**************************************
- Permutation of a map can be inverted
- *************************************)
-
-Let permutation_map_ex_aux :
- forall (A B : Set) (f : A -> B) l1 l2 l3,
- permutation l1 l2 ->
- l1 = map f l3 -> exists l4, permutation l4 l3 /\ l2 = map f l4.
-intros A1 B1 f l1 l2 l3 H; generalize l3; elim H; clear H l1 l2 l3.
-intros l3; case l3; simpl in |- *; auto.
-intros H; exists (nil (A:=A1)); auto.
-intros; discriminate.
-intros a0 l1 l2 H H0 l3; case l3; simpl in |- *; auto.
-intros; discriminate.
-intros a1 l H1; case (H0 l); auto.
-injection H1; auto.
-intros l5 (H2, H3); exists (a1 :: l5); split; simpl in |- *; auto.
-eq_tac; auto; injection H1; auto.
-intros a0 b l l3; case l3.
-intros; discriminate.
-intros a1 l0; case l0; simpl in |- *.
-intros; discriminate.
-intros a2 l1 H; exists (a2 :: a1 :: l1); split; simpl in |- *; auto.
-repeat eq_tac; injection H; auto.
-intros l1 l2 l3 H H0 H1 H2 l0 H3.
-case H0 with (1 := H3); auto.
-intros l4 (HH1, HH2).
-case H2 with (1 := HH2); auto.
-intros l5 (HH3, HH4); exists l5; split; auto.
-apply permutation_trans with (1 := HH3); auto.
-Qed.
-
-Theorem permutation_map_ex :
- forall (A B : Set) (f : A -> B) l1 l2,
- permutation (map f l1) l2 ->
- exists l3, permutation l3 l1 /\ l2 = map f l3.
-intros A0 B f l1 l2 H; apply permutation_map_ex_aux with (l1 := map f l1);
- auto.
-Qed.
-
-(**************************************
- Permutation is compatible with flat_map
- **************************************)
-
-Theorem permutation_flat_map :
- forall (A B : Set) (f : A -> list B) l1 l2,
- permutation l1 l2 -> permutation (flat_map f l1) (flat_map f l2).
-intros A B f l1 l2 H; elim H; simpl in |- *; auto.
-intros a b l; auto.
-repeat rewrite <- app_ass.
-apply permutation_app_comp; auto.
-intros k3 l4 l5 H0 H1 H2 H3; apply permutation_trans with (1 := H1); auto.
-Qed.
diff --git a/theories/Ints/List/ListAux.v b/theories/Ints/List/ListAux.v
deleted file mode 100644
index 5a6541c95..000000000
--- a/theories/Ints/List/ListAux.v
+++ /dev/null
@@ -1,272 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-(**********************************************************************
- Aux.v
-
- Auxillary functions & Theorems
- **********************************************************************)
-Require Export List.
-Require Export Arith.
-Require Export Tactic.
-Require Import Inverse_Image.
-Require Import Wf_nat.
-
-(**************************************
- Some properties on list operators: app, map,...
-**************************************)
-
-Section List.
-Variables (A : Set) (B : Set) (C : Set).
-Variable f : A -> B.
-
-(**************************************
- An induction theorem for list based on length
-**************************************)
-
-Theorem list_length_ind:
- forall (P : list A -> Prop),
- (forall (l1 : list A),
- (forall (l2 : list A), length l2 < length l1 -> P l2) -> P l1) ->
- forall (l : list A), P l.
-intros P H l;
- apply well_founded_ind with ( R := fun (x y : list A) => length x < length y );
- auto.
-apply wf_inverse_image with ( R := lt ); auto.
-apply lt_wf.
-Qed.
-
-Definition list_length_induction:
- forall (P : list A -> Set),
- (forall (l1 : list A),
- (forall (l2 : list A), length l2 < length l1 -> P l2) -> P l1) ->
- forall (l : list A), P l.
-intros P H l;
- apply well_founded_induction
- with ( R := fun (x y : list A) => length x < length y ); auto.
-apply wf_inverse_image with ( R := lt ); auto.
-apply lt_wf.
-Qed.
-
-Theorem in_ex_app:
- forall (a : A) (l : list A),
- In a l -> (exists l1 : list A , exists l2 : list A , l = l1 ++ (a :: l2) ).
-intros a l; elim l; clear l; simpl; auto.
-intros H; case H.
-intros a1 l H [H1|H1]; auto.
-exists (nil (A:=A)); exists l; simpl; auto.
-eq_tac; auto.
-case H; auto; intros l1 [l2 Hl2]; exists (a1 :: l1); exists l2; simpl; auto.
-eq_tac; auto.
-Qed.
-
-(**************************************
- Properties on app
-**************************************)
-
-Theorem length_app:
- forall (l1 l2 : list A), length (l1 ++ l2) = length l1 + length l2.
-intros l1; elim l1; simpl; auto.
-Qed.
-
-Theorem app_inv_head:
- forall (l1 l2 l3 : list A), l1 ++ l2 = l1 ++ l3 -> l2 = l3.
-intros l1; elim l1; simpl; auto.
-intros a l H l2 l3 H0; apply H; injection H0; auto.
-Qed.
-
-Theorem app_inv_tail:
- forall (l1 l2 l3 : list A), l2 ++ l1 = l3 ++ l1 -> l2 = l3.
-intros l1 l2; generalize l1; elim l2; clear l1 l2; simpl; auto.
-intros l1 l3; case l3; auto.
-intros b l H; absurd (length ((b :: l) ++ l1) <= length l1).
-simpl; rewrite length_app; auto with arith.
-rewrite <- H; auto with arith.
-intros a l H l1 l3; case l3.
-simpl; intros H1; absurd (length (a :: (l ++ l1)) <= length l1).
-simpl; rewrite length_app; auto with arith.
-rewrite H1; auto with arith.
-simpl; intros b l0 H0; injection H0.
-intros H1 H2; eq_tac; auto.
-apply H with ( 1 := H1 ); auto.
-Qed.
-
-Theorem app_inv_app:
- forall l1 l2 l3 l4 a,
- l1 ++ l2 = l3 ++ (a :: l4) ->
- (exists l5 : list A , l1 = l3 ++ (a :: l5) ) \/
- (exists l5 , l2 = l5 ++ (a :: l4) ).
-intros l1; elim l1; simpl; auto.
-intros l2 l3 l4 a H; right; exists l3; auto.
-intros a l H l2 l3 l4 a0; case l3; simpl.
-intros H0; left; exists l; eq_tac; injection H0; auto.
-intros b l0 H0; case (H l2 l0 l4 a0); auto.
-injection H0; auto.
-intros [l5 H1].
-left; exists l5; eq_tac; injection H0; auto.
-Qed.
-
-Theorem app_inv_app2:
- forall l1 l2 l3 l4 a b,
- l1 ++ l2 = l3 ++ (a :: (b :: l4)) ->
- (exists l5 : list A , l1 = l3 ++ (a :: (b :: l5)) ) \/
- ((exists l5 , l2 = l5 ++ (a :: (b :: l4)) ) \/
- l1 = l3 ++ (a :: nil) /\ l2 = b :: l4).
-intros l1; elim l1; simpl; auto.
-intros l2 l3 l4 a b H; right; left; exists l3; auto.
-intros a l H l2 l3 l4 a0 b; case l3; simpl.
-case l; simpl.
-intros H0; right; right; injection H0; split; auto.
-eq_tac; auto.
-intros b0 l0 H0; left; exists l0; injection H0; intros; (repeat eq_tac); auto.
-intros b0 l0 H0; case (H l2 l0 l4 a0 b); auto.
-injection H0; auto.
-intros [l5 HH1]; left; exists l5; eq_tac; auto; injection H0; auto.
-intros [H1|[H1 H2]]; auto.
-right; right; split; auto; eq_tac; auto; injection H0; auto.
-Qed.
-
-Theorem same_length_ex:
- forall (a : A) l1 l2 l3,
- length (l1 ++ (a :: l2)) = length l3 ->
- (exists l4 ,
- exists l5 ,
- exists b : B ,
- length l1 = length l4 /\ (length l2 = length l5 /\ l3 = l4 ++ (b :: l5)) ).
-intros a l1; elim l1; simpl; auto.
-intros l2 l3; case l3; simpl; (try (intros; discriminate)).
-intros b l H; exists (nil (A:=B)); exists l; exists b; (repeat (split; auto)).
-intros a0 l H l2 l3; case l3; simpl; (try (intros; discriminate)).
-intros b l0 H0.
-case (H l2 l0); auto.
-intros l4 [l5 [b1 [HH1 [HH2 HH3]]]].
-exists (b :: l4); exists l5; exists b1; (repeat (simpl; split; auto)).
-eq_tac; auto.
-Qed.
-
-(**************************************
- Properties on map
-**************************************)
-
-Theorem in_map_inv:
- forall (b : B) (l : list A),
- In b (map f l) -> (exists a : A , In a l /\ b = f a ).
-intros b l; elim l; simpl; auto.
-intros tmp; case tmp.
-intros a0 l0 H [H1|H1]; auto.
-exists a0; auto.
-case (H H1); intros a1 [H2 H3]; exists a1; auto.
-Qed.
-
-Theorem in_map_fst_inv:
- forall a (l : list (B * C)),
- In a (map (fst (B:=_)) l) -> (exists c , In (a, c) l ).
-intros a l; elim l; simpl; auto.
-intros H; case H.
-intros a0 l0 H [H0|H0]; auto.
-exists (snd a0); left; rewrite <- H0; case a0; simpl; auto.
-case H; auto; intros l1 Hl1; exists l1; auto.
-Qed.
-
-Theorem length_map: forall l, length (map f l) = length l.
-intros l; elim l; simpl; auto.
-Qed.
-
-Theorem map_app: forall l1 l2, map f (l1 ++ l2) = map f l1 ++ map f l2.
-intros l; elim l; simpl; auto.
-intros a l0 H l2; eq_tac; auto.
-Qed.
-
-Theorem map_length_decompose:
- forall l1 l2 l3 l4,
- length l1 = length l2 ->
- map f (app l1 l3) = app l2 l4 -> map f l1 = l2 /\ map f l3 = l4.
-intros l1; elim l1; simpl; auto; clear l1.
-intros l2; case l2; simpl; auto.
-intros; discriminate.
-intros a l1 Rec l2; case l2; simpl; clear l2; auto.
-intros; discriminate.
-intros b l2 l3 l4 H1 H2.
-injection H2; clear H2; intros H2 H3.
-case (Rec l2 l3 l4); auto.
-intros H4 H5; split; auto.
-eq_tac; auto.
-Qed.
-
-(**************************************
- Properties of flat_map
-**************************************)
-
-Theorem in_flat_map:
- forall (l : list B) (f : B -> list C) a b,
- In a (f b) -> In b l -> In a (flat_map f l).
-intros l g; elim l; simpl; auto.
-intros a l0 H a0 b H0 [H1|H1]; apply in_or_app; auto.
-left; rewrite H1; auto.
-right; apply H with ( b := b ); auto.
-Qed.
-
-Theorem in_flat_map_ex:
- forall (l : list B) (f : B -> list C) a,
- In a (flat_map f l) -> (exists b , In b l /\ In a (f b) ).
-intros l g; elim l; simpl; auto.
-intros a H; case H.
-intros a l0 H a0 H0; case in_app_or with ( 1 := H0 ); simpl; auto.
-intros H1; exists a; auto.
-intros H1; case H with ( 1 := H1 ).
-intros b [H2 H3]; exists b; simpl; auto.
-Qed.
-
-(**************************************
- Properties of fold_left
-**************************************)
-
-Theorem fold_left_invol:
- forall (f: A -> B -> A) (P: A -> Prop) l a,
- P a -> (forall x y, P x -> P (f x y)) -> P (fold_left f l a).
-intros f1 P l; elim l; simpl; auto.
-Qed.
-
-Theorem fold_left_invol_in:
- forall (f: A -> B -> A) (P: A -> Prop) l a b,
- In b l -> (forall x, P (f x b)) -> (forall x y, P x -> P (f x y)) ->
- P (fold_left f l a).
-intros f1 P l; elim l; simpl; auto.
-intros a1 b HH; case HH.
-intros a1 l1 Rec a2 b [V|V] V1 V2; subst; auto.
-apply fold_left_invol; auto.
-apply Rec with (b := b); auto.
-Qed.
-
-End List.
-
-
-(**************************************
- Propertie of list_prod
-**************************************)
-
-Theorem length_list_prod:
- forall (A : Set) (l1 l2 : list A),
- length (list_prod l1 l2) = length l1 * length l2.
-intros A l1 l2; elim l1; simpl; auto.
-intros a l H; rewrite length_app; rewrite length_map; rewrite H; auto.
-Qed.
-
-Theorem in_list_prod_inv:
- forall (A B : Set) a l1 l2,
- In a (list_prod l1 l2) ->
- (exists b : A , exists c : B , a = (b, c) /\ (In b l1 /\ In c l2) ).
-intros A B a l1 l2; elim l1; simpl; auto; clear l1.
-intros H; case H.
-intros a1 l1 H1 H2.
-case in_app_or with ( 1 := H2 ); intros H3; auto.
-case in_map_inv with ( 1 := H3 ); intros b1 [Hb1 Hb2]; auto.
-exists a1; exists b1; split; auto.
-case H1; auto; intros b1 [c1 [Hb1 [Hb2 Hb3]]].
-exists b1; exists c1; split; auto.
-Qed.
diff --git a/theories/Ints/List/UList.v b/theories/Ints/List/UList.v
deleted file mode 100644
index 5248a8b1f..000000000
--- a/theories/Ints/List/UList.v
+++ /dev/null
@@ -1,286 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-(***********************************************************************
- UList.v
-
- Definition of list with distinct elements
-
- Definition: ulist
-************************************************************************)
-Require Import List.
-Require Import Arith.
-Require Import Permutation.
-Require Import ListSet.
-
-Section UniqueList.
-Variable A : Set.
-Variable eqA_dec : forall (a b : A), ({ a = b }) + ({ a <> b }).
-(* A list is unique if there is not twice the same element in the list *)
-
-Inductive ulist : list A -> Prop :=
- ulist_nil: ulist nil
- | ulist_cons: forall a l, ~ In a l -> ulist l -> ulist (a :: l) .
-Hint Constructors ulist .
-(* Inversion theorem *)
-
-Theorem ulist_inv: forall a l, ulist (a :: l) -> ulist l.
-intros a l H; inversion H; auto.
-Qed.
-(* The append of two unique list is unique if the list are distinct *)
-
-Theorem ulist_app:
- forall l1 l2,
- ulist l1 ->
- ulist l2 -> (forall (a : A), In a l1 -> In a l2 -> False) -> ulist (l1 ++ l2).
-intros L1; elim L1; simpl; auto.
-intros a l H l2 H0 H1 H2; apply ulist_cons; simpl; auto.
-red; intros H3; case in_app_or with ( 1 := H3 ); auto; intros H4.
-inversion H0; auto.
-apply H2 with a; auto.
-apply H; auto.
-apply ulist_inv with ( 1 := H0 ); auto.
-intros a0 H3 H4; apply (H2 a0); auto.
-Qed.
-(* Iinversion theorem the appended list *)
-
-Theorem ulist_app_inv:
- forall l1 l2 (a : A), ulist (l1 ++ l2) -> In a l1 -> In a l2 -> False.
-intros l1; elim l1; simpl; auto.
-intros a l H l2 a0 H0 [H1|H1] H2.
-inversion H0 as [|a1 l0 H3 H4 H5]; auto.
-case H4; rewrite H1; auto with datatypes.
-apply (H l2 a0); auto.
-apply ulist_inv with ( 1 := H0 ); auto.
-Qed.
-(* Iinversion theorem the appended list *)
-
-Theorem ulist_app_inv_l: forall (l1 l2 : list A), ulist (l1 ++ l2) -> ulist l1.
-intros l1; elim l1; simpl; auto.
-intros a l H l2 H0.
-inversion H0 as [|il1 iH1 iH2 il2 [iH4 iH5]]; apply ulist_cons; auto.
-intros H5; case iH2; auto with datatypes.
-apply H with l2; auto.
-Qed.
-(* Iinversion theorem the appended list *)
-
-Theorem ulist_app_inv_r: forall (l1 l2 : list A), ulist (l1 ++ l2) -> ulist l2.
-intros l1; elim l1; simpl; auto.
-intros a l H l2 H0; inversion H0; auto.
-Qed.
-(* Uniqueness is decidable *)
-
-Definition ulist_dec: forall l, ({ ulist l }) + ({ ~ ulist l }).
-intros l; elim l; auto.
-intros a l1 [H|H]; auto.
-case (In_dec eqA_dec a l1); intros H2; auto.
-right; red; intros H1; inversion H1; auto.
-right; intros H1; case H; apply ulist_inv with ( 1 := H1 ).
-Defined.
-(* Uniqueness is compatible with permutation *)
-
-Theorem ulist_perm:
- forall (l1 l2 : list A), permutation l1 l2 -> ulist l1 -> ulist l2.
-intros l1 l2 H; elim H; clear H l1 l2; simpl; auto.
-intros a l1 l2 H0 H1 H2; apply ulist_cons; auto.
-inversion_clear H2 as [|ia il iH1 iH2 [iH3 iH4]]; auto.
-intros H3; case iH1;
- apply permutation_in with ( 1 := permutation_sym _ _ _ H0 ); auto.
-inversion H2; auto.
-intros a b L H0; apply ulist_cons; auto.
-inversion_clear H0 as [|ia il iH1 iH2]; auto.
-inversion_clear iH2 as [|ia il iH3 iH4]; auto.
-intros H; case H; auto.
-intros H1; case iH1; rewrite H1; simpl; auto.
-apply ulist_cons; auto.
-inversion_clear H0 as [|ia il iH1 iH2]; auto.
-intros H; case iH1; simpl; auto.
-inversion_clear H0 as [|ia il iH1 iH2]; auto.
-inversion iH2; auto.
-Qed.
-
-Theorem ulist_def:
- forall l a,
- In a l -> ulist l -> ~ (exists l1 , permutation l (a :: (a :: l1)) ).
-intros l a H H0 [l1 H1].
-absurd (ulist (a :: (a :: l1))); auto.
-intros H2; inversion_clear H2; simpl; auto with datatypes.
-apply ulist_perm with ( 1 := H1 ); auto.
-Qed.
-
-Theorem ulist_incl_permutation:
- forall (l1 l2 : list A),
- ulist l1 -> incl l1 l2 -> (exists l3 , permutation l2 (l1 ++ l3) ).
-intros l1; elim l1; simpl; auto.
-intros l2 H H0; exists l2; simpl; auto.
-intros a l H l2 H0 H1; auto.
-case (in_permutation_ex _ a l2); auto with datatypes.
-intros l3 Hl3.
-case (H l3); auto.
-apply ulist_inv with ( 1 := H0 ); auto.
-intros b Hb.
-assert (H2: In b (a :: l3)).
-apply permutation_in with ( 1 := permutation_sym _ _ _ Hl3 );
- auto with datatypes.
-simpl in H2 |-; case H2; intros H3; simpl; auto.
-inversion_clear H0 as [|c lc Hk1]; auto.
-case Hk1; subst a; auto.
-intros l4 H4; exists l4.
-apply permutation_trans with (a :: l3); auto.
-apply permutation_sym; auto.
-Qed.
-
-Theorem ulist_eq_permutation:
- forall (l1 l2 : list A),
- ulist l1 -> incl l1 l2 -> length l1 = length l2 -> permutation l1 l2.
-intros l1 l2 H1 H2 H3.
-case (ulist_incl_permutation l1 l2); auto.
-intros l3 H4.
-assert (H5: l3 = @nil A).
-generalize (permutation_length _ _ _ H4); rewrite length_app; rewrite H3.
-rewrite plus_comm; case l3; simpl; auto.
-intros a l H5; absurd (lt (length l2) (length l2)); auto with arith.
-pattern (length l2) at 2; rewrite H5; auto with arith.
-replace l1 with (app l1 l3); auto.
-apply permutation_sym; auto.
-rewrite H5; rewrite app_nil_end; auto.
-Qed.
-
-
-Theorem ulist_incl_length:
- forall (l1 l2 : list A), ulist l1 -> incl l1 l2 -> le (length l1) (length l2).
-intros l1 l2 H1 Hi; case ulist_incl_permutation with ( 2 := Hi ); auto.
-intros l3 Hl3; rewrite permutation_length with ( 1 := Hl3 ); auto.
-rewrite length_app; simpl; auto with arith.
-Qed.
-
-Theorem ulist_incl2_permutation:
- forall (l1 l2 : list A),
- ulist l1 -> ulist l2 -> incl l1 l2 -> incl l2 l1 -> permutation l1 l2.
-intros l1 l2 H1 H2 H3 H4.
-apply ulist_eq_permutation; auto.
-apply le_antisym; apply ulist_incl_length; auto.
-Qed.
-
-
-Theorem ulist_incl_length_strict:
- forall (l1 l2 : list A),
- ulist l1 -> incl l1 l2 -> ~ incl l2 l1 -> lt (length l1) (length l2).
-intros l1 l2 H1 Hi Hi0; case ulist_incl_permutation with ( 2 := Hi ); auto.
-intros l3 Hl3; rewrite permutation_length with ( 1 := Hl3 ); auto.
-rewrite length_app; simpl; auto with arith.
-generalize Hl3; case l3; simpl; auto with arith.
-rewrite <- app_nil_end; auto.
-intros H2; case Hi0; auto.
-intros a HH; apply permutation_in with ( 1 := H2 ); auto.
-intros a l Hl0; (rewrite plus_comm; simpl; rewrite plus_comm; auto with arith).
-Qed.
-
-Theorem in_inv_dec:
- forall (a b : A) l, In a (cons b l) -> a = b \/ ~ a = b /\ In a l.
-intros a b l H; case (eqA_dec a b); auto; intros H1.
-right; split; auto; inversion H; auto.
-case H1; auto.
-Qed.
-
-Theorem in_ex_app_first:
- forall (a : A) (l : list A),
- In a l ->
- (exists l1 : list A , exists l2 : list A , l = l1 ++ (a :: l2) /\ ~ In a l1 ).
-intros a l; elim l; clear l; auto.
-intros H; case H.
-intros a1 l H H1; auto.
-generalize (in_inv_dec _ _ _ H1); intros [H2|[H2 H3]].
-exists (nil (A:=A)); exists l; simpl; split; auto.
-eq_tac; auto.
-case H; auto; intros l1 [l2 [Hl2 Hl3]]; exists (a1 :: l1); exists l2; simpl;
- split; auto.
-eq_tac; auto.
-intros H4; case H4; auto.
-Qed.
-
-Theorem ulist_inv_ulist:
- forall (l : list A),
- ~ ulist l ->
- (exists a ,
- exists l1 ,
- exists l2 ,
- exists l3 , l = l1 ++ ((a :: l2) ++ (a :: l3)) /\ ulist (l1 ++ (a :: l2)) ).
-intros l; elim l using list_length_ind; clear l.
-intros l; case l; simpl; auto; clear l.
-intros Rec H0; case H0; auto.
-intros a l H H0.
-case (In_dec eqA_dec a l); intros H1; auto.
-case in_ex_app_first with ( 1 := H1 ); intros l1 [l2 [Hl1 Hl2]]; subst l.
-case (ulist_dec l1); intros H2.
-exists a; exists (@nil A); exists l1; exists l2; split; auto.
-simpl; apply ulist_cons; auto.
-case (H l1); auto.
-rewrite length_app; auto with arith.
-intros b [l3 [l4 [l5 [Hl3 Hl4]]]]; subst l1.
-exists b; exists (a :: l3); exists l4; exists (l5 ++ (a :: l2)); split; simpl;
- auto.
-(repeat (rewrite <- ass_app; simpl)); auto.
-apply ulist_cons; auto.
-contradict Hl2; auto.
-replace (l3 ++ (b :: (l4 ++ (b :: l5)))) with ((l3 ++ (b :: l4)) ++ (b :: l5));
- auto with datatypes.
-(repeat (rewrite <- ass_app; simpl)); auto.
-case (H l); auto; intros a1 [l1 [l2 [l3 [Hl3 Hl4]]]]; subst l.
-exists a1; exists (a :: l1); exists l2; exists l3; split; auto.
-simpl; apply ulist_cons; auto.
-contradict H1.
-replace (l1 ++ (a1 :: (l2 ++ (a1 :: l3))))
- with ((l1 ++ (a1 :: l2)) ++ (a1 :: l3)); auto with datatypes.
-(repeat (rewrite <- ass_app; simpl)); auto.
-Qed.
-
-Theorem incl_length_repetition:
- forall (l1 l2 : list A),
- incl l1 l2 ->
- lt (length l2) (length l1) ->
- (exists a ,
- exists ll1 ,
- exists ll2 ,
- exists ll3 ,
- l1 = ll1 ++ ((a :: ll2) ++ (a :: ll3)) /\ ulist (ll1 ++ (a :: ll2)) ).
-intros l1 l2 H H0; apply ulist_inv_ulist.
-intros H1; absurd (le (length l1) (length l2)); auto with arith.
-apply ulist_incl_length; auto.
-Qed.
-
-End UniqueList.
-Implicit Arguments ulist [A].
-Hint Constructors ulist .
-
-Theorem ulist_map:
- forall (A B : Set) (f : A -> B) l,
- (forall x y, (In x l) -> (In y l) -> f x = f y -> x = y) -> ulist l -> ulist (map f l).
-intros a b f l Hf Hl; generalize Hf; elim Hl; clear Hf; auto.
-simpl; auto.
-intros a1 l1 H1 H2 H3 Hf; simpl.
-apply ulist_cons; auto with datatypes.
-contradict H1.
-case in_map_inv with ( 1 := H1 ); auto with datatypes.
-intros b1 [Hb1 Hb2].
-replace a1 with b1; auto with datatypes.
-Qed.
-
-Theorem ulist_list_prod:
- forall (A : Set) (l1 l2 : list A),
- ulist l1 -> ulist l2 -> ulist (list_prod l1 l2).
-intros A l1 l2 Hl1 Hl2; elim Hl1; simpl; auto.
-intros a l H1 H2 H3; apply ulist_app; auto.
-apply ulist_map; auto.
-intros x y _ _ H; inversion H; auto.
-intros p Hp1 Hp2; case H1.
-case in_map_inv with ( 1 := Hp1 ); intros a1 [Ha1 Ha2]; auto.
-case in_list_prod_inv with ( 1 := Hp2 ); intros b1 [c1 [Hb1 [Hb2 Hb3]]]; auto.
-replace a with b1; auto.
-rewrite Ha2 in Hb1; injection Hb1; auto.
-Qed.
diff --git a/theories/Ints/List/ZProgression.v b/theories/Ints/List/ZProgression.v
deleted file mode 100644
index e4c15e38d..000000000
--- a/theories/Ints/List/ZProgression.v
+++ /dev/null
@@ -1,105 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-Require Export Iterator.
-Require Import ZArith.
-Require Export UList.
-Open Scope Z_scope.
-
-Theorem next_n_Z: forall n m, next_n Zsucc n m = n + Z_of_nat m.
-intros n m; generalize n; elim m; clear n m.
-intros n; simpl; auto with zarith.
-intros m H n.
-replace (n + Z_of_nat (S m)) with (Zsucc n + Z_of_nat m); auto with zarith.
-rewrite <- H; auto with zarith.
-rewrite inj_S; auto with zarith.
-Qed.
-
-Theorem Zprogression_end:
- forall n m,
- progression Zsucc n (S m) =
- app (progression Zsucc n m) (cons (n + Z_of_nat m) nil).
-intros n m; generalize n; elim m; clear n m.
-simpl; intros; apply f_equal2 with ( f := @cons Z ); auto with zarith.
-intros m1 Hm1 n1.
-apply trans_equal with (cons n1 (progression Zsucc (Zsucc n1) (S m1))); auto.
-rewrite Hm1.
-replace (Zsucc n1 + Z_of_nat m1) with (n1 + Z_of_nat (S m1)); auto with zarith.
-replace (Z_of_nat (S m1)) with (1 + Z_of_nat m1); auto with zarith.
-rewrite inj_S; auto with zarith.
-Qed.
-
-Theorem Zprogression_pred_end:
- forall n m,
- progression Zpred n (S m) =
- app (progression Zpred n m) (cons (n - Z_of_nat m) nil).
-intros n m; generalize n; elim m; clear n m.
-simpl; intros; apply f_equal2 with ( f := @cons Z ); auto with zarith.
-intros m1 Hm1 n1.
-apply trans_equal with (cons n1 (progression Zpred (Zpred n1) (S m1))); auto.
-rewrite Hm1.
-replace (Zpred n1 - Z_of_nat m1) with (n1 - Z_of_nat (S m1)); auto with zarith.
-replace (Z_of_nat (S m1)) with (1 + Z_of_nat m1); auto with zarith.
-unfold Zpred; ring.
-rewrite inj_S; auto with zarith.
-Qed.
-
-Theorem Zprogression_opp:
- forall n m,
- rev (progression Zsucc n m) = progression Zpred (n + Z_of_nat (pred m)) m.
-intros n m; generalize n; elim m; clear n m.
-simpl; auto.
-intros m Hm n.
-rewrite (Zprogression_end n); auto.
-rewrite distr_rev.
-rewrite Hm; simpl; auto.
-case m.
-simpl; auto.
-intros m1;
- replace (n + Z_of_nat (pred (S m1))) with (Zpred (n + Z_of_nat (S m1))); auto.
-rewrite inj_S; simpl; (unfold Zpred; unfold Zsucc); auto with zarith.
-Qed.
-
-Theorem Zprogression_le_init:
- forall n m p, In p (progression Zsucc n m) -> (n <= p).
-intros n m; generalize n; elim m; clear n m; simpl; auto.
-intros; contradiction.
-intros m H n p [H1|H1]; auto with zarith.
-generalize (H _ _ H1); auto with zarith.
-Qed.
-
-Theorem Zprogression_le_end:
- forall n m p, In p (progression Zsucc n m) -> (p < n + Z_of_nat m).
-intros n m; generalize n; elim m; clear n m; auto.
-intros; contradiction.
-intros m H n p H1; simpl in H1 |-; case H1; clear H1; intros H1;
- auto with zarith.
-subst n; auto with zarith.
-apply Zle_lt_trans with (p + 0); auto with zarith.
-apply Zplus_lt_compat_l; red; simpl; auto with zarith.
-apply Zlt_le_trans with (Zsucc n + Z_of_nat m); auto with zarith.
-rewrite inj_S; rewrite Zplus_succ_comm; auto with zarith.
-Qed.
-
-Theorem ulist_Zprogression: forall a n, ulist (progression Zsucc a n).
-intros a n; generalize a; elim n; clear a n; simpl; auto with zarith.
-intros n H1 a; apply ulist_cons; auto.
-intros H2; absurd (Zsucc a <= a); auto with zarith.
-apply Zprogression_le_init with ( 1 := H2 ).
-Qed.
-
-Theorem in_Zprogression:
- forall a b n, ( a <= b < a + Z_of_nat n ) -> In b (progression Zsucc a n).
-intros a b n; generalize a b; elim n; clear a b n; auto with zarith.
-simpl; auto with zarith.
-intros n H a b.
-replace (a + Z_of_nat (S n)) with (Zsucc a + Z_of_nat n); auto with zarith.
-intros [H1 H2]; simpl; auto with zarith.
-case (Zle_lt_or_eq _ _ H1); auto with zarith.
-rewrite inj_S; auto with zarith.
-Qed.
diff --git a/theories/Ints/Tactic.v b/theories/Ints/Tactic.v
deleted file mode 100644
index 6837f5922..000000000
--- a/theories/Ints/Tactic.v
+++ /dev/null
@@ -1,76 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-
-(**********************************************************************
- Tactic.v
- Useful tactics
- **********************************************************************)
-
-(**************************************
- A simple tactic to end a proof
-**************************************)
-Ltac finish := intros; auto; trivial; discriminate.
-
-
-(**************************************
- A tactic for proof by contradiction
- with contradict H
- H: ~A |- B gives |- A
- H: ~A |- ~ B gives H: B |- A
- H: A |- B gives |- ~ A
- H: A |- B gives |- ~ A
- H: A |- ~ B gives H: A |- ~ A
-**************************************)
-
-Ltac contradict name :=
- let term := type of name in (
- match term with
- (~_) =>
- match goal with
- |- ~ _ => let x := fresh in
- (intros x; case name;
- generalize x; clear x name;
- intro name)
- | |- _ => case name; clear name
- end
- | _ =>
- match goal with
- |- ~ _ => let x := fresh in
- (intros x; absurd term;
- [idtac | exact name]; generalize x; clear x name;
- intros name)
- | |- _ => generalize name; absurd term;
- [idtac | exact name]; clear name
- end
- end).
-
-
-(**************************************
- A tactic to use f_equal? theorems
-**************************************)
-
-Ltac eq_tac :=
- match goal with
- |- (?g _ = ?g _) => apply f_equal with (f := g)
- | |- (?g ?X _ = ?g ?X _) => apply f_equal with (f := g X)
- | |- (?g _ _ = ?g _ _) => apply f_equal2 with (f := g)
- | |- (?g ?X ?Y _ = ?g ?X ?Y _) => apply f_equal with (f := g X Y)
- | |- (?g ?X _ _ = ?g ?X _ _) => apply f_equal2 with (f := g X)
- | |- (?g _ _ _ = ?g _ _ _) => apply f_equal3 with (f := g)
- | |- (?g ?X ?Y ?Z _ = ?g ?X ?Y ?Z _) => apply f_equal with (f := g X Y Z)
- | |- (?g ?X ?Y _ _ = ?g ?X ?Y _ _) => apply f_equal2 with (f := g X Y)
- | |- (?g ?X _ _ _ = ?g ?X _ _ _) => apply f_equal3 with (f := g X)
- | |- (?g _ _ _ _ _ = ?g _ _ _ _) => apply f_equal4 with (f := g)
- end.
-
-(**************************************
- A stupid tactic that tries auto also after applying sym_equal
-**************************************)
-
-Ltac sauto := (intros; apply sym_equal; auto; fail) || auto.
diff --git a/theories/Ints/Z/IntsZmisc.v b/theories/Ints/Z/IntsZmisc.v
deleted file mode 100644
index e4dee2d4f..000000000
--- a/theories/Ints/Z/IntsZmisc.v
+++ /dev/null
@@ -1,185 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-Require Export ZArith.
-Open Local Scope Z_scope.
-
-Coercion Zpos : positive >-> Z.
-Coercion Z_of_N : N >-> Z.
-
-Lemma Zpos_plus : forall p q, Zpos (p + q) = p + q.
-Proof. intros;trivial. Qed.
-
-Lemma Zpos_mult : forall p q, Zpos (p * q) = p * q.
-Proof. intros;trivial. Qed.
-
-Lemma Zpos_xI_add : forall p, Zpos (xI p) = Zpos p + Zpos p + Zpos 1.
-Proof. intros p;rewrite Zpos_xI;ring. Qed.
-
-Lemma Zpos_xO_add : forall p, Zpos (xO p) = Zpos p + Zpos p.
-Proof. intros p;rewrite Zpos_xO;ring. Qed.
-
-Lemma Psucc_Zplus : forall p, Zpos (Psucc p) = p + 1.
-Proof. intros p;rewrite Zpos_succ_morphism;unfold Zsucc;trivial. Qed.
-
-Hint Rewrite Zpos_xI_add Zpos_xO_add Pplus_carry_spec
- Psucc_Zplus Zpos_plus : zmisc.
-
-Lemma Zlt_0_pos : forall p, 0 < Zpos p.
-Proof. unfold Zlt;trivial. Qed.
-
-
-Lemma Pminus_mask_carry_spec : forall p q,
- Pminus_mask_carry p q = Pminus_mask p (Psucc q).
-Proof.
- intros p q;generalize q p;clear q p.
- induction q;destruct p;simpl;try rewrite IHq;trivial.
- destruct p;trivial. destruct p;trivial.
-Qed.
-
-Hint Rewrite Pminus_mask_carry_spec : zmisc.
-
-Ltac zsimpl := autorewrite with zmisc.
-Ltac CaseEq t := generalize (refl_equal t);pattern t at -1;case t.
-Ltac generalizeclear H := generalize H;clear H.
-
-Lemma Pminus_mask_spec :
- forall p q,
- match Pminus_mask p q with
- | IsNul => Zpos p = Zpos q
- | IsPos k => Zpos p = q + k
- | IsNeq => p < q
- end.
-Proof with zsimpl;auto with zarith.
- induction p;destruct q;simpl;zsimpl;
- match goal with
- | [|- context [(Pminus_mask ?p1 ?q1)]] =>
- assert (H1 := IHp q1);destruct (Pminus_mask p1 q1)
- | _ => idtac
- end;simpl ...
- inversion H1 ... inversion H1 ...
- rewrite Psucc_Zplus in H1 ...
- clear IHp;induction p;simpl ...
- rewrite IHp;destruct (Pdouble_minus_one p) ...
- assert (H:= Zlt_0_pos q) ... assert (H:= Zlt_0_pos q) ...
-Qed.
-
-Definition PminusN x y :=
- match Pminus_mask x y with
- | IsPos k => Npos k
- | _ => N0
- end.
-
-Lemma PminusN_le : forall x y:positive, x <= y -> Z_of_N (PminusN y x) = y - x.
-Proof.
- intros x y Hle;unfold PminusN.
- assert (H := Pminus_mask_spec y x);destruct (Pminus_mask y x).
- rewrite H;unfold Z_of_N;auto with zarith.
- rewrite H;unfold Z_of_N;auto with zarith.
- elimtype False;omega.
-Qed.
-
-Lemma Ppred_Zminus : forall p, 1< Zpos p -> (p-1)%Z = Ppred p.
-Proof. destruct p;simpl;trivial. intros;elimtype False;omega. Qed.
-
-
-Open Local Scope positive_scope.
-
-Delimit Scope P_scope with P.
-Open Local Scope P_scope.
-
-Definition is_lt (n m : positive) :=
- match (n ?= m) Eq with
- | Lt => true
- | _ => false
- end.
-Infix "?<" := is_lt (at level 70, no associativity) : P_scope.
-
-Lemma is_lt_spec : forall n m, if n ?< m then (n < m)%Z else (m <= n)%Z.
-Proof.
- intros n m; unfold is_lt, Zlt, Zle, Zcompare.
- rewrite (ZC4 m n);destruct ((n ?= m) Eq);trivial;try (intro;discriminate).
-Qed.
-
-Definition is_eq a b :=
- match (a ?= b) Eq with
- | Eq => true
- | _ => false
- end.
-Infix "?=" := is_eq (at level 70, no associativity) : P_scope.
-
-Lemma is_eq_refl : forall n, n ?= n = true.
-Proof. intros n;unfold is_eq;rewrite Pcompare_refl;trivial. Qed.
-
-Lemma is_eq_eq : forall n m, n ?= m = true -> n = m.
-Proof.
- unfold is_eq;intros n m H; apply Pcompare_Eq_eq;
- destruct ((n ?= m)%positive Eq);trivial;try discriminate.
-Qed.
-
-Lemma is_eq_spec_pos : forall n m, if n ?= m then n = m else m <> n.
-Proof.
- intros n m; CaseEq (n ?= m);intro H.
- rewrite (is_eq_eq _ _ H);trivial.
- intro H1;rewrite H1 in H;rewrite is_eq_refl in H;discriminate H.
-Qed.
-
-Lemma is_eq_spec : forall n m, if n ?= m then Zpos n = m else Zpos m <> n.
-Proof.
- intros n m; CaseEq (n ?= m);intro H.
- rewrite (is_eq_eq _ _ H);trivial.
- intro H1;inversion H1.
- rewrite H2 in H;rewrite is_eq_refl in H;discriminate H.
-Qed.
-
-Definition is_Eq a b :=
- match a, b with
- | N0, N0 => true
- | Npos a', Npos b' => a' ?= b'
- | _, _ => false
- end.
-
-Lemma is_Eq_spec :
- forall n m, if is_Eq n m then Z_of_N n = m else Z_of_N m <> n.
-Proof.
- destruct n;destruct m;simpl;trivial;try (intro;discriminate).
- apply is_eq_spec.
-Qed.
-
-(* [times x y] return [x * y], a litle bit more efficiant *)
-Fixpoint times (x y : positive) {struct y} : positive :=
- match x, y with
- | xH, _ => y
- | _, xH => x
- | xO x', xO y' => xO (xO (times x' y'))
- | xO x', xI y' => xO (x' + xO (times x' y'))
- | xI x', xO y' => xO (y' + xO (times x' y'))
- | xI x', xI y' => xI (x' + y' + xO (times x' y'))
- end.
-
-Infix "*" := times : P_scope.
-
-Lemma times_Zmult : forall p q, Zpos (p * q)%P = (p * q)%Z.
-Proof.
- intros p q;generalize q p;clear p q.
- induction q;destruct p; unfold times; try fold (times p q);
- autorewrite with zmisc; try rewrite IHq; ring.
-Qed.
-
-Fixpoint square (x:positive) : positive :=
- match x with
- | xH => xH
- | xO x => xO (xO (square x))
- | xI x => xI (xO (square x + x))
- end.
-
-Lemma square_Zmult : forall x, Zpos (square x) = (x * x) %Z.
-Proof.
- induction x as [x IHx|x IHx |];unfold square;try (fold (square x));
- autorewrite with zmisc; try rewrite IHx; ring.
-Qed.
diff --git a/theories/Ints/Z/Pmod.v b/theories/Ints/Z/Pmod.v
deleted file mode 100644
index 1ea08b4fa..000000000
--- a/theories/Ints/Z/Pmod.v
+++ /dev/null
@@ -1,565 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-Require Import IntsZmisc.
-Open Local Scope P_scope.
-
-(* [div_eucl a b] return [(q,r)] such that a = q*b + r *)
-Fixpoint div_eucl (a b : positive) {struct a} : N * N :=
- match a with
- | xH => if 1 ?< b then (0%N, 1%N) else (1%N, 0%N)
- | xO a' =>
- let (q, r) := div_eucl a' b in
- match q, r with
- | N0, N0 => (0%N, 0%N) (* n'arrive jamais *)
- | N0, Npos r =>
- if (xO r) ?< b then (0%N, Npos (xO r))
- else (1%N,PminusN (xO r) b)
- | Npos q, N0 => (Npos (xO q), 0%N)
- | Npos q, Npos r =>
- if (xO r) ?< b then (Npos (xO q), Npos (xO r))
- else (Npos (xI q),PminusN (xO r) b)
- end
- | xI a' =>
- let (q, r) := div_eucl a' b in
- match q, r with
- | N0, N0 => (0%N, 0%N) (* Impossible *)
- | N0, Npos r =>
- if (xI r) ?< b then (0%N, Npos (xI r))
- else (1%N,PminusN (xI r) b)
- | Npos q, N0 => if 1 ?< b then (Npos (xO q), 1%N) else (Npos (xI q), 0%N)
- | Npos q, Npos r =>
- if (xI r) ?< b then (Npos (xO q), Npos (xI r))
- else (Npos (xI q),PminusN (xI r) b)
- end
- end.
-Infix "/" := div_eucl : P_scope.
-
-Open Scope Z_scope.
-Opaque Zmult.
-Lemma div_eucl_spec : forall a b,
- Zpos a = fst (a/b)%P * b + snd (a/b)%P
- /\ snd (a/b)%P < b.
-Proof with zsimpl;try apply Zlt_0_pos;try ((ring;fail) || omega).
- intros a b;generalize a;clear a;induction a;simpl;zsimpl;
- try (case IHa;clear IHa;repeat rewrite Zmult_0_l;zsimpl;intros H1 H2;
- try rewrite H1; destruct (a/b)%P as [q r];
- destruct q as [|q];destruct r as [|r];simpl in *;
- generalize H1 H2;clear H1 H2);repeat rewrite Zmult_0_l;
- repeat rewrite Zplus_0_r;
- zsimpl;simpl;intros;
- match goal with
- | [H : Zpos _ = 0 |- _] => discriminate H
- | [|- context [ ?xx ?< b ]] =>
- assert (H3 := is_lt_spec xx b);destruct (xx ?< b)
- | _ => idtac
- end;simpl;try assert(H4 := Zlt_0_pos r);split;repeat rewrite Zplus_0_r;
- try (generalize H3;zsimpl;intros);
- try (rewrite PminusN_le;trivial) ...
- assert (Zpos b = 1) ... rewrite H ...
- assert (H4 := Zlt_0_pos b); assert (Zpos b = 1) ...
-Qed.
-Transparent Zmult.
-
-(******** Definition du modulo ************)
-
-(* [mod a b] return [a] modulo [b] *)
-Fixpoint Pmod (a b : positive) {struct a} : N :=
- match a with
- | xH => if 1 ?< b then 1%N else 0%N
- | xO a' =>
- let r := Pmod a' b in
- match r with
- | N0 => 0%N
- | Npos r' =>
- if (xO r') ?< b then Npos (xO r')
- else PminusN (xO r') b
- end
- | xI a' =>
- let r := Pmod a' b in
- match r with
- | N0 => if 1 ?< b then 1%N else 0%N
- | Npos r' =>
- if (xI r') ?< b then Npos (xI r')
- else PminusN (xI r') b
- end
- end.
-
-Infix "mod" := Pmod (at level 40, no associativity) : P_scope.
-Open Local Scope P_scope.
-
-Lemma Pmod_div_eucl : forall a b, a mod b = snd (a/b).
-Proof with auto.
- intros a b;generalize a;clear a;induction a;simpl;
- try (rewrite IHa;
- assert (H1 := div_eucl_spec a b); destruct (a/b) as [q r];
- destruct q as [|q];destruct r as [|r];simpl in *;
- match goal with
- | [|- context [ ?xx ?< b ]] =>
- assert (H2 := is_lt_spec xx b);destruct (xx ?< b)
- | _ => idtac
- end;simpl) ...
- destruct H1 as [H3 H4];discriminate H3.
- destruct (1 ?< b);simpl ...
-Qed.
-
-Lemma mod1: forall a, a mod 1 = 0%N.
-Proof. induction a;simpl;try rewrite IHa;trivial. Qed.
-
-Lemma mod_a_a_0 : forall a, a mod a = N0.
-Proof.
- intros a;generalize (div_eucl_spec a a);rewrite <- Pmod_div_eucl.
- destruct (fst (a / a));unfold Z_of_N at 1.
- rewrite Zmult_0_l;intros (H1,H2);elimtype False;omega.
- assert (a<=p*a).
- pattern (Zpos a) at 1;rewrite <- (Zmult_1_l a).
- assert (H1:= Zlt_0_pos p);assert (H2:= Zle_0_pos a);
- apply Zmult_le_compat;trivial;try omega.
- destruct (a mod a)%P;auto with zarith.
- unfold Z_of_N;assert (H1:= Zlt_0_pos p0);intros (H2,H3);elimtype False;omega.
-Qed.
-
-Lemma mod_le_2r : forall (a b r: positive) (q:N),
- Zpos a = b*q + r -> b <= a -> r < b -> 2*r <= a.
-Proof.
- intros a b r q H0 H1 H2.
- assert (H3:=Zlt_0_pos a). assert (H4:=Zlt_0_pos b). assert (H5:=Zlt_0_pos r).
- destruct q as [|q]. rewrite Zmult_0_r in H0. elimtype False;omega.
- assert (H6:=Zlt_0_pos q). unfold Z_of_N in H0.
- assert (Zpos r = a - b*q). omega.
- simpl;zsimpl. pattern r at 2;rewrite H.
- assert (b <= b * q).
- pattern (Zpos b) at 1;rewrite <- (Zmult_1_r b).
- apply Zmult_le_compat;try omega.
- apply Zle_trans with (a - b * q + b). omega.
- apply Zle_trans with (a - b + b);omega.
-Qed.
-
-Lemma mod_lt : forall a b r, a mod b = Npos r -> r < b.
-Proof.
- intros a b r H;generalize (div_eucl_spec a b);rewrite <- Pmod_div_eucl;
- rewrite H;simpl;intros (H1,H2);omega.
-Qed.
-
-Lemma mod_le : forall a b r, a mod b = Npos r -> r <= b.
-Proof. intros a b r H;assert (H1:= mod_lt _ _ _ H);omega. Qed.
-
-Lemma mod_le_a : forall a b r, a mod b = r -> r <= a.
-Proof.
- intros a b r H;generalize (div_eucl_spec a b);rewrite <- Pmod_div_eucl;
- rewrite H;simpl;intros (H1,H2).
- assert (0 <= fst (a / b) * b).
- destruct (fst (a / b));simpl;auto with zarith.
- auto with zarith.
-Qed.
-
-Lemma lt_mod : forall a b, Zpos a < Zpos b -> (a mod b)%P = Npos a.
-Proof.
- intros a b H; rewrite Pmod_div_eucl. case (div_eucl_spec a b).
- assert (0 <= snd(a/b)). destruct (snd(a/b));simpl;auto with zarith.
- destruct (fst (a/b)).
- unfold Z_of_N at 1;rewrite Zmult_0_l;rewrite Zplus_0_l.
- destruct (snd (a/b));simpl; intros H1 H2;inversion H1;trivial.
- unfold Z_of_N at 1;assert (b <= p*b).
- pattern (Zpos b) at 1; rewrite <- (Zmult_1_l (Zpos b)).
- assert (H1 := Zlt_0_pos p);apply Zmult_le_compat;try omega.
- apply Zle_0_pos.
- intros;elimtype False;omega.
-Qed.
-
-Fixpoint gcd_log2 (a b c:positive) {struct c}: option positive :=
- match a mod b with
- | N0 => Some b
- | Npos r =>
- match b mod r, c with
- | N0, _ => Some r
- | Npos r', xH => None
- | Npos r', xO c' => gcd_log2 r r' c'
- | Npos r', xI c' => gcd_log2 r r' c'
- end
- end.
-
-Fixpoint egcd_log2 (a b c:positive) {struct c}:
- option (Z * Z * positive) :=
- match a/b with
- | (_, N0) => Some (0, 1, b)
- | (q, Npos r) =>
- match b/r, c with
- | (_, N0), _ => Some (1, -q, r)
- | (q', Npos r'), xH => None
- | (q', Npos r'), xO c' =>
- match egcd_log2 r r' c' with
- None => None
- | Some (u', v', w') =>
- let u := u' - v' * q' in
- Some (u, v' - q * u, w')
- end
- | (q', Npos r'), xI c' =>
- match egcd_log2 r r' c' with
- None => None
- | Some (u', v', w') =>
- let u := u' - v' * q' in
- Some (u, v' - q * u, w')
- end
- end
- end.
-
-Lemma egcd_gcd_log2: forall c a b,
- match egcd_log2 a b c, gcd_log2 a b c with
- None, None => True
- | Some (u,v,r), Some r' => r = r'
- | _, _ => False
- end.
-induction c; simpl; auto; try
- (intros a b; generalize (Pmod_div_eucl a b); case (a/b); simpl;
- intros q r1 H; subst; case (a mod b); auto;
- intros r; generalize (Pmod_div_eucl b r); case (b/r); simpl;
- intros q' r1 H; subst; case (b mod r); auto;
- intros r'; generalize (IHc r r'); case egcd_log2; auto;
- intros ((p1,p2),p3); case gcd_log2; auto).
-Qed.
-
-Ltac rw l :=
- match l with
- | (?r, ?r1) =>
- match type of r with
- True => rewrite <- r1
- | _ => rw r; rw r1
- end
- | ?r => rewrite r
- end.
-
-Lemma egcd_log2_ok: forall c a b,
- match egcd_log2 a b c with
- None => True
- | Some (u,v,r) => u * a + v * b = r
- end.
-induction c; simpl; auto;
- intros a b; generalize (div_eucl_spec a b); case (a/b);
- simpl fst; simpl snd; intros q r1; case r1; try (intros; ring);
- simpl; intros r (Hr1, Hr2); clear r1;
- generalize (div_eucl_spec b r); case (b/r);
- simpl fst; simpl snd; intros q' r1; case r1;
- try (intros; rewrite Hr1; ring);
- simpl; intros r' (Hr'1, Hr'2); clear r1; auto;
- generalize (IHc r r'); case egcd_log2; auto;
- intros ((u',v'),w'); case gcd_log2; auto; intros;
- rw ((I, H), Hr1, Hr'1); ring.
-Qed.
-
-
-Fixpoint log2 (a:positive) : positive :=
- match a with
- | xH => xH
- | xO a => Psucc (log2 a)
- | xI a => Psucc (log2 a)
- end.
-
-Lemma gcd_log2_1: forall a c, gcd_log2 a xH c = Some xH.
-Proof. destruct c;simpl;try rewrite mod1;trivial. Qed.
-
-Lemma log2_Zle :forall a b, Zpos a <= Zpos b -> log2 a <= log2 b.
-Proof with zsimpl;try omega.
- induction a;destruct b;zsimpl;intros;simpl ...
- assert (log2 a <= log2 b) ... apply IHa ...
- assert (log2 a <= log2 b) ... apply IHa ...
- assert (H1 := Zlt_0_pos a);elimtype False;omega.
- assert (log2 a <= log2 b) ... apply IHa ...
- assert (log2 a <= log2 b) ... apply IHa ...
- assert (H1 := Zlt_0_pos a);elimtype False;omega.
- assert (H1 := Zlt_0_pos (log2 b)) ...
- assert (H1 := Zlt_0_pos (log2 b)) ...
-Qed.
-
-Lemma log2_1_inv : forall a, Zpos (log2 a) = 1 -> a = xH.
-Proof.
- destruct a;simpl;zsimpl;intros;trivial.
- assert (H1:= Zlt_0_pos (log2 a));elimtype False;omega.
- assert (H1:= Zlt_0_pos (log2 a));elimtype False;omega.
-Qed.
-
-Lemma mod_log2 :
- forall a b r:positive, a mod b = Npos r -> b <= a -> log2 r + 1 <= log2 a.
-Proof.
- intros; cut (log2 (xO r) <= log2 a). simpl;zsimpl;trivial.
- apply log2_Zle.
- replace (Zpos (xO r)) with (2 * r)%Z;trivial.
- generalize (div_eucl_spec a b);rewrite <- Pmod_div_eucl;rewrite H.
- rewrite Zmult_comm;intros [H1 H2];apply mod_le_2r with b (fst (a/b));trivial.
-Qed.
-
-Lemma gcd_log2_None_aux :
- forall c a b, Zpos b <= Zpos a -> log2 b <= log2 c ->
- gcd_log2 a b c <> None.
-Proof.
- induction c;simpl;intros;
- (CaseEq (a mod b);[intros Heq|intros r Heq];try (intro;discriminate));
- (CaseEq (b mod r);[intros Heq'|intros r' Heq'];try (intro;discriminate)).
- apply IHc. apply mod_le with b;trivial.
- generalize H0 (mod_log2 _ _ _ Heq' (mod_le _ _ _ Heq));zsimpl;intros;omega.
- apply IHc. apply mod_le with b;trivial.
- generalize H0 (mod_log2 _ _ _ Heq' (mod_le _ _ _ Heq));zsimpl;intros;omega.
- assert (Zpos (log2 b) = 1).
- assert (H1 := Zlt_0_pos (log2 b));omega.
- rewrite (log2_1_inv _ H1) in Heq;rewrite mod1 in Heq;discriminate Heq.
-Qed.
-
-Lemma gcd_log2_None : forall a b, Zpos b <= Zpos a -> gcd_log2 a b b <> None.
-Proof. intros;apply gcd_log2_None_aux;auto with zarith. Qed.
-
-Lemma gcd_log2_Zle :
- forall c1 c2 a b, log2 c1 <= log2 c2 ->
- gcd_log2 a b c1 <> None -> gcd_log2 a b c2 = gcd_log2 a b c1.
-Proof with zsimpl;trivial;try omega.
- induction c1;destruct c2;simpl;intros;
- (destruct (a mod b) as [|r];[idtac | destruct (b mod r)]) ...
- apply IHc1;trivial. generalize H;zsimpl;intros;omega.
- apply IHc1;trivial. generalize H;zsimpl;intros;omega.
- elim H;destruct (log2 c1);trivial.
- apply IHc1;trivial. generalize H;zsimpl;intros;omega.
- apply IHc1;trivial. generalize H;zsimpl;intros;omega.
- elim H;destruct (log2 c1);trivial.
- elim H0;trivial. elim H0;trivial.
-Qed.
-
-Lemma gcd_log2_Zle_log :
- forall a b c, log2 b <= log2 c -> Zpos b <= Zpos a ->
- gcd_log2 a b c = gcd_log2 a b b.
-Proof.
- intros a b c H1 H2; apply gcd_log2_Zle; trivial.
- apply gcd_log2_None; trivial.
-Qed.
-
-Lemma gcd_log2_mod0 :
- forall a b c, a mod b = N0 -> gcd_log2 a b c = Some b.
-Proof. intros a b c H;destruct c;simpl;rewrite H;trivial. Qed.
-
-
-Require Import Zwf.
-
-Lemma Zwf_pos : well_founded (fun x y => Zpos x < Zpos y).
-Proof.
- unfold well_founded.
- assert (forall x a ,x = Zpos a -> Acc (fun x y : positive => x < y) a).
- intros x;assert (Hacc := Zwf_well_founded 0 x);induction Hacc;intros;subst x.
- constructor;intros. apply H0 with (Zpos y);trivial.
- split;auto with zarith.
- intros a;apply H with (Zpos a);trivial.
-Qed.
-
-Opaque Pmod.
-Lemma gcd_log2_mod : forall a b, Zpos b <= Zpos a ->
- forall r, a mod b = Npos r -> gcd_log2 a b b = gcd_log2 b r r.
-Proof.
- intros a b;generalize a;clear a; assert (Hacc := Zwf_pos b).
- induction Hacc; intros a Hle r Hmod.
- rename x into b. destruct b;simpl;rewrite Hmod.
- CaseEq (xI b mod r)%P;intros. rewrite gcd_log2_mod0;trivial.
- assert (H2 := mod_le _ _ _ H1);assert (H3 := mod_lt _ _ _ Hmod);
- assert (H4 := mod_le _ _ _ Hmod).
- rewrite (gcd_log2_Zle_log r p b);trivial.
- symmetry;apply H0;trivial.
- generalize (mod_log2 _ _ _ H1 H4);simpl;zsimpl;intros;omega.
- CaseEq (xO b mod r)%P;intros. rewrite gcd_log2_mod0;trivial.
- assert (H2 := mod_le _ _ _ H1);assert (H3 := mod_lt _ _ _ Hmod);
- assert (H4 := mod_le _ _ _ Hmod).
- rewrite (gcd_log2_Zle_log r p b);trivial.
- symmetry;apply H0;trivial.
- generalize (mod_log2 _ _ _ H1 H4);simpl;zsimpl;intros;omega.
- rewrite mod1 in Hmod;discriminate Hmod.
-Qed.
-
-Lemma gcd_log2_xO_Zle :
- forall a b, Zpos b <= Zpos a -> gcd_log2 a b (xO b) = gcd_log2 a b b.
-Proof.
- intros a b Hle;apply gcd_log2_Zle.
- simpl;zsimpl;auto with zarith.
- apply gcd_log2_None_aux;auto with zarith.
-Qed.
-
-Lemma gcd_log2_xO_Zlt :
- forall a b, Zpos a < Zpos b -> gcd_log2 a b (xO b) = gcd_log2 b a a.
-Proof.
- intros a b H;simpl. assert (Hlt := Zlt_0_pos a).
- assert (H0 := lt_mod _ _ H).
- rewrite H0;simpl.
- CaseEq (b mod a)%P;intros;simpl.
- symmetry;apply gcd_log2_mod0;trivial.
- assert (H2 := mod_lt _ _ _ H1).
- rewrite (gcd_log2_Zle_log a p b);auto with zarith.
- symmetry;apply gcd_log2_mod;auto with zarith.
- apply log2_Zle.
- replace (Zpos p) with (Z_of_N (Npos p));trivial.
- apply mod_le_a with a;trivial.
-Qed.
-
-Lemma gcd_log2_x0 : forall a b, gcd_log2 a b (xO b) <> None.
-Proof.
- intros;simpl;CaseEq (a mod b)%P;intros. intro;discriminate.
- CaseEq (b mod p)%P;intros. intro;discriminate.
- assert (H1 := mod_le_a _ _ _ H0). unfold Z_of_N in H1.
- assert (H2 := mod_le _ _ _ H0).
- apply gcd_log2_None_aux. trivial.
- apply log2_Zle. trivial.
-Qed.
-
-Lemma egcd_log2_x0 : forall a b, egcd_log2 a b (xO b) <> None.
-Proof.
-intros a b H; generalize (egcd_gcd_log2 (xO b) a b) (gcd_log2_x0 a b);
- rw H; case gcd_log2; auto.
-Qed.
-
-Definition gcd a b :=
- match gcd_log2 a b (xO b) with
- | Some p => p
- | None => (* can not appear *) 1%positive
- end.
-
-Definition egcd a b :=
- match egcd_log2 a b (xO b) with
- | Some p => p
- | None => (* can not appear *) (1,1,1%positive)
- end.
-
-
-Lemma gcd_mod0 : forall a b, (a mod b)%P = N0 -> gcd a b = b.
-Proof.
- intros a b H;unfold gcd.
- pattern (gcd_log2 a b (xO b)) at 1;
- rewrite (gcd_log2_mod0 _ _ (xO b) H);trivial.
-Qed.
-
-Lemma gcd1 : forall a, gcd a xH = xH.
-Proof. intros a;rewrite gcd_mod0;[trivial|apply mod1]. Qed.
-
-Lemma gcd_mod : forall a b r, (a mod b)%P = Npos r ->
- gcd a b = gcd b r.
-Proof.
- intros a b r H;unfold gcd.
- assert (log2 r <= log2 (xO r)). simpl;zsimpl;omega.
- assert (H1 := mod_lt _ _ _ H).
- pattern (gcd_log2 b r (xO r)) at 1; rewrite gcd_log2_Zle_log;auto with zarith.
- destruct (Z_lt_le_dec a b).
- pattern (gcd_log2 a b (xO b)) at 1; rewrite gcd_log2_xO_Zlt;trivial.
- rewrite (lt_mod _ _ z) in H;inversion H.
- assert (r <= b). omega.
- generalize (gcd_log2_None _ _ H2).
- destruct (gcd_log2 b r r);intros;trivial.
- assert (log2 b <= log2 (xO b)). simpl;zsimpl;omega.
- pattern (gcd_log2 a b (xO b)) at 1; rewrite gcd_log2_Zle_log;auto with zarith.
- pattern (gcd_log2 a b b) at 1;rewrite (gcd_log2_mod _ _ z _ H).
- assert (r <= b). omega.
- generalize (gcd_log2_None _ _ H3).
- destruct (gcd_log2 b r r);intros;trivial.
-Qed.
-
-Require Import ZArith.
-Require Import Znumtheory.
-
-Hint Rewrite Zpos_mult times_Zmult square_Zmult Psucc_Zplus: zmisc.
-
-Ltac mauto :=
- trivial;autorewrite with zmisc;trivial;auto with zarith.
-
-Lemma gcd_Zis_gcd : forall a b:positive, (Zis_gcd b a (gcd b a)%P).
-Proof with mauto.
- intros a;assert (Hacc := Zwf_pos a);induction Hacc;rename x into a;intros.
- generalize (div_eucl_spec b a)...
- rewrite <- (Pmod_div_eucl b a).
- CaseEq (b mod a)%P;[intros Heq|intros r Heq]; intros (H1,H2).
- simpl in H1;rewrite Zplus_0_r in H1.
- rewrite (gcd_mod0 _ _ Heq).
- constructor;mauto.
- apply Zdivide_intro with (fst (b/a)%P);trivial.
- rewrite (gcd_mod _ _ _ Heq).
- rewrite H1;apply Zis_gcd_sym.
- rewrite Zmult_comm;apply Zis_gcd_for_euclid2;simpl in *.
- apply Zis_gcd_sym;auto.
-Qed.
-
-Lemma egcd_Zis_gcd : forall a b:positive,
- let (uv,w) := egcd a b in
- let (u,v) := uv in
- u * a + v * b = w /\ (Zis_gcd b a w).
-Proof with mauto.
- intros a b; unfold egcd.
- generalize (egcd_log2_ok (xO b) a b) (egcd_gcd_log2 (xO b) a b)
- (egcd_log2_x0 a b) (gcd_Zis_gcd b a); unfold egcd, gcd.
- case egcd_log2; try (intros ((u,v),w)); case gcd_log2;
- try (intros; match goal with H: False |- _ => case H end);
- try (intros _ _ H1; case H1; auto; fail).
- intros; subst; split; try apply Zis_gcd_sym; auto.
-Qed.
-
-Definition Zgcd a b :=
- match a, b with
- | Z0, _ => b
- | _, Z0 => a
- | Zpos a, Zneg b => Zpos (gcd a b)
- | Zneg a, Zpos b => Zpos (gcd a b)
- | Zpos a, Zpos b => Zpos (gcd a b)
- | Zneg a, Zneg b => Zpos (gcd a b)
- end.
-
-
-Lemma Zgcd_is_gcd : forall x y, Zis_gcd x y (Zgcd x y).
-Proof.
- destruct x;destruct y;simpl.
- apply Zis_gcd_0.
- apply Zis_gcd_sym;apply Zis_gcd_0.
- apply Zis_gcd_sym;apply Zis_gcd_0.
- apply Zis_gcd_0.
- apply gcd_Zis_gcd.
- apply Zis_gcd_sym;apply Zis_gcd_minus;simpl;apply gcd_Zis_gcd.
- apply Zis_gcd_0.
- apply Zis_gcd_minus;simpl;apply Zis_gcd_sym;apply gcd_Zis_gcd.
- apply Zis_gcd_minus;apply Zis_gcd_minus;simpl;apply gcd_Zis_gcd.
-Qed.
-
-Definition Zegcd a b :=
- match a, b with
- | Z0, Z0 => (0,0,0)
- | Zpos _, Z0 => (1,0,a)
- | Zneg _, Z0 => (-1,0,-a)
- | Z0, Zpos _ => (0,1,b)
- | Z0, Zneg _ => (0,-1,-b)
- | Zpos a, Zneg b =>
- match egcd a b with (u,v,w) => (u,-v, Zpos w) end
- | Zneg a, Zpos b =>
- match egcd a b with (u,v,w) => (-u,v, Zpos w) end
- | Zpos a, Zpos b =>
- match egcd a b with (u,v,w) => (u,v, Zpos w) end
- | Zneg a, Zneg b =>
- match egcd a b with (u,v,w) => (-u,-v, Zpos w) end
- end.
-
-Lemma Zegcd_is_egcd : forall x y,
- match Zegcd x y with
- (u,v,w) => u * x + v * y = w /\ Zis_gcd x y w /\ 0 <= w
- end.
-Proof.
- assert (zx0: forall x, Zneg x = -x).
- simpl; auto.
- assert (zx1: forall x, -(-x) = x).
- intro x; case x; simpl; auto.
- destruct x;destruct y;simpl; try (split; [idtac|split]);
- auto; try (red; simpl; intros; discriminate);
- try (rewrite zx0; apply Zis_gcd_minus; try rewrite zx1; auto;
- apply Zis_gcd_minus; try rewrite zx1; simpl; auto);
- try apply Zis_gcd_0; try (apply Zis_gcd_sym;apply Zis_gcd_0);
- generalize (egcd_Zis_gcd p p0); case egcd; intros (u,v,w) (H1, H2);
- split; repeat rewrite zx0; try (rewrite <- H1; ring); auto;
- (split; [idtac | red; intros; discriminate]).
- apply Zis_gcd_sym; auto.
- apply Zis_gcd_sym; apply Zis_gcd_minus; rw zx1;
- apply Zis_gcd_sym; auto.
- apply Zis_gcd_minus; rw zx1; auto.
- apply Zis_gcd_minus; rw zx1; auto.
- apply Zis_gcd_minus; rw zx1; auto.
- apply Zis_gcd_sym; auto.
-Qed.
diff --git a/theories/Ints/Z/Ppow.v b/theories/Ints/Z/Ppow.v
deleted file mode 100644
index b4e4ca5ef..000000000
--- a/theories/Ints/Z/Ppow.v
+++ /dev/null
@@ -1,39 +0,0 @@
-Require Import ZArith.
-Require Import ZAux.
-
-Open Scope Z_scope.
-
-Fixpoint Ppow a z {struct z}:=
- match z with
- xH => a
- | xO z1 => let v := Ppow a z1 in (Pmult v v)
- | xI z1 => let v := Ppow a z1 in (Pmult a (Pmult v v))
- end.
-
-Theorem Ppow_correct: forall a z,
- Zpos (Ppow a z) = (Zpos a) ^ (Zpos z).
-intros a z; elim z; simpl Ppow; auto;
- try (intros z1 Hrec; repeat rewrite Zpos_mult_morphism; rewrite Hrec).
- rewrite Zpos_xI; rewrite Zpower_exp; auto with zarith.
- rewrite Zpower_exp_1; rewrite (Zmult_comm 2);
- try rewrite Zpower_mult; auto with zarith.
- change 2 with (1 + 1); rewrite Zpower_exp; auto with zarith.
- rewrite Zpower_exp_1; rewrite Zmult_comm; auto.
- apply Zle_ge; auto with zarith.
- rewrite Zpos_xO; rewrite (Zmult_comm 2);
- rewrite Zpower_mult; auto with zarith.
- change 2 with (1 + 1); rewrite Zpower_exp; auto with zarith.
- rewrite Zpower_exp_1; auto.
- rewrite Zpower_exp_1; auto.
-Qed.
-
-Theorem Ppow_plus: forall a z1 z2,
- Ppow a (z1 + z2) = ((Ppow a z1) * (Ppow a z2))%positive.
-intros a z1 z2.
- assert (tmp: forall x y, Zpos x = Zpos y -> x = y).
- intros x y H; injection H; auto.
- apply tmp.
- rewrite Zpos_mult_morphism; repeat rewrite Ppow_correct.
- rewrite Zpos_plus_distr; rewrite Zpower_exp; auto; red; simpl;
- intros; discriminate.
-Qed.
diff --git a/theories/Ints/Z/ZAux.v b/theories/Ints/Z/ZAux.v
index 83337ee50..8e4b1d64f 100644
--- a/theories/Ints/Z/ZAux.v
+++ b/theories/Ints/Z/ZAux.v
@@ -15,8 +15,13 @@
Require Import ArithRing.
Require Export ZArith.
Require Export Znumtheory.
-Require Export Tactic.
-(* Require Import MOmega. *)
+Require Export Zpow_facts.
+
+(* *** Nota Bene ***
+ All results that were general enough has been moved in ZArith.
+ Only remain here specialized lemmas and compatibility elements.
+ (P.L. 5/11/2007).
+*)
Open Local Scope Z_scope.
@@ -35,51 +40,18 @@ Hint Extern 2 (Zlt _ _) =>
| H: _ < ?p |- _ <= ?p => apply Zle_lt_trans with (2 := H)
end).
+
+Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
+
(**************************************
Properties of order and product
**************************************)
-Theorem Zmult_interval: forall p q, 0 < p * q -> 1 < p -> 0 < q < p * q.
-intros p q H1 H2; assert (0 < q).
-case (Zle_or_lt q 0); auto; intros H3; contradict H1; apply Zle_not_lt.
-rewrite <- (Zmult_0_r p).
-apply Zmult_le_compat_l; auto with zarith.
-split; auto.
-pattern q at 1; rewrite <- (Zmult_1_l q).
-apply Zmult_lt_compat_r; auto with zarith.
-Qed.
-
-Theorem Zmult_lt_compat: forall n m p q : Z, 0 < n <= p -> 0 < m < q -> n * m < p * q.
-intros n m p q (H1, H2) (H3, H4).
-apply Zle_lt_trans with (p * m).
-apply Zmult_le_compat_r; auto with zarith.
-apply Zmult_lt_compat_l; auto with zarith.
-Qed.
-
-Theorem Zle_square_mult: forall a b, 0 <= a <= b -> a * a <= b * b.
-intros a b (H1, H2); apply Zle_trans with (a * b); auto with zarith.
-Qed.
-
-Theorem Zlt_square_mult: forall a b, 0 <= a < b -> a * a < b * b.
-intros a b (H1, H2); apply Zle_lt_trans with (a * b); auto with zarith.
-apply Zmult_lt_compat_r; auto with zarith.
-Qed.
-
-Theorem Zlt_square_mult_inv: forall a b, 0 <= a -> 0 <= b -> a * a < b * b -> a < b.
-intros a b H1 H2 H3; case (Zle_or_lt b a); auto; intros H4; apply Zmult_lt_reg_r with a;
- contradict H3; apply Zle_not_lt; apply Zle_square_mult; auto.
-Qed.
-
-
-Theorem Zpower_2: forall x, x^2 = x * x.
-intros; ring.
-Qed.
-
Theorem beta_lex: forall a b c d beta,
a * beta + b <= c * beta + d ->
0 <= b < beta -> 0 <= d < beta ->
a <= c.
-Proof.
+ Proof.
intros a b c d beta H1 (H3, H4) (H5, H6).
assert (a - c < 1); auto with zarith.
apply Zmult_lt_reg_r with beta; auto with zarith.
@@ -175,1203 +147,161 @@ Theorem mult_add_ineq3: forall x y c cross beta,
apply Zle_trans with (1*beta+cross);auto with zarith.
Qed.
-
-(**************************************
- Properties of Z_nat
- **************************************)
-
-Theorem inj_eq_inv: forall (n m : nat), Z_of_nat n = Z_of_nat m -> n = m.
-intros n m H1; case (le_or_lt n m); auto with arith.
-intros H2; case (le_lt_or_eq _ _ H2); auto; intros H3.
-contradict H1; auto with zarith.
-intros H2; contradict H1; auto with zarith.
-Qed.
-
-Theorem inj_le_inv: forall (n m : nat), Z_of_nat n <= Z_of_nat m-> (n <= m)%nat.
-intros n m H1; case (le_or_lt n m); auto with arith.
-intros H2; contradict H1; auto with zarith.
-Qed.
-
-Theorem Z_of_nat_Zabs_nat:
- forall (z : Z), 0 <= z -> Z_of_nat (Zabs_nat z) = z.
-intros z; case z; simpl; auto with zarith.
-intros; apply sym_equal; apply Zpos_eq_Z_of_nat_o_nat_of_P; auto.
-intros p H1; contradict H1; simpl; auto with zarith.
-Qed.
-
-(**************************************
- Properties of Zabs
-**************************************)
-
-Theorem Zabs_square: forall a, a * a = Zabs a * Zabs a.
-intros a; rewrite <- Zabs_Zmult; apply sym_equal; apply Zabs_eq;
- auto with zarith.
-case (Zle_or_lt 0%Z a); auto with zarith.
-intros Ha; replace (a * a) with (- a * - a); auto with zarith.
-ring.
-Qed.
-
-(**************************************
- Properties of Zabs_nat
-**************************************)
-
-Theorem Z_of_nat_abs_le:
- forall x y, x <= y -> x + Z_of_nat (Zabs_nat (y - x)) = y.
-intros x y Hx1.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-Qed.
-
-Theorem Zabs_nat_Zsucc:
- forall p, 0 <= p -> Zabs_nat (Zsucc p) = S (Zabs_nat p).
-intros p Hp.
-apply inj_eq_inv.
-rewrite inj_S; (repeat rewrite Z_of_nat_Zabs_nat); auto with zarith.
-Qed.
-
-Theorem Zabs_nat_Z_of_nat: forall n, Zabs_nat (Z_of_nat n) = n.
-intros n1; apply inj_eq_inv; rewrite Z_of_nat_Zabs_nat; auto with zarith.
-Qed.
-
-
-(**************************************
- Properties Zsqrt_plain
-**************************************)
-
-Theorem Zsqrt_plain_is_pos: forall n, 0 <= n -> 0 <= Zsqrt_plain n.
-intros n m; case (Zsqrt_interval n); auto with zarith.
-intros H1 H2; case (Zle_or_lt 0 (Zsqrt_plain n)); auto.
-intros H3; contradict H2; apply Zle_not_lt.
-apply Zle_trans with ( 2 := H1 ).
-replace ((Zsqrt_plain n + 1) * (Zsqrt_plain n + 1))
- with (Zsqrt_plain n * Zsqrt_plain n + (2 * Zsqrt_plain n + 1));
- auto with zarith.
-ring.
-Qed.
-
-Theorem Zsqrt_square_id: forall a, 0 <= a -> Zsqrt_plain (a * a) = a.
-intros a H.
-generalize (Zsqrt_plain_is_pos (a * a)); auto with zarith; intros Haa.
-case (Zsqrt_interval (a * a)); auto with zarith.
-intros H1 H2.
-case (Zle_or_lt a (Zsqrt_plain (a * a))); intros H3; auto.
-case Zle_lt_or_eq with ( 1 := H3 ); auto; clear H3; intros H3.
-contradict H1; apply Zlt_not_le; auto with zarith.
-apply Zle_lt_trans with (a * Zsqrt_plain (a * a)); auto with zarith.
-apply Zmult_lt_compat_r; auto with zarith.
-contradict H2; apply Zle_not_lt; auto with zarith.
-apply Zmult_le_compat; auto with zarith.
-Qed.
-
-Theorem Zsqrt_le:
- forall p q, 0 <= p <= q -> Zsqrt_plain p <= Zsqrt_plain q.
-intros p q [H1 H2]; case Zle_lt_or_eq with ( 1 := H2 ); clear H2; intros H2.
-2:subst q; auto with zarith.
-case (Zle_or_lt (Zsqrt_plain p) (Zsqrt_plain q)); auto; intros H3.
-assert (Hp: (0 <= Zsqrt_plain q)).
-apply Zsqrt_plain_is_pos; auto with zarith.
-absurd (q <= p); auto with zarith.
-apply Zle_trans with ((Zsqrt_plain q + 1) * (Zsqrt_plain q + 1)).
-case (Zsqrt_interval q); auto with zarith.
-apply Zle_trans with (Zsqrt_plain p * Zsqrt_plain p); auto with zarith.
-apply Zmult_le_compat; auto with zarith.
-case (Zsqrt_interval p); auto with zarith.
-Qed.
+Hint Rewrite Zmult_1_r Zmult_0_r Zmult_1_l Zmult_0_l Zplus_0_l Zplus_0_r Zminus_0_r: rm10.
(**************************************
- Properties Zpower
+ Properties of Zdiv and Zmod
**************************************)
-Theorem Zpower_1: forall a, 0 <= a -> 1 ^ a = 1.
-intros a Ha; pattern a; apply natlike_ind; auto with zarith.
-intros x Hx Hx1; unfold Zsucc.
-rewrite Zpower_exp; auto with zarith.
-rewrite Hx1; simpl; auto.
-Qed.
-
-Theorem Zpower_exp_0: forall a, a ^ 0 = 1.
-simpl; unfold Zpower_pos; simpl; auto with zarith.
-Qed.
-
-Theorem Zpower_exp_1: forall a, a ^ 1 = a.
-simpl; unfold Zpower_pos; simpl; auto with zarith.
-Qed.
-
-Theorem Zpower_Zabs: forall a b, Zabs (a ^ b) = (Zabs a) ^ b.
-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.
-eq_tac; 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.
-intros p n H.
-unfold Zsucc; rewrite Zpower_exp; auto with zarith.
-rewrite Zpower_exp_1; apply Zmult_comm.
-Qed.
-
-Theorem Zpower_mult: forall p q r, 0 <= q -> 0 <= r -> p ^ (q * r) = (p ^ q) ^ r.
-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_exp_1; try rewrite <- Zpower_exp; try eq_tac; auto with zarith.
-ring.
-apply Zle_ge; replace 0 with (0 * r1); try apply Zmult_le_compat_r; auto.
-Qed.
-
-Theorem Zpower_lt_0: forall a b: Z, 0 < a -> 0 <= b-> 0 < a ^b.
-intros a b; case b; auto with zarith.
-simpl; intros; auto with zarith.
-2: intros p H H1; case H1; auto.
-intros p H1 H; generalize H; pattern (Zpos p); apply natlike_ind; auto.
-intros; case a; compute; auto.
-intros p1 H2 H3 _; unfold Zsucc; rewrite Zpower_exp; simpl; auto with zarith.
-apply Zmult_lt_O_compat; auto with zarith.
-generalize H1; case a; compute; intros; auto; discriminate.
-Qed.
-
-Theorem Zpower_le_monotone: forall a b c: Z, 0 < a -> 0 <= b <= c -> a ^ b <= a ^ c.
-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_lt_0; auto with zarith.
-apply Zlt_le_weak; apply Zpower_lt_0; auto with zarith.
-Qed.
-
-Theorem Zpower_lt_monotone: forall a b c: Z, 1 < a -> 0 <= b < c -> a ^ b < a ^ c.
-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_lt_0; auto with zarith.
-assert (0 < a ^ (c - b)); auto with zarith.
-apply Zpower_lt_0; auto with zarith.
-apply Zlt_le_trans with (a ^1); auto with zarith.
-rewrite Zpower_exp_1; auto with zarith.
-apply Zpower_le_monotone; auto with zarith.
-Qed.
+Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
+ Proof.
+ intros a b H H1;case (Z_mod_lt a b);auto with zarith;intros H2 H3;split;auto.
+ case (Zle_or_lt b a); intros H4; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
+ Qed.
-Theorem Zpower_nat_Zpower: forall p q, 0 <= q -> p ^ q = Zpower_nat p (Zabs_nat q).
-intros p1 q1; case q1; simpl.
-intros _; exact (refl_equal _).
-intros p2 _; apply Zpower_pos_nat.
-intros p2 H1; case H1; auto.
-Qed.
-Theorem Zgt_pow_1 : forall n m : Z, 0 < n -> 1 < m -> 1 < m ^ n.
-Proof.
-intros n m H1 H2.
-replace 1 with (m ^ 0) by apply Zpower_exp_0.
-apply Zpower_lt_monotone; auto with zarith.
-Qed.
+ Theorem Zmod_distr: forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
+ (2 ^a * r + t) mod (2 ^ b) = (2 ^a * r) mod (2 ^ b) + t.
+ Proof.
+ intros a b r t (H1, H2) H3 (H4, H5).
+ assert (t < 2 ^ b).
+ apply Zlt_le_trans with (1:= H5); auto with zarith.
+ apply Zpower_le_monotone; auto with zarith.
+ rewrite Zplus_mod; auto with zarith.
+ rewrite Zmod_small with (a := t); auto with zarith.
+ apply Zmod_small; auto with zarith.
+ split; auto with zarith.
+ assert (0 <= 2 ^a * r); auto with zarith.
+ apply Zplus_le_0_compat; auto with zarith.
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ auto with zarith.
+ pattern (2 ^ b) at 2; replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a);
+ try ring.
+ apply Zplus_le_lt_compat; auto with zarith.
+ replace b with ((b - a) + a); try ring.
+ rewrite Zpower_exp; auto with zarith.
+ pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
+ try rewrite <- Zmult_minus_distr_r.
+ rewrite (Zmult_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l;
+ auto with zarith.
+ rewrite (Zmult_comm (2 ^a)); apply Zmult_le_compat_r; auto with zarith.
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ auto with zarith.
+ Qed.
-(**************************************
- Properties Zmod
-**************************************)
-
-Theorem Zmod_mult:
- forall a b n, 0 < n -> (a * b) mod n = ((a mod n) * (b mod n)) mod n.
-intros a b n H.
-pattern a at 1; rewrite (Z_div_mod_eq a n); auto with zarith.
-pattern b at 1; rewrite (Z_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) +
- (((n * (a / n)) * (b / n) + (b mod n) * (a / n)) + (a mod n) * (b / n)) *
- n); auto with zarith.
-apply Z_mod_plus; auto with zarith.
-ring.
-Qed.
+ Theorem Zmod_shift_r:
+ forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
+ (r * 2 ^a + t) mod (2 ^ b) = (r * 2 ^a) mod (2 ^ b) + t.
+ Proof.
+ intros a b r t (H1, H2) H3 (H4, H5).
+ assert (t < 2 ^ b).
+ apply Zlt_le_trans with (1:= H5); auto with zarith.
+ apply Zpower_le_monotone; auto with zarith.
+ rewrite Zplus_mod; auto with zarith.
+ rewrite Zmod_small with (a := t); auto with zarith.
+ apply Zmod_small; auto with zarith.
+ split; auto with zarith.
+ assert (0 <= 2 ^a * r); auto with zarith.
+ apply Zplus_le_0_compat; auto with zarith.
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ auto with zarith.
+ pattern (2 ^ b) at 2;replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); try ring.
+ apply Zplus_le_lt_compat; auto with zarith.
+ replace b with ((b - a) + a); try ring.
+ rewrite Zpower_exp; auto with zarith.
+ pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
+ try rewrite <- Zmult_minus_distr_r.
+ repeat rewrite (fun x => Zmult_comm x (2 ^ a)); rewrite Zmult_mod_distr_l;
+ auto with zarith.
+ apply Zmult_le_compat_l; auto with zarith.
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ auto with zarith.
+ Qed.
-Theorem Zmod_plus:
- forall a b n, 0 < n -> (a + b) mod n = (a mod n + b mod n) mod n.
-intros a b n H.
-pattern a at 1; rewrite (Z_div_mod_eq a n); auto with zarith.
-pattern b at 1; rewrite (Z_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); auto with zarith.
-apply Z_mod_plus; auto with zarith.
-ring.
-Qed.
+ Theorem Zdiv_shift_r:
+ forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
+ (r * 2 ^a + t) / (2 ^ b) = (r * 2 ^a) / (2 ^ b).
+ Proof.
+ intros a b r t (H1, H2) H3 (H4, H5).
+ assert (Eq: t < 2 ^ b); auto with zarith.
+ apply Zlt_le_trans with (1 := H5); auto with zarith.
+ apply Zpower_le_monotone; auto with zarith.
+ pattern (r * 2 ^ a) at 1; rewrite Z_div_mod_eq with (b := 2 ^ b);
+ auto with zarith.
+ rewrite <- Zplus_assoc.
+ rewrite <- Zmod_shift_r; auto with zarith.
+ rewrite (Zmult_comm (2 ^ b)); rewrite Z_div_plus_full_l; auto with zarith.
+ rewrite (fun x y => @Zdiv_small (x mod y)); auto with zarith.
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ auto with zarith.
+ Qed.
+
+
+ Lemma shift_unshift_mod : forall n p a,
+ 0 <= a < 2^n ->
+ 0 <= p <= n ->
+ a * 2^p = a / 2^(n - p) * 2^n + (a*2^p) mod 2^n.
+ Proof.
+ intros n p a H1 H2.
+ pattern (a*2^p) at 1;replace (a*2^p) with
+ (a*2^p/2^n * 2^n + a*2^p mod 2^n).
+ 2:symmetry;rewrite (Zmult_comm (a*2^p/2^n));apply Z_div_mod_eq.
+ replace (a * 2 ^ p / 2 ^ n) with (a / 2 ^ (n - p));trivial.
+ replace (2^n) with (2^(n-p)*2^p).
+ symmetry;apply Zdiv_mult_cancel_r.
+ destruct H1;trivial.
+ cut (0 < 2^p); auto with zarith.
+ rewrite <- Zpower_exp.
+ replace (n-p+p) with n;trivial. ring.
+ omega. omega.
+ apply Zlt_gt. apply Zpower_gt_0;auto with zarith.
+ Qed.
-Theorem Zmod_mod: forall a n, 0 < n -> (a mod n) mod n = a mod n.
-intros a n H.
-pattern a at 2; rewrite (Z_div_mod_eq a n); auto with zarith.
-rewrite Zplus_comm; rewrite Zmult_comm.
-apply sym_equal; apply Z_mod_plus; auto with zarith.
-Qed.
+ Lemma div_le_0 : forall p x, 0 <= x -> 0 <= x / 2 ^ p.
+ Proof.
+ intros p x Hle;destruct (Z_le_gt_dec 0 p).
+ apply Zdiv_le_lower_bound;auto with zarith.
+ replace (2^p) with 0.
+ destruct x;compute;intro;discriminate.
+ destruct p;trivial;discriminate z.
+ Qed.
-Theorem Zmod_def_small: forall a n, 0 <= a < n -> a mod n = a.
-intros a n [H1 H2]; unfold Zmod.
-generalize (Z_div_mod a n); case (Zdiv_eucl a n).
-intros q r H3; case H3; clear H3; auto with zarith.
-auto with zarith.
-intros H4 [H5 H6].
-case (Zle_or_lt q (- 1)); intros H7.
-contradict H1; apply Zlt_not_le.
-subst a.
-apply Zle_lt_trans with (n * - 1 + r); auto with zarith.
-case (Zle_lt_or_eq 0 q); auto with zarith; intros H8.
-contradict H2; apply Zle_not_lt.
-apply Zle_trans with (n * 1 + r); auto with zarith.
-rewrite H4; auto with zarith.
-subst a; subst q; auto with zarith.
-Qed.
-
-Theorem Zmod_minus: forall a b n, 0 < n -> (a - b) mod n = (a mod n - b mod n) mod n.
-intros a b n H; replace (a - b) with (a + (-1) * b); auto with zarith.
-replace (a mod n - b mod n) with (a mod n + (-1) * (b mod n)); auto with zarith.
-rewrite Zmod_plus; auto with zarith.
-rewrite Zmod_mult; auto with zarith.
-rewrite (fun x y => Zmod_plus x ((-1) * y)); auto with zarith.
-rewrite Zmod_mult; auto with zarith.
-rewrite (fun x => Zmod_mult x (b mod n)); auto with zarith.
-repeat rewrite Zmod_mod; auto.
-Qed.
-
-Theorem Zmod_Zpower: forall p q n, 0 < n -> (p ^ q) mod n = ((p mod n) ^ q) mod n.
-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_exp_1; auto with zarith.
-rewrite (fun x => (Zmod_mult x p)); try rewrite Rec; auto.
-rewrite (fun x y => (Zmod_mult (x ^y))); try eq_tac; auto.
-eq_tac; auto; apply sym_equal; apply Zmod_mod; auto with zarith.
-generalize H1; case q; simpl; auto.
-intros; discriminate.
-Qed.
-
-Theorem Zmod_le: forall a n, 0 < n -> 0 <= a -> (Zmod a n) <= a.
-intros a n H1 H2; case (Zle_or_lt n a); intros H3.
-case (Z_mod_lt a n); auto with zarith.
-rewrite Zmod_def_small; auto with zarith.
-Qed.
-
-Lemma Zplus_mod_idemp_l: forall a b n, 0 < n -> (a mod n + b) mod n = (a + b) mod n.
-Proof.
-intros. rewrite Zmod_plus; auto.
-rewrite Zmod_mod; auto.
-symmetry; apply Zmod_plus; auto.
-Qed.
-
-Lemma Zplus_mod_idemp_r: forall a b n, 0 < n -> (b + a mod n) mod n = (b + a) mod n.
-Proof.
-intros a b n H; repeat rewrite (Zplus_comm b).
-apply Zplus_mod_idemp_l; auto.
-Qed.
-
-Lemma Zminus_mod_idemp_l: forall a b n, 0 < n -> (a mod n - b) mod n = (a - b) mod n.
-Proof.
-intros. rewrite Zmod_minus; auto.
-rewrite Zmod_mod; auto.
-symmetry; apply Zmod_minus; auto.
-Qed.
-
-Lemma Zminus_mod_idemp_r: forall a b n, 0 < n -> (a - b mod n) mod n = (a - b) mod n.
-Proof.
-intros. rewrite Zmod_minus; auto.
-rewrite Zmod_mod; auto.
-symmetry; apply Zmod_minus; auto.
-Qed.
-
-Lemma Zmult_mod_idemp_l: forall a b n, 0 < n -> (a mod n * b) mod n = (a * b) mod n.
-Proof.
-intros; rewrite Zmod_mult; auto.
-rewrite Zmod_mod; auto.
-symmetry; apply Zmod_mult; auto.
-Qed.
-
-Lemma Zmult_mod_idemp_r: forall a b n, 0 < n -> (b * (a mod n)) mod n = (b * a) mod n.
-Proof.
-intros a b n H; repeat rewrite (Zmult_comm b).
-apply Zmult_mod_idemp_l; auto.
-Qed.
-
-Lemma Zmod_div_mod: forall 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 Zmod_plus; auto.
-rewrite <- Zmult_assoc; rewrite Zmod_mult; auto.
-rewrite Z_mod_same; try rewrite Zmult_0_l; auto with zarith.
-rewrite (Zmod_def_small 0); auto with zarith.
-rewrite Zplus_0_l; rewrite Zmod_mod; auto with zarith.
-Qed.
-
-(** A better way to compute Zpower mod **)
-
-Fixpoint Zpow_mod_pos (a: Z) (m: positive) (n : Z) {struct m} : Z :=
- match m with
- | xH => a mod n
- | xO m' =>
- let z := Zpow_mod_pos a m' n in
- match z with
- | 0 => 0
- | _ => (z * z) mod n
- end
- | xI m' =>
- let z := Zpow_mod_pos a m' n in
- match z with
- | 0 => 0
- | _ => (z * z * a) mod n
- end
- end.
-
-Theorem Zpow_mod_pos_Zpower_pos_correct: forall a m n, 0 < n -> Zpow_mod_pos a m n = (Zpower_pos a m) mod n.
-intros a m; elim m; simpl; auto.
-intros p Rec n H1; rewrite xI_succ_xO; rewrite Pplus_one_succ_r; rewrite <- Pplus_diag; auto.
-repeat rewrite Zpower_pos_is_exp; auto.
-repeat rewrite Rec; auto.
-replace (Zpower_pos a 1) with a; auto.
-2: unfold Zpower_pos; simpl; auto with zarith.
-repeat rewrite (fun x => (Zmod_mult x a)); auto.
-rewrite (Zmod_mult (Zpower_pos a p)); auto.
-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 (Zmod_mult (Zpower_pos a p)); auto.
-case (Zpower_pos a p mod n); auto.
-unfold Zpower_pos; simpl; rewrite Zmult_1_r; auto with zarith.
-Qed.
-
-Definition Zpow_mod a m n := match m with 0 => 1 | Zpos p1 => Zpow_mod_pos a p1 n | Zneg p1 => 0 end.
-
-Theorem Zpow_mod_Zpower_correct: forall a m n, 1 < n -> 0 <= m -> Zpow_mod a m n = (a ^ m) mod n.
-intros a m n; case m; simpl.
-intros; apply sym_equal; apply Zmod_def_small; auto with zarith.
-intros; apply Zpow_mod_pos_Zpower_pos_correct; auto with zarith.
-intros p H H1; case H1; auto.
-Qed.
-
-(* A direct way to compute Zmod *)
-
-Fixpoint Zmod_POS (a : positive) (b : Z) {struct a} : Z :=
- match a with
- | 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)
- | 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
- end.
-
-Theorem Zmod_POS_correct: forall a b, Zmod_POS a b = (snd (Zdiv_eucl_POS a b)).
-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.
-Qed.
-
-Definition Zmodd a b :=
-match a with
-| Z0 => 0
-| Zpos a' =>
- match b with
- | Z0 => 0
- | Zpos _ => Zmod_POS a' b
- | Zneg b' =>
- let r := Zmod_POS a' (Zpos b') in
- match r with Z0 => 0 | _ => b + r end
- end
-| Zneg a' =>
- match b with
- | Z0 => 0
- | Zpos _ =>
- let r := Zmod_POS a' b in
- match r with Z0 => 0 | _ => b - r end
- | Zneg b' => - (Zmod_POS a' (Zpos b'))
- end
-end.
-
-Theorem Zmodd_correct: forall a b, Zmodd a b = Zmod a b.
-intros a b; unfold Zmod; 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.
-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.
-intros p1; rewrite Zmod_POS_correct; simpl; auto.
-case (Zdiv_eucl_POS p (Zpos p1)); auto.
-Qed.
-
-(**************************************
- Properties of Zdivide
-**************************************)
-
-Theorem Zmod_divide_minus: forall a b c : Z,
- 0 < b -> a mod b = c -> (b | a - c).
-Proof.
- intros a b c H H1; apply Zmod_divide; auto with zarith.
- rewrite Zmod_minus; auto.
- rewrite H1; pattern c at 1; rewrite <- (Zmod_def_small c b); auto with zarith.
- rewrite Zminus_diag; apply Zmod_def_small; auto with zarith.
- subst; apply Z_mod_lt; auto with zarith.
-Qed.
-
-Theorem Zdivide_mod_minus: forall a b c : Z,
- 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.
- replace a with ((a - c) + c); auto with zarith.
- rewrite Zmod_plus; auto with zarith.
- rewrite (Zdivide_mod (a -c) b); try rewrite Zplus_0_l; auto with zarith.
- rewrite Zmod_mod; try apply Zmod_def_small; auto with zarith.
-Qed.
-
-Theorem Zmod_closeby_eq: forall a b n, 0 <= a -> 0 <= b < n -> a - b < n -> a mod n = b -> a = b.
-Proof.
- intros a b n H H1 H2 H3.
- case (Zle_or_lt 0 (a - b)); intros H4.
- case Zle_lt_or_eq with (1 := H4); clear H4; intros H4; auto with zarith.
- absurd_hyp H2; auto.
- apply Zle_not_lt; apply Zdivide_le; auto with zarith.
- apply Zmod_divide_minus; auto with zarith.
- rewrite <- (Zmod_def_small a n); try split; auto with zarith.
-Qed.
-
-Theorem Zpower_divide: forall p q, 0 < q -> (p | p ^ q).
-Proof.
- intros p q H; exists (p ^(q - 1)).
- pattern p at 3; rewrite <- (Zpower_exp_1 p); rewrite <- Zpower_exp; try eq_tac; auto with zarith.
-Qed.
-
-
-(**************************************
- Properties of Zis_gcd
-**************************************)
-
-(* P.L. : See Numtheory.v *)
+ Lemma div_lt : forall p x y, 0 <= x < y -> x / 2^p < y.
+ Proof.
+ intros p x y H;destruct (Z_le_gt_dec 0 p).
+ apply Zdiv_lt_upper_bound;auto with zarith.
+ apply Zlt_le_trans with y;auto with zarith.
+ rewrite <- (Zmult_1_r y);apply Zmult_le_compat;auto with zarith.
+ assert (0 < 2^p);auto with zarith.
+ replace (2^p) with 0.
+ destruct x;change (0<y);auto with zarith.
+ destruct p;trivial;discriminate z.
+ Qed.
-(**************************************
- Properties rel_prime
-**************************************)
-
-Theorem rel_prime_sym: forall a b, rel_prime a b -> rel_prime b a.
-intros a b H; auto with zarith.
-red; apply Zis_gcd_sym; auto with zarith.
-Qed.
-
-Theorem rel_prime_le_prime:
- forall a p, prime p -> 1 <= a < p -> rel_prime a p.
-intros a p Hp [H1 H2].
-apply rel_prime_sym; apply prime_rel_prime; auto.
-intros [q Hq]; subst a.
-case (Zle_or_lt q 0); intros Hl.
-absurd (q * p <= 0 * p); auto with zarith.
-absurd (1 * p <= q * p); auto with zarith.
-Qed.
+ Theorem Zgcd_div_pos a b:
+ (0 < b)%Z -> (0 < Zgcd a b)%Z -> (0 < b / Zgcd a b)%Z.
+ Proof.
+ intros a b Ha Hg.
+ case (Zle_lt_or_eq 0 (b/Zgcd a b)); auto.
+ apply Z_div_pos; auto with zarith.
+ intros H; generalize Ha.
+ pattern b at 1; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto.
+ rewrite <- H; auto with zarith.
+ assert (F := (Zgcd_is_gcd a b)); inversion F; auto.
+ Qed.
-Definition rel_prime_dec:
- forall a b, ({ rel_prime a b }) + ({ ~ rel_prime a b }).
-intros a b; case (Z_eq_dec (Zgcd a b) 1); intros H1.
-left; red.
-rewrite <- H1; apply Zgcd_is_gcd.
-right; contradict H1.
-case (Zis_gcd_unique a b (Zgcd a b) 1); auto.
-apply Zgcd_is_gcd.
-intros H2; absurd (0 <= Zgcd a b); auto with zarith.
-generalize (Zgcd_is_pos a b); auto with zarith.
-Qed.
+(* For compatibility of scripts, weaker version of some lemmas of Zdiv *)
-Theorem rel_prime_mod_rev: forall p q, 0 < q -> rel_prime (p mod q) q -> rel_prime p q.
-intros p q H H0.
-rewrite (Z_div_mod_eq p q); auto with zarith.
-red.
-apply Zis_gcd_sym; apply Zis_gcd_for_euclid2; auto with zarith.
-Qed.
-
-Theorem rel_prime_div: forall p q r, rel_prime p q -> (r | p) -> rel_prime r q.
-intros p q r H (u, H1); subst.
-inversion_clear H as [H1 H2 H3].
-red; apply Zis_gcd_intro; try apply Zone_divide.
-intros x H4 H5; apply H3; auto.
-apply Zdivide_mult_r; auto.
-Qed.
-
-Theorem rel_prime_1: forall n, rel_prime 1 n.
-intros n; red; apply Zis_gcd_intro; auto.
-exists 1; auto with zarith.
-exists n; auto with zarith.
-Qed.
-
-Theorem not_rel_prime_0: forall n, 1 < n -> ~rel_prime 0 n.
-intros n H H1; absurd (n = 1 \/ n = -1).
-intros [H2 | H2]; subst; contradict H; auto with zarith.
-case (Zis_gcd_unique 0 n n 1); auto.
-apply Zis_gcd_intro; auto.
-exists 0; auto with zarith.
-exists 1; auto with zarith.
-Qed.
-
-Theorem rel_prime_mod: forall p q, 0 < q -> rel_prime p q -> rel_prime (p mod q) q.
-intros p q H H0.
-assert (H1: Bezout p q 1).
-apply rel_prime_bezout; auto.
-inversion_clear H1 as [q1 r1 H2].
-apply bezout_rel_prime.
-apply Bezout_intro with q1 (r1 + q1 * (p / q)).
-rewrite <- H2.
-pattern p at 3; rewrite (Z_div_mod_eq p q); try ring; auto with zarith.
-Qed.
-
-Theorem Zrel_prime_neq_mod_0: forall a b, 1 < b -> rel_prime a b -> a mod b <> 0.
+Lemma Zlt0_not_eq : forall n, 0<n -> n<>0.
Proof.
-intros a b H H1 H2.
-case (not_rel_prime_0 _ H).
-rewrite <- H2.
-apply rel_prime_mod; auto with zarith.
-Qed.
-
-Theorem rel_prime_Zpower_r: forall i p q, 0 < i -> rel_prime p q -> rel_prime p (q^i).
-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.
-apply rel_prime_mult; auto.
-case Zle_lt_or_eq with (1 := Hi); intros Hi1; subst; auto.
-rewrite Zpower_exp_0; apply rel_prime_sym; apply rel_prime_1.
-Qed.
-
-
-(**************************************
- Properties prime
-**************************************)
-
-Theorem not_prime_0: ~ prime 0.
-intros H1; case (prime_divisors _ H1 2); auto with zarith.
-Qed.
-
-
-Theorem not_prime_1: ~ prime 1.
-intros H1; absurd (1 < 1); auto with zarith.
-inversion H1; auto.
-Qed.
-
-Theorem prime_2: prime 2.
-apply prime_intro; auto with zarith.
-intros n [H1 H2]; case Zle_lt_or_eq with ( 1 := H1 ); auto with zarith;
- clear H1; intros H1.
-contradict H2; auto with zarith.
-subst n; red; auto with zarith.
-apply Zis_gcd_intro; auto with zarith.
-Qed.
-
-Theorem prime_3: prime 3.
-apply prime_intro; auto with zarith.
-intros n [H1 H2]; case Zle_lt_or_eq with ( 1 := H1 ); auto with zarith;
- clear H1; intros H1.
-case (Zle_lt_or_eq 2 n); auto with zarith; clear H1; intros H1.
-contradict H2; auto with zarith.
-subst n; red; auto with zarith.
-apply Zis_gcd_intro; auto with zarith.
-intros x [q1 Hq1] [q2 Hq2].
-exists (q2 - q1).
-apply trans_equal with (3 - 2); auto with zarith.
-rewrite Hq1; rewrite Hq2; ring.
-subst n; red; auto with zarith.
-apply Zis_gcd_intro; auto with zarith.
-Qed.
-
-Theorem prime_le_2: forall p, prime p -> 2 <= p.
-intros p Hp; inversion Hp; auto with zarith.
-Qed.
-
-Definition prime_dec_aux:
- forall p m,
- ({ forall n, 1 < n < m -> rel_prime n p }) +
- ({ exists n , 1 < n < m /\ ~ rel_prime n p }).
-intros p m.
-case (Z_lt_dec 1 m); intros H1.
-apply natlike_rec
- with
- ( P :=
- fun m =>
- ({ forall (n : Z), 1 < n < m -> rel_prime n p }) +
- ({ exists n : Z , 1 < n < m /\ ~ rel_prime n p }) );
auto with zarith.
-left; intros n [HH1 HH2]; contradict HH2; auto with zarith.
-intros x Hx Rec; case Rec.
-intros P1; case (rel_prime_dec x p); intros P2.
-left; intros n [HH1 HH2].
-case (Zgt_succ_gt_or_eq x n); auto with zarith.
-intros HH3; subst x; auto.
-case (Z_lt_dec 1 x); intros HH1.
-right; exists x; split; auto with zarith.
-left; intros n [HHH1 HHH2]; contradict HHH1; auto with zarith.
-intros tmp; right; case tmp; intros n [HH1 HH2]; exists n; auto with zarith.
-left; intros n [HH1 HH2]; contradict H1; auto with zarith.
-Defined.
-
-Theorem not_prime_divide:
- forall p, 1 < p -> ~ prime p -> exists n, 1 < n < p /\ (n | p) .
-intros p Hp Hp1.
-case (prime_dec_aux p p); intros H1.
-case Hp1; apply prime_intro; auto.
-intros n [Hn1 Hn2].
-case Zle_lt_or_eq with ( 1 := Hn1 ); auto with zarith.
-intros H2; subst n; red; apply Zis_gcd_intro; auto with zarith.
-case H1; intros n [Hn1 Hn2].
-generalize (Zgcd_is_pos n p); intros Hpos.
-case (Zle_lt_or_eq 0 (Zgcd n p)); auto with zarith; intros H3.
-case (Zle_lt_or_eq 1 (Zgcd n p)); auto with zarith; intros H4.
-exists (Zgcd n p); split; auto.
-split; auto.
-apply Zle_lt_trans with n; auto with zarith.
-generalize (Zgcd_is_gcd n p); intros tmp; inversion_clear tmp as [Hr1 Hr2 Hr3].
-case Hr1; intros q Hq.
-case (Zle_or_lt q 0); auto with zarith; intros Ht.
-absurd (n <= 0 * Zgcd n p) ; auto with zarith.
-pattern n at 1; rewrite Hq; auto with zarith.
-apply Zle_trans with (1 * Zgcd n p); auto with zarith.
-pattern n at 2; rewrite Hq; auto with zarith.
-generalize (Zgcd_is_gcd n p); intros Ht; inversion Ht; auto.
-case Hn2; red.
-rewrite H4; apply Zgcd_is_gcd.
-generalize (Zgcd_is_gcd n p); rewrite <- H3; intros tmp;
- inversion_clear tmp as [Hr1 Hr2 Hr3].
-absurd (n = 0); auto with zarith.
-case Hr1; auto with zarith.
-Defined.
-
-Definition prime_dec: forall p, ({ prime p }) + ({ ~ prime p }).
-intros p; case (Z_lt_dec 1 p); intros H1.
-case (prime_dec_aux p p); intros H2.
-left; apply prime_intro; auto.
-intros n [Hn1 Hn2]; case Zle_lt_or_eq with ( 1 := Hn1 ); auto.
-intros HH; subst n.
-red; apply Zis_gcd_intro; auto with zarith.
-right; intros H3; inversion_clear H3 as [Hp1 Hp2].
-case H2; intros n [Hn1 Hn2]; case Hn2; auto with zarith.
-right; intros H3; inversion_clear H3 as [Hp1 Hp2]; case H1; auto.
-Defined.
-
-
-Theorem prime_def:
- forall p, 1 < p -> (forall n, 1 < n < p -> ~ (n | p)) -> prime p.
-intros p H1 H2.
-apply prime_intro; auto.
-intros n H3.
-red; apply Zis_gcd_intro; auto with zarith.
-intros x H4 H5.
-case (Zle_lt_or_eq 0 (Zabs x)); auto with zarith; intros H6.
-case (Zle_lt_or_eq 1 (Zabs x)); auto with zarith; intros H7.
-case (Zle_lt_or_eq (Zabs x) p); auto with zarith.
-apply Zdivide_le; auto with zarith.
-apply Zdivide_Zabs_inv_l; auto.
-intros H8; case (H2 (Zabs x)); auto.
-apply Zdivide_Zabs_inv_l; auto.
-intros H8; subst p; absurd (Zabs x <= n); auto with zarith.
-apply Zdivide_le; auto with zarith.
-apply Zdivide_Zabs_inv_l; auto.
-rewrite H7; pattern (Zabs x); apply Zabs_intro; auto with zarith.
-absurd (0%Z = p); auto with zarith.
-cut (Zdivide (Zabs x) p).
-intros [q Hq]; subst p; rewrite <- H6; auto with zarith.
-apply Zdivide_Zabs_inv_l; auto.
-Qed.
-
-Theorem prime_inv_def: forall p n, prime p -> 1 < n < p -> ~ (n | p).
-intros p n H1 H2 H3.
-absurd (rel_prime n p); auto.
-unfold rel_prime; intros H4.
-case (Zis_gcd_unique n p n 1); auto with zarith.
-apply Zis_gcd_intro; auto with zarith.
-inversion H1; auto with zarith.
-Qed.
-
-Theorem square_not_prime: forall a, ~ prime (a * a).
-intros a; rewrite (Zabs_square a).
-case (Zle_lt_or_eq 0 (Zabs a)); auto with zarith; intros Hza1.
-case (Zle_lt_or_eq 1 (Zabs a)); auto with zarith; intros Hza2.
-intros Ha; case (prime_inv_def (Zabs a * Zabs a) (Zabs a)); auto.
-split; auto.
-pattern (Zabs a) at 1; replace (Zabs a) with (1 * Zabs a); auto with zarith.
-apply Zmult_lt_compat_r; auto with zarith.
-exists (Zabs a); auto.
-rewrite <- Hza2; simpl; apply not_prime_1.
-rewrite <- Hza1; simpl; apply not_prime_0.
Qed.
-Theorem prime_divide_prime_eq:
- forall p1 p2, prime p1 -> prime p2 -> Zdivide p1 p2 -> p1 = p2.
-intros p1 p2 Hp1 Hp2 Hp3.
-assert (Ha: 1 < p1).
-inversion Hp1; auto.
-assert (Ha1: 1 < p2).
-inversion Hp2; auto.
-case (Zle_lt_or_eq p1 p2); auto with zarith.
-apply Zdivide_le; auto with zarith.
-intros Hp4.
-case (prime_inv_def p2 p1); auto with zarith.
-Qed.
-
-Theorem Zdivide_div_prime_le_square: forall x, 1 < x -> ~prime x -> exists p, prime p /\ (p | x) /\ p * p <= x.
-intros x Hx; generalize Hx; pattern x; apply Z_lt_induction; auto with zarith.
-clear x Hx; intros x Rec H H1.
-case (not_prime_divide x); auto.
-intros x1 ((H2, H3), H4); case (prime_dec x1); intros H5.
-case (Zle_or_lt (x1 * x1) x); intros H6.
-exists x1; auto.
-case H4; clear H4; intros x2 H4; subst.
-assert (Hx2: x2 <= x1).
-case (Zle_or_lt x2 x1); auto; intros H8; contradict H6; apply Zle_not_lt.
-apply Zmult_le_compat_r; auto with zarith.
-case (prime_dec x2); intros H7.
-exists x2; repeat (split; auto with zarith).
-apply Zmult_le_compat_l; auto with zarith.
-apply Zle_trans with 2%Z; try apply prime_le_2; auto with zarith.
-case (Zle_or_lt 0 x2); intros H8.
-case Zle_lt_or_eq with (1 := H8); auto with zarith; clear H8; intros H8; subst; auto with zarith.
-case (Zle_lt_or_eq 1 x2); auto with zarith; clear H8; intros H8; subst; auto with zarith.
-case (Rec x2); try split; auto with zarith.
-intros x3 (H9, (H10, H11)).
-exists x3; repeat (split; auto with zarith).
-contradict H; apply Zle_not_lt; auto with zarith.
-apply Zle_trans with (0 * x1); auto with zarith.
-case (Rec x1); try split; auto with zarith.
-intros x3 (H9, (H10, H11)).
-exists x3; repeat (split; auto with zarith).
-apply Zdivide_trans with x1; auto with zarith.
-Qed.
-
-Theorem prime_div_prime: forall p q, prime p -> prime q -> (p | q) -> p = q.
-intros p q H H1 H2;
-assert (Hp: 0 < p); try apply Zlt_le_trans with 2; try apply prime_le_2; auto with zarith.
-assert (Hq: 0 < q); try apply Zlt_le_trans with 2; try apply prime_le_2; auto with zarith.
-case prime_divisors with (2 := H2); auto.
-intros H4; contradict Hp; subst; auto with zarith.
-intros [H4| [H4 | H4]]; subst; auto.
-contradict H; apply not_prime_1.
-contradict Hp; auto with zarith.
-Qed.
-
-Theorem prime_div_Zpower_prime: forall n p q, 0 <= n -> prime p -> prime q -> (p | q ^ n) -> p = q.
-intros n p q Hp Hq; generalize p q Hq; pattern n; apply natlike_ind; auto; clear n p q Hp Hq.
-intros p q Hp Hq; rewrite Zpower_exp_0.
-intros (r, H); subst.
-case (Zmult_interval p r); auto; try rewrite Zmult_comm.
-rewrite <- H; auto with zarith.
-apply Zlt_le_trans with 2; try apply prime_le_2; auto with zarith.
-rewrite <- H; intros H1 H2; contradict H2; auto with zarith.
-intros n1 H Rec p q Hp Hq; try rewrite Zpower_Zsucc; auto with zarith; intros H1.
-case prime_mult with (2 := H1); auto.
-intros H2; apply prime_div_prime; auto.
-Qed.
-
-Theorem rel_prime_Zpower: forall i j p q, 0 <= i -> 0 <= j -> rel_prime p q -> rel_prime (p^i) (q^j).
-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_exp_0; 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_exp_0; apply rel_prime_sym; apply rel_prime_1.
-Qed.
-
-Theorem prime_induction: forall (P: Z -> Prop), P 0 -> P 1 -> (forall p q, prime p -> P q -> P (p * q)) -> forall p, 0 <= p -> P p.
-intros P H H1 H2 p Hp.
-generalize Hp; pattern p; apply Z_lt_induction; auto; clear p Hp.
-intros p Rec Hp.
-case Zle_lt_or_eq with (1 := Hp); clear Hp; intros Hp; subst; auto.
-case (Zle_lt_or_eq 1 p); auto with zarith; clear Hp; intros Hp; subst; auto.
-case (prime_dec p); intros H3.
-rewrite <- (Zmult_1_r p); apply H2; auto.
- case (Zdivide_div_prime_le_square p); auto.
-intros q (Hq1, ((q2, Hq2), Hq3)); subst.
-case (Zmult_interval q q2).
-rewrite Zmult_comm; apply Zlt_trans with 1; auto with zarith.
-apply Zlt_le_trans with 2; auto with zarith; apply prime_le_2; auto.
-intros H4 H5; rewrite Zmult_comm; apply H2; auto.
-apply Rec; try split; auto with zarith.
-rewrite Zmult_comm; auto.
-Qed.
-
-Theorem div_power_max: forall p q, 1 < p -> 0 < q -> exists n, 0 <= n /\ (p ^n | q) /\ ~(p ^(1 + n) | q).
-intros p q H1 H2; generalize H2; pattern q; apply Z_lt_induction; auto with zarith; clear q H2.
-intros q Rec H2.
-case (Zdivide_dec p q); intros H3.
-case (Zdivide_Zdiv_lt_pos p q); auto with zarith; intros H4 H5.
-case (Rec (Zdiv q p)); auto with zarith.
-intros n (Ha1, (Ha2, Ha3)); exists (n + 1); split; auto with zarith; split.
-case Ha2; intros q1 Hq; exists q1.
-rewrite Zpower_exp; try rewrite Zpower_exp_1; auto with zarith.
-rewrite Zmult_assoc; rewrite <- Hq.
-rewrite Zmult_comm; apply Zdivide_Zdiv_eq; auto with zarith.
-intros (q1, Hu); case Ha3; exists q1.
-apply Zmult_reg_r with p; auto with zarith.
-rewrite (Zmult_comm (q / p)); rewrite <- Zdivide_Zdiv_eq; auto with zarith.
-apply trans_equal with (1 := Hu); repeat rewrite Zpower_exp; try rewrite Zpower_exp_1; auto with zarith.
-ring.
-exists 0; repeat split; try rewrite Zpower_exp_1; try rewrite Zpower_exp_0; auto with zarith.
-Qed.
-
-Theorem prime_divide_Zpower_Zdiv: forall m a p i, 0 <= i -> prime p -> (m | a) -> ~(m | (a/p)) -> (p^i | a) -> (p^i | m).
-intros m a p i Hi Hp (k, Hk) H (l, Hl); subst.
-case (Zle_lt_or_eq 0 i); auto with arith; intros Hi1; subst.
-assert (Hp0: 0 < p).
-apply Zlt_le_trans with 2; auto with zarith; apply prime_le_2; auto.
-case (Zdivide_dec p k); intros H1.
-case H1; intros k' H2; subst.
-case H; replace (k' * p * m) with ((k' * m) * p); try ring; rewrite Z_div_mult; auto with zarith.
-apply Gauss with k.
-exists l; rewrite Hl; ring.
-apply rel_prime_sym; apply rel_prime_Zpower_r; auto.
-apply rel_prime_sym; apply prime_rel_prime; auto.
-rewrite Zpower_exp_0; apply Zone_divide.
-Qed.
-
-
-Theorem Zdivide_Zpower: forall n m, 0 < n -> (forall p i, prime p -> 0 < i -> (p^i | n) -> (p^i | m)) -> (n | m).
-intros n m Hn; generalize m Hn; pattern n; apply prime_induction; auto with zarith; clear n m Hn.
-intros m H1; contradict H1; auto with zarith.
-intros p q H Rec m H1 H2.
-assert (H3: (p | m)).
-rewrite <- (Zpower_exp_1 p); apply H2; auto with zarith; rewrite Zpower_exp_1; apply Zdivide_factor_r.
-case (Zmult_interval p q); auto.
-apply Zlt_le_trans with 2; auto with zarith; apply prime_le_2; auto.
-case H3; intros k Hk; subst.
-intros Hq Hq1.
-rewrite (Zmult_comm k); apply Zmult_divide_compat_l.
-apply Rec; auto.
-intros p1 i Hp1 Hp2 Hp3.
-case (Z_eq_dec p p1); intros Hpp1; subst.
-case (H2 p1 (Zsucc i)); auto with zarith.
-rewrite Zpower_Zsucc; try apply Zmult_divide_compat_l; auto with zarith.
-intros q2 Hq2; exists q2.
-apply Zmult_reg_r with p1.
-contradict H; subst; apply not_prime_0.
-rewrite Hq2; rewrite Zpower_Zsucc; try ring; auto with zarith.
-apply Gauss with p.
-rewrite Zmult_comm; apply H2; auto.
-apply Zdivide_trans with (1:= Hp3).
-apply Zdivide_factor_l.
-apply rel_prime_sym; apply rel_prime_Zpower_r; auto.
-apply prime_rel_prime; auto.
-contradict Hpp1; apply prime_divide_prime_eq; auto.
-Qed.
-
-Theorem divide_prime_divide:
- forall a n m, 0 < a -> (n | m) -> (a | m) ->
- (forall p, prime p -> (p | a) -> ~(n | (m/p))) ->
- (a | n).
-intros a n m Ha Hnm Ham Hp.
-apply Zdivide_Zpower; auto.
-intros p i H1 H2 H3.
-apply prime_divide_Zpower_Zdiv with m; auto with zarith.
-apply Hp; auto; apply Zdivide_trans with (2 := H3); auto.
-apply Zpower_divide; auto.
-apply Zdivide_trans with (1 := H3); auto.
-Qed.
-
-Theorem prime_div_induction:
- forall (P: Z -> Prop) n,
- 0 < n ->
- (P 1) ->
- (forall p i, prime p -> 0 <= i -> (p^i | n) -> P (p^i)) ->
- (forall p q, rel_prime p q -> P p -> P q -> P (p * q)) ->
- forall m, 0 <= m -> (m | n) -> P m.
-intros P n P1 Hn H H1 m Hm.
-generalize Hm; pattern m; apply Z_lt_induction; auto; clear m Hm.
-intros m Rec Hm H2.
-case (prime_dec m); intros Hm1.
-rewrite <- Zpower_exp_1; apply H; auto with zarith.
-rewrite Zpower_exp_1; auto.
-case Zle_lt_or_eq with (1 := Hm); clear Hm; intros Hm; subst.
-2: contradict P1; case H2; intros; subst; auto with zarith.
-case (Zle_lt_or_eq 1 m); auto with zarith; clear Hm; intros Hm; subst; auto.
-case Zdivide_div_prime_le_square with m; auto.
-intros p (Hp1, (Hp2, Hp3)).
-case (div_power_max p m); auto with zarith.
-generalize (prime_le_2 p Hp1); auto with zarith.
-intros i (Hi, (Hi1, Hi2)).
-case Zle_lt_or_eq with (1 := Hi); clear Hi; intros Hi.
-assert (Hpi: 0 < p ^ i).
-apply Zpower_lt_0; auto with zarith.
-apply Zlt_le_trans with 2; try apply prime_le_2; auto with zarith.
-rewrite (Z_div_exact_2 m (p ^ i)); auto with zarith.
-apply H1; auto with zarith.
-apply rel_prime_sym; apply rel_prime_Zpower_r; auto.
-apply rel_prime_sym.
-apply prime_rel_prime; auto.
-contradict Hi2.
-case Hi1; intros; subst.
-rewrite Z_div_mult in Hi2; auto with zarith.
-case Hi2; intros q0 Hq0; subst.
-exists q0; rewrite Zpower_exp; try rewrite Zpower_exp_1; auto with zarith.
-apply H; auto with zarith.
-apply Zdivide_trans with (1 := Hi1); auto.
-apply Rec; auto with zarith.
-split; auto with zarith.
-apply Zge_le; apply Z_div_ge0; auto with zarith.
-apply Z_div_lt; auto with zarith.
-apply Zle_ge; apply Zle_trans with p.
-apply prime_le_2; auto.
-pattern p at 1; rewrite <- Zpower_exp_1; apply Zpower_le_monotone; auto with zarith.
-apply Zlt_le_trans with 2; try apply prime_le_2; auto with zarith.
-apply Zge_le; apply Z_div_ge0; auto with zarith.
-apply Zdivide_trans with (2 := H2); auto.
-exists (p ^ i); apply Z_div_exact_2; auto with zarith.
-apply Zdivide_mod; auto with zarith.
-apply Zdivide_mod; auto with zarith.
-case Hi2; rewrite <- Hi; rewrite Zplus_0_r; rewrite Zpower_exp_1; auto.
-Qed.
-
-(**************************************
- A tail recursive way of compute a^n
-**************************************)
-
-Fixpoint Zpower_tr_aux (z1 z2: Z) (n: nat) {struct n}: Z :=
- match n with O => z1 | (S n1) => Zpower_tr_aux (z2 * z1) z2 n1 end.
-
-Theorem Zpower_tr_aux_correct:
-forall z1 z2 n p, z1 = Zpower_nat z2 p -> Zpower_tr_aux z1 z2 n = Zpower_nat z2 (p + n).
-intros z1 z2 n; generalize z1; elim n; clear z1 n; simpl; auto.
-intros z1 p; rewrite plus_0_r; auto.
-intros n1 Rec z1 p H1.
-rewrite Rec with (p:= S p).
-rewrite <- plus_n_Sm; simpl; auto.
-pattern z2 at 1; replace z2 with (Zpower_nat z2 1).
-rewrite H1; rewrite <- Zpower_nat_is_exp; simpl; auto.
-unfold Zpower_nat; simpl; rewrite Zmult_1_r; auto.
-Qed.
-
-Definition Zpower_nat_tr := Zpower_tr_aux 1.
-
-Theorem Zpower_nat_tr_correct:
-forall z n, Zpower_nat_tr z n = Zpower_nat z n.
-intros z n; unfold Zpower_nat_tr.
-rewrite Zpower_tr_aux_correct with (p := 0%nat); auto.
-Qed.
-
-
-(**************************************
- Definition of Zsquare
-**************************************)
-
-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.
-
-Theorem Psquare_correct: (forall p, Psquare p = p * p)%positive.
-intros p; elim p; simpl; auto.
-intros p1 Rec; rewrite Rec.
-eq_tac.
-apply trans_equal with (xO p1 + xO (p1 * p1) )%positive; auto.
-rewrite (Pplus_comm (xO p1)); auto.
-rewrite Pmult_xI_permute_r; rewrite Pplus_assoc.
-eq_tac; auto.
-apply sym_equal; apply Pplus_diag.
-intros p1 Rec; rewrite Rec; simpl; auto.
-eq_tac; auto.
-apply sym_equal; apply Pmult_xO_permute_r.
-Qed.
-
-Definition Zsquare p :=
-match p with Z0 => Z0 | Zpos p => Zpos (Psquare p) | Zneg p => Zpos (Psquare p) end.
-
-Theorem Zsquare_correct: forall p, Zsquare p = p * p.
-intro p; case p; simpl; auto; intros p1; rewrite Psquare_correct; auto.
-Qed.
-
-(**************************************
- Some properties of Zpower
-**************************************)
-
-Theorem prime_power_2: forall x n, 0 <= n -> prime x -> (x | 2 ^ n) -> x = 2.
-intros x n H Hx; pattern n; apply natlike_ind; auto; clear n H.
-rewrite Zpower_exp_0.
-intros H1; absurd (x <= 1).
-apply Zlt_not_le; apply Zlt_le_trans with 2%Z; auto with zarith.
-apply prime_le_2; auto.
-apply Zdivide_le; auto with zarith.
-apply Zle_trans with 2%Z; try apply prime_le_2; auto with zarith.
-intros n1 H H1.
-unfold Zsucc; rewrite Zpower_exp; try rewrite Zpower_exp_1; auto with zarith.
-intros H2; case prime_mult with (2 := H2); auto.
-intros H3; case (Zle_lt_or_eq x 2); auto.
-apply Zdivide_le; auto with zarith.
-apply Zle_trans with 2%Z; try apply prime_le_2; auto with zarith.
-intros H4; contradict H4; apply Zle_not_lt.
-apply prime_le_2; auto with zarith.
-Qed.
-
-Theorem Zdivide_power_2: forall x n, 0 <= n -> 0 <= x -> (x | 2 ^ n) -> exists q, x = 2 ^ q.
-intros x n Hn H; generalize n H Hn; pattern x; apply Z_lt_induction; auto; clear x n H Hn.
-intros x Rec n H Hn H1.
-case Zle_lt_or_eq with (1 := H); auto; clear H; intros H; subst.
-case (Zle_lt_or_eq 1 x); auto with zarith; clear H; intros H; subst.
-case (prime_dec x); intros H2.
-exists 1; simpl; apply prime_power_2 with n; auto.
-case not_prime_divide with (2 := H2); auto.
-intros p1 ((H3, H4), (q1, Hq1)); subst.
-case (Rec p1) with n; auto with zarith.
-apply Zdivide_trans with (2 := H1); exists q1; auto with zarith.
-intros r1 Hr1.
-case (Rec q1) with n; auto with zarith.
-case (Zle_lt_or_eq 0 q1).
-apply Zmult_le_0_reg_r with p1; 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 H; auto with zarith.
-apply Zmult_le_0_reg_r with p1; auto with zarith.
-apply Zdivide_trans with (2 := H1); exists p1; auto with zarith.
-intros r2 Hr2; exists (r2 + r1); subst.
-apply sym_equal; apply Zpower_exp.
-generalize H; case r2; simpl; auto with zarith.
-intros; red; simpl; intros; discriminate.
-generalize H; case r1; simpl; auto with zarith.
-intros; red; simpl; intros; discriminate.
-exists 0; simpl; auto.
-case H1; intros q1; try rewrite Zmult_0_r; intros H2.
-absurd (0 < 0); auto with zarith.
-pattern 0 at 2; rewrite <- H2; auto with zarith.
-apply Zpower_lt_0; auto with zarith.
-Qed.
-
-
-(**************************************
- Some properties of Zodd and Zeven
-**************************************)
-
-Theorem Zeven_ex: forall p, Zeven p -> exists q, p = 2 * q.
-intros p; case p; simpl; auto.
-intros _; exists 0; auto.
-intros p1; case p1; try ((intros H; case H; fail) || intros z H; case H; fail).
-intros p2 _; exists (Zpos p2); auto.
-intros p1; case p1; try ((intros H; case H; fail) || intros z H; case H; fail).
-intros p2 _; exists (Zneg p2); auto.
-Qed.
-
-Theorem Zodd_ex: forall p, Zodd p -> exists q, p = 2 * q + 1.
-intros p HH; case (Zle_or_lt 0 p); intros HH1.
-exists (Zdiv2 p); apply Zodd_div2; auto with zarith.
-exists ((Zdiv2 p) - 1); pattern p at 1; rewrite Zodd_div2_neg; auto with zarith.
-Qed.
-
-Theorem Zeven_2p: forall p, Zeven (2 * p).
-intros p; case p; simpl; auto.
-Qed.
-
-Theorem Zodd_2p_plus_1: forall p, Zodd (2 * p + 1).
-intros p; case p; simpl; auto.
-intros p1; case p1; simpl; auto.
-Qed.
-
-Theorem Zeven_plus_Zodd_Zodd: forall z1 z2, Zeven z1 -> Zodd z2 -> Zodd (z1 + z2).
-intros z1 z2 HH1 HH2; case Zeven_ex with (1 := HH1); intros x HH3; try rewrite HH3; auto.
-case Zodd_ex with (1 := HH2); intros y HH4; try rewrite HH4; auto.
-replace (2 * x + (2 * y + 1)) with (2 * (x + y) + 1); try apply Zodd_2p_plus_1; auto with zarith.
-Qed.
-
-Theorem Zeven_plus_Zeven_Zeven: forall z1 z2, Zeven z1 -> Zeven z2 -> Zeven (z1 + z2).
-intros z1 z2 HH1 HH2; case Zeven_ex with (1 := HH1); intros x HH3; try rewrite HH3; auto.
-case Zeven_ex with (1 := HH2); intros y HH4; try rewrite HH4; auto.
-replace (2 * x + 2 * y) with (2 * (x + y)); try apply Zeven_2p; auto with zarith.
-Qed.
-
-Theorem Zodd_plus_Zeven_Zodd: forall z1 z2, Zodd z1 -> Zeven z2 -> Zodd (z1 + z2).
-intros z1 z2 HH1 HH2; rewrite Zplus_comm; apply Zeven_plus_Zodd_Zodd; auto.
-Qed.
-
-Theorem Zodd_plus_Zodd_Zeven: forall z1 z2, Zodd z1 -> Zodd z2 -> Zeven (z1 + z2).
-intros z1 z2 HH1 HH2; case Zodd_ex with (1 := HH1); intros x HH3; try rewrite HH3; auto.
-case Zodd_ex with (1 := HH2); intros y HH4; try rewrite HH4; auto.
-replace ((2 * x + 1) + (2 * y + 1)) with (2 * (x + y + 1)); try apply Zeven_2p; try ring.
-Qed.
-
-Theorem Zeven_mult_Zeven_l: forall z1 z2, Zeven z1 -> Zeven (z1 * z2).
-intros z1 z2 HH1; case Zeven_ex with (1 := HH1); intros x HH3; try rewrite HH3; auto.
-replace (2 * x * z2) with (2 * (x * z2)); try apply Zeven_2p; auto with zarith.
-Qed.
-
-Theorem Zeven_mult_Zeven_r: forall z1 z2, Zeven z2 -> Zeven (z1 * z2).
-intros z1 z2 HH1; case Zeven_ex with (1 := HH1); intros x HH3; try rewrite HH3; auto.
-replace (z1 * (2 * x)) with (2 * (x * z1)); try apply Zeven_2p; try ring.
-Qed.
-
-Theorem Zodd_mult_Zodd_Zodd: forall z1 z2, Zodd z1 -> Zodd z2 -> Zodd (z1 * z2).
-intros z1 z2 HH1 HH2; case Zodd_ex with (1 := HH1); intros x HH3; try rewrite HH3; auto.
-case Zodd_ex with (1 := HH2); intros y HH4; try rewrite HH4; auto.
-replace ((2 * x + 1) * (2 * y + 1)) with (2 * (2 * x * y + x + y) + 1); try apply Zodd_2p_plus_1; try ring.
-Qed.
-
-Definition Zmult_lt_0_compat := Zmult_lt_O_compat.
-
-Hint Rewrite Zmult_1_r Zmult_0_r Zmult_1_l Zmult_0_l Zplus_0_l Zplus_0_r Zminus_0_r: rm10.
-Hint Rewrite Zmult_plus_distr_r Zmult_plus_distr_l Zmult_minus_distr_r Zmult_minus_distr_l: distr.
-
-Theorem Zmult_lt_compat_bis:
- forall n m p q : Z, 0 <= n < p -> 0 <= m < q -> n * m < p * q.
-intros n m p q (H1, H2) (H3,H4).
-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; autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith.
-rewrite <- H5; autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith.
-Qed.
-
-
-Theorem nat_of_P_xO:
- forall p, nat_of_P (xO p) = (2 * nat_of_P p)%nat.
-intros p; unfold nat_of_P; simpl; rewrite Pmult_nat_2_mult_2_permute; auto with arith.
-Qed.
-
-Theorem nat_of_P_xI:
- forall p, nat_of_P (xI p) = (2 * nat_of_P p + 1)%nat.
-intros p; unfold nat_of_P; simpl; rewrite Pmult_nat_2_mult_2_permute; auto with arith.
-ring.
-Qed.
-
-Theorem nat_of_P_xH: nat_of_P xH = 1%nat.
-trivial.
-Qed.
-
-Hint Rewrite
- nat_of_P_xO nat_of_P_xI nat_of_P_xH
- nat_of_P_succ_morphism
- nat_of_P_plus_carry_morphism
- nat_of_P_plus_morphism
- nat_of_P_mult_morphism
- nat_of_P_minus_morphism: pos_morph.
-
-Ltac pos_tac :=
- match goal with |- ?X = ?Y =>
- assert (tmp: Zpos X = Zpos Y);
- [idtac; repeat rewrite Zpos_eq_Z_of_nat_o_nat_of_P; eq_tac | injection tmp; auto]
- end; autorewrite with pos_morph.
-
-(**************************************
- Bounded induction
-**************************************)
+Definition Zdiv_mult_cancel_r a b c H := Zdiv.Zdiv_mult_cancel_r a b c (Zlt0_not_eq _ H).
+Definition Zdiv_mult_cancel_l a b c H := Zdiv.Zdiv_mult_cancel_r a b c (Zlt0_not_eq _ H).
+Definition Z_div_plus_l a b c H := Zdiv.Z_div_plus_full_l a b c (Zlt0_not_eq _ H).
Theorem Zbounded_induction :
(forall Q : Z -> Prop, forall b : Z,
@@ -1392,4 +322,3 @@ right; auto with zarith.
unfold Q' in *; intros n H1 H2. destruct (H n H1) as [[H3 H4] | H3].
assumption. apply Zle_not_lt in H3. false_hyp H2 H3.
Qed.
-
diff --git a/theories/Ints/Z/ZDivModAux.v b/theories/Ints/Z/ZDivModAux.v
deleted file mode 100644
index 127d3cefa..000000000
--- a/theories/Ints/Z/ZDivModAux.v
+++ /dev/null
@@ -1,479 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-(**********************************************************************
- ZDivModAux.v
-
- Auxillary functions & Theorems for Zdiv and Zmod
- **********************************************************************)
-
-Require Export ZArith.
-Require Export Znumtheory.
-Require Export Tactic.
-Require Import ZAux.
-Require Import ZPowerAux.
-
-Open Local Scope Z_scope.
-
-Hint Extern 2 (Zle _ _) =>
- (match goal with
- |- Zpos _ <= Zpos _ => exact (refl_equal _)
- | H: _ <= ?p |- _ <= ?p => apply Zle_trans with (2 := H)
- | H: _ < ?p |- _ <= ?p =>
- apply Zlt_le_weak; apply Zle_lt_trans with (2 := H)
- end).
-
-Hint Extern 2 (Zlt _ _) =>
- (match goal with
- |- Zpos _ < Zpos _ => exact (refl_equal _)
-| H: _ <= ?p |- _ <= ?p => apply Zlt_le_trans with (2 := H)
-| H: _ < ?p |- _ <= ?p => apply Zle_lt_trans with (2 := H)
- end).
-
-Hint Resolve Zlt_gt Zle_ge: zarith.
-
-(**************************************
- Properties Zmod
-**************************************)
-
- Theorem Zmod_mult:
- forall a b n, 0 < n -> (a * b) mod n = ((a mod n) * (b mod n)) mod n.
- Proof.
- intros a b n H.
- pattern a at 1; rewrite (Z_div_mod_eq a n); auto with zarith.
- pattern b at 1; rewrite (Z_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) +
- (((n*(a/n)) * (b/n) + (b mod n)*(a / n)) + (a mod n) * (b / n)) * n);
- auto with zarith.
- apply Z_mod_plus; auto with zarith.
- ring.
- Qed.
-
- Theorem Zmod_plus:
- forall a b n, 0 < n -> (a + b) mod n = (a mod n + b mod n) mod n.
- Proof.
- intros a b n H.
- pattern a at 1; rewrite (Z_div_mod_eq a n); auto with zarith.
- pattern b at 1; rewrite (Z_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); auto with zarith.
- apply Z_mod_plus; auto with zarith.
- ring.
- Qed.
-
- Theorem Zmod_mod: forall a n, 0 < n -> (a mod n) mod n = a mod n.
- Proof.
- intros a n H.
- pattern a at 2; rewrite (Z_div_mod_eq a n); auto with zarith.
- rewrite Zplus_comm; rewrite Zmult_comm.
- apply sym_equal; apply Z_mod_plus; auto with zarith.
- Qed.
-
- Theorem Zmod_def_small: forall a n, 0 <= a < n -> a mod n = a.
- Proof.
- intros a n [H1 H2]; unfold Zmod.
- generalize (Z_div_mod a n); case (Zdiv_eucl a n).
- intros q r H3; case H3; clear H3; auto with zarith.
- intros H4 [H5 H6].
- case (Zle_or_lt q (- 1)); intros H7.
- contradict H1; apply Zlt_not_le.
- subst a.
- apply Zle_lt_trans with (n * - 1 + r); auto with zarith.
- case (Zle_lt_or_eq 0 q); auto with zarith; intros H8.
- contradict H2; apply Zle_not_lt.
- apply Zle_trans with (n * 1 + r); auto with zarith.
- rewrite H4; auto with zarith.
- subst a; subst q; auto with zarith.
- Qed.
-
- Theorem Zmod_minus:
- forall a b n, 0 < n -> (a - b) mod n = (a mod n - b mod n) mod n.
- Proof.
- intros a b n H; replace (a - b) with (a + (-1) * b); auto with zarith.
- replace (a mod n - b mod n) with (a mod n + (-1)*(b mod n));auto with zarith.
- rewrite Zmod_plus; auto with zarith.
- rewrite Zmod_mult; auto with zarith.
- rewrite (fun x y => Zmod_plus x ((-1) * y)); auto with zarith.
- rewrite Zmod_mult; auto with zarith.
- rewrite (fun x => Zmod_mult x (b mod n)); auto with zarith.
- repeat rewrite Zmod_mod; auto.
- Qed.
-
- Theorem Zmod_le: forall a n, 0 < n -> 0 <= a -> (Zmod a n) <= a.
- Proof.
- intros a n H1 H2; case (Zle_or_lt n a); intros H3.
- case (Z_mod_lt a n); auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
- Qed.
-
- Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
- Proof.
- intros a b H H1;case (Z_mod_lt a b);auto with zarith;intros H2 H3;split;auto.
- case (Zle_or_lt b a); intros H4; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
- Qed.
-
-
-(**************************************
- Properties of Zdivide
-**************************************)
-
- Theorem Zdiv_pos: forall a b, 0 < b -> 0 <= a -> 0 <= a / b.
- Proof.
- intros; apply Zge_le; apply Z_div_ge0; auto with zarith.
- Qed.
- Hint Resolve Zdiv_pos: zarith.
-
- 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.
- 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 Zmod_mult; case (Zmod_le_first ((c mod b) * (a mod b)) b);
- auto with zarith.
- apply Zmult_le_compat_r; auto with zarith.
- case (Zmod_le_first 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.
-
- Theorem Zdiv_unique:
- forall n d q r, 0 < d -> ( 0 <= r < d ) -> n = d * q + r -> q = n / d.
- Proof.
- intros n d q r H H0 H1.
- assert (H2: n = d * (n / d) + n mod d).
- apply Z_div_mod_eq; auto with zarith.
- assert (H3: 0 <= n mod d < d ).
- apply Z_mod_lt; auto with zarith.
- case (Ztrichotomy q (n / d)); auto.
- intros H4.
- absurd (n < n); auto with zarith.
- pattern n at 1; rewrite H1; rewrite H2.
- apply Zlt_le_trans with (d * (q + 1)); auto with zarith.
- rewrite Zmult_plus_distr_r; auto with zarith.
- apply Zle_trans with (d * (n / d)); auto with zarith.
- intros tmp; case tmp; auto; intros H4; clear tmp.
- absurd (n < n); auto with zarith.
- pattern n at 2; rewrite H1; rewrite H2.
- apply Zlt_le_trans with (d * (n / d + 1)); auto with zarith.
- rewrite Zmult_plus_distr_r; auto with zarith.
- apply Zle_trans with (d * q); auto with zarith.
- Qed.
-
- Theorem Zmod_unique:
- forall n d q r, 0 < d -> ( 0 <= r < d ) -> n = d * q + r -> r = n mod d.
- Proof.
- intros n d q r H H0 H1.
- assert (H2: n = d * (n / d) + n mod d).
- apply Z_div_mod_eq; auto with zarith.
- rewrite (Zdiv_unique n d q r) in H1; auto.
- apply (Zplus_reg_l (d * (n / d))); auto with zarith.
- Qed.
-
- Theorem Zmod_Zmult_compat_l: forall a b c,
- 0 < b -> 0 < c -> c * a mod (c * b) = c * (a mod b).
- Proof.
- intros a b c H2 H3.
- pattern a at 1; rewrite (Z_div_mod_eq a b); auto with zarith.
- rewrite Zplus_comm; rewrite Zmult_plus_distr_r.
- rewrite Zmult_assoc; rewrite (Zmult_comm (c * b)).
- rewrite Z_mod_plus; auto with zarith.
- apply Zmod_def_small; split; auto with zarith.
- apply Zmult_le_0_compat; auto with zarith.
- destruct (Z_mod_lt a b);auto with zarith.
- apply Zmult_lt_compat_l; auto with zarith.
- destruct (Z_mod_lt a b);auto with zarith.
- Qed.
-
- Theorem Zdiv_Zmult_compat_l:
- forall a b c, 0 <= a -> 0 < b -> 0 < c -> c * a / (c * b) = a / b.
- Proof.
- intros a b c H1 H2 H3; case (Z_mod_lt a b); auto with zarith; intros H4 H5.
- apply Zdiv_unique with (a mod b); auto with zarith.
- apply Zmult_reg_l with c; auto with zarith.
- rewrite Zmult_plus_distr_r; rewrite <- Zmod_Zmult_compat_l; auto with zarith.
- rewrite Zmult_assoc; apply Z_div_mod_eq; auto with zarith.
- Qed.
-
- Theorem Zdiv_0_l: forall a, 0 / a = 0.
- Proof.
- intros a; case a; auto.
- Qed.
-
- Theorem Zdiv_0_r: forall a, a / 0 = 0.
- Proof.
- intros a; case a; auto.
- Qed.
-
- Theorem Zmod_0_l: forall a, 0 mod a = 0.
- Proof.
- intros a; case a; auto.
- Qed.
-
- Theorem Zmod_0_r: forall a, a mod 0 = 0.
- Proof.
- intros a; case a; auto.
- Qed.
-
- Theorem Zdiv_le_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_le_reg_r with b; auto with zarith.
- apply Zle_trans with (2 := H3).
- 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.
-
- Theorem Zdiv_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 (Z_div_mod_eq a b); auto with zarith.
- rewrite (Zmult_comm b); case (Z_mod_lt a b); auto with zarith.
- Qed.
-
- Theorem Zdiv_le_lower_bound:
- forall a b q, 0 <= a -> 0 < b -> q * b <= a -> q <= a / b.
- Proof.
- intros a b q H1 H2 H3.
- assert (q < a / b + 1); auto with zarith.
- apply Zmult_lt_reg_r with b; auto with zarith.
- apply Zle_lt_trans with (1 := H3).
- pattern a at 1; rewrite (Z_div_mod_eq a b); auto with zarith.
- rewrite Zmult_plus_distr_l; rewrite (Zmult_comm b); case (Z_mod_lt a b);
- auto with zarith.
- Qed.
-
- Theorem Zmult_mod_distr_l:
- forall a b c, 0 < a -> 0 < c -> (a * b) mod (a * c) = a * (b mod c).
- Proof.
- intros a b c H Hc.
- apply sym_equal; apply Zmod_unique with (b / c); auto with zarith.
- apply Zmult_lt_0_compat; auto.
- case (Z_mod_lt b c); auto with zarith; intros; split; auto with zarith.
- apply Zmult_lt_compat_l; auto.
- rewrite <- Zmult_assoc; rewrite <- Zmult_plus_distr_r.
- rewrite <- Z_div_mod_eq; auto with zarith.
- Qed.
-
-
- Theorem Zmod_distr: forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
- (2 ^a * r + t) mod (2 ^ b) = (2 ^a * r) mod (2 ^ b) + t.
- Proof.
- intros a b r t (H1, H2) H3 (H4, H5).
- assert (t < 2 ^ b).
- apply Zlt_le_trans with (1:= H5); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- rewrite Zmod_plus; auto with zarith.
- rewrite Zmod_def_small with (a := t); auto with zarith.
- apply Zmod_def_small; auto with zarith.
- split; auto with zarith.
- assert (0 <= 2 ^a * r); auto with zarith.
- apply Zplus_le_0_compat; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- pattern (2 ^ b) at 2; replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a);
- try ring.
- apply Zplus_le_lt_compat; auto with zarith.
- replace b with ((b - a) + a); try ring.
- rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
- try rewrite <- Zmult_minus_distr_r.
- rewrite (Zmult_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l;
- auto with zarith.
- rewrite (Zmult_comm (2 ^a)); apply Zmult_le_compat_r; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- Qed.
-
- Theorem Zmult_mod_distr_r:
- forall a b c : Z, 0 < a -> 0 < c -> (b * a) mod (c * a) = (b mod c) * a.
- Proof.
- intros; repeat rewrite (fun x => (Zmult_comm x a)).
- apply Zmult_mod_distr_l; auto.
- Qed.
-
- Theorem Z_div_plus_l: forall a b c : Z, 0 < b -> (a * b + c) / b = a + c / b.
- Proof.
- intros a b c H; rewrite Zplus_comm; rewrite Z_div_plus;
- try apply Zplus_comm; auto with zarith.
- Qed.
-
- Theorem Zmod_shift_r:
- forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
- (r * 2 ^a + t) mod (2 ^ b) = (r * 2 ^a) mod (2 ^ b) + t.
- Proof.
- intros a b r t (H1, H2) H3 (H4, H5).
- assert (t < 2 ^ b).
- apply Zlt_le_trans with (1:= H5); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- rewrite Zmod_plus; auto with zarith.
- rewrite Zmod_def_small with (a := t); auto with zarith.
- apply Zmod_def_small; auto with zarith.
- split; auto with zarith.
- assert (0 <= 2 ^a * r); auto with zarith.
- apply Zplus_le_0_compat; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- pattern (2 ^ b) at 2;replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); try ring.
- apply Zplus_le_lt_compat; auto with zarith.
- replace b with ((b - a) + a); try ring.
- rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
- try rewrite <- Zmult_minus_distr_r.
- repeat rewrite (fun x => Zmult_comm x (2 ^ a)); rewrite Zmult_mod_distr_l;
- auto with zarith.
- apply Zmult_le_compat_l; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- Qed.
-
- Theorem Zdiv_lt_0: forall a b, 0 <= a < b -> a / b = 0.
- intros a b H; apply sym_equal; apply Zdiv_unique with a; auto with zarith.
- Qed.
-
- Theorem Zmod_mult_0: forall a b, 0 < b -> (a * b) mod b = 0.
- Proof.
- intros a b H; rewrite <- (Zplus_0_l (a * b)); rewrite Z_mod_plus;
- auto with zarith.
- Qed.
-
- Theorem Zdiv_shift_r:
- forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
- (r * 2 ^a + t) / (2 ^ b) = (r * 2 ^a) / (2 ^ b).
- Proof.
- intros a b r t (H1, H2) H3 (H4, H5).
- assert (Eq: t < 2 ^ b); auto with zarith.
- apply Zlt_le_trans with (1 := H5); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
- pattern (r * 2 ^ a) at 1; rewrite Z_div_mod_eq with (b := 2 ^ b);
- auto with zarith.
- rewrite <- Zplus_assoc.
- rewrite <- Zmod_shift_r; auto with zarith.
- rewrite (Zmult_comm (2 ^ b)); rewrite Z_div_plus_l; auto with zarith.
- rewrite (fun x y => @Zdiv_lt_0 (x mod y)); auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
- auto with zarith.
- Qed.
-
-
- Theorem Zpos_minus:
- forall a b, Zpos a < Zpos b -> Zpos (b- a) = Zpos b - Zpos a.
- Proof.
- intros a b H.
- repeat rewrite Zpos_eq_Z_of_nat_o_nat_of_P; autorewrite with pos_morph;
- auto with zarith.
- rewrite inj_minus1; auto with zarith.
- match goal with |- (?X <= ?Y)%nat =>
- case (le_or_lt X Y); auto; intro tmp; absurd (Z_of_nat X < Z_of_nat Y);
- try apply Zle_not_lt; auto with zarith
- end.
- repeat rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P; auto with zarith.
- generalize (Zlt_gt _ _ H); auto.
- Qed.
-
- Theorem Zdiv_Zmult_compat_r:
- forall a b c : Z, 0 <= a -> 0 < b -> 0 < c -> a * c / (b * c) = a / b.
- Proof.
- intros a b c H H1 H2; repeat rewrite (fun x => Zmult_comm x c);
- apply Zdiv_Zmult_compat_l; auto.
- Qed.
-
-
- Lemma shift_unshift_mod : forall n p a,
- 0 <= a < 2^n ->
- 0 <= p <= n ->
- a * 2^p = a / 2^(n - p) * 2^n + (a*2^p) mod 2^n.
- Proof.
- intros n p a H1 H2.
- pattern (a*2^p) at 1;replace (a*2^p) with
- (a*2^p/2^n * 2^n + a*2^p mod 2^n).
- 2:symmetry;rewrite (Zmult_comm (a*2^p/2^n));apply Z_div_mod_eq.
- replace (a * 2 ^ p / 2 ^ n) with (a / 2 ^ (n - p));trivial.
- replace (2^n) with (2^(n-p)*2^p).
- symmetry;apply Zdiv_Zmult_compat_r.
- destruct H1;trivial.
- apply Zpower_lt_0;auto with zarith.
- apply Zpower_lt_0;auto with zarith.
- rewrite <- Zpower_exp.
- replace (n-p+p) with n;trivial. ring.
- omega. omega.
- apply Zlt_gt. apply Zpower_lt_0;auto with zarith.
- Qed.
-
- Lemma Zdiv_Zdiv : forall a b c, 0 < b -> 0 < c -> (a/b)/c = a / (b*c).
- Proof.
- intros a b c H H0.
- pattern a at 2;rewrite (Z_div_mod_eq a b);auto with zarith.
- pattern (a/b) at 2;rewrite (Z_div_mod_eq (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));try ring.
- rewrite Z_div_plus_l;auto with zarith.
- rewrite (Zdiv_lt_0 (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.
- apply Zmult_lt_0_compat;auto with zarith.
- Qed.
-
- Lemma div_le_0 : forall p x, 0 <= x -> 0 <= x / 2 ^ p.
- Proof.
- intros p x Hle;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_le_lower_bound;auto with zarith.
- replace (2^p) with 0.
- destruct x;compute;intro;discriminate.
- destruct p;trivial;discriminate z.
- Qed.
-
- Lemma div_lt : forall p x y, 0 <= x < y -> x / 2^p < y.
- Proof.
- intros p x y H;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_lt_upper_bound;auto with zarith.
- apply Zlt_le_trans with y;auto with zarith.
- rewrite <- (Zmult_1_r y);apply Zmult_le_compat;auto with zarith.
- assert (0 < 2^p);auto with zarith.
- replace (2^p) with 0.
- destruct x;change (0<y);auto with zarith.
- destruct p;trivial;discriminate z.
- Qed.
-
-
- Theorem Zgcd_div_pos a b:
- (0 < b)%Z -> (0 < Zgcd a b)%Z -> (0 < b / Zgcd a b)%Z.
- intros a b Ha Hg.
- case (Zle_lt_or_eq 0 (b/Zgcd a b)); auto.
- apply Zdiv_pos; auto with zarith.
- intros H; generalize Ha.
- pattern b at 1; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto.
- rewrite <- H; auto with zarith.
- assert (F := (Zgcd_is_gcd a b)); inversion F; auto.
- Qed.
-
diff --git a/theories/Ints/Z/ZPowerAux.v b/theories/Ints/Z/ZPowerAux.v
deleted file mode 100644
index 151b9a73a..000000000
--- a/theories/Ints/Z/ZPowerAux.v
+++ /dev/null
@@ -1,203 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-(**********************************************************************
-
-
- ZPowerAux.v Auxillary functions & Theorems for Zpower
- **********************************************************************)
-
-Require Export ZArith.
-Require Export Znumtheory.
-Require Export Tactic.
-
-Open Local Scope Z_scope.
-
-Hint Extern 2 (Zle _ _) =>
- (match goal with
- |- Zpos _ <= Zpos _ => exact (refl_equal _)
-| H: _ <= ?p |- _ <= ?p => apply Zle_trans with (2 := H)
-| H: _ < ?p |- _ <= ?p => apply Zlt_le_weak; apply Zle_lt_trans with (2 := H)
- end).
-
-Hint Extern 2 (Zlt _ _) =>
- (match goal with
- |- Zpos _ < Zpos _ => exact (refl_equal _)
-| H: _ <= ?p |- _ <= ?p => apply Zlt_le_trans with (2 := H)
-| H: _ < ?p |- _ <= ?p => apply Zle_lt_trans with (2 := H)
- end).
-
-Hint Resolve Zlt_gt Zle_ge: zarith.
-
-(**************************************
- Properties Zpower
-**************************************)
-
-Theorem Zpower_1: forall a, 0 <= a -> 1 ^ a = 1.
-intros a Ha; pattern a; apply natlike_ind; auto with zarith.
-intros x Hx Hx1; unfold Zsucc.
-rewrite Zpower_exp; auto with zarith.
-rewrite Hx1; simpl; auto.
-Qed.
-
-Theorem Zpower_exp_0: forall a, a ^ 0 = 1.
-simpl; unfold Zpower_pos; simpl; auto with zarith.
-Qed.
-
-Theorem Zpower_exp_1: forall a, a ^ 1 = a.
-simpl; unfold Zpower_pos; simpl; auto with zarith.
-Qed.
-
-Theorem Zpower_Zabs: forall a b, Zabs (a ^ b) = (Zabs a) ^ b.
-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.
-eq_tac; 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.
-intros p n H.
-unfold Zsucc; rewrite Zpower_exp; auto with zarith.
-rewrite Zpower_exp_1; apply Zmult_comm.
-Qed.
-
-Theorem Zpower_mult: forall p q r, 0 <= q -> 0 <= r -> p ^ (q * r) = (p ^ q) ^
- r.
-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_exp_1; try rewrite <- Zpower_exp; try eq_tac;
-auto with zarith.
-ring.
-Qed.
-
-Theorem Zpower_lt_0: forall a b: Z, 0 < a -> 0 <= b-> 0 < a ^b.
-intros a b; case b; auto with zarith.
-simpl; intros; auto with zarith.
-2: intros p H H1; case H1; auto.
-intros p H1 H; generalize H; pattern (Zpos p); apply natlike_ind; auto.
-intros; case a; compute; auto.
-intros p1 H2 H3 _; unfold Zsucc; rewrite Zpower_exp; simpl; auto with zarith.
-apply Zmult_lt_O_compat; auto with zarith.
-generalize H1; case a; compute; intros; auto; discriminate.
-Qed.
-
-Theorem Zpower_le_monotone: forall a b c: Z, 0 < a -> 0 <= b <= c -> a ^ b <= a ^ c.
-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_lt_0; auto with zarith.
-apply Zlt_le_weak; apply Zpower_lt_0; auto with zarith.
-Qed.
-
-
-Theorem Zpower_le_0: forall a b: Z, 0 <= a -> 0 <= a ^b.
-intros a b; case b; 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.
-intros p1 H2 H3 _; unfold Zsucc; rewrite Zpower_exp; simpl; auto with zarith.
-apply Zmult_le_0_compat; auto with zarith.
-generalize H1; case a; compute; intros; auto; discriminate.
-Qed.
-
-Hint Resolve Zpower_le_0 Zpower_lt_0: zarith.
-
-Theorem Zpower_prod: forall p q r, 0 <= q -> 0 <= r -> (p * q) ^ r = p ^ r * q ^ r.
-intros p q r H1 H2; generalize H2; pattern r; apply natlike_ind; auto.
-intros r1 H3 H4 H5.
-unfold Zsucc; rewrite Zpower_exp; auto with zarith.
-rewrite H4; repeat (rewrite Zpower_exp_1 || rewrite Zpower_exp); auto with zarith; ring.
-Qed.
-
-Theorem Zpower_le_monotone_exp: forall a b c: Z, 0 <= c -> 0 <= a <= b -> a ^ c <= b ^ c.
-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_exp_1.
-apply Zle_trans with (a ^x * b); auto with zarith.
-Qed.
-
-Theorem Zpower_lt_monotone: forall a b c: Z, 1 < a -> 0 <= b < c -> a ^ b < a ^
- c.
-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.
-assert (0 < a ^ (c - b)); auto with zarith.
-apply Zlt_le_trans with (a ^1); auto with zarith.
-rewrite Zpower_exp_1; auto with zarith.
-apply Zpower_le_monotone; auto with zarith.
-Qed.
-
-Lemma Zpower_le_monotone_inv :
- forall 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_monotone_exp;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; elimtype False;omega.
- destruct c;try discriminate z0. simpl in H1. elimtype False;omega.
- assert (H4 := Zpower_lt_monotone a c b H). elimtype False;omega.
-Qed.
-
-
-Theorem Zpower_le_monotone2:
- forall a b c: Z, 0 < a -> b <= c -> a ^ b <= a ^ c.
-intros a b c H H2.
-destruct (Z_le_gt_dec 0 b).
-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.
-replace (a^b) with 0.
-destruct (Z_le_gt_dec 0 c).
-destruct (Zle_lt_or_eq _ _ z0).
-apply Zlt_le_weak;apply Zpower_lt_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 Zpower2_lt_lin: forall n,
- 0 <= n -> n < 2 ^ n.
-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_exp_1.
- 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_le_lin: forall n,
- 0 <= n -> n <= 2 ^ n.
-intros; apply Zlt_le_weak.
-apply Zpower2_lt_lin; auto.
-Qed.
diff --git a/theories/Ints/Z/ZSum.v b/theories/Ints/Z/ZSum.v
deleted file mode 100644
index bcde7386c..000000000
--- a/theories/Ints/Z/ZSum.v
+++ /dev/null
@@ -1,321 +0,0 @@
-
-(*************************************************************)
-(* This file is distributed under the terms of the *)
-(* GNU Lesser General Public License Version 2.1 *)
-(*************************************************************)
-(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
-(*************************************************************)
-
-(***********************************************************************
- Summation.v from Z to Z
- *********************************************************************)
-Require Import Arith.
-Require Import ArithRing.
-Require Import ListAux.
-Require Import ZArith.
-Require Import ZAux.
-Require Import Iterator.
-Require Import ZProgression.
-
-
-Open Scope Z_scope.
-(* Iterated Sum *)
-
-Definition Zsum :=
- fun n m f =>
- if Zle_bool n m
- then iter 0 f Zplus (progression Zsucc n (Zabs_nat ((1 + m) - n)))
- else iter 0 f Zplus (progression Zpred n (Zabs_nat ((1 + n) - m))).
-Hint Unfold Zsum .
-
-Lemma Zsum_nn: forall n f, Zsum n n f = f n.
-intros n f; unfold Zsum; rewrite Zle_bool_refl.
-replace ((1 + n) - n) with 1; auto with zarith.
-simpl; ring.
-Qed.
-
-Theorem permutation_rev: forall (A:Set) (l : list A), permutation (rev l) l.
-intros a l; elim l; simpl; auto.
-intros a1 l1 Hl1.
-apply permutation_trans with (cons a1 (rev l1)); auto.
-change (permutation (rev l1 ++ (a1 :: nil)) (app (cons a1 nil) (rev l1))); auto.
-Qed.
-
-Lemma Zsum_swap: forall (n m : Z) (f : Z -> Z), Zsum n m f = Zsum m n f.
-intros n m f; unfold Zsum.
-generalize (Zle_cases n m) (Zle_cases m n); case (Zle_bool n m);
- case (Zle_bool m n); auto with arith.
-intros; replace n with m; auto with zarith.
-3:intros H1 H2; contradict H2; auto with zarith.
-intros H1 H2; apply iter_permutation; auto with zarith.
-apply permutation_trans
- with (rev (progression Zsucc n (Zabs_nat ((1 + m) - n)))).
-apply permutation_sym; apply permutation_rev.
-rewrite Zprogression_opp; auto with zarith.
-replace (n + Z_of_nat (pred (Zabs_nat ((1 + m) - n)))) with m; auto.
-replace (Zabs_nat ((1 + m) - n)) with (S (Zabs_nat (m - n))); auto with zarith.
-simpl.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-replace ((1 + m) - n) with (1 + (m - n)); auto with zarith.
-cut (0 <= m - n); auto with zarith; unfold Zabs_nat.
-case (m - n); auto with zarith.
-intros p; case p; simpl; auto with zarith.
-intros p1 Hp1; rewrite nat_of_P_xO; rewrite nat_of_P_xI;
- rewrite nat_of_P_succ_morphism.
-simpl; repeat rewrite plus_0_r.
-repeat rewrite <- plus_n_Sm; simpl; auto.
-intros p H3; contradict H3; auto with zarith.
-intros H1 H2; apply iter_permutation; auto with zarith.
-apply permutation_trans
- with (rev (progression Zsucc m (Zabs_nat ((1 + n) - m)))).
-rewrite Zprogression_opp; auto with zarith.
-replace (m + Z_of_nat (pred (Zabs_nat ((1 + n) - m)))) with n; auto.
-replace (Zabs_nat ((1 + n) - m)) with (S (Zabs_nat (n - m))); auto with zarith.
-simpl.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-replace ((1 + n) - m) with (1 + (n - m)); auto with zarith.
-cut (0 <= n - m); auto with zarith; unfold Zabs_nat.
-case (n - m); auto with zarith.
-intros p; case p; simpl; auto with zarith.
-intros p1 Hp1; rewrite nat_of_P_xO; rewrite nat_of_P_xI;
- rewrite nat_of_P_succ_morphism.
-simpl; repeat rewrite plus_0_r.
-repeat rewrite <- plus_n_Sm; simpl; auto.
-intros p H3; contradict H3; auto with zarith.
-apply permutation_rev.
-Qed.
-
-Lemma Zsum_split_up:
- forall (n m p : Z) (f : Z -> Z),
- ( n <= m < p ) -> Zsum n p f = Zsum n m f + Zsum (m + 1) p f.
-intros n m p f [H H0].
-case (Zle_lt_or_eq _ _ H); clear H; intros H.
-unfold Zsum; (repeat rewrite Zle_imp_le_bool); auto with zarith.
-assert (H1: n < p).
-apply Zlt_trans with ( 1 := H ); auto with zarith.
-assert (H2: m < 1 + p).
-apply Zlt_trans with ( 1 := H0 ); auto with zarith.
-assert (H3: n < 1 + m).
-apply Zlt_trans with ( 1 := H ); auto with zarith.
-assert (H4: n < 1 + p).
-apply Zlt_trans with ( 1 := H1 ); auto with zarith.
-replace (Zabs_nat ((1 + p) - (m + 1)))
- with (minus (Zabs_nat ((1 + p) - n)) (Zabs_nat ((1 + m) - n))).
-apply iter_progression_app; auto with zarith.
-apply inj_le_inv.
-(repeat rewrite Z_of_nat_Zabs_nat); auto with zarith.
-rewrite next_n_Z; auto with zarith.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-apply inj_eq_inv; auto with zarith.
-rewrite inj_minus1; auto with zarith.
-(repeat rewrite Z_of_nat_Zabs_nat); auto with zarith.
-apply inj_le_inv.
-(repeat rewrite Z_of_nat_Zabs_nat); auto with zarith.
-subst m.
-rewrite Zsum_nn; auto with zarith.
-unfold Zsum; generalize (Zle_cases n p); generalize (Zle_cases (n + 1) p);
- case (Zle_bool n p); case (Zle_bool (n + 1) p); auto with zarith.
-intros H1 H2.
-replace (Zabs_nat ((1 + p) - n)) with (S (Zabs_nat (p - n))); auto with zarith.
-replace (n + 1) with (Zsucc n); auto with zarith.
-replace ((1 + p) - Zsucc n) with (p - n); auto with zarith.
-apply inj_eq_inv; auto with zarith.
-rewrite inj_S; (repeat rewrite Z_of_nat_Zabs_nat); auto with zarith.
-Qed.
-
-Lemma Zsum_S_left:
- forall (n m : Z) (f : Z -> Z), n < m -> Zsum n m f = f n + Zsum (n + 1) m f.
-intros n m f H; rewrite (Zsum_split_up n n m f); auto with zarith.
-rewrite Zsum_nn; auto with zarith.
-Qed.
-
-Lemma Zsum_S_right:
- forall (n m : Z) (f : Z -> Z),
- n <= m -> Zsum n (m + 1) f = Zsum n m f + f (m + 1).
-intros n m f H; rewrite (Zsum_split_up n m (m + 1) f); auto with zarith.
-rewrite Zsum_nn; auto with zarith.
-Qed.
-
-Lemma Zsum_split_down:
- forall (n m p : Z) (f : Z -> Z),
- ( p < m <= n ) -> Zsum n p f = Zsum n m f + Zsum (m - 1) p f.
-intros n m p f [H H0].
-case (Zle_lt_or_eq p (m - 1)); auto with zarith; intros H1.
-pattern m at 1; replace m with ((m - 1) + 1); auto with zarith.
-repeat rewrite (Zsum_swap n).
-rewrite (Zsum_swap (m - 1)).
-rewrite Zplus_comm.
-apply Zsum_split_up; auto with zarith.
-subst p.
-repeat rewrite (Zsum_swap n).
-rewrite Zsum_nn.
-unfold Zsum; (repeat rewrite Zle_imp_le_bool); auto with zarith.
-replace (Zabs_nat ((1 + n) - (m - 1))) with (S (Zabs_nat (n - (m - 1)))).
-rewrite Zplus_comm.
-replace (Zabs_nat ((1 + n) - m)) with (Zabs_nat (n - (m - 1))); auto with zarith.
-pattern m at 4; replace m with (Zsucc (m - 1)); auto with zarith.
-apply f_equal with ( f := Zabs_nat ); auto with zarith.
-apply inj_eq_inv; auto with zarith.
-rewrite inj_S.
-(repeat rewrite Z_of_nat_Zabs_nat); auto with zarith.
-Qed.
-
-
-Lemma Zsum_ext:
- forall (n m : Z) (f g : Z -> Z),
- n <= m ->
- (forall (x : Z), ( n <= x <= m ) -> f x = g x) -> Zsum n m f = Zsum n m g.
-intros n m f g HH H.
-unfold Zsum; auto.
-unfold Zsum; (repeat rewrite Zle_imp_le_bool); auto with zarith.
-apply iter_ext; auto with zarith.
-intros a H1; apply H; auto; split.
-apply Zprogression_le_init with ( 1 := H1 ).
-cut (a < Zsucc m); auto with zarith.
-replace (Zsucc m) with (n + Z_of_nat (Zabs_nat ((1 + m) - n))); auto with zarith.
-apply Zprogression_le_end; auto with zarith.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-Qed.
-
-Lemma Zsum_add:
- forall (n m : Z) (f g : Z -> Z),
- Zsum n m f + Zsum n m g = Zsum n m (fun (i : Z) => f i + g i).
-intros n m f g; unfold Zsum; case (Zle_bool n m); apply iter_comp;
- auto with zarith.
-Qed.
-
-Lemma Zsum_times:
- forall n m x f, x * Zsum n m f = Zsum n m (fun i=> x * f i).
-intros n m x f.
-unfold Zsum. case (Zle_bool n m); intros; apply iter_comp_const with (k := (fun y : Z => x * y)); auto with zarith.
-Qed.
-
-Lemma inv_Zsum:
- forall (P : Z -> Prop) (n m : Z) (f : Z -> Z),
- n <= m ->
- P 0 ->
- (forall (a b : Z), P a -> P b -> P (a + b)) ->
- (forall (x : Z), ( n <= x <= m ) -> P (f x)) -> P (Zsum n m f).
-intros P n m f HH H H0 H1.
-unfold Zsum; rewrite Zle_imp_le_bool; auto with zarith; apply iter_inv; auto.
-intros x H3; apply H1; auto; split.
-apply Zprogression_le_init with ( 1 := H3 ).
-cut (x < Zsucc m); auto with zarith.
-replace (Zsucc m) with (n + Z_of_nat (Zabs_nat ((1 + m) - n))); auto with zarith.
-apply Zprogression_le_end; auto with zarith.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-Qed.
-
-
-Lemma Zsum_pred:
- forall (n m : Z) (f : Z -> Z),
- Zsum n m f = Zsum (n + 1) (m + 1) (fun (i : Z) => f (Zpred i)).
-intros n m f.
-unfold Zsum.
-generalize (Zle_cases n m); generalize (Zle_cases (n + 1) (m + 1));
- case (Zle_bool n m); case (Zle_bool (n + 1) (m + 1)); auto with zarith.
-replace ((1 + (m + 1)) - (n + 1)) with ((1 + m) - n); auto with zarith.
-intros H1 H2; cut (exists c , c = Zabs_nat ((1 + m) - n) ).
-intros [c H3]; rewrite <- H3.
-generalize n; elim c; auto with zarith; clear H1 H2 H3 c n.
-intros c H n; simpl; eq_tac; auto with zarith.
-eq_tac; unfold Zpred; auto with zarith.
-replace (Zsucc (n + 1)) with (Zsucc n + 1); auto with zarith.
-exists (Zabs_nat ((1 + m) - n)); auto.
-replace ((1 + (n + 1)) - (m + 1)) with ((1 + n) - m); auto with zarith.
-intros H1 H2; cut (exists c , c = Zabs_nat ((1 + n) - m) ).
-intros [c H3]; rewrite <- H3.
-generalize n; elim c; auto with zarith; clear H1 H2 H3 c n.
-intros c H n; simpl; (eq_tac; auto with zarith).
-eq_tac; unfold Zpred; auto with zarith.
-replace (Zpred (n + 1)) with (Zpred n + 1); auto with zarith.
-unfold Zpred; auto with zarith.
-exists (Zabs_nat ((1 + n) - m)); auto.
-Qed.
-
-Theorem Zsum_c:
- forall (c p q : Z), p <= q -> Zsum p q (fun x => c) = ((1 + q) - p) * c.
-intros c p q Hq; unfold Zsum.
-rewrite Zle_imp_le_bool; auto with zarith.
-pattern ((1 + q) - p) at 2; rewrite <- Z_of_nat_Zabs_nat; auto with zarith.
-cut (exists r , r = Zabs_nat ((1 + q) - p) );
- [intros [r H1]; rewrite <- H1 | exists (Zabs_nat ((1 + q) - p))]; auto.
-generalize p; elim r; auto with zarith.
-intros n H p0; replace (Z_of_nat (S n)) with (Z_of_nat n + 1); auto with zarith.
-simpl; rewrite H; ring.
-rewrite inj_S; auto with zarith.
-Qed.
-
-Theorem Zsum_Zsum_f:
- forall (i j k l : Z) (f : Z -> Z -> Z),
- i <= j ->
- k < l ->
- Zsum i j (fun x => Zsum k (l + 1) (fun y => f x y)) =
- Zsum i j (fun x => Zsum k l (fun y => f x y) + f x (l + 1)).
-intros; apply Zsum_ext; intros; auto with zarith.
-rewrite Zsum_S_right; auto with zarith.
-Qed.
-
-Theorem Zsum_com:
- forall (i j k l : Z) (f : Z -> Z -> Z),
- Zsum i j (fun x => Zsum k l (fun y => f x y)) =
- Zsum k l (fun y => Zsum i j (fun x => f x y)).
-intros; unfold Zsum; case (Zle_bool i j); case (Zle_bool k l); apply iter_com;
- auto with zarith.
-Qed.
-
-Theorem Zsum_le:
- forall (n m : Z) (f g : Z -> Z),
- n <= m ->
- (forall (x : Z), ( n <= x <= m ) -> (f x <= g x )) ->
- (Zsum n m f <= Zsum n m g ).
-intros n m f g Hl H.
-unfold Zsum; rewrite Zle_imp_le_bool; auto with zarith.
-unfold Zsum;
- cut
- (forall x,
- In x (progression Zsucc n (Zabs_nat ((1 + m) - n))) -> ( f x <= g x )).
-elim (progression Zsucc n (Zabs_nat ((1 + m) - n))); simpl; auto with zarith.
-intros x H1; apply H; split.
-apply Zprogression_le_init with ( 1 := H1 ); auto.
-cut (x < m + 1); auto with zarith.
-replace (m + 1) with (n + Z_of_nat (Zabs_nat ((1 + m) - n))); auto with zarith.
-apply Zprogression_le_end; auto with zarith.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-Qed.
-
-Theorem iter_le:
-forall (f g: Z -> Z) l, (forall a, In a l -> f a <= g a) ->
- iter 0 f Zplus l <= iter 0 g Zplus l.
-intros f g l; elim l; simpl; auto with zarith.
-Qed.
-
-Theorem Zsum_lt:
- forall n m f g,
- (forall x, n <= x -> x <= m -> f x <= g x) ->
- (exists x, n <= x /\ x <= m /\ f x < g x) ->
- Zsum n m f < Zsum n m g.
-intros n m f g H (d, (Hd1, (Hd2, Hd3))); unfold Zsum; rewrite Zle_imp_le_bool; auto with zarith.
-cut (In d (progression Zsucc n (Zabs_nat (1 + m - n)))).
-cut (forall x, In x (progression Zsucc n (Zabs_nat (1 + m - n)))-> f x <= g x).
-elim (progression Zsucc n (Zabs_nat (1 + m - n))); simpl; auto with zarith.
-intros a l Rec H0 [H1 | H1]; subst; auto.
-apply Zle_lt_trans with (f d + iter 0 g Zplus l); auto with zarith.
-apply Zplus_le_compat_l.
-apply iter_le; auto.
-apply Zlt_le_trans with (f a + iter 0 g Zplus l); auto with zarith.
-intros x H1; apply H.
-apply Zprogression_le_init with ( 1 := H1 ); auto.
-cut (x < m + 1); auto with zarith.
-replace (m + 1) with (n + Z_of_nat (Zabs_nat ((1 + m) - n))); auto with zarith.
-apply Zprogression_le_end with ( 1 := H1 ); auto with arith.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-apply in_Zprogression.
-rewrite Z_of_nat_Zabs_nat; auto with zarith.
-Qed.
-
-Theorem Zsum_minus:
- forall n m f g, Zsum n m f - Zsum n m g = Zsum n m (fun x => f x - g x).
-intros n m f g; apply trans_equal with (Zsum n m f + (-1) * Zsum n m g); auto with zarith.
-rewrite Zsum_times; rewrite Zsum_add; auto with zarith.
-Qed.
diff --git a/theories/Ints/num/GenAdd.v b/theories/Ints/num/GenAdd.v
index 9d4c57902..77f5e2301 100644
--- a/theories/Ints/num/GenAdd.v
+++ b/theories/Ints/num/GenAdd.v
@@ -10,7 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Basic_type.
Require Import GenBase.
@@ -265,11 +264,11 @@ Section GenAdd.
Lemma spec_ww_succ : forall x, [[ww_succ x]] = ([[x]] + 1) mod wwB.
Proof.
destruct x as [ |xh xl];simpl.
- rewrite spec_ww_1;rewrite Zmod_def_small;trivial.
+ rewrite spec_ww_1;rewrite Zmod_small;trivial.
split;[intro;discriminate|apply wwB_pos].
rewrite <- Zplus_assoc;generalize (spec_w_succ_c xl);
destruct (w_succ_c xl) as[l|l];intro H;unfold interp_carry in H;rewrite <-H.
- rewrite Zmod_def_small;trivial.
+ rewrite Zmod_small;trivial.
rewrite wwB_wBwB;apply beta_mult;apply spec_to_Z.
assert ([|l|] = 0). clear spec_ww_1 spec_w_1 spec_w_0.
assert (H1:= spec_to_Z l); assert (H2:= spec_to_Z xl); omega.
@@ -281,10 +280,10 @@ Section GenAdd.
Lemma spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB.
Proof.
destruct x as [ |xh xl];intros y;simpl.
- rewrite Zmod_def_small;trivial. apply spec_ww_to_Z;trivial.
+ rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
destruct y as [ |yh yl].
change [[W0]] with 0;rewrite Zplus_0_r.
- rewrite Zmod_def_small;trivial.
+ rewrite Zmod_small;trivial.
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)).
simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]))
with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring.
diff --git a/theories/Ints/num/GenBase.v b/theories/Ints/num/GenBase.v
index 6b6376b69..a9936f1dd 100644
--- a/theories/Ints/num/GenBase.v
+++ b/theories/Ints/num/GenBase.v
@@ -10,8 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
-Require Import ZPowerAux.
Require Import Basic_type.
Require Import JMeq.
@@ -188,7 +186,7 @@ Section GenBase.
Lemma lt_0_wB : 0 < wB.
Proof.
- unfold base;apply Zpower_lt_0. unfold Zlt;reflexivity.
+ unfold base;apply Zpower_gt_0. unfold Zlt;reflexivity.
unfold Zle;intros H;discriminate H.
Qed.
@@ -208,7 +206,7 @@ Section GenBase.
Proof.
assert (H:= wB_pos);rewrite wwB_wBwB;rewrite <-(Zmult_1_r 1).
rewrite Zpower_2.
- apply Zmult_lt_compat;(split;[unfold Zlt;reflexivity|trivial]).
+ apply Zmult_lt_compat2;(split;[unfold Zlt;reflexivity|trivial]).
apply Zlt_le_weak;trivial.
Qed.
@@ -217,11 +215,11 @@ Section GenBase.
clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W
spec_to_Z;unfold base.
assert (2 ^ Zpos w_digits = 2 * (2 ^ (Zpos w_digits - 1))).
- pattern 2 at 2; rewrite <- Zpower_exp_1.
+ pattern 2 at 2; rewrite <- Zpower_1_r.
rewrite <- Zpower_exp; auto with zarith.
- eq_tac; auto with zarith.
+ f_equal; auto with zarith.
case w_digits; compute; intros; discriminate.
- rewrite H; eq_tac; auto with zarith.
+ rewrite H; f_equal; auto with zarith.
rewrite Zmult_comm; apply Z_div_mult; auto with zarith.
Qed.
@@ -239,17 +237,17 @@ Section GenBase.
(z*wB + [|x|]) mod wwB = (z mod wB)*wB + [|x|].
Proof.
intros z x.
- rewrite Zmod_plus. 2:apply lt_0_wwB.
+ rewrite Zplus_mod.
pattern wwB at 1;rewrite wwB_wBwB; rewrite Zpower_2.
rewrite Zmult_mod_distr_r;try apply lt_0_wB.
- rewrite (Zmod_def_small [|x|]).
- apply Zmod_def_small;rewrite wwB_wBwB;apply beta_mult;try apply spec_to_Z.
+ rewrite (Zmod_small [|x|]).
+ apply Zmod_small;rewrite wwB_wBwB;apply beta_mult;try apply spec_to_Z.
apply Z_mod_lt;apply Zlt_gt;apply lt_0_wB.
destruct (spec_to_Z x);split;trivial.
change [|x|] with (0*wB+[|x|]). rewrite wwB_wBwB.
rewrite Zpower_2;rewrite <- (Zplus_0_r (wB*wB));apply beta_lex_inv.
apply lt_0_wB. apply spec_to_Z. split;[apply Zle_refl | apply lt_0_wB].
- Qed.
+ Qed.
Lemma wB_div : forall x y, ([|x|] * wB + [|y|]) / wB = [|x|].
Proof.
@@ -321,7 +319,7 @@ Section GenBase.
apply Zmult_le_compat; auto with zarith.
apply Zle_trans with wB; auto with zarith.
unfold base.
- rewrite <- (ZAux.Zpower_exp_0 2).
+ rewrite <- (Zpower_0_r 2).
apply Zpower_le_monotone2; auto with zarith.
unfold base; auto with zarith.
Qed.
diff --git a/theories/Ints/num/GenDiv.v b/theories/Ints/num/GenDiv.v
index 6466c198f..237adb48b 100644
--- a/theories/Ints/num/GenDiv.v
+++ b/theories/Ints/num/GenDiv.v
@@ -10,8 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
-Require Import ZPowerAux.
Require Import Basic_type.
Require Import GenBase.
Require Import GenDivn1.
@@ -102,42 +100,41 @@ Section POS_MOD.
assert (F0: forall x y, x - y + y = x); auto with zarith.
intros w1 p; case (spec_to_w_Z p); intros HH1 HH2.
unfold ww_pos_mod; case w1.
- simpl; rewrite Zmod_def_small; split; auto with zarith.
+ simpl; rewrite Zmod_small; split; auto with zarith.
intros xh xl; generalize (spec_ww_compare p (w_0W w_zdigits));
case ww_compare;
rewrite spec_w_0W; rewrite spec_zdigits; fold wB;
intros H1.
rewrite H1; simpl ww_to_Z.
autorewrite with w_rewrite rm10.
- rewrite Zmod_plus; auto with zarith.
- rewrite Zmod_mult_0; auto with zarith.
+ rewrite Zplus_mod; auto with zarith.
+ rewrite Z_mod_mult; auto with zarith.
autorewrite with rm10.
rewrite Zmod_mod; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
autorewrite with w_rewrite rm10.
simpl ww_to_Z.
rewrite spec_pos_mod.
assert (HH0: [|low p|] = [[p]]).
rewrite spec_low.
- apply Zmod_def_small; auto with zarith.
+ apply Zmod_small; auto with zarith.
case (spec_to_w_Z p); intros HHH1 HHH2; split; auto with zarith.
apply Zlt_le_trans with (1 := H1).
- unfold base.
- apply Zpower2_le_lin; auto with zarith.
+ unfold base; apply Zpower2_le_lin; auto with zarith.
rewrite HH0.
- rewrite Zmod_plus; auto with zarith.
+ rewrite Zplus_mod; auto with zarith.
unfold base.
rewrite <- (F0 (Zpos w_digits) [[p]]).
rewrite Zpower_exp; auto with zarith.
rewrite Zmult_assoc.
- rewrite Zmod_mult_0; auto with zarith.
+ rewrite Z_mod_mult; auto with zarith.
autorewrite with w_rewrite rm10.
rewrite Zmod_mod; auto with zarith.
generalize (spec_ww_compare p ww_zdigits);
case ww_compare; rewrite spec_ww_zdigits;
rewrite spec_zdigits; intros H2.
replace (2^[[p]]) with wwB.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
unfold base; rewrite H2.
rewrite spec_ww_digits; auto.
assert (HH0: [|low (ww_sub p (w_0W w_zdigits))|] =
@@ -146,7 +143,7 @@ generalize (spec_ww_compare p ww_zdigits);
rewrite spec_ww_sub.
rewrite spec_w_0W; rewrite spec_zdigits.
rewrite <- Zmod_div_mod; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zlt_le_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
@@ -162,11 +159,11 @@ generalize (spec_ww_compare p ww_zdigits);
unfold base; rewrite <- Zmult_assoc; rewrite <- Zpower_exp;
auto with zarith.
rewrite F0; auto with zarith.
- rewrite <- Zplus_assoc; rewrite Zmod_plus; auto with zarith.
- rewrite Zmod_mult_0; auto with zarith.
+ rewrite <- Zplus_assoc; rewrite Zplus_mod; auto with zarith.
+ rewrite Z_mod_mult; auto with zarith.
autorewrite with rm10.
rewrite Zmod_mod; auto with zarith.
- apply sym_equal; apply Zmod_def_small; auto with zarith.
+ apply sym_equal; apply Zmod_small; auto with zarith.
case (spec_to_Z xh); intros U1 U2.
case (spec_to_Z xl); intros U3 U4.
split; auto with zarith.
@@ -186,7 +183,7 @@ generalize (spec_ww_compare p ww_zdigits);
case (spec_to_Z xl); unfold base; auto with zarith.
rewrite Zmult_minus_distr_r; rewrite <- Zpower_exp; auto with zarith.
rewrite F0; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
case (spec_to_w_Z (WW xh xl)); intros U1 U2.
split; auto with zarith.
apply Zlt_le_trans with (1:= U2).
@@ -349,7 +346,7 @@ Section GenDiv32.
simpl;rewrite spec_sub.
assert ([|a2|] - [|b2|] = wB*(-1) + ([|a2|] - [|b2|] + wB)). ring.
assert (0 <= [|a2|] - [|b2|] + wB < wB). omega.
- rewrite <-(Zmod_unique ([|a2|]-[|b2|]) wB (-1) ([|a2|]-[|b2|]+wB) U H1 H0).
+ rewrite <-(Zmod_unique ([|a2|]-[|b2|]) wB (-1) ([|a2|]-[|b2|]+wB) H1 H0).
rewrite wwB_wBwB;ring.
assert (U2 := wB_pos w_digits).
eapply spec_ww_add_c_cont with (P :=
@@ -433,8 +430,8 @@ Section GenDiv32.
rewrite <- H8 in H2;rewrite H2 in H7.
assert (0 < [|b1|]*wB). apply Zmult_lt_0_compat;zarith.
Spec_ww_to_Z r2. zarith.
- rewrite (Zmod_def_small ([|q|] -1));zarith.
- rewrite (Zmod_def_small ([|q|] -1 -1));zarith.
+ rewrite (Zmod_small ([|q|] -1));zarith.
+ rewrite (Zmod_small ([|q|] -1 -1));zarith.
assert ([[r2]] + ([|b1|] * wB + [|b2|]) =
wwB * 1 +
([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2 * ([|b1|] * wB + [|b2|]))).
@@ -455,11 +452,10 @@ Section GenDiv32.
wwB
1
([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2*([|b1|] * wB + [|b2|]))
- U1
H10 H8).
split. ring. zarith.
intros r2;repeat (rewrite spec_pred);simpl ww_to_Z;intros H7.
- rewrite (Zmod_def_small ([|q|] -1));zarith.
+ rewrite (Zmod_small ([|q|] -1));zarith.
split.
replace [[r2]] with ([[r1]] + ([|b1|] * wB + [|b2|]) -wwB).
rewrite H2; ring. rewrite <- H7; ring.
@@ -580,7 +576,7 @@ Section GenDiv21.
match goal with |-context [ww_compare ?Y ?Z] =>
generalize (spec_ww_compare Y Z); case (ww_compare Y Z)
end; simpl;try rewrite spec_ww_1;autorewrite with rm10; intros;zarith.
- rewrite spec_ww_sub;simpl. rewrite Zmod_def_small;zarith.
+ rewrite spec_ww_sub;simpl. rewrite Zmod_small;zarith.
split. ring.
assert (wwB <= 2*[[b]]);zarith.
rewrite wwB_div;zarith.
@@ -927,7 +923,7 @@ Section GenDivGt.
Spec_ww_to_Z (WW ah al).
rewrite spec_ww_sub;eauto.
simpl;rewrite spec_ww_1;rewrite Zmult_1_l;simpl.
- simpl ww_to_Z in Hgt, H, HH;rewrite Zmod_def_small;split;zarith.
+ simpl ww_to_Z in Hgt, H, HH;rewrite Zmod_small;split;zarith.
case (spec_to_Z (w_head0 bh)); auto with zarith.
assert ([|w_head0 bh|] < Zpos w_digits).
destruct (Z_lt_ge_dec [|w_head0 bh|] (Zpos w_digits));trivial.
@@ -951,8 +947,8 @@ Section GenDivGt.
assert (H5:=to_Z_div_minus_p bl HHHH).
rewrite Zmult_comm in Hh.
assert (2^[|w_head0 bh|] < wB). unfold base;apply Zpower_lt_monotone;zarith.
- unfold base in H0;rewrite Zmod_def_small;zarith.
- fold wB; rewrite (Zmod_def_small ([|bh|] * 2 ^ [|w_head0 bh|]));zarith.
+ unfold base in H0;rewrite Zmod_small;zarith.
+ fold wB; rewrite (Zmod_small ([|bh|] * 2 ^ [|w_head0 bh|]));zarith.
intros U1 U2 U3 V1 V2.
generalize (@spec_w_div32 (w_add_mul_div (w_head0 bh) w_0 ah)
(w_add_mul_div (w_head0 bh) ah al)
@@ -991,7 +987,7 @@ Section GenDivGt.
Zmult_0_l;rewrite Zplus_0_l.
replace [[ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry
_ww_zdigits (w_0W (w_head0 bh))) W0 r]] with ([[r]]/2^[|w_head0 bh|]).
- assert (0 < 2^[|w_head0 bh|]). apply Zpower_lt_0;zarith.
+ assert (0 < 2^[|w_head0 bh|]). apply Zpower_gt_0;zarith.
split.
rewrite <- (Z_div_mult [[WW ah al]] (2^[|w_head0 bh|]));zarith.
rewrite H1;rewrite Zmult_assoc;apply Z_div_plus_l;trivial.
@@ -1002,9 +998,9 @@ Section GenDivGt.
change (Zpos (xO (w_digits))) with (2*Zpos (w_digits));zarith.
simpl ww_to_Z;rewrite Zmult_0_l;rewrite Zplus_0_l.
rewrite spec_w_0W.
- rewrite (fun x y => Zmod_def_small (x-y)); auto with zarith.
+ rewrite (fun x y => Zmod_small (x-y)); auto with zarith.
ring_simplify (2 * Zpos w_digits - (2 * Zpos w_digits - [|w_head0 bh|])).
- rewrite Zmod_def_small;zarith.
+ rewrite Zmod_small;zarith.
split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith.
Spec_ww_to_Z r.
apply Zlt_le_trans with wwB;zarith.
@@ -1015,7 +1011,7 @@ Section GenDivGt.
apply Zpower2_lt_lin; auto with zarith.
rewrite spec_ww_sub; auto with zarith.
rewrite spec_ww_digits_; rewrite spec_w_0W.
- rewrite Zmod_def_small;zarith.
+ rewrite Zmod_small;zarith.
rewrite Zpos_xO; split; auto with zarith.
apply Zle_lt_trans with (2 * Zpos w_digits); auto with zarith.
unfold base, ww_digits; rewrite (Zpos_xO w_digits).
@@ -1412,7 +1408,7 @@ Section GenDiv.
intros a b Hpos;unfold ww_mod.
assert (H := spec_ww_compare a b);destruct (ww_compare a b).
simpl;apply Zmod_unique with 1;try rewrite H;zarith.
- Spec_ww_to_Z a;symmetry;apply Zmod_def_small;zarith.
+ Spec_ww_to_Z a;symmetry;apply Zmod_small;zarith.
apply spec_ww_mod_gt;trivial.
Qed.
diff --git a/theories/Ints/num/GenDivn1.v b/theories/Ints/num/GenDivn1.v
index 84bf247f7..7987741da 100644
--- a/theories/Ints/num/GenDivn1.v
+++ b/theories/Ints/num/GenDivn1.v
@@ -9,8 +9,7 @@
Set Implicit Arguments.
Require Import ZArith.
-Require Import ZDivModAux.
-Require Import ZPowerAux.
+Require Import ZAux.
Require Import Basic_type.
Require Import GenBase.
@@ -230,7 +229,7 @@ Section GENDIVN1.
with (2*Zpos (gen_digits w_digits n));auto with zarith.
replace (2 ^ (Zpos (gen_digits w_digits (S n)) - [|p|])) with
(2^(Zpos (gen_digits w_digits n) - [|p|])*2^Zpos (gen_digits w_digits n)).
- rewrite Zdiv_Zmult_compat_r;auto with zarith.
+ rewrite Zdiv_mult_cancel_r;auto with zarith.
rewrite Zmult_plus_distr_l with (p:= 2^[|p|]).
pattern ([!n|hl!] * 2^[|p|]) at 2;
rewrite (shift_unshift_mod (Zpos(gen_digits w_digits n))([|p|])([!n|hl!]));
@@ -254,7 +253,7 @@ Section GENDIVN1.
with (2^Zpos(gen_digits w_digits n) *2^Zpos(gen_digits w_digits n)).
rewrite (Zmult_comm (([!n|hh!] * 2 ^ [|p|] +
[!n|hl!] / 2 ^ (Zpos (gen_digits w_digits n) - [|p|])))).
- rewrite Zmod_Zmult_compat_l;auto with zarith.
+ rewrite Zmult_mod_distr_l;auto with zarith.
ring.
rewrite Zpower_exp;auto with zarith.
assert (0 < Zpos (gen_digits w_digits n)). unfold Zlt;reflexivity.
@@ -336,7 +335,7 @@ Section GENDIVN1.
replace (2 ^ (Zpos (gen_digits w_digits (S n)) - Zpos w_digits)) with
(2^(Zpos (gen_digits w_digits n) - Zpos w_digits) *
2^Zpos (gen_digits w_digits n)).
- rewrite Zdiv_Zmult_compat_r;auto with zarith.
+ rewrite Zdiv_mult_cancel_r;auto with zarith.
rewrite <- Zpower_exp;auto with zarith.
replace (Zpos (gen_digits w_digits n) - Zpos w_digits +
Zpos (gen_digits w_digits n)) with
@@ -373,7 +372,7 @@ Section GENDIVN1.
generalize (spec_compare (w_head0 b) w_0); case w_compare;
rewrite spec_0; intros H2; auto with zarith.
assert (Hv1: wB/2 <= [|b|]).
- generalize H0; rewrite H2; rewrite Zpower_exp_0;
+ generalize H0; rewrite H2; rewrite Zpower_0_r;
rewrite Zmult_1_l; auto.
assert (Hv2: [|w_0|] < [|b|]).
rewrite spec_0; auto.
@@ -394,7 +393,7 @@ Section GENDIVN1.
rewrite (spec_add_mul_div b w_0); auto with zarith.
rewrite spec_0;rewrite Zdiv_0_l; try omega.
rewrite Zplus_0_r; rewrite Zmult_comm.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
assert (H5 := spec_to_Z (high n a)).
assert
([|w_add_mul_div (w_head0 b) w_0 (high n a)|]
@@ -407,15 +406,15 @@ Section GENDIVN1.
apply Zlt_le_trans with wB;auto with zarith.
pattern wB at 1;replace wB with (wB*1);try ring.
apply Zmult_le_compat;auto with zarith.
- assert (H6 := Zpower_lt_0 2 (Zpos w_digits - [|w_head0 b|]));
+ assert (H6 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|]));
auto with zarith.
- rewrite Zmod_def_small;auto with zarith.
+ rewrite Zmod_small;auto with zarith.
apply Zdiv_lt_upper_bound;auto with zarith.
apply Zlt_le_trans with wB;auto with zarith.
apply Zle_trans with (2 ^ [|w_head0 b|] * [|b|] * 2).
rewrite <- wB_div_2; try omega.
apply Zmult_le_compat;auto with zarith.
- pattern 2 at 1;rewrite <- Zpower_exp_1.
+ pattern 2 at 1;rewrite <- Zpower_1_r.
apply Zpower_le_monotone;split;auto with zarith.
rewrite <- H4 in H0.
assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith.
@@ -430,13 +429,13 @@ Section GENDIVN1.
rewrite spec_0 in H7;rewrite Zmult_0_l in H7;rewrite Zplus_0_l in H7.
assert (([|high n a|] / 2 ^ (Zpos w_digits - [|w_head0 b|])) mod wB
= [!n|a!] / 2^(Zpos (gen_digits w_digits n) - [|w_head0 b|])).
- rewrite Zmod_def_small;auto with zarith.
+ rewrite Zmod_small;auto with zarith.
rewrite spec_high. rewrite Zdiv_Zdiv;auto with zarith.
rewrite <- Zpower_exp;auto with zarith.
replace (Zpos (gen_digits w_digits n) - Zpos w_digits +
(Zpos w_digits - [|w_head0 b|]))
with (Zpos (gen_digits w_digits n) - [|w_head0 b|]);trivial;ring.
- assert (H8 := Zpower_lt_0 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith.
+ assert (H8 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith.
split;auto with zarith.
apply Zle_lt_trans with ([|high n a|]);auto with zarith.
apply Zdiv_le_upper_bound;auto with zarith.
@@ -451,20 +450,20 @@ Section GENDIVN1.
rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l.
replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|])
with ([|w_head0 b|]).
- rewrite Zmod_def_small;auto with zarith.
+ rewrite Zmod_small;auto with zarith.
assert (H9 := spec_to_Z r).
split;auto with zarith.
apply Zle_lt_trans with ([|r|]);auto with zarith.
apply Zdiv_le_upper_bound;auto with zarith.
pattern ([|r|]) at 1;rewrite <- Zmult_1_r.
apply Zmult_le_compat;auto with zarith.
- assert (H10 := Zpower_lt_0 2 ([|w_head0 b|]));auto with zarith.
+ assert (H10 := Zpower_gt_0 2 ([|w_head0 b|]));auto with zarith.
rewrite spec_sub.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
case (spec_to_Z w_zdigits); auto with zarith.
rewrite spec_sub.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
case (spec_to_Z w_zdigits); auto with zarith.
case H7; intros H71 H72.
diff --git a/theories/Ints/num/GenLift.v b/theories/Ints/num/GenLift.v
index a7d8480ba..06c76067e 100644
--- a/theories/Ints/num/GenLift.v
+++ b/theories/Ints/num/GenLift.v
@@ -10,8 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZPowerAux.
-Require Import ZDivModAux.
Require Import Basic_type.
Require Import GenBase.
@@ -283,7 +281,7 @@ Section GenLift.
repeat rewrite <- Zmult_assoc.
apply f_equal2 with (f := Zmult); auto.
case (spec_to_Z (w_tail0 xl)); intros HH3 HH4.
- pattern 2 at 2; rewrite <- Zpower_exp_1.
+ pattern 2 at 2; rewrite <- Zpower_1_r.
lazy beta; repeat rewrite <- Zpower_exp; auto with zarith.
unfold base; apply f_equal with (f := Zpower 2); auto with zarith.
@@ -328,12 +326,12 @@ Section GenLift.
fold wB.
rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;rewrite <- Zplus_assoc.
rewrite <- Zpower_2.
- rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|]. apply lt_0_wwB.
+ rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|].
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xl yh)). ring.
simpl ww_to_Z; w_rewrite;zarith.
assert (HH0: [|low p|] = [[p]]).
rewrite spec_low.
- apply Zmod_def_small.
+ apply Zmod_small.
case (spec_to_w_Z p); intros HH1 HH2; split; auto.
generalize H1; unfold zdigits; rewrite spec_w_0W;
rewrite spec_zdigits; intros tmp.
@@ -370,7 +368,7 @@ Section GenLift.
rewrite spec_ww_sub.
unfold zdigits; rewrite spec_w_0W; rewrite spec_zdigits.
rewrite <- Zmod_div_mod; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zle_lt_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_lt_lin; auto with zarith.
@@ -407,7 +405,7 @@ Section GenLift.
([|xh|]*2^u*wB). 2:ring.
repeat rewrite <- Zplus_assoc.
rewrite (Zplus_comm ([|xh|] * 2 ^ u * wB)).
- rewrite Z_mod_plus;zarith. rewrite Zmod_mult_0;zarith.
+ rewrite Z_mod_plus;zarith. rewrite Z_mod_mult;zarith.
unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith.
unfold u; split;zarith.
split;zarith. unfold u; apply Zdiv_lt_upper_bound;zarith.
@@ -450,7 +448,7 @@ Section GenLift.
generalize H1; w_rewrite; rewrite spec_zdigits; clear H1; intros H1.
assert (HH0: [|low p|] = [[p]]).
rewrite spec_low.
- apply Zmod_def_small.
+ apply Zmod_small.
case (spec_to_w_Z p); intros HH1 HH2; split; auto.
apply Zlt_le_trans with (1 := H1).
unfold base; apply Zpower2_le_lin; auto with zarith.
@@ -466,7 +464,7 @@ Section GenLift.
rewrite spec_low.
rewrite spec_ww_sub; w_rewrite; intros H1.
rewrite <- Zmod_div_mod; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zle_lt_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_lt_lin; auto with zarith.
diff --git a/theories/Ints/num/GenMul.v b/theories/Ints/num/GenMul.v
index 7550781f1..5522e41bf 100644
--- a/theories/Ints/num/GenMul.v
+++ b/theories/Ints/num/GenMul.v
@@ -10,7 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Basic_type.
Require Import GenBase.
@@ -327,7 +326,7 @@ Section GenMul.
2:ring. rewrite <- H1;rewrite <- H;rewrite <- H0.
assert (H2 := sum_mul_carry _ _ _ _ _ _ H1).
destruct cc as [ | cch ccl]; simpl zn2z_to_Z; simpl ww_to_Z.
- rewrite spec_ww_add;rewrite spec_w_W0;rewrite Zmod_def_small;
+ rewrite spec_ww_add;rewrite spec_w_W0;rewrite Zmod_small;
rewrite wwB_wBwB. ring.
rewrite <- (Zplus_0_r ([|wc|]*wB));rewrite H;apply mult_add_ineq3;zarith.
simpl ww_to_Z in H1. assert (U:=spec_to_Z cch).
@@ -349,7 +348,7 @@ Section GenMul.
as [l|l];unfold interp_carry;rewrite spec_w_W0;try rewrite Zmult_1_l;
simpl zn2z_to_Z;
try rewrite spec_ww_add;try rewrite spec_ww_add_carry;rewrite spec_w_WW;
- rewrite Zmod_def_small;rewrite wwB_wBwB;intros.
+ rewrite Zmod_small;rewrite wwB_wBwB;intros.
rewrite H4;ring. rewrite H;apply mult_add_ineq2;zarith.
rewrite Zplus_assoc;rewrite Zmult_plus_distr_l.
rewrite Zmult_1_l;rewrite <- Zplus_assoc;rewrite H4;ring.
@@ -385,8 +384,8 @@ Section GenMul.
Lemma spec_w_2: [|w_2|] = 2.
unfold w_2; rewrite spec_w_add; rewrite spec_w_1; simpl.
- apply Zmod_def_small; split; auto with zarith.
- rewrite <- (Zpower_exp_1 2); unfold base; apply Zpower_lt_monotone; auto with zarith.
+ apply Zmod_small; split; auto with zarith.
+ rewrite <- (Zpower_1_r 2); unfold base; apply Zpower_lt_monotone; auto with zarith.
Qed.
Lemma kara_prod_aux : forall xh xl yh yl,
@@ -410,7 +409,7 @@ Section GenMul.
generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh;
try rewrite Hylh; try rewrite spec_w_0; try (ring; fail).
repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
split; auto with zarith.
simpl in Hz; rewrite Hz; rewrite H; rewrite H0.
rewrite kara_prod_aux; apply Zplus_le_0_compat; apply Zmult_le_0_compat; auto with zarith.
@@ -422,10 +421,10 @@ Section GenMul.
intros z1 Hz2
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2;
repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh;
try rewrite Hylh; try rewrite spec_w_0; try (ring; fail).
match goal with |- context[ww_add_c ?x ?y] =>
@@ -433,12 +432,12 @@ Section GenMul.
intros z1 Hz2
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2;
repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
split.
match goal with |- context[(?x - ?y) * (?z - ?t)] =>
replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring]
@@ -459,22 +458,22 @@ Section GenMul.
intros z1 Hz2
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l.
generalize Hz2; clear Hz2; unfold interp_carry.
repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
match goal with |- context[ww_add_c ?x ?y] =>
generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1;
intros z1 Hz2
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_2; unfold interp_carry in Hz2.
apply trans_equal with (wwB + (1 * wwB + [[z1]])).
ring.
rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh;
try rewrite Hylh; try rewrite spec_w_1; try (ring; fail).
match goal with |- context[ww_add_c ?x ?y] =>
@@ -482,25 +481,25 @@ Section GenMul.
intros z1 Hz2
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_2; unfold interp_carry in Hz2.
apply trans_equal with (wwB + (1 * wwB + [[z1]])).
ring.
rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
match goal with |- context[ww_sub_c ?x ?y] =>
generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1;
intros z1 Hz2
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l.
match goal with |- context[(?x - ?y) * (?z - ?t)] =>
replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring]
end.
generalize Hz2; clear Hz2; unfold interp_carry.
repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
- repeat rewrite Zmod_def_small; auto with zarith; try (ring; fail).
+ repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
Qed.
Lemma sub_carry : forall xh xl yh yl z,
@@ -539,21 +538,21 @@ Section GenMul.
assert (U1:= lt_0_wwB w_digits).
intros x y; case x; auto; intros xh xl.
case y; auto.
- simpl; rewrite Zmult_0_r; rewrite Zmod_def_small; auto with zarith.
+ simpl; rewrite Zmult_0_r; rewrite Zmod_small; auto with zarith.
intros yh yl;simpl.
repeat (rewrite spec_ww_add || rewrite spec_w_W0 || rewrite spec_w_mul_c
|| rewrite spec_w_add || rewrite spec_w_mul).
- rewrite <- Zmod_plus; auto with zarith.
+ rewrite <- Zplus_mod; auto with zarith.
repeat (rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r).
rewrite <- Zmult_mod_distr_r; auto with zarith.
rewrite <- Zpower_2; rewrite <- wwB_wBwB; auto with zarith.
- rewrite Zmod_plus; auto with zarith.
+ rewrite Zplus_mod; auto with zarith.
rewrite Zmod_mod; auto with zarith.
- rewrite <- Zmod_plus; auto with zarith.
+ rewrite <- Zplus_mod; auto with zarith.
match goal with |- ?X mod _ = _ =>
rewrite <- Z_mod_plus with (a := X) (b := [|xh|] * [|yh|])
end; auto with zarith.
- eq_tac; auto; rewrite wwB_wBwB; ring.
+ f_equal; auto; rewrite wwB_wBwB; ring.
Qed.
Lemma spec_ww_square_c : forall x, [||ww_square_c x||] = [[x]]*[[x]].
@@ -607,7 +606,7 @@ Section GenMul.
rewrite spec_w_0;trivial.
assert (U:=spec_w_add_c l r);destruct (w_add_c l r) as [lr|lr];unfold
interp_carry in U;try rewrite Zmult_1_l in H;simpl.
- rewrite U;ring. rewrite spec_w_succ. rewrite Zmod_def_small.
+ rewrite U;ring. rewrite spec_w_succ. rewrite Zmod_small.
rewrite <- Zplus_assoc;rewrite <- U;ring.
simpl in H;assert (H1:= Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)).
rewrite <- H in H1.
diff --git a/theories/Ints/num/GenSqrt.v b/theories/Ints/num/GenSqrt.v
index 07b11ad74..3a5b59b6d 100644
--- a/theories/Ints/num/GenSqrt.v
+++ b/theories/Ints/num/GenSqrt.v
@@ -10,8 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
-Require Import ZPowerAux.
Require Import Basic_type.
Require Import GenBase.
@@ -277,10 +275,10 @@ Section GenSqrt.
clear spec_more_than_1_digit.
intros x; case x; simpl ww_is_even.
simpl.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
intros w1 w2; simpl.
unfold base.
- rewrite Zmod_plus; auto with zarith.
+ rewrite Zplus_mod; auto with zarith.
rewrite (fun x y => (Zdivide_mod (x * y))); auto with zarith.
rewrite Zplus_0_l; rewrite Zmod_mod; auto with zarith.
apply spec_w_is_even; auto with zarith.
@@ -313,12 +311,12 @@ intros x; case x; simpl ww_is_even.
destruct (spec_to_Z a2);auto with zarith.
intros H1 H2; split.
unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith.
- rewrite H2; rewrite Zmod_def_small; auto with zarith.
+ rewrite H2; rewrite Zmod_small; auto with zarith.
ring.
destruct (spec_to_Z a2);auto with zarith.
rewrite spec_w_sub; auto with zarith.
destruct (spec_to_Z a2) as [H3 H4];auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
assert ([|a2|] < 2 * [|b|]); auto with zarith.
apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith.
@@ -332,7 +330,7 @@ intros x; case x; simpl ww_is_even.
intros H1.
assert (H2: [|w_sub a1 b|] < [|b|]).
rewrite spec_w_sub; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
assert ([|a1|] < 2 * [|b|]); auto with zarith.
apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith.
rewrite wB_div_2; auto.
@@ -344,7 +342,7 @@ intros x; case x; simpl ww_is_even.
auto
end.
intros w0 w1; replace [+|C1 w0|] with (wB + [|w0|]).
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
intros (H3, H4); split; auto.
rewrite Zmult_plus_distr_l.
rewrite <- Zplus_assoc; rewrite <- H3; ring.
@@ -366,7 +364,7 @@ intros x; case x; simpl ww_is_even.
intros w1.
assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1).
rewrite spec_pred; rewrite spec_w_zdigits.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zlt_le_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
@@ -375,7 +373,7 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[?X - ?Y] =>
replace (X - Y) with 1
end.
- rewrite Zpower_exp_1; rewrite Zmod_def_small; auto with zarith.
+ rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith.
destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
split; auto with zarith.
apply Zdiv_lt_upper_bound; auto with zarith.
@@ -388,7 +386,7 @@ intros x; case x; simpl ww_is_even.
intros w1.
assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1).
rewrite spec_pred; rewrite spec_w_zdigits.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zlt_le_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
@@ -397,7 +395,8 @@ intros x; case x; simpl ww_is_even.
replace (X - Y) with 1
end; rewrite Hp; try ring.
rewrite Zpos_minus; auto with zarith.
- rewrite Zpower_exp_1; rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmax_right; auto with zarith.
+ rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith.
destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
split; auto with zarith.
unfold base.
@@ -405,14 +404,14 @@ intros x; case x; simpl ww_is_even.
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp
end.
- rewrite Zpower_exp; try rewrite Zpower_exp_1; auto with zarith.
+ rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith.
assert (tmp: forall p, 1 + (p -1) - 1 = p - 1); auto with zarith;
rewrite tmp; clear tmp; auto with zarith.
match goal with |- ?X + ?Y < _ =>
assert (Y < X); auto with zarith
end.
apply Zdiv_lt_upper_bound; auto with zarith.
- pattern 2 at 2; rewrite <- Zpower_exp_1; rewrite <- Zpower_exp;
+ pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp;
auto with zarith.
assert (tmp: forall p, (p - 1) + 1 = p); auto with zarith;
rewrite tmp; clear tmp; auto with zarith.
@@ -422,7 +421,7 @@ intros x; case x; simpl ww_is_even.
[|w_add_mul_div w_1 w w_0|] = 2 * [|w|] mod wB.
intros w1.
autorewrite with w_rewrite rm10; auto with zarith.
- rewrite Zpower_exp_1; auto with zarith.
+ rewrite Zpower_1_r; auto with zarith.
rewrite Zmult_comm; auto.
Qed.
@@ -432,7 +431,7 @@ intros x; case x; simpl ww_is_even.
rewrite spec_ww_add_mul_div; auto with zarith.
autorewrite with w_rewrite rm10.
rewrite spec_w_0W; rewrite spec_w_1.
- rewrite Zpower_exp_1; auto with zarith.
+ rewrite Zpower_1_r; auto with zarith.
rewrite Zmult_comm; auto.
rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
red; simpl; intros; discriminate.
@@ -444,22 +443,18 @@ intros x; case x; simpl ww_is_even.
intros w1.
rewrite spec_ww_add_mul_div; auto with zarith.
rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
- rewrite Zpower_exp_1; auto with zarith.
- eq_tac; auto.
- rewrite Zmult_comm; eq_tac; auto.
+ rewrite Zpower_1_r; auto with zarith.
+ f_equal; auto.
+ rewrite Zmult_comm; f_equal; auto.
autorewrite with w_rewrite rm10.
unfold ww_digits, base.
apply sym_equal; apply Zdiv_unique with (r := 2 ^ (Zpos (ww_digits w_digits) - 1) -1);
auto with zarith.
- apply Zpower_lt_0; auto with zarith.
- match goal with |- 0 <= ?X - 1 =>
- assert (0 < X); auto with zarith; red; reflexivity
- end.
unfold ww_digits; split; auto with zarith.
match goal with |- 0 <= ?X - 1 =>
assert (0 < X); auto with zarith
end.
- apply Zpower_lt_0; auto with zarith.
+ apply Zpower_gt_0; auto with zarith.
match goal with |- 0 <= ?X - 1 =>
assert (0 < X); auto with zarith; red; reflexivity
end.
@@ -468,8 +463,8 @@ intros x; case x; simpl ww_is_even.
rewrite tmp; clear tmp.
assert (tmp: forall p, p + p = 2 * p); auto with zarith;
rewrite tmp; clear tmp.
- eq_tac; auto.
- pattern 2 at 2; rewrite <- Zpower_exp_1; rewrite <- Zpower_exp;
+ f_equal; auto.
+ pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp;
auto with zarith.
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite tmp; clear tmp; auto.
@@ -480,8 +475,8 @@ intros x; case x; simpl ww_is_even.
red; simpl; intros; discriminate.
Qed.
- Theorem Zmod_plus_one: forall a1 b1, 0 < b1 -> (a1 + b1) mod b1 = a1 mod b1.
- intros a1 b1 H; rewrite Zmod_plus; auto with zarith.
+ Theorem Zplus_mod_one: forall a1 b1, 0 < b1 -> (a1 + b1) mod b1 = a1 mod b1.
+ intros a1 b1 H; rewrite Zplus_mod; auto with zarith.
rewrite Z_mod_same; try rewrite Zplus_0_r; auto with zarith.
apply Zmod_mod; auto.
Qed.
@@ -544,9 +539,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_exp_1; auto with zarith
+ try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
+ rewrite Zmax_right; auto with zarith.
ring.
repeat rewrite C0_id.
rewrite spec_w_add_c; auto with zarith.
@@ -560,9 +556,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_exp_1; auto with zarith
+ try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
+ rewrite Zmax_right; auto with zarith.
ring.
repeat rewrite C1_plus_wB in H0.
rewrite C1_plus_wB.
@@ -573,8 +570,8 @@ intros x; case x; simpl ww_is_even.
end.
intros c w1; case c.
intros w2 (Hw1, Hw2); rewrite C0_id in Hw1.
- rewrite <- Zmod_plus_one in Hw1; auto with zarith.
- rewrite Zmod_def_small in Hw1; auto with zarith.
+ rewrite <- Zplus_mod_one in Hw1; auto with zarith.
+ rewrite Zmod_small in Hw1; auto with zarith.
match goal with |- context[w_is_even ?y] =>
generalize (spec_w_is_even y);
case (w_is_even y)
@@ -592,9 +589,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_exp_1; auto with zarith
+ try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
+ rewrite Zmax_right; auto with zarith.
ring.
repeat rewrite C0_id.
rewrite add_mult_div_2_plus_1.
@@ -610,9 +608,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_exp_1; auto with zarith
+ try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
+ rewrite Zmax_right; auto with zarith.
ring.
split; auto with zarith.
destruct (spec_to_Z b);auto with zarith.
@@ -620,8 +619,8 @@ intros x; case x; simpl ww_is_even.
destruct (spec_to_Z b);auto with zarith.
destruct (spec_to_Z b);auto with zarith.
intros w2; rewrite C1_plus_wB.
- rewrite <- Zmod_plus_one; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite <- Zplus_mod_one; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
intros (Hw1, Hw2).
match goal with |- context[w_is_even ?y] =>
generalize (spec_w_is_even y);
@@ -662,7 +661,7 @@ intros x; case x; simpl ww_is_even.
4 * (2 ^ (Zpos w_digits - 2))).
change 4 with (2 ^ 2).
rewrite <- Zpower_exp; auto with zarith.
- eq_tac; auto with zarith.
+ f_equal; auto with zarith.
rewrite H.
rewrite (fun x => (Zmult_comm 4 (2 ^x))).
rewrite Z_div_mult; auto with zarith.
@@ -670,7 +669,7 @@ intros x; case x; simpl ww_is_even.
Theorem Zsquare_mult: forall p, p ^ 2 = p * p.
intros p; change 2 with (1 + 1); rewrite Zpower_exp;
- try rewrite Zpower_exp_1; auto with zarith.
+ try rewrite Zpower_1_r; auto with zarith.
Qed.
Theorem Zsquare_pos: forall p, 0 <= p ^ 2.
@@ -817,7 +816,7 @@ intros x; case x; simpl ww_is_even.
rewrite spec_ww_pred.
rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]);
auto with zarith.
- intros Hz1; rewrite Zmod_def_small; auto with zarith.
+ intros Hz1; rewrite Zmod_small; auto with zarith.
match type of H5 with -?X + ?Y = ?Z =>
assert (V: Y = Z + X);
try (rewrite <- H5; ring)
@@ -866,7 +865,7 @@ intros x; case x; simpl ww_is_even.
rewrite spec_ww_add; auto with zarith.
rewrite spec_ww_pred; auto with zarith.
rewrite ww_add_mult_mult_2.
- assert (VV1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)).
+ rename V1 into VV1.
assert (VV2: 0 < [[WW w4 w5]]); auto with zarith.
apply Zlt_le_trans with (wB/ 2 * wB + 0); auto with zarith.
autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith.
@@ -886,7 +885,7 @@ intros x; case x; simpl ww_is_even.
simpl ww_to_Z; assert (V4 := spec_to_Z w5);auto with zarith.
rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]);
auto with zarith.
- intros Hz1; rewrite Zmod_def_small; auto with zarith.
+ intros Hz1; rewrite Zmod_small; auto with zarith.
match type of H5 with -?X + ?Y = ?Z =>
assert (V: Y = Z + X);
try (rewrite <- H5; ring)
@@ -895,7 +894,7 @@ intros x; case x; simpl ww_is_even.
assert (V1: Y = Z - 1);
[replace (Z - 1) with (X + (-X + Z -1));
[rewrite <- Hz1 | idtac]; ring
- | idtac]
+ | idtac]
end.
rewrite <- Zmod_unique with (q := 1) (r := -wwB + [[z1]] + [[z]]);
auto with zarith.
@@ -1161,24 +1160,24 @@ intros x; case x; simpl ww_is_even.
assert (Hp0: 0 < [[ww_head0 x]]).
generalize (spec_ww_is_even (ww_head0 x)); rewrite H1.
generalize Hv1; case [[ww_head0 x]].
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
intros; assert (0 < Zpos p); auto with zarith.
red; simpl; auto.
intros p H2; case H2; auto.
assert (Hp: [[ww_pred (ww_head0 x)]] = [[ww_head0 x]] - 1).
rewrite spec_ww_pred.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
intros H2; split.
generalize (spec_ww_is_even (ww_pred (ww_head0 x)));
case ww_is_even; auto.
rewrite Hp.
- rewrite Zmod_minus; auto with zarith.
- rewrite H2; repeat rewrite Zmod_def_small; auto with zarith.
+ rewrite Zminus_mod; auto with zarith.
+ rewrite H2; repeat rewrite Zmod_small; auto with zarith.
intros H3; rewrite Hp.
case (spec_ww_head0 x); auto; intros Hv3 Hv4.
assert (Hu: forall u, 0 < u -> 2 * 2 ^ (u - 1) = 2 ^u).
intros u Hu.
- pattern 2 at 1; rewrite <- Zpower_exp_1.
+ pattern 2 at 1; rewrite <- Zpower_1_r.
rewrite <- Zpower_exp; auto with zarith.
ring_simplify (1 + (u - 1)); auto with zarith.
split; auto with zarith.
@@ -1254,13 +1253,13 @@ intros x; case x; simpl ww_is_even.
generalize (spec_ww_add_mul_div x W0 (ww_head1 x) Hv2);
case ww_add_mul_div.
simpl ww_to_Z; autorewrite with w_rewrite rm10.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
intros H2; case (Zmult_integral _ _ (sym_equal H2)); clear H2; intros H2.
rewrite H2; unfold Zpower, Zpower_pos; simpl; auto with zarith.
match type of H2 with ?X = ?Y =>
absurd (Y < X); try (rewrite H2; auto with zarith; fail)
end.
- apply Zpower_lt_0; auto with zarith.
+ apply Zpower_gt_0; auto with zarith.
split; auto with zarith.
case Hp2; intros _ tmp; apply Zle_lt_trans with (2 := tmp);
clear tmp.
@@ -1271,7 +1270,7 @@ intros x; case x; simpl ww_is_even.
generalize (spec_ww_is_even (ww_head1 x)); rewrite Hp1;
intros tmp; rewrite tmp; rewrite Zplus_0_r; auto.
intros w0 w1; autorewrite with w_rewrite rm10.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
2: rewrite Zmult_comm; auto with zarith.
intros H2.
assert (V: wB/4 <= [|w0|]).
@@ -1287,7 +1286,7 @@ intros x; case x; simpl ww_is_even.
assert (Hv3: [[ww_pred ww_zdigits]]
= Zpos (xO w_digits) - 1).
rewrite spec_ww_pred; rewrite spec_ww_zdigits.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zlt_le_trans with (Zpos (xO w_digits)); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
@@ -1303,8 +1302,8 @@ intros x; case x; simpl ww_is_even.
simpl ww_to_Z; autorewrite with rm10.
rewrite Hv3.
ring_simplify (Zpos (xO w_digits) - (Zpos (xO w_digits) - 1)).
- rewrite Zpower_exp_1.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zpower_1_r.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zlt_le_trans with (1 := Hv4); auto with zarith.
unfold base; apply Zpower_le_monotone; auto with zarith.
@@ -1313,15 +1312,15 @@ intros x; case x; simpl ww_is_even.
assert (Hv6: [|low(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))|]
= [[ww_head1 x]]/2).
rewrite spec_low.
- rewrite Hv5; rewrite Zmod_def_small; auto with zarith.
+ rewrite Hv5; rewrite Zmod_small; auto with zarith.
rewrite spec_w_add_mul_div; auto with zarith.
rewrite spec_w_sub; auto with zarith.
rewrite spec_w_0.
simpl ww_to_Z; autorewrite with rm10.
rewrite Hv6; rewrite spec_w_zdigits.
- rewrite (fun x y => Zmod_def_small (x - y)).
+ rewrite (fun x y => Zmod_small (x - y)).
ring_simplify (Zpos w_digits - (Zpos w_digits - [[ww_head1 x]] / 2)).
- rewrite Zmod_def_small.
+ rewrite Zmod_small.
simpl ww_to_Z in H2; rewrite H2; auto with zarith.
intros (H4, H5); split.
apply Zmult_le_reg_r with (2 ^ [[ww_head1 x]]); auto with zarith.
@@ -1335,7 +1334,7 @@ intros x; case x; simpl ww_is_even.
assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2);
try (intros; repeat rewrite Zsquare_mult; ring);
rewrite tmp; clear tmp.
- apply ZPowerAux.Zpower_le_monotone_exp; auto with zarith.
+ apply Zpower_le_monotone3; auto with zarith.
split; auto with zarith.
pattern [|w2|] at 2;
rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]] / 2)));
@@ -1361,7 +1360,7 @@ intros x; case x; simpl ww_is_even.
assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2);
try (intros; repeat rewrite Zsquare_mult; ring);
rewrite tmp; clear tmp.
- apply ZPowerAux.Zpower_le_monotone_exp; auto with zarith.
+ apply Zpower_le_monotone3; auto with zarith.
split; auto with zarith.
pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]]/2)));
auto with zarith.
@@ -1375,7 +1374,7 @@ intros x; case x; simpl ww_is_even.
auto with zarith.
apply Zmult_le_compat_l; auto with zarith.
apply Zpower_le_monotone; auto with zarith.
- rewrite Zpower_exp_0; autorewrite with rm10; auto.
+ rewrite Zpower_0_r; autorewrite with rm10; auto.
split; auto with zarith.
rewrite Hv0 in Hv2; rewrite (Zpos_xO w_digits) in Hv2; auto with zarith.
apply Zle_lt_trans with (Zpos w_digits); auto with zarith.
@@ -1383,7 +1382,7 @@ intros x; case x; simpl ww_is_even.
rewrite spec_w_sub; auto with zarith.
rewrite Hv6; rewrite spec_w_zdigits; auto with zarith.
assert (Hv7: 0 < [[ww_head1 x]]/2); auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
assert ([[ww_head1 x]]/2 <= Zpos w_digits); auto with zarith.
apply Zmult_le_reg_r with 2; auto with zarith.
diff --git a/theories/Ints/num/GenSub.v b/theories/Ints/num/GenSub.v
index 43661edd5..9b0924853 100644
--- a/theories/Ints/num/GenSub.v
+++ b/theories/Ints/num/GenSub.v
@@ -10,7 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Basic_type.
Require Import GenBase.
@@ -221,7 +220,7 @@ Section GenSub.
rewrite H0;rewrite Zplus_0_r; rewrite Zpower_2;
rewrite Zmult_mod_distr_r;try apply lt_0_wB.
rewrite spec_opp;trivial.
- apply Zmod_unique with (q:= -1). apply lt_0_wwB.
+ apply Zmod_unique with (q:= -1).
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW (w_opp_carry xh) l)).
rewrite spec_opp_carry;rewrite wwB_wBwB;ring.
Qed.
@@ -295,12 +294,12 @@ Section GenSub.
Lemma spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB.
Proof.
destruct x as [ |xh xl];simpl.
- apply Zmod_unique with (-1). apply lt_0_wwB. apply spec_ww_to_Z;trivial.
+ apply Zmod_unique with (-1). apply spec_ww_to_Z;trivial.
rewrite spec_ww_Bm1;ring.
replace ([|xh|]*wB + [|xl|] - 1) with ([|xh|]*wB + ([|xl|] - 1)). 2:ring.
generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l];intro H;
unfold interp_carry in H;rewrite <- H;simpl ww_to_Z.
- rewrite Zmod_def_small. apply spec_w_WW.
+ rewrite Zmod_small. apply spec_w_WW.
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh l)).
rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
change ([|xh|] + -1) with ([|xh|] - 1).
@@ -313,7 +312,7 @@ Section GenSub.
Lemma spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
Proof.
destruct y as [ |yh yl];simpl.
- ring_simplify ([[x]] - 0);rewrite Zmod_def_small;trivial. apply spec_ww_to_Z;trivial.
+ ring_simplify ([[x]] - 0);rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
destruct x as [ |xh xl];simpl. exact (spec_ww_opp (WW yh yl)).
replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]))
with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|])). 2:ring.
@@ -331,7 +330,7 @@ Section GenSub.
destruct y as [ |yh yl];simpl.
ring_simplify ([[x]] - 0);exact (spec_ww_pred x).
destruct x as [ |xh xl];simpl.
- apply Zmod_unique with (-1). apply lt_0_wwB.
+ apply Zmod_unique with (-1).
apply spec_ww_to_Z;trivial.
fold (ww_opp_carry (WW yh yl)).
rewrite (spec_ww_opp_carry (WW yh yl));simpl ww_to_Z;ring.
diff --git a/theories/Ints/num/Nbasic.v b/theories/Ints/num/Nbasic.v
index f7731ae6a..0d85c92ed 100644
--- a/theories/Ints/num/Nbasic.v
+++ b/theories/Ints/num/Nbasic.v
@@ -1,6 +1,5 @@
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Basic_type.
Require Import Max.
Require Import GenBase.
@@ -28,13 +27,13 @@ assert (tmp: (forall p, 2 * p = p + p)%Z);
intros p1 Hp1; rewrite F; rewrite (Zpos_xO p1).
assert (tmp: (forall p, 2 * p = p + p)%Z);
try repeat rewrite tmp; auto with zarith.
-rewrite ZPowerAux.Zpower_exp_1; auto with zarith.
+rewrite Zpower_1_r; auto with zarith.
Qed.
Theorem plength_pred_correct: forall p, (Zpos p <= 2 ^ Zpos (plength (Ppred p)))%Z.
intros p; case (Psucc_pred p); intros H1.
subst; simpl plength.
-rewrite ZPowerAux.Zpower_exp_1; auto with zarith.
+rewrite Zpower_1_r; auto with zarith.
pattern p at 1; rewrite <- H1.
rewrite Zpos_succ_morphism; unfold Zsucc; auto with zarith.
generalize (plength_correct (Ppred p)); auto with zarith.
@@ -296,7 +295,7 @@ Section CompareRec.
Lemma base_xO: forall n, base (xO n) = (base n)^2.
Proof.
intros n1; unfold base.
- rewrite (Zpos_xO n1); rewrite Zmult_comm; rewrite ZAux.Zpower_mult; auto with zarith.
+ rewrite (Zpos_xO n1); rewrite Zmult_comm; rewrite Zpower_mult; auto with zarith.
Qed.
Let gen_to_Z_pos: forall n x, 0 <= gen_to_Z n x < gen_wB n :=
diff --git a/theories/Ints/num/Q0Make.v b/theories/Ints/num/Q0Make.v
index 326e62902..c39a63301 100644
--- a/theories/Ints/num/Q0Make.v
+++ b/theories/Ints/num/Q0Make.v
@@ -2,7 +2,6 @@ Require Import Bool.
Require Import ZArith.
Require Import Znumtheory.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Arith.
Require Export BigN.
Require Export BigZ.
diff --git a/theories/Ints/num/QbiMake.v b/theories/Ints/num/QbiMake.v
index 53fb65b2a..fdf864707 100644
--- a/theories/Ints/num/QbiMake.v
+++ b/theories/Ints/num/QbiMake.v
@@ -2,7 +2,6 @@ Require Import Bool.
Require Import ZArith.
Require Import Znumtheory.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Arith.
Require Export BigN.
Require Export BigZ.
diff --git a/theories/Ints/num/QifMake.v b/theories/Ints/num/QifMake.v
index add89898a..5f461aa59 100644
--- a/theories/Ints/num/QifMake.v
+++ b/theories/Ints/num/QifMake.v
@@ -2,7 +2,6 @@ Require Import Bool.
Require Import ZArith.
Require Import Znumtheory.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Arith.
Require Export BigN.
Require Export BigZ.
diff --git a/theories/Ints/num/QpMake.v b/theories/Ints/num/QpMake.v
index 3c859d0f1..906cf8d15 100644
--- a/theories/Ints/num/QpMake.v
+++ b/theories/Ints/num/QpMake.v
@@ -2,7 +2,6 @@ Require Import Bool.
Require Import ZArith.
Require Import Znumtheory.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Arith.
Require Export BigN.
Require Export BigZ.
diff --git a/theories/Ints/num/QvMake.v b/theories/Ints/num/QvMake.v
index eb97123da..e51291a04 100644
--- a/theories/Ints/num/QvMake.v
+++ b/theories/Ints/num/QvMake.v
@@ -2,7 +2,6 @@ Require Import Bool.
Require Import ZArith.
Require Import Znumtheory.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Arith.
Require Export BigN.
Require Export BigZ.
diff --git a/theories/Ints/num/ZMake.v b/theories/Ints/num/ZMake.v
index 0a2b5cd3f..d00f14ec6 100644
--- a/theories/Ints/num/ZMake.v
+++ b/theories/Ints/num/ZMake.v
@@ -1,6 +1,5 @@
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
Open Scope Z_scope.
diff --git a/theories/Ints/num/Zn2Z.v b/theories/Ints/num/Zn2Z.v
index db3f28b41..a244635ea 100644
--- a/theories/Ints/num/Zn2Z.v
+++ b/theories/Ints/num/Zn2Z.v
@@ -10,7 +10,6 @@ Set Implicit Arguments.
Require Import ZArith.
Require Import ZAux.
-Require Import ZDivModAux.
Require Import Basic_type.
Require Import GenBase.
Require Import GenAdd.
@@ -588,7 +587,7 @@ Section Zn2Z.
_ _ _ _ _ _ _ _ _ _ _ _ _ _ _);auto.
unfold w_Bm2, w_to_Z, w_pred, w_Bm1.
rewrite (spec_pred op_spec);rewrite (spec_Bm1 op_spec).
- unfold w_digits;rewrite Zmod_def_small. ring.
+ unfold w_digits;rewrite Zmod_small. ring.
assert (H:= wB_pos(znz_digits w_op)). omega.
exact (spec_WW op_spec). exact (spec_compare op_spec).
exact (spec_mul_c op_spec). exact (spec_div21 op_spec).
@@ -621,12 +620,12 @@ Section Zn2Z.
w_to_Z (low x) = [|x|] mod wB.
intros x; case x; simpl low.
unfold ww_to_Z, w_to_Z, w_0; rewrite (spec_0 op_spec); simpl.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
split; auto with zarith.
unfold wB, base; auto with zarith.
intros xh xl; simpl.
rewrite Zplus_comm; rewrite Z_mod_plus; auto with zarith.
- rewrite Zmod_def_small; auto with zarith.
+ rewrite Zmod_small; auto with zarith.
unfold wB, base; auto with zarith.
Qed.
diff --git a/theories/Ints/num/ZnZ.v b/theories/Ints/num/ZnZ.v
index 68f1092cd..61d1ced60 100644
--- a/theories/Ints/num/ZnZ.v
+++ b/theories/Ints/num/ZnZ.v
@@ -8,7 +8,6 @@
Set Implicit Arguments.
-Require Import Tactic.
Require Import ZArith.
Require Import Znumtheory.
Require Import Basic_type.
diff --git a/theories/Numbers/Integer/TreeMod/ZTreeMod.v b/theories/Numbers/Integer/TreeMod/ZTreeMod.v
index 2d63a22fa..7ee894ca4 100644
--- a/theories/Numbers/Integer/TreeMod/ZTreeMod.v
+++ b/theories/Numbers/Integer/TreeMod/ZTreeMod.v
@@ -79,7 +79,8 @@ Notation "x * y" := (NZtimes x y) : IntScope.
Theorem gt_wB_1 : 1 < wB.
Proof.
-unfold base. apply Zgt_pow_1; unfold Zlt; auto with zarith.
+unfold base.
+apply Zpower_gt_1; unfold Zlt; auto with zarith.
Qed.
Theorem gt_wB_0 : 0 < wB.
@@ -90,22 +91,22 @@ Qed.
Lemma NZsucc_mod_wB : forall n : Z, (n + 1) mod wB = ((n mod wB) + 1) mod wB.
Proof.
intro n.
-pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zmod_plus; [ | apply gt_wB_0].
+pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zplus_mod.
reflexivity.
-now rewrite Zmod_def_small; [ | split; [auto with zarith | apply gt_wB_1]].
+now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]].
Qed.
Lemma NZpred_mod_wB : forall n : Z, (n - 1) mod wB = ((n mod wB) - 1) mod wB.
Proof.
intro n.
-pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zmod_minus; [ | apply gt_wB_0].
+pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zminus_mod.
reflexivity.
-now rewrite Zmod_def_small; [ | split; [auto with zarith | apply gt_wB_1]].
+now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]].
Qed.
Lemma NZ_to_Z_mod : forall n : NZ, [| n |] mod wB = [| n |].
Proof.
-intro n; rewrite Zmod_def_small. reflexivity. apply w_spec.(spec_to_Z).
+intro n; rewrite Zmod_small. reflexivity. apply w_spec.(spec_to_Z).
Qed.
Theorem NZpred_succ : forall n : NZ, P (S n) == n.
@@ -147,7 +148,7 @@ setoid_replace (Z_to_NZ (n + 1)) with (S (Z_to_NZ n)) using relation NZeq. assum
unfold NZeq. rewrite w_spec.(spec_succ).
unfold NZ_to_Z, Z_to_NZ.
do 2 (rewrite znz_of_Z_correct; [ | exact w_spec | auto with zarith]).
-symmetry; apply Zmod_def_small; auto with zarith.
+symmetry; apply Zmod_small; auto with zarith.
Qed.
Lemma B_holds : forall n : Z, 0 <= n < wB -> B n.
@@ -172,7 +173,7 @@ End Induction.
Theorem NZplus_0_l : forall n : NZ, 0 + n == n.
Proof.
intro n; unfold NZplus, NZ0, NZeq. rewrite w_spec.(spec_add). rewrite w_spec.(spec_0).
-rewrite Zplus_0_l. rewrite Zmod_def_small; [reflexivity | apply w_spec.(spec_to_Z)].
+rewrite Zplus_0_l. rewrite Zmod_small; [reflexivity | apply w_spec.(spec_to_Z)].
Qed.
Theorem NZplus_succ_l : forall n m : NZ, (S n) + m == S (n + m).
@@ -180,7 +181,7 @@ Proof.
intros n m; unfold NZplus, NZsucc, NZeq. rewrite w_spec.(spec_add).
do 2 rewrite w_spec.(spec_succ). rewrite w_spec.(spec_add).
rewrite NZsucc_mod_wB. repeat rewrite Zplus_mod_idemp_l; try apply gt_wB_0.
-rewrite <- (Zplus_assoc ([| n |] mod wB) 1 [| m |]). rewrite Zplus_mod_idemp_l; [ |apply gt_wB_0].
+rewrite <- (Zplus_assoc ([| n |] mod wB) 1 [| m |]). rewrite Zplus_mod_idemp_l.
rewrite (Zplus_comm 1 [| m |]); now rewrite Zplus_assoc.
Qed.
@@ -194,8 +195,8 @@ Theorem NZminus_succ_r : forall n m : NZ, n - (S m) == P (n - m).
Proof.
intros n m; unfold NZminus, NZsucc, NZpred, NZeq.
rewrite w_spec.(spec_pred). do 2 rewrite w_spec.(spec_sub).
-rewrite w_spec.(spec_succ). rewrite Zminus_mod_idemp_r; [ | apply gt_wB_0].
-rewrite Zminus_mod_idemp_l; [ | apply gt_wB_0].
+rewrite w_spec.(spec_succ). rewrite Zminus_mod_idemp_r.
+rewrite Zminus_mod_idemp_l.
now replace ([|n|] - ([|m|] + 1))%Z with ([|n|] - [|m|] - 1)%Z by auto with zarith.
Qed.
@@ -209,7 +210,7 @@ Theorem NZtimes_succ_r : forall n m : NZ, n * (S m) == n * m + n.
Proof.
intros n m; unfold NZtimes, NZsucc, NZplus, NZeq. rewrite w_spec.(spec_mul).
rewrite w_spec.(spec_add). rewrite w_spec.(spec_mul). rewrite w_spec.(spec_succ).
-rewrite Zplus_mod_idemp_l; [ | apply gt_wB_0]. rewrite Zmult_mod_idemp_r; [ | apply gt_wB_0].
+rewrite Zplus_mod_idemp_l. rewrite Zmult_mod_idemp_r.
rewrite Zmult_plus_distr_r. now rewrite Zmult_1_r.
Qed.
diff --git a/theories/QArith/Qpower.v b/theories/QArith/Qpower.v
index 8fa680a72..54cf4bcbd 100644
--- a/theories/QArith/Qpower.v
+++ b/theories/QArith/Qpower.v
@@ -1,4 +1,4 @@
-Require Import Qfield Qreduction.
+Require Import Zpow_facts Qfield Qreduction.
Lemma Qpower_positive_1 : forall n, Qpower_positive 1 n == 1.
Proof.
diff --git a/theories/QArith/Qround.v b/theories/QArith/Qround.v
index c5b06af7f..2479c2259 100644
--- a/theories/QArith/Qround.v
+++ b/theories/QArith/Qround.v
@@ -35,7 +35,7 @@ Proof.
intros z.
unfold Qceiling.
simpl.
-rewrite Z_div_1.
+rewrite Zdiv_1_r.
auto with *.
Qed.
diff --git a/theories/ZArith/Zabs.v b/theories/ZArith/Zabs.v
index e43d68bfa..ab699f4a7 100644
--- a/theories/ZArith/Zabs.v
+++ b/theories/ZArith/Zabs.v
@@ -13,6 +13,8 @@ Require Import Arith_base.
Require Import BinPos.
Require Import BinInt.
Require Import Zorder.
+Require Import Zmax.
+Require Import Znat.
Require Import ZArith_dec.
Open Local Scope Z_scope.
@@ -63,6 +65,11 @@ 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;
@@ -96,16 +103,6 @@ Proof.
apply Zplus_le_compat; simpl in |- *; auto with zarith.
Qed.
-(** * A characterization of the sign function: *)
-
-Lemma Zsgn_spec : forall x:Z,
- 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.
-Qed.
-
(** * Absolute value and multiplication *)
Lemma Zsgn_Zabs : forall n:Z, n * Zsgn n = Zabs n.
@@ -123,25 +120,90 @@ Proof.
intros z1 z2; case z1; case z2; simpl in |- *; auto.
Qed.
-(** * Absolute value in nat is compatible with order *)
+Theorem Zabs_square : forall a, Zabs a * Zabs a = a * a.
+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 /\ n < m -> (Zabs_nat n < Zabs_nat m)%nat.
+ forall n m:Z, 0 <= n < m -> (Zabs_nat n < Zabs_nat m)%nat.
Proof.
- intros x y. case x; simpl in |- *. case y; simpl in |- *.
+ 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.
+Qed.
- intro. absurd (0 < 0). compute in |- *. intro H0. discriminate H0. intuition.
- intros. elim (ZL4 p). intros. rewrite H0. auto with arith.
- intros. elim (ZL4 p). intros. rewrite H0. auto with arith.
-
- case y; simpl in |- *.
- intros. absurd (Zpos p < 0). compute in |- *. intro H0. discriminate H0. intuition.
- intros. change (nat_of_P p > nat_of_P p0)%nat in |- *.
- apply nat_of_P_gt_Gt_compare_morphism.
- elim H; auto with arith. intro. exact (ZC2 p0 p).
+(** * Some results about the sign function. *)
- intros. absurd (Zpos p0 < Zneg p).
- compute in |- *. intro H0. discriminate H0. intuition.
+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.
+
+(** A characterization of the sign function: *)
- intros. absurd (0 <= Zneg p). compute in |- *. auto with arith. intuition.
+Lemma Zsgn_spec : forall x:Z,
+ 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.
Qed.
diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v
index 57c0af02f..258ae1d13 100644
--- a/theories/ZArith/Zdiv.v
+++ b/theories/ZArith/Zdiv.v
@@ -70,8 +70,21 @@ Unboxed Fixpoint Zdiv_eucl_POS (a:positive) (b:Z) {struct a} :
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|.
+ smaller or equal to a/b. And sgn(r)=sgn(b) and |r| < |b| (at least when
+ r is not null).
+*)
+
+(* 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.
+
+ * 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).
+ * 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 :=
@@ -96,7 +109,7 @@ Definition Zdiv_eucl (a b:Z) : Z * Z :=
(** 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.
@@ -108,20 +121,20 @@ 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).
-Eval Compute in `(Zdiv_eucl (-7) 3)`.
+Eval compute in (Zdiv_eucl (-7) 3).
-Eval Compute in `(Zdiv_eucl 7 (-3))`.
+Eval compute in (Zdiv_eucl 7 (-3)).
-Eval Compute in `(Zdiv_eucl (-7) (-3))`.
+Eval compute in (Zdiv_eucl (-7) (-3)).
*)
(** * Main division theorem *)
-(** First a lemma for positive *)
+(** First a lemma for two positive arguments *)
Lemma Z_div_mod_POS :
forall b:Z,
@@ -148,6 +161,7 @@ case (Zge_bool b 2); (intros; split; [ try ring | omega ]).
omega.
Qed.
+(** Then the usual situation of a positive [b] and no restriction on [a] *)
Theorem Z_div_mod :
forall a b:Z,
@@ -167,129 +181,299 @@ Proof.
intros [H1 H2].
split; trivial.
- replace (Zneg p0) with (- Zpos p0); [ rewrite H1; ring | trivial ].
+ change (Zneg p0) with (- Zpos p0); rewrite H1; ring.
intros p1 [H1 H2].
split; trivial.
- replace (Zneg p0) with (- Zpos p0); [ rewrite H1; ring | trivial ].
+ change (Zneg p0) with (- Zpos p0); rewrite H1; ring.
generalize (Zorder.Zgt_pos_0 p1); omega.
intros p1 [H1 H2].
split; trivial.
- replace (Zneg p0) with (- Zpos p0); [ rewrite H1; ring | trivial ].
+ change (Zneg p0) with (- Zpos p0); rewrite H1; ring.
generalize (Zorder.Zlt_neg_0 p1); omega.
intros; discriminate.
Qed.
-(** Existence theorems *)
+(** For stating the fully general result, let's give a short name
+ to the condition on the remainder. *)
-Theorem Zdiv_eucl_exist :
- forall b:Z,
- b > 0 ->
- forall a:Z, {qr : Z * Z | let (q, r) := qr in a = b * q + r /\ 0 <= r < b}.
+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.
+
+(* 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. *)
+
+Lemma Remainder_equiv : forall r b, Remainder r b <-> Remainder_alt r b.
Proof.
- intros b Hb a.
- exists (Zdiv_eucl a b).
- exact (Z_div_mod a b Hb).
+ intros; unfold Remainder, Remainder_alt; omega with *.
Qed.
-Implicit Arguments Zdiv_eucl_exist.
+Hint Unfold Remainder.
-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}.
+(** 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.
+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 *.
+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.
+Proof.
+ unfold Zmod; intros a b Hb; generalize (Z_div_mod a b Hb).
+ destruct Zdiv_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_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_div_mod_eq : forall a b:Z, b > 0 -> a = b*(a/b) + (a mod b).
+Proof.
+ intros; apply Z_div_mod_eq_full; auto with zarith.
+Qed.
+
+(** Existence theorem *)
+
+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.
- elim (Z_le_gt_dec 0 b); intro Hb'.
- cut (b > 0); [ intro Hb'' | omega ].
- rewrite Zabs_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.
- exists (- q, r).
- elim Hqr; intros.
- split.
- rewrite <- Zmult_opp_comm; assumption.
- rewrite Zabs_non_eq; [ assumption | omega ].
+ exists (Zdiv_eucl a b).
+ exact (Z_div_mod a b Hb).
Qed.
-Implicit Arguments Zdiv_eucl_extended.
+Implicit Arguments Zdiv_eucl_exist.
+
(** Uniqueness theorems *)
Theorem Zdiv_mod_unique :
forall b q1 q2 r1 r2:Z,
0 <= r1 < Zabs b -> 0 <= r2 < Zabs b ->
- b*q1 + r1=b*q2+r2 -> q1=q2/\r1=r2.
+ b*q1+r1 = b*q2+r2 -> q1=q2 /\ r1=r2.
Proof.
intros b q1 q2 r1 r2 Hr1 Hr2 H.
-assert (H0:=Zeq_minus _ _ H).
-replace (b * q1 + r1 - (b * q2 + r2))
- with (b*(q1-q2) - (r2-r1)) in H0 by ring.
-assert (H1:=Zminus_eq _ _ H0).
-destruct (Z_dec q1 q2).
- elim (Zlt_not_le (Zabs (r2 - r1)) (Zabs b)).
- apply Zabs_ind; omega.
- rewrite <- H1.
- replace (Zabs b) with (1*(Zabs b)) by ring.
- rewrite Zabs_Zmult.
- rewrite (Zmult_comm (Zabs b)).
- apply Zmult_le_compat_r; auto with *.
- change 1 with (Zsucc 0).
- apply Zlt_le_succ.
- destruct s; apply Zabs_ind; omega.
-split; try assumption.
-rewrite e in H1.
-ring_simplify in H1.
-omega.
+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.
-(** * Auxiliary lemmas about [Zdiv] and [Zmod] *)
+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.
-Lemma Z_div_mod_eq : forall a b:Z, b > 0 -> a = b * Zdiv a b + Zmod a b.
+Theorem Zdiv_unique_full:
+ forall a b q r, Remainder r b ->
+ a = b*q + r -> q = a/b.
Proof.
- unfold Zdiv, Zmod in |- *.
- intros a b Hb.
- generalize (Z_div_mod a b Hb).
- case Zdiv_eucl; tauto.
+ 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.
-Lemma Z_mod_lt : forall a b:Z, b > 0 -> 0 <= Zmod a b < b.
+Theorem Zdiv_unique:
+ forall a b q r, 0 <= r < b ->
+ a = b*q + r -> q = a/b.
Proof.
- unfold Zmod in |- *.
- intros a b Hb.
- generalize (Z_div_mod a b Hb).
- case (Zdiv_eucl a b); tauto.
+ intros; eapply Zdiv_unique_full; eauto.
Qed.
-Lemma Z_div_POS_ge0 :
- forall (b:Z) (a:positive), let (q, _) := Zdiv_eucl_POS a b in q >= 0.
+Theorem Zmod_unique_full:
+ forall a b q r, Remainder r b ->
+ a = b*q + r -> r = a mod 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; case (Zgt_bool b (2 * z0 + 1)); intros; omega.
- intro p; case (Zdiv_eucl_POS p b).
- intros; case (Zgt_bool b (2 * z0)); intros; omega.
- case (Zge_bool b 2); simpl in |- *; omega.
+ 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.
-Lemma Z_div_ge0 : forall a b:Z, b > 0 -> a >= 0 -> Zdiv a b >= 0.
+Theorem Zmod_unique:
+ forall a b q r, 0 <= r < b ->
+ a = b*q + r -> r = a mod b.
Proof.
- intros a b Hb; unfold Zdiv, Zdiv_eucl in |- *; case a; simpl in |- *; intros.
- case b; simpl in |- *; trivial.
- generalize Hb; case b; try trivial.
- auto with zarith.
- intros p0 Hp0; generalize (Z_div_POS_ge0 (Zpos p0) p).
- case (Zdiv_eucl_POS p (Zpos p0)); simpl in |- *; tauto.
- intros; discriminate.
- elim H; trivial.
+ intros; eapply Zmod_unique_full; eauto.
+Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma Zmod_0_l: forall a, 0 mod a = 0.
+Proof.
+ destruct a; simpl; auto.
+Qed.
+
+Lemma Zmod_0_r: forall a, a mod 0 = 0.
+Proof.
+ destruct a; simpl; auto.
Qed.
-Lemma Z_div_lt : forall a b:Z, b >= 2 -> a > 0 -> Zdiv a b < a.
+Lemma Zdiv_0_l: forall a, 0/a = 0.
+Proof.
+ destruct a; simpl; auto.
+Qed.
+
+Lemma Zdiv_0_r: forall a, a/0 = 0.
+Proof.
+ destruct a; simpl; auto.
+Qed.
+
+Lemma Zmod_1_r: forall a, a mod 1 = 0.
+Proof.
+ intros; symmetry; apply Zmod_unique with a; auto with zarith.
+Qed.
+
+Lemma Zdiv_1_r: forall a, a/1 = a.
+Proof.
+ intros; symmetry; apply Zdiv_unique with 0; auto with zarith.
+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.
+
+Lemma Zmod_1_l: forall a, 1 < a -> 1 mod a = 1.
+Proof.
+ intros; symmetry; apply Zmod_unique with 0; auto with zarith.
+Qed.
+
+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.
+
+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.
+
+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.
+
+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.
+
+(** * 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.
+
+Lemma Z_div_ge0: forall a b, b > 0 -> a >= 0 -> a/b >=0.
+Proof.
+ intros; generalize (Z_div_pos a b H); auto with zarith.
+Qed.
+
+(** As soon as the divisor is greater or equal than 2,
+ 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).
@@ -302,9 +486,23 @@ Proof.
auto with zarith.
Qed.
-(** * Other lemmas (now using the syntax for [Zdiv] and [Zmod]). *)
+(** A division of a small number by a bigger one yields zero. *)
-Lemma Z_div_ge : forall a b c:Z, c > 0 -> a >= b -> a / c >= b / c.
+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.
+
+(** 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.
+
+(** [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).
@@ -316,13 +514,8 @@ Proof.
intro.
absurd (b - a >= 1).
omega.
- rewrite H0.
- rewrite H2.
- assert
- (c * (b / c) + b mod c - (c * (a / c) + a mod c) =
- c * (b / c - a / c) + b mod c - a mod c).
- ring.
- rewrite H3.
+ 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.
@@ -332,180 +525,640 @@ Proof.
omega.
Qed.
-Lemma Z_div_le : forall a b c:Z, c > 0 -> a <= b -> a / c <= b / c.
+(** 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 *.
+ intros a b c H H0.
+ apply Zge_le.
+ apply Z_div_ge; auto with *.
Qed.
-Lemma Z_mod_plus : forall a b c:Z, c > 0 -> (a + b * c) mod c = a mod c.
+(** 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 c cPos.
- generalize (Z_div_mod_eq a c cPos).
- generalize (Z_mod_lt a c cPos).
- generalize (Z_div_mod_eq (a + b * c) c cPos).
- generalize (Z_mod_lt (a + b * c) c cPos).
- intros.
+ intros a b H; generalize (Z_div_mod_eq a b H) (Z_mod_lt a b H); omega.
+Qed.
- assert ((a + b * c) mod c - a mod c = c * (b + a / c - (a + b * c) / c)).
- replace ((a + b * c) mod c) with (a + b * c - c * ((a + b * c) / c)).
- replace (a mod c) with (a - c * (a / c)).
- ring.
- omega.
- omega.
- set (q := b + a / c - (a + b * c) / c) in *.
- apply (Zcase_sign q); intros.
- assert (c * q = 0).
- rewrite H4; ring.
- rewrite H5 in H3.
- omega.
+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.
- assert (c * q >= c).
- pattern c at 2 in |- *; replace c with (c * 1).
- apply Zmult_ge_compat_l; omega.
- ring.
- omega.
+(** 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.
+
+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.
+
+(** 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.
+
+(** Some additionnal inequalities about Zdiv. *)
+
+Theorem Zdiv_le_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_le_reg_r with b; auto with zarith.
+ apply Zle_trans with (2 := H3).
+ 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.
+
+Theorem Zdiv_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 (Z_div_mod_eq a b); auto with zarith.
+ rewrite (Zmult_comm b); case (Z_mod_lt a b); auto with zarith.
+Qed.
+
+Theorem Zdiv_le_lower_bound:
+ forall a b q, 0 <= a -> 0 < b -> q*b <= a -> q <= a/b.
+Proof.
+ intros a b q H1 H2 H3.
+ assert (q < a / b + 1); auto with zarith.
+ apply Zmult_lt_reg_r with b; auto with zarith.
+ apply Zle_lt_trans with (1 := H3).
+ pattern a at 1; rewrite (Z_div_mod_eq a b); auto with zarith.
+ rewrite Zmult_plus_distr_l; rewrite (Zmult_comm b); case (Z_mod_lt a b);
+ auto with zarith.
+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.
+
+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.
+
+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.
- assert (c * q <= - c).
- replace (- c) with (c * -1).
- apply Zmult_le_compat_l; omega.
+(** [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.
+
+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.
+
+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.
- omega.
Qed.
-Lemma Z_div_plus : forall a b c:Z, c > 0 -> (a + b * c) / c = a / c + b.
+Lemma Z_mod_nz_opp_full : forall a b:Z, a mod b <> 0 ->
+ (-a) mod b = b - (a mod b).
Proof.
- intros a b c cPos.
- generalize (Z_div_mod_eq a c cPos).
- generalize (Z_mod_lt a c cPos).
- generalize (Z_div_mod_eq (a + b * c) c cPos).
- generalize (Z_mod_lt (a + b * c) c cPos).
intros.
- apply Zmult_reg_l with c. omega.
- replace (c * ((a + b * c) / c)) with (a + b * c - (a + b * c) mod c).
- rewrite (Z_mod_plus a b c cPos).
- pattern a at 1 in |- *; rewrite H2.
- ring.
- pattern (a + b * c) at 1 in |- *; rewrite H0.
- ring.
+ assert (b<>0) by (swap 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.
-Lemma Zdiv_opp_opp : forall a b:Z, Zdiv (-a) (-b) = Zdiv a b.
+Lemma Z_mod_zero_opp_r : forall a b:Z, a mod b = 0 -> a mod (-b) = 0.
Proof.
-intros [|a|a] [|b|b];
- try reflexivity;
- unfold Zdiv;
- simpl;
- destruct (Zdiv_eucl_POS a (Zpos b)); destruct z0; try reflexivity.
+ intros.
+ rewrite <- (Zopp_involutive a).
+ rewrite Zmod_opp_opp.
+ rewrite Z_mod_zero_opp_full; auto.
Qed.
-Lemma Z_div_mult : forall a b:Z, b > 0 -> a * b / b = a.
- intros; replace (a * b) with (0 + a * b); auto.
- rewrite Z_div_plus; auto.
+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.
-Lemma Zdiv_mult_cancel_r : forall (a b c:Z),
- (c <> 0) -> (a*c/(b*c) = a/b).
+Lemma Z_div_zero_opp_full : forall a b:Z, a mod b = 0 -> (-a)/b = -(a/b).
Proof.
-assert (X: forall a b c, b > 0 -> c > 0 -> a * c / (b * c) = a / b).
+ 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.
+
+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 (swap 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.
+
+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.
+
+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.
+
+(** Cancellations. *)
+
+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.
- assert (X:=Z_div_mod (a*c) (b*c)).
- assert (Y:=Z_div_mod a b Hb).
- unfold Zdiv.
- destruct (Zdiv_eucl (a*c) (b*c)).
- destruct (Zdiv_eucl a b).
- destruct Y as [Y0 Y1].
- rewrite Y0 in X.
- clear Y0.
- destruct X as [X0 X1].
- auto with *.
- replace ((b*z1 + z2)* c) with
- ((b*c)*z1 + z2*c) in X0 by ring.
- destruct (Zdiv_mod_unique (b*c) z1 z (z2*c) z0); try rewrite Zabs_eq; auto with *.
- split; auto with *.
- apply Zmult_lt_compat_r; auto with *.
+ 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).
- rewrite <- (Zdiv_opp_opp a b).
- replace (b*c) with ((-b)*(-c)) by ring.
- replace (a*c) with ((-a)*(-c)) by ring.
- apply X; auto with *.
- rewrite <- (Zdiv_opp_opp a b).
- rewrite <- Zdiv_opp_opp.
- replace (-(b*c)) with ((-b)*c) by ring.
- replace (-(a*c)) with ((-a)*c) by ring.
- apply X; auto with *.
- rewrite <- Zdiv_opp_opp.
- replace (-(b*c)) with (b*(-c)) by ring.
- replace (-(a*c)) with (a*(-c)) by ring.
- apply X; auto with *.
- apply X; auto with *.
-rewrite Hb.
-destruct (a*c); destruct a; reflexivity.
-Qed.
-
-Lemma Z_mult_div_ge : forall a b:Z, b > 0 -> b * (a / b) <= a.
-Proof.
- intros a b bPos.
- generalize (Z_div_mod_eq a _ bPos); intros.
- generalize (Z_mod_lt a _ bPos); intros.
- pattern a at 2 in |- *; rewrite H.
- omega.
+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.
-Lemma Z_mod_same : forall a:Z, a > 0 -> a mod a = 0.
+Lemma Zdiv_mult_cancel_l : forall a b c:Z,
+ c<>0 -> (c*a)/(c*b) = a/b.
Proof.
- intros a aPos.
- generalize (Z_mod_plus 0 1 a aPos).
- replace (0 + 1 * a) with a.
intros.
- rewrite H.
- compute in |- *.
- trivial.
+ rewrite (Zmult_comm c a); rewrite (Zmult_comm c b).
+ apply Zdiv_mult_cancel_r; 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.
Qed.
-Lemma Z_div_same : forall a:Z, a > 0 -> a / a = 1.
+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.
+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.
+
+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.
+
+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.
+
+Theorem Zminus_mod: forall a b n,
+ (a - b) mod n = (a mod n - b mod n) mod n.
Proof.
- intros a aPos.
- generalize (Z_div_plus 0 1 a aPos).
- replace (0 + 1 * a) with a.
intros.
- rewrite H.
- compute in |- *.
- trivial.
+ replace (a - b) with (a + (-1) * b); auto with zarith.
+ replace (a mod n - b mod n) with (a mod n + (-1) * (b mod n)); auto with zarith.
+ rewrite Zplus_mod.
+ rewrite Zmult_mod.
+ rewrite Zplus_mod with (b:=(-1) * (b mod n)).
+ rewrite Zmult_mod.
+ rewrite Zmult_mod with (b:= b mod n).
+ repeat rewrite Zmod_mod; auto.
+Qed.
+
+Lemma Zplus_mod_idemp_l: forall a b n, (a mod n + b) mod n = (a + b) mod n.
+Proof.
+ intros; rewrite Zplus_mod, Zmod_mod, <- Zplus_mod; auto.
+Qed.
+
+Lemma Zplus_mod_idemp_r: forall a b n, (b + a mod n) mod n = (b + a) mod n.
+Proof.
+ intros; rewrite Zplus_mod, Zmod_mod, <- Zplus_mod; auto.
+Qed.
+
+Lemma Zminus_mod_idemp_l: forall a b n, (a mod n - b) mod n = (a - b) mod n.
+Proof.
+ intros; rewrite Zminus_mod, Zmod_mod, <- Zminus_mod; auto.
+Qed.
+
+Lemma Zminus_mod_idemp_r: forall a b n, (a - b mod n) mod n = (a - b) mod n.
+Proof.
+ intros; rewrite Zminus_mod, Zmod_mod, <- Zminus_mod; auto.
+Qed.
+
+Lemma Zmult_mod_idemp_l: forall a b n, (a mod n * b) mod n = (a * b) mod n.
+Proof.
+ intros; rewrite Zmult_mod, Zmod_mod, <- Zmult_mod; auto.
+Qed.
+
+Lemma Zmult_mod_idemp_r: forall a b n, (b * (a mod n)) mod n = (b * a) mod n.
+Proof.
+ intros; rewrite Zmult_mod, Zmod_mod, <- Zmult_mod; auto.
+Qed.
+
+Section EqualityModulo.
+ (** For a specific number n, equality modulo n is hence a nice setoid
+ equivalence, compatible with the usual operations. But for the
+ moment, Coq setoids cannot be parametric, and all this nice framework
+ will vanish at the end of this section. *)
+
+ Variable n:Z.
+
+ 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 : forall a b, a == b -> b == a.
+ Proof. unfold eqm; auto. Qed.
+
+ Lemma eqm_trans : forall a b c, a == b -> b == c -> a == c.
+ Proof. unfold eqm; eauto with *. Qed.
+
+ Add Relation Z eqm
+ reflexivity proved by eqm_refl
+ symmetry proved by eqm_sym
+ transitivity proved by eqm_trans as eqm_setoid.
+
+ Add Morphism Zplus : Zplus_eqm.
+ Proof.
+ unfold eqm; intros; rewrite Zplus_mod, H, H0, <- Zplus_mod; auto.
+ Qed.
+
+ Add Morphism Zminus : Zminus_eqm.
+ Proof.
+ unfold eqm; intros; rewrite Zminus_mod, H, H0, <- Zminus_mod; auto.
+ Qed.
+
+ Add Morphism Zmult : Zmult_eqm.
+ Proof.
+ unfold eqm; intros; rewrite Zmult_mod, H, H0, <- Zmult_mod; auto.
+ Qed.
+
+ Add Morphism Zopp : Zopp_eqm.
+ Proof.
+ intros; change (-x1 == -x2) with (0-x1 == 0-x2).
+ rewrite H; red; auto.
+ Qed.
+
+ Lemma Zmod_eqm : forall a, a mod n == a.
+ Proof.
+ unfold eqm; intros; apply Zmod_mod.
+ Qed.
+
+ (* Zmod and Zdiv are not full morphisms with respect to eqm.
+ For instance, take n=2. Then 3 == 1 but we don't have
+ 1 mod 3 == 1 mod 1 nor 1/3 == 1/1.
+ *)
+
+End EqualityModulo.
+
+Lemma Zdiv_Zdiv : forall a b c, b>0 -> c>0 -> (a/b)/c = a/(b*c).
+Proof.
+ intros a b c H H0.
+ 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.
Qed.
-Lemma Z_div_1 : forall z:Z, (z/1 = z)%Z.
+(** Unfortunately, the previous result isn't always true on negative numbers.
+ For instance: 3/(-2)/(-2) = 1 <> 0 = 3 / (-2*-2) *)
+
+(** A last inequality: *)
+
+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.
+ 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.
+
+(** * Compatibility *)
+
+(** Weaker results kept only for compatibility *)
+
+Lemma Z_mod_same : forall a, a > 0 -> a mod a = 0.
+Proof.
+ intros; apply Z_mod_same_full.
+Qed.
+
+Lemma Z_div_same : forall a, a > 0 -> a/a = 1.
+Proof.
+ intros; apply Z_div_same_full; auto with zarith.
+Qed.
+
+Lemma Z_div_mult : forall a b:Z, b > 0 -> (a*b)/b = a.
Proof.
-intros z.
-set (z':=z).
-unfold z' at 1.
-replace z with (0 + z*1)%Z by ring.
-rewrite (Z_div_plus 0 z 1);[reflexivity|constructor].
+ intros; apply Z_div_mult_full; auto with zarith.
Qed.
-Hint Resolve Z_div_1 : zarith.
+Lemma Z_mod_plus : forall a b c:Z, c > 0 -> (a + b * c) mod c = a mod c.
+Proof.
+ intros; apply Z_mod_plus_full; auto with zarith.
+Qed.
-Lemma Z_div_exact_1 : forall a b:Z, b > 0 -> a = b * (a / b) -> a mod b = 0.
- intros a b Hb; generalize (Z_div_mod a b Hb); unfold Zmod, Zdiv in |- *.
- case (Zdiv_eucl a b); intros q r; omega.
+Lemma Z_div_exact_1 : forall a b:Z, b > 0 -> a = b*(a/b) -> a mod b = 0.
+Proof.
+ intros; apply Z_div_exact_full_1; auto with zarith.
Qed.
-Lemma Z_div_exact_2 : forall a b:Z, b > 0 -> a mod b = 0 -> a = b * (a / b).
- intros a b Hb; generalize (Z_div_mod a b Hb); unfold Zmod, Zdiv in |- *.
- case (Zdiv_eucl a b); intros q r; omega.
+Lemma Z_div_exact_2 : forall a b:Z, b > 0 -> a mod b = 0 -> a = b*(a/b).
+Proof.
+ intros; apply Z_div_exact_full_2; auto with zarith.
Qed.
-Lemma Z_mod_zero_opp : forall a b:Z, b > 0 -> a mod b = 0 -> - a mod b = 0.
- intros a b Hb.
+Lemma Z_mod_zero_opp : forall a b:Z, b > 0 -> a mod b = 0 -> (-a) mod b = 0.
+Proof.
+ intros; apply Z_mod_zero_opp_full; auto with zarith.
+Qed.
+
+(** * A direct way to compute Zmod *)
+
+Fixpoint Zmod_POS (a : positive) (b : Z) {struct a} : Z :=
+ match a with
+ | 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)
+ | 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
+ end.
+
+Definition Zmod' a b :=
+ match a with
+ | Z0 => 0
+ | Zpos a' =>
+ match b with
+ | Z0 => 0
+ | Zpos _ => Zmod_POS a' b
+ | Zneg b' =>
+ let r := Zmod_POS a' (Zpos b') in
+ match r with Z0 => 0 | _ => b + r end
+ end
+ | Zneg a' =>
+ match b with
+ | Z0 => 0
+ | Zpos _ =>
+ let r := Zmod_POS a' b in
+ match r with Z0 => 0 | _ => b - r end
+ | Zneg b' => - (Zmod_POS a' (Zpos b'))
+ end
+ end.
+
+
+Theorem Zmod_POS_correct: forall a b, Zmod_POS a b = (snd (Zdiv_eucl_POS 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.
+Qed.
+
+Theorem Zmod'_correct: forall a b, Zmod' a b = Zmod a b.
+Proof.
+ intros a b; unfold Zmod; 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.
+ 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.
+ intros p1; rewrite Zmod_POS_correct; simpl; auto.
+ case (Zdiv_eucl_POS p (Zpos p1)); auto.
+Qed.
+
+
+(** Another convention is possible for division by negative numbers:
+ * quotient is always the biggest integer smaller than or equal to a/b
+ * remainder is hence always positive or null. *)
+
+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}.
+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 ].
+ cut (- b > 0); [ intro Hb'' | omega ].
+ elim (Zdiv_eucl_exist Hb'' a); intros qr.
+ elim qr; intros q r Hqr.
+ exists (- q, r).
+ elim Hqr; intros.
+ split.
+ rewrite <- Zmult_opp_comm; assumption.
+ rewrite Zabs_non_eq; [ assumption | omega ].
+Qed.
+
+Implicit Arguments Zdiv_eucl_extended.
+
+(** A third convention: Ocaml.
+
+ 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).
+
+*)
+
+Definition Odiv a b :=
+ Zdiv (Zabs a) (Zabs b) * Zsgn a * Zsgn b.
+
+Definition Omod a b :=
+ Zmod (Zabs a) (Zabs b) * Zsgn a.
+
+Lemma Odiv_Omod_eq : forall a b, b<>0 ->
+ a = b*(Odiv a b) + (Omod a b).
+Proof.
intros.
- rewrite Z_div_exact_2 with a b; auto.
- replace (- (b * (a / b))) with (0 + - (a / b) * b).
- rewrite Z_mod_plus; auto.
+ assert (Zabs b <> 0).
+ swap H.
+ destruct b; simpl in *; auto with zarith; inversion H0.
+ pattern a at 1; rewrite <- (Zabs_Zsgn a).
+ rewrite (Z_div_mod_eq_full (Zabs a) (Zabs b) H0).
+ unfold Odiv, Omod.
+ replace (b*(Zabs a/Zabs b * Zsgn a * Zsgn b)) with
+ ((b*Zsgn b)*(Zabs a/Zabs b)*(Zsgn a)) by ring.
+ rewrite Zsgn_Zabs.
ring.
Qed.
+
+Lemma Odiv_opp_l : forall a b, b<>0 -> Odiv (-a) b = - (Odiv a b).
+Proof.
+ intros; unfold Odiv; rewrite Zabs_Zopp, Zsgn_Zopp; ring.
+Qed.
+
+Lemma Odiv_opp_r : forall a b, b<>0 -> Odiv a (-b) = - (Odiv a b).
+Proof.
+ intros; unfold Odiv; rewrite Zabs_Zopp, Zsgn_Zopp; ring.
+Qed.
+
+Lemma Odiv_opp_opp : forall a b, b<>0 -> Odiv (-a) (-b) = Odiv a b.
+Proof.
+ intros; unfold Odiv; do 2 rewrite Zabs_Zopp, Zsgn_Zopp; ring.
+Qed.
+
+Lemma Omod_opp_l : forall a b, b<>0 -> Omod (-a) b = - (Omod a b).
+Proof.
+ intros; unfold Omod; rewrite Zabs_Zopp, Zsgn_Zopp; ring.
+Qed.
+
+Lemma Omod_opp_r : forall a b, b<>0 -> Omod a (-b) = Omod a b.
+Proof.
+ intros; unfold Omod; rewrite Zabs_Zopp; ring.
+Qed.
+
+Lemma Omod_opp_opp : forall a b, b<>0 -> Omod (-a) (-b) = - (Omod a b).
+Proof.
+ intros; unfold Omod; do 2 rewrite Zabs_Zopp; rewrite Zsgn_Zopp; ring.
+Qed.
diff --git a/theories/ZArith/Zeven.v b/theories/ZArith/Zeven.v
index 0e5a03dd8..a0a75cf1e 100644
--- a/theories/ZArith/Zeven.v
+++ b/theories/ZArith/Zeven.v
@@ -10,6 +10,8 @@
Require Import BinInt.
+Open Scope Z_scope.
+
(*******************************************************************)
(** About parity: even and odd predicates on Z, division by 2 on Z *)
@@ -135,14 +137,14 @@ Hint Unfold Zeven Zodd: zarith.
Definition Zdiv2 (z:Z) :=
match z with
- | Z0 => 0%Z
- | Zpos xH => 0%Z
+ | Z0 => 0
+ | Zpos xH => 0
| Zpos p => Zpos (Pdiv2 p)
- | Zneg xH => 0%Z
+ | Zneg xH => 0
| Zneg p => Zneg (Pdiv2 p)
end.
-Lemma Zeven_div2 : forall n:Z, Zeven n -> n = (2 * Zdiv2 n)%Z.
+Lemma Zeven_div2 : forall n:Z, Zeven n -> n = 2 * Zdiv2 n.
Proof.
intro x; destruct x.
auto with arith.
@@ -154,27 +156,27 @@ Proof.
intros. absurd (Zeven (-1)); red in |- *; auto with arith.
Qed.
-Lemma Zodd_div2 : forall n:Z, (n >= 0)%Z -> Zodd n -> n = (2 * Zdiv2 n + 1)%Z.
+Lemma Zodd_div2 : forall n:Z, n >= 0 -> Zodd n -> n = 2 * Zdiv2 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)%Z; red in |- *; auto with arith.
+ intros. absurd (Zneg p >= 0); red in |- *; auto with arith.
Qed.
Lemma Zodd_div2_neg :
- forall n:Z, (n <= 0)%Z -> Zodd n -> n = (2 * Zdiv2 n - 1)%Z.
+ forall n:Z, n <= 0 -> Zodd n -> n = 2 * Zdiv2 n - 1.
Proof.
intro x; destruct x.
intros. absurd (Zodd 0); red in |- *; auto with arith.
- intros. absurd (Zneg p >= 0)%Z; 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.
Qed.
Lemma Z_modulo_2 :
- forall n:Z, {y : Z | n = (2 * y)%Z} + {y : Z | n = (2 * y + 1)%Z}.
+ forall n:Z, {y : Z | n = 2 * y} + {y : Z | n = 2 * y + 1}.
Proof.
intros x.
elim (Zeven_odd_dec x); intro.
@@ -193,7 +195,7 @@ Qed.
Lemma Zsplit2 :
forall n:Z,
{p : Z * Z |
- let (x1, x2) := p in n = (x1 + x2)%Z /\ (x1 = x2 \/ x2 = (x1 + 1)%Z)}.
+ let (x1, x2) := p in n = x1 + x2 /\ (x1 = x2 \/ x2 = x1 + 1)}.
Proof.
intros x.
elim (Z_modulo_2 x); intros [y Hy]; rewrite Zmult_comm in Hy;
@@ -206,3 +208,109 @@ Proof.
right; reflexivity.
Qed.
+
+Theorem Zeven_ex: forall n, Zeven n -> exists m, n = 2 * m.
+Proof.
+ intro n; exists (Zdiv2 n); apply Zeven_div2; auto.
+Qed.
+
+Theorem Zodd_ex: forall n, Zodd n -> exists m, n = 2 * m + 1.
+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).
+Qed.
+
+Theorem Zeven_2p: forall p, Zeven (2 * p).
+Proof.
+ destruct p; simpl; auto.
+Qed.
+
+Theorem Zodd_2p_plus_1: forall p, Zodd (2 * p + 1).
+Proof.
+ destruct p; simpl; auto.
+ destruct p; simpl; auto.
+Qed.
+
+Theorem Zeven_plus_Zodd: forall a b,
+ Zeven a -> Zodd b -> Zodd (a + b).
+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.
+Qed.
+
+Theorem Zeven_plus_Zeven: forall a b,
+ Zeven a -> Zeven b -> Zeven (a + b).
+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.
+Qed.
+
+Theorem Zodd_plus_Zeven: forall a b,
+ Zodd a -> Zeven b -> Zodd (a + b).
+Proof.
+ intros a b H1 H2; rewrite Zplus_comm; apply Zeven_plus_Zodd; auto.
+Qed.
+
+Theorem Zodd_plus_Zodd: forall a b,
+ Zodd a -> Zodd b -> Zeven (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 * (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.
+Qed.
+
+Theorem Zeven_mult_Zeven_l: forall a b,
+ Zeven a -> Zeven (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.
+Qed.
+
+Theorem Zeven_mult_Zeven_r: forall a b,
+ 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.
+Qed.
+
+Hint Rewrite Zmult_plus_distr_r Zmult_plus_distr_l
+ Zplus_assoc Zmult_1_r Zmult_1_l : Zexpand.
+
+Theorem Zodd_mult_Zodd: forall 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.
+Qed.
+
+(* for compatibility *)
+Close Scope Z_scope.
diff --git a/theories/ZArith/Zmax.v b/theories/ZArith/Zmax.v
index fbc7bfafc..c21b4affb 100644
--- a/theories/ZArith/Zmax.v
+++ b/theories/ZArith/Zmax.v
@@ -46,6 +46,20 @@ Proof.
destruct (Zcompare x y); [ left | right | left ]; split; auto; discriminate.
Qed.
+Lemma Zmax_left : forall n m:Z, n>=m -> Zmax n m = n.
+Proof.
+ intros n m; unfold Zmax, Zge; destruct (n ?= m); auto.
+ intro H; elim H; auto.
+Qed.
+
+Lemma Zmax_right : forall n m:Z, n<=m -> Zmax n m = m.
+Proof.
+ intros n m; unfold Zmax, Zle.
+ generalize (Zcompare_Eq_eq n m).
+ destruct (n ?= m); auto.
+ intros _ H; elim H; auto.
+Qed.
+
(** * Least upper bound properties of max *)
Lemma Zle_max_l : forall n m:Z, n <= Zmax n m.
diff --git a/theories/ZArith/Znat.v b/theories/ZArith/Znat.v
index d9fb4e97c..ec37c1412 100644
--- a/theories/ZArith/Znat.v
+++ b/theories/ZArith/Znat.v
@@ -256,11 +256,6 @@ Proof.
rewrite Zpos_xO; simpl in |- *; rewrite Pplus_diag; reflexivity.
Qed.
-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.
-
(** Misc *)
Theorem intro_Z :
diff --git a/theories/ZArith/Znumtheory.v b/theories/ZArith/Znumtheory.v
index 262d7bac5..710c8aca0 100644
--- a/theories/ZArith/Znumtheory.v
+++ b/theories/ZArith/Znumtheory.v
@@ -293,6 +293,40 @@ Proof.
apply Zmult_lt_compat_r; auto with zarith.
Qed.
+Lemma Zmod_div_mod: forall 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.
+Qed.
+
+Lemma Zmod_divide_minus: forall a b c : Z, 0 < b ->
+ a mod b = c -> (b | a - c).
+Proof.
+ intros a b c H H1; apply Zmod_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.
+Qed.
+
+Lemma Zdivide_mod_minus: forall a b c : Z, 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.
+ 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.
+Qed.
+
(** * Greatest common divisor (gcd). *)
(** There is no unicity of the gcd; hence we define the predicate [gcd a b d]
@@ -574,6 +608,7 @@ Qed.
Lemma Zis_gcd_rel_prime :
forall a b g:Z,
b > 0 -> g >= 0 -> Zis_gcd a b g -> rel_prime (a / g) (b / g).
+Proof.
intros a b g; intros.
assert (g <> 0).
intro.
@@ -602,6 +637,68 @@ Lemma Zis_gcd_rel_prime :
exists q; auto with zarith.
Qed.
+Theorem rel_prime_sym: forall a b, rel_prime a b -> rel_prime b a.
+Proof.
+ intros a b H; auto with zarith.
+ red; apply Zis_gcd_sym; auto with zarith.
+Qed.
+
+Theorem rel_prime_div: forall p q r,
+ rel_prime p q -> (r | p) -> rel_prime r q.
+Proof.
+ intros p q r H (u, H1); subst.
+ inversion_clear H as [H1 H2 H3].
+ red; apply Zis_gcd_intro; try apply Zone_divide.
+ intros x H4 H5; apply H3; auto.
+ apply Zdivide_mult_r; auto.
+Qed.
+
+Theorem rel_prime_1: forall n, rel_prime 1 n.
+Proof.
+ intros n; red; apply Zis_gcd_intro; auto.
+ exists 1; auto with zarith.
+ exists n; auto with zarith.
+Qed.
+
+Theorem not_rel_prime_0: forall n, 1 < n -> ~ rel_prime 0 n.
+Proof.
+ intros n H H1; absurd (n = 1 \/ n = -1).
+ intros [H2 | H2]; subst; absurd_hyp H; auto with zarith.
+ case (Zis_gcd_unique 0 n n 1); auto.
+ apply Zis_gcd_intro; auto.
+ exists 0; auto with zarith.
+ exists 1; auto with zarith.
+Qed.
+
+Theorem rel_prime_mod: forall p q, 0 < q ->
+ rel_prime p q -> rel_prime (p mod q) q.
+Proof.
+ intros p q H H0.
+ assert (H1: Bezout p q 1).
+ apply rel_prime_bezout; auto.
+ inversion_clear H1 as [q1 r1 H2].
+ apply bezout_rel_prime.
+ apply Bezout_intro with q1 (r1 + q1 * (p / q)).
+ rewrite <- H2.
+ pattern p at 3; rewrite (Z_div_mod_eq p q); try ring; auto with zarith.
+Qed.
+
+Theorem rel_prime_mod_rev: forall p q, 0 < q ->
+ rel_prime (p mod q) q -> rel_prime p q.
+Proof.
+ intros p q H H0.
+ rewrite (Z_div_mod_eq p q); auto with zarith; red.
+ apply Zis_gcd_sym; apply Zis_gcd_for_euclid2; auto with zarith.
+Qed.
+
+Theorem Zrel_prime_neq_mod_0: forall a b, 1 < b -> rel_prime a b -> a mod b <> 0.
+Proof.
+ intros a b H H1 H2.
+ case (not_rel_prime_0 _ H).
+ rewrite <- H2.
+ apply rel_prime_mod; auto with zarith.
+Qed.
+
(** * Primality *)
Inductive prime (p:Z) : Prop :=
@@ -654,6 +751,20 @@ Qed.
Hint Resolve prime_rel_prime: zarith.
+(** As a consequence, a prime number is relatively prime with smaller numbers *)
+
+Theorem rel_prime_le_prime:
+ forall a p, prime p -> 1 <= a < p -> rel_prime a p.
+Proof.
+ intros a p Hp [H1 H2].
+ apply rel_prime_sym; apply prime_rel_prime; auto.
+ intros [q Hq]; subst a.
+ case (Zle_or_lt q 0); intros Hl.
+ absurd (q * p <= 0 * p); auto with zarith.
+ absurd (1 * p <= q * p); auto with zarith.
+Qed.
+
+
(** If a prime [p] divides [ab] then it divides either [a] or [b] *)
Lemma prime_mult :
@@ -664,6 +775,108 @@ Proof.
right; apply Gauss with a; auto with zarith.
Qed.
+Lemma not_prime_0: ~ prime 0.
+Proof.
+ intros H1; case (prime_divisors _ H1 2); auto with zarith.
+Qed.
+
+Lemma not_prime_1: ~ prime 1.
+Proof.
+ intros H1; absurd (1 < 1); auto with zarith.
+ inversion H1; auto.
+Qed.
+
+Lemma prime_2: prime 2.
+Proof.
+ apply prime_intro; auto with zarith.
+ intros n [H1 H2]; case Zle_lt_or_eq with ( 1 := H1 ); auto with zarith;
+ clear H1; intros H1.
+ absurd_hyp H2; auto with zarith.
+ subst n; red; auto with zarith.
+ apply Zis_gcd_intro; auto with zarith.
+Qed.
+
+Theorem prime_3: prime 3.
+Proof.
+ apply prime_intro; auto with zarith.
+ intros n [H1 H2]; case Zle_lt_or_eq with ( 1 := H1 ); auto with zarith;
+ clear H1; intros H1.
+ case (Zle_lt_or_eq 2 n); auto with zarith; clear H1; intros H1.
+ absurd_hyp H2; auto with zarith.
+ subst n; red; auto with zarith.
+ apply Zis_gcd_intro; auto with zarith.
+ intros x [q1 Hq1] [q2 Hq2].
+ exists (q2 - q1).
+ apply trans_equal with (3 - 2); auto with zarith.
+ rewrite Hq1; rewrite Hq2; ring.
+ subst n; red; auto with zarith.
+ apply Zis_gcd_intro; auto with zarith.
+Qed.
+
+Theorem prime_ge_2: forall p, prime p -> 2 <= p.
+Proof.
+ intros p Hp; inversion Hp; auto with zarith.
+Qed.
+
+Definition prime' p := 1<p /\ (forall n, 1<n<p -> ~ (n|p)).
+
+Theorem prime_alt:
+ forall p, prime' p <-> prime p.
+Proof.
+ split; destruct 1; intros.
+ (* prime -> prime' *)
+ constructor; auto; intros.
+ red; apply Zis_gcd_intro; auto with zarith; intros.
+ case (Zle_lt_or_eq 0 (Zabs x)); auto with zarith; intros H6.
+ case (Zle_lt_or_eq 1 (Zabs x)); auto with zarith; intros H7.
+ case (Zle_lt_or_eq (Zabs x) p); auto with zarith.
+ apply Zdivide_le; auto with zarith.
+ apply Zdivide_Zabs_inv_l; auto.
+ intros H8; case (H0 (Zabs x)); auto.
+ apply Zdivide_Zabs_inv_l; auto.
+ intros H8; subst p; absurd (Zabs x <= n); auto with zarith.
+ apply Zdivide_le; auto with zarith.
+ apply Zdivide_Zabs_inv_l; auto.
+ rewrite H7; pattern (Zabs x); apply Zabs_intro; auto with zarith.
+ absurd (0%Z = p); auto with zarith.
+ assert (x=0) by (destruct x; simpl in *; now auto).
+ subst x; elim H3; intro q; rewrite Zmult_0_r; auto.
+ (* prime' -> prime *)
+ split; auto; intros.
+ intros H2.
+ case (Zis_gcd_unique n p n 1); auto with zarith.
+ apply Zis_gcd_intro; auto with zarith.
+ apply H0; auto with zarith.
+Qed.
+
+Theorem square_not_prime: forall a, ~ prime (a * a).
+Proof.
+ intros a Ha.
+ rewrite <- (Zabs_square a) in Ha.
+ assert (0 <= Zabs a) by auto with zarith.
+ set (b:=Zabs a) in *; clearbody b.
+ rewrite <- prime_alt in Ha; destruct Ha.
+ case (Zle_lt_or_eq 0 b); auto with zarith; intros Hza1; [ | subst; omega].
+ case (Zle_lt_or_eq 1 b); auto with zarith; intros Hza2; [ | subst; omega].
+ assert (Hza3 := Zmult_lt_compat_r 1 b b Hza1 Hza2).
+ rewrite Zmult_1_l in Hza3.
+ elim (H1 _ (conj Hza2 Hza3)).
+ exists b; auto.
+Qed.
+
+Theorem prime_div_prime: forall p q,
+ prime p -> prime q -> (p | q) -> p = q.
+Proof.
+ intros p q H H1 H2;
+ assert (Hp: 0 < p); try apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+ assert (Hq: 0 < q); try apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+ case prime_divisors with (2 := H2); auto.
+ intros H4; absurd_hyp Hp; subst; auto with zarith.
+ intros [H4| [H4 | H4]]; subst; auto.
+ absurd_hyp H; auto; apply not_prime_1.
+ absurd_hyp Hp; auto with zarith.
+Qed.
+
(** We could obtain a [Zgcd] function via Euclid algorithm. But we propose
here a binary version of [Zgcd], faster and executable within Coq.
@@ -895,12 +1108,12 @@ Theorem Zgcd_div_swap0 : forall a b : Z,
0 < b ->
(a / Zgcd a b) * b = a * (b/Zgcd 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.
- repeat rewrite Zmult_assoc; f_equal.
- rewrite Zmult_comm.
- rewrite <- Zdivide_Zdiv_eq; auto.
+ 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.
+ repeat rewrite Zmult_assoc; f_equal.
+ rewrite Zmult_comm.
+ rewrite <- Zdivide_Zdiv_eq; auto.
Qed.
Theorem Zgcd_div_swap : forall a b c : Z,
@@ -908,16 +1121,100 @@ Theorem Zgcd_div_swap : forall a b c : Z,
0 < b ->
(c * a) / Zgcd a b * b = c * a * (b/Zgcd a b).
Proof.
- intros a b c 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.
- repeat rewrite Zmult_assoc; f_equal.
- rewrite Zdivide_Zdiv_eq_2; auto.
- repeat rewrite <- Zmult_assoc; f_equal.
- rewrite Zmult_comm.
- rewrite <- Zdivide_Zdiv_eq; auto.
+ intros a b c 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.
+ repeat rewrite Zmult_assoc; f_equal.
+ rewrite Zdivide_Zdiv_eq_2; auto.
+ repeat rewrite <- Zmult_assoc; f_equal.
+ rewrite Zmult_comm.
+ rewrite <- Zdivide_Zdiv_eq; auto.
+Qed.
+
+Theorem Zgcd_1_rel_prime : forall a b,
+ Zgcd a b = 1 <-> rel_prime a b.
+Proof.
+ unfold rel_prime; split; intro H.
+ rewrite <- H; apply Zgcd_is_gcd.
+ case (Zis_gcd_unique a b (Zgcd a b) 1); auto.
+ apply Zgcd_is_gcd.
+ intros H2; absurd (0 <= Zgcd a b); auto with zarith.
+ generalize (Zgcd_is_pos a b); auto with zarith.
Qed.
+Definition rel_prime_dec: forall a b,
+ { rel_prime a b }+{ ~ rel_prime a b }.
+Proof.
+ intros a b; case (Z_eq_dec (Zgcd a b) 1); intros H1.
+ left; apply -> Zgcd_1_rel_prime; auto.
+ right; swap H1; apply <- Zgcd_1_rel_prime; auto.
+Defined.
+
+Definition prime_dec_aux:
+ forall p m,
+ { forall n, 1 < n < m -> rel_prime n p } +
+ { exists n, 1 < n < m /\ ~ rel_prime n p }.
+Proof.
+ intros p m.
+ case (Z_lt_dec 1 m); intros H1;
+ [ | left; intros; elimtype False; omega ].
+ pattern m; apply natlike_rec; auto with zarith.
+ left; intros; elimtype False; omega.
+ intros x Hx IH; destruct IH as [F|E].
+ destruct (rel_prime_dec x p) as [Y|N].
+ left; intros n [HH1 HH2].
+ case (Zgt_succ_gt_or_eq x n); auto with zarith.
+ intros HH3; subst x; auto.
+ case (Z_lt_dec 1 x); intros HH1.
+ right; exists x; split; auto with zarith.
+ left; intros n [HHH1 HHH2]; absurd_hyp HHH1; auto with zarith.
+ right; destruct E as (n,((H0,H2),H3)); exists n; auto with zarith.
+Defined.
+
+Definition prime_dec: forall p, { prime p }+{ ~ prime p }.
+Proof.
+ intros p; case (Z_lt_dec 1 p); intros H1.
+ case (prime_dec_aux p p); intros H2.
+ left; apply prime_intro; auto.
+ intros n [Hn1 Hn2]; case Zle_lt_or_eq with ( 1 := Hn1 ); auto.
+ intros HH; subst n.
+ red; apply Zis_gcd_intro; auto with zarith.
+ right; intros H3; inversion_clear H3 as [Hp1 Hp2].
+ case H2; intros n [Hn1 Hn2]; case Hn2; auto with zarith.
+ right; intros H3; inversion_clear H3 as [Hp1 Hp2]; case H1; auto.
+Defined.
+
+Theorem not_prime_divide:
+ forall p, 1 < p -> ~ prime p -> exists n, 1 < n < p /\ (n | p).
+Proof.
+ intros p Hp Hp1.
+ case (prime_dec_aux p p); intros H1.
+ elim Hp1; constructor; auto.
+ intros n [Hn1 Hn2].
+ case Zle_lt_or_eq with ( 1 := Hn1 ); auto with zarith.
+ intros H2; subst n; red; apply Zis_gcd_intro; auto with zarith.
+ case H1; intros n [Hn1 Hn2].
+ generalize (Zgcd_is_pos n p); intros Hpos.
+ case (Zle_lt_or_eq 0 (Zgcd n p)); auto with zarith; intros H3.
+ case (Zle_lt_or_eq 1 (Zgcd n p)); auto with zarith; intros H4.
+ exists (Zgcd n p); split; auto.
+ split; auto.
+ apply Zle_lt_trans with n; auto with zarith.
+ generalize (Zgcd_is_gcd n p); intros tmp; inversion_clear tmp as [Hr1 Hr2 Hr3].
+ case Hr1; intros q Hq.
+ case (Zle_or_lt q 0); auto with zarith; intros Ht.
+ absurd (n <= 0 * Zgcd n p) ; auto with zarith.
+ pattern n at 1; rewrite Hq; auto with zarith.
+ apply Zle_trans with (1 * Zgcd n p); auto with zarith.
+ pattern n at 2; rewrite Hq; auto with zarith.
+ generalize (Zgcd_is_gcd n p); intros Ht; inversion Ht; auto.
+ case Hn2; red.
+ rewrite H4; apply Zgcd_is_gcd.
+ generalize (Zgcd_is_gcd n p); rewrite <- H3; intros tmp;
+ inversion_clear tmp as [Hr1 Hr2 Hr3].
+ 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. *)
diff --git a/theories/ZArith/Zorder.v b/theories/ZArith/Zorder.v
index 45369561d..454560b85 100644
--- a/theories/ZArith/Zorder.v
+++ b/theories/ZArith/Zorder.v
@@ -7,7 +7,7 @@
(************************************************************************)
(*i $Id$ i*)
-(** Binary Integers (Pierre Crégut (CNET, Lannion, France) *)
+(** Binary Integers (Pierre Crégut (CNET, Lannion, France) *)
Require Import BinPos.
Require Import BinInt.
@@ -997,5 +997,31 @@ Proof.
rewrite <- Zplus_assoc; rewrite Zplus_opp_l; rewrite Zplus_0_r; exact H.
Qed.
+Lemma Zmult_lt_compat:
+ forall n m p q : Z, 0 <= n < p -> 0 <= m < q -> n * m < p * q.
+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.
+Qed.
+
+Lemma Zmult_lt_compat2:
+ forall n m p q : Z, 0 < n <= p -> 0 < m < q -> n * m < p * q.
+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.
+Qed.
+
(** For compatibility *)
Notation Zlt_O_minus_lt := Zlt_0_minus_lt (only parsing).
diff --git a/theories/ZArith/Zpow_facts.v b/theories/ZArith/Zpow_facts.v
new file mode 100644
index 000000000..448fa8602
--- /dev/null
+++ b/theories/ZArith/Zpow_facts.v
@@ -0,0 +1,451 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+Require Import ZArith_base.
+Require Import ZArithRing.
+Require Import Zcomplements.
+Require Export Zpower.
+Require Import Zdiv.
+Require Import Znumtheory.
+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.
+
+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.
+
+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_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.
+
+
+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.
+
+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.
+
+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,
+ 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.
+
+Theorem Zpower_lt_monotone: forall 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.
+
+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.
+ 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; 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 Zmult_power: forall p q r, 0 <= q -> 0 <= r ->
+ (p*q)^r = p^r * q^r.
+Proof.
+ intros p q r H1 H2; generalize H2; pattern r; apply natlike_ind; auto.
+ intros r1 H3 H4 H5.
+ unfold Zsucc; rewrite Zpower_exp; auto with zarith.
+ rewrite H4; repeat rewrite Zpower_exp; auto with zarith; ring.
+Qed.
+
+Hint Resolve Zpower_ge_0 Zpower_gt_0: zarith.
+
+Theorem Zpower_le_monotone3: forall 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.
+
+Lemma Zpower_le_monotone_inv: forall 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; elimtype False;omega.
+ destruct c;try discriminate z0. simpl in H1. elimtype False;omega.
+ assert (H4 := Zpower_lt_monotone a c b H). elimtype False;omega.
+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.
+
+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_le_lin: forall n, 0 <= n -> n <= 2^n.
+Proof.
+ intros; apply Zlt_le_weak; apply Zpower2_lt_lin; auto.
+Qed.
+
+
+(** * Zpower and modulo *)
+
+Theorem Zpower_mod: forall 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.
+Qed.
+
+(** A direct way to compute Zpower modulo **)
+
+Fixpoint Zpow_mod_pos (a: Z)(m: positive)(n : Z) {struct m} : Z :=
+ match m with
+ | xH => a mod n
+ | xO m' =>
+ let z := Zpow_mod_pos a m' n in
+ match z with
+ | 0 => 0
+ | _ => (z * z) mod n
+ end
+ | xI m' =>
+ let z := Zpow_mod_pos a m' n in
+ match z with
+ | 0 => 0
+ | _ => (z * z * a) mod n
+ end
+ end.
+
+Definition Zpow_mod a m n :=
+ match m with
+ | 0 => 1
+ | 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.
+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.
+Qed.
+
+Theorem Zpow_mod_correct: forall a m n, 1 < n -> 0 <= m ->
+ 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.
+Qed.
+
+(* Complements about power and number theory. *)
+
+Lemma Zpower_divide: forall 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.
+Qed.
+
+Theorem rel_prime_Zpower_r: forall 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; absurd_hyp H; auto with zarith.
+ intros i Hi Rec _; rewrite Zpower_Zsucc; 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).
+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.
+Qed.
+
+Theorem prime_power_prime: forall 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.
+Qed.
+
+Theorem Zdivide_power_2: forall 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.
+ 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.
+ (* 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.
+ 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; absurd_hyp 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.
+ (* x = 1 *)
+ exists 0; rewrite Zpower_0_r; auto.
+ (* x = 0 *)
+ exists n; destruct H; rewrite Zmult_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.
diff --git a/theories/ZArith/Zpower.v b/theories/ZArith/Zpower.v
index a1963a965..f3f357de1 100644
--- a/theories/ZArith/Zpower.v
+++ b/theories/ZArith/Zpower.v
@@ -14,6 +14,8 @@ Require Import Omega.
Require Import Zcomplements.
Open Local Scope Z_scope.
+Infix "^" := Zpower : Z_scope.
+
(** * Definition of powers over [Z]*)
(** [Zpower_nat z n] is the n-th power of [z] when [n] is an unary
@@ -37,7 +39,7 @@ Qed.
(** This theorem shows that powers of unary and binary integers
are the same thing, modulo the function convert : [positive -> nat] *)
-Theorem Zpower_pos_nat :
+Lemma Zpower_pos_nat :
forall (z:Z) (p:positive), Zpower_pos z p = Zpower_nat z (nat_of_P p).
Proof.
intros; unfold Zpower_pos in |- *; unfold Zpower_nat in |- *;
@@ -48,7 +50,7 @@ Qed.
deduce that the function [[n:positive](Zpower_pos z n)] is a morphism
for [add : positive->positive] and [Zmult : Z->Z] *)
-Theorem Zpower_pos_is_exp :
+Lemma Zpower_pos_is_exp :
forall (n m:positive) (z:Z),
Zpower_pos z (n + m) = Zpower_pos z n * Zpower_pos z m.
Proof.
@@ -60,69 +62,19 @@ Proof.
apply Zpower_nat_is_exp.
Qed.
-Theorem Zpower_pos_1_r: forall x, Zpower_pos x 1 = x.
-Proof.
- intros x; unfold Zpower_pos; simpl; auto with zarith.
-Qed.
-
-Theorem 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.
-
-Theorem 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.
-
-Theorem 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.
-
-Infix "^" := Zpower : Z_scope.
-
Hint Immediate Zpower_nat_is_exp Zpower_pos_is_exp : zarith.
Hint Unfold Zpower_pos Zpower_nat: zarith.
-Lemma Zpower_exp :
+Theorem Zpower_exp :
forall x n m:Z, n >= 0 -> m >= 0 -> x ^ (n + m) = x ^ n * x ^ m.
Proof.
destruct n; destruct m; auto with zarith.
- simpl in |- *; intros; apply Zred_factor0.
- simpl in |- *; auto with zarith.
- intros; compute in H0; absurd (Datatypes.Lt = Datatypes.Lt); auto with zarith.
- intros; compute in H0; absurd (Datatypes.Lt = Datatypes.Lt); 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.
Qed.
-
Section Powers_of_2.
(** * Powers of 2 *)
diff --git a/theories/ZArith/Zsqrt.v b/theories/ZArith/Zsqrt.v
index 4a395e6f8..00da766d8 100644
--- a/theories/ZArith/Zsqrt.v
+++ b/theories/ZArith/Zsqrt.v
@@ -148,6 +148,7 @@ Definition Zsqrt_plain (x:Z) : Z :=
end.
(** A basic theorem about Zsqrt_plain *)
+
Theorem Zsqrt_interval :
forall n:Z,
0 <= n ->
@@ -162,3 +163,53 @@ Proof.
intros p Hle; elim Hle; auto.
Qed.
+(** Positivity *)
+
+Theorem Zsqrt_plain_is_pos: forall n, 0 <= n -> 0 <= Zsqrt_plain n.
+Proof.
+ intros n m; case (Zsqrt_interval n); auto with zarith.
+ intros H1 H2; case (Zle_or_lt 0 (Zsqrt_plain n)); auto.
+ intros H3; absurd_hyp H2; auto; apply Zle_not_lt.
+ apply Zle_trans with ( 2 := H1 ).
+ replace ((Zsqrt_plain n + 1) * (Zsqrt_plain n + 1))
+ with (Zsqrt_plain n * Zsqrt_plain n + (2 * Zsqrt_plain n + 1));
+ auto with zarith.
+ ring.
+Qed.
+
+(** Direct correctness on squares. *)
+
+Theorem Zsqrt_square_id: forall a, 0 <= a -> Zsqrt_plain (a * a) = a.
+Proof.
+ intros a H.
+ generalize (Zsqrt_plain_is_pos (a * a)); auto with zarith; intros Haa.
+ case (Zsqrt_interval (a * a)); auto with zarith.
+ intros H1 H2.
+ case (Zle_or_lt a (Zsqrt_plain (a * a))); intros H3; auto.
+ case Zle_lt_or_eq with (1:=H3); auto; clear H3; intros H3.
+ absurd_hyp H1; auto; apply Zlt_not_le; auto with zarith.
+ apply Zle_lt_trans with (a * Zsqrt_plain (a * a)); auto with zarith.
+ apply Zmult_lt_compat_r; auto with zarith.
+ absurd_hyp H2; auto; apply Zle_not_lt; auto with zarith.
+ apply Zmult_le_compat; auto with zarith.
+Qed.
+
+(** [Zsqrt_plain] is increasing *)
+
+Theorem Zsqrt_le:
+ forall p q, 0 <= p <= q -> Zsqrt_plain p <= Zsqrt_plain q.
+Proof.
+ intros p q [H1 H2]; case Zle_lt_or_eq with (1:=H2); clear H2; intros H2;
+ [ | subst q; auto with zarith].
+ case (Zle_or_lt (Zsqrt_plain p) (Zsqrt_plain q)); auto; intros H3.
+ assert (Hp: (0 <= Zsqrt_plain q)).
+ apply Zsqrt_plain_is_pos; auto with zarith.
+ absurd (q <= p); auto with zarith.
+ apply Zle_trans with ((Zsqrt_plain q + 1) * (Zsqrt_plain q + 1)).
+ case (Zsqrt_interval q); auto with zarith.
+ apply Zle_trans with (Zsqrt_plain p * Zsqrt_plain p); auto with zarith.
+ apply Zmult_le_compat; auto with zarith.
+ case (Zsqrt_interval p); auto with zarith.
+Qed.
+
+