From 9a6e3fe764dc2543dfa94de20fe5eec42d6be705 Mon Sep 17 00:00:00 2001 From: herbelin Date: Sat, 29 Nov 2003 17:28:49 +0000 Subject: Remplacement des fichiers .v ancienne syntaxe de theories, contrib et states par les fichiers nouvelle syntaxe git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@5027 85f007b7-540e-0410-9357-904b9bb8a0f7 --- theories/ZArith/BinInt.v | 1191 +++++++++++++++++++++------------------- theories/ZArith/Wf_Z.v | 260 ++++----- theories/ZArith/ZArith.v | 2 +- theories/ZArith/ZArith_base.v | 13 +- theories/ZArith/ZArith_dec.v | 341 ++++++------ theories/ZArith/Zabs.v | 144 +++-- theories/ZArith/Zbinary.v | 421 +++++++------- theories/ZArith/Zbool.v | 196 ++++--- theories/ZArith/Zcompare.v | 727 ++++++++++++------------ theories/ZArith/Zcomplements.v | 268 ++++----- theories/ZArith/Zdiv.v | 579 ++++++++++--------- theories/ZArith/Zeven.v | 260 +++++---- theories/ZArith/Zhints.v | 93 ++-- theories/ZArith/Zlogarithm.v | 309 +++++------ theories/ZArith/Zmin.v | 102 ++-- theories/ZArith/Zmisc.v | 205 ++----- theories/ZArith/Znat.v | 156 +++--- theories/ZArith/Znumtheory.v | 846 ++++++++++++++-------------- theories/ZArith/Zorder.v | 1006 +++++++++++++++++---------------- theories/ZArith/Zpower.v | 496 ++++++++--------- theories/ZArith/Zsqrt.v | 233 ++++---- theories/ZArith/Zsyntax.v | 278 ---------- theories/ZArith/Zwf.v | 76 +-- theories/ZArith/auxiliary.v | 251 +++------ theories/ZArith/fast_integer.v | 191 ------- theories/ZArith/zarith_aux.v | 151 ----- 26 files changed, 4048 insertions(+), 4747 deletions(-) delete mode 100644 theories/ZArith/Zsyntax.v delete mode 100644 theories/ZArith/fast_integer.v delete mode 100644 theories/ZArith/zarith_aux.v (limited to 'theories/ZArith') diff --git a/theories/ZArith/BinInt.v b/theories/ZArith/BinInt.v index 81cf647701..b6980123a2 100644 --- a/theories/ZArith/BinInt.v +++ b/theories/ZArith/BinInt.v @@ -14,176 +14,179 @@ Require Export BinPos. Require Export Pnat. -Require BinNat. -Require Plus. -Require Mult. +Require Import BinNat. +Require Import Plus. +Require Import Mult. (**********************************************************************) (** Binary integer numbers *) -Inductive Z : Set := - ZERO : Z | POS : positive -> Z | NEG : positive -> Z. +Inductive Z : Set := + | Z0 : Z + | Zpos : positive -> Z + | Zneg : positive -> Z. (** Declare Scope Z_scope with Key Z *) -Delimits Scope Z_scope with Z. +Delimit Scope Z_scope with Z. (** Automatically open scope positive_scope for the constructors of Z *) Bind Scope Z_scope with Z. -Arguments Scope POS [ positive_scope ]. -Arguments Scope NEG [ positive_scope ]. +Arguments Scope Zpos [positive_scope]. +Arguments Scope Zneg [positive_scope]. (** Subtraction of positive into Z *) -Definition Zdouble_plus_one [x:Z] := - Cases x of - | ZERO => (POS xH) - | (POS p) => (POS (xI p)) - | (NEG p) => (NEG (double_moins_un p)) +Definition Zdouble_plus_one (x:Z) := + match x with + | Z0 => Zpos 1 + | Zpos p => Zpos (xI p) + | Zneg p => Zneg (Pdouble_minus_one p) end. -Definition Zdouble_minus_one [x:Z] := - Cases x of - | ZERO => (NEG xH) - | (NEG p) => (NEG (xI p)) - | (POS p) => (POS (double_moins_un p)) +Definition Zdouble_minus_one (x:Z) := + match x with + | Z0 => Zneg 1 + | Zneg p => Zneg (xI p) + | Zpos p => Zpos (Pdouble_minus_one p) end. -Definition Zdouble [x:Z] := - Cases x of - | ZERO => ZERO - | (POS p) => (POS (xO p)) - | (NEG p) => (NEG (xO p)) - end. - -Fixpoint ZPminus [x,y:positive] : Z := - Cases x y of - | (xI x') (xI y') => (Zdouble (ZPminus x' y')) - | (xI x') (xO y') => (Zdouble_plus_one (ZPminus x' y')) - | (xI x') xH => (POS (xO x')) - | (xO x') (xI y') => (Zdouble_minus_one (ZPminus x' y')) - | (xO x') (xO y') => (Zdouble (ZPminus x' y')) - | (xO x') xH => (POS (double_moins_un x')) - | xH (xI y') => (NEG (xO y')) - | xH (xO y') => (NEG (double_moins_un y')) - | xH xH => ZERO +Definition Zdouble (x:Z) := + match x with + | Z0 => Z0 + | Zpos p => Zpos (xO p) + | Zneg p => Zneg (xO p) + end. + +Fixpoint ZPminus (x y:positive) {struct y} : Z := + match x, y with + | xI x', xI y' => Zdouble (ZPminus x' y') + | xI x', xO y' => Zdouble_plus_one (ZPminus x' y') + | xI x', xH => Zpos (xO x') + | xO x', xI y' => Zdouble_minus_one (ZPminus x' y') + | xO x', xO y' => Zdouble (ZPminus x' y') + | xO x', xH => Zpos (Pdouble_minus_one x') + | xH, xI y' => Zneg (xO y') + | xH, xO y' => Zneg (Pdouble_minus_one y') + | xH, xH => Z0 end. (** Addition on integers *) -Definition Zplus := [x,y:Z] - Cases x y of - ZERO y => y - | x ZERO => x - | (POS x') (POS y') => (POS (add x' y')) - | (POS x') (NEG y') => - Cases (compare x' y' EGAL) of - | EGAL => ZERO - | INFERIEUR => (NEG (true_sub y' x')) - | SUPERIEUR => (POS (true_sub x' y')) +Definition Zplus (x y:Z) := + match x, y with + | Z0, y => y + | x, Z0 => x + | Zpos x', Zpos y' => Zpos (x' + y') + | Zpos x', Zneg y' => + match (x' ?= y')%positive Eq with + | Eq => Z0 + | Lt => Zneg (y' - x') + | Gt => Zpos (x' - y') end - | (NEG x') (POS y') => - Cases (compare x' y' EGAL) of - | EGAL => ZERO - | INFERIEUR => (POS (true_sub y' x')) - | SUPERIEUR => (NEG (true_sub x' y')) + | Zneg x', Zpos y' => + match (x' ?= y')%positive Eq with + | Eq => Z0 + | Lt => Zpos (y' - x') + | Gt => Zneg (x' - y') end - | (NEG x') (NEG y') => (NEG (add x' y')) + | Zneg x', Zneg y' => Zneg (x' + y') end. -V8Infix "+" Zplus : Z_scope. +Infix "+" := Zplus : Z_scope. (** Opposite *) -Definition Zopp := [x:Z] - Cases x of - ZERO => ZERO - | (POS x) => (NEG x) - | (NEG x) => (POS x) - end. +Definition Zopp (x:Z) := + match x with + | Z0 => Z0 + | Zpos x => Zneg x + | Zneg x => Zpos x + end. -V8Notation "- x" := (Zopp x) : Z_scope. +Notation "- x" := (Zopp x) : Z_scope. (** Successor on integers *) -Definition Zs := [x:Z](Zplus x (POS xH)). +Definition Zsucc (x:Z) := (x + Zpos 1)%Z. (** Predecessor on integers *) -Definition Zpred := [x:Z](Zplus x (NEG xH)). +Definition Zpred (x:Z) := (x + Zneg 1)%Z. (** Subtraction on integers *) -Definition Zminus := [m,n:Z](Zplus m (Zopp n)). +Definition Zminus (m n:Z) := (m + - n)%Z. -V8Infix "-" Zminus : Z_scope. +Infix "-" := Zminus : Z_scope. (** Multiplication on integers *) -Definition Zmult := [x,y:Z] - Cases x y of - | ZERO _ => ZERO - | _ ZERO => ZERO - | (POS x') (POS y') => (POS (times x' y')) - | (POS x') (NEG y') => (NEG (times x' y')) - | (NEG x') (POS y') => (NEG (times x' y')) - | (NEG x') (NEG y') => (POS (times x' y')) +Definition Zmult (x y:Z) := + match x, y with + | Z0, _ => Z0 + | _, Z0 => Z0 + | Zpos x', Zpos y' => Zpos (x' * y') + | Zpos x', Zneg y' => Zneg (x' * y') + | Zneg x', Zpos y' => Zneg (x' * y') + | Zneg x', Zneg y' => Zpos (x' * y') end. -V8Infix "*" Zmult : Z_scope. +Infix "*" := Zmult : Z_scope. (** Comparison of integers *) -Definition Zcompare := [x,y:Z] - Cases x y of - | ZERO ZERO => EGAL - | ZERO (POS y') => INFERIEUR - | ZERO (NEG y') => SUPERIEUR - | (POS x') ZERO => SUPERIEUR - | (POS x') (POS y') => (compare x' y' EGAL) - | (POS x') (NEG y') => SUPERIEUR - | (NEG x') ZERO => INFERIEUR - | (NEG x') (POS y') => INFERIEUR - | (NEG x') (NEG y') => (Op (compare x' y' EGAL)) +Definition Zcompare (x y:Z) := + match x, y with + | Z0, Z0 => Eq + | Z0, Zpos y' => Lt + | Z0, Zneg y' => Gt + | Zpos x', Z0 => Gt + | Zpos x', Zpos y' => (x' ?= y')%positive Eq + | Zpos x', Zneg y' => Gt + | Zneg x', Z0 => Lt + | Zneg x', Zpos y' => Lt + | Zneg x', Zneg y' => CompOpp ((x' ?= y')%positive Eq) end. -V8Infix "?=" Zcompare (at level 70, no associativity) : Z_scope. +Infix "?=" := Zcompare (at level 70, no associativity) : Z_scope. -Tactic Definition ElimCompare com1 com2:= - Case (Dcompare (Zcompare com1 com2)); [ Idtac | - Let x = FreshId "H" In Intro x; Case x; Clear x ]. +Ltac elim_compare com1 com2 := + case (Dcompare (com1 ?= com2)%Z); + [ idtac | let x := fresh "H" in + (intro x; case x; clear x) ]. (** Sign function *) -Definition Zsgn [z:Z] : Z := - Cases z of - ZERO => ZERO - | (POS p) => (POS xH) - | (NEG p) => (NEG xH) +Definition Zsgn (z:Z) : Z := + match z with + | Z0 => Z0 + | Zpos p => Zpos 1 + | Zneg p => Zneg 1 end. (** Direct, easier to handle variants of successor and addition *) -Definition Zsucc' [x:Z] := - Cases x of - | ZERO => (POS xH) - | (POS x') => (POS (add_un x')) - | (NEG x') => (ZPminus xH x') +Definition Zsucc' (x:Z) := + match x with + | Z0 => Zpos 1 + | Zpos x' => Zpos (Psucc x') + | Zneg x' => ZPminus 1 x' end. -Definition Zpred' [x:Z] := - Cases x of - | ZERO => (NEG xH) - | (POS x') => (ZPminus x' xH) - | (NEG x') => (NEG (add_un x')) +Definition Zpred' (x:Z) := + match x with + | Z0 => Zneg 1 + | Zpos x' => ZPminus x' 1 + | Zneg x' => Zneg (Psucc x') end. -Definition Zplus' := [x,y:Z] - Cases x y of - ZERO y => y - | x ZERO => x - | (POS x') (POS y') => (POS (add x' y')) - | (POS x') (NEG y') => (ZPminus x' y') - | (NEG x') (POS y') => (ZPminus y' x') - | (NEG x') (NEG y') => (NEG (add x' y')) +Definition Zplus' (x y:Z) := + match x, y with + | Z0, y => y + | x, Z0 => x + | Zpos x', Zpos y' => Zpos (x' + y') + | Zpos x', Zneg y' => ZPminus x' y' + | Zneg x', Zpos y' => ZPminus y' x' + | Zneg x', Zneg y' => Zneg (x' + y') end. Open Local Scope Z_scope. @@ -191,74 +194,83 @@ Open Local Scope Z_scope. (**********************************************************************) (** Inductive specification of Z *) -Theorem Zind : (P:(Z ->Prop)) - (P ZERO) -> ((x:Z)(P x) ->(P (Zsucc' x))) -> ((x:Z)(P x) ->(P (Zpred' x))) -> - (z:Z)(P z). +Theorem Zind : + forall P:Z -> Prop, + P Z0 -> + (forall x:Z, P x -> P (Zsucc' x)) -> + (forall x:Z, P x -> P (Zpred' x)) -> forall n:Z, P n. Proof. -Intros P H0 Hs Hp z; NewDestruct z. - Assumption. - Apply Pind with P:=[p](P (POS p)). - Change (P (Zsucc' ZERO)); Apply Hs; Apply H0. - Intro n; Exact (Hs (POS n)). - Apply Pind with P:=[p](P (NEG p)). - Change (P (Zpred' ZERO)); Apply Hp; Apply H0. - Intro n; Exact (Hp (NEG n)). +intros P H0 Hs Hp z; destruct z. + assumption. + apply Pind with (P := fun p => P (Zpos p)). + change (P (Zsucc' Z0)) in |- *; apply Hs; apply H0. + intro n; exact (Hs (Zpos n)). + apply Pind with (P := fun p => P (Zneg p)). + change (P (Zpred' Z0)) in |- *; apply Hp; apply H0. + intro n; exact (Hp (Zneg n)). Qed. (**********************************************************************) (** Properties of opposite on binary integer numbers *) -Theorem Zopp_NEG : (x:positive) (Zopp (NEG x)) = (POS x). +Theorem Zopp_neg : forall p:positive, - Zneg p = Zpos p. Proof. -Reflexivity. +reflexivity. Qed. (** [opp] is involutive *) -Theorem Zopp_Zopp: (x:Z) (Zopp (Zopp x)) = x. +Theorem Zopp_involutive : forall n:Z, - - n = n. Proof. -Intro x; NewDestruct x; Reflexivity. +intro x; destruct x; reflexivity. Qed. (** Injectivity of the opposite *) -Theorem Zopp_intro : (x,y:Z) (Zopp x) = (Zopp y) -> x = y. +Theorem Zopp_inj : forall n m:Z, - n = - m -> n = m. Proof. -Intros x y;Case x;Case y;Simpl;Intros; [ - Trivial | Discriminate H | Discriminate H | Discriminate H -| Simplify_eq H; Intro E; Rewrite E; Trivial -| Discriminate H | Discriminate H | Discriminate H -| Simplify_eq H; Intro E; Rewrite E; Trivial ]. +intros x y; case x; case y; simpl in |- *; intros; + [ trivial + | discriminate H + | discriminate H + | discriminate H + | simplify_eq H; intro E; rewrite E; trivial + | discriminate H + | discriminate H + | discriminate H + | simplify_eq H; intro E; rewrite E; trivial ]. Qed. (**********************************************************************) (* Properties of the direct definition of successor and predecessor *) -Lemma Zpred'_succ' : (x:Z)(Zpred' (Zsucc' x))=x. +Lemma Zpred'_succ' : forall n:Z, Zpred' (Zsucc' n) = n. Proof. -Intro x; NewDestruct x; Simpl. - Reflexivity. -NewDestruct p; Simpl; Try Rewrite double_moins_un_add_un_xI; Reflexivity. -NewDestruct p; Simpl; Try Rewrite is_double_moins_un; Reflexivity. +intro x; destruct x; simpl in |- *. + reflexivity. +destruct p; simpl in |- *; try rewrite Pdouble_minus_one_o_succ_eq_xI; + reflexivity. +destruct p; simpl in |- *; try rewrite Psucc_o_double_minus_one_eq_xO; + reflexivity. Qed. -Lemma Zsucc'_discr : (x:Z)x<>(Zsucc' x). +Lemma Zsucc'_discr : forall n:Z, n <> Zsucc' n. Proof. -Intro x; NewDestruct x; Simpl. - Discriminate. - Injection; Apply add_un_discr. - NewDestruct p; Simpl. - Discriminate. - Intro H; Symmetry in H; Injection H; Apply double_moins_un_xO_discr. - Discriminate. +intro x; destruct x; simpl in |- *. + discriminate. + injection; apply Psucc_discr. + destruct p; simpl in |- *. + discriminate. + intro H; symmetry in H; injection H; apply double_moins_un_xO_discr. + discriminate. Qed. (**********************************************************************) (** Other properties of binary integer numbers *) -Lemma ZL0 : (S (S O))=(plus (S O) (S O)). +Lemma ZL0 : 2%nat = (1 + 1)%nat. Proof. -Reflexivity. +reflexivity. Qed. (**********************************************************************) @@ -266,740 +278,761 @@ Qed. (** zero is left neutral for addition *) -Theorem Zero_left: (x:Z) (Zplus ZERO x) = x. +Theorem Zplus_0_l : forall n:Z, Z0 + n = n. Proof. -Intro x; NewDestruct x; Reflexivity. +intro x; destruct x; reflexivity. Qed. (** zero is right neutral for addition *) -Theorem Zero_right: (x:Z) (Zplus x ZERO) = x. +Theorem Zplus_0_r : forall n:Z, n + Z0 = n. Proof. -Intro x; NewDestruct x; Reflexivity. +intro x; destruct x; reflexivity. Qed. (** addition is commutative *) -Theorem Zplus_sym: (x,y:Z) (Zplus x y) = (Zplus y x). +Theorem Zplus_comm : forall n m:Z, n + m = m + n. Proof. -Intro x;NewInduction x as [|p|p];Intro y; NewDestruct y as [|q|q];Simpl;Try Reflexivity. - Rewrite add_sym; Reflexivity. - Rewrite ZC4; NewDestruct (compare q p EGAL); Reflexivity. - Rewrite ZC4; NewDestruct (compare q p EGAL); Reflexivity. - Rewrite add_sym; Reflexivity. +intro x; induction x as [| p| p]; intro y; destruct y as [| q| q]; + simpl in |- *; try reflexivity. + rewrite Pplus_comm; reflexivity. + rewrite ZC4; destruct ((q ?= p)%positive Eq); reflexivity. + rewrite ZC4; destruct ((q ?= p)%positive Eq); reflexivity. + rewrite Pplus_comm; reflexivity. Qed. (** opposite distributes over addition *) -Theorem Zopp_Zplus: - (x,y:Z) (Zopp (Zplus x y)) = (Zplus (Zopp x) (Zopp y)). +Theorem Zopp_plus_distr : forall n m:Z, - (n + m) = - n + - m. Proof. -Intro x; NewDestruct x as [|p|p]; Intro y; NewDestruct y as [|q|q]; Simpl; - Reflexivity Orelse NewDestruct (compare p q EGAL); Reflexivity. +intro x; destruct x as [| p| p]; intro y; destruct y as [| q| q]; + simpl in |- *; reflexivity || destruct ((p ?= q)%positive Eq); + reflexivity. Qed. (** opposite is inverse for addition *) -Theorem Zplus_inverse_r: (x:Z) (Zplus x (Zopp x)) = ZERO. +Theorem Zplus_opp_r : forall n:Z, n + - n = Z0. Proof. -Intro x; NewDestruct x as [|p|p]; Simpl; [ - Reflexivity -| Rewrite (convert_compare_EGAL p); Reflexivity -| Rewrite (convert_compare_EGAL p); Reflexivity ]. +intro x; destruct x as [| p| p]; simpl in |- *; + [ reflexivity + | rewrite (Pcompare_refl p); reflexivity + | rewrite (Pcompare_refl p); reflexivity ]. Qed. -Theorem Zplus_inverse_l: (x:Z) (Zplus (Zopp x) x) = ZERO. +Theorem Zplus_opp_l : forall n:Z, - n + n = Z0. Proof. -Intro; Rewrite Zplus_sym; Apply Zplus_inverse_r. +intro; rewrite Zplus_comm; apply Zplus_opp_r. Qed. -Hints Local Resolve Zero_left Zero_right. +Hint Local Resolve Zplus_0_l Zplus_0_r. (** addition is associative *) Lemma weak_assoc : - (x,y:positive)(z:Z) (Zplus (POS x) (Zplus (POS y) z))= - (Zplus (Zplus (POS x) (POS y)) z). -Proof. -Intros x y z';Case z'; [ - Auto with arith -| Intros z;Simpl; Rewrite add_assoc;Auto with arith -| Intros z; Simpl; ElimPcompare y z; - Intros E0;Rewrite E0; - ElimPcompare '(add x y) 'z;Intros E1;Rewrite E1; [ - Absurd (compare (add x y) z EGAL)=EGAL; [ (* Case 1 *) - Rewrite convert_compare_SUPERIEUR; [ - Discriminate - | Rewrite convert_add; Rewrite (compare_convert_EGAL y z E0); - Elim (ZL4 x);Intros k E2;Rewrite E2; Simpl; Unfold gt lt; Apply le_n_S; - Apply le_plus_r ] - | Assumption ] - | Absurd (compare (add x y) z EGAL)=INFERIEUR; [ (* Case 2 *) - Rewrite convert_compare_SUPERIEUR; [ - Discriminate - | Rewrite convert_add; Rewrite (compare_convert_EGAL y z E0); - Elim (ZL4 x);Intros k E2;Rewrite E2; Simpl; Unfold gt lt; Apply le_n_S; - Apply le_plus_r] - | Assumption ] - | Rewrite (compare_convert_EGAL y z E0); (* Case 3 *) - Elim (sub_pos_SUPERIEUR (add x z) z);[ - Intros t H; Elim H;Intros H1 H2;Elim H2;Intros H3 H4; - Unfold true_sub; Rewrite H1; Cut x=t; [ - Intros E;Rewrite E;Auto with arith - | Apply simpl_add_r with z:=z; Rewrite <- H3; Rewrite add_sym; Trivial with arith ] - | Pattern 1 z; Rewrite <- (compare_convert_EGAL y z E0); Assumption ] - | Elim (sub_pos_SUPERIEUR z y); [ (* Case 4 *) - Intros k H;Elim H;Intros H1 H2;Elim H2;Intros H3 H4; Unfold 1 true_sub; - Rewrite H1; Cut x=k; [ - Intros E;Rewrite E; Rewrite (convert_compare_EGAL k); Trivial with arith - | Apply simpl_add_r with z:=y; Rewrite (add_sym k y); Rewrite H3; - Apply compare_convert_EGAL; Assumption ] - | Apply ZC2;Assumption] - | Elim (sub_pos_SUPERIEUR z y); [ (* Case 5 *) - Intros k H;Elim H;Intros H1 H2;Elim H2;Intros H3 H4; - Unfold 1 3 5 true_sub; Rewrite H1; - Cut (compare x k EGAL)=INFERIEUR; [ - Intros E2;Rewrite E2; Elim (sub_pos_SUPERIEUR k x); [ - Intros i H5;Elim H5;Intros H6 H7;Elim H7;Intros H8 H9; - Elim (sub_pos_SUPERIEUR z (add x y)); [ - Intros j H10;Elim H10;Intros H11 H12;Elim H12;Intros H13 H14; - Unfold true_sub ;Rewrite H6;Rewrite H11; Cut i=j; [ - Intros E;Rewrite E;Auto with arith - | Apply (simpl_add_l (add x y)); Rewrite H13; - Rewrite (add_sym x y); Rewrite <- add_assoc; Rewrite H8; - Assumption ] - | Apply ZC2; Assumption] - | Apply ZC2;Assumption] - | Apply convert_compare_INFERIEUR; - Apply simpl_lt_plus_l with p:=(convert y); - Do 2 Rewrite <- convert_add; Apply compare_convert_INFERIEUR; - Rewrite H3; Rewrite add_sym; Assumption ] - | Apply ZC2; Assumption ] - | Elim (sub_pos_SUPERIEUR z y); [ (* Case 6 *) - Intros k H;Elim H;Intros H1 H2;Elim H2;Intros H3 H4; - Elim (sub_pos_SUPERIEUR (add x y) z); [ - Intros i H5;Elim H5;Intros H6 H7;Elim H7;Intros H8 H9; - Unfold true_sub; Rewrite H1;Rewrite H6; - Cut (compare x k EGAL)=SUPERIEUR; [ - Intros H10;Elim (sub_pos_SUPERIEUR x k H10); - Intros j H11;Elim H11;Intros H12 H13;Elim H13;Intros H14 H15; - Rewrite H10; Rewrite H12; Cut i=j; [ - Intros H16;Rewrite H16;Auto with arith - | Apply (simpl_add_l (add z k)); Rewrite <- (add_assoc z k j); - Rewrite H14; Rewrite (add_sym z k); Rewrite <- add_assoc; - Rewrite H8; Rewrite (add_sym x y); Rewrite add_assoc; - Rewrite (add_sym k y); Rewrite H3; Trivial with arith] - | Apply convert_compare_SUPERIEUR; Unfold lt gt; - Apply simpl_lt_plus_l with p:=(convert y); - Do 2 Rewrite <- convert_add; Apply compare_convert_INFERIEUR; - Rewrite H3; Rewrite add_sym; Apply ZC1; Assumption ] - | Assumption ] - | Apply ZC2;Assumption ] - | Absurd (compare (add x y) z EGAL)=EGAL; [ (* Case 7 *) - Rewrite convert_compare_SUPERIEUR; [ - Discriminate - | Rewrite convert_add; Unfold gt;Apply lt_le_trans with m:=(convert y);[ - Apply compare_convert_INFERIEUR; Apply ZC1; Assumption - | Apply le_plus_r]] - | Assumption ] - | Absurd (compare (add x y) z EGAL)=INFERIEUR; [ (* Case 8 *) - Rewrite convert_compare_SUPERIEUR; [ - Discriminate - | Unfold gt; Apply lt_le_trans with m:=(convert y);[ - Exact (compare_convert_SUPERIEUR y z E0) - | Rewrite convert_add; Apply le_plus_r]] - | Assumption ] - | Elim sub_pos_SUPERIEUR with 1:=E0;Intros k H1; (* Case 9 *) - Elim sub_pos_SUPERIEUR with 1:=E1; Intros i H2;Elim H1;Intros H3 H4; - Elim H4;Intros H5 H6; Elim H2;Intros H7 H8;Elim H8;Intros H9 H10; - Unfold true_sub ;Rewrite H3;Rewrite H7; Cut (add x k)=i; [ - Intros E;Rewrite E;Auto with arith - | Apply (simpl_add_l z);Rewrite (add_sym x k); - Rewrite add_assoc; Rewrite H5;Rewrite H9; - Rewrite add_sym; Trivial with arith ]]]. -Qed. - -Hints Local Resolve weak_assoc. - -Theorem Zplus_assoc : - (n,m,p:Z) (Zplus n (Zplus m p))= (Zplus (Zplus n m) p). -Proof. -Intros x y z;Case x;Case y;Case z;Auto with arith; Intros; [ - Rewrite (Zplus_sym (NEG p0)); Rewrite weak_assoc; - Rewrite (Zplus_sym (Zplus (POS p1) (NEG p0))); Rewrite weak_assoc; - Rewrite (Zplus_sym (POS p1)); Trivial with arith -| Apply Zopp_intro; Do 4 Rewrite Zopp_Zplus; - Do 2 Rewrite Zopp_NEG; Rewrite Zplus_sym; Rewrite <- weak_assoc; - Rewrite (Zplus_sym (Zopp (POS p1))); - Rewrite (Zplus_sym (Zplus (POS p0) (Zopp (POS p1)))); - Rewrite (weak_assoc p); Rewrite weak_assoc; Rewrite (Zplus_sym (POS p0)); - Trivial with arith -| Rewrite Zplus_sym; Rewrite (Zplus_sym (POS p0) (POS p)); - Rewrite <- weak_assoc; Rewrite Zplus_sym; Rewrite (Zplus_sym (POS p0)); - Trivial with arith -| Apply Zopp_intro; Do 4 Rewrite Zopp_Zplus; - Do 2 Rewrite Zopp_NEG; Rewrite (Zplus_sym (Zopp (POS p0))); - Rewrite weak_assoc; Rewrite (Zplus_sym (Zplus (POS p1) (Zopp (POS p0)))); - Rewrite weak_assoc;Rewrite (Zplus_sym (POS p)); Trivial with arith -| Apply Zopp_intro; Do 4 Rewrite Zopp_Zplus; Do 2 Rewrite Zopp_NEG; - Apply weak_assoc -| Apply Zopp_intro; Do 4 Rewrite Zopp_Zplus; Do 2 Rewrite Zopp_NEG; - Apply weak_assoc] -. -Qed. - -V7only [Notation Zplus_assoc_l := Zplus_assoc.]. - -Lemma Zplus_assoc_r : (n,m,p:Z)(Zplus (Zplus n m) p) =(Zplus n (Zplus m p)). -Proof. -Intros; Symmetry; Apply Zplus_assoc. + forall (p q:positive) (n:Z), Zpos p + (Zpos q + n) = Zpos p + Zpos q + n. +Proof. +intros x y z'; case z'; + [ auto with arith + | intros z; simpl in |- *; rewrite Pplus_assoc; auto with arith + | intros z; simpl in |- *; ElimPcompare y z; intros E0; rewrite E0; + ElimPcompare (x + y)%positive z; intros E1; rewrite E1; + [ absurd ((x + y ?= z)%positive Eq = Eq); + [ (* Case 1 *) + rewrite nat_of_P_gt_Gt_compare_complement_morphism; + [ discriminate + | rewrite nat_of_P_plus_morphism; rewrite (Pcompare_Eq_eq y z E0); + elim (ZL4 x); intros k E2; rewrite E2; + simpl in |- *; unfold gt, lt in |- *; + apply le_n_S; apply le_plus_r ] + | assumption ] + | absurd ((x + y ?= z)%positive Eq = Lt); + [ (* Case 2 *) + rewrite nat_of_P_gt_Gt_compare_complement_morphism; + [ discriminate + | rewrite nat_of_P_plus_morphism; rewrite (Pcompare_Eq_eq y z E0); + elim (ZL4 x); intros k E2; rewrite E2; + simpl in |- *; unfold gt, lt in |- *; + apply le_n_S; apply le_plus_r ] + | assumption ] + | rewrite (Pcompare_Eq_eq y z E0); + (* Case 3 *) + elim (Pminus_mask_Gt (x + z) z); + [ intros t H; elim H; intros H1 H2; elim H2; intros H3 H4; + unfold Pminus in |- *; rewrite H1; cut (x = t); + [ intros E; rewrite E; auto with arith + | apply Pplus_reg_r with (r := z); rewrite <- H3; + rewrite Pplus_comm; trivial with arith ] + | pattern z at 1 in |- *; rewrite <- (Pcompare_Eq_eq y z E0); + assumption ] + | elim (Pminus_mask_Gt z y); + [ (* Case 4 *) + intros k H; elim H; intros H1 H2; elim H2; intros H3 H4; + unfold Pminus at 1 in |- *; rewrite H1; cut (x = k); + [ intros E; rewrite E; rewrite (Pcompare_refl k); + trivial with arith + | apply Pplus_reg_r with (r := y); rewrite (Pplus_comm k y); + rewrite H3; apply Pcompare_Eq_eq; assumption ] + | apply ZC2; assumption ] + | elim (Pminus_mask_Gt z y); + [ (* Case 5 *) + intros k H; elim H; intros H1 H2; elim H2; intros H3 H4; + unfold Pminus at 1 3 5 in |- *; rewrite H1; + cut ((x ?= k)%positive Eq = Lt); + [ intros E2; rewrite E2; elim (Pminus_mask_Gt k x); + [ intros i H5; elim H5; intros H6 H7; elim H7; intros H8 H9; + elim (Pminus_mask_Gt z (x + y)); + [ intros j H10; elim H10; intros H11 H12; elim H12; + intros H13 H14; unfold Pminus in |- *; + rewrite H6; rewrite H11; cut (i = j); + [ intros E; rewrite E; auto with arith + | apply (Pplus_reg_l (x + y)); rewrite H13; + rewrite (Pplus_comm x y); rewrite <- Pplus_assoc; + rewrite H8; assumption ] + | apply ZC2; assumption ] + | apply ZC2; assumption ] + | apply nat_of_P_lt_Lt_compare_complement_morphism; + apply plus_lt_reg_l with (p := nat_of_P y); + do 2 rewrite <- nat_of_P_plus_morphism; + apply nat_of_P_lt_Lt_compare_morphism; + rewrite H3; rewrite Pplus_comm; assumption ] + | apply ZC2; assumption ] + | elim (Pminus_mask_Gt z y); + [ (* Case 6 *) + intros k H; elim H; intros H1 H2; elim H2; intros H3 H4; + elim (Pminus_mask_Gt (x + y) z); + [ intros i H5; elim H5; intros H6 H7; elim H7; intros H8 H9; + unfold Pminus in |- *; rewrite H1; rewrite H6; + cut ((x ?= k)%positive Eq = Gt); + [ intros H10; elim (Pminus_mask_Gt x k H10); intros j H11; + elim H11; intros H12 H13; elim H13; + intros H14 H15; rewrite H10; rewrite H12; + cut (i = j); + [ intros H16; rewrite H16; auto with arith + | apply (Pplus_reg_l (z + k)); rewrite <- (Pplus_assoc z k j); + rewrite H14; rewrite (Pplus_comm z k); + rewrite <- Pplus_assoc; rewrite H8; + rewrite (Pplus_comm x y); rewrite Pplus_assoc; + rewrite (Pplus_comm k y); rewrite H3; + trivial with arith ] + | apply nat_of_P_gt_Gt_compare_complement_morphism; + unfold lt, gt in |- *; + apply plus_lt_reg_l with (p := nat_of_P y); + do 2 rewrite <- nat_of_P_plus_morphism; + apply nat_of_P_lt_Lt_compare_morphism; + rewrite H3; rewrite Pplus_comm; apply ZC1; + assumption ] + | assumption ] + | apply ZC2; assumption ] + | absurd ((x + y ?= z)%positive Eq = Eq); + [ (* Case 7 *) + rewrite nat_of_P_gt_Gt_compare_complement_morphism; + [ discriminate + | rewrite nat_of_P_plus_morphism; unfold gt in |- *; + apply lt_le_trans with (m := nat_of_P y); + [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; assumption + | apply le_plus_r ] ] + | assumption ] + | absurd ((x + y ?= z)%positive Eq = Lt); + [ (* Case 8 *) + rewrite nat_of_P_gt_Gt_compare_complement_morphism; + [ discriminate + | unfold gt in |- *; apply lt_le_trans with (m := nat_of_P y); + [ exact (nat_of_P_gt_Gt_compare_morphism y z E0) + | rewrite nat_of_P_plus_morphism; apply le_plus_r ] ] + | assumption ] + | elim Pminus_mask_Gt with (1 := E0); intros k H1; + (* Case 9 *) + elim Pminus_mask_Gt with (1 := E1); intros i H2; + elim H1; intros H3 H4; elim H4; intros H5 H6; + elim H2; intros H7 H8; elim H8; intros H9 H10; + unfold Pminus in |- *; rewrite H3; rewrite H7; + cut ((x + k)%positive = i); + [ intros E; rewrite E; auto with arith + | apply (Pplus_reg_l z); rewrite (Pplus_comm x k); rewrite Pplus_assoc; + rewrite H5; rewrite H9; rewrite Pplus_comm; + trivial with arith ] ] ]. +Qed. + +Hint Local Resolve weak_assoc. + +Theorem Zplus_assoc : forall n m p:Z, n + (m + p) = n + m + p. +Proof. +intros x y z; case x; case y; case z; auto with arith; intros; + [ rewrite (Zplus_comm (Zneg p0)); rewrite weak_assoc; + rewrite (Zplus_comm (Zpos p1 + Zneg p0)); rewrite weak_assoc; + rewrite (Zplus_comm (Zpos p1)); trivial with arith + | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg; + rewrite Zplus_comm; rewrite <- weak_assoc; + rewrite (Zplus_comm (- Zpos p1)); + rewrite (Zplus_comm (Zpos p0 + - Zpos p1)); rewrite (weak_assoc p); + rewrite weak_assoc; rewrite (Zplus_comm (Zpos p0)); + trivial with arith + | rewrite Zplus_comm; rewrite (Zplus_comm (Zpos p0) (Zpos p)); + rewrite <- weak_assoc; rewrite Zplus_comm; rewrite (Zplus_comm (Zpos p0)); + trivial with arith + | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg; + rewrite (Zplus_comm (- Zpos p0)); rewrite weak_assoc; + rewrite (Zplus_comm (Zpos p1 + - Zpos p0)); rewrite weak_assoc; + rewrite (Zplus_comm (Zpos p)); trivial with arith + | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg; + apply weak_assoc + | apply Zopp_inj; do 4 rewrite Zopp_plus_distr; do 2 rewrite Zopp_neg; + apply weak_assoc ]. +Qed. + + +Lemma Zplus_assoc_reverse : forall n m p:Z, n + m + p = n + (m + p). +Proof. +intros; symmetry in |- *; apply Zplus_assoc. Qed. (** Associativity mixed with commutativity *) -Theorem Zplus_permute : (n,m,p:Z) (Zplus n (Zplus m p))=(Zplus m (Zplus n p)). +Theorem Zplus_permute : forall n m p:Z, n + (m + p) = m + (n + p). Proof. -Intros n m p; -Rewrite Zplus_sym;Rewrite <- Zplus_assoc; Rewrite (Zplus_sym p n); Trivial with arith. +intros n m p; rewrite Zplus_comm; rewrite <- Zplus_assoc; + rewrite (Zplus_comm p n); trivial with arith. Qed. (** addition simplifies *) -Theorem Zsimpl_plus_l : (n,m,p:Z)(Zplus n m)=(Zplus n p)->m=p. -Intros n m p H; Cut (Zplus (Zopp n) (Zplus n m))=(Zplus (Zopp n) (Zplus n p));[ - Do 2 Rewrite -> Zplus_assoc; Rewrite -> (Zplus_sym (Zopp n) n); - Rewrite -> Zplus_inverse_r;Simpl; Trivial with arith -| Rewrite -> H; Trivial with arith ]. +Theorem Zplus_reg_l : forall n m p:Z, n + m = n + p -> m = p. +intros n m p H; cut (- n + (n + m) = - n + (n + p)); + [ do 2 rewrite Zplus_assoc; rewrite (Zplus_comm (- n) n); + rewrite Zplus_opp_r; simpl in |- *; trivial with arith + | rewrite H; trivial with arith ]. Qed. (** addition and successor permutes *) -Lemma Zplus_S_n: (x,y:Z) (Zplus (Zs x) y) = (Zs (Zplus x y)). +Lemma Zplus_succ_l : forall n m:Z, Zsucc n + m = Zsucc (n + m). Proof. -Intros x y; Unfold Zs; Rewrite (Zplus_sym (Zplus x y)); Rewrite Zplus_assoc; -Rewrite (Zplus_sym (POS xH)); Trivial with arith. +intros x y; unfold Zsucc in |- *; rewrite (Zplus_comm (x + y)); + rewrite Zplus_assoc; rewrite (Zplus_comm (Zpos 1)); + trivial with arith. Qed. -Lemma Zplus_n_Sm : (n,m:Z) (Zs (Zplus n m))=(Zplus n (Zs m)). +Lemma Zplus_succ_r : forall n m:Z, Zsucc (n + m) = n + Zsucc m. Proof. -Intros n m; Unfold Zs; Rewrite Zplus_assoc; Trivial with arith. +intros n m; unfold Zsucc in |- *; rewrite Zplus_assoc; trivial with arith. Qed. -Lemma Zplus_Snm_nSm : (n,m:Z)(Zplus (Zs n) m)=(Zplus n (Zs m)). +Lemma Zplus_succ_comm : forall n m:Z, Zsucc n + m = n + Zsucc m. Proof. -Unfold Zs ;Intros n m; Rewrite <- Zplus_assoc; Rewrite (Zplus_sym (POS xH)); -Trivial with arith. +unfold Zsucc in |- *; intros n m; rewrite <- Zplus_assoc; + rewrite (Zplus_comm (Zpos 1)); trivial with arith. Qed. (** Misc properties, usually redundant or non natural *) -Lemma Zplus_n_O : (n:Z) n=(Zplus n ZERO). +Lemma Zplus_0_r_reverse : forall n:Z, n = n + Z0. Proof. -Symmetry; Apply Zero_right. +symmetry in |- *; apply Zplus_0_r. Qed. -Lemma Zplus_unit_left : (n,m:Z) (Zplus n ZERO)=m -> n=m. +Lemma Zplus_0_simpl_l : forall n m:Z, n + Z0 = m -> n = m. Proof. -Intros n m; Rewrite Zero_right; Intro; Assumption. +intros n m; rewrite Zplus_0_r; intro; assumption. Qed. -Lemma Zplus_unit_right : (n,m:Z) n=(Zplus m ZERO) -> n=m. +Lemma Zplus_0_simpl_l_reverse : forall n m:Z, n = m + Z0 -> n = m. Proof. -Intros n m; Rewrite Zero_right; Intro; Assumption. +intros n m; rewrite Zplus_0_r; intro; assumption. Qed. -Lemma Zplus_simpl : (x,y,z,t:Z) x=y -> z=t -> (Zplus x z)=(Zplus y t). +Lemma Zplus_eq_compat : forall n m p q:Z, n = m -> p = q -> n + p = m + q. Proof. -Intros; Rewrite H; Rewrite H0; Reflexivity. +intros; rewrite H; rewrite H0; reflexivity. Qed. -Lemma Zplus_Zopp_expand : (x,y,z:Z) - (Zplus x (Zopp y))=(Zplus (Zplus x (Zopp z)) (Zplus z (Zopp y))). +Lemma Zplus_opp_expand : forall n m p:Z, n + - m = n + - p + (p + - m). Proof. -Intros x y z. -Rewrite <- (Zplus_assoc x). -Rewrite (Zplus_assoc (Zopp z)). -Rewrite Zplus_inverse_l. -Reflexivity. +intros x y z. +rewrite <- (Zplus_assoc x). +rewrite (Zplus_assoc (- z)). +rewrite Zplus_opp_l. +reflexivity. Qed. (**********************************************************************) (** Properties of successor and predecessor on binary integer numbers *) -Theorem Zn_Sn : (x:Z) ~ x=(Zs x). +Theorem Zsucc_discr : forall n:Z, n <> Zsucc n. Proof. -Intros n;Cut ~ZERO=(POS xH);[ - Unfold not ;Intros H1 H2;Apply H1;Apply (Zsimpl_plus_l n);Rewrite Zero_right; - Exact H2 -| Discriminate ]. +intros n; cut (Z0 <> Zpos 1); + [ unfold not in |- *; intros H1 H2; apply H1; apply (Zplus_reg_l n); + rewrite Zplus_0_r; exact H2 + | discriminate ]. Qed. -Theorem add_un_Zs : (x:positive) (POS (add_un x)) = (Zs (POS x)). +Theorem Zpos_succ_morphism : + forall p:positive, Zpos (Psucc p) = Zsucc (Zpos p). Proof. -Intro; Rewrite -> ZL12; Unfold Zs; Simpl; Trivial with arith. +intro; rewrite Pplus_one_succ_r; unfold Zsucc in |- *; simpl in |- *; + trivial with arith. Qed. (** successor and predecessor are inverse functions *) -Theorem Zs_pred : (n:Z) n=(Zs (Zpred n)). +Theorem Zsucc_pred : forall n:Z, n = Zsucc (Zpred n). Proof. -Intros n; Unfold Zs Zpred ;Rewrite <- Zplus_assoc; Simpl; Rewrite Zero_right; -Trivial with arith. +intros n; unfold Zsucc, Zpred in |- *; rewrite <- Zplus_assoc; simpl in |- *; + rewrite Zplus_0_r; trivial with arith. Qed. -Hints Immediate Zs_pred : zarith. +Hint Immediate Zsucc_pred: zarith. -Theorem Zpred_Sn : (x:Z) x=(Zpred (Zs x)). +Theorem Zpred_succ : forall n:Z, n = Zpred (Zsucc n). Proof. -Intros m; Unfold Zpred Zs; Rewrite <- Zplus_assoc; Simpl; -Rewrite Zplus_sym; Auto with arith. +intros m; unfold Zpred, Zsucc in |- *; rewrite <- Zplus_assoc; simpl in |- *; + rewrite Zplus_comm; auto with arith. Qed. -Theorem Zeq_add_S : (n,m:Z) (Zs n)=(Zs m) -> n=m. +Theorem Zsucc_inj : forall n m:Z, Zsucc n = Zsucc m -> n = m. Proof. -Intros n m H. -Change (Zplus (Zplus (NEG xH) (POS xH)) n)= - (Zplus (Zplus (NEG xH) (POS xH)) m); -Do 2 Rewrite <- Zplus_assoc; Do 2 Rewrite (Zplus_sym (POS xH)); -Unfold Zs in H;Rewrite H; Trivial with arith. +intros n m H. +change (Zneg 1 + Zpos 1 + n = Zneg 1 + Zpos 1 + m) in |- *; + do 2 rewrite <- Zplus_assoc; do 2 rewrite (Zplus_comm (Zpos 1)); + unfold Zsucc in H; rewrite H; trivial with arith. Qed. (** Misc properties, usually redundant or non natural *) -Lemma Zeq_S : (n,m:Z) n=m -> (Zs n)=(Zs m). +Lemma Zsucc_eq_compat : forall n m:Z, n = m -> Zsucc n = Zsucc m. Proof. -Intros n m H; Rewrite H; Reflexivity. +intros n m H; rewrite H; reflexivity. Qed. -Lemma Znot_eq_S : (n,m:Z) ~(n=m) -> ~((Zs n)=(Zs m)). +Lemma Zsucc_inj_contrapositive : forall n m:Z, n <> m -> Zsucc n <> Zsucc m. Proof. -Unfold not ;Intros n m H1 H2;Apply H1;Apply Zeq_add_S; Assumption. +unfold not in |- *; intros n m H1 H2; apply H1; apply Zsucc_inj; assumption. Qed. (**********************************************************************) (** Properties of subtraction on binary integer numbers *) -Lemma Zminus_0_r : (x:Z) (Zminus x ZERO)=x. +Lemma Zminus_0_r : forall n:Z, n - Z0 = n. Proof. -Intro; Unfold Zminus; Simpl;Rewrite Zero_right; Trivial with arith. +intro; unfold Zminus in |- *; simpl in |- *; rewrite Zplus_0_r; + trivial with arith. Qed. -Lemma Zminus_n_O : (x:Z) x=(Zminus x ZERO). +Lemma Zminus_0_l_reverse : forall n:Z, n = n - Z0. Proof. -Intro; Symmetry; Apply Zminus_0_r. +intro; symmetry in |- *; apply Zminus_0_r. Qed. -Lemma Zminus_diag : (n:Z)(Zminus n n)=ZERO. +Lemma Zminus_diag : forall n:Z, n - n = Z0. Proof. -Intro; Unfold Zminus; Rewrite Zplus_inverse_r; Trivial with arith. +intro; unfold Zminus in |- *; rewrite Zplus_opp_r; trivial with arith. Qed. -Lemma Zminus_n_n : (n:Z)(ZERO=(Zminus n n)). +Lemma Zminus_diag_reverse : forall n:Z, Z0 = n - n. Proof. -Intro; Symmetry; Apply Zminus_diag. +intro; symmetry in |- *; apply Zminus_diag. Qed. -Lemma Zplus_minus : (x,y,z:Z)(x=(Zplus y z))->(z=(Zminus x y)). +Lemma Zplus_minus_eq : forall n m p:Z, n = m + p -> p = n - m. Proof. -Intros n m p H;Unfold Zminus;Apply (Zsimpl_plus_l m); -Rewrite (Zplus_sym m (Zplus n (Zopp m))); Rewrite <- Zplus_assoc; -Rewrite Zplus_inverse_l; Rewrite Zero_right; Rewrite H; Trivial with arith. +intros n m p H; unfold Zminus in |- *; apply (Zplus_reg_l m); + rewrite (Zplus_comm m (n + - m)); rewrite <- Zplus_assoc; + rewrite Zplus_opp_l; rewrite Zplus_0_r; rewrite H; + trivial with arith. Qed. -Lemma Zminus_plus : (x,y:Z)(Zminus (Zplus x y) x)=y. +Lemma Zminus_plus : forall n m:Z, n + m - n = m. Proof. -Intros n m;Unfold Zminus ;Rewrite -> (Zplus_sym n m);Rewrite <- Zplus_assoc; -Rewrite -> Zplus_inverse_r; Apply Zero_right. +intros n m; unfold Zminus in |- *; rewrite (Zplus_comm n m); + rewrite <- Zplus_assoc; rewrite Zplus_opp_r; apply Zplus_0_r. Qed. -Lemma Zle_plus_minus : (n,m:Z) (Zplus n (Zminus m n))=m. +Lemma Zplus_minus : forall n m:Z, n + (m - n) = m. Proof. -Unfold Zminus; Intros n m; Rewrite Zplus_permute; Rewrite Zplus_inverse_r; -Apply Zero_right. +unfold Zminus in |- *; intros n m; rewrite Zplus_permute; rewrite Zplus_opp_r; + apply Zplus_0_r. Qed. -Lemma Zminus_Sn_m : (n,m:Z)((Zs (Zminus n m))=(Zminus (Zs n) m)). +Lemma Zminus_succ_l : forall n m:Z, Zsucc (n - m) = Zsucc n - m. Proof. -Intros n m;Unfold Zminus Zs; Rewrite (Zplus_sym n (Zopp m)); -Rewrite <- Zplus_assoc;Apply Zplus_sym. +intros n m; unfold Zminus, Zsucc in |- *; rewrite (Zplus_comm n (- m)); + rewrite <- Zplus_assoc; apply Zplus_comm. Qed. -Lemma Zminus_plus_simpl_l : - (x,y,z:Z)(Zminus (Zplus z x) (Zplus z y))=(Zminus x y). +Lemma Zminus_plus_simpl_l : forall n m p:Z, p + n - (p + m) = n - m. Proof. -Intros n m p;Unfold Zminus; Rewrite Zopp_Zplus; Rewrite Zplus_assoc; -Rewrite (Zplus_sym p); Rewrite <- (Zplus_assoc n p); Rewrite Zplus_inverse_r; -Rewrite Zero_right; Trivial with arith. +intros n m p; unfold Zminus in |- *; rewrite Zopp_plus_distr; + rewrite Zplus_assoc; rewrite (Zplus_comm p); rewrite <- (Zplus_assoc n p); + rewrite Zplus_opp_r; rewrite Zplus_0_r; trivial with arith. Qed. -Lemma Zminus_plus_simpl : - (x,y,z:Z)((Zminus x y)=(Zminus (Zplus z x) (Zplus z y))). +Lemma Zminus_plus_simpl_l_reverse : forall n m p:Z, n - m = p + n - (p + m). Proof. -Intros; Symmetry; Apply Zminus_plus_simpl_l. +intros; symmetry in |- *; apply Zminus_plus_simpl_l. Qed. -Lemma Zminus_Zplus_compatible : - (x,y,z:Z) (Zminus (Zplus x z) (Zplus y z)) = (Zminus x y). -Intros x y n. -Unfold Zminus. -Rewrite -> Zopp_Zplus. -Rewrite -> (Zplus_sym (Zopp y) (Zopp n)). -Rewrite -> Zplus_assoc. -Rewrite <- (Zplus_assoc x n (Zopp n)). -Rewrite -> (Zplus_inverse_r n). -Rewrite <- Zplus_n_O. -Reflexivity. +Lemma Zminus_plus_simpl_r : forall n m p:Z, n + p - (m + p) = n - m. +intros x y n. +unfold Zminus in |- *. +rewrite Zopp_plus_distr. +rewrite (Zplus_comm (- y) (- n)). +rewrite Zplus_assoc. +rewrite <- (Zplus_assoc x n (- n)). +rewrite (Zplus_opp_r n). +rewrite <- Zplus_0_r_reverse. +reflexivity. Qed. (** Misc redundant properties *) -V7only [Set Implicit Arguments.]. -Lemma Zeq_Zminus : (x,y:Z)x=y -> (Zminus x y)=ZERO. +Lemma Zeq_minus : forall n m:Z, n = m -> n - m = Z0. Proof. -Intros x y H; Rewrite H; Symmetry; Apply Zminus_n_n. +intros x y H; rewrite H; symmetry in |- *; apply Zminus_diag_reverse. Qed. -Lemma Zminus_Zeq : (x,y:Z)(Zminus x y)=ZERO -> x=y. +Lemma Zminus_eq : forall n m:Z, n - m = Z0 -> n = m. Proof. -Intros x y H; Rewrite <- (Zle_plus_minus y x); Rewrite H; Apply Zero_right. +intros x y H; rewrite <- (Zplus_minus y x); rewrite H; apply Zplus_0_r. Qed. -V7only [Unset Implicit Arguments.]. (**********************************************************************) (** Properties of multiplication on binary integer numbers *) (** One is neutral for multiplication *) -Theorem Zmult_1_n : (n:Z)(Zmult (POS xH) n)=n. +Theorem Zmult_1_l : forall n:Z, Zpos 1 * n = n. Proof. -Intro x; NewDestruct x; Reflexivity. +intro x; destruct x; reflexivity. Qed. -V7only [Notation Zmult_one := Zmult_1_n.]. -Theorem Zmult_n_1 : (n:Z)(Zmult n (POS xH))=n. +Theorem Zmult_1_r : forall n:Z, n * Zpos 1 = n. Proof. -Intro x; NewDestruct x; Simpl; Try Rewrite times_x_1; Reflexivity. +intro x; destruct x; simpl in |- *; try rewrite Pmult_1_r; reflexivity. Qed. (** Zero property of multiplication *) -Theorem Zero_mult_left: (x:Z) (Zmult ZERO x) = ZERO. +Theorem Zmult_0_l : forall n:Z, Z0 * n = Z0. Proof. -Intro x; NewDestruct x; Reflexivity. +intro x; destruct x; reflexivity. Qed. -Theorem Zero_mult_right: (x:Z) (Zmult x ZERO) = ZERO. +Theorem Zmult_0_r : forall n:Z, n * Z0 = Z0. Proof. -Intro x; NewDestruct x; Reflexivity. +intro x; destruct x; reflexivity. Qed. -Hints Local Resolve Zero_mult_left Zero_mult_right. +Hint Local Resolve Zmult_0_l Zmult_0_r. -Lemma Zmult_n_O : (n:Z) ZERO=(Zmult n ZERO). +Lemma Zmult_0_r_reverse : forall n:Z, Z0 = n * Z0. Proof. -Intro x; NewDestruct x; Reflexivity. +intro x; destruct x; reflexivity. Qed. (** Commutativity of multiplication *) -Theorem Zmult_sym : (x,y:Z) (Zmult x y) = (Zmult y x). +Theorem Zmult_comm : forall n m:Z, n * m = m * n. Proof. -Intros x y; NewDestruct x as [|p|p]; NewDestruct y as [|q|q]; Simpl; - Try Rewrite (times_sym p q); Reflexivity. +intros x y; destruct x as [| p| p]; destruct y as [| q| q]; simpl in |- *; + try rewrite (Pmult_comm p q); reflexivity. Qed. (** Associativity of multiplication *) -Theorem Zmult_assoc : - (x,y,z:Z) (Zmult x (Zmult y z))= (Zmult (Zmult x y) z). +Theorem Zmult_assoc : forall n m p:Z, n * (m * p) = n * m * p. Proof. -Intros x y z; NewDestruct x; NewDestruct y; NewDestruct z; Simpl; - Try Rewrite times_assoc; Reflexivity. +intros x y z; destruct x; destruct y; destruct z; simpl in |- *; + try rewrite Pmult_assoc; reflexivity. Qed. -V7only [Notation Zmult_assoc_l := Zmult_assoc.]. -Lemma Zmult_assoc_r : (n,m,p:Z)((Zmult (Zmult n m) p) = (Zmult n (Zmult m p))). +Lemma Zmult_assoc_reverse : forall n m p:Z, n * m * p = n * (m * p). Proof. -Intros n m p; Rewrite Zmult_assoc; Trivial with arith. +intros n m p; rewrite Zmult_assoc; trivial with arith. Qed. (** Associativity mixed with commutativity *) -Theorem Zmult_permute : (n,m,p:Z)(Zmult n (Zmult m p)) = (Zmult m (Zmult n p)). +Theorem Zmult_permute : forall n m p:Z, n * (m * p) = m * (n * p). Proof. -Intros x y z; Rewrite -> (Zmult_assoc y x z); Rewrite -> (Zmult_sym y x). -Apply Zmult_assoc. +intros x y z; rewrite (Zmult_assoc y x z); rewrite (Zmult_comm y x). +apply Zmult_assoc. Qed. (** Z is integral *) -Theorem Zmult_eq: (x,y:Z) ~(x=ZERO) -> (Zmult y x) = ZERO -> y = ZERO. +Theorem Zmult_integral_l : forall n m:Z, n <> Z0 -> m * n = Z0 -> m = Z0. Proof. -Intros x y; NewDestruct x as [|p|p]. - Intro H; Absurd ZERO=ZERO; Trivial. - Intros _ H; NewDestruct y as [|q|q]; Reflexivity Orelse Discriminate. - Intros _ H; NewDestruct y as [|q|q]; Reflexivity Orelse Discriminate. +intros x y; destruct x as [| p| p]. + intro H; absurd (Z0 = Z0); trivial. + intros _ H; destruct y as [| q| q]; reflexivity || discriminate. + intros _ H; destruct y as [| q| q]; reflexivity || discriminate. Qed. -V7only [Set Implicit Arguments.]. -Theorem Zmult_zero : (x,y:Z)(Zmult x y)=ZERO -> x=ZERO \/ y=ZERO. +Theorem Zmult_integral : forall n m:Z, n * m = Z0 -> n = Z0 \/ m = Z0. Proof. -Intros x y; NewDestruct x; NewDestruct y; Auto; Simpl; Intro H; Discriminate H. +intros x y; destruct x; destruct y; auto; simpl in |- *; intro H; + discriminate H. Qed. -V7only [Unset Implicit Arguments.]. -Lemma Zmult_1_inversion_l : - (x,y:Z) (Zmult x y)=(POS xH) -> x=(POS xH) \/ x=(NEG xH). +Lemma Zmult_1_inversion_l : + forall n m:Z, n * m = Zpos 1 -> n = Zpos 1 \/ n = Zneg 1. Proof. -Intros x y; NewDestruct x as [|p|p]; Intro; [ Discriminate | Left | Right ]; - (NewDestruct y as [|q|q]; Try Discriminate; - Simpl in H; Injection H; Clear H; Intro H; - Rewrite times_one_inversion_l with 1:=H; Reflexivity). +intros x y; destruct x as [| p| p]; intro; [ discriminate | left | right ]; + (destruct y as [| q| q]; try discriminate; simpl in H; injection H; clear H; + intro H; rewrite Pmult_1_inversion_l with (1 := H); + reflexivity). Qed. (** Multiplication and Opposite *) -Theorem Zopp_Zmult_l : (x,y:Z)(Zopp (Zmult x y)) = (Zmult (Zopp x) y). +Theorem Zopp_mult_distr_l : forall n m:Z, - (n * m) = - n * m. Proof. -Intros x y; NewDestruct x; NewDestruct y; Reflexivity. +intros x y; destruct x; destruct y; reflexivity. Qed. -Theorem Zopp_Zmult_r : (x,y:Z)(Zopp (Zmult x y)) = (Zmult x (Zopp y)). -Intros x y; Rewrite (Zmult_sym x y); Rewrite Zopp_Zmult_l; Apply Zmult_sym. +Theorem Zopp_mult_distr_r : forall n m:Z, - (n * m) = n * - m. +intros x y; rewrite (Zmult_comm x y); rewrite Zopp_mult_distr_l; + apply Zmult_comm. Qed. -Lemma Zopp_Zmult: (x,y:Z) (Zmult (Zopp x) y) = (Zopp (Zmult x y)). +Lemma Zopp_mult_distr_l_reverse : forall n m:Z, - n * m = - (n * m). Proof. -Intros x y; Symmetry; Apply Zopp_Zmult_l. +intros x y; symmetry in |- *; apply Zopp_mult_distr_l. Qed. -Theorem Zmult_Zopp_left : (x,y:Z)(Zmult (Zopp x) y) = (Zmult x (Zopp y)). -Intros x y; Rewrite Zopp_Zmult; Rewrite Zopp_Zmult_r; Trivial with arith. +Theorem Zmult_opp_comm : forall n m:Z, - n * m = n * - m. +intros x y; rewrite Zopp_mult_distr_l_reverse; rewrite Zopp_mult_distr_r; + trivial with arith. Qed. -Theorem Zmult_Zopp_Zopp: (x,y:Z) (Zmult (Zopp x) (Zopp y)) = (Zmult x y). +Theorem Zmult_opp_opp : forall n m:Z, - n * - m = n * m. Proof. -Intros x y; NewDestruct x; NewDestruct y; Reflexivity. +intros x y; destruct x; destruct y; reflexivity. Qed. -Theorem Zopp_one : (x:Z)(Zopp x)=(Zmult x (NEG xH)). -Intro x; NewInduction x; Intros; Rewrite Zmult_sym; Auto with arith. +Theorem Zopp_eq_mult_neg_1 : forall n:Z, - n = n * Zneg 1. +intro x; induction x; intros; rewrite Zmult_comm; auto with arith. Qed. (** Distributivity of multiplication over addition *) -Lemma weak_Zmult_plus_distr_r: - (x:positive)(y,z:Z) - (Zmult (POS x) (Zplus y z)) = (Zplus (Zmult (POS x) y) (Zmult (POS x) z)). +Lemma weak_Zmult_plus_distr_r : + forall (p:positive) (n m:Z), Zpos p * (n + m) = Zpos p * n + Zpos p * m. Proof. -Intros x y' z';Case y';Case z';Auto with arith;Intros y z; - (Simpl; Rewrite times_add_distr; Trivial with arith) -Orelse - (Simpl; ElimPcompare z y; Intros E0;Rewrite E0; [ - Rewrite (compare_convert_EGAL z y E0); - Rewrite (convert_compare_EGAL (times x y)); Trivial with arith - | Cut (compare (times x z) (times x y) EGAL)=INFERIEUR; [ - Intros E;Rewrite E; Rewrite times_true_sub_distr; [ - Trivial with arith - | Apply ZC2;Assumption ] - | Apply convert_compare_INFERIEUR;Do 2 Rewrite times_convert; - Elim (ZL4 x);Intros h H1;Rewrite H1;Apply lt_mult_left; - Exact (compare_convert_INFERIEUR z y E0)] - | Cut (compare (times x z) (times x y) EGAL)=SUPERIEUR; [ - Intros E;Rewrite E; Rewrite times_true_sub_distr; Auto with arith - | Apply convert_compare_SUPERIEUR; Unfold gt; Do 2 Rewrite times_convert; - Elim (ZL4 x);Intros h H1;Rewrite H1;Apply lt_mult_left; - Exact (compare_convert_SUPERIEUR z y E0) ]]). +intros x y' z'; case y'; case z'; auto with arith; intros y z; + (simpl in |- *; rewrite Pmult_plus_distr_l; trivial with arith) || + (simpl in |- *; ElimPcompare z y; intros E0; rewrite E0; + [ rewrite (Pcompare_Eq_eq z y E0); rewrite (Pcompare_refl (x * y)); + trivial with arith + | cut ((x * z ?= x * y)%positive Eq = Lt); + [ intros E; rewrite E; rewrite Pmult_minus_distr_l; + [ trivial with arith | apply ZC2; assumption ] + | apply nat_of_P_lt_Lt_compare_complement_morphism; + do 2 rewrite nat_of_P_mult_morphism; elim (ZL4 x); + intros h H1; rewrite H1; apply mult_S_lt_compat_l; + exact (nat_of_P_lt_Lt_compare_morphism z y E0) ] + | cut ((x * z ?= x * y)%positive Eq = Gt); + [ intros E; rewrite E; rewrite Pmult_minus_distr_l; auto with arith + | apply nat_of_P_gt_Gt_compare_complement_morphism; unfold gt in |- *; + do 2 rewrite nat_of_P_mult_morphism; elim (ZL4 x); + intros h H1; rewrite H1; apply mult_S_lt_compat_l; + exact (nat_of_P_gt_Gt_compare_morphism z y E0) ] ]). Qed. -Theorem Zmult_plus_distr_r: - (x,y,z:Z) (Zmult x (Zplus y z)) = (Zplus (Zmult x y) (Zmult x z)). +Theorem Zmult_plus_distr_r : forall n m p:Z, n * (m + p) = n * m + n * p. Proof. -Intros x y z; Case x; [ - Auto with arith -| Intros x';Apply weak_Zmult_plus_distr_r -| Intros p; Apply Zopp_intro; Rewrite Zopp_Zplus; - Do 3 Rewrite <- Zopp_Zmult; Rewrite Zopp_NEG; - Apply weak_Zmult_plus_distr_r ]. +intros x y z; case x; + [ auto with arith + | intros x'; apply weak_Zmult_plus_distr_r + | intros p; apply Zopp_inj; rewrite Zopp_plus_distr; + do 3 rewrite <- Zopp_mult_distr_l_reverse; rewrite Zopp_neg; + apply weak_Zmult_plus_distr_r ]. Qed. -Theorem Zmult_plus_distr_l : - (n,m,p:Z)((Zmult (Zplus n m) p)=(Zplus (Zmult n p) (Zmult m p))). +Theorem Zmult_plus_distr_l : forall n m p:Z, (n + m) * p = n * p + m * p. Proof. -Intros n m p;Rewrite Zmult_sym;Rewrite Zmult_plus_distr_r; -Do 2 Rewrite -> (Zmult_sym p); Trivial with arith. +intros n m p; rewrite Zmult_comm; rewrite Zmult_plus_distr_r; + do 2 rewrite (Zmult_comm p); trivial with arith. Qed. (** Distributivity of multiplication over subtraction *) -Lemma Zmult_Zminus_distr_l : - (x,y,z:Z)((Zmult (Zminus x y) z)=(Zminus (Zmult x z) (Zmult y z))). +Lemma Zmult_minus_distr_r : forall n m p:Z, (n - m) * p = n * p - m * p. Proof. -Intros x y z; Unfold Zminus. -Rewrite <- Zopp_Zmult. -Apply Zmult_plus_distr_l. +intros x y z; unfold Zminus in |- *. +rewrite <- Zopp_mult_distr_l_reverse. +apply Zmult_plus_distr_l. Qed. -V7only [Notation Zmult_minus_distr := Zmult_Zminus_distr_l.]. -Lemma Zmult_Zminus_distr_r : - (x,y,z:Z)(Zmult z (Zminus x y)) = (Zminus (Zmult z x) (Zmult z y)). +Lemma Zmult_minus_distr_l : forall n m p:Z, p * (n - m) = p * n - p * m. Proof. -Intros x y z; Rewrite (Zmult_sym z (Zminus x y)). -Rewrite (Zmult_sym z x). -Rewrite (Zmult_sym z y). -Apply Zmult_Zminus_distr_l. +intros x y z; rewrite (Zmult_comm z (x - y)). +rewrite (Zmult_comm z x). +rewrite (Zmult_comm z y). +apply Zmult_minus_distr_r. Qed. (** Simplification of multiplication for non-zero integers *) -V7only [Set Implicit Arguments.]. -Lemma Zmult_reg_left : (x,y,z:Z) z<>ZERO -> (Zmult z x)=(Zmult z y) -> x=y. +Lemma Zmult_reg_l : forall n m p:Z, p <> Z0 -> p * n = p * m -> n = m. Proof. -Intros x y z H H0. -Generalize (Zeq_Zminus H0). -Intro. -Apply Zminus_Zeq. -Rewrite <- Zmult_Zminus_distr_r in H1. -Clear H0; NewDestruct (Zmult_zero H1). -Contradiction. -Trivial. +intros x y z H H0. +generalize (Zeq_minus _ _ H0). +intro. +apply Zminus_eq. +rewrite <- Zmult_minus_distr_l in H1. +clear H0; destruct (Zmult_integral _ _ H1). +contradiction. +trivial. Qed. -Lemma Zmult_reg_right : (x,y,z:Z) z<>ZERO -> (Zmult x z)=(Zmult y z) -> x=y. +Lemma Zmult_reg_r : forall n m p:Z, p <> Z0 -> n * p = m * p -> n = m. Proof. -Intros x y z Hz. -Rewrite (Zmult_sym x z). -Rewrite (Zmult_sym y z). -Intro; Apply Zmult_reg_left with z; Assumption. +intros x y z Hz. +rewrite (Zmult_comm x z). +rewrite (Zmult_comm y z). +intro; apply Zmult_reg_l with z; assumption. Qed. -V7only [Unset Implicit Arguments.]. (** Addition and multiplication by 2 *) -Lemma Zplus_Zmult_2 : (x:Z) (Zplus x x) = (Zmult x (POS (xO xH))). +Lemma Zplus_diag_eq_mult_2 : forall n:Z, n + n = n * Zpos 2. Proof. -Intros x; Pattern 1 2 x ; Rewrite <- (Zmult_n_1 x); -Rewrite <- Zmult_plus_distr_r; Reflexivity. +intros x; pattern x at 1 2 in |- *; rewrite <- (Zmult_1_r x); + rewrite <- Zmult_plus_distr_r; reflexivity. Qed. (** Multiplication and successor *) -Lemma Zmult_succ_r : (n,m:Z) (Zmult n (Zs m))=(Zplus (Zmult n m) n). +Lemma Zmult_succ_r : forall n m:Z, n * Zsucc m = n * m + n. Proof. -Intros n m;Unfold Zs; Rewrite Zmult_plus_distr_r; -Rewrite (Zmult_sym n (POS xH));Rewrite Zmult_one; Trivial with arith. +intros n m; unfold Zsucc in |- *; rewrite Zmult_plus_distr_r; + rewrite (Zmult_comm n (Zpos 1)); rewrite Zmult_1_l; + trivial with arith. Qed. -Lemma Zmult_n_Sm : (n,m:Z) (Zplus (Zmult n m) n)=(Zmult n (Zs m)). +Lemma Zmult_succ_r_reverse : forall n m:Z, n * m + n = n * Zsucc m. Proof. -Intros; Symmetry; Apply Zmult_succ_r. +intros; symmetry in |- *; apply Zmult_succ_r. Qed. -Lemma Zmult_succ_l : (n,m:Z) (Zmult (Zs n) m)=(Zplus (Zmult n m) m). +Lemma Zmult_succ_l : forall n m:Z, Zsucc n * m = n * m + m. Proof. -Intros n m; Unfold Zs; Rewrite Zmult_plus_distr_l; Rewrite Zmult_1_n; -Trivial with arith. +intros n m; unfold Zsucc in |- *; rewrite Zmult_plus_distr_l; + rewrite Zmult_1_l; trivial with arith. Qed. -Lemma Zmult_Sm_n : (n,m:Z) (Zplus (Zmult n m) m)=(Zmult (Zs n) m). +Lemma Zmult_succ_l_reverse : forall n m:Z, n * m + m = Zsucc n * m. Proof. -Intros; Symmetry; Apply Zmult_succ_l. +intros; symmetry in |- *; apply Zmult_succ_l. Qed. (** Misc redundant properties *) -Lemma Z_eq_mult: - (x,y:Z) y = ZERO -> (Zmult y x) = ZERO. -Intros x y H; Rewrite H; Auto with arith. +Lemma Z_eq_mult : forall n m:Z, m = Z0 -> m * n = Z0. +intros x y H; rewrite H; auto with arith. Qed. (**********************************************************************) (** Relating binary positive numbers and binary integers *) -Lemma POS_xI : (p:positive) (POS (xI p))=(Zplus (Zmult (POS (xO xH)) (POS p)) (POS xH)). +Lemma Zpos_xI : forall p:positive, Zpos (xI p) = Zpos 2 * Zpos p + Zpos 1. Proof. -Intro; Apply refl_equal. +intro; apply refl_equal. Qed. -Lemma POS_xO : (p:positive) (POS (xO p))=(Zmult (POS (xO xH)) (POS p)). +Lemma Zpos_xO : forall p:positive, Zpos (xO p) = Zpos 2 * Zpos p. Proof. -Intro; Apply refl_equal. +intro; apply refl_equal. Qed. -Lemma NEG_xI : (p:positive) (NEG (xI p))=(Zminus (Zmult (POS (xO xH)) (NEG p)) (POS xH)). +Lemma Zneg_xI : forall p:positive, Zneg (xI p) = Zpos 2 * Zneg p - Zpos 1. Proof. -Intro; Apply refl_equal. +intro; apply refl_equal. Qed. -Lemma NEG_xO : (p:positive) (NEG (xO p))=(Zmult (POS (xO xH)) (NEG p)). +Lemma Zneg_xO : forall p:positive, Zneg (xO p) = Zpos 2 * Zneg p. Proof. -Reflexivity. +reflexivity. Qed. -Lemma POS_add : (p,p':positive)(POS (add p p'))=(Zplus (POS p) (POS p')). +Lemma Zpos_plus_distr : forall p q:positive, Zpos (p + q) = Zpos p + Zpos q. Proof. -Intros p p'; NewDestruct p; NewDestruct p'; Reflexivity. +intros p p'; destruct p; + [ destruct p' as [p0| p0| ] + | destruct p' as [p0| p0| ] + | destruct p' as [p| p| ] ]; reflexivity. Qed. -Lemma NEG_add : (p,p':positive)(NEG (add p p'))=(Zplus (NEG p) (NEG p')). +Lemma Zneg_plus_distr : forall p q:positive, Zneg (p + q) = Zneg p + Zneg q. Proof. -Intros p p'; NewDestruct p; NewDestruct p'; Reflexivity. +intros p p'; destruct p; + [ destruct p' as [p0| p0| ] + | destruct p' as [p0| p0| ] + | destruct p' as [p| p| ] ]; reflexivity. Qed. (**********************************************************************) (** Order relations *) -Definition Zlt := [x,y:Z](Zcompare x y) = INFERIEUR. -Definition Zgt := [x,y:Z](Zcompare x y) = SUPERIEUR. -Definition Zle := [x,y:Z]~(Zcompare x y) = SUPERIEUR. -Definition Zge := [x,y:Z]~(Zcompare x y) = INFERIEUR. -Definition Zne := [x,y:Z] ~(x=y). +Definition Zlt (x y:Z) := (x ?= y) = Lt. +Definition Zgt (x y:Z) := (x ?= y) = Gt. +Definition Zle (x y:Z) := (x ?= y) <> Gt. +Definition Zge (x y:Z) := (x ?= y) <> Lt. +Definition Zne (x y:Z) := x <> y. -V8Infix "<=" Zle : Z_scope. -V8Infix "<" Zlt : Z_scope. -V8Infix ">=" Zge : Z_scope. -V8Infix ">" Zgt : Z_scope. +Infix "<=" := Zle : Z_scope. +Infix "<" := Zlt : Z_scope. +Infix ">=" := Zge : Z_scope. +Infix ">" := Zgt : Z_scope. -V8Notation "x <= y <= z" := (Zle x y)/\(Zle y z) :Z_scope. -V8Notation "x <= y < z" := (Zle x y)/\(Zlt y z) :Z_scope. -V8Notation "x < y < z" := (Zlt x y)/\(Zlt y z) :Z_scope. -V8Notation "x < y <= z" := (Zlt x y)/\(Zle y z) :Z_scope. +Notation "x <= y <= z" := (x <= y /\ y <= z) : Z_scope. +Notation "x <= y < z" := (x <= y /\ y < z) : Z_scope. +Notation "x < y < z" := (x < y /\ y < z) : Z_scope. +Notation "x < y <= z" := (x < y /\ y <= z) : Z_scope. (**********************************************************************) (** Absolute value on integers *) -Definition absolu [x:Z] : nat := - Cases x of - ZERO => O - | (POS p) => (convert p) - | (NEG p) => (convert p) +Definition Zabs_nat (x:Z) : nat := + match x with + | Z0 => 0%nat + | Zpos p => nat_of_P p + | Zneg p => nat_of_P p end. -Definition Zabs [z:Z] : Z := - Cases z of - ZERO => ZERO - | (POS p) => (POS p) - | (NEG p) => (POS p) +Definition Zabs (z:Z) : Z := + match z with + | Z0 => Z0 + | Zpos p => Zpos p + | Zneg p => Zpos p end. (**********************************************************************) (** From [nat] to [Z] *) -Definition inject_nat := - [x:nat]Cases x of - O => ZERO - | (S y) => (POS (anti_convert y)) - end. +Definition Z_of_nat (x:nat) := + match x with + | O => Z0 + | S y => Zpos (P_of_succ_nat y) + end. -Require BinNat. +Require Import BinNat. -Definition entier_of_Z := - [z:Z]Cases z of ZERO => Nul | (POS p) => (Pos p) | (NEG p) => (Pos p) end. +Definition Zabs_N (z:Z) := + match z with + | Z0 => 0%N + | Zpos p => Npos p + | Zneg p => Npos p + end. -Definition Z_of_entier := - [x:entier]Cases x of Nul => ZERO | (Pos p) => (POS p) end. +Definition Z_of_N (x:N) := match x with + | N0 => Z0 + | Npos p => Zpos p + end. \ No newline at end of file diff --git a/theories/ZArith/Wf_Z.v b/theories/ZArith/Wf_Z.v index eecfc42b2d..4c2efceb17 100644 --- a/theories/ZArith/Wf_Z.v +++ b/theories/ZArith/Wf_Z.v @@ -8,14 +8,12 @@ (*i $Id$ i*) -Require BinInt. -Require Zcompare. -Require Zorder. -Require Znat. -Require Zmisc. -Require Zsyntax. -Require Wf_nat. -V7only [Import Z_scope.]. +Require Import BinInt. +Require Import Zcompare. +Require Import Zorder. +Require Import Znat. +Require Import Zmisc. +Require Import Wf_nat. Open Local Scope Z_scope. (** Our purpose is to write an induction shema for {0,1,2,...} @@ -36,86 +34,83 @@ Open Local Scope Z_scope. >> Then the diagram will be closed and the theorem proved. *) -Lemma inject_nat_complete : - (x:Z)`0 <= x` -> (EX n:nat | x=(inject_nat n)). -Intro x; NewDestruct x; Intros; -[ Exists O; Auto with arith -| Specialize (ZL4 p); Intros Hp; Elim Hp; Intros; - Exists (S x); Intros; Simpl; - Specialize (bij1 x); Intro Hx0; - Rewrite <- H0 in Hx0; - Apply f_equal with f:=POS; - Apply convert_intro; Auto with arith -| Absurd `0 <= (NEG p)`; - [ Unfold Zle; Simpl; Do 2 (Unfold not); Auto with arith - | Assumption] -]. +Lemma Z_of_nat_complete : + forall x:Z, 0 <= x -> exists n : nat | x = Z_of_nat n. +intro x; destruct x; intros; + [ exists 0%nat; auto with arith + | specialize (ZL4 p); intros Hp; elim Hp; intros; exists (S x); intros; + simpl in |- *; specialize (nat_of_P_o_P_of_succ_nat_eq_succ x); + intro Hx0; rewrite <- H0 in Hx0; apply f_equal with (f := Zpos); + apply nat_of_P_inj; auto with arith + | absurd (0 <= Zneg p); + [ unfold Zle in |- *; simpl in |- *; do 2 unfold not in |- *; + auto with arith + | assumption ] ]. Qed. -Lemma ZL4_inf: (y:positive) { h:nat | (convert y)=(S h) }. -Intro y; NewInduction y as [p H|p H1|]; [ - Elim H; Intros x H1; Exists (plus (S x) (S x)); - Unfold convert ;Simpl; Rewrite ZL0; Rewrite ZL2; Unfold convert in H1; - Rewrite H1; Auto with arith -| Elim H1;Intros x H2; Exists (plus x (S x)); Unfold convert; - Simpl; Rewrite ZL0; Rewrite ZL2;Unfold convert in H2; Rewrite H2; Auto with arith -| Exists O ;Auto with arith]. +Lemma ZL4_inf : forall y:positive, {h : nat | nat_of_P y = S h}. +intro y; induction y as [p H| p H1| ]; + [ elim H; intros x H1; exists (S x + S x)%nat; unfold nat_of_P in |- *; + simpl in |- *; rewrite ZL0; rewrite Pmult_nat_r_plus_morphism; + unfold nat_of_P in H1; rewrite H1; auto with arith + | elim H1; intros x H2; exists (x + S x)%nat; unfold nat_of_P in |- *; + simpl in |- *; rewrite ZL0; rewrite Pmult_nat_r_plus_morphism; + unfold nat_of_P in H2; rewrite H2; auto with arith + | exists 0%nat; auto with arith ]. Qed. -Lemma inject_nat_complete_inf : - (x:Z)`0 <= x` -> { n:nat | (x=(inject_nat n)) }. -Intro x; NewDestruct x; Intros; -[ Exists O; Auto with arith -| Specialize (ZL4_inf p); Intros Hp; Elim Hp; Intros x0 H0; - Exists (S x0); Intros; Simpl; - Specialize (bij1 x0); Intro Hx0; - Rewrite <- H0 in Hx0; - Apply f_equal with f:=POS; - Apply convert_intro; Auto with arith -| Absurd `0 <= (NEG p)`; - [ Unfold Zle; Simpl; Do 2 (Unfold not); Auto with arith - | Assumption] -]. +Lemma Z_of_nat_complete_inf : + forall x:Z, 0 <= x -> {n : nat | x = Z_of_nat n}. +intro x; destruct x; intros; + [ exists 0%nat; auto with arith + | specialize (ZL4_inf p); intros Hp; elim Hp; intros x0 H0; exists (S x0); + intros; simpl in |- *; specialize (nat_of_P_o_P_of_succ_nat_eq_succ x0); + intro Hx0; rewrite <- H0 in Hx0; apply f_equal with (f := Zpos); + apply nat_of_P_inj; auto with arith + | absurd (0 <= Zneg p); + [ unfold Zle in |- *; simpl in |- *; do 2 unfold not in |- *; + auto with arith + | assumption ] ]. Qed. -Lemma inject_nat_prop : - (P:Z->Prop)((n:nat)(P (inject_nat n))) -> - (x:Z) `0 <= x` -> (P x). -Intros P H x H0. -Specialize (inject_nat_complete x H0). -Intros Hn; Elim Hn; Intros. -Rewrite -> H1; Apply H. +Lemma Z_of_nat_prop : + forall P:Z -> Prop, + (forall n:nat, P (Z_of_nat n)) -> forall x:Z, 0 <= x -> P x. +intros P H x H0. +specialize (Z_of_nat_complete x H0). +intros Hn; elim Hn; intros. +rewrite H1; apply H. Qed. -Lemma inject_nat_set : - (P:Z->Set)((n:nat)(P (inject_nat n))) -> - (x:Z) `0 <= x` -> (P x). -Intros P H x H0. -Specialize (inject_nat_complete_inf x H0). -Intros Hn; Elim Hn; Intros. -Rewrite -> p; Apply H. +Lemma Z_of_nat_set : + forall P:Z -> Set, + (forall n:nat, P (Z_of_nat n)) -> forall x:Z, 0 <= x -> P x. +intros P H x H0. +specialize (Z_of_nat_complete_inf x H0). +intros Hn; elim Hn; intros. +rewrite p; apply H. Qed. -Lemma natlike_ind : (P:Z->Prop) (P `0`) -> - ((x:Z)(`0 <= x` -> (P x) -> (P (Zs x)))) -> - (x:Z) `0 <= x` -> (P x). -Intros P H H0 x H1; Apply inject_nat_prop; -[ Induction n; - [ Simpl; Assumption - | Intros; Rewrite -> (inj_S n0); - Exact (H0 (inject_nat n0) (ZERO_le_inj n0) H2) ] -| Assumption]. +Lemma natlike_ind : + forall P:Z -> Prop, + P 0 -> + (forall x:Z, 0 <= x -> P x -> P (Zsucc x)) -> forall x:Z, 0 <= x -> P x. +intros P H H0 x H1; apply Z_of_nat_prop; + [ simple induction n; + [ simpl in |- *; assumption + | intros; rewrite (inj_S n0); exact (H0 (Z_of_nat n0) (Zle_0_nat n0) H2) ] + | assumption ]. Qed. -Lemma natlike_rec : (P:Z->Set) (P `0`) -> - ((x:Z)(`0 <= x` -> (P x) -> (P (Zs x)))) -> - (x:Z) `0 <= x` -> (P x). -Intros P H H0 x H1; Apply inject_nat_set; -[ Induction n; - [ Simpl; Assumption - | Intros; Rewrite -> (inj_S n0); - Exact (H0 (inject_nat n0) (ZERO_le_inj n0) H2) ] -| Assumption]. +Lemma natlike_rec : + forall P:Z -> Set, + P 0 -> + (forall x:Z, 0 <= x -> P x -> P (Zsucc x)) -> forall x:Z, 0 <= x -> P x. +intros P H H0 x H1; apply Z_of_nat_set; + [ simple induction n; + [ simpl in |- *; assumption + | intros; rewrite (inj_S n0); exact (H0 (Z_of_nat n0) (Zle_0_nat n0) H2) ] + | assumption ]. Qed. Section Efficient_Rec. @@ -123,72 +118,87 @@ Section Efficient_Rec. (** [natlike_rec2] is the same as [natlike_rec], but with a different proof, designed to give a better extracted term. *) -Local R := [a,b:Z]`0<=a`/\`a (convert p) | ZERO => O | (NEG _) => O end. -Apply well_founded_lt_compat with f. -Unfold R f; Clear f R. -Intros x y; Case x; Intros; Elim H; Clear H. -Case y; Intros; Apply compare_convert_O Orelse Inversion H0. -Case y; Intros; Apply compare_convert_INFERIEUR Orelse Inversion H0; Auto. -Intros; Elim H; Auto. +set + (f := + fun z => + match z with + | Zpos p => nat_of_P p + | Z0 => 0%nat + | Zneg _ => 0%nat + end) in *. +apply well_founded_lt_compat with f. +unfold R, f in |- *; clear f R. +intros x y; case x; intros; elim H; clear H. +case y; intros; apply lt_O_nat_of_P || inversion H0. +case y; intros; apply nat_of_P_lt_Lt_compare_morphism || inversion H0; auto. +intros; elim H; auto. Qed. -Lemma natlike_rec2 : (P:Z->Type)(P `0`) -> - ((z:Z)`0<=z` -> (P z) -> (P (Zs z))) -> (z:Z)`0<=z` -> (P z). +Lemma natlike_rec2 : + forall P:Z -> Type, + P 0 -> + (forall z:Z, 0 <= z -> P z -> P (Zsucc z)) -> forall z:Z, 0 <= z -> P z. Proof. -Intros P Ho Hrec z; Pattern z; Apply (well_founded_induction_type Z R R_wf). -Intro x; Case x. -Trivial. -Intros. -Assert `0<=(Zpred (POS p))`. -Apply Zlt_ZERO_pred_le_ZERO; Unfold Zlt; Simpl; Trivial. -Rewrite Zs_pred. -Apply Hrec. -Auto. -Apply X; Auto; Unfold R; Intuition; Apply Zlt_pred_n_n. -Intros; Elim H; Simpl; Trivial. +intros P Ho Hrec z; pattern z in |- *; + apply (well_founded_induction_type R_wf). +intro x; case x. +trivial. +intros. +assert (0 <= Zpred (Zpos p)). +apply Zorder.Zlt_0_le_0_pred; unfold Zlt in |- *; simpl in |- *; trivial. +rewrite Zsucc_pred. +apply Hrec. +auto. +apply X; auto; unfold R in |- *; intuition; apply Zlt_pred. +intros; elim H; simpl in |- *; trivial. Qed. (** A variant of the previous using [Zpred] instead of [Zs]. *) -Lemma natlike_rec3 : (P:Z->Type)(P `0`) -> - ((z:Z)`0 (P (Zpred z)) -> (P z)) -> (z:Z)`0<=z` -> (P z). +Lemma natlike_rec3 : + forall P:Z -> Type, + P 0 -> + (forall z:Z, 0 < z -> P (Zpred z) -> P z) -> forall z:Z, 0 <= z -> P z. Proof. -Intros P Ho Hrec z; Pattern z; Apply (well_founded_induction_type Z R R_wf). -Intro x; Case x. -Trivial. -Intros; Apply Hrec. -Unfold Zlt; Trivial. -Assert `0<=(Zpred (POS p))`. -Apply Zlt_ZERO_pred_le_ZERO; Unfold Zlt; Simpl; Trivial. -Apply X; Auto; Unfold R; Intuition; Apply Zlt_pred_n_n. -Intros; Elim H; Simpl; Trivial. +intros P Ho Hrec z; pattern z in |- *; + apply (well_founded_induction_type R_wf). +intro x; case x. +trivial. +intros; apply Hrec. +unfold Zlt in |- *; trivial. +assert (0 <= Zpred (Zpos p)). +apply Zorder.Zlt_0_le_0_pred; unfold Zlt in |- *; simpl in |- *; trivial. +apply X; auto; unfold R in |- *; intuition; apply Zlt_pred. +intros; elim H; simpl in |- *; trivial. Qed. (** A more general induction principal using [Zlt]. *) -Lemma Z_lt_rec : (P:Z->Type) - ((x:Z)((y:Z)`0 <= y < x`->(P y))->(P x)) -> (x:Z)`0 <= x`->(P x). +Lemma Z_lt_rec : + forall P:Z -> Type, + (forall x:Z, (forall y:Z, 0 <= y < x -> P y) -> P x) -> + forall x:Z, 0 <= x -> P x. Proof. -Intros P Hrec z; Pattern z; Apply (well_founded_induction_type Z R R_wf). -Intro x; Case x; Intros. -Apply Hrec; Intros. -Assert H2: `0<0`. - Apply Zle_lt_trans with y; Intuition. -Inversion H2. -Firstorder. -Unfold Zle Zcompare in H; Elim H; Auto. +intros P Hrec z; pattern z in |- *; apply (well_founded_induction_type R_wf). +intro x; case x; intros. +apply Hrec; intros. +assert (H2 : 0 < 0). + apply Zle_lt_trans with y; intuition. +inversion H2. +firstorder. +unfold Zle, Zcompare in H; elim H; auto. Defined. -Lemma Z_lt_induction : - (P:Z->Prop) - ((x:Z)((y:Z)`0 <= y < x`->(P y))->(P x)) - -> (x:Z)`0 <= x`->(P x). +Lemma Z_lt_induction : + forall P:Z -> Prop, + (forall x:Z, (forall y:Z, 0 <= y < x -> P y) -> P x) -> + forall x:Z, 0 <= x -> P x. Proof. -Exact Z_lt_rec. +exact Z_lt_rec. Qed. -End Efficient_Rec. +End Efficient_Rec. \ No newline at end of file diff --git a/theories/ZArith/ZArith.v b/theories/ZArith/ZArith.v index f85b0bddd3..bc58a1a4b2 100644 --- a/theories/ZArith/ZArith.v +++ b/theories/ZArith/ZArith.v @@ -19,4 +19,4 @@ Require Export Zsqrt. Require Export Zpower. Require Export Zdiv. Require Export Zlogarithm. -Require Export Zbool. +Require Export Zbool. \ No newline at end of file diff --git a/theories/ZArith/ZArith_base.v b/theories/ZArith/ZArith_base.v index 97f4c3f3ed..ec57dda578 100644 --- a/theories/ZArith/ZArith_base.v +++ b/theories/ZArith/ZArith_base.v @@ -12,10 +12,6 @@ These are the basic modules, required by [Omega] and [Ring] for instance. The full library is [ZArith]. *) -V7only [ -Require Export fast_integer. -Require Export zarith_aux. -]. Require Export BinPos. Require Export BinNat. Require Export BinInt. @@ -26,14 +22,13 @@ Require Export Zmin. Require Export Zabs. Require Export Znat. Require Export auxiliary. -Require Export Zsyntax. Require Export ZArith_dec. Require Export Zbool. Require Export Zmisc. Require Export Wf_Z. -Hints Resolve Zle_n Zplus_sym Zplus_assoc Zmult_sym Zmult_assoc - Zero_left Zero_right Zmult_one Zplus_inverse_l Zplus_inverse_r - Zmult_plus_distr_l Zmult_plus_distr_r : zarith. +Hint Resolve Zle_refl Zplus_comm Zplus_assoc Zmult_comm Zmult_assoc Zplus_0_l + Zplus_0_r Zmult_1_l Zplus_opp_l Zplus_opp_r Zmult_plus_distr_l + Zmult_plus_distr_r: zarith. -Require Export Zhints. +Require Export Zhints. \ No newline at end of file diff --git a/theories/ZArith/ZArith_dec.v b/theories/ZArith/ZArith_dec.v index e8f83fe1a1..ed323a641b 100644 --- a/theories/ZArith/ZArith_dec.v +++ b/theories/ZArith/ZArith_dec.v @@ -8,236 +8,219 @@ (*i $Id$ i*) -Require Sumbool. +Require Import Sumbool. -Require BinInt. -Require Zorder. -Require Zcompare. -Require Zsyntax. -V7only [Import Z_scope.]. +Require Import BinInt. +Require Import Zorder. +Require Import Zcompare. Open Local Scope Z_scope. -Lemma Dcompare_inf : (r:relation) {r=EGAL} + {r=INFERIEUR} + {r=SUPERIEUR}. +Lemma Dcompare_inf : forall r:comparison, {r = Eq} + {r = Lt} + {r = Gt}. Proof. -Induction r; Auto with arith. +simple induction r; auto with arith. Defined. Lemma Zcompare_rec : - (P:Set)(x,y:Z) - ((Zcompare x y)=EGAL -> P) -> - ((Zcompare x y)=INFERIEUR -> P) -> - ((Zcompare x y)=SUPERIEUR -> P) -> - P. + forall (P:Set) (n m:Z), + ((n ?= m) = Eq -> P) -> ((n ?= m) = Lt -> P) -> ((n ?= m) = Gt -> P) -> P. Proof. -Intros P x y H1 H2 H3. -Elim (Dcompare_inf (Zcompare x y)). -Intro H. Elim H; Auto with arith. Auto with arith. +intros P x y H1 H2 H3. +elim (Dcompare_inf (x ?= y)). +intro H. elim H; auto with arith. auto with arith. Defined. Section decidability. -Variables x,y : Z. +Variables x y : Z. (** Decidability of equality on binary integers *) -Definition Z_eq_dec : {x=y}+{~x=y}. +Definition Z_eq_dec : {x = y} + {x <> y}. Proof. -Apply Zcompare_rec with x:=x y:=y. -Intro. Left. Elim (Zcompare_EGAL x y); Auto with arith. -Intro H3. Right. Elim (Zcompare_EGAL x y). Intros H1 H2. Unfold not. Intro H4. - Rewrite (H2 H4) in H3. Discriminate H3. -Intro H3. Right. Elim (Zcompare_EGAL x y). Intros H1 H2. Unfold not. Intro H4. - Rewrite (H2 H4) in H3. Discriminate H3. +apply Zcompare_rec with (n := x) (m := y). +intro. left. elim (Zcompare_Eq_iff_eq x y); auto with arith. +intro H3. right. elim (Zcompare_Eq_iff_eq x y). intros H1 H2. unfold not in |- *. intro H4. + rewrite (H2 H4) in H3. discriminate H3. +intro H3. right. elim (Zcompare_Eq_iff_eq x y). intros H1 H2. unfold not in |- *. intro H4. + rewrite (H2 H4) in H3. discriminate H3. Defined. (** Decidability of order on binary integers *) -Definition Z_lt_dec : {(Zlt x y)}+{~(Zlt x y)}. +Definition Z_lt_dec : {x < y} + {~ x < y}. Proof. -Unfold Zlt. -Apply Zcompare_rec with x:=x y:=y; Intro H. -Right. Rewrite H. Discriminate. -Left; Assumption. -Right. Rewrite H. Discriminate. +unfold Zlt in |- *. +apply Zcompare_rec with (n := x) (m := y); intro H. +right. rewrite H. discriminate. +left; assumption. +right. rewrite H. discriminate. Defined. -Definition Z_le_dec : {(Zle x y)}+{~(Zle x y)}. +Definition Z_le_dec : {x <= y} + {~ x <= y}. Proof. -Unfold Zle. -Apply Zcompare_rec with x:=x y:=y; Intro H. -Left. Rewrite H. Discriminate. -Left. Rewrite H. Discriminate. -Right. Tauto. +unfold Zle in |- *. +apply Zcompare_rec with (n := x) (m := y); intro H. +left. rewrite H. discriminate. +left. rewrite H. discriminate. +right. tauto. Defined. -Definition Z_gt_dec : {(Zgt x y)}+{~(Zgt x y)}. +Definition Z_gt_dec : {x > y} + {~ x > y}. Proof. -Unfold Zgt. -Apply Zcompare_rec with x:=x y:=y; Intro H. -Right. Rewrite H. Discriminate. -Right. Rewrite H. Discriminate. -Left; Assumption. +unfold Zgt in |- *. +apply Zcompare_rec with (n := x) (m := y); intro H. +right. rewrite H. discriminate. +right. rewrite H. discriminate. +left; assumption. Defined. -Definition Z_ge_dec : {(Zge x y)}+{~(Zge x y)}. +Definition Z_ge_dec : {x >= y} + {~ x >= y}. Proof. -Unfold Zge. -Apply Zcompare_rec with x:=x y:=y; Intro H. -Left. Rewrite H. Discriminate. -Right. Tauto. -Left. Rewrite H. Discriminate. +unfold Zge in |- *. +apply Zcompare_rec with (n := x) (m := y); intro H. +left. rewrite H. discriminate. +right. tauto. +left. rewrite H. discriminate. Defined. -Definition Z_lt_ge_dec : {`x < y`}+{`x >= y`}. +Definition Z_lt_ge_dec : {x < y} + {x >= y}. Proof. -Exact Z_lt_dec. +exact Z_lt_dec. Defined. -V7only [ (* From Zextensions *) ]. -Lemma Z_lt_le_dec: {`x < y`}+{`y <= x`}. +Lemma Z_lt_le_dec : {x < y} + {y <= x}. Proof. -Intros. -Elim Z_lt_ge_dec. -Intros; Left; Assumption. -Intros; Right; Apply Zge_le; Assumption. +intros. +elim Z_lt_ge_dec. +intros; left; assumption. +intros; right; apply Zge_le; assumption. Qed. -Definition Z_le_gt_dec : {`x <= y`}+{`x > y`}. +Definition Z_le_gt_dec : {x <= y} + {x > y}. Proof. -Elim Z_le_dec; Auto with arith. -Intro. Right. Apply not_Zle; Auto with arith. +elim Z_le_dec; auto with arith. +intro. right. apply Znot_le_gt; auto with arith. Defined. -Definition Z_gt_le_dec : {`x > y`}+{`x <= y`}. +Definition Z_gt_le_dec : {x > y} + {x <= y}. Proof. -Exact Z_gt_dec. +exact Z_gt_dec. Defined. -Definition Z_ge_lt_dec : {`x >= y`}+{`x < y`}. +Definition Z_ge_lt_dec : {x >= y} + {x < y}. Proof. -Elim Z_ge_dec; Auto with arith. -Intro. Right. Apply not_Zge; Auto with arith. +elim Z_ge_dec; auto with arith. +intro. right. apply Znot_ge_lt; auto with arith. Defined. -Definition Z_le_lt_eq_dec : `x <= y` -> {`x < y`}+{x=y}. +Definition Z_le_lt_eq_dec : x <= y -> {x < y} + {x = y}. Proof. -Intro H. -Apply Zcompare_rec with x:=x y:=y. -Intro. Right. Elim (Zcompare_EGAL x y); Auto with arith. -Intro. Left. Elim (Zcompare_EGAL x y); Auto with arith. -Intro H1. Absurd `x > y`; Auto with arith. +intro H. +apply Zcompare_rec with (n := x) (m := y). +intro. right. elim (Zcompare_Eq_iff_eq x y); auto with arith. +intro. left. elim (Zcompare_Eq_iff_eq x y); auto with arith. +intro H1. absurd (x > y); auto with arith. Defined. End decidability. (** Cotransitivity of order on binary integers *) -Lemma Zlt_cotrans:(n,m:Z)`n(p:Z){`n{`0{`x<0`}+{`y<0`}. -Proof. - Intros x y H; - Case (Zlt_cotrans `x+y` `0` H x); - Intro Hxy; - [ Right; - Apply Zsimpl_lt_plus_l with p:=`x`; - Rewrite Zero_right - | Left - ]; - Assumption. -Defined. - -Lemma not_Zeq_inf:(x,y:Z)`x<>y`->{`xy`}+{`x=y`}. -Proof. - Intros x y. - Case (Z_lt_ge_dec x y). - Intro H. - Left. - Left. - Assumption. - Intro H. - Generalize (Zge_le ? ? H). - Intro H0. - Case (Z_le_lt_eq_dec y x H0). - Intro H1. - Left. - Right. - Apply Zlt_gt. - Assumption. - Intro. - Right. - Symmetry. - Assumption. -Defined. - - -Lemma Z_dec':(x,y:Z){`x 0`)}. -Proof. -Exact [x:Z](Z_eq_dec x ZERO). -Defined. - -Definition Z_notzerop := [x:Z](sumbool_not ? ? (Z_zerop x)). - -Definition Z_noteq_dec := [x,y:Z](sumbool_not ? ? (Z_eq_dec x y)). +Lemma Zlt_cotrans : forall n m:Z, n < m -> forall p:Z, {n < p} + {p < m}. +Proof. + intros x y H z. + case (Z_lt_ge_dec x z). + intro. + left. + assumption. + intro. + right. + apply Zle_lt_trans with (m := x). + apply Zge_le. + assumption. + assumption. +Defined. + +Lemma Zlt_cotrans_pos : forall n m:Z, 0 < n + m -> {0 < n} + {0 < m}. +Proof. + intros x y H. + case (Zlt_cotrans 0 (x + y) H x). + intro. + left. + assumption. + intro. + right. + apply Zplus_lt_reg_l with (p := x). + rewrite Zplus_0_r. + assumption. +Defined. + +Lemma Zlt_cotrans_neg : forall n m:Z, n + m < 0 -> {n < 0} + {m < 0}. +Proof. + intros x y H; case (Zlt_cotrans (x + y) 0 H x); intro Hxy; + [ right; apply Zplus_lt_reg_l with (p := x); rewrite Zplus_0_r | left ]; + assumption. +Defined. + +Lemma not_Zeq_inf : forall n m:Z, n <> m -> {n < m} + {m < n}. +Proof. + intros x y H. + case Z_lt_ge_dec with x y. + intro. + left. + assumption. + intro H0. + generalize (Zge_le _ _ H0). + intro. + case (Z_le_lt_eq_dec _ _ H1). + intro. + right. + assumption. + intro. + apply False_rec. + apply H. + symmetry in |- *. + assumption. +Defined. + +Lemma Z_dec : forall n m:Z, {n < m} + {n > m} + {n = m}. +Proof. + intros x y. + case (Z_lt_ge_dec x y). + intro H. + left. + left. + assumption. + intro H. + generalize (Zge_le _ _ H). + intro H0. + case (Z_le_lt_eq_dec y x H0). + intro H1. + left. + right. + apply Zlt_gt. + assumption. + intro. + right. + symmetry in |- *. + assumption. +Defined. + + +Lemma Z_dec' : forall n m:Z, {n < m} + {m < n} + {n = m}. +Proof. + intros x y. + case (Z_eq_dec x y); intro H; + [ right; assumption | left; apply (not_Zeq_inf _ _ H) ]. +Defined. + + + +Definition Z_zerop : forall x:Z, {x = 0} + {x <> 0}. +Proof. +exact (fun x:Z => Z_eq_dec x 0). +Defined. + +Definition Z_notzerop (x:Z) := sumbool_not _ _ (Z_zerop x). + +Definition Z_noteq_dec (x y:Z) := sumbool_not _ _ (Z_eq_dec x y). \ No newline at end of file diff --git a/theories/ZArith/Zabs.v b/theories/ZArith/Zabs.v index 27c72c4d15..eff457fc5f 100644 --- a/theories/ZArith/Zabs.v +++ b/theories/ZArith/Zabs.v @@ -9,130 +9,120 @@ (** Binary Integers (Pierre Crégut (CNET, Lannion, France) *) -Require Arith. -Require BinPos. -Require BinInt. -Require Zorder. -Require Zsyntax. -Require ZArith_dec. - -V7only [Import nat_scope.]. +Require Import Arith. +Require Import BinPos. +Require Import BinInt. +Require Import Zorder. +Require Import ZArith_dec. + Open Local Scope Z_scope. (**********************************************************************) (** Properties of absolute value *) -Lemma Zabs_eq : (x:Z) (Zle ZERO x) -> (Zabs x)=x. -Intro x; NewDestruct x; Auto with arith. -Compute; Intros; Absurd SUPERIEUR=SUPERIEUR; Trivial with arith. +Lemma Zabs_eq : forall n:Z, 0 <= n -> Zabs n = n. +intro x; destruct x; auto with arith. +compute in |- *; intros; absurd (Gt = Gt); trivial with arith. Qed. -Lemma Zabs_non_eq : (x:Z) (Zle x ZERO) -> (Zabs x)=(Zopp x). +Lemma Zabs_non_eq : forall n:Z, n <= 0 -> Zabs n = - n. Proof. -Intro x; NewDestruct x; Auto with arith. -Compute; Intros; Absurd SUPERIEUR=SUPERIEUR; Trivial with arith. +intro x; destruct x; auto with arith. +compute in |- *; intros; absurd (Gt = Gt); trivial with arith. Qed. -V7only [ (* From Zdivides *) ]. -Theorem Zabs_Zopp: (z : Z) (Zabs (Zopp z)) = (Zabs z). +Theorem Zabs_Zopp : forall n:Z, Zabs (- n) = Zabs n. Proof. -Intros z; Case z; Simpl; Auto. +intros z; case z; simpl in |- *; auto. Qed. (** Proving a property of the absolute value by cases *) -Lemma Zabs_ind : - (P:Z->Prop)(x:Z)(`x >= 0` -> (P x)) -> (`x <= 0` -> (P `-x`)) -> - (P `|x|`). +Lemma Zabs_ind : + forall (P:Z -> Prop) (n:Z), + (n >= 0 -> P n) -> (n <= 0 -> P (- n)) -> P (Zabs n). Proof. -Intros P x H H0; Elim (Z_lt_ge_dec x `0`); Intro. -Assert `x<=0`. Apply Zlt_le_weak; Assumption. -Rewrite Zabs_non_eq. Apply H0. Assumption. Assumption. -Rewrite Zabs_eq. Apply H; Assumption. Apply Zge_le. Assumption. -Save. - -V7only [ (* From Zdivides *) ]. -Theorem Zabs_intro: (P : ?) (z : Z) (P (Zopp z)) -> (P z) -> (P (Zabs z)). -Intros P z; Case z; Simpl; Auto. +intros P x H H0; elim (Z_lt_ge_dec x 0); intro. +assert (x <= 0). apply Zlt_le_weak; assumption. +rewrite Zabs_non_eq. apply H0. assumption. assumption. +rewrite Zabs_eq. apply H; assumption. apply Zge_le. assumption. +Qed. + +Theorem Zabs_intro : forall P (n:Z), P (- n) -> P n -> P (Zabs n). +intros P z; case z; simpl in |- *; auto. Qed. -Definition Zabs_dec : (x:Z){x=(Zabs x)}+{x=(Zopp (Zabs x))}. +Definition Zabs_dec : forall x:Z, {x = Zabs x} + {x = - Zabs x}. Proof. -Intro x; NewDestruct x;Auto with arith. +intro x; destruct x; auto with arith. Defined. -Lemma Zabs_pos : (x:Z)(Zle ZERO (Zabs x)). -Intro x; NewDestruct x;Auto with arith; Compute; Intros H;Inversion H. +Lemma Zabs_pos : forall n:Z, 0 <= Zabs n. +intro x; destruct x; auto with arith; compute in |- *; intros H; inversion H. Qed. -V7only [ (* From Zdivides *) ]. -Theorem Zabs_eq_case: - (z1, z2 : Z) (Zabs z1) = (Zabs z2) -> z1 = z2 \/ z1 = (Zopp z2). +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; Auto; Try (Intros; Discriminate); - Intros p1 p2 H1; Injection H1; (Intros H2; Rewrite H2); Auto. +intros z1 z2; case z1; case z2; simpl in |- *; auto; + try (intros; discriminate); intros p1 p2 H1; injection H1; + (intros H2; rewrite H2); auto. Qed. (** Triangular inequality *) -Hints Local Resolve Zle_NEG_POS :zarith. +Hint Local Resolve Zle_neg_pos: zarith. -V7only [ (* From Zdivides *) ]. -Theorem Zabs_triangle: - (z1, z2 : Z) (Zle (Zabs (Zplus z1 z2)) (Zplus (Zabs z1) (Zabs z2))). +Theorem Zabs_triangle : forall n m:Z, Zabs (n + m) <= Zabs n + Zabs m. Proof. -Intros z1 z2; Case z1; Case z2; Try (Simpl; Auto with zarith; Fail). -Intros p1 p2; - Apply Zabs_intro - with P := [x : ?] (Zle x (Zplus (Zabs (POS p2)) (Zabs (NEG p1)))); - Try Rewrite Zopp_Zplus; Auto with zarith. -Apply Zle_plus_plus; Simpl; Auto with zarith. -Apply Zle_plus_plus; Simpl; Auto with zarith. -Intros p1 p2; - Apply Zabs_intro - with P := [x : ?] (Zle x (Zplus (Zabs (POS p2)) (Zabs (NEG p1)))); - Try Rewrite Zopp_Zplus; Auto with zarith. -Apply Zle_plus_plus; Simpl; Auto with zarith. -Apply Zle_plus_plus; Simpl; Auto with zarith. +intros z1 z2; case z1; case z2; try (simpl in |- *; auto with zarith; fail). +intros p1 p2; + apply Zabs_intro with (P := fun x => x <= Zabs (Zpos p2) + Zabs (Zneg p1)); + try rewrite Zopp_plus_distr; auto with zarith. +apply Zplus_le_compat; simpl in |- *; auto with zarith. +apply Zplus_le_compat; simpl in |- *; auto with zarith. +intros p1 p2; + apply Zabs_intro with (P := fun x => x <= Zabs (Zpos p2) + Zabs (Zneg p1)); + try rewrite Zopp_plus_distr; auto with zarith. +apply Zplus_le_compat; simpl in |- *; auto with zarith. +apply Zplus_le_compat; simpl in |- *; auto with zarith. Qed. (** Absolute value and multiplication *) -Lemma Zsgn_Zabs: (x:Z)(Zmult x (Zsgn x))=(Zabs x). +Lemma Zsgn_Zabs : forall n:Z, n * Zsgn n = Zabs n. Proof. -Intro x; NewDestruct x; Rewrite Zmult_sym; Auto with arith. +intro x; destruct x; rewrite Zmult_comm; auto with arith. Qed. -Lemma Zabs_Zsgn: (x:Z)(Zmult (Zabs x) (Zsgn x))=x. +Lemma Zabs_Zsgn : forall n:Z, Zabs n * Zsgn n = n. Proof. -Intro x; NewDestruct x; Rewrite Zmult_sym; Auto with arith. +intro x; destruct x; rewrite Zmult_comm; auto with arith. Qed. -V7only [ (* From Zdivides *) ]. -Theorem Zabs_Zmult: - (z1, z2 : Z) (Zabs (Zmult z1 z2)) = (Zmult (Zabs z1) (Zabs z2)). +Theorem Zabs_Zmult : forall n m:Z, Zabs (n * m) = Zabs n * Zabs m. Proof. -Intros z1 z2; Case z1; Case z2; Simpl; Auto. +intros z1 z2; case z1; case z2; simpl in |- *; auto. Qed. (** absolute value in nat is compatible with order *) -Lemma absolu_lt : (x,y:Z) (Zle ZERO x)/\(Zlt x y) -> (lt (absolu x) (absolu y)). +Lemma Zabs_nat_lt : + forall n m:Z, 0 <= n /\ n < m -> (Zabs_nat n < Zabs_nat m)%nat. Proof. -Intros x y. Case x; Simpl. Case y; Simpl. +intros x y. case x; simpl in |- *. case y; simpl in |- *. -Intro. Absurd (Zlt ZERO ZERO). Compute. 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. +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. -Intros. Absurd (Zlt (POS p) ZERO). Compute. Intro H0. Discriminate H0. Intuition. -Intros. Change (gt (convert p) (convert p0)). -Apply compare_convert_SUPERIEUR. -Elim H; Auto with arith. Intro. Exact (ZC2 p0 p). +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). -Intros. Absurd (Zlt (POS p0) (NEG p)). -Compute. Intro H0. Discriminate H0. Intuition. +intros. absurd (Zpos p0 < Zneg p). +compute in |- *. intro H0. discriminate H0. intuition. -Intros. Absurd (Zle ZERO (NEG p)). Compute. Auto with arith. Intuition. -Qed. +intros. absurd (0 <= Zneg p). compute in |- *. auto with arith. intuition. +Qed. \ No newline at end of file diff --git a/theories/ZArith/Zbinary.v b/theories/ZArith/Zbinary.v index 142bfdef60..cd8872dac5 100644 --- a/theories/ZArith/Zbinary.v +++ b/theories/ZArith/Zbinary.v @@ -11,10 +11,10 @@ (** Bit vectors interpreted as integers. Contribution by Jean Duprat (ENS Lyon). *) -Require Bvector. -Require ZArith. +Require Import Bvector. +Require Import ZArith. Require Export Zpower. -Require Omega. +Require Import Omega. (* L'évaluation des vecteurs de booléens se font à la fois en binaire et @@ -41,29 +41,29 @@ Le compl supérieure ou égale à un, le bit de signe étant évalué négativement. *) -Definition bit_value [b:bool] : Z := -Cases b of - | true => `1` - | false => `0` -end. +Definition bit_value (b:bool) : Z := + match b with + | true => 1%Z + | false => 0%Z + end. -Lemma binary_value : (n:nat) (Bvector n) -> Z. +Lemma binary_value : forall n:nat, Bvector n -> Z. Proof. - Induction n; Intros. - Exact `0`. + simple induction n; intros. + exact 0%Z. - Inversion H0. - Exact (Zplus (bit_value a) (Zmult `2` (H H2))). + inversion H0. + exact (bit_value a + 2 * H H2)%Z. Defined. -Lemma two_compl_value : (n:nat) (Bvector (S n)) -> Z. +Lemma two_compl_value : forall n:nat, Bvector (S n) -> Z. Proof. - Induction n; Intros. - Inversion H. - Exact (Zopp (bit_value a)). + simple induction n; intros. + inversion H. + exact (- bit_value a)%Z. - Inversion H0. - Exact (Zplus (bit_value a) (Zmult `2` (H H2))). + inversion H0. + exact (bit_value a + 2 * H H2)%Z. Defined. (* @@ -91,52 +91,50 @@ La valeur en compl avec Zmod2, le paramètre est la taille moins un. *) -Definition Zmod2 := [z:Z] Cases z of - | ZERO => `0` - | ((POS p)) => Cases p of - | (xI q) => (POS q) - | (xO q) => (POS q) - | xH => `0` - end - | ((NEG p)) => Cases p of - | (xI q) => `(NEG q) - 1` - | (xO q) => (NEG q) - | xH => `-1` - end - end. - -V7only [ -Notation double_moins_un_add_un := - [p](sym_eq ? ? ? (double_moins_un_add_un_xI p)). -]. - -Lemma Zmod2_twice : (z:Z) - `z = (2*(Zmod2 z) + (bit_value (Zodd_bool z)))`. +Definition Zmod2 (z:Z) := + match z with + | Z0 => 0%Z + | Zpos p => match p with + | xI q => Zpos q + | xO q => Zpos q + | xH => 0%Z + end + | Zneg p => + match p with + | xI q => (Zneg q - 1)%Z + | xO q => Zneg q + | xH => (-1)%Z + end + end. + + +Lemma Zmod2_twice : + forall z:Z, z = (2 * Zmod2 z + bit_value (Zeven.Zodd_bool z))%Z. Proof. - NewDestruct z; Simpl. - Trivial. + destruct z; simpl in |- *. + trivial. - NewDestruct p; Simpl; Trivial. + destruct p; simpl in |- *; trivial. - NewDestruct p; Simpl. - NewDestruct p as [p|p|]; Simpl. - Rewrite <- (double_moins_un_add_un_xI p); Trivial. + destruct p; simpl in |- *. + destruct p as [p| p| ]; simpl in |- *. + rewrite <- (Pdouble_minus_one_o_succ_eq_xI p); trivial. - Trivial. + trivial. - Trivial. + trivial. - Trivial. + trivial. - Trivial. -Save. + trivial. +Qed. -Lemma Z_to_binary : (n:nat) Z -> (Bvector n). +Lemma Z_to_binary : forall n:nat, Z -> Bvector n. Proof. - Induction n; Intros. - Exact Bnil. + simple induction n; intros. + exact Bnil. - Exact (Bcons (Zodd_bool H0) n0 (H (Zdiv2 H0))). + exact (Bcons (Zeven.Zodd_bool H0) n0 (H (Zeven.Zdiv2 H0))). Defined. (* @@ -148,12 +146,12 @@ Eval Compute in (Z_to_binary (5) `5`). : (Bvector (5)) *) -Lemma Z_to_two_compl : (n:nat) Z -> (Bvector (S n)). +Lemma Z_to_two_compl : forall n:nat, Z -> Bvector (S n). Proof. - Induction n; Intros. - Exact (Bcons (Zodd_bool H) (0) Bnil). + simple induction n; intros. + exact (Bcons (Zeven.Zodd_bool H) 0 Bnil). - Exact (Bcons (Zodd_bool H0) (S n0) (H (Zmod2 H0))). + exact (Bcons (Zeven.Zodd_bool H0) (S n0) (H (Zmod2 H0))). Defined. (* @@ -186,93 +184,97 @@ Utilise largement ZArith. Meriterait d'etre reecrite. *) -Lemma binary_value_Sn : (n:nat) (b:bool) (bv : (Bvector n)) - (binary_value (S n) (Vcons bool b n bv))=`(bit_value b) + 2*(binary_value n bv)`. +Lemma binary_value_Sn : + forall (n:nat) (b:bool) (bv:Bvector n), + binary_value (S n) (Vcons bool b n bv) = + (bit_value b + 2 * binary_value n bv)%Z. Proof. - Intros; Auto. -Save. + intros; auto. +Qed. -Lemma Z_to_binary_Sn : (n:nat) (b:bool) (z:Z) - `z>=0`-> - (Z_to_binary (S n) `(bit_value b) + 2*z`)=(Bcons b n (Z_to_binary n z)). +Lemma Z_to_binary_Sn : + forall (n:nat) (b:bool) (z:Z), + (z >= 0)%Z -> + Z_to_binary (S n) (bit_value b + 2 * z) = Bcons b n (Z_to_binary n z). Proof. - NewDestruct b; NewDestruct z; Simpl; Auto. - Intro H; Elim H; Trivial. -Save. + destruct b; destruct z; simpl in |- *; auto. + intro H; elim H; trivial. +Qed. -Lemma binary_value_pos : (n:nat) (bv:(Bvector n)) - `(binary_value n bv) >= 0`. +Lemma binary_value_pos : + forall (n:nat) (bv:Bvector n), (binary_value n bv >= 0)%Z. Proof. - NewInduction bv as [|a n v IHbv]; Simpl. - Omega. + induction bv as [| a n v IHbv]; simpl in |- *. + omega. - NewDestruct a; NewDestruct (binary_value n v); Simpl; Auto. - Auto with zarith. -Save. + destruct a; destruct (binary_value n v); simpl in |- *; auto. + auto with zarith. +Qed. -V7only [Notation add_un_double_moins_un_xO := is_double_moins_un.]. -Lemma two_compl_value_Sn : (n:nat) (bv : (Bvector (S n))) (b:bool) - (two_compl_value (S n) (Bcons b (S n) bv)) = - `(bit_value b) + 2*(two_compl_value n bv)`. +Lemma two_compl_value_Sn : + forall (n:nat) (bv:Bvector (S n)) (b:bool), + two_compl_value (S n) (Bcons b (S n) bv) = + (bit_value b + 2 * two_compl_value n bv)%Z. Proof. - Intros; Auto. -Save. + intros; auto. +Qed. -Lemma Z_to_two_compl_Sn : (n:nat) (b:bool) (z:Z) - (Z_to_two_compl (S n) `(bit_value b) + 2*z`) = - (Bcons b (S n) (Z_to_two_compl n z)). +Lemma Z_to_two_compl_Sn : + forall (n:nat) (b:bool) (z:Z), + Z_to_two_compl (S n) (bit_value b + 2 * z) = + Bcons b (S n) (Z_to_two_compl n z). Proof. - NewDestruct b; NewDestruct z as [|p|p]; Auto. - NewDestruct p as [p|p|]; Auto. - NewDestruct p as [p|p|]; Simpl; Auto. - Intros; Rewrite (add_un_double_moins_un_xO p); Trivial. -Save. - -Lemma Z_to_binary_Sn_z : (n:nat) (z:Z) - (Z_to_binary (S n) z)=(Bcons (Zodd_bool z) n (Z_to_binary n (Zdiv2 z))). + destruct b; destruct z as [| p| p]; auto. + destruct p as [p| p| ]; auto. + destruct p as [p| p| ]; simpl in |- *; auto. + intros; rewrite (Psucc_o_double_minus_one_eq_xO p); trivial. +Qed. + +Lemma Z_to_binary_Sn_z : + forall (n:nat) (z:Z), + Z_to_binary (S n) z = + Bcons (Zeven.Zodd_bool z) n (Z_to_binary n (Zeven.Zdiv2 z)). Proof. - Intros; Auto. -Save. + intros; auto. +Qed. -Lemma Z_div2_value : (z:Z) - ` z>=0 `-> - `(bit_value (Zodd_bool z))+2*(Zdiv2 z) = z`. +Lemma Z_div2_value : + forall z:Z, + (z >= 0)%Z -> (bit_value (Zeven.Zodd_bool z) + 2 * Zeven.Zdiv2 z)%Z = z. Proof. - NewDestruct z as [|p|p]; Auto. - NewDestruct p; Auto. - Intro H; Elim H; Trivial. -Save. + destruct z as [| p| p]; auto. + destruct p; auto. + intro H; elim H; trivial. +Qed. -Lemma Zdiv2_pos : (z:Z) - ` z >= 0 ` -> - `(Zdiv2 z) >= 0 `. +Lemma Pdiv2 : forall z:Z, (z >= 0)%Z -> (Zeven.Zdiv2 z >= 0)%Z. Proof. - NewDestruct z as [|p|p]. - Auto. + destruct z as [| p| p]. + auto. - NewDestruct p; Auto. - Simpl; Intros; Omega. + destruct p; auto. + simpl in |- *; intros; omega. - Intro H; Elim H; Trivial. + intro H; elim H; trivial. -Save. +Qed. -Lemma Zdiv2_two_power_nat : (z:Z) (n:nat) - ` z >= 0 ` -> - ` z < (two_power_nat (S n)) ` -> - `(Zdiv2 z) < (two_power_nat n) `. +Lemma Zdiv2_two_power_nat : + forall (z:Z) (n:nat), + (z >= 0)%Z -> + (z < two_power_nat (S n))%Z -> (Zeven.Zdiv2 z < two_power_nat n)%Z. Proof. - Intros. - Cut (Zlt (Zmult `2` (Zdiv2 z)) (Zmult `2` (two_power_nat n))); Intros. - Omega. + intros. + cut (2 * Zeven.Zdiv2 z < 2 * two_power_nat n)%Z; intros. + omega. - Rewrite <- two_power_nat_S. - NewDestruct (Zeven_odd_dec z); Intros. - Rewrite <- Zeven_div2; Auto. + rewrite <- two_power_nat_S. + destruct (Zeven.Zeven_odd_dec z); intros. + rewrite <- Zeven.Zeven_div2; auto. - Generalize (Zodd_div2 z H z0); Omega. -Save. + generalize (Zeven.Zodd_div2 z H z0); omega. +Qed. (* @@ -299,54 +301,54 @@ Proof. Save. *) -Lemma Z_to_two_compl_Sn_z : (n:nat) (z:Z) - (Z_to_two_compl (S n) z)=(Bcons (Zodd_bool z) (S n) (Z_to_two_compl n (Zmod2 z))). +Lemma Z_to_two_compl_Sn_z : + forall (n:nat) (z:Z), + Z_to_two_compl (S n) z = + Bcons (Zeven.Zodd_bool z) (S n) (Z_to_two_compl n (Zmod2 z)). Proof. - Intros; Auto. -Save. + intros; auto. +Qed. -Lemma Zeven_bit_value : (z:Z) - (Zeven z) -> - `(bit_value (Zodd_bool z))=0`. +Lemma Zeven_bit_value : + forall z:Z, Zeven.Zeven z -> bit_value (Zeven.Zodd_bool z) = 0%Z. Proof. - NewDestruct z; Unfold bit_value; Auto. - NewDestruct p; Tauto Orelse (Intro H; Elim H). - NewDestruct p; Tauto Orelse (Intro H; Elim H). -Save. + destruct z; unfold bit_value in |- *; auto. + destruct p; tauto || (intro H; elim H). + destruct p; tauto || (intro H; elim H). +Qed. -Lemma Zodd_bit_value : (z:Z) - (Zodd z) -> - `(bit_value (Zodd_bool z))=1`. +Lemma Zodd_bit_value : + forall z:Z, Zeven.Zodd z -> bit_value (Zeven.Zodd_bool z) = 1%Z. Proof. - NewDestruct z; Unfold bit_value; Auto. - Intros; Elim H. - NewDestruct p; Tauto Orelse (Intros; Elim H). - NewDestruct p; Tauto Orelse (Intros; Elim H). -Save. - -Lemma Zge_minus_two_power_nat_S : (n:nat) (z:Z) - `z >= (-(two_power_nat (S n)))`-> - `(Zmod2 z) >= (-(two_power_nat n))`. + destruct z; unfold bit_value in |- *; auto. + intros; elim H. + destruct p; tauto || (intros; elim H). + destruct p; tauto || (intros; elim H). +Qed. + +Lemma Zge_minus_two_power_nat_S : + forall (n:nat) (z:Z), + (z >= - two_power_nat (S n))%Z -> (Zmod2 z >= - two_power_nat n)%Z. Proof. - Intros n z; Rewrite (two_power_nat_S n). - Generalize (Zmod2_twice z). - NewDestruct (Zeven_odd_dec z) as [H|H]. - Rewrite (Zeven_bit_value z H); Intros; Omega. + intros n z; rewrite (two_power_nat_S n). + generalize (Zmod2_twice z). + destruct (Zeven.Zeven_odd_dec z) as [H| H]. + rewrite (Zeven_bit_value z H); intros; omega. - Rewrite (Zodd_bit_value z H); Intros; Omega. -Save. + rewrite (Zodd_bit_value z H); intros; omega. +Qed. -Lemma Zlt_two_power_nat_S : (n:nat) (z:Z) - `z < (two_power_nat (S n))`-> - `(Zmod2 z) < (two_power_nat n)`. +Lemma Zlt_two_power_nat_S : + forall (n:nat) (z:Z), + (z < two_power_nat (S n))%Z -> (Zmod2 z < two_power_nat n)%Z. Proof. - Intros n z; Rewrite (two_power_nat_S n). - Generalize (Zmod2_twice z). - NewDestruct (Zeven_odd_dec z) as [H|H]. - Rewrite (Zeven_bit_value z H); Intros; Omega. + intros n z; rewrite (two_power_nat_S n). + generalize (Zmod2_twice z). + destruct (Zeven.Zeven_odd_dec z) as [H| H]. + rewrite (Zeven_bit_value z H); intros; omega. - Rewrite (Zodd_bit_value z H); Intros; Omega. -Save. + rewrite (Zodd_bit_value z H); intros; omega. +Qed. End Z_BRIC_A_BRAC. @@ -358,68 +360,67 @@ r Elles utilisent les lemmes du bric-a-brac. *) -Lemma binary_to_Z_to_binary : (n:nat) (bv : (Bvector n)) - (Z_to_binary n (binary_value n bv))=bv. +Lemma binary_to_Z_to_binary : + forall (n:nat) (bv:Bvector n), Z_to_binary n (binary_value n bv) = bv. Proof. - NewInduction bv as [|a n bv IHbv]. - Auto. + induction bv as [| a n bv IHbv]. + auto. - Rewrite binary_value_Sn. - Rewrite Z_to_binary_Sn. - Rewrite IHbv; Trivial. + rewrite binary_value_Sn. + rewrite Z_to_binary_Sn. + rewrite IHbv; trivial. - Apply binary_value_pos. -Save. + apply binary_value_pos. +Qed. -Lemma two_compl_to_Z_to_two_compl : (n:nat) (bv : (Bvector n)) (b:bool) - (Z_to_two_compl n (two_compl_value n (Bcons b n bv)))= - (Bcons b n bv). +Lemma two_compl_to_Z_to_two_compl : + forall (n:nat) (bv:Bvector n) (b:bool), + Z_to_two_compl n (two_compl_value n (Bcons b n bv)) = Bcons b n bv. Proof. - NewInduction bv as [|a n bv IHbv]; Intro b. - NewDestruct b; Auto. - - Rewrite two_compl_value_Sn. - Rewrite Z_to_two_compl_Sn. - Rewrite IHbv; Trivial. -Save. - -Lemma Z_to_binary_to_Z : (n:nat) (z : Z) - `z >= 0 `-> - `z < (two_power_nat n) `-> - (binary_value n (Z_to_binary n z))=z. + induction bv as [| a n bv IHbv]; intro b. + destruct b; auto. + + rewrite two_compl_value_Sn. + rewrite Z_to_two_compl_Sn. + rewrite IHbv; trivial. +Qed. + +Lemma Z_to_binary_to_Z : + forall (n:nat) (z:Z), + (z >= 0)%Z -> + (z < two_power_nat n)%Z -> binary_value n (Z_to_binary n z) = z. Proof. - NewInduction n as [|n IHn]. - Unfold two_power_nat shift_nat; Simpl; Intros; Omega. + induction n as [| n IHn]. + unfold two_power_nat, shift_nat in |- *; simpl in |- *; intros; omega. - Intros; Rewrite Z_to_binary_Sn_z. - Rewrite binary_value_Sn. - Rewrite IHn. - Apply Z_div2_value; Auto. + intros; rewrite Z_to_binary_Sn_z. + rewrite binary_value_Sn. + rewrite IHn. + apply Z_div2_value; auto. - Apply Zdiv2_pos; Trivial. + apply Pdiv2; trivial. - Apply Zdiv2_two_power_nat; Trivial. -Save. + apply Zdiv2_two_power_nat; trivial. +Qed. -Lemma Z_to_two_compl_to_Z : (n:nat) (z : Z) - `z >= -(two_power_nat n) `-> - `z < (two_power_nat n) `-> - (two_compl_value n (Z_to_two_compl n z))=z. +Lemma Z_to_two_compl_to_Z : + forall (n:nat) (z:Z), + (z >= - two_power_nat n)%Z -> + (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = z. Proof. - NewInduction n as [|n IHn]. - Unfold two_power_nat shift_nat; Simpl; Intros. - Assert `z=-1`\/`z=0`. Omega. -Intuition; Subst z; Trivial. + induction n as [| n IHn]. + unfold two_power_nat, shift_nat in |- *; simpl in |- *; intros. + assert (z = (-1)%Z \/ z = 0%Z). omega. +intuition; subst z; trivial. - Intros; Rewrite Z_to_two_compl_Sn_z. - Rewrite two_compl_value_Sn. - Rewrite IHn. - Generalize (Zmod2_twice z); Omega. + intros; rewrite Z_to_two_compl_Sn_z. + rewrite two_compl_value_Sn. + rewrite IHn. + generalize (Zmod2_twice z); omega. - Apply Zge_minus_two_power_nat_S; Auto. + apply Zge_minus_two_power_nat_S; auto. - Apply Zlt_two_power_nat_S; Auto. -Save. + apply Zlt_two_power_nat_S; auto. +Qed. End COHERENT_VALUE. - diff --git a/theories/ZArith/Zbool.v b/theories/ZArith/Zbool.v index fcbdd1141f..a95218b637 100644 --- a/theories/ZArith/Zbool.v +++ b/theories/ZArith/Zbool.v @@ -8,151 +8,179 @@ (* $Id$ *) -Require BinInt. -Require Zeven. -Require Zorder. -Require Zcompare. -Require ZArith_dec. -Require Zsyntax. -Require Sumbool. +Require Import BinInt. +Require Import Zeven. +Require Import Zorder. +Require Import Zcompare. +Require Import ZArith_dec. +Require Import Sumbool. (** The decidability of equality and order relations over type [Z] give some boolean functions with the adequate specification. *) -Definition Z_lt_ge_bool := [x,y:Z](bool_of_sumbool (Z_lt_ge_dec x y)). -Definition Z_ge_lt_bool := [x,y:Z](bool_of_sumbool (Z_ge_lt_dec x y)). +Definition Z_lt_ge_bool (x y:Z) := bool_of_sumbool (Z_lt_ge_dec x y). +Definition Z_ge_lt_bool (x y:Z) := bool_of_sumbool (Z_ge_lt_dec x y). -Definition Z_le_gt_bool := [x,y:Z](bool_of_sumbool (Z_le_gt_dec x y)). -Definition Z_gt_le_bool := [x,y:Z](bool_of_sumbool (Z_gt_le_dec x y)). +Definition Z_le_gt_bool (x y:Z) := bool_of_sumbool (Z_le_gt_dec x y). +Definition Z_gt_le_bool (x y:Z) := bool_of_sumbool (Z_gt_le_dec x y). -Definition Z_eq_bool := [x,y:Z](bool_of_sumbool (Z_eq_dec x y)). -Definition Z_noteq_bool := [x,y:Z](bool_of_sumbool (Z_noteq_dec x y)). +Definition Z_eq_bool (x y:Z) := bool_of_sumbool (Z_eq_dec x y). +Definition Z_noteq_bool (x y:Z) := bool_of_sumbool (Z_noteq_dec x y). -Definition Zeven_odd_bool := [x:Z](bool_of_sumbool (Zeven_odd_dec x)). +Definition Zeven_odd_bool (x:Z) := bool_of_sumbool (Zeven_odd_dec x). (**********************************************************************) (** Boolean comparisons of binary integers *) -Definition Zle_bool := - [x,y:Z]Cases `x ?= y` of SUPERIEUR => false | _ => true end. -Definition Zge_bool := - [x,y:Z]Cases `x ?= y` of INFERIEUR => false | _ => true end. -Definition Zlt_bool := - [x,y:Z]Cases `x ?= y` of INFERIEUR => true | _ => false end. -Definition Zgt_bool := - [x,y:Z]Cases ` x ?= y` of SUPERIEUR => true | _ => false end. -Definition Zeq_bool := - [x,y:Z]Cases `x ?= y` of EGAL => true | _ => false end. -Definition Zneq_bool := - [x,y:Z]Cases `x ?= y` of EGAL => false | _ => true end. - -Lemma Zle_cases : (x,y:Z)if (Zle_bool x y) then `x<=y` else `x>y`. +Definition Zle_bool (x y:Z) := + match (x ?= y)%Z with + | Gt => false + | _ => true + end. +Definition Zge_bool (x y:Z) := + match (x ?= y)%Z with + | Lt => false + | _ => true + end. +Definition Zlt_bool (x y:Z) := + match (x ?= y)%Z with + | Lt => true + | _ => false + end. +Definition Zgt_bool (x y:Z) := + match (x ?= y)%Z with + | Gt => true + | _ => false + end. +Definition Zeq_bool (x y:Z) := + match (x ?= y)%Z with + | Eq => true + | _ => false + end. +Definition Zneq_bool (x y:Z) := + match (x ?= y)%Z with + | Eq => false + | _ => true + end. + +Lemma Zle_cases : + forall n m:Z, if Zle_bool n m then (n <= m)%Z else (n > m)%Z. Proof. -Intros x y; Unfold Zle_bool Zle Zgt. -Case (Zcompare x y); Auto; Discriminate. +intros x y; unfold Zle_bool, Zle, Zgt in |- *. +case (x ?= y)%Z; auto; discriminate. Qed. -Lemma Zlt_cases : (x,y:Z)if (Zlt_bool x y) then `x=y`. +Lemma Zlt_cases : + forall n m:Z, if Zlt_bool n m then (n < m)%Z else (n >= m)%Z. Proof. -Intros x y; Unfold Zlt_bool Zlt Zge. -Case (Zcompare x y); Auto; Discriminate. +intros x y; unfold Zlt_bool, Zlt, Zge in |- *. +case (x ?= y)%Z; auto; discriminate. Qed. -Lemma Zge_cases : (x,y:Z)if (Zge_bool x y) then `x>=y` else `x= m)%Z else (n < m)%Z. Proof. -Intros x y; Unfold Zge_bool Zge Zlt. -Case (Zcompare x y); Auto; Discriminate. +intros x y; unfold Zge_bool, Zge, Zlt in |- *. +case (x ?= y)%Z; auto; discriminate. Qed. -Lemma Zgt_cases : (x,y:Z)if (Zgt_bool x y) then `x>y` else `x<=y`. +Lemma Zgt_cases : + forall n m:Z, if Zgt_bool n m then (n > m)%Z else (n <= m)%Z. Proof. -Intros x y; Unfold Zgt_bool Zgt Zle. -Case (Zcompare x y); Auto; Discriminate. +intros x y; unfold Zgt_bool, Zgt, Zle in |- *. +case (x ?= y)%Z; auto; discriminate. Qed. (** Lemmas on [Zle_bool] used in contrib/graphs *) -Lemma Zle_bool_imp_le : (x,y:Z) (Zle_bool x y)=true -> (Zle x y). +Lemma Zle_bool_imp_le : forall n m:Z, Zle_bool n m = true -> (n <= m)%Z. Proof. - Unfold Zle_bool Zle. Intros x y. Unfold not. - Case (Zcompare x y); Intros; Discriminate. + unfold Zle_bool, Zle in |- *. intros x y. unfold not in |- *. + case (x ?= y)%Z; intros; discriminate. Qed. -Lemma Zle_imp_le_bool : (x,y:Z) (Zle x y) -> (Zle_bool x y)=true. +Lemma Zle_imp_le_bool : forall n m:Z, (n <= m)%Z -> Zle_bool n m = true. Proof. - Unfold Zle Zle_bool. Intros x y. Case (Zcompare x y); Trivial. Intro. Elim (H (refl_equal ? ?)). + unfold Zle, Zle_bool in |- *. intros x y. case (x ?= y)%Z; trivial. intro. elim (H (refl_equal _)). Qed. -Lemma Zle_bool_refl : (x:Z) (Zle_bool x x)=true. +Lemma Zle_bool_refl : forall n:Z, Zle_bool n n = true. Proof. - Intro. Apply Zle_imp_le_bool. Apply Zle_refl. Reflexivity. + intro. apply Zle_imp_le_bool. apply Zeq_le. reflexivity. Qed. -Lemma Zle_bool_antisym : (x,y:Z) (Zle_bool x y)=true -> (Zle_bool y x)=true -> x=y. +Lemma Zle_bool_antisym : + forall n m:Z, Zle_bool n m = true -> Zle_bool m n = true -> n = m. Proof. - Intros. Apply Zle_antisym. Apply Zle_bool_imp_le. Assumption. - Apply Zle_bool_imp_le. Assumption. + intros. apply Zle_antisym. apply Zle_bool_imp_le. assumption. + apply Zle_bool_imp_le. assumption. Qed. -Lemma Zle_bool_trans : (x,y,z:Z) (Zle_bool x y)=true -> (Zle_bool y z)=true -> (Zle_bool x z)=true. +Lemma Zle_bool_trans : + forall n m p:Z, + Zle_bool n m = true -> Zle_bool m p = true -> Zle_bool n p = true. Proof. - Intros x y z; Intros. Apply Zle_imp_le_bool. Apply Zle_trans with m:=y. Apply Zle_bool_imp_le. Assumption. - Apply Zle_bool_imp_le. Assumption. + intros x y z; intros. apply Zle_imp_le_bool. apply Zle_trans with (m := y). apply Zle_bool_imp_le. assumption. + apply Zle_bool_imp_le. assumption. Qed. -Definition Zle_bool_total : (x,y:Z) {(Zle_bool x y)=true}+{(Zle_bool y x)=true}. +Definition Zle_bool_total : + forall x y:Z, {Zle_bool x y = true} + {Zle_bool y x = true}. Proof. - Intros x y; Intros. Unfold Zle_bool. Cut (Zcompare x y)=SUPERIEUR<->(Zcompare y x)=INFERIEUR. - Case (Zcompare x y). Left . Reflexivity. - Left . Reflexivity. - Right . Rewrite (proj1 ? ? H (refl_equal ? ?)). Reflexivity. - Apply Zcompare_ANTISYM. + intros x y; intros. unfold Zle_bool in |- *. cut ((x ?= y)%Z = Gt <-> (y ?= x)%Z = Lt). + case (x ?= y)%Z. left. reflexivity. + left. reflexivity. + right. rewrite (proj1 H (refl_equal _)). reflexivity. + apply Zcompare_Gt_Lt_antisym. Defined. -Lemma Zle_bool_plus_mono : (x,y,z,t:Z) (Zle_bool x y)=true -> (Zle_bool z t)=true -> - (Zle_bool (Zplus x z) (Zplus y t))=true. +Lemma Zle_bool_plus_mono : + forall n m p q:Z, + Zle_bool n m = true -> + Zle_bool p q = true -> Zle_bool (n + p) (m + q) = true. Proof. - Intros. Apply Zle_imp_le_bool. Apply Zle_plus_plus. Apply Zle_bool_imp_le. Assumption. - Apply Zle_bool_imp_le. Assumption. + intros. apply Zle_imp_le_bool. apply Zplus_le_compat. apply Zle_bool_imp_le. assumption. + apply Zle_bool_imp_le. assumption. Qed. -Lemma Zone_pos : (Zle_bool `1` `0`)=false. +Lemma Zone_pos : Zle_bool 1 0 = false. Proof. - Reflexivity. + reflexivity. Qed. -Lemma Zone_min_pos : (x:Z) (Zle_bool x `0`)=false -> (Zle_bool `1` x)=true. +Lemma Zone_min_pos : forall n:Z, Zle_bool n 0 = false -> Zle_bool 1 n = true. Proof. - Intros x; Intros. Apply Zle_imp_le_bool. Change (Zle (Zs ZERO) x). Apply Zgt_le_S. Generalize H. - Unfold Zle_bool Zgt. Case (Zcompare x ZERO). Intro H0. Discriminate H0. - Intro H0. Discriminate H0. - Reflexivity. + intros x; intros. apply Zle_imp_le_bool. change (Zsucc 0 <= x)%Z in |- *. apply Zgt_le_succ. generalize H. + unfold Zle_bool, Zgt in |- *. case (x ?= 0)%Z. intro H0. discriminate H0. + intro H0. discriminate H0. + reflexivity. Qed. - Lemma Zle_is_le_bool : (x,y:Z) (Zle x y) <-> (Zle_bool x y)=true. + Lemma Zle_is_le_bool : forall n m:Z, (n <= m)%Z <-> Zle_bool n m = true. Proof. - Intros. Split. Intro. Apply Zle_imp_le_bool. Assumption. - Intro. Apply Zle_bool_imp_le. Assumption. + intros. split. intro. apply Zle_imp_le_bool. assumption. + intro. apply Zle_bool_imp_le. assumption. Qed. - Lemma Zge_is_le_bool : (x,y:Z) (Zge x y) <-> (Zle_bool y x)=true. + Lemma Zge_is_le_bool : forall n m:Z, (n >= m)%Z <-> Zle_bool m n = true. Proof. - Intros. Split. Intro. Apply Zle_imp_le_bool. Apply Zge_le. Assumption. - Intro. Apply Zle_ge. Apply Zle_bool_imp_le. Assumption. + intros. split. intro. apply Zle_imp_le_bool. apply Zge_le. assumption. + intro. apply Zle_ge. apply Zle_bool_imp_le. assumption. Qed. - Lemma Zlt_is_le_bool : (x,y:Z) (Zlt x y) <-> (Zle_bool x `y-1`)=true. + Lemma Zlt_is_le_bool : + forall n m:Z, (n < m)%Z <-> Zle_bool n (m - 1) = true. Proof. - Intros x y. Split. Intro. Apply Zle_imp_le_bool. Apply Zlt_n_Sm_le. Rewrite (Zs_pred y) in H. - Assumption. - Intro. Rewrite (Zs_pred y). Apply Zle_lt_n_Sm. Apply Zle_bool_imp_le. Assumption. + intros x y. split. intro. apply Zle_imp_le_bool. apply Zlt_succ_le. rewrite (Zsucc_pred y) in H. + assumption. + intro. rewrite (Zsucc_pred y). apply Zle_lt_succ. apply Zle_bool_imp_le. assumption. Qed. - Lemma Zgt_is_le_bool : (x,y:Z) (Zgt x y) <-> (Zle_bool y `x-1`)=true. + Lemma Zgt_is_le_bool : + forall n m:Z, (n > m)%Z <-> Zle_bool m (n - 1) = true. Proof. - Intros x y. Apply iff_trans with `y < x`. Split. Exact (Zgt_lt x y). - Exact (Zlt_gt y x). - Exact (Zlt_is_le_bool y x). + intros x y. apply iff_trans with (y < x)%Z. split. exact (Zgt_lt x y). + exact (Zlt_gt y x). + exact (Zlt_is_le_bool y x). Qed. - diff --git a/theories/ZArith/Zcompare.v b/theories/ZArith/Zcompare.v index 2383e90cb5..f7015089c5 100644 --- a/theories/ZArith/Zcompare.v +++ b/theories/ZArith/Zcompare.v @@ -10,11 +10,10 @@ Require Export BinPos. Require Export BinInt. -Require Zsyntax. -Require Lt. -Require Gt. -Require Plus. -Require Mult. +Require Import Lt. +Require Import Gt. +Require Import Plus. +Require Import Mult. Open Local Scope Z_scope. @@ -25,456 +24,478 @@ Open Local Scope Z_scope. (**********************************************************************) (** Comparison on integers *) -Lemma Zcompare_x_x : (x:Z) (Zcompare x x) = EGAL. +Lemma Zcompare_refl : forall n:Z, (n ?= n) = Eq. Proof. -Intro x; NewDestruct x as [|p|p]; Simpl; [ Reflexivity | Apply convert_compare_EGAL - | Rewrite convert_compare_EGAL; Reflexivity ]. +intro x; destruct x as [| p| p]; simpl in |- *; + [ reflexivity | apply Pcompare_refl | rewrite Pcompare_refl; reflexivity ]. Qed. -Lemma Zcompare_EGAL_eq : (x,y:Z) (Zcompare x y) = EGAL -> x = y. +Lemma Zcompare_Eq_eq : forall n m:Z, (n ?= m) = Eq -> n = m. Proof. -Intros x y; NewDestruct x as [|x'|x'];NewDestruct y as [|y'|y'];Simpl;Intro H; Reflexivity Orelse Try Discriminate H; [ - Rewrite (compare_convert_EGAL x' y' H); Reflexivity - | Rewrite (compare_convert_EGAL x' y'); [ - Reflexivity - | NewDestruct (compare x' y' EGAL); - Reflexivity Orelse Discriminate]]. +intros x y; destruct x as [| x'| x']; destruct y as [| y'| y']; simpl in |- *; + intro H; reflexivity || (try discriminate H); + [ rewrite (Pcompare_Eq_eq x' y' H); reflexivity + | rewrite (Pcompare_Eq_eq x' y'); + [ reflexivity + | destruct ((x' ?= y')%positive Eq); reflexivity || discriminate ] ]. Qed. -Lemma Zcompare_EGAL : (x,y:Z) (Zcompare x y) = EGAL <-> x = y. +Lemma Zcompare_Eq_iff_eq : forall n m:Z, (n ?= m) = Eq <-> n = m. Proof. -Intros x y;Split; Intro E; [ Apply Zcompare_EGAL_eq; Assumption - | Rewrite E; Apply Zcompare_x_x ]. +intros x y; split; intro E; + [ apply Zcompare_Eq_eq; assumption | rewrite E; apply Zcompare_refl ]. Qed. -Lemma Zcompare_antisym : - (x,y:Z)(Op (Zcompare x y)) = (Zcompare y x). +Lemma Zcompare_antisym : forall n m:Z, CompOpp (n ?= m) = (m ?= n). Proof. -Intros x y; NewDestruct x; NewDestruct y; Simpl; - Reflexivity Orelse Discriminate H Orelse - Rewrite Pcompare_antisym; Reflexivity. +intros x y; destruct x; destruct y; simpl in |- *; + reflexivity || discriminate H || rewrite Pcompare_antisym; + reflexivity. Qed. -Lemma Zcompare_ANTISYM : - (x,y:Z) (Zcompare x y) = SUPERIEUR <-> (Zcompare y x) = INFERIEUR. +Lemma Zcompare_Gt_Lt_antisym : forall n m:Z, (n ?= m) = Gt <-> (m ?= n) = Lt. Proof. -Intros x y; Split; Intro H; [ - Change INFERIEUR with (Op SUPERIEUR); - Rewrite <- Zcompare_antisym; Rewrite H; Reflexivity -| Change SUPERIEUR with (Op INFERIEUR); - Rewrite <- Zcompare_antisym; Rewrite H; Reflexivity ]. +intros x y; split; intro H; + [ change Lt with (CompOpp Gt) in |- *; rewrite <- Zcompare_antisym; + rewrite H; reflexivity + | change Gt with (CompOpp Lt) in |- *; rewrite <- Zcompare_antisym; + rewrite H; reflexivity ]. Qed. (** Transitivity of comparison *) -Lemma Zcompare_trans_SUPERIEUR : - (x,y,z:Z) (Zcompare x y) = SUPERIEUR -> - (Zcompare y z) = SUPERIEUR -> - (Zcompare x z) = SUPERIEUR. +Lemma Zcompare_Gt_trans : + forall n m p:Z, (n ?= m) = Gt -> (m ?= p) = Gt -> (n ?= p) = Gt. Proof. -Intros x y z;Case x;Case y;Case z; Simpl; -Try (Intros; Discriminate H Orelse Discriminate H0); -Auto with arith; [ - Intros p q r H H0;Apply convert_compare_SUPERIEUR; Unfold gt; - Apply lt_trans with m:=(convert q); - Apply compare_convert_INFERIEUR;Apply ZC1;Assumption -| Intros p q r; Do 3 Rewrite <- ZC4; Intros H H0; - Apply convert_compare_SUPERIEUR;Unfold gt;Apply lt_trans with m:=(convert q); - Apply compare_convert_INFERIEUR;Apply ZC1;Assumption ]. +intros x y z; case x; case y; case z; simpl in |- *; + try (intros; discriminate H || discriminate H0); auto with arith; + [ intros p q r H H0; apply nat_of_P_gt_Gt_compare_complement_morphism; + unfold gt in |- *; apply lt_trans with (m := nat_of_P q); + apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; + assumption + | intros p q r; do 3 rewrite <- ZC4; intros H H0; + apply nat_of_P_gt_Gt_compare_complement_morphism; + unfold gt in |- *; apply lt_trans with (m := nat_of_P q); + apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; + assumption ]. Qed. (** Comparison and opposite *) -Lemma Zcompare_Zopp : - (x,y:Z) (Zcompare x y) = (Zcompare (Zopp y) (Zopp x)). +Lemma Zcompare_opp : forall n m:Z, (n ?= m) = (- m ?= - n). Proof. -(Intros x y;Case x;Case y;Simpl;Auto with arith); -Intros;Rewrite <- ZC4;Trivial with arith. +intros x y; case x; case y; simpl in |- *; auto with arith; intros; + rewrite <- ZC4; trivial with arith. Qed. -Hints Local Resolve convert_compare_EGAL. +Hint Local Resolve Pcompare_refl. (** Comparison first-order specification *) -Lemma SUPERIEUR_POS : - (x,y:Z) (Zcompare x y) = SUPERIEUR -> - (EX h:positive |(Zplus x (Zopp y)) = (POS h)). +Lemma Zcompare_Gt_spec : + forall n m:Z, (n ?= m) = Gt -> exists h : positive | n + - m = Zpos h. Proof. -Intros x y;Case x;Case y; [ - Simpl; Intros H; Discriminate H -| Simpl; Intros p H; Discriminate H -| Intros p H; Exists p; Simpl; Auto with arith -| Intros p H; Exists p; Simpl; Auto with arith -| Intros q p H; Exists (true_sub p q); Unfold Zplus Zopp; - Unfold Zcompare in H; Rewrite H; Trivial with arith -| Intros q p H; Exists (add p q); Simpl; Trivial with arith -| Simpl; Intros p H; Discriminate H -| Simpl; Intros q p H; Discriminate H -| Unfold Zcompare; Intros q p; Rewrite <- ZC4; Intros H; Exists (true_sub q p); - Simpl; Rewrite (ZC1 q p H); Trivial with arith]. +intros x y; case x; case y; + [ simpl in |- *; intros H; discriminate H + | simpl in |- *; intros p H; discriminate H + | intros p H; exists p; simpl in |- *; auto with arith + | intros p H; exists p; simpl in |- *; auto with arith + | intros q p H; exists (p - q)%positive; unfold Zplus, Zopp in |- *; + unfold Zcompare in H; rewrite H; trivial with arith + | intros q p H; exists (p + q)%positive; simpl in |- *; trivial with arith + | simpl in |- *; intros p H; discriminate H + | simpl in |- *; intros q p H; discriminate H + | unfold Zcompare in |- *; intros q p; rewrite <- ZC4; intros H; + exists (q - p)%positive; simpl in |- *; rewrite (ZC1 q p H); + trivial with arith ]. Qed. (** Comparison and addition *) -Lemma weaken_Zcompare_Zplus_compatible : - ((n,m:Z) (p:positive) - (Zcompare (Zplus (POS p) n) (Zplus (POS p) m)) = (Zcompare n m)) -> - (x,y,z:Z) (Zcompare (Zplus z x) (Zplus z y)) = (Zcompare x y). +Lemma weaken_Zcompare_Zplus_compatible : + (forall (n m:Z) (p:positive), (Zpos p + n ?= Zpos p + m) = (n ?= m)) -> + forall n m p:Z, (p + n ?= p + m) = (n ?= m). Proof. -Intros H x y z; NewDestruct z; [ - Reflexivity -| Apply H -| Rewrite (Zcompare_Zopp x y); Rewrite Zcompare_Zopp; - Do 2 Rewrite Zopp_Zplus; Rewrite Zopp_NEG; Apply H ]. +intros H x y z; destruct z; + [ reflexivity + | apply H + | rewrite (Zcompare_opp x y); rewrite Zcompare_opp; + do 2 rewrite Zopp_plus_distr; rewrite Zopp_neg; + apply H ]. Qed. -Hints Local Resolve ZC4. +Hint Local Resolve ZC4. -Lemma weak_Zcompare_Zplus_compatible : - (x,y:Z) (z:positive) - (Zcompare (Zplus (POS z) x) (Zplus (POS z) y)) = (Zcompare x y). +Lemma weak_Zcompare_Zplus_compatible : + forall (n m:Z) (p:positive), (Zpos p + n ?= Zpos p + m) = (n ?= m). Proof. -Intros x y z;Case x;Case y;Simpl;Auto with arith; [ - Intros p;Apply convert_compare_INFERIEUR; Apply ZL17 -| Intros p;ElimPcompare z p;Intros E;Rewrite E;Auto with arith; - Apply convert_compare_SUPERIEUR; Rewrite true_sub_convert; [ Unfold gt ; - Apply ZL16 | Assumption ] -| Intros p;ElimPcompare z p; - Intros E;Auto with arith; Apply convert_compare_SUPERIEUR; - Unfold gt;Apply ZL17 -| Intros p q; - ElimPcompare q p; - Intros E;Rewrite E;[ - Rewrite (compare_convert_EGAL q p E); Apply convert_compare_EGAL - | Apply convert_compare_INFERIEUR;Do 2 Rewrite convert_add;Apply lt_reg_l; - Apply compare_convert_INFERIEUR with 1:=E - | Apply convert_compare_SUPERIEUR;Unfold gt ;Do 2 Rewrite convert_add; - Apply lt_reg_l;Exact (compare_convert_SUPERIEUR q p E) ] -| Intros p q; - ElimPcompare z p; - Intros E;Rewrite E;Auto with arith; - Apply convert_compare_SUPERIEUR; Rewrite true_sub_convert; [ - Unfold gt; Apply lt_trans with m:=(convert z); [Apply ZL16 | Apply ZL17] - | Assumption ] -| Intros p;ElimPcompare z p;Intros E;Rewrite E;Auto with arith; Simpl; - Apply convert_compare_INFERIEUR;Rewrite true_sub_convert;[Apply ZL16| - Assumption] -| Intros p q; - ElimPcompare z q; - Intros E;Rewrite E;Auto with arith; Simpl;Apply convert_compare_INFERIEUR; - Rewrite true_sub_convert;[ - Apply lt_trans with m:=(convert z) ;[Apply ZL16|Apply ZL17] - | Assumption] -| Intros p q; ElimPcompare z q; Intros E0;Rewrite E0; - ElimPcompare z p; Intros E1;Rewrite E1; ElimPcompare q p; - Intros E2;Rewrite E2;Auto with arith; [ - Absurd (compare q p EGAL)=INFERIEUR; [ - Rewrite <- (compare_convert_EGAL z q E0); - Rewrite <- (compare_convert_EGAL z p E1); - Rewrite (convert_compare_EGAL z); Discriminate - | Assumption ] - | Absurd (compare q p EGAL)=SUPERIEUR; [ - Rewrite <- (compare_convert_EGAL z q E0); - Rewrite <- (compare_convert_EGAL z p E1); - Rewrite (convert_compare_EGAL z);Discriminate - | Assumption] - | Absurd (compare z p EGAL)=INFERIEUR; [ - Rewrite (compare_convert_EGAL z q E0); - Rewrite <- (compare_convert_EGAL q p E2); - Rewrite (convert_compare_EGAL q);Discriminate - | Assumption ] - | Absurd (compare z p EGAL)=INFERIEUR; [ - Rewrite (compare_convert_EGAL z q E0); Rewrite E2;Discriminate - | Assumption] - | Absurd (compare z p EGAL)=SUPERIEUR;[ - Rewrite (compare_convert_EGAL z q E0); - Rewrite <- (compare_convert_EGAL q p E2); - Rewrite (convert_compare_EGAL q);Discriminate - | Assumption] - | Absurd (compare z p EGAL)=SUPERIEUR;[ - Rewrite (compare_convert_EGAL z q E0);Rewrite E2;Discriminate - | Assumption] - | Absurd (compare z q EGAL)=INFERIEUR;[ - Rewrite (compare_convert_EGAL z p E1); - Rewrite (compare_convert_EGAL q p E2); - Rewrite (convert_compare_EGAL p); Discriminate - | Assumption] - | Absurd (compare p q EGAL)=SUPERIEUR; [ - Rewrite <- (compare_convert_EGAL z p E1); - Rewrite E0; Discriminate - | Apply ZC2;Assumption ] - | Simpl; Rewrite (compare_convert_EGAL q p E2); - Rewrite (convert_compare_EGAL (true_sub p z)); Auto with arith - | Simpl; Rewrite <- ZC4; Apply convert_compare_SUPERIEUR; - Rewrite true_sub_convert; [ - Rewrite true_sub_convert; [ - Unfold gt; Apply simpl_lt_plus_l with p:=(convert z); - Rewrite le_plus_minus_r; [ - Rewrite le_plus_minus_r; [ - Apply compare_convert_INFERIEUR;Assumption - | Apply lt_le_weak; Apply compare_convert_INFERIEUR;Assumption ] - | Apply lt_le_weak; Apply compare_convert_INFERIEUR;Assumption ] - | Apply ZC2;Assumption ] - | Apply ZC2;Assumption ] - | Simpl; Rewrite <- ZC4; Apply convert_compare_INFERIEUR; - Rewrite true_sub_convert; [ - Rewrite true_sub_convert; [ - Apply simpl_lt_plus_l with p:=(convert z); - Rewrite le_plus_minus_r; [ - Rewrite le_plus_minus_r; [ - Apply compare_convert_INFERIEUR;Apply ZC1;Assumption - | Apply lt_le_weak; Apply compare_convert_INFERIEUR;Assumption ] - | Apply lt_le_weak; Apply compare_convert_INFERIEUR;Assumption ] - | Apply ZC2;Assumption] - | Apply ZC2;Assumption ] - | Absurd (compare z q EGAL)=INFERIEUR; [ - Rewrite (compare_convert_EGAL q p E2);Rewrite E1;Discriminate - | Assumption ] - | Absurd (compare q p EGAL)=INFERIEUR; [ - Cut (compare q p EGAL)=SUPERIEUR; [ - Intros E;Rewrite E;Discriminate - | Apply convert_compare_SUPERIEUR; Unfold gt; - Apply lt_trans with m:=(convert z); [ - Apply compare_convert_INFERIEUR;Apply ZC1;Assumption - | Apply compare_convert_INFERIEUR;Assumption ]] - | Assumption ] - | Absurd (compare z q EGAL)=SUPERIEUR; [ - Rewrite (compare_convert_EGAL z p E1); - Rewrite (compare_convert_EGAL q p E2); - Rewrite (convert_compare_EGAL p); Discriminate - | Assumption ] - | Absurd (compare z q EGAL)=SUPERIEUR; [ - Rewrite (compare_convert_EGAL z p E1); - Rewrite ZC1; [Discriminate | Assumption ] - | Assumption ] - | Absurd (compare z q EGAL)=SUPERIEUR; [ - Rewrite (compare_convert_EGAL q p E2); Rewrite E1; Discriminate - | Assumption ] - | Absurd (compare q p EGAL)=SUPERIEUR; [ - Rewrite ZC1; [ - Discriminate - | Apply convert_compare_SUPERIEUR; Unfold gt; - Apply lt_trans with m:=(convert z); [ - Apply compare_convert_INFERIEUR;Apply ZC1;Assumption - | Apply compare_convert_INFERIEUR;Assumption ]] - | Assumption ] - | Simpl; Rewrite (compare_convert_EGAL q p E2); Apply convert_compare_EGAL - | Simpl; Apply convert_compare_SUPERIEUR; Unfold gt; - Rewrite true_sub_convert; [ - Rewrite true_sub_convert; [ - Apply simpl_lt_plus_l with p:=(convert p); Rewrite le_plus_minus_r; [ - Rewrite plus_sym; Apply simpl_lt_plus_l with p:=(convert q); - Rewrite plus_assoc_l; Rewrite le_plus_minus_r; [ - Rewrite (plus_sym (convert q)); Apply lt_reg_l; - Apply compare_convert_INFERIEUR;Assumption - | Apply lt_le_weak; Apply compare_convert_INFERIEUR; - Apply ZC1;Assumption ] - | Apply lt_le_weak; Apply compare_convert_INFERIEUR;Apply ZC1; - Assumption ] - | Assumption ] - | Assumption ] - | Simpl; Apply convert_compare_INFERIEUR; Rewrite true_sub_convert; [ - Rewrite true_sub_convert; [ - Apply simpl_lt_plus_l with p:=(convert q); Rewrite le_plus_minus_r; [ - Rewrite plus_sym; Apply simpl_lt_plus_l with p:=(convert p); - Rewrite plus_assoc_l; Rewrite le_plus_minus_r; [ - Rewrite (plus_sym (convert p)); Apply lt_reg_l; - Apply compare_convert_INFERIEUR;Apply ZC1;Assumption - | Apply lt_le_weak; Apply compare_convert_INFERIEUR;Apply ZC1; - Assumption ] - | Apply lt_le_weak;Apply compare_convert_INFERIEUR;Apply ZC1;Assumption] - | Assumption] - | Assumption]]]. +intros x y z; case x; case y; simpl in |- *; auto with arith; + [ intros p; apply nat_of_P_lt_Lt_compare_complement_morphism; apply ZL17 + | intros p; ElimPcompare z p; intros E; rewrite E; auto with arith; + apply nat_of_P_gt_Gt_compare_complement_morphism; + rewrite nat_of_P_minus_morphism; + [ unfold gt in |- *; apply ZL16 | assumption ] + | intros p; ElimPcompare z p; intros E; auto with arith; + apply nat_of_P_gt_Gt_compare_complement_morphism; + unfold gt in |- *; apply ZL17 + | intros p q; ElimPcompare q p; intros E; rewrite E; + [ rewrite (Pcompare_Eq_eq q p E); apply Pcompare_refl + | apply nat_of_P_lt_Lt_compare_complement_morphism; + do 2 rewrite nat_of_P_plus_morphism; apply plus_lt_compat_l; + apply nat_of_P_lt_Lt_compare_morphism with (1 := E) + | apply nat_of_P_gt_Gt_compare_complement_morphism; unfold gt in |- *; + do 2 rewrite nat_of_P_plus_morphism; apply plus_lt_compat_l; + exact (nat_of_P_gt_Gt_compare_morphism q p E) ] + | intros p q; ElimPcompare z p; intros E; rewrite E; auto with arith; + apply nat_of_P_gt_Gt_compare_complement_morphism; + rewrite nat_of_P_minus_morphism; + [ unfold gt in |- *; apply lt_trans with (m := nat_of_P z); + [ apply ZL16 | apply ZL17 ] + | assumption ] + | intros p; ElimPcompare z p; intros E; rewrite E; auto with arith; + simpl in |- *; apply nat_of_P_lt_Lt_compare_complement_morphism; + rewrite nat_of_P_minus_morphism; [ apply ZL16 | assumption ] + | intros p q; ElimPcompare z q; intros E; rewrite E; auto with arith; + simpl in |- *; apply nat_of_P_lt_Lt_compare_complement_morphism; + rewrite nat_of_P_minus_morphism; + [ apply lt_trans with (m := nat_of_P z); [ apply ZL16 | apply ZL17 ] + | assumption ] + | intros p q; ElimPcompare z q; intros E0; rewrite E0; ElimPcompare z p; + intros E1; rewrite E1; ElimPcompare q p; intros E2; + rewrite E2; auto with arith; + [ absurd ((q ?= p)%positive Eq = Lt); + [ rewrite <- (Pcompare_Eq_eq z q E0); + rewrite <- (Pcompare_Eq_eq z p E1); rewrite (Pcompare_refl z); + discriminate + | assumption ] + | absurd ((q ?= p)%positive Eq = Gt); + [ rewrite <- (Pcompare_Eq_eq z q E0); + rewrite <- (Pcompare_Eq_eq z p E1); rewrite (Pcompare_refl z); + discriminate + | assumption ] + | absurd ((z ?= p)%positive Eq = Lt); + [ rewrite (Pcompare_Eq_eq z q E0); rewrite <- (Pcompare_Eq_eq q p E2); + rewrite (Pcompare_refl q); discriminate + | assumption ] + | absurd ((z ?= p)%positive Eq = Lt); + [ rewrite (Pcompare_Eq_eq z q E0); rewrite E2; discriminate + | assumption ] + | absurd ((z ?= p)%positive Eq = Gt); + [ rewrite (Pcompare_Eq_eq z q E0); rewrite <- (Pcompare_Eq_eq q p E2); + rewrite (Pcompare_refl q); discriminate + | assumption ] + | absurd ((z ?= p)%positive Eq = Gt); + [ rewrite (Pcompare_Eq_eq z q E0); rewrite E2; discriminate + | assumption ] + | absurd ((z ?= q)%positive Eq = Lt); + [ rewrite (Pcompare_Eq_eq z p E1); rewrite (Pcompare_Eq_eq q p E2); + rewrite (Pcompare_refl p); discriminate + | assumption ] + | absurd ((p ?= q)%positive Eq = Gt); + [ rewrite <- (Pcompare_Eq_eq z p E1); rewrite E0; discriminate + | apply ZC2; assumption ] + | simpl in |- *; rewrite (Pcompare_Eq_eq q p E2); + rewrite (Pcompare_refl (p - z)); auto with arith + | simpl in |- *; rewrite <- ZC4; + apply nat_of_P_gt_Gt_compare_complement_morphism; + rewrite nat_of_P_minus_morphism; + [ rewrite nat_of_P_minus_morphism; + [ unfold gt in |- *; apply plus_lt_reg_l with (p := nat_of_P z); + rewrite le_plus_minus_r; + [ rewrite le_plus_minus_r; + [ apply nat_of_P_lt_Lt_compare_morphism; assumption + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + assumption ] + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + assumption ] + | apply ZC2; assumption ] + | apply ZC2; assumption ] + | simpl in |- *; rewrite <- ZC4; + apply nat_of_P_lt_Lt_compare_complement_morphism; + rewrite nat_of_P_minus_morphism; + [ rewrite nat_of_P_minus_morphism; + [ apply plus_lt_reg_l with (p := nat_of_P z); + rewrite le_plus_minus_r; + [ rewrite le_plus_minus_r; + [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; + assumption + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + assumption ] + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + assumption ] + | apply ZC2; assumption ] + | apply ZC2; assumption ] + | absurd ((z ?= q)%positive Eq = Lt); + [ rewrite (Pcompare_Eq_eq q p E2); rewrite E1; discriminate + | assumption ] + | absurd ((q ?= p)%positive Eq = Lt); + [ cut ((q ?= p)%positive Eq = Gt); + [ intros E; rewrite E; discriminate + | apply nat_of_P_gt_Gt_compare_complement_morphism; + unfold gt in |- *; apply lt_trans with (m := nat_of_P z); + [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; assumption + | apply nat_of_P_lt_Lt_compare_morphism; assumption ] ] + | assumption ] + | absurd ((z ?= q)%positive Eq = Gt); + [ rewrite (Pcompare_Eq_eq z p E1); rewrite (Pcompare_Eq_eq q p E2); + rewrite (Pcompare_refl p); discriminate + | assumption ] + | absurd ((z ?= q)%positive Eq = Gt); + [ rewrite (Pcompare_Eq_eq z p E1); rewrite ZC1; + [ discriminate | assumption ] + | assumption ] + | absurd ((z ?= q)%positive Eq = Gt); + [ rewrite (Pcompare_Eq_eq q p E2); rewrite E1; discriminate + | assumption ] + | absurd ((q ?= p)%positive Eq = Gt); + [ rewrite ZC1; + [ discriminate + | apply nat_of_P_gt_Gt_compare_complement_morphism; + unfold gt in |- *; apply lt_trans with (m := nat_of_P z); + [ apply nat_of_P_lt_Lt_compare_morphism; apply ZC1; assumption + | apply nat_of_P_lt_Lt_compare_morphism; assumption ] ] + | assumption ] + | simpl in |- *; rewrite (Pcompare_Eq_eq q p E2); apply Pcompare_refl + | simpl in |- *; apply nat_of_P_gt_Gt_compare_complement_morphism; + unfold gt in |- *; rewrite nat_of_P_minus_morphism; + [ rewrite nat_of_P_minus_morphism; + [ apply plus_lt_reg_l with (p := nat_of_P p); + rewrite le_plus_minus_r; + [ rewrite plus_comm; apply plus_lt_reg_l with (p := nat_of_P q); + rewrite plus_assoc; rewrite le_plus_minus_r; + [ rewrite (plus_comm (nat_of_P q)); apply plus_lt_compat_l; + apply nat_of_P_lt_Lt_compare_morphism; + assumption + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + apply ZC1; assumption ] + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + apply ZC1; assumption ] + | assumption ] + | assumption ] + | simpl in |- *; apply nat_of_P_lt_Lt_compare_complement_morphism; + rewrite nat_of_P_minus_morphism; + [ rewrite nat_of_P_minus_morphism; + [ apply plus_lt_reg_l with (p := nat_of_P q); + rewrite le_plus_minus_r; + [ rewrite plus_comm; apply plus_lt_reg_l with (p := nat_of_P p); + rewrite plus_assoc; rewrite le_plus_minus_r; + [ rewrite (plus_comm (nat_of_P p)); apply plus_lt_compat_l; + apply nat_of_P_lt_Lt_compare_morphism; + apply ZC1; assumption + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + apply ZC1; assumption ] + | apply lt_le_weak; apply nat_of_P_lt_Lt_compare_morphism; + apply ZC1; assumption ] + | assumption ] + | assumption ] ] ]. Qed. -Lemma Zcompare_Zplus_compatible : - (x,y,z:Z) (Zcompare (Zplus z x) (Zplus z y)) = (Zcompare x y). +Lemma Zcompare_plus_compat : forall n m p:Z, (p + n ?= p + m) = (n ?= m). Proof. -Exact (weaken_Zcompare_Zplus_compatible weak_Zcompare_Zplus_compatible). +exact (weaken_Zcompare_Zplus_compatible weak_Zcompare_Zplus_compatible). Qed. -Lemma Zcompare_Zplus_compatible2 : - (r:relation)(x,y,z,t:Z) - (Zcompare x y) = r -> (Zcompare z t) = r -> - (Zcompare (Zplus x z) (Zplus y t)) = r. +Lemma Zplus_compare_compat : + forall (r:comparison) (n m p q:Z), + (n ?= m) = r -> (p ?= q) = r -> (n + p ?= m + q) = r. Proof. -Intros r x y z t; Case r; [ - Intros H1 H2; Elim (Zcompare_EGAL x y); Elim (Zcompare_EGAL z t); - Intros H3 H4 H5 H6; Rewrite H3; [ - Rewrite H5; [ Elim (Zcompare_EGAL (Zplus y t) (Zplus y t)); Auto with arith | Auto with arith ] - | Auto with arith ] -| Intros H1 H2; Elim (Zcompare_ANTISYM (Zplus y t) (Zplus x z)); - Intros H3 H4; Apply H3; - Apply Zcompare_trans_SUPERIEUR with y:=(Zplus y z) ; [ - Rewrite Zcompare_Zplus_compatible; - Elim (Zcompare_ANTISYM t z); Auto with arith - | Do 2 Rewrite <- (Zplus_sym z); - Rewrite Zcompare_Zplus_compatible; - Elim (Zcompare_ANTISYM y x); Auto with arith] -| Intros H1 H2; - Apply Zcompare_trans_SUPERIEUR with y:=(Zplus x t) ; [ - Rewrite Zcompare_Zplus_compatible; Assumption - | Do 2 Rewrite <- (Zplus_sym t); - Rewrite Zcompare_Zplus_compatible; Assumption]]. +intros r x y z t; case r; + [ intros H1 H2; elim (Zcompare_Eq_iff_eq x y); elim (Zcompare_Eq_iff_eq z t); + intros H3 H4 H5 H6; rewrite H3; + [ rewrite H5; + [ elim (Zcompare_Eq_iff_eq (y + t) (y + t)); auto with arith + | auto with arith ] + | auto with arith ] + | intros H1 H2; elim (Zcompare_Gt_Lt_antisym (y + t) (x + z)); intros H3 H4; + apply H3; apply Zcompare_Gt_trans with (m := y + z); + [ rewrite Zcompare_plus_compat; elim (Zcompare_Gt_Lt_antisym t z); + auto with arith + | do 2 rewrite <- (Zplus_comm z); rewrite Zcompare_plus_compat; + elim (Zcompare_Gt_Lt_antisym y x); auto with arith ] + | intros H1 H2; apply Zcompare_Gt_trans with (m := x + t); + [ rewrite Zcompare_plus_compat; assumption + | do 2 rewrite <- (Zplus_comm t); rewrite Zcompare_plus_compat; + assumption ] ]. Qed. -Lemma Zcompare_Zs_SUPERIEUR : (x:Z)(Zcompare (Zs x) x)=SUPERIEUR. +Lemma Zcompare_succ_Gt : forall n:Z, (Zsucc n ?= n) = Gt. Proof. -Intro x; Unfold Zs; Pattern 2 x; Rewrite <- (Zero_right x); -Rewrite Zcompare_Zplus_compatible;Reflexivity. +intro x; unfold Zsucc in |- *; pattern x at 2 in |- *; + rewrite <- (Zplus_0_r x); rewrite Zcompare_plus_compat; + reflexivity. Qed. -Lemma Zcompare_et_un: - (x,y:Z) (Zcompare x y)=SUPERIEUR <-> - ~(Zcompare x (Zplus y (POS xH)))=INFERIEUR. +Lemma Zcompare_Gt_not_Lt : forall n m:Z, (n ?= m) = Gt <-> (n ?= m + 1) <> Lt. Proof. -Intros x y; Split; [ - Intro H; (ElimCompare 'x '(Zplus y (POS xH)));[ - Intro H1; Rewrite H1; Discriminate - | Intros H1; Elim SUPERIEUR_POS with 1:=H; Intros h H2; - Absurd (gt (convert h) O) /\ (lt (convert h) (S O)); [ - Unfold not ;Intros H3;Elim H3;Intros H4 H5; Absurd (gt (convert h) O); [ - Unfold gt ;Apply le_not_lt; Apply le_S_n; Exact H5 - | Assumption] - | Split; [ - Elim (ZL4 h); Intros i H3;Rewrite H3; Apply gt_Sn_O - | Change (lt (convert h) (convert xH)); - Apply compare_convert_INFERIEUR; - Change (Zcompare (POS h) (POS xH))=INFERIEUR; - Rewrite <- H2; Rewrite <- [m,n:Z](Zcompare_Zplus_compatible m n y); - Rewrite (Zplus_sym x);Rewrite Zplus_assoc; Rewrite Zplus_inverse_r; - Simpl; Exact H1 ]] - | Intros H1;Rewrite -> H1;Discriminate ] -| Intros H; (ElimCompare 'x '(Zplus y (POS xH))); [ - Intros H1;Elim (Zcompare_EGAL x (Zplus y (POS xH))); Intros H2 H3; - Rewrite (H2 H1); Exact (Zcompare_Zs_SUPERIEUR y) - | Intros H1;Absurd (Zcompare x (Zplus y (POS xH)))=INFERIEUR;Assumption - | Intros H1; Apply Zcompare_trans_SUPERIEUR with y:=(Zs y); - [ Exact H1 | Exact (Zcompare_Zs_SUPERIEUR y)]]]. +intros x y; split; + [ intro H; elim_compare x (y + 1); + [ intro H1; rewrite H1; discriminate + | intros H1; elim Zcompare_Gt_spec with (1 := H); intros h H2; + absurd ((nat_of_P h > 0)%nat /\ (nat_of_P h < 1)%nat); + [ unfold not in |- *; intros H3; elim H3; intros H4 H5; + absurd (nat_of_P h > 0)%nat; + [ unfold gt in |- *; apply le_not_lt; apply le_S_n; exact H5 + | assumption ] + | split; + [ elim (ZL4 h); intros i H3; rewrite H3; apply gt_Sn_O + | change (nat_of_P h < nat_of_P 1)%nat in |- *; + apply nat_of_P_lt_Lt_compare_morphism; + change ((Zpos h ?= 1) = Lt) in |- *; rewrite <- H2; + rewrite <- (fun m n:Z => Zcompare_plus_compat m n y); + rewrite (Zplus_comm x); rewrite Zplus_assoc; + rewrite Zplus_opp_r; simpl in |- *; exact H1 ] ] + | intros H1; rewrite H1; discriminate ] + | intros H; elim_compare x (y + 1); + [ intros H1; elim (Zcompare_Eq_iff_eq x (y + 1)); intros H2 H3; + rewrite (H2 H1); exact (Zcompare_succ_Gt y) + | intros H1; absurd ((x ?= y + 1) = Lt); assumption + | intros H1; apply Zcompare_Gt_trans with (m := Zsucc y); + [ exact H1 | exact (Zcompare_succ_Gt y) ] ] ]. Qed. (** Successor and comparison *) -Lemma Zcompare_n_S : (n,m:Z)(Zcompare (Zs n) (Zs m)) = (Zcompare n m). +Lemma Zcompare_succ_compat : forall n m:Z, (Zsucc n ?= Zsucc m) = (n ?= m). Proof. -Intros n m;Unfold Zs ;Do 2 Rewrite -> [t:Z](Zplus_sym t (POS xH)); -Rewrite -> Zcompare_Zplus_compatible;Auto with arith. +intros n m; unfold Zsucc in |- *; do 2 rewrite (fun t:Z => Zplus_comm t 1); + rewrite Zcompare_plus_compat; auto with arith. Qed. (** Multiplication and comparison *) -Lemma Zcompare_Zmult_compatible : - (x:positive)(y,z:Z) - (Zcompare (Zmult (POS x) y) (Zmult (POS x) z)) = (Zcompare y z). +Lemma Zcompare_mult_compat : + forall (p:positive) (n m:Z), (Zpos p * n ?= Zpos p * m) = (n ?= m). Proof. -Intros x; NewInduction x as [p H|p H|]; [ - Intros y z; - Cut (POS (xI p))=(Zplus (Zplus (POS p) (POS p)) (POS xH)); [ - Intros E; Rewrite E; Do 4 Rewrite Zmult_plus_distr_l; - Do 2 Rewrite Zmult_one; - Apply Zcompare_Zplus_compatible2; [ - Apply Zcompare_Zplus_compatible2; Apply H - | Trivial with arith] - | Simpl; Rewrite (add_x_x p); Trivial with arith] -| Intros y z; Cut (POS (xO p))=(Zplus (POS p) (POS p)); [ - Intros E; Rewrite E; Do 2 Rewrite Zmult_plus_distr_l; - Apply Zcompare_Zplus_compatible2; Apply H - | Simpl; Rewrite (add_x_x p); Trivial with arith] - | Intros y z; Do 2 Rewrite Zmult_one; Trivial with arith]. +intros x; induction x as [p H| p H| ]; + [ intros y z; cut (Zpos (xI p) = Zpos p + Zpos p + 1); + [ intros E; rewrite E; do 4 rewrite Zmult_plus_distr_l; + do 2 rewrite Zmult_1_l; apply Zplus_compare_compat; + [ apply Zplus_compare_compat; apply H | trivial with arith ] + | simpl in |- *; rewrite (Pplus_diag p); trivial with arith ] + | intros y z; cut (Zpos (xO p) = Zpos p + Zpos p); + [ intros E; rewrite E; do 2 rewrite Zmult_plus_distr_l; + apply Zplus_compare_compat; apply H + | simpl in |- *; rewrite (Pplus_diag p); trivial with arith ] + | intros y z; do 2 rewrite Zmult_1_l; trivial with arith ]. Qed. (** Reverting [x ?= y] to trichotomy *) -Lemma rename : (A:Set)(P:A->Prop)(x:A) ((y:A)(x=y)->(P y)) -> (P x). +Lemma rename : + forall (A:Set) (P:A -> Prop) (x:A), (forall y:A, x = y -> P y) -> P x. Proof. -Auto with arith. +auto with arith. Qed. Lemma Zcompare_elim : - (c1,c2,c3:Prop)(x,y:Z) - ((x=y) -> c1) ->(`x c2) ->(`x>y`-> c3) - -> Cases (Zcompare x y) of EGAL => c1 | INFERIEUR => c2 | SUPERIEUR => c3 end. + forall (c1 c2 c3:Prop) (n m:Z), + (n = m -> c1) -> + (n < m -> c2) -> + (n > m -> c3) -> match n ?= m with + | Eq => c1 + | Lt => c2 + | Gt => c3 + end. Proof. -Intros c1 c2 c3 x y; Intros. -Apply rename with x:=(Zcompare x y); Intro r; Elim r; -[ Intro; Apply H; Apply (Zcompare_EGAL_eq x y); Assumption -| Unfold Zlt in H0; Assumption -| Unfold Zgt in H1; Assumption ]. +intros c1 c2 c3 x y; intros. +apply rename with (x := x ?= y); intro r; elim r; + [ intro; apply H; apply (Zcompare_Eq_eq x y); assumption + | unfold Zlt in H0; assumption + | unfold Zgt in H1; assumption ]. Qed. -Lemma Zcompare_eq_case : - (c1,c2,c3:Prop)(x,y:Z) c1 -> x=y -> - Cases (Zcompare x y) of EGAL => c1 | INFERIEUR => c2 | SUPERIEUR => c3 end. +Lemma Zcompare_eq_case : + forall (c1 c2 c3:Prop) (n m:Z), + c1 -> n = m -> match n ?= m with + | Eq => c1 + | Lt => c2 + | Gt => c3 + end. Proof. -Intros c1 c2 c3 x y; Intros. -Rewrite H0; Rewrite (Zcompare_x_x). -Assumption. +intros c1 c2 c3 x y; intros. +rewrite H0; rewrite Zcompare_refl. +assumption. Qed. (** Decompose an egality between two [?=] relations into 3 implications *) Lemma Zcompare_egal_dec : - (x1,y1,x2,y2:Z) - (`x1`x2((Zcompare x1 y1)=EGAL -> (Zcompare x2 y2)=EGAL) - ->(`x1>y1`->`x2>y2`)->(Zcompare x1 y1)=(Zcompare x2 y2). + forall n m p q:Z, + (n < m -> p < q) -> + ((n ?= m) = Eq -> (p ?= q) = Eq) -> + (n > m -> p > q) -> (n ?= m) = (p ?= q). Proof. -Intros x1 y1 x2 y2. -Unfold Zgt; Unfold Zlt; -Case (Zcompare x1 y1); Case (Zcompare x2 y2); Auto with arith; Symmetry; Auto with arith. +intros x1 y1 x2 y2. +unfold Zgt in |- *; unfold Zlt in |- *; case (x1 ?= y1); case (x2 ?= y2); + auto with arith; symmetry in |- *; auto with arith. Qed. (** Relating [x ?= y] to [Zle], [Zlt], [Zge] or [Zgt] *) -Lemma Zle_Zcompare : - (x,y:Z)`x<=y` -> - Cases (Zcompare x y) of EGAL => True | INFERIEUR => True | SUPERIEUR => False end. +Lemma Zle_compare : + forall n m:Z, + n <= m -> match n ?= m with + | Eq => True + | Lt => True + | Gt => False + end. Proof. -Intros x y; Unfold Zle; Elim (Zcompare x y); Auto with arith. +intros x y; unfold Zle in |- *; elim (x ?= y); auto with arith. Qed. -Lemma Zlt_Zcompare : - (x,y:Z)`x - Cases (Zcompare x y) of EGAL => False | INFERIEUR => True | SUPERIEUR => False end. +Lemma Zlt_compare : + forall n m:Z, + n < m -> match n ?= m with + | Eq => False + | Lt => True + | Gt => False + end. Proof. -Intros x y; Unfold Zlt; Elim (Zcompare x y); Intros; Discriminate Orelse Trivial with arith. +intros x y; unfold Zlt in |- *; elim (x ?= y); intros; + discriminate || trivial with arith. Qed. -Lemma Zge_Zcompare : - (x,y:Z)`x>=y`-> - Cases (Zcompare x y) of EGAL => True | INFERIEUR => False | SUPERIEUR => True end. +Lemma Zge_compare : + forall n m:Z, + n >= m -> match n ?= m with + | Eq => True + | Lt => False + | Gt => True + end. Proof. -Intros x y; Unfold Zge; Elim (Zcompare x y); Auto with arith. +intros x y; unfold Zge in |- *; elim (x ?= y); auto with arith. Qed. -Lemma Zgt_Zcompare : - (x,y:Z)`x>y` -> - Cases (Zcompare x y) of EGAL => False | INFERIEUR => False | SUPERIEUR => True end. +Lemma Zgt_compare : + forall n m:Z, + n > m -> match n ?= m with + | Eq => False + | Lt => False + | Gt => True + end. Proof. -Intros x y; Unfold Zgt; Elim (Zcompare x y); Intros; Discriminate Orelse Trivial with arith. +intros x y; unfold Zgt in |- *; elim (x ?= y); intros; + discriminate || trivial with arith. Qed. (**********************************************************************) (* Other properties *) -V7only [Set Implicit Arguments.]. -Lemma Zcompare_Zmult_left : (x,y,z:Z)`z>0` -> `x ?= y`=`z*x ?= z*y`. +Lemma Zmult_compare_compat_l : + forall n m p:Z, p > 0 -> (n ?= m) = (p * n ?= p * m). Proof. -Intros x y z H; NewDestruct z. - Discriminate H. - Rewrite Zcompare_Zmult_compatible; Reflexivity. - Discriminate H. +intros x y z H; destruct z. + discriminate H. + rewrite Zcompare_mult_compat; reflexivity. + discriminate H. Qed. -Lemma Zcompare_Zmult_right : (x,y,z:Z)` z>0` -> `x ?= y`=`x*z ?= y*z`. +Lemma Zmult_compare_compat_r : + forall n m p:Z, p > 0 -> (n ?= m) = (n * p ?= m * p). Proof. -Intros x y z H; -Rewrite (Zmult_sym x z); -Rewrite (Zmult_sym y z); -Apply Zcompare_Zmult_left; Assumption. +intros x y z H; rewrite (Zmult_comm x z); rewrite (Zmult_comm y z); + apply Zmult_compare_compat_l; assumption. Qed. -V7only [Unset Implicit Arguments.]. - diff --git a/theories/ZArith/Zcomplements.v b/theories/ZArith/Zcomplements.v index 8d27f81d2d..01e8d4870b 100644 --- a/theories/ZArith/Zcomplements.v +++ b/theories/ZArith/Zcomplements.v @@ -8,39 +8,38 @@ (*i $Id$ i*) -Require ZArithRing. -Require ZArith_base. -Require Omega. -Require Wf_nat. -V7only [Import Z_scope.]. +Require Import ZArithRing. +Require Import ZArith_base. +Require Import Omega. +Require Import Wf_nat. Open Local Scope Z_scope. -V7only [Set Implicit Arguments.]. (**********************************************************************) (** About parity *) -Lemma two_or_two_plus_one : (x:Z) { y:Z | `x = 2*y`}+{ y:Z | `x = 2*y+1`}. +Lemma two_or_two_plus_one : + forall n:Z, {y : Z | n = 2 * y} + {y : Z | n = 2 * y + 1}. Proof. -Intro x; NewDestruct x. -Left ; Split with ZERO; Reflexivity. +intro x; destruct x. +left; split with 0; reflexivity. -NewDestruct p. -Right ; Split with (POS p); Reflexivity. +destruct p. +right; split with (Zpos p); reflexivity. -Left ; Split with (POS p); Reflexivity. +left; split with (Zpos p); reflexivity. -Right ; Split with ZERO; Reflexivity. +right; split with 0; reflexivity. -NewDestruct p. -Right ; Split with (NEG (add xH p)). -Rewrite NEG_xI. -Rewrite NEG_add. -Omega. +destruct p. +right; split with (Zneg (1 + p)). +rewrite BinInt.Zneg_xI. +rewrite BinInt.Zneg_plus_distr. +omega. -Left ; Split with (NEG p); Reflexivity. +left; split with (Zneg p); reflexivity. -Right ; Split with `-1`; Reflexivity. +right; split with (-1); reflexivity. Qed. (**********************************************************************) @@ -49,164 +48,165 @@ Qed. Easy to compute: replace all "1" of the binary representation by "0", except the first "1" (or the first one :-) *) -Fixpoint floor_pos [a : positive] : positive := - Cases a of - | xH => xH - | (xO a') => (xO (floor_pos a')) - | (xI b') => (xO (floor_pos b')) +Fixpoint floor_pos (a:positive) : positive := + match a with + | xH => 1%positive + | xO a' => xO (floor_pos a') + | xI b' => xO (floor_pos b') end. -Definition floor := [a:positive](POS (floor_pos a)). +Definition floor (a:positive) := Zpos (floor_pos a). -Lemma floor_gt0 : (x:positive) `(floor x) > 0`. +Lemma floor_gt0 : forall p:positive, floor p > 0. Proof. -Intro. -Compute. -Trivial. +intro. +compute in |- *. +trivial. Qed. -Lemma floor_ok : (a:positive) - `(floor a) <= (POS a) < 2*(floor a)`. +Lemma floor_ok : forall p:positive, floor p <= Zpos p < 2 * floor p. Proof. -Unfold floor. -Intro a; NewInduction a as [p|p|]. - -Simpl. -Repeat Rewrite POS_xI. -Rewrite (POS_xO (xO (floor_pos p))). -Rewrite (POS_xO (floor_pos p)). -Omega. - -Simpl. -Repeat Rewrite POS_xI. -Rewrite (POS_xO (xO (floor_pos p))). -Rewrite (POS_xO (floor_pos p)). -Rewrite (POS_xO p). -Omega. - -Simpl; Omega. +unfold floor in |- *. +intro a; induction a as [p| p| ]. + +simpl in |- *. +repeat rewrite BinInt.Zpos_xI. +rewrite (BinInt.Zpos_xO (xO (floor_pos p))). +rewrite (BinInt.Zpos_xO (floor_pos p)). +omega. + +simpl in |- *. +repeat rewrite BinInt.Zpos_xI. +rewrite (BinInt.Zpos_xO (xO (floor_pos p))). +rewrite (BinInt.Zpos_xO (floor_pos p)). +rewrite (BinInt.Zpos_xO p). +omega. + +simpl in |- *; omega. Qed. (**********************************************************************) (** Two more induction principles over [Z]. *) -Theorem Z_lt_abs_rec : (P: Z -> Set) - ((n: Z) ((m: Z) `|m|<|n|` -> (P m)) -> (P n)) -> (p: Z) (P p). +Theorem Z_lt_abs_rec : + forall P:Z -> Set, + (forall n:Z, (forall m:Z, Zabs m < Zabs n -> P m) -> P n) -> + forall n:Z, P n. Proof. -Intros P HP p. -LetTac Q:=[z]`0<=z`->(P z)*(P `-z`). -Cut (Q `|p|`);[Intros|Apply (Z_lt_rec Q);Auto with zarith]. -Elim (Zabs_dec p);Intro eq;Rewrite eq;Elim H;Auto with zarith. -Unfold Q;Clear Q;Intros. -Apply pair;Apply HP. -Rewrite Zabs_eq;Auto;Intros. -Elim (H `|m|`);Intros;Auto with zarith. -Elim (Zabs_dec m);Intro eq;Rewrite eq;Trivial. -Rewrite Zabs_non_eq;Auto with zarith. -Rewrite Zopp_Zopp;Intros. -Elim (H `|m|`);Intros;Auto with zarith. -Elim (Zabs_dec m);Intro eq;Rewrite eq;Trivial. +intros P HP p. +set (Q := fun z => 0 <= z -> P z * P (- z)) in *. +cut (Q (Zabs p)); [ intros | apply (Z_lt_rec Q); auto with zarith ]. +elim (Zabs_dec p); intro eq; rewrite eq; elim H; auto with zarith. +unfold Q in |- *; clear Q; intros. +apply pair; apply HP. +rewrite Zabs_eq; auto; intros. +elim (H (Zabs m)); intros; auto with zarith. +elim (Zabs_dec m); intro eq; rewrite eq; trivial. +rewrite Zabs_non_eq; auto with zarith. +rewrite Zopp_involutive; intros. +elim (H (Zabs m)); intros; auto with zarith. +elim (Zabs_dec m); intro eq; rewrite eq; trivial. Qed. -Theorem Z_lt_abs_induction : (P: Z -> Prop) - ((n: Z) ((m: Z) `|m|<|n|` -> (P m)) -> (P n)) -> (p: Z) (P p). +Theorem Z_lt_abs_induction : + forall P:Z -> Prop, + (forall n:Z, (forall m:Z, Zabs m < Zabs n -> P m) -> P n) -> + forall n:Z, P n. Proof. -Intros P HP p. -LetTac Q:=[z]`0<=z`->(P z) /\ (P `-z`). -Cut (Q `|p|`);[Intros|Apply (Z_lt_induction Q);Auto with zarith]. -Elim (Zabs_dec p);Intro eq;Rewrite eq;Elim H;Auto with zarith. -Unfold Q;Clear Q;Intros. -Split;Apply HP. -Rewrite Zabs_eq;Auto;Intros. -Elim (H `|m|`);Intros;Auto with zarith. -Elim (Zabs_dec m);Intro eq;Rewrite eq;Trivial. -Rewrite Zabs_non_eq;Auto with zarith. -Rewrite Zopp_Zopp;Intros. -Elim (H `|m|`);Intros;Auto with zarith. -Elim (Zabs_dec m);Intro eq;Rewrite eq;Trivial. +intros P HP p. +set (Q := fun z => 0 <= z -> P z /\ P (- z)) in *. +cut (Q (Zabs p)); [ intros | apply (Z_lt_induction Q); auto with zarith ]. +elim (Zabs_dec p); intro eq; rewrite eq; elim H; auto with zarith. +unfold Q in |- *; clear Q; intros. +split; apply HP. +rewrite Zabs_eq; auto; intros. +elim (H (Zabs m)); intros; auto with zarith. +elim (Zabs_dec m); intro eq; rewrite eq; trivial. +rewrite Zabs_non_eq; auto with zarith. +rewrite Zopp_involutive; intros. +elim (H (Zabs m)); intros; auto with zarith. +elim (Zabs_dec m); intro eq; rewrite eq; trivial. Qed. -V7only [Unset Implicit Arguments.]. (** To do case analysis over the sign of [z] *) -Lemma Zcase_sign : (x:Z)(P:Prop) - (`x=0` -> P) -> - (`x>0` -> P) -> - (`x<0` -> P) -> P. +Lemma Zcase_sign : + forall (n:Z) (P:Prop), (n = 0 -> P) -> (n > 0 -> P) -> (n < 0 -> P) -> P. Proof. -Intros x P Hzero Hpos Hneg. -Induction x. -Apply Hzero; Trivial. -Apply Hpos; Apply POS_gt_ZERO. -Apply Hneg; Apply NEG_lt_ZERO. -Save. - -Lemma sqr_pos : (x:Z)`x*x >= 0`. +intros x P Hzero Hpos Hneg. +induction x as [| p| p]. +apply Hzero; trivial. +apply Hpos; apply Zorder.Zgt_pos_0. +apply Hneg; apply Zorder.Zlt_neg_0. +Qed. + +Lemma sqr_pos : forall n:Z, n * n >= 0. Proof. -Intro x. -Apply (Zcase_sign x `x*x >= 0`). -Intros H; Rewrite H; Omega. -Intros H; Replace `0` with `0*0`. -Apply Zge_Zmult_pos_compat; Omega. -Omega. -Intros H; Replace `0` with `0*0`. -Replace `x*x` with `(-x)*(-x)`. -Apply Zge_Zmult_pos_compat; Omega. -Ring. -Omega. -Save. +intro x. +apply (Zcase_sign x (x * x >= 0)). +intros H; rewrite H; omega. +intros H; replace 0 with (0 * 0). +apply Zmult_ge_compat; omega. +omega. +intros H; replace 0 with (0 * 0). +replace (x * x) with (- x * - x). +apply Zmult_ge_compat; omega. +ring. +omega. +Qed. (**********************************************************************) (** A list length in Z, tail recursive. *) -Require PolyList. +Require Import List. -Fixpoint Zlength_aux [acc: Z; A:Set; l:(list A)] : Z := Cases l of - nil => acc - | (cons _ l) => (Zlength_aux (Zs acc) A l) -end. +Fixpoint Zlength_aux (acc:Z) (A:Set) (l:list A) {struct l} : Z := + match l with + | nil => acc + | _ :: l => Zlength_aux (Zsucc acc) A l + end. -Definition Zlength := (Zlength_aux 0). -Implicits Zlength [1]. +Definition Zlength := Zlength_aux 0. +Implicit Arguments Zlength [A]. Section Zlength_properties. -Variable A:Set. +Variable A : Set. -Implicit Variable Type l:(list A). +Implicit Type l : list A. -Lemma Zlength_correct : (l:(list A))(Zlength l)=(inject_nat (length l)). +Lemma Zlength_correct : forall l, Zlength l = Z_of_nat (length l). Proof. -Assert (l:(list A))(acc:Z)(Zlength_aux acc A l)=acc+(inject_nat (length l)). -Induction l. -Simpl; Auto with zarith. -Intros; Simpl (length (cons a l0)); Rewrite inj_S. -Simpl; Rewrite H; Auto with zarith. -Unfold Zlength; Intros; Rewrite H; Auto. +assert (forall l (acc:Z), Zlength_aux acc A l = acc + Z_of_nat (length l)). +simple induction l. +simpl in |- *; auto with zarith. +intros; simpl (length (a :: l0)) in |- *; rewrite Znat.inj_S. +simpl in |- *; rewrite H; auto with zarith. +unfold Zlength in |- *; intros; rewrite H; auto. Qed. -Lemma Zlength_nil : (Zlength 1!A (nil A))=0. +Lemma Zlength_nil : Zlength (A:=A) nil = 0. Proof. -Auto. +auto. Qed. -Lemma Zlength_cons : (x:A)(l:(list A))(Zlength (cons x l))=(Zs (Zlength l)). +Lemma Zlength_cons : forall (x:A) l, Zlength (x :: l) = Zsucc (Zlength l). Proof. -Intros; Do 2 Rewrite Zlength_correct. -Simpl (length (cons x l)); Rewrite inj_S; Auto. +intros; do 2 rewrite Zlength_correct. +simpl (length (x :: l)) in |- *; rewrite Znat.inj_S; auto. Qed. -Lemma Zlength_nil_inv : (l:(list A))(Zlength l)=0 -> l=(nil ?). +Lemma Zlength_nil_inv : forall l, Zlength l = 0 -> l = nil. Proof. -Intro l; Rewrite Zlength_correct. -Case l; Auto. -Intros x l'; Simpl (length (cons x l')). -Rewrite inj_S. -Intros; ElimType False; Generalize (ZERO_le_inj (length l')); Omega. +intro l; rewrite Zlength_correct. +case l; auto. +intros x l'; simpl (length (x :: l')) in |- *. +rewrite Znat.inj_S. +intros; elimtype False; generalize (Zle_0_nat (length l')); omega. Qed. End Zlength_properties. -Implicits Zlength_correct [1]. -Implicits Zlength_cons [1]. -Implicits Zlength_nil_inv [1]. +Implicit Arguments Zlength_correct [A]. +Implicit Arguments Zlength_cons [A]. +Implicit Arguments Zlength_nil_inv [A]. \ No newline at end of file diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v index ee6987215c..7738e868cd 100644 --- a/theories/ZArith/Zdiv.v +++ b/theories/ZArith/Zdiv.v @@ -20,11 +20,10 @@ Then only after proves the main required property. *) Require Export ZArith_base. -Require Zbool. -Require Omega. -Require ZArithRing. -Require Zcomplements. -V7only [Import Z_scope.]. +Require Import Zbool. +Require Import Omega. +Require Import ZArithRing. +Require Import Zcomplements. Open Local Scope Z_scope. (** @@ -37,16 +36,19 @@ Open Local Scope Z_scope. *) -Fixpoint Zdiv_eucl_POS [a:positive] : Z -> Z*Z := [b:Z] - Cases a of - | xH => if `(Zge_bool b 2)` then `(0,1)` else `(1,0)` - | (xO a') => - let (q,r) = (Zdiv_eucl_POS a' b) in - [r':=`2*r`] if `(Zgt_bool b r')` then `(2*q,r')` else `(2*q+1,r'-b)` - | (xI a') => - let (q,r) = (Zdiv_eucl_POS a' b) in - [r':=`2*r+1`] if `(Zgt_bool b r')` then `(2*q,r')` else `(2*q+1,r'-b)` - end. +Fixpoint Zdiv_eucl_POS (a:positive) (b:Z) {struct a} : + Z * Z := + match a with + | xH => if Zge_bool b 2 then (0, 1) else (1, 0) + | xO a' => + let (q, r) := Zdiv_eucl_POS a' b in + let r' := 2 * r in + if Zgt_bool b r' then (2 * q, r') else (2 * q + 1, r' - b) + | xI a' => + let (q, r) := Zdiv_eucl_POS a' b in + let r' := 2 * r + 1 in + if Zgt_bool b r' then (2 * q, r') else (2 * q + 1, r' - b) + end. (** @@ -78,33 +80,32 @@ Fixpoint Zdiv_eucl_POS [a:positive] : Z -> Z*Z := [b:Z] *) -Definition Zdiv_eucl [a,b:Z] : Z*Z := - Cases a b of - | ZERO _ => `(0,0)` - | _ ZERO => `(0,0)` - | (POS a') (POS _) => (Zdiv_eucl_POS a' b) - | (NEG a') (POS _) => - let (q,r) = (Zdiv_eucl_POS a' b) in - Cases r of - | ZERO => `(-q,0)` - | _ => `(-(q+1),b-r)` +Definition Zdiv_eucl (a b:Z) : Z * Z := + match a, b with + | Z0, _ => (0, 0) + | _, Z0 => (0, 0) + | Zpos a', Zpos _ => Zdiv_eucl_POS a' b + | Zneg a', Zpos _ => + let (q, r) := Zdiv_eucl_POS a' b in + match r with + | Z0 => (- q, 0) + | _ => (- (q + 1), b - r) end - | (NEG a') (NEG b') => - let (q,r) = (Zdiv_eucl_POS a' (POS b')) in `(q,-r)` - | (POS a') (NEG b') => - let (q,r) = (Zdiv_eucl_POS a' (POS b')) in - Cases r of - | ZERO => `(-q,0)` - | _ => `(-(q+1),b+r)` + | Zneg a', Zneg b' => let (q, r) := Zdiv_eucl_POS a' (Zpos b') in (q, - r) + | Zpos a', Zneg b' => + let (q, r) := Zdiv_eucl_POS a' (Zpos b') in + match r with + | Z0 => (- q, 0) + | _ => (- (q + 1), b + r) end - end. + end. (** Division and modulo are projections of [Zdiv_eucl] *) -Definition Zdiv [a,b:Z] : Z := let (q,_) = (Zdiv_eucl a b) in q. +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. +Definition Zmod (a b:Z) : Z := let (_, r) := Zdiv_eucl a b in r. (* Tests: @@ -127,306 +128,296 @@ Eval Compute in `(Zdiv_eucl (-7) (-3))`. *) -Lemma Z_div_mod_POS : (b:Z)`b > 0` -> (a:positive) - let (q,r)=(Zdiv_eucl_POS a b) in `(POS a) = b*q + r`/\`0<=r 0 -> + forall a:positive, + let (q, r) := Zdiv_eucl_POS a b in Zpos a = b * q + r /\ 0 <= r < b. Proof. -Induction a; Unfold Zdiv_eucl_POS; Fold Zdiv_eucl_POS. - -Intro p; Case (Zdiv_eucl_POS p b); Intros q r (H0,H1). -Generalize (Zgt_cases b `2*r+1`). -Case (Zgt_bool b `2*r+1`); -(Rewrite POS_xI; Rewrite H0; Split ; [ Ring | Omega ]). - -Intros p; Case (Zdiv_eucl_POS p b); Intros q r (H0,H1). -Generalize (Zgt_cases b `2*r`). -Case (Zgt_bool b `2*r`); - Rewrite POS_xO; Change (POS (xO p)) with `2*(POS p)`; - Rewrite H0; (Split; [Ring | Omega]). - -Generalize (Zge_cases b `2`). -Case (Zge_bool b `2`); (Intros; Split; [Ring | Omega ]). -Omega. +simple induction a; unfold Zdiv_eucl_POS in |- *; fold Zdiv_eucl_POS in |- *. + +intro p; case (Zdiv_eucl_POS p b); intros q r [H0 H1]. +generalize (Zgt_cases b (2 * r + 1)). +case (Zgt_bool b (2 * r + 1)); + (rewrite BinInt.Zpos_xI; rewrite H0; split; [ ring | omega ]). + +intros p; case (Zdiv_eucl_POS p b); intros q r [H0 H1]. +generalize (Zgt_cases b (2 * r)). +case (Zgt_bool b (2 * r)); rewrite BinInt.Zpos_xO; + change (Zpos (xO p)) with (2 * Zpos p) in |- *; rewrite H0; + (split; [ ring | omega ]). + +generalize (Zge_cases b 2). +case (Zge_bool b 2); (intros; split; [ ring | omega ]). +omega. Qed. -Theorem Z_div_mod : (a,b:Z)`b > 0` -> - let (q,r) = (Zdiv_eucl a b) in `a = b*q + r` /\ `0<=r 0 -> let (q, r) := Zdiv_eucl a b in a = b * q + r /\ 0 <= r < b. Proof. -Intros a b; Case a; Case b; Try (Simpl; Intros; Omega). -Unfold Zdiv_eucl; Intros; Apply Z_div_mod_POS; Trivial. +intros a b; case a; case b; try (simpl in |- *; intros; omega). +unfold Zdiv_eucl in |- *; intros; apply Z_div_mod_POS; trivial. -Intros; Discriminate. +intros; discriminate. -Intros. -Generalize (Z_div_mod_POS (POS p) H p0). -Unfold Zdiv_eucl. -Case (Zdiv_eucl_POS p0 (POS p)). -Intros z z0. -Case z0. +intros. +generalize (Z_div_mod_POS (Zpos p) H p0). +unfold Zdiv_eucl in |- *. +case (Zdiv_eucl_POS p0 (Zpos p)). +intros z z0. +case z0. -Intros [H1 H2]. -Split; Trivial. -Replace (NEG p0) with `-(POS p0)`; [ Rewrite H1; Ring | Trivial ]. +intros [H1 H2]. +split; trivial. +replace (Zneg p0) with (- Zpos p0); [ rewrite H1; ring | trivial ]. -Intros p1 [H1 H2]. -Split; Trivial. -Replace (NEG p0) with `-(POS p0)`; [ Rewrite H1; Ring | Trivial ]. -Generalize (POS_gt_ZERO p1); Omega. +intros p1 [H1 H2]. +split; trivial. +replace (Zneg p0) with (- Zpos p0); [ rewrite H1; ring | trivial ]. +generalize (Zorder.Zgt_pos_0 p1); omega. -Intros p1 [H1 H2]. -Split; Trivial. -Replace (NEG p0) with `-(POS p0)`; [ Rewrite H1; Ring | Trivial ]. -Generalize (NEG_lt_ZERO p1); Omega. +intros p1 [H1 H2]. +split; trivial. +replace (Zneg p0) with (- Zpos p0); [ rewrite H1; ring | trivial ]. +generalize (Zorder.Zlt_neg_0 p1); omega. -Intros; Discriminate. +intros; discriminate. Qed. (** Existence theorems *) -Theorem Zdiv_eucl_exist : (b:Z)`b > 0` -> (a:Z) - { qr:Z*Z | let (q,r)=qr in `a=b*q+r` /\ `0 <= r < b` }. +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}. Proof. -Intros b Hb a. -Exists (Zdiv_eucl a b). -Exact (Z_div_mod a b Hb). +intros b Hb a. +exists (Zdiv_eucl a b). +exact (Z_div_mod a b Hb). Qed. -Implicits Zdiv_eucl_exist. +Implicit Arguments Zdiv_eucl_exist. -Theorem Zdiv_eucl_extended : (b:Z)`b <> 0` -> (a:Z) - { qr:Z*Z | let (q,r)=qr in `a=b*q+r` /\ `0 <= r < |b|` }. +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 (pair ? ? `-q` r). -Elim Hqr;Intros. -Split. -Rewrite <- Zmult_Zopp_left;Assumption. -Rewrite Zabs_non_eq;[Assumption|Omega]. +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. -Implicits Zdiv_eucl_extended. +Implicit Arguments Zdiv_eucl_extended. (** Auxiliary lemmas about [Zdiv] and [Zmod] *) -Lemma Z_div_mod_eq : (a,b:Z)`b > 0` -> `a = b * (Zdiv a b) + (Zmod a b)`. +Lemma Z_div_mod_eq : forall a b:Z, b > 0 -> a = b * Zdiv a b + Zmod a b. Proof. -Unfold Zdiv Zmod. -Intros a b Hb. -Generalize (Z_div_mod a b Hb). -Case (Zdiv_eucl); Tauto. -Save. +unfold Zdiv, Zmod in |- *. +intros a b Hb. +generalize (Z_div_mod a b Hb). +case Zdiv_eucl; tauto. +Qed. -Lemma Z_mod_lt : (a,b:Z)`b > 0` -> `0 <= (Zmod a b) < b`. +Lemma Z_mod_lt : forall a b:Z, b > 0 -> 0 <= Zmod a b < b. Proof. -Unfold Zmod. -Intros a b Hb. -Generalize (Z_div_mod a b Hb). -Case (Zdiv_eucl a b); Tauto. -Save. - -Lemma Z_div_POS_ge0 : (b:Z)(a:positive) - let (q,_) = (Zdiv_eucl_POS a b) in `q >= 0`. +unfold Zmod in |- *. +intros a b Hb. +generalize (Z_div_mod a b Hb). +case (Zdiv_eucl a b); tauto. +Qed. + +Lemma Z_div_POS_ge0 : + forall (b:Z) (a:positive), let (q, _) := Zdiv_eucl_POS a b in q >= 0. Proof. -Induction a; Unfold Zdiv_eucl_POS; Fold Zdiv_eucl_POS. -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; Omega. -Save. - -Lemma Z_div_ge0 : (a,b:Z)`b > 0` -> `a >= 0` -> `(Zdiv a b) >= 0`. +simple induction a; unfold Zdiv_eucl_POS in |- *; fold Zdiv_eucl_POS in |- *. +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. +Qed. + +Lemma Z_div_ge0 : forall a b:Z, b > 0 -> a >= 0 -> Zdiv a b >= 0. Proof. -Intros a b Hb; Unfold Zdiv Zdiv_eucl; Case a; Simpl; Intros. -Case b; Simpl; Trivial. -Generalize Hb; Case b; Try Trivial. -Auto with zarith. -Intros p0 Hp0; Generalize (Z_div_POS_ge0 (POS p0) p). -Case (Zdiv_eucl_POS p (POS p0)); Simpl; Tauto. -Intros; Discriminate. -Elim H; Trivial. -Save. - -Lemma Z_div_lt : (a,b:Z)`b >= 2` -> `a > 0` -> `(Zdiv a b) < a`. +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. +Qed. + +Lemma Z_div_lt : forall a b:Z, b >= 2 -> a > 0 -> Zdiv a b < a. Proof. -Intros. Cut `b > 0`; [Intro Hb | Omega]. -Generalize (Z_div_mod a b Hb). -Cut `a >= 0`; [Intro Ha | Omega]. -Generalize (Z_div_ge0 a b Hb Ha). -Unfold Zdiv; Case (Zdiv_eucl a b); Intros q r H1 [H2 H3]. -Cut `a >= 2*q` -> `q < a`; [ Intro h; Apply h; Clear h | Intros; Omega ]. -Apply Zge_trans with `b*q`. -Omega. -Auto with zarith. -Save. +intros. cut (b > 0); [ intro Hb | omega ]. +generalize (Z_div_mod a b Hb). +cut (a >= 0); [ intro Ha | omega ]. +generalize (Z_div_ge0 a b Hb Ha). +unfold Zdiv in |- *; case (Zdiv_eucl a b); intros q r H1 [H2 H3]. +cut (a >= 2 * q -> q < a); [ intro h; apply h; clear h | intros; omega ]. +apply Zge_trans with (b * q). +omega. +auto with zarith. +Qed. (** Syntax *) -V7only[ -Grammar znatural expr2 : constr := - expr_div [ expr2($p) "/" expr2($c) ] -> [ (Zdiv $p $c) ] -| expr_mod [ expr2($p) "%" expr2($c) ] -> [ (Zmod $p $c) ] -. - -Syntax constr - level 6: - Zdiv [ (Zdiv $n1 $n2) ] - -> [ [ "`"(ZEXPR $n1):E "/" [0 0] (ZEXPR $n2):L "`"] ] - | Zmod [ (Zmod $n1 $n2) ] - -> [ [ "`"(ZEXPR $n1):E "%" [0 0] (ZEXPR $n2):L "`"] ] - | Zdiv_inside - [ << (ZEXPR <<(Zdiv $n1 $n2)>>) >> ] - -> [ (ZEXPR $n1):E "/" [0 0] (ZEXPR $n2):L ] - | Zmod_inside - [ << (ZEXPR <<(Zmod $n1 $n2)>>) >> ] - -> [ (ZEXPR $n1):E " %" [1 0] (ZEXPR $n2):L ] -. -]. - - -Infix 3 "/" Zdiv (no associativity) : Z_scope V8only. -Infix 3 "mod" Zmod (no associativity) : Z_scope. + + +Infix "/" := Zdiv : Z_scope. +Infix "mod" := Zmod (at level 40, no associativity) : Z_scope. (** Other lemmas (now using the syntax for [Zdiv] and [Zmod]). *) -Lemma Z_div_ge : (a,b,c:Z)`c > 0`->`a >= b`->`a/c >= b/c`. +Lemma Z_div_ge : forall a b c:Z, c > 0 -> a >= b -> a / c >= b / c. Proof. -Intros a b c cPos aGeb. -Generalize (Z_div_mod_eq a c cPos). -Generalize (Z_mod_lt a c cPos). -Generalize (Z_div_mod_eq b c cPos). -Generalize (Z_mod_lt b c cPos). -Intros. -Elim (Z_ge_lt_dec `a/c` `b/c`); Trivial. -Intro. -Absurd `b-a >= 1`. -Omega. -Rewrite -> H0. -Rewrite -> H2. -Assert `c*(b/c)+b % c-(c*(a/c)+a % c) = c*(b/c - a/c) + b % c - a % c`. -Ring. -Rewrite H3. -Assert `c*(b/c-a/c) >= c*1`. -Apply Zge_Zmult_pos_left. -Omega. -Omega. -Assert `c*1=c`. -Ring. -Omega. -Save. - -Lemma Z_mod_plus : (a,b,c:Z)`c > 0`->`(a+b*c) % c = a % c`. +intros a b c cPos aGeb. +generalize (Z_div_mod_eq a c cPos). +generalize (Z_mod_lt a c cPos). +generalize (Z_div_mod_eq b c cPos). +generalize (Z_mod_lt b c cPos). +intros. +elim (Z_ge_lt_dec (a / c) (b / c)); trivial. +intro. +absurd (b - a >= 1). +omega. +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. +assert (c * (b / c - a / c) >= c * 1). +apply Zmult_ge_compat_l. +omega. +omega. +assert (c * 1 = c). +ring. +omega. +Qed. + +Lemma Z_mod_plus : forall a b c:Z, c > 0 -> (a + b * c) mod c = a mod c. 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. - -Assert `(a+b*c) % c - a % c = c*(b+a/c - (a+b*c)/c)`. -Replace `(a+b*c) % c` with `a+b*c - c*((a+b*c)/c)`. -Replace `a % c` with `a - c*(a/c)`. -Ring. -Omega. -Omega. -LetTac q := `b+a/c-(a+b*c)/c`. -Apply (Zcase_sign q); Intros. -Assert `c*q=0`. -Rewrite H4; Ring. -Rewrite H5 in H3. -Omega. - -Assert `c*q >= c`. -Pattern 2 c; Replace c with `c*1`. -Apply Zge_Zmult_pos_left; Omega. -Ring. -Omega. - -Assert `c*q <= -c`. -Replace `-c` with `c*(-1)`. -Apply Zle_Zmult_pos_left; Omega. -Ring. -Omega. -Save. - -Lemma Z_div_plus : (a,b,c:Z)`c > 0`->`(a+b*c)/c = a/c+b`. +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. + +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. + +assert (c * q >= c). +pattern c at 2 in |- *; replace c with (c * 1). +apply Zmult_ge_compat_l; omega. +ring. +omega. + +assert (c * q <= - c). +replace (- c) with (c * -1). +apply Zmult_le_compat_l; omega. +ring. +omega. +Qed. + +Lemma Z_div_plus : forall a b c:Z, c > 0 -> (a + b * c) / c = a / c + 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_left with c. Omega. -Replace `c*((a+b*c)/c)` with `a+b*c-(a+b*c) % c`. -Rewrite (Z_mod_plus a b c cPos). -Pattern 1 a; Rewrite H2. -Ring. -Pattern 1 `a+b*c`; Rewrite H0. -Ring. -Save. - -Lemma Z_div_mult : (a,b:Z)`b > 0`->`(a*b)/b = a`. -Intros; Replace `a*b` with `0+a*b`; Auto. -Rewrite Z_div_plus; Auto. -Save. - -Lemma Z_mult_div_ge : (a,b:Z)`b>0`->`b*(a/b) <= a`. +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. +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. +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 2 a; Rewrite H. -Omega. -Save. - -Lemma Z_mod_same : (a:Z)`a>0`->`a % a=0`. +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. +Qed. + +Lemma Z_mod_same : forall a:Z, a > 0 -> a mod a = 0. Proof. -Intros a aPos. -Generalize (Z_mod_plus `0` `1` a aPos). -Replace `0+1*a` with `a`. -Intros. -Rewrite H. -Compute. -Trivial. -Ring. -Save. - -Lemma Z_div_same : (a:Z)`a>0`->`a/a=1`. +intros a aPos. +generalize (Z_mod_plus 0 1 a aPos). +replace (0 + 1 * a) with a. +intros. +rewrite H. +compute in |- *. +trivial. +ring. +Qed. + +Lemma Z_div_same : forall a:Z, a > 0 -> a / a = 1. Proof. -Intros a aPos. -Generalize (Z_div_plus `0` `1` a aPos). -Replace `0+1*a` with `a`. -Intros. -Rewrite H. -Compute. -Trivial. -Ring. -Save. - -Lemma Z_div_exact_1 : (a,b:Z)`b>0` -> `a = b*(a/b)` -> `a%b = 0`. -Intros a b Hb; Generalize (Z_div_mod a b Hb); Unfold Zmod Zdiv. -Case (Zdiv_eucl a b); Intros q r; Omega. -Save. - -Lemma Z_div_exact_2 : (a,b:Z)`b>0` -> `a%b = 0` -> `a = b*(a/b)`. -Intros a b Hb; Generalize (Z_div_mod a b Hb); Unfold Zmod Zdiv. -Case (Zdiv_eucl a b); Intros q r; Omega. -Save. - -Lemma Z_mod_zero_opp : (a,b:Z)`b>0` -> `a%b = 0` -> `(-a)%b = 0`. -Intros a b Hb. -Intros. -Rewrite Z_div_exact_2 with a b; Auto. -Replace `-(b*(a/b))` with `0+(-(a/b))*b`. -Rewrite Z_mod_plus; Auto. -Ring. -Save. +intros a aPos. +generalize (Z_div_plus 0 1 a aPos). +replace (0 + 1 * a) with a. +intros. +rewrite H. +compute in |- *. +trivial. +ring. +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. +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. +Qed. + +Lemma Z_mod_zero_opp : forall a b:Z, b > 0 -> a mod b = 0 -> - a mod b = 0. +intros a b Hb. +intros. +rewrite Z_div_exact_2 with a b; auto. +replace (- (b * (a / b))) with (0 + - (a / b) * b). +rewrite Z_mod_plus; auto. +ring. +Qed. diff --git a/theories/ZArith/Zeven.v b/theories/ZArith/Zeven.v index e22dc20f65..728e16da9e 100644 --- a/theories/ZArith/Zeven.v +++ b/theories/ZArith/Zeven.v @@ -8,8 +8,7 @@ (*i $Id$ i*) -Require BinInt. -Require Zsyntax. +Require Import BinInt. (**********************************************************************) (** About parity: even and odd predicates on Z, division by 2 on Z *) @@ -17,168 +16,189 @@ Require Zsyntax. (**********************************************************************) (** [Zeven], [Zodd], [Zdiv2] and their related properties *) -Definition Zeven := - [z:Z]Cases z of ZERO => True - | (POS (xO _)) => True - | (NEG (xO _)) => True - | _ => False - end. - -Definition Zodd := - [z:Z]Cases z of (POS xH) => True - | (NEG xH) => True - | (POS (xI _)) => True - | (NEG (xI _)) => True - | _ => False - end. - -Definition Zeven_bool := - [z:Z]Cases z of ZERO => true - | (POS (xO _)) => true - | (NEG (xO _)) => true - | _ => false - end. - -Definition Zodd_bool := - [z:Z]Cases z of ZERO => false - | (POS (xO _)) => false - | (NEG (xO _)) => false - | _ => true - end. - -Definition Zeven_odd_dec : (z:Z) { (Zeven z) }+{ (Zodd z) }. +Definition Zeven (z:Z) := + match z with + | Z0 => True + | Zpos (xO _) => True + | Zneg (xO _) => True + | _ => False + end. + +Definition Zodd (z:Z) := + match z with + | Zpos xH => True + | Zneg xH => True + | Zpos (xI _) => True + | Zneg (xI _) => True + | _ => False + end. + +Definition Zeven_bool (z:Z) := + match z with + | Z0 => true + | Zpos (xO _) => true + | Zneg (xO _) => true + | _ => false + end. + +Definition Zodd_bool (z:Z) := + match z with + | Z0 => false + | Zpos (xO _) => false + | Zneg (xO _) => false + | _ => true + end. + +Definition Zeven_odd_dec : forall z:Z, {Zeven z} + {Zodd z}. Proof. - Intro z. Case z; - [ Left; Compute; Trivial - | Intro p; Case p; Intros; - (Right; Compute; Exact I) Orelse (Left; Compute; Exact I) - | Intro p; Case p; Intros; - (Right; Compute; Exact I) Orelse (Left; Compute; Exact I) ]. + intro z. case z; + [ left; compute in |- *; trivial + | intro p; case p; intros; + (right; compute in |- *; exact I) || (left; compute in |- *; exact I) + | intro p; case p; intros; + (right; compute in |- *; exact I) || (left; compute in |- *; exact I) ]. Defined. -Definition Zeven_dec : (z:Z) { (Zeven z) }+{ ~(Zeven z) }. +Definition Zeven_dec : forall z:Z, {Zeven z} + {~ Zeven z}. Proof. - Intro z. Case z; - [ Left; Compute; Trivial - | Intro p; Case p; Intros; - (Left; Compute; Exact I) Orelse (Right; Compute; Trivial) - | Intro p; Case p; Intros; - (Left; Compute; Exact I) Orelse (Right; Compute; Trivial) ]. + intro z. case z; + [ left; compute in |- *; trivial + | intro p; case p; intros; + (left; compute in |- *; exact I) || (right; compute in |- *; trivial) + | intro p; case p; intros; + (left; compute in |- *; exact I) || (right; compute in |- *; trivial) ]. Defined. -Definition Zodd_dec : (z:Z) { (Zodd z) }+{ ~(Zodd z) }. +Definition Zodd_dec : forall z:Z, {Zodd z} + {~ Zodd z}. Proof. - Intro z. Case z; - [ Right; Compute; Trivial - | Intro p; Case p; Intros; - (Left; Compute; Exact I) Orelse (Right; Compute; Trivial) - | Intro p; Case p; Intros; - (Left; Compute; Exact I) Orelse (Right; Compute; Trivial) ]. + intro z. case z; + [ right; compute in |- *; trivial + | intro p; case p; intros; + (left; compute in |- *; exact I) || (right; compute in |- *; trivial) + | intro p; case p; intros; + (left; compute in |- *; exact I) || (right; compute in |- *; trivial) ]. Defined. -Lemma Zeven_not_Zodd : (z:Z)(Zeven z) -> ~(Zodd z). +Lemma Zeven_not_Zodd : forall n:Z, Zeven n -> ~ Zodd n. Proof. - Intro z; NewDestruct z; [ Idtac | NewDestruct p | NewDestruct p ]; Compute; Trivial. + intro z; destruct z; [ idtac | destruct p | destruct p ]; compute in |- *; + trivial. Qed. -Lemma Zodd_not_Zeven : (z:Z)(Zodd z) -> ~(Zeven z). +Lemma Zodd_not_Zeven : forall n:Z, Zodd n -> ~ Zeven n. Proof. - Intro z; NewDestruct z; [ Idtac | NewDestruct p | NewDestruct p ]; Compute; Trivial. + intro z; destruct z; [ idtac | destruct p | destruct p ]; compute in |- *; + trivial. Qed. -Lemma Zeven_Sn : (z:Z)(Zodd z) -> (Zeven (Zs z)). +Lemma Zeven_Sn : forall n:Z, Zodd n -> Zeven (Zsucc n). Proof. - Intro z; NewDestruct z; Unfold Zs; [ Idtac | NewDestruct p | NewDestruct p ]; Simpl; Trivial. - Unfold double_moins_un; Case p; Simpl; Auto. + intro z; destruct z; unfold Zsucc in |- *; + [ idtac | destruct p | destruct p ]; simpl in |- *; + trivial. + unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto. Qed. -Lemma Zodd_Sn : (z:Z)(Zeven z) -> (Zodd (Zs z)). +Lemma Zodd_Sn : forall n:Z, Zeven n -> Zodd (Zsucc n). Proof. - Intro z; NewDestruct z; Unfold Zs; [ Idtac | NewDestruct p | NewDestruct p ]; Simpl; Trivial. - Unfold double_moins_un; Case p; Simpl; Auto. + intro z; destruct z; unfold Zsucc in |- *; + [ idtac | destruct p | destruct p ]; simpl in |- *; + trivial. + unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto. Qed. -Lemma Zeven_pred : (z:Z)(Zodd z) -> (Zeven (Zpred z)). +Lemma Zeven_pred : forall n:Z, Zodd n -> Zeven (Zpred n). Proof. - Intro z; NewDestruct z; Unfold Zpred; [ Idtac | NewDestruct p | NewDestruct p ]; Simpl; Trivial. - Unfold double_moins_un; Case p; Simpl; Auto. + intro z; destruct z; unfold Zpred in |- *; + [ idtac | destruct p | destruct p ]; simpl in |- *; + trivial. + unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto. Qed. -Lemma Zodd_pred : (z:Z)(Zeven z) -> (Zodd (Zpred z)). +Lemma Zodd_pred : forall n:Z, Zeven n -> Zodd (Zpred n). Proof. - Intro z; NewDestruct z; Unfold Zpred; [ Idtac | NewDestruct p | NewDestruct p ]; Simpl; Trivial. - Unfold double_moins_un; Case p; Simpl; Auto. + intro z; destruct z; unfold Zpred in |- *; + [ idtac | destruct p | destruct p ]; simpl in |- *; + trivial. + unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto. Qed. -Hints Unfold Zeven Zodd : zarith. +Hint Unfold Zeven Zodd: zarith. (**********************************************************************) (** [Zdiv2] is defined on all [Z], but notice that for odd negative integers it is not the euclidean quotient: in that case we have [n = 2*(n/2)-1] *) -Definition Zdiv2 := - [z:Z]Cases z of ZERO => ZERO - | (POS xH) => ZERO - | (POS p) => (POS (Zdiv2_pos p)) - | (NEG xH) => ZERO - | (NEG p) => (NEG (Zdiv2_pos p)) - end. +Definition Zdiv2 (z:Z) := + match z with + | Z0 => 0%Z + | Zpos xH => 0%Z + | Zpos p => Zpos (Pdiv2 p) + | Zneg xH => 0%Z + | Zneg p => Zneg (Pdiv2 p) + end. -Lemma Zeven_div2 : (x:Z) (Zeven x) -> `x = 2*(Zdiv2 x)`. +Lemma Zeven_div2 : forall n:Z, Zeven n -> n = (2 * Zdiv2 n)%Z. Proof. -Intro x; NewDestruct x. -Auto with arith. -NewDestruct p; Auto with arith. -Intros. Absurd (Zeven (POS (xI p))); Red; Auto with arith. -Intros. Absurd (Zeven `1`); Red; Auto with arith. -NewDestruct p; Auto with arith. -Intros. Absurd (Zeven (NEG (xI p))); Red; Auto with arith. -Intros. Absurd (Zeven `-1`); Red; Auto with arith. +intro x; destruct x. +auto with arith. +destruct p; auto with arith. +intros. absurd (Zeven (Zpos (xI p))); red in |- *; auto with arith. +intros. absurd (Zeven 1); red in |- *; auto with arith. +destruct p; auto with arith. +intros. absurd (Zeven (Zneg (xI p))); red in |- *; auto with arith. +intros. absurd (Zeven (-1)); red in |- *; auto with arith. Qed. -Lemma Zodd_div2 : (x:Z) `x >= 0` -> (Zodd x) -> `x = 2*(Zdiv2 x)+1`. +Lemma Zodd_div2 : forall n:Z, (n >= 0)%Z -> Zodd n -> n = (2 * Zdiv2 n + 1)%Z. Proof. -Intro x; NewDestruct x. -Intros. Absurd (Zodd `0`); Red; Auto with arith. -NewDestruct p; Auto with arith. -Intros. Absurd (Zodd (POS (xO p))); Red; Auto with arith. -Intros. Absurd `(NEG p) >= 0`; Red; Auto with arith. +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. Qed. -Lemma Zodd_div2_neg : (x:Z) `x <= 0` -> (Zodd x) -> `x = 2*(Zdiv2 x)-1`. +Lemma Zodd_div2_neg : + forall n:Z, (n <= 0)%Z -> Zodd n -> n = (2 * Zdiv2 n - 1)%Z. Proof. -Intro x; NewDestruct x. -Intros. Absurd (Zodd `0`); Red; Auto with arith. -Intros. Absurd `(NEG p) >= 0`; Red; Auto with arith. -NewDestruct p; Auto with arith. -Intros. Absurd (Zodd (NEG (xO p))); Red; Auto with arith. +intro x; destruct x. +intros. absurd (Zodd 0); red in |- *; auto with arith. +intros. absurd (Zneg p >= 0)%Z; 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 : (x:Z) { y:Z | `x=2*y` }+{ y:Z | `x=2*y+1` }. +Lemma Z_modulo_2 : + forall n:Z, {y : Z | n = (2 * y)%Z} + {y : Z | n = (2 * y + 1)%Z}. Proof. -Intros x. -Elim (Zeven_odd_dec x); Intro. -Left. Split with (Zdiv2 x). Exact (Zeven_div2 x a). -Right. Generalize b; Clear b; Case x. -Intro b; Inversion b. -Intro p; Split with (Zdiv2 (POS p)). Apply (Zodd_div2 (POS p)); Trivial. -Unfold Zge Zcompare; Simpl; Discriminate. -Intro p; Split with (Zdiv2 (Zpred (NEG p))). -Pattern 1 (NEG p); Rewrite (Zs_pred (NEG p)). -Pattern 1 (Zpred (NEG p)); Rewrite (Zeven_div2 (Zpred (NEG p))). -Reflexivity. -Apply Zeven_pred; Assumption. +intros x. +elim (Zeven_odd_dec x); intro. +left. split with (Zdiv2 x). exact (Zeven_div2 x a). +right. generalize b; clear b; case x. +intro b; inversion b. +intro p; split with (Zdiv2 (Zpos p)). apply (Zodd_div2 (Zpos p)); trivial. +unfold Zge, Zcompare in |- *; simpl in |- *; discriminate. +intro p; split with (Zdiv2 (Zpred (Zneg p))). +pattern (Zneg p) at 1 in |- *; rewrite (Zsucc_pred (Zneg p)). +pattern (Zpred (Zneg p)) at 1 in |- *; rewrite (Zeven_div2 (Zpred (Zneg p))). +reflexivity. +apply Zeven_pred; assumption. Qed. -Lemma Zsplit2 : (x:Z) { p : Z*Z | let (x1,x2)=p in (`x=x1+x2` /\ (x1=x2 \/ `x2=x1+1`)) }. +Lemma Zsplit2 : + forall n:Z, + {p : Z * Z | + let (x1, x2) := p in n = (x1 + x2)%Z /\ (x1 = x2 \/ x2 = (x1 + 1)%Z)}. Proof. -Intros x. -Elim (Z_modulo_2 x); Intros (y,Hy); Rewrite Zmult_sym in Hy; Rewrite <- Zplus_Zmult_2 in Hy. -Exists (y,y); Split. -Assumption. -Left; Reflexivity. -Exists (y,`y+1`); Split. -Rewrite Zplus_assoc; Assumption. -Right; Reflexivity. -Qed. +intros x. +elim (Z_modulo_2 x); intros [y Hy]; rewrite Zmult_comm in Hy; + rewrite <- Zplus_diag_eq_mult_2 in Hy. +exists (y, y); split. +assumption. +left; reflexivity. +exists (y, (y + 1)%Z); split. +rewrite Zplus_assoc; assumption. +right; reflexivity. +Qed. \ No newline at end of file diff --git a/theories/ZArith/Zhints.v b/theories/ZArith/Zhints.v index 6eb668a5ae..5cce66fc59 100644 --- a/theories/ZArith/Zhints.v +++ b/theories/ZArith/Zhints.v @@ -27,81 +27,80 @@ (* Lemmas involving positive and compare are not taken into account *) -Require BinInt. -Require Zorder. -Require Zmin. -Require Zabs. -Require Zcompare. -Require Znat. -Require auxiliary. -Require Zsyntax. -Require Zmisc. -Require Wf_Z. +Require Import BinInt. +Require Import Zorder. +Require Import Zmin. +Require Import Zabs. +Require Import Zcompare. +Require Import Znat. +Require Import auxiliary. +Require Import Zmisc. +Require Import Wf_Z. (**********************************************************************) (* Simplification lemmas *) (* No subgoal or smaller subgoals *) -Hints Resolve +Hint Resolve (* A) Reversible simplification lemmas (no loss of information) *) (* Should clearly declared as hints *) (* Lemmas ending by eq *) - Zeq_S (* :(n,m:Z)`n = m`->`(Zs n) = (Zs m)` *) + Zsucc_eq_compat (* :(n,m:Z)`n = m`->`(Zs n) = (Zs m)` *) (* Lemmas ending by Zgt *) - Zgt_n_S (* :(n,m:Z)`m > n`->`(Zs m) > (Zs n)` *) - Zgt_Sn_n (* :(n:Z)`(Zs n) > n` *) - POS_gt_ZERO (* :(p:positive)`(POS p) > 0` *) - Zgt_reg_l (* :(n,m,p:Z)`n > m`->`p+n > p+m` *) - Zgt_reg_r (* :(n,m,p:Z)`n > m`->`n+p > m+p` *) + Zsucc_gt_compat (* :(n,m:Z)`m > n`->`(Zs m) > (Zs n)` *) + Zgt_succ (* :(n:Z)`(Zs n) > n` *) + Zorder.Zgt_pos_0 (* :(p:positive)`(POS p) > 0` *) + Zplus_gt_compat_l (* :(n,m,p:Z)`n > m`->`p+n > p+m` *) + Zplus_gt_compat_r (* :(n,m,p:Z)`n > m`->`n+p > m+p` *) (* Lemmas ending by Zlt *) - Zlt_n_Sn (* :(n:Z)`n < (Zs n)` *) - Zlt_n_S (* :(n,m:Z)`n < m`->`(Zs n) < (Zs m)` *) - Zlt_pred_n_n (* :(n:Z)`(Zpred n) < n` *) - Zlt_reg_l (* :(n,m,p:Z)`n < m`->`p+n < p+m` *) - Zlt_reg_r (* :(n,m,p:Z)`n < m`->`n+p < m+p` *) + Zlt_succ (* :(n:Z)`n < (Zs n)` *) + Zsucc_lt_compat (* :(n,m:Z)`n < m`->`(Zs n) < (Zs m)` *) + Zlt_pred (* :(n:Z)`(Zpred n) < n` *) + Zplus_lt_compat_l (* :(n,m,p:Z)`n < m`->`p+n < p+m` *) + Zplus_lt_compat_r (* :(n,m,p:Z)`n < m`->`n+p < m+p` *) (* Lemmas ending by Zle *) - ZERO_le_inj (* :(n:nat)`0 <= (inject_nat n)` *) - ZERO_le_POS (* :(p:positive)`0 <= (POS p)` *) - Zle_n (* :(n:Z)`n <= n` *) - Zle_n_Sn (* :(n:Z)`n <= (Zs n)` *) - Zle_n_S (* :(n,m:Z)`m <= n`->`(Zs m) <= (Zs n)` *) - Zle_pred_n (* :(n:Z)`(Zpred n) <= n` *) + Zle_0_nat (* :(n:nat)`0 <= (inject_nat n)` *) + Zorder.Zle_0_pos (* :(p:positive)`0 <= (POS p)` *) + Zle_refl (* :(n:Z)`n <= n` *) + Zle_succ (* :(n:Z)`n <= (Zs n)` *) + Zsucc_le_compat (* :(n,m:Z)`m <= n`->`(Zs m) <= (Zs n)` *) + Zle_pred (* :(n:Z)`(Zpred n) <= n` *) Zle_min_l (* :(n,m:Z)`(Zmin n m) <= n` *) Zle_min_r (* :(n,m:Z)`(Zmin n m) <= m` *) - Zle_reg_l (* :(n,m,p:Z)`n <= m`->`p+n <= p+m` *) - Zle_reg_r (* :(a,b,c:Z)`a <= b`->`a+c <= b+c` *) + Zplus_le_compat_l (* :(n,m,p:Z)`n <= m`->`p+n <= p+m` *) + Zplus_le_compat_r (* :(a,b,c:Z)`a <= b`->`a+c <= b+c` *) Zabs_pos (* :(x:Z)`0 <= |x|` *) (* B) Irreversible simplification lemmas : Probably to be declared as *) (* hints, when no other simplification is possible *) (* Lemmas ending by eq *) - Z_eq_mult (* :(x,y:Z)`y = 0`->`y*x = 0` *) - Zplus_simpl (* :(n,m,p,q:Z)`n = m`->`p = q`->`n+p = m+q` *) + BinInt.Z_eq_mult (* :(x,y:Z)`y = 0`->`y*x = 0` *) + Zplus_eq_compat (* :(n,m,p,q:Z)`n = m`->`p = q`->`n+p = m+q` *) (* Lemmas ending by Zge *) - Zge_Zmult_pos_right (* :(a,b,c:Z)`a >= b`->`c >= 0`->`a*c >= b*c` *) - Zge_Zmult_pos_left (* :(a,b,c:Z)`a >= b`->`c >= 0`->`c*a >= c*b` *) - Zge_Zmult_pos_compat (* : - (a,b,c,d:Z)`a >= c`->`b >= d`->`c >= 0`->`d >= 0`->`a*b >= c*d` *) + Zorder.Zmult_ge_compat_r (* :(a,b,c:Z)`a >= b`->`c >= 0`->`a*c >= b*c` *) + Zorder.Zmult_ge_compat_l (* :(a,b,c:Z)`a >= b`->`c >= 0`->`c*a >= c*b` *) + Zorder.Zmult_ge_compat (* : + (a,b,c,d:Z)`a >= c`->`b >= d`->`c >= 0`->`d >= 0`->`a*b >= c*d` *) (* Lemmas ending by Zlt *) - Zgt_ZERO_mult (* :(a,b:Z)`a > 0`->`b > 0`->`a*b > 0` *) - Zlt_S (* :(n,m:Z)`n < m`->`n < (Zs m)` *) + Zorder.Zmult_gt_0_compat (* :(a,b:Z)`a > 0`->`b > 0`->`a*b > 0` *) + Zlt_lt_succ (* :(n,m:Z)`n < m`->`n < (Zs m)` *) (* Lemmas ending by Zle *) - Zle_ZERO_mult (* :(x,y:Z)`0 <= x`->`0 <= y`->`0 <= x*y` *) - Zle_Zmult_pos_right (* :(a,b,c:Z)`a <= b`->`0 <= c`->`a*c <= b*c` *) - Zle_Zmult_pos_left (* :(a,b,c:Z)`a <= b`->`0 <= c`->`c*a <= c*b` *) - OMEGA2 (* :(x,y:Z)`0 <= x`->`0 <= y`->`0 <= x+y` *) - Zle_le_S (* :(x,y:Z)`x <= y`->`x <= (Zs y)` *) - Zle_plus_plus (* :(n,m,p,q:Z)`n <= m`->`p <= q`->`n+p <= m+q` *) - -: zarith. + Zorder.Zmult_le_0_compat (* :(x,y:Z)`0 <= x`->`0 <= y`->`0 <= x*y` *) + Zorder.Zmult_le_compat_r (* :(a,b,c:Z)`a <= b`->`0 <= c`->`a*c <= b*c` *) + Zorder.Zmult_le_compat_l (* :(a,b,c:Z)`a <= b`->`0 <= c`->`c*a <= c*b` *) + Zplus_le_0_compat (* :(x,y:Z)`0 <= x`->`0 <= y`->`0 <= x+y` *) + Zle_le_succ (* :(x,y:Z)`x <= y`->`x <= (Zs y)` *) + Zplus_le_compat (* :(n,m,p,q:Z)`n <= m`->`p <= q`->`n+p <= m+q` *) + + : zarith. (**********************************************************************) (* Reversible lemmas relating operators *) @@ -384,4 +383,4 @@ inj_minus2: (x,y:nat)(gt y x)->`(inject_nat (minus x y)) = 0` Zred_factor5: (x,y:Z)`x*0+y = y` *) -(*i*) +(*i*) \ No newline at end of file diff --git a/theories/ZArith/Zlogarithm.v b/theories/ZArith/Zlogarithm.v index 2879fefe8b..ba6d21c4d8 100644 --- a/theories/ZArith/Zlogarithm.v +++ b/theories/ZArith/Zlogarithm.v @@ -20,161 +20,161 @@ - [Log_nearest]: [y= (Log_nearest x) iff 2^(y-1/2) < x <= 2^(y+1/2)] i.e. [Log_nearest x] is the integer nearest from [Log x] *) -Require ZArith_base. -Require Omega. -Require Zcomplements. -Require Zpower. -V7only [Import Z_scope.]. +Require Import ZArith_base. +Require Import Omega. +Require Import Zcomplements. +Require Import Zpower. Open Local Scope Z_scope. Section Log_pos. (* Log of positive integers *) (** First we build [log_inf] and [log_sup] *) -Fixpoint log_inf [p:positive] : Z := - Cases p of - xH => `0` (* 1 *) - | (xO q) => (Zs (log_inf q)) (* 2n *) - | (xI q) => (Zs (log_inf q)) (* 2n+1 *) +Fixpoint log_inf (p:positive) : Z := + match p with + | xH => 0 (* 1 *) + | xO q => Zsucc (log_inf q) (* 2n *) + | xI q => Zsucc (log_inf q) (* 2n+1 *) end. -Fixpoint log_sup [p:positive] : Z := - Cases p of - xH => `0` (* 1 *) - | (xO n) => (Zs (log_sup n)) (* 2n *) - | (xI n) => (Zs (Zs (log_inf n))) (* 2n+1 *) +Fixpoint log_sup (p:positive) : Z := + match p with + | xH => 0 (* 1 *) + | xO n => Zsucc (log_sup n) (* 2n *) + | xI n => Zsucc (Zsucc (log_inf n)) (* 2n+1 *) end. -Hints Unfold log_inf log_sup. +Hint Unfold log_inf log_sup. (** Then we give the specifications of [log_inf] and [log_sup] and prove their validity *) (*i Hints Resolve ZERO_le_S : zarith. i*) -Hints Resolve Zle_trans : zarith. - -Theorem log_inf_correct : (x:positive) ` 0 <= (log_inf x)` /\ - ` (two_p (log_inf x)) <= (POS x) < (two_p (Zs (log_inf x)))`. -Induction x; Intros; Simpl; -[ Elim H; Intros Hp HR; Clear H; Split; - [ Auto with zarith - | Conditional (Apply Zle_le_S; Trivial) Rewrite two_p_S with x:=(Zs (log_inf p)); - Conditional Trivial Rewrite two_p_S; - Conditional Trivial Rewrite two_p_S in HR; - Rewrite (POS_xI p); Omega ] -| Elim H; Intros Hp HR; Clear H; Split; - [ Auto with zarith - | Conditional (Apply Zle_le_S; Trivial) Rewrite two_p_S with x:=(Zs (log_inf p)); - Conditional Trivial Rewrite two_p_S; - Conditional Trivial Rewrite two_p_S in HR; - Rewrite (POS_xO p); Omega ] -| Unfold two_power_pos; Unfold shift_pos; Simpl; Omega -]. +Hint Resolve Zle_trans: zarith. + +Theorem log_inf_correct : + forall x:positive, + 0 <= log_inf x /\ two_p (log_inf x) <= Zpos x < two_p (Zsucc (log_inf x)). +simple induction x; intros; simpl in |- *; + [ elim H; intros Hp HR; clear H; split; + [ auto with zarith + | conditional apply Zle_le_succ; trivial rewrite + two_p_S with (x := Zsucc (log_inf p)); + conditional trivial rewrite two_p_S; + conditional trivial rewrite two_p_S in HR; rewrite (BinInt.Zpos_xI p); + omega ] + | elim H; intros Hp HR; clear H; split; + [ auto with zarith + | conditional apply Zle_le_succ; trivial rewrite + two_p_S with (x := Zsucc (log_inf p)); + conditional trivial rewrite two_p_S; + conditional trivial rewrite two_p_S in HR; rewrite (BinInt.Zpos_xO p); + omega ] + | unfold two_power_pos in |- *; unfold shift_pos in |- *; simpl in |- *; + omega ]. Qed. -Definition log_inf_correct1 := - [p:positive](proj1 ? ? (log_inf_correct p)). -Definition log_inf_correct2 := - [p:positive](proj2 ? ? (log_inf_correct p)). +Definition log_inf_correct1 (p:positive) := proj1 (log_inf_correct p). +Definition log_inf_correct2 (p:positive) := proj2 (log_inf_correct p). Opaque log_inf_correct1 log_inf_correct2. -Hints Resolve log_inf_correct1 log_inf_correct2 : zarith. +Hint Resolve log_inf_correct1 log_inf_correct2: zarith. -Lemma log_sup_correct1 : (p:positive)` 0 <= (log_sup p)`. -Induction p; Intros; Simpl; Auto with zarith. +Lemma log_sup_correct1 : forall p:positive, 0 <= log_sup p. +simple induction p; intros; simpl in |- *; auto with zarith. Qed. (** For every [p], either [p] is a power of two and [(log_inf p)=(log_sup p)] either [(log_sup p)=(log_inf p)+1] *) -Theorem log_sup_log_inf : (p:positive) - IF (POS p)=(two_p (log_inf p)) - then (POS p)=(two_p (log_sup p)) - else ` (log_sup p)=(Zs (log_inf p))`. - -Induction p; Intros; -[ Elim H; Right; Simpl; - Rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0)); - Rewrite POS_xI; Unfold Zs; Omega -| Elim H; Clear H; Intro Hif; - [ Left; Simpl; - Rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0)); - Rewrite (two_p_S (log_sup p0) (log_sup_correct1 p0)); - Rewrite <- (proj1 ? ? Hif); Rewrite <- (proj2 ? ? Hif); - Auto - | Right; Simpl; - Rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0)); - Rewrite POS_xO; Unfold Zs; Omega ] -| Left; Auto ]. +Theorem log_sup_log_inf : + forall p:positive, + IF Zpos p = two_p (log_inf p) then Zpos p = two_p (log_sup p) + else log_sup p = Zsucc (log_inf p). + +simple induction p; intros; + [ elim H; right; simpl in |- *; + rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0)); + rewrite BinInt.Zpos_xI; unfold Zsucc in |- *; omega + | elim H; clear H; intro Hif; + [ left; simpl in |- *; + rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0)); + rewrite (two_p_S (log_sup p0) (log_sup_correct1 p0)); + rewrite <- (proj1 Hif); rewrite <- (proj2 Hif); + auto + | right; simpl in |- *; + rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0)); + rewrite BinInt.Zpos_xO; unfold Zsucc in |- *; + omega ] + | left; auto ]. Qed. -Theorem log_sup_correct2 : (x:positive) - ` (two_p (Zpred (log_sup x))) < (POS x) <= (two_p (log_sup x))`. +Theorem log_sup_correct2 : + forall x:positive, two_p (Zpred (log_sup x)) < Zpos x <= two_p (log_sup x). -Intro. -Elim (log_sup_log_inf x). +intro. +elim (log_sup_log_inf x). (* x is a power of two and [log_sup = log_inf] *) -Intros (E1,E2); Rewrite E2. -Split ; [ Apply two_p_pred; Apply log_sup_correct1 | Apply Zle_n ]. -Intros (E1,E2); Rewrite E2. -Rewrite <- (Zpred_Sn (log_inf x)). -Generalize (log_inf_correct2 x); Omega. +intros [E1 E2]; rewrite E2. +split; [ apply two_p_pred; apply log_sup_correct1 | apply Zle_refl ]. +intros [E1 E2]; rewrite E2. +rewrite <- (Zpred_succ (log_inf x)). +generalize (log_inf_correct2 x); omega. Qed. -Lemma log_inf_le_log_sup : - (p:positive) `(log_inf p) <= (log_sup p)`. -Induction p; Simpl; Intros; Omega. +Lemma log_inf_le_log_sup : forall p:positive, log_inf p <= log_sup p. +simple induction p; simpl in |- *; intros; omega. Qed. -Lemma log_sup_le_Slog_inf : - (p:positive) `(log_sup p) <= (Zs (log_inf p))`. -Induction p; Simpl; Intros; Omega. +Lemma log_sup_le_Slog_inf : forall p:positive, log_sup p <= Zsucc (log_inf p). +simple induction p; simpl in |- *; intros; omega. Qed. (** Now it's possible to specify and build the [Log] rounded to the nearest *) -Fixpoint log_near[x:positive] : Z := - Cases x of - xH => `0` - | (xO xH) => `1` - | (xI xH) => `2` - | (xO y) => (Zs (log_near y)) - | (xI y) => (Zs (log_near y)) +Fixpoint log_near (x:positive) : Z := + match x with + | xH => 0 + | xO xH => 1 + | xI xH => 2 + | xO y => Zsucc (log_near y) + | xI y => Zsucc (log_near y) end. -Theorem log_near_correct1 : (p:positive)` 0 <= (log_near p)`. -Induction p; Simpl; Intros; -[Elim p0; Auto with zarith | Elim p0; Auto with zarith | Trivial with zarith ]. -Intros; Apply Zle_le_S. -Generalize H0; Elim p1; Intros; Simpl; - [ Assumption | Assumption | Apply ZERO_le_POS ]. -Intros; Apply Zle_le_S. -Generalize H0; Elim p1; Intros; Simpl; - [ Assumption | Assumption | Apply ZERO_le_POS ]. +Theorem log_near_correct1 : forall p:positive, 0 <= log_near p. +simple induction p; simpl in |- *; intros; + [ elim p0; auto with zarith + | elim p0; auto with zarith + | trivial with zarith ]. +intros; apply Zle_le_succ. +generalize H0; elim p1; intros; simpl in |- *; + [ assumption | assumption | apply Zorder.Zle_0_pos ]. +intros; apply Zle_le_succ. +generalize H0; elim p1; intros; simpl in |- *; + [ assumption | assumption | apply Zorder.Zle_0_pos ]. Qed. -Theorem log_near_correct2: (p:positive) - (log_near p)=(log_inf p) -\/(log_near p)=(log_sup p). -Induction p. -Intros p0 [Einf|Esup]. -Simpl. Rewrite Einf. -Case p0; [Left | Left | Right]; Reflexivity. -Simpl; Rewrite Esup. -Elim (log_sup_log_inf p0). -Generalize (log_inf_le_log_sup p0). -Generalize (log_sup_le_Slog_inf p0). -Case p0; Auto with zarith. -Intros; Omega. -Case p0; Intros; Auto with zarith. -Intros p0 [Einf|Esup]. -Simpl. -Repeat Rewrite Einf. -Case p0; Intros; Auto with zarith. -Simpl. -Repeat Rewrite Esup. -Case p0; Intros; Auto with zarith. -Auto. +Theorem log_near_correct2 : + forall p:positive, log_near p = log_inf p \/ log_near p = log_sup p. +simple induction p. +intros p0 [Einf| Esup]. +simpl in |- *. rewrite Einf. +case p0; [ left | left | right ]; reflexivity. +simpl in |- *; rewrite Esup. +elim (log_sup_log_inf p0). +generalize (log_inf_le_log_sup p0). +generalize (log_sup_le_Slog_inf p0). +case p0; auto with zarith. +intros; omega. +case p0; intros; auto with zarith. +intros p0 [Einf| Esup]. +simpl in |- *. +repeat rewrite Einf. +case p0; intros; auto with zarith. +simpl in |- *. +repeat rewrite Esup. +case p0; intros; auto with zarith. +auto. Qed. (*i****************** @@ -205,61 +205,55 @@ Section divers. (** Number of significative digits. *) -Definition N_digits := - [x:Z]Cases x of - (POS p) => (log_inf p) - | (NEG p) => (log_inf p) - | ZERO => `0` - end. - -Lemma ZERO_le_N_digits : (x:Z) ` 0 <= (N_digits x)`. -Induction x; Simpl; -[ Apply Zle_n -| Exact log_inf_correct1 -| Exact log_inf_correct1]. +Definition N_digits (x:Z) := + match x with + | Zpos p => log_inf p + | Zneg p => log_inf p + | Z0 => 0 + end. + +Lemma ZERO_le_N_digits : forall x:Z, 0 <= N_digits x. +simple induction x; simpl in |- *; + [ apply Zle_refl | exact log_inf_correct1 | exact log_inf_correct1 ]. Qed. -Lemma log_inf_shift_nat : - (n:nat)(log_inf (shift_nat n xH))=(inject_nat n). -Induction n; Intros; -[ Try Trivial -| Rewrite -> inj_S; Rewrite <- H; Reflexivity]. +Lemma log_inf_shift_nat : forall n:nat, log_inf (shift_nat n 1) = Z_of_nat n. +simple induction n; intros; + [ try trivial | rewrite Znat.inj_S; rewrite <- H; reflexivity ]. Qed. -Lemma log_sup_shift_nat : - (n:nat)(log_sup (shift_nat n xH))=(inject_nat n). -Induction n; Intros; -[ Try Trivial -| Rewrite -> inj_S; Rewrite <- H; Reflexivity]. +Lemma log_sup_shift_nat : forall n:nat, log_sup (shift_nat n 1) = Z_of_nat n. +simple induction n; intros; + [ try trivial | rewrite Znat.inj_S; rewrite <- H; reflexivity ]. Qed. (** [Is_power p] means that p is a power of two *) -Fixpoint Is_power[p:positive] : Prop := - Cases p of - xH => True - | (xO q) => (Is_power q) - | (xI q) => False +Fixpoint Is_power (p:positive) : Prop := + match p with + | xH => True + | xO q => Is_power q + | xI q => False end. Lemma Is_power_correct : - (p:positive) (Is_power p) <-> (Ex [y:nat](p=(shift_nat y xH))). - -Split; -[ Elim p; - [ Simpl; Tauto - | Simpl; Intros; Generalize (H H0); Intro H1; Elim H1; Intros y0 Hy0; - Exists (S y0); Rewrite Hy0; Reflexivity - | Intro; Exists O; Reflexivity] -| Intros; Elim H; Intros; Rewrite -> H0; Elim x; Intros; Simpl; Trivial]. + forall p:positive, Is_power p <-> ( exists y : nat | p = shift_nat y 1). + +split; + [ elim p; + [ simpl in |- *; tauto + | simpl in |- *; intros; generalize (H H0); intro H1; elim H1; + intros y0 Hy0; exists (S y0); rewrite Hy0; reflexivity + | intro; exists 0%nat; reflexivity ] + | intros; elim H; intros; rewrite H0; elim x; intros; simpl in |- *; trivial ]. Qed. -Lemma Is_power_or : (p:positive) (Is_power p)\/~(Is_power p). -Induction p; -[ Intros; Right; Simpl; Tauto -| Intros; Elim H; - [ Intros; Left; Simpl; Exact H0 - | Intros; Right; Simpl; Exact H0] -| Left; Simpl; Trivial]. +Lemma Is_power_or : forall p:positive, Is_power p \/ ~ Is_power p. +simple induction p; + [ intros; right; simpl in |- *; tauto + | intros; elim H; + [ intros; left; simpl in |- *; exact H0 + | intros; right; simpl in |- *; exact H0 ] + | left; simpl in |- *; trivial ]. Qed. End divers. @@ -269,4 +263,3 @@ End divers. - diff --git a/theories/ZArith/Zmin.v b/theories/ZArith/Zmin.v index 01192c3bce..deab633924 100644 --- a/theories/ZArith/Zmin.v +++ b/theories/ZArith/Zmin.v @@ -9,94 +9,98 @@ (** Binary Integers (Pierre Crégut (CNET, Lannion, France) *) -Require Arith. -Require BinInt. -Require Zcompare. -Require Zorder. +Require Import Arith. +Require Import BinInt. +Require Import Zcompare. +Require Import Zorder. Open Local Scope Z_scope. (**********************************************************************) (** Minimum on binary integer numbers *) -Definition Zmin := [n,m:Z] - Cases (Zcompare n m) of - EGAL => n - | INFERIEUR => n - | SUPERIEUR => m - end. +Definition Zmin (n m:Z) := + match n ?= m return Z with + | Eq => n + | Lt => n + | Gt => m + end. (** Properties of minimum on binary integer numbers *) -Lemma Zmin_SS : (n,m:Z)((Zs (Zmin n m))=(Zmin (Zs n) (Zs m))). +Lemma Zmin_SS : forall n m:Z, Zsucc (Zmin n m) = Zmin (Zsucc n) (Zsucc m). Proof. -Intros n m;Unfold Zmin; Rewrite (Zcompare_n_S n m); -(ElimCompare 'n 'm);Intros E;Rewrite E;Auto with arith. +intros n m; unfold Zmin in |- *; rewrite (Zcompare_succ_compat n m); + elim_compare n m; intros E; rewrite E; auto with arith. Qed. -Lemma Zle_min_l : (n,m:Z)(Zle (Zmin n m) n). +Lemma Zle_min_l : forall n m:Z, Zmin n m <= n. Proof. -Intros n m;Unfold Zmin ; (ElimCompare 'n 'm);Intros E;Rewrite -> E; - [ Apply Zle_n | Apply Zle_n | Apply Zlt_le_weak; Apply Zgt_lt;Exact E ]. +intros n m; unfold Zmin in |- *; elim_compare n m; intros E; rewrite E; + [ apply Zle_refl + | apply Zle_refl + | apply Zlt_le_weak; apply Zgt_lt; exact E ]. Qed. -Lemma Zle_min_r : (n,m:Z)(Zle (Zmin n m) m). +Lemma Zle_min_r : forall n m:Z, Zmin n m <= m. Proof. -Intros n m;Unfold Zmin ; (ElimCompare 'n 'm);Intros E;Rewrite -> E;[ - Unfold Zle ;Rewrite -> E;Discriminate -| Unfold Zle ;Rewrite -> E;Discriminate -| Apply Zle_n ]. +intros n m; unfold Zmin in |- *; elim_compare n m; intros E; rewrite E; + [ unfold Zle in |- *; rewrite E; discriminate + | unfold Zle in |- *; rewrite E; discriminate + | apply Zle_refl ]. Qed. -Lemma Zmin_case : (n,m:Z)(P:Z->Set)(P n)->(P m)->(P (Zmin n m)). +Lemma Zmin_case : forall (n m:Z) (P:Z -> Set), P n -> P m -> P (Zmin n m). Proof. -Intros n m P H1 H2; Unfold Zmin; Case (Zcompare n m);Auto with arith. +intros n m P H1 H2; unfold Zmin in |- *; case (n ?= m); auto with arith. Qed. -Lemma Zmin_or : (n,m:Z)(Zmin n m)=n \/ (Zmin n m)=m. +Lemma Zmin_or : forall n m:Z, Zmin n m = n \/ Zmin n m = m. Proof. -Unfold Zmin; Intros; Elim (Zcompare n m); Auto. +unfold Zmin in |- *; intros; elim (n ?= m); auto. Qed. -Lemma Zmin_n_n : (n:Z) (Zmin n n)=n. +Lemma Zmin_n_n : forall n:Z, Zmin n n = n. Proof. -Unfold Zmin; Intros; Elim (Zcompare n n); Auto. +unfold Zmin in |- *; intros; elim (n ?= n); auto. Qed. -Lemma Zmin_plus : - (x,y,n:Z)(Zmin (Zplus x n) (Zplus y n))=(Zplus (Zmin x y) n). +Lemma Zmin_plus : forall n m p:Z, Zmin (n + p) (m + p) = Zmin n m + p. Proof. -Intros x y n; Unfold Zmin. -Rewrite (Zplus_sym x n); -Rewrite (Zplus_sym y n); -Rewrite (Zcompare_Zplus_compatible x y n). -Case (Zcompare x y); Apply Zplus_sym. +intros x y n; unfold Zmin in |- *. +rewrite (Zplus_comm x n); rewrite (Zplus_comm y n); + rewrite (Zcompare_plus_compat x y n). +case (x ?= y); apply Zplus_comm. Qed. (**********************************************************************) (** Maximum of two binary integer numbers *) -V7only [ (* From Zdivides *) ]. -Definition Zmax := - [a, b : ?] Cases (Zcompare a b) of INFERIEUR => b | _ => a end. +Definition Zmax a b := match a ?= b with + | Lt => b + | _ => a + end. (** Properties of maximum on binary integer numbers *) -Tactic Definition CaseEq name := -Generalize (refl_equal ? name); Pattern -1 name; Case name. +Ltac CaseEq name := + generalize (refl_equal name); pattern name at -1 in |- *; case name. -Theorem Zmax1: (a, b : ?) (Zle a (Zmax a b)). +Theorem Zmax1 : forall a b, a <= Zmax a b. Proof. -Intros a b; Unfold Zmax; (CaseEq '(Zcompare a b)); Simpl; Auto with zarith. -Unfold Zle; Intros H; Rewrite H; Red; Intros; Discriminate. +intros a b; unfold Zmax in |- *; CaseEq (a ?= b); simpl in |- *; + auto with zarith. +unfold Zle in |- *; intros H; rewrite H; red in |- *; intros; discriminate. Qed. -Theorem Zmax2: (a, b : ?) (Zle b (Zmax a b)). +Theorem Zmax2 : forall a b, b <= Zmax a b. Proof. -Intros a b; Unfold Zmax; (CaseEq '(Zcompare a b)); Simpl; Auto with zarith. -Intros H; - (Case (Zle_or_lt b a); Auto; Unfold Zlt; Rewrite H; Intros; Discriminate). -Intros H; - (Case (Zle_or_lt b a); Auto; Unfold Zlt; Rewrite H; Intros; Discriminate). +intros a b; unfold Zmax in |- *; CaseEq (a ?= b); simpl in |- *; + auto with zarith. +intros H; + (case (Zle_or_lt b a); auto; unfold Zlt in |- *; rewrite H; intros; + discriminate). +intros H; + (case (Zle_or_lt b a); auto; unfold Zlt in |- *; rewrite H; intros; + discriminate). Qed. - diff --git a/theories/ZArith/Zmisc.v b/theories/ZArith/Zmisc.v index a8bbcfc00a..0ad0ef288f 100644 --- a/theories/ZArith/Zmisc.v +++ b/theories/ZArith/Zmisc.v @@ -8,181 +8,90 @@ (*i $Id$ i*) -Require BinInt. -Require Zcompare. -Require Zorder. -Require Zsyntax. -Require Bool. -V7only [Import Z_scope.]. +Require Import BinInt. +Require Import Zcompare. +Require Import Zorder. +Require Import Bool. Open Local Scope Z_scope. (**********************************************************************) (** Iterators *) (** [n]th iteration of the function [f] *) -Fixpoint iter_nat[n:nat] : (A:Set)(f:A->A)A->A := - [A:Set][f:A->A][x:A] - Cases n of - O => x - | (S n') => (f (iter_nat n' A f x)) - end. +Fixpoint iter_nat (n:nat) (A:Set) (f:A -> A) (x:A) {struct n} : A := + match n with + | O => x + | S n' => f (iter_nat n' A f x) + end. -Fixpoint iter_pos[n:positive] : (A:Set)(f:A->A)A->A := - [A:Set][f:A->A][x:A] - Cases n of - xH => (f x) - | (xO n') => (iter_pos n' A f (iter_pos n' A f x)) - | (xI n') => (f (iter_pos n' A f (iter_pos n' A f x))) - end. +Fixpoint iter_pos (n:positive) (A:Set) (f:A -> A) (x:A) {struct n} : A := + match n with + | xH => f x + | xO n' => iter_pos n' A f (iter_pos n' A f x) + | xI n' => f (iter_pos n' A f (iter_pos n' A f x)) + end. -Definition iter := - [n:Z][A:Set][f:A->A][x:A]Cases n of - ZERO => x - | (POS p) => (iter_pos p A f x) - | (NEG p) => x +Definition iter (n:Z) (A:Set) (f:A -> A) (x:A) := + match n with + | Z0 => x + | Zpos p => iter_pos p A f x + | Zneg p => x end. Theorem iter_nat_plus : - (n,m:nat)(A:Set)(f:A->A)(x:A) - (iter_nat (plus n m) A f x)=(iter_nat n A f (iter_nat m A f x)). + forall (n m:nat) (A:Set) (f:A -> A) (x:A), + iter_nat (n + m) A f x = iter_nat n A f (iter_nat m A f x). Proof. -Induction n; -[ Simpl; Auto with arith -| Intros; Simpl; Apply f_equal with f:=f; Apply H -]. +simple induction n; + [ simpl in |- *; auto with arith + | intros; simpl in |- *; apply f_equal with (f := f); apply H ]. Qed. -Theorem iter_convert : (n:positive)(A:Set)(f:A->A)(x:A) - (iter_pos n A f x) = (iter_nat (convert n) A f x). +Theorem iter_nat_of_P : + forall (p:positive) (A:Set) (f:A -> A) (x:A), + iter_pos p A f x = iter_nat (nat_of_P p) A f x. Proof. -Intro n; NewInduction n as [p H|p H|]; -[ Intros; Simpl; Rewrite -> (H A f x); - Rewrite -> (H A f (iter_nat (convert p) A f x)); - Rewrite -> (ZL6 p); Symmetry; Apply f_equal with f:=f; - Apply iter_nat_plus -| Intros; Unfold convert; Simpl; Rewrite -> (H A f x); - Rewrite -> (H A f (iter_nat (convert p) A f x)); - Rewrite -> (ZL6 p); Symmetry; - Apply iter_nat_plus -| Simpl; Auto with arith -]. +intro n; induction n as [p H| p H| ]; + [ intros; simpl in |- *; rewrite (H A f x); + rewrite (H A f (iter_nat (nat_of_P p) A f x)); + rewrite (ZL6 p); symmetry in |- *; apply f_equal with (f := f); + apply iter_nat_plus + | intros; unfold nat_of_P in |- *; simpl in |- *; rewrite (H A f x); + rewrite (H A f (iter_nat (nat_of_P p) A f x)); + rewrite (ZL6 p); symmetry in |- *; apply iter_nat_plus + | simpl in |- *; auto with arith ]. Qed. -Theorem iter_pos_add : - (n,m:positive)(A:Set)(f:A->A)(x:A) - (iter_pos (add n m) A f x)=(iter_pos n A f (iter_pos m A f x)). +Theorem iter_pos_plus : + forall (p q:positive) (A:Set) (f:A -> A) (x:A), + iter_pos (p + q) A f x = iter_pos p A f (iter_pos q A f x). Proof. -Intros n m; Intros. -Rewrite -> (iter_convert m A f x). -Rewrite -> (iter_convert n A f (iter_nat (convert m) A f x)). -Rewrite -> (iter_convert (add n m) A f x). -Rewrite -> (convert_add n m). -Apply iter_nat_plus. +intros n m; intros. +rewrite (iter_nat_of_P m A f x). +rewrite (iter_nat_of_P n A f (iter_nat (nat_of_P m) A f x)). +rewrite (iter_nat_of_P (n + m) A f x). +rewrite (nat_of_P_plus_morphism n m). +apply iter_nat_plus. Qed. (** Preservation of invariants : if [f : A->A] preserves the invariant [Inv], then the iterates of [f] also preserve it. *) Theorem iter_nat_invariant : - (n:nat)(A:Set)(f:A->A)(Inv:A->Prop) - ((x:A)(Inv x)->(Inv (f x)))->(x:A)(Inv x)->(Inv (iter_nat n A f x)). + forall (n:nat) (A:Set) (f:A -> A) (Inv:A -> Prop), + (forall x:A, Inv x -> Inv (f x)) -> + forall x:A, Inv x -> Inv (iter_nat n A f x). Proof. -Induction n; Intros; -[ Trivial with arith -| Simpl; Apply H0 with x:=(iter_nat n0 A f x); Apply H; Trivial with arith]. +simple induction n; intros; + [ trivial with arith + | simpl in |- *; apply H0 with (x := iter_nat n0 A f x); apply H; + trivial with arith ]. Qed. Theorem iter_pos_invariant : - (n:positive)(A:Set)(f:A->A)(Inv:A->Prop) - ((x:A)(Inv x)->(Inv (f x)))->(x:A)(Inv x)->(Inv (iter_pos n A f x)). + forall (p:positive) (A:Set) (f:A -> A) (Inv:A -> Prop), + (forall x:A, Inv x -> Inv (f x)) -> + forall x:A, Inv x -> Inv (iter_pos p A f x). Proof. -Intros; Rewrite iter_convert; Apply iter_nat_invariant; Trivial with arith. +intros; rewrite iter_nat_of_P; apply iter_nat_invariant; trivial with arith. Qed. - -V7only [ -(* Compatibility *) -Require Zbool. -Require Zeven. -Require Zabs. -Require Zmin. -Notation rename := rename. -Notation POS_xI := POS_xI. -Notation POS_xO := POS_xO. -Notation NEG_xI := NEG_xI. -Notation NEG_xO := NEG_xO. -Notation POS_add := POS_add. -Notation NEG_add := NEG_add. -Notation Zle_cases := Zle_cases. -Notation Zlt_cases := Zlt_cases. -Notation Zge_cases := Zge_cases. -Notation Zgt_cases := Zgt_cases. -Notation POS_gt_ZERO := POS_gt_ZERO. -Notation ZERO_le_POS := ZERO_le_POS. -Notation Zlt_ZERO_pred_le_ZERO := Zlt_ZERO_pred_le_ZERO. -Notation NEG_lt_ZERO := NEG_lt_ZERO. -Notation Zeven_not_Zodd := Zeven_not_Zodd. -Notation Zodd_not_Zeven := Zodd_not_Zeven. -Notation Zeven_Sn := Zeven_Sn. -Notation Zodd_Sn := Zodd_Sn. -Notation Zeven_pred := Zeven_pred. -Notation Zodd_pred := Zodd_pred. -Notation Zeven_div2 := Zeven_div2. -Notation Zodd_div2 := Zodd_div2. -Notation Zodd_div2_neg := Zodd_div2_neg. -Notation Z_modulo_2 := Z_modulo_2. -Notation Zsplit2 := Zsplit2. -Notation Zminus_Zplus_compatible := Zminus_Zplus_compatible. -Notation Zcompare_egal_dec := Zcompare_egal_dec. -Notation Zcompare_elim := Zcompare_elim. -Notation Zcompare_x_x := Zcompare_x_x. -Notation Zlt_not_eq := Zlt_not_eq. -Notation Zcompare_eq_case := Zcompare_eq_case. -Notation Zle_Zcompare := Zle_Zcompare. -Notation Zlt_Zcompare := Zlt_Zcompare. -Notation Zge_Zcompare := Zge_Zcompare. -Notation Zgt_Zcompare := Zgt_Zcompare. -Notation Zmin_plus := Zmin_plus. -Notation absolu_lt := absolu_lt. -Notation Zle_bool_imp_le := Zle_bool_imp_le. -Notation Zle_imp_le_bool := Zle_imp_le_bool. -Notation Zle_bool_refl := Zle_bool_refl. -Notation Zle_bool_antisym := Zle_bool_antisym. -Notation Zle_bool_trans := Zle_bool_trans. -Notation Zle_bool_plus_mono := Zle_bool_plus_mono. -Notation Zone_pos := Zone_pos. -Notation Zone_min_pos := Zone_min_pos. -Notation Zle_is_le_bool := Zle_is_le_bool. -Notation Zge_is_le_bool := Zge_is_le_bool. -Notation Zlt_is_le_bool := Zlt_is_le_bool. -Notation Zgt_is_le_bool := Zgt_is_le_bool. -Notation Zle_plus_swap := Zle_plus_swap. -Notation Zge_iff_le := Zge_iff_le. -Notation Zlt_plus_swap := Zlt_plus_swap. -Notation Zgt_iff_lt := Zgt_iff_lt. -Notation Zeq_plus_swap := Zeq_plus_swap. -(* Definitions *) -Notation entier_of_Z := entier_of_Z. -Notation Z_of_entier := Z_of_entier. -Notation Zle_bool := Zle_bool. -Notation Zge_bool := Zge_bool. -Notation Zlt_bool := Zlt_bool. -Notation Zgt_bool := Zgt_bool. -Notation Zeq_bool := Zeq_bool. -Notation Zneq_bool := Zneq_bool. -Notation Zeven := Zeven. -Notation Zodd := Zodd. -Notation Zeven_bool := Zeven_bool. -Notation Zodd_bool := Zodd_bool. -Notation Zeven_odd_dec := Zeven_odd_dec. -Notation Zeven_dec := Zeven_dec. -Notation Zodd_dec := Zodd_dec. -Notation Zdiv2_pos := Zdiv2_pos. -Notation Zdiv2 := Zdiv2. -Notation Zle_bool_total := Zle_bool_total. -Export Zbool. -Export Zeven. -Export Zabs. -Export Zmin. -Export Zorder. -Export Zcompare. -]. diff --git a/theories/ZArith/Znat.v b/theories/ZArith/Znat.v index fe53fce90e..d9bc4d1b2b 100644 --- a/theories/ZArith/Znat.v +++ b/theories/ZArith/Znat.v @@ -11,128 +11,128 @@ (** Binary Integers (Pierre Crégut, CNET, Lannion, France) *) Require Export Arith. -Require BinPos. -Require BinInt. -Require Zcompare. -Require Zorder. -Require Decidable. -Require Peano_dec. +Require Import BinPos. +Require Import BinInt. +Require Import Zcompare. +Require Import Zorder. +Require Import Decidable. +Require Import Peano_dec. Require Export Compare_dec. Open Local Scope Z_scope. -Definition neq := [x,y:nat] ~(x=y). +Definition neq (x y:nat) := x <> y. (**********************************************************************) (** Properties of the injection from nat into Z *) -Theorem inj_S : (y:nat) (inject_nat (S y)) = (Zs (inject_nat y)). +Theorem inj_S : forall n:nat, Z_of_nat (S n) = Zsucc (Z_of_nat n). Proof. -Intro y; NewInduction y as [|n H]; [ - Unfold Zs ; Simpl; Trivial with arith -| Change (POS (add_un (anti_convert n)))=(Zs (inject_nat (S n))); - Rewrite add_un_Zs; Trivial with arith]. +intro y; induction y as [| n H]; + [ unfold Zsucc in |- *; simpl in |- *; trivial with arith + | change (Zpos (Psucc (P_of_succ_nat n)) = Zsucc (Z_of_nat (S n))) in |- *; + rewrite Zpos_succ_morphism; trivial with arith ]. Qed. -Theorem inj_plus : - (x,y:nat) (inject_nat (plus x y)) = (Zplus (inject_nat x) (inject_nat y)). +Theorem inj_plus : forall n m:nat, Z_of_nat (n + m) = Z_of_nat n + Z_of_nat m. Proof. -Intro x; NewInduction x as [|n H]; Intro y; NewDestruct y as [|m]; [ - Simpl; Trivial with arith -| Simpl; Trivial with arith -| Simpl; Rewrite <- plus_n_O; Trivial with arith -| Change (inject_nat (S (plus n (S m))))= - (Zplus (inject_nat (S n)) (inject_nat (S m))); - Rewrite inj_S; Rewrite H; Do 2 Rewrite inj_S; Rewrite Zplus_S_n; Trivial with arith]. +intro x; induction x as [| n H]; intro y; destruct y as [| m]; + [ simpl in |- *; trivial with arith + | simpl in |- *; trivial with arith + | simpl in |- *; rewrite <- plus_n_O; trivial with arith + | change (Z_of_nat (S (n + S m)) = Z_of_nat (S n) + Z_of_nat (S m)) in |- *; + rewrite inj_S; rewrite H; do 2 rewrite inj_S; rewrite Zplus_succ_l; + trivial with arith ]. Qed. -Theorem inj_mult : - (x,y:nat) (inject_nat (mult x y)) = (Zmult (inject_nat x) (inject_nat y)). +Theorem inj_mult : forall n m:nat, Z_of_nat (n * m) = Z_of_nat n * Z_of_nat m. Proof. -Intro x; NewInduction x as [|n H]; [ - Simpl; Trivial with arith -| Intro y; Rewrite -> inj_S; Rewrite <- Zmult_Sm_n; - Rewrite <- H;Rewrite <- inj_plus; Simpl; Rewrite plus_sym; Trivial with arith]. +intro x; induction x as [| n H]; + [ simpl in |- *; trivial with arith + | intro y; rewrite inj_S; rewrite <- Zmult_succ_l_reverse; rewrite <- H; + rewrite <- inj_plus; simpl in |- *; rewrite plus_comm; + trivial with arith ]. Qed. -Theorem inj_neq: - (x,y:nat) (neq x y) -> (Zne (inject_nat x) (inject_nat y)). +Theorem inj_neq : forall n m:nat, neq n m -> Zne (Z_of_nat n) (Z_of_nat m). Proof. -Unfold neq Zne not ; Intros x y H1 H2; Apply H1; Generalize H2; -Case x; Case y; Intros; [ - Auto with arith -| Discriminate H0 -| Discriminate H0 -| Simpl in H0; Injection H0; Do 2 Rewrite <- bij1; Intros E; Rewrite E; Auto with arith]. +unfold neq, Zne, not in |- *; intros x y H1 H2; apply H1; generalize H2; + case x; case y; intros; + [ auto with arith + | discriminate H0 + | discriminate H0 + | simpl in H0; injection H0; + do 2 rewrite <- nat_of_P_o_P_of_succ_nat_eq_succ; + intros E; rewrite E; auto with arith ]. Qed. -Theorem inj_le: - (x,y:nat) (le x y) -> (Zle (inject_nat x) (inject_nat y)). +Theorem inj_le : forall n m:nat, (n <= m)%nat -> Z_of_nat n <= Z_of_nat m. Proof. -Intros x y; Intros H; Elim H; [ - Unfold Zle ; Elim (Zcompare_EGAL (inject_nat x) (inject_nat x)); - Intros H1 H2; Rewrite H2; [ Discriminate | Trivial with arith] -| Intros m H1 H2; Apply Zle_trans with (inject_nat m); - [Assumption | Rewrite inj_S; Apply Zle_n_Sn]]. +intros x y; intros H; elim H; + [ unfold Zle in |- *; elim (Zcompare_Eq_iff_eq (Z_of_nat x) (Z_of_nat x)); + intros H1 H2; rewrite H2; [ discriminate | trivial with arith ] + | intros m H1 H2; apply Zle_trans with (Z_of_nat m); + [ assumption | rewrite inj_S; apply Zle_succ ] ]. Qed. -Theorem inj_lt: (x,y:nat) (lt x y) -> (Zlt (inject_nat x) (inject_nat y)). +Theorem inj_lt : forall n m:nat, (n < m)%nat -> Z_of_nat n < Z_of_nat m. Proof. -Intros x y H; Apply Zgt_lt; Apply Zle_S_gt; Rewrite <- inj_S; Apply inj_le; -Exact H. +intros x y H; apply Zgt_lt; apply Zlt_succ_gt; rewrite <- inj_S; apply inj_le; + exact H. Qed. -Theorem inj_gt: (x,y:nat) (gt x y) -> (Zgt (inject_nat x) (inject_nat y)). +Theorem inj_gt : forall n m:nat, (n > m)%nat -> Z_of_nat n > Z_of_nat m. Proof. -Intros x y H; Apply Zlt_gt; Apply inj_lt; Exact H. +intros x y H; apply Zlt_gt; apply inj_lt; exact H. Qed. -Theorem inj_ge: (x,y:nat) (ge x y) -> (Zge (inject_nat x) (inject_nat y)). +Theorem inj_ge : forall n m:nat, (n >= m)%nat -> Z_of_nat n >= Z_of_nat m. Proof. -Intros x y H; Apply Zle_ge; Apply inj_le; Apply H. +intros x y H; apply Zle_ge; apply inj_le; apply H. Qed. -Theorem inj_eq: (x,y:nat) x=y -> (inject_nat x) = (inject_nat y). +Theorem inj_eq : forall n m:nat, n = m -> Z_of_nat n = Z_of_nat m. Proof. -Intros x y H; Rewrite H; Trivial with arith. +intros x y H; rewrite H; trivial with arith. Qed. -Theorem intro_Z : - (x:nat) (EX y:Z | (inject_nat x)=y /\ - (Zle ZERO (Zplus (Zmult y (POS xH)) ZERO))). +Theorem intro_Z : + forall n:nat, exists y : Z | Z_of_nat n = y /\ 0 <= y * 1 + 0. Proof. -Intros x; Exists (inject_nat x); Split; [ - Trivial with arith -| Rewrite Zmult_sym; Rewrite Zmult_one; Rewrite Zero_right; - Unfold Zle ; Elim x; Intros;Simpl; Discriminate ]. +intros x; exists (Z_of_nat x); split; + [ trivial with arith + | rewrite Zmult_comm; rewrite Zmult_1_l; rewrite Zplus_0_r; + unfold Zle in |- *; elim x; intros; simpl in |- *; + discriminate ]. Qed. Theorem inj_minus1 : - (x,y:nat) (le y x) -> - (inject_nat (minus x y)) = (Zminus (inject_nat x) (inject_nat y)). + forall n m:nat, (m <= n)%nat -> Z_of_nat (n - m) = Z_of_nat n - Z_of_nat m. Proof. -Intros x y H; Apply (Zsimpl_plus_l (inject_nat y)); Unfold Zminus ; -Rewrite Zplus_permute; Rewrite Zplus_inverse_r; Rewrite <- inj_plus; -Rewrite <- (le_plus_minus y x H);Rewrite Zero_right; Trivial with arith. +intros x y H; apply (Zplus_reg_l (Z_of_nat y)); unfold Zminus in |- *; + rewrite Zplus_permute; rewrite Zplus_opp_r; rewrite <- inj_plus; + rewrite <- (le_plus_minus y x H); rewrite Zplus_0_r; + trivial with arith. Qed. -Theorem inj_minus2: (x,y:nat) (gt y x) -> (inject_nat (minus x y)) = ZERO. +Theorem inj_minus2 : forall n m:nat, (m > n)%nat -> Z_of_nat (n - m) = 0. Proof. -Intros x y H; Rewrite inj_minus_aux; [ Trivial with arith | Apply gt_not_le; Assumption]. +intros x y H; rewrite not_le_minus_0; + [ trivial with arith | apply gt_not_le; assumption ]. Qed. -V7only [ (* From Zdivides *) ]. -Theorem POS_inject: (x : positive) (POS x) = (inject_nat (convert x)). +Theorem Zpos_eq_Z_of_nat_o_nat_of_P : + forall p:positive, Zpos p = Z_of_nat (nat_of_P p). Proof. -Intros x; Elim x; Simpl; Auto. -Intros p H; Rewrite ZL6. -Apply f_equal with f := POS. -Apply convert_intro. -Rewrite bij1; Unfold convert; Simpl. -Rewrite ZL6; Auto. -Intros p H; Unfold convert; Simpl. -Rewrite ZL6; Simpl. -Rewrite inj_plus; Repeat Rewrite <- H. -Rewrite POS_xO; Simpl; Rewrite add_x_x; Reflexivity. +intros x; elim x; simpl in |- *; auto. +intros p H; rewrite ZL6. +apply f_equal with (f := Zpos). +apply nat_of_P_inj. +rewrite nat_of_P_o_P_of_succ_nat_eq_succ; unfold nat_of_P in |- *; + simpl in |- *. +rewrite ZL6; auto. +intros p H; unfold nat_of_P in |- *; simpl in |- *. +rewrite ZL6; simpl in |- *. +rewrite inj_plus; repeat rewrite <- H. +rewrite Zpos_xO; simpl in |- *; rewrite Pplus_diag; reflexivity. Qed. - diff --git a/theories/ZArith/Znumtheory.v b/theories/ZArith/Znumtheory.v index dfe1c31fda..ed6272c445 100644 --- a/theories/ZArith/Znumtheory.v +++ b/theories/ZArith/Znumtheory.v @@ -8,11 +8,10 @@ (*i $Id$ i*) -Require ZArith_base. -Require ZArithRing. -Require Zcomplements. -Require Zdiv. -V7only [Import Z_scope.]. +Require Import ZArith_base. +Require Import ZArithRing. +Require Import Zcomplements. +Require Import Zdiv. Open Local Scope Z_scope. (** This file contains some notions of number theory upon Z numbers: @@ -26,176 +25,173 @@ Open Local Scope Z_scope. (** * Divisibility *) -Inductive Zdivide [a,b:Z] : Prop := - Zdivide_intro : (q:Z) `b = q * a` -> (Zdivide a b). +Inductive Zdivide (a b:Z) : Prop := + Zdivide_intro : forall q:Z, b = q * a -> Zdivide a b. (** Syntax for divisibility *) -Notation "( a | b )" := (Zdivide a b) - (at level 0, a,b at level 10) : Z_scope - V8only (at level 0, a,b at level 250). +Notation "( a | b )" := (Zdivide a b) (at level 0, a, b at level 250) : + Z_scope. (** Results concerning divisibility*) -Lemma Zdivide_refl : (a:Z) (a|a). +Lemma Zdivide_refl : forall a:Z, (a | a). Proof. -Intros; Apply Zdivide_intro with `1`; Ring. -Save. +intros; apply Zdivide_intro with 1; ring. +Qed. -Lemma Zone_divide : (a:Z) (1|a). +Lemma Zone_divide : forall a:Z, (1 | a). Proof. -Intros; Apply Zdivide_intro with `a`; Ring. -Save. +intros; apply Zdivide_intro with a; ring. +Qed. -Lemma Zdivide_0 : (a:Z) (a|0). +Lemma Zdivide_0 : forall a:Z, (a | 0). Proof. -Intros; Apply Zdivide_intro with `0`; Ring. -Save. +intros; apply Zdivide_intro with 0; ring. +Qed. -Hints Resolve Zdivide_refl Zone_divide Zdivide_0 : zarith. +Hint Resolve Zdivide_refl Zone_divide Zdivide_0: zarith. -Lemma Zdivide_mult_left : (a,b,c:Z) (a|b) -> (`c*a`|`c*b`). +Lemma Zmult_divide_compat_l : forall a b c:Z, (a | b) -> (c * a | c * b). Proof. -Induction 1; Intros; Apply Zdivide_intro with q. -Rewrite H0; Ring. -Save. +simple induction 1; intros; apply Zdivide_intro with q. +rewrite H0; ring. +Qed. -Lemma Zdivide_mult_right : (a,b,c:Z) (a|b) -> (`a*c`|`b*c`). +Lemma Zmult_divide_compat_r : forall a b c:Z, (a | b) -> (a * c | b * c). Proof. -Intros a b c; Rewrite (Zmult_sym a c); Rewrite (Zmult_sym b c). -Apply Zdivide_mult_left; Trivial. -Save. +intros a b c; rewrite (Zmult_comm a c); rewrite (Zmult_comm b c). +apply Zmult_divide_compat_l; trivial. +Qed. -Hints Resolve Zdivide_mult_left Zdivide_mult_right : zarith. +Hint Resolve Zmult_divide_compat_l Zmult_divide_compat_r: zarith. -Lemma Zdivide_plus : (a,b,c:Z) (a|b) -> (a|c) -> (a|`b+c`). +Lemma Zdivide_plus_r : forall a b c:Z, (a | b) -> (a | c) -> (a | b + c). Proof. -Induction 1; Intros q Hq; Induction 1; Intros q' Hq'. -Apply Zdivide_intro with `q+q'`. -Rewrite Hq; Rewrite Hq'; Ring. -Save. +simple induction 1; intros q Hq; simple induction 1; intros q' Hq'. +apply Zdivide_intro with (q + q'). +rewrite Hq; rewrite Hq'; ring. +Qed. -Lemma Zdivide_opp : (a,b:Z) (a|b) -> (a|`-b`). +Lemma Zdivide_opp_r : forall a b:Z, (a | b) -> (a | - b). Proof. -Induction 1; Intros; Apply Zdivide_intro with `-q`. -Rewrite H0; Ring. -Save. +simple induction 1; intros; apply Zdivide_intro with (- q). +rewrite H0; ring. +Qed. -Lemma Zdivide_opp_rev : (a,b:Z) (a|`-b`) -> (a| b). +Lemma Zdivide_opp_r_rev : forall a b:Z, (a | - b) -> (a | b). Proof. -Intros; Replace b with `-(-b)`. Apply Zdivide_opp; Trivial. Ring. -Save. +intros; replace b with (- - b). apply Zdivide_opp_r; trivial. ring. +Qed. -Lemma Zdivide_opp_left : (a,b:Z) (a|b) -> (`-a`|b). +Lemma Zdivide_opp_l : forall a b:Z, (a | b) -> (- a | b). Proof. -Induction 1; Intros; Apply Zdivide_intro with `-q`. -Rewrite H0; Ring. -Save. +simple induction 1; intros; apply Zdivide_intro with (- q). +rewrite H0; ring. +Qed. -Lemma Zdivide_opp_left_rev : (a,b:Z) (`-a`|b) -> (a|b). +Lemma Zdivide_opp_l_rev : forall a b:Z, (- a | b) -> (a | b). Proof. -Intros; Replace a with `-(-a)`. Apply Zdivide_opp_left; Trivial. Ring. -Save. +intros; replace a with (- - a). apply Zdivide_opp_l; trivial. ring. +Qed. -Lemma Zdivide_minus : (a,b,c:Z) (a|b) -> (a|c) -> (a|`b-c`). +Lemma Zdivide_minus_l : forall a b c:Z, (a | b) -> (a | c) -> (a | b - c). Proof. -Induction 1; Intros q Hq; Induction 1; Intros q' Hq'. -Apply Zdivide_intro with `q-q'`. -Rewrite Hq; Rewrite Hq'; Ring. -Save. +simple induction 1; intros q Hq; simple induction 1; intros q' Hq'. +apply Zdivide_intro with (q - q'). +rewrite Hq; rewrite Hq'; ring. +Qed. -Lemma Zdivide_left : (a,b,c:Z) (a|b) -> (a|`b*c`). +Lemma Zdivide_mult_l : forall a b c:Z, (a | b) -> (a | b * c). Proof. -Induction 1; Intros q Hq; Apply Zdivide_intro with `q*c`. -Rewrite Hq; Ring. -Save. +simple induction 1; intros q Hq; apply Zdivide_intro with (q * c). +rewrite Hq; ring. +Qed. -Lemma Zdivide_right : (a,b,c:Z) (a|c) -> (a|`b*c`). +Lemma Zdivide_mult_r : forall a b c:Z, (a | c) -> (a | b * c). Proof. -Induction 1; Intros q Hq; Apply Zdivide_intro with `q*b`. -Rewrite Hq; Ring. -Save. +simple induction 1; intros q Hq; apply Zdivide_intro with (q * b). +rewrite Hq; ring. +Qed. -Lemma Zdivide_a_ab : (a,b:Z) (a|`a*b`). +Lemma Zdivide_factor_r : forall a b:Z, (a | a * b). Proof. -Intros; Apply Zdivide_intro with b; Ring. -Save. +intros; apply Zdivide_intro with b; ring. +Qed. -Lemma Zdivide_a_ba : (a,b:Z) (a|`b*a`). +Lemma Zdivide_factor_l : forall a b:Z, (a | b * a). Proof. -Intros; Apply Zdivide_intro with b; Ring. -Save. +intros; apply Zdivide_intro with b; ring. +Qed. -Hints Resolve Zdivide_plus Zdivide_opp Zdivide_opp_rev - Zdivide_opp_left Zdivide_opp_left_rev - Zdivide_minus Zdivide_left Zdivide_right - Zdivide_a_ab Zdivide_a_ba : zarith. +Hint Resolve Zdivide_plus_r Zdivide_opp_r Zdivide_opp_r_rev Zdivide_opp_l + Zdivide_opp_l_rev Zdivide_minus_l Zdivide_mult_l Zdivide_mult_r + Zdivide_factor_r Zdivide_factor_l: zarith. (** Auxiliary result. *) -Lemma Zmult_one : - (x,y:Z) `x>=0` -> `x*y=1` -> `x=1`. +Lemma Zmult_one : forall x y:Z, x >= 0 -> x * y = 1 -> x = 1. Proof. -Intros x y H H0; NewDestruct (Zmult_1_inversion_l ? ? H0) as [Hpos|Hneg]. - Assumption. - Rewrite Hneg in H; Simpl in H. - Contradiction (Zle_not_lt `0` `-1`). - Apply Zge_le; Assumption. - Apply NEG_lt_ZERO. -Save. +intros x y H H0; destruct (Zmult_1_inversion_l _ _ H0) as [Hpos| Hneg]. + assumption. + rewrite Hneg in H; simpl in H. + contradiction (Zle_not_lt 0 (-1)). + apply Zge_le; assumption. + apply Zorder.Zlt_neg_0. +Qed. (** Only [1] and [-1] divide [1]. *) -Lemma Zdivide_1 : (x:Z) (x|1) -> `x=1` \/ `x=-1`. +Lemma Zdivide_1 : forall x:Z, (x | 1) -> x = 1 \/ x = -1. Proof. -Induction 1; Intros. -Elim (Z_lt_ge_dec `0` x); [Left|Right]. -Apply Zmult_one with q; Auto with zarith; Rewrite H0; Ring. -Assert `(-x) = 1`; Auto with zarith. -Apply Zmult_one with (-q); Auto with zarith; Rewrite H0; Ring. -Save. +simple induction 1; intros. +elim (Z_lt_ge_dec 0 x); [ left | right ]. +apply Zmult_one with q; auto with zarith; rewrite H0; ring. +assert (- x = 1); auto with zarith. +apply Zmult_one with (- q); auto with zarith; rewrite H0; ring. +Qed. (** If [a] divides [b] and [b] divides [a] then [a] is [b] or [-b]. *) -Lemma Zdivide_antisym : (a,b:Z) (a|b) -> (b|a) -> `a=b` \/ `a=-b`. -Proof. -Induction 1; Intros. -Inversion H1. -Rewrite H0 in H2; Clear H H1. -Case (Z_zerop a); Intro. -Left; Rewrite H0; Rewrite e; Ring. -Assert Hqq0: `q0*q = 1`. -Apply Zmult_reg_left with a. -Assumption. -Ring. -Pattern 2 a; Rewrite H2; Ring. -Assert (q|1). -Rewrite <- Hqq0; Auto with zarith. -Elim (Zdivide_1 q H); Intros. -Rewrite H1 in H0; Left; Omega. -Rewrite H1 in H0; Right; Omega. -Save. +Lemma Zdivide_antisym : forall a b:Z, (a | b) -> (b | a) -> a = b \/ a = - b. +Proof. +simple induction 1; intros. +inversion H1. +rewrite H0 in H2; clear H H1. +case (Z_zerop a); intro. +left; rewrite H0; rewrite e; ring. +assert (Hqq0 : q0 * q = 1). +apply Zmult_reg_l with a. +assumption. +ring. +pattern a at 2 in |- *; rewrite H2; ring. +assert (q | 1). +rewrite <- Hqq0; auto with zarith. +elim (Zdivide_1 q H); intros. +rewrite H1 in H0; left; omega. +rewrite H1 in H0; right; omega. +Qed. (** If [a] divides [b] and [b<>0] then [|a| <= |b|]. *) -Lemma Zdivide_bounds : (a,b:Z) (a|b) -> `b<>0` -> `|a| <= |b|`. -Proof. -Induction 1; Intros. -Assert `|b|=|q|*|a|`. - Subst; Apply Zabs_Zmult. -Rewrite H2. -Assert H3 := (Zabs_pos q). -Assert H4 := (Zabs_pos a). -Assert `|q|*|a|>=1*|a|`; Auto with zarith. -Apply Zge_Zmult_pos_compat; Auto with zarith. -Elim (Z_lt_ge_dec `|q|` `1`); [ Intros | Auto with zarith ]. -Assert `|q|=0`. - Omega. -Assert `q=0`. - Rewrite <- (Zabs_Zsgn q). -Rewrite H5; Auto with zarith. -Subst q; Omega. -Save. +Lemma Zdivide_bounds : forall a b:Z, (a | b) -> b <> 0 -> Zabs a <= Zabs b. +Proof. +simple induction 1; intros. +assert (Zabs b = Zabs q * Zabs a). + subst; apply Zabs_Zmult. +rewrite H2. +assert (H3 := Zabs_pos q). +assert (H4 := Zabs_pos a). +assert (Zabs q * Zabs a >= 1 * Zabs a); auto with zarith. +apply Zmult_ge_compat; auto with zarith. +elim (Z_lt_ge_dec (Zabs q) 1); [ intros | auto with zarith ]. +assert (Zabs q = 0). + omega. +assert (q = 0). + rewrite <- (Zabs_Zsgn q). +rewrite H5; auto with zarith. +subst q; omega. +Qed. (** * Greatest common divisor (gcd). *) @@ -203,53 +199,54 @@ Save. expressing that [d] is a gcd of [a] and [b]. (We show later that the [gcd] is actually unique if we discard its sign.) *) -Inductive gcd [a,b,d:Z] : Prop := - gcd_intro : - (d|a) -> (d|b) -> ((x:Z) (x|a) -> (x|b) -> (x|d)) -> (gcd a b d). +Inductive Zis_gcd (a b d:Z) : Prop := + Zis_gcd_intro : + (d | a) -> + (d | b) -> (forall x:Z, (x | a) -> (x | b) -> (x | d)) -> Zis_gcd a b d. (** Trivial properties of [gcd] *) -Lemma gcd_sym : (a,b,d:Z)(gcd a b d) -> (gcd b a d). +Lemma Zis_gcd_sym : forall a b d:Z, Zis_gcd a b d -> Zis_gcd b a d. Proof. -Induction 1; Constructor; Intuition. -Save. +simple induction 1; constructor; intuition. +Qed. -Lemma gcd_0 : (a:Z)(gcd a `0` a). +Lemma Zis_gcd_0 : forall a:Z, Zis_gcd a 0 a. Proof. -Constructor; Auto with zarith. -Save. +constructor; auto with zarith. +Qed. -Lemma gcd_minus :(a,b,d:Z)(gcd a `-b` d) -> (gcd b a d). +Lemma Zis_gcd_minus : forall a b d:Z, Zis_gcd a (- b) d -> Zis_gcd b a d. Proof. -Induction 1; Constructor; Intuition. -Save. +simple induction 1; constructor; intuition. +Qed. -Lemma gcd_opp :(a,b,d:Z)(gcd a b d) -> (gcd b a `-d`). +Lemma Zis_gcd_opp : forall a b d:Z, Zis_gcd a b d -> Zis_gcd b a (- d). Proof. -Induction 1; Constructor; Intuition. -Save. +simple induction 1; constructor; intuition. +Qed. -Hints Resolve gcd_sym gcd_0 gcd_minus gcd_opp : zarith. +Hint Resolve Zis_gcd_sym Zis_gcd_0 Zis_gcd_minus Zis_gcd_opp: zarith. (** * Extended Euclid algorithm. *) (** Euclid's algorithm to compute the [gcd] mainly relies on the following property. *) -Lemma gcd_for_euclid : - (a,b,d,q:Z) (gcd b `a-q*b` d) -> (gcd a b d). +Lemma Zis_gcd_for_euclid : + forall a b d q:Z, Zis_gcd b (a - q * b) d -> Zis_gcd a b d. Proof. -Induction 1; Constructor; Intuition. -Replace a with `a-q*b+q*b`. Auto with zarith. Ring. -Save. +simple induction 1; constructor; intuition. +replace a with (a - q * b + q * b). auto with zarith. ring. +Qed. -Lemma gcd_for_euclid2 : - (b,d,q,r:Z) (gcd r b d) -> (gcd b `b*q+r` d). +Lemma Zis_gcd_for_euclid2 : + forall b d q r:Z, Zis_gcd r b d -> Zis_gcd b (b * q + r) d. Proof. -Induction 1; Constructor; Intuition. -Apply H2; Auto. -Replace r with `b*q+r-b*q`. Auto with zarith. Ring. -Save. +simple induction 1; constructor; intuition. +apply H2; auto. +replace r with (b * q + r - b * q). auto with zarith. ring. +Qed. (** We implement the extended version of Euclid's algorithm, i.e. the one computing Bezout's coefficients as it computes @@ -258,14 +255,14 @@ Save. Section extended_euclid_algorithm. -Variable a,b : Z. +Variables a b : Z. (** The specification of Euclid's algorithm is the existence of [u], [v] and [d] such that [ua+vb=d] and [(gcd a b d)]. *) Inductive Euclid : Set := - Euclid_intro : - (u,v,d:Z) `u*a+v*b=d` -> (gcd a b d) -> Euclid. + Euclid_intro : + forall u v d:Z, u * a + v * b = d -> Zis_gcd a b d -> Euclid. (** The recursive part of Euclid's algorithm uses well-founded recursion of non-negative integers. It maintains 6 integers @@ -274,356 +271,371 @@ Inductive Euclid : Set := *) Lemma euclid_rec : - (v3:Z) `0 <= v3` -> (u1,u2,u3,v1,v2:Z) `u1*a+u2*b=u3` -> `v1*a+v2*b=v3` -> - ((d:Z)(gcd u3 v3 d) -> (gcd a b d)) -> Euclid. -Proof. -Intros v3 Hv3; Generalize Hv3; Pattern v3. -Apply Z_lt_rec. -Clear v3 Hv3; Intros. -Elim (Z_zerop x); Intro. -Apply Euclid_intro with u:=u1 v:=u2 d:=u3. -Assumption. -Apply H2. -Rewrite a0; Auto with zarith. -LetTac q := (Zdiv u3 x). -Assert Hq: `0 <= u3-q*x < x`. -Replace `u3-q*x` with `u3%x`. -Apply Z_mod_lt; Omega. -Assert xpos : `x > 0`. Omega. -Generalize (Z_div_mod_eq u3 x xpos). -Unfold q. -Intro eq; Pattern 2 u3; Rewrite eq; Ring. -Apply (H `u3-q*x` Hq (proj1 ? ? Hq) v1 v2 x `u1-q*v1` `u2-q*v2`). -Tauto. -Replace `(u1-q*v1)*a+(u2-q*v2)*b` with `(u1*a+u2*b)-q*(v1*a+v2*b)`. -Rewrite H0; Rewrite H1; Trivial. -Ring. -Intros; Apply H2. -Apply gcd_for_euclid with q; Assumption. -Assumption. -Save. + forall v3:Z, + 0 <= v3 -> + forall u1 u2 u3 v1 v2:Z, + u1 * a + u2 * b = u3 -> + v1 * a + v2 * b = v3 -> + (forall d:Z, Zis_gcd u3 v3 d -> Zis_gcd a b d) -> Euclid. +Proof. +intros v3 Hv3; generalize Hv3; pattern v3 in |- *. +apply Z_lt_rec. +clear v3 Hv3; intros. +elim (Z_zerop x); intro. +apply Euclid_intro with (u := u1) (v := u2) (d := u3). +assumption. +apply H2. +rewrite a0; auto with zarith. +set (q := u3 / x) in *. +assert (Hq : 0 <= u3 - q * x < x). +replace (u3 - q * x) with (u3 mod x). +apply Z_mod_lt; omega. +assert (xpos : x > 0). omega. +generalize (Z_div_mod_eq u3 x xpos). +unfold q in |- *. +intro eq; pattern u3 at 2 in |- *; rewrite eq; ring. +apply (H (u3 - q * x) Hq (proj1 Hq) v1 v2 x (u1 - q * v1) (u2 - q * v2)). +tauto. +replace ((u1 - q * v1) * a + (u2 - q * v2) * b) with + (u1 * a + u2 * b - q * (v1 * a + v2 * b)). +rewrite H0; rewrite H1; trivial. +ring. +intros; apply H2. +apply Zis_gcd_for_euclid with q; assumption. +assumption. +Qed. (** We get Euclid's algorithm by applying [euclid_rec] on [1,0,a,0,1,b] when [b>=0] and [1,0,a,0,-1,-b] when [b<0]. *) Lemma euclid : Euclid. Proof. -Case (Z_le_gt_dec `0` b); Intro. -Intros; Apply euclid_rec with u1:=`1` u2:=`0` u3:=a - v1:=`0` v2:=`1` v3:=b; -Auto with zarith; Ring. -Intros; Apply euclid_rec with u1:=`1` u2:=`0` u3:=a - v1:=`0` v2:=`-1` v3:=`-b`; -Auto with zarith; Try Ring. -Save. +case (Z_le_gt_dec 0 b); intro. +intros; + apply euclid_rec with + (u1 := 1) (u2 := 0) (u3 := a) (v1 := 0) (v2 := 1) (v3 := b); + auto with zarith; ring. +intros; + apply euclid_rec with + (u1 := 1) (u2 := 0) (u3 := a) (v1 := 0) (v2 := -1) (v3 := - b); + auto with zarith; try ring. +Qed. End extended_euclid_algorithm. -Theorem gcd_uniqueness_apart_sign : - (a,b,d,d':Z) (gcd a b d) -> (gcd a b d') -> `d = d'` \/ `d = -d'`. +Theorem Zis_gcd_uniqueness_apart_sign : + forall a b d d':Z, Zis_gcd a b d -> Zis_gcd a b d' -> d = d' \/ d = - d'. Proof. -Induction 1. -Intros H1 H2 H3; Induction 1; Intros. -Generalize (H3 d' H4 H5); Intro Hd'd. -Generalize (H6 d H1 H2); Intro Hdd'. -Exact (Zdivide_antisym d d' Hdd' Hd'd). -Save. +simple induction 1. +intros H1 H2 H3; simple induction 1; intros. +generalize (H3 d' H4 H5); intro Hd'd. +generalize (H6 d H1 H2); intro Hdd'. +exact (Zdivide_antisym d d' Hdd' Hd'd). +Qed. (** * Bezout's coefficients *) -Inductive Bezout [a,b,d:Z] : Prop := - Bezout_intro : (u,v:Z) `u*a + v*b = d` -> (Bezout a b d). +Inductive Bezout (a b d:Z) : Prop := + Bezout_intro : forall u v:Z, u * a + v * b = d -> Bezout a b d. (** Existence of Bezout's coefficients for the [gcd] of [a] and [b] *) -Lemma gcd_bezout : (a,b,d:Z) (gcd a b d) -> (Bezout a b d). +Lemma Zis_gcd_bezout : forall a b d:Z, Zis_gcd a b d -> Bezout a b d. Proof. -Intros a b d Hgcd. -Elim (euclid a b); Intros u v d0 e g. -Generalize (gcd_uniqueness_apart_sign a b d d0 Hgcd g). -Intro H; Elim H; Clear H; Intros. -Apply Bezout_intro with u v. -Rewrite H; Assumption. -Apply Bezout_intro with `-u` `-v`. -Rewrite H; Rewrite <- e; Ring. -Save. +intros a b d Hgcd. +elim (euclid a b); intros u v d0 e g. +generalize (Zis_gcd_uniqueness_apart_sign a b d d0 Hgcd g). +intro H; elim H; clear H; intros. +apply Bezout_intro with u v. +rewrite H; assumption. +apply Bezout_intro with (- u) (- v). +rewrite H; rewrite <- e; ring. +Qed. (** gcd of [ca] and [cb] is [c gcd(a,b)]. *) -Lemma gcd_mult : (a,b,c,d:Z) (gcd a b d) -> (gcd `c*a` `c*b` `c*d`). -Proof. -Intros a b c d; Induction 1; Constructor; Intuition. -Elim (gcd_bezout a b d H); Intros. -Elim H3; Intros. -Elim H4; Intros. -Apply Zdivide_intro with `u*q+v*q0`. -Rewrite <- H5. -Replace `c*(u*a+v*b)` with `u*(c*a)+v*(c*b)`. -Rewrite H6; Rewrite H7; Ring. -Ring. -Save. +Lemma Zis_gcd_mult : + forall a b c d:Z, Zis_gcd a b d -> Zis_gcd (c * a) (c * b) (c * d). +Proof. +intros a b c d; simple induction 1; constructor; intuition. +elim (Zis_gcd_bezout a b d H); intros. +elim H3; intros. +elim H4; intros. +apply Zdivide_intro with (u * q + v * q0). +rewrite <- H5. +replace (c * (u * a + v * b)) with (u * (c * a) + v * (c * b)). +rewrite H6; rewrite H7; ring. +ring. +Qed. (** We could obtain a [Zgcd] function via [euclid]. But we propose here a more direct version of a [Zgcd], with better extraction (no bezout coeffs). *) -Definition Zgcd_pos : (a:Z)`0<=a` -> (b:Z) - { g:Z | `0<=a` -> (gcd a b g) /\ `g>=0` }. -Proof. -Intros a Ha. -Apply (Z_lt_rec [a:Z](b:Z) { g:Z | `0<=a` -> (gcd a b g) /\`g>=0` }); Try Assumption. -Intro x; Case x. -Intros _ b; Exists (Zabs b). - Elim (Z_le_lt_eq_dec ? ? (Zabs_pos b)). - Intros H0; Split. - Apply Zabs_ind. - Intros; Apply gcd_sym; Apply gcd_0; Auto. - Intros; Apply gcd_opp; Apply gcd_0; Auto. - Auto with zarith. +Definition Zgcd_pos : + forall a:Z, + 0 <= a -> forall b:Z, {g : Z | 0 <= a -> Zis_gcd a b g /\ g >= 0}. +Proof. +intros a Ha. +apply + (Z_lt_rec + (fun a:Z => forall b:Z, {g : Z | 0 <= a -> Zis_gcd a b g /\ g >= 0})); + try assumption. +intro x; case x. +intros _ b; exists (Zabs b). + elim (Z_le_lt_eq_dec _ _ (Zabs_pos b)). + intros H0; split. + apply Zabs_ind. + intros; apply Zis_gcd_sym; apply Zis_gcd_0; auto. + intros; apply Zis_gcd_opp; apply Zis_gcd_0; auto. + auto with zarith. - Intros H0; Rewrite <- H0. - Rewrite <- (Zabs_Zsgn b); Rewrite <- H0; Simpl. - Split; [Apply gcd_0|Idtac];Auto with zarith. + intros H0; rewrite <- H0. + rewrite <- (Zabs_Zsgn b); rewrite <- H0; simpl in |- *. + split; [ apply Zis_gcd_0 | idtac ]; auto with zarith. -Intros p Hrec b. -Generalize (Z_div_mod b (POS p)). -Case (Zdiv_eucl b (POS p)); Intros q r Hqr. -Elim Hqr; Clear Hqr; Intros; Auto with zarith. -Elim (Hrec r H0 (POS p)); Intros g Hgkl. -Inversion_clear H0. -Elim (Hgkl H1); Clear Hgkl; Intros H3 H4. -Exists g; Intros. -Split; Auto. -Rewrite H. -Apply gcd_for_euclid2; Auto. - -Intros p Hrec b. -Exists `0`; Intros. -Elim H; Auto. +intros p Hrec b. +generalize (Z_div_mod b (Zpos p)). +case (Zdiv_eucl b (Zpos p)); intros q r Hqr. +elim Hqr; clear Hqr; intros; auto with zarith. +elim (Hrec r H0 (Zpos p)); intros g Hgkl. +inversion_clear H0. +elim (Hgkl H1); clear Hgkl; intros H3 H4. +exists g; intros. +split; auto. +rewrite H. +apply Zis_gcd_for_euclid2; auto. + +intros p Hrec b. +exists 0; intros. +elim H; auto. Defined. -Definition Zgcd_spec : (a,b:Z){ g : Z | (gcd a b g) /\ `g>=0` }. +Definition Zgcd_spec : forall a b:Z, {g : Z | Zis_gcd a b g /\ g >= 0}. Proof. -Intros a; Case (Z_gt_le_dec `0` a). -Intros; Assert `0 <= -a`. -Omega. -Elim (Zgcd_pos `-a` H b); Intros g Hgkl. -Exists g. -Intuition. -Intros Ha b; Elim (Zgcd_pos a Ha b); Intros g; Exists g; Intuition. +intros a; case (Z_gt_le_dec 0 a). +intros; assert (0 <= - a). +omega. +elim (Zgcd_pos (- a) H b); intros g Hgkl. +exists g. +intuition. +intros Ha b; elim (Zgcd_pos a Ha b); intros g; exists g; intuition. Defined. -Definition Zgcd := [a,b:Z](let (g,_) = (Zgcd_spec a b) in g). +Definition Zgcd (a b:Z) := let (g, _) := Zgcd_spec a b in g. -Lemma Zgcd_is_pos : (a,b:Z)`(Zgcd a b) >=0`. -Intros a b; Unfold Zgcd; Case (Zgcd_spec a b); Tauto. +Lemma Zgcd_is_pos : forall a b:Z, Zgcd a b >= 0. +intros a b; unfold Zgcd in |- *; case (Zgcd_spec a b); tauto. Qed. -Lemma Zgcd_is_gcd : (a,b:Z)(gcd a b (Zgcd a b)). -Intros a b; Unfold Zgcd; Case (Zgcd_spec a b); Tauto. +Lemma Zgcd_is_gcd : forall a b:Z, Zis_gcd a b (Zgcd a b). +intros a b; unfold Zgcd in |- *; case (Zgcd_spec a b); tauto. Qed. (** * Relative primality *) -Definition rel_prime [a,b:Z] : Prop := (gcd a b `1`). +Definition rel_prime (a b:Z) : Prop := Zis_gcd a b 1. (** Bezout's theorem: [a] and [b] are relatively prime if and only if there exist [u] and [v] such that [ua+vb = 1]. *) -Lemma rel_prime_bezout : - (a,b:Z) (rel_prime a b) -> (Bezout a b `1`). +Lemma rel_prime_bezout : forall a b:Z, rel_prime a b -> Bezout a b 1. Proof. -Intros a b; Exact (gcd_bezout a b `1`). -Save. +intros a b; exact (Zis_gcd_bezout a b 1). +Qed. -Lemma bezout_rel_prime : - (a,b:Z) (Bezout a b `1`) -> (rel_prime a b). +Lemma bezout_rel_prime : forall a b:Z, Bezout a b 1 -> rel_prime a b. Proof. -Induction 1; Constructor; Auto with zarith. -Intros. Rewrite <- H0; Auto with zarith. -Save. +simple induction 1; constructor; auto with zarith. +intros. rewrite <- H0; auto with zarith. +Qed. (** Gauss's theorem: if [a] divides [bc] and if [a] and [b] are relatively prime, then [a] divides [c]. *) -Theorem Gauss : (a,b,c:Z) (a |`b*c`) -> (rel_prime a b) -> (a | c). +Theorem Gauss : forall a b c:Z, (a | b * c) -> rel_prime a b -> (a | c). Proof. -Intros. Elim (rel_prime_bezout a b H0); Intros. -Replace c with `c*1`; [ Idtac | Ring ]. -Rewrite <- H1. -Replace `c*(u*a+v*b)` with `(c*u)*a + v*(b*c)`; [ EAuto with zarith | Ring ]. -Save. +intros. elim (rel_prime_bezout a b H0); intros. +replace c with (c * 1); [ idtac | ring ]. +rewrite <- H1. +replace (c * (u * a + v * b)) with (c * u * a + v * (b * c)); + [ eauto with zarith | ring ]. +Qed. (** If [a] is relatively prime to [b] and [c], then it is to [bc] *) -Lemma rel_prime_mult : - (a,b,c:Z) (rel_prime a b) -> (rel_prime a c) -> (rel_prime a `b*c`). -Proof. -Intros a b c Hb Hc. -Elim (rel_prime_bezout a b Hb); Intros. -Elim (rel_prime_bezout a c Hc); Intros. -Apply bezout_rel_prime. -Apply Bezout_intro with u:=`u*u0*a+v0*c*u+u0*v*b` v:=`v*v0`. -Rewrite <- H. -Replace `u*a+v*b` with `(u*a+v*b) * 1`; [ Idtac | Ring ]. -Rewrite <- H0. -Ring. -Save. +Lemma rel_prime_mult : + forall a b c:Z, rel_prime a b -> rel_prime a c -> rel_prime a (b * c). +Proof. +intros a b c Hb Hc. +elim (rel_prime_bezout a b Hb); intros. +elim (rel_prime_bezout a c Hc); intros. +apply bezout_rel_prime. +apply Bezout_intro with + (u := u * u0 * a + v0 * c * u + u0 * v * b) (v := v * v0). +rewrite <- H. +replace (u * a + v * b) with ((u * a + v * b) * 1); [ idtac | ring ]. +rewrite <- H0. +ring. +Qed. Lemma rel_prime_cross_prod : - (a,b,c,d:Z) (rel_prime a b) -> (rel_prime c d) -> `b>0` -> `d>0` -> - `a*d = b*c` -> (a=c /\ b=d). -Proof. -Intros a b c d; Intros. -Elim (Zdivide_antisym b d). -Split; Auto with zarith. -Rewrite H4 in H3. -Rewrite Zmult_sym in H3. -Apply Zmult_reg_left with d; Auto with zarith. -Intros; Omega. -Apply Gauss with a. -Rewrite H3. -Auto with zarith. -Red; Auto with zarith. -Apply Gauss with c. -Rewrite Zmult_sym. -Rewrite <- H3. -Auto with zarith. -Red; Auto with zarith. -Save. + forall a b c d:Z, + rel_prime a b -> + rel_prime c d -> b > 0 -> d > 0 -> a * d = b * c -> a = c /\ b = d. +Proof. +intros a b c d; intros. +elim (Zdivide_antisym b d). +split; auto with zarith. +rewrite H4 in H3. +rewrite Zmult_comm in H3. +apply Zmult_reg_l with d; auto with zarith. +intros; omega. +apply Gauss with a. +rewrite H3. +auto with zarith. +red in |- *; auto with zarith. +apply Gauss with c. +rewrite Zmult_comm. +rewrite <- H3. +auto with zarith. +red in |- *; auto with zarith. +Qed. (** After factorization by a gcd, the original numbers are relatively prime. *) -Lemma gcd_rel_prime : - (a,b,g:Z)`b>0` -> `g>=0`-> (gcd a b g) -> (rel_prime `a/g` `b/g`). -Intros a b g; Intros. -Assert `g <> 0`. - Intro. - Elim H1; Intros. - Elim H4; Intros. - Rewrite H2 in H6; Subst b; Omega. -Unfold rel_prime. -Elim (Zgcd_spec `a/g` `b/g`); Intros g' (H3,H4). -Assert H5 := (gcd_mult ? ? g ? H3). -Rewrite <- Z_div_exact_2 in H5; Auto with zarith. -Rewrite <- Z_div_exact_2 in H5; Auto with zarith. -Elim (gcd_uniqueness_apart_sign ? ? ? ? H1 H5). -Intros; Rewrite (!Zmult_reg_left `1` g' g); Auto with zarith. -Intros; Rewrite (!Zmult_reg_left `1` `-g'` g); Auto with zarith. -Pattern 1 g; Rewrite H6; Ring. - -Elim H1; Intros. -Elim H7; Intros. -Rewrite H9. -Replace `q*g` with `0+q*g`. -Rewrite Z_mod_plus. -Compute; Auto. -Omega. -Ring. - -Elim H1; Intros. -Elim H6; Intros. -Rewrite H9. -Replace `q*g` with `0+q*g`. -Rewrite Z_mod_plus. -Compute; Auto. -Omega. -Ring. -Save. +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). +intros a b g; intros. +assert (g <> 0). + intro. + elim H1; intros. + elim H4; intros. + rewrite H2 in H6; subst b; omega. +unfold rel_prime in |- *. +elim (Zgcd_spec (a / g) (b / g)); intros g' [H3 H4]. +assert (H5 := Zis_gcd_mult _ _ g _ H3). +rewrite <- Z_div_exact_2 in H5; auto with zarith. +rewrite <- Z_div_exact_2 in H5; auto with zarith. +elim (Zis_gcd_uniqueness_apart_sign _ _ _ _ H1 H5). +intros; rewrite (Zmult_reg_l 1 g' g); auto with zarith. +intros; rewrite (Zmult_reg_l 1 (- g') g); auto with zarith. +pattern g at 1 in |- *; rewrite H6; ring. + +elim H1; intros. +elim H7; intros. +rewrite H9. +replace (q * g) with (0 + q * g). +rewrite Z_mod_plus. +compute in |- *; auto. +omega. +ring. + +elim H1; intros. +elim H6; intros. +rewrite H9. +replace (q * g) with (0 + q * g). +rewrite Z_mod_plus. +compute in |- *; auto. +omega. +ring. +Qed. (** * Primality *) -Inductive prime [p:Z] : Prop := - prime_intro : - `1 < p` -> ((n:Z) `1 <= n < p` -> (rel_prime n p)) -> (prime p). +Inductive prime (p:Z) : Prop := + prime_intro : + 1 < p -> (forall n:Z, 1 <= n < p -> rel_prime n p) -> prime p. (** The sole divisors of a prime number [p] are [-1], [1], [p] and [-p]. *) -Lemma prime_divisors : - (p:Z) (prime p) -> - (a:Z) (a|p) -> `a = -1` \/ `a = 1` \/ a = p \/ `a = -p`. -Proof. -Induction 1; Intros. -Assert `a = (-p)`\/`-p forall a:Z, (a | p) -> a = -1 \/ a = 1 \/ a = p \/ a = - p. +Proof. +simple induction 1; intros. +assert + (a = - p \/ - p < a < -1 \/ a = -1 \/ a = 0 \/ a = 1 \/ 1 < a < p \/ a = p). +assert (Zabs a <= Zabs p). apply Zdivide_bounds; [ assumption | omega ]. +generalize H3. +pattern (Zabs a) in |- *; apply Zabs_ind; pattern (Zabs p) in |- *; + apply Zabs_ind; intros; omega. +intuition idtac. (* -p < a < -1 *) -Absurd (rel_prime `-a` p); Intuition. -Inversion H3. -Assert (`-a` | `-a`); Auto with zarith. -Assert (`-a` | p); Auto with zarith. -Generalize (H8 `-a` H9 H10); Intuition Idtac. -Generalize (Zdivide_1 `-a` H11); Intuition. +absurd (rel_prime (- a) p); intuition. +inversion H3. +assert (- a | - a); auto with zarith. +assert (- a | p); auto with zarith. +generalize (H8 (- a) H9 H10); intuition idtac. +generalize (Zdivide_1 (- a) H11); intuition. (* a = 0 *) -Inversion H2. Subst a; Omega. +inversion H2. subst a; omega. (* 1 < a < p *) -Absurd (rel_prime a p); Intuition. -Inversion H3. -Assert (a | a); Auto with zarith. -Assert (a | p); Auto with zarith. -Generalize (H8 a H9 H10); Intuition Idtac. -Generalize (Zdivide_1 a H11); Intuition. -Save. +absurd (rel_prime a p); intuition. +inversion H3. +assert (a | a); auto with zarith. +assert (a | p); auto with zarith. +generalize (H8 a H9 H10); intuition idtac. +generalize (Zdivide_1 a H11); intuition. +Qed. (** A prime number is relatively prime with any number it does not divide *) -Lemma prime_rel_prime : - (p:Z) (prime p) -> (a:Z) ~ (p|a) -> (rel_prime p a). +Lemma prime_rel_prime : + forall p:Z, prime p -> forall a:Z, ~ (p | a) -> rel_prime p a. Proof. -Induction 1; Intros. -Constructor; Intuition. -Elim (prime_divisors p H x H3); Intuition; Subst; Auto with zarith. -Absurd (p | a); Auto with zarith. -Absurd (p | a); Intuition. -Save. +simple induction 1; intros. +constructor; intuition. +elim (prime_divisors p H x H3); intuition; subst; auto with zarith. +absurd (p | a); auto with zarith. +absurd (p | a); intuition. +Qed. -Hints Resolve prime_rel_prime : zarith. +Hint Resolve prime_rel_prime: zarith. (** [Zdivide] can be expressed using [Zmod]. *) -Lemma Zmod_Zdivide : (a,b:Z) `b>0` -> `a%b = 0` -> (b|a). -Intros a b H H0. -Apply Zdivide_intro with `(a/b)`. -Pattern 1 a; Rewrite (Z_div_mod_eq a b H). -Rewrite H0; Ring. -Save. +Lemma Zmod_divide : forall a b:Z, b > 0 -> a mod b = 0 -> (b | a). +intros a b H H0. +apply Zdivide_intro with (a / b). +pattern a at 1 in |- *; rewrite (Z_div_mod_eq a b H). +rewrite H0; ring. +Qed. -Lemma Zdivide_Zmod : (a,b:Z) `b>0` -> (b|a) -> `a%b = 0`. -Intros a b; Destruct 2; Intros; Subst. -Change `q*b` with `0+q*b`. -Rewrite Z_mod_plus; Auto. -Save. +Lemma Zdivide_mod : forall a b:Z, b > 0 -> (b | a) -> a mod b = 0. +intros a b; simple destruct 2; intros; subst. +change (q * b) with (0 + q * b) in |- *. +rewrite Z_mod_plus; auto. +Qed. (** [Zdivide] is hence decidable *) -Lemma Zdivide_dec : (a,b:Z) { (a|b) } + { ~ (a|b) }. +Lemma Zdivide_dec : forall a b:Z, {(a | b)} + {~ (a | b)}. Proof. -Intros a b; Elim (Ztrichotomy_inf a `0`). +intros a b; elim (Ztrichotomy_inf a 0). (* a<0 *) -Intros H; Elim H; Intros. -Case (Z_eq_dec `b%(-a)` `0`). -Left; Apply Zdivide_opp_left_rev; Apply Zmod_Zdivide; Auto with zarith. -Intro H1; Right; Intro; Elim H1; Apply Zdivide_Zmod; Auto with zarith. +intros H; elim H; intros. +case (Z_eq_dec (b mod - a) 0). +left; apply Zdivide_opp_l_rev; apply Zmod_divide; auto with zarith. +intro H1; right; intro; elim H1; apply Zdivide_mod; auto with zarith. (* a=0 *) -Case (Z_eq_dec b `0`); Intro. -Left; Subst; Auto with zarith. -Right; Subst; Intro H0; Inversion H0; Omega. +case (Z_eq_dec b 0); intro. +left; subst; auto with zarith. +right; subst; intro H0; inversion H0; omega. (* a>0 *) -Intro H; Case (Z_eq_dec `b%a` `0`). -Left; Apply Zmod_Zdivide; Auto with zarith. -Intro H1; Right; Intro; Elim H1; Apply Zdivide_Zmod; Auto with zarith. -Save. +intro H; case (Z_eq_dec (b mod a) 0). +left; apply Zmod_divide; auto with zarith. +intro H1; right; intro; elim H1; apply Zdivide_mod; auto with zarith. +Qed. (** If a prime [p] divides [ab] then it divides either [a] or [b] *) -Lemma prime_mult : - (p:Z) (prime p) -> (a,b:Z) (p | `a*b`) -> (p | a) \/ (p | b). +Lemma prime_mult : + forall p:Z, prime p -> forall a b:Z, (p | a * b) -> (p | a) \/ (p | b). Proof. -Intro p; Induction 1; Intros. -Case (Zdivide_dec p a); Intuition. -Right; Apply Gauss with a; Auto with zarith. -Save. - +intro p; simple induction 1; intros. +case (Zdivide_dec p a); intuition. +right; apply Gauss with a; auto with zarith. +Qed. diff --git a/theories/ZArith/Zorder.v b/theories/ZArith/Zorder.v index bfe56b82e5..eeb9f681b0 100644 --- a/theories/ZArith/Zorder.v +++ b/theories/ZArith/Zorder.v @@ -9,961 +9,957 @@ (** Binary Integers (Pierre Crégut (CNET, Lannion, France) *) -Require BinPos. -Require BinInt. -Require Arith. -Require Decidable. -Require Zsyntax. -Require Zcompare. - -V7only [Import nat_scope.]. +Require Import BinPos. +Require Import BinInt. +Require Import Arith. +Require Import Decidable. +Require Import Zcompare. + Open Local Scope Z_scope. -Implicit Variable Type x,y,z:Z. +Implicit Types x y z : Z. (**********************************************************************) (** Properties of the order relations on binary integers *) (** Trichotomy *) -Theorem Ztrichotomy_inf : (m,n:Z) {`mn`}. +Theorem Ztrichotomy_inf : forall n m:Z, {n < m} + {n = m} + {n > m}. Proof. -Unfold Zgt Zlt; Intros m n; Assert H:=(refl_equal ? (Zcompare m n)). - LetTac x := (Zcompare m n) in 2 H Goal. - NewDestruct x; - [Left; Right;Rewrite Zcompare_EGAL_eq with 1:=H - | Left; Left - | Right ]; Reflexivity. +unfold Zgt, Zlt in |- *; intros m n; assert (H := refl_equal (m ?= n)). + set (x := m ?= n) in H at 2 |- *. + destruct x; + [ left; right; rewrite Zcompare_Eq_eq with (1 := H) | left; left | right ]; + reflexivity. Qed. -Theorem Ztrichotomy : (m,n:Z) `mn`. +Theorem Ztrichotomy : forall n m:Z, n < m \/ n = m \/ n > m. Proof. - Intros m n; NewDestruct (Ztrichotomy_inf m n) as [[Hlt|Heq]|Hgt]; - [Left | Right; Left |Right; Right]; Assumption. + intros m n; destruct (Ztrichotomy_inf m n) as [[Hlt| Heq]| Hgt]; + [ left | right; left | right; right ]; assumption. Qed. (**********************************************************************) (** Decidability of equality and order on Z *) -Theorem dec_eq: (x,y:Z) (decidable (x=y)). +Theorem dec_eq : forall n m:Z, decidable (n = m). Proof. -Intros x y; Unfold decidable ; Elim (Zcompare_EGAL x y); -Intros H1 H2; Elim (Dcompare (Zcompare x y)); [ - Tauto - | Intros H3; Right; Unfold not ; Intros H4; - Elim H3; Rewrite (H2 H4); Intros H5; Discriminate H5]. +intros x y; unfold decidable in |- *; elim (Zcompare_Eq_iff_eq x y); + intros H1 H2; elim (Dcompare (x ?= y)); + [ tauto + | intros H3; right; unfold not in |- *; intros H4; elim H3; rewrite (H2 H4); + intros H5; discriminate H5 ]. Qed. -Theorem dec_Zne: (x,y:Z) (decidable (Zne x y)). +Theorem dec_Zne : forall n m:Z, decidable (Zne n m). Proof. -Intros x y; Unfold decidable Zne ; Elim (Zcompare_EGAL x y). -Intros H1 H2; Elim (Dcompare (Zcompare x y)); - [ Right; Rewrite H1; Auto - | Left; Unfold not; Intro; Absurd (Zcompare x y)=EGAL; - [ Elim H; Intros HR; Rewrite HR; Discriminate - | Auto]]. +intros x y; unfold decidable, Zne in |- *; elim (Zcompare_Eq_iff_eq x y). +intros H1 H2; elim (Dcompare (x ?= y)); + [ right; rewrite H1; auto + | left; unfold not in |- *; intro; absurd ((x ?= y) = Eq); + [ elim H; intros HR; rewrite HR; discriminate | auto ] ]. Qed. -Theorem dec_Zle: (x,y:Z) (decidable `x<=y`). +Theorem dec_Zle : forall n m:Z, decidable (n <= m). Proof. -Intros x y; Unfold decidable Zle ; Elim (Zcompare x y); [ - Left; Discriminate - | Left; Discriminate - | Right; Unfold not ; Intros H; Apply H; Trivial with arith]. +intros x y; unfold decidable, Zle in |- *; elim (x ?= y); + [ left; discriminate + | left; discriminate + | right; unfold not in |- *; intros H; apply H; trivial with arith ]. Qed. -Theorem dec_Zgt: (x,y:Z) (decidable `x>y`). +Theorem dec_Zgt : forall n m:Z, decidable (n > m). Proof. -Intros x y; Unfold decidable Zgt ; Elim (Zcompare x y); - [ Right; Discriminate | Right; Discriminate | Auto with arith]. +intros x y; unfold decidable, Zgt in |- *; elim (x ?= y); + [ right; discriminate | right; discriminate | auto with arith ]. Qed. -Theorem dec_Zge: (x,y:Z) (decidable `x>=y`). +Theorem dec_Zge : forall n m:Z, decidable (n >= m). Proof. -Intros x y; Unfold decidable Zge ; Elim (Zcompare x y); [ - Left; Discriminate -| Right; Unfold not ; Intros H; Apply H; Trivial with arith -| Left; Discriminate]. +intros x y; unfold decidable, Zge in |- *; elim (x ?= y); + [ left; discriminate + | right; unfold not in |- *; intros H; apply H; trivial with arith + | left; discriminate ]. Qed. -Theorem dec_Zlt: (x,y:Z) (decidable `x `x m -> n < m \/ m < n. Proof. -Intros x y; Elim (Dcompare (Zcompare x y)); [ - Intros H1 H2; Absurd x=y; [ Assumption | Elim (Zcompare_EGAL x y); Auto with arith] -| Unfold Zlt ; Intros H; Elim H; Intros H1; - [Auto with arith | Right; Elim (Zcompare_ANTISYM x y); Auto with arith]]. +intros x y; elim (Dcompare (x ?= y)); + [ intros H1 H2; absurd (x = y); + [ assumption | elim (Zcompare_Eq_iff_eq x y); auto with arith ] + | unfold Zlt in |- *; intros H; elim H; intros H1; + [ auto with arith + | right; elim (Zcompare_Gt_Lt_antisym x y); auto with arith ] ]. Qed. (** Relating strict and large orders *) -Lemma Zgt_lt : (m,n:Z) `m>n` -> `n m -> m < n. Proof. -Unfold Zgt Zlt ;Intros m n H; Elim (Zcompare_ANTISYM m n); Auto with arith. +unfold Zgt, Zlt in |- *; intros m n H; elim (Zcompare_Gt_Lt_antisym m n); + auto with arith. Qed. -Lemma Zlt_gt : (m,n:Z) `m `n>m`. +Lemma Zlt_gt : forall n m:Z, n < m -> m > n. Proof. -Unfold Zgt Zlt ;Intros m n H; Elim (Zcompare_ANTISYM n m); Auto with arith. +unfold Zgt, Zlt in |- *; intros m n H; elim (Zcompare_Gt_Lt_antisym n m); + auto with arith. Qed. -Lemma Zge_le : (m,n:Z) `m>=n` -> `n<=m`. +Lemma Zge_le : forall n m:Z, n >= m -> m <= n. Proof. -Intros m n; Change ~`m ~`n>m`; -Unfold not; Intros H1 H2; Apply H1; Apply Zgt_lt; Assumption. +intros m n; change (~ m < n -> ~ n > m) in |- *; unfold not in |- *; + intros H1 H2; apply H1; apply Zgt_lt; assumption. Qed. -Lemma Zle_ge : (m,n:Z) `m<=n` -> `n>=m`. +Lemma Zle_ge : forall n m:Z, n <= m -> m >= n. Proof. -Intros m n; Change ~`m>n`-> ~`n n -> ~ n < m) in |- *; unfold not in |- *; + intros H1 H2; apply H1; apply Zlt_gt; assumption. Qed. -Lemma Zle_not_gt : (n,m:Z)`n<=m` -> ~`n>m`. +Lemma Zle_not_gt : forall n m:Z, n <= m -> ~ n > m. Proof. -Trivial. +trivial. Qed. -Lemma Zgt_not_le : (n,m:Z)`n>m` -> ~`n<=m`. +Lemma Zgt_not_le : forall n m:Z, n > m -> ~ n <= m. Proof. -Intros n m H1 H2; Apply H2; Assumption. +intros n m H1 H2; apply H2; assumption. Qed. -Lemma Zle_not_lt : (n,m:Z)`n<=m` -> ~`m ~ m < n. Proof. -Intros n m H1 H2. -Assert H3:=(Zlt_gt ? ? H2). -Apply Zle_not_gt with n m; Assumption. +intros n m H1 H2. +assert (H3 := Zlt_gt _ _ H2). +apply Zle_not_gt with n m; assumption. Qed. -Lemma Zlt_not_le : (n,m:Z)`n ~`m<=n`. +Lemma Zlt_not_le : forall n m:Z, n < m -> ~ m <= n. Proof. -Intros n m H1 H2. -Apply Zle_not_lt with m n; Assumption. +intros n m H1 H2. +apply Zle_not_lt with m n; assumption. Qed. -Lemma not_Zge : (x,y:Z) ~`x>=y` -> `x= m -> n < m. Proof. -Unfold Zge Zlt ; Intros x y H; Apply dec_not_not; - [ Exact (dec_Zlt x y) | Assumption]. +unfold Zge, Zlt in |- *; intros x y H; apply dec_not_not; + [ exact (dec_Zlt x y) | assumption ]. Qed. -Lemma not_Zlt : (x,y:Z) ~`x `x>=y`. +Lemma Znot_lt_ge : forall n m:Z, ~ n < m -> n >= m. Proof. -Unfold Zlt Zge; Auto with arith. +unfold Zlt, Zge in |- *; auto with arith. Qed. -Lemma not_Zgt : (x,y:Z)~`x>y` -> `x<=y`. +Lemma Znot_gt_le : forall n m:Z, ~ n > m -> n <= m. Proof. -Trivial. +trivial. Qed. -Lemma not_Zle : (x,y:Z) ~`x<=y` -> `x>y`. +Lemma Znot_le_gt : forall n m:Z, ~ n <= m -> n > m. Proof. -Unfold Zle Zgt ; Intros x y H; Apply dec_not_not; - [ Exact (dec_Zgt x y) | Assumption]. +unfold Zle, Zgt in |- *; intros x y H; apply dec_not_not; + [ exact (dec_Zgt x y) | assumption ]. Qed. -Lemma Zge_iff_le : (x,y:Z) `x>=y` <-> `y<=x`. +Lemma Zge_iff_le : forall n m:Z, n >= m <-> m <= n. Proof. - Intros x y; Intros. Split. Intro. Apply Zge_le. Assumption. - Intro. Apply Zle_ge. Assumption. + intros x y; intros. split. intro. apply Zge_le. assumption. + intro. apply Zle_ge. assumption. Qed. -Lemma Zgt_iff_lt : (x,y:Z) `x>y` <-> `y m <-> m < n. Proof. - Intros x y. Split. Intro. Apply Zgt_lt. Assumption. - Intro. Apply Zlt_gt. Assumption. + intros x y. split. intro. apply Zgt_lt. assumption. + intro. apply Zlt_gt. assumption. Qed. (** Reflexivity *) -Lemma Zle_n : (n:Z) (Zle n n). +Lemma Zle_refl : forall n:Z, n <= n. Proof. -Intros n; Unfold Zle; Rewrite (Zcompare_x_x n); Discriminate. +intros n; unfold Zle in |- *; rewrite (Zcompare_refl n); discriminate. Qed. -Lemma Zle_refl : (n,m:Z) n=m -> `n<=m`. +Lemma Zeq_le : forall n m:Z, n = m -> n <= m. Proof. -Intros; Rewrite H; Apply Zle_n. +intros; rewrite H; apply Zle_refl. Qed. -Hints Resolve Zle_n : zarith. +Hint Resolve Zle_refl: zarith. (** Antisymmetry *) -Lemma Zle_antisym : (n,m:Z)`n<=m`->`m<=n`->n=m. +Lemma Zle_antisym : forall n m:Z, n <= m -> m <= n -> n = m. Proof. -Intros n m H1 H2; NewDestruct (Ztrichotomy n m) as [Hlt|[Heq|Hgt]]. - Absurd `m>n`; [ Apply Zle_not_gt | Apply Zlt_gt]; Assumption. - Assumption. - Absurd `n>m`; [ Apply Zle_not_gt | Idtac]; Assumption. +intros n m H1 H2; destruct (Ztrichotomy n m) as [Hlt| [Heq| Hgt]]. + absurd (m > n); [ apply Zle_not_gt | apply Zlt_gt ]; assumption. + assumption. + absurd (n > m); [ apply Zle_not_gt | idtac ]; assumption. Qed. (** Asymmetry *) -Lemma Zgt_not_sym : (n,m:Z)`n>m` -> ~`m>n`. +Lemma Zgt_asym : forall n m:Z, n > m -> ~ m > n. Proof. -Unfold Zgt ;Intros n m H; Elim (Zcompare_ANTISYM n m); Intros H1 H2; -Rewrite -> H1; [ Discriminate | Assumption ]. +unfold Zgt in |- *; intros n m H; elim (Zcompare_Gt_Lt_antisym n m); + intros H1 H2; rewrite H1; [ discriminate | assumption ]. Qed. -Lemma Zlt_not_sym : (n,m:Z)`n ~`m ~ m < n. Proof. -Intros n m H H1; -Assert H2:`m>n`. Apply Zlt_gt; Assumption. -Assert H3: `n>m`. Apply Zlt_gt; Assumption. -Apply Zgt_not_sym with m n; Assumption. +intros n m H H1; assert (H2 : m > n). apply Zlt_gt; assumption. +assert (H3 : n > m). apply Zlt_gt; assumption. +apply Zgt_asym with m n; assumption. Qed. (** Irreflexivity *) -Lemma Zgt_antirefl : (n:Z)~`n>n`. +Lemma Zgt_irrefl : forall n:Z, ~ n > n. Proof. -Intros n H; Apply (Zgt_not_sym n n H H). +intros n H; apply (Zgt_asym n n H H). Qed. -Lemma Zlt_n_n : (n:Z)~`n ~x=y. +Lemma Zlt_not_eq : forall n m:Z, n < m -> n <> m. Proof. -Unfold not; Intros x y H H0. -Rewrite H0 in H. -Apply (Zlt_n_n ? H). +unfold not in |- *; intros x y H H0. +rewrite H0 in H. +apply (Zlt_irrefl _ H). Qed. (** Large = strict or equal *) -Lemma Zlt_le_weak : (n,m:Z)`n`n<=m`. +Lemma Zlt_le_weak : forall n m:Z, n < m -> n <= m. Proof. -Intros n m Hlt; Apply not_Zgt; Apply Zgt_not_sym; Apply Zlt_gt; Assumption. +intros n m Hlt; apply Znot_gt_le; apply Zgt_asym; apply Zlt_gt; assumption. Qed. -Lemma Zle_lt_or_eq : (n,m:Z)`n<=m`->(`n n < m \/ n = m. Proof. -Intros n m H; NewDestruct (Ztrichotomy n m) as [Hlt|[Heq|Hgt]]; [ - Left; Assumption -| Right; Assumption -| Absurd `n>m`; [Apply Zle_not_gt|Idtac]; Assumption ]. +intros n m H; destruct (Ztrichotomy n m) as [Hlt| [Heq| Hgt]]; + [ left; assumption + | right; assumption + | absurd (n > m); [ apply Zle_not_gt | idtac ]; assumption ]. Qed. (** Dichotomy *) -Lemma Zle_or_lt : (n,m:Z)`n<=m`\/`mm`->`m>p`->`n>p`. +Lemma Zgt_trans : forall n m p:Z, n > m -> m > p -> n > p. Proof. -Exact Zcompare_trans_SUPERIEUR. +exact Zcompare_Gt_trans. Qed. -Lemma Zlt_trans : (n,m,p:Z)`n`m`n m < p -> n < p. Proof. -Intros n m p H1 H2; Apply Zgt_lt; Apply Zgt_trans with m:= m; -Apply Zlt_gt; Assumption. +intros n m p H1 H2; apply Zgt_lt; apply Zgt_trans with (m := m); apply Zlt_gt; + assumption. Qed. (** Mixed transitivity *) -Lemma Zle_gt_trans : (n,m,p:Z)`m<=n`->`m>p`->`n>p`. +Lemma Zle_gt_trans : forall n m p:Z, m <= n -> m > p -> n > p. Proof. -Intros n m p H1 H2; NewDestruct (Zle_lt_or_eq m n H1) as [Hlt|Heq]; [ - Apply Zgt_trans with m; [Apply Zlt_gt; Assumption | Assumption ] -| Rewrite <- Heq; Assumption ]. +intros n m p H1 H2; destruct (Zle_lt_or_eq m n H1) as [Hlt| Heq]; + [ apply Zgt_trans with m; [ apply Zlt_gt; assumption | assumption ] + | rewrite <- Heq; assumption ]. Qed. -Lemma Zgt_le_trans : (n,m,p:Z)`n>m`->`p<=m`->`n>p`. +Lemma Zgt_le_trans : forall n m p:Z, n > m -> p <= m -> n > p. Proof. -Intros n m p H1 H2; NewDestruct (Zle_lt_or_eq p m H2) as [Hlt|Heq]; [ - Apply Zgt_trans with m; [Assumption|Apply Zlt_gt; Assumption] -| Rewrite Heq; Assumption ]. +intros n m p H1 H2; destruct (Zle_lt_or_eq p m H2) as [Hlt| Heq]; + [ apply Zgt_trans with m; [ assumption | apply Zlt_gt; assumption ] + | rewrite Heq; assumption ]. Qed. -Lemma Zlt_le_trans : (n,m,p:Z)`n`m<=p`->`n m <= p -> n < p. +intros n m p H1 H2; apply Zgt_lt; apply Zle_gt_trans with (m := m); + [ assumption | apply Zlt_gt; assumption ]. Qed. -Lemma Zle_lt_trans : (n,m,p:Z)`n<=m`->`m`n m < p -> n < p. Proof. -Intros n m p H1 H2;Apply Zgt_lt;Apply Zgt_le_trans with m:=m; - [ Apply Zlt_gt;Assumption | Assumption ]. +intros n m p H1 H2; apply Zgt_lt; apply Zgt_le_trans with (m := m); + [ apply Zlt_gt; assumption | assumption ]. Qed. (** Transitivity of large orders *) -Lemma Zle_trans : (n,m,p:Z)`n<=m`->`m<=p`->`n<=p`. +Lemma Zle_trans : forall n m p:Z, n <= m -> m <= p -> n <= p. Proof. -Intros n m p H1 H2; Apply not_Zgt. -Intro Hgt; Apply Zle_not_gt with n m. Assumption. -Exact (Zgt_le_trans n p m Hgt H2). +intros n m p H1 H2; apply Znot_gt_le. +intro Hgt; apply Zle_not_gt with n m. assumption. +exact (Zgt_le_trans n p m Hgt H2). Qed. -Lemma Zge_trans : (n, m, p : Z) `n>=m` -> `m>=p` -> `n>=p`. +Lemma Zge_trans : forall n m p:Z, n >= m -> m >= p -> n >= p. Proof. -Intros n m p H1 H2. -Apply Zle_ge. -Apply Zle_trans with m; Apply Zge_le; Trivial. +intros n m p H1 H2. +apply Zle_ge. +apply Zle_trans with m; apply Zge_le; trivial. Qed. -Hints Resolve Zle_trans : zarith. +Hint Resolve Zle_trans: zarith. (** Compatibility of successor wrt to order *) -Lemma Zle_n_S : (n,m:Z) `m<=n` -> `(Zs m)<=(Zs n)`. +Lemma Zsucc_le_compat : forall n m:Z, m <= n -> Zsucc m <= Zsucc n. Proof. -Unfold Zle not ;Intros m n H1 H2; Apply H1; -Rewrite <- (Zcompare_Zplus_compatible n m (POS xH)); -Do 2 Rewrite (Zplus_sym (POS xH)); Exact H2. +unfold Zle, not in |- *; intros m n H1 H2; apply H1; + rewrite <- (Zcompare_plus_compat n m 1); do 2 rewrite (Zplus_comm 1); + exact H2. Qed. -Lemma Zgt_n_S : (n,m:Z)`m>n` -> `(Zs m)>(Zs n)`. +Lemma Zsucc_gt_compat : forall n m:Z, m > n -> Zsucc m > Zsucc n. Proof. -Unfold Zgt; Intros n m H; Rewrite Zcompare_n_S; Auto with arith. +unfold Zgt in |- *; intros n m H; rewrite Zcompare_succ_compat; + auto with arith. Qed. -Lemma Zlt_n_S : (n,m:Z)`n`(Zs n)<(Zs m)`. +Lemma Zsucc_lt_compat : forall n m:Z, n < m -> Zsucc n < Zsucc m. Proof. -Intros n m H;Apply Zgt_lt;Apply Zgt_n_S;Apply Zlt_gt; Assumption. +intros n m H; apply Zgt_lt; apply Zsucc_gt_compat; apply Zlt_gt; assumption. Qed. -Hints Resolve Zle_n_S : zarith. +Hint Resolve Zsucc_le_compat: zarith. (** Simplification of successor wrt to order *) -Lemma Zgt_S_n : (n,p:Z)`(Zs p)>(Zs n)`->`p>n`. +Lemma Zsucc_gt_reg : forall n m:Z, Zsucc m > Zsucc n -> m > n. Proof. -Unfold Zs Zgt;Intros n p;Do 2 Rewrite -> [m:Z](Zplus_sym m (POS xH)); -Rewrite -> (Zcompare_Zplus_compatible p n (POS xH));Trivial with arith. +unfold Zsucc, Zgt in |- *; intros n p; + do 2 rewrite (fun m:Z => Zplus_comm m 1); + rewrite (Zcompare_plus_compat p n 1); trivial with arith. Qed. -Lemma Zle_S_n : (n,m:Z) `(Zs m)<=(Zs n)` -> `m<=n`. +Lemma Zsucc_le_reg : forall n m:Z, Zsucc m <= Zsucc n -> m <= n. Proof. -Unfold Zle not ;Intros m n H1 H2;Apply H1; -Unfold Zs ;Do 2 Rewrite <- (Zplus_sym (POS xH)); -Rewrite -> (Zcompare_Zplus_compatible n m (POS xH));Assumption. +unfold Zle, not in |- *; intros m n H1 H2; apply H1; unfold Zsucc in |- *; + do 2 rewrite <- (Zplus_comm 1); rewrite (Zcompare_plus_compat n m 1); + assumption. Qed. -Lemma Zlt_S_n : (n,m:Z)`(Zs n)<(Zs m)`->`n n < m. Proof. -Intros n m H;Apply Zgt_lt;Apply Zgt_S_n;Apply Zlt_gt; Assumption. +intros n m H; apply Zgt_lt; apply Zsucc_gt_reg; apply Zlt_gt; assumption. Qed. (** Compatibility of addition wrt to order *) -Lemma Zgt_reg_l : (n,m,p:Z)`n>m`->`p+n>p+m`. +Lemma Zplus_gt_compat_l : forall n m p:Z, n > m -> p + n > p + m. Proof. -Unfold Zgt; Intros n m p H; Rewrite (Zcompare_Zplus_compatible n m p); -Assumption. +unfold Zgt in |- *; intros n m p H; rewrite (Zcompare_plus_compat n m p); + assumption. Qed. -Lemma Zgt_reg_r : (n,m,p:Z)`n>m`->`n+p>m+p`. +Lemma Zplus_gt_compat_r : forall n m p:Z, n > m -> n + p > m + p. Proof. -Intros n m p H; Rewrite (Zplus_sym n p); Rewrite (Zplus_sym m p); Apply Zgt_reg_l; Trivial. +intros n m p H; rewrite (Zplus_comm n p); rewrite (Zplus_comm m p); + apply Zplus_gt_compat_l; trivial. Qed. -Lemma Zle_reg_l : (n,m,p:Z)`n<=m`->`p+n<=p+m`. +Lemma Zplus_le_compat_l : forall n m p:Z, n <= m -> p + n <= p + m. Proof. -Intros n m p; Unfold Zle not ;Intros H1 H2;Apply H1; -Rewrite <- (Zcompare_Zplus_compatible n m p); Assumption. +intros n m p; unfold Zle, not in |- *; intros H1 H2; apply H1; + rewrite <- (Zcompare_plus_compat n m p); assumption. Qed. -Lemma Zle_reg_r : (n,m,p:Z) `n<=m`->`n+p<=m+p`. +Lemma Zplus_le_compat_r : forall n m p:Z, n <= m -> n + p <= m + p. Proof. -Intros a b c;Do 2 Rewrite [n:Z](Zplus_sym n c); Exact (Zle_reg_l a b c). +intros a b c; do 2 rewrite (fun n:Z => Zplus_comm n c); + exact (Zplus_le_compat_l a b c). Qed. -Lemma Zlt_reg_l : (n,m,p:Z)`n`p+n p + n < p + m. Proof. -Unfold Zlt ;Intros n m p; Rewrite Zcompare_Zplus_compatible;Trivial with arith. +unfold Zlt in |- *; intros n m p; rewrite Zcompare_plus_compat; + trivial with arith. Qed. -Lemma Zlt_reg_r : (n,m,p:Z)`n`n+p n + p < m + p. Proof. -Intros n m p H; Rewrite (Zplus_sym n p); Rewrite (Zplus_sym m p); Apply Zlt_reg_l; Trivial. +intros n m p H; rewrite (Zplus_comm n p); rewrite (Zplus_comm m p); + apply Zplus_lt_compat_l; trivial. Qed. -Lemma Zlt_le_reg : (a,b,c,d:Z) `a`c<=d`->`a+c p <= q -> n + p < m + q. Proof. -Intros a b c d H0 H1. -Apply Zlt_le_trans with (Zplus b c). -Apply Zlt_reg_r; Trivial. -Apply Zle_reg_l; Trivial. +intros a b c d H0 H1. +apply Zlt_le_trans with (b + c). +apply Zplus_lt_compat_r; trivial. +apply Zplus_le_compat_l; trivial. Qed. -Lemma Zle_lt_reg : (a,b,c,d:Z) `a<=b`->`c`a+c p < q -> n + p < m + q. Proof. -Intros a b c d H0 H1. -Apply Zle_lt_trans with (Zplus b c). -Apply Zle_reg_r; Trivial. -Apply Zlt_reg_l; Trivial. +intros a b c d H0 H1. +apply Zle_lt_trans with (b + c). +apply Zplus_le_compat_r; trivial. +apply Zplus_lt_compat_l; trivial. Qed. -Lemma Zle_plus_plus : (n,m,p,q:Z) `n<=m`->(Zle p q)->`n+p<=m+q`. +Lemma Zplus_le_compat : forall n m p q:Z, n <= m -> p <= q -> n + p <= m + q. Proof. -Intros n m p q; Intros H1 H2;Apply Zle_trans with m:=(Zplus n q); [ - Apply Zle_reg_l;Assumption | Apply Zle_reg_r;Assumption ]. +intros n m p q; intros H1 H2; apply Zle_trans with (m := n + q); + [ apply Zplus_le_compat_l; assumption + | apply Zplus_le_compat_r; assumption ]. Qed. -V7only [Set Implicit Arguments.]. -Lemma Zlt_Zplus : (x1,x2,y1,y2:Z)`x1 < x2` -> `y1 < y2` -> `x1 + y1 < x2 + y2`. -Intros; Apply Zle_lt_reg. Apply Zlt_le_weak; Assumption. Assumption. +Lemma Zplus_lt_compat : forall n m p q:Z, n < m -> p < q -> n + p < m + q. +intros; apply Zplus_le_lt_compat. apply Zlt_le_weak; assumption. assumption. Qed. -V7only [Unset Implicit Arguments.]. (** Compatibility of addition wrt to being positive *) -Lemma Zle_0_plus : (x,y:Z) `0<=x` -> `0<=y` -> `0<=x+y`. +Lemma Zplus_le_0_compat : forall n m:Z, 0 <= n -> 0 <= m -> 0 <= n + m. Proof. -Intros x y H1 H2;Rewrite <- (Zero_left ZERO); Apply Zle_plus_plus; Assumption. +intros x y H1 H2; rewrite <- (Zplus_0_l 0); apply Zplus_le_compat; assumption. Qed. (** Simplification of addition wrt to order *) -Lemma Zsimpl_gt_plus_l : (n,m,p:Z)`p+n>p+m`->`n>m`. +Lemma Zplus_gt_reg_l : forall n m p:Z, p + n > p + m -> n > m. Proof. -Unfold Zgt; Intros n m p H; - Rewrite <- (Zcompare_Zplus_compatible n m p); Assumption. +unfold Zgt in |- *; intros n m p H; rewrite <- (Zcompare_plus_compat n m p); + assumption. Qed. -Lemma Zsimpl_gt_plus_r : (n,m,p:Z)`n+p>m+p`->`n>m`. +Lemma Zplus_gt_reg_r : forall n m p:Z, n + p > m + p -> n > m. Proof. -Intros n m p H; Apply Zsimpl_gt_plus_l with p. -Rewrite (Zplus_sym p n); Rewrite (Zplus_sym p m); Trivial. +intros n m p H; apply Zplus_gt_reg_l with p. +rewrite (Zplus_comm p n); rewrite (Zplus_comm p m); trivial. Qed. -Lemma Zsimpl_le_plus_l : (n,m,p:Z)`p+n<=p+m`->`n<=m`. +Lemma Zplus_le_reg_l : forall n m p:Z, p + n <= p + m -> n <= m. Proof. -Intros n m p; Unfold Zle not ;Intros H1 H2;Apply H1; -Rewrite (Zcompare_Zplus_compatible n m p); Assumption. +intros n m p; unfold Zle, not in |- *; intros H1 H2; apply H1; + rewrite (Zcompare_plus_compat n m p); assumption. Qed. -Lemma Zsimpl_le_plus_r : (n,m,p:Z)`n+p<=m+p`->`n<=m`. +Lemma Zplus_le_reg_r : forall n m p:Z, n + p <= m + p -> n <= m. Proof. -Intros n m p H; Apply Zsimpl_le_plus_l with p. -Rewrite (Zplus_sym p n); Rewrite (Zplus_sym p m); Trivial. +intros n m p H; apply Zplus_le_reg_l with p. +rewrite (Zplus_comm p n); rewrite (Zplus_comm p m); trivial. Qed. -Lemma Zsimpl_lt_plus_l : (n,m,p:Z)`p+n`n n < m. Proof. -Unfold Zlt ;Intros n m p; - Rewrite Zcompare_Zplus_compatible;Trivial with arith. +unfold Zlt in |- *; intros n m p; rewrite Zcompare_plus_compat; + trivial with arith. Qed. -Lemma Zsimpl_lt_plus_r : (n,m,p:Z)`n+p`n n < m. Proof. -Intros n m p H; Apply Zsimpl_lt_plus_l with p. -Rewrite (Zplus_sym p n); Rewrite (Zplus_sym p m); Trivial. +intros n m p H; apply Zplus_lt_reg_l with p. +rewrite (Zplus_comm p n); rewrite (Zplus_comm p m); trivial. Qed. (** Special base instances of order *) -Lemma Zgt_Sn_n : (n:Z)`(Zs n)>n`. +Lemma Zgt_succ : forall n:Z, Zsucc n > n. Proof. -Exact Zcompare_Zs_SUPERIEUR. +exact Zcompare_succ_Gt. Qed. -Lemma Zle_Sn_n : (n:Z)~`(Zs n)<=n`. +Lemma Znot_le_succ : forall n:Z, ~ Zsucc n <= n. Proof. -Intros n; Apply Zgt_not_le; Apply Zgt_Sn_n. +intros n; apply Zgt_not_le; apply Zgt_succ. Qed. -Lemma Zlt_n_Sn : (n:Z)`n<(Zs n)`. +Lemma Zlt_succ : forall n:Z, n < Zsucc n. Proof. -Intro n; Apply Zgt_lt; Apply Zgt_Sn_n. +intro n; apply Zgt_lt; apply Zgt_succ. Qed. -Lemma Zlt_pred_n_n : (n:Z)`(Zpred n)n`->`(Zs n)<=p`. +Lemma Zgt_le_succ : forall n m:Z, m > n -> Zsucc n <= m. Proof. -Unfold Zgt Zle; Intros n p H; Elim (Zcompare_et_un p n); Intros H1 H2; -Unfold not ;Intros H3; Unfold not in H1; Apply H1; [ - Assumption -| Elim (Zcompare_ANTISYM (Zplus n (POS xH)) p);Intros H4 H5;Apply H4;Exact H3]. +unfold Zgt, Zle in |- *; intros n p H; elim (Zcompare_Gt_not_Lt p n); + intros H1 H2; unfold not in |- *; intros H3; unfold not in H1; + apply H1; + [ assumption + | elim (Zcompare_Gt_Lt_antisym (n + 1) p); intros H4 H5; apply H4; exact H3 ]. Qed. -Lemma Zle_gt_S : (n,p:Z)`n<=p`->`(Zs p)>n`. +Lemma Zlt_gt_succ : forall n m:Z, n <= m -> Zsucc m > n. Proof. -Intros n p H; Apply Zgt_le_trans with p. - Apply Zgt_Sn_n. - Assumption. +intros n p H; apply Zgt_le_trans with p. + apply Zgt_succ. + assumption. Qed. -Lemma Zle_lt_n_Sm : (n,m:Z)`n<=m`->`n<(Zs m)`. +Lemma Zle_lt_succ : forall n m:Z, n <= m -> n < Zsucc m. Proof. -Intros n m H; Apply Zgt_lt; Apply Zle_gt_S; Assumption. +intros n m H; apply Zgt_lt; apply Zlt_gt_succ; assumption. Qed. -Lemma Zlt_le_S : (n,p:Z)`n`(Zs n)<=p`. +Lemma Zlt_le_succ : forall n m:Z, n < m -> Zsucc n <= m. Proof. -Intros n p H; Apply Zgt_le_S; Apply Zlt_gt; Assumption. +intros n p H; apply Zgt_le_succ; apply Zlt_gt; assumption. Qed. -Lemma Zgt_S_le : (n,p:Z)`(Zs p)>n`->`n<=p`. +Lemma Zgt_succ_le : forall n m:Z, Zsucc m > n -> n <= m. Proof. -Intros n p H;Apply Zle_S_n; Apply Zgt_le_S; Assumption. +intros n p H; apply Zsucc_le_reg; apply Zgt_le_succ; assumption. Qed. -Lemma Zlt_n_Sm_le : (n,m:Z)`n<(Zs m)`->`n<=m`. +Lemma Zlt_succ_le : forall n m:Z, n < Zsucc m -> n <= m. Proof. -Intros n m H; Apply Zgt_S_le; Apply Zlt_gt; Assumption. +intros n m H; apply Zgt_succ_le; apply Zlt_gt; assumption. Qed. -Lemma Zle_S_gt : (n,m:Z) `(Zs n)<=m` -> `m>n`. +Lemma Zlt_succ_gt : forall n m:Z, Zsucc n <= m -> m > n. Proof. -Intros n m H;Apply Zle_gt_trans with m:=(Zs n); - [ Assumption | Apply Zgt_Sn_n ]. +intros n m H; apply Zle_gt_trans with (m := Zsucc n); + [ assumption | apply Zgt_succ ]. Qed. (** Weakening order *) -Lemma Zle_n_Sn : (n:Z)`n<=(Zs n)`. +Lemma Zle_succ : forall n:Z, n <= Zsucc n. Proof. -Intros n; Apply Zgt_S_le;Apply Zgt_trans with m:=(Zs n) ;Apply Zgt_Sn_n. +intros n; apply Zgt_succ_le; apply Zgt_trans with (m := Zsucc n); + apply Zgt_succ. Qed. -Hints Resolve Zle_n_Sn : zarith. +Hint Resolve Zle_succ: zarith. -Lemma Zle_pred_n : (n:Z)`(Zpred n)<=n`. +Lemma Zle_pred : forall n:Z, Zpred n <= n. Proof. -Intros n;Pattern 2 n ;Rewrite Zs_pred; Apply Zle_n_Sn. +intros n; pattern n at 2 in |- *; rewrite Zsucc_pred; apply Zle_succ. Qed. -Lemma Zlt_S : (n,m:Z)`n`n<(Zs m)`. -Intros n m H;Apply Zgt_lt; Apply Zgt_trans with m:=m; [ - Apply Zgt_Sn_n -| Apply Zlt_gt; Assumption ]. +Lemma Zlt_lt_succ : forall n m:Z, n < m -> n < Zsucc m. +intros n m H; apply Zgt_lt; apply Zgt_trans with (m := m); + [ apply Zgt_succ | apply Zlt_gt; assumption ]. Qed. -Lemma Zle_le_S : (x,y:Z)`x<=y`->`x<=(Zs y)`. +Lemma Zle_le_succ : forall n m:Z, n <= m -> n <= Zsucc m. Proof. -Intros x y H. -Apply Zle_trans with y; Trivial with zarith. +intros x y H. +apply Zle_trans with y; trivial with zarith. Qed. -Lemma Zle_trans_S : (n,m:Z)`(Zs n)<=m`->`n<=m`. +Lemma Zle_succ_le : forall n m:Z, Zsucc n <= m -> n <= m. Proof. -Intros n m H;Apply Zle_trans with m:=(Zs n); [ Apply Zle_n_Sn | Assumption ]. +intros n m H; apply Zle_trans with (m := Zsucc n); + [ apply Zle_succ | assumption ]. Qed. -Hints Resolve Zle_le_S : zarith. +Hint Resolve Zle_le_succ: zarith. (** Relating order wrt successor and order wrt predecessor *) -Lemma Zgt_pred : (n,p:Z)`p>(Zs n)`->`(Zpred p)>n`. +Lemma Zgt_succ_pred : forall n m:Z, m > Zsucc n -> Zpred m > n. Proof. -Unfold Zgt Zs Zpred ;Intros n p H; -Rewrite <- [x,y:Z](Zcompare_Zplus_compatible x y (POS xH)); -Rewrite (Zplus_sym p); Rewrite Zplus_assoc; Rewrite [x:Z](Zplus_sym x n); -Simpl; Assumption. +unfold Zgt, Zsucc, Zpred in |- *; intros n p H; + rewrite <- (fun x y => Zcompare_plus_compat x y 1); + rewrite (Zplus_comm p); rewrite Zplus_assoc; + rewrite (fun x => Zplus_comm x n); simpl in |- *; + assumption. Qed. -Lemma Zlt_pred : (n,p:Z)`(Zs n)`n<(Zpred p)`. +Lemma Zlt_succ_pred : forall n m:Z, Zsucc n < m -> n < Zpred m. Proof. -Intros n p H;Apply Zlt_S_n; Rewrite <- Zs_pred; Assumption. +intros n p H; apply Zsucc_lt_reg; rewrite <- Zsucc_pred; assumption. Qed. (** Relating strict order and large order on positive *) -Lemma Zlt_ZERO_pred_le_ZERO : (n:Z) `0 `0<=(Zpred n)`. -Intros x H. -Rewrite (Zs_pred x) in H. -Apply Zgt_S_le. -Apply Zlt_gt. -Assumption. +Lemma Zlt_0_le_0_pred : forall n:Z, 0 < n -> 0 <= Zpred n. +intros x H. +rewrite (Zsucc_pred x) in H. +apply Zgt_succ_le. +apply Zlt_gt. +assumption. Qed. -V7only [Set Implicit Arguments.]. -Lemma Zgt0_le_pred : (y:Z) `y > 0` -> `0 <= (Zpred y)`. -Intros; Apply Zlt_ZERO_pred_le_ZERO; Apply Zgt_lt. Assumption. +Lemma Zgt_0_le_0_pred : forall n:Z, n > 0 -> 0 <= Zpred n. +intros; apply Zlt_0_le_0_pred; apply Zgt_lt. assumption. Qed. -V7only [Unset Implicit Arguments.]. (** Special cases of ordered integers *) -V7only [ (* Relevance confirmed from Zdivides *) ]. -Lemma Z_O_1: `0<1`. +Lemma Zlt_0_1 : 0 < 1. Proof. -Change `0<(Zs 0)`. Apply Zlt_n_Sn. +change (0 < Zsucc 0) in |- *. apply Zlt_succ. Qed. -Lemma Zle_0_1: `0<=1`. +Lemma Zle_0_1 : 0 <= 1. Proof. -Change `0<=(Zs 0)`. Apply Zle_n_Sn. +change (0 <= Zsucc 0) in |- *. apply Zle_succ. Qed. -V7only [ (* Relevance confirmed from Zdivides *) ]. -Lemma Zle_NEG_POS: (p,q:positive) `(NEG p)<=(POS q)`. +Lemma Zle_neg_pos : forall p q:positive, Zneg p <= Zpos q. Proof. -Intros p; Red; Simpl; Red; Intros H; Discriminate. +intros p; red in |- *; simpl in |- *; red in |- *; intros H; discriminate. Qed. -Lemma POS_gt_ZERO : (p:positive) `(POS p)>0`. -Unfold Zgt; Trivial. +Lemma Zgt_pos_0 : forall p:positive, Zpos p > 0. +unfold Zgt in |- *; trivial. Qed. (* weaker but useful (in [Zpower] for instance) *) -Lemma ZERO_le_POS : (p:positive) `0<=(POS p)`. -Intro; Unfold Zle; Discriminate. +Lemma Zle_0_pos : forall p:positive, 0 <= Zpos p. +intro; unfold Zle in |- *; discriminate. Qed. -Lemma NEG_lt_ZERO : (p:positive)`(NEG p)<0`. -Unfold Zlt; Trivial. +Lemma Zlt_neg_0 : forall p:positive, Zneg p < 0. +unfold Zlt in |- *; trivial. Qed. -Lemma ZERO_le_inj : - (n:nat) `0 <= (inject_nat n)`. -Induction n; Simpl; Intros; -[ Apply Zle_n -| Unfold Zle; Simpl; Discriminate]. +Lemma Zle_0_nat : forall n:nat, 0 <= Z_of_nat n. +simple induction n; simpl in |- *; intros; + [ apply Zle_refl | unfold Zle in |- *; simpl in |- *; discriminate ]. Qed. -Hints Immediate Zle_refl : zarith. +Hint Immediate Zeq_le: zarith. (** Transitivity using successor *) -Lemma Zgt_trans_S : (n,m,p:Z)`(Zs n)>m`->`m>p`->`n>p`. +Lemma Zge_trans_succ : forall n m p:Z, Zsucc n > m -> m > p -> n > p. Proof. -Intros n m p H1 H2;Apply Zle_gt_trans with m:=m; - [ Apply Zgt_S_le; Assumption | Assumption ]. +intros n m p H1 H2; apply Zle_gt_trans with (m := m); + [ apply Zgt_succ_le; assumption | assumption ]. Qed. (** Derived lemma *) -Lemma Zgt_S : (n,m:Z)`(Zs n)>m`->(`n>m`\/(m=n)). +Lemma Zgt_succ_gt_or_eq : forall n m:Z, Zsucc n > m -> n > m \/ m = n. Proof. -Intros n m H. -Assert Hle : `m<=n`. - Apply Zgt_S_le; Assumption. -NewDestruct (Zle_lt_or_eq ? ? Hle) as [Hlt|Heq]. - Left; Apply Zlt_gt; Assumption. - Right; Assumption. +intros n m H. +assert (Hle : m <= n). + apply Zgt_succ_le; assumption. +destruct (Zle_lt_or_eq _ _ Hle) as [Hlt| Heq]. + left; apply Zlt_gt; assumption. + right; assumption. Qed. (** Compatibility of multiplication by a positive wrt to order *) -V7only [Set Implicit Arguments.]. -Lemma Zle_Zmult_pos_right : (a,b,c : Z) `a<=b` -> `0<=c` -> `a*c<=b*c`. +Lemma Zmult_le_compat_r : forall n m p:Z, n <= m -> 0 <= p -> n * p <= m * p. Proof. -Intros a b c H H0; NewDestruct c. - Do 2 Rewrite Zero_mult_right; Assumption. - Rewrite (Zmult_sym a); Rewrite (Zmult_sym b). - Unfold Zle; Rewrite Zcompare_Zmult_compatible; Assumption. - Unfold Zle in H0; Contradiction H0; Reflexivity. +intros a b c H H0; destruct c. + do 2 rewrite Zmult_0_r; assumption. + rewrite (Zmult_comm a); rewrite (Zmult_comm b). + unfold Zle in |- *; rewrite Zcompare_mult_compat; assumption. + unfold Zle in H0; contradiction H0; reflexivity. Qed. -Lemma Zle_Zmult_pos_left : (a,b,c : Z) `a<=b` -> `0<=c` -> `c*a<=c*b`. +Lemma Zmult_le_compat_l : forall n m p:Z, n <= m -> 0 <= p -> p * n <= p * m. Proof. -Intros a b c H1 H2; Rewrite (Zmult_sym c a);Rewrite (Zmult_sym c b). -Apply Zle_Zmult_pos_right; Trivial. +intros a b c H1 H2; rewrite (Zmult_comm c a); rewrite (Zmult_comm c b). +apply Zmult_le_compat_r; trivial. Qed. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_lt_compat_r : (x,y,z:Z)`0 `x < y` -> `x*z < y*z`. +Lemma Zmult_lt_compat_r : forall n m p:Z, 0 < p -> n < m -> n * p < m * p. Proof. -Intros x y z H H0; NewDestruct z. - Contradiction (Zlt_n_n `0`). - Rewrite (Zmult_sym x); Rewrite (Zmult_sym y). - Unfold Zlt; Rewrite Zcompare_Zmult_compatible; Assumption. - Discriminate H. -Save. +intros x y z H H0; destruct z. + contradiction (Zlt_irrefl 0). + rewrite (Zmult_comm x); rewrite (Zmult_comm y). + unfold Zlt in |- *; rewrite Zcompare_mult_compat; assumption. + discriminate H. +Qed. -Lemma Zgt_Zmult_right : (x,y,z:Z)`z>0` -> `x > y` -> `x*z > y*z`. +Lemma Zmult_gt_compat_r : forall n m p:Z, p > 0 -> n > m -> n * p > m * p. Proof. -Intros x y z; Intros; Apply Zlt_gt; Apply Zmult_lt_compat_r; - Apply Zgt_lt; Assumption. +intros x y z; intros; apply Zlt_gt; apply Zmult_lt_compat_r; apply Zgt_lt; + assumption. Qed. -Lemma Zlt_Zmult_right : (x,y,z:Z)`z>0` -> `x < y` -> `x*z < y*z`. +Lemma Zmult_gt_0_lt_compat_r : + forall n m p:Z, p > 0 -> n < m -> n * p < m * p. Proof. -Intros x y z; Intros; Apply Zmult_lt_compat_r; - [Apply Zgt_lt; Assumption | Assumption]. +intros x y z; intros; apply Zmult_lt_compat_r; + [ apply Zgt_lt; assumption | assumption ]. Qed. -Lemma Zle_Zmult_right : (x,y,z:Z)`z>0` -> `x <= y` -> `x*z <= y*z`. +Lemma Zmult_gt_0_le_compat_r : + forall n m p:Z, p > 0 -> n <= m -> n * p <= m * p. Proof. -Intros x y z Hz Hxy. -Elim (Zle_lt_or_eq x y Hxy). -Intros; Apply Zlt_le_weak. -Apply Zlt_Zmult_right; Trivial. -Intros; Apply Zle_refl. -Rewrite H; Trivial. +intros x y z Hz Hxy. +elim (Zle_lt_or_eq x y Hxy). +intros; apply Zlt_le_weak. +apply Zmult_gt_0_lt_compat_r; trivial. +intros; apply Zeq_le. +rewrite H; trivial. Qed. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_lt_0_le_compat_r : (x,y,z:Z)`0 < z`->`x <= y`->`x*z <= y*z`. +Lemma Zmult_lt_0_le_compat_r : + forall n m p:Z, 0 < p -> n <= m -> n * p <= m * p. Proof. -Intros x y z; Intros; Apply Zle_Zmult_right; Try Apply Zlt_gt; Assumption. +intros x y z; intros; apply Zmult_gt_0_le_compat_r; try apply Zlt_gt; + assumption. Qed. -Lemma Zlt_Zmult_left : (x,y,z:Z)`z>0` -> `x < y` -> `z*x < z*y`. +Lemma Zmult_gt_0_lt_compat_l : + forall n m p:Z, p > 0 -> n < m -> p * n < p * m. Proof. -Intros x y z; Intros. -Rewrite (Zmult_sym z x); Rewrite (Zmult_sym z y); -Apply Zlt_Zmult_right; Assumption. +intros x y z; intros. +rewrite (Zmult_comm z x); rewrite (Zmult_comm z y); + apply Zmult_gt_0_lt_compat_r; assumption. Qed. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_lt_compat_l : (x,y,z:Z)`0 `x < y` -> `z*x < z*y`. +Lemma Zmult_lt_compat_l : forall n m p:Z, 0 < p -> n < m -> p * n < p * m. Proof. -Intros x y z; Intros. -Rewrite (Zmult_sym z x); Rewrite (Zmult_sym z y); -Apply Zlt_Zmult_right; Try Apply Zlt_gt; Assumption. -Save. +intros x y z; intros. +rewrite (Zmult_comm z x); rewrite (Zmult_comm z y); + apply Zmult_gt_0_lt_compat_r; try apply Zlt_gt; assumption. +Qed. -Lemma Zgt_Zmult_left : (x,y,z:Z)`z>0` -> `x > y` -> `z*x > z*y`. +Lemma Zmult_gt_compat_l : forall n m p:Z, p > 0 -> n > m -> p * n > p * m. Proof. -Intros x y z; Intros; -Rewrite (Zmult_sym z x); Rewrite (Zmult_sym z y); -Apply Zgt_Zmult_right; Assumption. +intros x y z; intros; rewrite (Zmult_comm z x); rewrite (Zmult_comm z y); + apply Zmult_gt_compat_r; assumption. Qed. -Lemma Zge_Zmult_pos_right : (a,b,c : Z) `a>=b` -> `c>=0` -> `a*c>=b*c`. +Lemma Zmult_ge_compat_r : forall n m p:Z, n >= m -> p >= 0 -> n * p >= m * p. Proof. -Intros a b c H1 H2; Apply Zle_ge. -Apply Zle_Zmult_pos_right; Apply Zge_le; Trivial. +intros a b c H1 H2; apply Zle_ge. +apply Zmult_le_compat_r; apply Zge_le; trivial. Qed. -Lemma Zge_Zmult_pos_left : (a,b,c : Z) `a>=b` -> `c>=0` -> `c*a>=c*b`. +Lemma Zmult_ge_compat_l : forall n m p:Z, n >= m -> p >= 0 -> p * n >= p * m. Proof. -Intros a b c H1 H2; Apply Zle_ge. -Apply Zle_Zmult_pos_left; Apply Zge_le; Trivial. +intros a b c H1 H2; apply Zle_ge. +apply Zmult_le_compat_l; apply Zge_le; trivial. Qed. -Lemma Zge_Zmult_pos_compat : - (a,b,c,d : Z) `a>=c` -> `b>=d` -> `c>=0` -> `d>=0` -> `a*b>=c*d`. +Lemma Zmult_ge_compat : + forall n m p q:Z, n >= p -> m >= q -> p >= 0 -> q >= 0 -> n * m >= p * q. Proof. -Intros a b c d H0 H1 H2 H3. -Apply Zge_trans with (Zmult a d). -Apply Zge_Zmult_pos_left; Trivial. -Apply Zge_trans with c; Trivial. -Apply Zge_Zmult_pos_right; Trivial. +intros a b c d H0 H1 H2 H3. +apply Zge_trans with (a * d). +apply Zmult_ge_compat_l; trivial. +apply Zge_trans with c; trivial. +apply Zmult_ge_compat_r; trivial. Qed. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_le_compat: (a, b, c, d : Z) - `a<=c` -> `b<=d` -> `0<=a` -> `0<=b` -> `a*b<=c*d`. +Lemma Zmult_le_compat : + forall n m p q:Z, n <= p -> m <= q -> 0 <= n -> 0 <= m -> n * m <= p * q. Proof. -Intros a b c d H0 H1 H2 H3. -Apply Zle_trans with (Zmult c b). -Apply Zle_Zmult_pos_right; Assumption. -Apply Zle_Zmult_pos_left. -Assumption. -Apply Zle_trans with a; Assumption. +intros a b c d H0 H1 H2 H3. +apply Zle_trans with (c * b). +apply Zmult_le_compat_r; assumption. +apply Zmult_le_compat_l. +assumption. +apply Zle_trans with a; assumption. Qed. (** Simplification of multiplication by a positive wrt to being positive *) -Lemma Zlt_Zmult_right2 : (x,y,z:Z)`z>0` -> `x*z < y*z` -> `x < y`. +Lemma Zmult_gt_0_lt_reg_r : forall n m p:Z, p > 0 -> n * p < m * p -> n < m. Proof. -Intros x y z; Intros; NewDestruct z. - Contradiction (Zgt_antirefl `0`). - Rewrite (Zmult_sym x) in H0; Rewrite (Zmult_sym y) in H0. - Unfold Zlt in H0; Rewrite Zcompare_Zmult_compatible in H0; Assumption. - Discriminate H. +intros x y z; intros; destruct z. + contradiction (Zgt_irrefl 0). + rewrite (Zmult_comm x) in H0; rewrite (Zmult_comm y) in H0. + unfold Zlt in H0; rewrite Zcompare_mult_compat in H0; assumption. + discriminate H. Qed. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_lt_reg_r : (a, b, c : Z) `0 `a*c `a n * p < m * p -> n < m. Proof. -Intros a b c H0 H1. -Apply Zlt_Zmult_right2 with c; Try Apply Zlt_gt; Assumption. +intros a b c H0 H1. +apply Zmult_gt_0_lt_reg_r with c; try apply Zlt_gt; assumption. Qed. -Lemma Zle_mult_simpl : (a,b,c:Z)`c>0`->`a*c<=b*c`->`a<=b`. +Lemma Zmult_le_reg_r : forall n m p:Z, p > 0 -> n * p <= m * p -> n <= m. Proof. -Intros x y z Hz Hxy. -Elim (Zle_lt_or_eq `x*z` `y*z` Hxy). -Intros; Apply Zlt_le_weak. -Apply Zlt_Zmult_right2 with z; Trivial. -Intros; Apply Zle_refl. -Apply Zmult_reg_right with z. - Intro. Rewrite H0 in Hz. Contradiction (Zgt_antirefl `0`). -Assumption. +intros x y z Hz Hxy. +elim (Zle_lt_or_eq (x * z) (y * z) Hxy). +intros; apply Zlt_le_weak. +apply Zmult_gt_0_lt_reg_r with z; trivial. +intros; apply Zeq_le. +apply Zmult_reg_r with z. + intro. rewrite H0 in Hz. contradiction (Zgt_irrefl 0). +assumption. Qed. -V7only [Notation Zle_Zmult_right2 := Zle_mult_simpl. -(* Zle_Zmult_right2 : (x,y,z:Z)`z>0` -> `x*z <= y*z` -> `x <= y`. *) -]. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_lt_0_le_reg_r: (x,y,z:Z)`0 `x*z <= y*z`->`x <= y`. -Intros x y z; Intros ; Apply Zle_mult_simpl with z. -Try Apply Zlt_gt; Assumption. -Assumption. +Lemma Zmult_lt_0_le_reg_r : forall n m p:Z, 0 < p -> n * p <= m * p -> n <= m. +intros x y z; intros; apply Zmult_le_reg_r with z. +try apply Zlt_gt; assumption. +assumption. Qed. -V7only [Unset Implicit Arguments.]. -Lemma Zge_mult_simpl : (a,b,c:Z) `c>0`->`a*c>=b*c`->`a>=b`. -Intros a b c H1 H2; Apply Zle_ge; Apply Zle_mult_simpl with c; Trivial. -Apply Zge_le; Trivial. +Lemma Zmult_ge_reg_r : forall n m p:Z, p > 0 -> n * p >= m * p -> n >= m. +intros a b c H1 H2; apply Zle_ge; apply Zmult_le_reg_r with c; trivial. +apply Zge_le; trivial. Qed. -Lemma Zgt_mult_simpl : (a,b,c:Z) `c>0`->`a*c>b*c`->`a>b`. -Intros a b c H1 H2; Apply Zlt_gt; Apply Zlt_Zmult_right2 with c; Trivial. -Apply Zgt_lt; Trivial. +Lemma Zmult_gt_reg_r : forall n m p:Z, p > 0 -> n * p > m * p -> n > m. +intros a b c H1 H2; apply Zlt_gt; apply Zmult_gt_0_lt_reg_r with c; trivial. +apply Zgt_lt; trivial. Qed. (** Compatibility of multiplication by a positive wrt to being positive *) -Lemma Zle_ZERO_mult : (x,y:Z) `0<=x` -> `0<=y` -> `0<=x*y`. +Lemma Zmult_le_0_compat : forall n m:Z, 0 <= n -> 0 <= m -> 0 <= n * m. Proof. -Intros x y; Case x. -Intros; Rewrite Zero_mult_left; Trivial. -Intros p H1; Unfold Zle. - Pattern 2 ZERO ; Rewrite <- (Zero_mult_right (POS p)). - Rewrite Zcompare_Zmult_compatible; Trivial. -Intros p H1 H2; Absurd (Zgt ZERO (NEG p)); Trivial. -Unfold Zgt; Simpl; Auto with zarith. +intros x y; case x. +intros; rewrite Zmult_0_l; trivial. +intros p H1; unfold Zle in |- *. + pattern 0 at 2 in |- *; rewrite <- (Zmult_0_r (Zpos p)). + rewrite Zcompare_mult_compat; trivial. +intros p H1 H2; absurd (0 > Zneg p); trivial. +unfold Zgt in |- *; simpl in |- *; auto with zarith. Qed. -Lemma Zgt_ZERO_mult: (a,b:Z) `a>0`->`b>0`->`a*b>0`. +Lemma Zmult_gt_0_compat : forall n m:Z, n > 0 -> m > 0 -> n * m > 0. Proof. -Intros x y; Case x. -Intros H; Discriminate H. -Intros p H1; Unfold Zgt; -Pattern 2 ZERO ; Rewrite <- (Zero_mult_right (POS p)). - Rewrite Zcompare_Zmult_compatible; Trivial. -Intros p H; Discriminate H. +intros x y; case x. +intros H; discriminate H. +intros p H1; unfold Zgt in |- *; pattern 0 at 2 in |- *; + rewrite <- (Zmult_0_r (Zpos p)). + rewrite Zcompare_mult_compat; trivial. +intros p H; discriminate H. Qed. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_lt_O_compat : (a, b : Z) `0 `0 `0 0 < m -> 0 < n * m. +intros a b apos bpos. +apply Zgt_lt. +apply Zmult_gt_0_compat; try apply Zlt_gt; assumption. Qed. -Lemma Zle_mult: (x,y:Z) `x>0` -> `0<=y` -> `0<=(Zmult y x)`. +Lemma Zmult_gt_0_le_0_compat : forall n m:Z, n > 0 -> 0 <= m -> 0 <= m * n. Proof. -Intros x y H1 H2; Apply Zle_ZERO_mult; Trivial. -Apply Zlt_le_weak; Apply Zgt_lt; Trivial. +intros x y H1 H2; apply Zmult_le_0_compat; trivial. +apply Zlt_le_weak; apply Zgt_lt; trivial. Qed. (** Simplification of multiplication by a positive wrt to being positive *) -Lemma Zmult_le: (x,y:Z) `x>0` -> `0<=(Zmult y x)` -> `0<=y`. +Lemma Zmult_le_0_reg_r : forall n m:Z, n > 0 -> 0 <= m * n -> 0 <= m. Proof. -Intros x y; Case x; [ - Simpl; Unfold Zgt ; Simpl; Intros H; Discriminate H -| Intros p H1; Unfold Zle; Rewrite -> Zmult_sym; - Pattern 1 ZERO ; Rewrite <- (Zero_mult_right (POS p)); - Rewrite Zcompare_Zmult_compatible; Auto with arith -| Intros p; Unfold Zgt ; Simpl; Intros H; Discriminate H]. +intros x y; case x; + [ simpl in |- *; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H + | intros p H1; unfold Zle in |- *; rewrite Zmult_comm; + pattern 0 at 1 in |- *; rewrite <- (Zmult_0_r (Zpos p)); + rewrite Zcompare_mult_compat; auto with arith + | intros p; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H ]. Qed. -Lemma Zmult_lt: (x,y:Z) `x>0` -> `0 `0 0 -> 0 < m * n -> 0 < m. Proof. -Intros x y; Case x; [ - Simpl; Unfold Zgt ; Simpl; Intros H; Discriminate H -| Intros p H1; Unfold Zlt; Rewrite -> Zmult_sym; - Pattern 1 ZERO ; Rewrite <- (Zero_mult_right (POS p)); - Rewrite Zcompare_Zmult_compatible; Auto with arith -| Intros p; Unfold Zgt ; Simpl; Intros H; Discriminate H]. +intros x y; case x; + [ simpl in |- *; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H + | intros p H1; unfold Zlt in |- *; rewrite Zmult_comm; + pattern 0 at 1 in |- *; rewrite <- (Zmult_0_r (Zpos p)); + rewrite Zcompare_mult_compat; auto with arith + | intros p; unfold Zgt in |- *; simpl in |- *; intros H; discriminate H ]. Qed. -V7only [ (* Relevance confirmed from Zextensions *) ]. -Lemma Zmult_lt_0_reg_r : (x,y:Z)`0 < x`->`0 < y*x`->`0 < y`. +Lemma Zmult_lt_0_reg_r : forall n m:Z, 0 < n -> 0 < m * n -> 0 < m. Proof. -Intros x y; Intros; EApply Zmult_lt with x ; Try Apply Zlt_gt; Assumption. +intros x y; intros; eapply Zmult_gt_0_lt_0_reg_r with x; try apply Zlt_gt; + assumption. Qed. -Lemma Zmult_gt: (x,y:Z) `x>0` -> `x*y>0` -> `y>0`. +Lemma Zmult_gt_0_reg_l : forall n m:Z, n > 0 -> n * m > 0 -> m > 0. Proof. -Intros x y; Case x. - Intros H; Discriminate H. - Intros p H1; Unfold Zgt. - Pattern 1 ZERO ; Rewrite <- (Zero_mult_right (POS p)). - Rewrite Zcompare_Zmult_compatible; Trivial. -Intros p H; Discriminate H. +intros x y; case x. + intros H; discriminate H. + intros p H1; unfold Zgt in |- *. + pattern 0 at 1 in |- *; rewrite <- (Zmult_0_r (Zpos p)). + rewrite Zcompare_mult_compat; trivial. +intros p H; discriminate H. Qed. (** Simplification of square wrt order *) -Lemma Zgt_square_simpl: (x, y : Z) `x>=0` -> `y>=0` -> `x*x>y*y` -> `x>y`. +Lemma Zgt_square_simpl : + forall n m:Z, n >= 0 -> m >= 0 -> n * n > m * m -> n > m. Proof. -Intros x y H0 H1 H2. -Case (dec_Zlt y x). -Intro; Apply Zlt_gt; Trivial. -Intros H3; Cut (Zge y x). -Intros H. -Elim Zgt_not_le with 1 := H2. -Apply Zge_le. -Apply Zge_Zmult_pos_compat; Auto. -Apply not_Zlt; Trivial. +intros x y H0 H1 H2. +case (dec_Zlt y x). +intro; apply Zlt_gt; trivial. +intros H3; cut (y >= x). +intros H. +elim Zgt_not_le with (1 := H2). +apply Zge_le. +apply Zmult_ge_compat; auto. +apply Znot_lt_ge; trivial. Qed. -Lemma Zlt_square_simpl: (x,y:Z) `0<=x` -> `0<=y` -> `y*y `y 0 <= m -> m * m < n * n -> m < n. Proof. -Intros x y H0 H1 H2. -Apply Zgt_lt. -Apply Zgt_square_simpl; Try Apply Zle_ge; Try Apply Zlt_gt; Assumption. +intros x y H0 H1 H2. +apply Zgt_lt. +apply Zgt_square_simpl; try apply Zle_ge; try apply Zlt_gt; assumption. Qed. (** Equivalence between inequalities *) -Lemma Zle_plus_swap : (x,y,z:Z) `x+z<=y` <-> `x<=y-z`. +Lemma Zle_plus_swap : forall n m p:Z, n + p <= m <-> n <= m - p. Proof. - Intros x y z; Intros. Split. Intro. Rewrite <- (Zero_right x). Rewrite <- (Zplus_inverse_r z). - Rewrite Zplus_assoc_l. Exact (Zle_reg_r ? ? ? H). - Intro. Rewrite <- (Zero_right y). Rewrite <- (Zplus_inverse_l z). Rewrite Zplus_assoc_l. - Apply Zle_reg_r. Assumption. + intros x y z; intros. split. intro. rewrite <- (Zplus_0_r x). rewrite <- (Zplus_opp_r z). + rewrite Zplus_assoc. exact (Zplus_le_compat_r _ _ _ H). + intro. rewrite <- (Zplus_0_r y). rewrite <- (Zplus_opp_l z). rewrite Zplus_assoc. + apply Zplus_le_compat_r. assumption. Qed. -Lemma Zlt_plus_swap : (x,y,z:Z) `x+z `x n < m - p. Proof. - Intros x y z; Intros. Split. Intro. Unfold Zminus. Rewrite Zplus_sym. Rewrite <- (Zero_left x). - Rewrite <- (Zplus_inverse_l z). Rewrite Zplus_assoc_r. Apply Zlt_reg_l. Rewrite Zplus_sym. - Assumption. - Intro. Rewrite Zplus_sym. Rewrite <- (Zero_left y). Rewrite <- (Zplus_inverse_r z). - Rewrite Zplus_assoc_r. Apply Zlt_reg_l. Rewrite Zplus_sym. Assumption. + intros x y z; intros. split. intro. unfold Zminus in |- *. rewrite Zplus_comm. rewrite <- (Zplus_0_l x). + rewrite <- (Zplus_opp_l z). rewrite Zplus_assoc_reverse. apply Zplus_lt_compat_l. rewrite Zplus_comm. + assumption. + intro. rewrite Zplus_comm. rewrite <- (Zplus_0_l y). rewrite <- (Zplus_opp_r z). + rewrite Zplus_assoc_reverse. apply Zplus_lt_compat_l. rewrite Zplus_comm. assumption. Qed. -Lemma Zeq_plus_swap : (x,y,z:Z)`x+z=y` <-> `x=y-z`. +Lemma Zeq_plus_swap : forall n m p:Z, n + p = m <-> n = m - p. Proof. -Intros x y z; Intros. Split. Intro. Apply Zplus_minus. Symmetry. Rewrite Zplus_sym. - Assumption. -Intro. Rewrite H. Unfold Zminus. Rewrite Zplus_assoc_r. - Rewrite Zplus_inverse_l. Apply Zero_right. +intros x y z; intros. split. intro. apply Zplus_minus_eq. symmetry in |- *. rewrite Zplus_comm. + assumption. +intro. rewrite H. unfold Zminus in |- *. rewrite Zplus_assoc_reverse. + rewrite Zplus_opp_l. apply Zplus_0_r. Qed. -Lemma Zlt_minus : (n,m:Z)`0`n-m n - m < n. Proof. -Intros n m H; Apply Zsimpl_lt_plus_l with p:=m; Rewrite Zle_plus_minus; -Pattern 1 n ;Rewrite <- (Zero_right n); Rewrite (Zplus_sym m n); -Apply Zlt_reg_l; Assumption. +intros n m H; apply Zplus_lt_reg_l with (p := m); rewrite Zplus_minus; + pattern n at 1 in |- *; rewrite <- (Zplus_0_r n); + rewrite (Zplus_comm m n); apply Zplus_lt_compat_l; + assumption. Qed. -Lemma Zlt_O_minus_lt : (n,m:Z)`0`m m < n. Proof. -Intros n m H; Apply Zsimpl_lt_plus_l with p:=(Zopp m); Rewrite Zplus_inverse_l; -Rewrite Zplus_sym;Exact H. -Qed. +intros n m H; apply Zplus_lt_reg_l with (p := - m); rewrite Zplus_opp_l; + rewrite Zplus_comm; exact H. +Qed. \ No newline at end of file diff --git a/theories/ZArith/Zpower.v b/theories/ZArith/Zpower.v index 73e8a08da3..c19ef4499d 100644 --- a/theories/ZArith/Zpower.v +++ b/theories/ZArith/Zpower.v @@ -8,10 +8,9 @@ (*i $Id$ i*) -Require ZArith_base. -Require Omega. -Require Zcomplements. -V7only [Import Z_scope.]. +Require Import ZArith_base. +Require Import Omega. +Require Import Zcomplements. Open Local Scope Z_scope. Section section1. @@ -19,86 +18,85 @@ Section section1. (** [Zpower_nat z n] is the n-th power of [z] when [n] is an unary integer (type [nat]) and [z] a signed integer (type [Z]) *) -Definition Zpower_nat := - [z:Z][n:nat] (iter_nat n Z ([x:Z]` z * x `) `1`). +Definition Zpower_nat (z:Z) (n:nat) := iter_nat n Z (fun x:Z => z * x) 1. (** [Zpower_nat_is_exp] says [Zpower_nat] is a morphism for [plus : nat->nat] and [Zmult : Z->Z] *) -Lemma Zpower_nat_is_exp : - (n,m:nat)(z:Z) - `(Zpower_nat z (plus n m)) = (Zpower_nat z n)*(Zpower_nat z m)`. +Lemma Zpower_nat_is_exp : + forall (n m:nat) (z:Z), + Zpower_nat z (n + m) = Zpower_nat z n * Zpower_nat z m. -Intros; Elim n; -[ Simpl; Elim (Zpower_nat z m); Auto with zarith -| Unfold Zpower_nat; Intros; Simpl; Rewrite H; - Apply Zmult_assoc]. +intros; elim n; + [ simpl in |- *; elim (Zpower_nat z m); auto with zarith + | unfold Zpower_nat in |- *; intros; simpl in |- *; rewrite H; + apply Zmult_assoc ]. Qed. (** [Zpower_pos z n] is the n-th power of [z] when [n] is an binary integer (type [positive]) and [z] a signed integer (type [Z]) *) -Definition Zpower_pos := - [z:Z][n:positive] (iter_pos n Z ([x:Z]`z * x`) `1`). +Definition Zpower_pos (z:Z) (n:positive) := iter_pos n Z (fun x:Z => z * x) 1. (** This theorem shows that powers of unary and binary integers are the same thing, modulo the function convert : [positive -> nat] *) -Theorem Zpower_pos_nat : - (z:Z)(p:positive)(Zpower_pos z p) = (Zpower_nat z (convert p)). +Theorem Zpower_pos_nat : + forall (z:Z) (p:positive), Zpower_pos z p = Zpower_nat z (nat_of_P p). -Intros; Unfold Zpower_pos; Unfold Zpower_nat; Apply iter_convert. +intros; unfold Zpower_pos in |- *; unfold Zpower_nat in |- *; + apply iter_nat_of_P. Qed. (** Using the theorem [Zpower_pos_nat] and the lemma [Zpower_nat_is_exp] we deduce that the function [[n:positive](Zpower_pos z n)] is a morphism for [add : positive->positive] and [Zmult : Z->Z] *) -Theorem Zpower_pos_is_exp : - (n,m:positive)(z:Z) - ` (Zpower_pos z (add n m)) = (Zpower_pos z n)*(Zpower_pos z m)`. +Theorem Zpower_pos_is_exp : + forall (n m:positive) (z:Z), + Zpower_pos z (n + m) = Zpower_pos z n * Zpower_pos z m. -Intros. -Rewrite -> (Zpower_pos_nat z n). -Rewrite -> (Zpower_pos_nat z m). -Rewrite -> (Zpower_pos_nat z (add n m)). -Rewrite -> (convert_add n m). -Apply Zpower_nat_is_exp. +intros. +rewrite (Zpower_pos_nat z n). +rewrite (Zpower_pos_nat z m). +rewrite (Zpower_pos_nat z (n + m)). +rewrite (nat_of_P_plus_morphism n m). +apply Zpower_nat_is_exp. Qed. -Definition Zpower := - [x,y:Z]Cases y of - (POS p) => (Zpower_pos x p) - | ZERO => `1` - | (NEG p) => `0` +Definition Zpower (x y:Z) := + match y with + | Zpos p => Zpower_pos x p + | Z0 => 1 + | Zneg p => 0 end. -Infix "^" Zpower (at level 2, left associativity) : Z_scope V8only. +Infix "^" := Zpower : Z_scope. -Hints Immediate Zpower_nat_is_exp : zarith. -Hints Immediate Zpower_pos_is_exp : zarith. -Hints Unfold Zpower_pos : zarith. -Hints Unfold Zpower_nat : zarith. +Hint Immediate Zpower_nat_is_exp: zarith. +Hint Immediate Zpower_pos_is_exp: zarith. +Hint Unfold Zpower_pos: zarith. +Hint Unfold Zpower_nat: zarith. -Lemma Zpower_exp : (x:Z)(n,m:Z) - `n >= 0` -> `m >= 0` -> `(Zpower x (n+m))=(Zpower x n)*(Zpower x m)`. -NewDestruct n; NewDestruct m; Auto with zarith. -Simpl; Intros; Apply Zred_factor0. -Simpl; Auto with zarith. -Intros; Compute in H0; Absurd INFERIEUR=INFERIEUR; Auto with zarith. -Intros; Compute in H0; Absurd INFERIEUR=INFERIEUR; Auto with zarith. +Lemma Zpower_exp : + forall x n m:Z, n >= 0 -> m >= 0 -> x ^ (n + m) = x ^ n * x ^ m. +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. Qed. End section1. (* Exporting notation "^" *) -Infix "^" Zpower (at level 2, left associativity) : Z_scope V8only. +Infix "^" := Zpower : Z_scope. -Hints Immediate Zpower_nat_is_exp : zarith. -Hints Immediate Zpower_pos_is_exp : zarith. -Hints Unfold Zpower_pos : zarith. -Hints Unfold Zpower_nat : zarith. +Hint Immediate Zpower_nat_is_exp: zarith. +Hint Immediate Zpower_pos_is_exp: zarith. +Hint Unfold Zpower_pos: zarith. +Hint Unfold Zpower_nat: zarith. Section Powers_of_2. @@ -109,100 +107,96 @@ Section Powers_of_2. (** [shift n m] computes [2^n * m], or [m] shifted by [n] positions *) -Definition shift_nat := - [n:nat][z:positive](iter_nat n positive xO z). -Definition shift_pos := - [n:positive][z:positive](iter_pos n positive xO z). -Definition shift := - [n:Z][z:positive] - Cases n of - ZERO => z - | (POS p) => (iter_pos p positive xO z) - | (NEG p) => z +Definition shift_nat (n:nat) (z:positive) := iter_nat n positive xO z. +Definition shift_pos (n z:positive) := iter_pos n positive xO z. +Definition shift (n:Z) (z:positive) := + match n with + | Z0 => z + | Zpos p => iter_pos p positive xO z + | Zneg p => z end. -Definition two_power_nat := [n:nat] (POS (shift_nat n xH)). -Definition two_power_pos := [x:positive] (POS (shift_pos x xH)). +Definition two_power_nat (n:nat) := Zpos (shift_nat n 1). +Definition two_power_pos (x:positive) := Zpos (shift_pos x 1). -Lemma two_power_nat_S : - (n:nat)` (two_power_nat (S n)) = 2*(two_power_nat n)`. -Intro; Simpl; Apply refl_equal. +Lemma two_power_nat_S : + forall n:nat, two_power_nat (S n) = 2 * two_power_nat n. +intro; simpl in |- *; apply refl_equal. Qed. Lemma shift_nat_plus : - (n,m:nat)(x:positive) - (shift_nat (plus n m) x)=(shift_nat n (shift_nat m x)). + forall (n m:nat) (x:positive), + shift_nat (n + m) x = shift_nat n (shift_nat m x). -Intros; Unfold shift_nat; Apply iter_nat_plus. +intros; unfold shift_nat in |- *; apply iter_nat_plus. Qed. Theorem shift_nat_correct : - (n:nat)(x:positive)(POS (shift_nat n x))=`(Zpower_nat 2 n)*(POS x)`. - -Unfold shift_nat; Induction n; -[ Simpl; Trivial with zarith -| Intros; Replace (Zpower_nat `2` (S n0)) with `2 * (Zpower_nat 2 n0)`; -[ Rewrite <- Zmult_assoc; Rewrite <- (H x); Simpl; Reflexivity -| Auto with zarith ] -]. + forall (n:nat) (x:positive), Zpos (shift_nat n x) = Zpower_nat 2 n * Zpos x. + +unfold shift_nat in |- *; simple induction n; + [ simpl in |- *; trivial with zarith + | intros; replace (Zpower_nat 2 (S n0)) with (2 * Zpower_nat 2 n0); + [ rewrite <- Zmult_assoc; rewrite <- (H x); simpl in |- *; reflexivity + | auto with zarith ] ]. Qed. Theorem two_power_nat_correct : - (n:nat)(two_power_nat n)=(Zpower_nat `2` n). + forall n:nat, two_power_nat n = Zpower_nat 2 n. -Intro n. -Unfold two_power_nat. -Rewrite -> (shift_nat_correct n). -Omega. +intro n. +unfold two_power_nat in |- *. +rewrite (shift_nat_correct n). +omega. Qed. (** Second we show that [two_power_pos] and [two_power_nat] are the same *) -Lemma shift_pos_nat : (p:positive)(x:positive) - (shift_pos p x)=(shift_nat (convert p) x). +Lemma shift_pos_nat : + forall p x:positive, shift_pos p x = shift_nat (nat_of_P p) x. -Unfold shift_pos. -Unfold shift_nat. -Intros; Apply iter_convert. +unfold shift_pos in |- *. +unfold shift_nat in |- *. +intros; apply iter_nat_of_P. Qed. -Lemma two_power_pos_nat : - (p:positive) (two_power_pos p)=(two_power_nat (convert p)). +Lemma two_power_pos_nat : + forall p:positive, two_power_pos p = two_power_nat (nat_of_P p). -Intro; Unfold two_power_pos; Unfold two_power_nat. -Apply f_equal with f:=POS. -Apply shift_pos_nat. +intro; unfold two_power_pos in |- *; unfold two_power_nat in |- *. +apply f_equal with (f := Zpos). +apply shift_pos_nat. Qed. (** Then we deduce that [two_power_pos] is also correct *) Theorem shift_pos_correct : - (p,x:positive) ` (POS (shift_pos p x)) = (Zpower_pos 2 p) * (POS x)`. + forall p x:positive, Zpos (shift_pos p x) = Zpower_pos 2 p * Zpos x. -Intros. -Rewrite -> (shift_pos_nat p x). -Rewrite -> (Zpower_pos_nat `2` p). -Apply shift_nat_correct. +intros. +rewrite (shift_pos_nat p x). +rewrite (Zpower_pos_nat 2 p). +apply shift_nat_correct. Qed. -Theorem two_power_pos_correct : - (x:positive) (two_power_pos x)=(Zpower_pos `2` x). +Theorem two_power_pos_correct : + forall x:positive, two_power_pos x = Zpower_pos 2 x. -Intro. -Rewrite -> two_power_pos_nat. -Rewrite -> Zpower_pos_nat. -Apply two_power_nat_correct. +intro. +rewrite two_power_pos_nat. +rewrite Zpower_pos_nat. +apply two_power_nat_correct. Qed. (** Some consequences *) Theorem two_power_pos_is_exp : - (x,y:positive) (two_power_pos (add x y)) - =(Zmult (two_power_pos x) (two_power_pos y)). -Intros. -Rewrite -> (two_power_pos_correct (add x y)). -Rewrite -> (two_power_pos_correct x). -Rewrite -> (two_power_pos_correct y). -Apply Zpower_pos_is_exp. + forall x y:positive, + two_power_pos (x + y) = two_power_pos x * two_power_pos y. +intros. +rewrite (two_power_pos_correct (x + y)). +rewrite (two_power_pos_correct x). +rewrite (two_power_pos_correct y). +apply Zpower_pos_is_exp. Qed. (** The exponentiation [z -> 2^z] for [z] a signed integer. @@ -211,80 +205,71 @@ Qed. 3 contructors [ Zero | Pos positive -> | minus_infty] but it's more complexe and not so useful. *) -Definition two_p := - [x:Z]Cases x of - ZERO => `1` - | (POS y) => (two_power_pos y) - | (NEG y) => `0` - end. +Definition two_p (x:Z) := + match x with + | Z0 => 1 + | Zpos y => two_power_pos y + | Zneg y => 0 + end. Theorem two_p_is_exp : - (x,y:Z) ` 0 <= x` -> ` 0 <= y` -> - ` (two_p (x+y)) = (two_p x)*(two_p y)`. -Induction x; -[ Induction y; Simpl; Auto with zarith -| Induction y; - [ Unfold two_p; Rewrite -> (Zmult_sym (two_power_pos p) `1`); - Rewrite -> (Zmult_one (two_power_pos p)); Auto with zarith - | Unfold Zplus; Unfold two_p; - Intros; Apply two_power_pos_is_exp - | Intros; Unfold Zle in H0; Unfold Zcompare in H0; - Absurd SUPERIEUR=SUPERIEUR; Trivial with zarith - ] -| Induction y; - [ Simpl; Auto with zarith - | Intros; Unfold Zle in H; Unfold Zcompare in H; - Absurd (SUPERIEUR=SUPERIEUR); Trivial with zarith - | Intros; Unfold Zle in H; Unfold Zcompare in H; - Absurd (SUPERIEUR=SUPERIEUR); Trivial with zarith - ] -]. + forall x y:Z, 0 <= x -> 0 <= y -> two_p (x + y) = two_p x * two_p y. +simple induction x; + [ simple induction y; simpl in |- *; auto with zarith + | simple induction y; + [ unfold two_p in |- *; rewrite (Zmult_comm (two_power_pos p) 1); + rewrite (Zmult_1_l (two_power_pos p)); auto with zarith + | unfold Zplus in |- *; unfold two_p in |- *; intros; + apply two_power_pos_is_exp + | intros; unfold Zle in H0; unfold Zcompare in H0; + absurd (Datatypes.Gt = Datatypes.Gt); trivial with zarith ] + | simple induction y; + [ simpl in |- *; auto with zarith + | intros; unfold Zle in H; unfold Zcompare in H; + absurd (Datatypes.Gt = Datatypes.Gt); trivial with zarith + | intros; unfold Zle in H; unfold Zcompare in H; + absurd (Datatypes.Gt = Datatypes.Gt); trivial with zarith ] ]. Qed. -Lemma two_p_gt_ZERO : (x:Z) ` 0 <= x` -> ` (two_p x) > 0`. -Induction x; Intros; -[ Simpl; Omega -| Simpl; Unfold two_power_pos; Apply POS_gt_ZERO -| Absurd ` 0 <= (NEG p)`; - [ Simpl; Unfold Zle; Unfold Zcompare; - Do 2 Unfold not; Auto with zarith - | Assumption ] -]. +Lemma two_p_gt_ZERO : forall x:Z, 0 <= x -> two_p x > 0. +simple induction x; intros; + [ simpl in |- *; omega + | simpl in |- *; unfold two_power_pos in |- *; apply Zorder.Zgt_pos_0 + | absurd (0 <= Zneg p); + [ simpl in |- *; unfold Zle in |- *; unfold Zcompare in |- *; + do 2 unfold not in |- *; auto with zarith + | assumption ] ]. Qed. -Lemma two_p_S : (x:Z) ` 0 <= x` -> - `(two_p (Zs x)) = 2 * (two_p x)`. -Intros; Unfold Zs. -Rewrite (two_p_is_exp x `1` H (ZERO_le_POS xH)). -Apply Zmult_sym. +Lemma two_p_S : forall x:Z, 0 <= x -> two_p (Zsucc x) = 2 * two_p x. +intros; unfold Zsucc in |- *. +rewrite (two_p_is_exp x 1 H (Zorder.Zle_0_pos 1)). +apply Zmult_comm. Qed. -Lemma two_p_pred : - (x:Z)` 0 <= x` -> ` (two_p (Zpred x)) < (two_p x)`. -Intros; Apply natlike_ind -with P:=[x:Z]` (two_p (Zpred x)) < (two_p x)`; -[ Simpl; Unfold Zlt; Auto with zarith -| Intros; Elim (Zle_lt_or_eq `0` x0 H0); - [ Intros; - Replace (two_p (Zpred (Zs x0))) - with (two_p (Zs (Zpred x0))); - [ Rewrite -> (two_p_S (Zpred x0)); - [ Rewrite -> (two_p_S x0); - [ Omega - | Assumption] - | Apply Zlt_ZERO_pred_le_ZERO; Assumption] - | Rewrite <- (Zs_pred x0); Rewrite <- (Zpred_Sn x0); Trivial with zarith] - | Intro Hx0; Rewrite <- Hx0; Simpl; Unfold Zlt; Auto with zarith] -| Assumption]. +Lemma two_p_pred : forall x:Z, 0 <= x -> two_p (Zpred x) < two_p x. +intros; apply natlike_ind with (P := fun x:Z => two_p (Zpred x) < two_p x); + [ simpl in |- *; unfold Zlt in |- *; auto with zarith + | intros; elim (Zle_lt_or_eq 0 x0 H0); + [ intros; + replace (two_p (Zpred (Zsucc x0))) with (two_p (Zsucc (Zpred x0))); + [ rewrite (two_p_S (Zpred x0)); + [ rewrite (two_p_S x0); [ omega | assumption ] + | apply Zorder.Zlt_0_le_0_pred; assumption ] + | rewrite <- (Zsucc_pred x0); rewrite <- (Zpred_succ x0); + trivial with zarith ] + | intro Hx0; rewrite <- Hx0; simpl in |- *; unfold Zlt in |- *; + auto with zarith ] + | assumption ]. Qed. -Lemma Zlt_lt_double : (x,y:Z) ` 0 <= x < y` -> ` x < 2*y`. -Intros; Omega. Qed. +Lemma Zlt_lt_double : forall x y:Z, 0 <= x < y -> x < 2 * y. +intros; omega. Qed. End Powers_of_2. -Hints Resolve two_p_gt_ZERO : zarith. -Hints Immediate two_p_pred two_p_S : zarith. +Hint Resolve two_p_gt_ZERO: zarith. +Hint Immediate two_p_pred two_p_S: zarith. Section power_div_with_rest. @@ -293,102 +278,95 @@ Section power_div_with_rest. [n = 2^p.q + r] and [0 <= r < 2^p] *) (** Invariant: [d*q + r = d'*q + r /\ d' = 2*d /\ 0<= r < d /\ 0 <= r' < d'] *) -Definition Zdiv_rest_aux := - [qrd:(Z*Z)*Z] - let (qr,d)=qrd in let (q,r)=qr in - (Cases q of - ZERO => ` (0, r)` - | (POS xH) => ` (0, d + r)` - | (POS (xI n)) => ` ((POS n), d + r)` - | (POS (xO n)) => ` ((POS n), r)` - | (NEG xH) => ` (-1, d + r)` - | (NEG (xI n)) => ` ((NEG n) - 1, d + r)` - | (NEG (xO n)) => ` ((NEG n), r)` - end, ` 2*d`). - -Definition Zdiv_rest := - [x:Z][p:positive]let (qr,d)=(iter_pos p ? Zdiv_rest_aux ((x,`0`),`1`)) in qr. +Definition Zdiv_rest_aux (qrd:Z * Z * Z) := + let (qr, d) := qrd in + let (q, r) := qr in + (match q with + | Z0 => (0, r) + | Zpos xH => (0, d + r) + | Zpos (xI n) => (Zpos n, d + r) + | Zpos (xO n) => (Zpos n, r) + | Zneg xH => (-1, d + r) + | Zneg (xI n) => (Zneg n - 1, d + r) + | Zneg (xO n) => (Zneg n, r) + end, 2 * d). + +Definition Zdiv_rest (x:Z) (p:positive) := + let (qr, d) := iter_pos p _ Zdiv_rest_aux (x, 0, 1) in qr. Lemma Zdiv_rest_correct1 : - (x:Z)(p:positive) - let (qr,d)=(iter_pos p ? Zdiv_rest_aux ((x,`0`),`1`)) in d=(two_power_pos p). - -Intros x p; -Rewrite (iter_convert p ? Zdiv_rest_aux ((x,`0`),`1`)); -Rewrite (two_power_pos_nat p); -Elim (convert p); Simpl; -[ Trivial with zarith -| Intro n; Rewrite (two_power_nat_S n); - Unfold 2 Zdiv_rest_aux; - Elim (iter_nat n (Z*Z)*Z Zdiv_rest_aux ((x,`0`),`1`)); - NewDestruct a; Intros; Apply f_equal with f:=[z:Z]`2*z`; Assumption ]. + forall (x:Z) (p:positive), + let (qr, d) := iter_pos p _ Zdiv_rest_aux (x, 0, 1) in d = two_power_pos p. + +intros x p; rewrite (iter_nat_of_P p _ Zdiv_rest_aux (x, 0, 1)); + rewrite (two_power_pos_nat p); elim (nat_of_P p); + simpl in |- *; + [ trivial with zarith + | intro n; rewrite (two_power_nat_S n); unfold Zdiv_rest_aux at 2 in |- *; + elim (iter_nat n (Z * Z * Z) Zdiv_rest_aux (x, 0, 1)); + destruct a; intros; apply f_equal with (f := fun z:Z => 2 * z); + assumption ]. Qed. Lemma Zdiv_rest_correct2 : - (x:Z)(p:positive) - let (qr,d)=(iter_pos p ? Zdiv_rest_aux ((x,`0`),`1`)) in - let (q,r)=qr in - ` x=q*d + r` /\ ` 0 <= r < d`. - -Intros; Apply iter_pos_invariant with - f:=Zdiv_rest_aux - Inv:=[qrd:(Z*Z)*Z]let (qr,d)=qrd in let (q,r)=qr in - ` x=q*d + r` /\ ` 0 <= r < d`; -[ Intro x0; Elim x0; Intro y0; Elim y0; - Intros q r d; Unfold Zdiv_rest_aux; - Elim q; - [ Omega - | NewDestruct p0; - [ Rewrite POS_xI; Intro; Elim H; Intros; Split; - [ Rewrite H0; Rewrite Zplus_assoc; - Rewrite Zmult_plus_distr_l; - Rewrite Zmult_1_n; Rewrite Zmult_assoc; - Rewrite (Zmult_sym (POS p0) `2`); Apply refl_equal - | Omega ] - | Rewrite POS_xO; Intro; Elim H; Intros; Split; - [ Rewrite H0; - Rewrite Zmult_assoc; Rewrite (Zmult_sym (POS p0) `2`); - Apply refl_equal - | Omega ] - | Omega ] - | NewDestruct p0; - [ Rewrite NEG_xI; Unfold Zminus; Intro; Elim H; Intros; Split; - [ Rewrite H0; Rewrite Zplus_assoc; - Apply f_equal with f:=[z:Z]`z+r`; - Do 2 (Rewrite Zmult_plus_distr_l); - Rewrite Zmult_assoc; - Rewrite (Zmult_sym (NEG p0) `2`); - Rewrite <- Zplus_assoc; - Apply f_equal with f:=[z:Z]`2 * (NEG p0) * d + z`; - Omega - | Omega ] - | Rewrite NEG_xO; Unfold Zminus; Intro; Elim H; Intros; Split; - [ Rewrite H0; - Rewrite Zmult_assoc; Rewrite (Zmult_sym (NEG p0) `2`); - Apply refl_equal - | Omega ] - | Omega ] ] -| Omega]. + forall (x:Z) (p:positive), + let (qr, d) := iter_pos p _ Zdiv_rest_aux (x, 0, 1) in + let (q, r) := qr in x = q * d + r /\ 0 <= r < d. + +intros; + apply iter_pos_invariant with + (f := Zdiv_rest_aux) + (Inv := fun qrd:Z * Z * Z => + let (qr, d) := qrd in + let (q, r) := qr in x = q * d + r /\ 0 <= r < d); + [ intro x0; elim x0; intro y0; elim y0; intros q r d; + unfold Zdiv_rest_aux in |- *; elim q; + [ omega + | destruct p0; + [ rewrite BinInt.Zpos_xI; intro; elim H; intros; split; + [ rewrite H0; rewrite Zplus_assoc; rewrite Zmult_plus_distr_l; + rewrite Zmult_1_l; rewrite Zmult_assoc; + rewrite (Zmult_comm (Zpos p0) 2); apply refl_equal + | omega ] + | rewrite BinInt.Zpos_xO; intro; elim H; intros; split; + [ rewrite H0; rewrite Zmult_assoc; rewrite (Zmult_comm (Zpos p0) 2); + apply refl_equal + | omega ] + | omega ] + | destruct p0; + [ rewrite BinInt.Zneg_xI; unfold Zminus in |- *; intro; elim H; intros; + split; + [ rewrite H0; rewrite Zplus_assoc; + apply f_equal with (f := fun z:Z => z + r); + do 2 rewrite Zmult_plus_distr_l; rewrite Zmult_assoc; + rewrite (Zmult_comm (Zneg p0) 2); rewrite <- Zplus_assoc; + apply f_equal with (f := fun z:Z => 2 * Zneg p0 * d + z); + omega + | omega ] + | rewrite BinInt.Zneg_xO; unfold Zminus in |- *; intro; elim H; intros; + split; + [ rewrite H0; rewrite Zmult_assoc; rewrite (Zmult_comm (Zneg p0) 2); + apply refl_equal + | omega ] + | omega ] ] + | omega ]. Qed. -Inductive Set Zdiv_rest_proofs[x:Z; p:positive] := - Zdiv_rest_proof : (q:Z)(r:Z) - `x = q * (two_power_pos p) + r` - -> `0 <= r` - -> `r < (two_power_pos p)` - -> (Zdiv_rest_proofs x p). - -Lemma Zdiv_rest_correct : - (x:Z)(p:positive)(Zdiv_rest_proofs x p). -Intros x p. -Generalize (Zdiv_rest_correct1 x p); Generalize (Zdiv_rest_correct2 x p). -Elim (iter_pos p (Z*Z)*Z Zdiv_rest_aux ((x,`0`),`1`)). -Induction a. -Intros. -Elim H; Intros H1 H2; Clear H. -Rewrite -> H0 in H1; Rewrite -> H0 in H2; -Elim H2; Intros; -Apply Zdiv_rest_proof with q:=a0 r:=b; Assumption. +Inductive Zdiv_rest_proofs (x:Z) (p:positive) : Set := + Zdiv_rest_proof : + forall q r:Z, + x = q * two_power_pos p + r -> + 0 <= r -> r < two_power_pos p -> Zdiv_rest_proofs x p. + +Lemma Zdiv_rest_correct : forall (x:Z) (p:positive), Zdiv_rest_proofs x p. +intros x p. +generalize (Zdiv_rest_correct1 x p); generalize (Zdiv_rest_correct2 x p). +elim (iter_pos p (Z * Z * Z) Zdiv_rest_aux (x, 0, 1)). +simple induction a. +intros. +elim H; intros H1 H2; clear H. +rewrite H0 in H1; rewrite H0 in H2; elim H2; intros; + apply Zdiv_rest_proof with (q := a0) (r := b); assumption. Qed. -End power_div_with_rest. +End power_div_with_rest. \ No newline at end of file diff --git a/theories/ZArith/Zsqrt.v b/theories/ZArith/Zsqrt.v index b8040335e8..f560050804 100644 --- a/theories/ZArith/Zsqrt.v +++ b/theories/ZArith/Zsqrt.v @@ -8,10 +8,9 @@ (* $Id$ *) -Require Omega. +Require Import Omega. Require Export ZArith_base. Require Export ZArithRing. -V7only [Import Z_scope.]. Open Local Scope Z_scope. (**********************************************************************) @@ -19,118 +18,146 @@ Open Local Scope Z_scope. (** The following tactic replaces all instances of (POS (xI ...)) by `2*(POS ...)+1` , but only when ... is not made only with xO, XI, or xH. *) -Tactic Definition compute_POS := - Match Context With - | [|- [(POS (xI ?1))]] -> - (Match ?1 With - | [[xH]] -> Fail - | _ -> Rewrite (POS_xI ?1)) - | [|- [(POS (xO ?1))]] -> - (Match ?1 With - | [[xH]] -> Fail - | _ -> Rewrite (POS_xO ?1)). +Ltac compute_POS := + match goal with + | |- context [(Zpos (xI ?X1))] => + match constr:X1 with + | context [1%positive] => fail + | _ => rewrite (BinInt.Zpos_xI X1) + end + | |- context [(Zpos (xO ?X1))] => + match constr:X1 with + | context [1%positive] => fail + | _ => rewrite (BinInt.Zpos_xO X1) + end + end. -Inductive sqrt_data [n : Z] : Set := - c_sqrt: (s, r :Z)`n=s*s+r`->`0<=r<=2*s`->(sqrt_data n) . +Inductive sqrt_data (n:Z) : Set := + c_sqrt : forall s r:Z, n = s * s + r -> 0 <= r <= 2 * s -> sqrt_data n. -Definition sqrtrempos: (p : positive) (sqrt_data (POS p)). -Refine (Fix sqrtrempos { - sqrtrempos [p : positive] : (sqrt_data (POS p)) := - <[p : ?] (sqrt_data (POS p))> Cases p of - xH => (c_sqrt `1` `1` `0` ? ?) - | (xO xH) => (c_sqrt `2` `1` `1` ? ?) - | (xI xH) => (c_sqrt `3` `1` `2` ? ?) - | (xO (xO p')) => - Cases (sqrtrempos p') of - (c_sqrt s' r' Heq Hint) => - Cases (Z_le_gt_dec `4*s'+1` `4*r'`) of - (left Hle) => - (c_sqrt (POS (xO (xO p'))) `2*s'+1` `4*r'-(4*s'+1)` ? ?) - | (right Hgt) => - (c_sqrt (POS (xO (xO p'))) `2*s'` `4*r'` ? ?) - end - end - | (xO (xI p')) => - Cases (sqrtrempos p') of - (c_sqrt s' r' Heq Hint) => - Cases - (Z_le_gt_dec `4*s'+1` `4*r'+2`) of - (left Hle) => - (c_sqrt - (POS (xO (xI p'))) `2*s'+1` `4*r'+2-(4*s'+1)` ? ?) - | (right Hgt) => - (c_sqrt (POS (xO (xI p'))) `2*s'` `4*r'+2` ? ?) - end - end - | (xI (xO p')) => - Cases (sqrtrempos p') of - (c_sqrt s' r' Heq Hint) => - Cases - (Z_le_gt_dec `4*s'+1` `4*r'+1`) of - (left Hle) => - (c_sqrt - (POS (xI (xO p'))) `2*s'+1` `4*r'+1-(4*s'+1)` ? ?) - | (right Hgt) => - (c_sqrt (POS (xI (xO p'))) `2*s'` `4*r'+1` ? ?) - end - end - | (xI (xI p')) => - Cases (sqrtrempos p') of - (c_sqrt s' r' Heq Hint) => - Cases - (Z_le_gt_dec `4*s'+1` `4*r'+3`) of - (left Hle) => - (c_sqrt - (POS (xI (xI p'))) `2*s'+1` `4*r'+3-(4*s'+1)` ? ?) - | (right Hgt) => - (c_sqrt (POS (xI (xI p'))) `2*s'` `4*r'+3` ? ?) - end - end +Definition sqrtrempos : forall p:positive, sqrt_data (Zpos p). +refine + (fix sqrtrempos (p:positive) : sqrt_data (Zpos p) := + match p return sqrt_data (Zpos p) with + | xH => c_sqrt 1 1 0 _ _ + | xO xH => c_sqrt 2 1 1 _ _ + | xI xH => c_sqrt 3 1 2 _ _ + | xO (xO p') => + match sqrtrempos p' with + | c_sqrt s' r' Heq Hint => + match Z_le_gt_dec (4 * s' + 1) (4 * r') with + | left Hle => + c_sqrt (Zpos (xO (xO p'))) (2 * s' + 1) + (4 * r' - (4 * s' + 1)) _ _ + | right Hgt => c_sqrt (Zpos (xO (xO p'))) (2 * s') (4 * r') _ _ end - }); Clear sqrtrempos; Repeat compute_POS; - Try (Try Rewrite Heq; Ring; Fail); Try Omega. + end + | xO (xI p') => + match sqrtrempos p' with + | c_sqrt s' r' Heq Hint => + match Z_le_gt_dec (4 * s' + 1) (4 * r' + 2) with + | left Hle => + c_sqrt (Zpos (xO (xI p'))) (2 * s' + 1) + (4 * r' + 2 - (4 * s' + 1)) _ _ + | right Hgt => + c_sqrt (Zpos (xO (xI p'))) (2 * s') (4 * r' + 2) _ _ + end + end + | xI (xO p') => + match sqrtrempos p' with + | c_sqrt s' r' Heq Hint => + match Z_le_gt_dec (4 * s' + 1) (4 * r' + 1) with + | left Hle => + c_sqrt (Zpos (xI (xO p'))) (2 * s' + 1) + (4 * r' + 1 - (4 * s' + 1)) _ _ + | right Hgt => + c_sqrt (Zpos (xI (xO p'))) (2 * s') (4 * r' + 1) _ _ + end + end + | xI (xI p') => + match sqrtrempos p' with + | c_sqrt s' r' Heq Hint => + match Z_le_gt_dec (4 * s' + 1) (4 * r' + 3) with + | left Hle => + c_sqrt (Zpos (xI (xI p'))) (2 * s' + 1) + (4 * r' + 3 - (4 * s' + 1)) _ _ + | right Hgt => + c_sqrt (Zpos (xI (xI p'))) (2 * s') (4 * r' + 3) _ _ + end + end + end); clear sqrtrempos; repeat compute_POS; + try (try rewrite Heq; ring; fail); try omega. Defined. (** Define with integer input, but with a strong (readable) specification. *) -Definition Zsqrt : (x:Z)`0<=x`->{s:Z & {r:Z | x=`s*s+r` /\ `s*s<=x<(s+1)*(s+1)`}}. -Refine [x] - <[x:Z]`0<=x`->{s:Z & {r:Z | x=`s*s+r` /\ `s*s<=x<(s+1)*(s+1)`}}>Cases x of - (POS p) => [h]Cases (sqrtrempos p) of - (c_sqrt s r Heq Hint) => - (existS ? [s:Z]{r:Z | `(POS p)=s*s+r` /\ - `s*s<=(POS p)<(s+1)*(s+1)`} - s - (exist Z [r:Z]((POS p)=`s*s+r` /\ `s*s<=(POS p)<(s+1)*(s+1)`) - r ?)) - end - | (NEG p) => [h](False_rec - {s:Z & {r:Z | - (NEG p)=`s*s+r` /\ `s*s<=(NEG p)<(s+1)*(s+1)`}} - (h (refl_equal ? SUPERIEUR))) - | ZERO => [h](existS ? [s:Z]{r:Z | `0=s*s+r` /\ `s*s<=0<(s+1)*(s+1)`} - `0` (exist Z [r:Z](`0=0*0+r`/\`0*0<=0<(0+1)*(0+1)`) - `0` ?)) - end;Try Omega. -Split;[Omega|Rewrite Heq;Ring `(s+1)*(s+1)`;Omega]. +Definition Zsqrt : + forall x:Z, + 0 <= x -> + {s : Z & {r : Z | x = s * s + r /\ s * s <= x < (s + 1) * (s + 1)}}. +refine + (fun x => + match + x + return + 0 <= x -> + {s : Z & {r : Z | x = s * s + r /\ s * s <= x < (s + 1) * (s + 1)}} + with + | Zpos p => + fun h => + match sqrtrempos p with + | c_sqrt s r Heq Hint => + existS + (fun s:Z => + {r : Z | + Zpos p = s * s + r /\ s * s <= Zpos p < (s + 1) * (s + 1)}) + s + (exist + (fun r:Z => + Zpos p = s * s + r /\ + s * s <= Zpos p < (s + 1) * (s + 1)) r _) + end + | Zneg p => + fun h => + False_rec + {s : Z & + {r : Z | + Zneg p = s * s + r /\ s * s <= Zneg p < (s + 1) * (s + 1)}} + (h (refl_equal Datatypes.Gt)) + | Z0 => + fun h => + existS + (fun s:Z => + {r : Z | 0 = s * s + r /\ s * s <= 0 < (s + 1) * (s + 1)}) 0 + (exist + (fun r:Z => 0 = 0 * 0 + r /\ 0 * 0 <= 0 < (0 + 1) * (0 + 1)) 0 + _) + end); try omega. +split; [ omega | rewrite Heq; ring ((s + 1) * (s + 1)); omega ]. Defined. (** Define a function of type Z->Z that computes the integer square root, but only for positive numbers, and 0 for others. *) -Definition Zsqrt_plain : Z->Z := - [x]Cases x of - (POS p)=>Cases (Zsqrt (POS p) (ZERO_le_POS p)) of (existS s _) => s end - |(NEG p)=>`0` - |ZERO=>`0` - end. +Definition Zsqrt_plain (x:Z) : Z := + match x with + | Zpos p => + match Zsqrt (Zpos p) (Zorder.Zle_0_pos p) with + | existS s _ => s + end + | Zneg p => 0 + | Z0 => 0 + end. (** A basic theorem about Zsqrt_plain *) -Theorem Zsqrt_interval :(x:Z)`0<=x`-> - `(Zsqrt_plain x)*(Zsqrt_plain x)<= x < ((Zsqrt_plain x)+1)*((Zsqrt_plain x)+1)`. -Intros x;Case x. -Unfold Zsqrt_plain;Omega. -Intros p;Unfold Zsqrt_plain;Case (Zsqrt (POS p) (ZERO_le_POS p)). -Intros s (r,(Heq,Hint)) Hle;Assumption. -Intros p Hle;Elim Hle;Auto. +Theorem Zsqrt_interval : + forall n:Z, + 0 <= n -> + Zsqrt_plain n * Zsqrt_plain n <= n < + (Zsqrt_plain n + 1) * (Zsqrt_plain n + 1). +intros x; case x. +unfold Zsqrt_plain in |- *; omega. +intros p; unfold Zsqrt_plain in |- *; + case (Zsqrt (Zpos p) (Zorder.Zle_0_pos p)). +intros s [r [Heq Hint]] Hle; assumption. +intros p Hle; elim Hle; auto. Qed. - diff --git a/theories/ZArith/Zsyntax.v b/theories/ZArith/Zsyntax.v deleted file mode 100644 index 5c226b3fce..0000000000 --- a/theories/ZArith/Zsyntax.v +++ /dev/null @@ -1,278 +0,0 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* [$id] - -with number := - -with negnumber := - -with formula : constr := - form_expr [ expr($p) ] -> [$p] -(*| form_eq [ expr($p) "=" expr($c) ] -> [ (eq Z $p $c) ]*) -| form_eq [ expr($p) "=" expr($c) ] -> [ (Coq.Init.Logic.eq ? $p $c) ] -| form_le [ expr($p) "<=" expr($c) ] -> [ (Zle $p $c) ] -| form_lt [ expr($p) "<" expr($c) ] -> [ (Zlt $p $c) ] -| form_ge [ expr($p) ">=" expr($c) ] -> [ (Zge $p $c) ] -| form_gt [ expr($p) ">" expr($c) ] -> [ (Zgt $p $c) ] -(*| form_eq_eq [ expr($p) "=" expr($c) "=" expr($c1) ] - -> [ (eq Z $p $c)/\(eq Z $c $c1) ]*) -| form_eq_eq [ expr($p) "=" expr($c) "=" expr($c1) ] - -> [ (Coq.Init.Logic.eq ? $p $c)/\(Coq.Init.Logic.eq ? $c $c1) ] -| form_le_le [ expr($p) "<=" expr($c) "<=" expr($c1) ] - -> [ (Zle $p $c)/\(Zle $c $c1) ] -| form_le_lt [ expr($p) "<=" expr($c) "<" expr($c1) ] - -> [ (Zle $p $c)/\(Zlt $c $c1) ] -| form_lt_le [ expr($p) "<" expr($c) "<=" expr($c1) ] - -> [ (Zlt $p $c)/\(Zle $c $c1) ] -| form_lt_lt [ expr($p) "<" expr($c) "<" expr($c1) ] - -> [ (Zlt $p $c)/\(Zlt $c $c1) ] -(*| form_neq [ expr($p) "<>" expr($c) ] -> [ ~(Coq.Init.Logic.eq Z $p $c) ]*) -| form_neq [ expr($p) "<>" expr($c) ] -> [ ~(Coq.Init.Logic.eq ? $p $c) ] -| form_comp [ expr($p) "?=" expr($c) ] -> [ (Zcompare $p $c) ] - -with expr : constr := - expr_plus [ expr($p) "+" expr($c) ] -> [ (Zplus $p $c) ] -| expr_minus [ expr($p) "-" expr($c) ] -> [ (Zminus $p $c) ] -| expr2 [ expr2($e) ] -> [$e] - -with expr2 : constr := - expr_mult [ expr2($p) "*" expr2($c) ] -> [ (Zmult $p $c) ] -| expr1 [ expr1($e) ] -> [$e] - -with expr1 : constr := - expr_abs [ "|" expr($c) "|" ] -> [ (Zabs $c) ] -| expr0 [ expr0($e) ] -> [$e] - -with expr0 : constr := - expr_id [ constr:global($c) ] -> [ $c ] -| expr_com [ "[" constr:constr($c) "]" ] -> [$c] -| expr_appl [ "(" application($a) ")" ] -> [$a] -| expr_num [ number($s) ] -> [$s ] -| expr_negnum [ "-" negnumber($n) ] -> [ $n ] -| expr_inv [ "-" expr0($c) ] -> [ (Zopp $c) ] -| expr_meta [ zmeta($m) ] -> [ $m ] - -with zmeta := -| rimpl [ "?" ] -> [ ? ] -| rmeta0 [ "?" "0" ] -> [ ?0 ] -| rmeta1 [ "?" "1" ] -> [ ?1 ] -| rmeta2 [ "?" "2" ] -> [ ?2 ] -| rmeta3 [ "?" "3" ] -> [ ?3 ] -| rmeta4 [ "?" "4" ] -> [ ?4 ] -| rmeta5 [ "?" "5" ] -> [ ?5 ] - -with application : constr := - apply [ application($p) expr($c1) ] -> [ ($p $c1) ] -| apply_inject_nat [ "inject_nat" constr:constr($c1) ] -> [ (inject_nat $c1) ] -| pair [ expr($p) "," expr($c) ] -> [ ($p, $c) ] -| appl0 [ expr($a) ] -> [$a] -. - -Grammar constr constr0 := - z_in_com [ "`" znatural:formula($c) "`" ] -> [$c]. - -Grammar constr pattern := - z_in_pattern [ "`" prim:bigint($c) "`" ] -> [ 'Z: $c ' ]. - -(* The symbols "`" "`" must be printed just once at the top of the expressions, - to avoid printings like |``x` + `y`` < `45`| - for |x + y < 45|. - So when a Z-expression is to be printed, its sub-expresssions are - enclosed into an ast (ZEXPR \$subexpr), which is printed like \$subexpr - but without symbols "`" "`" around. - - There is just one problem: NEG and Zopp have the same printing rules. - If Zopp is opaque, we may not be able to solve a goal like - ` -5 = -5 ` by reflexivity. (In fact, this precise Goal is solved - by the Reflexivity tactic, but more complex problems may arise - - SOLUTION : Print (Zopp 5) for constants and -x for variables *) - -Syntax constr - level 0: - Zle [ (Zle $n1 $n2) ] -> - [[ "`" (ZEXPR $n1) [1 0] "<= " (ZEXPR $n2) "`"]] - | Zlt [ (Zlt $n1 $n2) ] -> - [[ "`" (ZEXPR $n1) [1 0] "< " (ZEXPR $n2) "`" ]] - | Zge [ (Zge $n1 $n2) ] -> - [[ "`" (ZEXPR $n1) [1 0] ">= " (ZEXPR $n2) "`" ]] - | Zgt [ (Zgt $n1 $n2) ] -> - [[ "`" (ZEXPR $n1) [1 0] "> " (ZEXPR $n2) "`" ]] - | Zcompare [<<(Zcompare $n1 $n2)>>] -> - [[ "`" (ZEXPR $n1) [1 0] "?= " (ZEXPR $n2) "`" ]] - | Zeq [ (eq Z $n1 $n2) ] -> - [[ "`" (ZEXPR $n1) [1 0] "= " (ZEXPR $n2)"`"]] - | Zneq [ ~(eq Z $n1 $n2) ] -> - [[ "`" (ZEXPR $n1) [1 0] "<> " (ZEXPR $n2) "`"]] - | Zle_Zle [ (Zle $n1 $n2)/\(Zle $n2 $n3) ] -> - [[ "`" (ZEXPR $n1) [1 0] "<= " (ZEXPR $n2) - [1 0] "<= " (ZEXPR $n3) "`"]] - | Zle_Zlt [ (Zle $n1 $n2)/\(Zlt $n2 $n3) ] -> - [[ "`" (ZEXPR $n1) [1 0] "<= " (ZEXPR $n2) - [1 0] "< " (ZEXPR $n3) "`"]] - | Zlt_Zle [ (Zlt $n1 $n2)/\(Zle $n2 $n3) ] -> - [[ "`" (ZEXPR $n1) [1 0] "< " (ZEXPR $n2) - [1 0] "<= " (ZEXPR $n3) "`"]] - | Zlt_Zlt [ (Zlt $n1 $n2)/\(Zlt $n2 $n3) ] -> - [[ "`" (ZEXPR $n1) [1 0] "< " (ZEXPR $n2) - [1 0] "< " (ZEXPR $n3) "`"]] - | ZZero_v7 [ ZERO ] -> [ "`0`" ] - | ZPos_v7 [ (POS $r) ] -> [$r:"positive_printer":9] - | ZNeg_v7 [ (NEG $r) ] -> [$r:"negative_printer":9] - ; - - level 7: - Zplus [ (Zplus $n1 $n2) ] - -> [ [ "`" (ZEXPR $n1):E "+" [0 0] (ZEXPR $n2):L "`"] ] - | Zminus [ (Zminus $n1 $n2) ] - -> [ [ "`" (ZEXPR $n1):E "-" [0 0] (ZEXPR $n2):L "`"] ] - ; - - level 6: - Zmult [ (Zmult $n1 $n2) ] - -> [ [ "`" (ZEXPR $n1):E "*" [0 0] (ZEXPR $n2):L "`"] ] - ; - - level 8: - Zopp [ (Zopp $n1) ] -> [ [ "`" "-" (ZEXPR $n1):E "`"] ] - | Zopp_POS [ (Zopp (POS $r)) ] -> - [ [ "`(" "Zopp" [1 0] $r:"positive_printer_inside" ")`"] ] - | Zopp_ZERO [ (Zopp ZERO) ] -> [ [ "`(" "Zopp" [1 0] "0" ")`"] ] - | Zopp_NEG [ (Zopp (NEG $r)) ] -> - [ [ "`(" "Zopp" [1 0] "(" $r:"negative_printer_inside" "))`"] ] - ; - - level 4: - Zabs [ (Zabs $n1) ] -> [ [ "`|" (ZEXPR $n1):E "|`"] ] - ; - - level 0: - escape_inside [ << (ZEXPR $r) >> ] -> [ "[" $r:E "]" ] - ; - - level 4: - Zappl_inside [ << (ZEXPR (APPLIST $h ($LIST $t))) >> ] - -> [ [ "("(ZEXPR $h):E [1 0] (ZAPPLINSIDETAIL ($LIST $t)):E ")"] ] - | Zappl_inject_nat [ << (ZEXPR (APPLIST <> $n)) >> ] - -> [ [ "(inject_nat" [1 1] $n:L ")"] ] - | Zappl_inside_tail [ << (ZAPPLINSIDETAIL $h ($LIST $t)) >> ] - -> [(ZEXPR $h):E [1 0] (ZAPPLINSIDETAIL ($LIST $t)):E] - | Zappl_inside_one [ << (ZAPPLINSIDETAIL $e) >> ] ->[(ZEXPR $e):E] - | pair_inside [ << (ZEXPR <<(pair $s1 $s2 $z1 $z2)>>) >> ] - -> [ [ "("(ZEXPR $z1):E "," [1 0] (ZEXPR $z2):E ")"] ] - ; - - level 3: - var_inside [ << (ZEXPR ($VAR $i)) >> ] -> [$i] - | secvar_inside [ << (ZEXPR (SECVAR $i)) >> ] -> [(SECVAR $i)] - | const_inside [ << (ZEXPR (CONST $c)) >> ] -> [(CONST $c)] - | mutind_inside [ << (ZEXPR (MUTIND $i $n)) >> ] - -> [(MUTIND $i $n)] - | mutconstruct_inside [ << (ZEXPR (MUTCONSTRUCT $c1 $c2 $c3)) >> ] - -> [ (MUTCONSTRUCT $c1 $c2 $c3) ] - - | O_inside [ << (ZEXPR << O >>) >> ] -> [ "O" ] (* To shunt Arith printer *) - - (* Added by JCF, 9/3/98; updated HH, 11/9/01 *) - | implicit_head_inside [ << (ZEXPR (APPLISTEXPL ($LIST $c))) >> ] - -> [ (APPLIST ($LIST $c)) ] - | implicit_arg_inside [ << (ZEXPR (EXPL "!" $n $c)) >> ] -> [ ] - - ; - - level 7: - Zplus_inside - [ << (ZEXPR <<(Zplus $n1 $n2)>>) >> ] - -> [ (ZEXPR $n1):E "+" [0 0] (ZEXPR $n2):L ] - | Zminus_inside - [ << (ZEXPR <<(Zminus $n1 $n2)>>) >> ] - -> [ (ZEXPR $n1):E "-" [0 0] (ZEXPR $n2):L ] - ; - - level 6: - Zmult_inside - [ << (ZEXPR <<(Zmult $n1 $n2)>>) >> ] - -> [ (ZEXPR $n1):E "*" [0 0] (ZEXPR $n2):L ] - ; - - level 5: - Zopp_inside [ << (ZEXPR <<(Zopp $n1)>>) >> ] -> [ "(-" (ZEXPR $n1):E ")" ] - ; - - level 10: - Zopp_POS_inside [ << (ZEXPR <<(Zopp (POS $r))>>) >> ] -> - [ [ "Zopp" [1 0] $r:"positive_printer_inside" ] ] - | Zopp_ZERO_inside [ << (ZEXPR <<(Zopp ZERO)>>) >> ] -> - [ [ "Zopp" [1 0] "0"] ] - | Zopp_NEG_inside [ << (ZEXPR <<(Zopp (NEG $r))>>) >> ] -> - [ [ "Zopp" [1 0] $r:"negative_printer_inside" ] ] - ; - - level 4: - Zabs_inside [ << (ZEXPR <<(Zabs $n1)>>) >> ] -> [ "|" (ZEXPR $n1) "|"] - ; - - level 0: - ZZero_inside [ << (ZEXPR <>) >> ] -> ["0"] - | ZPos_inside [ << (ZEXPR <<(POS $p)>>) >>] -> - [$p:"positive_printer_inside":9] - | ZNeg_inside [ << (ZEXPR <<(NEG $p)>>) >>] -> - [$p:"negative_printer_inside":9] -. -]. - -V7only[ -(* For parsing/printing based on scopes *) -Module Z_scope. - -Infix LEFTA 4 "+" Zplus : Z_scope. -Infix LEFTA 4 "-" Zminus : Z_scope. -Infix LEFTA 3 "*" Zmult : Z_scope. -Notation "- x" := (Zopp x) (at level 0): Z_scope V8only. -Infix NONA 5 "<=" Zle : Z_scope. -Infix NONA 5 "<" Zlt : Z_scope. -Infix NONA 5 ">=" Zge : Z_scope. -Infix NONA 5 ">" Zgt : Z_scope. -Infix NONA 5 "?=" Zcompare : Z_scope. -Notation "x <= y <= z" := (Zle x y)/\(Zle y z) - (at level 5, y at level 4):Z_scope - V8only (at level 70, y at next level). -Notation "x <= y < z" := (Zle x y)/\(Zlt y z) - (at level 5, y at level 4):Z_scope - V8only (at level 70, y at next level). -Notation "x < y < z" := (Zlt x y)/\(Zlt y z) - (at level 5, y at level 4):Z_scope - V8only (at level 70, y at next level). -Notation "x < y <= z" := (Zlt x y)/\(Zle y z) - (at level 5, y at level 4):Z_scope - V8only (at level 70, y at next level). -Notation "x = y = z" := x=y/\y=z : Z_scope - V8only (at level 70, y at next level). - -(* Now a polymorphic notation -Notation "x <> y" := ~(eq Z x y) (at level 5, no associativity) : Z_scope. -*) - -(* Notation "| x |" (Zabs x) : Z_scope.(* "|" conflicts with THENS *)*) - -(* Overwrite the printing of "`x = y`" *) -Syntax constr level 0: - Zeq [ (eq Z $n1 $n2) ] -> [[ $n1 [1 0] "= " $n2 ]]. - -Open Scope Z_scope. - -End Z_scope. -]. diff --git a/theories/ZArith/Zwf.v b/theories/ZArith/Zwf.v index 5468f82cc8..7f91b0f6f7 100644 --- a/theories/ZArith/Zwf.v +++ b/theories/ZArith/Zwf.v @@ -8,10 +8,9 @@ (* $Id$ *) -Require ZArith_base. +Require Import ZArith_base. Require Export Wf_nat. -Require Omega. -V7only [Import Z_scope.]. +Require Import Omega. Open Local Scope Z_scope. (** Well-founded relations on Z. *) @@ -21,7 +20,7 @@ Open Local Scope Z_scope. [x (Zwf c) y] iff [x < y & c <= y] *) -Definition Zwf := [c:Z][x,y:Z] `c <= y` /\ `x < y`. +Definition Zwf (c x y:Z) := c <= y /\ x < y. (** and we prove that [(Zwf c)] is well founded *) @@ -32,34 +31,34 @@ Variable c : Z. (** The proof of well-foundness is classic: we do the proof by induction on a measure in nat, which is here [|x-c|] *) -Local f := [z:Z](absolu (Zminus z c)). +Let f (z:Z) := Zabs_nat (z - c). -Lemma Zwf_well_founded : (well_founded Z (Zwf c)). -Red; Intros. -Assert (n:nat)(a:Z)(lt (f a) n)\/(`a (Acc Z (Zwf c) a). -Clear a; Induction n; Intros. +Lemma Zwf_well_founded : well_founded (Zwf c). +red in |- *; intros. +assert (forall (n:nat) (a:Z), (f a < n)%nat \/ a < c -> Acc (Zwf c) a). +clear a; simple induction n; intros. (** n= 0 *) -Case H; Intros. -Case (lt_n_O (f a)); Auto. -Apply Acc_intro; Unfold Zwf; Intros. -Assert False;Omega Orelse Contradiction. +case H; intros. +case (lt_n_O (f a)); auto. +apply Acc_intro; unfold Zwf in |- *; intros. +assert False; omega || contradiction. (** inductive case *) -Case H0; Clear H0; Intro; Auto. -Apply Acc_intro; Intros. -Apply H. -Unfold Zwf in H1. -Case (Zle_or_lt c y); Intro; Auto with zarith. -Left. -Red in H0. -Apply lt_le_trans with (f a); Auto with arith. -Unfold f. -Apply absolu_lt; Omega. -Apply (H (S (f a))); Auto. -Save. +case H0; clear H0; intro; auto. +apply Acc_intro; intros. +apply H. +unfold Zwf in H1. +case (Zle_or_lt c y); intro; auto with zarith. +left. +red in H0. +apply lt_le_trans with (f a); auto with arith. +unfold f in |- *. +apply Zabs.Zabs_nat_lt; omega. +apply (H (S (f a))); auto. +Qed. End wf_proof. -Hints Resolve Zwf_well_founded : datatypes v62. +Hint Resolve Zwf_well_founded: datatypes v62. (** We also define the other family of relations: @@ -67,7 +66,7 @@ Hints Resolve Zwf_well_founded : datatypes v62. [x (Zwf_up c) y] iff [y < x <= c] *) -Definition Zwf_up := [c:Z][x,y:Z] `y < x <= c`. +Definition Zwf_up (c x y:Z) := y < x <= c. (** and we prove that [(Zwf_up c)] is well founded *) @@ -78,19 +77,20 @@ Variable c : Z. (** The proof of well-foundness is classic: we do the proof by induction on a measure in nat, which is here [|c-x|] *) -Local f := [z:Z](absolu (Zminus c z)). +Let f (z:Z) := Zabs_nat (c - z). -Lemma Zwf_up_well_founded : (well_founded Z (Zwf_up c)). +Lemma Zwf_up_well_founded : well_founded (Zwf_up c). Proof. -Apply well_founded_lt_compat with f:=f. -Unfold Zwf_up f. -Intros. -Apply absolu_lt. -Unfold Zminus. Split. -Apply Zle_left; Intuition. -Apply Zlt_reg_l; Unfold Zlt; Rewrite <- Zcompare_Zopp; Intuition. -Save. +apply well_founded_lt_compat with (f := f). +unfold Zwf_up, f in |- *. +intros. +apply Zabs.Zabs_nat_lt. +unfold Zminus in |- *. split. +apply Zle_left; intuition. +apply Zplus_lt_compat_l; unfold Zlt in |- *; rewrite <- Zcompare_opp; + intuition. +Qed. End wf_proof_up. -Hints Resolve Zwf_up_well_founded : datatypes v62. +Hint Resolve Zwf_up_well_founded: datatypes v62. \ No newline at end of file diff --git a/theories/ZArith/auxiliary.v b/theories/ZArith/auxiliary.v index 3f713c5edc..50c22b1b4e 100644 --- a/theories/ZArith/auxiliary.v +++ b/theories/ZArith/auxiliary.v @@ -11,10 +11,10 @@ (** Binary Integers (Pierre Crégut, CNET, Lannion, France) *) Require Export Arith. -Require BinInt. -Require Zorder. -Require Decidable. -Require Peano_dec. +Require Import BinInt. +Require Import Zorder. +Require Import Decidable. +Require Import Peano_dec. Require Export Compare_dec. Open Local Scope Z_scope. @@ -22,198 +22,129 @@ Open Local Scope Z_scope. (**********************************************************************) (** Moving terms from one side to the other of an inequality *) -Theorem Zne_left : (x,y:Z) (Zne x y) -> (Zne (Zplus x (Zopp y)) ZERO). +Theorem Zne_left : forall n m:Z, Zne n m -> Zne (n + - m) 0. Proof. -Intros x y; Unfold Zne; Unfold not; Intros H1 H2; Apply H1; -Apply Zsimpl_plus_l with (Zopp y); Rewrite Zplus_inverse_l; Rewrite Zplus_sym; -Trivial with arith. +intros x y; unfold Zne in |- *; unfold not in |- *; intros H1 H2; apply H1; + apply Zplus_reg_l with (- y); rewrite Zplus_opp_l; + rewrite Zplus_comm; trivial with arith. Qed. -Theorem Zegal_left : (x,y:Z) (x=y) -> (Zplus x (Zopp y)) = ZERO. +Theorem Zegal_left : forall n m:Z, n = m -> n + - m = 0. Proof. -Intros x y H; -Apply (Zsimpl_plus_l y);Rewrite -> Zplus_permute; -Rewrite -> Zplus_inverse_r;Do 2 Rewrite -> Zero_right;Assumption. +intros x y H; apply (Zplus_reg_l y); rewrite Zplus_permute; + rewrite Zplus_opp_r; do 2 rewrite Zplus_0_r; assumption. Qed. -Theorem Zle_left : (x,y:Z) (Zle x y) -> (Zle ZERO (Zplus y (Zopp x))). +Theorem Zle_left : forall n m:Z, n <= m -> 0 <= m + - n. Proof. -Intros x y H; Replace ZERO with (Zplus x (Zopp x)). -Apply Zle_reg_r; Trivial. -Apply Zplus_inverse_r. +intros x y H; replace 0 with (x + - x). +apply Zplus_le_compat_r; trivial. +apply Zplus_opp_r. Qed. -Theorem Zle_left_rev : (x,y:Z) (Zle ZERO (Zplus y (Zopp x))) - -> (Zle x y). +Theorem Zle_left_rev : forall n m:Z, 0 <= m + - n -> n <= m. Proof. -Intros x y H; Apply Zsimpl_le_plus_r with (Zopp x). -Rewrite Zplus_inverse_r; Trivial. +intros x y H; apply Zplus_le_reg_r with (- x). +rewrite Zplus_opp_r; trivial. Qed. -Theorem Zlt_left_rev : (x,y:Z) (Zlt ZERO (Zplus y (Zopp x))) - -> (Zlt x y). +Theorem Zlt_left_rev : forall n m:Z, 0 < m + - n -> n < m. Proof. -Intros x y H; Apply Zsimpl_lt_plus_r with (Zopp x). -Rewrite Zplus_inverse_r; Trivial. +intros x y H; apply Zplus_lt_reg_r with (- x). +rewrite Zplus_opp_r; trivial. Qed. -Theorem Zlt_left : - (x,y:Z) (Zlt x y) -> (Zle ZERO (Zplus (Zplus y (NEG xH)) (Zopp x))). +Theorem Zlt_left : forall n m:Z, n < m -> 0 <= m + -1 + - n. Proof. -Intros x y H; Apply Zle_left; Apply Zle_S_n; -Change (Zle (Zs x) (Zs (Zpred y))); Rewrite <- Zs_pred; Apply Zlt_le_S; -Assumption. +intros x y H; apply Zle_left; apply Zsucc_le_reg; + change (Zsucc x <= Zsucc (Zpred y)) in |- *; rewrite <- Zsucc_pred; + apply Zlt_le_succ; assumption. Qed. -Theorem Zlt_left_lt : - (x,y:Z) (Zlt x y) -> (Zlt ZERO (Zplus y (Zopp x))). +Theorem Zlt_left_lt : forall n m:Z, n < m -> 0 < m + - n. Proof. -Intros x y H; Replace ZERO with (Zplus x (Zopp x)). -Apply Zlt_reg_r; Trivial. -Apply Zplus_inverse_r. +intros x y H; replace 0 with (x + - x). +apply Zplus_lt_compat_r; trivial. +apply Zplus_opp_r. Qed. -Theorem Zge_left : (x,y:Z) (Zge x y) -> (Zle ZERO (Zplus x (Zopp y))). +Theorem Zge_left : forall n m:Z, n >= m -> 0 <= n + - m. Proof. -Intros x y H; Apply Zle_left; Apply Zge_le; Assumption. +intros x y H; apply Zle_left; apply Zge_le; assumption. Qed. -Theorem Zgt_left : - (x,y:Z) (Zgt x y) -> (Zle ZERO (Zplus (Zplus x (NEG xH)) (Zopp y))). +Theorem Zgt_left : forall n m:Z, n > m -> 0 <= n + -1 + - m. Proof. -Intros x y H; Apply Zlt_left; Apply Zgt_lt; Assumption. +intros x y H; apply Zlt_left; apply Zgt_lt; assumption. Qed. -Theorem Zgt_left_gt : - (x,y:Z) (Zgt x y) -> (Zgt (Zplus x (Zopp y)) ZERO). +Theorem Zgt_left_gt : forall n m:Z, n > m -> n + - m > 0. Proof. -Intros x y H; Replace ZERO with (Zplus y (Zopp y)). -Apply Zgt_reg_r; Trivial. -Apply Zplus_inverse_r. +intros x y H; replace 0 with (y + - y). +apply Zplus_gt_compat_r; trivial. +apply Zplus_opp_r. Qed. -Theorem Zgt_left_rev : (x,y:Z) (Zgt (Zplus x (Zopp y)) ZERO) - -> (Zgt x y). +Theorem Zgt_left_rev : forall n m:Z, n + - m > 0 -> n > m. Proof. -Intros x y H; Apply Zsimpl_gt_plus_r with (Zopp y). -Rewrite Zplus_inverse_r; Trivial. +intros x y H; apply Zplus_gt_reg_r with (- y). +rewrite Zplus_opp_r; trivial. Qed. (**********************************************************************) (** Factorization lemmas *) -Theorem Zred_factor0 : (x:Z) x = (Zmult x (POS xH)). -Intro x; Rewrite (Zmult_n_1 x); Reflexivity. +Theorem Zred_factor0 : forall n:Z, n = n * 1. +intro x; rewrite (Zmult_1_r x); reflexivity. Qed. -Theorem Zred_factor1 : (x:Z) (Zplus x x) = (Zmult x (POS (xO xH))). +Theorem Zred_factor1 : forall n:Z, n + n = n * 2. Proof. -Exact Zplus_Zmult_2. -Qed. - -Theorem Zred_factor2 : - (x,y:Z) (Zplus x (Zmult x y)) = (Zmult x (Zplus (POS xH) y)). - -Intros x y; Pattern 1 x ; Rewrite <- (Zmult_n_1 x); -Rewrite <- Zmult_plus_distr_r; Trivial with arith. -Qed. - -Theorem Zred_factor3 : - (x,y:Z) (Zplus (Zmult x y) x) = (Zmult x (Zplus (POS xH) y)). - -Intros x y; Pattern 2 x ; Rewrite <- (Zmult_n_1 x); -Rewrite <- Zmult_plus_distr_r; Rewrite Zplus_sym; Trivial with arith. -Qed. -Theorem Zred_factor4 : - (x,y,z:Z) (Zplus (Zmult x y) (Zmult x z)) = (Zmult x (Zplus y z)). -Intros x y z; Symmetry; Apply Zmult_plus_distr_r. -Qed. - -Theorem Zred_factor5 : (x,y:Z) (Zplus (Zmult x ZERO) y) = y. - -Intros x y; Rewrite <- Zmult_n_O;Auto with arith. -Qed. - -Theorem Zred_factor6 : (x:Z) x = (Zplus x ZERO). - -Intro; Rewrite Zero_right; Trivial with arith. -Qed. - -Theorem Zle_mult_approx: - (x,y,z:Z) (Zgt x ZERO) -> (Zgt z ZERO) -> (Zle ZERO y) -> - (Zle ZERO (Zplus (Zmult y x) z)). - -Intros x y z H1 H2 H3; Apply Zle_trans with m:=(Zmult y x) ; [ - Apply Zle_mult; Assumption -| Pattern 1 (Zmult y x) ; Rewrite <- Zero_right; Apply Zle_reg_l; - Apply Zlt_le_weak; Apply Zgt_lt; Assumption]. -Qed. - -Theorem Zmult_le_approx: - (x,y,z:Z) (Zgt x ZERO) -> (Zgt x z) -> - (Zle ZERO (Zplus (Zmult y x) z)) -> (Zle ZERO y). - -Intros x y z H1 H2 H3; Apply Zlt_n_Sm_le; Apply Zmult_lt with x; [ - Assumption - | Apply Zle_lt_trans with 1:=H3 ; Rewrite <- Zmult_Sm_n; - Apply Zlt_reg_l; Apply Zgt_lt; Assumption]. - -Qed. - -V7only [ -(* Compatibility *) -Require Znat. -Require Zcompare. -Notation neq := neq. -Notation Zne := Zne. -Notation OMEGA2 := Zle_0_plus. -Notation add_un_Zs := add_un_Zs. -Notation inj_S := inj_S. -Notation Zplus_S_n := Zplus_S_n. -Notation inj_plus := inj_plus. -Notation inj_mult := inj_mult. -Notation inj_neq := inj_neq. -Notation inj_le := inj_le. -Notation inj_lt := inj_lt. -Notation inj_gt := inj_gt. -Notation inj_ge := inj_ge. -Notation inj_eq := inj_eq. -Notation intro_Z := intro_Z. -Notation inj_minus1 := inj_minus1. -Notation inj_minus2 := inj_minus2. -Notation dec_eq := dec_eq. -Notation dec_Zne := dec_Zne. -Notation dec_Zle := dec_Zle. -Notation dec_Zgt := dec_Zgt. -Notation dec_Zge := dec_Zge. -Notation dec_Zlt := dec_Zlt. -Notation dec_eq_nat := dec_eq_nat. -Notation not_Zge := not_Zge. -Notation not_Zlt := not_Zlt. -Notation not_Zle := not_Zle. -Notation not_Zgt := not_Zgt. -Notation not_Zeq := not_Zeq. -Notation Zopp_one := Zopp_one. -Notation Zopp_Zmult_r := Zopp_Zmult_r. -Notation Zmult_Zopp_left := Zmult_Zopp_left. -Notation Zopp_Zmult_l := Zopp_Zmult_l. -Notation Zcompare_Zplus_compatible2 := Zcompare_Zplus_compatible2. -Notation Zcompare_Zmult_compatible := Zcompare_Zmult_compatible. -Notation Zmult_eq := Zmult_eq. -Notation Z_eq_mult := Z_eq_mult. -Notation Zmult_le := Zmult_le. -Notation Zle_ZERO_mult := Zle_ZERO_mult. -Notation Zgt_ZERO_mult := Zgt_ZERO_mult. -Notation Zle_mult := Zle_mult. -Notation Zmult_lt := Zmult_lt. -Notation Zmult_gt := Zmult_gt. -Notation Zle_Zmult_pos_right := Zle_Zmult_pos_right. -Notation Zle_Zmult_pos_left := Zle_Zmult_pos_left. -Notation Zge_Zmult_pos_right := Zge_Zmult_pos_right. -Notation Zge_Zmult_pos_left := Zge_Zmult_pos_left. -Notation Zge_Zmult_pos_compat := Zge_Zmult_pos_compat. -Notation Zle_mult_simpl := Zle_mult_simpl. -Notation Zge_mult_simpl := Zge_mult_simpl. -Notation Zgt_mult_simpl := Zgt_mult_simpl. -Notation Zgt_square_simpl := Zgt_square_simpl. -]. +exact Zplus_diag_eq_mult_2. +Qed. + +Theorem Zred_factor2 : forall n m:Z, n + n * m = n * (1 + m). + +intros x y; pattern x at 1 in |- *; rewrite <- (Zmult_1_r x); + rewrite <- Zmult_plus_distr_r; trivial with arith. +Qed. + +Theorem Zred_factor3 : forall n m:Z, n * m + n = n * (1 + m). + +intros x y; pattern x at 2 in |- *; rewrite <- (Zmult_1_r x); + rewrite <- Zmult_plus_distr_r; rewrite Zplus_comm; + trivial with arith. +Qed. +Theorem Zred_factor4 : forall n m p:Z, n * m + n * p = n * (m + p). +intros x y z; symmetry in |- *; apply Zmult_plus_distr_r. +Qed. + +Theorem Zred_factor5 : forall n m:Z, n * 0 + m = m. + +intros x y; rewrite <- Zmult_0_r_reverse; auto with arith. +Qed. + +Theorem Zred_factor6 : forall n:Z, n = n + 0. + +intro; rewrite Zplus_0_r; trivial with arith. +Qed. + +Theorem Zle_mult_approx : + forall n m p:Z, n > 0 -> p > 0 -> 0 <= m -> 0 <= m * n + p. + +intros x y z H1 H2 H3; apply Zle_trans with (m := y * x); + [ apply Zmult_gt_0_le_0_compat; assumption + | pattern (y * x) at 1 in |- *; rewrite <- Zplus_0_r; + apply Zplus_le_compat_l; apply Zlt_le_weak; apply Zgt_lt; + assumption ]. +Qed. + +Theorem Zmult_le_approx : + forall n m p:Z, n > 0 -> n > p -> 0 <= m * n + p -> 0 <= m. + +intros x y z H1 H2 H3; apply Zlt_succ_le; apply Zmult_gt_0_lt_0_reg_r with x; + [ assumption + | apply Zle_lt_trans with (1 := H3); rewrite <- Zmult_succ_l_reverse; + apply Zplus_lt_compat_l; apply Zgt_lt; assumption ]. + +Qed. diff --git a/theories/ZArith/fast_integer.v b/theories/ZArith/fast_integer.v deleted file mode 100644 index 81b69037f8..0000000000 --- a/theories/ZArith/fast_integer.v +++ /dev/null @@ -1,191 +0,0 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* positive;y:positive](times x y). -Notation times1_convert := - [x,y:positive;_:positive->positive](times_convert x y). - -Notation Z := Z. -Notation POS := POS. -Notation NEG := NEG. -Notation ZERO := ZERO. -Notation Zero_left := Zero_left. -Notation Zopp_Zopp := Zopp_Zopp. -Notation Zero_right := Zero_right. -Notation Zplus_inverse_r := Zplus_inverse_r. -Notation Zopp_Zplus := Zopp_Zplus. -Notation Zplus_sym := Zplus_sym. -Notation Zplus_inverse_l := Zplus_inverse_l. -Notation Zopp_intro := Zopp_intro. -Notation Zopp_NEG := Zopp_NEG. -Notation weak_assoc := weak_assoc. -Notation Zplus_assoc := Zplus_assoc. -Notation Zplus_simpl := Zplus_simpl. -Notation Zmult_sym := Zmult_sym. -Notation Zmult_assoc := Zmult_assoc. -Notation Zmult_one := Zmult_one. -Notation lt_mult_left := lt_mult_left. (* Mult*) -Notation Zero_mult_left := Zero_mult_left. -Notation Zero_mult_right := Zero_mult_right. -Notation Zopp_Zmult := Zopp_Zmult. -Notation Zmult_Zopp_Zopp := Zmult_Zopp_Zopp. -Notation weak_Zmult_plus_distr_r := weak_Zmult_plus_distr_r. -Notation Zmult_plus_distr_r := Zmult_plus_distr_r. -Notation Zcompare_EGAL := Zcompare_EGAL. -Notation Zcompare_ANTISYM := Zcompare_ANTISYM. -Notation le_minus := le_minus. -Notation Zcompare_Zopp := Zcompare_Zopp. -Notation weaken_Zcompare_Zplus_compatible := weaken_Zcompare_Zplus_compatible. -Notation weak_Zcompare_Zplus_compatible := weak_Zcompare_Zplus_compatible. -Notation Zcompare_Zplus_compatible := Zcompare_Zplus_compatible. -Notation Zcompare_trans_SUPERIEUR := Zcompare_trans_SUPERIEUR. -Notation SUPERIEUR_POS := SUPERIEUR_POS. -Export Datatypes. -Export BinPos. -Export BinNat. -Export BinInt. -Export Zcompare. -Export Mult. -]. diff --git a/theories/ZArith/zarith_aux.v b/theories/ZArith/zarith_aux.v deleted file mode 100644 index 61a712b92f..0000000000 --- a/theories/ZArith/zarith_aux.v +++ /dev/null @@ -1,151 +0,0 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(*