summaryrefslogtreecommitdiff
path: root/theories/QArith
diff options
context:
space:
mode:
Diffstat (limited to 'theories/QArith')
-rw-r--r--theories/QArith/QArith.v10
-rw-r--r--theories/QArith/QArith_base.v76
-rw-r--r--theories/QArith/QOrderedType.v10
-rw-r--r--theories/QArith/Qabs.v25
-rw-r--r--theories/QArith/Qcabs.v14
-rw-r--r--theories/QArith/Qcanon.v14
-rw-r--r--theories/QArith/Qfield.v10
-rw-r--r--theories/QArith/Qminmax.v10
-rw-r--r--theories/QArith/Qpower.v14
-rw-r--r--theories/QArith/Qreals.v44
-rw-r--r--theories/QArith/Qreduction.v36
-rw-r--r--theories/QArith/Qring.v10
-rw-r--r--theories/QArith/Qround.v24
-rw-r--r--theories/QArith/vo.itarget13
14 files changed, 170 insertions, 140 deletions
diff --git a/theories/QArith/QArith.v b/theories/QArith/QArith.v
index 5ad08b65..81390082 100644
--- a/theories/QArith/QArith.v
+++ b/theories/QArith/QArith.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Export QArith_base.
diff --git a/theories/QArith/QArith_base.v b/theories/QArith/QArith_base.v
index 62304876..35706e7f 100644
--- a/theories/QArith/QArith_base.v
+++ b/theories/QArith/QArith_base.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Export ZArith.
@@ -171,6 +173,26 @@ Proof.
auto with qarith.
Qed.
+Lemma Qeq_bool_comm x y: Qeq_bool x y = Qeq_bool y x.
+Proof.
+ apply eq_true_iff_eq. rewrite !Qeq_bool_iff. now symmetry.
+Qed.
+
+Lemma Qeq_bool_refl x: Qeq_bool x x = true.
+Proof.
+ rewrite Qeq_bool_iff. now reflexivity.
+Qed.
+
+Lemma Qeq_bool_sym x y: Qeq_bool x y = true -> Qeq_bool y x = true.
+Proof.
+ rewrite !Qeq_bool_iff. now symmetry.
+Qed.
+
+Lemma Qeq_bool_trans x y z: Qeq_bool x y = true -> Qeq_bool y z = true -> Qeq_bool x z = true.
+Proof.
+ rewrite !Qeq_bool_iff; apply Qeq_trans.
+Qed.
+
Hint Resolve Qnot_eq_sym : qarith.
(** * Addition, multiplication and opposite *)
@@ -205,9 +227,7 @@ Infix "/" := Qdiv : Q_scope.
(** A light notation for [Zpos] *)
-Notation " ' x " := (Zpos x) (at level 20, no associativity) : Z_scope.
-
-Lemma Qmake_Qdiv a b : a#b==inject_Z a/inject_Z ('b).
+Lemma Qmake_Qdiv a b : a#b==inject_Z a/inject_Z (Zpos b).
Proof.
unfold Qeq. simpl. ring.
Qed.
@@ -220,9 +240,9 @@ Proof.
Open Scope Z_scope.
intros (p1, p2) (q1, q2) H (r1, r2) (s1, s2) H0; simpl in *.
simpl_mult; ring_simplify.
- replace (p1 * 'r2 * 'q2) with (p1 * 'q2 * 'r2) by ring.
+ replace (p1 * Zpos r2 * Zpos q2) with (p1 * Zpos q2 * Zpos r2) by ring.
rewrite H.
- replace (r1 * 'p2 * 'q2 * 's2) with (r1 * 's2 * 'p2 * 'q2) by ring.
+ replace (r1 * Zpos p2 * Zpos q2 * Zpos s2) with (r1 * Zpos s2 * Zpos p2 * Zpos q2) by ring.
rewrite H0.
ring.
Close Scope Z_scope.
@@ -233,7 +253,7 @@ Proof.
unfold Qeq, Qopp; simpl.
Open Scope Z_scope.
intros x y H; simpl.
- replace (- Qnum x * ' Qden y) with (- (Qnum x * ' Qden y)) by ring.
+ replace (- Qnum x * Zpos (Qden y)) with (- (Qnum x * Zpos (Qden y))) by ring.
rewrite H; ring.
Close Scope Z_scope.
Qed.
@@ -250,9 +270,9 @@ Proof.
Open Scope Z_scope.
intros (p1, p2) (q1, q2) H (r1, r2) (s1, s2) H0; simpl in *.
intros; simpl_mult; ring_simplify.
- replace (q1 * s1 * 'p2) with (q1 * 'p2 * s1) by ring.
+ replace (q1 * s1 * Zpos p2) with (q1 * Zpos p2 * s1) by ring.
rewrite <- H.
- replace (p1 * r1 * 'q2 * 's2) with (r1 * 's2 * p1 * 'q2) by ring.
+ replace (p1 * r1 * Zpos q2 * Zpos s2) with (r1 * Zpos s2 * p1 * Zpos q2) by ring.
rewrite H0.
ring.
Close Scope Z_scope.
@@ -283,13 +303,13 @@ Proof.
unfold Qeq, Qcompare.
Open Scope Z_scope.
intros (p1,p2) (q1,q2) H (r1,r2) (s1,s2) H'; simpl in *.
- rewrite <- (Zcompare_mult_compat (q2*s2) (p1*'r2)).
- rewrite <- (Zcompare_mult_compat (p2*r2) (q1*'s2)).
- change ('(q2*s2)) with ('q2 * 's2).
- change ('(p2*r2)) with ('p2 * 'r2).
- replace ('q2 * 's2 * (p1*'r2)) with ((p1*'q2)*'r2*'s2) by ring.
+ rewrite <- (Zcompare_mult_compat (q2*s2) (p1*Zpos r2)).
+ rewrite <- (Zcompare_mult_compat (p2*r2) (q1*Zpos s2)).
+ change (Zpos (q2*s2)) with (Zpos q2 * Zpos s2).
+ change (Zpos (p2*r2)) with (Zpos p2 * Zpos r2).
+ replace (Zpos q2 * Zpos s2 * (p1*Zpos r2)) with ((p1*Zpos q2)*Zpos r2*Zpos s2) by ring.
rewrite H.
- replace ('q2 * 's2 * (r1*'p2)) with ((r1*'s2)*'q2*'p2) by ring.
+ replace (Zpos q2 * Zpos s2 * (r1*Zpos p2)) with ((r1*Zpos s2)*Zpos q2*Zpos p2) by ring.
rewrite H'.
f_equal; ring.
Close Scope Z_scope.
@@ -550,8 +570,8 @@ Lemma Qle_trans : forall x y z, x<=y -> y<=z -> x<=z.
Proof.
unfold Qle; intros (x1, x2) (y1, y2) (z1, z2); simpl; intros.
Open Scope Z_scope.
- apply Z.mul_le_mono_pos_r with ('y2); [easy|].
- apply Z.le_trans with (y1 * 'x2 * 'z2).
+ apply Z.mul_le_mono_pos_r with (Zpos y2); [easy|].
+ apply Z.le_trans with (y1 * Zpos x2 * Zpos z2).
- rewrite Z.mul_shuffle0. now apply Z.mul_le_mono_pos_r.
- rewrite Z.mul_shuffle0, (Z.mul_shuffle0 z1).
now apply Z.mul_le_mono_pos_r.
@@ -598,8 +618,8 @@ Lemma Qle_lt_trans : forall x y z, x<=y -> y<z -> x<z.
Proof.
unfold Qle, Qlt; intros (x1, x2) (y1, y2) (z1, z2); simpl; intros.
Open Scope Z_scope.
- apply Z.mul_lt_mono_pos_r with ('y2); [easy|].
- apply Z.le_lt_trans with (y1 * 'x2 * 'z2).
+ apply Z.mul_lt_mono_pos_r with (Zpos y2); [easy|].
+ apply Z.le_lt_trans with (y1 * Zpos x2 * Zpos z2).
- rewrite Z.mul_shuffle0. now apply Z.mul_le_mono_pos_r.
- rewrite Z.mul_shuffle0, (Z.mul_shuffle0 z1).
now apply Z.mul_lt_mono_pos_r.
@@ -610,8 +630,8 @@ Lemma Qlt_le_trans : forall x y z, x<y -> y<=z -> x<z.
Proof.
unfold Qle, Qlt; intros (x1, x2) (y1, y2) (z1, z2); simpl; intros.
Open Scope Z_scope.
- apply Z.mul_lt_mono_pos_r with ('y2); [easy|].
- apply Z.lt_le_trans with (y1 * 'x2 * 'z2).
+ apply Z.mul_lt_mono_pos_r with (Zpos y2); [easy|].
+ apply Z.lt_le_trans with (y1 * Zpos x2 * Zpos z2).
- rewrite Z.mul_shuffle0. now apply Z.mul_lt_mono_pos_r.
- rewrite Z.mul_shuffle0, (Z.mul_shuffle0 z1).
now apply Z.mul_le_mono_pos_r.
@@ -701,9 +721,9 @@ Proof.
match goal with |- ?a <= ?b => ring_simplify a b end.
rewrite Z.add_comm.
apply Z.add_le_mono.
- match goal with |- ?a <= ?b => ring_simplify z1 t1 ('z2) ('t2) a b end.
+ match goal with |- ?a <= ?b => ring_simplify z1 t1 (Zpos z2) (Zpos t2) a b end.
auto with zarith.
- match goal with |- ?a <= ?b => ring_simplify x1 y1 ('x2) ('y2) a b end.
+ match goal with |- ?a <= ?b => ring_simplify x1 y1 (Zpos x2) (Zpos y2) a b end.
auto with zarith.
Close Scope Z_scope.
Qed.
@@ -718,9 +738,9 @@ Proof.
match goal with |- ?a < ?b => ring_simplify a b end.
rewrite Z.add_comm.
apply Z.add_le_lt_mono.
- match goal with |- ?a <= ?b => ring_simplify z1 t1 ('z2) ('t2) a b end.
+ match goal with |- ?a <= ?b => ring_simplify z1 t1 (Zpos z2) (Zpos t2) a b end.
auto with zarith.
- match goal with |- ?a < ?b => ring_simplify x1 y1 ('x2) ('y2) a b end.
+ match goal with |- ?a < ?b => ring_simplify x1 y1 (Zpos x2) (Zpos y2) a b end.
do 2 (apply Z.mul_lt_mono_pos_r;try easy).
Close Scope Z_scope.
Qed.
diff --git a/theories/QArith/QOrderedType.v b/theories/QArith/QOrderedType.v
index 25e98f0b..37b4b298 100644
--- a/theories/QArith/QOrderedType.v
+++ b/theories/QArith/QOrderedType.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Import QArith_base Equalities Orders OrdersTac.
diff --git a/theories/QArith/Qabs.v b/theories/QArith/Qabs.v
index c60d0451..31eb41bc 100644
--- a/theories/QArith/Qabs.v
+++ b/theories/QArith/Qabs.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Export QArith.
@@ -26,8 +28,8 @@ intros [xn xd] [yn yd] H.
simpl.
unfold Qeq in *.
simpl in *.
-change (' yd)%Z with (Z.abs (' yd)).
-change (' xd)%Z with (Z.abs (' xd)).
+change (Zpos yd)%Z with (Z.abs (Zpos yd)).
+change (Zpos xd)%Z with (Z.abs (Zpos xd)).
repeat rewrite <- Z.abs_mul.
congruence.
Qed.
@@ -86,8 +88,8 @@ unfold Qplus.
unfold Qle.
simpl.
apply Z.mul_le_mono_nonneg_r;auto with *.
-change (' yd)%Z with (Z.abs (' yd)).
-change (' xd)%Z with (Z.abs (' xd)).
+change (Zpos yd)%Z with (Z.abs (Zpos yd)).
+change (Zpos xd)%Z with (Z.abs (Zpos xd)).
repeat rewrite <- Z.abs_mul.
apply Z.abs_triangle.
Qed.
@@ -100,6 +102,13 @@ rewrite Z.abs_mul.
reflexivity.
Qed.
+Lemma Qabs_Qinv : forall q, Qabs (/ q) == / (Qabs q).
+Proof.
+ intros [n d]; simpl.
+ unfold Qinv.
+ case_eq n; intros; simpl in *; apply Qeq_refl.
+Qed.
+
Lemma Qabs_Qminus x y: Qabs (x - y) = Qabs (y - x).
Proof.
unfold Qminus, Qopp. simpl.
diff --git a/theories/QArith/Qcabs.v b/theories/QArith/Qcabs.v
index c0ababff..f45868a7 100644
--- a/theories/QArith/Qcabs.v
+++ b/theories/QArith/Qcabs.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** * An absolute value for normalized rational numbers. *)
@@ -22,7 +24,7 @@ Lemma Qcabs_canon (x : Q) : Qred x = x -> Qred (Qabs x) = Qabs x.
Proof. intros H; now rewrite (Qred_abs x), H. Qed.
Definition Qcabs (x:Qc) : Qc := {| canon := Qcabs_canon x (canon x) |}.
-Notation "[ q ]" := (Qcabs q) (q at next level, format "[ q ]") : Qc_scope.
+Notation "[ q ]" := (Qcabs q) : Qc_scope.
Ltac Qc_unfolds :=
unfold Qcabs, Qcminus, Qcopp, Qcplus, Qcmult, Qcle, Q2Qc, this.
@@ -126,4 +128,4 @@ Proof.
destruct (proj1 (Qcabs_Qcle_condition x 0)) as [A B].
+ rewrite H; apply Qcle_refl.
+ apply Qcle_antisym; auto.
-Qed. \ No newline at end of file
+Qed.
diff --git a/theories/QArith/Qcanon.v b/theories/QArith/Qcanon.v
index 9f9651d8..1510a7b8 100644
--- a/theories/QArith/Qcanon.v
+++ b/theories/QArith/Qcanon.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Import Field.
@@ -28,7 +30,7 @@ Lemma Qred_identity :
Proof.
intros (a,b) H; simpl in *.
rewrite <- Z.ggcd_gcd in H.
- generalize (Z.ggcd_correct_divisors a ('b)).
+ generalize (Z.ggcd_correct_divisors a (Zpos b)).
destruct Z.ggcd as (g,(aa,bb)); simpl in *; subst.
rewrite !Z.mul_1_l. now intros (<-,<-).
Qed.
@@ -37,7 +39,7 @@ Lemma Qred_identity2 :
forall q:Q, Qred q = q -> Z.gcd (Qnum q) (QDen q) = 1%Z.
Proof.
intros (a,b) H; simpl in *.
- generalize (Z.gcd_nonneg a ('b)) (Z.ggcd_correct_divisors a ('b)).
+ generalize (Z.gcd_nonneg a (Zpos b)) (Z.ggcd_correct_divisors a (Zpos b)).
rewrite <- Z.ggcd_gcd.
destruct Z.ggcd as (g,(aa,bb)); simpl in *.
injection H as <- <-. intros H (_,H').
diff --git a/theories/QArith/Qfield.v b/theories/QArith/Qfield.v
index bbaf6027..6cbb491b 100644
--- a/theories/QArith/Qfield.v
+++ b/theories/QArith/Qfield.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Export Field.
diff --git a/theories/QArith/Qminmax.v b/theories/QArith/Qminmax.v
index 86584d9e..264b2f92 100644
--- a/theories/QArith/Qminmax.v
+++ b/theories/QArith/Qminmax.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Import QArith_base Orders QOrderedType GenericMinMax.
diff --git a/theories/QArith/Qpower.v b/theories/QArith/Qpower.v
index af89d300..01078220 100644
--- a/theories/QArith/Qpower.v
+++ b/theories/QArith/Qpower.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Import Zpow_facts Qfield Qreduction.
@@ -88,7 +90,7 @@ rewrite Qinv_power.
reflexivity.
Qed.
-Lemma Qinv_power_n : forall n p, (1#p)^n == /(inject_Z ('p))^n.
+Lemma Qinv_power_n : forall n p, (1#p)^n == /(inject_Z (Zpos p))^n.
Proof.
intros n p.
rewrite Qmake_Qdiv.
@@ -188,7 +190,7 @@ unfold Z.succ.
rewrite Zpower_exp; auto with *; try discriminate.
rewrite Qpower_plus' by discriminate.
rewrite <- IHn by discriminate.
-replace (a^'n*a^1)%Z with (a^'n*a)%Z by ring.
+replace (a^Zpos n*a^1)%Z with (a^Zpos n*a)%Z by ring.
ring_simplify.
reflexivity.
Qed.
diff --git a/theories/QArith/Qreals.v b/theories/QArith/Qreals.v
index 048e409c..c8329625 100644
--- a/theories/QArith/Qreals.v
+++ b/theories/QArith/Qreals.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Export Rbase.
@@ -15,7 +17,8 @@ Definition Q2R (x : Q) : R := (IZR (Qnum x) * / IZR (QDen x))%R.
Lemma IZR_nz : forall p : positive, IZR (Zpos p) <> 0%R.
Proof.
-intros; apply not_O_IZR; auto with qarith.
+intros.
+now apply not_O_IZR.
Qed.
Hint Resolve IZR_nz Rmult_integral_contrapositive.
@@ -48,8 +51,7 @@ assert ((X1 * Y2)%R = (Y1 * X2)%R).
apply IZR_eq; auto.
clear H.
field_simplify_eq; auto.
-ring_simplify X1 Y2 (Y2 * X1)%R.
-rewrite H0; ring.
+rewrite H0; ring.
Qed.
Lemma Rle_Qle : forall x y : Q, (Q2R x <= Q2R y)%R -> x<=y.
@@ -66,10 +68,8 @@ replace (X1 * Y2)%R with (X1 * / X2 * (X2 * Y2))%R; try (field; auto).
replace (Y1 * X2)%R with (Y1 * / Y2 * (X2 * Y2))%R; try (field; auto).
apply Rmult_le_compat_r; auto.
apply Rmult_le_pos.
-unfold X2; replace 0%R with (IZR 0); auto; apply IZR_le;
- auto with zarith.
-unfold Y2; replace 0%R with (IZR 0); auto; apply IZR_le;
- auto with zarith.
+now apply IZR_le.
+now apply IZR_le.
Qed.
Lemma Qle_Rle : forall x y : Q, x<=y -> (Q2R x <= Q2R y)%R.
@@ -88,10 +88,8 @@ replace (X1 * / X2)%R with (X1 * Y2 * (/ X2 * / Y2))%R; try (field; auto).
replace (Y1 * / Y2)%R with (Y1 * X2 * (/ X2 * / Y2))%R; try (field; auto).
apply Rmult_le_compat_r; auto.
apply Rmult_le_pos; apply Rlt_le; apply Rinv_0_lt_compat.
-unfold X2; replace 0%R with (IZR 0); auto; apply IZR_lt; red;
- auto with zarith.
-unfold Y2; replace 0%R with (IZR 0); auto; apply IZR_lt; red;
- auto with zarith.
+now apply IZR_lt.
+now apply IZR_lt.
Qed.
Lemma Rlt_Qlt : forall x y : Q, (Q2R x < Q2R y)%R -> x<y.
@@ -108,10 +106,8 @@ replace (X1 * Y2)%R with (X1 * / X2 * (X2 * Y2))%R; try (field; auto).
replace (Y1 * X2)%R with (Y1 * / Y2 * (X2 * Y2))%R; try (field; auto).
apply Rmult_lt_compat_r; auto.
apply Rmult_lt_0_compat.
-unfold X2; replace 0%R with (IZR 0); auto; apply IZR_lt; red;
- auto with zarith.
-unfold Y2; replace 0%R with (IZR 0); auto; apply IZR_lt; red;
- auto with zarith.
+now apply IZR_lt.
+now apply IZR_lt.
Qed.
Lemma Qlt_Rlt : forall x y : Q, x<y -> (Q2R x < Q2R y)%R.
@@ -130,10 +126,8 @@ replace (X1 * / X2)%R with (X1 * Y2 * (/ X2 * / Y2))%R; try (field; auto).
replace (Y1 * / Y2)%R with (Y1 * X2 * (/ X2 * / Y2))%R; try (field; auto).
apply Rmult_lt_compat_r; auto.
apply Rmult_lt_0_compat; apply Rinv_0_lt_compat.
-unfold X2; replace 0%R with (IZR 0); auto; apply IZR_lt; red;
- auto with zarith.
-unfold Y2; replace 0%R with (IZR 0); auto; apply IZR_lt; red;
- auto with zarith.
+now apply IZR_lt.
+now apply IZR_lt.
Qed.
Lemma Q2R_plus : forall x y : Q, Q2R (x+y) = (Q2R x + Q2R y)%R.
@@ -173,8 +167,8 @@ unfold Qinv, Q2R, Qeq; intros (x1, x2). case x1; unfold Qnum, Qden.
simpl; intros; elim H; trivial.
intros; field; auto.
intros;
- change (IZR (Zneg x2)) with (- IZR (' x2))%R;
- change (IZR (Zneg p)) with (- IZR (' p))%R;
+ change (IZR (Zneg x2)) with (- IZR (Zpos x2))%R;
+ change (IZR (Zneg p)) with (- IZR (Zpos p))%R;
simpl; field; (*auto 8 with real.*)
repeat split; auto; auto with real.
Qed.
diff --git a/theories/QArith/Qreduction.v b/theories/QArith/Qreduction.v
index 131214f5..17307c82 100644
--- a/theories/QArith/Qreduction.v
+++ b/theories/QArith/Qreduction.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(** Normalisation functions for rational numbers. *)
@@ -11,22 +13,22 @@
Require Export QArith_base.
Require Import Znumtheory.
-Notation Z2P := Z.to_pos (compat "8.3").
-Notation Z2P_correct := Z2Pos.id (compat "8.3").
+Notation Z2P := Z.to_pos (only parsing).
+Notation Z2P_correct := Z2Pos.id (only parsing).
(** Simplification of fractions using [Z.gcd].
This version can compute within Coq. *)
Definition Qred (q:Q) :=
let (q1,q2) := q in
- let (r1,r2) := snd (Z.ggcd q1 ('q2))
+ let (r1,r2) := snd (Z.ggcd q1 (Zpos q2))
in r1#(Z.to_pos r2).
Lemma Qred_correct : forall q, (Qred q) == q.
Proof.
unfold Qred, Qeq; intros (n,d); simpl.
- generalize (Z.ggcd_gcd n ('d)) (Z.gcd_nonneg n ('d))
- (Z.ggcd_correct_divisors n ('d)).
+ generalize (Z.ggcd_gcd n (Zpos d)) (Z.gcd_nonneg n (Zpos d))
+ (Z.ggcd_correct_divisors n (Zpos d)).
destruct (Z.ggcd n (Zpos d)) as (g,(nn,dd)); simpl.
Open Scope Z_scope.
intros Hg LE (Hn,Hd). rewrite Hd, Hn.
@@ -43,13 +45,13 @@ Proof.
unfold Qred, Qeq in *; simpl in *.
Open Scope Z_scope.
intros H.
- generalize (Z.ggcd_gcd a ('b)) (Zgcd_is_gcd a ('b))
- (Z.gcd_nonneg a ('b)) (Z.ggcd_correct_divisors a ('b)).
+ generalize (Z.ggcd_gcd a (Zpos b)) (Zgcd_is_gcd a (Zpos b))
+ (Z.gcd_nonneg a (Zpos b)) (Z.ggcd_correct_divisors a (Zpos b)).
destruct (Z.ggcd a (Zpos b)) as (g,(aa,bb)).
simpl. intros <- Hg1 Hg2 (Hg3,Hg4).
assert (Hg0 : g <> 0) by (intro; now subst g).
- generalize (Z.ggcd_gcd c ('d)) (Zgcd_is_gcd c ('d))
- (Z.gcd_nonneg c ('d)) (Z.ggcd_correct_divisors c ('d)).
+ generalize (Z.ggcd_gcd c (Zpos d)) (Zgcd_is_gcd c (Zpos d))
+ (Z.gcd_nonneg c (Zpos d)) (Z.ggcd_correct_divisors c (Zpos d)).
destruct (Z.ggcd c (Zpos d)) as (g',(cc,dd)).
simpl. intros <- Hg'1 Hg'2 (Hg'3,Hg'4).
assert (Hg'0 : g' <> 0) by (intro; now subst g').
@@ -101,7 +103,7 @@ Proof.
- apply Qred_complete.
Qed.
-Add Morphism Qred : Qred_comp.
+Add Morphism Qred with signature (Qeq ==> Qeq) as Qred_comp.
Proof.
intros. now rewrite !Qred_correct.
Qed.
@@ -125,19 +127,19 @@ Proof.
intros; unfold Qminus'; apply Qred_correct; auto.
Qed.
-Add Morphism Qplus' : Qplus'_comp.
+Add Morphism Qplus' with signature (Qeq ==> Qeq ==> Qeq) as Qplus'_comp.
Proof.
intros; unfold Qplus'.
rewrite H, H0; auto with qarith.
Qed.
-Add Morphism Qmult' : Qmult'_comp.
+Add Morphism Qmult' with signature (Qeq ==> Qeq ==> Qeq) as Qmult'_comp.
Proof.
intros; unfold Qmult'.
rewrite H, H0; auto with qarith.
Qed.
-Add Morphism Qminus' : Qminus'_comp.
+Add Morphism Qminus' with signature (Qeq ==> Qeq ==> Qeq) as Qminus'_comp.
Proof.
intros; unfold Qminus'.
rewrite H, H0; auto with qarith.
diff --git a/theories/QArith/Qring.v b/theories/QArith/Qring.v
index da11c2b1..7f972d56 100644
--- a/theories/QArith/Qring.v
+++ b/theories/QArith/Qring.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Export Qfield.
diff --git a/theories/QArith/Qround.v b/theories/QArith/Qround.v
index 0ed6d557..7c5ddbb6 100644
--- a/theories/QArith/Qround.v
+++ b/theories/QArith/Qround.v
@@ -1,9 +1,11 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Require Import QArith.
@@ -78,11 +80,11 @@ unfold Qlt.
simpl.
replace (n*1)%Z with n by ring.
ring_simplify.
-replace (n / ' d * ' d + ' d)%Z with
- (('d * (n / 'd) + n mod 'd) + 'd - n mod 'd)%Z by ring.
+replace (n / Zpos d * Zpos d + Zpos d)%Z with
+ ((Zpos d * (n / Zpos d) + n mod Zpos d) + Zpos d - n mod Zpos d)%Z by ring.
rewrite <- Z_div_mod_eq; auto with*.
rewrite <- Z.lt_add_lt_sub_r.
-destruct (Z_mod_lt n ('d)); auto with *.
+destruct (Z_mod_lt n (Zpos d)); auto with *.
Qed.
Hint Resolve Qlt_floor : qarith.
@@ -105,9 +107,9 @@ Proof.
intros [xn xd] [yn yd] Hxy.
unfold Qle in *.
simpl in *.
-rewrite <- (Zdiv_mult_cancel_r xn ('xd) ('yd)); auto with *.
-rewrite <- (Zdiv_mult_cancel_r yn ('yd) ('xd)); auto with *.
-rewrite (Z.mul_comm ('yd) ('xd)).
+rewrite <- (Zdiv_mult_cancel_r xn (Zpos xd) (Zpos yd)); auto with *.
+rewrite <- (Zdiv_mult_cancel_r yn (Zpos yd) (Zpos xd)); auto with *.
+rewrite (Z.mul_comm (Zpos yd) (Zpos xd)).
apply Z_div_le; auto with *.
Qed.
@@ -141,7 +143,7 @@ Qed.
Lemma Zdiv_Qdiv (n m: Z): (n / m)%Z = Qfloor (n / m).
Proof.
unfold Qfloor. intros. simpl.
- destruct m as [?|?|p]; simpl.
+ destruct m as [ | | p]; simpl.
now rewrite Zdiv_0_r, Z.mul_0_r.
now rewrite Z.mul_1_r.
rewrite <- Z.opp_eq_mul_m1.
diff --git a/theories/QArith/vo.itarget b/theories/QArith/vo.itarget
deleted file mode 100644
index b550b471..00000000
--- a/theories/QArith/vo.itarget
+++ /dev/null
@@ -1,13 +0,0 @@
-Qabs.vo
-QArith_base.vo
-QArith.vo
-Qcanon.vo
-Qcabs.vo
-Qfield.vo
-Qpower.vo
-Qreals.vo
-Qreduction.vo
-Qring.vo
-Qround.vo
-QOrderedType.vo
-Qminmax.vo \ No newline at end of file