aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2012-07-05 16:56:16 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2012-07-05 16:56:16 +0000
commitfc2613e871dffffa788d90044a81598f671d0a3b (patch)
treef6f308b3d6b02e1235446b2eb4a2d04b135a0462
parentf93f073df630bb46ddd07802026c0326dc72dafd (diff)
ZArith + other : favor the use of modern names instead of compat notations
- For instance, refl_equal --> eq_refl - Npos, Zpos, Zneg now admit more uniform qualified aliases N.pos, Z.pos, Z.neg. - A new module BinInt.Pos2Z with results about injections from positive to Z - A result about Z.pow pushed in the generic layer - Zmult_le_compat_{r,l} --> Z.mul_le_mono_nonneg_{r,l} - Using tactic Z.le_elim instead of Zle_lt_or_eq - Some cleanup in ring, field, micromega (use of "Equivalence", "Proper" ...) - Some adaptions in QArith (for instance changed Qpower.Qpower_decomp) - In ZMake and ZMake, functor parameters are now named NN and ZZ instead of N and Z for avoiding confusions git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@15515 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--doc/RecTutorial/RecTutorial.tex28
-rw-r--r--doc/RecTutorial/RecTutorial.v14
-rw-r--r--doc/faq/FAQ.tex16
-rw-r--r--doc/faq/interval_discr.v14
-rw-r--r--doc/refman/Classes.tex2
-rw-r--r--doc/refman/Natural.tex2
-rw-r--r--doc/refman/Omega.tex2
-rw-r--r--doc/refman/Polynom.tex2
-rw-r--r--doc/refman/Program.tex2
-rw-r--r--doc/refman/RefMan-coi.tex2
-rw-r--r--doc/refman/RefMan-gal.tex14
-rw-r--r--doc/refman/RefMan-lib.tex40
-rw-r--r--doc/refman/RefMan-ltac.tex6
-rw-r--r--doc/refman/RefMan-oth.tex2
-rw-r--r--doc/refman/RefMan-tacex.tex6
-rw-r--r--doc/refman/Setoid.tex6
-rw-r--r--interp/constrextern.ml2
-rw-r--r--plugins/cc/cctac.ml6
-rw-r--r--plugins/extraction/ExtrOcamlZBigInt.v6
-rw-r--r--plugins/extraction/ExtrOcamlZInt.v2
-rw-r--r--plugins/micromega/Env.v151
-rw-r--r--plugins/micromega/EnvRing.v1179
-rw-r--r--plugins/micromega/MExtraction.v2
-rw-r--r--plugins/micromega/Psatz.v6
-rw-r--r--plugins/micromega/QMicromega.v8
-rw-r--r--plugins/micromega/RMicromega.v28
-rw-r--r--plugins/micromega/RingMicromega.v48
-rw-r--r--plugins/micromega/ZCoeff.v14
-rw-r--r--plugins/micromega/ZMicromega.v214
-rw-r--r--plugins/micromega/coq_micromega.ml21
-rw-r--r--plugins/nsatz/Nsatz.v34
-rw-r--r--plugins/omega/Omega.v6
-rw-r--r--plugins/omega/OmegaLemmas.v266
-rw-r--r--plugins/omega/PreOmega.v351
-rw-r--r--plugins/omega/coq_omega.ml57
-rw-r--r--plugins/ring/LegacyNArithRing.v23
-rw-r--r--plugins/ring/LegacyZArithRing.v6
-rw-r--r--plugins/ring/Ring_abstract.v6
-rw-r--r--plugins/ring/Ring_normalize.v6
-rw-r--r--plugins/ring/Setoid_ring_normalize.v6
-rw-r--r--plugins/ring/ring.ml6
-rw-r--r--plugins/romega/ReflOmegaCore.v93
-rw-r--r--plugins/rtauto/Bintree.v14
-rw-r--r--plugins/rtauto/refl_tauto.ml2
-rw-r--r--plugins/setoid_ring/ArithRing.v8
-rw-r--r--plugins/setoid_ring/BinList.v75
-rw-r--r--plugins/setoid_ring/Cring.v25
-rw-r--r--plugins/setoid_ring/Field_tac.v4
-rw-r--r--plugins/setoid_ring/Field_theory.v183
-rw-r--r--plugins/setoid_ring/InitialRing.v34
-rw-r--r--plugins/setoid_ring/Integral_domain.v5
-rw-r--r--plugins/setoid_ring/Ncring.v33
-rw-r--r--plugins/setoid_ring/Ncring_initial.v54
-rw-r--r--plugins/setoid_ring/Ncring_polynom.v109
-rw-r--r--plugins/setoid_ring/Ncring_tac.v8
-rw-r--r--plugins/setoid_ring/RealField.v12
-rw-r--r--plugins/setoid_ring/Ring_polynom.v556
-rw-r--r--plugins/setoid_ring/Ring_tac.v7
-rw-r--r--plugins/setoid_ring/Ring_theory.v291
-rw-r--r--plugins/setoid_ring/Rings_Z.v2
-rw-r--r--plugins/setoid_ring/ZArithRing.v4
-rw-r--r--tactics/tactics.ml2
-rw-r--r--test-suite/bugs/closed/shouldsucceed/1414.v4
-rw-r--r--test-suite/bugs/closed/shouldsucceed/1784.v2
-rw-r--r--test-suite/bugs/closed/shouldsucceed/1844.v2
-rw-r--r--test-suite/bugs/closed/shouldsucceed/1935.v2
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2127.v4
-rw-r--r--test-suite/complexity/ring2.v6
-rw-r--r--test-suite/ideal-features/eapply_evar.v2
-rw-r--r--test-suite/micromega/example.v10
-rw-r--r--test-suite/micromega/square.v18
-rw-r--r--test-suite/success/Hints.v26
-rw-r--r--test-suite/success/MatchFail.v4
-rw-r--r--test-suite/success/OmegaPre.v16
-rw-r--r--test-suite/success/ProgramWf.v4
-rw-r--r--test-suite/success/ROmegaPre.v16
-rw-r--r--test-suite/success/RecTutorial.v10
-rw-r--r--test-suite/success/Scopes.v2
-rw-r--r--test-suite/success/apply.v6
-rw-r--r--test-suite/success/decl_mode.v10
-rw-r--r--test-suite/success/fix.v4
-rw-r--r--test-suite/success/searchabout.v2
-rw-r--r--test-suite/success/specialize.v28
-rw-r--r--test-suite/success/unicode_utf8.v2
-rw-r--r--theories/Arith/Compare_dec.v2
-rw-r--r--theories/Arith/Factorial.v2
-rw-r--r--theories/Bool/Bool.v4
-rw-r--r--theories/FSets/FSetBridge.v4
-rw-r--r--theories/FSets/FSetEqProperties.v8
-rw-r--r--theories/FSets/FSetFacts.v6
-rw-r--r--theories/FSets/FSetProperties.v2
-rw-r--r--theories/Init/Datatypes.v18
-rw-r--r--theories/Init/Logic_Type.v8
-rw-r--r--theories/Init/Notations.v2
-rw-r--r--theories/Init/Peano.v18
-rw-r--r--theories/Init/Specif.v26
-rw-r--r--theories/Lists/List.v29
-rw-r--r--theories/Lists/ListTactics.v2
-rw-r--r--theories/Lists/StreamMemo.v27
-rw-r--r--theories/Logic/ClassicalFacts.v14
-rw-r--r--theories/Logic/Classical_Prop.v2
-rw-r--r--theories/Logic/Diaconescu.v4
-rw-r--r--theories/Logic/EqdepFacts.v12
-rw-r--r--theories/Logic/Eqdep_dec.v24
-rw-r--r--theories/Logic/ProofIrrelevanceFacts.v2
-rw-r--r--theories/MSets/MSetEqProperties.v8
-rw-r--r--theories/MSets/MSetPositive.v4
-rw-r--r--theories/MSets/MSetProperties.v2
-rw-r--r--theories/MSets/MSetRBT.v2
-rw-r--r--theories/NArith/BinNat.v42
-rw-r--r--theories/NArith/BinNatDef.v88
-rw-r--r--theories/NArith/Ndec.v441
-rw-r--r--theories/NArith/Ndigits.v181
-rw-r--r--theories/NArith/Ndist.v54
-rw-r--r--theories/Numbers/BigNumPrelude.v161
-rw-r--r--theories/Numbers/Cyclic/Abstract/CyclicAxioms.v44
-rw-r--r--theories/Numbers/Cyclic/Abstract/NZCyclic.v12
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v48
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v113
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v28
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v270
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v78
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v184
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v86
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v391
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v24
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v2
-rw-r--r--theories/Numbers/Cyclic/Int31/Cyclic31.v607
-rw-r--r--theories/Numbers/Cyclic/Int31/Int31.v16
-rw-r--r--theories/Numbers/Cyclic/Int31/Ring31.v2
-rw-r--r--theories/Numbers/Cyclic/ZModulo/ZModulo.v189
-rw-r--r--theories/Numbers/Integer/Abstract/ZBits.v2
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivFloor.v2
-rw-r--r--theories/Numbers/Integer/Abstract/ZPow.v11
-rw-r--r--theories/Numbers/Integer/BigZ/BigZ.v8
-rw-r--r--theories/Numbers/Integer/BigZ/ZMake.v452
-rw-r--r--theories/Numbers/Integer/Binary/ZBinary.v2
-rw-r--r--theories/Numbers/Integer/NatPairs/ZNatPairs.v2
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v4
-rw-r--r--theories/Numbers/NatInt/NZAxioms.v2
-rw-r--r--theories/Numbers/NatInt/NZMulOrder.v4
-rw-r--r--theories/Numbers/Natural/Abstract/NBits.v2
-rw-r--r--theories/Numbers/Natural/BigN/BigN.v2
-rw-r--r--theories/Numbers/Natural/BigN/NMake.v362
-rw-r--r--theories/Numbers/Natural/BigN/Nbasic.v118
-rw-r--r--theories/Numbers/Natural/Binary/NBinary.v4
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v8
-rw-r--r--theories/Numbers/Rational/BigQ/BigQ.v8
-rw-r--r--theories/Numbers/Rational/BigQ/QMake.v483
-rw-r--r--theories/PArith/BinPos.v23
-rw-r--r--theories/PArith/Pnat.v19
-rw-r--r--theories/Program/Subset.v2
-rw-r--r--theories/QArith/QArith_base.v253
-rw-r--r--theories/QArith/Qabs.v24
-rw-r--r--theories/QArith/Qcanon.v18
-rw-r--r--theories/QArith/Qfield.v4
-rw-r--r--theories/QArith/Qpower.v86
-rw-r--r--theories/QArith/Qreduction.v164
-rw-r--r--theories/QArith/Qround.v22
-rw-r--r--theories/Reals/AltSeries.v2
-rw-r--r--theories/Reals/Exp_prop.v2
-rw-r--r--theories/Reals/RIneq.v76
-rw-r--r--theories/Reals/R_Ifp.v12
-rw-r--r--theories/Reals/Ranalysis2.v2
-rw-r--r--theories/Reals/Raxioms.v4
-rw-r--r--theories/Reals/Rbasic_fun.v4
-rw-r--r--theories/Reals/Rderiv.v14
-rw-r--r--theories/Reals/Rfunctions.v106
-rw-r--r--theories/Reals/RiemannInt_SF.v2
-rw-r--r--theories/Reals/Rlimit.v4
-rw-r--r--theories/Reals/Rpower.v6
-rw-r--r--theories/Reals/Rseries.v10
-rw-r--r--theories/Reals/Rsqrt_def.v2
-rw-r--r--theories/Reals/Rtopology.v4
-rw-r--r--theories/Reals/Rtrigo_alt.v4
-rw-r--r--theories/Reals/Rtrigo_calc.v2
-rw-r--r--theories/Reals/Rtrigo_def.v12
-rw-r--r--theories/Reals/Rtrigo_fun.v10
-rw-r--r--theories/Reals/Rtrigo_reg.v4
-rw-r--r--theories/Reals/SeqProp.v10
-rw-r--r--theories/Reals/Sqrt_reg.v2
-rw-r--r--theories/Relations/Relation_Operators.v6
-rw-r--r--theories/Sorting/PermutSetoid.v2
-rw-r--r--theories/Strings/Ascii.v12
-rw-r--r--theories/Structures/DecidableTypeEx.v6
-rw-r--r--theories/Structures/OrderedTypeEx.v160
-rw-r--r--theories/Unicode/Utf8.v2
-rw-r--r--theories/Wellfounded/Lexicographic_Product.v2
-rw-r--r--theories/ZArith/BinInt.v642
-rw-r--r--theories/ZArith/BinIntDef.v245
-rw-r--r--theories/ZArith/Int.v16
-rw-r--r--theories/ZArith/Wf_Z.v10
-rw-r--r--theories/ZArith/ZArith_dec.v83
-rw-r--r--theories/ZArith/Zabs.v8
-rw-r--r--theories/ZArith/Zbool.v2
-rw-r--r--theories/ZArith/Zcompare.v6
-rw-r--r--theories/ZArith/Zcomplements.v6
-rw-r--r--theories/ZArith/Zdigits.v30
-rw-r--r--theories/ZArith/Zdiv.v66
-rw-r--r--theories/ZArith/Zeven.v4
-rw-r--r--theories/ZArith/Zgcd_alt.v239
-rw-r--r--theories/ZArith/Zhints.v93
-rw-r--r--theories/ZArith/Zlogarithm.v60
-rw-r--r--theories/ZArith/Zmax.v117
-rw-r--r--theories/ZArith/Zmin.v88
-rw-r--r--theories/ZArith/Zmisc.v5
-rw-r--r--theories/ZArith/Znat.v44
-rw-r--r--theories/ZArith/Znumtheory.v280
-rw-r--r--theories/ZArith/Zorder.v12
-rw-r--r--theories/ZArith/Zpow_alt.v2
-rw-r--r--theories/ZArith/Zpow_def.v2
-rw-r--r--theories/ZArith/Zpow_facts.v10
-rw-r--r--theories/ZArith/Zpower.v12
-rw-r--r--theories/ZArith/Zquot.v349
-rw-r--r--theories/ZArith/Zsqrt_compat.v59
-rw-r--r--theories/ZArith/Zwf.v19
216 files changed, 5779 insertions, 6911 deletions
diff --git a/doc/RecTutorial/RecTutorial.tex b/doc/RecTutorial/RecTutorial.tex
index f2cb383e0..857ba84d7 100644
--- a/doc/RecTutorial/RecTutorial.tex
+++ b/doc/RecTutorial/RecTutorial.tex
@@ -560,11 +560,11 @@ as it can be infered from $a$.
\begin{alltt}
Print eq.
\it{} Inductive eq (A : Type) (x : A) : A \arrow{} Prop :=
- refl_equal : x = x
+ eq_refl : x = x
For eq: Argument A is implicit
-For refl_equal: Argument A is implicit
+For eq_refl: Argument A is implicit
For eq: Argument scopes are [type_scope _ _]
-For refl_equal: Argument scopes are [type_scope _]
+For eq_refl: Argument scopes are [type_scope _]
\end{alltt}
Notice also that the first parameter $A$ of \texttt{eq} has type
@@ -581,15 +581,15 @@ Proof.
reflexivity.
Qed.
-Lemma eq_proof_proof : refl_equal (2*6) = refl_equal (3*4).
+Lemma eq_proof_proof : eq_refl (2*6) = eq_refl (3*4).
Proof.
reflexivity.
Qed.
Print eq_proof_proof.
\it eq_proof_proof =
-refl_equal (refl_equal (3 * 4))
- : refl_equal (2 * 6) = refl_equal (3 * 4)
+eq_refl (eq_refl (3 * 4))
+ : eq_refl (2 * 6) = eq_refl (3 * 4)
\tt
Lemma eq_lt_le : ( 2 < 4) = (3 {\coqle} 4).
@@ -942,10 +942,10 @@ predecessor = fun n : nat {\funarrow}
\textbf{| O {\funarrow}}
exist (fun m : nat {\funarrow} 0 = 0 {\coqand} m = 0 {\coqor} 0 = S m) 0
(or_introl (0 = 1)
- (conj (refl_equal 0) (refl_equal 0)))
+ (conj (eq_refl 0) (eq_refl 0)))
\textbf{| S n0 {\funarrow}}
exist (fun m : nat {\funarrow} S n0 = 0 {\coqand} m = 0 {\coqor} S n0 = S m) n0
- (or_intror (S n0 = 0 {\coqand} n0 = 0) (refl_equal (S n0)))
+ (or_intror (S n0 = 0 {\coqand} n0 = 0) (eq_refl (S n0)))
\textbf{end} : {\prodsym} n : nat, \textbf{pred_spec n}
\end{alltt}
@@ -1084,7 +1084,7 @@ The following term is a proof of ``~$Q\;a\, \arrow{}\, Q\;b$~''.
\begin{alltt}
fun H : Q a {\funarrow}
match \(\pi\) in (_ = y) return Q y with
- refl_equal {\funarrow} H
+ eq_refl {\funarrow} H
end
\end{alltt}
Notice the header of the \texttt{match} construct.
@@ -1552,13 +1552,13 @@ node, a tree of height 1 and a tree of height 2:
\begin{alltt}
Definition isingle l := inode l (fun i {\funarrow} ileaf).
-Definition t1 := inode 0 (fun n {\funarrow} isingle (Z_of_nat n)).
+Definition t1 := inode 0 (fun n {\funarrow} isingle (Z.of_nat n)).
Definition t2 :=
inode 0
(fun n : nat {\funarrow}
- inode (Z_of_nat n)
- (fun p {\funarrow} isingle (Z_of_nat (n*p)))).
+ inode (Z.of_nat n)
+ (fun p {\funarrow} isingle (Z.of_nat (n*p)))).
\end{alltt}
@@ -1572,7 +1572,7 @@ appear:
Inductive itree_le : itree{\arrow} itree {\arrow} Prop :=
| le_leaf : {\prodsym} t, itree_le ileaf t
| le_node : {\prodsym} l l' s s',
- Zle l l' {\arrow}
+ Z.le l l' {\arrow}
({\prodsym} i, {\exsym} j:nat, itree_le (s i) (s' j)){\arrow}
itree_le (inode l s) (inode l' s').
@@ -1597,7 +1597,7 @@ the type of \texttt{itree\_le}, does not present this problem:
Inductive itree_le' : itree{\arrow} itree {\arrow} Prop :=
| le_leaf' : {\prodsym} t, itree_le' ileaf t
| le_node' : {\prodsym} l l' s s' g,
- Zle l l' {\arrow}
+ Z.le l l' {\arrow}
({\prodsym} i, itree_le' (s i) (s' (g i))) {\arrow}
itree_le' (inode l s) (inode l' s').
diff --git a/doc/RecTutorial/RecTutorial.v b/doc/RecTutorial/RecTutorial.v
index 28aaf7520..8cfeebc28 100644
--- a/doc/RecTutorial/RecTutorial.v
+++ b/doc/RecTutorial/RecTutorial.v
@@ -83,7 +83,7 @@ Proof.
Qed.
Print eq_3_3.
-Lemma eq_proof_proof : refl_equal (2*6) = refl_equal (3*4).
+Lemma eq_proof_proof : eq_refl (2*6) = eq_refl (3*4).
Proof.
reflexivity.
Qed.
@@ -241,7 +241,7 @@ Section equality_elimination.
(Q : A -> Type).
Check (fun H : Q a =>
match p in (eq _ y) return Q y with
- refl_equal => H
+ eq_refl => H
end).
End equality_elimination.
@@ -377,18 +377,18 @@ Inductive itree : Set :=
Definition isingle l := inode l (fun i => ileaf).
-Definition t1 := inode 0 (fun n => isingle (Z_of_nat (2*n))).
+Definition t1 := inode 0 (fun n => isingle (Z.of_nat (2*n))).
Definition t2 := inode 0
(fun n : nat =>
- inode (Z_of_nat n)
- (fun p => isingle (Z_of_nat (n*p)))).
+ inode (Z.of_nat n)
+ (fun p => isingle (Z.of_nat (n*p)))).
Inductive itree_le : itree-> itree -> Prop :=
| le_leaf : forall t, itree_le ileaf t
| le_node : forall l l' s s',
- Zle l l' ->
+ Z.le l l' ->
(forall i, exists j:nat, itree_le (s i) (s' j)) ->
itree_le (inode l s) (inode l' s').
@@ -423,7 +423,7 @@ Qed.
Inductive itree_le' : itree-> itree -> Prop :=
| le_leaf' : forall t, itree_le' ileaf t
| le_node' : forall l l' s s' g,
- Zle l l' ->
+ Z.le l l' ->
(forall i, itree_le' (s i) (s' (g i))) ->
itree_le' (inode l s) (inode l' s').
diff --git a/doc/faq/FAQ.tex b/doc/faq/FAQ.tex
index b63f3ee26..5ce5e0436 100644
--- a/doc/faq/FAQ.tex
+++ b/doc/faq/FAQ.tex
@@ -545,7 +545,7 @@ dependent elimination of reflexive equality proofs.
\begin{coq_example*}
Axiom Streicher_K :
forall (A:Type) (x:A) (P: x=x -> Prop),
- P (refl_equal x) -> forall p: x=x, P p.
+ P (eq_refl x) -> forall p: x=x, P p.
\end{coq_example*}
In the general case, axiom $K$ is an independent statement of the
@@ -563,7 +563,7 @@ Axiom UIP : forall (A:Set) (x y:A) (p1 p2: x=y), p1 = p2.
Axiom $K$ is also equivalent to {\em Uniqueness of Reflexive Identity Proofs} \cite{HofStr98}
\begin{coq_example*}
-Axiom UIP_refl : forall (A:Set) (x:A) (p: x=x), p = refl_equal x.
+Axiom UIP_refl : forall (A:Set) (x:A) (p: x=x), p = eq_refl x.
\end{coq_example*}
Axiom $K$ is also equivalent to
@@ -2108,7 +2108,7 @@ Yes, because equality is decidable on {\tt nat}. Here is the proof.
Require Import Eqdep_dec.
Require Import Peano_dec.
Theorem K_nat :
- forall (x:nat) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p.
+ forall (x:nat) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p.
Proof.
intros; apply K_dec_set with (p := p).
apply eq_nat_dec.
@@ -2139,16 +2139,16 @@ Theorem le_uniqueness_proof : forall (n m : nat) (p q : n <= m), p = q.
Proof.
induction p using le_ind'; intro q.
replace (le_n n) with
- (eq_rect _ (fun n0 => n <= n0) (le_n n) _ (refl_equal n)).
+ (eq_rect _ (fun n0 => n <= n0) (le_n n) _ eq_refl).
2:reflexivity.
- generalize (refl_equal n).
+ generalize (eq_refl n).
pattern n at 2 4 6 10, q; case q; [intro | intros m l e].
rewrite <- eq_rect_eq_nat; trivial.
contradiction (le_Sn_n m); rewrite <- e; assumption.
replace (le_S n m p) with
- (eq_rect _ (fun n0 => n <= n0) (le_S n m p) _ (refl_equal (S m))).
+ (eq_rect _ (fun n0 => n <= n0) (le_S n m p) _ eq_refl).
2:reflexivity.
- generalize (refl_equal (S m)).
+ generalize (eq_refl (S m)).
pattern (S m) at 1 3 4 6, q; case q; [intro Heq | intros m0 l HeqS].
contradiction (le_Sn_n m); rewrite Heq; assumption.
injection HeqS; intro Heq; generalize l HeqS.
@@ -2536,7 +2536,7 @@ existential variables.
Lemma example_show_existentials : forall a b c:nat, a=b -> b=c -> a=c.
Proof.
intros.
-eapply trans_equal.
+eapply eq_trans.
Show Existentials.
eassumption.
assumption.
diff --git a/doc/faq/interval_discr.v b/doc/faq/interval_discr.v
index ed2c0e37e..671dc988a 100644
--- a/doc/faq/interval_discr.v
+++ b/doc/faq/interval_discr.v
@@ -32,16 +32,16 @@ Theorem le_uniqueness_proof : forall (n m : nat) (p q : n <= m), p = q.
Proof.
induction p using le_ind'; intro q.
replace (le_n n) with
- (eq_rect _ (fun n0 => n <= n0) (le_n n) _ (refl_equal n)).
+ (eq_rect _ (fun n0 => n <= n0) (le_n n) _ eq_refl).
2:reflexivity.
- generalize (refl_equal n).
+ generalize (eq_refl n).
pattern n at 2 4 6 10, q; case q; [intro | intros m l e].
rewrite <- eq_rect_eq_nat; trivial.
contradiction (le_Sn_n m); rewrite <- e; assumption.
replace (le_S n m p) with
- (eq_rect _ (fun n0 => n <= n0) (le_S n m p) _ (refl_equal (S m))).
+ (eq_rect _ (fun n0 => n <= n0) (le_S n m p) _ eq_refl).
2:reflexivity.
- generalize (refl_equal (S m)).
+ generalize (eq_refl (S m)).
pattern (S m) at 1 3 4 6, q; case q; [intro Heq | intros m0 l HeqS].
contradiction (le_Sn_n m); rewrite Heq; assumption.
injection HeqS; intro Heq; generalize l HeqS.
@@ -216,7 +216,7 @@ Lemma inj_restrict :
Proof.
intros A f x y z Hfinj Hneqx Hfy Hfx Heq.
assert (f z <> f x).
- apply sym_not_eq.
+ apply not_eq_sym.
intro Heqf.
apply Hneqx.
apply Hfinj.
@@ -292,7 +292,7 @@ destruct (le_lt_dec (f xSn) (f y)) as [Hlefy|Hgefy].
assert (Heq : x = y).
apply Hfinj.
assert (f xSn <> f y).
- apply sym_not_eq.
+ apply not_eq_sym.
intro Heqf.
apply Hneqy.
apply Hfinj.
@@ -302,7 +302,7 @@ assert (Heq : x = y).
apply le_O_n.
apply le_neq_lt; assumption.
assert (f xSn <> f x).
- apply sym_not_eq.
+ apply not_eq_sym.
intro Heqf.
apply Hneqx.
apply Hfinj.
diff --git a/doc/refman/Classes.tex b/doc/refman/Classes.tex
index 20ff649aa..f7c4bd5ca 100644
--- a/doc/refman/Classes.tex
+++ b/doc/refman/Classes.tex
@@ -71,7 +71,7 @@ Leibniz equality on some type. An example implementation is:
Instance unit_EqDec : EqDec unit :=
{ eqb x y := true ;
eqb_leibniz x y H :=
- match x, y return x = y with tt, tt => refl_equal tt end }.
+ match x, y return x = y with tt, tt => eq_refl tt end }.
\end{coq_example*}
If one does not give all the members in the \texttt{Instance}
diff --git a/doc/refman/Natural.tex b/doc/refman/Natural.tex
index 9a9abe5df..f33c0d356 100644
--- a/doc/refman/Natural.tex
+++ b/doc/refman/Natural.tex
@@ -158,7 +158,7 @@ Add Natural Implicit constr1.
By default, the proposition (or predicate) constructors
\verb=conj=, \verb=or_introl=, \verb=or_intror=, \verb=ex_intro=,
-\verb=exT_intro=, \verb=refl_equal=, \verb=refl_eqT= and \verb=exist=
+\verb=eq_refl= and \verb=exist=
\noindent are declared implicit. Note that declaring implicit the
constructor of a datatype (i.e. an inductive type of type \verb=Set=)
diff --git a/doc/refman/Omega.tex b/doc/refman/Omega.tex
index b9e899ce8..213c05061 100644
--- a/doc/refman/Omega.tex
+++ b/doc/refman/Omega.tex
@@ -51,7 +51,7 @@ on atomic formulas. Atomic formulas are built from the predicates
and in expressions of type \verb=Z=, {\tt omega} recognizes
\begin{quote}
-\verb!+, -, *, Zsucc!, and constants.
+\verb!+, -, *, Z.succ!, and constants.
\end{quote}
All expressions of type \verb=nat= or \verb=Z= not built on these
diff --git a/doc/refman/Polynom.tex b/doc/refman/Polynom.tex
index 3898bf4c4..664c0d3ff 100644
--- a/doc/refman/Polynom.tex
+++ b/doc/refman/Polynom.tex
@@ -927,7 +927,7 @@ Open Scope Z_scope.
Goal forall x y z:Z, x + 3 + y + y * z = x + 3 + y + z * y.
\end{coq_example}
\begin{coq_example*}
-intros; rewrite (Zmult_comm y z); reflexivity.
+intros; rewrite (Z.mul_comm y z); reflexivity.
Save toto.
\end{coq_example*}
\begin{coq_example}
diff --git a/doc/refman/Program.tex b/doc/refman/Program.tex
index 96073d71a..fee070fd6 100644
--- a/doc/refman/Program.tex
+++ b/doc/refman/Program.tex
@@ -53,7 +53,7 @@ will be first rewrote to:
(match x as y return (x = y -> _) with
| 0 => fun H : x = 0 -> t
| S n => fun H : x = S n -> u
- end) (refl_equal n).
+ end) (eq_refl n).
\end{coq_example*}
This permits to get the proper equalities in the context of proof
diff --git a/doc/refman/RefMan-coi.tex b/doc/refman/RefMan-coi.tex
index e609fce82..d75b9eb93 100644
--- a/doc/refman/RefMan-coi.tex
+++ b/doc/refman/RefMan-coi.tex
@@ -277,7 +277,7 @@ an infinite object of type
definition.
\begin{coq_example}
CoFixpoint eqproof (s1 s2:Stream A) : EqSt s1 (conc s1 s2) :=
- eqst s1 (conc s1 s2) (refl_equal (hd A (conc s1 s2)))
+ eqst s1 (conc s1 s2) (eq_refl (hd A (conc s1 s2)))
(eqproof (tl A s1) s2).
\end{coq_example}
\begin{coq_eval}
diff --git a/doc/refman/RefMan-gal.tex b/doc/refman/RefMan-gal.tex
index a9feb34c5..6fa7596db 100644
--- a/doc/refman/RefMan-gal.tex
+++ b/doc/refman/RefMan-gal.tex
@@ -566,16 +566,16 @@ clause where {\ident} is dependent in the return type.
For instance, in the following example:
\begin{coq_example*}
Inductive bool : Type := true : bool | false : bool.
-Inductive eq (A:Type) (x:A) : A -> Prop := refl_equal : eq A x x.
+Inductive eq (A:Type) (x:A) : A -> Prop := eq_refl : eq A x x.
Inductive or (A:Prop) (B:Prop) : Prop :=
| or_introl : A -> or A B
| or_intror : B -> or A B.
Definition bool_case (b:bool) : or (eq bool b true) (eq bool b false)
:= match b as x return or (eq bool x true) (eq bool x false) with
| true => or_introl (eq bool true true) (eq bool true false)
- (refl_equal bool true)
+ (eq_refl bool true)
| false => or_intror (eq bool false true) (eq bool false false)
- (refl_equal bool false)
+ (eq_refl bool false)
end.
\end{coq_example*}
the branches have respective types {\tt or (eq bool true true) (eq
@@ -591,9 +591,9 @@ same meaning as the previous one.
Definition bool_case (b:bool) : or (eq bool b true) (eq bool b false)
:= match b return or (eq bool b true) (eq bool b false) with
| true => or_introl (eq bool true true) (eq bool true false)
- (refl_equal bool true)
+ (eq_refl bool true)
| false => or_intror (eq bool false true) (eq bool false false)
- (refl_equal bool false)
+ (eq_refl bool false)
end.
\end{coq_example*}
@@ -621,9 +621,9 @@ the return type is not dependent on them.
For instance, in the following example:
\begin{coq_example*}
-Definition sym_equal (A:Type) (x y:A) (H:eq A x y) : eq A y x :=
+Definition eq_sym (A:Type) (x y:A) (H:eq A x y) : eq A y x :=
match H in eq _ _ z return eq A z x with
- | refl_equal => refl_equal A x
+ | eq_refl => eq_refl A x
end.
\end{coq_example*}
the type of the branch has type {\tt eq~A~x~x} because the third
diff --git a/doc/refman/RefMan-lib.tex b/doc/refman/RefMan-lib.tex
index 31c6fef4a..26c564e52 100644
--- a/doc/refman/RefMan-lib.tex
+++ b/doc/refman/RefMan-lib.tex
@@ -220,11 +220,11 @@ define \verb:eq: as the smallest reflexive relation, and it is also
equivalent to Leibniz' equality.
\ttindex{eq}
-\ttindex{refl\_equal}
+\ttindex{eq\_refl}
\begin{coq_example*}
Inductive eq (A:Type) (x:A) : A -> Prop :=
- refl_equal : eq A x x.
+ eq_refl : eq A x x.
\end{coq_example*}
\subsubsection[Lemmas]{Lemmas\label{PreludeLemmas}}
@@ -239,8 +239,8 @@ Theorem absurd : forall A C:Prop, A -> ~ A -> C.
\begin{coq_eval}
Abort.
\end{coq_eval}
-\ttindex{sym\_eq}
-\ttindex{trans\_eq}
+\ttindex{eq\_sym}
+\ttindex{eq\_trans}
\ttindex{f\_equal}
\ttindex{sym\_not\_eq}
\begin{coq_example*}
@@ -248,10 +248,10 @@ Section equality.
Variables A B : Type.
Variable f : A -> B.
Variables x y z : A.
-Theorem sym_eq : x = y -> y = x.
-Theorem trans_eq : x = y -> y = z -> x = z.
+Theorem eq_sym : x = y -> y = x.
+Theorem eq_trans : x = y -> y = z -> x = z.
Theorem f_equal : x = y -> f x = f y.
-Theorem sym_not_eq : x <> y -> y <> x.
+Theorem not_eq_sym : x <> y -> y <> x.
\end{coq_example*}
\begin{coq_eval}
Abort.
@@ -280,7 +280,7 @@ Abort.
\end{coq_eval}
%Abort (for now predefined eq_rect)
\begin{coq_example*}
-Hint Immediate sym_eq sym_not_eq : core.
+Hint Immediate eq_sym not_eq_sym : core.
\end{coq_example*}
\ttindex{f\_equal$i$}
@@ -864,22 +864,22 @@ module {\tt ZArith} and opening scope {\tt Z\_scope}.
\begin{tabular}{l|l|l|l}
Notation & Interpretation & Precedence & Associativity\\
\hline
-\verb!_ < _! & {\tt Zlt} &&\\
-\verb!x <= y! & {\tt Zle} &&\\
-\verb!_ > _! & {\tt Zgt} &&\\
-\verb!x >= y! & {\tt Zge} &&\\
+\verb!_ < _! & {\tt Z.lt} &&\\
+\verb!x <= y! & {\tt Z.le} &&\\
+\verb!_ > _! & {\tt Z.gt} &&\\
+\verb!x >= y! & {\tt Z.ge} &&\\
\verb!x < y < z! & {\tt x < y \verb!/\! y < z} &&\\
\verb!x < y <= z! & {\tt x < y \verb!/\! y <= z} &&\\
\verb!x <= y < z! & {\tt x <= y \verb!/\! y < z} &&\\
\verb!x <= y <= z! & {\tt x <= y \verb!/\! y <= z} &&\\
-\verb!_ ?= _! & {\tt Zcompare} & 70 & no\\
-\verb!_ + _! & {\tt Zplus} &&\\
-\verb!_ - _! & {\tt Zminus} &&\\
-\verb!_ * _! & {\tt Zmult} &&\\
-\verb!_ / _! & {\tt Zdiv} &&\\
-\verb!_ mod _! & {\tt Zmod} & 40 & no \\
-\verb!- _! & {\tt Zopp} &&\\
-\verb!_ ^ _! & {\tt Zpower} &&\\
+\verb!_ ?= _! & {\tt Z.compare} & 70 & no\\
+\verb!_ + _! & {\tt Z.add} &&\\
+\verb!_ - _! & {\tt Z.sub} &&\\
+\verb!_ * _! & {\tt Z.mul} &&\\
+\verb!_ / _! & {\tt Z.div} &&\\
+\verb!_ mod _! & {\tt Z.modulo} & 40 & no \\
+\verb!- _! & {\tt Z.opp} &&\\
+\verb!_ ^ _! & {\tt Z.pow} &&\\
\end{tabular}
\end{center}
\caption{Definition of the scope for integer arithmetics ({\tt Z\_scope})}
diff --git a/doc/refman/RefMan-ltac.tex b/doc/refman/RefMan-ltac.tex
index d7f00584e..f10b9c3ee 100644
--- a/doc/refman/RefMan-ltac.tex
+++ b/doc/refman/RefMan-ltac.tex
@@ -607,7 +607,7 @@ Ltac f x :=
match x with
context f [S ?X] =>
idtac X; (* To display the evaluation order *)
- assert (p := refl_equal 1 : X=1); (* To filter the case X=1 *)
+ assert (p := eq_refl 1 : X=1); (* To filter the case X=1 *)
let x:= context f[O] in assert (x=O) (* To observe the context *)
end.
Goal True.
@@ -1205,7 +1205,7 @@ Axiom AR_unit : (A -> unit) = unit.
Axiom AL_unit : (unit -> A) = A.
Lemma Cons : B = C -> A * B = A * C.
Proof.
-intro Heq; rewrite Heq; apply refl_equal.
+intro Heq; rewrite Heq; reflexivity.
Qed.
End Iso_axioms.
\end{coq_example*}
@@ -1272,7 +1272,7 @@ Ltac assoc := repeat rewrite <- Ass.
\begin{coq_example}
Ltac DoCompare n :=
match goal with
- | [ |- (?A = ?A) ] => apply refl_equal
+ | [ |- (?A = ?A) ] => reflexivity
| [ |- (?A * ?B = ?A * ?C) ] =>
apply Cons; let newn := Length B in DoCompare newn
| [ |- (?A * ?B = ?C) ] =>
diff --git a/doc/refman/RefMan-oth.tex b/doc/refman/RefMan-oth.tex
index f81811430..4208787fc 100644
--- a/doc/refman/RefMan-oth.tex
+++ b/doc/refman/RefMan-oth.tex
@@ -266,7 +266,7 @@ may be enclosed by optional {\tt [ ]} delimiters.
Require Import ZArith.
\end{coq_example*}
\begin{coq_example}
-SearchAbout Zmult Zplus "distr".
+SearchAbout Z.mul Z.add "distr".
SearchAbout "+"%Z "*"%Z "distr" -positive -Prop.
SearchAbout (?x * _ + ?x * _)%Z outside OmegaLemmas.
\end{coq_example}
diff --git a/doc/refman/RefMan-tacex.tex b/doc/refman/RefMan-tacex.tex
index 8330a434b..91ff3d5ec 100644
--- a/doc/refman/RefMan-tacex.tex
+++ b/doc/refman/RefMan-tacex.tex
@@ -1129,7 +1129,7 @@ red; intros (x, (y, Hy)).
elim (Hy 0); elim (Hy 1); elim (Hy 2); intros;
match goal with
| [_:(?a = ?b),_:(?a = ?c) |- _ ] =>
- cut (b = c); [ discriminate | apply trans_equal with a; auto ]
+ cut (b = c); [ discriminate | transitivity a; auto ]
end.
Qed.
\end{coq_example*}
@@ -1379,7 +1379,7 @@ Axiom AR_unit : (A -> unit) = unit.
Axiom AL_unit : (unit -> A) = A.
Lemma Cons : B = C -> A * B = A * C.
Proof.
-intro Heq; rewrite Heq; apply refl_equal.
+intro Heq; rewrite Heq; reflexivity.
Qed.
End Iso_axioms.
\end{coq_example*}
@@ -1439,7 +1439,7 @@ Ltac assoc := repeat rewrite <- Ass.
\begin{coq_example}
Ltac DoCompare n :=
match goal with
- | [ |- (?A = ?A) ] => apply refl_equal
+ | [ |- (?A = ?A) ] => reflexivity
| [ |- (?A * ?B = ?A * ?C) ] =>
apply Cons; let newn := Length B in
DoCompare newn
diff --git a/doc/refman/Setoid.tex b/doc/refman/Setoid.tex
index 8e1bb10c9..c0913135c 100644
--- a/doc/refman/Setoid.tex
+++ b/doc/refman/Setoid.tex
@@ -371,9 +371,9 @@ the replaced term occurs in a covariant position.
\begin{cscexample}[Covariance and contravariance]
Suppose that division over real numbers has been defined as a
-morphism of signature \texttt{Zdiv: Zlt ++> Zlt -{}-> Zlt} (i.e.
-\texttt{Zdiv} is increasing in its first argument, but decreasing on the
-second one). Let \texttt{<} denotes \texttt{Zlt}.
+morphism of signature \texttt{Z.div: Z.lt ++> Z.lt -{}-> Z.lt} (i.e.
+\texttt{Z.div} is increasing in its first argument, but decreasing on the
+second one). Let \texttt{<} denotes \texttt{Z.lt}.
Under the hypothesis \texttt{H: x < y} we have
\texttt{k < x / y -> k < x / x}, but not
\texttt{k < y / x -> k < x / x}.
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 86308b6a0..7396fc44c 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -321,7 +321,7 @@ let make_notation_gen loc ntn mknot mkprim destprim l =
if has_curly_brackets ntn
then expand_curly_brackets loc mknot ntn l
else match ntn,List.map destprim l with
- (* Special case to avoid writing "- 3" for e.g. (Zopp 3) *)
+ (* Special case to avoid writing "- 3" for e.g. (Z.opp 3) *)
| "- _", [Some (Numeral p)] when Bigint.is_strictly_pos p ->
mknot (loc,ntn,([mknot (loc,"( _ )",l)]))
| _ ->
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 204af93d5..ab53fec6f 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -34,11 +34,11 @@ let _f_equal = constant ["Init";"Logic"] "f_equal"
let _eq_rect = constant ["Init";"Logic"] "eq_rect"
-let _refl_equal = constant ["Init";"Logic"] "refl_equal"
+let _refl_equal = constant ["Init";"Logic"] "eq_refl"
-let _sym_eq = constant ["Init";"Logic"] "sym_eq"
+let _sym_eq = constant ["Init";"Logic"] "eq_sym"
-let _trans_eq = constant ["Init";"Logic"] "trans_eq"
+let _trans_eq = constant ["Init";"Logic"] "eq_trans"
let _eq = constant ["Init";"Logic"] "eq"
diff --git a/plugins/extraction/ExtrOcamlZBigInt.v b/plugins/extraction/ExtrOcamlZBigInt.v
index 12607b3ad..27fce8eab 100644
--- a/plugins/extraction/ExtrOcamlZBigInt.v
+++ b/plugins/extraction/ExtrOcamlZBigInt.v
@@ -75,13 +75,13 @@ Extract Constant Z.compare => "Big.compare_case Eq Lt Gt".
Extract Constant Z.of_N => "fun p -> p".
Extract Constant Z.abs_N => "Big.abs".
-(** Zdiv and Zmod are quite complex to define in terms of (/) and (mod).
+(** Z.div and Z.modulo are quite complex to define in terms of (/) and (mod).
For the moment we don't even try *)
(** Test:
Require Import ZArith NArith.
Extraction "/tmp/test.ml"
- Pplus Ppred Pminus Pmult Pcompare Npred Nminus Ndiv Nmod Ncompare
- Zplus Zmult BinInt.Zcompare Z_of_N Zabs_N Zdiv.Zdiv Zmod.
+ Pos.add Pos.pred Pos.sub Pos.mul Pos.compare N.pred N.sub N.div N.modulo N.compare
+ Z.add Z.mul Z.compare Z.of_N Z.abs_N Z.div Z.modulo.
*)
diff --git a/plugins/extraction/ExtrOcamlZInt.v b/plugins/extraction/ExtrOcamlZInt.v
index 55ba0ca1c..9566c0186 100644
--- a/plugins/extraction/ExtrOcamlZInt.v
+++ b/plugins/extraction/ExtrOcamlZInt.v
@@ -74,7 +74,7 @@ Extract Constant Z.compare =>
Extract Constant Z.of_N => "fun p -> p".
Extract Constant Z.abs_N => "abs".
-(** Zdiv and Zmod are quite complex to define in terms of (/) and (mod).
+(** Z.div and Z.modulo are quite complex to define in terms of (/) and (mod).
For the moment we don't even try *)
diff --git a/plugins/micromega/Env.v b/plugins/micromega/Env.v
index 5f6c60beb..2ebdf3072 100644
--- a/plugins/micromega/Env.v
+++ b/plugins/micromega/Env.v
@@ -12,10 +12,9 @@
(* *)
(************************************************************************)
-Require Import ZArith.
-Require Import Coq.Arith.Max.
-Require Import List.
+Require Import BinInt List.
Set Implicit Arguments.
+Local Open Scope positive_scope.
Section S.
@@ -23,154 +22,78 @@ Section S.
Definition Env := positive -> D.
- Definition jump (j:positive) (e:Env) := fun x => e (Pplus x j).
+ Definition jump (j:positive) (e:Env) := fun x => e (x+j).
- Definition nth (n:positive) (e : Env ) := e n.
+ Definition nth (n:positive) (e:Env) := e n.
- Definition hd (x:D) (e: Env) := nth xH e.
+ Definition hd (e:Env) := nth 1 e.
- Definition tail (e: Env) := jump xH e.
+ Definition tail (e:Env) := jump 1 e.
- Lemma psucc : forall p, (match p with
- | xI y' => xO (Psucc y')
- | xO y' => xI y'
- | 1%positive => 2%positive
- end) = (p+1)%positive.
+ Lemma jump_add i j l x : jump (i + j) l x = jump i (jump j l) x.
Proof.
- destruct p.
- auto with zarith.
- rewrite xI_succ_xO.
- auto with zarith.
- reflexivity.
+ unfold jump. f_equal. apply Pos.add_assoc.
Qed.
- Lemma jump_Pplus : forall i j l,
- forall x, jump (i + j) l x = jump i (jump j l) x.
- Proof.
- unfold jump.
- intros.
- rewrite Pplus_assoc.
- reflexivity.
- Qed.
-
- Lemma jump_simpl : forall p l,
- forall x, jump p l x =
+ Lemma jump_simpl p l x :
+ jump p l x =
match p with
| xH => tail l x
| xO p => jump p (jump p l) x
| xI p => jump p (jump p (tail l)) x
end.
Proof.
- destruct p ; unfold tail ; intros ; repeat rewrite <- jump_Pplus.
- (* xI p = p + p + 1 *)
- rewrite xI_succ_xO.
- rewrite Pplus_diag.
- rewrite <- Pplus_one_succ_r.
- reflexivity.
- (* xO p = p + p *)
- rewrite Pplus_diag.
- reflexivity.
- reflexivity.
+ destruct p; unfold tail; rewrite <- ?jump_add; f_equal;
+ now rewrite Pos.add_diag.
Qed.
- Ltac jump_s :=
- repeat
- match goal with
- | |- context [jump xH ?e] => rewrite (jump_simpl xH)
- | |- context [jump (xO ?p) ?e] => rewrite (jump_simpl (xO p))
- | |- context [jump (xI ?p) ?e] => rewrite (jump_simpl (xI p))
- end.
-
- Lemma jump_tl : forall j l, forall x, tail (jump j l) x = jump j (tail l) x.
+ Lemma jump_tl j l x : tail (jump j l) x = jump j (tail l) x.
Proof.
- unfold tail.
- intros.
- repeat rewrite <- jump_Pplus.
- rewrite Pplus_comm.
- reflexivity.
+ unfold tail. rewrite <- !jump_add. f_equal. apply Pos.add_comm.
Qed.
- Lemma jump_Psucc : forall j l,
- forall x, (jump (Psucc j) l x) = (jump 1 (jump j l) x).
+ Lemma jump_succ j l x : jump (Pos.succ j) l x = jump 1 (jump j l) x.
Proof.
- intros.
- rewrite <- jump_Pplus.
- rewrite Pplus_one_succ_r.
- rewrite Pplus_comm.
- reflexivity.
+ rewrite <- jump_add. f_equal. symmetry. apply Pos.add_1_l.
Qed.
- Lemma jump_Pdouble_minus_one : forall i l,
- forall x, (jump (Pdouble_minus_one i) (tail l)) x = (jump i (jump i l)) x.
+ Lemma jump_pred_double i l x :
+ jump (Pos.pred_double i) (tail l) x = jump i (jump i l) x.
Proof.
- unfold tail.
- intros.
- repeat rewrite <- jump_Pplus.
- rewrite <- Pplus_one_succ_r.
- rewrite Psucc_o_double_minus_one_eq_xO.
- rewrite Pplus_diag.
- reflexivity.
+ unfold tail. rewrite <- !jump_add. f_equal.
+ now rewrite Pos.add_1_r, Pos.succ_pred_double, Pos.add_diag.
Qed.
- Lemma jump_x0_tail : forall p l, forall x, jump (xO p) (tail l) x = jump (xI p) l x.
- Proof.
- intros.
- unfold jump.
- unfold tail.
- unfold jump.
- rewrite <- Pplus_assoc.
- simpl.
- reflexivity.
- Qed.
-
- Lemma nth_spec : forall p l x,
+ Lemma nth_spec p l :
nth p l =
match p with
- | xH => hd x l
+ | xH => hd l
| xO p => nth p (jump p l)
| xI p => nth p (jump p (tail l))
end.
Proof.
- unfold nth.
- destruct p.
- intros.
- unfold jump, tail.
- unfold jump.
- rewrite Pplus_diag.
- rewrite xI_succ_xO.
- simpl.
- reflexivity.
- unfold jump.
- rewrite Pplus_diag.
- reflexivity.
- unfold hd.
- unfold nth.
- reflexivity.
+ unfold hd, nth, tail, jump.
+ destruct p; f_equal; now rewrite Pos.add_diag.
Qed.
-
- Lemma nth_jump : forall p l x, nth p (tail l) = hd x (jump p l).
+ Lemma nth_jump p l : nth p (tail l) = hd (jump p l).
Proof.
- unfold tail.
- unfold hd.
- unfold jump.
- unfold nth.
- intros.
- rewrite Pplus_comm.
- reflexivity.
+ unfold hd, nth, tail, jump. f_equal. apply Pos.add_comm.
Qed.
- Lemma nth_Pdouble_minus_one :
- forall p l, nth (Pdouble_minus_one p) (tail l) = nth p (jump p l).
+ Lemma nth_pred_double p l :
+ nth (Pos.pred_double p) (tail l) = nth p (jump p l).
Proof.
- intros.
- unfold tail.
- unfold nth, jump.
- rewrite Pplus_diag.
- rewrite <- Psucc_o_double_minus_one_eq_xO.
- rewrite Pplus_one_succ_r.
- reflexivity.
+ unfold nth, tail, jump. f_equal.
+ now rewrite Pos.add_1_r, Pos.succ_pred_double, Pos.add_diag.
Qed.
End S.
+Ltac jump_simpl :=
+ repeat
+ match goal with
+ | |- appcontext [jump xH] => rewrite (jump_simpl xH)
+ | |- appcontext [jump (xO ?p)] => rewrite (jump_simpl (xO p))
+ | |- appcontext [jump (xI ?p)] => rewrite (jump_simpl (xI p))
+ end.
diff --git a/plugins/micromega/EnvRing.v b/plugins/micromega/EnvRing.v
index 309ebdef1..c404919af 100644
--- a/plugins/micromega/EnvRing.v
+++ b/plugins/micromega/EnvRing.v
@@ -11,15 +11,10 @@
Set Implicit Arguments.
-Require Import Setoid.
-Require Import BinList.
-Require Import Env.
-Require Import BinPos.
-Require Import BinNat.
-Require Import BinInt.
+Require Import Setoid Morphisms Env BinPos BinNat BinInt.
Require Export Ring_theory.
-Open Local Scope positive_scope.
+Local Open Scope positive_scope.
Import RingSyntax.
Section MakeRingPol.
@@ -30,7 +25,7 @@ Section MakeRingPol.
Variable req : R -> R -> Prop.
(* Ring properties *)
- Variable Rsth : Setoid_Theory R req.
+ Variable Rsth : Equivalence req.
Variable Reqe : ring_eq_ext radd rmul ropp req.
Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req.
@@ -42,35 +37,55 @@ Section MakeRingPol.
Variable CRmorph : ring_morph rO rI radd rmul rsub ropp req
cO cI cadd cmul csub copp ceqb phi.
- (* Power coefficients *)
+ (* Power coefficients *)
Variable Cpow : Type.
Variable Cp_phi : N -> Cpow.
Variable rpow : R -> Cpow -> R.
Variable pow_th : power_theory rI rmul req Cp_phi rpow.
-
(* R notations *)
Notation "0" := rO. Notation "1" := rI.
- Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y).
- Notation "x - y " := (rsub x y). Notation "- x" := (ropp x).
- Notation "x == y" := (req x y).
+ Infix "+" := radd. Infix "*" := rmul.
+ Infix "-" := rsub. Notation "- x" := (ropp x).
+ Infix "==" := req.
+ Infix "^" := (pow_pos rmul).
(* C notations *)
- Notation "x +! y" := (cadd x y). Notation "x *! y " := (cmul x y).
- Notation "x -! y " := (csub x y). Notation "-! x" := (copp x).
- Notation " x ?=! y" := (ceqb x y). Notation "[ x ]" := (phi x).
-
- (* Usefull tactics *)
- Add Setoid R req Rsth as R_set1.
- Ltac rrefl := gen_reflexivity Rsth.
- Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed.
- Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed.
- Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed.
- Add Morphism rsub : rsub_ext. exact (ARsub_ext Rsth Reqe ARth). Qed.
+ Infix "+!" := cadd. Infix "*!" := cmul.
+ Infix "-! " := csub. Notation "-! x" := (copp x).
+ Infix "?=!" := ceqb. Notation "[ x ]" := (phi x).
+
+ (* Useful tactics *)
+ Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed.
+ Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed.
+ Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed.
+ Add Morphism rsub : rsub_ext. exact (ARsub_ext Rsth Reqe ARth). Qed.
Ltac rsimpl := gen_srewrite Rsth Reqe ARth.
+
Ltac add_push := gen_add_push radd Rsth Reqe ARth.
Ltac mul_push := gen_mul_push rmul Rsth Reqe ARth.
+ Ltac add_permut_rec t :=
+ match t with
+ | ?x + ?y => add_permut_rec y || add_permut_rec x
+ | _ => add_push t; apply (Radd_ext Reqe); [|reflexivity]
+ end.
+
+ Ltac add_permut :=
+ repeat (reflexivity ||
+ match goal with |- ?t == _ => add_permut_rec t end).
+
+ Ltac mul_permut_rec t :=
+ match t with
+ | ?x * ?y => mul_permut_rec y || mul_permut_rec x
+ | _ => mul_push t; apply (Rmul_ext Reqe); [|reflexivity]
+ end.
+
+ Ltac mul_permut :=
+ repeat (reflexivity ||
+ match goal with |- ?t == _ => mul_permut_rec t end).
+
+
(* Definition of multivariable polynomials with coefficients in C :
Type [Pol] represents [X1 ... Xn].
The representation is Horner's where a [n] variable polynomial
@@ -117,19 +132,19 @@ Section MakeRingPol.
| _, _ => false
end.
- Notation " P ?== P' " := (Peq P P').
+ Infix "?==" := Peq.
Definition mkPinj j P :=
match P with
| Pc _ => P
- | Pinj j' Q => Pinj ((j + j'):positive) Q
+ | Pinj j' Q => Pinj (j + j') Q
| _ => Pinj j P
end.
Definition mkPinj_pred j P:=
match j with
| xH => P
- | xO j => Pinj (Pdouble_minus_one j) P
+ | xO j => Pinj (Pos.pred_double j) P
| xI j => Pinj (xO j) P
end.
@@ -157,14 +172,14 @@ Section MakeRingPol.
(** Addition et subtraction *)
- Fixpoint PaddC (P:Pol) (c:C) {struct P} : Pol :=
+ Fixpoint PaddC (P:Pol) (c:C) : Pol :=
match P with
| Pc c1 => Pc (c1 +! c)
| Pinj j Q => Pinj j (PaddC Q c)
| PX P i Q => PX P i (PaddC Q c)
end.
- Fixpoint PsubC (P:Pol) (c:C) {struct P} : Pol :=
+ Fixpoint PsubC (P:Pol) (c:C) : Pol :=
match P with
| Pc c1 => Pc (c1 -! c)
| Pinj j Q => Pinj j (PsubC Q c)
@@ -176,11 +191,11 @@ Section MakeRingPol.
Variable Pop : Pol -> Pol -> Pol.
Variable Q : Pol.
- Fixpoint PaddI (j:positive) (P:Pol){struct P} : Pol :=
+ Fixpoint PaddI (j:positive) (P:Pol) : Pol :=
match P with
| Pc c => mkPinj j (PaddC Q c)
| Pinj j' Q' =>
- match ZPminus j' j with
+ match Z.pos_sub j' j with
| Zpos k => mkPinj j (Pop (Pinj k Q') Q)
| Z0 => mkPinj j (Pop Q' Q)
| Zneg k => mkPinj j' (PaddI k Q')
@@ -188,16 +203,16 @@ Section MakeRingPol.
| PX P i Q' =>
match j with
| xH => PX P i (Pop Q' Q)
- | xO j => PX P i (PaddI (Pdouble_minus_one j) Q')
+ | xO j => PX P i (PaddI (Pos.pred_double j) Q')
| xI j => PX P i (PaddI (xO j) Q')
end
end.
- Fixpoint PsubI (j:positive) (P:Pol){struct P} : Pol :=
+ Fixpoint PsubI (j:positive) (P:Pol) : Pol :=
match P with
| Pc c => mkPinj j (PaddC (--Q) c)
| Pinj j' Q' =>
- match ZPminus j' j with
+ match Z.pos_sub j' j with
| Zpos k => mkPinj j (Pop (Pinj k Q') Q)
| Z0 => mkPinj j (Pop Q' Q)
| Zneg k => mkPinj j' (PsubI k Q')
@@ -205,41 +220,41 @@ Section MakeRingPol.
| PX P i Q' =>
match j with
| xH => PX P i (Pop Q' Q)
- | xO j => PX P i (PsubI (Pdouble_minus_one j) Q')
+ | xO j => PX P i (PsubI (Pos.pred_double j) Q')
| xI j => PX P i (PsubI (xO j) Q')
end
end.
Variable P' : Pol.
- Fixpoint PaddX (i':positive) (P:Pol) {struct P} : Pol :=
+ Fixpoint PaddX (i':positive) (P:Pol) : Pol :=
match P with
| Pc c => PX P' i' P
| Pinj j Q' =>
match j with
| xH => PX P' i' Q'
- | xO j => PX P' i' (Pinj (Pdouble_minus_one j) Q')
+ | xO j => PX P' i' (Pinj (Pos.pred_double j) Q')
| xI j => PX P' i' (Pinj (xO j) Q')
end
| PX P i Q' =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Pop (PX P k P0) P') i' Q'
| Z0 => mkPX (Pop P P') i Q'
| Zneg k => mkPX (PaddX k P) i Q'
end
end.
- Fixpoint PsubX (i':positive) (P:Pol) {struct P} : Pol :=
+ Fixpoint PsubX (i':positive) (P:Pol) : Pol :=
match P with
| Pc c => PX (--P') i' P
| Pinj j Q' =>
match j with
| xH => PX (--P') i' Q'
- | xO j => PX (--P') i' (Pinj (Pdouble_minus_one j) Q')
+ | xO j => PX (--P') i' (Pinj (Pos.pred_double j) Q')
| xI j => PX (--P') i' (Pinj (xO j) Q')
end
| PX P i Q' =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Pop (PX P k P0) P') i' Q'
| Z0 => mkPX (Pop P P') i Q'
| Zneg k => mkPX (PsubX k P) i Q'
@@ -259,18 +274,18 @@ Section MakeRingPol.
| Pinj j Q =>
match j with
| xH => PX P' i' (Padd Q Q')
- | xO j => PX P' i' (Padd (Pinj (Pdouble_minus_one j) Q) Q')
+ | xO j => PX P' i' (Padd (Pinj (Pos.pred_double j) Q) Q')
| xI j => PX P' i' (Padd (Pinj (xO j) Q) Q')
end
| PX P i Q =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Padd (PX P k P0) P') i' (Padd Q Q')
| Z0 => mkPX (Padd P P') i (Padd Q Q')
| Zneg k => mkPX (PaddX Padd P' k P) i (Padd Q Q')
end
end
end.
- Notation "P ++ P'" := (Padd P P').
+ Infix "++" := Padd.
Fixpoint Psub (P P': Pol) {struct P'} : Pol :=
match P' with
@@ -282,22 +297,22 @@ Section MakeRingPol.
| Pinj j Q =>
match j with
| xH => PX (--P') i' (Psub Q Q')
- | xO j => PX (--P') i' (Psub (Pinj (Pdouble_minus_one j) Q) Q')
+ | xO j => PX (--P') i' (Psub (Pinj (Pos.pred_double j) Q) Q')
| xI j => PX (--P') i' (Psub (Pinj (xO j) Q) Q')
end
| PX P i Q =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Psub (PX P k P0) P') i' (Psub Q Q')
| Z0 => mkPX (Psub P P') i (Psub Q Q')
| Zneg k => mkPX (PsubX Psub P' k P) i (Psub Q Q')
end
end
end.
- Notation "P -- P'" := (Psub P P').
+ Infix "--" := Psub.
(** Multiplication *)
- Fixpoint PmulC_aux (P:Pol) (c:C) {struct P} : Pol :=
+ Fixpoint PmulC_aux (P:Pol) (c:C) : Pol :=
match P with
| Pc c' => Pc (c' *! c)
| Pinj j Q => mkPinj j (PmulC_aux Q c)
@@ -311,11 +326,11 @@ Section MakeRingPol.
Section PmulI.
Variable Pmul : Pol -> Pol -> Pol.
Variable Q : Pol.
- Fixpoint PmulI (j:positive) (P:Pol) {struct P} : Pol :=
+ Fixpoint PmulI (j:positive) (P:Pol) : Pol :=
match P with
| Pc c => mkPinj j (PmulC Q c)
| Pinj j' Q' =>
- match ZPminus j' j with
+ match Z.pos_sub j' j with
| Zpos k => mkPinj j (Pmul (Pinj k Q') Q)
| Z0 => mkPinj j (Pmul Q' Q)
| Zneg k => mkPinj j' (PmulI k Q')
@@ -323,13 +338,12 @@ Section MakeRingPol.
| PX P' i' Q' =>
match j with
| xH => mkPX (PmulI xH P') i' (Pmul Q' Q)
- | xO j' => mkPX (PmulI j P') i' (PmulI (Pdouble_minus_one j') Q')
+ | xO j' => mkPX (PmulI j P') i' (PmulI (Pos.pred_double j') Q')
| xI j' => mkPX (PmulI j P') i' (PmulI (xO j') Q')
end
end.
End PmulI.
-(* A symmetric version of the multiplication *)
Fixpoint Pmul (P P'' : Pol) {struct P''} : Pol :=
match P'' with
@@ -342,7 +356,7 @@ Section MakeRingPol.
let QQ' :=
match j with
| xH => Pmul Q Q'
- | xO j => Pmul (Pinj (Pdouble_minus_one j) Q) Q'
+ | xO j => Pmul (Pinj (Pos.pred_double j) Q) Q'
| xI j => Pmul (Pinj (xO j) Q) Q'
end in
mkPX (Pmul P P') i' QQ'
@@ -355,25 +369,7 @@ Section MakeRingPol.
end
end.
-(* Non symmetric *)
-(*
- Fixpoint Pmul_aux (P P' : Pol) {struct P'} : Pol :=
- match P' with
- | Pc c' => PmulC P c'
- | Pinj j' Q' => PmulI Pmul_aux Q' j' P
- | PX P' i' Q' =>
- (mkPX (Pmul_aux P P') i' P0) ++ (PmulI Pmul_aux Q' xH P)
- end.
-
- Definition Pmul P P' :=
- match P with
- | Pc c => PmulC P' c
- | Pinj j Q => PmulI Pmul_aux Q j P'
- | PX P i Q =>
- (mkPX (Pmul_aux P P') i P0) ++ (PmulI Pmul_aux Q xH P')
- end.
-*)
- Notation "P ** P'" := (Pmul P P').
+ Infix "**" := Pmul.
Fixpoint Psquare (P:Pol) : Pol :=
match P with
@@ -388,26 +384,35 @@ Section MakeRingPol.
(** Monomial **)
+ (** A monomial is X1^k1...Xi^ki. Its representation
+ is a simplified version of the polynomial representation:
+
+ - [mon0] correspond to the polynom [P1].
+ - [(zmon j M)] corresponds to [(Pinj j ...)],
+ i.e. skip j variable indices.
+ - [(vmon i M)] is X^i*M with X the current variable,
+ its corresponds to (PX P1 i ...)]
+ *)
+
Inductive Mon: Set :=
- mon0: Mon
+ | mon0: Mon
| zmon: positive -> Mon -> Mon
| vmon: positive -> Mon -> Mon.
- Fixpoint Mphi(l:Env R) (M: Mon) {struct M} : R :=
+ Fixpoint Mphi (l:Env R)(M: Mon) : R :=
match M with
- mon0 => rI
- | zmon j M1 => Mphi (jump j l) M1
- | vmon i M1 =>
- let x := hd 0 l in
- let xi := pow_pos rmul x i in
- (Mphi (tail l) M1) * xi
+ | mon0 => rI
+ | zmon j M1 => Mphi (jump j l) M1
+ | vmon i M1 => Mphi (tail l) M1 * (hd l) ^ i
end.
+ Notation "M @@ l" := (Mphi l M) (at level 10, no associativity).
+
Definition mkZmon j M :=
match M with mon0 => mon0 | _ => zmon j M end.
Definition zmon_pred j M :=
- match j with xH => M | _ => mkZmon (Ppred j) M end.
+ match j with xH => M | _ => mkZmon (Pos.pred j) M end.
Definition mkVmon i M :=
match M with
@@ -416,7 +421,7 @@ Section MakeRingPol.
| vmon i' m => vmon (i+i') m
end.
- Fixpoint MFactor (P: Pol) (M: Mon) {struct P}: Pol * Pol :=
+ Fixpoint MFactor (P: Pol) (M: Mon) : Pol * Pol :=
match P, M with
_, mon0 => (Pc cO, P)
| Pc _, _ => (P, Pc cO)
@@ -453,7 +458,7 @@ Section MakeRingPol.
| _ => Some (Padd Q1 (Pmul P2 R1))
end.
- Fixpoint PNSubst1 (P1: Pol) (M1: Mon) (P2: Pol) (n: nat) {struct n}: Pol :=
+ Fixpoint PNSubst1 (P1: Pol) (M1: Mon) (P2: Pol) (n: nat) : Pol :=
match POneSubst P1 M1 P2 with
Some P3 => match n with S n1 => PNSubst1 P3 M1 P2 n1 | _ => P3 end
| _ => P1
@@ -465,14 +470,13 @@ Section MakeRingPol.
| _ => None
end.
- Fixpoint PSubstL1 (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) {struct LM1}:
- Pol :=
+ Fixpoint PSubstL1 (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) : Pol :=
match LM1 with
cons (M1,P2) LM2 => PSubstL1 (PNSubst1 P1 M1 P2 n) LM2 n
| _ => P1
end.
- Fixpoint PSubstL (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) {struct LM1}: option Pol :=
+ Fixpoint PSubstL (P1: Pol) (LM1: list (Mon * Pol)) (n: nat) : option Pol :=
match LM1 with
cons (M1,P2) LM2 =>
match PNSubst P1 M1 P2 n with
@@ -482,7 +486,7 @@ Section MakeRingPol.
| _ => None
end.
- Fixpoint PNSubstL (P1: Pol) (LM1: list (Mon * Pol)) (m n: nat) {struct m}: Pol :=
+ Fixpoint PNSubstL (P1: Pol) (LM1: list (Mon * Pol)) (m n: nat) : Pol :=
match PSubstL P1 LM1 n with
Some P3 => match m with S m1 => PNSubstL P3 LM1 m1 n | _ => P3 end
| _ => P1
@@ -490,146 +494,112 @@ Section MakeRingPol.
(** Evaluation of a polynomial towards R *)
- Fixpoint Pphi(l:Env R) (P:Pol) {struct P} : R :=
+ Fixpoint Pphi(l:Env R) (P:Pol) : R :=
match P with
| Pc c => [c]
| Pinj j Q => Pphi (jump j l) Q
- | PX P i Q =>
- let x := hd 0 l in
- let xi := pow_pos rmul x i in
- (Pphi l P) * xi + (Pphi (tail l) Q)
+ | PX P i Q => Pphi l P * (hd l) ^ i + Pphi (tail l) Q
end.
Reserved Notation "P @ l " (at level 10, no associativity).
Notation "P @ l " := (Pphi l P).
+
(** Proofs *)
- Lemma ZPminus_spec : forall x y,
- match ZPminus x y with
- | Z0 => x = y
- | Zpos k => x = (y + k)%positive
- | Zneg k => y = (x + k)%positive
+
+ Ltac destr_pos_sub :=
+ match goal with |- context [Z.pos_sub ?x ?y] =>
+ generalize (Z.pos_sub_discr x y); destruct (Z.pos_sub x y)
end.
+
+ Lemma Peq_ok P P' : (P ?== P') = true -> forall l, P@l == P'@ l.
+ Proof.
+ revert P';induction P;destruct P';simpl; intros H l; try easy.
+ - now apply (morph_eq CRmorph).
+ - destruct (Pos.compare_spec p p0); [ subst | easy | easy ].
+ now rewrite IHP.
+ - specialize (IHP1 P'1); specialize (IHP2 P'2).
+ destruct (Pos.compare_spec p p0); [ subst | easy | easy ].
+ destruct (P2 ?== P'1); [|easy].
+ rewrite H in *.
+ now rewrite IHP1, IHP2.
+ Qed.
+
+ Lemma Peq_spec P P' :
+ BoolSpec (forall l, P@l == P'@l) True (P ?== P').
Proof.
- induction x;destruct y.
- replace (ZPminus (xI x) (xI y)) with (Zdouble (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial.
- replace (ZPminus (xI x) (xO y)) with (Zdouble_plus_one (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_plus_one;rewrite H;trivial.
- apply Pplus_xI_double_minus_one.
- simpl;trivial.
- replace (ZPminus (xO x) (xI y)) with (Zdouble_minus_one (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_minus_one;rewrite H;trivial.
- apply Pplus_xI_double_minus_one.
- replace (ZPminus (xO x) (xO y)) with (Zdouble (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial.
- replace (ZPminus (xO x) xH) with (Zpos (Pdouble_minus_one x));trivial.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO;trivial.
- replace (ZPminus xH (xI y)) with (Zneg (xO y));trivial.
- replace (ZPminus xH (xO y)) with (Zneg (Pdouble_minus_one y));trivial.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO;trivial.
- simpl;trivial.
+ generalize (Peq_ok P P'). destruct (P ?== P'); auto.
Qed.
- Lemma Peq_ok : forall P P',
- (P ?== P') = true -> forall l, P@l == P'@ l.
+ Lemma Pphi0 l : P0@l == 0.
Proof.
- induction P;destruct P';simpl;intros;try discriminate;trivial.
- apply (morph_eq CRmorph);trivial.
- assert (H1 := Pos.compare_eq p p0); destruct (p ?= p0);
- try discriminate H.
- rewrite (IHP P' H); rewrite H1;trivial;rrefl.
- assert (H1 := Pos.compare_eq p p0); destruct (p ?= p0);
- try discriminate H.
- rewrite H1;trivial. clear H1.
- assert (H1 := IHP1 P'1);assert (H2 := IHP2 P'2);
- destruct (P2 ?== P'1);[destruct (P3 ?== P'2); [idtac|discriminate H]
- |discriminate H].
- rewrite (H1 H);rewrite (H2 H);rrefl.
+ simpl;apply (morph0 CRmorph).
Qed.
- Lemma Pphi0 : forall l, P0@l == 0.
+ Lemma Pphi1 l : P1@l == 1.
Proof.
- intros;simpl;apply (morph0 CRmorph).
+ simpl;apply (morph1 CRmorph).
Qed.
-Lemma env_morph : forall p e1 e2, (forall x, e1 x = e2 x) ->
- p @ e1 = p @ e2.
+Lemma env_morph p e1 e2 :
+ (forall x, e1 x = e2 x) -> p @ e1 = p @ e2.
Proof.
- induction p ; simpl.
- reflexivity.
- intros.
- apply IHp.
- intros.
- unfold jump.
- apply H.
- intros.
- rewrite (IHp1 e1 e2) ; auto.
- rewrite (IHp2 (tail e1) (tail e2)) ; auto.
- unfold hd. unfold nth. rewrite H. reflexivity.
- unfold tail. unfold jump. intros ; apply H.
+ revert e1 e2. induction p ; simpl.
+ - reflexivity.
+ - intros e1 e2 EQ. apply IHp. intros. apply EQ.
+ - intros e1 e2 EQ. f_equal; [f_equal|].
+ + now apply IHp1.
+ + f_equal. apply EQ.
+ + apply IHp2. intros; apply EQ.
Qed.
-Lemma Pjump_Pplus : forall P i j l, P @ (jump (i + j) l ) = P @ (jump j (jump i l)).
+Lemma Pjump_add P i j l :
+ P @ (jump (i + j) l) = P @ (jump j (jump i l)).
Proof.
- intros. apply env_morph. intros. rewrite <- jump_Pplus.
- rewrite Pplus_comm.
- reflexivity.
+ apply env_morph. intros. rewrite <- jump_add. f_equal.
+ apply Pos.add_comm.
Qed.
-Lemma Pjump_xO_tail : forall P p l,
+Lemma Pjump_xO_tail P p l :
P @ (jump (xO p) (tail l)) = P @ (jump (xI p) l).
Proof.
- intros.
- apply env_morph.
- intros.
- rewrite (@jump_simpl R (xI p)).
- rewrite (@jump_simpl R (xO p)).
- reflexivity.
+ apply env_morph. intros. now jump_simpl.
Qed.
-Lemma Pjump_Pdouble_minus_one : forall P p l,
- P @ (jump (Pdouble_minus_one p) (tail l)) = P @ (jump (xO p) l).
+Lemma Pjump_pred_double P p l :
+ P @ (jump (Pos.pred_double p) (tail l)) = P @ (jump (xO p) l).
Proof.
- intros.
- apply env_morph.
- intros.
- rewrite jump_Pdouble_minus_one.
- rewrite (@jump_simpl R (xO p)).
- reflexivity.
+ apply env_morph. intros.
+ rewrite jump_pred_double. now jump_simpl.
Qed.
-
-
- Lemma Pphi1 : forall l, P1@l == 1.
+ Lemma mkPinj_ok j l P : (mkPinj j P)@l == P@(jump j l).
Proof.
- intros;simpl;apply (morph1 CRmorph).
+ destruct P;simpl;rsimpl.
+ now rewrite Pjump_add.
Qed.
- Lemma mkPinj_ok : forall j l P, (mkPinj j P)@l == P@(jump j l).
+ Lemma pow_pos_add x i j : x^(j + i) == x^i * x^j.
Proof.
- intros j l p;destruct p;simpl;rsimpl.
- rewrite Pjump_Pplus.
- reflexivity.
+ rewrite Pos.add_comm.
+ apply (pow_pos_add Rsth Reqe.(Rmul_ext) ARth.(ARmul_assoc)).
Qed.
- Let pow_pos_Pplus :=
- pow_pos_Pplus rmul Rsth Reqe.(Rmul_ext) ARth.(ARmul_comm) ARth.(ARmul_assoc).
-
- Lemma mkPX_ok : forall l P i Q,
- (mkPX P i Q)@l == P@l*(pow_pos rmul (hd 0 l) i) + Q@(tail l).
+ Lemma ceqb_spec c c' : BoolSpec ([c] == [c']) True (c ?=! c').
Proof.
- intros l P i Q;unfold mkPX.
- destruct P;try (simpl;rrefl).
- assert (H := morph_eq CRmorph c cO);destruct (c ?=! cO);simpl;try rrefl.
- rewrite (H (refl_equal true));rewrite (morph0 CRmorph).
- rewrite mkPinj_ok;rsimpl;simpl;rrefl.
- assert (H := @Peq_ok P3 P0);destruct (P3 ?== P0);simpl;try rrefl.
- rewrite (H (refl_equal true));trivial.
- rewrite Pphi0. rewrite pow_pos_Pplus;rsimpl.
+ generalize (morph_eq CRmorph c c').
+ destruct (c ?=! c'); auto.
Qed.
+ Lemma mkPX_ok l P i Q :
+ (mkPX P i Q)@l == P@l * (hd l)^i + Q@(tail l).
+ Proof.
+ unfold mkPX. destruct P.
+ - case ceqb_spec; intros H; simpl; try reflexivity.
+ rewrite H, (morph0 CRmorph), mkPinj_ok; rsimpl.
+ - reflexivity.
+ - case Peq_spec; intros H; simpl; try reflexivity.
+ rewrite H, Pphi0, Pos.add_comm, pow_pos_add; rsimpl.
+ Qed.
Ltac Esimpl :=
repeat (progress (
@@ -647,43 +617,42 @@ Qed.
end));
rsimpl; simpl.
- Lemma PaddC_ok : forall c P l, (PaddC P c)@l == P@l + [c].
+ Lemma PaddC_ok c P l : (PaddC P c)@l == P@l + [c].
Proof.
- induction P;simpl;intros;Esimpl;trivial.
+ revert l;induction P;simpl;intros;Esimpl;trivial.
rewrite IHP2;rsimpl.
Qed.
- Lemma PsubC_ok : forall c P l, (PsubC P c)@l == P@l - [c].
+ Lemma PsubC_ok c P l : (PsubC P c)@l == P@l - [c].
Proof.
- induction P;simpl;intros.
- Esimpl.
- rewrite IHP;rsimpl.
- rewrite IHP2;rsimpl.
+ revert l;induction P;simpl;intros.
+ - Esimpl.
+ - rewrite IHP;rsimpl.
+ - rewrite IHP2;rsimpl.
Qed.
- Lemma PmulC_aux_ok : forall c P l, (PmulC_aux P c)@l == P@l * [c].
+ Lemma PmulC_aux_ok c P l : (PmulC_aux P c)@l == P@l * [c].
Proof.
- induction P;simpl;intros;Esimpl;trivial.
- rewrite IHP1;rewrite IHP2;rsimpl.
- mul_push ([c]);rrefl.
+ revert l;induction P;simpl;intros;Esimpl;trivial.
+ rewrite IHP1, IHP2;rsimpl. add_permut. mul_permut.
Qed.
- Lemma PmulC_ok : forall c P l, (PmulC P c)@l == P@l * [c].
+ Lemma PmulC_ok c P l : (PmulC P c)@l == P@l * [c].
Proof.
- intros c P l; unfold PmulC.
- assert (H:= morph_eq CRmorph c cO);destruct (c ?=! cO).
- rewrite (H (refl_equal true));Esimpl.
- assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI).
- rewrite (H1 (refl_equal true));Esimpl.
- apply PmulC_aux_ok.
+ unfold PmulC.
+ case ceqb_spec; intros H.
+ - rewrite H; Esimpl.
+ - case ceqb_spec; intros H'.
+ + rewrite H'; Esimpl.
+ + apply PmulC_aux_ok.
Qed.
- Lemma Popp_ok : forall P l, (--P)@l == - P@l.
+ Lemma Popp_ok P l : (--P)@l == - P@l.
Proof.
- induction P;simpl;intros.
- Esimpl.
- apply IHP.
- rewrite IHP1;rewrite IHP2;rsimpl.
+ revert l;induction P;simpl;intros.
+ - Esimpl.
+ - apply IHP.
+ - rewrite IHP1, IHP2;rsimpl.
Qed.
Ltac Esimpl2 :=
@@ -696,520 +665,273 @@ Qed.
| |- context [(--?P)@?l] => rewrite (Popp_ok P l)
end)); Esimpl.
-
-
-
- Lemma Padd_ok : forall P' P l, (P ++ P')@l == P@l + P'@l.
- Proof.
- induction P';simpl;intros;Esimpl2.
- generalize P p l;clear P p l.
- induction P;simpl;intros.
- Esimpl2;apply (ARadd_comm ARth).
- assert (H := ZPminus_spec p p0);destruct (ZPminus p p0).
- rewrite H;Esimpl. rewrite IHP';rrefl.
- rewrite H;Esimpl. rewrite IHP';Esimpl.
- rewrite Pjump_Pplus. rrefl.
- rewrite H;Esimpl. rewrite IHP.
- rewrite Pjump_Pplus. rrefl.
- destruct p0;simpl.
- rewrite IHP2;simpl. rsimpl.
- rewrite Pjump_xO_tail. Esimpl.
- rewrite IHP2;simpl.
- rewrite Pjump_Pdouble_minus_one.
- rsimpl.
- rewrite IHP'.
- rsimpl.
- destruct P;simpl.
- Esimpl2;add_push [c];rrefl.
- destruct p0;simpl;Esimpl2.
- rewrite IHP'2;simpl.
- rewrite Pjump_xO_tail.
- rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl.
- rewrite IHP'2;simpl.
- rewrite Pjump_Pdouble_minus_one. rsimpl.
- add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl.
- rewrite IHP'2;rsimpl.
- unfold tail.
- add_push (P @ (jump 1 l));rrefl.
- assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2.
- rewrite IHP'1;rewrite IHP'2;rsimpl.
- add_push (P3 @ (tail l));rewrite H;rrefl.
- rewrite IHP'1;rewrite IHP'2;simpl;Esimpl.
- rewrite H;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;rsimpl.
- add_push (P3 @ (tail l));rrefl.
- assert (forall P k l,
- (PaddX Padd P'1 k P) @ l == P@l + P'1@l * pow_pos rmul (hd 0 l) k).
- induction P;simpl;intros;try apply (ARadd_comm ARth).
- destruct p2; simpl; try apply (ARadd_comm ARth).
- rewrite Pjump_xO_tail.
- apply (ARadd_comm ARth).
- rewrite Pjump_Pdouble_minus_one.
- apply (ARadd_comm ARth).
- assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2.
- rewrite IHP'1;rsimpl; rewrite H1;add_push (P5 @ (tail l0));rrefl.
- rewrite IHP'1;simpl;Esimpl.
- rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;Esimpl.
- add_push (P5 @ (tail l0));rrefl.
- rewrite IHP1;rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;rsimpl.
- add_push (P5 @ (tail l0));rrefl.
- rewrite H0;rsimpl.
- add_push (P3 @ (tail l)).
- rewrite H;rewrite Pplus_comm.
- rewrite IHP'2;rewrite pow_pos_Pplus;rsimpl.
- add_push (P3 @ (tail l));rrefl.
- Qed.
-
- Lemma Psub_ok : forall P' P l, (P -- P')@l == P@l - P'@l.
+ Lemma PaddX_ok P' P k l :
+ (forall P l, (P++P')@l == P@l + P'@l) ->
+ (PaddX Padd P' k P) @ l == P@l + P'@l * (hd l)^k.
Proof.
- induction P';simpl;intros;Esimpl2;trivial.
- generalize P p l;clear P p l.
- induction P;simpl;intros.
- Esimpl2;apply (ARadd_comm ARth).
- assert (H := ZPminus_spec p p0);destruct (ZPminus p p0).
- rewrite H;Esimpl. rewrite IHP';rsimpl.
- rewrite H;Esimpl. rewrite IHP';Esimpl.
- rewrite <- Pjump_Pplus;rewrite Pplus_comm;rrefl.
- rewrite H;Esimpl. rewrite IHP.
- rewrite <- Pjump_Pplus;rewrite Pplus_comm;rrefl.
- destruct p0;simpl.
- rewrite IHP2;simpl; try rewrite Pjump_xO_tail ; rsimpl.
- rewrite IHP2;simpl.
- rewrite Pjump_Pdouble_minus_one;rsimpl.
- unfold tail ; rsimpl.
- rewrite IHP';rsimpl.
- destruct P;simpl.
- repeat rewrite Popp_ok;Esimpl2;rsimpl;add_push [c];try rrefl.
- destruct p0;simpl;Esimpl2.
- rewrite IHP'2;simpl;rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));trivial.
- rewrite Pjump_xO_tail.
- add_push (P @ ((jump (xI p0) l)));rrefl.
- rewrite IHP'2;simpl;rewrite Pjump_Pdouble_minus_one;rsimpl.
- add_push (- (P'1 @ l * pow_pos rmul (hd 0 l) p));rrefl.
- unfold tail.
- rewrite IHP'2;rsimpl;add_push (P @ (jump 1 l));rrefl.
- assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2.
- rewrite IHP'1; rewrite IHP'2;rsimpl.
- add_push (P3 @ (tail l));rewrite H;rrefl.
- rewrite IHP'1; rewrite IHP'2;rsimpl;simpl;Esimpl.
- rewrite H;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;rsimpl.
- add_push (P3 @ (tail l));rrefl.
- assert (forall P k l,
- (PsubX Psub P'1 k P) @ l == P@l + - P'1@l * pow_pos rmul (hd 0 l) k).
- induction P;simpl;intros.
- rewrite Popp_ok;rsimpl;apply (ARadd_comm ARth);trivial.
- destruct p2;simpl; rewrite Popp_ok;rsimpl.
- rewrite Pjump_xO_tail.
- apply (ARadd_comm ARth);trivial.
- rewrite Pjump_Pdouble_minus_one.
- apply (ARadd_comm ARth);trivial.
- apply (ARadd_comm ARth);trivial.
- assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2;rsimpl.
- rewrite IHP'1;rsimpl;add_push (P5 @ (tail l0));rewrite H1;rrefl.
- rewrite IHP'1;rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;Esimpl.
- add_push (P5 @ (tail l0));rrefl.
- rewrite IHP1;rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;rsimpl.
- add_push (P5 @ (tail l0));rrefl.
- rewrite H0;rsimpl.
- rewrite IHP'2;rsimpl;add_push (P3 @ (tail l)).
- rewrite H;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;rsimpl.
+ intros IHP'.
+ revert k l. induction P;simpl;intros.
+ - add_permut.
+ - destruct p; simpl;
+ rewrite ?Pjump_xO_tail, ?Pjump_pred_double; add_permut.
+ - destr_pos_sub; intros ->;Esimpl2.
+ + rewrite IHP';rsimpl. add_permut.
+ + rewrite IHP', pow_pos_add;simpl;Esimpl. add_permut.
+ + rewrite IHP1, pow_pos_add;rsimpl. add_permut.
Qed.
-(* Proof for the symmetric version *)
- Lemma PmulI_ok :
- forall P',
- (forall (P : Pol) (l : Env R), (Pmul P P') @ l == P @ l * P' @ l) ->
- forall (P : Pol) (p : positive) (l : Env R),
- (PmulI Pmul P' p P) @ l == P @ l * P' @ (jump p l).
+ Lemma Padd_ok P' P l : (P ++ P')@l == P@l + P'@l.
Proof.
- induction P;simpl;intros.
- Esimpl2;apply (ARmul_comm ARth).
- assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2.
- rewrite H1; rewrite H;rrefl.
- rewrite H1; rewrite H.
- rewrite Pjump_Pplus;simpl;rrefl.
- rewrite H1.
- rewrite Pjump_Pplus;rewrite IHP;rrefl.
- destruct p0;Esimpl2.
- rewrite IHP1;rewrite IHP2;rsimpl.
- rewrite Pjump_xO_tail.
- mul_push (pow_pos rmul (hd 0 l) p);rrefl.
- rewrite IHP1;rewrite IHP2;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p); rewrite Pjump_Pdouble_minus_one.
- rrefl.
- rewrite IHP1;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p).
- rewrite H;rrefl.
+ revert P l; induction P';simpl;intros;Esimpl2.
+ - revert p l; induction P;simpl;intros.
+ + Esimpl2; add_permut.
+ + destr_pos_sub; intros ->;Esimpl.
+ * now rewrite IHP'.
+ * rewrite IHP';Esimpl. now rewrite Pjump_add.
+ * rewrite IHP. now rewrite Pjump_add.
+ + destruct p0;simpl.
+ * rewrite IHP2;simpl. rsimpl. rewrite Pjump_xO_tail. Esimpl.
+ * rewrite IHP2;simpl. rewrite Pjump_pred_double. rsimpl.
+ * rewrite IHP'. rsimpl.
+ - destruct P;simpl.
+ + Esimpl2. add_permut.
+ + destruct p0;simpl;Esimpl2; rewrite IHP'2; simpl.
+ * rewrite Pjump_xO_tail. rsimpl. add_permut.
+ * rewrite Pjump_pred_double. rsimpl. add_permut.
+ * rsimpl. unfold tail. add_permut.
+ + destr_pos_sub; intros ->; Esimpl2.
+ * rewrite IHP'1, IHP'2;rsimpl. add_permut.
+ * rewrite IHP'1, IHP'2;simpl;Esimpl.
+ rewrite pow_pos_add;rsimpl. add_permut.
+ * rewrite PaddX_ok by trivial; rsimpl.
+ rewrite IHP'2, pow_pos_add; rsimpl. add_permut.
Qed.
-(*
- Lemma PmulI_ok :
- forall P',
- (forall (P : Pol) (l : list R), (Pmul_aux P P') @ l == P @ l * P' @ l) ->
- forall (P : Pol) (p : positive) (l : list R),
- (PmulI Pmul_aux P' p P) @ l == P @ l * P' @ (jump p l).
+ Lemma PsubX_ok P' P k l :
+ (forall P l, (P--P')@l == P@l - P'@l) ->
+ (PsubX Psub P' k P) @ l == P@l - P'@l * (hd l)^k.
Proof.
- induction P;simpl;intros.
- Esimpl2;apply (ARmul_comm ARth).
- assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2.
- rewrite H1; rewrite H;rrefl.
- rewrite H1; rewrite H.
- rewrite Pplus_comm.
- rewrite jump_Pplus;simpl;rrefl.
- rewrite H1;rewrite Pplus_comm.
- rewrite jump_Pplus;rewrite IHP;rrefl.
- destruct p0;Esimpl2.
- rewrite IHP1;rewrite IHP2;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p);rrefl.
- rewrite IHP1;rewrite IHP2;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p); rewrite jump_Pdouble_minus_one;rrefl.
- rewrite IHP1;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p).
- rewrite H;rrefl.
+ intros IHP'.
+ revert k l. induction P;simpl;intros.
+ - rewrite Popp_ok;rsimpl; add_permut.
+ - destruct p; simpl;
+ rewrite Popp_ok;rsimpl;
+ rewrite ?Pjump_xO_tail, ?Pjump_pred_double; add_permut.
+ - destr_pos_sub; intros ->; Esimpl2; rsimpl.
+ + rewrite IHP';rsimpl. add_permut.
+ + rewrite IHP', pow_pos_add;simpl;Esimpl. add_permut.
+ + rewrite IHP1, pow_pos_add;rsimpl. add_permut.
Qed.
- Lemma Pmul_aux_ok : forall P' P l,(Pmul_aux P P')@l == P@l * P'@l.
+ Lemma Psub_ok P' P l : (P -- P')@l == P@l - P'@l.
Proof.
- induction P';simpl;intros.
- Esimpl2;trivial.
- apply PmulI_ok;trivial.
- rewrite Padd_ok;Esimpl2.
- rewrite (PmulI_ok P'2 IHP'2). rewrite IHP'1. rrefl.
+ revert P l; induction P';simpl;intros;Esimpl2.
+ - revert p l; induction P;simpl;intros.
+ + Esimpl2; add_permut.
+ + destr_pos_sub; intros ->;Esimpl.
+ * rewrite IHP';rsimpl.
+ * rewrite IHP';Esimpl. now rewrite Pjump_add.
+ * rewrite IHP. now rewrite Pjump_add.
+ + destruct p0;simpl.
+ * rewrite IHP2;simpl. rsimpl. rewrite Pjump_xO_tail. Esimpl.
+ * rewrite IHP2;simpl. rewrite Pjump_pred_double. rsimpl.
+ * rewrite IHP'. rsimpl.
+ - destruct P;simpl.
+ + Esimpl2; add_permut.
+ + destruct p0;simpl;Esimpl2; rewrite IHP'2; simpl.
+ * rewrite Pjump_xO_tail. rsimpl. add_permut.
+ * rewrite Pjump_pred_double. rsimpl. add_permut.
+ * rsimpl. unfold tail. add_permut.
+ + destr_pos_sub; intros ->; Esimpl2.
+ * rewrite IHP'1, IHP'2;rsimpl. add_permut.
+ * rewrite IHP'1, IHP'2;simpl;Esimpl.
+ rewrite pow_pos_add;rsimpl. add_permut.
+ * rewrite PsubX_ok by trivial;rsimpl.
+ rewrite IHP'2, pow_pos_add;rsimpl. add_permut.
Qed.
-*)
-(* Proof for the symmetric version *)
- Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
+ Lemma PmulI_ok P' :
+ (forall P l, (Pmul P P') @ l == P @ l * P' @ l) ->
+ forall P p l, (PmulI Pmul P' p P) @ l == P @ l * P' @ (jump p l).
Proof.
- intros P P';generalize P;clear P;induction P';simpl;intros.
- apply PmulC_ok. apply PmulI_ok;trivial.
- destruct P.
- rewrite (ARmul_comm ARth);Esimpl2;Esimpl2.
- Esimpl2. rewrite IHP'1;Esimpl2.
- assert (match p0 with
- | xI j => Pinj (xO j) P ** P'2
- | xO j => Pinj (Pdouble_minus_one j) P ** P'2
- | 1 => P ** P'2
- end @ (tail l) == P @ (jump p0 l) * P'2 @ (tail l)).
- destruct p0;rewrite IHP'2;Esimpl.
- rewrite Pjump_xO_tail. reflexivity.
- rewrite Pjump_Pdouble_minus_one;Esimpl.
- rewrite H;Esimpl.
- rewrite Padd_ok; Esimpl2. rewrite Padd_ok; Esimpl2.
- repeat (rewrite IHP'1 || rewrite IHP'2);simpl.
- rewrite PmulI_ok;trivial.
- unfold tail.
- mul_push (P'1@l). simpl. mul_push (P'2 @ (jump 1 l)). Esimpl.
+ intros IHP'.
+ induction P;simpl;intros.
+ - Esimpl2; mul_permut.
+ - destr_pos_sub; intros ->;Esimpl2.
+ + now rewrite IHP'.
+ + now rewrite IHP', Pjump_add.
+ + now rewrite IHP, Pjump_add.
+ - destruct p0;Esimpl2; rewrite ?IHP1, ?IHP2; rsimpl.
+ + rewrite Pjump_xO_tail. f_equiv. mul_permut.
+ + rewrite Pjump_pred_double. f_equiv. mul_permut.
+ + rewrite IHP'. f_equiv. mul_permut.
Qed.
-(*
-Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
+ Lemma Pmul_ok P P' l : (P**P')@l == P@l * P'@l.
Proof.
- destruct P;simpl;intros.
- Esimpl2;apply (ARmul_comm ARth).
- rewrite (PmulI_ok P (Pmul_aux_ok P)).
- apply (ARmul_comm ARth).
- rewrite Padd_ok; Esimpl2.
- rewrite (PmulI_ok P3 (Pmul_aux_ok P3));trivial.
- rewrite Pmul_aux_ok;mul_push (P' @ l).
- rewrite (ARmul_comm ARth (P' @ l));rrefl.
+ revert P l;induction P';simpl;intros.
+ - apply PmulC_ok.
+ - apply PmulI_ok;trivial.
+ - destruct P.
+ + rewrite (ARmul_comm ARth). Esimpl2; Esimpl2.
+ + Esimpl2. rewrite IHP'1;Esimpl2. f_equiv.
+ destruct p0;rewrite IHP'2;Esimpl.
+ * now rewrite Pjump_xO_tail.
+ * rewrite Pjump_pred_double; Esimpl.
+ + rewrite Padd_ok, !mkPX_ok, Padd_ok, !mkPX_ok,
+ !IHP'1, !IHP'2, PmulI_ok; trivial. Esimpl2.
+ unfold tail.
+ add_permut; f_equiv; mul_permut.
Qed.
-*)
- Lemma Psquare_ok : forall P l, (Psquare P)@l == P@l * P@l.
+ Lemma Psquare_ok P l : (Psquare P)@l == P@l * P@l.
Proof.
- induction P;simpl;intros;Esimpl2.
- apply IHP. rewrite Padd_ok. rewrite Pmul_ok;Esimpl2.
- rewrite IHP1;rewrite IHP2.
- mul_push (pow_pos rmul (hd 0 l) p). mul_push (P2@l).
- rrefl.
+ revert l;induction P;simpl;intros;Esimpl2.
+ - apply IHP.
+ - rewrite Padd_ok, Pmul_ok;Esimpl2.
+ rewrite IHP1, IHP2.
+ mul_push ((hd l)^p). now mul_push (P2@l).
Qed.
- Lemma Mphi_morph : forall P env env', (forall x, env x = env' x ) ->
- Mphi env P = Mphi env' P.
+ Lemma Mphi_morph M e1 e2 :
+ (forall x, e1 x = e2 x) -> M @@ e1 = M @@ e2.
Proof.
- induction P ; simpl.
- reflexivity.
- intros.
- apply IHP.
- intros.
- unfold jump.
- apply H.
- (**)
- intros.
- replace (Mphi (tail env) P) with (Mphi (tail env') P).
- unfold hd. unfold nth.
- rewrite H.
- reflexivity.
- apply IHP.
- unfold tail,jump.
- intros. symmetry. apply H.
+ revert e1 e2; induction M; simpl; intros e1 e2 EQ; trivial.
+ - apply IHM. intros; apply EQ.
+ - f_equal.
+ * apply IHM. intros; apply EQ.
+ * f_equal. apply EQ.
Qed.
-Lemma Mjump_xO_tail : forall M p l,
- Mphi (jump (xO p) (tail l)) M = Mphi (jump (xI p) l) M.
+Lemma Mjump_xO_tail M p l :
+ M @@ (jump (xO p) (tail l)) = M @@ (jump (xI p) l).
Proof.
- intros.
- apply Mphi_morph.
- intros.
- rewrite (@jump_simpl R (xI p)).
- rewrite (@jump_simpl R (xO p)).
- reflexivity.
+ apply Mphi_morph. intros. now jump_simpl.
Qed.
-Lemma Mjump_Pdouble_minus_one : forall M p l,
- Mphi (jump (Pdouble_minus_one p) (tail l)) M = Mphi (jump (xO p) l) M.
+Lemma Mjump_pred_double M p l :
+ M @@ (jump (Pos.pred_double p) (tail l)) = M @@ (jump (xO p) l).
Proof.
- intros.
- apply Mphi_morph.
- intros.
- rewrite jump_Pdouble_minus_one.
- rewrite (@jump_simpl R (xO p)).
- reflexivity.
+ apply Mphi_morph. intros.
+ rewrite jump_pred_double. now jump_simpl.
Qed.
-Lemma Mjump_Pplus : forall M i j l, Mphi (jump (i + j) l ) M = Mphi (jump j (jump i l)) M.
+Lemma Mjump_add M i j l :
+ M @@ (jump (i + j) l) = M @@ (jump j (jump i l)).
Proof.
- intros. apply Mphi_morph. intros. rewrite <- jump_Pplus.
- rewrite Pplus_comm.
- reflexivity.
+ apply Mphi_morph. intros. now rewrite <- jump_add, Pos.add_comm.
Qed.
-
-
- Lemma mkZmon_ok: forall M j l,
- Mphi l (mkZmon j M) == Mphi l (zmon j M).
- intros M j l; case M; simpl; intros; rsimpl.
+ Lemma mkZmon_ok M j l :
+ (mkZmon j M) @@ l == (zmon j M) @@ l.
+ Proof.
+ destruct M; simpl; rsimpl.
Qed.
- Lemma zmon_pred_ok : forall M j l,
- Mphi (tail l) (zmon_pred j M) == Mphi l (zmon j M).
+ Lemma zmon_pred_ok M j l :
+ (zmon_pred j M) @@ (tail l) == (zmon j M) @@ l.
Proof.
- destruct j; simpl;intros l; rsimpl.
- rewrite mkZmon_ok;rsimpl.
- simpl.
- rewrite Mjump_xO_tail.
- reflexivity.
- rewrite mkZmon_ok;simpl.
- rewrite Mjump_Pdouble_minus_one; rsimpl.
+ destruct j; simpl; rewrite ?mkZmon_ok; simpl; rsimpl.
+ - now rewrite Mjump_xO_tail.
+ - rewrite Mjump_pred_double; rsimpl.
Qed.
- Lemma mkVmon_ok : forall M i l, Mphi l (mkVmon i M) == Mphi l M*pow_pos rmul (hd 0 l) i.
+ Lemma mkVmon_ok M i l :
+ (mkVmon i M)@@l == M@@l * (hd l)^i.
Proof.
destruct M;simpl;intros;rsimpl.
- rewrite zmon_pred_ok;simpl;rsimpl.
- rewrite Pplus_comm;rewrite pow_pos_Pplus;rsimpl.
+ - rewrite zmon_pred_ok;simpl;rsimpl.
+ - rewrite pow_pos_add;rsimpl.
Qed.
+ Ltac destr_mfactor R S := match goal with
+ | H : context [MFactor ?P _] |- context [MFactor ?P ?M] =>
+ specialize (H M); destruct MFactor as (R,S)
+ end.
- Lemma Mphi_ok: forall P M l,
- let (Q,R) := MFactor P M in
- P@l == Q@l + (Mphi l M) * (R@l).
+ Lemma Mphi_ok P M l :
+ let (Q,R) := MFactor P M in
+ P@l == Q@l + M@@l * R@l.
Proof.
- intros P; elim P; simpl; auto; clear P.
- intros c M l; case M; simpl; auto; try intro p; try intro m;
- try rewrite (morph0 CRmorph); rsimpl.
-
- intros i P Hrec M l; case M; simpl; clear M.
- rewrite (morph0 CRmorph); rsimpl.
- intros j M.
- case_eq (i ?= j); intros He; simpl.
- rewrite (Pos.compare_eq _ _ He).
- generalize (Hrec M (jump j l)); case (MFactor P M);
- simpl; intros P2 Q2 H; repeat rewrite mkPinj_ok; auto.
- generalize (Hrec (zmon (j -i) M) (jump i l));
- case (MFactor P (zmon (j -i) M)); simpl.
- intros P2 Q2 H; repeat rewrite mkPinj_ok; auto.
- rewrite <- (Pplus_minus _ _ (ZC2 _ _ He)).
- rewrite Mjump_Pplus; auto.
- rewrite (morph0 CRmorph); rsimpl.
- intros P2 m; rewrite (morph0 CRmorph); rsimpl.
-
- intros P2 Hrec1 i Q2 Hrec2 M l; case M; simpl; auto.
- rewrite (morph0 CRmorph); rsimpl.
- intros j M1.
- generalize (Hrec1 (zmon j M1) l);
- case (MFactor P2 (zmon j M1)).
- intros R1 S1 H1.
- generalize (Hrec2 (zmon_pred j M1) (tail l));
- case (MFactor Q2 (zmon_pred j M1)); simpl.
- intros R2 S2 H2; rewrite H1; rewrite H2.
- repeat rewrite mkPX_ok; simpl.
- rsimpl.
- apply radd_ext; rsimpl.
- rewrite (ARadd_comm ARth); rsimpl.
- apply radd_ext; rsimpl.
- rewrite (ARadd_comm ARth); rsimpl.
- rewrite zmon_pred_ok;rsimpl.
- intros j M1.
- case_eq (i ?= j); intros He; simpl.
- rewrite (Pos.compare_eq _ _ He).
- generalize (Hrec1 (mkZmon xH M1) l); case (MFactor P2 (mkZmon xH M1));
- simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto.
- rewrite H; rewrite mkPX_ok; rsimpl.
- repeat (rewrite <-(ARadd_assoc ARth)).
- apply radd_ext; rsimpl.
- rewrite (ARadd_comm ARth); rsimpl.
- apply radd_ext; rsimpl.
- repeat (rewrite <-(ARmul_assoc ARth)).
- rewrite mkZmon_ok.
- apply rmul_ext; rsimpl.
- rewrite (ARmul_comm ARth); rsimpl.
- generalize (Hrec1 (vmon (j - i) M1) l);
- case (MFactor P2 (vmon (j - i) M1));
- simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto.
- rewrite H; rsimpl; repeat rewrite mkPinj_ok; auto.
- rewrite mkPX_ok; rsimpl.
- repeat (rewrite <-(ARadd_assoc ARth)).
- apply radd_ext; rsimpl.
- rewrite (ARadd_comm ARth); rsimpl.
- apply radd_ext; rsimpl.
- repeat (rewrite <-(ARmul_assoc ARth)).
- apply rmul_ext; rsimpl.
- rewrite (ARmul_comm ARth); rsimpl.
- apply rmul_ext; rsimpl.
- rewrite <- pow_pos_Pplus.
- rewrite (Pplus_minus _ _ (ZC2 _ _ He)); rsimpl.
- generalize (Hrec1 (mkZmon 1 M1) l);
- case (MFactor P2 (mkZmon 1 M1));
- simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto.
- rewrite H; rsimpl.
- rewrite mkPX_ok; rsimpl.
- repeat (rewrite <-(ARadd_assoc ARth)).
- apply radd_ext; rsimpl.
- rewrite (ARadd_comm ARth); rsimpl.
- apply radd_ext; rsimpl.
- rewrite mkZmon_ok.
- repeat (rewrite <-(ARmul_assoc ARth)).
- apply rmul_ext; rsimpl.
- rewrite (ARmul_comm ARth); rsimpl.
- rewrite mkPX_ok; simpl; rsimpl.
- rewrite (morph0 CRmorph); rsimpl.
- repeat (rewrite <-(ARmul_assoc ARth)).
- rewrite (ARmul_comm ARth (Q3@l)); rsimpl.
- apply rmul_ext; rsimpl.
- rewrite <- pow_pos_Pplus.
- rewrite (Pplus_minus _ _ He); rsimpl.
+ revert M l; induction P; destruct M; intros l; simpl; auto; Esimpl.
+ - case Pos.compare_spec; intros He; simpl.
+ * destr_mfactor R1 S1. now rewrite IHP, He, !mkPinj_ok.
+ * destr_mfactor R1 S1. rewrite IHP; simpl.
+ now rewrite !mkPinj_ok, <- Mjump_add, Pos.add_comm, Pos.sub_add.
+ * Esimpl.
+ - destr_mfactor R1 S1. destr_mfactor R2 S2.
+ rewrite IHP1, IHP2, !mkPX_ok, zmon_pred_ok; simpl; rsimpl.
+ add_permut.
+ - case Pos.compare_spec; intros He; simpl; destr_mfactor R1 S1;
+ rewrite ?He, IHP1, mkPX_ok, ?mkZmon_ok; simpl; rsimpl;
+ unfold tail; add_permut; mul_permut.
+ * rewrite <- pow_pos_add, Pos.add_comm, Pos.sub_add by trivial; rsimpl.
+ * rewrite mkPX_ok; Esimpl2. mul_permut.
+ rewrite <- pow_pos_add, Pos.sub_add by trivial; rsimpl.
Qed.
-(* Proof for the symmetric version *)
-
- Lemma POneSubst_ok: forall P1 M1 P2 P3 l,
- POneSubst P1 M1 P2 = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l.
+ Lemma POneSubst_ok P1 M1 P2 P3 l :
+ POneSubst P1 M1 P2 = Some P3 -> M1@@l == P2@l ->
+ P1@l == P3@l.
Proof.
- intros P2 M1 P3 P4 l; unfold POneSubst.
- generalize (Mphi_ok P2 M1 l); case (MFactor P2 M1); simpl; auto.
- intros Q1 R1; case R1.
- intros c H; rewrite H.
- generalize (morph_eq CRmorph c cO);
- case (c ?=! cO); simpl; auto.
- intros H1 H2; rewrite H1; auto; rsimpl.
- discriminate.
- intros _ H1 H2; injection H1; intros; subst.
- rewrite H2; rsimpl.
- (* new version *)
- rewrite Padd_ok; rewrite PmulC_ok; rsimpl.
- intros i P5 H; rewrite H.
- intros HH H1; injection HH; intros; subst; rsimpl.
- rewrite Padd_ok; rewrite PmulI_ok by (intros;apply Pmul_ok). rewrite H1; rsimpl.
- intros i P5 P6 H1 H2 H3; rewrite H1; rewrite H3.
- assert (P4 = Q1 ++ P3 ** PX i P5 P6).
- injection H2; intros; subst;trivial.
- rewrite H;rewrite Padd_ok;rewrite Pmul_ok;rsimpl.
-Qed.
-(*
- Lemma POneSubst_ok: forall P1 M1 P2 P3 l,
- POneSubst P1 M1 P2 = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l.
-Proof.
- intros P2 M1 P3 P4 l; unfold POneSubst.
- generalize (Mphi_ok P2 M1 l); case (MFactor P2 M1); simpl; auto.
- intros Q1 R1; case R1.
- intros c H; rewrite H.
- generalize (morph_eq CRmorph c cO);
- case (c ?=! cO); simpl; auto.
- intros H1 H2; rewrite H1; auto; rsimpl.
- discriminate.
- intros _ H1 H2; injection H1; intros; subst.
- rewrite H2; rsimpl.
- rewrite Padd_ok; rewrite Pmul_ok; rsimpl.
- intros i P5 H; rewrite H.
- intros HH H1; injection HH; intros; subst; rsimpl.
- rewrite Padd_ok; rewrite Pmul_ok. rewrite H1; rsimpl.
- intros i P5 P6 H1 H2 H3; rewrite H1; rewrite H3.
- injection H2; intros; subst; rsimpl.
- rewrite Padd_ok.
- rewrite Pmul_ok; rsimpl.
+ unfold POneSubst.
+ assert (H := Mphi_ok P1). destr_mfactor R1 S1. rewrite H; clear H.
+ intros EQ EQ'.
+ assert (EQ'' : P3 = R1 ++ P2 ** S1).
+ { revert EQ. destruct S1; try now injection 1.
+ case ceqb_spec; try discriminate. now injection 2. }
+ rewrite EQ', EQ''; clear EQ EQ' EQ''.
+ rewrite Padd_ok, Pmul_ok; rsimpl.
Qed.
-*)
- Lemma PNSubst1_ok: forall n P1 M1 P2 l,
- Mphi l M1 == P2@l -> P1@l == (PNSubst1 P1 M1 P2 n)@l.
+
+ Lemma PNSubst1_ok n P1 M1 P2 l :
+ M1@@l == P2@l -> P1@l == (PNSubst1 P1 M1 P2 n)@l.
Proof.
- intros n; elim n; simpl; auto.
- intros P2 M1 P3 l H.
- generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l);
- case (POneSubst P2 M1 P3); [idtac | intros; rsimpl].
- intros P4 Hrec; rewrite (Hrec P4); auto; rsimpl.
- intros n1 Hrec P2 M1 P3 l H.
- generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l);
- case (POneSubst P2 M1 P3); [idtac | intros; rsimpl].
- intros P4 Hrec1; rewrite (Hrec1 P4); auto; rsimpl.
+ revert P1. induction n; simpl; intros P1;
+ generalize (POneSubst_ok P1 M1 P2); destruct POneSubst;
+ intros; rewrite <- ?IHn; auto; reflexivity.
Qed.
- Lemma PNSubst_ok: forall n P1 M1 P2 l P3,
- PNSubst P1 M1 P2 n = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l.
+ Lemma PNSubst_ok n P1 M1 P2 l P3 :
+ PNSubst P1 M1 P2 n = Some P3 -> M1@@l == P2@l -> P1@l == P3@l.
Proof.
- intros n P2 M1 P3 l P4; unfold PNSubst.
- generalize (fun P4 => @POneSubst_ok P2 M1 P3 P4 l);
- case (POneSubst P2 M1 P3); [idtac | intros; discriminate].
- intros P5 H1; case n; try (intros; discriminate).
- intros n1 H2; injection H2; intros; subst.
- rewrite <- PNSubst1_ok; auto.
+ unfold PNSubst.
+ assert (H := POneSubst_ok P1 M1 P2); destruct POneSubst; try discriminate.
+ destruct n; [discriminate | injection 1; intros EQ'].
+ intros. rewrite <- EQ', <- PNSubst1_ok; auto.
Qed.
- Fixpoint MPcond (LM1: list (Mon * Pol)) (l: Env R) {struct LM1} : Prop :=
- match LM1 with
- cons (M1,P2) LM2 => (Mphi l M1 == P2@l) /\ (MPcond LM2 l)
- | _ => True
- end.
+ Fixpoint MPcond (LM1: list (Mon * Pol)) (l: Env R) : Prop :=
+ match LM1 with
+ | cons (M1,P2) LM2 => (Mphi l M1 == P2@l) /\ (MPcond LM2 l)
+ | _ => True
+ end.
- Lemma PSubstL1_ok: forall n LM1 P1 l,
- MPcond LM1 l -> P1@l == (PSubstL1 P1 LM1 n)@l.
+ Lemma PSubstL1_ok n LM1 P1 l :
+ MPcond LM1 l -> P1@l == (PSubstL1 P1 LM1 n)@l.
Proof.
- intros n LM1; elim LM1; simpl; auto.
- intros; rsimpl.
- intros (M2,P2) LM2 Hrec P3 l [H H1].
- rewrite <- Hrec; auto.
- apply PNSubst1_ok; auto.
+ revert P1; induction LM1 as [|(M2,P2) LM2 IH]; simpl; intros.
+ - reflexivity.
+ - rewrite <- IH by intuition. now apply PNSubst1_ok.
Qed.
- Lemma PSubstL_ok: forall n LM1 P1 P2 l,
- PSubstL P1 LM1 n = Some P2 -> MPcond LM1 l -> P1@l == P2@l.
+ Lemma PSubstL_ok n LM1 P1 P2 l :
+ PSubstL P1 LM1 n = Some P2 -> MPcond LM1 l -> P1@l == P2@l.
Proof.
- intros n LM1; elim LM1; simpl; auto.
- intros; discriminate.
- intros (M2,P2) LM2 Hrec P3 P4 l.
- generalize (PNSubst_ok n P3 M2 P2); case (PNSubst P3 M2 P2 n).
- intros P5 H0 H1 [H2 H3]; injection H1; intros; subst.
- rewrite <- PSubstL1_ok; auto.
- intros l1 H [H1 H2]; auto.
+ revert P1. induction LM1 as [|(M2,P2') LM2 IH]; simpl; intros.
+ - discriminate.
+ - assert (H':=PNSubst_ok n P3 M2 P2'). destruct PNSubst.
+ * injection H; intros <-. rewrite <- PSubstL1_ok; intuition.
+ * now apply IH.
Qed.
- Lemma PNSubstL_ok: forall m n LM1 P1 l,
- MPcond LM1 l -> P1@l == (PNSubstL P1 LM1 m n)@l.
+ Lemma PNSubstL_ok m n LM1 P1 l :
+ MPcond LM1 l -> P1@l == (PNSubstL P1 LM1 m n)@l.
Proof.
- intros m; elim m; simpl; auto.
- intros n LM1 P2 l H; generalize (fun P3 => @PSubstL_ok n LM1 P2 P3 l);
- case (PSubstL P2 LM1 n); intros; rsimpl; auto.
- intros m1 Hrec n LM1 P2 l H.
- generalize (fun P3 => @PSubstL_ok n LM1 P2 P3 l);
- case (PSubstL P2 LM1 n); intros; rsimpl; auto.
- rewrite <- Hrec; auto.
+ revert LM1 P1. induction m; simpl; intros;
+ assert (H' := PSubstL_ok n LM1 P2); destruct PSubstL;
+ auto; try reflexivity.
+ rewrite <- IHm; auto.
Qed.
(** Definition of polynomial expressions *)
@@ -1228,7 +950,7 @@ Proof.
(** evaluation of polynomial expressions towards R *)
- Fixpoint PEeval (l:Env R) (pe:PExpr) {struct pe} : R :=
+ Fixpoint PEeval (l:Env R) (pe:PExpr) : R :=
match pe with
| PEc c => phi c
| PEX j => nth j l
@@ -1241,60 +963,27 @@ Proof.
(** Correctness proofs *)
- Lemma mkX_ok : forall p l, nth p l == (mk_X p) @ l.
+ Lemma mkX_ok p l : nth p l == (mk_X p) @ l.
Proof.
destruct p;simpl;intros;Esimpl;trivial.
rewrite nth_spec ; auto.
unfold hd.
- rewrite <- nth_Pdouble_minus_one.
- rewrite (nth_jump (Pdouble_minus_one p) l 1).
- reflexivity.
+ now rewrite <- nth_pred_double, nth_jump.
Qed.
Ltac Esimpl3 :=
repeat match goal with
| |- context [(?P1 ++ ?P2)@?l] => rewrite (Padd_ok P2 P1 l)
| |- context [(?P1 -- ?P2)@?l] => rewrite (Psub_ok P2 P1 l)
- end;Esimpl2;try rrefl;try apply (ARadd_comm ARth).
-
-(* Power using the chinise algorithm *)
-(*Section POWER.
- Variable subst_l : Pol -> Pol.
- Fixpoint Ppow_pos (P:Pol) (p:positive){struct p} : Pol :=
- match p with
- | xH => P
- | xO p => subst_l (Psquare (Ppow_pos P p))
- | xI p => subst_l (Pmul P (Psquare (Ppow_pos P p)))
- end.
-
- Definition Ppow_N P n :=
- match n with
- | N0 => P1
- | Npos p => Ppow_pos P p
- end.
-
- Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) ->
- forall P p, (Ppow_pos P p)@l == (pow_pos Pmul P p)@l.
- Proof.
- intros l subst_l_ok P.
- induction p;simpl;intros;try rrefl;try rewrite subst_l_ok.
- repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl.
- repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl.
- Qed.
-
- Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) ->
- forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l.
- Proof. destruct n;simpl. rrefl. apply Ppow_pos_ok. trivial. Qed.
-
- End POWER. *)
+ end;Esimpl2;try reflexivity;try apply (ARadd_comm ARth).
Section POWER.
Variable subst_l : Pol -> Pol.
- Fixpoint Ppow_pos (res P:Pol) (p:positive){struct p} : Pol :=
+ Fixpoint Ppow_pos (res P:Pol) (p:positive) : Pol :=
match p with
- | xH => subst_l (Pmul res P)
+ | xH => subst_l (res ** P)
| xO p => Ppow_pos (Ppow_pos res P p) P p
- | xI p => subst_l (Pmul (Ppow_pos (Ppow_pos res P p) P p) P)
+ | xI p => subst_l ((Ppow_pos (Ppow_pos res P p) P p) ** P)
end.
Definition Ppow_N P n :=
@@ -1303,17 +992,23 @@ Section POWER.
| Npos p => Ppow_pos P1 P p
end.
- Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) ->
- forall res P p, (Ppow_pos res P p)@l == res@l * (pow_pos Pmul P p)@l.
+ Lemma Ppow_pos_ok l :
+ (forall P, subst_l P@l == P@l) ->
+ forall res P p, (Ppow_pos res P p)@l == res@l * (pow_pos Pmul P p)@l.
Proof.
- intros l subst_l_ok res P p. generalize res;clear res.
- induction p;simpl;intros;try rewrite subst_l_ok; repeat rewrite Pmul_ok;repeat rewrite IHp.
- rsimpl. mul_push (P@l);rsimpl. rsimpl. rrefl.
+ intros subst_l_ok res P p. revert res.
+ induction p;simpl;intros; rewrite ?subst_l_ok, ?Pmul_ok, ?IHp;
+ mul_permut.
Qed.
- Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) ->
- forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l.
- Proof. destruct n;simpl. rrefl. rewrite Ppow_pos_ok. trivial. Esimpl. auto. Qed.
+ Lemma Ppow_N_ok l :
+ (forall P, subst_l P@l == P@l) ->
+ forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l.
+ Proof.
+ destruct n;simpl.
+ - reflexivity.
+ - rewrite Ppow_pos_ok by trivial. Esimpl.
+ Qed.
End POWER.
@@ -1342,57 +1037,23 @@ Section POWER.
Definition norm_subst pe := subst_l (norm_aux pe).
- (*
- Fixpoint norm_subst (pe:PExpr) : Pol :=
- match pe with
- | PEc c => Pc c
- | PEX j => subst_l (mk_X j)
- | PEadd (PEopp pe1) pe2 => Psub (norm_subst pe2) (norm_subst pe1)
- | PEadd pe1 (PEopp pe2) =>
- Psub (norm_subst pe1) (norm_subst pe2)
- | PEadd pe1 pe2 => Padd (norm_subst pe1) (norm_subst pe2)
- | PEsub pe1 pe2 => Psub (norm_subst pe1) (norm_subst pe2)
- | PEmul pe1 pe2 => Pmul_subst (norm_subst pe1) (norm_subst pe2)
- | PEopp pe1 => Popp (norm_subst pe1)
- | PEpow pe1 n => Ppow_subst (norm_subst pe1) n
- end.
+ Lemma norm_aux_opp pe :
+ norm_aux (PEopp pe) = Popp (norm_aux pe).
+ Proof. reflexivity. Qed.
- Lemma norm_subst_spec :
- forall l pe, MPcond lmp l ->
- PEeval l pe == (norm_subst pe)@l.
- Proof.
- intros;assert (subst_l_ok:forall P, (subst_l P)@l == P@l).
- unfold subst_l;intros.
- rewrite <- PNSubstL_ok;trivial. rrefl.
- assert (Pms_ok:forall P1 P2, (Pmul_subst P1 P2)@l == P1@l*P2@l).
- intros;unfold Pmul_subst;rewrite subst_l_ok;rewrite Pmul_ok;rrefl.
- induction pe;simpl;Esimpl3.
- rewrite subst_l_ok;apply mkX_ok.
- rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3.
- rewrite IHpe1;rewrite IHpe2;rrefl.
- rewrite Pms_ok;rewrite IHpe1;rewrite IHpe2;rrefl.
- rewrite IHpe;rrefl.
- unfold Ppow_subst. rewrite Ppow_N_ok. trivial.
- rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3.
- induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok;
- repeat rewrite Pmul_ok;rrefl.
- Qed.
-*)
- Lemma norm_aux_spec :
- forall l pe, (*MPcond lmp l ->*)
- PEeval l pe == (norm_aux pe)@l.
+ Lemma norm_aux_spec l pe : PEeval l pe == (norm_aux pe)@l.
Proof.
- intros.
- induction pe;simpl;Esimpl3.
- apply mkX_ok.
- rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3.
- rewrite IHpe1;rewrite IHpe2;rrefl.
- rewrite IHpe1;rewrite IHpe2. rewrite Pmul_ok. rrefl.
- rewrite IHpe;rrefl.
- rewrite Ppow_N_ok by reflexivity.
- rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3.
- induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok;
- repeat rewrite Pmul_ok;rrefl.
+ induction pe.
+ - reflexivity.
+ - apply mkX_ok.
+ - simpl PEeval. rewrite IHpe1, IHpe2.
+ simpl. destruct pe1, pe2; Esimpl3.
+ - simpl. now rewrite IHpe1, IHpe2, Psub_ok.
+ - simpl. now rewrite IHpe1, IHpe2, Pmul_ok.
+ - simpl. rewrite IHpe. Esimpl2.
+ - simpl. rewrite Ppow_N_ok by reflexivity.
+ rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl2.
+ induction p;simpl; now rewrite ?IHp, ?IHpe, ?Pms_ok, ?Pmul_ok.
Qed.
diff --git a/plugins/micromega/MExtraction.v b/plugins/micromega/MExtraction.v
index 19a98f876..dcaccaa9f 100644
--- a/plugins/micromega/MExtraction.v
+++ b/plugins/micromega/MExtraction.v
@@ -51,7 +51,7 @@ Extract Constant Rinv => "fun x -> 1 / x".
Extraction "micromega.ml"
List.map simpl_cone (*map_cone indexes*)
denorm Qpower
- n_of_Z N_of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find.
+ n_of_Z N.of_nat ZTautoChecker ZWeakChecker QTautoChecker RTautoChecker find.
diff --git a/plugins/micromega/Psatz.v b/plugins/micromega/Psatz.v
index 7f6cf79be..114ac0ab4 100644
--- a/plugins/micromega/Psatz.v
+++ b/plugins/micromega/Psatz.v
@@ -81,14 +81,14 @@ Ltac lra :=
first [ psatzl R | psatzl Q ].
Ltac lia :=
- zify ; unfold Zsucc in * ;
- (*cbv delta - [Zplus Zminus Zopp Zmult Zpower Zgt Zge Zle Zlt iff not] ;*) xlia ;
+ zify ; unfold Z.succ in * ;
+ (*cbv delta - [Z.add Z.sub Z.opp Z.mul Z.pow Z.gt Z.ge Z.le Z.lt iff not] ;*) xlia ;
intros __wit __varmap __ff ;
change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
apply (ZTautoChecker_sound __ff __wit); vm_compute ; reflexivity.
Ltac nia :=
- zify ; unfold Zsucc in * ;
+ zify ; unfold Z.succ in * ;
xnlia ;
intros __wit __varmap __ff ;
change (Tauto.eval_f (Zeval_formula (@find Z Z0 __varmap)) __ff) ;
diff --git a/plugins/micromega/QMicromega.v b/plugins/micromega/QMicromega.v
index f64504a54..74961f1b5 100644
--- a/plugins/micromega/QMicromega.v
+++ b/plugins/micromega/QMicromega.v
@@ -60,7 +60,7 @@ Proof.
Qed.
-(*Definition Zeval_expr := eval_pexpr 0 Zplus Zmult Zminus Zopp (fun x => x) (fun x => Z_of_N x) (Zpower).*)
+(*Definition Zeval_expr := eval_pexpr 0 Z.add Z.mul Z.sub Z.opp (fun x => x) (fun x => Z.of_N x) (Z.pow).*)
Require Import EnvRing.
Fixpoint Qeval_expr (env: PolEnv Q) (e: PExpr Q) : Q :=
@@ -71,7 +71,7 @@ Fixpoint Qeval_expr (env: PolEnv Q) (e: PExpr Q) : Q :=
| PEsub pe1 pe2 => (Qeval_expr env pe1) - (Qeval_expr env pe2)
| PEmul pe1 pe2 => (Qeval_expr env pe1) * (Qeval_expr env pe2)
| PEopp pe1 => - (Qeval_expr env pe1)
- | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z_of_N n)
+ | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z.of_N n)
end.
Lemma Qeval_expr_simpl : forall env e,
@@ -83,7 +83,7 @@ Lemma Qeval_expr_simpl : forall env e,
| PEsub pe1 pe2 => (Qeval_expr env pe1) - (Qeval_expr env pe2)
| PEmul pe1 pe2 => (Qeval_expr env pe1) * (Qeval_expr env pe2)
| PEopp pe1 => - (Qeval_expr env pe1)
- | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z_of_N n)
+ | PEpow pe1 n => Qpower (Qeval_expr env pe1) (Z.of_N n)
end.
Proof.
destruct e ; reflexivity.
@@ -91,7 +91,7 @@ Qed.
Definition Qeval_expr' := eval_pexpr Qplus Qmult Qminus Qopp (fun x => x) (fun x => x) (pow_N 1 Qmult).
-Lemma QNpower : forall r n, r ^ Z_of_N n = pow_N 1 Qmult r n.
+Lemma QNpower : forall r n, r ^ Z.of_N n = pow_N 1 Qmult r n.
Proof.
destruct n ; reflexivity.
Qed.
diff --git a/plugins/micromega/RMicromega.v b/plugins/micromega/RMicromega.v
index 2be99da1e..e575ed29e 100644
--- a/plugins/micromega/RMicromega.v
+++ b/plugins/micromega/RMicromega.v
@@ -85,17 +85,17 @@ Qed.
Ltac INR_nat_of_P :=
match goal with
- | H : context[INR (nat_of_P ?X)] |- _ =>
+ | H : context[INR (Pos.to_nat ?X)] |- _ =>
revert H ;
let HH := fresh in
- assert (HH := pos_INR_nat_of_P X) ; revert HH ; generalize (INR (nat_of_P X))
- | |- context[INR (nat_of_P ?X)] =>
+ assert (HH := pos_INR_nat_of_P X) ; revert HH ; generalize (INR (Pos.to_nat X))
+ | |- context[INR (Pos.to_nat ?X)] =>
let HH := fresh in
- assert (HH := pos_INR_nat_of_P X) ; revert HH ; generalize (INR (nat_of_P X))
+ assert (HH := pos_INR_nat_of_P X) ; revert HH ; generalize (INR (Pos.to_nat X))
end.
Ltac add_eq expr val := set (temp := expr) ;
- generalize (refl_equal temp) ;
+ generalize (eq_refl temp) ;
unfold temp at 1 ; generalize temp ; intro val ; clear temp.
Ltac Rinv_elim :=
@@ -210,7 +210,7 @@ Proof.
rewrite plus_IZR in *.
rewrite mult_IZR in *.
simpl.
- rewrite nat_of_P_mult_morphism.
+ rewrite Pos2Nat.inj_mul.
rewrite mult_INR.
rewrite mult_IZR.
simpl.
@@ -244,7 +244,7 @@ Proof.
simpl.
repeat rewrite mult_IZR.
simpl.
- rewrite nat_of_P_mult_morphism.
+ rewrite Pos2Nat.inj_mul.
rewrite mult_INR.
repeat INR_nat_of_P.
intros. field ; split ; apply Rlt_neq ; auto.
@@ -275,7 +275,7 @@ Proof.
apply Rlt_neq ; auto.
simpl in H.
exfalso.
- rewrite Pmult_comm in H.
+ rewrite Pos.mul_comm in H.
compute in H.
discriminate.
Qed.
@@ -291,7 +291,7 @@ Proof.
destruct x.
unfold Qopp.
simpl.
- rewrite Zopp_involutive.
+ rewrite Z.opp_involutive.
reflexivity.
Qed.
@@ -348,7 +348,7 @@ Proof.
intros.
assert ( 0 > x \/ 0 < x)%Q.
destruct x ; unfold Qlt, Qeq in * ; simpl in *.
- rewrite Zmult_1_r in *.
+ rewrite Z.mul_1_r in *.
destruct Qnum ; simpl in * ; intuition auto.
right. reflexivity.
left ; reflexivity.
@@ -379,7 +379,7 @@ Proof.
Qed.
-Notation to_nat := N.to_nat. (*Nnat.nat_of_N*)
+Notation to_nat := N.to_nat.
Lemma QSORaddon :
@SORaddon R
@@ -471,7 +471,7 @@ Definition INZ (n:N) : R :=
| Npos p => IZR (Zpos p)
end.
-Definition Reval_expr := eval_pexpr Rplus Rmult Rminus Ropp R_of_Rcst nat_of_N pow.
+Definition Reval_expr := eval_pexpr Rplus Rmult Rminus Ropp R_of_Rcst N.to_nat pow.
Definition Reval_op2 (o:Op2) : R -> R -> Prop :=
@@ -490,10 +490,10 @@ Definition Reval_formula (e: PolEnv R) (ff : Formula Rcst) :=
Definition Reval_formula' :=
- eval_sformula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt nat_of_N pow R_of_Rcst.
+ eval_sformula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt N.to_nat pow R_of_Rcst.
Definition QReval_formula :=
- eval_formula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt IQR nat_of_N pow .
+ eval_formula Rplus Rmult Rminus Ropp (@eq R) Rle Rlt IQR N.to_nat pow .
Lemma Reval_formula_compat : forall env f, Reval_formula env f <-> Reval_formula' env f.
Proof.
diff --git a/plugins/micromega/RingMicromega.v b/plugins/micromega/RingMicromega.v
index 4af650861..48bf3e2ac 100644
--- a/plugins/micromega/RingMicromega.v
+++ b/plugins/micromega/RingMicromega.v
@@ -142,7 +142,7 @@ Qed.
Definition PolC := Pol C. (* polynomials in generalized Horner form, defined in Ring_polynom or EnvRing *)
Definition PolEnv := Env R. (* For interpreting PolC *)
Definition eval_pol (env : PolEnv) (p:PolC) : R :=
- Pphi 0 rplus rtimes phi env p.
+ Pphi rplus rtimes phi env p.
Inductive Op1 : Set := (* relations with 0 *)
| Equal (* == 0 *)
@@ -320,7 +320,7 @@ Definition map_option2 (A B C : Type) (f : A -> B -> option C)
Arguments map_option2 [A B C] f o o'.
-Definition Rops_wd := mk_reqe rplus rtimes ropp req
+Definition Rops_wd := mk_reqe (*rplus rtimes ropp req*)
sor.(SORplus_wd)
sor.(SORtimes_wd)
sor.(SORopp_wd).
@@ -469,17 +469,11 @@ Fixpoint ge_bool (n m : nat) : bool :=
end
end.
-Lemma ge_bool_cases : forall n m, (if ge_bool n m then n >= m else n < m)%nat.
+Lemma ge_bool_cases : forall n m,
+ (if ge_bool n m then n >= m else n < m)%nat.
Proof.
- induction n ; simpl.
- destruct m ; simpl.
- constructor.
- omega.
- destruct m.
- constructor.
- omega.
- generalize (IHn m).
- destruct (ge_bool n m) ; omega.
+ induction n; destruct m ; simpl; auto with arith.
+ specialize (IHn m). destruct (ge_bool); auto with arith.
Qed.
@@ -593,7 +587,7 @@ Definition paddC := PaddC cplus.
Definition psubC := PsubC cminus.
Definition PsubC_ok : forall c P env, eval_pol env (psubC P c) == eval_pol env P - [c] :=
- let Rops_wd := mk_reqe rplus rtimes ropp req
+ let Rops_wd := mk_reqe (*rplus rtimes ropp req*)
sor.(SORplus_wd)
sor.(SORtimes_wd)
sor.(SORopp_wd) in
@@ -601,7 +595,7 @@ Definition PsubC_ok : forall c P env, eval_pol env (psubC P c) == eval_pol env
addon.(SORrm).
Definition PaddC_ok : forall c P env, eval_pol env (paddC P c) == eval_pol env P + [c] :=
- let Rops_wd := mk_reqe rplus rtimes ropp req
+ let Rops_wd := mk_reqe (*rplus rtimes ropp req*)
sor.(SORplus_wd)
sor.(SORtimes_wd)
sor.(SORopp_wd) in
@@ -882,13 +876,14 @@ Qed.
Fixpoint xdenorm (jmp : positive) (p: Pol C) : PExpr C :=
match p with
| Pc c => PEc c
- | Pinj j p => xdenorm (Pplus j jmp ) p
+ | Pinj j p => xdenorm (Pos.add j jmp ) p
| PX p j q => PEadd
(PEmul (xdenorm jmp p) (PEpow (PEX _ jmp) (Npos j)))
- (xdenorm (Psucc jmp) q)
+ (xdenorm (Pos.succ jmp) q)
end.
-Lemma xdenorm_correct : forall p i env, eval_pol (jump i env) p == eval_pexpr env (xdenorm (Psucc i) p).
+Lemma xdenorm_correct : forall p i env,
+ eval_pol (jump i env) p == eval_pexpr env (xdenorm (Pos.succ i) p).
Proof.
unfold eval_pol.
induction p.
@@ -896,22 +891,21 @@ Proof.
(* Pinj *)
simpl.
intros.
- rewrite Pplus_succ_permute_r.
+ rewrite Pos.add_succ_r.
rewrite <- IHp.
symmetry.
- rewrite Pplus_comm.
- rewrite Pjump_Pplus. reflexivity.
+ rewrite Pos.add_comm.
+ rewrite Pjump_add. reflexivity.
(* PX *)
simpl.
intros.
- rewrite <- IHp1.
- rewrite <- IHp2.
+ rewrite <- IHp1, <- IHp2.
unfold Env.tail , Env.hd.
- rewrite <- Pjump_Pplus.
- rewrite <- Pplus_one_succ_r.
+ rewrite <- Pjump_add.
+ rewrite Pos.add_1_r.
unfold Env.nth.
unfold jump at 2.
- rewrite Pplus_one_succ_l.
+ rewrite <- Pos.add_1_l.
rewrite addon.(SORpower).(rpow_pow_N).
unfold pow_N. ring.
Qed.
@@ -924,14 +918,14 @@ Proof.
induction p.
reflexivity.
simpl.
- rewrite <- Pplus_one_succ_r.
+ rewrite Pos.add_1_r.
apply xdenorm_correct.
simpl.
intros.
rewrite IHp1.
unfold Env.tail.
rewrite xdenorm_correct.
- change (Psucc xH) with 2%positive.
+ change (Pos.succ xH) with 2%positive.
rewrite addon.(SORpower).(rpow_pow_N).
simpl. reflexivity.
Qed.
diff --git a/plugins/micromega/ZCoeff.v b/plugins/micromega/ZCoeff.v
index 2bf3d8c35..b43ce6f04 100644
--- a/plugins/micromega/ZCoeff.v
+++ b/plugins/micromega/ZCoeff.v
@@ -109,7 +109,7 @@ Qed.
Lemma Zring_morph :
ring_morph 0 1 rplus rtimes rminus ropp req
- 0%Z 1%Z Zplus Zmult Zminus Zopp
+ 0%Z 1%Z Z.add Z.mul Z.sub Z.opp
Zeq_bool gen_order_phi_Z.
Proof.
exact (gen_phiZ_morph sor.(SORsetoid) ring_ops_wd sor.(SORrt)).
@@ -122,7 +122,7 @@ try apply (Rplus_pos_pos sor); try apply (Rtimes_pos_pos sor); try apply (Rplus_
try apply (Rlt_0_1 sor); assumption.
Qed.
-Lemma phi_pos1_succ : forall x : positive, phi_pos1 (Psucc x) == 1 + phi_pos1 x.
+Lemma phi_pos1_succ : forall x : positive, phi_pos1 (Pos.succ x) == 1 + phi_pos1 x.
Proof.
exact (ARgen_phiPOS_Psucc sor.(SORsetoid) ring_ops_wd
(Rth_ARth sor.(SORsetoid) ring_ops_wd sor.(SORrt))).
@@ -130,7 +130,7 @@ Qed.
Lemma clt_pos_morph : forall x y : positive, (x < y)%positive -> phi_pos1 x < phi_pos1 y.
Proof.
-intros x y H. pattern y; apply Plt_ind with x.
+intros x y H. pattern y; apply Pos.lt_ind with x.
rewrite phi_pos1_succ; apply (Rlt_succ_r sor).
clear y H; intros y _ H. rewrite phi_pos1_succ. now apply (Rlt_lt_succ sor).
assumption.
@@ -150,9 +150,9 @@ apply -> (Ropp_lt_mono sor); apply clt_pos_morph.
red. now rewrite Pos.compare_antisym.
Qed.
-Lemma Zcleb_morph : forall x y : Z, Zle_bool x y = true -> [x] <= [y].
+Lemma Zcleb_morph : forall x y : Z, Z.leb x y = true -> [x] <= [y].
Proof.
-unfold Zle_bool; intros x y H.
+unfold Z.leb; intros x y H.
case_eq (x ?= y)%Z; intro H1; rewrite H1 in H.
le_equal. apply Zring_morph.(morph_eq). unfold Zeq_bool; now rewrite H1.
le_less. now apply clt_morph.
@@ -162,9 +162,9 @@ Qed.
Lemma Zcneqb_morph : forall x y : Z, Zeq_bool x y = false -> [x] ~= [y].
Proof.
intros x y H. unfold Zeq_bool in H.
-case_eq (Zcompare x y); intro H1; rewrite H1 in *; (discriminate || clear H).
+case_eq (Z.compare x y); intro H1; rewrite H1 in *; (discriminate || clear H).
apply (Rlt_neq sor). now apply clt_morph.
-fold (x > y)%Z in H1. rewrite Zgt_iff_lt in H1.
+fold (x > y)%Z in H1. rewrite Z.gt_lt_iff in H1.
apply (Rneq_symm sor). apply (Rlt_neq sor). now apply clt_morph.
Qed.
diff --git a/plugins/micromega/ZMicromega.v b/plugins/micromega/ZMicromega.v
index 461f53b54..b65774406 100644
--- a/plugins/micromega/ZMicromega.v
+++ b/plugins/micromega/ZMicromega.v
@@ -34,20 +34,20 @@ Require Import EnvRing.
Open Scope Z_scope.
-Lemma Zsor : SOR 0 1 Zplus Zmult Zminus Zopp (@eq Z) Zle Zlt.
+Lemma Zsor : SOR 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt.
Proof.
constructor ; intros ; subst ; try (intuition (auto with zarith)).
apply Zsth.
apply Zth.
- destruct (Ztrichotomy n m) ; intuition (auto with zarith).
- apply Zmult_lt_0_compat ; auto.
+ destruct (Z.lt_trichotomy n m) ; intuition.
+ apply Z.mul_pos_pos ; auto.
Qed.
Lemma ZSORaddon :
- SORaddon 0 1 Zplus Zmult Zminus Zopp (@eq Z) Zle (* ring elements *)
- 0%Z 1%Z Zplus Zmult Zminus Zopp (* coefficients *)
- Zeq_bool Zle_bool
- (fun x => x) (fun x => x) (pow_N 1 Zmult).
+ SORaddon 0 1 Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le (* ring elements *)
+ 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (* coefficients *)
+ Zeq_bool Z.leb
+ (fun x => x) (fun x => x) (pow_N 1 Z.mul).
Proof.
constructor.
constructor ; intros ; try reflexivity.
@@ -65,20 +65,20 @@ Fixpoint Zeval_expr (env : PolEnv Z) (e: PExpr Z) : Z :=
| PEX x => env x
| PEadd e1 e2 => Zeval_expr env e1 + Zeval_expr env e2
| PEmul e1 e2 => Zeval_expr env e1 * Zeval_expr env e2
- | PEpow e1 n => Zpower (Zeval_expr env e1) (Z_of_N n)
+ | PEpow e1 n => Z.pow (Zeval_expr env e1) (Z.of_N n)
| PEsub e1 e2 => (Zeval_expr env e1) - (Zeval_expr env e2)
- | PEopp e => Zopp (Zeval_expr env e)
+ | PEopp e => Z.opp (Zeval_expr env e)
end.
-Definition eval_expr := eval_pexpr Zplus Zmult Zminus Zopp (fun x => x) (fun x => x) (pow_N 1 Zmult).
+Definition eval_expr := eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x => x) (fun x => x) (pow_N 1 Z.mul).
-Lemma ZNpower : forall r n, r ^ Z_of_N n = pow_N 1 Zmult r n.
+Lemma ZNpower : forall r n, r ^ Z.of_N n = pow_N 1 Z.mul r n.
Proof.
destruct n.
reflexivity.
simpl.
- unfold Zpower_pos.
- replace (pow_pos Zmult r p) with (1 * (pow_pos Zmult r p)) by ring.
+ unfold Z.pow_pos.
+ replace (pow_pos Z.mul r p) with (1 * (pow_pos Z.mul r p)) by ring.
generalize 1.
induction p; simpl ; intros ; repeat rewrite IHp ; ring.
Qed.
@@ -94,10 +94,10 @@ Definition Zeval_op2 (o : Op2) : Z -> Z -> Prop :=
match o with
| OpEq => @eq Z
| OpNEq => fun x y => ~ x = y
-| OpLe => Zle
-| OpGe => Zge
-| OpLt => Zlt
-| OpGt => Zgt
+| OpLe => Z.le
+| OpGe => Z.ge
+| OpLt => Z.lt
+| OpGt => Z.gt
end.
Definition Zeval_formula (env : PolEnv Z) (f : Formula Z):=
@@ -105,23 +105,23 @@ Definition Zeval_formula (env : PolEnv Z) (f : Formula Z):=
(Zeval_op2 op) (Zeval_expr env lhs) (Zeval_expr env rhs).
Definition Zeval_formula' :=
- eval_formula Zplus Zmult Zminus Zopp (@eq Z) Zle Zlt (fun x => x) (fun x => x) (pow_N 1 Zmult).
+ eval_formula Z.add Z.mul Z.sub Z.opp (@eq Z) Z.le Z.lt (fun x => x) (fun x => x) (pow_N 1 Z.mul).
Lemma Zeval_formula_compat : forall env f, Zeval_formula env f <-> Zeval_formula' env f.
Proof.
destruct f ; simpl.
rewrite Zeval_expr_compat. rewrite Zeval_expr_compat.
unfold eval_expr.
- generalize (eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x)
- (fun x : N => x) (pow_N 1 Zmult) env Flhs).
- generalize ((eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x)
- (fun x : N => x) (pow_N 1 Zmult) env Frhs)).
+ generalize (eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x)
+ (fun x : N => x) (pow_N 1 Z.mul) env Flhs).
+ generalize ((eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x)
+ (fun x : N => x) (pow_N 1 Z.mul) env Frhs)).
destruct Fop ; simpl; intros ; intuition (auto with zarith).
Qed.
Definition eval_nformula :=
- eval_nformula 0 Zplus Zmult (@eq Z) Zle Zlt (fun x => x) .
+ eval_nformula 0 Z.add Z.mul (@eq Z) Z.le Z.lt (fun x => x) .
Definition Zeval_op1 (o : Op1) : Z -> Prop :=
match o with
@@ -140,7 +140,7 @@ Qed.
Definition ZWitness := Psatz Z.
-Definition ZWeakChecker := check_normalised_formulas 0 1 Zplus Zmult Zeq_bool Zle_bool.
+Definition ZWeakChecker := check_normalised_formulas 0 1 Z.add Z.mul Zeq_bool Z.leb.
Lemma ZWeakChecker_sound : forall (l : list (NFormula Z)) (cm : ZWitness),
ZWeakChecker l cm = true ->
@@ -154,13 +154,13 @@ Proof.
exact H.
Qed.
-Definition psub := psub Z0 Zplus Zminus Zopp Zeq_bool.
+Definition psub := psub Z0 Z.add Z.sub Z.opp Zeq_bool.
-Definition padd := padd Z0 Zplus Zeq_bool.
+Definition padd := padd Z0 Z.add Zeq_bool.
-Definition norm := norm 0 1 Zplus Zmult Zminus Zopp Zeq_bool.
+Definition norm := norm 0 1 Z.add Z.mul Z.sub Z.opp Zeq_bool.
-Definition eval_pol := eval_pol 0 Zplus Zmult (fun x => x).
+Definition eval_pol := eval_pol Z.add Z.mul (fun x => x).
Lemma eval_pol_sub : forall env lhs rhs, eval_pol env (psub lhs rhs) = eval_pol env lhs - eval_pol env rhs.
Proof.
@@ -211,10 +211,10 @@ Proof.
repeat rewrite eval_pol_add;
repeat rewrite <- eval_pol_norm ; simpl in *;
unfold eval_expr;
- generalize ( eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x)
- (fun x : N => x) (pow_N 1 Zmult) env lhs);
- generalize (eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x)
- (fun x : N => x) (pow_N 1 Zmult) env rhs) ; intros z1 z2 ; intros ; subst;
+ generalize ( eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x)
+ (fun x : N => x) (pow_N 1 Z.mul) env lhs);
+ generalize (eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x)
+ (fun x : N => x) (pow_N 1 Z.mul) env rhs) ; intros z1 z2 ; intros ; subst;
intuition (auto with zarith).
Transparent padd.
Qed.
@@ -248,17 +248,17 @@ Proof.
repeat rewrite eval_pol_add;
repeat rewrite <- eval_pol_norm ; simpl in *;
unfold eval_expr;
- generalize ( eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x)
- (fun x : N => x) (pow_N 1 Zmult) env lhs);
- generalize (eval_pexpr Zplus Zmult Zminus Zopp (fun x : Z => x)
- (fun x : N => x) (pow_N 1 Zmult) env rhs) ; intros z1 z2 ; intros ; subst;
+ generalize ( eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x)
+ (fun x : N => x) (pow_N 1 Z.mul) env lhs);
+ generalize (eval_pexpr Z.add Z.mul Z.sub Z.opp (fun x : Z => x)
+ (fun x : N => x) (pow_N 1 Z.mul) env rhs) ; intros z1 z2 ; intros ; subst;
intuition (auto with zarith).
Transparent padd.
Qed.
-Definition Zunsat := check_inconsistent 0 Zeq_bool Zle_bool.
+Definition Zunsat := check_inconsistent 0 Zeq_bool Z.leb.
-Definition Zdeduce := nformula_plus_nformula 0 Zplus Zeq_bool.
+Definition Zdeduce := nformula_plus_nformula 0 Z.add Zeq_bool.
Definition ZweakTautoChecker (w: list ZWitness) (f : BFormula (Formula Z)) : bool :=
@@ -270,7 +270,7 @@ Require Import Zdiv.
Open Scope Z_scope.
Definition ceiling (a b:Z) : Z :=
- let (q,r) := Zdiv_eucl a b in
+ let (q,r) := Z.div_eucl a b in
match r with
| Z0 => q
| _ => q + 1
@@ -279,47 +279,38 @@ Definition ceiling (a b:Z) : Z :=
Require Import Znumtheory.
-Lemma Zdivide_ceiling : forall a b, (b | a) -> ceiling a b = Zdiv a b.
+Lemma Zdivide_ceiling : forall a b, (b | a) -> ceiling a b = Z.div a b.
Proof.
unfold ceiling.
intros.
apply Zdivide_mod in H.
- case_eq (Zdiv_eucl a b).
+ case_eq (Z.div_eucl a b).
intros.
change z with (fst (z,z0)).
rewrite <- H0.
- change (fst (Zdiv_eucl a b)) with (Zdiv a b).
+ change (fst (Z.div_eucl a b)) with (Z.div a b).
change z0 with (snd (z,z0)).
rewrite <- H0.
- change (snd (Zdiv_eucl a b)) with (Zmod a b).
+ change (snd (Z.div_eucl a b)) with (Z.modulo a b).
rewrite H.
reflexivity.
Qed.
-Lemma narrow_interval_lower_bound : forall a b x, a > 0 -> a * x >= b -> x >= ceiling b a.
+Lemma narrow_interval_lower_bound a b x :
+ a > 0 -> a * x >= b -> x >= ceiling b a.
Proof.
+ rewrite !Z.ge_le_iff.
unfold ceiling.
- intros.
- generalize (Z_div_mod b a H).
- destruct (Zdiv_eucl b a).
- intros.
- destruct H1.
- destruct H2.
- subst.
- destruct (Ztrichotomy z0 0) as [ HH1 | [HH2 | HH3]]; destruct z0 ; try auto with zarith ; try discriminate.
- assert (HH :x >= z \/ x < z) by (destruct (Ztrichotomy x z) ; auto with zarith).
- destruct HH ;auto.
- generalize (Zmult_lt_compat_l _ _ _ H3 H1).
- auto with zarith.
- clear H2.
- assert (HH :x >= z +1 \/ x <= z) by (destruct (Ztrichotomy x z) ; intuition (auto with zarith)).
- destruct HH ;auto.
- assert (0 < a) by auto with zarith.
- generalize (Zmult_lt_0_le_compat_r _ _ _ H2 H1).
- intros.
- rewrite Zmult_comm in H4.
- rewrite (Zmult_comm z) in H4.
- auto with zarith.
+ intros Ha H.
+ generalize (Z_div_mod b a Ha).
+ destruct (Z.div_eucl b a) as (q,r). intros (->,(H1,H2)).
+ destruct r as [|r|r].
+ - rewrite Z.add_0_r in H.
+ apply Z.mul_le_mono_pos_l in H; auto with zarith.
+ - assert (0 < Z.pos r) by easy.
+ rewrite Z.add_1_r, Z.le_succ_l.
+ apply Z.mul_lt_mono_pos_l with a; auto with zarith.
+ - now elim H1.
Qed.
(** NB: narrow_interval_upper_bound is Zdiv.Zdiv_le_lower_bound *)
@@ -360,7 +351,7 @@ Proof.
destruct x ; simpl ; intuition congruence.
Qed.
-Definition ZgcdM (x y : Z) := Zmax (Zgcd x y) 1.
+Definition ZgcdM (x y : Z) := Z.max (Z.gcd x y) 1.
Fixpoint Zgcd_pol (p : PolC Z) : (Z * Z) :=
@@ -378,7 +369,7 @@ Fixpoint Zgcd_pol (p : PolC Z) : (Z * Z) :=
Fixpoint Zdiv_pol (p:PolC Z) (x:Z) : PolC Z :=
match p with
- | Pc c => Pc (Zdiv c x)
+ | Pc c => Pc (Z.div c x)
| Pinj j p => Pinj j (Zdiv_pol p x)
| PX p j q => PX (Zdiv_pol p x) j (Zdiv_pol q x)
end.
@@ -421,10 +412,10 @@ Proof.
intros.
simpl.
unfold ZgcdM.
- generalize (Zgcd_is_pos z1 z2).
- generalize (Zmax_spec (Zgcd z1 z2) 1).
- generalize (Zgcd_is_pos (Zmax (Zgcd z1 z2) 1) z).
- generalize (Zmax_spec (Zgcd (Zmax (Zgcd z1 z2) 1) z) 1).
+ generalize (Z.gcd_nonneg z1 z2).
+ generalize (Zmax_spec (Z.gcd z1 z2) 1).
+ generalize (Z.gcd_nonneg (Z.max (Z.gcd z1 z2) 1) z).
+ generalize (Zmax_spec (Z.gcd (Z.max (Z.gcd z1 z2) 1) z) 1).
auto with zarith.
Qed.
@@ -433,7 +424,7 @@ Proof.
intros.
induction H.
constructor.
- apply Zdivide_trans with (1:= H0) ; assumption.
+ apply Z.divide_trans with (1:= H0) ; assumption.
constructor. auto.
constructor ; auto.
Qed.
@@ -444,20 +435,20 @@ Proof.
exists c. ring.
Qed.
-Lemma Zgcd_minus : forall a b c, (a | c - b ) -> (Zgcd a b | c).
+Lemma Zgcd_minus : forall a b c, (a | c - b ) -> (Z.gcd a b | c).
Proof.
intros a b c (q,Hq).
destruct (Zgcd_is_gcd a b) as [(a',Ha) (b',Hb) _].
- set (g:=Zgcd a b) in *; clearbody g.
+ set (g:=Z.gcd a b) in *; clearbody g.
exists (q * a' + b').
- symmetry in Hq. rewrite <- Zeq_plus_swap in Hq.
+ symmetry in Hq. rewrite <- Z.add_move_r in Hq.
rewrite <- Hq, Hb, Ha. ring.
Qed.
Lemma Zdivide_pol_sub : forall p a b,
- 0 < Zgcd a b ->
- Zdivide_pol a (PsubC Zminus p b) ->
- Zdivide_pol (Zgcd a b) p.
+ 0 < Z.gcd a b ->
+ Zdivide_pol a (PsubC Z.sub p b) ->
+ Zdivide_pol (Z.gcd a b) p.
Proof.
induction p.
simpl.
@@ -477,7 +468,7 @@ Proof.
Qed.
Lemma Zdivide_pol_sub_0 : forall p a,
- Zdivide_pol a (PsubC Zminus p 0) ->
+ Zdivide_pol a (PsubC Z.sub p 0) ->
Zdivide_pol a p.
Proof.
induction p.
@@ -496,7 +487,7 @@ Qed.
Lemma Zgcd_pol_div : forall p g c,
- Zgcd_pol p = (g, c) -> Zdivide_pol g (PsubC Zminus p c).
+ Zgcd_pol p = (g, c) -> Zdivide_pol g (PsubC Z.sub p c).
Proof.
induction p ; simpl.
(* Pc *)
@@ -511,12 +502,12 @@ Proof.
case_eq (Zgcd_pol p1) ; case_eq (Zgcd_pol p3) ; intros.
inv H1.
unfold ZgcdM at 1.
- destruct (Zmax_spec (Zgcd (ZgcdM z1 z2) z) 1) as [HH1 | HH1];
+ destruct (Zmax_spec (Z.gcd (ZgcdM z1 z2) z) 1) as [HH1 | HH1];
destruct HH1 as [HH1 HH1'] ; rewrite HH1'.
constructor.
apply Zdivide_pol_Zdivide with (x:= ZgcdM z1 z2).
unfold ZgcdM.
- destruct (Zmax_spec (Zgcd z1 z2) 1) as [HH2 | HH2].
+ destruct (Zmax_spec (Z.gcd z1 z2) 1) as [HH2 | HH2].
destruct HH2.
rewrite H2.
apply Zdivide_pol_sub ; auto.
@@ -524,9 +515,9 @@ Proof.
destruct HH2. rewrite H2.
apply Zdivide_pol_one.
unfold ZgcdM in HH1. unfold ZgcdM.
- destruct (Zmax_spec (Zgcd z1 z2) 1) as [HH2 | HH2].
+ destruct (Zmax_spec (Z.gcd z1 z2) 1) as [HH2 | HH2].
destruct HH2. rewrite H2 in *.
- destruct (Zgcd_is_gcd (Zgcd z1 z2) z); auto.
+ destruct (Zgcd_is_gcd (Z.gcd z1 z2) z); auto.
destruct HH2. rewrite H2.
destruct (Zgcd_is_gcd 1 z); auto.
apply Zdivide_pol_Zdivide with (x:= z).
@@ -539,7 +530,7 @@ Qed.
-Lemma Zgcd_pol_correct_lt : forall p env g c, Zgcd_pol p = (g,c) -> 0 < g -> eval_pol env p = g * (eval_pol env (Zdiv_pol (PsubC Zminus p c) g)) + c.
+Lemma Zgcd_pol_correct_lt : forall p env g c, Zgcd_pol p = (g,c) -> 0 < g -> eval_pol env p = g * (eval_pol env (Zdiv_pol (PsubC Z.sub p c) g)) + c.
Proof.
intros.
rewrite <- Zdiv_pol_correct ; auto.
@@ -553,8 +544,8 @@ Qed.
Definition makeCuttingPlane (p : PolC Z) : PolC Z * Z :=
let (g,c) := Zgcd_pol p in
- if Zgt_bool g Z0
- then (Zdiv_pol (PsubC Zminus p c) g , Zopp (ceiling (Zopp c) g))
+ if Z.gtb g Z0
+ then (Zdiv_pol (PsubC Z.sub p c) g , Z.opp (ceiling (Z.opp c) g))
else (p,Z0).
@@ -562,13 +553,13 @@ Definition genCuttingPlane (f : NFormula Z) : option (PolC Z * Z * Op1) :=
let (e,op) := f in
match op with
| Equal => let (g,c) := Zgcd_pol e in
- if andb (Zgt_bool g Z0) (andb (negb (Zeq_bool c Z0)) (negb (Zeq_bool (Zgcd g c) g)))
+ if andb (Z.gtb g Z0) (andb (negb (Zeq_bool c Z0)) (negb (Zeq_bool (Z.gcd g c) g)))
then None (* inconsistent *)
else (* Could be optimised Zgcd_pol is recomputed *)
let (p,c) := makeCuttingPlane e in
Some (p,c,Equal)
| NonEqual => Some (e,Z0,op)
- | Strict => let (p,c) := makeCuttingPlane (PsubC Zminus e 1) in
+ | Strict => let (p,c) := makeCuttingPlane (PsubC Z.sub e 1) in
Some (p,c,NonStrict)
| NonStrict => let (p,c) := makeCuttingPlane e in
Some (p,c,NonStrict)
@@ -595,7 +586,7 @@ Qed.
Definition eval_Psatz : list (NFormula Z) -> ZWitness -> option (NFormula Z) :=
- eval_Psatz 0 1 Zplus Zmult Zeq_bool Zle_bool.
+ eval_Psatz 0 1 Z.add Z.mul Zeq_bool Z.leb.
Definition valid_cut_sign (op:Op1) :=
@@ -634,9 +625,9 @@ Fixpoint ZChecker (l:list (NFormula Z)) (pf : ZArithProof) {struct pf} : bool :
(fix label (pfs:list ZArithProof) :=
fun lb ub =>
match pfs with
- | nil => if Zgt_bool lb ub then true else false
- | pf::rsr => andb (ZChecker ((psub e1 (Pc lb), Equal) :: l) pf) (label rsr (Zplus lb 1%Z) ub)
- end) pf (Zopp z1) z2
+ | nil => if Z.gtb lb ub then true else false
+ | pf::rsr => andb (ZChecker ((psub e1 (Pc lb), Equal) :: l) pf) (label rsr (Z.add lb 1%Z) ub)
+ end) pf (Z.opp z1) z2
else false
| _ , _ => true
end
@@ -710,12 +701,12 @@ Proof.
unfold makeCuttingPlane in H0.
revert H0.
case_eq (Zgcd_pol e) ; intros g c0.
- generalize (Zgt_cases g 0) ; destruct (Zgt_bool g 0).
+ generalize (Zgt_cases g 0) ; destruct (Z.gtb g 0).
intros.
inv H2.
- change (RingMicromega.eval_pol 0 Zplus Zmult (fun x : Z => x)) with eval_pol in *.
+ change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in *.
apply Zgcd_pol_correct_lt with (env:=env) in H1.
- generalize (narrow_interval_lower_bound g (- c0) (eval_pol env (Zdiv_pol (PsubC Zminus e c0) g)) H0).
+ generalize (narrow_interval_lower_bound g (- c0) (eval_pol env (Zdiv_pol (PsubC Z.sub e c0) g)) H0).
auto with zarith.
auto with zarith.
(* g <= 0 *)
@@ -733,7 +724,7 @@ Proof.
(* Equal *)
destruct p as [[e' z] op].
case_eq (Zgcd_pol e) ; intros g c.
- case_eq (Zgt_bool g 0 && (negb (Zeq_bool c 0) && negb (Zeq_bool (Zgcd g c) g))) ; [discriminate|].
+ case_eq (Z.gtb g 0 && (negb (Zeq_bool c 0) && negb (Zeq_bool (Z.gcd g c) g))) ; [discriminate|].
case_eq (makeCuttingPlane e).
intros.
inv H3.
@@ -741,7 +732,7 @@ Proof.
rewrite H1 in H.
revert H.
change (eval_pol env e = 0) in H2.
- case_eq (Zgt_bool g 0).
+ case_eq (Z.gtb g 0).
intros.
rewrite <- Zgt_is_gt_bool in H.
rewrite Zgcd_pol_correct_lt with (1:= H1) in H2; auto with zarith.
@@ -749,7 +740,7 @@ Proof.
change (eval_pol env (padd e' (Pc z)) = 0).
inv H3.
rewrite eval_pol_add.
- set (x:=eval_pol env (Zdiv_pol (PsubC Zminus e c) g)) in *; clearbody x.
+ set (x:=eval_pol env (Zdiv_pol (PsubC Z.sub e c) g)) in *; clearbody x.
simpl.
rewrite andb_false_iff in H0.
destruct H0.
@@ -759,8 +750,7 @@ Proof.
rewrite negb_false_iff in H0.
apply Zeq_bool_eq in H0.
subst. simpl.
- rewrite Zplus_0_r in H2.
- apply Zmult_integral in H2.
+ rewrite Z.add_0_r, Z.mul_eq_0 in H2.
intuition auto with zarith.
rewrite negb_false_iff in H0.
apply Zeq_bool_eq in H0.
@@ -769,7 +759,7 @@ Proof.
inv HH.
apply Zdivide_opp_r in H4.
rewrite Zdivide_ceiling ; auto.
- apply Zeq_minus.
+ apply Z.sub_move_0_r.
apply Z.div_unique_exact ; auto with zarith.
intros.
unfold nformula_of_cutting_plane.
@@ -789,7 +779,7 @@ Proof.
simpl. auto with zarith.
(* Strict *)
destruct p as [[e' z] op].
- case_eq (makeCuttingPlane (PsubC Zminus e 1)).
+ case_eq (makeCuttingPlane (PsubC Z.sub e 1)).
intros.
inv H1.
apply makeCuttingPlane_ns_sound with (env:=env) (2:= H).
@@ -813,7 +803,7 @@ Proof.
destruct f.
destruct o.
case_eq (Zgcd_pol p) ; intros g c.
- case_eq (Zgt_bool g 0 && (negb (Zeq_bool c 0) && negb (Zeq_bool (Zgcd g c) g))).
+ case_eq (Z.gtb g 0 && (negb (Zeq_bool c 0) && negb (Zeq_bool (Z.gcd g c) g))).
intros.
flatten_bool.
rewrite negb_true_iff in H5.
@@ -823,16 +813,16 @@ Proof.
apply Zeq_bool_neq in H.
change (eval_pol env p = 0) in H2.
rewrite Zgcd_pol_correct_lt with (1:= H0) in H2; auto with zarith.
- set (x:=eval_pol env (Zdiv_pol (PsubC Zminus p c) g)) in *; clearbody x.
+ set (x:=eval_pol env (Zdiv_pol (PsubC Z.sub p c) g)) in *; clearbody x.
contradict H5.
apply Zis_gcd_gcd; auto with zarith.
constructor; auto with zarith.
exists (-x).
- rewrite <- Zopp_mult_distr_l, Zmult_comm; auto with zarith.
+ rewrite Z.mul_opp_l, Z.mul_comm; auto with zarith.
(**)
destruct (makeCuttingPlane p); discriminate.
discriminate.
- destruct (makeCuttingPlane (PsubC Zminus p 1)) ; discriminate.
+ destruct (makeCuttingPlane (PsubC Z.sub p 1)) ; discriminate.
destruct (makeCuttingPlane p) ; discriminate.
Qed.
@@ -920,7 +910,7 @@ Proof.
unfold nformula_of_cutting_plane in HCutR.
unfold eval_nformula in HCutR.
unfold RingMicromega.eval_nformula in HCutR.
- change (RingMicromega.eval_pol 0 Zplus Zmult (fun x : Z => x)) with eval_pol in HCutR.
+ change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in HCutR.
unfold eval_op1 in HCutR.
destruct op1 ; simpl in Hop1 ; try discriminate;
rewrite eval_pol_add in HCutR; simpl in HCutR; auto with zarith.
@@ -933,7 +923,7 @@ Proof.
unfold nformula_of_cutting_plane in HCutL.
unfold eval_nformula in HCutL.
unfold RingMicromega.eval_nformula in HCutL.
- change (RingMicromega.eval_pol 0 Zplus Zmult (fun x : Z => x)) with eval_pol in HCutL.
+ change (RingMicromega.eval_pol Z.add Z.mul (fun x : Z => x)) with eval_pol in HCutL.
unfold eval_op1 in HCutL.
rewrite eval_pol_add in HCutL. simpl in HCutL.
destruct op2 ; simpl in Hop2 ; try discriminate ; omega.
@@ -944,14 +934,14 @@ Proof.
intros.
assert (HH :forall x, -z1 <= x <= z2 -> exists pr,
(In pr pf /\
- ZChecker ((PsubC Zminus p1 x,Equal) :: l) pr = true)%Z).
+ ZChecker ((PsubC Z.sub p1 x,Equal) :: l) pr = true)%Z).
clear HZ0 Hop1 Hop2 HCutL HCutR H0 H1.
revert Hfix.
generalize (-z1). clear z1. intro z1.
revert z1 z2.
induction pf;simpl ;intros.
generalize (Zgt_cases z1 z2).
- destruct (Zgt_bool z1 z2).
+ destruct (Z.gtb z1 z2).
intros.
apply False_ind ; omega.
discriminate.
@@ -972,7 +962,7 @@ Proof.
zify. omega.
(*/asser *)
destruct (HH _ H1) as [pr [Hin Hcheker]].
- assert (make_impl (eval_nformula env) ((PsubC Zminus p1 (eval_pol env p1),Equal) :: l) False).
+ assert (make_impl (eval_nformula env) ((PsubC Z.sub p1 (eval_pol env p1),Equal) :: l) False).
apply (H pr);auto.
apply in_bdepth ; auto.
rewrite <- make_conj_impl in H2.
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index 68dc0319f..48b72f34b 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -301,6 +301,8 @@ struct
["Coq";"Reals" ; "Rpow_def"] ;
]
+ let z_modules = [["Coq";"ZArith";"BinInt"]]
+
(**
* Initialization : a large amount of Caml symbols are derived from
* ZMicromega.v
@@ -310,6 +312,7 @@ struct
let constant = gen_constant_in_modules "ZMicromega" coq_modules
let bin_constant = gen_constant_in_modules "ZMicromega" bin_module
let r_constant = gen_constant_in_modules "ZMicromega" r_modules
+ let z_constant = gen_constant_in_modules "ZMicromega" z_modules
(* let constant = gen_constant_in_modules "Omicron" coq_modules *)
let coq_and = lazy (init_constant "and")
@@ -372,17 +375,17 @@ struct
let coq_cutProof = lazy (constant "CutProof")
let coq_enumProof = lazy (constant "EnumProof")
- let coq_Zgt = lazy (constant "Zgt")
- let coq_Zge = lazy (constant "Zge")
- let coq_Zle = lazy (constant "Zle")
- let coq_Zlt = lazy (constant "Zlt")
+ let coq_Zgt = lazy (z_constant "Z.gt")
+ let coq_Zge = lazy (z_constant "Z.ge")
+ let coq_Zle = lazy (z_constant "Z.le")
+ let coq_Zlt = lazy (z_constant "Z.lt")
let coq_Eq = lazy (init_constant "eq")
- let coq_Zplus = lazy (constant "Zplus")
- let coq_Zminus = lazy (constant "Zminus")
- let coq_Zopp = lazy (constant "Zopp")
- let coq_Zmult = lazy (constant "Zmult")
- let coq_Zpower = lazy (constant "Zpower")
+ let coq_Zplus = lazy (z_constant "Z.add")
+ let coq_Zminus = lazy (z_constant "Z.sub")
+ let coq_Zopp = lazy (z_constant "Z.opp")
+ let coq_Zmult = lazy (z_constant "Z.mul")
+ let coq_Zpower = lazy (z_constant "Z.pow")
let coq_Qgt = lazy (constant "Qgt")
let coq_Qge = lazy (constant "Qge")
diff --git a/plugins/nsatz/Nsatz.v b/plugins/nsatz/Nsatz.v
index 9a0c9090f..25255dd0d 100644
--- a/plugins/nsatz/Nsatz.v
+++ b/plugins/nsatz/Nsatz.v
@@ -64,15 +64,15 @@ Definition PEZ := PExpr Z.
Definition P0Z : PolZ := P0 (C:=Z) 0%Z.
Definition PolZadd : PolZ -> PolZ -> PolZ :=
- @Padd Z 0%Z Zplus Zeq_bool.
+ @Padd Z 0%Z Z.add Zeq_bool.
Definition PolZmul : PolZ -> PolZ -> PolZ :=
- @Pmul Z 0%Z 1%Z Zplus Zmult Zeq_bool.
+ @Pmul Z 0%Z 1%Z Z.add Z.mul Zeq_bool.
Definition PolZeq := @Peq Z Zeq_bool.
Definition norm :=
- @norm_aux Z 0%Z 1%Z Zplus Zmult Zminus Zopp Zeq_bool.
+ @norm_aux Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool.
Fixpoint mult_l (la : list PEZ) (lp: list PolZ) : PolZ :=
match la, lp with
@@ -100,16 +100,16 @@ Definition PhiR : list R -> PolZ -> R :=
Definition PEevalR : list R -> PEZ -> R :=
PEeval ring0 add mul sub opp
(gen_phiZ ring0 ring1 add mul opp)
- nat_of_N pow.
+ N.to_nat pow.
Lemma P0Z_correct : forall l, PhiR l P0Z = 0.
Proof. trivial. Qed.
Lemma Rext: ring_eq_ext add mul opp _==_.
-apply mk_reqe. intros. rewrite H ; rewrite H0; cring.
- intros. rewrite H; rewrite H0; cring.
-intros. rewrite H; cring. Qed.
-
+Proof.
+constructor; solve_proper.
+Qed.
+
Lemma Rset : Setoid_Theory R _==_.
apply ring_setoid.
Qed.
@@ -144,8 +144,8 @@ unfold PolZmul, PhiR. intros.
Qed.
Lemma R_power_theory
- : Ring_theory.power_theory ring1 mul _==_ nat_of_N pow.
-apply Ring_theory.mkpow_th. unfold pow. intros. rewrite Nnat.N_of_nat_of_N.
+ : Ring_theory.power_theory ring1 mul _==_ N.to_nat pow.
+apply Ring_theory.mkpow_th. unfold pow. intros. rewrite Nnat.N2Nat.id.
reflexivity. Qed.
Lemma norm_correct :
@@ -241,9 +241,9 @@ Fixpoint interpret3 t fv {struct t}: R :=
| (PEopp t1) =>
let v1 := interpret3 t1 fv in (-v1)
| (PEpow t1 t2) =>
- let v1 := interpret3 t1 fv in pow v1 (nat_of_N t2)
+ let v1 := interpret3 t1 fv in pow v1 (N.to_nat t2)
| (PEc t1) => (IZR1 t1)
- | (PEX n) => List.nth (pred (nat_of_P n)) fv 0
+ | (PEX n) => List.nth (pred (Pos.to_nat n)) fv 0
end.
@@ -308,9 +308,9 @@ Ltac nsatz_call radicalmax info nparam p lp kont :=
lazymatch n with
| 0%N => fail
| _ =>
- (let r := eval compute in (Nminus radicalmax (Npred n)) in
+ (let r := eval compute in (N.sub radicalmax (N.pred n)) in
nsatz_call_n info nparam p r lp kont) ||
- let n' := eval compute in (Npred n) in try_n n'
+ let n' := eval compute in (N.pred n) in try_n n'
end in
try_n radicalmax.
@@ -343,7 +343,7 @@ Ltac get_lpol g :=
end.
Ltac nsatz_generic radicalmax info lparam lvar :=
- let nparam := eval compute in (Z_of_nat (List.length lparam)) in
+ let nparam := eval compute in (Z.of_nat (List.length lparam)) in
match goal with
|- ?g => let lb := lterm_goal g in
match (match lvar with
@@ -397,7 +397,7 @@ Ltac nsatz_generic radicalmax info lparam lvar :=
(*simpl*) idtac;
repeat (split;[assumption|idtac]); exact I
| (*simpl in Hg2;*) (*simpl*) idtac;
- apply Rintegral_domain_pow with (interpret3 c fv) (nat_of_N r);
+ apply Rintegral_domain_pow with (interpret3 c fv) (N.to_nat r);
(*simpl*) idtac;
try apply integral_domain_one_zero;
try apply integral_domain_minus_one_zero;
@@ -502,7 +502,7 @@ omega.
Qed.
Instance Zcri: (Cring (Rr:=Zr)).
-red. exact Zmult_comm. Defined.
+red. exact Z.mul_comm. Defined.
Instance Zdi : (Integral_domain (Rcr:=Zcri)).
constructor.
diff --git a/plugins/omega/Omega.v b/plugins/omega/Omega.v
index 3f9d0f448..6cdfea43f 100644
--- a/plugins/omega/Omega.v
+++ b/plugins/omega/Omega.v
@@ -19,9 +19,9 @@ Require Export OmegaLemmas.
Require Export PreOmega.
Declare ML Module "omega_plugin".
-Hint Resolve Zle_refl Zplus_comm Zplus_assoc Zmult_comm Zmult_assoc Zplus_0_l
- Zplus_0_r Zmult_1_l Zplus_opp_l Zplus_opp_r Zmult_plus_distr_l
- Zmult_plus_distr_r: zarith.
+Hint Resolve Z.le_refl Z.add_comm Z.add_assoc Z.mul_comm Z.mul_assoc Z.add_0_l
+ Z.add_0_r Z.mul_1_l Z.add_opp_diag_l Z.add_opp_diag_r Z.mul_add_distr_r
+ Z.mul_add_distr_l: zarith.
Require Export Zhints.
diff --git a/plugins/omega/OmegaLemmas.v b/plugins/omega/OmegaLemmas.v
index 5b6f4670f..1872f5766 100644
--- a/plugins/omega/OmegaLemmas.v
+++ b/plugins/omega/OmegaLemmas.v
@@ -6,232 +6,192 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-Require Import ZArith_base.
-Open Local Scope Z_scope.
+Require Import BinInt Znat.
+Local Open Scope Z_scope.
(** Factorization lemmas *)
-Theorem Zred_factor0 : forall n:Z, n = n * 1.
- intro x; rewrite (Zmult_1_r x); reflexivity.
+Theorem Zred_factor0 n : n = n * 1.
+Proof.
+ now Z.nzsimpl.
Qed.
-Theorem Zred_factor1 : forall n:Z, n + n = n * 2.
+Theorem Zred_factor1 n : n + n = n * 2.
Proof.
- exact Zplus_diag_eq_mult_2.
+ rewrite Z.mul_comm. apply Z.add_diag.
Qed.
-Theorem Zred_factor2 : forall n m:Z, n + n * m = n * (1 + m).
+Theorem Zred_factor2 n m : n + n * m = n * (1 + m).
Proof.
- intros x y; pattern x at 1 in |- *; rewrite <- (Zmult_1_r x);
- rewrite <- Zmult_plus_distr_r; trivial with arith.
+ rewrite Z.mul_add_distr_l; now Z.nzsimpl.
Qed.
-Theorem Zred_factor3 : forall n m:Z, n * m + n = n * (1 + m).
+Theorem Zred_factor3 n m : n * m + n = n * (1 + m).
Proof.
- intros x y; pattern x at 2 in |- *; rewrite <- (Zmult_1_r x);
- rewrite <- Zmult_plus_distr_r; rewrite Zplus_comm;
- trivial with arith.
+ now Z.nzsimpl.
Qed.
-Theorem Zred_factor4 : forall n m p:Z, n * m + n * p = n * (m + p).
+Theorem Zred_factor4 n m p : n * m + n * p = n * (m + p).
Proof.
- intros x y z; symmetry in |- *; apply Zmult_plus_distr_r.
+ symmetry; apply Z.mul_add_distr_l.
Qed.
-Theorem Zred_factor5 : forall n m:Z, n * 0 + m = m.
+Theorem Zred_factor5 n m : n * 0 + m = m.
Proof.
- intros x y; rewrite <- Zmult_0_r_reverse; auto with arith.
+ now Z.nzsimpl.
Qed.
-Theorem Zred_factor6 : forall n:Z, n = n + 0.
+Theorem Zred_factor6 n : n = n + 0.
Proof.
- intro; rewrite Zplus_0_r; trivial with arith.
+ now Z.nzsimpl.
Qed.
(** Other specific variants of theorems dedicated for the Omega tactic *)
Lemma new_var : forall x : Z, exists y : Z, x = y.
-intros x; exists x; trivial with arith.
+Proof.
+intros x; now exists x.
Qed.
-Lemma OMEGA1 : forall x y : Z, x = y -> 0 <= x -> 0 <= y.
-intros x y H; rewrite H; auto with arith.
+Lemma OMEGA1 x y : x = y -> 0 <= x -> 0 <= y.
+Proof.
+now intros ->.
Qed.
-Lemma OMEGA2 : forall x y : Z, 0 <= x -> 0 <= y -> 0 <= x + y.
-exact Zplus_le_0_compat.
+Lemma OMEGA2 x y : 0 <= x -> 0 <= y -> 0 <= x + y.
+Proof.
+Z.order_pos.
Qed.
-Lemma OMEGA3 : forall x y k : Z, k > 0 -> x = y * k -> x = 0 -> y = 0.
-
-intros x y k H1 H2 H3; apply (Zmult_integral_l k);
- [ unfold not in |- *; intros H4; absurd (k > 0);
- [ rewrite H4; unfold Zgt in |- *; simpl in |- *; discriminate
- | assumption ]
- | rewrite <- H2; assumption ].
+Lemma OMEGA3 x y k : k > 0 -> x = y * k -> x = 0 -> y = 0.
+Proof.
+intros LT -> EQ. apply Z.mul_eq_0 in EQ. destruct EQ; now subst.
Qed.
-Lemma OMEGA4 : forall x y z : Z, x > 0 -> y > x -> z * y + x <> 0.
-
-unfold not in |- *; intros x y z H1 H2 H3; cut (y > 0);
- [ intros H4; cut (0 <= z * y + x);
- [ intros H5; generalize (Zmult_le_approx y z x H4 H2 H5); intros H6;
- absurd (z * y + x > 0);
- [ rewrite H3; unfold Zgt in |- *; simpl in |- *; discriminate
- | apply Zle_gt_trans with x;
- [ pattern x at 1 in |- *; rewrite <- (Zplus_0_l x);
- apply Zplus_le_compat_r; rewrite Zmult_comm;
- generalize H4; unfold Zgt in |- *; case y;
- [ simpl in |- *; intros H7; discriminate H7
- | intros p H7; rewrite <- (Zmult_0_r (Zpos p));
- unfold Zle in |- *; rewrite Zcompare_mult_compat;
- exact H6
- | simpl in |- *; intros p H7; discriminate H7 ]
- | assumption ] ]
- | rewrite H3; unfold Zle in |- *; simpl in |- *; discriminate ]
- | apply Zgt_trans with x; [ assumption | assumption ] ].
+Lemma OMEGA4 x y z : x > 0 -> y > x -> z * y + x <> 0.
+Proof.
+Z.swap_greater. intros Hx Hxy.
+rewrite Z.add_move_0_l, <- Z.mul_opp_l.
+destruct (Z.lt_trichotomy (-z) 1) as [LT|[->|GT]].
+- intro. revert LT. apply Z.le_ngt, (Z.le_succ_l 0).
+ apply Z.mul_pos_cancel_r with y; Z.order.
+- Z.nzsimpl. Z.order.
+- rewrite (Z.mul_lt_mono_pos_r y), Z.mul_1_l in GT; Z.order.
Qed.
-Lemma OMEGA5 : forall x y z : Z, x = 0 -> y = 0 -> x + y * z = 0.
-
-intros x y z H1 H2; rewrite H1; rewrite H2; simpl in |- *; trivial with arith.
+Lemma OMEGA5 x y z : x = 0 -> y = 0 -> x + y * z = 0.
+Proof.
+now intros -> ->.
Qed.
-Lemma OMEGA6 : forall x y z : Z, 0 <= x -> y = 0 -> 0 <= x + y * z.
-
-intros x y z H1 H2; rewrite H2; simpl in |- *; rewrite Zplus_0_r; assumption.
+Lemma OMEGA6 x y z : 0 <= x -> y = 0 -> 0 <= x + y * z.
+Proof.
+intros H ->. now Z.nzsimpl.
Qed.
-Lemma OMEGA7 :
- forall x y z t : Z, z > 0 -> t > 0 -> 0 <= x -> 0 <= y -> 0 <= x * z + y * t.
-
-intros x y z t H1 H2 H3 H4; rewrite <- (Zplus_0_l 0); apply Zplus_le_compat;
- apply Zmult_gt_0_le_0_compat; assumption.
+Lemma OMEGA7 x y z t :
+ z > 0 -> t > 0 -> 0 <= x -> 0 <= y -> 0 <= x * z + y * t.
+Proof.
+intros. Z.swap_greater. Z.order_pos.
Qed.
-Lemma OMEGA8 : forall x y : Z, 0 <= x -> 0 <= y -> x = - y -> x = 0.
-
-intros x y H1 H2 H3; elim (Zle_lt_or_eq 0 x H1);
- [ intros H4; absurd (0 < x);
- [ change (0 >= x) in |- *; apply Zle_ge; apply Zplus_le_reg_l with y;
- rewrite H3; rewrite Zplus_opp_r; rewrite Zplus_0_r;
- assumption
- | assumption ]
- | intros H4; rewrite H4; trivial with arith ].
+Lemma OMEGA8 x y : 0 <= x -> 0 <= y -> x = - y -> x = 0.
+Proof.
+intros H1 H2 H3. rewrite <- Z.opp_nonpos_nonneg in H2. Z.order.
Qed.
-Lemma OMEGA9 : forall x y z t : Z, y = 0 -> x = z -> y + (- x + z) * t = 0.
-
-intros x y z t H1 H2; rewrite H2; rewrite Zplus_opp_l; rewrite Zmult_0_l;
- rewrite Zplus_0_r; assumption.
+Lemma OMEGA9 x y z t : y = 0 -> x = z -> y + (- x + z) * t = 0.
+Proof.
+intros. subst. now rewrite Z.add_opp_diag_l.
Qed.
-Lemma OMEGA10 :
- forall v c1 c2 l1 l2 k1 k2 : Z,
+Lemma OMEGA10 v c1 c2 l1 l2 k1 k2 :
(v * c1 + l1) * k1 + (v * c2 + l2) * k2 =
v * (c1 * k1 + c2 * k2) + (l1 * k1 + l2 * k2).
-
-intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
- repeat rewrite Zmult_assoc; repeat elim Zplus_assoc;
- rewrite (Zplus_permute (l1 * k1) (v * c2 * k2)); trivial with arith.
+Proof.
+rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc.
+rewrite <- !Z.add_assoc. f_equal. apply Z.add_shuffle3.
Qed.
-Lemma OMEGA11 :
- forall v1 c1 l1 l2 k1 : Z,
+Lemma OMEGA11 v1 c1 l1 l2 k1 :
(v1 * c1 + l1) * k1 + l2 = v1 * (c1 * k1) + (l1 * k1 + l2).
-
-intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
- repeat rewrite Zmult_assoc; repeat elim Zplus_assoc;
- trivial with arith.
+Proof.
+rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc.
+now rewrite Z.add_assoc.
Qed.
-Lemma OMEGA12 :
- forall v2 c2 l1 l2 k2 : Z,
+Lemma OMEGA12 v2 c2 l1 l2 k2 :
l1 + (v2 * c2 + l2) * k2 = v2 * (c2 * k2) + (l1 + l2 * k2).
-
-intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
- repeat rewrite Zmult_assoc; repeat elim Zplus_assoc;
- rewrite Zplus_permute; trivial with arith.
+Proof.
+rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc.
+apply Z.add_shuffle3.
Qed.
-Lemma OMEGA13 :
- forall (v l1 l2 : Z) (x : positive),
+Lemma OMEGA13 (v l1 l2 : Z) (x : positive) :
v * Zpos x + l1 + (v * Zneg x + l2) = l1 + l2.
-
-intros; rewrite Zplus_assoc; rewrite (Zplus_comm (v * Zpos x) l1);
- rewrite (Zplus_assoc_reverse l1); rewrite <- Zmult_plus_distr_r;
- rewrite <- Zopp_neg; rewrite (Zplus_comm (- Zneg x) (Zneg x));
- rewrite Zplus_opp_r; rewrite Zmult_0_r; rewrite Zplus_0_r;
- trivial with arith.
+Proof.
+ rewrite Z.add_shuffle1.
+ rewrite <- Z.mul_add_distr_l, <- Pos2Z.opp_neg, Z.add_opp_diag_r.
+ now Z.nzsimpl.
Qed.
-Lemma OMEGA14 :
- forall (v l1 l2 : Z) (x : positive),
+Lemma OMEGA14 (v l1 l2 : Z) (x : positive) :
v * Zneg x + l1 + (v * Zpos x + l2) = l1 + l2.
-
-intros; rewrite Zplus_assoc; rewrite (Zplus_comm (v * Zneg x) l1);
- rewrite (Zplus_assoc_reverse l1); rewrite <- Zmult_plus_distr_r;
- rewrite <- Zopp_neg; rewrite Zplus_opp_r; rewrite Zmult_0_r;
- rewrite Zplus_0_r; trivial with arith.
+Proof.
+ rewrite Z.add_shuffle1.
+ rewrite <- Z.mul_add_distr_l, <- Pos2Z.opp_neg, Z.add_opp_diag_r.
+ now Z.nzsimpl.
Qed.
-Lemma OMEGA15 :
- forall v c1 c2 l1 l2 k2 : Z,
- v * c1 + l1 + (v * c2 + l2) * k2 = v * (c1 + c2 * k2) + (l1 + l2 * k2).
-intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
- repeat rewrite Zmult_assoc; repeat elim Zplus_assoc;
- rewrite (Zplus_permute l1 (v * c2 * k2)); trivial with arith.
+Lemma OMEGA15 v c1 c2 l1 l2 k2 :
+ v * c1 + l1 + (v * c2 + l2) * k2 = v * (c1 + c2 * k2) + (l1 + l2 * k2).
+Proof.
+ rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc.
+ apply Z.add_shuffle1.
Qed.
-Lemma OMEGA16 : forall v c l k : Z, (v * c + l) * k = v * (c * k) + l * k.
-
-intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
- repeat rewrite Zmult_assoc; repeat elim Zplus_assoc;
- trivial with arith.
+Lemma OMEGA16 v c l k : (v * c + l) * k = v * (c * k) + l * k.
+Proof.
+ now rewrite ?Z.mul_add_distr_r, ?Z.mul_add_distr_l, ?Z.mul_assoc.
Qed.
-Lemma OMEGA17 : forall x y z : Z, Zne x 0 -> y = 0 -> Zne (x + y * z) 0.
-
-unfold Zne, not in |- *; intros x y z H1 H2 H3; apply H1;
- apply Zplus_reg_l with (y * z); rewrite Zplus_comm;
- rewrite H3; rewrite H2; auto with arith.
+Lemma OMEGA17 x y z : Zne x 0 -> y = 0 -> Zne (x + y * z) 0.
+Proof.
+ unfold Zne, not. intros NE EQ. subst. now Z.nzsimpl.
Qed.
-Lemma OMEGA18 : forall x y k : Z, x = y * k -> Zne x 0 -> Zne y 0.
-
-unfold Zne, not in |- *; intros x y k H1 H2 H3; apply H2; rewrite H1;
- rewrite H3; auto with arith.
+Lemma OMEGA18 x y k : x = y * k -> Zne x 0 -> Zne y 0.
+Proof.
+ unfold Zne, not. intros. subst; auto.
Qed.
-Lemma OMEGA19 : forall x : Z, Zne x 0 -> 0 <= x + -1 \/ 0 <= x * -1 + -1.
-
-unfold Zne in |- *; intros x H; elim (Zle_or_lt 0 x);
- [ intros H1; elim Zle_lt_or_eq with (1 := H1);
- [ intros H2; left; change (0 <= Zpred x) in |- *; apply Zsucc_le_reg;
- rewrite <- Zsucc_pred; apply Zlt_le_succ; assumption
- | intros H2; absurd (x = 0); auto with arith ]
- | intros H1; right; rewrite <- Zopp_eq_mult_neg_1; rewrite Zplus_comm;
- apply Zle_left; apply Zsucc_le_reg; simpl in |- *;
- apply Zlt_le_succ; auto with arith ].
+Lemma OMEGA19 x : Zne x 0 -> 0 <= x + -1 \/ 0 <= x * -1 + -1.
+Proof.
+ unfold Zne. intros Hx. apply Z.lt_gt_cases in Hx.
+ destruct Hx as [LT|GT].
+ - right. change (-1) with (-(1)).
+ rewrite Z.mul_opp_r, <- Z.opp_add_distr. Z.nzsimpl.
+ rewrite Z.opp_nonneg_nonpos. now apply Z.le_succ_l.
+ - left. now apply Z.lt_le_pred.
Qed.
-Lemma OMEGA20 : forall x y z : Z, Zne x 0 -> y = 0 -> Zne (x + y * z) 0.
-
-unfold Zne, not in |- *; intros x y z H1 H2 H3; apply H1; rewrite H2 in H3;
- simpl in H3; rewrite Zplus_0_r in H3; trivial with arith.
+Lemma OMEGA20 x y z : Zne x 0 -> y = 0 -> Zne (x + y * z) 0.
+Proof.
+ unfold Zne, not. intros H1 H2 H3; apply H1; rewrite H2 in H3;
+ simpl in H3; rewrite Z.add_0_r in H3; trivial with arith.
Qed.
Definition fast_Zplus_comm (x y : Z) (P : Z -> Prop)
- (H : P (y + x)) := eq_ind_r P H (Zplus_comm x y).
+ (H : P (y + x)) := eq_ind_r P H (Z.add_comm x y).
Definition fast_Zplus_assoc_reverse (n m p : Z) (P : Z -> Prop)
(H : P (n + (m + p))) := eq_ind_r P H (Zplus_assoc_reverse n m p).
Definition fast_Zplus_assoc (n m p : Z) (P : Z -> Prop)
- (H : P (n + m + p)) := eq_ind_r P H (Zplus_assoc n m p).
+ (H : P (n + m + p)) := eq_ind_r P H (Z.add_assoc n m p).
Definition fast_Zplus_permute (n m p : Z) (P : Z -> Prop)
- (H : P (m + (n + p))) := eq_ind_r P H (Zplus_permute n m p).
+ (H : P (m + (n + p))) := eq_ind_r P H (Z.add_shuffle3 n m p).
Definition fast_OMEGA10 (v c1 c2 l1 l2 k1 k2 : Z) (P : Z -> Prop)
(H : P (v * (c1 * k1 + c2 * k2) + (l1 * k1 + l2 * k2))) :=
@@ -259,24 +219,24 @@ Definition fast_Zred_factor0 (x : Z) (P : Z -> Prop)
(H : P (x * 1)) := eq_ind_r P H (Zred_factor0 x).
Definition fast_Zopp_eq_mult_neg_1 (x : Z) (P : Z -> Prop)
- (H : P (x * -1)) := eq_ind_r P H (Zopp_eq_mult_neg_1 x).
+ (H : P (x * -1)) := eq_ind_r P H (Z.opp_eq_mul_m1 x).
Definition fast_Zmult_comm (x y : Z) (P : Z -> Prop)
- (H : P (y * x)) := eq_ind_r P H (Zmult_comm x y).
+ (H : P (y * x)) := eq_ind_r P H (Z.mul_comm x y).
Definition fast_Zopp_plus_distr (x y : Z) (P : Z -> Prop)
- (H : P (- x + - y)) := eq_ind_r P H (Zopp_plus_distr x y).
+ (H : P (- x + - y)) := eq_ind_r P H (Z.opp_add_distr x y).
Definition fast_Zopp_involutive (x : Z) (P : Z -> Prop) (H : P x) :=
- eq_ind_r P H (Zopp_involutive x).
+ eq_ind_r P H (Z.opp_involutive x).
Definition fast_Zopp_mult_distr_r (x y : Z) (P : Z -> Prop)
(H : P (x * - y)) := eq_ind_r P H (Zopp_mult_distr_r x y).
Definition fast_Zmult_plus_distr_l (n m p : Z) (P : Z -> Prop)
- (H : P (n * p + m * p)) := eq_ind_r P H (Zmult_plus_distr_l n m p).
+ (H : P (n * p + m * p)) := eq_ind_r P H (Z.mul_add_distr_r n m p).
Definition fast_Zmult_opp_comm (x y : Z) (P : Z -> Prop)
- (H : P (x * - y)) := eq_ind_r P H (Zmult_opp_comm x y).
+ (H : P (x * - y)) := eq_ind_r P H (Z.mul_opp_comm x y).
Definition fast_Zmult_assoc_reverse (n m p : Z) (P : Z -> Prop)
(H : P (n * (m * p))) := eq_ind_r P H (Zmult_assoc_reverse n m p).
@@ -300,8 +260,8 @@ Definition fast_Zred_factor6 (x : Z) (P : Z -> Prop)
(H : P (x + 0)) := eq_ind_r P H (Zred_factor6 x).
Theorem intro_Z :
- forall n:nat, exists y : Z, Z_of_nat n = y /\ 0 <= y * 1 + 0.
+ forall n:nat, exists y : Z, Z.of_nat n = y /\ 0 <= y * 1 + 0.
Proof.
- intros n; exists (Z_of_nat n); split; trivial.
- rewrite Zmult_1_r, Zplus_0_r. apply Zle_0_nat.
+ intros n; exists (Z.of_nat n); split; trivial.
+ rewrite Z.mul_1_r, Z.add_0_r. apply Nat2Z.is_nonneg.
Qed.
diff --git a/plugins/omega/PreOmega.v b/plugins/omega/PreOmega.v
index 46fd5682d..bc08deaf9 100644
--- a/plugins/omega/PreOmega.v
+++ b/plugins/omega/PreOmega.v
@@ -1,4 +1,12 @@
-Require Import Arith Max Min ZArith_base NArith Nnat.
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Arith Max Min BinInt BinNat Znat Nnat.
Open Local Scope Z_scope.
@@ -15,16 +23,16 @@ Open Local Scope Z_scope.
- { eq, le, lt, ge, gt } on { Z, positive, N, nat }
Recognized operations:
- - on Z: Zmin, Zmax, Zabs, Zsgn are translated in term of <= < =
- - on nat: + * - S O pred min max nat_of_P nat_of_N Zabs_nat
- - on positive: Zneg Zpos xI xO xH + * - Psucc Ppred Pmin Pmax P_of_succ_nat
- - on N: N0 Npos + * - Nsucc Nmin Nmax N_of_nat Zabs_N
+ - on Z: Z.min, Z.max, Z.abs, Z.sgn are translated in term of <= < =
+ - on nat: + * - S O pred min max Pos.to_nat N.to_nat Z.abs_nat
+ - on positive: Zneg Zpos xI xO xH + * - Pos.succ Pos.pred Pos.min Pos.max Pos.of_succ_nat
+ - on N: N0 Npos + * - N.succ N.min N.max N.of_nat Z.abs_N
*)
-(** I) translation of Zmax, Zmin, Zabs, Zsgn into recognized equations *)
+(** I) translation of Z.max, Z.min, Z.abs, Z.sgn into recognized equations *)
Ltac zify_unop_core t thm a :=
(* Let's introduce the specification theorem for t *)
@@ -48,7 +56,7 @@ Ltac zify_unop t thm a :=
end.
Ltac zify_unop_nored t thm a :=
- (* in this version, we don't try to reduce the unop (that can be (Zplus x)) *)
+ (* in this version, we don't try to reduce the unop (that can be (Z.add x)) *)
let isz := isZcst a in
match isz with
| true => zify_unop_core t thm a
@@ -72,14 +80,14 @@ Ltac zify_binop t thm a b:=
Ltac zify_op_1 :=
match goal with
- | |- context [ Zmax ?a ?b ] => zify_binop Zmax Zmax_spec a b
- | H : context [ Zmax ?a ?b ] |- _ => zify_binop Zmax Zmax_spec a b
- | |- context [ Zmin ?a ?b ] => zify_binop Zmin Zmin_spec a b
- | H : context [ Zmin ?a ?b ] |- _ => zify_binop Zmin Zmin_spec a b
- | |- context [ Zsgn ?a ] => zify_unop Zsgn Zsgn_spec a
- | H : context [ Zsgn ?a ] |- _ => zify_unop Zsgn Zsgn_spec a
- | |- context [ Zabs ?a ] => zify_unop Zabs Zabs_spec a
- | H : context [ Zabs ?a ] |- _ => zify_unop Zabs Zabs_spec a
+ | |- context [ Z.max ?a ?b ] => zify_binop Z.max Z.max_spec a b
+ | H : context [ Z.max ?a ?b ] |- _ => zify_binop Z.max Z.max_spec a b
+ | |- context [ Z.min ?a ?b ] => zify_binop Z.min Z.min_spec a b
+ | H : context [ Z.min ?a ?b ] |- _ => zify_binop Z.min Z.min_spec a b
+ | |- context [ Z.sgn ?a ] => zify_unop Z.sgn Z.sgn_spec a
+ | H : context [ Z.sgn ?a ] |- _ => zify_unop Z.sgn Z.sgn_spec a
+ | |- context [ Z.abs ?a ] => zify_unop Z.abs Z.abs_spec a
+ | H : context [ Z.abs ?a ] |- _ => zify_unop Z.abs Z.abs_spec a
end.
Ltac zify_op := repeat zify_op_1.
@@ -91,100 +99,95 @@ Ltac zify_op := repeat zify_op_1.
(** II) Conversion from nat to Z *)
-Definition Z_of_nat' := Z_of_nat.
+Definition Z_of_nat' := Z.of_nat.
Ltac hide_Z_of_nat t :=
- let z := fresh "z" in set (z:=Z_of_nat t) in *;
- change Z_of_nat with Z_of_nat' in z;
+ let z := fresh "z" in set (z:=Z.of_nat t) in *;
+ change Z.of_nat with Z_of_nat' in z;
unfold z in *; clear z.
Ltac zify_nat_rel :=
match goal with
(* I: equalities *)
- | H : (@eq nat ?a ?b) |- _ => generalize (inj_eq _ _ H); clear H; intro H
- | |- (@eq nat ?a ?b) => apply (inj_eq_rev a b)
- | H : context [ @eq nat ?a ?b ] |- _ => rewrite (inj_eq_iff a b) in H
- | |- context [ @eq nat ?a ?b ] => rewrite (inj_eq_iff a b)
+ | |- (@eq nat ?a ?b) => apply (Nat2Z.inj a b) (* shortcut *)
+ | H : context [ @eq nat ?a ?b ] |- _ => rewrite <- (Nat2Z.inj_iff a b) in H
+ | |- context [ @eq nat ?a ?b ] => rewrite <- (Nat2Z.inj_iff a b)
(* II: less than *)
- | H : (lt ?a ?b) |- _ => generalize (inj_lt _ _ H); clear H; intro H
- | |- (lt ?a ?b) => apply (inj_lt_rev a b)
- | H : context [ lt ?a ?b ] |- _ => rewrite (inj_lt_iff a b) in H
- | |- context [ lt ?a ?b ] => rewrite (inj_lt_iff a b)
+ | H : context [ lt ?a ?b ] |- _ => rewrite (Nat2Z.inj_lt a b) in H
+ | |- context [ lt ?a ?b ] => rewrite (Nat2Z.inj_lt a b)
(* III: less or equal *)
- | H : (le ?a ?b) |- _ => generalize (inj_le _ _ H); clear H; intro H
- | |- (le ?a ?b) => apply (inj_le_rev a b)
- | H : context [ le ?a ?b ] |- _ => rewrite (inj_le_iff a b) in H
- | |- context [ le ?a ?b ] => rewrite (inj_le_iff a b)
+ | H : context [ le ?a ?b ] |- _ => rewrite (Nat2Z.inj_le a b) in H
+ | |- context [ le ?a ?b ] => rewrite (Nat2Z.inj_le a b)
(* IV: greater than *)
- | H : (gt ?a ?b) |- _ => generalize (inj_gt _ _ H); clear H; intro H
- | |- (gt ?a ?b) => apply (inj_gt_rev a b)
- | H : context [ gt ?a ?b ] |- _ => rewrite (inj_gt_iff a b) in H
- | |- context [ gt ?a ?b ] => rewrite (inj_gt_iff a b)
+ | H : context [ gt ?a ?b ] |- _ => rewrite (Nat2Z.inj_gt a b) in H
+ | |- context [ gt ?a ?b ] => rewrite (Nat2Z.inj_gt a b)
(* V: greater or equal *)
- | H : (ge ?a ?b) |- _ => generalize (inj_ge _ _ H); clear H; intro H
- | |- (ge ?a ?b) => apply (inj_ge_rev a b)
- | H : context [ ge ?a ?b ] |- _ => rewrite (inj_ge_iff a b) in H
- | |- context [ ge ?a ?b ] => rewrite (inj_ge_iff a b)
+ | H : context [ ge ?a ?b ] |- _ => rewrite (Nat2Z.inj_ge a b) in H
+ | |- context [ ge ?a ?b ] => rewrite (Nat2Z.inj_ge a b)
end.
Ltac zify_nat_op :=
match goal with
(* misc type conversions: positive/N/Z to nat *)
- | H : context [ Z_of_nat (nat_of_P ?a) ] |- _ => rewrite <- (Zpos_eq_Z_of_nat_o_nat_of_P a) in H
- | |- context [ Z_of_nat (nat_of_P ?a) ] => rewrite <- (Zpos_eq_Z_of_nat_o_nat_of_P a)
- | H : context [ Z_of_nat (nat_of_N ?a) ] |- _ => rewrite (Z_of_nat_of_N a) in H
- | |- context [ Z_of_nat (nat_of_N ?a) ] => rewrite (Z_of_nat_of_N a)
- | H : context [ Z_of_nat (Zabs_nat ?a) ] |- _ => rewrite (inj_Zabs_nat a) in H
- | |- context [ Z_of_nat (Zabs_nat ?a) ] => rewrite (inj_Zabs_nat a)
-
- (* plus -> Zplus *)
- | H : context [ Z_of_nat (plus ?a ?b) ] |- _ => rewrite (inj_plus a b) in H
- | |- context [ Z_of_nat (plus ?a ?b) ] => rewrite (inj_plus a b)
-
- (* min -> Zmin *)
- | H : context [ Z_of_nat (min ?a ?b) ] |- _ => rewrite (inj_min a b) in H
- | |- context [ Z_of_nat (min ?a ?b) ] => rewrite (inj_min a b)
-
- (* max -> Zmax *)
- | H : context [ Z_of_nat (max ?a ?b) ] |- _ => rewrite (inj_max a b) in H
- | |- context [ Z_of_nat (max ?a ?b) ] => rewrite (inj_max a b)
-
- (* minus -> Zmax (Zminus ... ...) 0 *)
- | H : context [ Z_of_nat (minus ?a ?b) ] |- _ => rewrite (inj_minus a b) in H
- | |- context [ Z_of_nat (minus ?a ?b) ] => rewrite (inj_minus a b)
-
- (* pred -> minus ... -1 -> Zmax (Zminus ... -1) 0 *)
- | H : context [ Z_of_nat (pred ?a) ] |- _ => rewrite (pred_of_minus a) in H
- | |- context [ Z_of_nat (pred ?a) ] => rewrite (pred_of_minus a)
-
- (* mult -> Zmult and a positivity hypothesis *)
- | H : context [ Z_of_nat (mult ?a ?b) ] |- _ =>
- pose proof (Zle_0_nat (mult a b)); rewrite (inj_mult a b) in *
- | |- context [ Z_of_nat (mult ?a ?b) ] =>
- pose proof (Zle_0_nat (mult a b)); rewrite (inj_mult a b) in *
+ | H : context [ Z.of_nat (Pos.to_nat ?a) ] |- _ => rewrite (positive_nat_Z a) in H
+ | |- context [ Z.of_nat (Pos.to_nat ?a) ] => rewrite (positive_nat_Z a)
+ | H : context [ Z.of_nat (N.to_nat ?a) ] |- _ => rewrite (N_nat_Z a) in H
+ | |- context [ Z.of_nat (N.to_nat ?a) ] => rewrite (N_nat_Z a)
+ | H : context [ Z.of_nat (Z.abs_nat ?a) ] |- _ => rewrite (Zabs2Nat.id_abs a) in H
+ | |- context [ Z.of_nat (Z.abs_nat ?a) ] => rewrite (Zabs2Nat.id_abs a)
+
+ (* plus -> Z.add *)
+ | H : context [ Z.of_nat (plus ?a ?b) ] |- _ => rewrite (Nat2Z.inj_add a b) in H
+ | |- context [ Z.of_nat (plus ?a ?b) ] => rewrite (Nat2Z.inj_add a b)
+
+ (* min -> Z.min *)
+ | H : context [ Z.of_nat (min ?a ?b) ] |- _ => rewrite (Nat2Z.inj_min a b) in H
+ | |- context [ Z.of_nat (min ?a ?b) ] => rewrite (Nat2Z.inj_min a b)
+
+ (* max -> Z.max *)
+ | H : context [ Z.of_nat (max ?a ?b) ] |- _ => rewrite (Nat2Z.inj_max a b) in H
+ | |- context [ Z.of_nat (max ?a ?b) ] => rewrite (Nat2Z.inj_max a b)
+
+ (* minus -> Z.max (Z.sub ... ...) 0 *)
+ | H : context [ Z.of_nat (minus ?a ?b) ] |- _ => rewrite (Nat2Z.inj_sub_max a b) in H
+ | |- context [ Z.of_nat (minus ?a ?b) ] => rewrite (Nat2Z.inj_sub_max a b)
+
+ (* pred -> minus ... -1 -> Z.max (Z.sub ... -1) 0 *)
+ | H : context [ Z.of_nat (pred ?a) ] |- _ => rewrite (pred_of_minus a) in H
+ | |- context [ Z.of_nat (pred ?a) ] => rewrite (pred_of_minus a)
+
+ (* mult -> Z.mul and a positivity hypothesis *)
+ | H : context [ Z.of_nat (mult ?a ?b) ] |- _ =>
+ pose proof (Nat2Z.is_nonneg (mult a b));
+ rewrite (Nat2Z.inj_mul a b) in *
+ | |- context [ Z.of_nat (mult ?a ?b) ] =>
+ pose proof (Nat2Z.is_nonneg (mult a b));
+ rewrite (Nat2Z.inj_mul a b) in *
(* O -> Z0 *)
- | H : context [ Z_of_nat O ] |- _ => simpl (Z_of_nat O) in H
- | |- context [ Z_of_nat O ] => simpl (Z_of_nat O)
+ | H : context [ Z.of_nat O ] |- _ => simpl (Z.of_nat O) in H
+ | |- context [ Z.of_nat O ] => simpl (Z.of_nat O)
- (* S -> number or Zsucc *)
- | H : context [ Z_of_nat (S ?a) ] |- _ =>
+ (* S -> number or Z.succ *)
+ | H : context [ Z.of_nat (S ?a) ] |- _ =>
let isnat := isnatcst a in
match isnat with
- | true => simpl (Z_of_nat (S a)) in H
- | _ => rewrite (inj_S a) in H
+ | true => simpl (Z.of_nat (S a)) in H
+ | _ => rewrite (Nat2Z.inj_succ a) in H
end
- | |- context [ Z_of_nat (S ?a) ] =>
+ | |- context [ Z.of_nat (S ?a) ] =>
let isnat := isnatcst a in
match isnat with
- | true => simpl (Z_of_nat (S a))
- | _ => rewrite (inj_S a)
+ | true => simpl (Z.of_nat (S a))
+ | _ => rewrite (Nat2Z.inj_succ a)
end
(* atoms of type nat : we add a positivity condition (if not already there) *)
- | _ : 0 <= Z_of_nat ?a |- _ => hide_Z_of_nat a
- | _ : context [ Z_of_nat ?a ] |- _ => pose proof (Zle_0_nat a); hide_Z_of_nat a
- | |- context [ Z_of_nat ?a ] => pose proof (Zle_0_nat a); hide_Z_of_nat a
+ | _ : 0 <= Z.of_nat ?a |- _ => hide_Z_of_nat a
+ | _ : context [ Z.of_nat ?a ] |- _ =>
+ pose proof (Nat2Z.is_nonneg a); hide_Z_of_nat a
+ | |- context [ Z.of_nat ?a ] =>
+ pose proof (Nat2Z.is_nonneg a); hide_Z_of_nat a
end.
Ltac zify_nat := repeat zify_nat_rel; repeat zify_nat_op; unfold Z_of_nat' in *.
@@ -205,10 +208,9 @@ Ltac hide_Zpos t :=
Ltac zify_positive_rel :=
match goal with
(* I: equalities *)
- | H : (@eq positive ?a ?b) |- _ => generalize (Zpos_eq _ _ H); clear H; intro H
- | |- (@eq positive ?a ?b) => apply (Zpos_eq_rev a b)
- | H : context [ @eq positive ?a ?b ] |- _ => rewrite (Zpos_eq_iff a b) in H
- | |- context [ @eq positive ?a ?b ] => rewrite (Zpos_eq_iff a b)
+ | |- (@eq positive ?a ?b) => apply Pos2Z.inj
+ | H : context [ @eq positive ?a ?b ] |- _ => rewrite <- (Pos2Z.inj_iff a b) in H
+ | |- context [ @eq positive ?a ?b ] => rewrite <- (Pos2Z.inj_iff a b)
(* II: less than *)
| H : context [ (?a < ?b)%positive ] |- _ => change (a<b)%positive with (Zpos a<Zpos b) in H
| |- context [ (?a < ?b)%positive ] => change (a<b)%positive with (Zpos a<Zpos b)
@@ -240,64 +242,66 @@ Ltac zify_positive_op :=
end
(* misc type conversions: nat to positive *)
- | H : context [ Zpos (P_of_succ_nat ?a) ] |- _ => rewrite (Zpos_P_of_succ_nat a) in H
- | |- context [ Zpos (P_of_succ_nat ?a) ] => rewrite (Zpos_P_of_succ_nat a)
+ | H : context [ Zpos (Pos.of_succ_nat ?a) ] |- _ => rewrite (Zpos_P_of_succ_nat a) in H
+ | |- context [ Zpos (Pos.of_succ_nat ?a) ] => rewrite (Zpos_P_of_succ_nat a)
- (* Pplus -> Zplus *)
- | H : context [ Zpos (Pplus ?a ?b) ] |- _ => change (Zpos (Pplus a b)) with (Zplus (Zpos a) (Zpos b)) in H
- | |- context [ Zpos (Pplus ?a ?b) ] => change (Zpos (Pplus a b)) with (Zplus (Zpos a) (Zpos b))
+ (* Pos.add -> Z.add *)
+ | H : context [ Zpos (?a + ?b) ] |- _ => change (Zpos (a+b)) with (Zpos a + Zpos b) in H
+ | |- context [ Zpos (?a + ?b) ] => change (Zpos (a+b)) with (Zpos a + Zpos b)
- (* Pmin -> Zmin *)
- | H : context [ Zpos (Pmin ?a ?b) ] |- _ => rewrite (Zpos_min a b) in H
- | |- context [ Zpos (Pmin ?a ?b) ] => rewrite (Zpos_min a b)
+ (* Pos.min -> Z.min *)
+ | H : context [ Zpos (Pos.min ?a ?b) ] |- _ => rewrite (Pos2Z.inj_min a b) in H
+ | |- context [ Zpos (Pos.min ?a ?b) ] => rewrite (Pos2Z.inj_min a b)
- (* Pmax -> Zmax *)
- | H : context [ Zpos (Pmax ?a ?b) ] |- _ => rewrite (Zpos_max a b) in H
- | |- context [ Zpos (Pmax ?a ?b) ] => rewrite (Zpos_max a b)
+ (* Pos.max -> Z.max *)
+ | H : context [ Zpos (Pos.max ?a ?b) ] |- _ => rewrite (Pos2Z.inj_max a b) in H
+ | |- context [ Zpos (Pos.max ?a ?b) ] => rewrite (Pos2Z.inj_max a b)
- (* Pminus -> Zmax 1 (Zminus ... ...) *)
- | H : context [ Zpos (Pminus ?a ?b) ] |- _ => rewrite (Zpos_minus a b) in H
- | |- context [ Zpos (Pminus ?a ?b) ] => rewrite (Zpos_minus a b)
+ (* Pos.sub -> Z.max 1 (Z.sub ... ...) *)
+ | H : context [ Zpos (Pos.sub ?a ?b) ] |- _ => rewrite (Pos2Z.inj_sub a b) in H
+ | |- context [ Zpos (Pos.sub ?a ?b) ] => rewrite (Pos2Z.inj_sub a b)
- (* Psucc -> Zsucc *)
- | H : context [ Zpos (Psucc ?a) ] |- _ => rewrite (Zpos_succ_morphism a) in H
- | |- context [ Zpos (Psucc ?a) ] => rewrite (Zpos_succ_morphism a)
+ (* Pos.succ -> Z.succ *)
+ | H : context [ Zpos (Pos.succ ?a) ] |- _ => rewrite (Pos2Z.inj_succ a) in H
+ | |- context [ Zpos (Pos.succ ?a) ] => rewrite (Pos2Z.inj_succ a)
- (* Ppred -> Pminus ... -1 -> Zmax 1 (Zminus ... - 1) *)
- | H : context [ Zpos (Ppred ?a) ] |- _ => rewrite (Ppred_minus a) in H
- | |- context [ Zpos (Ppred ?a) ] => rewrite (Ppred_minus a)
+ (* Pos.pred -> Pos.sub ... -1 -> Z.max 1 (Z.sub ... - 1) *)
+ | H : context [ Zpos (Pos.pred ?a) ] |- _ => rewrite <- (Pos.sub_1_r a) in H
+ | |- context [ Zpos (Pos.pred ?a) ] => rewrite <- (Pos.sub_1_r a)
- (* Pmult -> Zmult and a positivity hypothesis *)
- | H : context [ Zpos (Pmult ?a ?b) ] |- _ =>
- pose proof (Zgt_pos_0 (Pmult a b)); rewrite (Zpos_mult_morphism a b) in *
- | |- context [ Zpos (Pmult ?a ?b) ] =>
- pose proof (Zgt_pos_0 (Pmult a b)); rewrite (Zpos_mult_morphism a b) in *
+ (* Pos.mul -> Z.mul and a positivity hypothesis *)
+ | H : context [ Zpos (?a * ?b) ] |- _ =>
+ pose proof (Pos2Z.is_pos (Pos.mul a b));
+ change (Zpos (a*b)) with (Zpos a * Zpos b) in *
+ | |- context [ Zpos (?a * ?b) ] =>
+ pose proof (Pos2Z.is_pos (Pos.mul a b));
+ change (Zpos (a*b)) with (Zpos a * Zpos b) in *
(* xO *)
| H : context [ Zpos (xO ?a) ] |- _ =>
let isp := isPcst a in
match isp with
| true => change (Zpos (xO a)) with (Zpos' (xO a)) in H
- | _ => rewrite (Zpos_xO a) in H
+ | _ => rewrite (Pos2Z.inj_xO a) in H
end
| |- context [ Zpos (xO ?a) ] =>
let isp := isPcst a in
match isp with
| true => change (Zpos (xO a)) with (Zpos' (xO a))
- | _ => rewrite (Zpos_xO a)
+ | _ => rewrite (Pos2Z.inj_xO a)
end
(* xI *)
| H : context [ Zpos (xI ?a) ] |- _ =>
let isp := isPcst a in
match isp with
| true => change (Zpos (xI a)) with (Zpos' (xI a)) in H
- | _ => rewrite (Zpos_xI a) in H
+ | _ => rewrite (Pos2Z.inj_xI a) in H
end
| |- context [ Zpos (xI ?a) ] =>
let isp := isPcst a in
match isp with
| true => change (Zpos (xI a)) with (Zpos' (xI a))
- | _ => rewrite (Zpos_xI a)
+ | _ => rewrite (Pos2Z.inj_xI a)
end
(* xI : nothing to do, just prevent adding a useless positivity condition *)
@@ -305,9 +309,9 @@ Ltac zify_positive_op :=
| |- context [ Zpos xH ] => hide_Zpos xH
(* atoms of type positive : we add a positivity condition (if not already there) *)
- | _ : Zpos ?a > 0 |- _ => hide_Zpos a
- | _ : context [ Zpos ?a ] |- _ => pose proof (Zgt_pos_0 a); hide_Zpos a
- | |- context [ Zpos ?a ] => pose proof (Zgt_pos_0 a); hide_Zpos a
+ | _ : 0 < Zpos ?a |- _ => hide_Zpos a
+ | _ : context [ Zpos ?a ] |- _ => pose proof (Pos2Z.is_pos a); hide_Zpos a
+ | |- context [ Zpos ?a ] => pose proof (Pos2Z.is_pos a); hide_Zpos a
end.
Ltac zify_positive :=
@@ -319,84 +323,75 @@ Ltac zify_positive :=
(* IV) conversion from N to Z *)
-Definition Z_of_N' := Z_of_N.
+Definition Z_of_N' := Z.of_N.
Ltac hide_Z_of_N t :=
- let z := fresh "z" in set (z:=Z_of_N t) in *;
- change Z_of_N with Z_of_N' in z;
+ let z := fresh "z" in set (z:=Z.of_N t) in *;
+ change Z.of_N with Z_of_N' in z;
unfold z in *; clear z.
Ltac zify_N_rel :=
match goal with
(* I: equalities *)
- | H : (@eq N ?a ?b) |- _ => generalize (Z_of_N_eq _ _ H); clear H; intro H
- | |- (@eq N ?a ?b) => apply (Z_of_N_eq_rev a b)
- | H : context [ @eq N ?a ?b ] |- _ => rewrite (Z_of_N_eq_iff a b) in H
- | |- context [ @eq N ?a ?b ] => rewrite (Z_of_N_eq_iff a b)
+ | |- (@eq N ?a ?b) => apply (N2Z.inj a b) (* shortcut *)
+ | H : context [ @eq N ?a ?b ] |- _ => rewrite <- (N2Z.inj_iff a b) in H
+ | |- context [ @eq N ?a ?b ] => rewrite <- (N2Z.inj_iff a b)
(* II: less than *)
- | H : (?a < ?b)%N |- _ => generalize (Z_of_N_lt _ _ H); clear H; intro H
- | |- (?a < ?b)%N => apply (Z_of_N_lt_rev a b)
- | H : context [ (?a < ?b)%N ] |- _ => rewrite (Z_of_N_lt_iff a b) in H
- | |- context [ (?a < ?b)%N ] => rewrite (Z_of_N_lt_iff a b)
+ | H : context [ (?a < ?b)%N ] |- _ => rewrite (N2Z.inj_lt a b) in H
+ | |- context [ (?a < ?b)%N ] => rewrite (N2Z.inj_lt a b)
(* III: less or equal *)
- | H : (?a <= ?b)%N |- _ => generalize (Z_of_N_le _ _ H); clear H; intro H
- | |- (?a <= ?b)%N => apply (Z_of_N_le_rev a b)
- | H : context [ (?a <= ?b)%N ] |- _ => rewrite (Z_of_N_le_iff a b) in H
- | |- context [ (?a <= ?b)%N ] => rewrite (Z_of_N_le_iff a b)
+ | H : context [ (?a <= ?b)%N ] |- _ => rewrite (N2Z.inj_le a b) in H
+ | |- context [ (?a <= ?b)%N ] => rewrite (N2Z.inj_le a b)
(* IV: greater than *)
- | H : (?a > ?b)%N |- _ => generalize (Z_of_N_gt _ _ H); clear H; intro H
- | |- (?a > ?b)%N => apply (Z_of_N_gt_rev a b)
- | H : context [ (?a > ?b)%N ] |- _ => rewrite (Z_of_N_gt_iff a b) in H
- | |- context [ (?a > ?b)%N ] => rewrite (Z_of_N_gt_iff a b)
+ | H : context [ (?a > ?b)%N ] |- _ => rewrite (N2Z.inj_gt a b) in H
+ | |- context [ (?a > ?b)%N ] => rewrite (N2Z.inj_gt a b)
(* V: greater or equal *)
- | H : (?a >= ?b)%N |- _ => generalize (Z_of_N_ge _ _ H); clear H; intro H
- | |- (?a >= ?b)%N => apply (Z_of_N_ge_rev a b)
- | H : context [ (?a >= ?b)%N ] |- _ => rewrite (Z_of_N_ge_iff a b) in H
- | |- context [ (?a >= ?b)%N ] => rewrite (Z_of_N_ge_iff a b)
+ | H : context [ (?a >= ?b)%N ] |- _ => rewrite (N2Z.inj_ge a b) in H
+ | |- context [ (?a >= ?b)%N ] => rewrite (N2Z.inj_ge a b)
end.
Ltac zify_N_op :=
match goal with
(* misc type conversions: nat to positive *)
- | H : context [ Z_of_N (N_of_nat ?a) ] |- _ => rewrite (Z_of_N_of_nat a) in H
- | |- context [ Z_of_N (N_of_nat ?a) ] => rewrite (Z_of_N_of_nat a)
- | H : context [ Z_of_N (Zabs_N ?a) ] |- _ => rewrite (Z_of_N_abs a) in H
- | |- context [ Z_of_N (Zabs_N ?a) ] => rewrite (Z_of_N_abs a)
- | H : context [ Z_of_N (Npos ?a) ] |- _ => rewrite (Z_of_N_pos a) in H
- | |- context [ Z_of_N (Npos ?a) ] => rewrite (Z_of_N_pos a)
- | H : context [ Z_of_N N0 ] |- _ => change (Z_of_N N0) with Z0 in H
- | |- context [ Z_of_N N0 ] => change (Z_of_N N0) with Z0
-
- (* Nplus -> Zplus *)
- | H : context [ Z_of_N (Nplus ?a ?b) ] |- _ => rewrite (Z_of_N_plus a b) in H
- | |- context [ Z_of_N (Nplus ?a ?b) ] => rewrite (Z_of_N_plus a b)
-
- (* Nmin -> Zmin *)
- | H : context [ Z_of_N (Nmin ?a ?b) ] |- _ => rewrite (Z_of_N_min a b) in H
- | |- context [ Z_of_N (Nmin ?a ?b) ] => rewrite (Z_of_N_min a b)
-
- (* Nmax -> Zmax *)
- | H : context [ Z_of_N (Nmax ?a ?b) ] |- _ => rewrite (Z_of_N_max a b) in H
- | |- context [ Z_of_N (Nmax ?a ?b) ] => rewrite (Z_of_N_max a b)
-
- (* Nminus -> Zmax 0 (Zminus ... ...) *)
- | H : context [ Z_of_N (Nminus ?a ?b) ] |- _ => rewrite (Z_of_N_minus a b) in H
- | |- context [ Z_of_N (Nminus ?a ?b) ] => rewrite (Z_of_N_minus a b)
-
- (* Nsucc -> Zsucc *)
- | H : context [ Z_of_N (Nsucc ?a) ] |- _ => rewrite (Z_of_N_succ a) in H
- | |- context [ Z_of_N (Nsucc ?a) ] => rewrite (Z_of_N_succ a)
-
- (* Nmult -> Zmult and a positivity hypothesis *)
- | H : context [ Z_of_N (Nmult ?a ?b) ] |- _ =>
- pose proof (Z_of_N_le_0 (Nmult a b)); rewrite (Z_of_N_mult a b) in *
- | |- context [ Z_of_N (Nmult ?a ?b) ] =>
- pose proof (Z_of_N_le_0 (Nmult a b)); rewrite (Z_of_N_mult a b) in *
+ | H : context [ Z.of_N (N.of_nat ?a) ] |- _ => rewrite (nat_N_Z a) in H
+ | |- context [ Z.of_N (N.of_nat ?a) ] => rewrite (nat_N_Z a)
+ | H : context [ Z.of_N (Z.abs_N ?a) ] |- _ => rewrite (N2Z.inj_abs_N a) in H
+ | |- context [ Z.of_N (Z.abs_N ?a) ] => rewrite (N2Z.inj_abs_N a)
+ | H : context [ Z.of_N (Npos ?a) ] |- _ => rewrite (N2Z.inj_pos a) in H
+ | |- context [ Z.of_N (Npos ?a) ] => rewrite (N2Z.inj_pos a)
+ | H : context [ Z.of_N N0 ] |- _ => change (Z.of_N N0) with Z0 in H
+ | |- context [ Z.of_N N0 ] => change (Z.of_N N0) with Z0
+
+ (* N.add -> Z.add *)
+ | H : context [ Z.of_N (N.add ?a ?b) ] |- _ => rewrite (N2Z.inj_add a b) in H
+ | |- context [ Z.of_N (N.add ?a ?b) ] => rewrite (N2Z.inj_add a b)
+
+ (* N.min -> Z.min *)
+ | H : context [ Z.of_N (N.min ?a ?b) ] |- _ => rewrite (N2Z.inj_min a b) in H
+ | |- context [ Z.of_N (N.min ?a ?b) ] => rewrite (N2Z.inj_min a b)
+
+ (* N.max -> Z.max *)
+ | H : context [ Z.of_N (N.max ?a ?b) ] |- _ => rewrite (N2Z.inj_max a b) in H
+ | |- context [ Z.of_N (N.max ?a ?b) ] => rewrite (N2Z.inj_max a b)
+
+ (* N.sub -> Z.max 0 (Z.sub ... ...) *)
+ | H : context [ Z.of_N (N.sub ?a ?b) ] |- _ => rewrite (N2Z.inj_sub_max a b) in H
+ | |- context [ Z.of_N (N.sub ?a ?b) ] => rewrite (N2Z.inj_sub_max a b)
+
+ (* N.succ -> Z.succ *)
+ | H : context [ Z.of_N (N.succ ?a) ] |- _ => rewrite (N2Z.inj_succ a) in H
+ | |- context [ Z.of_N (N.succ ?a) ] => rewrite (N2Z.inj_succ a)
+
+ (* N.mul -> Z.mul and a positivity hypothesis *)
+ | H : context [ Z.of_N (N.mul ?a ?b) ] |- _ =>
+ pose proof (N2Z.is_nonneg (N.mul a b)); rewrite (N2Z.inj_mul a b) in *
+ | |- context [ Z.of_N (N.mul ?a ?b) ] =>
+ pose proof (N2Z.is_nonneg (N.mul a b)); rewrite (N2Z.inj_mul a b) in *
(* atoms of type N : we add a positivity condition (if not already there) *)
- | _ : 0 <= Z_of_N ?a |- _ => hide_Z_of_N a
- | _ : context [ Z_of_N ?a ] |- _ => pose proof (Z_of_N_le_0 a); hide_Z_of_N a
- | |- context [ Z_of_N ?a ] => pose proof (Z_of_N_le_0 a); hide_Z_of_N a
+ | _ : 0 <= Z.of_N ?a |- _ => hide_Z_of_N a
+ | _ : context [ Z.of_N ?a ] |- _ => pose proof (N2Z.is_nonneg a); hide_Z_of_N a
+ | |- context [ Z.of_N ?a ] => pose proof (N2Z.is_nonneg a); hide_Z_of_N a
end.
Ltac zify_N := repeat zify_N_rel; repeat zify_N_op; unfold Z_of_N' in *.
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index 2a98cae53..9fa18e7dc 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -173,6 +173,9 @@ let init_constant = gen_constant_in_modules "Omega" init_modules
let constant = gen_constant_in_modules "Omega" coq_modules
let z_constant = gen_constant_in_modules "Omega" [["Coq";"ZArith"]]
+let zbase_constant =
+ gen_constant_in_modules "Omega" [["Coq";"ZArith";"BinInt"]]
+
(* Zarith *)
let coq_xH = lazy (constant "xH")
@@ -184,20 +187,20 @@ let coq_Zneg = lazy (constant "Zneg")
let coq_Z = lazy (constant "Z")
let coq_comparison = lazy (constant "comparison")
let coq_Gt = lazy (constant "Gt")
-let coq_Zplus = lazy (constant "Zplus")
-let coq_Zmult = lazy (constant "Zmult")
-let coq_Zopp = lazy (constant "Zopp")
-let coq_Zminus = lazy (constant "Zminus")
-let coq_Zsucc = lazy (constant "Zsucc")
-let coq_Zpred = lazy (constant "Zpred")
-let coq_Zgt = lazy (constant "Zgt")
-let coq_Zle = lazy (constant "Zle")
-let coq_Z_of_nat = lazy (constant "Z_of_nat")
-let coq_inj_plus = lazy (constant "inj_plus")
-let coq_inj_mult = lazy (constant "inj_mult")
-let coq_inj_minus1 = lazy (constant "inj_minus1")
+let coq_Zplus = lazy (zbase_constant "Z.add")
+let coq_Zmult = lazy (zbase_constant "Z.mul")
+let coq_Zopp = lazy (zbase_constant "Z.opp")
+let coq_Zminus = lazy (zbase_constant "Z.sub")
+let coq_Zsucc = lazy (zbase_constant "Z.succ")
+let coq_Zpred = lazy (zbase_constant "Z.pred")
+let coq_Zgt = lazy (zbase_constant "Z.gt")
+let coq_Zle = lazy (zbase_constant "Z.le")
+let coq_Z_of_nat = lazy (zbase_constant "Z.of_nat")
+let coq_inj_plus = lazy (z_constant "Nat2Z.inj_add")
+let coq_inj_mult = lazy (z_constant "Nat2Z.inj_mul")
+let coq_inj_minus1 = lazy (z_constant "Nat2Z.inj_sub")
let coq_inj_minus2 = lazy (constant "inj_minus2")
-let coq_inj_S = lazy (z_constant "inj_S")
+let coq_inj_S = lazy (z_constant "Nat2Z.inj_succ")
let coq_inj_le = lazy (z_constant "Znat.inj_le")
let coq_inj_lt = lazy (z_constant "Znat.inj_lt")
let coq_inj_ge = lazy (z_constant "Znat.inj_ge")
@@ -253,10 +256,10 @@ let coq_Zle_left = lazy (constant "Zle_left")
let coq_new_var = lazy (constant "new_var")
let coq_intro_Z = lazy (constant "intro_Z")
-let coq_dec_eq = lazy (constant "dec_eq")
+let coq_dec_eq = lazy (zbase_constant "Z.eq_decidable")
let coq_dec_Zne = lazy (constant "dec_Zne")
-let coq_dec_Zle = lazy (constant "dec_Zle")
-let coq_dec_Zlt = lazy (constant "dec_Zlt")
+let coq_dec_Zle = lazy (zbase_constant "Z.le_decidable")
+let coq_dec_Zlt = lazy (zbase_constant "Z.lt_decidable")
let coq_dec_Zgt = lazy (constant "dec_Zgt")
let coq_dec_Zge = lazy (constant "dec_Zge")
@@ -268,10 +271,10 @@ let coq_Znot_ge_lt = lazy (constant "Znot_ge_lt")
let coq_Znot_gt_le = lazy (constant "Znot_gt_le")
let coq_neq = lazy (constant "neq")
let coq_Zne = lazy (constant "Zne")
-let coq_Zle = lazy (constant "Zle")
-let coq_Zgt = lazy (constant "Zgt")
-let coq_Zge = lazy (constant "Zge")
-let coq_Zlt = lazy (constant "Zlt")
+let coq_Zle = lazy (zbase_constant "Z.le")
+let coq_Zgt = lazy (zbase_constant "Z.gt")
+let coq_Zge = lazy (zbase_constant "Z.ge")
+let coq_Zlt = lazy (zbase_constant "Z.lt")
(* Peano/Datatypes *)
let coq_le = lazy (init_constant "le")
@@ -329,13 +332,13 @@ let evaluable_ref_of_constr s c = match kind_of_term (Lazy.force c) with
EvalConstRef kn
| _ -> anomaly ("Coq_omega: "^s^" is not an evaluable constant")
-let sp_Zsucc = lazy (evaluable_ref_of_constr "Zsucc" coq_Zsucc)
-let sp_Zpred = lazy (evaluable_ref_of_constr "Zpred" coq_Zpred)
-let sp_Zminus = lazy (evaluable_ref_of_constr "Zminus" coq_Zminus)
-let sp_Zle = lazy (evaluable_ref_of_constr "Zle" coq_Zle)
-let sp_Zgt = lazy (evaluable_ref_of_constr "Zgt" coq_Zgt)
-let sp_Zge = lazy (evaluable_ref_of_constr "Zge" coq_Zge)
-let sp_Zlt = lazy (evaluable_ref_of_constr "Zlt" coq_Zlt)
+let sp_Zsucc = lazy (evaluable_ref_of_constr "Z.succ" coq_Zsucc)
+let sp_Zpred = lazy (evaluable_ref_of_constr "Z.pred" coq_Zpred)
+let sp_Zminus = lazy (evaluable_ref_of_constr "Z.sub" coq_Zminus)
+let sp_Zle = lazy (evaluable_ref_of_constr "Z.le" coq_Zle)
+let sp_Zgt = lazy (evaluable_ref_of_constr "Z.gt" coq_Zgt)
+let sp_Zge = lazy (evaluable_ref_of_constr "Z.ge" coq_Zge)
+let sp_Zlt = lazy (evaluable_ref_of_constr "Z.lt" coq_Zlt)
let sp_not = lazy (evaluable_ref_of_constr "not" (lazy (build_coq_not ())))
let mk_var v = mkVar (id_of_string v)
diff --git a/plugins/ring/LegacyNArithRing.v b/plugins/ring/LegacyNArithRing.v
index 5dcd6d840..a69cf0957 100644
--- a/plugins/ring/LegacyNArithRing.v
+++ b/plugins/ring/LegacyNArithRing.v
@@ -22,23 +22,22 @@ Definition Neq (n m:N) :=
Lemma Neq_prop : forall n m:N, Is_true (Neq n m) -> n = m.
intros n m H; unfold Neq in H.
- apply Ncompare_Eq_eq.
+ apply N.compare_eq.
destruct (n ?= m)%N; [ reflexivity | contradiction | contradiction ].
Qed.
-Definition NTheory : Semi_Ring_Theory Nplus Nmult 1%N 0%N Neq.
+Definition NTheory : Semi_Ring_Theory N.add N.mul 1%N 0%N Neq.
split.
- apply Nplus_comm.
- apply Nplus_assoc.
- apply Nmult_comm.
- apply Nmult_assoc.
- apply Nplus_0_l.
- apply Nmult_1_l.
- apply Nmult_0_l.
- apply Nmult_plus_distr_r.
-(* apply Nplus_reg_l.*)
+ apply N.add_comm.
+ apply N.add_assoc.
+ apply N.mul_comm.
+ apply N.mul_assoc.
+ apply N.add_0_l.
+ apply N.mul_1_l.
+ apply N.mul_0_l.
+ apply N.mul_add_distr_r.
apply Neq_prop.
Qed.
Add Legacy Semi Ring
- N Nplus Nmult 1%N 0%N Neq NTheory [ Npos 0%N xO xI 1%positive ].
+ N N.add N.mul 1%N 0%N Neq NTheory [ Npos 0%N xO xI 1%positive ].
diff --git a/plugins/ring/LegacyZArithRing.v b/plugins/ring/LegacyZArithRing.v
index 5845062dd..5c702c90e 100644
--- a/plugins/ring/LegacyZArithRing.v
+++ b/plugins/ring/LegacyZArithRing.v
@@ -21,15 +21,15 @@ Definition Zeq (x y:Z) :=
Lemma Zeq_prop : forall x y:Z, Is_true (Zeq x y) -> x = y.
intros x y H; unfold Zeq in H.
- apply Zcompare_Eq_eq.
+ apply Z.compare_eq.
destruct (x ?= y)%Z; [ reflexivity | contradiction | contradiction ].
Qed.
-Definition ZTheory : Ring_Theory Zplus Zmult 1%Z 0%Z Zopp Zeq.
+Definition ZTheory : Ring_Theory Z.add Z.mul 1%Z 0%Z Z.opp Zeq.
split; intros; eauto with zarith.
apply Zeq_prop; assumption.
Qed.
(* NatConstants and NatTheory are defined in Ring_theory.v *)
-Add Legacy Ring Z Zplus Zmult 1%Z 0%Z Zopp Zeq ZTheory
+Add Legacy Ring Z Z.add Z.mul 1%Z 0%Z Z.opp Zeq ZTheory
[ Zpos Zneg 0%Z xO xI 1%positive ].
diff --git a/plugins/ring/Ring_abstract.v b/plugins/ring/Ring_abstract.v
index 1763d70a6..535894160 100644
--- a/plugins/ring/Ring_abstract.v
+++ b/plugins/ring/Ring_abstract.v
@@ -137,8 +137,7 @@ Hint Resolve (SR_plus_zero_right2 T).
Hint Resolve (SR_mult_one_right T).
Hint Resolve (SR_mult_one_right2 T).
(*Hint Resolve (SR_plus_reg_right T).*)
-Hint Resolve refl_equal sym_equal trans_equal.
-(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hint Resolve eq_refl eq_sym eq_trans.
Hint Immediate T.
Remark iacs_aux_ok :
@@ -446,8 +445,7 @@ Hint Resolve (Th_plus_zero_right2 T).
Hint Resolve (Th_mult_one_right T).
Hint Resolve (Th_mult_one_right2 T).
(*Hint Resolve (Th_plus_reg_right T).*)
-Hint Resolve refl_equal sym_equal trans_equal.
-(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hint Resolve eq_refl eq_sym eq_trans.
Hint Immediate T.
Lemma isacs_aux_ok :
diff --git a/plugins/ring/Ring_normalize.v b/plugins/ring/Ring_normalize.v
index c6dff3e00..9755d71e1 100644
--- a/plugins/ring/Ring_normalize.v
+++ b/plugins/ring/Ring_normalize.v
@@ -365,8 +365,7 @@ Hint Resolve (SR_plus_zero_right2 T).
Hint Resolve (SR_mult_one_right T).
Hint Resolve (SR_mult_one_right2 T).
(*Hint Resolve (SR_plus_reg_right T).*)
-Hint Resolve refl_equal sym_equal trans_equal.
-(* Hints Resolve refl_eqT sym_eqT trans_eqT. *)
+Hint Resolve eq_refl eq_sym eq_trans.
Hint Immediate T.
Lemma varlist_eq_prop : forall x y:varlist, Is_true (varlist_eq x y) -> x = y.
@@ -794,8 +793,7 @@ Hint Resolve (Th_plus_zero_right2 T).
Hint Resolve (Th_mult_one_right T).
Hint Resolve (Th_mult_one_right2 T).
(*Hint Resolve (Th_plus_reg_right T).*)
-Hint Resolve refl_equal sym_equal trans_equal.
-(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hint Resolve eq_refl eq_sym eq_trans.
Hint Immediate T.
(*** Definitions *)
diff --git a/plugins/ring/Setoid_ring_normalize.v b/plugins/ring/Setoid_ring_normalize.v
index ad75a8a4d..94b36e246 100644
--- a/plugins/ring/Setoid_ring_normalize.v
+++ b/plugins/ring/Setoid_ring_normalize.v
@@ -387,8 +387,7 @@ Hint Resolve (SSR_plus_zero_right2 S T).
Hint Resolve (SSR_mult_one_right S T).
Hint Resolve (SSR_mult_one_right2 S T).
Hint Resolve (SSR_plus_reg_right S T).
-Hint Resolve refl_equal sym_equal trans_equal.
-(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hint Resolve eq_refl eq_sym eq_trans.
Hint Immediate T.
Lemma varlist_eq_prop : forall x y:varlist, Is_true (varlist_eq x y) -> x = y.
@@ -1052,8 +1051,7 @@ Hint Resolve (STh_plus_zero_right2 S T).
Hint Resolve (STh_mult_one_right S T).
Hint Resolve (STh_mult_one_right2 S T).
Hint Resolve (STh_plus_reg_right S plus_morph T).
-Hint Resolve refl_equal sym_equal trans_equal.
-(*Hints Resolve refl_eqT sym_eqT trans_eqT.*)
+Hint Resolve eq_refl eq_sym eq_trans.
Hint Immediate T.
diff --git a/plugins/ring/ring.ml b/plugins/ring/ring.ml
index ab6ee1573..22d5adb18 100644
--- a/plugins/ring/ring.ml
+++ b/plugins/ring/ring.ml
@@ -451,7 +451,7 @@ let build_polynom gl th lc =
mkLApp(coq_Pplus, [|th.th_a; aux c1; aux c2 |])
| App (binop, [|c1; c2|]) when safe_pf_conv_x gl binop th.th_mult ->
mkLApp(coq_Pmult, [|th.th_a; aux c1; aux c2 |])
- (* The special case of Zminus *)
+ (* The special case of Z.sub *)
| App (binop, [|c1; c2|])
when safe_pf_conv_x gl c
(mkApp (th.th_plus, [|c1; mkApp(unbox th.th_opp, [|c2|])|])) ->
@@ -569,7 +569,7 @@ let build_apolynom gl th lc =
mkLApp(coq_APplus, [| aux c1; aux c2 |])
| App (binop, [|c1; c2|]) when safe_pf_conv_x gl binop th.th_mult ->
mkLApp(coq_APmult, [| aux c1; aux c2 |])
- (* The special case of Zminus *)
+ (* The special case of Z.sub *)
| App (binop, [|c1; c2|])
when safe_pf_conv_x gl c
(mkApp(th.th_plus, [|c1; mkApp(unbox th.th_opp,[|c2|]) |])) ->
@@ -630,7 +630,7 @@ let build_setpolynom gl th lc =
mkLApp(coq_SetPplus, [|th.th_a; aux c1; aux c2 |])
| App (binop, [|c1; c2|]) when safe_pf_conv_x gl binop th.th_mult ->
mkLApp(coq_SetPmult, [|th.th_a; aux c1; aux c2 |])
- (* The special case of Zminus *)
+ (* The special case of Z.sub *)
| App (binop, [|c1; c2|])
when safe_pf_conv_x gl c
(mkApp(th.th_plus, [|c1; mkApp(unbox th.th_opp,[|c2|])|])) ->
diff --git a/plugins/romega/ReflOmegaCore.v b/plugins/romega/ReflOmegaCore.v
index 56ae921ed..dad368931 100644
--- a/plugins/romega/ReflOmegaCore.v
+++ b/plugins/romega/ReflOmegaCore.v
@@ -86,73 +86,50 @@ Module Z_as_Int <: Int.
Definition int := Z.
Definition zero := 0.
Definition one := 1.
- Definition plus := Zplus.
- Definition opp := Zopp.
- Definition minus := Zminus.
- Definition mult := Zmult.
+ Definition plus := Z.add.
+ Definition opp := Z.opp.
+ Definition minus := Z.sub.
+ Definition mult := Z.mul.
Lemma ring : @ring_theory int zero one plus mult minus opp (@eq int).
Proof.
constructor.
- exact Zplus_0_l.
- exact Zplus_comm.
- exact Zplus_assoc.
- exact Zmult_1_l.
- exact Zmult_comm.
- exact Zmult_assoc.
- exact Zmult_plus_distr_l.
- unfold minus, Zminus; auto.
- exact Zplus_opp_r.
+ exact Z.add_0_l.
+ exact Z.add_comm.
+ exact Z.add_assoc.
+ exact Z.mul_1_l.
+ exact Z.mul_comm.
+ exact Z.mul_assoc.
+ exact Z.mul_add_distr_r.
+ unfold minus, Z.sub; auto.
+ exact Z.add_opp_diag_r.
Qed.
- Definition le := Zle.
- Definition lt := Zlt.
- Definition ge := Zge.
- Definition gt := Zgt.
- Lemma le_lt_iff : forall i j, (i<=j) <-> ~(j<i).
- Proof.
- split; intros.
- apply Zle_not_lt; auto.
- rewrite <- Zge_iff_le.
- apply Znot_lt_ge; auto.
- Qed.
- Definition ge_le_iff := Zge_iff_le.
- Definition gt_lt_iff := Zgt_iff_lt.
+ Definition le := Z.le.
+ Definition lt := Z.lt.
+ Definition ge := Z.ge.
+ Definition gt := Z.gt.
+ Definition le_lt_iff := Z.le_ngt.
+ Definition ge_le_iff := Z.ge_le_iff.
+ Definition gt_lt_iff := Z.gt_lt_iff.
- Definition lt_trans := Zlt_trans.
- Definition lt_not_eq := Zlt_not_eq.
+ Definition lt_trans := Z.lt_trans.
+ Definition lt_not_eq := Z.lt_neq.
- Definition lt_0_1 := Zlt_0_1.
- Definition plus_le_compat := Zplus_le_compat.
+ Definition lt_0_1 := Z.lt_0_1.
+ Definition plus_le_compat := Z.add_le_mono.
Definition mult_lt_compat_l := Zmult_lt_compat_l.
- Lemma opp_le_compat : forall i j, i<=j -> (-j)<=(-i).
- Proof.
- unfold Zle; intros; rewrite <- Zcompare_opp; auto.
- Qed.
+ Lemma opp_le_compat i j : i<=j -> (-j)<=(-i).
+ Proof. apply -> Z.opp_le_mono. Qed.
- Definition compare := Zcompare.
- Definition compare_Eq := Zcompare_Eq_iff_eq.
- Lemma compare_Lt : forall i j, compare i j = Lt <-> i<j.
- Proof. intros; unfold compare, Zlt; intuition. Qed.
- Lemma compare_Gt : forall i j, compare i j = Gt <-> i>j.
- Proof. intros; unfold compare, Zgt; intuition. Qed.
+ Definition compare := Z.compare.
+ Definition compare_Eq := Z.compare_eq_iff.
+ Lemma compare_Lt i j : compare i j = Lt <-> i<j.
+ Proof. reflexivity. Qed.
+ Lemma compare_Gt i j : compare i j = Gt <-> i>j.
+ Proof. reflexivity. Qed.
- Lemma le_lt_int : forall x y, x<y <-> x<=y+-(1).
- Proof.
- intros; split; intros.
- generalize (Zlt_left _ _ H); simpl; intros.
- apply Zle_left_rev; auto.
- apply Zlt_0_minus_lt.
- generalize (Zplus_le_lt_compat x (y+-1) (-x) (-x+1) H).
- rewrite Zplus_opp_r.
- rewrite <-Zplus_assoc.
- rewrite (Zplus_permute (-1)).
- simpl in *.
- rewrite Zplus_0_r.
- intro H'; apply H'.
- replace (-x+1) with (Zsucc (-x)); auto.
- apply Zlt_succ.
- Qed.
+ Definition le_lt_int := Z.lt_le_pred.
End Z_as_Int.
@@ -2192,7 +2169,7 @@ Proof.
auto; case (nth_hyps j l);
auto; intros t3 t4; case t3; auto;
simpl in |- *; intros z z' H1 H2;
- generalize (refl_equal (interp_term e (fusion_cancel t (t2 + t4)%term)));
+ generalize (eq_refl (interp_term e (fusion_cancel t (t2 + t4)%term)));
pattern (fusion_cancel t (t2 + t4)%term) at 2 3 in |- *;
case (fusion_cancel t (t2 + t4)%term); simpl in |- *;
auto; intro k; elim (fusion_cancel_stable t); simpl in |- *.
@@ -2370,7 +2347,7 @@ Proof.
unfold valid1, exact_divide in |- *; intros k1 k2 t ep e p1;
Simplify; simpl; auto; subst;
rewrite <- scalar_norm_stable; simpl; intros;
- [ destruct (mult_integral _ _ (sym_eq H0)); intuition
+ [ destruct (mult_integral _ _ (eq_sym H0)); intuition
| contradict H0; rewrite <- H0, mult_0_l; auto
].
Qed.
diff --git a/plugins/rtauto/Bintree.v b/plugins/rtauto/Bintree.v
index 77f8f8345..8f024a15f 100644
--- a/plugins/rtauto/Bintree.v
+++ b/plugins/rtauto/Bintree.v
@@ -15,14 +15,14 @@ Open Scope positive_scope.
Ltac clean := try (simpl; congruence).
Lemma Gt_Psucc: forall p q,
- (p ?= Psucc q) = Gt -> (p ?= q) = Gt.
+ (p ?= Pos.succ q) = Gt -> (p ?= q) = Gt.
Proof.
intros. rewrite <- Pos.compare_succ_succ.
now apply Pos.lt_gt, Pos.lt_lt_succ, Pos.gt_lt.
Qed.
Lemma Psucc_Gt : forall p,
- (Psucc p ?= p) = Gt.
+ (Pos.succ p ?= p) = Gt.
Proof.
intros. apply Pos.lt_gt, Pos.lt_succ_diag_r.
Qed.
@@ -181,7 +181,7 @@ mkStore {index:positive;contents:Tree}.
Definition empty := mkStore xH Tempty.
Definition push a S :=
-mkStore (Psucc (index S)) (Tadd (index S) a (contents S)).
+mkStore (Pos.succ (index S)) (Tadd (index S) a (contents S)).
Definition get i S := Tget i (contents S).
@@ -214,7 +214,7 @@ intros a S.
rewrite Tget_Tadd.
rewrite Psucc_Gt.
intro W.
-change (get (Psucc (index S)) S =PNone).
+change (get (Pos.succ (index S)) S =PNone).
apply get_Full_Gt; auto.
apply Psucc_Gt.
Qed.
@@ -248,7 +248,7 @@ forall x, get i S = PSome x ->
Proof.
intros i a S F x H.
case_eq (i ?= index S);intro test.
-rewrite (Pcompare_Eq_eq _ _ test) in H.
+rewrite (Pos.compare_eq _ _ test) in H.
rewrite (get_Full_Eq _ F) in H;congruence.
rewrite <- H.
rewrite (get_push_Full i a).
@@ -260,13 +260,13 @@ Qed.
Lemma Full_index_one_empty : forall S, Full S -> index S = 1 -> S=empty.
intros [ind cont] F one; inversion F.
reflexivity.
-simpl index in one;assert (h:=Psucc_not_one (index S)).
+simpl index in one;assert (h:=Pos.succ_not_1 (index S)).
congruence.
Qed.
Lemma push_not_empty: forall a S, (push a S) <> empty.
intros a [ind cont];unfold push,empty.
-simpl;intro H;injection H; intros _ ; apply Psucc_not_one.
+simpl;intro H;injection H; intros _ ; apply Pos.succ_not_1.
Qed.
Fixpoint In (x:A) (S:Store) (F:Full S) {struct F}: Prop :=
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index 8db267641..2fc80e1ff 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -33,7 +33,7 @@ let data_constant =
Coqlib.gen_constant "refl_tauto" ["Init";"Datatypes"]
let l_true_equals_true =
- lazy (mkApp(logic_constant "refl_equal",
+ lazy (mkApp(logic_constant "eq_refl",
[|data_constant "bool";data_constant "true"|]))
let pos_constant =
diff --git a/plugins/setoid_ring/ArithRing.v b/plugins/setoid_ring/ArithRing.v
index 06822ae16..8e234db74 100644
--- a/plugins/setoid_ring/ArithRing.v
+++ b/plugins/setoid_ring/ArithRing.v
@@ -21,17 +21,17 @@ Lemma natSRth : semi_ring_theory O (S O) plus mult (@eq nat).
Lemma nat_morph_N :
semi_morph 0 1 plus mult (eq (A:=nat))
- 0%N 1%N N.add N.mul N.eqb nat_of_N.
+ 0%N 1%N N.add N.mul N.eqb N.to_nat.
Proof.
constructor;trivial.
- exact nat_of_Nplus.
- exact nat_of_Nmult.
+ exact N2Nat.inj_add.
+ exact N2Nat.inj_mul.
intros x y H. apply N.eqb_eq in H. now subst.
Qed.
Ltac natcst t :=
match isnatcst t with
- true => constr:(N_of_nat t)
+ true => constr:(N.of_nat t)
| _ => constr:InitialRing.NotConstant
end.
diff --git a/plugins/setoid_ring/BinList.v b/plugins/setoid_ring/BinList.v
index 7128280a0..2e317d784 100644
--- a/plugins/setoid_ring/BinList.v
+++ b/plugins/setoid_ring/BinList.v
@@ -6,11 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-Set Implicit Arguments.
Require Import BinPos.
Require Export List.
-Require Export ListTactics.
-Open Local Scope positive_scope.
+Set Implicit Arguments.
+Local Open Scope positive_scope.
Section MakeBinList.
Variable A : Type.
@@ -18,76 +17,64 @@ Section MakeBinList.
Fixpoint jump (p:positive) (l:list A) {struct p} : list A :=
match p with
- | xH => tail l
+ | xH => tl l
| xO p => jump p (jump p l)
- | xI p => jump p (jump p (tail l))
+ | xI p => jump p (jump p (tl l))
end.
Fixpoint nth (p:positive) (l:list A) {struct p} : A:=
match p with
| xH => hd default l
| xO p => nth p (jump p l)
- | xI p => nth p (jump p (tail l))
+ | xI p => nth p (jump p (tl l))
end.
- Lemma jump_tl : forall j l, tail (jump j l) = jump j (tail l).
+ Lemma jump_tl : forall j l, tl (jump j l) = jump j (tl l).
Proof.
- induction j;simpl;intros.
- repeat rewrite IHj;trivial.
- repeat rewrite IHj;trivial.
- trivial.
+ induction j;simpl;intros; now rewrite ?IHj.
Qed.
- Lemma jump_Psucc : forall j l,
- (jump (Psucc j) l) = (jump 1 (jump j l)).
+ Lemma jump_succ : forall j l,
+ jump (Pos.succ j) l = jump 1 (jump j l).
Proof.
induction j;simpl;intros.
- repeat rewrite IHj;simpl;repeat rewrite jump_tl;trivial.
- repeat rewrite jump_tl;trivial.
- trivial.
+ - rewrite !IHj; simpl; now rewrite !jump_tl.
+ - now rewrite !jump_tl.
+ - trivial.
Qed.
- Lemma jump_Pplus : forall i j l,
- (jump (i + j) l) = (jump i (jump j l)).
+ Lemma jump_add : forall i j l,
+ jump (i + j) l = jump i (jump j l).
Proof.
- induction i;intros.
- rewrite xI_succ_xO;rewrite Pplus_one_succ_r.
- rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
- repeat rewrite IHi.
- rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite jump_Psucc;trivial.
- rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
- repeat rewrite IHi;trivial.
- rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite jump_Psucc;trivial.
+ induction i using Pos.peano_ind; intros.
+ - now rewrite Pos.add_1_l, jump_succ.
+ - now rewrite Pos.add_succ_l, !jump_succ, IHi.
Qed.
- Lemma jump_Pdouble_minus_one : forall i l,
- (jump (Pdouble_minus_one i) (tail l)) = (jump i (jump i l)).
+ Lemma jump_pred_double : forall i l,
+ jump (Pos.pred_double i) (tl l) = jump i (jump i l).
Proof.
induction i;intros;simpl.
- repeat rewrite jump_tl;trivial.
- rewrite IHi. do 2 rewrite <- jump_tl;rewrite IHi;trivial.
- trivial.
+ - now rewrite !jump_tl.
+ - now rewrite IHi, <- 2 jump_tl, IHi.
+ - trivial.
Qed.
-
- Lemma nth_jump : forall p l, nth p (tail l) = hd default (jump p l).
+ Lemma nth_jump : forall p l, nth p (tl l) = hd default (jump p l).
Proof.
induction p;simpl;intros.
- rewrite <-jump_tl;rewrite IHp;trivial.
- rewrite <-jump_tl;rewrite IHp;trivial.
- trivial.
+ - now rewrite <-jump_tl, IHp.
+ - now rewrite <-jump_tl, IHp.
+ - trivial.
Qed.
- Lemma nth_Pdouble_minus_one :
- forall p l, nth (Pdouble_minus_one p) (tail l) = nth p (jump p l).
+ Lemma nth_pred_double :
+ forall p l, nth (Pos.pred_double p) (tl l) = nth p (jump p l).
Proof.
induction p;simpl;intros.
- repeat rewrite jump_tl;trivial.
- rewrite jump_Pdouble_minus_one.
- repeat rewrite <- jump_tl;rewrite IHp;trivial.
- trivial.
+ - now rewrite !jump_tl.
+ - now rewrite jump_pred_double, <- !jump_tl, IHp.
+ - trivial.
Qed.
End MakeBinList.
-
-
diff --git a/plugins/setoid_ring/Cring.v b/plugins/setoid_ring/Cring.v
index 3d6e53fcd..592efbf6d 100644
--- a/plugins/setoid_ring/Cring.v
+++ b/plugins/setoid_ring/Cring.v
@@ -42,10 +42,9 @@ Section cring.
Context {R:Type}`{Rr:Cring R}.
Lemma cring_eq_ext: ring_eq_ext _+_ _*_ -_ _==_.
-intros. apply mk_reqe;intros.
-rewrite H. rewrite H0. reflexivity.
-rewrite H. rewrite H0. reflexivity.
- rewrite H. reflexivity. Defined.
+Proof.
+intros. apply mk_reqe; solve_proper.
+Defined.
Lemma cring_almost_ring_theory:
almost_ring_theory (R:=R) zero one _+_ _*_ _-_ -_ _==_.
@@ -64,11 +63,11 @@ rewrite ring_sub_def ; reflexivity. Defined.
Lemma cring_morph:
ring_morph zero one _+_ _*_ _-_ -_ _==_
- 0%Z 1%Z Zplus Zmult Zminus Zopp Zeq_bool
+ 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool
Ncring_initial.gen_phiZ.
intros. apply mkmorph ; intros; simpl; try reflexivity.
rewrite Ncring_initial.gen_phiZ_add; reflexivity.
-rewrite ring_sub_def. unfold Zminus. rewrite Ncring_initial.gen_phiZ_add.
+rewrite ring_sub_def. unfold Z.sub. rewrite Ncring_initial.gen_phiZ_add.
rewrite Ncring_initial.gen_phiZ_opp; reflexivity.
rewrite Ncring_initial.gen_phiZ_mul; reflexivity.
rewrite Ncring_initial.gen_phiZ_opp; reflexivity.
@@ -80,7 +79,7 @@ Lemma cring_power_theory :
intros; apply Ring_theory.mkpow_th. reflexivity. Defined.
Lemma cring_div_theory:
- div_theory _==_ Zplus Zmult Ncring_initial.gen_phiZ Z.quotrem.
+ div_theory _==_ Z.add Z.mul Ncring_initial.gen_phiZ Z.quotrem.
intros. apply InitialRing.Ztriv_div_th. unfold Setoid_Theory.
simpl. apply ring_setoid. Defined.
@@ -102,7 +101,7 @@ Ltac cring_gen :=
ring_setoid
cring_eq_ext
cring_almost_ring_theory
- Z 0%Z 1%Z Zplus Zmult Zminus Zopp Zeq_bool
+ Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool
Ncring_initial.gen_phiZ
cring_morph
N
@@ -126,7 +125,7 @@ Ltac cring:=
cring_compute.
Instance Zcri: (Cring (Rr:=Zr)).
-red. exact Zmult_comm. Defined.
+red. exact Z.mul_comm. Defined.
(* Cring_simplify *)
@@ -136,7 +135,7 @@ Ltac cring_simplify_aux lterm fv lexpr hyp :=
match lexpr with
| ?e::?le =>
let t := constr:(@Ring_polynom.norm_subst
- Z 0%Z 1%Z Zplus Zmult Zminus Zopp Zeq_bool Z.quotrem O nil e) in
+ Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool Z.quotrem O nil e) in
let te :=
constr:(@Ring_polynom.Pphi_dev
_ 0 1 _+_ _*_ _-_ -_
@@ -149,7 +148,7 @@ Ltac cring_simplify_aux lterm fv lexpr hyp :=
let t':= fresh "t" in
pose (t' := nft);
assert (eq1 : t = t');
- [vm_cast_no_check (refl_equal t')|
+ [vm_cast_no_check (eq_refl t')|
let eq2 := fresh "ring" in
assert (eq2:(@Ring_polynom.PEeval
_ zero _+_ _*_ _-_ -_ Z Ncring_initial.gen_phiZ N (fun n:N => n)
@@ -159,7 +158,7 @@ Ltac cring_simplify_aux lterm fv lexpr hyp :=
ring_setoid
cring_eq_ext
cring_almost_ring_theory
- Z 0%Z 1%Z Zplus Zmult Zminus Zopp Zeq_bool
+ Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Zeq_bool
Ncring_initial.gen_phiZ
cring_morph
N
@@ -169,7 +168,7 @@ Ltac cring_simplify_aux lterm fv lexpr hyp :=
Z.quotrem
cring_div_theory
get_signZ get_signZ_th
- O nil fv I nil (refl_equal nil) );
+ O nil fv I nil (eq_refl nil) );
intro eq3; apply eq3; reflexivity|
match hyp with
| 1%nat => rewrite eq2
diff --git a/plugins/setoid_ring/Field_tac.v b/plugins/setoid_ring/Field_tac.v
index da42bbd95..056203926 100644
--- a/plugins/setoid_ring/Field_tac.v
+++ b/plugins/setoid_ring/Field_tac.v
@@ -447,7 +447,7 @@ Ltac prove_field_eqn ope FLD fv expr :=
pose (res' := res);
let lemma := get_L1 FLD in
let lemma :=
- constr:(lemma O fv List.nil expr' res' I List.nil (refl_equal _)) in
+ constr:(lemma O fv List.nil expr' res' I List.nil (eq_refl _)) in
let ty := type of lemma in
let lhs := match ty with
forall _, ?lhs=_ -> _ => lhs
@@ -487,7 +487,7 @@ Ltac reduce_field_expr ope kont FLD fv expr :=
kont c.
(* Hack to let a Ltac return a term in the context of a primitive tactic *)
-Ltac return_term x := generalize (refl_equal x).
+Ltac return_term x := generalize (eq_refl x).
Ltac get_term :=
match goal with
| |- ?x = _ -> _ => x
diff --git a/plugins/setoid_ring/Field_theory.v b/plugins/setoid_ring/Field_theory.v
index 40138526d..17595639b 100644
--- a/plugins/setoid_ring/Field_theory.v
+++ b/plugins/setoid_ring/Field_theory.v
@@ -7,7 +7,7 @@
(************************************************************************)
Require Ring.
-Import Ring_polynom Ring_tac Ring_theory InitialRing Setoid List.
+Import Ring_polynom Ring_tac Ring_theory InitialRing Setoid List Morphisms.
Require Import ZArith_base.
(*Require Import Omega.*)
Set Implicit Arguments.
@@ -27,7 +27,7 @@ Section MakeFieldPol.
Notation "x == y" := (req x y) (at level 70, no associativity).
(* Equality properties *)
- Variable Rsth : Setoid_Theory R req.
+ Variable Rsth : Equivalence req.
Variable Reqe : ring_eq_ext radd rmul ropp req.
Variable SRinv_ext : forall p q, p == q -> / p == / q.
@@ -75,7 +75,6 @@ Qed.
(* Useful tactics *)
- Add Setoid R req Rsth as R_set1.
Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed.
Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed.
Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed.
@@ -116,6 +115,7 @@ Notation NPphi_pow := (Pphi_pow rO rI radd rmul rsub ropp cO cI ceqb phi Cp_phi
(* add abstract semi-ring to help with some proofs *)
Add Ring Rring : (ARth_SRth ARth).
+Local Hint Extern 2 (_ == _) => f_equiv.
(* additional ring properties *)
@@ -135,6 +135,7 @@ Qed.
***************************************************************************)
Theorem rdiv_simpl: forall p q, ~ q == 0 -> q * (p / q) == p.
+Proof.
intros p q H.
rewrite rdiv_def in |- *.
transitivity (/ q * q * p); [ ring | idtac ].
@@ -142,35 +143,32 @@ rewrite rinv_l in |- *; auto.
Qed.
Hint Resolve rdiv_simpl .
-Theorem SRdiv_ext:
- forall p1 p2, p1 == p2 -> forall q1 q2, q1 == q2 -> p1 / q1 == p2 / q2.
-intros p1 p2 H q1 q2 H0.
+Instance SRdiv_ext: Proper (req ==> req ==> req) rdiv.
+Proof.
+intros p1 p2 Ep q1 q2 Eq.
transitivity (p1 * / q1); auto.
transitivity (p2 * / q2); auto.
Qed.
-Hint Resolve SRdiv_ext .
-
- Add Morphism rdiv : rdiv_ext. exact SRdiv_ext. Qed.
+Hint Resolve SRdiv_ext.
Lemma rmul_reg_l : forall p q1 q2,
~ p == 0 -> p * q1 == p * q2 -> q1 == q2.
-intros.
-rewrite <- (@rdiv_simpl q1 p) in |- *; trivial.
-rewrite <- (@rdiv_simpl q2 p) in |- *; trivial.
-repeat rewrite rdiv_def in |- *.
-repeat rewrite (ARmul_assoc ARth) in |- *.
-auto.
+Proof.
+intros p q1 q2 H EQ.
+rewrite <- (@rdiv_simpl q1 p) by trivial.
+rewrite <- (@rdiv_simpl q2 p) by trivial.
+rewrite !rdiv_def, !(ARmul_assoc ARth).
+now rewrite EQ.
Qed.
Theorem field_is_integral_domain : forall r1 r2,
~ r1 == 0 -> ~ r2 == 0 -> ~ r1 * r2 == 0.
Proof.
-red in |- *; intros.
-apply H0.
+intros r1 r2 H1 H2. contradict H2.
transitivity (1 * r2); auto.
transitivity (/ r1 * r1 * r2); auto.
-rewrite <- (ARmul_assoc ARth) in |- *.
-rewrite H1 in |- *.
+rewrite <- (ARmul_assoc ARth).
+rewrite H2.
apply ARmul_0_r with (1 := Rsth) (2 := ARth).
Qed.
@@ -205,12 +203,12 @@ Proof.
intros r1 r2 r3 r4 H H0.
assert (~ r2 * r4 == 0) by complete (apply field_is_integral_domain; trivial).
apply rmul_reg_l with (r2 * r4); trivial.
-rewrite rdiv_simpl in |- *; trivial.
-rewrite (ARdistr_r Rsth Reqe ARth) in |- *.
+rewrite rdiv_simpl; trivial.
+rewrite (ARdistr_r Rsth Reqe ARth).
apply (Radd_ext Reqe).
- transitivity (r2 * (r1 / r2) * r4); [ ring | auto ].
- transitivity (r2 * (r4 * (r3 / r4))); auto.
- transitivity (r2 * r3); auto.
+- transitivity (r2 * (r1 / r2) * r4); [ ring | auto ].
+- transitivity (r2 * (r4 * (r3 / r4))); auto.
+ transitivity (r2 * r3); auto.
Qed.
@@ -235,25 +233,26 @@ apply (Radd_ext Reqe).
Qed.
Theorem rdiv5: forall r1 r2, - (r1 / r2) == - r1 / r2.
+Proof.
intros r1 r2.
transitivity (- (r1 * / r2)); auto.
transitivity (- r1 * / r2); auto.
Qed.
Hint Resolve rdiv5 .
-Theorem rdiv3:
- forall r1 r2 r3 r4,
+Theorem rdiv3 r1 r2 r3 r4 :
~ r2 == 0 ->
~ r4 == 0 ->
r1 / r2 - r3 / r4 == (r1 * r4 - r3 * r2) / (r2 * r4).
-intros r1 r2 r3 r4 H H0.
+Proof.
+intros H2 H4.
assert (~ r2 * r4 == 0) by (apply field_is_integral_domain; trivial).
transitivity (r1 / r2 + - (r3 / r4)); auto.
transitivity (r1 / r2 + - r3 / r4); auto.
-transitivity ((r1 * r4 + - r3 * r2) / (r2 * r4)); auto.
+transitivity ((r1 * r4 + - r3 * r2) / (r2 * r4)).
apply rdiv2; auto.
-apply SRdiv_ext; auto.
-transitivity (r1 * r4 + - (r3 * r2)); symmetry; auto.
+f_equiv.
+transitivity (r1 * r4 + - (r3 * r2)); auto.
Qed.
@@ -410,14 +409,7 @@ Qed.
Add Morphism (pow_N rI rmul) with signature req ==> eq ==> req as pow_N_morph.
intros x y H [|p];simpl;auto. apply pow_morph;trivial.
Qed.
-(*
-Lemma rpow_morph : forall x y n, x == y ->rpow x (Cp_phi n) == rpow y (Cp_phi n).
-Proof.
- intros; repeat rewrite pow_th.(rpow_pow_N).
- destruct n;simpl. apply eq_refl.
- induction p;simpl;try rewrite IHp;try rewrite H; apply eq_refl.
-Qed.
-*)
+
Theorem PExpr_eq_semi_correct:
forall l e1 e2, PExpr_eq e1 e2 = true -> NPEeval l e1 == NPEeval l e2.
intros l e1; elim e1.
@@ -705,10 +697,10 @@ Fixpoint isIn (e1:PExpr C) (p1:positive)
end
end
| PEpow e3 N0 => None
- | PEpow e3 (Npos p3) => isIn e1 p1 e3 (Pmult p3 p2)
+ | PEpow e3 (Npos p3) => isIn e1 p1 e3 (Pos.mul p3 p2)
| _ =>
if PExpr_eq e1 e2 then
- match Zminus (Zpos p1) (Zpos p2) with
+ match Z.pos_sub p1 p2 with
| Zpos p => Some (Npos p, PEc cI)
| Z0 => Some (N0, PEc cI)
| Zneg p => Some (N0, NPEpow e2 (Npos p))
@@ -719,21 +711,19 @@ Fixpoint isIn (e1:PExpr C) (p1:positive)
Definition ZtoN z := match z with Zpos p => Npos p | _ => N0 end.
Definition NtoZ n := match n with Npos p => Zpos p | _ => Z0 end.
- Notation pow_pos_plus := (Ring_theory.pow_pos_Pplus _ Rsth Reqe.(Rmul_ext)
- ARth.(ARmul_comm) ARth.(ARmul_assoc)).
+ Notation pow_pos_add :=
+ (Ring_theory.pow_pos_add Rsth Reqe.(Rmul_ext) ARth.(ARmul_assoc)).
- Lemma Z_pos_sub_gt : forall p q, (p > q)%positive ->
+ Lemma Z_pos_sub_gt p q : (p > q)%positive ->
Z.pos_sub p q = Zpos (p - q).
- Proof.
- intros. apply Z.pos_sub_gt. now apply Pos.gt_lt.
- Qed.
+ Proof. intros; now apply Z.pos_sub_gt, Pos.gt_lt. Qed.
Ltac simpl_pos_sub := rewrite ?Z_pos_sub_gt in * by assumption.
Lemma isIn_correct_aux : forall l e1 e2 p1 p2,
match
(if PExpr_eq e1 e2 then
- match Zminus (Zpos p1) (Zpos p2) with
+ match Z.sub (Zpos p1) (Zpos p2) with
| Zpos p => Some (Npos p, PEc cI)
| Z0 => Some (N0, PEc cI)
| Zneg p => Some (N0, NPEpow e2 (Npos p))
@@ -750,33 +740,28 @@ Proof.
intros l e1 e2 p1 p2; generalize (PExpr_eq_semi_correct l e1 e2);
case (PExpr_eq e1 e2); simpl; auto; intros H.
rewrite Z.pos_sub_spec.
- case_eq ((p1 ?= p2)%positive);intros;simpl.
- repeat rewrite pow_th.(rpow_pow_N);simpl. split. 2:refine (refl_equal _).
- rewrite (Pcompare_Eq_eq _ _ H0).
- rewrite H by trivial. ring [ (morph1 CRmorph)].
- fold (p2 - p1 =? 1)%positive.
- fold (NPEpow e2 (Npos (p2 - p1))).
- rewrite NPEpow_correct;simpl.
- repeat rewrite pow_th.(rpow_pow_N);simpl.
- rewrite H;trivial. split. 2:refine (refl_equal _).
- rewrite <- pow_pos_plus; rewrite Pplus_minus;auto. apply ZC2;trivial.
- repeat rewrite pow_th.(rpow_pow_N);simpl.
- rewrite H;trivial.
- change (Z.pos_sub p1 (p1-p2)) with (Zpos p1 - Zpos (p1 -p2))%Z.
- replace (Zpos (p1 - p2)) with (Zpos p1 - Zpos p2)%Z.
- split.
- repeat rewrite Zth.(Rsub_def). rewrite (Ring_theory.Ropp_add Zsth Zeqe Zth).
- rewrite Zplus_assoc, Z.add_opp_diag_r. simpl.
- ring [ (morph1 CRmorph)].
- assert (Zpos p1 > 0 /\ Zpos p2 > 0)%Z. split;refine (refl_equal _).
- apply Zplus_gt_reg_l with (Zpos p2).
- rewrite Zplus_minus. change (Zpos p2 + Zpos p1 > 0 + Zpos p1)%Z.
- apply Zplus_gt_compat_r. refine (refl_equal _).
- simpl. now simpl_pos_sub.
+ case Pos.compare_spec;intros;simpl.
+ - repeat rewrite pow_th.(rpow_pow_N);simpl. split. 2:reflexivity.
+ subst. rewrite H by trivial. ring [ (morph1 CRmorph)].
+ - fold (p2 - p1 =? 1)%positive.
+ fold (NPEpow e2 (Npos (p2 - p1))).
+ rewrite NPEpow_correct;simpl.
+ repeat rewrite pow_th.(rpow_pow_N);simpl.
+ rewrite H;trivial. split. 2:reflexivity.
+ rewrite <- pow_pos_add. now rewrite Pos.add_comm, Pos.sub_add.
+ - repeat rewrite pow_th.(rpow_pow_N);simpl.
+ rewrite H;trivial.
+ rewrite Z.pos_sub_gt by now apply Pos.sub_decr.
+ replace (p1 - (p1 - p2))%positive with p2;
+ [| rewrite Pos.sub_sub_distr, Pos.add_comm;
+ auto using Pos.add_sub, Pos.sub_decr ].
+ split.
+ simpl. ring [ (morph1 CRmorph)].
+ now apply Z.lt_gt, Pos.sub_decr.
Qed.
Lemma pow_pos_pow_pos : forall x p1 p2, pow_pos rmul (pow_pos rmul x p1) p2 == pow_pos rmul x (p1*p2).
-induction p1;simpl;intros;repeat rewrite pow_pos_mul;repeat rewrite pow_pos_plus;simpl.
+induction p1;simpl;intros;repeat rewrite pow_pos_mul;repeat rewrite pow_pos_add;simpl.
ring [(IHp1 p2)]. ring [(IHp1 p2)]. auto.
Qed.
@@ -808,8 +793,9 @@ destruct n.
(pow_pos rmul (NPEeval l e1) p4 * NPEeval l p5) ==
pow_pos rmul (NPEeval l e1) p4 * pow_pos rmul (NPEeval l e1) (p1 - p4) *
NPEeval l p3 *NPEeval l p5) by ring. rewrite H;clear H.
- rewrite <- pow_pos_plus. rewrite Pplus_minus.
- split. symmetry;apply ARth.(ARmul_assoc). refine (refl_equal _). trivial.
+ rewrite <- pow_pos_add.
+ rewrite Pos.add_comm, Pos.sub_add by (now apply Z.gt_lt in H4).
+ split. symmetry;apply ARth.(ARmul_assoc). reflexivity.
repeat rewrite pow_th.(rpow_pow_N);simpl.
intros (H1,H2) (H3,H4).
simpl_pos_sub. simpl in H1, H3.
@@ -822,15 +808,15 @@ destruct n.
(pow_pos rmul (NPEeval l e1) (p4 - p6) * NPEeval l p5) ==
pow_pos rmul (NPEeval l e1) (p1 - p4) * pow_pos rmul (NPEeval l e1) (p4 - p6) *
NPEeval l p3 * NPEeval l p5) by ring. rewrite H0;clear H0.
- rewrite <- pow_pos_plus.
+ rewrite <- pow_pos_add.
replace (p1 - p4 + (p4 - p6))%positive with (p1 - p6)%positive.
rewrite NPEmul_correct. simpl;ring.
assert
(Zpos p1 - Zpos p6 = Zpos p1 - Zpos p4 + (Zpos p4 - Zpos p6))%Z.
change ((Zpos p1 - Zpos p6)%Z = (Zpos p1 + (- Zpos p4) + (Zpos p4 +(- Zpos p6)))%Z).
- rewrite <- Zplus_assoc. rewrite (Zplus_assoc (- Zpos p4)).
+ rewrite <- Z.add_assoc. rewrite (Z.add_assoc (- Zpos p4)).
simpl. rewrite Z.pos_sub_diag. simpl. reflexivity.
- unfold Zminus, Zopp in H0. simpl in H0.
+ unfold Z.sub, Z.opp in H0. simpl in H0.
simpl_pos_sub. inversion H0; trivial.
simpl. repeat rewrite pow_th.(rpow_pow_N).
intros H1 (H2,H3). simpl_pos_sub.
@@ -875,7 +861,7 @@ Fixpoint split_aux (e1: PExpr C) (p:positive) (e2:PExpr C) {struct e1}: rsplit :
(NPEmul (common r1) (common r2))
(right r2)
| PEpow e3 N0 => mk_rsplit (PEc cI) (PEc cI) e2
- | PEpow e3 (Npos p3) => split_aux e3 (Pmult p3 p) e2
+ | PEpow e3 (Npos p3) => split_aux e3 (Pos.mul p3 p) e2
| _ =>
match isIn e1 p e2 xH with
| Some (N0,e3) => mk_rsplit (PEc cI) (NPEpow e1 (Npos p)) e3
@@ -903,7 +889,8 @@ Proof.
repeat rewrite pow_th.(rpow_pow_N);simpl).
intros (H, Hgt);split;try ring [H CRmorph.(morph1)].
intros (H, Hgt). simpl_pos_sub. simpl in H;split;try ring [H].
- rewrite <- pow_pos_plus. rewrite Pplus_minus. reflexivity. trivial.
+ apply Z.gt_lt in Hgt.
+ now rewrite <- pow_pos_add, Pos.add_comm, Pos.sub_add.
simpl;intros. repeat rewrite NPEmul_correct;simpl.
rewrite NPEpow_correct;simpl. split;ring [CRmorph.(morph1)].
Qed.
@@ -1319,9 +1306,9 @@ apply Fnorm_crossproduct; trivial.
match goal with
[ |- NPEeval l ?x == NPEeval l ?y] =>
rewrite (ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec
- O nil l I (refl_equal nil) x (refl_equal (Nnorm O nil x)));
+ O nil l I Logic.eq_refl x Logic.eq_refl);
rewrite (ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec
- O nil l I (refl_equal nil) y (refl_equal (Nnorm O nil y)))
+ O nil l I Logic.eq_refl y Logic.eq_refl)
end.
trivial.
Qed.
@@ -1352,15 +1339,15 @@ rewrite <-(
let x := PEmul (num (Fnorm fe1))
(rsplit_right (split (denum (Fnorm fe1)) (denum (Fnorm fe2)))) in
ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
- Hlpe (refl_equal (Nmk_monpol_list lpe))
- x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+ Hlpe Logic.eq_refl
+ x Logic.eq_refl) in Hcrossprod.
rewrite <-(
let x := (PEmul (num (Fnorm fe2))
(rsplit_left
(split (denum (Fnorm fe1)) (denum (Fnorm fe2))))) in
ring_rw_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
- Hlpe (refl_equal (Nmk_monpol_list lpe))
- x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+ Hlpe Logic.eq_refl
+ x Logic.eq_refl) in Hcrossprod.
simpl in Hcrossprod.
rewrite Hcrossprod in |- *.
reflexivity.
@@ -1392,15 +1379,15 @@ rewrite <-(
let x := PEmul (num (Fnorm fe1))
(rsplit_right (split (denum (Fnorm fe1)) (denum (Fnorm fe2)))) in
ring_rw_pow_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
- Hlpe (refl_equal (Nmk_monpol_list lpe))
- x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+ Hlpe Logic.eq_refl
+ x Logic.eq_refl) in Hcrossprod.
rewrite <-(
let x := (PEmul (num (Fnorm fe2))
(rsplit_left
(split (denum (Fnorm fe1)) (denum (Fnorm fe2))))) in
ring_rw_pow_correct Rsth Reqe ARth CRmorph pow_th cdiv_th get_sign_spec n lpe l
- Hlpe (refl_equal (Nmk_monpol_list lpe))
- x (refl_equal (Nnorm n (Nmk_monpol_list lpe) x))) in Hcrossprod.
+ Hlpe Logic.eq_refl
+ x Logic.eq_refl) in Hcrossprod.
simpl in Hcrossprod.
rewrite Hcrossprod in |- *.
reflexivity.
@@ -1756,7 +1743,7 @@ Hypothesis gen_phiPOS_not_0 : forall p, ~ gen_phiPOS1 rI radd rmul p == 0.
Lemma add_inj_r : forall p x y,
gen_phiPOS1 rI radd rmul p + x == gen_phiPOS1 rI radd rmul p + y -> x==y.
intros p x y.
-elim p using Pind; simpl in |- *; intros.
+elim p using Pos.peano_ind; simpl; intros.
apply S_inj; trivial.
apply H.
apply S_inj.
@@ -1775,18 +1762,16 @@ case (Pos.compare_spec x y).
intros.
elim gen_phiPOS_not_0 with (y - x)%positive.
apply add_inj_r with x.
- symmetry in |- *.
- rewrite (ARadd_0_r Rsth ARth) in |- *.
- rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth) in |- *.
- rewrite Pplus_minus in |- *; trivial.
- now apply Pos.lt_gt.
+ symmetry.
+ rewrite (ARadd_0_r Rsth ARth).
+ rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth).
+ now rewrite Pos.add_comm, Pos.sub_add.
intros.
elim gen_phiPOS_not_0 with (x - y)%positive.
apply add_inj_r with y.
- rewrite (ARadd_0_r Rsth ARth) in |- *.
- rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth) in |- *.
- rewrite Pplus_minus in |- *; trivial.
- now apply Pos.lt_gt.
+ rewrite (ARadd_0_r Rsth ARth).
+ rewrite <- (ARgen_phiPOS_add Rsth Reqe ARth).
+ now rewrite Pos.add_comm, Pos.sub_add.
Qed.
@@ -1897,7 +1882,7 @@ Lemma gen_phiZ_complete : forall x y,
intros.
replace y with x.
unfold Zeq_bool in |- *.
- rewrite Zcompare_refl in |- *; trivial.
+ rewrite Z.compare_refl in |- *; trivial.
apply gen_phiZ_inj; trivial.
Qed.
diff --git a/plugins/setoid_ring/InitialRing.v b/plugins/setoid_ring/InitialRing.v
index 763dbe7b9..bc0f888ce 100644
--- a/plugins/setoid_ring/InitialRing.v
+++ b/plugins/setoid_ring/InitialRing.v
@@ -27,14 +27,14 @@ Definition NotConstant := false.
Lemma Zsth : Setoid_Theory Z (@eq Z).
Proof (Eqsth Z).
-Lemma Zeqe : ring_eq_ext Zplus Zmult Zopp (@eq Z).
-Proof (Eq_ext Zplus Zmult Zopp).
+Lemma Zeqe : ring_eq_ext Z.add Z.mul Z.opp (@eq Z).
+Proof (Eq_ext Z.add Z.mul Z.opp).
-Lemma Zth : ring_theory Z0 (Zpos xH) Zplus Zmult Zminus Zopp (@eq Z).
+Lemma Zth : ring_theory Z0 (Zpos xH) Z.add Z.mul Z.sub Z.opp (@eq Z).
Proof.
- constructor. exact Zplus_0_l. exact Zplus_comm. exact Zplus_assoc.
- exact Zmult_1_l. exact Zmult_comm. exact Zmult_assoc.
- exact Zmult_plus_distr_l. trivial. exact Zminus_diag.
+ constructor. exact Z.add_0_l. exact Z.add_comm. exact Z.add_assoc.
+ exact Z.mul_1_l. exact Z.mul_comm. exact Z.mul_assoc.
+ exact Z.mul_add_distr_r. trivial. exact Z.sub_diag.
Qed.
(** Two generic morphisms from Z to (abrbitrary) rings, *)
@@ -92,12 +92,12 @@ Section ZMORPHISM.
| _ => None
end.
- Lemma get_signZ_th : sign_theory Zopp Zeq_bool get_signZ.
+ Lemma get_signZ_th : sign_theory Z.opp Zeq_bool get_signZ.
Proof.
constructor.
destruct c;intros;try discriminate.
injection H;clear H;intros H1;subst c'.
- simpl. unfold Zeq_bool. rewrite Zcompare_refl. trivial.
+ simpl. unfold Zeq_bool. rewrite Z.compare_refl. trivial.
Qed.
@@ -116,7 +116,7 @@ Section ZMORPHISM.
Qed.
Lemma ARgen_phiPOS_Psucc : forall x,
- gen_phiPOS1 (Psucc x) == 1 + (gen_phiPOS1 x).
+ gen_phiPOS1 (Pos.succ x) == 1 + (gen_phiPOS1 x).
Proof.
induction x;simpl;norm.
rewrite IHx;norm.
@@ -127,7 +127,7 @@ Section ZMORPHISM.
gen_phiPOS1 (x + y) == (gen_phiPOS1 x) + (gen_phiPOS1 y).
Proof.
induction x;destruct y;simpl;norm.
- rewrite Pplus_carry_spec.
+ rewrite Pos.add_carry_spec.
rewrite ARgen_phiPOS_Psucc.
rewrite IHx;norm.
add_push (gen_phiPOS1 y);add_push 1;rrefl.
@@ -208,10 +208,10 @@ Section ZMORPHISM.
(*proof that [.] satisfies morphism specifications*)
Lemma gen_phiZ_morph :
ring_morph 0 1 radd rmul rsub ropp req Z0 (Zpos xH)
- Zplus Zmult Zminus Zopp Zeq_bool gen_phiZ.
+ Z.add Z.mul Z.sub Z.opp Zeq_bool gen_phiZ.
Proof.
assert ( SRmorph : semi_morph 0 1 radd rmul req Z0 (Zpos xH)
- Zplus Zmult Zeq_bool gen_phiZ).
+ Z.add Z.mul Zeq_bool gen_phiZ).
apply mkRmorph;simpl;try rrefl.
apply gen_phiZ_add. apply gen_phiZ_mul. apply gen_Zeqb_ok.
apply (Smorph_morph Rsth Reqe Rth Zth SRmorph gen_phiZ_ext).
@@ -741,10 +741,10 @@ Ltac gen_ring_sign morph sspec :=
Ltac default_div_spec set reqe arth morph :=
match type of morph with
| @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
- Z ?c0 ?c1 Zplus Zmult ?csub ?copp ?ceq_b ?phi =>
+ Z ?c0 ?c1 Z.add Z.mul ?csub ?copp ?ceq_b ?phi =>
constr:(mkhypo (Ztriv_div_th set phi))
| @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
- N ?c0 ?c1 Nplus Nmult ?csub ?copp ?ceq_b ?phi =>
+ N ?c0 ?c1 N.add N.mul ?csub ?copp ?ceq_b ?phi =>
constr:(mkhypo (Ntriv_div_th set phi))
| @ring_morph ?R ?r0 ?rI ?radd ?rmul ?rsub ?ropp ?req
?C ?c0 ?c1 ?cadd ?cmul ?csub ?copp ?ceq_b ?phi =>
@@ -836,7 +836,7 @@ Ltac isPcst t :=
| xO ?p => isPcst p
| xH => constr:true
(* nat -> positive *)
- | P_of_succ_nat ?n => isnatcst n
+ | Pos.of_succ_nat ?n => isnatcst n
| _ => constr:false
end.
@@ -853,9 +853,9 @@ Ltac isZcst t :=
| Zpos ?p => isPcst p
| Zneg ?p => isPcst p
(* injection nat -> Z *)
- | Z_of_nat ?n => isnatcst n
+ | Z.of_nat ?n => isnatcst n
(* injection N -> Z *)
- | Z_of_N ?n => isNcst n
+ | Z.of_N ?n => isNcst n
(* *)
| _ => constr:false
end.
diff --git a/plugins/setoid_ring/Integral_domain.v b/plugins/setoid_ring/Integral_domain.v
index 5a224e38e..0c16fe1a3 100644
--- a/plugins/setoid_ring/Integral_domain.v
+++ b/plugins/setoid_ring/Integral_domain.v
@@ -19,7 +19,7 @@ rewrite H0. rewrite <- H. cring.
Qed.
-Definition pow (r : R) (n : nat) := Ring_theory.pow_N 1 mul r (N_of_nat n).
+Definition pow (r : R) (n : nat) := Ring_theory.pow_N 1 mul r (N.of_nat n).
Lemma pow_not_zero: forall p n, pow p n == 0 -> p == 0.
induction n. unfold pow; simpl. intros. absurd (1 == 0).
@@ -29,9 +29,8 @@ intros.
case (integral_domain_product p (pow p n) H). trivial. trivial.
unfold pow; simpl.
clear IHn. induction n; simpl; try cring.
- rewrite Ring_theory.pow_pos_Psucc. cring. apply ring_setoid.
+ rewrite Ring_theory.pow_pos_succ. cring. apply ring_setoid.
apply ring_mult_comp.
-apply cring_mul_comm.
apply ring_mul_assoc.
Qed.
diff --git a/plugins/setoid_ring/Ncring.v b/plugins/setoid_ring/Ncring.v
index 9a30fa47e..c093716b3 100644
--- a/plugins/setoid_ring/Ncring.v
+++ b/plugins/setoid_ring/Ncring.v
@@ -106,9 +106,10 @@ Context {R:Type}`{Rr:Ring R}.
(* Powers *)
- Lemma pow_pos_comm : forall x j, x * pow_pos x j == pow_pos x j * x.
+Lemma pow_pos_comm : forall x j, x * pow_pos x j == pow_pos x j * x.
+Proof.
induction j; simpl. rewrite <- ring_mul_assoc.
-rewrite <- ring_mul_assoc.
+rewrite <- ring_mul_assoc.
rewrite <- IHj. rewrite (ring_mul_assoc (pow_pos x j) x (pow_pos x j)).
rewrite <- IHj. rewrite <- ring_mul_assoc. reflexivity.
rewrite <- ring_mul_assoc. rewrite <- IHj.
@@ -116,10 +117,10 @@ rewrite ring_mul_assoc. rewrite IHj.
rewrite <- ring_mul_assoc. rewrite IHj. reflexivity. reflexivity.
Qed.
- Lemma pow_pos_Psucc : forall x j, pow_pos x (Psucc j) == x * pow_pos x j.
- Proof.
- induction j; simpl.
- rewrite IHj.
+Lemma pow_pos_succ : forall x j, pow_pos x (Pos.succ j) == x * pow_pos x j.
+Proof.
+induction j; simpl.
+ rewrite IHj.
rewrite <- (ring_mul_assoc x (pow_pos x j) (x * pow_pos x j)).
rewrite (ring_mul_assoc (pow_pos x j) x (pow_pos x j)).
rewrite <- pow_pos_comm.
@@ -127,20 +128,20 @@ rewrite <- ring_mul_assoc. reflexivity.
reflexivity. reflexivity.
Qed.
- Lemma pow_pos_Pplus : forall x i j,
- pow_pos x (i + j) == pow_pos x i * pow_pos x j.
- Proof.
+Lemma pow_pos_add : forall x i j,
+ pow_pos x (i + j) == pow_pos x i * pow_pos x j.
+Proof.
intro x;induction i;intros.
- rewrite xI_succ_xO;rewrite Pplus_one_succ_r.
- rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
+ rewrite Pos.xI_succ_xO;rewrite <- Pos.add_1_r.
+ rewrite <- Pos.add_diag;repeat rewrite <- Pos.add_assoc.
repeat rewrite IHi.
- rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;
- rewrite pow_pos_Psucc.
+ rewrite Pos.add_comm;rewrite Pos.add_1_r;
+ rewrite pow_pos_succ.
simpl;repeat rewrite ring_mul_assoc. reflexivity.
- rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
+ rewrite <- Pos.add_diag;repeat rewrite <- Pos.add_assoc.
repeat rewrite IHi. rewrite ring_mul_assoc. reflexivity.
- rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite pow_pos_Psucc.
- simpl. reflexivity.
+ rewrite Pos.add_comm;rewrite Pos.add_1_r;rewrite pow_pos_succ.
+ simpl. reflexivity.
Qed.
Definition id_phi_N (x:N) : N := x.
diff --git a/plugins/setoid_ring/Ncring_initial.v b/plugins/setoid_ring/Ncring_initial.v
index 3c79f7d9b..90fd82054 100644
--- a/plugins/setoid_ring/Ncring_initial.v
+++ b/plugins/setoid_ring/Ncring_initial.v
@@ -27,20 +27,17 @@ Definition NotConstant := false.
(** Z is a ring and a setoid*)
-Lemma Zsth : Setoid_Theory Z (@eq Z).
-constructor;red;intros;subst;trivial.
-Qed.
+Lemma Zsth : Equivalence (@eq Z).
+Proof. exact Z.eq_equiv. Qed.
-Instance Zops:@Ring_ops Z 0%Z 1%Z Zplus Zmult Zminus Zopp (@eq Z).
+Instance Zops:@Ring_ops Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (@eq Z).
Instance Zr: (@Ring _ _ _ _ _ _ _ _ Zops).
-constructor;
-try (try apply Zsth;
- try (unfold respectful, Proper; unfold equality; unfold eq_notation in *;
- intros; try rewrite H; try rewrite H0; reflexivity)).
- exact Zplus_comm. exact Zplus_assoc.
- exact Zmult_1_l. exact Zmult_1_r. exact Zmult_assoc.
- exact Zmult_plus_distr_l. intros; apply Zmult_plus_distr_r. exact Zminus_diag.
+Proof.
+constructor; try apply Zsth; try solve_proper.
+ exact Z.add_comm. exact Z.add_assoc.
+ exact Z.mul_1_l. exact Z.mul_1_r. exact Z.mul_assoc.
+ exact Z.mul_add_distr_r. intros; apply Z.mul_add_distr_l. exact Z.sub_diag.
Defined.
(*Instance ZEquality: @Equality Z:= (@eq Z).*)
@@ -102,7 +99,7 @@ Ltac rsimpl := simpl.
Qed.
Lemma ARgen_phiPOS_Psucc : forall x,
- gen_phiPOS1 (Psucc x) == 1 + (gen_phiPOS1 x).
+ gen_phiPOS1 (Pos.succ x) == 1 + (gen_phiPOS1 x).
Proof.
induction x;rsimpl;norm.
rewrite IHx. gen_rewrite. add_push 1. reflexivity.
@@ -112,7 +109,7 @@ Ltac rsimpl := simpl.
gen_phiPOS1 (x + y) == (gen_phiPOS1 x) + (gen_phiPOS1 y).
Proof.
induction x;destruct y;simpl;norm.
- rewrite Pplus_carry_spec.
+ rewrite Pos.add_carry_spec.
rewrite ARgen_phiPOS_Psucc.
rewrite IHx;norm.
add_push (gen_phiPOS1 y);add_push 1;reflexivity.
@@ -152,20 +149,13 @@ Ltac rsimpl := simpl.
== gen_phiPOS1 x + -gen_phiPOS1 y.
Proof.
intros x y.
- rewrite Z.pos_sub_spec.
- assert (HH0 := Pminus_mask_Gt x y). unfold Pos.gt in HH0.
- assert (HH1 := Pminus_mask_Gt y x). unfold Pos.gt in HH1.
- rewrite Pos.compare_antisym in HH1.
- destruct (Pos.compare_spec x y) as [HH|HH|HH].
- subst. rewrite ring_opp_def;reflexivity.
- destruct HH1 as [h [HHeq1 [HHeq2 HHor]]];trivial.
- unfold Pminus; rewrite HHeq1;rewrite <- HHeq2.
- rewrite ARgen_phiPOS_add;simpl;norm.
- rewrite ring_opp_def;norm.
- destruct HH0 as [h [HHeq1 [HHeq2 HHor]]];trivial.
- unfold Pminus; rewrite HHeq1;rewrite <- HHeq2.
- rewrite ARgen_phiPOS_add;simpl;norm.
- add_push (gen_phiPOS1 h). rewrite ring_opp_def ; norm.
+ generalize (Z.pos_sub_discr x y).
+ destruct (Z.pos_sub x y) as [|p|p]; intros; subst.
+ - now rewrite ring_opp_def.
+ - rewrite ARgen_phiPOS_add;simpl;norm.
+ add_push (gen_phiPOS1 p). rewrite ring_opp_def;norm.
+ - rewrite ARgen_phiPOS_add;simpl;norm.
+ rewrite ring_opp_def;norm.
Qed.
Lemma match_compOpp : forall x (B:Type) (be bl bg:B),
@@ -206,16 +196,14 @@ Lemma gen_phiZ_opp : forall x, [- x] == - [x].
Global Instance gen_phiZ_morph :
(@Ring_morphism (Z:Type) R _ _ _ _ _ _ _ Zops Zr _ _ _ _ _ _ _ _ _ gen_phiZ) . (* beurk!*)
apply Build_Ring_morphism; simpl;try reflexivity.
- apply gen_phiZ_add. intros. rewrite ring_sub_def.
-replace (Zminus x y) with (x + (-y))%Z. rewrite gen_phiZ_add.
-rewrite gen_phiZ_opp. rewrite ring_sub_def. reflexivity.
+ apply gen_phiZ_add. intros. rewrite ring_sub_def.
+replace (x-y)%Z with (x + (-y))%Z.
+now rewrite gen_phiZ_add, gen_phiZ_opp, ring_sub_def.
reflexivity.
- apply gen_phiZ_mul. apply gen_phiZ_opp. apply gen_phiZ_ext.
+ apply gen_phiZ_mul. apply gen_phiZ_opp. apply gen_phiZ_ext.
Defined.
End ZMORPHISM.
Instance multiplication_phi_ring{R:Type}`{Ring R} : Multiplication :=
{multiplication x y := (gen_phiZ x) * y}.
-
-
diff --git a/plugins/setoid_ring/Ncring_polynom.v b/plugins/setoid_ring/Ncring_polynom.v
index 17dd2daa3..3c2900c39 100644
--- a/plugins/setoid_ring/Ncring_polynom.v
+++ b/plugins/setoid_ring/Ncring_polynom.v
@@ -52,7 +52,7 @@ Instance equalityb_coef : Equalityb C :=
match P, P' with
| Pc c, Pc c' => c =? c'
| PX P i n Q, PX P' i' n' Q' =>
- match Pcompare i i' Eq, Pcompare n n' Eq with
+ match Pos.compare i i', Pos.compare n n' with
| Eq, Eq => if Peq P P' then Peq Q Q' else false
| _,_ => false
end
@@ -67,7 +67,7 @@ Instance equalityb_pol : Equalityb Pol :=
match P with
| Pc c => if c =? 0 then Q else PX P i n Q
| PX P' i' n' Q' =>
- match Pcompare i i' Eq with
+ match Pos.compare i i' with
| Eq => if Q' =? P0 then PX P' i (n + n') Q else PX P i n Q
| _ => PX P i n Q
end
@@ -109,13 +109,13 @@ Fixpoint PaddX (i n:positive)(Q:Pol){struct Q}:=
match Q with
| Pc c => mkPX P i n Q
| PX P' i' n' Q' =>
- match Pcompare i i' Eq with
+ match Pos.compare i i' with
| (* i > i' *)
Gt => mkPX P i n Q
| (* i < i' *)
Lt => mkPX P' i' n' (PaddX i n Q')
| (* i = i' *)
- Eq => match ZPminus n n' with
+ Eq => match Z.pos_sub n n' with
| (* n > n' *)
Zpos k => mkPX (PaddX i k P') i' n' Q'
| (* n = n' *)
@@ -178,61 +178,25 @@ Definition Psub(P P':Pol):= P ++ (--P').
Reserved Notation "P @ l " (at level 10, no associativity).
Notation "P @ l " := (Pphi l P).
+
(** Proofs *)
- Lemma ZPminus_spec : forall x y,
- match ZPminus x y with
- | Z0 => x = y
- | Zpos k => x = (y + k)%positive
- | Zneg k => y = (x + k)%positive
+
+ Ltac destr_pos_sub H :=
+ match goal with |- context [Z.pos_sub ?x ?y] =>
+ assert (H := Z.pos_sub_discr x y); destruct (Z.pos_sub x y)
end.
- Proof.
- induction x;destruct y.
- replace (ZPminus (xI x) (xI y)) with (Zdouble (ZPminus x y));trivial.
- assert (Hh := IHx y);destruct (ZPminus x y);unfold Zdouble;
-rewrite Hh;trivial.
- replace (ZPminus (xI x) (xO y)) with (Zdouble_plus_one (ZPminus x y));
-trivial.
- assert (Hh := IHx y);destruct (ZPminus x y);unfold Zdouble_plus_one;
-rewrite Hh;trivial.
- apply Pplus_xI_double_minus_one.
- simpl;trivial.
- replace (ZPminus (xO x) (xI y)) with (Zdouble_minus_one (ZPminus x y));
-trivial.
- assert (Hh := IHx y);destruct (ZPminus x y);unfold Zdouble_minus_one;
-rewrite Hh;trivial.
- apply Pplus_xI_double_minus_one.
- replace (ZPminus (xO x) (xO y)) with (Zdouble (ZPminus x y));trivial.
- assert (Hh := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite Hh;
-trivial.
- replace (ZPminus (xO x) xH) with (Zpos (Pdouble_minus_one x));trivial.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO;trivial.
- replace (ZPminus xH (xI y)) with (Zneg (xO y));trivial.
- replace (ZPminus xH (xO y)) with (Zneg (Pdouble_minus_one y));trivial.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO;trivial.
- simpl;trivial.
- Qed.
Lemma Peq_ok : forall P P',
(P =? P') = true -> forall l, P@l == P'@ l.
Proof.
- induction P;destruct P';simpl;intros;try discriminate;trivial.
- apply ring_morphism_eq.
- apply Ceqb_eq ;trivial.
- assert (H1h := IHP1 P'1);assert (H2h := IHP2 P'2).
- simpl in H1h. destruct (Peq P2 P'1). simpl in H2h;
-destruct (Peq P3 P'2).
- rewrite (H1h);trivial . rewrite (H2h);trivial.
-assert (H3h := Pcompare_Eq_eq p p1);
- destruct (Pos.compare_cont p p1 Eq);
-assert (H4h := Pcompare_Eq_eq p0 p2);
-destruct (Pos.compare_cont p0 p2 Eq); try (discriminate H).
- rewrite H3h;trivial. rewrite H4h;trivial. reflexivity.
- destruct (Pos.compare_cont p p1 Eq); destruct (Pos.compare_cont p0 p2 Eq);
- try (discriminate H).
- destruct (Pos.compare_cont p p1 Eq); destruct (Pos.compare_cont p0 p2 Eq);
- try (discriminate H).
+ induction P;destruct P';simpl;intros ;try easy.
+ - now apply ring_morphism_eq, Ceqb_eq.
+ - specialize (IHP1 P'1). specialize (IHP2 P'2).
+ simpl in IHP1, IHP2.
+ destruct (Pos.compare_spec p p1); try discriminate;
+ destruct (Pos.compare_spec p0 p2); try discriminate.
+ destruct (Peq P2 P'1); try discriminate.
+ subst; now rewrite IHP1, IHP2.
Qed.
Lemma Pphi0 : forall l, P0@l == 0.
@@ -255,12 +219,12 @@ destruct (Pos.compare_cont p0 p2 Eq); try (discriminate H).
simpl; case_eq (Ceqb c 0);simpl;try reflexivity.
intros.
rewrite Hh. rewrite ring_morphism0.
- rsimpl. apply Ceqb_eq. trivial. assert (Hh1 := Pcompare_Eq_eq i p);
-destruct (Pos.compare_cont i p Eq).
+ rsimpl. apply Ceqb_eq. trivial.
+ destruct (Pos.compare_spec i p).
assert (Hh := @Peq_ok P3 P0). case_eq (P3=? P0). intro. simpl.
rewrite Hh.
- rewrite Pphi0. rsimpl. rewrite Pplus_comm. rewrite pow_pos_Pplus;rsimpl.
-rewrite Hh1;trivial. reflexivity. trivial. intros. simpl. reflexivity. simpl. reflexivity.
+ rewrite Pphi0. rsimpl. rewrite Pos.add_comm. rewrite pow_pos_add;rsimpl.
+ subst;trivial. reflexivity. trivial. intros. simpl. reflexivity. simpl. reflexivity.
simpl. reflexivity.
Qed.
@@ -331,13 +295,13 @@ Lemma PaddXPX: forall P i n Q,
match Q with
| Pc c => mkPX P i n Q
| PX P' i' n' Q' =>
- match Pcompare i i' Eq with
+ match Pos.compare i i' with
| (* i > i' *)
Gt => mkPX P i n Q
| (* i < i' *)
Lt => mkPX P' i' n' (PaddX Padd P i n Q')
| (* i = i' *)
- Eq => match ZPminus n n' with
+ Eq => match Z.pos_sub n n' with
| (* n > n' *)
Zpos k => mkPX (PaddX Padd P i k P') i' n' Q'
| (* n = n' *)
@@ -359,17 +323,17 @@ Lemma PaddX_ok2 : forall P2,
induction P2;simpl;intros. split. intros. apply PaddCl_ok.
induction P. unfold PaddX. intros. rewrite mkPX_ok.
simpl. rsimpl.
-intros. simpl. assert (Hh := Pcompare_Eq_eq k p);
- destruct (Pos.compare_cont k p Eq).
- assert (H1h := ZPminus_spec n p0);destruct (ZPminus n p0). Esimpl2.
+intros. simpl.
+ destruct (Pos.compare_spec k p) as [Hh|Hh|Hh].
+ destr_pos_sub H1h. Esimpl2.
rewrite Hh; trivial. rewrite H1h. reflexivity.
simpl. rewrite mkPX_ok. rewrite IHP1. Esimpl2.
- rewrite Pplus_comm in H1h.
+ rewrite Pos.add_comm in H1h.
rewrite H1h.
-rewrite pow_pos_Pplus. Esimpl2.
+rewrite pow_pos_add. Esimpl2.
rewrite Hh; trivial. reflexivity.
-rewrite mkPX_ok. rewrite PaddCl_ok. Esimpl2. rewrite Pplus_comm in H1h.
-rewrite H1h. Esimpl2. rewrite pow_pos_Pplus. Esimpl2.
+rewrite mkPX_ok. rewrite PaddCl_ok. Esimpl2. rewrite Pos.add_comm in H1h.
+rewrite H1h. Esimpl2. rewrite pow_pos_add. Esimpl2.
rewrite Hh; trivial. reflexivity.
rewrite mkPX_ok. rewrite IHP2. Esimpl2.
rewrite (ring_add_comm (P2 @ l * pow_pos (nth 0 p l) p0)
@@ -382,19 +346,18 @@ split. intros. rewrite H0. rewrite H1.
Esimpl2.
induction P. unfold PaddX. intros. rewrite mkPX_ok. simpl. reflexivity.
intros. rewrite PaddXPX.
-assert (H3h := Pcompare_Eq_eq k p1);
- destruct (Pos.compare_cont k p1 Eq).
-assert (H4h := ZPminus_spec n p2);destruct (ZPminus n p2).
+destruct (Pos.compare_spec k p1) as [H3h|H3h|H3h].
+destr_pos_sub H4h.
rewrite mkPX_ok. simpl. rewrite H0. rewrite H1. Esimpl2.
rewrite H4h. rewrite H3h;trivial. reflexivity.
rewrite mkPX_ok. rewrite IHP1. Esimpl2. rewrite H3h;trivial.
-rewrite Pplus_comm in H4h.
-rewrite H4h. rewrite pow_pos_Pplus. Esimpl2.
+rewrite Pos.add_comm in H4h.
+rewrite H4h. rewrite pow_pos_add. Esimpl2.
rewrite mkPX_ok. simpl. rewrite H0. rewrite H1.
rewrite mkPX_ok.
Esimpl2. rewrite H3h;trivial.
- rewrite Pplus_comm in H4h.
-rewrite H4h. rewrite pow_pos_Pplus. Esimpl2.
+ rewrite Pos.add_comm in H4h.
+rewrite H4h. rewrite pow_pos_add. Esimpl2.
rewrite mkPX_ok. simpl. rewrite IHP2. Esimpl2.
gen_add_push (P2 @ l * pow_pos (nth 0 p1 l) p2). try reflexivity.
rewrite mkPX_ok. simpl. reflexivity.
diff --git a/plugins/setoid_ring/Ncring_tac.v b/plugins/setoid_ring/Ncring_tac.v
index d7cee83a5..6da54576d 100644
--- a/plugins/setoid_ring/Ncring_tac.v
+++ b/plugins/setoid_ring/Ncring_tac.v
@@ -104,7 +104,7 @@ Instance reify_pow (R:Type) `{Ring R}
Instance reify_var (R:Type) t lvar i
`{nth R t lvar i}
`{Rr: Ring (T:=R)}
- : reify (Rr:= Rr) (PEX Z (P_of_succ_nat i))lvar t
+ : reify (Rr:= Rr) (PEX Z (Pos.of_succ_nat i))lvar t
| 100.
Class reifylist (R:Type)`{Rr:Ring (T:=R)} (lexpr:list (PExpr Z)) (lvar:list R)
@@ -202,7 +202,7 @@ Ltac ring_simplify_aux lterm fv lexpr hyp :=
match lexpr with
| ?e::?le => (* e:PExpr Z est la réification de t0:R *)
let t := constr:(@Ncring_polynom.norm_subst
- Z 0%Z 1%Z Zplus Zmult Zminus Zopp (@eq Z) Zops Zeq_bool e) in
+ Z 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (@eq Z) Zops Zeq_bool e) in
(* t:Pol Z *)
let te :=
constr:(@Ncring_polynom.Pphi Z
@@ -212,13 +212,13 @@ Ltac ring_simplify_aux lterm fv lexpr hyp :=
let t':= fresh "t" in
pose (t' := nft);
assert (eq1 : t = t');
- [vm_cast_no_check (refl_equal t')|
+ [vm_cast_no_check (eq_refl t')|
let eq2 := fresh "ring" in
assert (eq2:(@Ncring_polynom.PEeval Z
_ 0 1 _+_ _*_ _-_ -_ _==_ _ Ncring_initial.gen_phiZ N (fun n:N => n)
(@Ring_theory.pow_N _ 1 multiplication) fv e) == te);
[apply (@Ncring_polynom.norm_subst_ok
- Z _ 0%Z 1%Z Zplus Zmult Zminus Zopp (@eq Z)
+ Z _ 0%Z 1%Z Z.add Z.mul Z.sub Z.opp (@eq Z)
_ _ 0 1 _+_ _*_ _-_ -_ _==_ _ _ Ncring_initial.gen_phiZ _
(@comm _ 0 1 _+_ _*_ _-_ -_ _==_ _ _) _ Zeqb_ok);
apply mkpow_th; reflexivity
diff --git a/plugins/setoid_ring/RealField.v b/plugins/setoid_ring/RealField.v
index 56473adb9..14c4270f9 100644
--- a/plugins/setoid_ring/RealField.v
+++ b/plugins/setoid_ring/RealField.v
@@ -111,19 +111,19 @@ Proof.
intros n0 H' m; rewrite H'; auto with real.
Qed.
-Lemma R_power_theory : power_theory 1%R Rmult (eq (A:=R)) nat_of_N pow.
+Lemma R_power_theory : power_theory 1%R Rmult (@eq R) N.to_nat pow.
Proof.
constructor. destruct n. reflexivity.
- simpl. induction p;simpl.
- rewrite ZL6. rewrite Rdef_pow_add;rewrite IHp. reflexivity.
- unfold nat_of_P;simpl;rewrite ZL6;rewrite Rdef_pow_add;rewrite IHp;trivial.
- rewrite Rmult_comm;apply Rmult_1_l.
+ simpl. induction p.
+ - rewrite Pos2Nat.inj_xI. simpl. now rewrite plus_0_r, Rdef_pow_add, IHp.
+ - rewrite Pos2Nat.inj_xO. simpl. now rewrite plus_0_r, Rdef_pow_add, IHp.
+ - simpl. rewrite Rmult_comm;apply Rmult_1_l.
Qed.
Ltac Rpow_tac t :=
match isnatcst t with
| false => constr:(InitialRing.NotConstant)
- | _ => constr:(N_of_nat t)
+ | _ => constr:(N.of_nat t)
end.
Add Field RField : Rfield
diff --git a/plugins/setoid_ring/Ring_polynom.v b/plugins/setoid_ring/Ring_polynom.v
index b722a31b6..4b4332e17 100644
--- a/plugins/setoid_ring/Ring_polynom.v
+++ b/plugins/setoid_ring/Ring_polynom.v
@@ -7,14 +7,10 @@
(************************************************************************)
Set Implicit Arguments.
-Require Import Setoid.
-Require Import BinList.
-Require Import BinPos.
-Require Import BinNat.
-Require Import BinInt.
+Require Import Setoid BinList BinPos BinNat BinInt.
Require Export Ring_theory.
-Open Local Scope positive_scope.
+Local Open Scope positive_scope.
Import RingSyntax.
Section MakeRingPol.
@@ -25,7 +21,7 @@ Section MakeRingPol.
Variable req : R -> R -> Prop.
(* Ring properties *)
- Variable Rsth : Setoid_Theory R req.
+ Variable Rsth : Equivalence req.
Variable Reqe : ring_eq_ext radd rmul ropp req.
Variable ARth : almost_ring_theory rO rI radd rmul rsub ropp req.
@@ -60,8 +56,6 @@ Section MakeRingPol.
Notation " x ?=! y" := (ceqb x y). Notation "[ x ]" := (phi x).
(* Useful tactics *)
- Add Setoid R req Rsth as R_set1.
- Ltac rrefl := gen_reflexivity Rsth.
Add Morphism radd : radd_ext. exact (Radd_ext Reqe). Qed.
Add Morphism rmul : rmul_ext. exact (Rmul_ext Reqe). Qed.
Add Morphism ropp : ropp_ext. exact (Ropp_ext Reqe). Qed.
@@ -128,7 +122,7 @@ Section MakeRingPol.
Definition mkPinj_pred j P:=
match j with
| xH => P
- | xO j => Pinj (Pdouble_minus_one j) P
+ | xO j => Pinj (Pos.pred_double j) P
| xI j => Pinj (xO j) P
end.
@@ -179,7 +173,7 @@ Section MakeRingPol.
match P with
| Pc c => mkPinj j (PaddC Q c)
| Pinj j' Q' =>
- match ZPminus j' j with
+ match Z.pos_sub j' j with
| Zpos k => mkPinj j (Pop (Pinj k Q') Q)
| Z0 => mkPinj j (Pop Q' Q)
| Zneg k => mkPinj j' (PaddI k Q')
@@ -187,7 +181,7 @@ Section MakeRingPol.
| PX P i Q' =>
match j with
| xH => PX P i (Pop Q' Q)
- | xO j => PX P i (PaddI (Pdouble_minus_one j) Q')
+ | xO j => PX P i (PaddI (Pos.pred_double j) Q')
| xI j => PX P i (PaddI (xO j) Q')
end
end.
@@ -196,7 +190,7 @@ Section MakeRingPol.
match P with
| Pc c => mkPinj j (PaddC (--Q) c)
| Pinj j' Q' =>
- match ZPminus j' j with
+ match Z.pos_sub j' j with
| Zpos k => mkPinj j (Pop (Pinj k Q') Q)
| Z0 => mkPinj j (Pop Q' Q)
| Zneg k => mkPinj j' (PsubI k Q')
@@ -204,7 +198,7 @@ Section MakeRingPol.
| PX P i Q' =>
match j with
| xH => PX P i (Pop Q' Q)
- | xO j => PX P i (PsubI (Pdouble_minus_one j) Q')
+ | xO j => PX P i (PsubI (Pos.pred_double j) Q')
| xI j => PX P i (PsubI (xO j) Q')
end
end.
@@ -217,11 +211,11 @@ Section MakeRingPol.
| Pinj j Q' =>
match j with
| xH => PX P' i' Q'
- | xO j => PX P' i' (Pinj (Pdouble_minus_one j) Q')
+ | xO j => PX P' i' (Pinj (Pos.pred_double j) Q')
| xI j => PX P' i' (Pinj (xO j) Q')
end
| PX P i Q' =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Pop (PX P k P0) P') i' Q'
| Z0 => mkPX (Pop P P') i Q'
| Zneg k => mkPX (PaddX k P) i Q'
@@ -234,11 +228,11 @@ Section MakeRingPol.
| Pinj j Q' =>
match j with
| xH => PX (--P') i' Q'
- | xO j => PX (--P') i' (Pinj (Pdouble_minus_one j) Q')
+ | xO j => PX (--P') i' (Pinj (Pos.pred_double j) Q')
| xI j => PX (--P') i' (Pinj (xO j) Q')
end
| PX P i Q' =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Pop (PX P k P0) P') i' Q'
| Z0 => mkPX (Pop P P') i Q'
| Zneg k => mkPX (PsubX k P) i Q'
@@ -258,11 +252,11 @@ Section MakeRingPol.
| Pinj j Q =>
match j with
| xH => PX P' i' (Padd Q Q')
- | xO j => PX P' i' (Padd (Pinj (Pdouble_minus_one j) Q) Q')
+ | xO j => PX P' i' (Padd (Pinj (Pos.pred_double j) Q) Q')
| xI j => PX P' i' (Padd (Pinj (xO j) Q) Q')
end
| PX P i Q =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Padd (PX P k P0) P') i' (Padd Q Q')
| Z0 => mkPX (Padd P P') i (Padd Q Q')
| Zneg k => mkPX (PaddX Padd P' k P) i (Padd Q Q')
@@ -281,11 +275,11 @@ Section MakeRingPol.
| Pinj j Q =>
match j with
| xH => PX (--P') i' (Psub Q Q')
- | xO j => PX (--P') i' (Psub (Pinj (Pdouble_minus_one j) Q) Q')
+ | xO j => PX (--P') i' (Psub (Pinj (Pos.pred_double j) Q) Q')
| xI j => PX (--P') i' (Psub (Pinj (xO j) Q) Q')
end
| PX P i Q =>
- match ZPminus i i' with
+ match Z.pos_sub i i' with
| Zpos k => mkPX (Psub (PX P k P0) P') i' (Psub Q Q')
| Z0 => mkPX (Psub P P') i (Psub Q Q')
| Zneg k => mkPX (PsubX Psub P' k P) i (Psub Q Q')
@@ -314,7 +308,7 @@ Section MakeRingPol.
match P with
| Pc c => mkPinj j (PmulC Q c)
| Pinj j' Q' =>
- match ZPminus j' j with
+ match Z.pos_sub j' j with
| Zpos k => mkPinj j (Pmul (Pinj k Q') Q)
| Z0 => mkPinj j (Pmul Q' Q)
| Zneg k => mkPinj j' (PmulI k Q')
@@ -322,13 +316,12 @@ Section MakeRingPol.
| PX P' i' Q' =>
match j with
| xH => mkPX (PmulI xH P') i' (Pmul Q' Q)
- | xO j' => mkPX (PmulI j P') i' (PmulI (Pdouble_minus_one j') Q')
+ | xO j' => mkPX (PmulI j P') i' (PmulI (Pos.pred_double j') Q')
| xI j' => mkPX (PmulI j P') i' (PmulI (xO j') Q')
end
end.
End PmulI.
-(* A symmetric version of the multiplication *)
Fixpoint Pmul (P P'' : Pol) {struct P''} : Pol :=
match P'' with
@@ -341,7 +334,7 @@ Section MakeRingPol.
let QQ' :=
match j with
| xH => Pmul Q Q'
- | xO j => Pmul (Pinj (Pdouble_minus_one j) Q) Q'
+ | xO j => Pmul (Pinj (Pos.pred_double j) Q) Q'
| xI j => Pmul (Pinj (xO j) Q) Q'
end in
mkPX (Pmul P P') i' QQ'
@@ -354,24 +347,6 @@ Section MakeRingPol.
end
end.
-(* Non symmetric *)
-(*
- Fixpoint Pmul_aux (P P' : Pol) {struct P'} : Pol :=
- match P' with
- | Pc c' => PmulC P c'
- | Pinj j' Q' => PmulI Pmul_aux Q' j' P
- | PX P' i' Q' =>
- (mkPX (Pmul_aux P P') i' P0) ++ (PmulI Pmul_aux Q' xH P)
- end.
-
- Definition Pmul P P' :=
- match P with
- | Pc c => PmulC P' c
- | Pinj j Q => PmulI Pmul_aux Q j P'
- | PX P i Q =>
- (mkPX (Pmul_aux P P') i P0) ++ (PmulI Pmul_aux Q xH P')
- end.
-*)
Notation "P ** P'" := (Pmul P P').
Fixpoint Psquare (P:Pol) : Pol :=
@@ -406,7 +381,7 @@ Section MakeRingPol.
match M with mon0 => mon0 | _ => zmon j M end.
Definition zmon_pred j M :=
- match j with xH => M | _ => mkZmon (Ppred j) M end.
+ match j with xH => M | _ => mkZmon (Pos.pred j) M end.
Definition mkVmon i M :=
match M with
@@ -517,51 +492,26 @@ Section MakeRingPol.
Reserved Notation "P @ l " (at level 10, no associativity).
Notation "P @ l " := (Pphi l P).
+
(** Proofs *)
- Lemma ZPminus_spec : forall x y,
- match ZPminus x y with
- | Z0 => x = y
- | Zpos k => x = (y + k)%positive
- | Zneg k => y = (x + k)%positive
+
+ Ltac destr_pos_sub H :=
+ match goal with |- context [Z.pos_sub ?x ?y] =>
+ assert (H := Z.pos_sub_discr x y); destruct (Z.pos_sub x y)
end.
- Proof.
- induction x;destruct y.
- replace (ZPminus (xI x) (xI y)) with (Zdouble (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial.
- replace (ZPminus (xI x) (xO y)) with (Zdouble_plus_one (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_plus_one;rewrite H;trivial.
- apply Pplus_xI_double_minus_one.
- simpl;trivial.
- replace (ZPminus (xO x) (xI y)) with (Zdouble_minus_one (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble_minus_one;rewrite H;trivial.
- apply Pplus_xI_double_minus_one.
- replace (ZPminus (xO x) (xO y)) with (Zdouble (ZPminus x y));trivial.
- assert (H := IHx y);destruct (ZPminus x y);unfold Zdouble;rewrite H;trivial.
- replace (ZPminus (xO x) xH) with (Zpos (Pdouble_minus_one x));trivial.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO;trivial.
- replace (ZPminus xH (xI y)) with (Zneg (xO y));trivial.
- replace (ZPminus xH (xO y)) with (Zneg (Pdouble_minus_one y));trivial.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO;trivial.
- simpl;trivial.
- Qed.
Lemma Peq_ok : forall P P',
(P ?== P') = true -> forall l, P@l == P'@ l.
Proof.
- induction P;destruct P';simpl;intros;try discriminate;trivial.
- apply (morph_eq CRmorph);trivial.
- assert (H1 := Pos.compare_eq p p0); destruct (p ?= p0);
- try discriminate H.
- rewrite (IHP P' H); rewrite H1;trivial;rrefl.
- assert (H1 := Pos.compare_eq p p0); destruct (p ?= p0);
- try discriminate H.
- rewrite H1;trivial. clear H1.
- assert (H1 := IHP1 P'1);assert (H2 := IHP2 P'2);
- destruct (P2 ?== P'1);[destruct (P3 ?== P'2); [idtac|discriminate H]
- |discriminate H].
- rewrite (H1 H);rewrite (H2 H);rrefl.
+ induction P;destruct P';simpl; intros H l; try easy.
+ - now apply (morph_eq CRmorph).
+ - destruct (Pos.compare_spec p p0); [ subst | easy | easy ].
+ now rewrite IHP.
+ - specialize (IHP1 P'1); specialize (IHP2 P'2).
+ destruct (Pos.compare_spec p p0); [ subst | easy | easy ].
+ destruct (P2 ?== P'1); [|easy].
+ rewrite H in *.
+ now rewrite IHP1, IHP2.
Qed.
Lemma Pphi0 : forall l, P0@l == 0.
@@ -577,23 +527,23 @@ Section MakeRingPol.
Lemma mkPinj_ok : forall j l P, (mkPinj j P)@l == P@(jump j l).
Proof.
intros j l p;destruct p;simpl;rsimpl.
- rewrite <-jump_Pplus;rewrite Pplus_comm;rrefl.
+ rewrite <-jump_add;rewrite Pos.add_comm;reflexivity.
Qed.
- Let pow_pos_Pplus :=
- pow_pos_Pplus rmul Rsth Reqe.(Rmul_ext) ARth.(ARmul_comm) ARth.(ARmul_assoc).
+ Let pow_pos_add :=
+ pow_pos_add Rsth Reqe.(Rmul_ext) ARth.(ARmul_assoc).
Lemma mkPX_ok : forall l P i Q,
(mkPX P i Q)@l == P@l*(pow_pos rmul (hd 0 l) i) + Q@(tail l).
Proof.
intros l P i Q;unfold mkPX.
- destruct P;try (simpl;rrefl).
- assert (H := morph_eq CRmorph c cO);destruct (c ?=! cO);simpl;try rrefl.
- rewrite (H (refl_equal true));rewrite (morph0 CRmorph).
- rewrite mkPinj_ok;rsimpl;simpl;rrefl.
- assert (H := @Peq_ok P3 P0);destruct (P3 ?== P0);simpl;try rrefl.
- rewrite (H (refl_equal true));trivial.
- rewrite Pphi0. rewrite pow_pos_Pplus;rsimpl.
+ destruct P;try (simpl;reflexivity).
+ assert (H := morph_eq CRmorph c cO);destruct (c ?=! cO);simpl;try reflexivity.
+ rewrite (H (eq_refl true));rewrite (morph0 CRmorph).
+ rewrite mkPinj_ok;rsimpl;simpl;reflexivity.
+ assert (H := @Peq_ok P3 P0);destruct (P3 ?== P0);simpl;try reflexivity.
+ rewrite (H (eq_refl true));trivial.
+ rewrite Pphi0. rewrite pow_pos_add;rsimpl.
Qed.
Ltac Esimpl :=
@@ -636,16 +586,16 @@ Section MakeRingPol.
Proof.
induction P;simpl;intros;Esimpl;trivial.
rewrite IHP1;rewrite IHP2;rsimpl.
- mul_push ([c]);rrefl.
+ mul_push ([c]);reflexivity.
Qed.
Lemma PmulC_ok : forall c P l, (PmulC P c)@l == P@l * [c].
Proof.
intros c P l; unfold PmulC.
assert (H:= morph_eq CRmorph c cO);destruct (c ?=! cO).
- rewrite (H (refl_equal true));Esimpl.
+ rewrite (H (eq_refl true));Esimpl.
assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI).
- rewrite (H1 (refl_equal true));Esimpl.
+ rewrite (H1 (eq_refl true));Esimpl.
apply PmulC_aux_ok.
Qed.
@@ -673,51 +623,51 @@ Section MakeRingPol.
generalize P p l;clear P p l.
induction P;simpl;intros.
Esimpl2;apply (ARadd_comm ARth).
- assert (H := ZPminus_spec p p0);destruct (ZPminus p p0).
- rewrite H;Esimpl. rewrite IHP';rrefl.
+ destr_pos_sub H.
+ rewrite H;Esimpl. rewrite IHP';reflexivity.
rewrite H;Esimpl. rewrite IHP';Esimpl.
- rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl.
+ rewrite <- jump_add;rewrite Pos.add_comm;reflexivity.
rewrite H;Esimpl. rewrite IHP.
- rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl.
+ rewrite <- jump_add;rewrite Pos.add_comm;reflexivity.
destruct p0;simpl.
rewrite IHP2;simpl;rsimpl.
rewrite IHP2;simpl.
- rewrite jump_Pdouble_minus_one;rsimpl.
+ rewrite jump_pred_double;rsimpl.
rewrite IHP';rsimpl.
destruct P;simpl.
- Esimpl2;add_push [c];rrefl.
+ Esimpl2;add_push [c];reflexivity.
destruct p0;simpl;Esimpl2.
rewrite IHP'2;simpl.
- rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl.
+ rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));reflexivity.
rewrite IHP'2;simpl.
- rewrite jump_Pdouble_minus_one;rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));rrefl.
- rewrite IHP'2;rsimpl. add_push (P @ (tail l));rrefl.
- assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2.
+ rewrite jump_pred_double;rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));reflexivity.
+ rewrite IHP'2;rsimpl. add_push (P @ (tail l));reflexivity.
+ destr_pos_sub H; Esimpl2.
rewrite IHP'1;rewrite IHP'2;rsimpl.
- add_push (P3 @ (tail l));rewrite H;rrefl.
+ add_push (P3 @ (tail l));rewrite H;reflexivity.
rewrite IHP'1;rewrite IHP'2;simpl;Esimpl.
- rewrite H;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;rsimpl.
- add_push (P3 @ (tail l));rrefl.
+ rewrite H;rewrite Pos.add_comm.
+ rewrite pow_pos_add;rsimpl.
+ add_push (P3 @ (tail l));reflexivity.
assert (forall P k l,
(PaddX Padd P'1 k P) @ l == P@l + P'1@l * pow_pos rmul (hd 0 l) k).
induction P;simpl;intros;try apply (ARadd_comm ARth).
destruct p2;simpl;try apply (ARadd_comm ARth).
- rewrite jump_Pdouble_minus_one;apply (ARadd_comm ARth).
- assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2.
- rewrite IHP'1;rsimpl; rewrite H1;add_push (P5 @ (tail l0));rrefl.
+ rewrite jump_pred_double;apply (ARadd_comm ARth).
+ destr_pos_sub H1; Esimpl2.
+ rewrite IHP'1;rsimpl; rewrite H1;add_push (P5 @ (tail l0));reflexivity.
rewrite IHP'1;simpl;Esimpl.
- rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;Esimpl.
- add_push (P5 @ (tail l0));rrefl.
- rewrite IHP1;rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;rsimpl.
- add_push (P5 @ (tail l0));rrefl.
+ rewrite H1;rewrite Pos.add_comm.
+ rewrite pow_pos_add;simpl;Esimpl.
+ add_push (P5 @ (tail l0));reflexivity.
+ rewrite IHP1;rewrite H1;rewrite Pos.add_comm.
+ rewrite pow_pos_add;simpl;rsimpl.
+ add_push (P5 @ (tail l0));reflexivity.
rewrite H0;rsimpl.
add_push (P3 @ (tail l)).
- rewrite H;rewrite Pplus_comm.
- rewrite IHP'2;rewrite pow_pos_Pplus;rsimpl.
- add_push (P3 @ (tail l));rrefl.
+ rewrite H;rewrite Pos.add_comm.
+ rewrite IHP'2;rewrite pow_pos_add;rsimpl.
+ add_push (P3 @ (tail l));reflexivity.
Qed.
Lemma Psub_ok : forall P' P l, (P -- P')@l == P@l - P'@l.
@@ -726,54 +676,53 @@ Section MakeRingPol.
generalize P p l;clear P p l.
induction P;simpl;intros.
Esimpl2;apply (ARadd_comm ARth).
- assert (H := ZPminus_spec p p0);destruct (ZPminus p p0).
+ destr_pos_sub H.
rewrite H;Esimpl. rewrite IHP';rsimpl.
rewrite H;Esimpl. rewrite IHP';Esimpl.
- rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl.
+ rewrite <- jump_add;rewrite Pos.add_comm;reflexivity.
rewrite H;Esimpl. rewrite IHP.
- rewrite <- jump_Pplus;rewrite Pplus_comm;rrefl.
+ rewrite <- jump_add;rewrite Pos.add_comm;reflexivity.
destruct p0;simpl.
rewrite IHP2;simpl;rsimpl.
rewrite IHP2;simpl.
- rewrite jump_Pdouble_minus_one;rsimpl.
+ rewrite jump_pred_double;rsimpl.
rewrite IHP';rsimpl.
destruct P;simpl.
- repeat rewrite Popp_ok;Esimpl2;rsimpl;add_push [c];try rrefl.
+ repeat rewrite Popp_ok;Esimpl2;rsimpl;add_push [c];try reflexivity.
destruct p0;simpl;Esimpl2.
rewrite IHP'2;simpl;rsimpl;add_push (P'1@l * (pow_pos rmul (hd 0 l) p));trivial.
- add_push (P @ (jump p0 (jump p0 (tail l))));rrefl.
- rewrite IHP'2;simpl;rewrite jump_Pdouble_minus_one;rsimpl.
- add_push (- (P'1 @ l * pow_pos rmul (hd 0 l) p));rrefl.
- rewrite IHP'2;rsimpl;add_push (P @ (tail l));rrefl.
- assert (H := ZPminus_spec p0 p);destruct (ZPminus p0 p);Esimpl2.
+ add_push (P @ (jump p0 (jump p0 (tail l))));reflexivity.
+ rewrite IHP'2;simpl;rewrite jump_pred_double;rsimpl.
+ add_push (- (P'1 @ l * pow_pos rmul (hd 0 l) p));reflexivity.
+ rewrite IHP'2;rsimpl;add_push (P @ (tail l));reflexivity.
+ destr_pos_sub H; Esimpl2.
rewrite IHP'1; rewrite IHP'2;rsimpl.
- add_push (P3 @ (tail l));rewrite H;rrefl.
+ add_push (P3 @ (tail l));rewrite H;reflexivity.
rewrite IHP'1; rewrite IHP'2;rsimpl;simpl;Esimpl.
- rewrite H;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;rsimpl.
- add_push (P3 @ (tail l));rrefl.
+ rewrite H;rewrite Pos.add_comm.
+ rewrite pow_pos_add;rsimpl.
+ add_push (P3 @ (tail l));reflexivity.
assert (forall P k l,
(PsubX Psub P'1 k P) @ l == P@l + - P'1@l * pow_pos rmul (hd 0 l) k).
induction P;simpl;intros.
rewrite Popp_ok;rsimpl;apply (ARadd_comm ARth);trivial.
destruct p2;simpl;rewrite Popp_ok;rsimpl.
apply (ARadd_comm ARth);trivial.
- rewrite jump_Pdouble_minus_one;apply (ARadd_comm ARth);trivial.
+ rewrite jump_pred_double;apply (ARadd_comm ARth);trivial.
apply (ARadd_comm ARth);trivial.
- assert (H1 := ZPminus_spec p2 k);destruct (ZPminus p2 k);Esimpl2;rsimpl.
- rewrite IHP'1;rsimpl;add_push (P5 @ (tail l0));rewrite H1;rrefl.
- rewrite IHP'1;rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;Esimpl.
- add_push (P5 @ (tail l0));rrefl.
- rewrite IHP1;rewrite H1;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;simpl;rsimpl.
- add_push (P5 @ (tail l0));rrefl.
+ destr_pos_sub H1; Esimpl2; rsimpl.
+ rewrite IHP'1;rsimpl;add_push (P5 @ (tail l0));rewrite H1;reflexivity.
+ rewrite IHP'1;rewrite H1;rewrite Pos.add_comm.
+ rewrite pow_pos_add;simpl;Esimpl.
+ add_push (P5 @ (tail l0));reflexivity.
+ rewrite IHP1;rewrite H1;rewrite Pos.add_comm.
+ rewrite pow_pos_add;simpl;rsimpl.
+ add_push (P5 @ (tail l0));reflexivity.
rewrite H0;rsimpl.
rewrite IHP'2;rsimpl;add_push (P3 @ (tail l)).
- rewrite H;rewrite Pplus_comm.
- rewrite pow_pos_Pplus;rsimpl.
+ rewrite H;rewrite Pos.add_comm.
+ rewrite pow_pos_add;rsimpl.
Qed.
-(* Proof for the symmetriv version *)
Lemma PmulI_ok :
forall P',
@@ -783,60 +732,23 @@ Section MakeRingPol.
Proof.
induction P;simpl;intros.
Esimpl2;apply (ARmul_comm ARth).
- assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2.
- rewrite H1; rewrite H;rrefl.
+ destr_pos_sub H1; Esimpl2.
+ rewrite H1; rewrite H;reflexivity.
rewrite H1; rewrite H.
- rewrite Pplus_comm.
- rewrite jump_Pplus;simpl;rrefl.
- rewrite H1;rewrite Pplus_comm.
- rewrite jump_Pplus;rewrite IHP;rrefl.
+ rewrite Pos.add_comm.
+ rewrite jump_add;simpl;reflexivity.
+ rewrite H1;rewrite Pos.add_comm.
+ rewrite jump_add;rewrite IHP;reflexivity.
destruct p0;Esimpl2.
rewrite IHP1;rewrite IHP2;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p);rrefl.
+ mul_push (pow_pos rmul (hd 0 l) p);reflexivity.
rewrite IHP1;rewrite IHP2;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p); rewrite jump_Pdouble_minus_one;rrefl.
+ mul_push (pow_pos rmul (hd 0 l) p); rewrite jump_pred_double;reflexivity.
rewrite IHP1;simpl;rsimpl.
mul_push (pow_pos rmul (hd 0 l) p).
- rewrite H;rrefl.
+ rewrite H;reflexivity.
Qed.
-(*
- Lemma PmulI_ok :
- forall P',
- (forall (P : Pol) (l : list R), (Pmul_aux P P') @ l == P @ l * P' @ l) ->
- forall (P : Pol) (p : positive) (l : list R),
- (PmulI Pmul_aux P' p P) @ l == P @ l * P' @ (jump p l).
- Proof.
- induction P;simpl;intros.
- Esimpl2;apply (ARmul_comm ARth).
- assert (H1 := ZPminus_spec p p0);destruct (ZPminus p p0);Esimpl2.
- rewrite H1; rewrite H;rrefl.
- rewrite H1; rewrite H.
- rewrite Pplus_comm.
- rewrite jump_Pplus;simpl;rrefl.
- rewrite H1;rewrite Pplus_comm.
- rewrite jump_Pplus;rewrite IHP;rrefl.
- destruct p0;Esimpl2.
- rewrite IHP1;rewrite IHP2;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p);rrefl.
- rewrite IHP1;rewrite IHP2;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p); rewrite jump_Pdouble_minus_one;rrefl.
- rewrite IHP1;simpl;rsimpl.
- mul_push (pow_pos rmul (hd 0 l) p).
- rewrite H;rrefl.
- Qed.
-
- Lemma Pmul_aux_ok : forall P' P l,(Pmul_aux P P')@l == P@l * P'@l.
- Proof.
- induction P';simpl;intros.
- Esimpl2;trivial.
- apply PmulI_ok;trivial.
- rewrite Padd_ok;Esimpl2.
- rewrite (PmulI_ok P'2 IHP'2). rewrite IHP'1. rrefl.
- Qed.
-*)
-
-(* Proof for the symmetric version *)
Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
Proof.
intros P P';generalize P;clear P;induction P';simpl;intros.
@@ -846,11 +758,11 @@ Section MakeRingPol.
Esimpl2. rewrite IHP'1;Esimpl2.
assert (match p0 with
| xI j => Pinj (xO j) P ** P'2
- | xO j => Pinj (Pdouble_minus_one j) P ** P'2
+ | xO j => Pinj (Pos.pred_double j) P ** P'2
| 1 => P ** P'2
end @ (tail l) == P @ (jump p0 l) * P'2 @ (tail l)).
destruct p0;simpl;rewrite IHP'2;Esimpl.
- rewrite jump_Pdouble_minus_one;Esimpl.
+ rewrite jump_pred_double;Esimpl.
rewrite H;Esimpl.
rewrite Padd_ok; Esimpl2. rewrite Padd_ok; Esimpl2.
repeat (rewrite IHP'1 || rewrite IHP'2);simpl.
@@ -858,27 +770,13 @@ Section MakeRingPol.
mul_push (P'1@l). simpl. mul_push (P'2 @ (tail l)). Esimpl.
Qed.
-(*
-Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
- Proof.
- destruct P;simpl;intros.
- Esimpl2;apply (ARmul_comm ARth).
- rewrite (PmulI_ok P (Pmul_aux_ok P)).
- apply (ARmul_comm ARth).
- rewrite Padd_ok; Esimpl2.
- rewrite (PmulI_ok P3 (Pmul_aux_ok P3));trivial.
- rewrite Pmul_aux_ok;mul_push (P' @ l).
- rewrite (ARmul_comm ARth (P' @ l));rrefl.
- Qed.
-*)
-
Lemma Psquare_ok : forall P l, (Psquare P)@l == P@l * P@l.
Proof.
induction P;simpl;intros;Esimpl2.
apply IHP. rewrite Padd_ok. rewrite Pmul_ok;Esimpl2.
rewrite IHP1;rewrite IHP2.
mul_push (pow_pos rmul (hd 0 l) p). mul_push (P2@l).
- rrefl.
+ reflexivity.
Qed.
@@ -892,14 +790,14 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
Proof.
destruct j; simpl;intros auto; rsimpl.
rewrite mkZmon_ok;rsimpl.
- rewrite mkZmon_ok;simpl. rewrite jump_Pdouble_minus_one; rsimpl.
+ rewrite mkZmon_ok;simpl. rewrite jump_pred_double; rsimpl.
Qed.
Lemma mkVmon_ok : forall M i l, Mphi l (mkVmon i M) == Mphi l M*pow_pos rmul (hd 0 l) i.
Proof.
destruct M;simpl;intros;rsimpl.
rewrite zmon_pred_ok;simpl;rsimpl.
- rewrite Pplus_comm;rewrite pow_pos_Pplus;rsimpl.
+ rewrite Pos.add_comm;rewrite pow_pos_add;rsimpl.
Qed.
Lemma Mcphi_ok: forall P c l,
@@ -930,9 +828,9 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
P@l == Q@l + (phi c) * (Mphi l M) * (R@l).
Proof.
intros P; elim P; simpl; auto; clear P.
- intros c (c1, M) l; case M; simpl; auto.
+ - intros c (c1, M) l; case M; simpl; auto.
assert (H1:= morph_eq CRmorph c1 cI);destruct (c1 ?=! cI).
- rewrite (H1 (refl_equal true));Esimpl.
+ rewrite (H1 (eq_refl true));Esimpl.
try rewrite (morph0 CRmorph); rsimpl.
generalize (div_th.(div_eucl_th) c c1); case (cdiv c c1).
intros q r H; rewrite H; clear H H1.
@@ -940,39 +838,39 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
rewrite (ARadd_comm ARth); rsimpl.
intros p m; Esimpl.
intros p m; Esimpl.
- intros i P Hrec (c,M) l; case M; simpl; clear M.
- assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI).
- rewrite (H1 (refl_equal true));Esimpl.
+ - intros i P Hrec (c,M) l; case M; simpl; clear M.
+ + assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI).
+ rewrite (H1 (eq_refl true));Esimpl.
Esimpl.
- generalize (Mcphi_ok P c (jump i l)); case CFactor.
- intros R1 Q1 HH; rewrite HH; Esimpl.
- intros j M.
- case_eq (i ?= j); intros He; simpl.
- rewrite (Pos.compare_eq _ _ He).
+ generalize (Mcphi_ok P c (jump i l)); case CFactor.
+ intros R1 Q1 HH; rewrite HH; Esimpl.
+ + intros j M.
+ case Pos.compare_spec; intros He; simpl.
+ * rewrite He.
generalize (Hrec (c, M) (jump j l)); case (MFactor P c M);
simpl; intros P2 Q2 H; repeat rewrite mkPinj_ok; auto.
- generalize (Hrec (c, (zmon (j -i) M)) (jump i l));
+ * generalize (Hrec (c, (zmon (j -i) M)) (jump i l));
case (MFactor P c (zmon (j -i) M)); simpl.
intros P2 Q2 H; repeat rewrite mkPinj_ok; auto.
- rewrite <- (Pplus_minus _ _ (ZC2 _ _ He)).
- rewrite Pplus_comm; rewrite jump_Pplus; auto.
- rewrite (morph0 CRmorph); rsimpl.
- intros P2 m; rewrite (morph0 CRmorph); rsimpl.
-
- intros P2 Hrec1 i Q2 Hrec2 (c, M) l; case M; simpl; auto.
- assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI).
- rewrite (H1 (refl_equal true));Esimpl.
+ rewrite <- (Pos.sub_add _ _ He).
+ rewrite jump_add; auto.
+ * rewrite (morph0 CRmorph); rsimpl.
+ + intros P2 m; rewrite (morph0 CRmorph); rsimpl.
+
+ - intros P2 Hrec1 i Q2 Hrec2 (c, M) l; case M; simpl; auto.
+ + assert (H1:= morph_eq CRmorph c cI);destruct (c ?=! cI).
+ rewrite (H1 (eq_refl true));Esimpl.
Esimpl.
- generalize (Mcphi_ok P2 c l); case CFactor.
- intros S1 S2 HS.
- generalize (Mcphi_ok Q2 c (tail l)); case CFactor.
- intros S3 S4 HS1; Esimpl; rewrite HS; rewrite HS1.
- rsimpl.
- apply (Radd_ext Reqe); rsimpl.
- repeat rewrite <- (ARadd_assoc ARth).
- apply (Radd_ext Reqe); rsimpl.
- rewrite (ARadd_comm ARth); rsimpl.
- intros j M1.
+ generalize (Mcphi_ok P2 c l); case CFactor.
+ intros S1 S2 HS.
+ generalize (Mcphi_ok Q2 c (tail l)); case CFactor.
+ intros S3 S4 HS1; Esimpl; rewrite HS; rewrite HS1.
+ rsimpl.
+ apply (Radd_ext Reqe); rsimpl.
+ repeat rewrite <- (ARadd_assoc ARth).
+ apply (Radd_ext Reqe); rsimpl.
+ rewrite (ARadd_comm ARth); rsimpl.
+ + intros j M1.
generalize (Hrec1 (c,zmon j M1) l);
case (MFactor P2 c (zmon j M1)).
intros R1 S1 H1.
@@ -986,9 +884,9 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
apply radd_ext; rsimpl.
rewrite (ARadd_comm ARth); rsimpl.
rewrite zmon_pred_ok;rsimpl.
- intros j M1.
- case_eq (i ?= j); intros He; simpl.
- rewrite (Pos.compare_eq _ _ He).
+ + intros j M1.
+ case Pos.compare_spec; intros He; simpl.
+ * rewrite He.
generalize (Hrec1 (c, mkZmon xH M1) l); case (MFactor P2 c (mkZmon xH M1));
simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto.
rewrite H; rewrite mkPX_ok; rsimpl.
@@ -1002,7 +900,7 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
repeat (rewrite <-(ARmul_assoc ARth)).
apply rmul_ext; rsimpl.
rewrite (ARmul_comm ARth); rsimpl.
- generalize (Hrec1 (c, vmon (j - i) M1) l);
+ * generalize (Hrec1 (c, vmon (j - i) M1) l);
case (MFactor P2 c (vmon (j - i) M1));
simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto.
rewrite H; rsimpl; repeat rewrite mkPinj_ok; auto.
@@ -1018,9 +916,9 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
rewrite <- (ARmul_comm ARth (Mphi (tail l) M1)); rsimpl.
repeat (rewrite <-(ARmul_assoc ARth)).
apply rmul_ext; rsimpl.
- rewrite <- pow_pos_Pplus.
- rewrite (Pplus_minus _ _ (ZC2 _ _ He)); rsimpl.
- generalize (Hrec1 (c, mkZmon 1 M1) l);
+ rewrite <- pow_pos_add.
+ rewrite Pos.add_comm, Pos.sub_add by trivial; rsimpl.
+ * generalize (Hrec1 (c, mkZmon 1 M1) l);
case (MFactor P2 c (mkZmon 1 M1));
simpl; intros P3 Q3 H; repeat rewrite mkPinj_ok; auto.
rewrite H; rsimpl.
@@ -1041,12 +939,10 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
rewrite (ARmul_comm ARth); rsimpl.
repeat (rewrite <- (ARmul_assoc ARth)).
apply rmul_ext; rsimpl.
- rewrite <- pow_pos_Pplus.
- rewrite (Pplus_minus _ _ He); rsimpl.
+ rewrite <- pow_pos_add.
+ rewrite Pos.add_comm, Pos.sub_add by trivial; rsimpl.
Qed.
-(* Proof for the symmetric version *)
-
Lemma POneSubst_ok: forall P1 M1 P2 P3 l,
POneSubst P1 M1 P2 = Some P3 -> phi (fst M1) * Mphi l (snd M1) == P2@l -> P1@l == P3@l.
Proof.
@@ -1070,30 +966,7 @@ Lemma Pmul_ok : forall P P' l, (P**P')@l == P@l * P'@l.
injection H2; intros; subst;trivial.
rewrite H;rewrite Padd_ok;rewrite Pmul_ok;rsimpl.
Qed.
-(*
- Lemma POneSubst_ok: forall P1 M1 P2 P3 l,
- POneSubst P1 M1 P2 = Some P3 -> Mphi l M1 == P2@l -> P1@l == P3@l.
-Proof.
- intros P2 M1 P3 P4 l; unfold POneSubst.
- generalize (Mphi_ok P2 M1 l); case (MFactor P2 M1); simpl; auto.
- intros Q1 R1; case R1.
- intros c H; rewrite H.
- generalize (morph_eq CRmorph c cO);
- case (c ?=! cO); simpl; auto.
- intros H1 H2; rewrite H1; auto; rsimpl.
- discriminate.
- intros _ H1 H2; injection H1; intros; subst.
- rewrite H2; rsimpl.
- rewrite Padd_ok; rewrite Pmul_ok; rsimpl.
- intros i P5 H; rewrite H.
- intros HH H1; injection HH; intros; subst; rsimpl.
- rewrite Padd_ok; rewrite Pmul_ok. rewrite H1; rsimpl.
- intros i P5 P6 H1 H2 H3; rewrite H1; rewrite H3.
- injection H2; intros; subst; rsimpl.
- rewrite Padd_ok.
- rewrite Pmul_ok; rsimpl.
- Qed.
-*)
+
Lemma PNSubst1_ok: forall n P1 M1 P2 l,
[fst M1] * Mphi l (snd M1) == P2@l -> P1@l == (PNSubst1 P1 M1 P2 n)@l.
Proof.
@@ -1193,47 +1066,16 @@ Strategy expand [PEeval].
Lemma mkX_ok : forall p l, nth 0 p l == (mk_X p) @ l.
Proof.
destruct p;simpl;intros;Esimpl;trivial.
- rewrite <-jump_tl;rewrite nth_jump;rrefl.
+ rewrite <-jump_tl;rewrite nth_jump;reflexivity.
rewrite <- nth_jump.
- rewrite nth_Pdouble_minus_one;rrefl.
+ rewrite nth_pred_double;reflexivity.
Qed.
Ltac Esimpl3 :=
repeat match goal with
| |- context [(?P1 ++ ?P2)@?l] => rewrite (Padd_ok P2 P1 l)
| |- context [(?P1 -- ?P2)@?l] => rewrite (Psub_ok P2 P1 l)
- end;Esimpl2;try rrefl;try apply (ARadd_comm ARth).
-
-(* Power using the chinise algorithm *)
-(*Section POWER.
- Variable subst_l : Pol -> Pol.
- Fixpoint Ppow_pos (P:Pol) (p:positive){struct p} : Pol :=
- match p with
- | xH => P
- | xO p => subst_l (Psquare (Ppow_pos P p))
- | xI p => subst_l (Pmul P (Psquare (Ppow_pos P p)))
- end.
-
- Definition Ppow_N P n :=
- match n with
- | N0 => P1
- | Npos p => Ppow_pos P p
- end.
-
- Lemma Ppow_pos_ok : forall l, (forall P, subst_l P@l == P@l) ->
- forall P p, (Ppow_pos P p)@l == (pow_pos Pmul P p)@l.
- Proof.
- intros l subst_l_ok P.
- induction p;simpl;intros;try rrefl;try rewrite subst_l_ok.
- repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl.
- repeat rewrite Pmul_ok;rewrite Psquare_ok;rewrite IHp;rrefl.
- Qed.
-
- Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) ->
- forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l.
- Proof. destruct n;simpl. rrefl. apply Ppow_pos_ok. trivial. Qed.
-
- End POWER. *)
+ end;Esimpl2;try reflexivity;try apply (ARadd_comm ARth).
Section POWER.
Variable subst_l : Pol -> Pol.
@@ -1255,12 +1097,12 @@ Section POWER.
Proof.
intros l subst_l_ok res P p. generalize res;clear res.
induction p;simpl;intros;try rewrite subst_l_ok; repeat rewrite Pmul_ok;repeat rewrite IHp.
- rsimpl. mul_push (P@l);rsimpl. rsimpl. rrefl.
+ rsimpl. mul_push (P@l);rsimpl. rsimpl. reflexivity.
Qed.
Lemma Ppow_N_ok : forall l, (forall P, subst_l P@l == P@l) ->
forall P n, (Ppow_N P n)@l == (pow_N P1 Pmul P n)@l.
- Proof. destruct n;simpl. rrefl. rewrite Ppow_pos_ok by trivial. Esimpl. Qed.
+ Proof. destruct n;simpl. reflexivity. rewrite Ppow_pos_ok by trivial. Esimpl. Qed.
End POWER.
@@ -1289,42 +1131,6 @@ Section POWER.
Definition norm_subst pe := subst_l (norm_aux pe).
- (*
- Fixpoint norm_subst (pe:PExpr) : Pol :=
- match pe with
- | PEc c => Pc c
- | PEX j => subst_l (mk_X j)
- | PEadd (PEopp pe1) pe2 => Psub (norm_subst pe2) (norm_subst pe1)
- | PEadd pe1 (PEopp pe2) =>
- Psub (norm_subst pe1) (norm_subst pe2)
- | PEadd pe1 pe2 => Padd (norm_subst pe1) (norm_subst pe2)
- | PEsub pe1 pe2 => Psub (norm_subst pe1) (norm_subst pe2)
- | PEmul pe1 pe2 => Pmul_subst (norm_subst pe1) (norm_subst pe2)
- | PEopp pe1 => Popp (norm_subst pe1)
- | PEpow pe1 n => Ppow_subst (norm_subst pe1) n
- end.
-
- Lemma norm_subst_spec :
- forall l pe, MPcond lmp l ->
- PEeval l pe == (norm_subst pe)@l.
- Proof.
- intros;assert (subst_l_ok:forall P, (subst_l P)@l == P@l).
- unfold subst_l;intros.
- rewrite <- PNSubstL_ok;trivial. rrefl.
- assert (Pms_ok:forall P1 P2, (Pmul_subst P1 P2)@l == P1@l*P2@l).
- intros;unfold Pmul_subst;rewrite subst_l_ok;rewrite Pmul_ok;rrefl.
- induction pe;simpl;Esimpl3.
- rewrite subst_l_ok;apply mkX_ok.
- rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3.
- rewrite IHpe1;rewrite IHpe2;rrefl.
- rewrite Pms_ok;rewrite IHpe1;rewrite IHpe2;rrefl.
- rewrite IHpe;rrefl.
- unfold Ppow_subst. rewrite Ppow_N_ok. trivial.
- rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3.
- induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok;
- repeat rewrite Pmul_ok;rrefl.
- Qed.
-*)
Lemma norm_aux_spec :
forall l pe, MPcond lmp l ->
PEeval l pe == (norm_aux pe)@l.
@@ -1333,13 +1139,13 @@ Section POWER.
induction pe;simpl;Esimpl3.
apply mkX_ok.
rewrite IHpe1;rewrite IHpe2;destruct pe1;destruct pe2;Esimpl3.
- rewrite IHpe1;rewrite IHpe2;rrefl.
- rewrite IHpe1;rewrite IHpe2. rewrite Pmul_ok. rrefl.
- rewrite IHpe;rrefl.
- rewrite Ppow_N_ok by (intros;rrefl).
+ rewrite IHpe1;rewrite IHpe2;reflexivity.
+ rewrite IHpe1;rewrite IHpe2. rewrite Pmul_ok. reflexivity.
+ rewrite IHpe;reflexivity.
+ rewrite Ppow_N_ok by (intros;reflexivity).
rewrite pow_th.(rpow_pow_N). destruct n0;Esimpl3.
induction p;simpl;try rewrite IHp;try rewrite IHpe;repeat rewrite Pms_ok;
- repeat rewrite Pmul_ok;rrefl.
+ repeat rewrite Pmul_ok;reflexivity.
Qed.
Lemma norm_subst_spec :
@@ -1519,7 +1325,7 @@ Section POWER.
| Pinj j Q =>
add_mult_dev rP Q (jump j fv) N0 (add_pow_list (hd 0 fv) n lm)
| PX P i Q =>
- let rP := add_mult_dev rP P fv (Nplus (Npos i) n) lm in
+ let rP := add_mult_dev rP P fv (N.add (Npos i) n) lm in
if Q ?== P0 then rP
else add_mult_dev rP Q (tail fv) N0 (add_pow_list (hd 0 fv) n lm)
end.
@@ -1531,7 +1337,7 @@ Section POWER.
| Pc c => mkmult_c c (add_pow_list (hd 0 fv) n lm)
| Pinj j Q => mult_dev Q (jump j fv) N0 (add_pow_list (hd 0 fv) n lm)
| PX P i Q =>
- let rP := mult_dev P fv (Nplus (Npos i) n) lm in
+ let rP := mult_dev P fv (N.add (Npos i) n) lm in
if Q ?== P0 then rP
else
let lmq := add_pow_list (hd 0 fv) n lm in
@@ -1575,7 +1381,7 @@ Section POWER.
(forall l lr : list (R * positive), r_list_pow (rev_append l lr) == r_list_pow lr * r_list_pow l).
induction l;intros;simpl;Esimpl.
destruct a;rewrite IHl;Esimpl.
- rewrite (ARmul_comm ARth (pow_pos rmul r p)). rrefl.
+ rewrite (ARmul_comm ARth (pow_pos rmul r p)). reflexivity.
intros;unfold rev'. rewrite H;simpl;Esimpl.
Qed.
@@ -1630,13 +1436,13 @@ Qed.
change match n with
| N0 => Npos p
| Npos q => Npos (p + q)
- end with (Nplus (Npos p) n);trivial.
+ end with (N.add (Npos p) n);trivial.
assert (H := Peq_ok P3 P0).
destruct (P3 ?== P0).
- rewrite (H (refl_equal true)).
- rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl.
+ rewrite (H (eq_refl true)).
+ rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl.
rewrite IHP2.
- rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl.
+ rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl.
Qed.
Lemma mult_dev_ok : forall P fv n lm,
@@ -1653,13 +1459,13 @@ Qed.
change match n with
| N0 => Npos p
| Npos q => Npos (p + q)
- end with (Nplus (Npos p) n);trivial.
+ end with (N.add (Npos p) n);trivial.
assert (H := Peq_ok P3 P0).
destruct (P3 ?== P0).
- rewrite (H (refl_equal true)).
- rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl.
+ rewrite (H (eq_refl true)).
+ rewrite IHP1. destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl.
rewrite add_mult_dev_ok. rewrite IHP1; rewrite add_pow_list_ok.
- destruct n;simpl;Esimpl;rewrite pow_pos_Pplus;Esimpl.
+ destruct n;simpl;Esimpl;rewrite pow_pos_add;Esimpl.
Qed.
Lemma Pphi_avoid_ok : forall P fv, Pphi_avoid fv P == P@fv.
@@ -1687,7 +1493,7 @@ Qed.
Lemma Pphi_pow_ok : forall P fv, Pphi_pow fv P == P@fv.
Proof.
- unfold Pphi_pow;intros;apply Pphi_avoid_ok;intros;try rewrite local_mkpow_ok;rrefl.
+ unfold Pphi_pow;intros;apply Pphi_avoid_ok;intros;try rewrite local_mkpow_ok;reflexivity.
Qed.
Lemma ring_rw_pow_correct : forall n lH l,
@@ -1711,14 +1517,14 @@ Qed.
Definition mkpow x p :=
match p with
| xH => x
- | xO p => mkmult_pow x x (Pdouble_minus_one p)
+ | xO p => mkmult_pow x x (Pos.pred_double p)
| xI p => mkmult_pow x x (xO p)
end.
Definition mkopp_pow x p :=
match p with
| xH => -x
- | xO p => mkmult_pow (-x) x (Pdouble_minus_one p)
+ | xO p => mkmult_pow (-x) x (Pos.pred_double p)
| xI p => mkmult_pow (-x) x (xO p)
end.
@@ -1737,9 +1543,9 @@ Qed.
repeat rewrite mkmult_pow_ok;Esimpl.
rewrite mkmult_pow_ok;Esimpl.
pattern x at 1;replace x with (pow_pos rmul x 1).
- rewrite <- pow_pos_Pplus.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO.
+ rewrite <- pow_pos_add.
+ rewrite Pos.add_1_l.
+ rewrite Pos.succ_pred_double.
simpl;Esimpl.
trivial.
Qed.
@@ -1750,9 +1556,9 @@ Qed.
repeat rewrite mkmult_pow_ok;Esimpl.
rewrite mkmult_pow_ok;Esimpl.
pattern x at 1;replace x with (pow_pos rmul x 1).
- rewrite <- pow_pos_Pplus.
- rewrite <- Pplus_one_succ_l.
- rewrite Psucc_o_double_minus_one_eq_xO.
+ rewrite <- pow_pos_add.
+ rewrite Pos.add_1_l.
+ rewrite Pos.succ_pred_double.
simpl;Esimpl.
trivial.
Qed.
diff --git a/plugins/setoid_ring/Ring_tac.v b/plugins/setoid_ring/Ring_tac.v
index d33e9a82a..7a7ffcfdc 100644
--- a/plugins/setoid_ring/Ring_tac.v
+++ b/plugins/setoid_ring/Ring_tac.v
@@ -3,6 +3,7 @@ Require Import Setoid.
Require Import BinPos.
Require Import Ring_polynom.
Require Import BinList.
+Require Export ListTactics.
Require Import InitialRing.
Require Import Quote.
Declare ML Module "newring_plugin".
@@ -14,7 +15,7 @@ Ltac compute_assertion eqn t' t :=
let nft := eval vm_compute in t in
pose (t' := nft);
assert (eqn : t = t');
- [vm_cast_no_check (refl_equal t')|idtac].
+ [vm_cast_no_check (eq_refl t')|idtac].
Ltac relation_carrier req :=
let ty := type of req in
@@ -340,7 +341,7 @@ Ltac Ring RNG lemma lH :=
|| idtac "can not automatically proof hypothesis :";
idtac " maybe a left member of a hypothesis is not a monomial")
| vm_compute;
- (exact (refl_equal true) || fail "not a valid ring equation")]).
+ (exact (eq_refl true) || fail "not a valid ring equation")]).
Ltac Ring_norm_gen f RNG lemma lH rl :=
let mkFV := get_RingFV RNG in
@@ -385,7 +386,7 @@ Ltac Ring_simplify_gen f RNG lH rl :=
let lemma := get_SimplifyLemma RNG in
let l := fresh "to_rewrite" in
pose (l:= rl);
- generalize (refl_equal l);
+ generalize (eq_refl l);
unfold l at 2;
get_Pre RNG ();
let rl :=
diff --git a/plugins/setoid_ring/Ring_theory.v b/plugins/setoid_ring/Ring_theory.v
index ab9925528..f1f419662 100644
--- a/plugins/setoid_ring/Ring_theory.v
+++ b/plugins/setoid_ring/Ring_theory.v
@@ -6,9 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-Require Import Setoid.
-Require Import BinPos.
-Require Import BinNat.
+Require Import Setoid Morphisms BinPos BinNat.
Set Implicit Arguments.
@@ -35,48 +33,42 @@ Section Power.
Variable rI : R.
Variable rmul : R -> R -> R.
Variable req : R -> R -> Prop.
- Variable Rsth : Setoid_Theory R req.
- Notation "x * y " := (rmul x y).
- Notation "x == y" := (req x y).
+ Variable Rsth : Equivalence req.
+ Infix "*" := rmul.
+ Infix "==" := req.
- Hypothesis mul_ext :
- forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2.
- Hypothesis mul_comm : forall x y, x * y == y * x.
+ Hypothesis mul_ext : Proper (req ==> req ==> req) rmul.
Hypothesis mul_assoc : forall x y z, x * (y * z) == (x * y) * z.
- Add Setoid R req Rsth as R_set_Power.
- Add Morphism rmul : rmul_ext_Power. exact mul_ext. Qed.
-
- Fixpoint pow_pos (x:R) (i:positive) {struct i}: R :=
+ Fixpoint pow_pos (x:R) (i:positive) : R :=
match i with
| xH => x
- | xO i => let p := pow_pos x i in rmul p p
- | xI i => let p := pow_pos x i in rmul x (rmul p p)
+ | xO i => let p := pow_pos x i in p * p
+ | xI i => let p := pow_pos x i in x * (p * p)
end.
- Lemma pow_pos_Psucc : forall x j, pow_pos x (Psucc j) == x * pow_pos x j.
+ Lemma pow_pos_swap x j : pow_pos x j * x == x * pow_pos x j.
Proof.
- induction j;simpl.
- rewrite IHj.
- rewrite (mul_comm x (pow_pos x j *pow_pos x j)).
- setoid_rewrite (mul_comm x (pow_pos x j)) at 2.
- repeat rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
- repeat rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
- apply (Seq_refl _ _ Rsth).
+ induction j; simpl; rewrite <- ?mul_assoc.
+ - f_equiv. now do 2 (rewrite IHj, mul_assoc).
+ - now do 2 (rewrite IHj, mul_assoc).
+ - reflexivity.
Qed.
- Lemma pow_pos_Pplus : forall x i j, pow_pos x (i + j) == pow_pos x i * pow_pos x j.
+ Lemma pow_pos_succ x j :
+ pow_pos x (Pos.succ j) == x * pow_pos x j.
Proof.
- intro x;induction i;intros.
- rewrite xI_succ_xO;rewrite Pplus_one_succ_r.
- rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
- repeat rewrite IHi.
- rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite pow_pos_Psucc.
- simpl;repeat rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
- rewrite <- Pplus_diag;repeat rewrite <- Pplus_assoc.
- repeat rewrite IHi;rewrite mul_assoc. apply (Seq_refl _ _ Rsth).
- rewrite Pplus_comm;rewrite <- Pplus_one_succ_r;rewrite pow_pos_Psucc;
- simpl. apply (Seq_refl _ _ Rsth).
+ induction j; simpl; try reflexivity.
+ rewrite IHj, <- mul_assoc; f_equiv.
+ now rewrite mul_assoc, pow_pos_swap, mul_assoc.
+ Qed.
+
+ Lemma pow_pos_add x i j :
+ pow_pos x (i + j) == pow_pos x i * pow_pos x j.
+ Proof.
+ induction i using Pos.peano_ind.
+ - now rewrite Pos.add_1_l, pow_pos_succ.
+ - now rewrite Pos.add_succ_l, !pow_pos_succ, IHi, mul_assoc.
Qed.
Definition pow_N (x:R) (p:N) :=
@@ -87,9 +79,9 @@ Section Power.
Definition id_phi_N (x:N) : N := x.
- Lemma pow_N_pow_N : forall x n, pow_N x (id_phi_N n) == pow_N x n.
+ Lemma pow_N_pow_N x n : pow_N x (id_phi_N n) == pow_N x n.
Proof.
- intros; apply (Seq_refl _ _ Rsth).
+ reflexivity.
Qed.
End Power.
@@ -98,19 +90,18 @@ Section DEFINITIONS.
Variable R : Type.
Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
Variable req : R -> R -> Prop.
- Notation "0" := rO. Notation "1" := rI.
- Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y).
- Notation "x - y " := (rsub x y). Notation "- x" := (ropp x).
- Notation "x == y" := (req x y).
+ Notation "0" := rO. Notation "1" := rI.
+ Infix "==" := req. Infix "+" := radd. Infix "*" := rmul.
+ Infix "-" := rsub. Notation "- x" := (ropp x).
(** Semi Ring *)
Record semi_ring_theory : Prop := mk_srt {
SRadd_0_l : forall n, 0 + n == n;
- SRadd_comm : forall n m, n + m == m + n ;
+ SRadd_comm : forall n m, n + m == m + n ;
SRadd_assoc : forall n m p, n + (m + p) == (n + m) + p;
SRmul_1_l : forall n, 1*n == n;
SRmul_0_l : forall n, 0*n == 0;
- SRmul_comm : forall n m, n*m == m*n;
+ SRmul_comm : forall n m, n*m == m*n;
SRmul_assoc : forall n m p, n*(m*p) == (n*m)*p;
SRdistr_l : forall n m p, (n + m)*p == n*p + m*p
}.
@@ -119,11 +110,11 @@ Section DEFINITIONS.
(*Almost ring are no ring : Ropp_def is missing **)
Record almost_ring_theory : Prop := mk_art {
ARadd_0_l : forall x, 0 + x == x;
- ARadd_comm : forall x y, x + y == y + x;
+ ARadd_comm : forall x y, x + y == y + x;
ARadd_assoc : forall x y z, x + (y + z) == (x + y) + z;
ARmul_1_l : forall x, 1 * x == x;
ARmul_0_l : forall x, 0 * x == 0;
- ARmul_comm : forall x y, x * y == y * x;
+ ARmul_comm : forall x y, x * y == y * x;
ARmul_assoc : forall x y z, x * (y * z) == (x * y) * z;
ARdistr_l : forall x y z, (x + y) * z == (x * z) + (y * z);
ARopp_mul_l : forall x y, -(x * y) == -x * y;
@@ -134,10 +125,10 @@ Section DEFINITIONS.
(** Ring *)
Record ring_theory : Prop := mk_rt {
Radd_0_l : forall x, 0 + x == x;
- Radd_comm : forall x y, x + y == y + x;
+ Radd_comm : forall x y, x + y == y + x;
Radd_assoc : forall x y z, x + (y + z) == (x + y) + z;
Rmul_1_l : forall x, 1 * x == x;
- Rmul_comm : forall x y, x * y == y * x;
+ Rmul_comm : forall x y, x * y == y * x;
Rmul_assoc : forall x y z, x * (y * z) == (x * y) * z;
Rdistr_l : forall x y z, (x + y) * z == (x * z) + (y * z);
Rsub_def : forall x y, x - y == x + -y;
@@ -148,19 +139,15 @@ Section DEFINITIONS.
Record sring_eq_ext : Prop := mk_seqe {
(* SRing operators are compatible with equality *)
- SRadd_ext :
- forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 + y1 == x2 + y2;
- SRmul_ext :
- forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2
+ SRadd_ext : Proper (req ==> req ==> req) radd;
+ SRmul_ext : Proper (req ==> req ==> req) rmul
}.
Record ring_eq_ext : Prop := mk_reqe {
(* Ring operators are compatible with equality *)
- Radd_ext :
- forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 + y1 == x2 + y2;
- Rmul_ext :
- forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 * y1 == x2 * y2;
- Ropp_ext : forall x1 x2, x1 == x2 -> -x1 == -x2
+ Radd_ext : Proper (req ==> req ==> req) radd;
+ Rmul_ext : Proper (req ==> req ==> req) rmul;
+ Ropp_ext : Proper (req ==> req) ropp
}.
(** Interpretation morphisms definition*)
@@ -170,9 +157,9 @@ Section DEFINITIONS.
Variable ceqb : C->C->bool.
(* [phi] est un morphisme de [C] dans [R] *)
Variable phi : C -> R.
- Notation "x +! y" := (cadd x y). Notation "x -! y " := (csub x y).
- Notation "x *! y " := (cmul x y). Notation "-! x" := (copp x).
- Notation "x ?=! y" := (ceqb x y). Notation "[ x ]" := (phi x).
+ Infix "+!" := cadd. Infix "-!" := csub.
+ Infix "*!" := cmul. Notation "-! x" := (copp x).
+ Infix "?=!" := ceqb. Notation "[ x ]" := (phi x).
(*for semi rings*)
Record semi_morph : Prop := mkRmorph {
@@ -216,15 +203,13 @@ Section DEFINITIONS.
End MORPHISM.
(** Identity is a morphism *)
- Variable Rsth : Setoid_Theory R req.
- Add Setoid R req Rsth as R_setoid1.
+ Variable Rsth : Equivalence req.
Variable reqb : R->R->bool.
Hypothesis morph_req : forall x y, (reqb x y) = true -> x == y.
Definition IDphi (x:R) := x.
Lemma IDmorph : ring_morph rO rI radd rmul rsub ropp reqb IDphi.
Proof.
- apply (mkmorph rO rI radd rmul rsub ropp reqb IDphi);intros;unfold IDphi;
- try apply (Seq_refl _ _ Rsth);auto.
+ now apply (mkmorph rO rI radd rmul rsub ropp reqb IDphi).
Qed.
(** Specification of the power function *)
@@ -239,35 +224,31 @@ Section DEFINITIONS.
End POWER.
- Definition pow_N_th := mkpow_th id_phi_N (pow_N rI rmul) (pow_N_pow_N rI rmul Rsth).
+ Definition pow_N_th :=
+ mkpow_th id_phi_N (pow_N rI rmul) (pow_N_pow_N rI rmul Rsth).
End DEFINITIONS.
-
-
Section ALMOST_RING.
Variable R : Type.
Variable (rO rI : R) (radd rmul rsub: R->R->R) (ropp : R -> R).
Variable req : R -> R -> Prop.
- Notation "0" := rO. Notation "1" := rI.
- Notation "x + y" := (radd x y). Notation "x * y " := (rmul x y).
- Notation "x - y " := (rsub x y). Notation "- x" := (ropp x).
- Notation "x == y" := (req x y).
+ Notation "0" := rO. Notation "1" := rI.
+ Infix "==" := req. Infix "+" := radd. Infix "* " := rmul.
+ Infix "-" := rsub. Notation "- x" := (ropp x).
(** Leibniz equality leads to a setoid theory and is extensional*)
- Lemma Eqsth : Setoid_Theory R (@eq R).
- Proof. constructor;red;intros;subst;trivial. Qed.
+ Lemma Eqsth : Equivalence (@eq R).
+ Proof. exact eq_equivalence. Qed.
Lemma Eq_s_ext : sring_eq_ext radd rmul (@eq R).
- Proof. constructor;intros;subst;trivial. Qed.
+ Proof. constructor;solve_proper. Qed.
Lemma Eq_ext : ring_eq_ext radd rmul ropp (@eq R).
- Proof. constructor;intros;subst;trivial. Qed.
+ Proof. constructor;solve_proper. Qed.
- Variable Rsth : Setoid_Theory R req.
- Add Setoid R req Rsth as R_setoid2.
- Ltac sreflexivity := apply (Seq_refl _ _ Rsth).
+ Variable Rsth : Equivalence req.
Section SEMI_RING.
Variable SReqe : sring_eq_ext radd rmul req.
@@ -282,23 +263,24 @@ Section ALMOST_RING.
Definition SRsub x y := x + -y. Notation "x - y " := (SRsub x y).
Lemma SRopp_ext : forall x y, x == y -> -x == -y.
- Proof. intros x y H;exact H. Qed.
+ Proof. intros x y H; exact H. Qed.
Lemma SReqe_Reqe : ring_eq_ext radd rmul SRopp req.
Proof.
- constructor. exact (SRadd_ext SReqe). exact (SRmul_ext SReqe).
- exact SRopp_ext.
+ constructor.
+ - exact (SRadd_ext SReqe).
+ - exact (SRmul_ext SReqe).
+ - exact SRopp_ext.
Qed.
Lemma SRopp_mul_l : forall x y, -(x * y) == -x * y.
- Proof. intros;sreflexivity. Qed.
+ Proof. reflexivity. Qed.
Lemma SRopp_add : forall x y, -(x + y) == -x + -y.
- Proof. intros;sreflexivity. Qed.
-
+ Proof. reflexivity. Qed.
Lemma SRsub_def : forall x y, x - y == x + -y.
- Proof. intros;sreflexivity. Qed.
+ Proof. reflexivity. Qed.
Lemma SRth_ARth : almost_ring_theory 0 1 radd rmul SRsub SRopp req.
Proof (mk_art 0 1 radd rmul SRsub SRopp req
@@ -315,7 +297,7 @@ Section ALMOST_RING.
Definition SRIDmorph : ring_morph 0 1 radd rmul SRsub SRopp req
0 1 radd rmul SRsub SRopp reqb (@IDphi R).
Proof.
- apply mkmorph;intros;try sreflexivity. unfold IDphi;auto.
+ now apply mkmorph.
Qed.
(* a semi_morph can be extended to a ring_morph for the almost_ring derived
@@ -331,9 +313,7 @@ Section ALMOST_RING.
ring_morph rO rI radd rmul SRsub SRopp req
cO cI cadd cmul cadd (fun x => x) ceqb phi.
Proof.
- case Smorph; intros; constructor; auto.
- unfold SRopp in |- *; intros.
- setoid_reflexivity.
+ case Smorph; now constructor.
Qed.
End SEMI_RING.
@@ -347,31 +327,28 @@ Section ALMOST_RING.
Variable Rth : ring_theory 0 1 radd rmul rsub ropp req.
(** Rings are almost rings*)
- Lemma Rmul_0_l : forall x, 0 * x == 0.
+ Lemma Rmul_0_l x : 0 * x == 0.
Proof.
- intro x; setoid_replace (0*x) with ((0+1)*x + -x).
- rewrite (Radd_0_l Rth); rewrite (Rmul_1_l Rth).
- rewrite (Ropp_def Rth);sreflexivity.
+ setoid_replace (0*x) with ((0+1)*x + -x).
+ now rewrite (Radd_0_l Rth), (Rmul_1_l Rth), (Ropp_def Rth).
- rewrite (Rdistr_l Rth);rewrite (Rmul_1_l Rth).
- rewrite <- (Radd_assoc Rth); rewrite (Ropp_def Rth).
- rewrite (Radd_comm Rth); rewrite (Radd_0_l Rth);sreflexivity.
+ rewrite (Rdistr_l Rth), (Rmul_1_l Rth).
+ rewrite <- (Radd_assoc Rth), (Ropp_def Rth).
+ now rewrite (Radd_comm Rth), (Radd_0_l Rth).
Qed.
- Lemma Ropp_mul_l : forall x y, -(x * y) == -x * y.
+ Lemma Ropp_mul_l x y : -(x * y) == -x * y.
Proof.
- intros x y;rewrite <-(Radd_0_l Rth (- x * y)).
- rewrite (Radd_comm Rth).
- rewrite <-(Ropp_def Rth (x*y)).
- rewrite (Radd_assoc Rth).
- rewrite <- (Rdistr_l Rth).
- rewrite (Rth.(Radd_comm) (-x));rewrite (Ropp_def Rth).
- rewrite Rmul_0_l;rewrite (Radd_0_l Rth);sreflexivity.
+ rewrite <-(Radd_0_l Rth (- x * y)).
+ rewrite (Radd_comm Rth), <-(Ropp_def Rth (x*y)).
+ rewrite (Radd_assoc Rth), <- (Rdistr_l Rth).
+ rewrite (Rth.(Radd_comm) (-x)), (Ropp_def Rth).
+ now rewrite Rmul_0_l, (Radd_0_l Rth).
Qed.
- Lemma Ropp_add : forall x y, -(x + y) == -x + -y.
+ Lemma Ropp_add x y : -(x + y) == -x + -y.
Proof.
- intros x y;rewrite <- ((Radd_0_l Rth) (-(x+y))).
+ rewrite <- ((Radd_0_l Rth) (-(x+y))).
rewrite <- ((Ropp_def Rth) x).
rewrite <- ((Radd_0_l Rth) (x + - x + - (x + y))).
rewrite <- ((Ropp_def Rth) y).
@@ -383,17 +360,17 @@ Section ALMOST_RING.
rewrite ((Radd_comm Rth) y).
rewrite <- ((Radd_assoc Rth) (- x)).
rewrite ((Radd_assoc Rth) y).
- rewrite ((Radd_comm Rth) y);rewrite (Ropp_def Rth).
- rewrite ((Radd_comm Rth) (-x) 0);rewrite (Radd_0_l Rth).
- apply (Radd_comm Rth).
+ rewrite ((Radd_comm Rth) y), (Ropp_def Rth).
+ rewrite ((Radd_comm Rth) (-x) 0), (Radd_0_l Rth).
+ now apply (Radd_comm Rth).
Qed.
- Lemma Ropp_opp : forall x, - -x == x.
+ Lemma Ropp_opp x : - -x == x.
Proof.
- intros x; rewrite <- (Radd_0_l Rth (- -x)).
+ rewrite <- (Radd_0_l Rth (- -x)).
rewrite <- (Ropp_def Rth x).
- rewrite <- (Radd_assoc Rth); rewrite (Ropp_def Rth).
- rewrite ((Radd_comm Rth) x);apply (Radd_0_l Rth).
+ rewrite <- (Radd_assoc Rth), (Ropp_def Rth).
+ rewrite ((Radd_comm Rth) x); now apply (Radd_0_l Rth).
Qed.
Lemma Rth_ARth : almost_ring_theory 0 1 radd rmul rsub ropp req.
@@ -407,10 +384,10 @@ Section ALMOST_RING.
Variable (cO cI : C) (cadd cmul csub: C->C->C) (copp : C -> C).
Variable (ceq : C -> C -> Prop) (ceqb : C -> C -> bool).
Variable phi : C -> R.
- Notation "x +! y" := (cadd x y). Notation "x *! y " := (cmul x y).
- Notation "x -! y " := (csub x y). Notation "-! x" := (copp x).
- Notation "x ?=! y" := (ceqb x y). Notation "[ x ]" := (phi x).
- Variable Csth : Setoid_Theory C ceq.
+ Infix "+!" := cadd. Infix "*!" := cmul.
+ Infix "-!" := csub. Notation "-! x" := (copp x).
+ Notation "?=!" := ceqb. Notation "[ x ]" := (phi x).
+ Variable Csth : Equivalence ceq.
Variable Ceqe : ring_eq_ext cadd cmul copp ceq.
Add Setoid C ceq Csth as C_setoid.
Add Morphism cadd : cadd_ext. exact (Radd_ext Ceqe). Qed.
@@ -420,9 +397,9 @@ Section ALMOST_RING.
Variable Smorph : semi_morph 0 1 radd rmul req cO cI cadd cmul ceqb phi.
Variable phi_ext : forall x y, ceq x y -> [x] == [y].
Add Morphism phi : phi_ext1. exact phi_ext. Qed.
- Lemma Smorph_opp : forall x, [-!x] == -[x].
+ Lemma Smorph_opp x : [-!x] == -[x].
Proof.
- intros x;rewrite <- (Rth.(Radd_0_l) [-!x]).
+ rewrite <- (Rth.(Radd_0_l) [-!x]).
rewrite <- ((Ropp_def Rth) [x]).
rewrite ((Radd_comm Rth) [x]).
rewrite <- (Radd_assoc Rth).
@@ -430,17 +407,18 @@ Section ALMOST_RING.
rewrite (Ropp_def Cth).
rewrite (Smorph0 Smorph).
rewrite (Radd_comm Rth (-[x])).
- apply (Radd_0_l Rth);sreflexivity.
+ now apply (Radd_0_l Rth).
Qed.
- Lemma Smorph_sub : forall x y, [x -! y] == [x] - [y].
+ Lemma Smorph_sub x y : [x -! y] == [x] - [y].
Proof.
- intros x y; rewrite (Rsub_def Cth);rewrite (Rsub_def Rth).
- rewrite (Smorph_add Smorph);rewrite Smorph_opp;sreflexivity.
+ rewrite (Rsub_def Cth), (Rsub_def Rth).
+ now rewrite (Smorph_add Smorph), Smorph_opp.
Qed.
- Lemma Smorph_morph : ring_morph 0 1 radd rmul rsub ropp req
- cO cI cadd cmul csub copp ceqb phi.
+ Lemma Smorph_morph :
+ ring_morph 0 1 radd rmul rsub ropp req
+ cO cI cadd cmul csub copp ceqb phi.
Proof
(mkmorph 0 1 radd rmul rsub ropp req cO cI cadd cmul csub copp ceqb phi
(Smorph0 Smorph) (Smorph1 Smorph)
@@ -458,17 +436,11 @@ elim ARth; intros.
constructor; trivial.
Qed.
- Lemma ARsub_ext :
- forall x1 x2, x1 == x2 -> forall y1 y2, y1 == y2 -> x1 - y1 == x2 - y2.
+ Instance ARsub_ext : Proper (req ==> req ==> req) rsub.
Proof.
- intros.
- setoid_replace (x1 - y1) with (x1 + -y1).
- setoid_replace (x2 - y2) with (x2 + -y2).
- rewrite H;rewrite H0;sreflexivity.
- apply (ARsub_def ARth).
- apply (ARsub_def ARth).
+ intros x1 x2 Ex y1 y2 Ey.
+ now rewrite !(ARsub_def ARth), Ex, Ey.
Qed.
- Add Morphism rsub : rsub_ext. exact ARsub_ext. Qed.
Ltac mrewrite :=
repeat first
@@ -479,64 +451,56 @@ Qed.
| rewrite (ARmul_0_l ARth)
| rewrite <- ((ARmul_comm ARth) 0)
| rewrite (ARdistr_l ARth)
- | sreflexivity
+ | reflexivity
| match goal with
| |- context [?z * (?x + ?y)] => rewrite ((ARmul_comm ARth) z (x+y))
end].
- Lemma ARadd_0_r : forall x, (x + 0) == x.
- Proof. intros; mrewrite. Qed.
+ Lemma ARadd_0_r x : x + 0 == x.
+ Proof. mrewrite. Qed.
- Lemma ARmul_1_r : forall x, x * 1 == x.
- Proof. intros;mrewrite. Qed.
+ Lemma ARmul_1_r x : x * 1 == x.
+ Proof. mrewrite. Qed.
- Lemma ARmul_0_r : forall x, x * 0 == 0.
- Proof. intros;mrewrite. Qed.
+ Lemma ARmul_0_r x : x * 0 == 0.
+ Proof. mrewrite. Qed.
- Lemma ARdistr_r : forall x y z, z * (x + y) == z*x + z*y.
+ Lemma ARdistr_r x y z : z * (x + y) == z*x + z*y.
Proof.
- intros;mrewrite.
- repeat rewrite (ARth.(ARmul_comm) z);sreflexivity.
+ mrewrite. now rewrite !(ARth.(ARmul_comm) z).
Qed.
- Lemma ARadd_assoc1 : forall x y z, (x + y) + z == (y + z) + x.
+ Lemma ARadd_assoc1 x y z : (x + y) + z == (y + z) + x.
Proof.
- intros;rewrite <-(ARth.(ARadd_assoc) x).
- rewrite (ARth.(ARadd_comm) x);sreflexivity.
+ now rewrite <-(ARth.(ARadd_assoc) x), (ARth.(ARadd_comm) x).
Qed.
- Lemma ARadd_assoc2 : forall x y z, (y + x) + z == (y + z) + x.
+ Lemma ARadd_assoc2 x y z : (y + x) + z == (y + z) + x.
Proof.
- intros; repeat rewrite <- (ARadd_assoc ARth);
- rewrite ((ARadd_comm ARth) x); sreflexivity.
+ now rewrite <- !(ARadd_assoc ARth), ((ARadd_comm ARth) x).
Qed.
- Lemma ARmul_assoc1 : forall x y z, (x * y) * z == (y * z) * x.
+ Lemma ARmul_assoc1 x y z : (x * y) * z == (y * z) * x.
Proof.
- intros;rewrite <-((ARmul_assoc ARth) x).
- rewrite ((ARmul_comm ARth) x);sreflexivity.
+ now rewrite <- ((ARmul_assoc ARth) x), ((ARmul_comm ARth) x).
Qed.
- Lemma ARmul_assoc2 : forall x y z, (y * x) * z == (y * z) * x.
+ Lemma ARmul_assoc2 x y z : (y * x) * z == (y * z) * x.
Proof.
- intros; repeat rewrite <- (ARmul_assoc ARth);
- rewrite ((ARmul_comm ARth) x); sreflexivity.
+ now rewrite <- !(ARmul_assoc ARth), ((ARmul_comm ARth) x).
Qed.
- Lemma ARopp_mul_r : forall x y, - (x * y) == x * -y.
+ Lemma ARopp_mul_r x y : - (x * y) == x * -y.
Proof.
- intros;rewrite ((ARmul_comm ARth) x y);
- rewrite (ARopp_mul_l ARth); apply (ARmul_comm ARth).
+ rewrite ((ARmul_comm ARth) x y), (ARopp_mul_l ARth).
+ now apply (ARmul_comm ARth).
Qed.
Lemma ARopp_zero : -0 == 0.
Proof.
- rewrite <- (ARmul_0_r 0); rewrite (ARopp_mul_l ARth).
- repeat rewrite ARmul_0_r; sreflexivity.
+ now rewrite <- (ARmul_0_r 0), (ARopp_mul_l ARth), !ARmul_0_r.
Qed.
-
-
End ALMOST_RING.
@@ -611,6 +575,8 @@ Ltac gen_add_push add Rsth Reqe ARth x :=
progress rewrite (ARadd_assoc2 Rsth Reqe ARth x y z)
| |- context [add (add x ?y) ?z] =>
progress rewrite (ARadd_assoc1 Rsth ARth x y z)
+ | |- context [(add x ?y)] =>
+ progress rewrite (ARadd_comm ARth x y)
end).
Ltac gen_mul_push mul Rsth Reqe ARth x :=
@@ -619,5 +585,6 @@ Ltac gen_mul_push mul Rsth Reqe ARth x :=
progress rewrite (ARmul_assoc2 Rsth Reqe ARth x y z)
| |- context [mul (mul x ?y) ?z] =>
progress rewrite (ARmul_assoc1 Rsth ARth x y z)
+ | |- context [(mul x ?y)] =>
+ progress rewrite (ARmul_comm ARth x y)
end).
-
diff --git a/plugins/setoid_ring/Rings_Z.v b/plugins/setoid_ring/Rings_Z.v
index 889048653..58a4d7ea6 100644
--- a/plugins/setoid_ring/Rings_Z.v
+++ b/plugins/setoid_ring/Rings_Z.v
@@ -3,7 +3,7 @@ Require Export Integral_domain.
Require Export Ncring_initial.
Instance Zcri: (Cring (Rr:=Zr)).
-red. exact Zmult_comm. Defined.
+red. exact Z.mul_comm. Defined.
Lemma Z_one_zero: 1%Z <> 0%Z.
omega.
diff --git a/plugins/setoid_ring/ZArithRing.v b/plugins/setoid_ring/ZArithRing.v
index d3ed36ee9..281ffa229 100644
--- a/plugins/setoid_ring/ZArithRing.v
+++ b/plugins/setoid_ring/ZArithRing.v
@@ -39,14 +39,14 @@ Ltac Zpower_neg :=
repeat match goal with
| [|- ?G] =>
match G with
- | context c [Zpower _ (Zneg _)] =>
+ | context c [Z.pow _ (Zneg _)] =>
let t := context c [Z0] in
change t
end
end.
Add Ring Zr : Zth
- (decidable Zeq_bool_eq, constants [Zcst], preprocess [Zpower_neg;unfold Zsucc],
+ (decidable Zeq_bool_eq, constants [Zcst], preprocess [Zpower_neg;unfold Z.succ],
power_tac Zpower_theory [Zpow_tac],
(* The two following option are not needed, it is the default chose when the set of
coefficiant is usual ring Z *)
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 0d32ed46f..ad44171da 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -3424,7 +3424,7 @@ let intros_symmetry =
(* Transitivity tactics *)
-(* This tactic first tries to apply a constant named trans_eq, where eq
+(* This tactic first tries to apply a constant named eq_trans, where eq
is the name of the equality predicate. If this constant is not
defined and the conclusion is a=b, it solves the goal doing
Cut x1=x2;
diff --git a/test-suite/bugs/closed/shouldsucceed/1414.v b/test-suite/bugs/closed/shouldsucceed/1414.v
index 495a16bca..ee9e2504a 100644
--- a/test-suite/bugs/closed/shouldsucceed/1414.v
+++ b/test-suite/bugs/closed/shouldsucceed/1414.v
@@ -26,13 +26,13 @@ Program Fixpoint union
| t1,Leaf => t1
| Node l1 v1 r1 h1, Node l2 v2 r2 h2 =>
if (Z_ge_lt_dec h1 h2) then
- if (Z_eq_dec h2 1)
+ if (Z.eq_dec h2 1)
then add v2 s
else
let (l2', r2') := split v1 u in
join (union l1 l2' _ _ _ _) v1 (union r1 (snd r2') _ _ _ _)
else
- if (Z_eq_dec h1 1)
+ if (Z.eq_dec h1 1)
then add v1 s
else
let (l1', r1') := split v2 u in
diff --git a/test-suite/bugs/closed/shouldsucceed/1784.v b/test-suite/bugs/closed/shouldsucceed/1784.v
index 718b0e861..fb2f0ca90 100644
--- a/test-suite/bugs/closed/shouldsucceed/1784.v
+++ b/test-suite/bugs/closed/shouldsucceed/1784.v
@@ -58,7 +58,7 @@ Program Fixpoint lt_dec (x y:sv) { struct x } : {slt x y}+{~slt x y} :=
match x with
| I x =>
match y with
- | I y => if (Z_eq_dec x y) then in_left else in_right
+ | I y => if (Z.eq_dec x y) then in_left else in_right
| S ys => in_right
end
| S xs =>
diff --git a/test-suite/bugs/closed/shouldsucceed/1844.v b/test-suite/bugs/closed/shouldsucceed/1844.v
index 5627612f6..17eeb3529 100644
--- a/test-suite/bugs/closed/shouldsucceed/1844.v
+++ b/test-suite/bugs/closed/shouldsucceed/1844.v
@@ -1,6 +1,6 @@
Require Import ZArith.
-Definition zeq := Z_eq_dec.
+Definition zeq := Z.eq_dec.
Definition update (A: Set) (x: Z) (v: A) (s: Z -> A) : Z -> A :=
fun y => if zeq x y then v else s y.
diff --git a/test-suite/bugs/closed/shouldsucceed/1935.v b/test-suite/bugs/closed/shouldsucceed/1935.v
index 72396d490..d58376198 100644
--- a/test-suite/bugs/closed/shouldsucceed/1935.v
+++ b/test-suite/bugs/closed/shouldsucceed/1935.v
@@ -13,7 +13,7 @@ Qed.
Require Import ZArith.
-Definition f'' (a:bool) := if a then eq (A:= Z) else Zlt.
+Definition f'' (a:bool) := if a then eq (A:= Z) else Z.lt.
Lemma f_refl'' : forall n , f'' true n n.
Proof.
diff --git a/test-suite/bugs/closed/shouldsucceed/2127.v b/test-suite/bugs/closed/shouldsucceed/2127.v
index 0fc854b6f..142ada268 100644
--- a/test-suite/bugs/closed/shouldsucceed/2127.v
+++ b/test-suite/bugs/closed/shouldsucceed/2127.v
@@ -1,8 +1,8 @@
-(* Check that "apply refl_equal" is not exported as an interactive
+(* Check that "apply eq_refl" is not exported as an interactive
tactic but as a statically globalized one *)
(* (this is a simplification of the original bug report) *)
Module A.
-Hint Rewrite sym_equal using apply refl_equal : foo.
+Hint Rewrite eq_sym using apply eq_refl : foo.
End A.
diff --git a/test-suite/complexity/ring2.v b/test-suite/complexity/ring2.v
index 6945edc88..c3634f640 100644
--- a/test-suite/complexity/ring2.v
+++ b/test-suite/complexity/ring2.v
@@ -3,7 +3,7 @@
Require Import BinInt Zbool.
-Definition Zplus x y :=
+Definition Z.add x y :=
match x with
| 0%Z => y
| Zpos x' =>
@@ -32,7 +32,7 @@ end.
Require Import Ring.
-Lemma Zth : ring_theory Z0 (Zpos xH) Zplus Zmult Zminus Zopp (@eq Z).
+Lemma Zth : ring_theory Z0 (Zpos xH) Z.add Z.mul Z.sub Z.opp (@eq Z).
Admitted.
Ltac Zcst t :=
@@ -45,7 +45,7 @@ Add Ring Zr : Zth
(decidable Zeq_bool_eq, constants [Zcst]).
Open Scope Z_scope.
-Infix "+" := Zplus : Z_scope.
+Infix "+" := Z.add : Z_scope.
Goal forall a, a+a+a+a+a+a+a+a+a+a+a+a+a = a*13.
Timeout 5 Time intro; ring.
diff --git a/test-suite/ideal-features/eapply_evar.v b/test-suite/ideal-features/eapply_evar.v
index 547860bf1..bb61afb88 100644
--- a/test-suite/ideal-features/eapply_evar.v
+++ b/test-suite/ideal-features/eapply_evar.v
@@ -4,6 +4,6 @@
and not "O = O" *)
Lemma eapply_evar : O=O -> 0=O.
-intro H; eapply trans_equal;
+intro H; eapply eq_trans;
[apply H | match goal with |- ?x = ?x => reflexivity end].
Qed.
diff --git a/test-suite/micromega/example.v b/test-suite/micromega/example.v
index f5a953566..4a1231b35 100644
--- a/test-suite/micromega/example.v
+++ b/test-suite/micromega/example.v
@@ -77,13 +77,13 @@ Definition rbound2 (C:Z -> Z -> Z) : Prop :=
Lemma bounded_drift : forall s t p q C D, s <= t /\ correct p t /\ correct q t /\
rbound1 C /\ rbound2 C /\ rbound1 D /\ rbound2 D ->
- Zabs (C p t - D q t) <= Zabs (C p s - D q s) + 2 * rho * (t- s).
+ Z.abs (C p t - D q t) <= Z.abs (C p s - D q s) + 2 * rho * (t- s).
Proof.
intros.
- generalize (Zabs_eq (C p t - D q t)).
- generalize (Zabs_non_eq (C p t - D q t)).
- generalize (Zabs_eq (C p s -D q s)).
- generalize (Zabs_non_eq (C p s - D q s)).
+ generalize (Z.abs_eq (C p t - D q t)).
+ generalize (Z.abs_neq (C p t - D q t)).
+ generalize (Z.abs_eq (C p s -D q s)).
+ generalize (Z.abs_neq (C p s - D q s)).
unfold rbound2 in H.
unfold rbound1 in H.
intuition.
diff --git a/test-suite/micromega/square.v b/test-suite/micromega/square.v
index 4c00ffe4a..6fa547358 100644
--- a/test-suite/micromega/square.v
+++ b/test-suite/micromega/square.v
@@ -9,17 +9,17 @@
Require Import ZArith Zwf Psatz QArith.
Open Scope Z_scope.
-Lemma Zabs_square : forall x, (Zabs x)^2 = x^2.
+Lemma Z.abs_square : forall x, (Z.abs x)^2 = x^2.
Proof.
intros ; case (Zabs_dec x) ; intros ; psatz Z 2.
Qed.
-Hint Resolve Zabs_pos Zabs_square.
+Hint Resolve Z.abs_nonneg Z.abs_square.
Lemma integer_statement : ~exists n, exists p, n^2 = 2*p^2 /\ n <> 0.
Proof.
-intros [n [p [Heq Hnz]]]; pose (n' := Zabs n); pose (p':=Zabs p).
-assert (facts : 0 <= Zabs n /\ 0 <= Zabs p /\ Zabs n^2=n^2
- /\ Zabs p^2 = p^2) by auto.
+intros [n [p [Heq Hnz]]]; pose (n' := Z.abs n); pose (p':=Z.abs p).
+assert (facts : 0 <= Z.abs n /\ 0 <= Z.abs p /\ Z.abs n^2=n^2
+ /\ Z.abs p^2 = p^2) by auto.
assert (H : (0 < n' /\ 0 <= p' /\ n' ^2 = 2* p' ^2)) by
(destruct facts as [Hf1 [Hf2 [Hf3 Hf4]]]; unfold n', p' ; psatz Z 2).
generalize p' H; elim n' using (well_founded_ind (Zwf_well_founded 0)); clear.
@@ -35,7 +35,7 @@ Lemma QnumZpower : forall x : Q, Qnum (x ^ 2)%Q = ((Qnum x) ^ 2) %Z.
Proof.
intros.
destruct x.
- cbv beta iota zeta delta - [Zmult].
+ cbv beta iota zeta delta - [Z.mul].
ring.
Qed.
@@ -45,15 +45,15 @@ Proof.
intros.
destruct x.
simpl.
- unfold Zpower_pos.
+ unfold Z.pow_pos.
simpl.
- rewrite Pmult_1_r.
+ rewrite Pos.mul_1_r.
reflexivity.
Qed.
Theorem sqrt2_not_rational : ~exists x:Q, x^2==2#1.
Proof.
- unfold Qeq; intros [x]; simpl (Qden (2#1)); rewrite Zmult_1_r.
+ unfold Qeq; intros [x]; simpl (Qden (2#1)); rewrite Z.mul_1_r.
intros HQeq.
assert (Heq : (Qnum x ^ 2 = 2 * ' Qden x ^ 2%Q)%Z) by
(rewrite QnumZpower in HQeq ; rewrite QdenZpower in HQeq ; auto).
diff --git a/test-suite/success/Hints.v b/test-suite/success/Hints.v
index 071fb9579..cc8cec470 100644
--- a/test-suite/success/Hints.v
+++ b/test-suite/success/Hints.v
@@ -2,26 +2,26 @@
(* Checks that qualified names are accepted *)
(* New-style syntax *)
-Hint Resolve refl_equal: core arith.
-Hint Immediate trans_equal.
-Hint Unfold sym_equal: core.
+Hint Resolve eq_refl: core arith.
+Hint Immediate eq_trans.
+Hint Unfold eq_sym: core.
Hint Constructors eq: foo bar.
-Hint Extern 3 (_ = _) => apply refl_equal: foo bar.
+Hint Extern 3 (_ = _) => apply eq_refl: foo bar.
(* Old-style syntax *)
-Hint Resolve refl_equal sym_equal.
-Hint Resolve refl_equal sym_equal: foo.
-Hint Immediate refl_equal sym_equal.
-Hint Immediate refl_equal sym_equal: foo.
-Hint Unfold fst sym_equal.
-Hint Unfold fst sym_equal: foo.
+Hint Resolve eq_refl eq_sym.
+Hint Resolve eq_refl eq_sym: foo.
+Hint Immediate eq_refl eq_sym.
+Hint Immediate eq_refl eq_sym: foo.
+Hint Unfold fst eq_sym.
+Hint Unfold fst eq_sym: foo.
(* Checks that local names are accepted *)
Section A.
Remark Refl : forall (A : Set) (x : A), x = x.
- Proof. exact @refl_equal. Defined.
- Definition Sym := sym_equal.
- Let Trans := trans_equal.
+ Proof. exact @eq_refl. Defined.
+ Definition Sym := eq_sym.
+ Let Trans := eq_trans.
Hint Resolve Refl: foo.
Hint Resolve Sym: bar.
diff --git a/test-suite/success/MatchFail.v b/test-suite/success/MatchFail.v
index d43cd0e87..7069bba43 100644
--- a/test-suite/success/MatchFail.v
+++ b/test-suite/success/MatchFail.v
@@ -12,13 +12,13 @@ Ltac compute_POS :=
let v := constr:X1 in
match constr:v with
| 1%positive => fail 1
- | _ => rewrite (BinInt.Zpos_xI v)
+ | _ => rewrite (BinInt.Pos2Z.inj_xI v)
end
| |- context [(Zpos (xO ?X1))] =>
let v := constr:X1 in
match constr:v with
| 1%positive => fail 1
- | _ => rewrite (BinInt.Zpos_xO v)
+ | _ => rewrite (BinInt.Pos2Z.inj_xO v)
end
end.
diff --git a/test-suite/success/OmegaPre.v b/test-suite/success/OmegaPre.v
index f4996734b..17531064c 100644
--- a/test-suite/success/OmegaPre.v
+++ b/test-suite/success/OmegaPre.v
@@ -14,38 +14,38 @@ Open Scope Z_scope.
(* zify_op *)
-Goal forall a:Z, Zmax a a = a.
+Goal forall a:Z, Z.max a a = a.
intros.
omega with *.
Qed.
-Goal forall a b:Z, Zmax a b = Zmax b a.
+Goal forall a b:Z, Z.max a b = Z.max b a.
intros.
omega with *.
Qed.
-Goal forall a b c:Z, Zmax a (Zmax b c) = Zmax (Zmax a b) c.
+Goal forall a b c:Z, Z.max a (Z.max b c) = Z.max (Z.max a b) c.
intros.
omega with *.
Qed.
-Goal forall a b:Z, Zmax a b + Zmin a b = a + b.
+Goal forall a b:Z, Z.max a b + Z.min a b = a + b.
intros.
omega with *.
Qed.
-Goal forall a:Z, (Zabs a)*(Zsgn a) = a.
+Goal forall a:Z, (Z.abs a)*(Z.sgn a) = a.
intros.
zify.
intuition; subst; omega. (* pure multiplication: omega alone can't do it *)
Qed.
-Goal forall a:Z, Zabs a = a -> a >= 0.
+Goal forall a:Z, Z.abs a = a -> a >= 0.
intros.
omega with *.
Qed.
-Goal forall a:Z, Zsgn a = a -> a = 1 \/ a = 0 \/ a = -1.
+Goal forall a:Z, Z.sgn a = a -> a = 1 \/ a = 0 \/ a = -1.
intros.
omega with *.
Qed.
@@ -119,7 +119,7 @@ Qed.
(* mix of datatypes *)
-Goal forall p, Z_of_N (N_of_nat (nat_of_N (Npos p))) = Zpos p.
+Goal forall p, Z.of_N (N.of_nat (N.to_nat (Npos p))) = Zpos p.
intros.
omega with *.
Qed.
diff --git a/test-suite/success/ProgramWf.v b/test-suite/success/ProgramWf.v
index f8d3ad2c3..28ab33af3 100644
--- a/test-suite/success/ProgramWf.v
+++ b/test-suite/success/ProgramWf.v
@@ -22,14 +22,14 @@ Program Fixpoint merge (n m : nat) {measure (n + m) (lt)} : nat :=
Print merge.
-Print Zlt.
+Print Z.lt.
Print Zwf.
Open Local Scope Z_scope.
Program Fixpoint Zwfrec (n m : Z) {measure (n + m) (Zwf 0)} : Z :=
match n ?= m with
- | Lt => Zwfrec n (Zpred m)
+ | Lt => Zwfrec n (Z.pred m)
| _ => 0
end.
diff --git a/test-suite/success/ROmegaPre.v b/test-suite/success/ROmegaPre.v
index bd473fa60..fa659273e 100644
--- a/test-suite/success/ROmegaPre.v
+++ b/test-suite/success/ROmegaPre.v
@@ -14,38 +14,38 @@ Open Scope Z_scope.
(* zify_op *)
-Goal forall a:Z, Zmax a a = a.
+Goal forall a:Z, Z.max a a = a.
intros.
romega with *.
Qed.
-Goal forall a b:Z, Zmax a b = Zmax b a.
+Goal forall a b:Z, Z.max a b = Z.max b a.
intros.
romega with *.
Qed.
-Goal forall a b c:Z, Zmax a (Zmax b c) = Zmax (Zmax a b) c.
+Goal forall a b c:Z, Z.max a (Z.max b c) = Z.max (Z.max a b) c.
intros.
romega with *.
Qed.
-Goal forall a b:Z, Zmax a b + Zmin a b = a + b.
+Goal forall a b:Z, Z.max a b + Z.min a b = a + b.
intros.
romega with *.
Qed.
-Goal forall a:Z, (Zabs a)*(Zsgn a) = a.
+Goal forall a:Z, (Z.abs a)*(Z.sgn a) = a.
intros.
zify.
intuition; subst; romega. (* pure multiplication: omega alone can't do it *)
Qed.
-Goal forall a:Z, Zabs a = a -> a >= 0.
+Goal forall a:Z, Z.abs a = a -> a >= 0.
intros.
romega with *.
Qed.
-Goal forall a:Z, Zsgn a = a -> a = 1 \/ a = 0 \/ a = -1.
+Goal forall a:Z, Z.sgn a = a -> a = 1 \/ a = 0 \/ a = -1.
intros.
romega with *.
Qed.
@@ -119,7 +119,7 @@ Qed.
(* mix of datatypes *)
-Goal forall p, Z_of_N (N_of_nat (nat_of_N (Npos p))) = Zpos p.
+Goal forall p, Z.of_N (N.of_nat (N.to_nat (Npos p))) = Zpos p.
intros.
romega with *.
Qed.
diff --git a/test-suite/success/RecTutorial.v b/test-suite/success/RecTutorial.v
index 64048fe24..459645f6f 100644
--- a/test-suite/success/RecTutorial.v
+++ b/test-suite/success/RecTutorial.v
@@ -378,18 +378,18 @@ Inductive itree : Set :=
Definition isingle l := inode l (fun i => ileaf).
-Definition t1 := inode 0 (fun n => isingle (Z_of_nat (2*n))).
+Definition t1 := inode 0 (fun n => isingle (Z.of_nat (2*n))).
Definition t2 := inode 0
(fun n : nat =>
- inode (Z_of_nat n)
- (fun p => isingle (Z_of_nat (n*p)))).
+ inode (Z.of_nat n)
+ (fun p => isingle (Z.of_nat (n*p)))).
Inductive itree_le : itree-> itree -> Prop :=
| le_leaf : forall t, itree_le ileaf t
| le_node : forall l l' s s',
- Zle l l' ->
+ Z.le l l' ->
(forall i, exists j:nat, itree_le (s i) (s' j)) ->
itree_le (inode l s) (inode l' s').
@@ -424,7 +424,7 @@ Qed.
Inductive itree_le' : itree-> itree -> Prop :=
| le_leaf' : forall t, itree_le' ileaf t
| le_node' : forall l l' s s' g,
- Zle l l' ->
+ Z.le l l' ->
(forall i, itree_le' (s i) (s' (g i))) ->
itree_le' (inode l s) (inode l' s').
diff --git a/test-suite/success/Scopes.v b/test-suite/success/Scopes.v
index 5b37dbd4e..997fedd38 100644
--- a/test-suite/success/Scopes.v
+++ b/test-suite/success/Scopes.v
@@ -3,7 +3,7 @@
Require Import ZArith.
Module A.
-Definition opp := Zopp.
+Definition opp := Z.opp.
End A.
Check (A.opp 3).
diff --git a/test-suite/success/apply.v b/test-suite/success/apply.v
index d3c761019..0d8bf5560 100644
--- a/test-suite/success/apply.v
+++ b/test-suite/success/apply.v
@@ -8,8 +8,8 @@ Qed.
Require Import ZArith.
Goal (forall x y z, ~ z <= 0 -> x * z < y * z -> x <= y)%Z.
-intros; apply Znot_le_gt, Zgt_lt in H.
-apply Zmult_lt_reg_r, Zlt_le_weak in H0; auto.
+intros; apply Znot_le_gt, Z.gt_lt in H.
+apply Zmult_lt_reg_r, Z.lt_le_incl in H0; auto.
Qed.
(* Test application under tuples *)
@@ -266,7 +266,7 @@ Qed.
(* This works because unfold calls clos_norm_flags which calls nf_evar *)
Lemma eapply_evar_unfold : let x:=O in O=x -> 0=O.
-intros x H; eapply trans_equal;
+intros x H; eapply eq_trans;
[apply H | unfold x;match goal with |- ?x = ?x => reflexivity end].
Qed.
diff --git a/test-suite/success/decl_mode.v b/test-suite/success/decl_mode.v
index bc1757fd5..52575eca5 100644
--- a/test-suite/success/decl_mode.v
+++ b/test-suite/success/decl_mode.v
@@ -138,7 +138,7 @@ Coercion IZR: Z >->R.*)
Open Scope R_scope.
Lemma square_abs_square:
- forall p,(INR (Zabs_nat p) * INR (Zabs_nat p)) = (IZR p * IZR p).
+ forall p,(INR (Z.abs_nat p) * INR (Z.abs_nat p)) = (IZR p * IZR p).
proof.
assume p:Z.
per cases on p.
@@ -147,7 +147,7 @@ proof.
suppose it is (Zpos z).
thus thesis.
suppose it is (Zneg z).
- have ((INR (Zabs_nat (Zneg z)) * INR (Zabs_nat (Zneg z))) =
+ have ((INR (Z.abs_nat (Zneg z)) * INR (Z.abs_nat (Zneg z))) =
(IZR (Zpos z) * IZR (Zpos z))).
~= ((- IZR (Zpos z)) * (- IZR (Zpos z))).
thus ~= (IZR (Zneg z) * IZR (Zneg z)).
@@ -165,15 +165,15 @@ proof.
have H_in_R:(INR q<>0:>R) by H.
have triv:((IZR p/INR q* INR q) =IZR p :>R) by * using field.
have sqrt2:((sqrt (INR 2%nat) * sqrt (INR 2%nat))= INR 2%nat:>R) by sqrt_def.
- have (INR (Zabs_nat p * Zabs_nat p)
- = (INR (Zabs_nat p) * INR (Zabs_nat p)))
+ have (INR (Z.abs_nat p * Z.abs_nat p)
+ = (INR (Z.abs_nat p) * INR (Z.abs_nat p)))
by mult_INR.
~= (IZR p* IZR p) by square_abs_square.
~= ((IZR p/INR q*INR q)*(IZR p/INR q*INR q)) by triv. (* we have to factor because field is too weak *)
~= ((IZR p/INR q)*(IZR p/INR q)*(INR q*INR q)) using ring.
~= (sqrt (INR 2%nat) * sqrt (INR 2%nat)*(INR q*INR q)) by H0.
~= (INR (2%nat * (q*q))) by sqrt2,mult_INR.
- then (Zabs_nat p * Zabs_nat p = 2* (q * q))%nat.
+ then (Z.abs_nat p * Z.abs_nat p = 2* (q * q))%nat.
~= ((q*q)+(q*q))%nat.
~= (Div2.double (q*q)).
then (q=0%nat) by main_theorem.
diff --git a/test-suite/success/fix.v b/test-suite/success/fix.v
index b25b502cf..23e3caf8a 100644
--- a/test-suite/success/fix.v
+++ b/test-suite/success/fix.v
@@ -9,12 +9,12 @@ Inductive rBoolOp : Set :=
| rAnd : rBoolOp
| rEq : rBoolOp.
-Definition rlt (a b : rNat) : Prop := Pcompare a b Eq = Lt.
+Definition rlt (a b : rNat) : Prop := Pos.compare_cont a b Eq = Lt.
Definition rltDec : forall m n : rNat, {rlt m n} + {rlt n m \/ m = n}.
intros n m; generalize (nat_of_P_lt_Lt_compare_morphism n m);
generalize (nat_of_P_gt_Gt_compare_morphism n m);
- generalize (Pcompare_Eq_eq n m); case (Pcompare n m Eq).
+ generalize (Pcompare_Eq_eq n m); case (Pos.compare_cont n m Eq).
intros H' H'0 H'1; right; right; auto.
intros H' H'0 H'1; left; unfold rlt in |- *.
apply nat_of_P_lt_Lt_compare_complement_morphism; auto.
diff --git a/test-suite/success/searchabout.v b/test-suite/success/searchabout.v
index d9ade3144..9edfd8255 100644
--- a/test-suite/success/searchabout.v
+++ b/test-suite/success/searchabout.v
@@ -55,6 +55,6 @@ SearchAbout [-"*"%nat "+"%nat] outside Logic.
Require Import ZArith.
-SearchAbout Zmult Zplus "distr".
+SearchAbout Z.mul Z.add "distr".
SearchAbout "+"%Z "*"%Z "distr" -positive -Prop.
SearchAbout (?x * _ + ?x * _)%Z outside OmegaLemmas.
diff --git a/test-suite/success/specialize.v b/test-suite/success/specialize.v
index 578373217..c5f032beb 100644
--- a/test-suite/success/specialize.v
+++ b/test-suite/success/specialize.v
@@ -5,20 +5,20 @@ intros.
(* "compatibility" mode: specializing a global name
means a kind of generalize *)
-specialize trans_equal. intros _.
-specialize trans_equal with (1:=H)(2:=H0). intros _.
-specialize trans_equal with (x:=a)(y:=b)(z:=c). intros _.
-specialize trans_equal with (1:=H)(z:=c). intros _.
-specialize trans_equal with nat a b c. intros _.
-specialize (@trans_equal nat). intros _.
-specialize (@trans_equal _ a b c). intros _.
-specialize (trans_equal (x:=a)). intros _.
-specialize (trans_equal (x:=a)(y:=b)). intros _.
-specialize (trans_equal H H0). intros _.
-specialize (trans_equal H0 (z:=b)). intros _.
+specialize eq_trans. intros _.
+specialize eq_trans with (1:=H)(2:=H0). intros _.
+specialize eq_trans with (x:=a)(y:=b)(z:=c). intros _.
+specialize eq_trans with (1:=H)(z:=c). intros _.
+specialize eq_trans with nat a b c. intros _.
+specialize (@eq_trans nat). intros _.
+specialize (@eq_trans _ a b c). intros _.
+specialize (eq_trans (x:=a)). intros _.
+specialize (eq_trans (x:=a)(y:=b)). intros _.
+specialize (eq_trans H H0). intros _.
+specialize (eq_trans H0 (z:=b)). intros _.
(* local "in place" specialization *)
-assert (Eq:=trans_equal).
+assert (Eq:=eq_trans).
specialize Eq.
specialize Eq with (1:=H)(2:=H0). Undo.
@@ -38,10 +38,10 @@ specialize (Eq _ _ _ b H0). Undo.
presque ok... *)
(* 2) echoue moins lorsque zero premise de mangé *)
-specialize trans_equal with (1:=Eq). (* mal typé !! *)
+specialize eq_trans with (1:=Eq). (* mal typé !! *)
(* 3) *)
-specialize trans_equal with _ a b c. intros _.
+specialize eq_trans with _ a b c. intros _.
(* Anomaly: Evar ?88 was not declared. Please report. *)
*)
diff --git a/test-suite/success/unicode_utf8.v b/test-suite/success/unicode_utf8.v
index 8b7764e5c..42e32cccc 100644
--- a/test-suite/success/unicode_utf8.v
+++ b/test-suite/success/unicode_utf8.v
@@ -75,7 +75,7 @@ Notation "x ≤ y" := (x<=y) (at level 70, no associativity).
Require Import ZArith.
Open Scope Z_scope.
-Locate "≤". (* still le, not Zle *)
+Locate "≤". (* still le, not Z.le *)
Notation "x ≤ y" := (x<=y) (at level 70, no associativity).
Locate "≤".
Close Scope Z_scope.
diff --git a/theories/Arith/Compare_dec.v b/theories/Arith/Compare_dec.v
index 360d760a5..cc2d8efb5 100644
--- a/theories/Arith/Compare_dec.v
+++ b/theories/Arith/Compare_dec.v
@@ -138,7 +138,7 @@ Proof.
Qed.
-(** A ternary comparison function in the spirit of [Zcompare]. *)
+(** A ternary comparison function in the spirit of [Z.compare]. *)
Fixpoint nat_compare n m :=
match n, m with
diff --git a/theories/Arith/Factorial.v b/theories/Arith/Factorial.v
index 146546dc1..e20626029 100644
--- a/theories/Arith/Factorial.v
+++ b/theories/Arith/Factorial.v
@@ -29,7 +29,7 @@ Qed.
Lemma fact_neq_0 : forall n:nat, fact n <> 0.
Proof.
intro.
- apply sym_not_eq.
+ apply not_eq_sym.
apply lt_O_neq.
apply lt_O_fact.
Qed.
diff --git a/theories/Bool/Bool.v b/theories/Bool/Bool.v
index d5d11ceaa..92a584256 100644
--- a/theories/Bool/Bool.v
+++ b/theories/Bool/Bool.v
@@ -615,12 +615,12 @@ Proof.
Qed.
Hint Resolve absurd_eq_true.
-(* A specific instance of trans_eq that preserves compatibility with
+(* A specific instance of eq_trans that preserves compatibility with
old hint bool_2 *)
Lemma trans_eq_bool : forall x y z:bool, x = y -> y = z -> x = z.
Proof.
- apply trans_eq.
+ apply eq_trans.
Qed.
Hint Resolve trans_eq_bool.
diff --git a/theories/FSets/FSetBridge.v b/theories/FSets/FSetBridge.v
index 25ce5577e..f0a588379 100644
--- a/theories/FSets/FSetBridge.v
+++ b/theories/FSets/FSetBridge.v
@@ -165,7 +165,7 @@ Module DepOfNodep (Import M: S) <: Sdep with Module E := M.E.
case (for_all (fdec Pdec) s); unfold For_all in |- *; [ left | right ];
intros.
assert (compat_bool E.eq (fdec Pdec)); auto.
- generalize (H0 H3 (refl_equal _) _ H2).
+ generalize (H0 H3 Logic.eq_refl _ H2).
unfold fdec in |- *.
case (Pdec x); intuition.
inversion H4.
@@ -230,7 +230,7 @@ Module DepOfNodep (Import M: S) <: Sdep with Module E := M.E.
unfold fdec in |- *; case (Pdec x); intuition.
change ((fun x => negb (fdec Pdec x)) x = true) in |- *.
apply (filter_2 (s:=s) (x:=x)); auto.
- set (b := fdec Pdec x) in *; generalize (refl_equal b);
+ set (b := fdec Pdec x) in *; generalize (Logic.eq_refl b);
pattern b at -1 in |- *; case b; unfold b in |- *;
[ left | right ].
elim (H4 x); intros _ B; apply B; auto with set.
diff --git a/theories/FSets/FSetEqProperties.v b/theories/FSets/FSetEqProperties.v
index 755bc7dd0..ac495c040 100644
--- a/theories/FSets/FSetEqProperties.v
+++ b/theories/FSets/FSetEqProperties.v
@@ -206,7 +206,7 @@ intros.
generalize (@choose_1 s) (@choose_2 s).
destruct (choose s);intros.
exists e;auto with set.
-generalize (H1 (refl_equal None)); clear H1.
+generalize (H1 Logic.eq_refl); clear H1.
intros; rewrite (is_empty_1 H1) in H; discriminate.
Qed.
@@ -631,7 +631,7 @@ destruct (choose (filter f s)).
intros H0 _; apply exists_1; auto.
exists e; generalize (H0 e); rewrite filter_iff; auto.
intros _ H0.
-rewrite (is_empty_1 (H0 (refl_equal None))) in H; auto; discriminate.
+rewrite (is_empty_1 (H0 Logic.eq_refl)) in H; auto; discriminate.
Qed.
Lemma partition_filter_1:
@@ -881,8 +881,8 @@ generalize (@add_filter_1 f Hf s0 (add x s0) x) (@add_filter_2 f Hf s0 (add x s0
assert (~ In x (filter f s0)).
intro H1; rewrite (mem_1 (filter_1 Hf H1)) in H; discriminate H.
case (f x); simpl; intros.
-rewrite (MP.cardinal_2 H1 (H2 (refl_equal true) (MP.Add_add s0 x))); auto.
-rewrite <- (MP.Equal_cardinal (H3 (refl_equal false) (MP.Add_add s0 x))); auto.
+rewrite (MP.cardinal_2 H1 (H2 Logic.eq_refl (MP.Add_add s0 x))); auto.
+rewrite <- (MP.Equal_cardinal (H3 Logic.eq_refl (MP.Add_add s0 x))); auto.
intros; rewrite fold_empty;auto.
rewrite MP.cardinal_1; auto.
unfold Empty; intros.
diff --git a/theories/FSets/FSetFacts.v b/theories/FSets/FSetFacts.v
index f473b3342..b240ede4d 100644
--- a/theories/FSets/FSetFacts.v
+++ b/theories/FSets/FSetFacts.v
@@ -315,11 +315,11 @@ symmetry.
rewrite <- H1; intros a Ha.
rewrite <- (H a) in Ha.
destruct H0 as (_,H0).
-exact (H0 (refl_equal true) _ Ha).
+exact (H0 Logic.eq_refl _ Ha).
rewrite <- H0; intros a Ha.
rewrite (H a) in Ha.
destruct H1 as (_,H1).
-exact (H1 (refl_equal true) _ Ha).
+exact (H1 Logic.eq_refl _ Ha).
Qed.
Instance Empty_m : Proper (Equal ==> iff) Empty.
@@ -489,5 +489,3 @@ End WFacts_fun.
Module WFacts (M:WS) := WFacts_fun M.E M.
Module Facts := WFacts.
-
-
diff --git a/theories/FSets/FSetProperties.v b/theories/FSets/FSetProperties.v
index 8651a96f7..25b042ca9 100644
--- a/theories/FSets/FSetProperties.v
+++ b/theories/FSets/FSetProperties.v
@@ -823,7 +823,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
rewrite (inter_subset_equal H).
generalize (@cardinal_inv_1 (diff s' s)).
destruct (cardinal (diff s' s)).
- intro H2; destruct (H2 (refl_equal _) x).
+ intro H2; destruct (H2 Logic.eq_refl x).
set_iff; auto.
intros _.
change (0 + cardinal s < S n + cardinal s).
diff --git a/theories/Init/Datatypes.v b/theories/Init/Datatypes.v
index 387ead7e0..5828fc3f8 100644
--- a/theories/Init/Datatypes.v
+++ b/theories/Init/Datatypes.v
@@ -344,14 +344,14 @@ Definition id : ID := fun A x => x.
(* Compatibility *)
-Notation prodT := prod (only parsing).
-Notation pairT := pair (only parsing).
-Notation prodT_rect := prod_rect (only parsing).
-Notation prodT_rec := prod_rec (only parsing).
-Notation prodT_ind := prod_ind (only parsing).
-Notation fstT := fst (only parsing).
-Notation sndT := snd (only parsing).
-Notation prodT_uncurry := prod_uncurry (only parsing).
-Notation prodT_curry := prod_curry (only parsing).
+Notation prodT := prod (compat "8.2").
+Notation pairT := pair (compat "8.2").
+Notation prodT_rect := prod_rect (compat "8.2").
+Notation prodT_rec := prod_rec (compat "8.2").
+Notation prodT_ind := prod_ind (compat "8.2").
+Notation fstT := fst (compat "8.2").
+Notation sndT := snd (compat "8.2").
+Notation prodT_uncurry := prod_uncurry (compat "8.2").
+Notation prodT_curry := prod_curry (compat "8.2").
(* end hide *)
diff --git a/theories/Init/Logic_Type.v b/theories/Init/Logic_Type.v
index 2a8335762..c244ac24f 100644
--- a/theories/Init/Logic_Type.v
+++ b/theories/Init/Logic_Type.v
@@ -66,7 +66,7 @@ Defined.
Hint Immediate identity_sym not_identity_sym: core v62.
-Notation refl_id := identity_refl (only parsing).
-Notation sym_id := identity_sym (only parsing).
-Notation trans_id := identity_trans (only parsing).
-Notation sym_not_id := not_identity_sym (only parsing).
+Notation refl_id := identity_refl (compat "8.3").
+Notation sym_id := identity_sym (compat "8.3").
+Notation trans_id := identity_trans (compat "8.3").
+Notation sym_not_id := not_identity_sym (compat "8.3").
diff --git a/theories/Init/Notations.v b/theories/Init/Notations.v
index 1171fcca2..3183282d1 100644
--- a/theories/Init/Notations.v
+++ b/theories/Init/Notations.v
@@ -10,7 +10,7 @@
(** Notations for propositional connectives *)
-Reserved Notation "x -> y" (at level 90, right associativity, y at level 200).
+Reserved Notation "x -> y" (at level 99, right associativity, y at level 200).
Reserved Notation "x <-> y" (at level 95, no associativity).
Reserved Notation "x /\ y" (at level 80, right associativity).
Reserved Notation "x \/ y" (at level 85, right associativity).
diff --git a/theories/Init/Peano.v b/theories/Init/Peano.v
index c3716eaa6..cbb960ceb 100644
--- a/theories/Init/Peano.v
+++ b/theories/Init/Peano.v
@@ -115,8 +115,8 @@ Qed.
(** Standard associated names *)
-Notation plus_0_r_reverse := plus_n_O (only parsing).
-Notation plus_succ_r_reverse := plus_n_Sm (only parsing).
+Notation plus_0_r_reverse := plus_n_O (compat "8.2").
+Notation plus_succ_r_reverse := plus_n_Sm (compat "8.2").
(** Multiplication *)
@@ -132,22 +132,22 @@ Hint Resolve (f_equal2 mult): core.
Lemma mult_n_O : forall n:nat, 0 = n * 0.
Proof.
- induction n; simpl in |- *; auto.
+ induction n; simpl; auto.
Qed.
Hint Resolve mult_n_O: core.
Lemma mult_n_Sm : forall n m:nat, n * m + n = n * S m.
Proof.
- intros; induction n as [| p H]; simpl in |- *; auto.
- destruct H; rewrite <- plus_n_Sm; apply (f_equal S).
- pattern m at 1 3 in |- *; elim m; simpl in |- *; auto.
+ intros; induction n as [| p H]; simpl; auto.
+ destruct H; rewrite <- plus_n_Sm; apply eq_S.
+ pattern m at 1 3; elim m; simpl; auto.
Qed.
Hint Resolve mult_n_Sm: core.
(** Standard associated names *)
-Notation mult_0_r_reverse := mult_n_O (only parsing).
-Notation mult_succ_r_reverse := mult_n_Sm (only parsing).
+Notation mult_0_r_reverse := mult_n_O (compat "8.2").
+Notation mult_succ_r_reverse := mult_n_Sm (compat "8.2").
(** Truncated subtraction: [m-n] is [0] if [n>=m] *)
@@ -219,7 +219,7 @@ Theorem nat_double_ind :
(forall n m:nat, R n m -> R (S n) (S m)) -> forall n m:nat, R n m.
Proof.
induction n; auto.
- destruct m as [| n0]; auto.
+ destruct m; auto.
Qed.
(** Maximum and minimum : definitions and specifications *)
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index 5b7afc991..d669df52d 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -226,16 +226,16 @@ Hint Resolve exist exist2 existT existT2: core.
(* Compatibility *)
-Notation sigS := sigT (only parsing).
-Notation existS := existT (only parsing).
-Notation sigS_rect := sigT_rect (only parsing).
-Notation sigS_rec := sigT_rec (only parsing).
-Notation sigS_ind := sigT_ind (only parsing).
-Notation projS1 := projT1 (only parsing).
-Notation projS2 := projT2 (only parsing).
-
-Notation sigS2 := sigT2 (only parsing).
-Notation existS2 := existT2 (only parsing).
-Notation sigS2_rect := sigT2_rect (only parsing).
-Notation sigS2_rec := sigT2_rec (only parsing).
-Notation sigS2_ind := sigT2_ind (only parsing).
+Notation sigS := sigT (compat "8.2").
+Notation existS := existT (compat "8.2").
+Notation sigS_rect := sigT_rect (compat "8.2").
+Notation sigS_rec := sigT_rec (compat "8.2").
+Notation sigS_ind := sigT_ind (compat "8.2").
+Notation projS1 := projT1 (compat "8.2").
+Notation projS2 := projT2 (compat "8.2").
+
+Notation sigS2 := sigT2 (compat "8.2").
+Notation existS2 := existT2 (compat "8.2").
+Notation sigS2_rect := sigT2_rect (compat "8.2").
+Notation sigS2_rec := sigT2_rec (compat "8.2").
+Notation sigS2_ind := sigT2_ind (compat "8.2").
diff --git a/theories/Lists/List.v b/theories/Lists/List.v
index 7296d07ec..5e161ffab 100644
--- a/theories/Lists/List.v
+++ b/theories/Lists/List.v
@@ -619,30 +619,21 @@ Section Elts.
end.
(** Compatibility of count_occ with operations on list *)
- Theorem count_occ_In : forall (l : list A) (x : A), In x l <-> count_occ l x > 0.
+ Theorem count_occ_In (l : list A) (x : A) : In x l <-> count_occ l x > 0.
Proof.
- induction l as [|y l].
- simpl; intros; split; [destruct 1 | apply gt_irrefl].
- simpl. intro x; destruct (eq_dec y x) as [Heq|Hneq].
- rewrite Heq; intuition.
- pose (IHl x). intuition.
+ induction l as [|y l]; simpl.
+ - split; [destruct 1 | apply gt_irrefl].
+ - destruct eq_dec as [->|Hneq]; rewrite IHl; intuition.
Qed.
- Theorem count_occ_inv_nil : forall (l : list A), (forall x:A, count_occ l x = 0) <-> l = [].
+ Theorem count_occ_inv_nil (l : list A) :
+ (forall x:A, count_occ l x = 0) <-> l = [].
Proof.
split.
- (* Case -> *)
- induction l as [|x l].
- trivial.
- intro H.
- elim (O_S (count_occ l x)).
- apply sym_eq.
- generalize (H x).
- simpl. destruct (eq_dec x x) as [|HF].
- trivial.
- elim HF; reflexivity.
- (* Case <- *)
- intro H; rewrite H; simpl; reflexivity.
+ - induction l as [|x l]; trivial.
+ intros H. specialize (H x). simpl in H.
+ destruct eq_dec as [_|NEQ]; [discriminate|now elim NEQ].
+ - now intros ->.
Qed.
Lemma count_occ_nil : forall (x : A), count_occ [] x = 0.
diff --git a/theories/Lists/ListTactics.v b/theories/Lists/ListTactics.v
index 3343aa6f3..447f420ea 100644
--- a/theories/Lists/ListTactics.v
+++ b/theories/Lists/ListTactics.v
@@ -60,7 +60,7 @@ Ltac Find_at a l :=
match l with
| nil => fail 100 "anomaly: Find_at"
| a :: _ => eval compute in n
- | _ :: ?l => find (Psucc n) l
+ | _ :: ?l => find (Pos.succ n) l
end
in find 1%positive l.
diff --git a/theories/Lists/StreamMemo.v b/theories/Lists/StreamMemo.v
index 45490c623..b6f2d5ed2 100644
--- a/theories/Lists/StreamMemo.v
+++ b/theories/Lists/StreamMemo.v
@@ -32,10 +32,10 @@ Fixpoint memo_get (n:nat) (l:Stream A) : A :=
Theorem memo_get_correct: forall n, memo_get n memo_list = f n.
Proof.
assert (F1: forall n m, memo_get n (memo_make m) = f (n + m)).
- induction n as [| n Hrec]; try (intros m; refine (refl_equal _)).
+{ induction n as [| n Hrec]; try (intros m; reflexivity).
intros m; simpl; rewrite Hrec.
- rewrite plus_n_Sm; auto.
-intros n; apply trans_equal with (f (n + 0)); try exact (F1 n 0).
+ rewrite plus_n_Sm; auto. }
+intros n; transitivity (f (n + 0)); try exact (F1 n 0).
rewrite <- plus_n_O; auto.
Qed.
@@ -57,11 +57,10 @@ Definition imemo_list := let f0 := f 0 in
Theorem imemo_get_correct: forall n, memo_get n imemo_list = f n.
Proof.
-assert (F1: forall n m,
- memo_get n (imemo_make (f m)) = f (S (n + m))).
- induction n as [| n Hrec]; try (intros m; exact (sym_equal (Hg_correct m))).
- simpl; intros m; rewrite <- Hg_correct; rewrite Hrec; rewrite <- plus_n_Sm; auto.
-destruct n as [| n]; try apply refl_equal.
+assert (F1: forall n m, memo_get n (imemo_make (f m)) = f (S (n + m))).
+{ induction n as [| n Hrec]; try (intros m; exact (eq_sym (Hg_correct m))).
+ simpl; intros m; rewrite <- Hg_correct, Hrec, <- plus_n_Sm; auto. }
+destruct n as [| n]; try reflexivity.
unfold imemo_list; simpl; rewrite F1.
rewrite <- plus_n_O; auto.
Qed.
@@ -82,7 +81,7 @@ Inductive memo_val: Type :=
Fixpoint is_eq (n m : nat) : {n = m} + {True} :=
match n, m return {n = m} + {True} with
- | 0, 0 =>left True (refl_equal 0)
+ | 0, 0 =>left True (eq_refl 0)
| 0, S m1 => right (0 = S m1) I
| S n1, 0 => right (S n1 = 0) I
| S n1, S m1 =>
@@ -98,7 +97,7 @@ match v with
match is_eq n m with
| left H =>
match H in (eq _ y) return (A y -> A n) with
- | refl_equal => fun v1 : A n => v1
+ | eq_refl => fun v1 : A n => v1
end
| right _ => fun _ : A m => f n
end x
@@ -115,7 +114,7 @@ Proof.
intros n; unfold dmemo_get, dmemo_list.
rewrite (memo_get_correct memo_val mf n); simpl.
case (is_eq n n); simpl; auto; intros e.
-assert (e = refl_equal n).
+assert (e = eq_refl n).
apply eq_proofs_unicity.
induction x as [| x Hx]; destruct y as [| y].
left; auto.
@@ -144,7 +143,7 @@ Proof.
intros n; unfold dmemo_get, dimemo_list.
rewrite (imemo_get_correct memo_val mf mg); simpl.
case (is_eq n n); simpl; auto; intros e.
-assert (e = refl_equal n).
+assert (e = eq_refl n).
apply eq_proofs_unicity.
induction x as [| x Hx]; destruct y as [| y].
left; auto.
@@ -169,11 +168,11 @@ Open Scope Z_scope.
Fixpoint tfact (n: nat) :=
match n with
| O => 1
- | S n1 => Z_of_nat n * tfact n1
+ | S n1 => Z.of_nat n * tfact n1
end.
Definition lfact_list :=
- dimemo_list _ tfact (fun n z => (Z_of_nat (S n) * z)).
+ dimemo_list _ tfact (fun n z => (Z.of_nat (S n) * z)).
Definition lfact n := dmemo_get _ tfact n lfact_list.
diff --git a/theories/Logic/ClassicalFacts.v b/theories/Logic/ClassicalFacts.v
index bcec657aa..7d5cec74d 100644
--- a/theories/Logic/ClassicalFacts.v
+++ b/theories/Logic/ClassicalFacts.v
@@ -191,7 +191,7 @@ Section Proof_irrelevance_gen.
intros Ext Ind.
case (ext_prop_fixpoint Ext bool true); intros G Gfix.
set (neg := fun b:bool => bool_elim bool false true b).
- generalize (refl_equal (G neg)).
+ generalize (eq_refl (G neg)).
pattern (G neg) at 1 in |- *.
apply Ind with (b := G neg); intro Heq.
rewrite (bool_elim_redl bool false true).
@@ -228,9 +228,9 @@ Section Proof_irrelevance_Prop_Ext_CC.
Definition FalseP : BoolP := fun C c1 c2 => c2.
Definition BoolP_elim C c1 c2 (b:BoolP) := b C c1 c2.
Definition BoolP_elim_redl (C:Prop) (c1 c2:C) :
- c1 = BoolP_elim C c1 c2 TrueP := refl_equal c1.
+ c1 = BoolP_elim C c1 c2 TrueP := eq_refl c1.
Definition BoolP_elim_redr (C:Prop) (c1 c2:C) :
- c2 = BoolP_elim C c1 c2 FalseP := refl_equal c2.
+ c2 = BoolP_elim C c1 c2 FalseP := eq_refl c2.
Definition BoolP_dep_induction :=
forall P:BoolP -> Prop, P TrueP -> P FalseP -> forall b:BoolP, P b.
@@ -263,9 +263,9 @@ Section Proof_irrelevance_CIC.
| trueP : boolP
| falseP : boolP.
Definition boolP_elim_redl (C:Prop) (c1 c2:C) :
- c1 = boolP_ind C c1 c2 trueP := refl_equal c1.
+ c1 = boolP_ind C c1 c2 trueP := eq_refl c1.
Definition boolP_elim_redr (C:Prop) (c1 c2:C) :
- c2 = boolP_ind C c1 c2 falseP := refl_equal c2.
+ c2 = boolP_ind C c1 c2 falseP := eq_refl c2.
Scheme boolP_indd := Induction for boolP Sort Prop.
Lemma ext_prop_dep_proof_irrel_cic : prop_extensionality -> proof_irrelevance.
@@ -392,9 +392,9 @@ Section Proof_irrelevance_CCI.
Hypothesis em : forall A:Prop, A \/ ~ A.
Definition or_elim_redl (A B C:Prop) (f:A -> C) (g:B -> C)
- (a:A) : f a = or_ind f g (or_introl B a) := refl_equal (f a).
+ (a:A) : f a = or_ind f g (or_introl B a) := eq_refl (f a).
Definition or_elim_redr (A B C:Prop) (f:A -> C) (g:B -> C)
- (b:B) : g b = or_ind f g (or_intror A b) := refl_equal (g b).
+ (b:B) : g b = or_ind f g (or_intror A b) := eq_refl (g b).
Scheme or_indd := Induction for or Sort Prop.
Theorem proof_irrelevance_cci : forall (B:Prop) (b1 b2:B), b1 = b2.
diff --git a/theories/Logic/Classical_Prop.v b/theories/Logic/Classical_Prop.v
index d2b35da29..5d7764e7e 100644
--- a/theories/Logic/Classical_Prop.v
+++ b/theories/Logic/Classical_Prop.v
@@ -112,7 +112,7 @@ Module Eq_rect_eq.
Lemma eq_rect_eq :
forall (U:Type) (p:U) (Q:U -> Type) (x:Q p) (h:p = p), x = eq_rect p Q x p h.
Proof.
-intros; rewrite proof_irrelevance with (p1:=h) (p2:=refl_equal p); reflexivity.
+intros; rewrite proof_irrelevance with (p1:=h) (p2:=eq_refl p); reflexivity.
Qed.
End Eq_rect_eq.
diff --git a/theories/Logic/Diaconescu.v b/theories/Logic/Diaconescu.v
index 8569e55ec..d1327af3e 100644
--- a/theories/Logic/Diaconescu.v
+++ b/theories/Logic/Diaconescu.v
@@ -188,8 +188,8 @@ Lemma projT1_injective : a1=a2 -> a1'=a2'.
Proof.
intro Heq ; unfold a1', a2', A'.
rewrite Heq.
- replace (or_introl (a2=a2) (refl_equal a2))
- with (or_intror (a2=a2) (refl_equal a2)).
+ replace (or_introl (a2=a2) (eq_refl a2))
+ with (or_intror (a2=a2) (eq_refl a2)).
reflexivity.
apply proof_irrelevance.
Qed.
diff --git a/theories/Logic/EqdepFacts.v b/theories/Logic/EqdepFacts.v
index d84cd8240..4d0560bd7 100644
--- a/theories/Logic/EqdepFacts.v
+++ b/theories/Logic/EqdepFacts.v
@@ -101,7 +101,7 @@ Section Dependent_Equality.
forall (p q:U) (x:P p) (y:P q), eq_dep p x q y -> eq_dep1 p x q y.
Proof.
destruct 1.
- apply eq_dep1_intro with (refl_equal p).
+ apply eq_dep1_intro with (eq_refl p).
simpl; trivial.
Qed.
@@ -121,7 +121,7 @@ Proof.
apply eq_dep_intro.
Qed.
-Notation eq_sigS_eq_dep := eq_sigT_eq_dep (only parsing). (* Compatibility *)
+Notation eq_sigS_eq_dep := eq_sigT_eq_dep (compat "8.2"). (* Compatibility *)
Lemma eq_dep_eq_sigT :
forall (U:Type) (P:U -> Type) (p q:U) (x:P p) (y:P q),
@@ -250,12 +250,12 @@ Section Equivalences.
(** Uniqueness of Reflexive Identity Proofs *)
Definition UIP_refl_ :=
- forall (x:U) (p:x = x), p = refl_equal x.
+ forall (x:U) (p:x = x), p = eq_refl x.
(** Streicher's axiom K *)
Definition Streicher_K_ :=
- forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p.
+ forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p.
(** Injectivity of Dependent Equality is a consequence of *)
(** Invariance by Substitution of Reflexive Equality Proof *)
@@ -389,14 +389,14 @@ Proof (eq_dep_eq__UIP U eq_dep_eq).
(** Uniqueness of Reflexive Identity Proofs is a direct instance of UIP *)
-Lemma UIP_refl : forall (x:U) (p:x = x), p = refl_equal x.
+Lemma UIP_refl : forall (x:U) (p:x = x), p = eq_refl x.
Proof (UIP__UIP_refl U UIP).
(** Streicher's axiom K is a direct consequence of Uniqueness of
Reflexive Identity Proofs *)
Lemma Streicher_K :
- forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p.
+ forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p.
Proof (UIP_refl__Streicher_K U UIP_refl).
End Axioms.
diff --git a/theories/Logic/Eqdep_dec.v b/theories/Logic/Eqdep_dec.v
index 59088aa75..2ed5d428c 100644
--- a/theories/Logic/Eqdep_dec.v
+++ b/theories/Logic/Eqdep_dec.v
@@ -9,7 +9,7 @@
(* Created by Bruno Barras, Jan 1998 *)
(* Made a module instance for EqdepFacts by Hugo Herbelin, Mar 2006 *)
-(** We prove that there is only one proof of [x=x], i.e [refl_equal x].
+(** We prove that there is only one proof of [x=x], i.e [eq_refl x].
This holds if the equality upon the set of [x] is decidable.
A corollary of this theorem is the equality of the right projections
of two equal dependent pairs.
@@ -43,7 +43,7 @@ Section EqdepDec.
Let comp (x y y':A) (eq1:x = y) (eq2:x = y') : y = y' :=
eq_ind _ (fun a => a = y') eq2 _ eq1.
- Remark trans_sym_eq : forall (x y:A) (u:x = y), comp u u = refl_equal y.
+ Remark trans_sym_eq : forall (x y:A) (u:x = y), comp u u = eq_refl y.
Proof.
intros.
case u; trivial.
@@ -69,7 +69,7 @@ Section EqdepDec.
Qed.
- Let nu_inv (y:A) (v:x = y) : x = y := comp (nu (refl_equal x)) v.
+ Let nu_inv (y:A) (v:x = y) : x = y := comp (nu (eq_refl x)) v.
Remark nu_left_inv : forall (y:A) (u:x = y), nu_inv (nu u) = u.
@@ -90,10 +90,10 @@ Section EqdepDec.
Qed.
Theorem K_dec :
- forall P:x = x -> Prop, P (refl_equal x) -> forall p:x = x, P p.
+ forall P:x = x -> Prop, P (eq_refl x) -> forall p:x = x, P p.
Proof.
intros.
- elim eq_proofs_unicity with x (refl_equal x) p.
+ elim eq_proofs_unicity with x (eq_refl x) p.
trivial.
Qed.
@@ -135,7 +135,7 @@ Require Import EqdepFacts.
Theorem K_dec_type :
forall A:Type,
(forall x y:A, {x = y} + {x <> y}) ->
- forall (x:A) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p.
+ forall (x:A) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p.
Proof.
intros A eq_dec x P H p.
elim p using K_dec; intros.
@@ -146,7 +146,7 @@ Qed.
Theorem K_dec_set :
forall A:Set,
(forall x y:A, {x = y} + {x <> y}) ->
- forall (x:A) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p.
+ forall (x:A) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p.
Proof fun A => K_dec_type (A:=A).
(** We deduce the [eq_rect_eq] axiom for (decidable) types *)
@@ -212,13 +212,13 @@ Module DecidableEqDep (M:DecidableType).
(** Uniqueness of Reflexive Identity Proofs *)
- Lemma UIP_refl : forall (x:U) (p:x = x), p = refl_equal x.
+ Lemma UIP_refl : forall (x:U) (p:x = x), p = eq_refl x.
Proof (UIP__UIP_refl U UIP).
(** Streicher's axiom K *)
Lemma Streicher_K :
- forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p.
+ forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p.
Proof (K_dec_type eq_dec).
(** Injectivity of equality on dependent pairs in [Type] *)
@@ -281,13 +281,13 @@ Module DecidableEqDepSet (M:DecidableSet).
(** Uniqueness of Reflexive Identity Proofs *)
- Lemma UIP_refl : forall (x:U) (p:x = x), p = refl_equal x.
+ Lemma UIP_refl : forall (x:U) (p:x = x), p = eq_refl x.
Proof N.UIP_refl.
(** Streicher's axiom K *)
Lemma Streicher_K :
- forall (x:U) (P:x = x -> Prop), P (refl_equal x) -> forall p:x = x, P p.
+ forall (x:U) (P:x = x -> Prop), P (eq_refl x) -> forall p:x = x, P p.
Proof N.Streicher_K.
(** Proof-irrelevance on subsets of decidable sets *)
@@ -301,7 +301,7 @@ Module DecidableEqDepSet (M:DecidableSet).
Lemma inj_pair2 :
forall (P:U -> Type) (p:U) (x y:P p),
- existS P p x = existS P p y -> x = y.
+ existT P p x = existT P p y -> x = y.
Proof eq_dep_eq__inj_pair2 U N.eq_dep_eq.
(** Injectivity of equality on dependent pairs with second component
diff --git a/theories/Logic/ProofIrrelevanceFacts.v b/theories/Logic/ProofIrrelevanceFacts.v
index 6accc4807..776010c03 100644
--- a/theories/Logic/ProofIrrelevanceFacts.v
+++ b/theories/Logic/ProofIrrelevanceFacts.v
@@ -25,7 +25,7 @@ Module ProofIrrelevanceTheory (M:ProofIrrelevance).
forall (U:Type) (p:U) (Q:U -> Type) (x:Q p) (h:p = p),
x = eq_rect p Q x p h.
Proof.
- intros; rewrite M.proof_irrelevance with (p1:=h) (p2:=refl_equal p).
+ intros; rewrite M.proof_irrelevance with (p1:=h) (p2:=eq_refl p).
reflexivity.
Qed.
End Eq_rect_eq.
diff --git a/theories/MSets/MSetEqProperties.v b/theories/MSets/MSetEqProperties.v
index 2e7da404e..4f0d93fb9 100644
--- a/theories/MSets/MSetEqProperties.v
+++ b/theories/MSets/MSetEqProperties.v
@@ -206,7 +206,7 @@ intros.
generalize (@choose_1 s) (@choose_2 s).
destruct (choose s);intros.
exists e;auto with set.
-generalize (H1 (refl_equal None)); clear H1.
+generalize (H1 (eq_refl None)); clear H1.
intros; rewrite (is_empty_1 H1) in H; discriminate.
Qed.
@@ -631,7 +631,7 @@ destruct (choose (filter f s)).
intros H0 _; apply exists_1; auto.
exists e; generalize (H0 e); rewrite filter_iff; auto.
intros _ H0.
-rewrite (is_empty_1 (H0 (refl_equal None))) in H; auto; discriminate.
+rewrite (is_empty_1 (H0 (eq_refl None))) in H; auto; discriminate.
Qed.
Lemma partition_filter_1:
@@ -879,8 +879,8 @@ generalize (@add_filter_1 f Hf s0 (add x s0) x) (@add_filter_2 f Hf s0 (add x s0
assert (~ In x (filter f s0)).
intro H1; rewrite (mem_1 (filter_1 Hf H1)) in H; discriminate H.
case (f x); simpl; intros.
-rewrite (MP.cardinal_2 H1 (H2 (refl_equal true) (MP.Add_add s0 x))); auto.
-rewrite <- (MP.Equal_cardinal (H3 (refl_equal false) (MP.Add_add s0 x))); auto.
+rewrite (MP.cardinal_2 H1 (H2 (eq_refl true) (MP.Add_add s0 x))); auto.
+rewrite <- (MP.Equal_cardinal (H3 (eq_refl false) (MP.Add_add s0 x))); auto.
intros; rewrite fold_empty;auto.
rewrite MP.cardinal_1; auto.
unfold Empty; intros.
diff --git a/theories/MSets/MSetPositive.v b/theories/MSets/MSetPositive.v
index e83ac27d6..e500602fd 100644
--- a/theories/MSets/MSetPositive.v
+++ b/theories/MSets/MSetPositive.v
@@ -36,8 +36,8 @@ Local Unset Boolean Equality Schemes.
Module PositiveOrderedTypeBits <: UsualOrderedType.
Definition t:=positive.
Include HasUsualEq <+ UsualIsEq.
- Definition eqb := Peqb.
- Definition eqb_eq := Peqb_eq.
+ Definition eqb := Pos.eqb.
+ Definition eqb_eq := Pos.eqb_eq.
Include HasEqBool2Dec.
Fixpoint bits_lt (p q:positive) : Prop :=
diff --git a/theories/MSets/MSetProperties.v b/theories/MSets/MSetProperties.v
index 0f24d76a9..396067b57 100644
--- a/theories/MSets/MSetProperties.v
+++ b/theories/MSets/MSetProperties.v
@@ -831,7 +831,7 @@ Module WPropertiesOn (Import E : DecidableType)(M : WSetsOn E).
rewrite (inter_subset_equal H).
generalize (@cardinal_inv_1 (diff s' s)).
destruct (cardinal (diff s' s)).
- intro H2; destruct (H2 (refl_equal _) x).
+ intro H2; destruct (H2 (eq_refl _) x).
set_iff; auto.
intros _.
change (0 + cardinal s < S n + cardinal s).
diff --git a/theories/MSets/MSetRBT.v b/theories/MSets/MSetRBT.v
index b53c03920..35d8245a7 100644
--- a/theories/MSets/MSetRBT.v
+++ b/theories/MSets/MSetRBT.v
@@ -282,7 +282,7 @@ Fixpoint treeify_aux (pred:bool)(n: positive) : treeify_t :=
Fixpoint plength (l:list elt) := match l with
| nil => 1%positive
- | _::l => Psucc (plength l)
+ | _::l => Pos.succ (plength l)
end.
Definition treeify (l:list elt) :=
diff --git a/theories/NArith/BinNat.v b/theories/NArith/BinNat.v
index 6128b9740..046670f7b 100644
--- a/theories/NArith/BinNat.v
+++ b/theories/NArith/BinNat.v
@@ -76,7 +76,7 @@ Defined.
(** Discrimination principle *)
-Definition discr n : { p:positive | n = Npos p } + { n = N0 }.
+Definition discr n : { p:positive | n = pos p } + { n = 0 }.
Proof.
destruct n; auto.
left; exists p; auto.
@@ -87,12 +87,12 @@ Defined.
Definition binary_rect (P:N -> Type) (f0 : P 0)
(f2 : forall n, P n -> P (double n))
(fS2 : forall n, P n -> P (succ_double n)) (n : N) : P n :=
- let P' p := P (Npos p) in
- let f2' p := f2 (Npos p) in
- let fS2' p := fS2 (Npos p) in
+ let P' p := P (pos p) in
+ let f2' p := f2 (pos p) in
+ let fS2' p := fS2 (pos p) in
match n with
| 0 => f0
- | Npos p => positive_rect P' fS2' f2' (fS2 0 f0) p
+ | pos p => positive_rect P' fS2' f2' (fS2 0 f0) p
end.
Definition binary_rec (P:N -> Set) := binary_rect P.
@@ -103,11 +103,11 @@ Definition binary_ind (P:N -> Prop) := binary_rect P.
Definition peano_rect
(P : N -> Type) (f0 : P 0)
(f : forall n : N, P n -> P (succ n)) (n : N) : P n :=
-let P' p := P (Npos p) in
-let f' p := f (Npos p) in
+let P' p := P (pos p) in
+let f' p := f (pos p) in
match n with
| 0 => f0
-| Npos p => Pos.peano_rect P' (f 0 f0) f' p
+| pos p => Pos.peano_rect P' (f 0 f0) f' p
end.
Theorem peano_rect_base P a f : peano_rect P a f 0 = a.
@@ -140,12 +140,12 @@ Qed.
(** Properties of mixed successor and predecessor. *)
-Lemma pos_pred_spec p : Pos.pred_N p = pred (Npos p).
+Lemma pos_pred_spec p : Pos.pred_N p = pred (pos p).
Proof.
now destruct p.
Qed.
-Lemma succ_pos_spec n : Npos (succ_pos n) = succ n.
+Lemma succ_pos_spec n : pos (succ_pos n) = succ n.
Proof.
now destruct n.
Qed.
@@ -155,7 +155,7 @@ Proof.
destruct n. trivial. apply Pos.pred_N_succ.
Qed.
-Lemma succ_pos_pred p : succ (Pos.pred_N p) = Npos p.
+Lemma succ_pos_pred p : succ (Pos.pred_N p) = pos p.
Proof.
destruct p; simpl; trivial. f_equal. apply Pos.succ_pred_double.
Qed.
@@ -472,7 +472,7 @@ Lemma log2_spec n : 0 < n ->
2^(log2 n) <= n < 2^(succ (log2 n)).
Proof.
destruct n as [|[p|p|]]; discriminate || intros _; simpl; split.
- apply (size_le (Npos p)).
+ apply (size_le (pos p)).
apply Pos.size_gt.
apply Pos.size_le.
apply Pos.size_gt.
@@ -494,7 +494,7 @@ Proof.
trivial.
destruct p; simpl; split; try easy.
intros (m,H). now destruct m.
- now exists (Npos p).
+ now exists (pos p).
intros (m,H). now destruct m.
Qed.
@@ -504,7 +504,7 @@ Proof.
split. discriminate.
intros (m,H). now destruct m.
destruct p; simpl; split; try easy.
- now exists (Npos p).
+ now exists (pos p).
intros (m,H). now destruct m.
now exists 0.
Qed.
@@ -512,19 +512,19 @@ Qed.
(** Specification of the euclidean division *)
Theorem pos_div_eucl_spec (a:positive)(b:N) :
- let (q,r) := pos_div_eucl a b in Npos a = q * b + r.
+ let (q,r) := pos_div_eucl a b in pos a = q * b + r.
Proof.
induction a; cbv beta iota delta [pos_div_eucl]; fold pos_div_eucl; cbv zeta.
(* a~1 *)
destruct pos_div_eucl as (q,r).
- change (Npos a~1) with (succ_double (Npos a)).
+ change (pos a~1) with (succ_double (pos a)).
rewrite IHa, succ_double_add, double_mul.
case leb_spec; intros H; trivial.
rewrite succ_double_mul, <- add_assoc. f_equal.
now rewrite (add_comm b), sub_add.
(* a~0 *)
destruct pos_div_eucl as (q,r).
- change (Npos a~0) with (double (Npos a)).
+ change (pos a~0) with (double (pos a)).
rewrite IHa, double_add, double_mul.
case leb_spec; intros H; trivial.
rewrite succ_double_mul, <- add_assoc. f_equal.
@@ -537,7 +537,7 @@ Theorem div_eucl_spec a b :
let (q,r) := div_eucl a b in a = b * q + r.
Proof.
destruct a as [|a], b as [|b]; unfold div_eucl; trivial.
- generalize (pos_div_eucl_spec a (Npos b)).
+ generalize (pos_div_eucl_spec a (pos b)).
destruct pos_div_eucl. now rewrite mul_comm.
Qed.
@@ -664,7 +664,7 @@ Proof.
destruct (Pos.gcd_greatest p q r) as (u,H).
exists s. now inversion Hs.
exists t. now inversion Ht.
- exists (Npos u). simpl; now f_equal.
+ exists (pos u). simpl; now f_equal.
Qed.
Lemma gcd_nonneg a b : 0 <= gcd a b.
@@ -862,7 +862,7 @@ Program Definition testbit_wd : Proper (eq==>eq==>Logic.eq) testbit := _.
Theorem bi_induction :
forall A : N -> Prop, Proper (Logic.eq==>iff) A ->
- A N0 -> (forall n, A n <-> A (succ n)) -> forall n : N, A n.
+ A 0 -> (forall n, A n <-> A (succ n)) -> forall n : N, A n.
Proof.
intros A A_wd A0 AS. apply peano_rect. assumption. intros; now apply -> AS.
Qed.
@@ -1018,7 +1018,7 @@ Notation N_rect := N_rect (only parsing).
Notation N_rec := N_rec (only parsing).
Notation N_ind := N_ind (only parsing).
Notation N0 := N0 (only parsing).
-Notation Npos := Npos (only parsing).
+Notation Npos := N.pos (only parsing).
Notation Ndiscr := N.discr (compat "8.3").
Notation Ndouble_plus_one := N.succ_double (compat "8.3").
diff --git a/theories/NArith/BinNatDef.v b/theories/NArith/BinNatDef.v
index b95c9b4bc..76526a5ce 100644
--- a/theories/NArith/BinNatDef.v
+++ b/theories/NArith/BinNatDef.v
@@ -19,6 +19,10 @@ Module N.
Definition t := N.
+(** ** Nicer name [N.pos] for contructor [Npos] *)
+
+Notation pos := Npos.
+
(** ** Constants *)
Definition zero := 0.
@@ -30,7 +34,7 @@ Definition two := 2.
Definition succ_double x :=
match x with
| 0 => 1
- | Npos p => Npos p~1
+ | pos p => pos p~1
end.
(** ** Operation [x -> 2*x] *)
@@ -38,7 +42,7 @@ Definition succ_double x :=
Definition double n :=
match n with
| 0 => 0
- | Npos p => Npos p~0
+ | pos p => pos p~0
end.
(** ** Successor *)
@@ -46,7 +50,7 @@ Definition double n :=
Definition succ n :=
match n with
| 0 => 1
- | Npos p => Npos (Pos.succ p)
+ | pos p => pos (Pos.succ p)
end.
(** ** Predecessor *)
@@ -54,15 +58,15 @@ Definition succ n :=
Definition pred n :=
match n with
| 0 => 0
- | Npos p => Pos.pred_N p
+ | pos p => Pos.pred_N p
end.
(** ** The successor of a [N] can be seen as a [positive] *)
Definition succ_pos (n : N) : positive :=
match n with
- | N0 => 1%positive
- | Npos p => Pos.succ p
+ | 0 => 1%positive
+ | pos p => Pos.succ p
end.
(** ** Addition *)
@@ -71,7 +75,7 @@ Definition add n m :=
match n, m with
| 0, _ => m
| _, 0 => n
- | Npos p, Npos q => Npos (p + q)
+ | pos p, pos q => pos (p + q)
end.
Infix "+" := add : N_scope.
@@ -82,9 +86,9 @@ Definition sub n m :=
match n, m with
| 0, _ => 0
| n, 0 => n
-| Npos n', Npos m' =>
+| pos n', pos m' =>
match Pos.sub_mask n' m' with
- | IsPos p => Npos p
+ | IsPos p => pos p
| _ => 0
end
end.
@@ -97,7 +101,7 @@ Definition mul n m :=
match n, m with
| 0, _ => 0
| _, 0 => 0
- | Npos p, Npos q => Npos (p * q)
+ | pos p, pos q => pos (p * q)
end.
Infix "*" := mul : N_scope.
@@ -107,9 +111,9 @@ Infix "*" := mul : N_scope.
Definition compare n m :=
match n, m with
| 0, 0 => Eq
- | 0, Npos m' => Lt
- | Npos n', 0 => Gt
- | Npos n', Npos m' => (n' ?= m')%positive
+ | 0, pos m' => Lt
+ | pos n', 0 => Gt
+ | pos n', pos m' => (n' ?= m')%positive
end.
Infix "?=" := compare (at level 70, no associativity) : N_scope.
@@ -119,7 +123,7 @@ Infix "?=" := compare (at level 70, no associativity) : N_scope.
Fixpoint eqb n m :=
match n, m with
| 0, 0 => true
- | Npos p, Npos q => Pos.eqb p q
+ | pos p, pos q => Pos.eqb p q
| _, _ => false
end.
@@ -151,8 +155,8 @@ Definition div2 n :=
match n with
| 0 => 0
| 1 => 0
- | Npos (p~0) => Npos p
- | Npos (p~1) => Npos p
+ | pos (p~0) => pos p
+ | pos (p~1) => pos p
end.
(** Parity *)
@@ -160,7 +164,7 @@ Definition div2 n :=
Definition even n :=
match n with
| 0 => true
- | Npos (xO _) => true
+ | pos (xO _) => true
| _ => false
end.
@@ -172,7 +176,7 @@ Definition pow n p :=
match p, n with
| 0, _ => 1
| _, 0 => 0
- | Npos p, Npos q => Npos (q^p)
+ | pos p, pos q => pos (q^p)
end.
Infix "^" := pow : N_scope.
@@ -182,7 +186,7 @@ Infix "^" := pow : N_scope.
Definition square n :=
match n with
| 0 => 0
- | Npos p => Npos (Pos.square p)
+ | pos p => pos (Pos.square p)
end.
(** Base-2 logarithm *)
@@ -191,8 +195,8 @@ Definition log2 n :=
match n with
| 0 => 0
| 1 => 0
- | Npos (p~0) => Npos (Pos.size p)
- | Npos (p~1) => Npos (Pos.size p)
+ | pos (p~0) => pos (Pos.size p)
+ | pos (p~1) => pos (Pos.size p)
end.
(** How many digits in a number ?
@@ -202,13 +206,13 @@ Definition log2 n :=
Definition size n :=
match n with
| 0 => 0
- | Npos p => Npos (Pos.size p)
+ | pos p => pos (Pos.size p)
end.
Definition size_nat n :=
match n with
| 0 => O
- | Npos p => Pos.size_nat p
+ | pos p => Pos.size_nat p
end.
(** Euclidean division *)
@@ -233,7 +237,7 @@ Definition div_eucl (a b:N) : N * N :=
match a, b with
| 0, _ => (0, 0)
| _, 0 => (0, a)
- | Npos na, _ => pos_div_eucl na b
+ | pos na, _ => pos_div_eucl na b
end.
Definition div a b := fst (div_eucl a b).
@@ -248,7 +252,7 @@ Definition gcd a b :=
match a, b with
| 0, _ => b
| _, 0 => a
- | Npos p, Npos q => Npos (Pos.gcd p q)
+ | pos p, pos q => pos (Pos.gcd p q)
end.
(** Generalized Gcd, also computing rests of [a] and [b] after
@@ -258,9 +262,9 @@ Definition ggcd a b :=
match a, b with
| 0, _ => (b,(0,1))
| _, 0 => (a,(1,0))
- | Npos p, Npos q =>
+ | pos p, pos q =>
let '(g,(aa,bb)) := Pos.ggcd p q in
- (Npos g, (Npos aa, Npos bb))
+ (pos g, (pos aa, pos bb))
end.
(** Square root *)
@@ -268,17 +272,17 @@ Definition ggcd a b :=
Definition sqrtrem n :=
match n with
| 0 => (0, 0)
- | Npos p =>
+ | pos p =>
match Pos.sqrtrem p with
- | (s, IsPos r) => (Npos s, Npos r)
- | (s, _) => (Npos s, 0)
+ | (s, IsPos r) => (pos s, pos r)
+ | (s, _) => (pos s, 0)
end
end.
Definition sqrt n :=
match n with
| 0 => 0
- | Npos p => Npos (Pos.sqrt p)
+ | pos p => pos (Pos.sqrt p)
end.
(** Operation over bits of a [N] number. *)
@@ -289,7 +293,7 @@ Definition lor n m :=
match n, m with
| 0, _ => m
| _, 0 => n
- | Npos p, Npos q => Npos (Pos.lor p q)
+ | pos p, pos q => pos (Pos.lor p q)
end.
(** Logical [and] *)
@@ -298,7 +302,7 @@ Definition land n m :=
match n, m with
| 0, _ => 0
| _, 0 => 0
- | Npos p, Npos q => Pos.land p q
+ | pos p, pos q => Pos.land p q
end.
(** Logical [diff] *)
@@ -307,7 +311,7 @@ Fixpoint ldiff n m :=
match n, m with
| 0, _ => 0
| _, 0 => n
- | Npos p, Npos q => Pos.ldiff p q
+ | pos p, pos q => Pos.ldiff p q
end.
(** [xor] *)
@@ -316,7 +320,7 @@ Definition lxor n m :=
match n, m with
| 0, _ => m
| _, 0 => n
- | Npos p, Npos q => Pos.lxor p q
+ | pos p, pos q => Pos.lxor p q
end.
(** Shifts *)
@@ -327,13 +331,13 @@ Definition shiftr_nat (a:N)(n:nat) := nat_iter n div2 a.
Definition shiftl a n :=
match a with
| 0 => 0
- | Npos a => Npos (Pos.shiftl a n)
+ | pos a => pos (Pos.shiftl a n)
end.
Definition shiftr a n :=
match n with
| 0 => a
- | Npos p => Pos.iter p div2 a
+ | pos p => Pos.iter p div2 a
end.
(** Checking whether a particular bit is set or not *)
@@ -341,7 +345,7 @@ Definition shiftr a n :=
Definition testbit_nat (a:N) :=
match a with
| 0 => fun _ => false
- | Npos p => Pos.testbit_nat p
+ | pos p => Pos.testbit_nat p
end.
(** Same, but with index in N *)
@@ -349,7 +353,7 @@ Definition testbit_nat (a:N) :=
Definition testbit a n :=
match a with
| 0 => false
- | Npos p => Pos.testbit p n
+ | pos p => Pos.testbit p n
end.
(** Translation from [N] to [nat] and back. *)
@@ -357,13 +361,13 @@ Definition testbit a n :=
Definition to_nat (a:N) :=
match a with
| 0 => O
- | Npos p => Pos.to_nat p
+ | pos p => Pos.to_nat p
end.
Definition of_nat (n:nat) :=
match n with
| O => 0
- | S n' => Npos (Pos.of_succ_nat n')
+ | S n' => pos (Pos.of_succ_nat n')
end.
(** Iteration of a function *)
@@ -371,7 +375,7 @@ Definition of_nat (n:nat) :=
Definition iter (n:N) {A} (f:A->A) (x:A) : A :=
match n with
| 0 => x
- | Npos p => Pos.iter p f x
+ | pos p => Pos.iter p f x
end.
End N. \ No newline at end of file
diff --git a/theories/NArith/Ndec.v b/theories/NArith/Ndec.v
index f2ee29cc0..ee0b2ebc4 100644
--- a/theories/NArith/Ndec.v
+++ b/theories/NArith/Ndec.v
@@ -15,315 +15,238 @@ Require Import Pnat.
Require Import Nnat.
Require Import Ndigits.
-(** A boolean equality over [N] *)
+Local Open Scope N_scope.
-Notation Peqb := Peqb (only parsing). (* Now in [BinPos] *)
-Notation Neqb := Neqb (only parsing). (* Now in [BinNat] *)
+(** Obsolete results about boolean comparisons over [N],
+ kept for compatibility with IntMap and SMC. *)
-Notation Peqb_correct := Peqb_refl (only parsing).
+Notation Peqb := Pos.eqb (compat "8.3").
+Notation Neqb := N.eqb (compat "8.3").
+Notation Peqb_correct := Pos.eqb_refl (compat "8.3").
+Notation Neqb_correct := N.eqb_refl (compat "8.3").
+Notation Neqb_comm := N.eqb_sym (compat "8.3").
-Lemma Peqb_complete : forall p p', Peqb p p' = true -> p = p'.
-Proof.
- intros. now apply (Peqb_eq p p').
-Qed.
+Lemma Peqb_complete p p' : Pos.eqb p p' = true -> p = p'.
+Proof. now apply Pos.eqb_eq. Qed.
-Lemma Peqb_Pcompare : forall p p', Peqb p p' = true -> Pos.compare p p' = Eq.
-Proof.
- intros. now rewrite Pos.compare_eq_iff, <- Peqb_eq.
-Qed.
-
-Lemma Pcompare_Peqb : forall p p', Pos.compare p p' = Eq -> Peqb p p' = true.
-Proof.
- intros; now rewrite Peqb_eq, <- Pos.compare_eq_iff.
-Qed.
+Lemma Peqb_Pcompare p p' : Pos.eqb p p' = true -> Pos.compare p p' = Eq.
+Proof. now rewrite Pos.compare_eq_iff, <- Pos.eqb_eq. Qed.
-Lemma Neqb_correct : forall n, Neqb n n = true.
-Proof.
- intros; now rewrite Neqb_eq.
-Qed.
+Lemma Pcompare_Peqb p p' : Pos.compare p p' = Eq -> Pos.eqb p p' = true.
+Proof. now rewrite Pos.eqb_eq, <- Pos.compare_eq_iff. Qed.
-Lemma Neqb_Ncompare : forall n n', Neqb n n' = true -> Ncompare n n' = Eq.
-Proof.
- intros; now rewrite Ncompare_eq_correct, <- Neqb_eq.
-Qed.
+Lemma Neqb_Ncompare n n' : N.eqb n n' = true -> N.compare n n' = Eq.
+Proof. now rewrite N.compare_eq_iff, <- N.eqb_eq. Qed.
-Lemma Ncompare_Neqb : forall n n', Ncompare n n' = Eq -> Neqb n n' = true.
-Proof.
- intros; now rewrite Neqb_eq, <- Ncompare_eq_correct.
-Qed.
+Lemma Ncompare_Neqb n n' : N.compare n n' = Eq -> N.eqb n n' = true.
+Proof. now rewrite N.eqb_eq, <- N.compare_eq_iff. Qed.
-Lemma Neqb_complete : forall a a', Neqb a a' = true -> a = a'.
-Proof.
- intros; now rewrite <- Neqb_eq.
-Qed.
+Lemma Neqb_complete n n' : N.eqb n n' = true -> n = n'.
+Proof. now apply N.eqb_eq. Qed.
-Lemma Neqb_comm : forall a a', Neqb a a' = Neqb a' a.
+Lemma Nxor_eq_true n n' : N.lxor n n' = 0 -> N.eqb n n' = true.
Proof.
- intros; apply eq_true_iff_eq. rewrite 2 Neqb_eq; auto with *.
+ intro H. apply N.lxor_eq in H. subst. apply N.eqb_refl.
Qed.
-Lemma Nxor_eq_true :
- forall a a', Nxor a a' = N0 -> Neqb a a' = true.
-Proof.
- intros. rewrite (Nxor_eq a a' H). apply Neqb_correct.
-Qed.
+Ltac eqb2eq := rewrite <- ?not_true_iff_false in *; rewrite ?N.eqb_eq in *.
-Lemma Nxor_eq_false :
- forall a a' p, Nxor a a' = Npos p -> Neqb a a' = false.
+Lemma Nxor_eq_false n n' p :
+ N.lxor n n' = N.pos p -> N.eqb n n' = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb a a')). intro H0.
- rewrite (Neqb_complete a a' H0) in H.
- rewrite (Nxor_nilpotent a') in H. discriminate H.
- trivial.
+ intros. eqb2eq. intro. subst. now rewrite N.lxor_nilpotent in *.
Qed.
-Lemma Nodd_not_double :
- forall a,
- Nodd a -> forall a0, Neqb (Ndouble a0) a = false.
+Lemma Nodd_not_double a :
+ Nodd a -> forall a0, N.eqb (N.double a0) a = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb (Ndouble a0) a)). intro H0.
- rewrite <- (Neqb_complete _ _ H0) in H.
- unfold Nodd in H.
- rewrite (Ndouble_bit0 a0) in H. discriminate H.
- trivial.
+ intros. eqb2eq. intros <-.
+ unfold Nodd in *. now rewrite Ndouble_bit0 in *.
Qed.
-Lemma Nnot_div2_not_double :
- forall a a0,
- Neqb (Ndiv2 a) a0 = false -> Neqb a (Ndouble a0) = false.
+Lemma Nnot_div2_not_double a a0 :
+ N.eqb (N.div2 a) a0 = false -> N.eqb a (N.double a0) = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb (Ndouble a0) a)). intro H0.
- rewrite <- (Neqb_complete _ _ H0) in H. rewrite (Ndouble_div2 a0) in H.
- rewrite (Neqb_correct a0) in H. discriminate H.
- intro. rewrite Neqb_comm. assumption.
+ intros H. eqb2eq. contradict H. subst. apply N.div2_double.
Qed.
-Lemma Neven_not_double_plus_one :
- forall a,
- Neven a -> forall a0, Neqb (Ndouble_plus_one a0) a = false.
+Lemma Neven_not_double_plus_one a :
+ Neven a -> forall a0, N.eqb (N.succ_double a0) a = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb (Ndouble_plus_one a0) a)). intro H0.
- rewrite <- (Neqb_complete _ _ H0) in H.
- unfold Neven in H.
- rewrite (Ndouble_plus_one_bit0 a0) in H.
- discriminate H.
- trivial.
+ intros. eqb2eq. intros <-.
+ unfold Neven in *. now rewrite Ndouble_plus_one_bit0 in *.
Qed.
-Lemma Nnot_div2_not_double_plus_one :
- forall a a0,
- Neqb (Ndiv2 a) a0 = false -> Neqb (Ndouble_plus_one a0) a = false.
+Lemma Nnot_div2_not_double_plus_one a a0 :
+ N.eqb (N.div2 a) a0 = false -> N.eqb (N.succ_double a0) a = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb a (Ndouble_plus_one a0))). intro H0.
- rewrite (Neqb_complete _ _ H0) in H. rewrite (Ndouble_plus_one_div2 a0) in H.
- rewrite (Neqb_correct a0) in H. discriminate H.
- intro H0. rewrite Neqb_comm. assumption.
+ intros H. eqb2eq. contradict H. subst. apply N.div2_succ_double.
Qed.
-Lemma Nbit0_neq :
- forall a a',
- Nbit0 a = false -> Nbit0 a' = true -> Neqb a a' = false.
+Lemma Nbit0_neq a a' :
+ N.odd a = false -> N.odd a' = true -> N.eqb a a' = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb a a')). intro H1.
- rewrite (Neqb_complete _ _ H1) in H.
- rewrite H in H0. discriminate H0.
- trivial.
+ intros. eqb2eq. now intros <-.
Qed.
-Lemma Ndiv2_eq :
- forall a a', Neqb a a' = true -> Neqb (Ndiv2 a) (Ndiv2 a') = true.
+Lemma Ndiv2_eq a a' :
+ N.eqb a a' = true -> N.eqb (N.div2 a) (N.div2 a') = true.
Proof.
- intros. cut (a = a'). intros. rewrite H0. apply Neqb_correct.
- apply Neqb_complete. exact H.
+ intros. eqb2eq. now subst.
Qed.
-Lemma Ndiv2_neq :
- forall a a',
- Neqb (Ndiv2 a) (Ndiv2 a') = false -> Neqb a a' = false.
+Lemma Ndiv2_neq a a' :
+ N.eqb (N.div2 a) (N.div2 a') = false -> N.eqb a a' = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb a a')). intro H0.
- rewrite (Neqb_complete _ _ H0) in H.
- rewrite (Neqb_correct (Ndiv2 a')) in H. discriminate H.
- trivial.
+ intros H. eqb2eq. contradict H. now subst.
Qed.
-Lemma Ndiv2_bit_eq :
- forall a a',
- Nbit0 a = Nbit0 a' -> Ndiv2 a = Ndiv2 a' -> a = a'.
+Lemma Ndiv2_bit_eq a a' :
+ N.odd a = N.odd a' -> N.div2 a = N.div2 a' -> a = a'.
Proof.
- intros. apply Nbit_faithful. unfold eqf in |- *. destruct n.
- rewrite Nbit0_correct. rewrite Nbit0_correct. assumption.
- rewrite <- Ndiv2_correct. rewrite <- Ndiv2_correct.
- rewrite H0. reflexivity.
+ intros H H'; now rewrite (N.div2_odd a), (N.div2_odd a'), H, H'.
Qed.
-Lemma Ndiv2_bit_neq :
- forall a a',
- Neqb a a' = false ->
- Nbit0 a = Nbit0 a' -> Neqb (Ndiv2 a) (Ndiv2 a') = false.
+Lemma Ndiv2_bit_neq a a' :
+ N.eqb a a' = false ->
+ N.odd a = N.odd a' -> N.eqb (N.div2 a) (N.div2 a') = false.
Proof.
- intros. elim (sumbool_of_bool (Neqb (Ndiv2 a) (Ndiv2 a'))). intro H1.
- rewrite (Ndiv2_bit_eq _ _ H0 (Neqb_complete _ _ H1)) in H.
- rewrite (Neqb_correct a') in H. discriminate H.
- trivial.
+ intros H H'. eqb2eq. contradict H. now apply Ndiv2_bit_eq.
Qed.
-Lemma Nneq_elim :
- forall a a',
- Neqb a a' = false ->
- Nbit0 a = negb (Nbit0 a') \/
- Neqb (Ndiv2 a) (Ndiv2 a') = false.
+Lemma Nneq_elim a a' :
+ N.eqb a a' = false ->
+ N.odd a = negb (N.odd a') \/
+ N.eqb (N.div2 a) (N.div2 a') = false.
Proof.
- intros. cut (Nbit0 a = Nbit0 a' \/ Nbit0 a = negb (Nbit0 a')).
+ intros. cut (N.odd a = N.odd a' \/ N.odd a = negb (N.odd a')).
intros. elim H0. intro. right. apply Ndiv2_bit_neq. assumption.
assumption.
intro. left. assumption.
- case (Nbit0 a); case (Nbit0 a'); auto.
+ case (N.odd a), (N.odd a'); auto.
Qed.
-Lemma Ndouble_or_double_plus_un :
- forall a,
- {a0 : N | a = Ndouble a0} + {a1 : N | a = Ndouble_plus_one a1}.
+Lemma Ndouble_or_double_plus_un a :
+ {a0 : N | a = N.double a0} + {a1 : N | a = N.succ_double a1}.
Proof.
- intro. elim (sumbool_of_bool (Nbit0 a)). intro H. right. split with (Ndiv2 a).
- rewrite (Ndiv2_double_plus_one a H). reflexivity.
- intro H. left. split with (Ndiv2 a). rewrite (Ndiv2_double a H). reflexivity.
+ elim (sumbool_of_bool (N.odd a)); intros H; [right|left];
+ exists (N.div2 a); symmetry;
+ apply Ndiv2_double_plus_one || apply Ndiv2_double; auto.
Qed.
-(** A boolean order on [N] *)
+(** An inefficient boolean order on [N]. Please use [N.leb] instead now. *)
-Definition Nleb (a b:N) := leb (nat_of_N a) (nat_of_N b).
+Definition Nleb (a b:N) := leb (N.to_nat a) (N.to_nat b).
-Lemma Nleb_Nle : forall a b, Nleb a b = true <-> Nle a b.
+Lemma Nleb_alt a b : Nleb a b = N.leb a b.
Proof.
- intros; unfold Nle; rewrite nat_of_Ncompare.
- unfold Nleb; apply leb_compare.
+ unfold Nleb.
+ now rewrite eq_iff_eq_true, N.leb_le, leb_compare, <- N2Nat.inj_compare.
Qed.
-Lemma Nleb_refl : forall a, Nleb a a = true.
-Proof.
- intro. unfold Nleb in |- *. apply leb_correct. apply le_n.
-Qed.
+Lemma Nleb_Nle a b : Nleb a b = true <-> a <= b.
+Proof. now rewrite Nleb_alt, N.leb_le. Qed.
-Lemma Nleb_antisym :
- forall a b, Nleb a b = true -> Nleb b a = true -> a = b.
-Proof.
- unfold Nleb in |- *. intros. rewrite <- (N_of_nat_of_N a). rewrite <- (N_of_nat_of_N b).
- rewrite (le_antisym _ _ (leb_complete _ _ H) (leb_complete _ _ H0)). reflexivity.
-Qed.
+Lemma Nleb_refl a : Nleb a a = true.
+Proof. rewrite Nleb_Nle; apply N.le_refl. Qed.
-Lemma Nleb_trans :
- forall a b c, Nleb a b = true -> Nleb b c = true -> Nleb a c = true.
-Proof.
- unfold Nleb in |- *. intros. apply leb_correct. apply le_trans with (m := nat_of_N b).
- apply leb_complete. assumption.
- apply leb_complete. assumption.
-Qed.
+Lemma Nleb_antisym a b : Nleb a b = true -> Nleb b a = true -> a = b.
+Proof. rewrite !Nleb_Nle. apply N.le_antisymm. Qed.
+
+Lemma Nleb_trans a b c : Nleb a b = true -> Nleb b c = true -> Nleb a c = true.
+Proof. rewrite !Nleb_Nle. apply N.le_trans. Qed.
-Lemma Nleb_ltb_trans :
- forall a b c,
- Nleb a b = true -> Nleb c b = false -> Nleb c a = false.
+Lemma Nleb_ltb_trans a b c :
+ Nleb a b = true -> Nleb c b = false -> Nleb c a = false.
Proof.
- unfold Nleb in |- *. intros. apply leb_correct_conv. apply le_lt_trans with (m := nat_of_N b).
+ unfold Nleb. intros. apply leb_correct_conv.
+ apply le_lt_trans with (m := N.to_nat b).
apply leb_complete. assumption.
apply leb_complete_conv. assumption.
Qed.
-Lemma Nltb_leb_trans :
- forall a b c,
- Nleb b a = false -> Nleb b c = true -> Nleb c a = false.
+Lemma Nltb_leb_trans a b c :
+ Nleb b a = false -> Nleb b c = true -> Nleb c a = false.
Proof.
- unfold Nleb in |- *. intros. apply leb_correct_conv. apply lt_le_trans with (m := nat_of_N b).
+ unfold Nleb. intros. apply leb_correct_conv.
+ apply lt_le_trans with (m := N.to_nat b).
apply leb_complete_conv. assumption.
apply leb_complete. assumption.
Qed.
-Lemma Nltb_trans :
- forall a b c,
- Nleb b a = false -> Nleb c b = false -> Nleb c a = false.
+Lemma Nltb_trans a b c :
+ Nleb b a = false -> Nleb c b = false -> Nleb c a = false.
Proof.
- unfold Nleb in |- *. intros. apply leb_correct_conv. apply lt_trans with (m := nat_of_N b).
+ unfold Nleb. intros. apply leb_correct_conv.
+ apply lt_trans with (m := N.to_nat b).
apply leb_complete_conv. assumption.
apply leb_complete_conv. assumption.
Qed.
-Lemma Nltb_leb_weak : forall a b:N, Nleb b a = false -> Nleb a b = true.
+Lemma Nltb_leb_weak a b : Nleb b a = false -> Nleb a b = true.
Proof.
- unfold Nleb in |- *. intros. apply leb_correct. apply lt_le_weak.
+ unfold Nleb. intros. apply leb_correct. apply lt_le_weak.
apply leb_complete_conv. assumption.
Qed.
-Lemma Nleb_double_mono :
- forall a b,
- Nleb a b = true -> Nleb (Ndouble a) (Ndouble b) = true.
+Lemma Nleb_double_mono a b :
+ Nleb a b = true -> Nleb (N.double a) (N.double b) = true.
Proof.
- unfold Nleb in |- *. intros. rewrite nat_of_Ndouble. rewrite nat_of_Ndouble. apply leb_correct.
- simpl in |- *. apply plus_le_compat. apply leb_complete. assumption.
- apply plus_le_compat. apply leb_complete. assumption.
- apply le_n.
+ unfold Nleb. intros. rewrite !N2Nat.inj_double. apply leb_correct.
+ apply mult_le_compat_l. now apply leb_complete.
Qed.
-Lemma Nleb_double_plus_one_mono :
- forall a b,
- Nleb a b = true ->
- Nleb (Ndouble_plus_one a) (Ndouble_plus_one b) = true.
+Lemma Nleb_double_plus_one_mono a b :
+ Nleb a b = true ->
+ Nleb (N.succ_double a) (N.succ_double b) = true.
Proof.
- unfold Nleb in |- *. intros. rewrite nat_of_Ndouble_plus_one. rewrite nat_of_Ndouble_plus_one.
- apply leb_correct. apply le_n_S. simpl in |- *. apply plus_le_compat. apply leb_complete.
- assumption.
- apply plus_le_compat. apply leb_complete. assumption.
- apply le_n.
+ unfold Nleb. intros. rewrite !N2Nat.inj_succ_double. apply leb_correct.
+ apply le_n_S, mult_le_compat_l. now apply leb_complete.
Qed.
-Lemma Nleb_double_mono_conv :
- forall a b,
- Nleb (Ndouble a) (Ndouble b) = true -> Nleb a b = true.
+Lemma Nleb_double_mono_conv a b :
+ Nleb (N.double a) (N.double b) = true -> Nleb a b = true.
Proof.
- unfold Nleb in |- *. intros a b. rewrite nat_of_Ndouble. rewrite nat_of_Ndouble. intro.
- apply leb_correct. apply (mult_S_le_reg_l 1). apply leb_complete. assumption.
+ unfold Nleb. rewrite !N2Nat.inj_double. intro. apply leb_correct.
+ apply (mult_S_le_reg_l 1). now apply leb_complete.
Qed.
-Lemma Nleb_double_plus_one_mono_conv :
- forall a b,
- Nleb (Ndouble_plus_one a) (Ndouble_plus_one b) = true ->
+Lemma Nleb_double_plus_one_mono_conv a b :
+ Nleb (N.succ_double a) (N.succ_double b) = true ->
Nleb a b = true.
Proof.
- unfold Nleb in |- *. intros a b. rewrite nat_of_Ndouble_plus_one. rewrite nat_of_Ndouble_plus_one.
- intro. apply leb_correct. apply (mult_S_le_reg_l 1). apply le_S_n. apply leb_complete.
- assumption.
+ unfold Nleb. rewrite !N2Nat.inj_succ_double. intro. apply leb_correct.
+ apply (mult_S_le_reg_l 1). apply le_S_n. now apply leb_complete.
Qed.
-Lemma Nltb_double_mono :
- forall a b,
- Nleb a b = false -> Nleb (Ndouble a) (Ndouble b) = false.
+Lemma Nltb_double_mono a b :
+ Nleb a b = false -> Nleb (N.double a) (N.double b) = false.
Proof.
- intros. elim (sumbool_of_bool (Nleb (Ndouble a) (Ndouble b))). intro H0.
+ intros. elim (sumbool_of_bool (Nleb (N.double a) (N.double b))). intro H0.
rewrite (Nleb_double_mono_conv _ _ H0) in H. discriminate H.
trivial.
Qed.
-Lemma Nltb_double_plus_one_mono :
- forall a b,
- Nleb a b = false ->
- Nleb (Ndouble_plus_one a) (Ndouble_plus_one b) = false.
+Lemma Nltb_double_plus_one_mono a b :
+ Nleb a b = false ->
+ Nleb (N.succ_double a) (N.succ_double b) = false.
Proof.
- intros. elim (sumbool_of_bool (Nleb (Ndouble_plus_one a) (Ndouble_plus_one b))). intro H0.
+ intros. elim (sumbool_of_bool (Nleb (N.succ_double a) (N.succ_double b))).
+ intro H0.
rewrite (Nleb_double_plus_one_mono_conv _ _ H0) in H. discriminate H.
trivial.
Qed.
-Lemma Nltb_double_mono_conv :
- forall a b,
- Nleb (Ndouble a) (Ndouble b) = false -> Nleb a b = false.
+Lemma Nltb_double_mono_conv a b :
+ Nleb (N.double a) (N.double b) = false -> Nleb a b = false.
Proof.
- intros. elim (sumbool_of_bool (Nleb a b)). intro H0. rewrite (Nleb_double_mono _ _ H0) in H.
- discriminate H.
+ intros. elim (sumbool_of_bool (Nleb a b)). intro H0.
+ rewrite (Nleb_double_mono _ _ H0) in H. discriminate H.
trivial.
Qed.
-Lemma Nltb_double_plus_one_mono_conv :
- forall a b,
- Nleb (Ndouble_plus_one a) (Ndouble_plus_one b) = false ->
+Lemma Nltb_double_plus_one_mono_conv a b :
+ Nleb (N.succ_double a) (N.succ_double b) = false ->
Nleb a b = false.
Proof.
intros. elim (sumbool_of_bool (Nleb a b)). intro H0.
@@ -331,110 +254,52 @@ Proof.
trivial.
Qed.
-(* Nleb and Ncompare *)
+(* Nleb and N.compare *)
-(* NB: No need to prove that Nleb a b = true <-> Ncompare a b <> Gt,
+(* NB: No need to prove that Nleb a b = true <-> N.compare a b <> Gt,
this statement is in fact Nleb_Nle! *)
-Lemma Nltb_Ncompare : forall a b,
- Nleb a b = false <-> Ncompare a b = Gt.
+Lemma Nltb_Ncompare a b : Nleb a b = false <-> N.compare a b = Gt.
Proof.
- intros.
- assert (IFF : forall x:bool, x = false <-> ~ x = true)
- by (destruct x; intuition).
- rewrite IFF, Nleb_Nle; unfold Nle.
- destruct (Ncompare a b); split; intro H; auto;
- elim H; discriminate.
+ now rewrite N.compare_nle_iff, <- Nleb_Nle, not_true_iff_false.
Qed.
-Lemma Ncompare_Gt_Nltb : forall a b,
- Ncompare a b = Gt -> Nleb a b = false.
-Proof.
- intros; apply <- Nltb_Ncompare; auto.
-Qed.
+Lemma Ncompare_Gt_Nltb a b : N.compare a b = Gt -> Nleb a b = false.
+Proof. apply <- Nltb_Ncompare; auto. Qed.
-Lemma Ncompare_Lt_Nltb : forall a b,
- Ncompare a b = Lt -> Nleb b a = false.
+Lemma Ncompare_Lt_Nltb a b : N.compare a b = Lt -> Nleb b a = false.
Proof.
- intros a b H.
- rewrite Nltb_Ncompare, <- Ncompare_antisym, H; auto.
+ intros H. rewrite Nltb_Ncompare, N.compare_antisym, H; auto.
Qed.
-(* An alternate [min] function over [N] *)
+(* Old results about [N.min] *)
-Definition Nmin' (a b:N) := if Nleb a b then a else b.
+Notation Nmin_choice := N.min_dec (compat "8.3").
-Lemma Nmin_Nmin' : forall a b, Nmin a b = Nmin' a b.
-Proof.
- unfold Nmin, Nmin', Nleb; intros.
- rewrite nat_of_Ncompare.
- generalize (leb_compare (nat_of_N a) (nat_of_N b));
- destruct (nat_compare (nat_of_N a) (nat_of_N b));
- destruct (leb (nat_of_N a) (nat_of_N b)); intuition.
- lapply H1; intros; discriminate.
- lapply H1; intros; discriminate.
-Qed.
+Lemma Nmin_le_1 a b : Nleb (N.min a b) a = true.
+Proof. rewrite Nleb_Nle. apply N.le_min_l. Qed.
-Lemma Nmin_choice : forall a b, {Nmin a b = a} + {Nmin a b = b}.
-Proof.
- unfold Nmin in *; intros; destruct (Ncompare a b); auto.
-Qed.
+Lemma Nmin_le_2 a b : Nleb (N.min a b) b = true.
+Proof. rewrite Nleb_Nle. apply N.le_min_r. Qed.
-Lemma Nmin_le_1 : forall a b, Nleb (Nmin a b) a = true.
-Proof.
- intros; rewrite Nmin_Nmin'.
- unfold Nmin'; elim (sumbool_of_bool (Nleb a b)). intro H. rewrite H.
- apply Nleb_refl.
- intro H. rewrite H. apply Nltb_leb_weak. assumption.
-Qed.
+Lemma Nmin_le_3 a b c : Nleb a (N.min b c) = true -> Nleb a b = true.
+Proof. rewrite !Nleb_Nle. apply N.min_glb_l. Qed.
-Lemma Nmin_le_2 : forall a b, Nleb (Nmin a b) b = true.
-Proof.
- intros; rewrite Nmin_Nmin'.
- unfold Nmin'; elim (sumbool_of_bool (Nleb a b)). intro H. rewrite H. assumption.
- intro H. rewrite H. apply Nleb_refl.
-Qed.
+Lemma Nmin_le_4 a b c : Nleb a (N.min b c) = true -> Nleb a c = true.
+Proof. rewrite !Nleb_Nle. apply N.min_glb_r. Qed.
-Lemma Nmin_le_3 :
- forall a b c, Nleb a (Nmin b c) = true -> Nleb a b = true.
-Proof.
- intros; rewrite Nmin_Nmin' in *.
- unfold Nmin' in *; elim (sumbool_of_bool (Nleb b c)). intro H0. rewrite H0 in H.
- assumption.
- intro H0. rewrite H0 in H. apply Nltb_leb_weak. apply Nleb_ltb_trans with (b := c); assumption.
-Qed.
+Lemma Nmin_le_5 a b c :
+ Nleb a b = true -> Nleb a c = true -> Nleb a (N.min b c) = true.
+Proof. rewrite !Nleb_Nle. apply N.min_glb. Qed.
-Lemma Nmin_le_4 :
- forall a b c, Nleb a (Nmin b c) = true -> Nleb a c = true.
+Lemma Nmin_lt_3 a b c : Nleb (N.min b c) a = false -> Nleb b a = false.
Proof.
- intros; rewrite Nmin_Nmin' in *.
- unfold Nmin' in *; elim (sumbool_of_bool (Nleb b c)). intro H0. rewrite H0 in H.
- apply Nleb_trans with (b := b); assumption.
- intro H0. rewrite H0 in H. assumption.
-Qed.
-
-Lemma Nmin_le_5 :
- forall a b c,
- Nleb a b = true -> Nleb a c = true -> Nleb a (Nmin b c) = true.
-Proof.
- intros. elim (Nmin_choice b c). intro H1. rewrite H1. assumption.
- intro H1. rewrite H1. assumption.
-Qed.
-
-Lemma Nmin_lt_3 :
- forall a b c, Nleb (Nmin b c) a = false -> Nleb b a = false.
-Proof.
- intros; rewrite Nmin_Nmin' in *.
- unfold Nmin' in *. intros. elim (sumbool_of_bool (Nleb b c)). intro H0. rewrite H0 in H.
- assumption.
- intro H0. rewrite H0 in H. apply Nltb_trans with (b := c); assumption.
+ rewrite <- !not_true_iff_false, !Nleb_Nle.
+ rewrite N.min_le_iff; auto.
Qed.
-Lemma Nmin_lt_4 :
- forall a b c, Nleb (Nmin b c) a = false -> Nleb c a = false.
+Lemma Nmin_lt_4 a b c : Nleb (N.min b c) a = false -> Nleb c a = false.
Proof.
- intros; rewrite Nmin_Nmin' in *.
- unfold Nmin' in *. elim (sumbool_of_bool (Nleb b c)). intro H0. rewrite H0 in H.
- apply Nltb_leb_trans with (b := b); assumption.
- intro H0. rewrite H0 in H. assumption.
+ rewrite <- !not_true_iff_false, !Nleb_Nle.
+ rewrite N.min_le_iff; auto.
Qed.
diff --git a/theories/NArith/Ndigits.v b/theories/NArith/Ndigits.v
index 06f47d719..77364e5ea 100644
--- a/theories/NArith/Ndigits.v
+++ b/theories/NArith/Ndigits.v
@@ -31,25 +31,25 @@ Notation Nxor_nilpotent := N.lxor_nilpotent (compat "8.3").
either with index in [N] or in [nat]. *)
Lemma Ptestbit_Pbit :
- forall p n, Pos.testbit p (N.of_nat n) = Pbit p n.
+ forall p n, Pos.testbit p (N.of_nat n) = Pos.testbit_nat p n.
Proof.
induction p as [p IH|p IH| ]; intros [|n]; simpl; trivial;
- rewrite <- IH; f_equal; rewrite (pred_Sn n) at 2; now rewrite N_of_pred.
+ rewrite <- IH; f_equal; rewrite (pred_Sn n) at 2; now rewrite Nat2N.inj_pred.
Qed.
-Lemma Ntestbit_Nbit : forall a n, N.testbit a (N.of_nat n) = Nbit a n.
+Lemma Ntestbit_Nbit : forall a n, N.testbit a (N.of_nat n) = N.testbit_nat a n.
Proof.
destruct a. trivial. apply Ptestbit_Pbit.
Qed.
Lemma Pbit_Ptestbit :
- forall p n, Pbit p (N.to_nat n) = Pos.testbit p n.
+ forall p n, Pos.testbit_nat p (N.to_nat n) = Pos.testbit p n.
Proof.
intros; now rewrite <- Ptestbit_Pbit, N2Nat.id.
Qed.
Lemma Nbit_Ntestbit :
- forall a n, Nbit a (N.to_nat n) = N.testbit a n.
+ forall a n, N.testbit_nat a (N.to_nat n) = N.testbit a n.
Proof.
destruct a. trivial. apply Pbit_Ptestbit.
Qed.
@@ -73,7 +73,7 @@ Lemma Nshiftr_nat_equiv :
Proof.
intros a [|n]; simpl. unfold N.shiftr_nat.
trivial.
- symmetry. apply iter_nat_of_P.
+ symmetry. apply Pos2Nat.inj_iter.
Qed.
Lemma Nshiftr_equiv_nat :
@@ -99,7 +99,7 @@ Qed.
(** Correctness proofs for shifts, nat version *)
Lemma Nshiftr_nat_spec : forall a n m,
- Nbit (N.shiftr_nat a n) m = Nbit a (m+n).
+ N.testbit_nat (N.shiftr_nat a n) m = N.testbit_nat a (m+n).
Proof.
induction n; intros m.
now rewrite <- plus_n_O.
@@ -108,7 +108,7 @@ Proof.
Qed.
Lemma Nshiftl_nat_spec_high : forall a n m, (n<=m)%nat ->
- Nbit (N.shiftl_nat a n) m = Nbit a (m-n).
+ N.testbit_nat (N.shiftl_nat a n) m = N.testbit_nat a (m-n).
Proof.
induction n; intros m H.
now rewrite <- minus_n_O.
@@ -118,7 +118,7 @@ Proof.
Qed.
Lemma Nshiftl_nat_spec_low : forall a n m, (m<n)%nat ->
- Nbit (N.shiftl_nat a n) m = false.
+ N.testbit_nat (N.shiftl_nat a n) m = false.
Proof.
induction n; intros m H. inversion H.
rewrite Nshiftl_nat_S.
@@ -151,52 +151,52 @@ Proof.
rewrite 2 Pshiftl_nat_S. now f_equal.
Qed.
-(** Semantics of bitwise operations with respect to [Nbit] *)
+(** Semantics of bitwise operations with respect to [N.testbit_nat] *)
Lemma Pxor_semantics p p' n :
- Nbit (Pos.lxor p p') n = xorb (Pbit p n) (Pbit p' n).
+ N.testbit_nat (Pos.lxor p p') n = xorb (Pos.testbit_nat p n) (Pos.testbit_nat p' n).
Proof.
rewrite <- Ntestbit_Nbit, <- !Ptestbit_Pbit. apply N.pos_lxor_spec.
Qed.
Lemma Nxor_semantics a a' n :
- Nbit (N.lxor a a') n = xorb (Nbit a n) (Nbit a' n).
+ N.testbit_nat (N.lxor a a') n = xorb (N.testbit_nat a n) (N.testbit_nat a' n).
Proof.
rewrite <- !Ntestbit_Nbit. apply N.lxor_spec.
Qed.
Lemma Por_semantics p p' n :
- Pbit (Pos.lor p p') n = (Pbit p n) || (Pbit p' n).
+ Pos.testbit_nat (Pos.lor p p') n = (Pos.testbit_nat p n) || (Pos.testbit_nat p' n).
Proof.
rewrite <- !Ptestbit_Pbit. apply N.pos_lor_spec.
Qed.
Lemma Nor_semantics a a' n :
- Nbit (N.lor a a') n = (Nbit a n) || (Nbit a' n).
+ N.testbit_nat (N.lor a a') n = (N.testbit_nat a n) || (N.testbit_nat a' n).
Proof.
rewrite <- !Ntestbit_Nbit. apply N.lor_spec.
Qed.
Lemma Pand_semantics p p' n :
- Nbit (Pos.land p p') n = (Pbit p n) && (Pbit p' n).
+ N.testbit_nat (Pos.land p p') n = (Pos.testbit_nat p n) && (Pos.testbit_nat p' n).
Proof.
rewrite <- Ntestbit_Nbit, <- !Ptestbit_Pbit. apply N.pos_land_spec.
Qed.
Lemma Nand_semantics a a' n :
- Nbit (N.land a a') n = (Nbit a n) && (Nbit a' n).
+ N.testbit_nat (N.land a a') n = (N.testbit_nat a n) && (N.testbit_nat a' n).
Proof.
rewrite <- !Ntestbit_Nbit. apply N.land_spec.
Qed.
Lemma Pdiff_semantics p p' n :
- Nbit (Pos.ldiff p p') n = (Pbit p n) && negb (Pbit p' n).
+ N.testbit_nat (Pos.ldiff p p') n = (Pos.testbit_nat p n) && negb (Pos.testbit_nat p' n).
Proof.
rewrite <- Ntestbit_Nbit, <- !Ptestbit_Pbit. apply N.pos_ldiff_spec.
Qed.
Lemma Ndiff_semantics a a' n :
- Nbit (N.ldiff a a') n = (Nbit a n) && negb (Nbit a' n).
+ N.testbit_nat (N.ldiff a a') n = (N.testbit_nat a n) && negb (N.testbit_nat a' n).
Proof.
rewrite <- !Ntestbit_Nbit. apply N.ldiff_spec.
Qed.
@@ -213,13 +213,13 @@ Local Infix "==" := eqf (at level 70, no associativity).
Local Notation Step H := (fun n => H (S n)).
-Lemma Pbit_faithful_0 : forall p, ~(Pbit p == (fun _ => false)).
+Lemma Pbit_faithful_0 : forall p, ~(Pos.testbit_nat p == (fun _ => false)).
Proof.
induction p as [p IHp|p IHp| ]; intros H; try discriminate (H O).
apply (IHp (Step H)).
Qed.
-Lemma Pbit_faithful : forall p p', Pbit p == Pbit p' -> p = p'.
+Lemma Pbit_faithful : forall p p', Pos.testbit_nat p == Pos.testbit_nat p' -> p = p'.
Proof.
induction p as [p IHp|p IHp| ]; intros [p'|p'|] H; trivial;
try discriminate (H O).
@@ -229,7 +229,7 @@ Proof.
symmetry in H. destruct (Pbit_faithful_0 _ (Step H)).
Qed.
-Lemma Nbit_faithful : forall n n', Nbit n == Nbit n' -> n = n'.
+Lemma Nbit_faithful : forall n n', N.testbit_nat n == N.testbit_nat n' -> n = n'.
Proof.
intros [|p] [|p'] H; trivial.
symmetry in H. destruct (Pbit_faithful_0 _ H).
@@ -237,7 +237,7 @@ Proof.
f_equal. apply Pbit_faithful, H.
Qed.
-Lemma Nbit_faithful_iff : forall n n', Nbit n == Nbit n' <-> n = n'.
+Lemma Nbit_faithful_iff : forall n n', N.testbit_nat n == N.testbit_nat n' <-> n = n'.
Proof.
split. apply Nbit_faithful. intros; now subst.
Qed.
@@ -249,28 +249,28 @@ Local Close Scope N_scope.
Notation Nbit0 := N.odd (compat "8.3").
-Definition Nodd (n:N) := Nbit0 n = true.
-Definition Neven (n:N) := Nbit0 n = false.
+Definition Nodd (n:N) := N.odd n = true.
+Definition Neven (n:N) := N.odd n = false.
-Lemma Nbit0_correct : forall n:N, Nbit n 0 = Nbit0 n.
+Lemma Nbit0_correct : forall n:N, N.testbit_nat n 0 = N.odd n.
Proof.
destruct n; trivial.
destruct p; trivial.
Qed.
-Lemma Ndouble_bit0 : forall n:N, Nbit0 (Ndouble n) = false.
+Lemma Ndouble_bit0 : forall n:N, N.odd (N.double n) = false.
Proof.
destruct n; trivial.
Qed.
Lemma Ndouble_plus_one_bit0 :
- forall n:N, Nbit0 (Ndouble_plus_one n) = true.
+ forall n:N, N.odd (N.succ_double n) = true.
Proof.
destruct n; trivial.
Qed.
Lemma Ndiv2_double :
- forall n:N, Neven n -> Ndouble (Ndiv2 n) = n.
+ forall n:N, Neven n -> N.double (N.div2 n) = n.
Proof.
destruct n. trivial. destruct p. intro H. discriminate H.
intros. reflexivity.
@@ -278,7 +278,7 @@ Proof.
Qed.
Lemma Ndiv2_double_plus_one :
- forall n:N, Nodd n -> Ndouble_plus_one (Ndiv2 n) = n.
+ forall n:N, Nodd n -> N.succ_double (N.div2 n) = n.
Proof.
destruct n. intro. discriminate H.
destruct p. intros. reflexivity.
@@ -287,31 +287,31 @@ Proof.
Qed.
Lemma Ndiv2_correct :
- forall (a:N) (n:nat), Nbit (Ndiv2 a) n = Nbit a (S n).
+ forall (a:N) (n:nat), N.testbit_nat (N.div2 a) n = N.testbit_nat a (S n).
Proof.
destruct a; trivial.
destruct p; trivial.
Qed.
Lemma Nxor_bit0 :
- forall a a':N, Nbit0 (Nxor a a') = xorb (Nbit0 a) (Nbit0 a').
+ forall a a':N, N.odd (N.lxor a a') = xorb (N.odd a) (N.odd a').
Proof.
intros. rewrite <- Nbit0_correct, (Nxor_semantics a a' O).
rewrite Nbit0_correct, Nbit0_correct. reflexivity.
Qed.
Lemma Nxor_div2 :
- forall a a':N, Ndiv2 (Nxor a a') = Nxor (Ndiv2 a) (Ndiv2 a').
+ forall a a':N, N.div2 (N.lxor a a') = N.lxor (N.div2 a) (N.div2 a').
Proof.
intros. apply Nbit_faithful. unfold eqf. intro.
- rewrite (Nxor_semantics (Ndiv2 a) (Ndiv2 a') n), Ndiv2_correct, (Nxor_semantics a a' (S n)).
+ rewrite (Nxor_semantics (N.div2 a) (N.div2 a') n), Ndiv2_correct, (Nxor_semantics a a' (S n)).
rewrite 2! Ndiv2_correct.
reflexivity.
Qed.
Lemma Nneg_bit0 :
forall a a':N,
- Nbit0 (Nxor a a') = true -> Nbit0 a = negb (Nbit0 a').
+ N.odd (N.lxor a a') = true -> N.odd a = negb (N.odd a').
Proof.
intros.
rewrite <- true_xorb, <- H, Nxor_bit0, xorb_assoc,
@@ -320,24 +320,24 @@ Proof.
Qed.
Lemma Nneg_bit0_1 :
- forall a a':N, Nxor a a' = Npos 1 -> Nbit0 a = negb (Nbit0 a').
+ forall a a':N, N.lxor a a' = Npos 1 -> N.odd a = negb (N.odd a').
Proof.
intros. apply Nneg_bit0. rewrite H. reflexivity.
Qed.
Lemma Nneg_bit0_2 :
forall (a a':N) (p:positive),
- Nxor a a' = Npos (xI p) -> Nbit0 a = negb (Nbit0 a').
+ N.lxor a a' = Npos (xI p) -> N.odd a = negb (N.odd a').
Proof.
intros. apply Nneg_bit0. rewrite H. reflexivity.
Qed.
Lemma Nsame_bit0 :
forall (a a':N) (p:positive),
- Nxor a a' = Npos (xO p) -> Nbit0 a = Nbit0 a'.
+ N.lxor a a' = Npos (xO p) -> N.odd a = N.odd a'.
Proof.
- intros. rewrite <- (xorb_false (Nbit0 a)).
- assert (H0: Nbit0 (Npos (xO p)) = false) by reflexivity.
+ intros. rewrite <- (xorb_false (N.odd a)).
+ assert (H0: N.odd (Npos (xO p)) = false) by reflexivity.
rewrite <- H0, <- H, Nxor_bit0, <- xorb_assoc, xorb_nilpotent, false_xorb.
reflexivity.
Qed.
@@ -346,77 +346,77 @@ Qed.
Fixpoint Nless_aux (a a':N) (p:positive) : bool :=
match p with
- | xO p' => Nless_aux (Ndiv2 a) (Ndiv2 a') p'
- | _ => andb (negb (Nbit0 a)) (Nbit0 a')
+ | xO p' => Nless_aux (N.div2 a) (N.div2 a') p'
+ | _ => andb (negb (N.odd a)) (N.odd a')
end.
Definition Nless (a a':N) :=
- match Nxor a a' with
+ match N.lxor a a' with
| N0 => false
| Npos p => Nless_aux a a' p
end.
Lemma Nbit0_less :
forall a a',
- Nbit0 a = false -> Nbit0 a' = true -> Nless a a' = true.
+ N.odd a = false -> N.odd a' = true -> Nless a a' = true.
Proof.
- intros. destruct (Ndiscr (Nxor a a')) as [(p,H2)|H1]. unfold Nless.
+ intros. destruct (N.discr (N.lxor a a')) as [(p,H2)|H1]. unfold Nless.
rewrite H2. destruct p. simpl. rewrite H, H0. reflexivity.
- assert (H1: Nbit0 (Nxor a a') = false) by (rewrite H2; reflexivity).
+ assert (H1: N.odd (N.lxor a a') = false) by (rewrite H2; reflexivity).
rewrite (Nxor_bit0 a a'), H, H0 in H1. discriminate H1.
simpl. rewrite H, H0. reflexivity.
- assert (H2: Nbit0 (Nxor a a') = false) by (rewrite H1; reflexivity).
+ assert (H2: N.odd (N.lxor a a') = false) by (rewrite H1; reflexivity).
rewrite (Nxor_bit0 a a'), H, H0 in H2. discriminate H2.
Qed.
Lemma Nbit0_gt :
forall a a',
- Nbit0 a = true -> Nbit0 a' = false -> Nless a a' = false.
+ N.odd a = true -> N.odd a' = false -> Nless a a' = false.
Proof.
- intros. destruct (Ndiscr (Nxor a a')) as [(p,H2)|H1]. unfold Nless.
+ intros. destruct (N.discr (N.lxor a a')) as [(p,H2)|H1]. unfold Nless.
rewrite H2. destruct p. simpl. rewrite H, H0. reflexivity.
- assert (H1: Nbit0 (Nxor a a') = false) by (rewrite H2; reflexivity).
+ assert (H1: N.odd (N.lxor a a') = false) by (rewrite H2; reflexivity).
rewrite (Nxor_bit0 a a'), H, H0 in H1. discriminate H1.
simpl. rewrite H, H0. reflexivity.
- assert (H2: Nbit0 (Nxor a a') = false) by (rewrite H1; reflexivity).
+ assert (H2: N.odd (N.lxor a a') = false) by (rewrite H1; reflexivity).
rewrite (Nxor_bit0 a a'), H, H0 in H2. discriminate H2.
Qed.
Lemma Nless_not_refl : forall a, Nless a a = false.
Proof.
- intro. unfold Nless. rewrite (Nxor_nilpotent a). reflexivity.
+ intro. unfold Nless. rewrite (N.lxor_nilpotent a). reflexivity.
Qed.
Lemma Nless_def_1 :
- forall a a', Nless (Ndouble a) (Ndouble a') = Nless a a'.
+ forall a a', Nless (N.double a) (N.double a') = Nless a a'.
Proof.
destruct a; destruct a'. reflexivity.
trivial.
unfold Nless. simpl. destruct p; trivial.
- unfold Nless. simpl. destruct (Pxor p p0). reflexivity.
+ unfold Nless. simpl. destruct (Pos.lxor p p0). reflexivity.
trivial.
Qed.
Lemma Nless_def_2 :
forall a a',
- Nless (Ndouble_plus_one a) (Ndouble_plus_one a') = Nless a a'.
+ Nless (N.succ_double a) (N.succ_double a') = Nless a a'.
Proof.
destruct a; destruct a'. reflexivity.
trivial.
unfold Nless. simpl. destruct p; trivial.
- unfold Nless. simpl. destruct (Pxor p p0). reflexivity.
+ unfold Nless. simpl. destruct (Pos.lxor p p0). reflexivity.
trivial.
Qed.
Lemma Nless_def_3 :
- forall a a', Nless (Ndouble a) (Ndouble_plus_one a') = true.
+ forall a a', Nless (N.double a) (N.succ_double a') = true.
Proof.
intros. apply Nbit0_less. apply Ndouble_bit0.
apply Ndouble_plus_one_bit0.
Qed.
Lemma Nless_def_4 :
- forall a a', Nless (Ndouble_plus_one a) (Ndouble a') = false.
+ forall a a', Nless (N.succ_double a) (N.double a') = false.
Proof.
intros. apply Nbit0_gt. apply Ndouble_plus_one_bit0.
apply Ndouble_bit0.
@@ -425,7 +425,7 @@ Qed.
Lemma Nless_z : forall a, Nless a N0 = false.
Proof.
induction a. reflexivity.
- unfold Nless. rewrite (Nxor_neutral_right (Npos p)). induction p; trivial.
+ unfold Nless. rewrite (N.lxor_0_r (Npos p)). induction p; trivial.
Qed.
Lemma N0_less_1 :
@@ -445,26 +445,26 @@ Lemma Nless_trans :
forall a a' a'',
Nless a a' = true -> Nless a' a'' = true -> Nless a a'' = true.
Proof.
- induction a as [|a IHa|a IHa] using N_ind_double; intros a' a'' H H0.
+ induction a as [|a IHa|a IHa] using N.binary_ind; intros a' a'' H H0.
case_eq (Nless N0 a'') ; intros Heqn. trivial.
rewrite (N0_less_2 a'' Heqn), (Nless_z a') in H0. discriminate H0.
- induction a' as [|a' _|a' _] using N_ind_double.
- rewrite (Nless_z (Ndouble a)) in H. discriminate H.
+ induction a' as [|a' _|a' _] using N.binary_ind.
+ rewrite (Nless_z (N.double a)) in H. discriminate H.
rewrite (Nless_def_1 a a') in H.
- induction a'' using N_ind_double.
- rewrite (Nless_z (Ndouble a')) in H0. discriminate H0.
+ induction a'' using N.binary_ind.
+ rewrite (Nless_z (N.double a')) in H0. discriminate H0.
rewrite (Nless_def_1 a' a'') in H0. rewrite (Nless_def_1 a a'').
exact (IHa _ _ H H0).
apply Nless_def_3.
- induction a'' as [|a'' _|a'' _] using N_ind_double.
- rewrite (Nless_z (Ndouble_plus_one a')) in H0. discriminate H0.
+ induction a'' as [|a'' _|a'' _] using N.binary_ind.
+ rewrite (Nless_z (N.succ_double a')) in H0. discriminate H0.
rewrite (Nless_def_4 a' a'') in H0. discriminate H0.
apply Nless_def_3.
- induction a' as [|a' _|a' _] using N_ind_double.
- rewrite (Nless_z (Ndouble_plus_one a)) in H. discriminate H.
+ induction a' as [|a' _|a' _] using N.binary_ind.
+ rewrite (Nless_z (N.succ_double a)) in H. discriminate H.
rewrite (Nless_def_4 a a') in H. discriminate H.
- induction a'' using N_ind_double.
- rewrite (Nless_z (Ndouble_plus_one a')) in H0. discriminate H0.
+ induction a'' using N.binary_ind.
+ rewrite (Nless_z (N.succ_double a')) in H0. discriminate H0.
rewrite (Nless_def_4 a' a'') in H0. discriminate H0.
rewrite (Nless_def_2 a' a'') in H0. rewrite (Nless_def_2 a a') in H.
rewrite (Nless_def_2 a a''). exact (IHa _ _ H H0).
@@ -473,17 +473,17 @@ Qed.
Lemma Nless_total :
forall a a', {Nless a a' = true} + {Nless a' a = true} + {a = a'}.
Proof.
- induction a using N_rec_double; intro a'.
+ induction a using N.binary_rec; intro a'.
case_eq (Nless N0 a') ; intros Heqb. left. left. auto.
right. rewrite (N0_less_2 a' Heqb). reflexivity.
- induction a' as [|a' _|a' _] using N_rec_double.
- case_eq (Nless N0 (Ndouble a)) ; intros Heqb. left. right. auto.
+ induction a' as [|a' _|a' _] using N.binary_rec.
+ case_eq (Nless N0 (N.double a)) ; intros Heqb. left. right. auto.
right. exact (N0_less_2 _ Heqb).
rewrite 2!Nless_def_1. destruct (IHa a') as [ | ->].
left. assumption.
right. reflexivity.
left. left. apply Nless_def_3.
- induction a' as [|a' _|a' _] using N_rec_double.
+ induction a' as [|a' _|a' _] using N.binary_rec.
left. right. destruct a; reflexivity.
left. right. apply Nless_def_3.
rewrite 2!Nless_def_2. destruct (IHa a') as [ | ->].
@@ -497,15 +497,15 @@ Notation Nsize := N.size_nat (compat "8.3").
(** conversions between N and bit vectors. *)
-Fixpoint P2Bv (p:positive) : Bvector (Psize p) :=
- match p return Bvector (Psize p) with
+Fixpoint P2Bv (p:positive) : Bvector (Pos.size_nat p) :=
+ match p return Bvector (Pos.size_nat p) with
| xH => Bvect_true 1%nat
- | xO p => Bcons false (Psize p) (P2Bv p)
- | xI p => Bcons true (Psize p) (P2Bv p)
+ | xO p => Bcons false (Pos.size_nat p) (P2Bv p)
+ | xI p => Bcons true (Pos.size_nat p) (P2Bv p)
end.
-Definition N2Bv (n:N) : Bvector (Nsize n) :=
- match n as n0 return Bvector (Nsize n0) with
+Definition N2Bv (n:N) : Bvector (N.size_nat n) :=
+ match n as n0 return Bvector (N.size_nat n0) with
| N0 => Bnil
| Npos p => P2Bv p
end.
@@ -513,8 +513,8 @@ Definition N2Bv (n:N) : Bvector (Nsize n) :=
Fixpoint Bv2N (n:nat)(bv:Bvector n) : N :=
match bv with
| Vector.nil => N0
- | Vector.cons false n bv => Ndouble (Bv2N n bv)
- | Vector.cons true n bv => Ndouble_plus_one (Bv2N n bv)
+ | Vector.cons false n bv => N.double (Bv2N n bv)
+ | Vector.cons true n bv => N.succ_double (Bv2N n bv)
end.
Lemma Bv2N_N2Bv : forall n, Bv2N _ (N2Bv n) = n.
@@ -528,7 +528,7 @@ Qed.
bit vector has some zeros on its right, they will disappear during
the return [Bv2N] translation: *)
-Lemma Bv2N_Nsize : forall n (bv:Bvector n), Nsize (Bv2N n bv) <= n.
+Lemma Bv2N_Nsize : forall n (bv:Bvector n), N.size_nat (Bv2N n bv) <= n.
Proof.
induction bv; intros.
auto.
@@ -543,7 +543,7 @@ Qed.
Lemma Bv2N_Nsize_1 : forall n (bv:Bvector (S n)),
Bsign _ bv = true <->
- Nsize (Bv2N _ bv) = (S n).
+ N.size_nat (Bv2N _ bv) = (S n).
Proof.
apply Vector.rectS ; intros ; simpl.
destruct a ; compute ; split ; intros x ; now inversion x.
@@ -567,7 +567,7 @@ Fixpoint N2Bv_gen (n:nat)(a:N) : Bvector n :=
(** The first [N2Bv] is then a special case of [N2Bv_gen] *)
-Lemma N2Bv_N2Bv_gen : forall (a:N), N2Bv a = N2Bv_gen (Nsize a) a.
+Lemma N2Bv_N2Bv_gen : forall (a:N), N2Bv a = N2Bv_gen (N.size_nat a) a.
Proof.
destruct a; simpl.
auto.
@@ -578,7 +578,7 @@ Qed.
[a] plus some zeros. *)
Lemma N2Bv_N2Bv_gen_above : forall (a:N)(k:nat),
- N2Bv_gen (Nsize a + k) a = Vector.append (N2Bv a) (Bvect_false k).
+ N2Bv_gen (N.size_nat a + k) a = Vector.append (N2Bv a) (Bvect_false k).
Proof.
destruct a; simpl.
destruct k; simpl; auto.
@@ -603,7 +603,7 @@ Qed.
(** accessing some precise bits. *)
Lemma Nbit0_Blow : forall n, forall (bv:Bvector (S n)),
- Nbit0 (Bv2N _ bv) = Blow _ bv.
+ N.odd (Bv2N _ bv) = Blow _ bv.
Proof.
apply Vector.caseS.
intros.
@@ -616,7 +616,7 @@ Qed.
Notation Bnth := (@Vector.nth_order bool).
Lemma Bnth_Nbit : forall n (bv:Bvector n) p (H:p<n),
- Bnth bv H = Nbit (Bv2N _ bv) p.
+ Bnth bv H = N.testbit_nat (Bv2N _ bv) p.
Proof.
induction bv; intros.
inversion H.
@@ -626,7 +626,7 @@ destruct p ; simpl.
simpl in * ; destruct (Bv2N n bv); destruct h; simpl in *; auto.
Qed.
-Lemma Nbit_Nsize : forall n p, Nsize n <= p -> Nbit n p = false.
+Lemma Nbit_Nsize : forall n p, N.size_nat n <= p -> N.testbit_nat n p = false.
Proof.
destruct n as [|n].
simpl; auto.
@@ -635,7 +635,8 @@ inversion H.
inversion H.
Qed.
-Lemma Nbit_Bth: forall n p (H:p < Nsize n), Nbit n p = Bnth (N2Bv n) H.
+Lemma Nbit_Bth: forall n p (H:p < N.size_nat n),
+ N.testbit_nat n p = Bnth (N2Bv n) H.
Proof.
destruct n as [|n].
inversion H.
@@ -646,7 +647,7 @@ Qed.
(** Binary bitwise operations are the same in the two worlds. *)
Lemma Nxor_BVxor : forall n (bv bv' : Bvector n),
- Bv2N _ (BVxor _ bv bv') = Nxor (Bv2N _ bv) (Bv2N _ bv').
+ Bv2N _ (BVxor _ bv bv') = N.lxor (Bv2N _ bv) (Bv2N _ bv').
Proof.
apply Vector.rect2 ; intros.
now simpl.
diff --git a/theories/NArith/Ndist.v b/theories/NArith/Ndist.v
index 22adc5050..7097159c7 100644
--- a/theories/NArith/Ndist.v
+++ b/theories/NArith/Ndist.v
@@ -38,7 +38,7 @@ Qed.
Lemma Nplength_zeros :
forall (a:N) (n:nat),
- Nplength a = ni n -> forall k:nat, k < n -> Nbit a k = false.
+ Nplength a = ni n -> forall k:nat, k < n -> N.testbit_nat a k = false.
Proof.
simple induction a; trivial.
simple induction p. simple induction n. intros. inversion H1.
@@ -46,33 +46,33 @@ Proof.
intros. simpl in H1. discriminate H1.
simple induction k. trivial.
generalize H0. case n. intros. inversion H3.
- intros. simpl in |- *. unfold Nbit in H. apply (H n0). simpl in H1. inversion H1. reflexivity.
+ intros. simpl in |- *. unfold N.testbit_nat in H. apply (H n0). simpl in H1. inversion H1. reflexivity.
exact (lt_S_n n1 n0 H3).
simpl in |- *. intros n H. inversion H. intros. inversion H0.
Qed.
Lemma Nplength_one :
- forall (a:N) (n:nat), Nplength a = ni n -> Nbit a n = true.
+ forall (a:N) (n:nat), Nplength a = ni n -> N.testbit_nat a n = true.
Proof.
simple induction a. intros. inversion H.
simple induction p. intros. simpl in H0. inversion H0. reflexivity.
- intros. simpl in H0. inversion H0. simpl in |- *. unfold Nbit in H. apply H. reflexivity.
+ intros. simpl in H0. inversion H0. simpl in |- *. unfold N.testbit_nat in H. apply H. reflexivity.
intros. simpl in H. inversion H. reflexivity.
Qed.
Lemma Nplength_first_one :
forall (a:N) (n:nat),
- (forall k:nat, k < n -> Nbit a k = false) ->
- Nbit a n = true -> Nplength a = ni n.
+ (forall k:nat, k < n -> N.testbit_nat a k = false) ->
+ N.testbit_nat a n = true -> Nplength a = ni n.
Proof.
simple induction a. intros. simpl in H0. discriminate H0.
simple induction p. intros. generalize H0. case n. intros. reflexivity.
- intros. absurd (Nbit (Npos (xI p0)) 0 = false). trivial with bool.
+ intros. absurd (N.testbit_nat (Npos (xI p0)) 0 = false). trivial with bool.
auto with bool arith.
intros. generalize H0 H1. case n. intros. simpl in H3. discriminate H3.
intros. simpl in |- *. unfold Nplength in H.
cut (ni (Pplength p0) = ni n0). intro. inversion H4. reflexivity.
- apply H. intros. change (Nbit (Npos (xO p0)) (S k) = false) in |- *. apply H2. apply lt_n_S. exact H4.
+ apply H. intros. change (N.testbit_nat (Npos (xO p0)) (S k) = false) in |- *. apply H2. apply lt_n_S. exact H4.
exact H3.
intro. case n. trivial.
intros. simpl in H0. discriminate H0.
@@ -222,27 +222,27 @@ Qed.
Lemma Nplength_lb :
forall (a:N) (n:nat),
- (forall k:nat, k < n -> Nbit a k = false) -> ni_le (ni n) (Nplength a).
+ (forall k:nat, k < n -> N.testbit_nat a k = false) -> ni_le (ni n) (Nplength a).
Proof.
simple induction a. intros. exact (ni_min_inf_r (ni n)).
intros. unfold Nplength in |- *. apply le_ni_le. case (le_or_lt n (Pplength p)). trivial.
- intro. absurd (Nbit (Npos p) (Pplength p) = false).
+ intro. absurd (N.testbit_nat (Npos p) (Pplength p) = false).
rewrite
(Nplength_one (Npos p) (Pplength p)
- (refl_equal (Nplength (Npos p)))).
+ (eq_refl (Nplength (Npos p)))).
discriminate.
apply H. exact H0.
Qed.
Lemma Nplength_ub :
- forall (a:N) (n:nat), Nbit a n = true -> ni_le (Nplength a) (ni n).
+ forall (a:N) (n:nat), N.testbit_nat a n = true -> ni_le (Nplength a) (ni n).
Proof.
simple induction a. intros. discriminate H.
intros. unfold Nplength in |- *. apply le_ni_le. case (le_or_lt (Pplength p) n). trivial.
- intro. absurd (Nbit (Npos p) n = true).
+ intro. absurd (N.testbit_nat (Npos p) n = true).
rewrite
(Nplength_zeros (Npos p) (Pplength p)
- (refl_equal (Nplength (Npos p))) n H0).
+ (eq_refl (Nplength (Npos p))) n H0).
discriminate.
exact H.
Qed.
@@ -255,26 +255,26 @@ Qed.
Instead of working with $d$, we work with $pd$, namely
[Npdist]: *)
-Definition Npdist (a a':N) := Nplength (Nxor a a').
+Definition Npdist (a a':N) := Nplength (N.lxor a a').
(** d is a distance, so $d(a,a')=0$ iff $a=a'$; this means that
$pd(a,a')=infty$ iff $a=a'$: *)
Lemma Npdist_eq_1 : forall a:N, Npdist a a = infty.
Proof.
- intros. unfold Npdist in |- *. rewrite Nxor_nilpotent. reflexivity.
+ intros. unfold Npdist in |- *. rewrite N.lxor_nilpotent. reflexivity.
Qed.
Lemma Npdist_eq_2 : forall a a':N, Npdist a a' = infty -> a = a'.
Proof.
- intros. apply Nxor_eq. apply Nplength_infty. exact H.
+ intros. apply N.lxor_eq. apply Nplength_infty. exact H.
Qed.
(** $d$ is a distance, so $d(a,a')=d(a',a)$: *)
Lemma Npdist_comm : forall a a':N, Npdist a a' = Npdist a' a.
Proof.
- unfold Npdist in |- *. intros. rewrite Nxor_comm. reflexivity.
+ unfold Npdist in |- *. intros. rewrite N.lxor_comm. reflexivity.
Qed.
(** $d$ is an ultrametric distance, that is, not only $d(a,a')\leq
@@ -292,21 +292,21 @@ Qed.
Lemma Nplength_ultra_1 :
forall a a':N,
ni_le (Nplength a) (Nplength a') ->
- ni_le (Nplength a) (Nplength (Nxor a a')).
+ ni_le (Nplength a) (Nplength (N.lxor a a')).
Proof.
simple induction a. intros. unfold ni_le in H. unfold Nplength at 1 3 in H.
rewrite (ni_min_inf_l (Nplength a')) in H.
rewrite (Nplength_infty a' H). simpl in |- *. apply ni_le_refl.
intros. unfold Nplength at 1 in |- *. apply Nplength_lb. intros.
- cut (forall a'':N, Nxor (Npos p) a' = a'' -> Nbit a'' k = false).
+ cut (forall a'':N, N.lxor (Npos p) a' = a'' -> N.testbit_nat a'' k = false).
intros. apply H1. reflexivity.
intro a''. case a''. intro. reflexivity.
intros. rewrite <- H1. rewrite (Nxor_semantics (Npos p) a' k).
rewrite
(Nplength_zeros (Npos p) (Pplength p)
- (refl_equal (Nplength (Npos p))) k H0).
+ (eq_refl (Nplength (Npos p))) k H0).
generalize H. case a'. trivial.
- intros. cut (Nbit (Npos p1) k = false). intros. rewrite H3. reflexivity.
+ intros. cut (N.testbit_nat (Npos p1) k = false). intros. rewrite H3. reflexivity.
apply Nplength_zeros with (n := Pplength p1). reflexivity.
apply (lt_le_trans k (Pplength p) (Pplength p1)). exact H0.
apply ni_le_le. exact H2.
@@ -314,14 +314,14 @@ Qed.
Lemma Nplength_ultra :
forall a a':N,
- ni_le (ni_min (Nplength a) (Nplength a')) (Nplength (Nxor a a')).
+ ni_le (ni_min (Nplength a) (Nplength a')) (Nplength (N.lxor a a')).
Proof.
intros. case (ni_le_total (Nplength a) (Nplength a')). intro.
cut (ni_min (Nplength a) (Nplength a') = Nplength a).
intro. rewrite H0. apply Nplength_ultra_1. exact H.
exact H.
intro. cut (ni_min (Nplength a) (Nplength a') = Nplength a').
- intro. rewrite H0. rewrite Nxor_comm. apply Nplength_ultra_1. exact H.
+ intro. rewrite H0. rewrite N.lxor_comm. apply Nplength_ultra_1. exact H.
rewrite ni_min_comm. exact H.
Qed.
@@ -329,8 +329,8 @@ Lemma Npdist_ultra :
forall a a' a'':N,
ni_le (ni_min (Npdist a a'') (Npdist a'' a')) (Npdist a a').
Proof.
- intros. unfold Npdist in |- *. cut (Nxor (Nxor a a'') (Nxor a'' a') = Nxor a a').
+ intros. unfold Npdist in |- *. cut (N.lxor (N.lxor a a'') (N.lxor a'' a') = N.lxor a a').
intro. rewrite <- H. apply Nplength_ultra.
- rewrite Nxor_assoc. rewrite <- (Nxor_assoc a'' a'' a'). rewrite Nxor_nilpotent.
- rewrite Nxor_neutral_left. reflexivity.
+ rewrite N.lxor_assoc. rewrite <- (N.lxor_assoc a'' a'' a'). rewrite N.lxor_nilpotent.
+ rewrite N.lxor_0_l. reflexivity.
Qed.
diff --git a/theories/Numbers/BigNumPrelude.v b/theories/Numbers/BigNumPrelude.v
index 26850688e..bf9238010 100644
--- a/theories/Numbers/BigNumPrelude.v
+++ b/theories/Numbers/BigNumPrelude.v
@@ -30,7 +30,7 @@ Declare ML Module "numbers_syntax_plugin".
Local Open Scope Z_scope.
-(* For compatibility of scripts, weaker version of some lemmas of Zdiv *)
+(* For compatibility of scripts, weaker version of some lemmas of Z.div *)
Lemma Zlt0_not_eq : forall n, 0<n -> n<>0.
Proof.
@@ -43,22 +43,22 @@ Definition Z_div_plus_l a b c H := Zdiv.Z_div_plus_full_l a b c (Zlt0_not_eq _ H
(* Automation *)
-Hint Extern 2 (Zle _ _) =>
+Hint Extern 2 (Z.le _ _) =>
(match goal with
- |- Zpos _ <= Zpos _ => exact (refl_equal _)
-| H: _ <= ?p |- _ <= ?p => apply Zle_trans with (2 := H)
-| H: _ < ?p |- _ <= ?p => apply Zlt_le_weak; apply Zle_lt_trans with (2 := H)
+ |- Zpos _ <= Zpos _ => exact (eq_refl _)
+| H: _ <= ?p |- _ <= ?p => apply Z.le_trans with (2 := H)
+| H: _ < ?p |- _ <= ?p => apply Z.lt_le_incl; apply Z.le_lt_trans with (2 := H)
end).
-Hint Extern 2 (Zlt _ _) =>
+Hint Extern 2 (Z.lt _ _) =>
(match goal with
- |- Zpos _ < Zpos _ => exact (refl_equal _)
-| H: _ <= ?p |- _ <= ?p => apply Zlt_le_trans with (2 := H)
-| H: _ < ?p |- _ <= ?p => apply Zle_lt_trans with (2 := H)
+ |- Zpos _ < Zpos _ => exact (eq_refl _)
+| H: _ <= ?p |- _ <= ?p => apply Z.lt_le_trans with (2 := H)
+| H: _ < ?p |- _ <= ?p => apply Z.le_lt_trans with (2 := H)
end).
-Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
+Hint Resolve Z.lt_gt Z.le_ge Z_div_pos: zarith.
(**************************************
Properties of order and product
@@ -71,9 +71,9 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
Proof.
intros a b c d beta H1 (H3, H4) (H5, H6).
assert (a - c < 1); auto with zarith.
- apply Zmult_lt_reg_r with beta; auto with zarith.
- apply Zle_lt_trans with (d - b); auto with zarith.
- rewrite Zmult_minus_distr_r; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with beta; auto with zarith.
+ apply Z.le_lt_trans with (d - b); auto with zarith.
+ rewrite Z.mul_sub_distr_r; auto with zarith.
Qed.
Theorem beta_lex_inv: forall a b c d beta,
@@ -82,15 +82,15 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
a * beta + b < c * beta + d.
Proof.
intros a b c d beta H1 (H3, H4) (H5, H6).
- case (Zle_or_lt (c * beta + d) (a * beta + b)); auto with zarith.
- intros H7; contradict H1;apply Zle_not_lt;apply beta_lex with (1 := H7);auto.
+ case (Z.le_gt_cases (c * beta + d) (a * beta + b)); auto with zarith.
+ intros H7. contradict H1. apply Z.le_ngt. apply beta_lex with (1 := H7); auto.
Qed.
Lemma beta_mult : forall h l beta,
0 <= h < beta -> 0 <= l < beta -> 0 <= h*beta+l < beta^2.
Proof.
intros h l beta H1 H2;split. auto with zarith.
- rewrite <- (Zplus_0_r (beta^2)); rewrite Zpower_2;
+ rewrite <- (Z.add_0_r (beta^2)); rewrite Z.pow_2_r;
apply beta_lex_inv;auto with zarith.
Qed.
@@ -98,9 +98,9 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
forall b x y, 0 <= x < b -> 0 <= y < b -> 0 <= x * y <= b^2 - 2*b + 1.
Proof.
intros b x y (Hx1,Hx2) (Hy1,Hy2);split;auto with zarith.
- apply Zle_trans with ((b-1)*(b-1)).
- apply Zmult_le_compat;auto with zarith.
- apply Zeq_le; ring.
+ apply Z.le_trans with ((b-1)*(b-1)).
+ apply Z.mul_le_mono_nonneg;auto with zarith.
+ apply Z.eq_le_incl; ring.
Qed.
Lemma sum_mul_carry : forall xh xl yh yl wc cc beta,
@@ -129,11 +129,10 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
Proof.
intros x y cross beta HH HH1 HH2.
split; auto with zarith.
- apply Zle_lt_trans with ((beta-1)*(beta-1)+(beta-1)); auto with zarith.
- apply Zplus_le_compat; auto with zarith.
- apply Zmult_le_compat; auto with zarith.
- repeat (rewrite Zmult_minus_distr_l || rewrite Zmult_minus_distr_r);
- rewrite Zpower_2; auto with zarith.
+ apply Z.le_lt_trans with ((beta-1)*(beta-1)+(beta-1)); auto with zarith.
+ apply Z.add_le_mono; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
+ rewrite ?Z.mul_sub_distr_l, ?Z.mul_sub_distr_r, Z.pow_2_r; auto with zarith.
Qed.
Theorem mult_add_ineq2: forall x y c cross beta,
@@ -144,11 +143,10 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
Proof.
intros x y c cross beta HH HH1 HH2.
split; auto with zarith.
- apply Zle_lt_trans with ((beta-1)*(beta-1)+(2*beta-2));auto with zarith.
- apply Zplus_le_compat; auto with zarith.
- apply Zmult_le_compat; auto with zarith.
- repeat (rewrite Zmult_minus_distr_l || rewrite Zmult_minus_distr_r);
- rewrite Zpower_2; auto with zarith.
+ apply Z.le_lt_trans with ((beta-1)*(beta-1)+(2*beta-2));auto with zarith.
+ apply Z.add_le_mono; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
+ rewrite ?Z.mul_sub_distr_l, ?Z.mul_sub_distr_r, Z.pow_2_r; auto with zarith.
Qed.
Theorem mult_add_ineq3: forall x y c cross beta,
@@ -161,20 +159,20 @@ Theorem mult_add_ineq3: forall x y c cross beta,
intros x y c cross beta HH HH1 HH2 HH3.
apply mult_add_ineq2;auto with zarith.
split;auto with zarith.
- apply Zle_trans with (1*beta+cross);auto with zarith.
+ apply Z.le_trans with (1*beta+cross);auto with zarith.
Qed.
-Hint Rewrite Zmult_1_r Zmult_0_r Zmult_1_l Zmult_0_l Zplus_0_l Zplus_0_r Zminus_0_r: rm10.
+Hint Rewrite Z.mul_1_r Z.mul_0_r Z.mul_1_l Z.mul_0_l Z.add_0_l Z.add_0_r Z.sub_0_r: rm10.
(**************************************
- Properties of Zdiv and Zmod
+ Properties of Z.div and Z.modulo
**************************************)
Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Proof.
intros a b H H1;case (Z_mod_lt a b);auto with zarith;intros H2 H3;split;auto.
- case (Zle_or_lt b a); intros H4; auto with zarith.
+ case (Z.le_gt_cases b a); intros H4; auto with zarith.
rewrite Zmod_small; auto with zarith.
Qed.
@@ -184,26 +182,26 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Proof.
intros a b r t (H1, H2) H3 (H4, H5).
assert (t < 2 ^ b).
- apply Zlt_le_trans with (1:= H5); auto with zarith.
+ apply Z.lt_le_trans with (1:= H5); auto with zarith.
apply Zpower_le_monotone; auto with zarith.
rewrite Zplus_mod; auto with zarith.
rewrite Zmod_small with (a := t); auto with zarith.
apply Zmod_small; auto with zarith.
split; auto with zarith.
assert (0 <= 2 ^a * r); auto with zarith.
- apply Zplus_le_0_compat; auto with zarith.
+ apply Z.add_nonneg_nonneg; auto with zarith.
match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
pattern (2 ^ b) at 2; replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a);
try ring.
- apply Zplus_le_lt_compat; auto with zarith.
+ apply Z.add_le_lt_mono; auto with zarith.
replace b with ((b - a) + a); try ring.
rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
- try rewrite <- Zmult_minus_distr_r.
- rewrite (Zmult_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l;
+ pattern (2 ^a) at 4; rewrite <- (Z.mul_1_l (2 ^a));
+ try rewrite <- Z.mul_sub_distr_r.
+ rewrite (Z.mul_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l;
auto with zarith.
- rewrite (Zmult_comm (2 ^a)); apply Zmult_le_compat_r; auto with zarith.
+ rewrite (Z.mul_comm (2 ^a)); apply Z.mul_le_mono_nonneg_r; auto with zarith.
match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
Qed.
@@ -214,25 +212,25 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Proof.
intros a b r t (H1, H2) H3 (H4, H5).
assert (t < 2 ^ b).
- apply Zlt_le_trans with (1:= H5); auto with zarith.
+ apply Z.lt_le_trans with (1:= H5); auto with zarith.
apply Zpower_le_monotone; auto with zarith.
rewrite Zplus_mod; auto with zarith.
rewrite Zmod_small with (a := t); auto with zarith.
apply Zmod_small; auto with zarith.
split; auto with zarith.
assert (0 <= 2 ^a * r); auto with zarith.
- apply Zplus_le_0_compat; auto with zarith.
+ apply Z.add_nonneg_nonneg; auto with zarith.
match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
pattern (2 ^ b) at 2;replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); try ring.
- apply Zplus_le_lt_compat; auto with zarith.
+ apply Z.add_le_lt_mono; auto with zarith.
replace b with ((b - a) + a); try ring.
rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
- try rewrite <- Zmult_minus_distr_r.
- repeat rewrite (fun x => Zmult_comm x (2 ^ a)); rewrite Zmult_mod_distr_l;
+ pattern (2 ^a) at 4; rewrite <- (Z.mul_1_l (2 ^a));
+ try rewrite <- Z.mul_sub_distr_r.
+ repeat rewrite (fun x => Z.mul_comm x (2 ^ a)); rewrite Zmult_mod_distr_l;
auto with zarith.
- apply Zmult_le_compat_l; auto with zarith.
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
Qed.
@@ -243,13 +241,13 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Proof.
intros a b r t (H1, H2) H3 (H4, H5).
assert (Eq: t < 2 ^ b); auto with zarith.
- apply Zlt_le_trans with (1 := H5); auto with zarith.
+ apply Z.lt_le_trans with (1 := H5); auto with zarith.
apply Zpower_le_monotone; auto with zarith.
pattern (r * 2 ^ a) at 1; rewrite Z_div_mod_eq with (b := 2 ^ b);
auto with zarith.
- rewrite <- Zplus_assoc.
+ rewrite <- Z.add_assoc.
rewrite <- Zmod_shift_r; auto with zarith.
- rewrite (Zmult_comm (2 ^ b)); rewrite Z_div_plus_full_l; auto with zarith.
+ rewrite (Z.mul_comm (2 ^ b)); rewrite Z_div_plus_full_l; auto with zarith.
rewrite (fun x y => @Zdiv_small (x mod y)); auto with zarith.
match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
@@ -264,7 +262,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
intros n p a H1 H2.
pattern (a*2^p) at 1;replace (a*2^p) with
(a*2^p/2^n * 2^n + a*2^p mod 2^n).
- 2:symmetry;rewrite (Zmult_comm (a*2^p/2^n));apply Z_div_mod_eq.
+ 2:symmetry;rewrite (Z.mul_comm (a*2^p/2^n));apply Z_div_mod_eq.
replace (a * 2 ^ p / 2 ^ n) with (a / 2 ^ (n - p));trivial.
replace (2^n) with (2^(n-p)*2^p).
symmetry;apply Zdiv_mult_cancel_r.
@@ -273,7 +271,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
rewrite <- Zpower_exp.
replace (n-p+p) with n;trivial. ring.
omega. omega.
- apply Zlt_gt. apply Zpower_gt_0;auto with zarith.
+ apply Z.lt_gt. apply Z.pow_pos_nonneg;auto with zarith.
Qed.
@@ -284,15 +282,15 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
intros.
rewrite Zmod_small.
rewrite Zmod_eq by (auto with zarith).
- unfold Zminus at 1.
+ unfold Z.sub at 1.
rewrite Z_div_plus_l by (auto with zarith).
assert (2^n = 2^(n-p)*2^p).
rewrite <- Zpower_exp by (auto with zarith).
replace (n-p+p) with n; auto with zarith.
rewrite H0.
rewrite <- Zdiv_Zdiv, Z_div_mult by (auto with zarith).
- rewrite (Zmult_comm (2^(n-p))), Zmult_assoc.
- rewrite Zopp_mult_distr_l.
+ rewrite (Z.mul_comm (2^(n-p))), Z.mul_assoc.
+ rewrite <- Z.mul_opp_l.
rewrite Z_div_mult by (auto with zarith).
symmetry; apply Zmod_eq; auto with zarith.
@@ -301,9 +299,9 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
split.
apply Z_div_pos; auto with zarith.
apply Zdiv_lt_upper_bound; auto with zarith.
- apply Zlt_le_trans with (2^n); auto with zarith.
- rewrite <- (Zmult_1_r (2^n)) at 1.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.lt_le_trans with (2^n); auto with zarith.
+ rewrite <- (Z.mul_1_r (2^n)) at 1.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
cut (0 < 2 ^ (n-p)); auto with zarith.
Qed.
@@ -320,8 +318,8 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Proof.
intros p x y H;destruct (Z_le_gt_dec 0 p).
apply Zdiv_lt_upper_bound;auto with zarith.
- apply Zlt_le_trans with y;auto with zarith.
- rewrite <- (Zmult_1_r y);apply Zmult_le_compat;auto with zarith.
+ apply Z.lt_le_trans with y;auto with zarith.
+ rewrite <- (Z.mul_1_r y);apply Z.mul_le_mono_nonneg;auto with zarith.
assert (0 < 2^p);auto with zarith.
replace (2^p) with 0.
destruct x;change (0<y);auto with zarith.
@@ -329,15 +327,13 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Qed.
Theorem Zgcd_div_pos a b:
- 0 < b -> 0 < Zgcd a b -> 0 < b / Zgcd a b.
+ 0 < b -> 0 < Z.gcd a b -> 0 < b / Z.gcd a b.
Proof.
- intros Ha Hg.
- case (Zle_lt_or_eq 0 (b/Zgcd a b)); auto.
- apply Z_div_pos; auto with zarith.
- intros H; generalize Ha.
- pattern b at 1; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto.
- rewrite <- H; auto with zarith.
- assert (F := (Zgcd_is_gcd a b)); inversion F; auto.
+ intros Hb Hg.
+ assert (H : 0 <= b / Z.gcd a b) by (apply Z.div_pos; auto with zarith).
+ Z.le_elim H; trivial.
+ rewrite (Zdivide_Zdiv_eq (Z.gcd a b) b), <- H, Z.mul_0_r in Hb;
+ auto using Z.gcd_divide_r with zarith.
Qed.
Theorem Zdiv_neg a b:
@@ -347,7 +343,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
assert (b > 0) by omega.
generalize (Z_mult_div_ge a _ H); intros.
assert (b * (a / b) < 0)%Z.
- apply Zle_lt_trans with a; auto with zarith.
+ apply Z.le_lt_trans with a; auto with zarith.
destruct b; try (compute in Hb; discriminate).
destruct (a/Zpos p)%Z.
compute in H1; discriminate.
@@ -355,20 +351,20 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
compute; auto.
Qed.
- Lemma Zdiv_gcd_zero : forall a b, b / Zgcd a b = 0 -> b <> 0 ->
- Zgcd a b = 0.
+ Lemma Zdiv_gcd_zero : forall a b, b / Z.gcd a b = 0 -> b <> 0 ->
+ Z.gcd a b = 0.
Proof.
intros.
generalize (Zgcd_is_gcd a b); destruct 1.
destruct H2 as (k,Hk).
generalize H; rewrite Hk at 1.
- destruct (Z_eq_dec (Zgcd a b) 0) as [H'|H']; auto.
+ destruct (Z.eq_dec (Z.gcd a b) 0) as [H'|H']; auto.
rewrite Z_div_mult_full; auto.
intros; subst k; simpl in *; subst b; elim H0; auto.
Qed.
Lemma Zgcd_mult_rel_prime : forall a b c,
- Zgcd a c = 1 -> Zgcd b c = 1 -> Zgcd (a*b) c = 1.
+ Z.gcd a c = 1 -> Z.gcd b c = 1 -> Z.gcd (a*b) c = 1.
Proof.
intros.
rewrite Zgcd_1_rel_prime in *.
@@ -396,23 +392,20 @@ intros Q b Q0 QS.
set (Q' := fun n => (n < b /\ Q n) \/ (b <= n)).
assert (H : forall n, 0 <= n -> Q' n).
apply natlike_rec2; unfold Q'.
-destruct (Zle_or_lt b 0) as [H | H]. now right. left; now split.
+destruct (Z.le_gt_cases b 0) as [H | H]. now right. left; now split.
intros n H IH. destruct IH as [[IH1 IH2] | IH].
-destruct (Zle_or_lt (b - 1) n) as [H1 | H1].
+destruct (Z.le_gt_cases (b - 1) n) as [H1 | H1].
right; auto with zarith.
left. split; [auto with zarith | now apply (QS n)].
right; auto with zarith.
unfold Q' in *; intros n H1 H2. destruct (H n H1) as [[H3 H4] | H3].
-assumption. apply Zle_not_lt in H3. false_hyp H2 H3.
+assumption. now apply Z.le_ngt in H3.
Qed.
-Lemma Zsquare_le : forall x, x <= x*x.
+Lemma Zsquare_le x : x <= x*x.
Proof.
-intros.
-destruct (Z_lt_le_dec 0 x).
-pattern x at 1; rewrite <- (Zmult_1_l x).
-apply Zmult_le_compat; auto with zarith.
-apply Zle_trans with 0; auto with zarith.
-rewrite <- Zmult_opp_opp.
-apply Zmult_le_0_compat; auto with zarith.
+destruct (Z.lt_ge_cases 0 x).
+- rewrite <- Z.mul_1_l at 1.
+ rewrite <- Z.mul_le_mono_pos_r; auto with zarith.
+- pose proof (Z.square_nonneg x); auto with zarith.
Qed.
diff --git a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
index 59656eedb..1cfefd3b9 100644
--- a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
+++ b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
@@ -111,7 +111,7 @@ Module ZnZ.
(* Conversion functions with Z *)
spec_to_Z : forall x, 0 <= [| x |] < wB;
spec_of_pos : forall p,
- Zpos p = (Z_of_N (fst (of_pos p)))*wB + [|(snd (of_pos p))|];
+ Zpos p = (Z.of_N (fst (of_pos p)))*wB + [|(snd (of_pos p))|];
spec_zdigits : [| zdigits |] = Zpos digits;
spec_more_than_1_digit: 1 < Zpos digits;
@@ -284,11 +284,11 @@ Module ZnZ.
generalize (spec_of_pos p).
case (of_pos p); intros n w1; simpl.
case n; simpl Npos; auto with zarith.
- intros p1 Hp1; contradict Hp; apply Zle_not_lt.
+ intros p1 Hp1; contradict Hp; apply Z.le_ngt.
replace (base digits) with (1 * base digits + 0) by ring.
rewrite Hp1.
- apply Zplus_le_compat.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.add_le_mono.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
case p1; simpl; intros; red; simpl; intros; discriminate.
unfold base; auto with zarith.
case (spec_to_Z w1); auto with zarith.
@@ -305,7 +305,7 @@ Module ZnZ.
Proof.
intros p; case p; simpl; try rewrite spec_0; auto.
intros; rewrite of_pos_correct; auto with zarith.
- intros p1 (H1, _); contradict H1; apply Zlt_not_le; red; simpl; auto.
+ intros p1 (H1, _); contradict H1; apply Z.lt_nge; red; simpl; auto.
Qed.
End Of_Z.
@@ -346,46 +346,46 @@ Ltac zify := unfold eq in *; autorewrite with cyclic.
Lemma add_0_l : forall x, 0 + x == x.
Proof.
-intros. zify. rewrite Zplus_0_l.
+intros. zify. rewrite Z.add_0_l.
apply Zmod_small. apply ZnZ.spec_to_Z.
Qed.
Lemma add_comm : forall x y, x + y == y + x.
Proof.
-intros. zify. now rewrite Zplus_comm.
+intros. zify. now rewrite Z.add_comm.
Qed.
Lemma add_assoc : forall x y z, x + (y + z) == x + y + z.
Proof.
-intros. zify. now rewrite Zplus_mod_idemp_r, Zplus_mod_idemp_l, Zplus_assoc.
+intros. zify. now rewrite Zplus_mod_idemp_r, Zplus_mod_idemp_l, Z.add_assoc.
Qed.
Lemma mul_1_l : forall x, 1 * x == x.
Proof.
-intros. zify. rewrite Zmult_1_l.
+intros. zify. rewrite Z.mul_1_l.
apply Zmod_small. apply ZnZ.spec_to_Z.
Qed.
Lemma mul_comm : forall x y, x * y == y * x.
Proof.
-intros. zify. now rewrite Zmult_comm.
+intros. zify. now rewrite Z.mul_comm.
Qed.
Lemma mul_assoc : forall x y z, x * (y * z) == x * y * z.
Proof.
-intros. zify. now rewrite Zmult_mod_idemp_r, Zmult_mod_idemp_l, Zmult_assoc.
+intros. zify. now rewrite Zmult_mod_idemp_r, Zmult_mod_idemp_l, Z.mul_assoc.
Qed.
Lemma mul_add_distr_r : forall x y z, (x+y)*z == x*z + y*z.
Proof.
-intros. zify. now rewrite <- Zplus_mod, Zmult_mod_idemp_l, Zmult_plus_distr_l.
+intros. zify. now rewrite <- Zplus_mod, Zmult_mod_idemp_l, Z.mul_add_distr_r.
Qed.
Lemma add_opp_r : forall x y, x + - y == x-y.
Proof.
-intros. zify. rewrite <- Zminus_mod_idemp_r. unfold Zminus.
-destruct (Z_eq_dec ([|y|] mod wB) 0) as [EQ|NEQ].
-rewrite Z_mod_zero_opp_full, EQ, 2 Zplus_0_r; auto.
+intros. zify. rewrite <- Zminus_mod_idemp_r. unfold Z.sub.
+destruct (Z.eq_dec ([|y|] mod wB) 0) as [EQ|NEQ].
+rewrite Z_mod_zero_opp_full, EQ, 2 Z.add_0_r; auto.
rewrite Z_mod_nz_opp_full by auto.
rewrite <- Zplus_mod_idemp_r, <- Zminus_mod_idemp_l.
rewrite Z_mod_same_full. simpl. now rewrite Zplus_mod_idemp_r.
@@ -393,7 +393,7 @@ Qed.
Lemma add_opp_diag_r : forall x, x + - x == 0.
Proof.
-intros. red. rewrite add_opp_r. zify. now rewrite Zminus_diag, Zmod_0_l.
+intros. red. rewrite add_opp_r. zify. now rewrite Z.sub_diag, Zmod_0_l.
Qed.
Lemma CyclicRing : ring_theory 0 1 ZnZ.add ZnZ.mul ZnZ.sub ZnZ.opp eq.
@@ -413,19 +413,9 @@ Lemma eqb_eq : forall x y, eqb x y = true <-> x == y.
Proof.
intros. unfold eqb, eq.
rewrite ZnZ.spec_compare.
- case Zcompare_spec; intuition; try discriminate.
+ case Z.compare_spec; intuition; try discriminate.
Qed.
-(* POUR HUGO:
-Lemma eqb_eq : forall x y, eqb x y = true <-> x == y.
-Proof.
- intros. unfold eqb, eq. generalize (ZnZ.spec_compare x y).
- case (ZnZ.compare x y); intuition; try discriminate.
- (* BUG ?! using destruct instead of case won't work:
- it gives 3 subcases, but ZnZ.compare x y is still there in them! *)
-Qed.
-*)
-
Lemma eqb_correct : forall x y, eqb x y = true -> x==y.
Proof. now apply eqb_eq. Qed.
diff --git a/theories/Numbers/Cyclic/Abstract/NZCyclic.v b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
index c52cbe105..82e4ad13f 100644
--- a/theories/Numbers/Cyclic/Abstract/NZCyclic.v
+++ b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
@@ -69,7 +69,7 @@ Program Instance mul_wd : Proper (eq ==> eq ==> eq) mul.
Theorem gt_wB_1 : 1 < wB.
Proof.
-unfold base. apply Zpower_gt_1; unfold Zlt; auto with zarith.
+unfold base. apply Zpower_gt_1; unfold Z.lt; auto with zarith.
Qed.
Theorem gt_wB_0 : 0 < wB.
@@ -161,20 +161,20 @@ End Induction.
Theorem add_0_l : forall n, 0 + n == n.
Proof.
intro n. zify.
-rewrite Zplus_0_l. apply Zmod_small. apply ZnZ.spec_to_Z.
+rewrite Z.add_0_l. apply Zmod_small. apply ZnZ.spec_to_Z.
Qed.
Theorem add_succ_l : forall n m, (S n) + m == S (n + m).
Proof.
intros n m. zify.
rewrite succ_mod_wB. repeat rewrite Zplus_mod_idemp_l; try apply gt_wB_0.
-rewrite <- (Zplus_assoc ([| n |] mod wB) 1 [| m |]). rewrite Zplus_mod_idemp_l.
-rewrite (Zplus_comm 1 [| m |]); now rewrite Zplus_assoc.
+rewrite <- (Z.add_assoc ([| n |] mod wB) 1 [| m |]). rewrite Zplus_mod_idemp_l.
+rewrite (Z.add_comm 1 [| m |]); now rewrite Z.add_assoc.
Qed.
Theorem sub_0_r : forall n, n - 0 == n.
Proof.
-intro n. zify. rewrite Zminus_0_r. apply NZ_to_Z_mod.
+intro n. zify. rewrite Z.sub_0_r. apply NZ_to_Z_mod.
Qed.
Theorem sub_succ_r : forall n m, n - (S m) == P (n - m).
@@ -192,7 +192,7 @@ Qed.
Theorem mul_succ_l : forall n m, (S n) * m == n * m + m.
Proof.
intros n m. zify. rewrite Zplus_mod_idemp_l, Zmult_mod_idemp_l.
-now rewrite Zmult_plus_distr_l, Zmult_1_l.
+now rewrite Z.mul_add_distr_r, Z.mul_1_l.
Qed.
Definition t := t.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
index deb216ddc..a9b976adb 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
@@ -182,7 +182,7 @@ Section DoubleAdd.
destruct x as [ |xh xl];simpl. apply spec_ww_1.
generalize (spec_w_succ_c xl);destruct (w_succ_c xl) as [l|l];
intro H;unfold interp_carry in H. simpl;rewrite H;ring.
- rewrite <- Zplus_assoc;rewrite <- H;rewrite Zmult_1_l.
+ rewrite <- Z.add_assoc;rewrite <- H;rewrite Z.mul_1_l.
assert ([|l|] = 0). generalize (spec_to_Z xl)(spec_to_Z l);omega.
rewrite H0;generalize (spec_w_succ_c xh);destruct (w_succ_c xh) as [h|h];
intro H1;unfold interp_carry in H1.
@@ -195,19 +195,19 @@ Section DoubleAdd.
Lemma spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]].
Proof.
destruct x as [ |xh xl];simpl;trivial.
- destruct y as [ |yh yl];simpl. rewrite Zplus_0_r;trivial.
+ destruct y as [ |yh yl];simpl. rewrite Z.add_0_r;trivial.
replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]))
with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring.
generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l];
intros H;unfold interp_carry in H;rewrite <- H.
generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h];
intros H1;unfold interp_carry in *;rewrite <- H1. trivial.
- repeat rewrite Zmult_1_l;rewrite spec_w_WW;rewrite wwB_wBwB; ring.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ repeat rewrite Z.mul_1_l;rewrite spec_w_WW;rewrite wwB_wBwB; ring.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh)
as [h|h]; intros H1;unfold interp_carry in *;rewrite <- H1.
simpl;ring.
- repeat rewrite Zmult_1_l;rewrite wwB_wBwB;rewrite spec_w_WW;ring.
+ repeat rewrite Z.mul_1_l;rewrite wwB_wBwB;rewrite spec_w_WW;ring.
Qed.
Section Cont.
@@ -221,23 +221,23 @@ Section DoubleAdd.
destruct x as [ |xh xl];simpl;trivial.
apply spec_f0;trivial.
destruct y as [ |yh yl];simpl.
- apply spec_f0;simpl;rewrite Zplus_0_r;trivial.
+ apply spec_f0;simpl;rewrite Z.add_0_r;trivial.
generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l];
intros H;unfold interp_carry in H.
generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h];
intros H1;unfold interp_carry in *.
apply spec_f0. simpl;rewrite H;rewrite H1;ring.
apply spec_f1. simpl;rewrite spec_w_WW;rewrite H.
- rewrite Zplus_assoc;rewrite wwB_wBwB. rewrite Zpower_2; rewrite <- Zmult_plus_distr_l.
- rewrite Zmult_1_l in H1;rewrite H1;ring.
+ rewrite Z.add_assoc;rewrite wwB_wBwB. rewrite Z.pow_2_r; rewrite <- Z.mul_add_distr_r.
+ rewrite Z.mul_1_l in H1;rewrite H1;ring.
generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh)
as [h|h]; intros H1;unfold interp_carry in *.
- apply spec_f0;simpl;rewrite H1. rewrite Zmult_plus_distr_l.
- rewrite <- Zplus_assoc;rewrite H;ring.
+ apply spec_f0;simpl;rewrite H1. rewrite Z.mul_add_distr_r.
+ rewrite <- Z.add_assoc;rewrite H;ring.
apply spec_f1. simpl;rewrite spec_w_WW;rewrite wwB_wBwB.
- rewrite Zplus_assoc; rewrite Zpower_2; rewrite <- Zmult_plus_distr_l.
- rewrite Zmult_1_l in H1;rewrite H1. rewrite Zmult_plus_distr_l.
- rewrite <- Zplus_assoc;rewrite H;ring.
+ rewrite Z.add_assoc; rewrite Z.pow_2_r; rewrite <- Z.mul_add_distr_r.
+ rewrite Z.mul_1_l in H1;rewrite H1. rewrite Z.mul_add_distr_r.
+ rewrite <- Z.add_assoc;rewrite H;ring.
Qed.
End Cont.
@@ -248,19 +248,19 @@ Section DoubleAdd.
destruct x as [ |xh xl];intro y;simpl.
exact (spec_ww_succ_c y).
destruct y as [ |yh yl];simpl.
- rewrite Zplus_0_r;exact (spec_ww_succ_c (WW xh xl)).
+ rewrite Z.add_0_r;exact (spec_ww_succ_c (WW xh xl)).
replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1)
with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring.
generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl)
as [l|l];intros H;unfold interp_carry in H;rewrite <- H.
generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h];
intros H1;unfold interp_carry in H1;rewrite <- H1. trivial.
- unfold interp_carry;repeat rewrite Zmult_1_l;simpl;rewrite wwB_wBwB;ring.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ unfold interp_carry;repeat rewrite Z.mul_1_l;simpl;rewrite wwB_wBwB;ring.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh)
as [h|h];intros H1;unfold interp_carry in H1;rewrite <- H1. trivial.
unfold interp_carry;rewrite spec_w_WW;
- repeat rewrite Zmult_1_l;simpl;rewrite wwB_wBwB;ring.
+ repeat rewrite Z.mul_1_l;simpl;rewrite wwB_wBwB;ring.
Qed.
Lemma spec_ww_succ : forall x, [[ww_succ x]] = ([[x]] + 1) mod wwB.
@@ -268,14 +268,14 @@ Section DoubleAdd.
destruct x as [ |xh xl];simpl.
rewrite spec_ww_1;rewrite Zmod_small;trivial.
split;[intro;discriminate|apply wwB_pos].
- rewrite <- Zplus_assoc;generalize (spec_w_succ_c xl);
+ rewrite <- Z.add_assoc;generalize (spec_w_succ_c xl);
destruct (w_succ_c xl) as[l|l];intro H;unfold interp_carry in H;rewrite <-H.
rewrite Zmod_small;trivial.
rewrite wwB_wBwB;apply beta_mult;apply spec_to_Z.
assert ([|l|] = 0). clear spec_ww_1 spec_w_1 spec_w_0.
assert (H1:= spec_to_Z l); assert (H2:= spec_to_Z xl); omega.
- rewrite H0;rewrite Zplus_0_r;rewrite <- Zmult_plus_distr_l;rewrite wwB_wBwB.
- rewrite Zpower_2; rewrite Zmult_mod_distr_r;try apply lt_0_wB.
+ rewrite H0;rewrite Z.add_0_r;rewrite <- Z.mul_add_distr_r;rewrite wwB_wBwB.
+ rewrite Z.pow_2_r; rewrite Zmult_mod_distr_r;try apply lt_0_wB.
rewrite spec_w_W0;rewrite spec_w_succ;trivial.
Qed.
@@ -284,7 +284,7 @@ Section DoubleAdd.
destruct x as [ |xh xl];intros y;simpl.
rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
destruct y as [ |yh yl].
- change [[W0]] with 0;rewrite Zplus_0_r.
+ change [[W0]] with 0;rewrite Z.add_0_r.
rewrite Zmod_small;trivial.
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)).
simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]))
@@ -292,7 +292,7 @@ Section DoubleAdd.
generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l];
unfold interp_carry;intros H;simpl;rewrite <- H.
rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial.
Qed.
@@ -302,13 +302,13 @@ Section DoubleAdd.
destruct x as [ |xh xl];intros y;simpl.
exact (spec_ww_succ y).
destruct y as [ |yh yl].
- change [[W0]] with 0;rewrite Zplus_0_r. exact (spec_ww_succ (WW xh xl)).
+ change [[W0]] with 0;rewrite Z.add_0_r. exact (spec_ww_succ (WW xh xl)).
simpl;replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1)
with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring.
generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl)
as [l|l];unfold interp_carry;intros H;rewrite <- H;simpl ww_to_Z.
rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial.
Qed.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
index e6c5a0e04..3eda130fb 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
@@ -161,13 +161,13 @@ Section DoubleBase.
Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
Variable spec_w_compare : forall x y,
- w_compare x y = Zcompare [|x|] [|y|].
+ w_compare x y = Z.compare [|x|] [|y|].
Lemma wwB_wBwB : wwB = wB^2.
Proof.
- unfold base, ww_digits;rewrite Zpower_2; rewrite (Zpos_xO w_digits).
+ unfold base, ww_digits;rewrite Z.pow_2_r; rewrite (Pos2Z.inj_xO w_digits).
replace (2 * Zpos w_digits) with (Zpos w_digits + Zpos w_digits).
- apply Zpower_exp; unfold Zge;simpl;intros;discriminate.
+ apply Zpower_exp; unfold Z.ge;simpl;intros;discriminate.
ring.
Qed.
@@ -179,28 +179,28 @@ Section DoubleBase.
Lemma lt_0_wB : 0 < wB.
Proof.
- unfold base;apply Zpower_gt_0. unfold Zlt;reflexivity.
- unfold Zle;intros H;discriminate H.
+ unfold base;apply Z.pow_pos_nonneg. unfold Z.lt;reflexivity.
+ unfold Z.le;intros H;discriminate H.
Qed.
Lemma lt_0_wwB : 0 < wwB.
- Proof. rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_lt_0_compat;apply lt_0_wB. Qed.
+ Proof. rewrite wwB_wBwB; rewrite Z.pow_2_r; apply Z.mul_pos_pos;apply lt_0_wB. Qed.
Lemma wB_pos: 1 < wB.
Proof.
- unfold base;apply Zlt_le_trans with (2^1). unfold Zlt;reflexivity.
- apply Zpower_le_monotone. unfold Zlt;reflexivity.
- split;unfold Zle;intros H. discriminate H.
+ unfold base;apply Z.lt_le_trans with (2^1). unfold Z.lt;reflexivity.
+ apply Zpower_le_monotone. unfold Z.lt;reflexivity.
+ split;unfold Z.le;intros H. discriminate H.
clear spec_w_0W w_0W spec_w_Bm1 spec_to_Z spec_w_WW w_WW.
destruct w_digits; discriminate H.
Qed.
Lemma wwB_pos: 1 < wwB.
Proof.
- assert (H:= wB_pos);rewrite wwB_wBwB;rewrite <-(Zmult_1_r 1).
- rewrite Zpower_2.
- apply Zmult_lt_compat2;(split;[unfold Zlt;reflexivity|trivial]).
- apply Zlt_le_weak;trivial.
+ assert (H:= wB_pos);rewrite wwB_wBwB;rewrite <-(Z.mul_1_r 1).
+ rewrite Z.pow_2_r.
+ apply Zmult_lt_compat2;(split;[unfold Z.lt;reflexivity|trivial]).
+ apply Z.lt_le_incl;trivial.
Qed.
Theorem wB_div_2: 2 * (wB / 2) = wB.
@@ -208,22 +208,22 @@ Section DoubleBase.
clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W
spec_to_Z;unfold base.
assert (2 ^ Zpos w_digits = 2 * (2 ^ (Zpos w_digits - 1))).
- pattern 2 at 2; rewrite <- Zpower_1_r.
+ pattern 2 at 2; rewrite <- Z.pow_1_r.
rewrite <- Zpower_exp; auto with zarith.
f_equal; auto with zarith.
case w_digits; compute; intros; discriminate.
rewrite H; f_equal; auto with zarith.
- rewrite Zmult_comm; apply Z_div_mult; auto with zarith.
+ rewrite Z.mul_comm; apply Z_div_mult; auto with zarith.
Qed.
Theorem wwB_div_2 : wwB / 2 = wB / 2 * wB.
Proof.
clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W
spec_to_Z.
- rewrite wwB_wBwB; rewrite Zpower_2.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
pattern wB at 1; rewrite <- wB_div_2; auto.
- rewrite <- Zmult_assoc.
- repeat (rewrite (Zmult_comm 2); rewrite Z_div_mult); auto with zarith.
+ rewrite <- Z.mul_assoc.
+ repeat (rewrite (Z.mul_comm 2); rewrite Z_div_mult); auto with zarith.
Qed.
Lemma mod_wwB : forall z x,
@@ -231,15 +231,15 @@ Section DoubleBase.
Proof.
intros z x.
rewrite Zplus_mod.
- pattern wwB at 1;rewrite wwB_wBwB; rewrite Zpower_2.
+ pattern wwB at 1;rewrite wwB_wBwB; rewrite Z.pow_2_r.
rewrite Zmult_mod_distr_r;try apply lt_0_wB.
rewrite (Zmod_small [|x|]).
apply Zmod_small;rewrite wwB_wBwB;apply beta_mult;try apply spec_to_Z.
- apply Z_mod_lt;apply Zlt_gt;apply lt_0_wB.
+ apply Z_mod_lt;apply Z.lt_gt;apply lt_0_wB.
destruct (spec_to_Z x);split;trivial.
change [|x|] with (0*wB+[|x|]). rewrite wwB_wBwB.
- rewrite Zpower_2;rewrite <- (Zplus_0_r (wB*wB));apply beta_lex_inv.
- apply lt_0_wB. apply spec_to_Z. split;[apply Zle_refl | apply lt_0_wB].
+ rewrite Z.pow_2_r;rewrite <- (Z.add_0_r (wB*wB));apply beta_lex_inv.
+ apply lt_0_wB. apply spec_to_Z. split;[apply Z.le_refl | apply lt_0_wB].
Qed.
Lemma wB_div : forall x y, ([|x|] * wB + [|y|]) / wB = [|x|].
@@ -265,29 +265,29 @@ Section DoubleBase.
clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
unfold base;apply Zpower_lt_monotone;auto with zarith.
assert (0 < Zpos w_digits). compute;reflexivity.
- unfold ww_digits;rewrite Zpos_xO;auto with zarith.
+ unfold ww_digits;rewrite Pos2Z.inj_xO;auto with zarith.
Qed.
Lemma w_to_Z_wwB : forall x, x < wB -> x < wwB.
Proof.
- intros x H;apply Zlt_trans with wB;trivial;apply lt_wB_wwB.
+ intros x H;apply Z.lt_trans with wB;trivial;apply lt_wB_wwB.
Qed.
Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB.
Proof.
clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
destruct x as [ |h l];simpl.
- split;[apply Zle_refl|apply lt_0_wwB].
+ split;[apply Z.le_refl|apply lt_0_wwB].
assert (H:=spec_to_Z h);assert (L:=spec_to_Z l);split.
- apply Zplus_le_0_compat;auto with zarith.
- rewrite <- (Zplus_0_r wwB);rewrite wwB_wBwB; rewrite Zpower_2;
+ apply Z.add_nonneg_nonneg;auto with zarith.
+ rewrite <- (Z.add_0_r wwB);rewrite wwB_wBwB; rewrite Z.pow_2_r;
apply beta_lex_inv;auto with zarith.
Qed.
Lemma double_wB_wwB : forall n, double_wB n * double_wB n = double_wB (S n).
Proof.
intros n;unfold double_wB;simpl.
- unfold base. rewrite Pshiftl_nat_S, (Zpos_xO (_ << _)).
+ unfold base. rewrite Pshiftl_nat_S, (Pos2Z.inj_xO (_ << _)).
replace (2 * Zpos (w_digits << n)) with
(Zpos (w_digits << n) + Zpos (w_digits << n)) by ring.
symmetry; apply Zpower_exp;intro;discriminate.
@@ -306,14 +306,14 @@ Section DoubleBase.
intros n; elim n; clear n; auto.
unfold double_wB, "<<"; auto with zarith.
intros n H1; rewrite <- double_wB_wwB.
- apply Zle_trans with (wB * 1).
- rewrite Zmult_1_r; apply Zle_refl.
- apply Zmult_le_compat; auto with zarith.
- apply Zle_trans with wB; auto with zarith.
- unfold base.
- rewrite <- (Zpower_0_r 2).
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.le_trans with (wB * 1).
+ rewrite Z.mul_1_r; apply Z.le_refl.
unfold base; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
+ apply Z.le_trans with wB; auto with zarith.
+ unfold base.
+ rewrite <- (Z.pow_0_r 2).
+ apply Z.pow_le_mono_r; auto with zarith.
Qed.
Lemma spec_double_to_Z :
@@ -326,9 +326,9 @@ Section DoubleBase.
unfold double_wB,base;split;auto with zarith.
assert (U0:= IHn w0);assert (U1:= IHn w1).
split;auto with zarith.
- apply Zlt_le_trans with ((double_wB n - 1) * double_wB n + double_wB n).
+ apply Z.lt_le_trans with ((double_wB n - 1) * double_wB n + double_wB n).
assert (double_to_Z n w0*double_wB n <= (double_wB n - 1)*double_wB n).
- apply Zmult_le_compat_r;auto with zarith.
+ apply Z.mul_le_mono_nonneg_r;auto with zarith.
auto with zarith.
rewrite <- double_wB_wwB.
replace ((double_wB n - 1) * double_wB n + double_wB n) with (double_wB n * double_wB n);
@@ -342,22 +342,19 @@ Section DoubleBase.
clear spec_w_1 spec_w_Bm1.
intros n; elim n; auto; clear n.
intros n Hrec x; case x; clear x; auto.
- intros xx yy H1; simpl in H1.
- assert (F1: [!n | xx!] = 0).
- case (Zle_lt_or_eq 0 ([!n | xx!])); auto.
- case (spec_double_to_Z n xx); auto.
- intros F2.
- assert (F3 := double_wB_more_digits n).
- assert (F4: 0 <= [!n | yy!]).
- case (spec_double_to_Z n yy); auto.
+ intros xx yy; simpl.
+ destruct (spec_double_to_Z n xx) as [F1 _]. Z.le_elim F1.
+ - (* 0 < [!n | xx!] *)
+ intros; exfalso.
+ assert (F3 := double_wB_more_digits n).
+ destruct (spec_double_to_Z n yy) as [F4 _].
assert (F5: 1 * wB <= [!n | xx!] * double_wB n);
auto with zarith.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
unfold base; auto with zarith.
- simpl get_low; simpl double_to_Z.
- generalize H1; clear H1.
- rewrite F1; rewrite Zmult_0_l; rewrite Zplus_0_l.
- intros H1; apply Hrec; auto.
+ - (* 0 = [!n | xx!] *)
+ rewrite <- F1; rewrite Z.mul_0_l, Z.add_0_l.
+ intros; apply Hrec; auto.
Qed.
Lemma spec_double_WW : forall n (h l : word w n),
@@ -399,36 +396,36 @@ Section DoubleBase.
Ltac comp2ord := match goal with
| |- Lt = (?x ?= ?y) => symmetry; change (x < y)
- | |- Gt = (?x ?= ?y) => symmetry; change (x > y); apply Zlt_gt
+ | |- Gt = (?x ?= ?y) => symmetry; change (x > y); apply Z.lt_gt
end.
Lemma spec_ww_compare : forall x y,
- ww_compare x y = Zcompare [[x]] [[y]].
+ ww_compare x y = Z.compare [[x]] [[y]].
Proof.
destruct x as [ |xh xl];destruct y as [ |yh yl];simpl;trivial.
(* 1st case *)
rewrite 2 spec_w_compare, spec_w_0.
- destruct (Zcompare_spec 0 [|yh|]) as [H|H|H].
+ destruct (Z.compare_spec 0 [|yh|]) as [H|H|H].
rewrite <- H;simpl. reflexivity.
symmetry. change (0 < [|yh|]*wB+[|yl|]).
change 0 with (0*wB+0). rewrite <- spec_w_0 at 2.
apply wB_lex_inv;trivial.
- absurd (0 <= [|yh|]). apply Zlt_not_le; trivial.
+ absurd (0 <= [|yh|]). apply Z.lt_nge; trivial.
destruct (spec_to_Z yh);trivial.
(* 2nd case *)
rewrite 2 spec_w_compare, spec_w_0.
- destruct (Zcompare_spec [|xh|] 0) as [H|H|H].
+ destruct (Z.compare_spec [|xh|] 0) as [H|H|H].
rewrite H;simpl;reflexivity.
- absurd (0 <= [|xh|]). apply Zlt_not_le; trivial.
+ absurd (0 <= [|xh|]). apply Z.lt_nge; trivial.
destruct (spec_to_Z xh);trivial.
comp2ord.
change 0 with (0*wB+0). rewrite <- spec_w_0 at 2.
apply wB_lex_inv;trivial.
(* 3rd case *)
rewrite 2 spec_w_compare.
- destruct (Zcompare_spec [|xh|] [|yh|]) as [H|H|H].
+ destruct (Z.compare_spec [|xh|] [|yh|]) as [H|H|H].
rewrite H.
- symmetry. apply Zcompare_plus_compat.
+ symmetry. apply Z.add_compare_mono_l.
comp2ord. apply wB_lex_inv;trivial.
comp2ord. apply wB_lex_inv;trivial.
Qed.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
index 00a840520..55ecefa15 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
@@ -390,21 +390,21 @@ Section Z_2nZ.
Proof. refine (spec_ww_to_Z w_digits w_to_Z _);auto. Qed.
Let spec_ww_of_pos : forall p,
- Zpos p = (Z_of_N (fst (ww_of_pos p)))*wwB + [|(snd (ww_of_pos p))|].
+ Zpos p = (Z.of_N (fst (ww_of_pos p)))*wwB + [|(snd (ww_of_pos p))|].
Proof.
unfold ww_of_pos;intros.
rewrite (ZnZ.spec_of_pos p). unfold w_of_pos.
case (ZnZ.of_pos p); intros. simpl.
destruct n; simpl ZnZ.to_Z.
simpl;unfold w_to_Z,w_0; rewrite ZnZ.spec_0;trivial.
- unfold Z_of_N.
+ unfold Z.of_N.
rewrite (ZnZ.spec_of_pos p0).
case (ZnZ.of_pos p0); intros. simpl.
- unfold fst, snd,Z_of_N, to_Z, wB, w_digits, w_to_Z, w_WW.
+ unfold fst, snd,Z.of_N, to_Z, wB, w_digits, w_to_Z, w_WW.
rewrite ZnZ.spec_WW.
replace wwB with (wB*wB).
unfold wB,w_to_Z,w_digits;destruct n;ring.
- symmetry. rewrite <- Zpower_2; exact (wwB_wBwB w_digits).
+ symmetry. rewrite <- Z.pow_2_r; exact (wwB_wBwB w_digits).
Qed.
Let spec_ww_0 : [|W0|] = 0.
@@ -417,7 +417,7 @@ Section Z_2nZ.
Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed.
Let spec_ww_compare :
- forall x y, compare x y = Zcompare [|x|] [|y|].
+ forall x y, compare x y = Z.compare [|x|] [|y|].
Proof.
refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);auto.
Qed.
@@ -575,9 +575,9 @@ Section Z_2nZ.
unfold w_add_c; case ZnZ.add_c; unfold interp_carry; simpl ww_to_Z.
intros w0 Hw0; simpl; unfold w_to_Z; rewrite Hw0.
unfold w_0; rewrite ZnZ.spec_0; simpl; auto with zarith.
- intros w0; rewrite Zmult_1_l; simpl.
+ intros w0; rewrite Z.mul_1_l; simpl.
unfold w_to_Z, w_1; rewrite ZnZ.spec_1; auto with zarith.
- rewrite Zmult_1_l; auto.
+ rewrite Z.mul_1_l; auto.
Qed.
Let spec_low: forall x,
@@ -585,7 +585,7 @@ Section Z_2nZ.
intros x; case x; simpl low.
unfold ww_to_Z, w_to_Z, w_0; rewrite ZnZ.spec_0; simpl; auto.
intros xh xl; simpl.
- rewrite Zplus_comm; rewrite Z_mod_plus; auto with zarith.
+ rewrite Z.add_comm; rewrite Z_mod_plus; auto with zarith.
rewrite Zmod_small; auto with zarith.
unfold wB, base; auto with zarith.
Qed.
@@ -597,7 +597,7 @@ Section Z_2nZ.
rewrite spec_add2.
unfold w_to_Z, w_zdigits, w_digits.
rewrite ZnZ.spec_zdigits; auto.
- rewrite Zpos_xO; auto with zarith.
+ rewrite Pos2Z.inj_xO; auto with zarith.
Qed.
@@ -605,7 +605,7 @@ Section Z_2nZ.
Proof.
refine (spec_ww_head00 w_0 w_0W
w_compare w_head0 w_add2 w_zdigits _ww_zdigits
- w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); auto.
+ w_to_Z _ _ _ (eq_refl _ww_digits) _ _ _ _); auto.
exact ZnZ.spec_head00.
exact ZnZ.spec_zdigits.
Qed.
@@ -623,7 +623,7 @@ Section Z_2nZ.
Proof.
refine (spec_ww_tail00 w_0 w_0W
w_compare w_tail0 w_add2 w_zdigits _ww_zdigits
- w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); wwauto.
+ w_to_Z _ _ _ (eq_refl _ww_digits) _ _ _ _); wwauto.
exact ZnZ.spec_tail00.
exact ZnZ.spec_zdigits.
Qed.
@@ -749,7 +749,7 @@ refine
| false => [|x|] mod 2 = 1
end.
Proof.
- refine (@spec_ww_is_even t w_is_even w_0 w_1 w_Bm1 w_digits _ _ _ _ _); auto.
+ refine (@spec_ww_is_even t w_is_even w_digits _ _ ).
exact ZnZ.spec_is_even.
Qed.
@@ -798,7 +798,7 @@ refine
exact ZnZ.spec_zdigits.
unfold w_to_Z, w_zdigits.
rewrite ZnZ.spec_zdigits.
- rewrite <- Zpos_xO; exact spec_ww_digits.
+ rewrite <- Pos2Z.inj_xO; exact spec_ww_digits.
Qed.
Global Instance mk_zn2z_specs_karatsuba : ZnZ.Specs mk_zn2z_ops_karatsuba.
@@ -811,7 +811,7 @@ refine
exact ZnZ.spec_zdigits.
unfold w_to_Z, w_zdigits.
rewrite ZnZ.spec_zdigits.
- rewrite <- Zpos_xO; exact spec_ww_digits.
+ rewrite <- Pos2Z.inj_xO; exact spec_ww_digits.
Qed.
End Z_2nZ.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
index 0cb6848e3..40bce95b4 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
@@ -80,7 +80,7 @@ Section POS_MOD.
Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
Variable spec_ww_compare : forall x y,
- ww_compare x y = Zcompare [[x]] [[y]].
+ ww_compare x y = Z.compare [[x]] [[y]].
Variable spec_ww_sub: forall x y,
[[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
@@ -100,7 +100,7 @@ Section POS_MOD.
unfold ww_pos_mod; case w1.
simpl; rewrite Zmod_small; split; auto with zarith.
intros xh xl; rewrite spec_ww_compare.
- case Zcompare_spec;
+ case Z.compare_spec;
rewrite spec_w_0W; rewrite spec_zdigits; fold wB;
intros H1.
rewrite H1; simpl ww_to_Z.
@@ -117,19 +117,19 @@ Section POS_MOD.
rewrite spec_low.
apply Zmod_small; auto with zarith.
case (spec_to_w_Z p); intros HHH1 HHH2; split; auto with zarith.
- apply Zlt_le_trans with (1 := H1).
+ apply Z.lt_le_trans with (1 := H1).
unfold base; apply Zpower2_le_lin; auto with zarith.
rewrite HH0.
rewrite Zplus_mod; auto with zarith.
unfold base.
rewrite <- (F0 (Zpos w_digits) [[p]]).
rewrite Zpower_exp; auto with zarith.
- rewrite Zmult_assoc.
+ rewrite Z.mul_assoc.
rewrite Z_mod_mult; auto with zarith.
autorewrite with w_rewrite rm10.
rewrite Zmod_mod; auto with zarith.
rewrite spec_ww_compare.
- case Zcompare_spec; rewrite spec_ww_zdigits;
+ case Z.compare_spec; rewrite spec_ww_zdigits;
rewrite spec_zdigits; intros H2.
replace (2^[[p]]) with wwB.
rewrite Zmod_small; auto with zarith.
@@ -143,52 +143,52 @@ Section POS_MOD.
rewrite <- Zmod_div_mod; auto with zarith.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
- apply Zlt_le_trans with (Zpos w_digits); auto with zarith.
+ apply Z.lt_le_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
exists wB; unfold base; rewrite <- Zpower_exp; auto with zarith.
rewrite spec_ww_digits;
- apply f_equal with (f := Zpower 2); rewrite Zpos_xO; auto with zarith.
+ apply f_equal with (f := Z.pow 2); rewrite Pos2Z.inj_xO; auto with zarith.
simpl ww_to_Z; autorewrite with w_rewrite.
rewrite spec_pos_mod; rewrite HH0.
pattern [|xh|] at 2;
rewrite Z_div_mod_eq with (b := 2 ^ ([[p]] - Zpos w_digits));
auto with zarith.
- rewrite (fun x => (Zmult_comm (2 ^ x))); rewrite Zmult_plus_distr_l.
- unfold base; rewrite <- Zmult_assoc; rewrite <- Zpower_exp;
+ rewrite (fun x => (Z.mul_comm (2 ^ x))); rewrite Z.mul_add_distr_r.
+ unfold base; rewrite <- Z.mul_assoc; rewrite <- Zpower_exp;
auto with zarith.
rewrite F0; auto with zarith.
- rewrite <- Zplus_assoc; rewrite Zplus_mod; auto with zarith.
+ rewrite <- Z.add_assoc; rewrite Zplus_mod; auto with zarith.
rewrite Z_mod_mult; auto with zarith.
autorewrite with rm10.
rewrite Zmod_mod; auto with zarith.
- apply sym_equal; apply Zmod_small; auto with zarith.
+ symmetry; apply Zmod_small; auto with zarith.
case (spec_to_Z xh); intros U1 U2.
case (spec_to_Z xl); intros U3 U4.
split; auto with zarith.
- apply Zplus_le_0_compat; auto with zarith.
- apply Zmult_le_0_compat; auto with zarith.
+ apply Z.add_nonneg_nonneg; auto with zarith.
+ apply Z.mul_nonneg_nonneg; auto with zarith.
match goal with |- 0 <= ?X mod ?Y =>
case (Z_mod_lt X Y); auto with zarith
end.
match goal with |- ?X mod ?Y * ?U + ?Z < ?T =>
- apply Zle_lt_trans with ((Y - 1) * U + Z );
+ apply Z.le_lt_trans with ((Y - 1) * U + Z );
[case (Z_mod_lt X Y); auto with zarith | idtac]
end.
match goal with |- ?X * ?U + ?Y < ?Z =>
- apply Zle_lt_trans with (X * U + (U - 1))
+ apply Z.le_lt_trans with (X * U + (U - 1))
end.
- apply Zplus_le_compat_l; auto with zarith.
+ apply Z.add_le_mono_l; auto with zarith.
case (spec_to_Z xl); unfold base; auto with zarith.
- rewrite Zmult_minus_distr_r; rewrite <- Zpower_exp; auto with zarith.
+ rewrite Z.mul_sub_distr_r; rewrite <- Zpower_exp; auto with zarith.
rewrite F0; auto with zarith.
rewrite Zmod_small; auto with zarith.
case (spec_to_w_Z (WW xh xl)); intros U1 U2.
split; auto with zarith.
- apply Zlt_le_trans with (1:= U2).
+ apply Z.lt_le_trans with (1:= U2).
unfold base; rewrite spec_ww_digits.
apply Zpower_le_monotone; auto with zarith.
split; auto with zarith.
- rewrite Zpos_xO; auto with zarith.
+ rewrite Pos2Z.inj_xO; auto with zarith.
Qed.
End POS_MOD.
@@ -260,7 +260,7 @@ Section DoubleDiv32.
Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
Variable spec_compare :
- forall x y, w_compare x y = Zcompare [|x|] [|y|].
+ forall x y, w_compare x y = Z.compare [|x|] [|y|].
Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
Variable spec_w_add_carry_c :
forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1.
@@ -290,14 +290,14 @@ Section DoubleDiv32.
assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x).
Theorem wB_div2: forall x, wB/2 <= x -> wB <= 2 * x.
- intros x H; rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith.
+ intros x H; rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; auto with zarith.
Qed.
Lemma Zmult_lt_0_reg_r_2 : forall n m : Z, 0 <= n -> 0 < m * n -> 0 < m.
Proof.
- intros n m H1 H2;apply Zmult_lt_0_reg_r with n;trivial.
- destruct (Zle_lt_or_eq _ _ H1);trivial.
- subst;rewrite Zmult_0_r in H2;discriminate H2.
+ intros n m H1 H2;apply Z.mul_pos_cancel_r with n;trivial.
+ Z.le_elim H1; trivial.
+ subst;rewrite Z.mul_0_r in H2;discriminate H2.
Qed.
Theorem spec_w_div32 : forall a1 a2 a3 b1 b2,
@@ -311,7 +311,7 @@ Section DoubleDiv32.
intros a1 a2 a3 b1 b2 Hle Hlt.
assert (U:= lt_0_wB w_digits); assert (U1:= lt_0_wwB w_digits).
Spec_w_to_Z a1;Spec_w_to_Z a2;Spec_w_to_Z a3;Spec_w_to_Z b1;Spec_w_to_Z b2.
- rewrite wwB_wBwB; rewrite Zpower_2; rewrite Zmult_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r; rewrite Z.mul_assoc;rewrite <- Z.mul_add_distr_r.
change (w_div32 a1 a2 a3 b1 b2) with
match w_compare a1 b1 with
| Lt =>
@@ -332,7 +332,7 @@ Section DoubleDiv32.
(WW (w_sub a2 b2) a3) (WW b1 b2)
| Gt => (w_0, W0) (* cas absurde *)
end.
- rewrite spec_compare. case Zcompare_spec; intro Hcmp.
+ rewrite spec_compare. case Z.compare_spec; intro Hcmp.
simpl in Hlt.
rewrite Hcmp in Hlt;assert ([|a2|] < [|b2|]). omega.
assert ([[WW (w_sub a2 b2) a3]] = ([|a2|]-[|b2|])*wB + [|a3|] + wwB).
@@ -351,17 +351,17 @@ Section DoubleDiv32.
rewrite H0;intros r.
repeat
(rewrite spec_ww_add;eauto || rewrite spec_w_Bm1 || rewrite spec_w_Bm2);
- simpl ww_to_Z;try rewrite Zmult_1_l;intros H1.
+ simpl ww_to_Z;try rewrite Z.mul_1_l;intros H1.
assert (0<= ([[r]] + ([|b1|] * wB + [|b2|])) - wwB < [|b1|] * wB + [|b2|]).
Spec_ww_to_Z r;split;zarith.
rewrite H1.
assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB).
- rewrite wwB_wBwB; rewrite Zpower_2; zarith.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r; zarith.
assert (-wwB < ([|a2|] - [|b2|]) * wB + [|a3|] < 0).
- split. apply Zlt_le_trans with (([|a2|] - [|b2|]) * wB);zarith.
+ split. apply Z.lt_le_trans with (([|a2|] - [|b2|]) * wB);zarith.
rewrite wwB_wBwB;replace (-(wB^2)) with (-wB*wB);[zarith | ring].
- apply Zmult_lt_compat_r;zarith.
- apply Zle_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith.
+ apply Z.mul_lt_mono_pos_r;zarith.
+ apply Z.le_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith.
replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with
(([|a2|] - [|b2|] + 1) * wB + - 1);[zarith | ring].
assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith.
@@ -376,13 +376,13 @@ Section DoubleDiv32.
Spec_ww_to_Z (WW b1 b2). simpl in HH4;zarith.
rewrite H0;intros r;repeat
(rewrite spec_w_Bm1 || rewrite spec_w_Bm2);
- simpl ww_to_Z;try rewrite Zmult_1_l;intros H1.
+ simpl ww_to_Z;try rewrite Z.mul_1_l;intros H1.
assert ([[r]]=([|a2|]-[|b2|])*wB+[|a3|]+([|b1|]*wB+[|b2|])). zarith.
split. rewrite H2;rewrite Hcmp;ring.
split. Spec_ww_to_Z r;zarith.
rewrite H2.
assert (([|a2|] - [|b2|]) * wB + [|a3|] < 0);zarith.
- apply Zle_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith.
+ apply Z.le_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith.
replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with
(([|a2|] - [|b2|] + 1) * wB + - 1);[zarith|ring].
assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith.
@@ -400,7 +400,7 @@ Section DoubleDiv32.
rewrite H1.
split. ring. split.
rewrite <- H1;destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z r1);trivial.
- apply Zle_lt_trans with ([|r|] * wB + [|a3|]).
+ apply Z.le_lt_trans with ([|r|] * wB + [|a3|]).
assert ( 0 <= [|q|] * [|b2|]);zarith.
apply beta_lex_inv;zarith.
assert ([[r1]] = [|r|] * wB + [|a3|] - [|q|] * [|b2|] + wwB).
@@ -418,10 +418,10 @@ Section DoubleDiv32.
intros r2;repeat (rewrite spec_pred || rewrite spec_ww_add;eauto);
simpl ww_to_Z;intros H7.
assert (0 < [|q|] - 1).
- assert (1 <= [|q|]). zarith.
- destruct (Zle_lt_or_eq _ _ H6);zarith.
- rewrite <- H8 in H2;rewrite H2 in H7.
- assert (0 < [|b1|]*wB). apply Zmult_lt_0_compat;zarith.
+ assert (H6 : 1 <= [|q|]) by zarith.
+ Z.le_elim H6;zarith.
+ rewrite <- H6 in H2;rewrite H2 in H7.
+ assert (0 < [|b1|]*wB). apply Z.mul_pos_pos;zarith.
Spec_ww_to_Z r2. zarith.
rewrite (Zmod_small ([|q|] -1));zarith.
rewrite (Zmod_small ([|q|] -1 -1));zarith.
@@ -439,7 +439,7 @@ Section DoubleDiv32.
< wwB). split;try omega.
replace (2*([|b1|]*wB+[|b2|])) with ((2*[|b1|])*wB+2*[|b2|]). 2:ring.
assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB).
- rewrite wwB_wBwB; rewrite Zpower_2; zarith. omega.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r; zarith. omega.
rewrite <- (Zmod_unique
([[r2]] + ([|b1|] * wB + [|b2|]))
wwB
@@ -534,13 +534,13 @@ Section DoubleDiv21.
0 <= [[r]] < [|b1|] * wB + [|b2|].
Variable spec_ww_1 : [[ww_1]] = 1.
Variable spec_ww_compare : forall x y,
- ww_compare x y = Zcompare [[x]] [[y]].
+ ww_compare x y = Z.compare [[x]] [[y]].
Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
Theorem wwB_div: wwB = 2 * (wwB / 2).
Proof.
- rewrite wwB_div_2; rewrite Zmult_assoc; rewrite wB_div_2; auto.
- rewrite <- Zpower_2; apply wwB_wBwB.
+ rewrite wwB_div_2; rewrite Z.mul_assoc; rewrite wB_div_2; auto.
+ rewrite <- Z.pow_2_r; apply wwB_wBwB.
Qed.
Ltac Spec_w_to_Z x :=
@@ -562,7 +562,7 @@ Section DoubleDiv21.
Spec_ww_to_Z b; assert (Eq: 0 < [[b]]). Spec_ww_to_Z a1;omega.
generalize Hlt H ;clear Hlt H;case a1.
intros H1 H2;simpl in H1;Spec_ww_to_Z a2.
- rewrite spec_ww_compare. case Zcompare_spec;
+ rewrite spec_ww_compare. case Z.compare_spec;
simpl;try rewrite spec_ww_1;autorewrite with rm10; intros;zarith.
rewrite spec_ww_sub;simpl. rewrite Zmod_small;zarith.
split. ring.
@@ -570,32 +570,32 @@ Section DoubleDiv21.
rewrite wwB_div;zarith.
intros a1h a1l. Spec_w_to_Z a1h;Spec_w_to_Z a1l. Spec_ww_to_Z a2.
destruct a2 as [ |a3 a4];
- (destruct b as [ |b1 b2];[unfold Zle in Eq;discriminate Eq|idtac]);
+ (destruct b as [ |b1 b2];[unfold Z.le in Eq;discriminate Eq|idtac]);
try (Spec_w_to_Z a3; Spec_w_to_Z a4); Spec_w_to_Z b1; Spec_w_to_Z b2;
intros Hlt H; match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] =>
generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U);
intros q1 r H0
end; (assert (Eq1: wB / 2 <= [|b1|]);[
apply (@beta_lex (wB / 2) 0 [|b1|] [|b2|] wB); auto with zarith;
- autorewrite with rm10;repeat rewrite (Zmult_comm wB);
+ autorewrite with rm10;repeat rewrite (Z.mul_comm wB);
rewrite <- wwB_div_2; trivial
| generalize (H0 Eq1 Hlt);clear H0;destruct r as [ |r1 r2];simpl;
- try rewrite spec_w_0; try rewrite spec_w_0W;repeat rewrite Zplus_0_r;
+ try rewrite spec_w_0; try rewrite spec_w_0W;repeat rewrite Z.add_0_r;
intros (H1,H2) ]).
- split;[rewrite wwB_wBwB; rewrite Zpower_2 | trivial].
- rewrite Zmult_assoc;rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;
- rewrite <- Zpower_2; rewrite <- wwB_wBwB;rewrite H1;ring.
+ split;[rewrite wwB_wBwB; rewrite Z.pow_2_r | trivial].
+ rewrite Z.mul_assoc;rewrite Z.mul_add_distr_r;rewrite <- Z.mul_assoc;
+ rewrite <- Z.pow_2_r; rewrite <- wwB_wBwB;rewrite H1;ring.
destruct H2 as (H2,H3);match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] =>
generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U);
intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end.
split;[rewrite wwB_wBwB | trivial].
- rewrite Zpower_2.
- rewrite Zmult_assoc;rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;
- rewrite <- Zpower_2.
+ rewrite Z.pow_2_r.
+ rewrite Z.mul_assoc;rewrite Z.mul_add_distr_r;rewrite <- Z.mul_assoc;
+ rewrite <- Z.pow_2_r.
rewrite <- wwB_wBwB;rewrite H1.
- rewrite spec_w_0 in H4;rewrite Zplus_0_r in H4.
- repeat rewrite Zmult_plus_distr_l. rewrite <- (Zmult_assoc [|r1|]).
- rewrite <- Zpower_2; rewrite <- wwB_wBwB;rewrite H4;simpl;ring.
+ rewrite spec_w_0 in H4;rewrite Z.add_0_r in H4.
+ repeat rewrite Z.mul_add_distr_r. rewrite <- (Z.mul_assoc [|r1|]).
+ rewrite <- Z.pow_2_r; rewrite <- wwB_wBwB;rewrite H4;simpl;ring.
split;[rewrite wwB_wBwB | split;zarith].
replace (([|a1h|] * wB + [|a1l|]) * wB^2 + ([|a3|] * wB + [|a4|]))
with (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB+ [|a4|]).
@@ -793,7 +793,7 @@ Section DoubleDivGt.
Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|].
Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
Variable spec_compare :
- forall x y, w_compare x y = Zcompare [|x|] [|y|].
+ forall x y, w_compare x y = Z.compare [|x|] [|y|].
Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0.
Variable spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|].
@@ -893,42 +893,42 @@ Section DoubleDivGt.
end in [[WW ah al]]=[[q]]*[[WW bh bl]]+[[r]] /\ 0 <=[[r]]< [[WW bh bl]]).
assert (Hh := spec_head0 Hpos).
lazy zeta.
- rewrite spec_compare; case Zcompare_spec;
+ rewrite spec_compare; case Z.compare_spec;
rewrite spec_w_0; intros HH.
- generalize Hh; rewrite HH; simpl Zpower;
- rewrite Zmult_1_l; intros (HH1, HH2); clear HH.
+ generalize Hh; rewrite HH; simpl Z.pow;
+ rewrite Z.mul_1_l; intros (HH1, HH2); clear HH.
assert (wwB <= 2*[[WW bh bl]]).
- apply Zle_trans with (2*[|bh|]*wB).
- rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat_r; zarith.
- rewrite <- wB_div_2; apply Zmult_le_compat_l; zarith.
- simpl ww_to_Z;rewrite Zmult_plus_distr_r;rewrite Zmult_assoc.
+ apply Z.le_trans with (2*[|bh|]*wB).
+ rewrite wwB_wBwB; rewrite Z.pow_2_r; apply Z.mul_le_mono_nonneg_r; zarith.
+ rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; zarith.
+ simpl ww_to_Z;rewrite Z.mul_add_distr_l;rewrite Z.mul_assoc.
Spec_w_to_Z bl;zarith.
Spec_ww_to_Z (WW ah al).
rewrite spec_ww_sub;eauto.
- simpl;rewrite spec_ww_1;rewrite Zmult_1_l;simpl.
+ simpl;rewrite spec_ww_1;rewrite Z.mul_1_l;simpl.
simpl ww_to_Z in Hgt, H, HH;rewrite Zmod_small;split;zarith.
case (spec_to_Z (w_head0 bh)); auto with zarith.
assert ([|w_head0 bh|] < Zpos w_digits).
destruct (Z_lt_ge_dec [|w_head0 bh|] (Zpos w_digits));trivial.
exfalso.
assert (2 ^ [|w_head0 bh|] * [|bh|] >= wB);auto with zarith.
- apply Zle_ge; replace wB with (wB * 1);try ring.
- Spec_w_to_Z bh;apply Zmult_le_compat;zarith.
+ apply Z.le_ge; replace wB with (wB * 1);try ring.
+ Spec_w_to_Z bh;apply Z.mul_le_mono_nonneg;zarith.
unfold base;apply Zpower_le_monotone;zarith.
assert (HHHH : 0 < [|w_head0 bh|] < Zpos w_digits); auto with zarith.
- assert (Hb:= Zlt_le_weak _ _ H).
+ assert (Hb:= Z.lt_le_incl _ _ H).
generalize (spec_add_mul_div w_0 ah Hb)
(spec_add_mul_div ah al Hb)
(spec_add_mul_div al w_0 Hb)
(spec_add_mul_div bh bl Hb)
(spec_add_mul_div bl w_0 Hb);
- rewrite spec_w_0; repeat rewrite Zmult_0_l;repeat rewrite Zplus_0_l;
- rewrite Zdiv_0_l;repeat rewrite Zplus_0_r.
+ rewrite spec_w_0; repeat rewrite Z.mul_0_l;repeat rewrite Z.add_0_l;
+ rewrite Zdiv_0_l;repeat rewrite Z.add_0_r.
Spec_w_to_Z ah;Spec_w_to_Z bh.
unfold base;repeat rewrite Zmod_shift_r;zarith.
assert (H3:=to_Z_div_minus_p ah HHHH);assert(H4:=to_Z_div_minus_p al HHHH);
assert (H5:=to_Z_div_minus_p bl HHHH).
- rewrite Zmult_comm in Hh.
+ rewrite Z.mul_comm in Hh.
assert (2^[|w_head0 bh|] < wB). unfold base;apply Zpower_lt_monotone;zarith.
unfold base in H0;rewrite Zmod_small;zarith.
fold wB; rewrite (Zmod_small ([|bh|] * 2 ^ [|w_head0 bh|]));zarith.
@@ -943,15 +943,15 @@ Section DoubleDivGt.
(w_add_mul_div (w_head0 bh) al w_0)
(w_add_mul_div (w_head0 bh) bh bl)
(w_add_mul_div (w_head0 bh) bl w_0)) as (q,r).
- rewrite V1;rewrite V2. rewrite Zmult_plus_distr_l.
- rewrite <- (Zplus_assoc ([|bh|] * 2 ^ [|w_head0 bh|] * wB)).
+ rewrite V1;rewrite V2. rewrite Z.mul_add_distr_r.
+ rewrite <- (Z.add_assoc ([|bh|] * 2 ^ [|w_head0 bh|] * wB)).
unfold base;rewrite <- shift_unshift_mod;zarith. fold wB.
replace ([|bh|] * 2 ^ [|w_head0 bh|] * wB + [|bl|] * 2 ^ [|w_head0 bh|]) with
([[WW bh bl]] * 2^[|w_head0 bh|]). 2:simpl;ring.
- fold wwB. rewrite wwB_wBwB. rewrite Zpower_2. rewrite U1;rewrite U2;rewrite U3.
- rewrite Zmult_assoc. rewrite Zmult_plus_distr_l.
- rewrite (Zplus_assoc ([|ah|] / 2^(Zpos(w_digits) - [|w_head0 bh|])*wB * wB)).
- rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc.
+ fold wwB. rewrite wwB_wBwB. rewrite Z.pow_2_r. rewrite U1;rewrite U2;rewrite U3.
+ rewrite Z.mul_assoc. rewrite Z.mul_add_distr_r.
+ rewrite (Z.add_assoc ([|ah|] / 2^(Zpos(w_digits) - [|w_head0 bh|])*wB * wB)).
+ rewrite <- Z.mul_add_distr_r. rewrite <- Z.add_assoc.
unfold base;repeat rewrite <- shift_unshift_mod;zarith. fold wB.
replace ([|ah|] * 2 ^ [|w_head0 bh|] * wB + [|al|] * 2 ^ [|w_head0 bh|]) with
([[WW ah al]] * 2^[|w_head0 bh|]). 2:simpl;ring.
@@ -962,42 +962,42 @@ Section DoubleDivGt.
unfold base.
replace (2^Zpos (w_digits)) with (2^(Zpos (w_digits) - 1)*2).
rewrite Z_div_mult;zarith. rewrite <- Zpower_exp;zarith.
- apply Zlt_le_trans with wB;zarith.
+ apply Z.lt_le_trans with wB;zarith.
unfold base;apply Zpower_le_monotone;zarith.
pattern 2 at 2;replace 2 with (2^1);trivial.
rewrite <- Zpower_exp;zarith. ring_simplify (Zpos (w_digits) - 1 + 1);trivial.
change [[WW w_0 q]] with ([|w_0|]*wB+[|q|]);rewrite spec_w_0;rewrite
- Zmult_0_l;rewrite Zplus_0_l.
+ Z.mul_0_l;rewrite Z.add_0_l.
replace [[ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry
_ww_zdigits (w_0W (w_head0 bh))) W0 r]] with ([[r]]/2^[|w_head0 bh|]).
- assert (0 < 2^[|w_head0 bh|]). apply Zpower_gt_0;zarith.
+ assert (0 < 2^[|w_head0 bh|]). apply Z.pow_pos_nonneg;zarith.
split.
rewrite <- (Z_div_mult [[WW ah al]] (2^[|w_head0 bh|]));zarith.
- rewrite H1;rewrite Zmult_assoc;apply Z_div_plus_l;trivial.
+ rewrite H1;rewrite Z.mul_assoc;apply Z_div_plus_l;trivial.
split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith.
rewrite spec_ww_add_mul_div.
rewrite spec_ww_sub; auto with zarith.
rewrite spec_ww_digits_.
change (Zpos (xO (w_digits))) with (2*Zpos (w_digits));zarith.
- simpl ww_to_Z;rewrite Zmult_0_l;rewrite Zplus_0_l.
+ simpl ww_to_Z;rewrite Z.mul_0_l;rewrite Z.add_0_l.
rewrite spec_w_0W.
rewrite (fun x y => Zmod_small (x-y)); auto with zarith.
ring_simplify (2 * Zpos w_digits - (2 * Zpos w_digits - [|w_head0 bh|])).
rewrite Zmod_small;zarith.
split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith.
Spec_ww_to_Z r.
- apply Zlt_le_trans with wwB;zarith.
- rewrite <- (Zmult_1_r wwB);apply Zmult_le_compat;zarith.
+ apply Z.lt_le_trans with wwB;zarith.
+ rewrite <- (Z.mul_1_r wwB);apply Z.mul_le_mono_nonneg;zarith.
split; auto with zarith.
- apply Zle_lt_trans with (2 * Zpos w_digits); auto with zarith.
- unfold base, ww_digits; rewrite (Zpos_xO w_digits).
+ apply Z.le_lt_trans with (2 * Zpos w_digits); auto with zarith.
+ unfold base, ww_digits; rewrite (Pos2Z.inj_xO w_digits).
apply Zpower2_lt_lin; auto with zarith.
rewrite spec_ww_sub; auto with zarith.
rewrite spec_ww_digits_; rewrite spec_w_0W.
rewrite Zmod_small;zarith.
- rewrite Zpos_xO; split; auto with zarith.
- apply Zle_lt_trans with (2 * Zpos w_digits); auto with zarith.
- unfold base, ww_digits; rewrite (Zpos_xO w_digits).
+ rewrite Pos2Z.inj_xO; split; auto with zarith.
+ apply Z.le_lt_trans with (2 * Zpos w_digits); auto with zarith.
+ unfold base, ww_digits; rewrite (Pos2Z.inj_xO w_digits).
apply Zpower2_lt_lin; auto with zarith.
Qed.
@@ -1037,9 +1037,9 @@ Section DoubleDivGt.
assert (H2:=spec_div_gt Hgt Hpos);destruct (w_div_gt al bl).
repeat rewrite spec_w_0W;simpl;rewrite spec_w_0;simpl;trivial.
clear H.
- rewrite spec_compare; case Zcompare_spec; intros Hcmp.
+ rewrite spec_compare; case Z.compare_spec; intros Hcmp.
rewrite spec_w_0 in Hcmp. change [[WW bh bl]] with ([|bh|]*wB+[|bl|]).
- rewrite <- Hcmp;rewrite Zmult_0_l;rewrite Zplus_0_l.
+ rewrite <- Hcmp;rewrite Z.mul_0_l;rewrite Z.add_0_l.
simpl in Hpos;rewrite <- Hcmp in Hpos;simpl in Hpos.
assert (H2:= @spec_double_divn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div
w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0
@@ -1079,7 +1079,7 @@ Section DoubleDivGt.
rewrite spec_mod_gt;trivial.
assert (H:=spec_div_gt Hgt Hpos).
destruct (w_div_gt a b) as (q,r);simpl.
- rewrite Zmult_comm in H;destruct H.
+ rewrite Z.mul_comm in H;destruct H.
symmetry;apply Zmod_unique with [|q|];trivial.
Qed.
@@ -1132,7 +1132,7 @@ Section DoubleDivGt.
rewrite spec_w_0W;rewrite spec_w_mod_gt_eq;trivial.
destruct (w_div_gt al bl);simpl;rewrite spec_w_0W;trivial.
clear H.
- rewrite spec_compare; case Zcompare_spec; intros H2.
+ rewrite spec_compare; case Z.compare_spec; intros H2.
rewrite (@spec_double_modn1_aux w w_zdigits w_0 w_WW w_head0 w_add_mul_div
w_div21 w_compare w_sub w_to_Z spec_w_0 spec_compare 1 (WW ah al) bl).
destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1
@@ -1149,7 +1149,7 @@ Section DoubleDivGt.
rewrite (spec_ww_mod_gt_eq a b Hgt Hpos).
destruct (ww_div_gt a b)as(q,r);destruct H.
apply Zmod_unique with[[q]];simpl;trivial.
- rewrite Zmult_comm;trivial.
+ rewrite Z.mul_comm;trivial.
Qed.
Lemma Zis_gcd_mod : forall a b d,
@@ -1206,13 +1206,13 @@ Section DoubleDivGt.
| Gt => W0 (* absurde *)
end).
rewrite spec_compare, spec_w_0.
- case Zcompare_spec; intros Hbh.
+ case Z.compare_spec; intros Hbh.
simpl ww_to_Z in *. rewrite <- Hbh.
- rewrite Zmult_0_l;rewrite Zplus_0_l.
+ rewrite Z.mul_0_l;rewrite Z.add_0_l.
rewrite spec_compare, spec_w_0.
- case Zcompare_spec; intros Hbl.
+ case Z.compare_spec; intros Hbl.
rewrite <- Hbl;apply Zis_gcd_0.
- simpl;rewrite spec_w_0;rewrite Zmult_0_l;rewrite Zplus_0_l.
+ simpl;rewrite spec_w_0;rewrite Z.mul_0_l;rewrite Z.add_0_l.
apply Zis_gcd_mod;zarith.
change ([|ah|] * wB + [|al|]) with (double_to_Z w_digits w_to_Z 1 (WW ah al)).
rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div
@@ -1220,19 +1220,19 @@ Section DoubleDivGt.
spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hbl).
apply spec_gcd_gt.
rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial.
- apply Zlt_gt;match goal with | |- ?x mod ?y < ?y =>
+ apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
destruct (Z_mod_lt x y);zarith end.
Spec_w_to_Z bl;exfalso;omega.
assert (H:= spec_ww_mod_gt_aux _ _ _ Hgt Hbh).
assert (H2 : 0 < [[WW bh bl]]).
- simpl;Spec_w_to_Z bl. apply Zlt_le_trans with ([|bh|]*wB);zarith.
- apply Zmult_lt_0_compat;zarith.
+ simpl;Spec_w_to_Z bl. apply Z.lt_le_trans with ([|bh|]*wB);zarith.
+ apply Z.mul_pos_pos;zarith.
apply Zis_gcd_mod;trivial. rewrite <- H.
simpl in *;destruct (ww_mod_gt_aux ah al bh bl) as [ |mh ml].
simpl;apply Zis_gcd_0;zarith.
- rewrite spec_compare, spec_w_0; case Zcompare_spec; intros Hmh.
+ rewrite spec_compare, spec_w_0; case Z.compare_spec; intros Hmh.
simpl;rewrite <- Hmh;simpl.
- rewrite spec_compare, spec_w_0; case Zcompare_spec; intros Hml.
+ rewrite spec_compare, spec_w_0; case Z.compare_spec; intros Hml.
rewrite <- Hml;simpl;apply Zis_gcd_0.
simpl; rewrite spec_w_0; simpl.
apply Zis_gcd_mod;zarith.
@@ -1242,38 +1242,38 @@ Section DoubleDivGt.
spec_div21 spec_compare spec_sub 1 (WW bh bl) ml Hml).
apply spec_gcd_gt.
rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial.
- apply Zlt_gt;match goal with | |- ?x mod ?y < ?y =>
+ apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
destruct (Z_mod_lt x y);zarith end.
Spec_w_to_Z ml;exfalso;omega.
assert ([[WW bh bl]] > [[WW mh ml]]).
- rewrite H;simpl; apply Zlt_gt;match goal with | |- ?x mod ?y < ?y =>
+ rewrite H;simpl; apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
destruct (Z_mod_lt x y);zarith end.
assert (H1:= spec_ww_mod_gt_aux _ _ _ H0 Hmh).
assert (H3 : 0 < [[WW mh ml]]).
- simpl;Spec_w_to_Z ml. apply Zlt_le_trans with ([|mh|]*wB);zarith.
- apply Zmult_lt_0_compat;zarith.
+ simpl;Spec_w_to_Z ml. apply Z.lt_le_trans with ([|mh|]*wB);zarith.
+ apply Z.mul_pos_pos;zarith.
apply Zis_gcd_mod;zarith. simpl in *;rewrite <- H1.
destruct (ww_mod_gt_aux bh bl mh ml) as [ |rh rl]. simpl; apply Zis_gcd_0.
simpl;apply Hcont. simpl in H1;rewrite H1.
- apply Zlt_gt;match goal with | |- ?x mod ?y < ?y =>
+ apply Z.lt_gt;match goal with | |- ?x mod ?y < ?y =>
destruct (Z_mod_lt x y);zarith end.
- apply Zle_trans with (2^n/2).
+ apply Z.le_trans with (2^n/2).
apply Zdiv_le_lower_bound;zarith.
- apply Zle_trans with ([|bh|] * wB + [|bl|]);zarith.
- assert (H3' := Z_div_mod_eq [[WW bh bl]] [[WW mh ml]] (Zlt_gt _ _ H3)).
- assert (H4' : 0 <= [[WW bh bl]]/[[WW mh ml]]).
- apply Zge_le;apply Z_div_ge0;zarith. simpl in *;rewrite H1.
+ apply Z.le_trans with ([|bh|] * wB + [|bl|]);zarith.
+ assert (H3' := Z_div_mod_eq [[WW bh bl]] [[WW mh ml]] (Z.lt_gt _ _ H3)).
+ assert (H4 : 0 <= [[WW bh bl]]/[[WW mh ml]]).
+ apply Z.ge_le;apply Z_div_ge0;zarith. simpl in *;rewrite H1.
pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3'.
- destruct (Zle_lt_or_eq _ _ H4').
+ Z.le_elim H4.
assert (H6' : [[WW bh bl]] mod [[WW mh ml]] =
[[WW bh bl]] - [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])).
simpl;pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3';ring. simpl in H6'.
assert ([[WW mh ml]] <= [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])).
- simpl;pattern ([|mh|]*wB+[|ml|]) at 1;rewrite <- Zmult_1_r;zarith.
+ simpl;pattern ([|mh|]*wB+[|ml|]) at 1;rewrite <- Z.mul_1_r;zarith.
simpl in *;assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in H8;
zarith.
assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in *;zarith.
- rewrite <- H4 in H3';rewrite Zmult_0_r in H3';simpl in H3';zarith.
+ rewrite <- H4 in H3';rewrite Z.mul_0_r in H3';simpl in H3';zarith.
pattern n at 1;replace n with (n-1+1);try ring.
rewrite Zpower_exp;zarith. change (2^1) with 2.
rewrite Z_div_mult;zarith.
@@ -1295,27 +1295,27 @@ Section DoubleDivGt.
[[ww_gcd_gt_aux p cont ah al bh bl]].
Proof.
induction p;intros cont n Hcont ah al bh bl Hgt Hs;simpl ww_gcd_gt_aux.
- assert (0 < Zpos p). unfold Zlt;reflexivity.
+ assert (0 < Zpos p). unfold Z.lt;reflexivity.
apply spec_ww_gcd_gt_aux_body with (n := Zpos (xI p) + n);
- trivial;rewrite Zpos_xI.
+ trivial;rewrite Pos2Z.inj_xI.
intros. apply IHp with (n := Zpos p + n);zarith.
intros. apply IHp with (n := n );zarith.
- apply Zle_trans with (2 ^ (2* Zpos p + 1+ n -1));zarith.
- apply Zpower_le_monotone2;zarith.
- assert (0 < Zpos p). unfold Zlt;reflexivity.
+ apply Z.le_trans with (2 ^ (2* Zpos p + 1+ n -1));zarith.
+ apply Z.pow_le_mono_r;zarith.
+ assert (0 < Zpos p). unfold Z.lt;reflexivity.
apply spec_ww_gcd_gt_aux_body with (n := Zpos (xO p) + n );trivial.
- rewrite (Zpos_xO p).
+ rewrite (Pos2Z.inj_xO p).
intros. apply IHp with (n := Zpos p + n - 1);zarith.
intros. apply IHp with (n := n -1 );zarith.
intros;apply Hcont;zarith.
- apply Zle_trans with (2^(n-1));zarith.
- apply Zpower_le_monotone2;zarith.
- apply Zle_trans with (2 ^ (Zpos p + n -1));zarith.
- apply Zpower_le_monotone2;zarith.
- apply Zle_trans with (2 ^ (2*Zpos p + n -1));zarith.
- apply Zpower_le_monotone2;zarith.
+ apply Z.le_trans with (2^(n-1));zarith.
+ apply Z.pow_le_mono_r;zarith.
+ apply Z.le_trans with (2 ^ (Zpos p + n -1));zarith.
+ apply Z.pow_le_mono_r;zarith.
+ apply Z.le_trans with (2 ^ (2*Zpos p + n -1));zarith.
+ apply Z.pow_le_mono_r;zarith.
apply spec_ww_gcd_gt_aux_body with (n := n+1);trivial.
- rewrite Zplus_comm;trivial.
+ rewrite Z.add_comm;trivial.
ring_simplify (n + 1 - 1);trivial.
Qed.
@@ -1353,7 +1353,7 @@ Section DoubleDiv.
Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
Variable spec_ww_1 : [[ww_1]] = 1.
Variable spec_ww_compare : forall x y,
- ww_compare x y = Zcompare [[x]] [[y]].
+ ww_compare x y = Z.compare [[x]] [[y]].
Variable spec_ww_div_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] ->
let (q,r) := ww_div_gt a b in
[[a]] = [[q]] * [[b]] + [[r]] /\
@@ -1375,7 +1375,7 @@ Section DoubleDiv.
0 <= [[r]] < [[b]].
Proof.
intros a b Hpos;unfold ww_div.
- rewrite spec_ww_compare; case Zcompare_spec; intros.
+ rewrite spec_ww_compare; case Z.compare_spec; intros.
simpl;rewrite spec_ww_1;split;zarith.
simpl;split;[ring|Spec_ww_to_Z a;zarith].
apply spec_ww_div_gt;auto with zarith.
@@ -1385,7 +1385,7 @@ Section DoubleDiv.
[[ww_mod a b]] = [[a]] mod [[b]].
Proof.
intros a b Hpos;unfold ww_mod.
- rewrite spec_ww_compare; case Zcompare_spec; intros.
+ rewrite spec_ww_compare; case Z.compare_spec; intros.
simpl;apply Zmod_unique with 1;try rewrite H;zarith.
Spec_ww_to_Z a;symmetry;apply Zmod_small;zarith.
apply spec_ww_mod_gt;auto with zarith.
@@ -1406,7 +1406,7 @@ Section DoubleDiv.
Variable spec_w_0 : [|w_0|] = 0.
Variable spec_w_1 : [|w_1|] = 1.
Variable spec_compare :
- forall x y, w_compare x y = Zcompare [|x|] [|y|].
+ forall x y, w_compare x y = Z.compare [|x|] [|y|].
Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0.
Variable spec_gcd_gt : forall a b, [|a|] > [|b|] ->
Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|].
@@ -1439,7 +1439,7 @@ Section DoubleDiv.
assert (H1:= beta_lex _ _ _ _ _ Hle (spec_to_Z yl) H).
Spec_w_to_Z yh;zarith.
unfold gcd_cont; rewrite spec_compare, spec_w_1.
- case Zcompare_spec; intros Hcmpy.
+ case Z.compare_spec; intros Hcmpy.
simpl;rewrite H;simpl;
rewrite spec_ww_1;rewrite <- Hcmpy;apply Zis_gcd_mod;zarith.
rewrite <- (Zmod_unique ([|xh|]*wB+[|xl|]) 1 ([|xh|]*wB+[|xl|]) 0);zarith.
@@ -1485,7 +1485,7 @@ Section DoubleDiv.
Spec_w_to_Z bh;assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt.
rewrite H1;simpl;auto. clear H.
apply spec_gcd_gt_fix with (n:= 0);trivial.
- rewrite Zplus_0_r;rewrite spec_ww_digits_.
+ rewrite Z.add_0_r;rewrite spec_ww_digits_.
change (2 ^ Zpos (xO w_digits)) with wwB. Spec_ww_to_Z (WW bh bl);zarith.
Qed.
@@ -1498,7 +1498,7 @@ Section DoubleDiv.
| Eq => a
| Lt => ww_gcd_gt b a
end).
- rewrite spec_ww_compare; case Zcompare_spec; intros Hcmp.
+ rewrite spec_ww_compare; case Z.compare_spec; intros Hcmp.
Spec_ww_to_Z b;rewrite Hcmp.
apply Zis_gcd_for_euclid with 1;zarith.
ring_simplify ([[b]] - 1 * [[b]]). apply Zis_gcd_0;zarith.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
index 062282f2e..8a8e90c3c 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
@@ -62,7 +62,7 @@ Section GENDIVN1.
[|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\
0 <= [|r|] < [|b|].
Variable spec_compare :
- forall x y, w_compare x y = Zcompare [|x|] [|y|].
+ forall x y, w_compare x y = Z.compare [|x|] [|y|].
Variable spec_sub: forall x y,
[|w_sub x y|] = ([|x|] - [|y|]) mod wB.
@@ -107,8 +107,8 @@ Section GENDIVN1.
destruct H4;split;trivial.
rewrite spec_double_WW;trivial.
rewrite <- double_wB_wwB.
- rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
- rewrite H0;rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc.
+ rewrite Z.mul_assoc;rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
+ rewrite H0;rewrite Z.mul_add_distr_r;rewrite <- Z.add_assoc.
rewrite H4;ring.
Qed.
@@ -160,7 +160,7 @@ Section GENDIVN1.
Lemma p_lt_double_digits : forall n, [|p|] <= Zpos (w_digits << n).
Proof.
induction n;simpl. trivial.
- case (spec_to_Z p); rewrite Pshiftl_nat_S, Zpos_xO;auto with zarith.
+ case (spec_to_Z p); rewrite Pshiftl_nat_S, Pos2Z.inj_xO;auto with zarith.
Qed.
Lemma spec_double_divn1_p : forall n r h l,
@@ -225,11 +225,11 @@ Section GENDIVN1.
replace (2 ^ (Zpos (w_digits << (S n)) - [|p|])) with
(2^(Zpos (w_digits << n) - [|p|])*2^Zpos (w_digits << n)).
rewrite Zdiv_mult_cancel_r;auto with zarith.
- rewrite Zmult_plus_distr_l with (p:= 2^[|p|]).
+ rewrite Z.mul_add_distr_r with (p:= 2^[|p|]).
pattern ([!n|hl!] * 2^[|p|]) at 2;
rewrite (shift_unshift_mod (Zpos(w_digits << n))([|p|])([!n|hl!]));
auto with zarith.
- rewrite Zplus_assoc.
+ rewrite Z.add_assoc.
replace
([!n|hh!] * 2^Zpos (w_digits << n)* 2^[|p|] +
([!n|hl!] / 2^(Zpos (w_digits << n)-[|p|])*
@@ -238,7 +238,7 @@ Section GENDIVN1.
(([!n|hh!] *2^[|p|] + double_to_Z w_digits w_to_Z n hl /
2^(Zpos (w_digits << n)-[|p|]))
* 2^Zpos(w_digits << n));try (ring;fail).
- rewrite <- Zplus_assoc.
+ rewrite <- Z.add_assoc.
rewrite <- (Zmod_shift_r ([|p|]));auto with zarith.
replace
(2 ^ Zpos (w_digits << n) * 2 ^ Zpos (w_digits << n)) with
@@ -246,12 +246,12 @@ Section GENDIVN1.
rewrite (Zmod_shift_r (Zpos (w_digits << n)));auto with zarith.
replace (2 ^ (Zpos (w_digits << n) + Zpos (w_digits << n)))
with (2^Zpos(w_digits << n) *2^Zpos(w_digits << n)).
- rewrite (Zmult_comm (([!n|hh!] * 2 ^ [|p|] +
+ rewrite (Z.mul_comm (([!n|hh!] * 2 ^ [|p|] +
[!n|hl!] / 2 ^ (Zpos (w_digits << n) - [|p|])))).
rewrite Zmult_mod_distr_l;auto with zarith.
ring.
rewrite Zpower_exp;auto with zarith.
- assert (0 < Zpos (w_digits << n)). unfold Zlt;reflexivity.
+ assert (0 < Zpos (w_digits << n)). unfold Z.lt;reflexivity.
auto with zarith.
apply Z_mod_lt;auto with zarith.
rewrite Zpower_exp;auto with zarith.
@@ -320,7 +320,7 @@ Section GENDIVN1.
replace (Zpos w_digits - Zpos w_digits) with 0;try ring.
simpl. rewrite <- (Zdiv_unique [|x|] 1 [|x|] 0);auto with zarith.
assert (U2 := spec_double_digits n).
- assert (U3 : 0 < Zpos w_digits). exact (refl_equal Lt).
+ assert (U3 : 0 < Zpos w_digits). exact (eq_refl Lt).
destruct x;unfold high;fold high.
unfold double_to_Z,zn2z_to_Z;rewrite spec_0.
rewrite Zdiv_0_l;trivial.
@@ -365,30 +365,30 @@ Section GENDIVN1.
intros n a b H. unfold double_divn1.
case (spec_head0 H); intros H0 H1.
case (spec_to_Z (w_head0 b)); intros HH1 HH2.
- rewrite spec_compare; case Zcompare_spec;
+ rewrite spec_compare; case Z.compare_spec;
rewrite spec_0; intros H2; auto with zarith.
assert (Hv1: wB/2 <= [|b|]).
- generalize H0; rewrite H2; rewrite Zpower_0_r;
- rewrite Zmult_1_l; auto.
+ generalize H0; rewrite H2; rewrite Z.pow_0_r;
+ rewrite Z.mul_1_l; auto.
assert (Hv2: [|w_0|] < [|b|]).
rewrite spec_0; auto.
generalize (spec_double_divn1_0 Hv1 n a Hv2).
- rewrite spec_0;rewrite Zmult_0_l; rewrite Zplus_0_l; auto.
+ rewrite spec_0;rewrite Z.mul_0_l; rewrite Z.add_0_l; auto.
contradict H2; auto with zarith.
assert (HHHH : 0 < [|w_head0 b|]); auto with zarith.
assert ([|w_head0 b|] < Zpos w_digits).
- case (Zle_or_lt (Zpos w_digits) [|w_head0 b|]); auto; intros HH.
+ case (Z.le_gt_cases (Zpos w_digits) [|w_head0 b|]); auto; intros HH.
assert (2 ^ [|w_head0 b|] < wB).
- apply Zle_lt_trans with (2 ^ [|w_head0 b|] * [|b|]);auto with zarith.
+ apply Z.le_lt_trans with (2 ^ [|w_head0 b|] * [|b|]);auto with zarith.
replace (2 ^ [|w_head0 b|]) with (2^[|w_head0 b|] * 1);try (ring;fail).
- apply Zmult_le_compat;auto with zarith.
+ apply Z.mul_le_mono_nonneg;auto with zarith.
assert (wB <= 2^[|w_head0 b|]).
unfold base;apply Zpower_le_monotone;auto with zarith. omega.
assert ([|w_add_mul_div (w_head0 b) b w_0|] =
2 ^ [|w_head0 b|] * [|b|]).
rewrite (spec_add_mul_div b w_0); auto with zarith.
rewrite spec_0;rewrite Zdiv_0_l; try omega.
- rewrite Zplus_0_r; rewrite Zmult_comm.
+ rewrite Z.add_0_r; rewrite Z.mul_comm.
rewrite Zmod_small; auto with zarith.
assert (H5 := spec_to_Z (high n a)).
assert
@@ -396,21 +396,21 @@ Section GENDIVN1.
<[|w_add_mul_div (w_head0 b) b w_0|]).
rewrite H4.
rewrite spec_add_mul_div;auto with zarith.
- rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l.
+ rewrite spec_0;rewrite Z.mul_0_l;rewrite Z.add_0_l.
assert (([|high n a|]/2^(Zpos w_digits - [|w_head0 b|])) < wB).
apply Zdiv_lt_upper_bound;auto with zarith.
- apply Zlt_le_trans with wB;auto with zarith.
+ apply Z.lt_le_trans with wB;auto with zarith.
pattern wB at 1;replace wB with (wB*1);try ring.
- apply Zmult_le_compat;auto with zarith.
- assert (H6 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|]));
+ apply Z.mul_le_mono_nonneg;auto with zarith.
+ assert (H6 := Z.pow_pos_nonneg 2 (Zpos w_digits - [|w_head0 b|]));
auto with zarith.
rewrite Zmod_small;auto with zarith.
apply Zdiv_lt_upper_bound;auto with zarith.
- apply Zlt_le_trans with wB;auto with zarith.
- apply Zle_trans with (2 ^ [|w_head0 b|] * [|b|] * 2).
+ apply Z.lt_le_trans with wB;auto with zarith.
+ apply Z.le_trans with (2 ^ [|w_head0 b|] * [|b|] * 2).
rewrite <- wB_div_2; try omega.
- apply Zmult_le_compat;auto with zarith.
- pattern 2 at 1;rewrite <- Zpower_1_r.
+ apply Z.mul_le_mono_nonneg;auto with zarith.
+ pattern 2 at 1;rewrite <- Z.pow_1_r.
apply Zpower_le_monotone;split;auto with zarith.
rewrite <- H4 in H0.
assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith.
@@ -420,9 +420,9 @@ Section GENDIVN1.
(double_0 w_0 n)) as (q,r).
assert (U:= spec_double_digits n).
rewrite spec_double_0 in H7;trivial;rewrite Zdiv_0_l in H7.
- rewrite Zplus_0_r in H7.
+ rewrite Z.add_0_r in H7.
rewrite spec_add_mul_div in H7;auto with zarith.
- rewrite spec_0 in H7;rewrite Zmult_0_l in H7;rewrite Zplus_0_l in H7.
+ rewrite spec_0 in H7;rewrite Z.mul_0_l in H7;rewrite Z.add_0_l in H7.
assert (([|high n a|] / 2 ^ (Zpos w_digits - [|w_head0 b|])) mod wB
= [!n|a!] / 2^(Zpos (w_digits << n) - [|w_head0 b|])).
rewrite Zmod_small;auto with zarith.
@@ -431,29 +431,29 @@ Section GENDIVN1.
replace (Zpos (w_digits << n) - Zpos w_digits +
(Zpos w_digits - [|w_head0 b|]))
with (Zpos (w_digits << n) - [|w_head0 b|]);trivial;ring.
- assert (H8 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith.
+ assert (H8 := Z.pow_pos_nonneg 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith.
split;auto with zarith.
- apply Zle_lt_trans with ([|high n a|]);auto with zarith.
+ apply Z.le_lt_trans with ([|high n a|]);auto with zarith.
apply Zdiv_le_upper_bound;auto with zarith.
- pattern ([|high n a|]) at 1;rewrite <- Zmult_1_r.
- apply Zmult_le_compat;auto with zarith.
+ pattern ([|high n a|]) at 1;rewrite <- Z.mul_1_r.
+ apply Z.mul_le_mono_nonneg;auto with zarith.
rewrite H8 in H7;unfold double_wB,base in H7.
rewrite <- shift_unshift_mod in H7;auto with zarith.
rewrite H4 in H7.
assert ([|w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r|]
= [|r|]/2^[|w_head0 b|]).
rewrite spec_add_mul_div.
- rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l.
+ rewrite spec_0;rewrite Z.mul_0_l;rewrite Z.add_0_l.
replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|])
with ([|w_head0 b|]).
rewrite Zmod_small;auto with zarith.
assert (H9 := spec_to_Z r).
split;auto with zarith.
- apply Zle_lt_trans with ([|r|]);auto with zarith.
+ apply Z.le_lt_trans with ([|r|]);auto with zarith.
apply Zdiv_le_upper_bound;auto with zarith.
- pattern ([|r|]) at 1;rewrite <- Zmult_1_r.
- apply Zmult_le_compat;auto with zarith.
- assert (H10 := Zpower_gt_0 2 ([|w_head0 b|]));auto with zarith.
+ pattern ([|r|]) at 1;rewrite <- Z.mul_1_r.
+ apply Z.mul_le_mono_nonneg;auto with zarith.
+ assert (H10 := Z.pow_pos_nonneg 2 ([|w_head0 b|]));auto with zarith.
rewrite spec_sub.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
@@ -475,7 +475,7 @@ Section GENDIVN1.
auto with zarith.
rewrite H9.
apply Zdiv_lt_upper_bound;auto with zarith.
- rewrite Zmult_comm;auto with zarith.
+ rewrite Z.mul_comm;auto with zarith.
exact (spec_double_to_Z w_digits w_to_Z spec_to_Z n a).
Qed.
@@ -498,7 +498,7 @@ Section GENDIVN1.
double_modn1 n a b = snd (double_divn1 n a b).
Proof.
intros n a b;unfold double_divn1,double_modn1.
- rewrite spec_compare; case Zcompare_spec;
+ rewrite spec_compare; case Z.compare_spec;
rewrite spec_0; intros H2; auto with zarith.
apply spec_double_modn1_0.
apply spec_double_modn1_0.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
index a6a0fc8e3..7086d0fd7 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
@@ -104,9 +104,9 @@ Section DoubleLift.
Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB.
Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
Variable spec_compare : forall x y,
- w_compare x y = Zcompare [|x|] [|y|].
+ w_compare x y = Z.compare [|x|] [|y|].
Variable spec_ww_compare : forall x y,
- ww_compare x y = Zcompare [[x]] [[y]].
+ ww_compare x y = Z.compare [[x]] [[y]].
Variable spec_ww_digits : ww_Digits = xO w_digits.
Variable spec_w_head00 : forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits.
Variable spec_w_head0 : forall x, 0 < [|x|] ->
@@ -140,20 +140,20 @@ Section DoubleLift.
case (spec_to_Z xh); intros Hx1 Hx2.
case (spec_to_Z xl); intros Hy1 Hy2.
assert (F1: [|xh|] = 0).
- case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3.
- absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- apply Zlt_le_trans with (1 := Hy3); auto with zarith.
- pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]).
- apply Zplus_le_compat_r; auto with zarith.
- case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3.
- absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith.
- apply Zmult_lt_0_compat; auto with zarith.
- rewrite spec_compare. case Zcompare_spec.
+ { Z.le_elim Hy1; auto.
+ - absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ apply Z.lt_le_trans with (1 := Hy1); auto with zarith.
+ pattern [|xl|] at 1; rewrite <- (Z.add_0_l [|xl|]).
+ apply Z.add_le_mono_r; auto with zarith.
+ - Z.le_elim Hx1; auto.
+ absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ rewrite <- Hy1; rewrite Z.add_0_r; auto with zarith.
+ apply Z.mul_pos_pos; auto with zarith. }
+ rewrite spec_compare. case Z.compare_spec.
intros H; simpl.
rewrite spec_w_add; rewrite spec_w_head00.
rewrite spec_zdigits; rewrite spec_ww_digits.
- rewrite Zpos_xO; auto with zarith.
+ rewrite Pos2Z.inj_xO; auto with zarith.
rewrite F1 in Hx; auto with zarith.
rewrite spec_w_0; auto with zarith.
rewrite spec_w_0; auto with zarith.
@@ -163,43 +163,43 @@ Section DoubleLift.
wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB.
Proof.
clear spec_ww_zdigits.
- rewrite wwB_div_2;rewrite Zmult_comm;rewrite wwB_wBwB.
+ rewrite wwB_div_2;rewrite Z.mul_comm;rewrite wwB_wBwB.
assert (U:= lt_0_wB w_digits); destruct x as [ |xh xl];simpl ww_to_Z;intros H.
- unfold Zlt in H;discriminate H.
- rewrite spec_compare, spec_w_0. case Zcompare_spec; intros H0.
- rewrite <- H0 in *. simpl Zplus. simpl in H.
+ unfold Z.lt in H;discriminate H.
+ rewrite spec_compare, spec_w_0. case Z.compare_spec; intros H0.
+ rewrite <- H0 in *. simpl Z.add. simpl in H.
case (spec_to_Z w_zdigits);
case (spec_to_Z (w_head0 xl)); intros HH1 HH2 HH3 HH4.
rewrite spec_w_add.
rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith.
case (spec_w_head0 H); intros H1 H2.
- rewrite Zpower_2; fold wB; rewrite <- Zmult_assoc; split.
- apply Zmult_le_compat_l; auto with zarith.
- apply Zmult_lt_compat_l; auto with zarith.
+ rewrite Z.pow_2_r; fold wB; rewrite <- Z.mul_assoc; split.
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
+ apply Z.mul_lt_mono_pos_l; auto with zarith.
assert (H1 := spec_w_head0 H0).
rewrite spec_w_0W.
split.
- rewrite Zmult_plus_distr_r;rewrite Zmult_assoc.
- apply Zle_trans with (2 ^ [|w_head0 xh|] * [|xh|] * wB).
- rewrite Zmult_comm; zarith.
+ rewrite Z.mul_add_distr_l;rewrite Z.mul_assoc.
+ apply Z.le_trans with (2 ^ [|w_head0 xh|] * [|xh|] * wB).
+ rewrite Z.mul_comm; zarith.
assert (0 <= 2 ^ [|w_head0 xh|] * [|xl|]);zarith.
- assert (H2:=spec_to_Z xl);apply Zmult_le_0_compat;zarith.
+ assert (H2:=spec_to_Z xl);apply Z.mul_nonneg_nonneg;zarith.
case (spec_to_Z (w_head0 xh)); intros H2 _.
generalize ([|w_head0 xh|]) H1 H2;clear H1 H2;
intros p H1 H2.
assert (Eq1 : 2^p < wB).
- rewrite <- (Zmult_1_r (2^p));apply Zle_lt_trans with (2^p*[|xh|]);zarith.
+ rewrite <- (Z.mul_1_r (2^p));apply Z.le_lt_trans with (2^p*[|xh|]);zarith.
assert (Eq2: p < Zpos w_digits).
- destruct (Zle_or_lt (Zpos w_digits) p);trivial;contradict Eq1.
- apply Zle_not_lt;unfold base;apply Zpower_le_monotone;zarith.
+ destruct (Z.le_gt_cases (Zpos w_digits) p);trivial;contradict Eq1.
+ apply Z.le_ngt;unfold base;apply Zpower_le_monotone;zarith.
assert (Zpos w_digits = p + (Zpos w_digits - p)). ring.
- rewrite Zpower_2.
+ rewrite Z.pow_2_r.
unfold base at 2;rewrite H3;rewrite Zpower_exp;zarith.
- rewrite <- Zmult_assoc; apply Zmult_lt_compat_l; zarith.
- rewrite <- (Zplus_0_r (2^(Zpos w_digits - p)*wB));apply beta_lex_inv;zarith.
- apply Zmult_lt_reg_r with (2 ^ p); zarith.
+ rewrite <- Z.mul_assoc; apply Z.mul_lt_mono_pos_l; zarith.
+ rewrite <- (Z.add_0_r (2^(Zpos w_digits - p)*wB));apply beta_lex_inv;zarith.
+ apply Z.mul_lt_mono_pos_r with (2 ^ p); zarith.
rewrite <- Zpower_exp;zarith.
- rewrite Zmult_comm;ring_simplify (Zpos w_digits - p + p);fold wB;zarith.
+ rewrite Z.mul_comm;ring_simplify (Zpos w_digits - p + p);fold wB;zarith.
assert (H1 := spec_to_Z xh);zarith.
Qed.
@@ -211,22 +211,22 @@ Section DoubleLift.
case (spec_to_Z xh); intros Hx1 Hx2.
case (spec_to_Z xl); intros Hy1 Hy2.
assert (F1: [|xh|] = 0).
- case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3.
- absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- apply Zlt_le_trans with (1 := Hy3); auto with zarith.
- pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]).
- apply Zplus_le_compat_r; auto with zarith.
- case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3.
- absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
- rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith.
- apply Zmult_lt_0_compat; auto with zarith.
+ { Z.le_elim Hy1; auto.
+ - absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ apply Z.lt_le_trans with (1 := Hy1); auto with zarith.
+ pattern [|xl|] at 1; rewrite <- (Z.add_0_l [|xl|]).
+ apply Z.add_le_mono_r; auto with zarith.
+ - Z.le_elim Hx1; auto.
+ absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ rewrite <- Hy1; rewrite Z.add_0_r; auto with zarith.
+ apply Z.mul_pos_pos; auto with zarith. }
assert (F2: [|xl|] = 0).
rewrite F1 in Hx; auto with zarith.
- rewrite spec_compare; case Zcompare_spec.
+ rewrite spec_compare; case Z.compare_spec.
intros H; simpl.
rewrite spec_w_add; rewrite spec_w_tail00; auto.
rewrite spec_zdigits; rewrite spec_ww_digits.
- rewrite Zpos_xO; auto with zarith.
+ rewrite Pos2Z.inj_xO; auto with zarith.
rewrite spec_w_0; auto with zarith.
rewrite spec_w_0; auto with zarith.
Qed.
@@ -236,51 +236,51 @@ Section DoubleLift.
Proof.
clear spec_ww_zdigits.
destruct x as [ |xh xl];simpl ww_to_Z;intros H.
- unfold Zlt in H;discriminate H.
- rewrite spec_compare, spec_w_0. case Zcompare_spec; intros H0.
- rewrite <- H0; rewrite Zplus_0_r.
+ unfold Z.lt in H;discriminate H.
+ rewrite spec_compare, spec_w_0. case Z.compare_spec; intros H0.
+ rewrite <- H0; rewrite Z.add_0_r.
case (spec_to_Z (w_tail0 xh)); intros HH1 HH2.
- generalize H; rewrite <- H0; rewrite Zplus_0_r; clear H; intros H.
+ generalize H; rewrite <- H0; rewrite Z.add_0_r; clear H; intros H.
case (@spec_w_tail0 xh).
- apply Zmult_lt_reg_r with wB; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with wB; auto with zarith.
unfold base; auto with zarith.
intros z (Hz1, Hz2); exists z; split; auto.
- rewrite spec_w_add; rewrite (fun x => Zplus_comm [|x|]).
+ rewrite spec_w_add; rewrite (fun x => Z.add_comm [|x|]).
rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith.
- rewrite Zmult_assoc; rewrite <- Hz2; auto.
+ rewrite Z.mul_assoc; rewrite <- Hz2; auto.
case (spec_to_Z (w_tail0 xh)); intros HH1 HH2.
case (spec_w_tail0 H0); intros z (Hz1, Hz2).
assert (Hp: [|w_tail0 xl|] < Zpos w_digits).
- case (Zle_or_lt (Zpos w_digits) [|w_tail0 xl|]); auto; intros H1.
+ case (Z.le_gt_cases (Zpos w_digits) [|w_tail0 xl|]); auto; intros H1.
absurd (2 ^ (Zpos w_digits) <= 2 ^ [|w_tail0 xl|]).
- apply Zlt_not_le.
+ apply Z.lt_nge.
case (spec_to_Z xl); intros HH3 HH4.
- apply Zle_lt_trans with (2 := HH4).
- apply Zle_trans with (1 * 2 ^ [|w_tail0 xl|]); auto with zarith.
+ apply Z.le_lt_trans with (2 := HH4).
+ apply Z.le_trans with (1 * 2 ^ [|w_tail0 xl|]); auto with zarith.
rewrite Hz2.
- apply Zmult_le_compat_r; auto with zarith.
+ apply Z.mul_le_mono_nonneg_r; auto with zarith.
apply Zpower_le_monotone; auto with zarith.
exists ([|xh|] * (2 ^ ((Zpos w_digits - [|w_tail0 xl|]) - 1)) + z); split.
- apply Zplus_le_0_compat; auto.
- apply Zmult_le_0_compat; auto with zarith.
+ apply Z.add_nonneg_nonneg; auto.
+ apply Z.mul_nonneg_nonneg; auto with zarith.
case (spec_to_Z xh); auto.
rewrite spec_w_0W.
- rewrite (Zmult_plus_distr_r 2); rewrite <- Zplus_assoc.
- rewrite Zmult_plus_distr_l; rewrite <- Hz2.
- apply f_equal2 with (f := Zplus); auto.
- rewrite (Zmult_comm 2).
- repeat rewrite <- Zmult_assoc.
- apply f_equal2 with (f := Zmult); auto.
+ rewrite (Z.mul_add_distr_l 2); rewrite <- Z.add_assoc.
+ rewrite Z.mul_add_distr_r; rewrite <- Hz2.
+ apply f_equal2 with (f := Z.add); auto.
+ rewrite (Z.mul_comm 2).
+ repeat rewrite <- Z.mul_assoc.
+ apply f_equal2 with (f := Z.mul); auto.
case (spec_to_Z (w_tail0 xl)); intros HH3 HH4.
- pattern 2 at 2; rewrite <- Zpower_1_r.
+ pattern 2 at 2; rewrite <- Z.pow_1_r.
lazy beta; repeat rewrite <- Zpower_exp; auto with zarith.
- unfold base; apply f_equal with (f := Zpower 2); auto with zarith.
+ unfold base; apply f_equal with (f := Z.pow 2); auto with zarith.
contradict H0; case (spec_to_Z xl); auto with zarith.
Qed.
- Hint Rewrite Zdiv_0_l Zmult_0_l Zplus_0_l Zmult_0_r Zplus_0_r
+ Hint Rewrite Zdiv_0_l Z.mul_0_l Z.add_0_l Z.mul_0_r Z.add_0_r
spec_w_W0 spec_w_0W spec_w_WW spec_w_0
(wB_div w_digits w_to_Z spec_to_Z)
(wB_div_plus w_digits w_to_Z spec_to_Z) : w_rewrite.
@@ -304,20 +304,20 @@ Section DoubleLift.
intros xh xl yh yl p zdigits;assert (HwwB := wwB_pos w_digits).
case (spec_to_w_Z p); intros Hv1 Hv2.
replace (Zpos (xO w_digits)) with (Zpos w_digits + Zpos w_digits).
- 2 : rewrite Zpos_xO;ring.
+ 2 : rewrite Pos2Z.inj_xO;ring.
replace (Zpos w_digits + Zpos w_digits - [[p]]) with
(Zpos w_digits + (Zpos w_digits - [[p]])). 2:ring.
intros Hp; assert (Hxh := spec_to_Z xh);assert (Hxl:=spec_to_Z xl);
assert (Hx := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl));
simpl in Hx;assert (Hyh := spec_to_Z yh);assert (Hyl:=spec_to_Z yl);
assert (Hy:=spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW yh yl));simpl in Hy.
- rewrite spec_ww_compare; case Zcompare_spec; intros H1.
+ rewrite spec_ww_compare; case Z.compare_spec; intros H1.
rewrite H1; unfold zdigits; rewrite spec_w_0W.
- rewrite spec_zdigits; rewrite Zminus_diag; rewrite Zplus_0_r.
+ rewrite spec_zdigits; rewrite Z.sub_diag; rewrite Z.add_0_r.
simpl ww_to_Z; w_rewrite;zarith.
fold wB.
- rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;rewrite <- Zplus_assoc.
- rewrite <- Zpower_2.
+ rewrite Z.mul_add_distr_r;rewrite <- Z.mul_assoc;rewrite <- Z.add_assoc.
+ rewrite <- Z.pow_2_r.
rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|].
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xl yh)). ring.
simpl ww_to_Z; w_rewrite;zarith.
@@ -327,7 +327,7 @@ Section DoubleLift.
case (spec_to_w_Z p); intros HH1 HH2; split; auto.
generalize H1; unfold zdigits; rewrite spec_w_0W;
rewrite spec_zdigits; intros tmp.
- apply Zlt_le_trans with (1 := tmp).
+ apply Z.lt_le_trans with (1 := tmp).
unfold base.
apply Zpower2_le_lin; auto with zarith.
2: generalize H1; unfold zdigits; rewrite spec_w_0W;
@@ -338,16 +338,16 @@ Section DoubleLift.
rewrite HH0; auto with zarith.
repeat rewrite spec_w_add_mul_div with (1 := HH).
rewrite HH0.
- rewrite Zmult_plus_distr_l.
+ rewrite Z.mul_add_distr_r.
pattern ([|xl|] * 2 ^ [[p]]) at 2;
rewrite shift_unshift_mod with (n:= Zpos w_digits);fold wB;zarith.
replace ([|xh|] * wB * 2^[[p]]) with ([|xh|] * 2^[[p]] * wB). 2:ring.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r. rewrite <- Z.add_assoc.
unfold base at 5;rewrite <- Zmod_shift_r;zarith.
unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits));
fold wB;fold wwB;zarith.
- rewrite wwB_wBwB;rewrite Zpower_2; rewrite Zmult_mod_distr_r;zarith.
- unfold ww_digits;rewrite Zpos_xO;zarith. apply Z_mod_lt;zarith.
+ rewrite wwB_wBwB;rewrite Z.pow_2_r; rewrite Zmult_mod_distr_r;zarith.
+ unfold ww_digits;rewrite Pos2Z.inj_xO;zarith. apply Z_mod_lt;zarith.
split;zarith. apply Zdiv_lt_upper_bound;zarith.
rewrite <- Zpower_exp;zarith.
ring_simplify ([[p]] + (Zpos w_digits - [[p]]));fold wB;zarith.
@@ -362,10 +362,10 @@ Section DoubleLift.
rewrite <- Zmod_div_mod; auto with zarith.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
- apply Zle_lt_trans with (Zpos w_digits); auto with zarith.
+ apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_lt_lin; auto with zarith.
exists wB; unfold base.
- unfold ww_digits; rewrite (Zpos_xO w_digits).
+ unfold ww_digits; rewrite (Pos2Z.inj_xO w_digits).
rewrite <- Zpower_exp; auto with zarith.
apply f_equal with (f := fun x => 2 ^ x); auto with zarith.
assert (HH: [|low (ww_sub p zdigits)|] <= Zpos w_digits).
@@ -378,25 +378,25 @@ Section DoubleLift.
pattern wB at 5;replace wB with
(2^(([[p]] - Zpos w_digits)
+ (Zpos w_digits - ([[p]] - Zpos w_digits)))).
- rewrite Zpower_exp;zarith. rewrite Zmult_assoc.
+ rewrite Zpower_exp;zarith. rewrite Z.mul_assoc.
rewrite Z_div_plus_l;zarith.
rewrite shift_unshift_mod with (a:= [|yh|]) (p:= [[p]] - Zpos w_digits)
(n := Zpos w_digits);zarith. fold wB.
set (u := [[p]] - Zpos w_digits).
replace [[p]] with (u + Zpos w_digits);zarith.
- rewrite Zpower_exp;zarith. rewrite Zmult_assoc. fold wB.
- repeat rewrite Zplus_assoc. rewrite <- Zmult_plus_distr_l.
- repeat rewrite <- Zplus_assoc.
+ rewrite Zpower_exp;zarith. rewrite Z.mul_assoc. fold wB.
+ repeat rewrite Z.add_assoc. rewrite <- Z.mul_add_distr_r.
+ repeat rewrite <- Z.add_assoc.
unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits));
fold wB;fold wwB;zarith.
unfold base;rewrite Zmod_shift_r with (a:= Zpos w_digits)
(b:= Zpos w_digits);fold wB;fold wwB;zarith.
- rewrite wwB_wBwB; rewrite Zpower_2; rewrite Zmult_mod_distr_r;zarith.
- rewrite Zmult_plus_distr_l.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r; rewrite Zmult_mod_distr_r;zarith.
+ rewrite Z.mul_add_distr_r.
replace ([|xh|] * wB * 2 ^ u) with
([|xh|]*2^u*wB). 2:ring.
- repeat rewrite <- Zplus_assoc.
- rewrite (Zplus_comm ([|xh|] * 2 ^ u * wB)).
+ repeat rewrite <- Z.add_assoc.
+ rewrite (Z.add_comm ([|xh|] * 2 ^ u * wB)).
rewrite Z_mod_plus;zarith. rewrite Z_mod_mult;zarith.
unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith.
unfold u; split;zarith.
@@ -404,7 +404,7 @@ Section DoubleLift.
rewrite <- Zpower_exp;zarith.
fold u.
ring_simplify (u + (Zpos w_digits - u)); fold
- wB;zarith. unfold ww_digits;rewrite Zpos_xO;zarith.
+ wB;zarith. unfold ww_digits;rewrite Pos2Z.inj_xO;zarith.
unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith.
unfold u; split;zarith.
unfold u; split;zarith.
@@ -434,14 +434,14 @@ Section DoubleLift.
clear H1;w_rewrite);simpl ww_add_mul_div.
replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial].
intros Heq;rewrite <- Heq;clear Heq; auto.
- rewrite spec_ww_compare. case Zcompare_spec; intros H1; w_rewrite.
+ rewrite spec_ww_compare. case Z.compare_spec; intros H1; w_rewrite.
rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith.
generalize H1; w_rewrite; rewrite spec_zdigits; clear H1; intros H1.
assert (HH0: [|low p|] = [[p]]).
rewrite spec_low.
apply Zmod_small.
case (spec_to_w_Z p); intros HH1 HH2; split; auto.
- apply Zlt_le_trans with (1 := H1).
+ apply Z.lt_le_trans with (1 := H1).
unfold base; apply Zpower2_le_lin; auto with zarith.
rewrite HH0; auto with zarith.
replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial].
@@ -449,7 +449,7 @@ Section DoubleLift.
generalize (spec_ww_compare p (w_0W w_zdigits));
case ww_compare; intros H1; w_rewrite.
rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith.
- rewrite Zpos_xO in H;zarith.
+ rewrite Pos2Z.inj_xO in H;zarith.
assert (HH: [|low (ww_sub p (w_0W w_zdigits)) |] = [[p]] - Zpos w_digits).
symmetry in H1; change ([[p]] > [[w_0W w_zdigits]]) in H1.
revert H1.
@@ -458,12 +458,12 @@ Section DoubleLift.
rewrite <- Zmod_div_mod; auto with zarith.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
- apply Zle_lt_trans with (Zpos w_digits); auto with zarith.
+ apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_lt_lin; auto with zarith.
unfold base; auto with zarith.
unfold base; auto with zarith.
exists wB; unfold base.
- unfold ww_digits; rewrite (Zpos_xO w_digits).
+ unfold ww_digits; rewrite (Pos2Z.inj_xO w_digits).
rewrite <- Zpower_exp; auto with zarith.
apply f_equal with (f := fun x => 2 ^ x); auto with zarith.
case (spec_to_Z xh); auto with zarith.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
index 0032d2c3f..ceee1869d 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
@@ -246,7 +246,7 @@ Section DoubleMul.
Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB.
Variable spec_w_0W : forall l, [[w_0W l]] = [|l|].
Variable spec_w_compare :
- forall x y, w_compare x y = Zcompare [|x|] [|y|].
+ forall x y, w_compare x y = Z.compare [|x|] [|y|].
Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB.
Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB.
@@ -325,7 +325,7 @@ Section DoubleMul.
destruct cc as [ | cch ccl]; simpl zn2z_to_Z; simpl ww_to_Z.
rewrite spec_ww_add;rewrite spec_w_W0;rewrite Zmod_small;
rewrite wwB_wBwB. ring.
- rewrite <- (Zplus_0_r ([|wc|]*wB));rewrite H;apply mult_add_ineq3;zarith.
+ rewrite <- (Z.add_0_r ([|wc|]*wB));rewrite H;apply mult_add_ineq3;zarith.
simpl ww_to_Z in H1. assert (U:=spec_to_Z cch).
assert ([|wc|]*wB + [|cch|] <= 2*wB - 3).
destruct (Z_le_gt_dec ([|wc|]*wB + [|cch|]) (2*wB - 3));trivial.
@@ -335,21 +335,21 @@ Section DoubleMul.
assert (H5 := Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)).
omega.
generalize H3;clear H3;rewrite <- H1.
- rewrite Zplus_assoc; rewrite Zpower_2; rewrite Zmult_assoc;
- rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc; rewrite Z.pow_2_r; rewrite Z.mul_assoc;
+ rewrite <- Z.mul_add_distr_r.
assert (((2 * wB - 4) + 2)*wB <= ([|wc|] * wB + [|cch|])*wB).
- apply Zmult_le_compat;zarith.
- rewrite Zmult_plus_distr_l in H3.
+ apply Z.mul_le_mono_nonneg;zarith.
+ rewrite Z.mul_add_distr_r in H3.
intros. assert (U2 := spec_to_Z ccl);omega.
generalize (spec_ww_add_c (w_W0 ccl) ll);destruct (ww_add_c (w_W0 ccl) ll)
- as [l|l];unfold interp_carry;rewrite spec_w_W0;try rewrite Zmult_1_l;
+ as [l|l];unfold interp_carry;rewrite spec_w_W0;try rewrite Z.mul_1_l;
simpl zn2z_to_Z;
try rewrite spec_ww_add;try rewrite spec_ww_add_carry;rewrite spec_w_WW;
rewrite Zmod_small;rewrite wwB_wBwB;intros.
rewrite H4;ring. rewrite H;apply mult_add_ineq2;zarith.
- rewrite Zplus_assoc;rewrite Zmult_plus_distr_l.
- rewrite Zmult_1_l;rewrite <- Zplus_assoc;rewrite H4;ring.
- repeat rewrite <- Zplus_assoc;rewrite H;apply mult_add_ineq2;zarith.
+ rewrite Z.add_assoc;rewrite Z.mul_add_distr_r.
+ rewrite Z.mul_1_l;rewrite <- Z.add_assoc;rewrite H4;ring.
+ repeat rewrite <- Z.add_assoc;rewrite H;apply mult_add_ineq2;zarith.
Qed.
Lemma spec_double_mul_c : forall cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w,
@@ -361,7 +361,7 @@ Section DoubleMul.
forall x y, [||double_mul_c cross x y||] = [[x]] * [[y]].
Proof.
intros cross Hcross x y;destruct x as [ |xh xl];simpl;trivial.
- destruct y as [ |yh yl];simpl. rewrite Zmult_0_r;trivial.
+ destruct y as [ |yh yl];simpl. rewrite Z.mul_0_r;trivial.
assert (H1:= spec_w_mul_c xh yh);assert (H2:= spec_w_mul_c xl yl).
generalize (Hcross _ _ _ _ _ _ H1 H2).
destruct (cross xh xl yh yl (w_mul_c xh yh) (w_mul_c xl yl)) as (wc,cc).
@@ -382,7 +382,7 @@ Section DoubleMul.
Lemma spec_w_2: [|w_2|] = 2.
unfold w_2; rewrite spec_w_add; rewrite spec_w_1; simpl.
apply Zmod_small; split; auto with zarith.
- rewrite <- (Zpower_1_r 2); unfold base; apply Zpower_lt_monotone; auto with zarith.
+ rewrite <- (Z.pow_1_r 2); unfold base; apply Zpower_lt_monotone; auto with zarith.
Qed.
Lemma kara_prod_aux : forall xh xl yh yl,
@@ -401,19 +401,19 @@ Section DoubleMul.
assert (Hyh := (spec_to_Z yh)); assert (Hyl := (spec_to_Z yl)).
generalize (spec_ww_add_c hh ll); case (ww_add_c hh ll);
intros z Hz; rewrite <- Hz; unfold interp_carry; assert (Hz1 := (spec_ww_to_Z z)).
- rewrite spec_w_compare; case Zcompare_spec; intros Hxlh;
+ rewrite spec_w_compare; case Z.compare_spec; intros Hxlh;
try rewrite Hxlh; try rewrite spec_w_0; try (ring; fail).
- rewrite spec_w_compare; case Zcompare_spec; intros Hylh.
+ rewrite spec_w_compare; case Z.compare_spec; intros Hylh.
rewrite Hylh; rewrite spec_w_0; try (ring; fail).
rewrite spec_w_0; try (ring; fail).
repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
split; auto with zarith.
simpl in Hz; rewrite Hz; rewrite H; rewrite H0.
- rewrite kara_prod_aux; apply Zplus_le_0_compat; apply Zmult_le_0_compat; auto with zarith.
- apply Zle_lt_trans with ([[z]]-0); auto with zarith.
- unfold Zminus; apply Zplus_le_compat_l; apply Zle_left_rev; simpl; rewrite Zopp_involutive.
- apply Zmult_le_0_compat; auto with zarith.
+ rewrite kara_prod_aux; apply Z.add_nonneg_nonneg; apply Z.mul_nonneg_nonneg; auto with zarith.
+ apply Z.le_lt_trans with ([[z]]-0); auto with zarith.
+ unfold Z.sub; apply Z.add_le_mono_l; apply Z.le_0_sub; simpl; rewrite Z.opp_involutive.
+ apply Z.mul_nonneg_nonneg; auto with zarith.
match goal with |- context[ww_add_c ?x ?y] =>
generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0;
intros z1 Hz2
@@ -423,7 +423,7 @@ Section DoubleMul.
rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2;
repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_compare; case Zcompare_spec; intros Hylh.
+ rewrite spec_w_compare; case Z.compare_spec; intros Hylh.
rewrite Hylh; rewrite spec_w_0; try (ring; fail).
match goal with |- context[ww_add_c ?x ?y] =>
generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0;
@@ -442,15 +442,15 @@ Section DoubleMul.
replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring]
end.
simpl in Hz; rewrite Hz; rewrite H; rewrite H0.
- rewrite kara_prod_aux; apply Zplus_le_0_compat; apply Zmult_le_0_compat; auto with zarith.
- apply Zle_lt_trans with ([[z]]-0); auto with zarith.
- unfold Zminus; apply Zplus_le_compat_l; apply Zle_left_rev; simpl; rewrite Zopp_involutive.
- apply Zmult_le_0_compat; auto with zarith.
+ rewrite kara_prod_aux; apply Z.add_nonneg_nonneg; apply Z.mul_nonneg_nonneg; auto with zarith.
+ apply Z.le_lt_trans with ([[z]]-0); auto with zarith.
+ unfold Z.sub; apply Z.add_le_mono_l; apply Z.le_0_sub; simpl; rewrite Z.opp_involutive.
+ apply Z.mul_nonneg_nonneg; auto with zarith.
(** there is a carry in hh + ll **)
- rewrite Zmult_1_l.
- rewrite spec_w_compare; case Zcompare_spec; intros Hxlh;
+ rewrite Z.mul_1_l.
+ rewrite spec_w_compare; case Z.compare_spec; intros Hxlh;
try rewrite Hxlh; try rewrite spec_w_1; try (ring; fail).
- rewrite spec_w_compare; case Zcompare_spec; intros Hylh;
+ rewrite spec_w_compare; case Z.compare_spec; intros Hylh;
try rewrite Hylh; try rewrite spec_w_1; try (ring; fail).
match goal with |- context[ww_sub_c ?x ?y] =>
generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1;
@@ -458,7 +458,7 @@ Section DoubleMul.
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l.
+ rewrite spec_w_0; rewrite Z.mul_0_l; rewrite Z.add_0_l.
generalize Hz2; clear Hz2; unfold interp_carry.
repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
@@ -469,11 +469,11 @@ Section DoubleMul.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_2; unfold interp_carry in Hz2.
- apply trans_equal with (wwB + (1 * wwB + [[z1]])).
+ transitivity (wwB + (1 * wwB + [[z1]])).
ring.
rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_compare; case Zcompare_spec; intros Hylh;
+ rewrite spec_w_compare; case Z.compare_spec; intros Hylh;
try rewrite Hylh; try rewrite spec_w_1; try (ring; fail).
match goal with |- context[ww_add_c ?x ?y] =>
generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1;
@@ -482,7 +482,7 @@ Section DoubleMul.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
rewrite spec_w_2; unfold interp_carry in Hz2.
- apply trans_equal with (wwB + (1 * wwB + [[z1]])).
+ transitivity (wwB + (1 * wwB + [[z1]])).
ring.
rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
@@ -492,7 +492,7 @@ Section DoubleMul.
end.
simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c).
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
- rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l.
+ rewrite spec_w_0; rewrite Z.mul_0_l; rewrite Z.add_0_l.
match goal with |- context[(?x - ?y) * (?z - ?t)] =>
replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring]
end.
@@ -513,7 +513,7 @@ Section DoubleMul.
rewrite <- wwB_wBwB;intros H1 H2.
assert (H3 := wB_pos w_digits).
assert (2*wB <= wwB).
- rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat;zarith.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r; apply Z.mul_le_mono_nonneg;zarith.
omega.
Qed.
@@ -537,14 +537,14 @@ Section DoubleMul.
assert (U1:= lt_0_wwB w_digits).
intros x y; case x; auto; intros xh xl.
case y; auto.
- simpl; rewrite Zmult_0_r; rewrite Zmod_small; auto with zarith.
+ simpl; rewrite Z.mul_0_r; rewrite Zmod_small; auto with zarith.
intros yh yl;simpl.
repeat (rewrite spec_ww_add || rewrite spec_w_W0 || rewrite spec_w_mul_c
|| rewrite spec_w_add || rewrite spec_w_mul).
rewrite <- Zplus_mod; auto with zarith.
- repeat (rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r).
+ repeat (rewrite Z.mul_add_distr_r || rewrite Z.mul_add_distr_l).
rewrite <- Zmult_mod_distr_r; auto with zarith.
- rewrite <- Zpower_2; rewrite <- wwB_wBwB; auto with zarith.
+ rewrite <- Z.pow_2_r; rewrite <- wwB_wBwB; auto with zarith.
rewrite Zplus_mod; auto with zarith.
rewrite Zmod_mod; auto with zarith.
rewrite <- Zplus_mod; auto with zarith.
@@ -564,10 +564,10 @@ Section DoubleMul.
apply (spec_mul_aux xh xl xh xl wc cc);trivial.
generalize Heq (spec_ww_add_c (w_mul_c xh xl) (w_mul_c xh xl));clear Heq.
rewrite spec_w_mul_c;destruct (ww_add_c (w_mul_c xh xl) (w_mul_c xh xl));
- unfold interp_carry;try rewrite Zmult_1_l;intros Heq Heq';inversion Heq;
- rewrite (Zmult_comm [|xl|]);subst.
- rewrite spec_w_0;rewrite Zmult_0_l;rewrite Zplus_0_l;trivial.
- rewrite spec_w_1;rewrite Zmult_1_l;rewrite <- wwB_wBwB;trivial.
+ unfold interp_carry;try rewrite Z.mul_1_l;intros Heq Heq';inversion Heq;
+ rewrite (Z.mul_comm [|xl|]);subst.
+ rewrite spec_w_0;rewrite Z.mul_0_l;rewrite Z.add_0_l;trivial.
+ rewrite spec_w_1;rewrite Z.mul_1_l;rewrite <- wwB_wBwB;trivial.
Qed.
Section DoubleMulAddn1Proof.
@@ -589,8 +589,8 @@ Section DoubleMul.
assert(H:=IHn xl y r);destruct (double_mul_add_n1 w_mul_add n xl y r)as(rl,l).
assert(U:=IHn xh y rl);destruct(double_mul_add_n1 w_mul_add n xh y rl)as(rh,h).
rewrite <- double_wB_wwB. rewrite spec_double_WW;simpl;trivial.
- rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc;rewrite <- H.
- rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.mul_add_distr_r;rewrite <- Z.add_assoc;rewrite <- H.
+ rewrite Z.mul_assoc;rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
rewrite U;ring.
Qed.
@@ -604,9 +604,9 @@ Section DoubleMul.
destruct (w_mul_c x y) as [ |h l];simpl;rewrite <- H.
rewrite spec_w_0;trivial.
assert (U:=spec_w_add_c l r);destruct (w_add_c l r) as [lr|lr];unfold
- interp_carry in U;try rewrite Zmult_1_l in H;simpl.
+ interp_carry in U;try rewrite Z.mul_1_l in H;simpl.
rewrite U;ring. rewrite spec_w_succ. rewrite Zmod_small.
- rewrite <- Zplus_assoc;rewrite <- U;ring.
+ rewrite <- Z.add_assoc;rewrite <- U;ring.
simpl in H;assert (H1:= Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)).
rewrite <- H in H1.
assert (H2:=spec_to_Z h);split;zarith.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
index b073d6bed..1318c6123 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
@@ -219,7 +219,7 @@ Section DoubleSqrt.
Variable spec_w_is_even : forall x,
if w_is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1.
Variable spec_w_compare : forall x y,
- w_compare x y = Zcompare [|x|] [|y|].
+ w_compare x y = Z.compare [|x|] [|y|].
Variable spec_w_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
Variable spec_w_square_c : forall x, [[ w_square_c x]] = [|x|] * [|x|].
Variable spec_w_div21 : forall a1 a2 b,
@@ -232,7 +232,7 @@ Section DoubleSqrt.
[|p|] <= Zpos w_digits ->
[| w_add_mul_div p x y |] =
([|x|] * (2 ^ [|p|]) +
- [|y|] / (Zpower 2 ((Zpos w_digits) - [|p|]))) mod wB.
+ [|y|] / (Z.pow 2 ((Zpos w_digits) - [|p|]))) mod wB.
Variable spec_ww_add_mul_div : forall x y p,
[[p]] <= Zpos (xO w_digits) ->
[[ ww_add_mul_div p x y ]] =
@@ -251,7 +251,7 @@ Section DoubleSqrt.
Variable spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB.
Variable spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]].
Variable spec_ww_compare : forall x y,
- ww_compare x y = Zcompare [[x]] [[y]].
+ ww_compare x y = Z.compare [[x]] [[y]].
Variable spec_ww_head0 : forall x, 0 < [[x]] ->
wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB.
Variable spec_low: forall x, [|low x|] = [[x]] mod wB.
@@ -272,10 +272,9 @@ intros x; case x; simpl ww_is_even.
unfold base.
rewrite Zplus_mod; auto with zarith.
rewrite (fun x y => (Zdivide_mod (x * y))); auto with zarith.
- rewrite Zplus_0_l; rewrite Zmod_mod; auto with zarith.
+ rewrite Z.add_0_l; rewrite Zmod_mod; auto with zarith.
apply spec_w_is_even; auto with zarith.
- apply Zdivide_mult_r; apply Zpower_divide; auto with zarith.
- red; simpl; auto.
+ apply Z.divide_mul_r; apply Zpower_divide; auto with zarith.
Qed.
@@ -286,10 +285,10 @@ intros x; case x; simpl ww_is_even.
intros a1 a2 b Hb; unfold w_div21c.
assert (H: 0 < [|b|]); auto with zarith.
assert (U := wB_pos w_digits).
- apply Zlt_le_trans with (2 := Hb); auto with zarith.
- apply Zlt_le_trans with 1; auto with zarith.
+ apply Z.lt_le_trans with (2 := Hb); auto with zarith.
+ apply Z.lt_le_trans with 1; auto with zarith.
apply Zdiv_le_lower_bound; auto with zarith.
- rewrite !spec_w_compare. repeat case Zcompare_spec.
+ rewrite !spec_w_compare. repeat case Z.compare_spec.
intros H1 H2; split.
unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith.
rewrite H1; rewrite H2; ring.
@@ -308,7 +307,7 @@ intros x; case x; simpl ww_is_even.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
assert ([|a2|] < 2 * [|b|]); auto with zarith.
- apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith.
+ apply Z.lt_le_trans with (2 * (wB / 2)); auto with zarith.
rewrite wB_div_2; auto.
intros H1.
match goal with |- context[w_div21 ?y ?z ?t] =>
@@ -321,7 +320,7 @@ intros x; case x; simpl ww_is_even.
rewrite spec_w_sub; auto with zarith.
rewrite Zmod_small; auto with zarith.
assert ([|a1|] < 2 * [|b|]); auto with zarith.
- apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith.
+ apply Z.lt_le_trans with (2 * (wB / 2)); auto with zarith.
rewrite wB_div_2; auto.
destruct (spec_to_Z a1);auto with zarith.
destruct (spec_to_Z a1);auto with zarith.
@@ -333,11 +332,11 @@ intros x; case x; simpl ww_is_even.
intros w0 w1; replace [+|C1 w0|] with (wB + [|w0|]).
rewrite Zmod_small; auto with zarith.
intros (H3, H4); split; auto.
- rewrite Zmult_plus_distr_l.
- rewrite <- Zplus_assoc; rewrite <- H3; ring.
+ rewrite Z.mul_add_distr_r.
+ rewrite <- Z.add_assoc; rewrite <- H3; ring.
split; auto with zarith.
assert ([|a1|] < 2 * [|b|]); auto with zarith.
- apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith.
+ apply Z.lt_le_trans with (2 * (wB / 2)); auto with zarith.
rewrite wB_div_2; auto.
destruct (spec_to_Z a1);auto with zarith.
destruct (spec_to_Z a1);auto with zarith.
@@ -355,14 +354,14 @@ intros x; case x; simpl ww_is_even.
rewrite spec_pred; rewrite spec_w_zdigits.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
- apply Zlt_le_trans with (Zpos w_digits); auto with zarith.
+ apply Z.lt_le_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
rewrite spec_w_add_mul_div; auto with zarith.
autorewrite with w_rewrite rm10.
match goal with |- context[?X - ?Y] =>
replace (X - Y) with 1
end.
- rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith.
+ rewrite Z.pow_1_r; rewrite Zmod_small; auto with zarith.
destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
split; auto with zarith.
apply Zdiv_lt_upper_bound; auto with zarith.
@@ -377,15 +376,15 @@ intros x; case x; simpl ww_is_even.
rewrite spec_pred; rewrite spec_w_zdigits.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
- apply Zlt_le_trans with (Zpos w_digits); auto with zarith.
+ apply Z.lt_le_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
autorewrite with w_rewrite rm10; auto with zarith.
match goal with |- context[?X - ?Y] =>
replace (X - Y) with 1
end; rewrite Hp; try ring.
- rewrite Zpos_minus; auto with zarith.
- rewrite Zmax_right; auto with zarith.
- rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith.
+ rewrite Pos2Z.inj_sub_max; auto with zarith.
+ rewrite Z.max_r; auto with zarith.
+ rewrite Z.pow_1_r; rewrite Zmod_small; auto with zarith.
destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
split; auto with zarith.
unfold base.
@@ -393,14 +392,14 @@ intros x; case x; simpl ww_is_even.
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp
end.
- rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith.
+ rewrite Zpower_exp; try rewrite Z.pow_1_r; auto with zarith.
assert (tmp: forall p, 1 + (p -1) - 1 = p - 1); auto with zarith;
rewrite tmp; clear tmp; auto with zarith.
match goal with |- ?X + ?Y < _ =>
assert (Y < X); auto with zarith
end.
apply Zdiv_lt_upper_bound; auto with zarith.
- pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp;
+ pattern 2 at 2; rewrite <- Z.pow_1_r; rewrite <- Zpower_exp;
auto with zarith.
assert (tmp: forall p, (p - 1) + 1 = p); auto with zarith;
rewrite tmp; clear tmp; auto with zarith.
@@ -410,8 +409,8 @@ intros x; case x; simpl ww_is_even.
[|w_add_mul_div w_1 w w_0|] = 2 * [|w|] mod wB.
intros w1.
autorewrite with w_rewrite rm10; auto with zarith.
- rewrite Zpower_1_r; auto with zarith.
- rewrite Zmult_comm; auto.
+ rewrite Z.pow_1_r; auto with zarith.
+ rewrite Z.mul_comm; auto.
Qed.
Theorem ww_add_mult_mult_2: forall w,
@@ -420,8 +419,8 @@ intros x; case x; simpl ww_is_even.
rewrite spec_ww_add_mul_div; auto with zarith.
autorewrite with w_rewrite rm10.
rewrite spec_w_0W; rewrite spec_w_1.
- rewrite Zpower_1_r; auto with zarith.
- rewrite Zmult_comm; auto.
+ rewrite Z.pow_1_r; auto with zarith.
+ rewrite Z.mul_comm; auto.
rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
red; simpl; intros; discriminate.
Qed.
@@ -432,18 +431,18 @@ intros x; case x; simpl ww_is_even.
intros w1.
rewrite spec_ww_add_mul_div; auto with zarith.
rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
- rewrite Zpower_1_r; auto with zarith.
+ rewrite Z.pow_1_r; auto with zarith.
f_equal; auto.
- rewrite Zmult_comm; f_equal; auto.
+ rewrite Z.mul_comm; f_equal; auto.
autorewrite with w_rewrite rm10.
unfold ww_digits, base.
- apply sym_equal; apply Zdiv_unique with (r := 2 ^ (Zpos (ww_digits w_digits) - 1) -1);
+ symmetry; apply Zdiv_unique with (r := 2 ^ (Zpos (ww_digits w_digits) - 1) -1);
auto with zarith.
unfold ww_digits; split; auto with zarith.
match goal with |- 0 <= ?X - 1 =>
assert (0 < X); auto with zarith
end.
- apply Zpower_gt_0; auto with zarith.
+ apply Z.pow_pos_nonneg; auto with zarith.
match goal with |- 0 <= ?X - 1 =>
assert (0 < X); auto with zarith; red; reflexivity
end.
@@ -453,7 +452,7 @@ intros x; case x; simpl ww_is_even.
assert (tmp: forall p, p + p = 2 * p); auto with zarith;
rewrite tmp; clear tmp.
f_equal; auto.
- pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp;
+ pattern 2 at 2; rewrite <- Z.pow_1_r; rewrite <- Zpower_exp;
auto with zarith.
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite tmp; clear tmp; auto.
@@ -466,7 +465,7 @@ intros x; case x; simpl ww_is_even.
Theorem Zplus_mod_one: forall a1 b1, 0 < b1 -> (a1 + b1) mod b1 = a1 mod b1.
intros a1 b1 H; rewrite Zplus_mod; auto with zarith.
- rewrite Z_mod_same; try rewrite Zplus_0_r; auto with zarith.
+ rewrite Z_mod_same; try rewrite Z.add_0_r; auto with zarith.
apply Zmod_mod; auto.
Qed.
@@ -481,8 +480,8 @@ intros x; case x; simpl ww_is_even.
intros a1 a2 b H.
assert (HH: 0 < [|b|]); auto with zarith.
assert (U := wB_pos w_digits).
- apply Zlt_le_trans with (2 := H); auto with zarith.
- apply Zlt_le_trans with 1; auto with zarith.
+ apply Z.lt_le_trans with (2 := H); auto with zarith.
+ apply Z.lt_le_trans with 1; auto with zarith.
apply Zdiv_le_lower_bound; auto with zarith.
unfold w_div2s; case a1; intros w0 H0.
match goal with |- context[w_div21c ?y ?z ?t] =>
@@ -528,10 +527,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_1_r; auto with zarith
+ try rewrite Z.pow_1_r; auto with zarith
end.
- rewrite Zpos_minus; auto with zarith.
- rewrite Zmax_right; auto with zarith.
+ rewrite Pos2Z.inj_sub_max; auto with zarith.
+ rewrite Z.max_r; auto with zarith.
ring.
repeat rewrite C0_id.
rewrite spec_w_add_c; auto with zarith.
@@ -545,10 +544,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_1_r; auto with zarith
+ try rewrite Z.pow_1_r; auto with zarith
end.
- rewrite Zpos_minus; auto with zarith.
- rewrite Zmax_right; auto with zarith.
+ rewrite Pos2Z.inj_sub_max; auto with zarith.
+ rewrite Z.max_r; auto with zarith.
ring.
repeat rewrite C1_plus_wB in H0.
rewrite C1_plus_wB.
@@ -570,7 +569,7 @@ intros x; case x; simpl ww_is_even.
rewrite add_mult_div_2_plus_1.
replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
auto with zarith.
- rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc.
+ rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
rewrite Hw1.
pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
auto with zarith.
@@ -578,10 +577,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_1_r; auto with zarith
+ try rewrite Z.pow_1_r; auto with zarith
end.
- rewrite Zpos_minus; auto with zarith.
- rewrite Zmax_right; auto with zarith.
+ rewrite Pos2Z.inj_sub_max; auto with zarith.
+ rewrite Z.max_r; auto with zarith.
ring.
repeat rewrite C0_id.
rewrite add_mult_div_2_plus_1.
@@ -589,7 +588,7 @@ intros x; case x; simpl ww_is_even.
intros H1; split; auto with zarith.
replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
auto with zarith.
- rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc.
+ rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
rewrite Hw1.
pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
auto with zarith.
@@ -597,10 +596,10 @@ intros x; case x; simpl ww_is_even.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
- try rewrite Zpower_1_r; auto with zarith
+ try rewrite Z.pow_1_r; auto with zarith
end.
- rewrite Zpos_minus; auto with zarith.
- rewrite Zmax_right; auto with zarith.
+ rewrite Pos2Z.inj_sub_max; auto with zarith.
+ rewrite Z.max_r; auto with zarith.
ring.
split; auto with zarith.
destruct (spec_to_Z b);auto with zarith.
@@ -620,7 +619,7 @@ intros x; case x; simpl ww_is_even.
rewrite add_mult_div_2.
replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
auto with zarith.
- rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc.
+ rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
rewrite Hw1.
pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
auto with zarith.
@@ -631,7 +630,7 @@ intros x; case x; simpl ww_is_even.
rewrite add_mult_div_2.
replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB));
auto with zarith.
- rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc.
+ rewrite Z.mul_add_distr_r; rewrite <- Z.add_assoc.
rewrite Hw1.
pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2);
auto with zarith.
@@ -652,20 +651,20 @@ intros x; case x; simpl ww_is_even.
rewrite <- Zpower_exp; auto with zarith.
f_equal; auto with zarith.
rewrite H.
- rewrite (fun x => (Zmult_comm 4 (2 ^x))).
+ rewrite (fun x => (Z.mul_comm 4 (2 ^x))).
rewrite Z_div_mult; auto with zarith.
Qed.
Theorem Zsquare_mult: forall p, p ^ 2 = p * p.
intros p; change 2 with (1 + 1); rewrite Zpower_exp;
- try rewrite Zpower_1_r; auto with zarith.
+ try rewrite Z.pow_1_r; auto with zarith.
Qed.
Theorem Zsquare_pos: forall p, 0 <= p ^ 2.
- intros p; case (Zle_or_lt 0 p); intros H1.
- rewrite Zsquare_mult; apply Zmult_le_0_compat; auto with zarith.
+ intros p; case (Z.le_gt_cases 0 p); intros H1.
+ rewrite Zsquare_mult; apply Z.mul_nonneg_nonneg; auto with zarith.
rewrite Zsquare_mult; replace (p * p) with ((- p) * (- p)); try ring.
- apply Zmult_le_0_compat; auto with zarith.
+ apply Z.mul_nonneg_nonneg; auto with zarith.
Qed.
Lemma spec_split: forall x,
@@ -676,13 +675,12 @@ intros x; case x; simpl ww_is_even.
Theorem mult_wwB: forall x y, [|x|] * [|y|] < wwB.
Proof.
- intros x y; rewrite wwB_wBwB; rewrite Zpower_2.
+ intros x y; rewrite wwB_wBwB; rewrite Z.pow_2_r.
generalize (spec_to_Z x); intros U.
generalize (spec_to_Z y); intros U1.
- apply Zle_lt_trans with ((wB -1 ) * (wB - 1)); auto with zarith.
- apply Zmult_le_compat; auto with zarith.
- repeat (rewrite Zmult_minus_distr_r || rewrite Zmult_minus_distr_l);
- auto with zarith.
+ apply Z.le_lt_trans with ((wB -1 ) * (wB - 1)); auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
+ rewrite ?Z.mul_sub_distr_l, ?Z.mul_sub_distr_r; auto with zarith.
Qed.
Hint Resolve mult_wwB.
@@ -697,22 +695,22 @@ intros x; case x; simpl ww_is_even.
end; simpl fst; simpl snd.
intros w0 w1 Hw0 w2 w3 Hw1.
assert (U: wB/4 <= [|w2|]).
- case (Zle_or_lt (wB / 4) [|w2|]); auto; intros H1.
- contradict H; apply Zlt_not_le.
- rewrite wwB_wBwB; rewrite Zpower_2.
- pattern wB at 1; rewrite <- wB_div_4; rewrite <- Zmult_assoc;
- rewrite Zmult_comm.
+ case (Z.le_gt_cases (wB / 4) [|w2|]); auto; intros H1.
+ contradict H; apply Z.lt_nge.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
+ pattern wB at 1; rewrite <- wB_div_4; rewrite <- Z.mul_assoc;
+ rewrite Z.mul_comm.
rewrite Z_div_mult; auto with zarith.
rewrite <- Hw1.
match goal with |- _ < ?X =>
- pattern X; rewrite <- Zplus_0_r; apply beta_lex_inv;
+ pattern X; rewrite <- Z.add_0_r; apply beta_lex_inv;
auto with zarith
end.
destruct (spec_to_Z w3);auto with zarith.
generalize (@spec_w_sqrt2 w2 w3 U); case (w_sqrt2 w2 w3).
intros w4 c (H1, H2).
assert (U1: wB/2 <= [|w4|]).
- case (Zle_or_lt (wB/2) [|w4|]); auto with zarith.
+ case (Z.le_gt_cases (wB/2) [|w4|]); auto with zarith.
intros U1.
assert (U2 : [|w4|] <= wB/2 -1); auto with zarith.
assert (U3 : [|w4|] ^ 2 <= wB/4 * wB - wB + 1); auto with zarith.
@@ -720,19 +718,19 @@ intros x; case x; simpl ww_is_even.
rewrite Zsquare_mult;
replace Y with ((wB/2 - 1) * (wB/2 -1))
end.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
destruct (spec_to_Z w4);auto with zarith.
destruct (spec_to_Z w4);auto with zarith.
pattern wB at 4 5; rewrite <- wB_div_2.
- rewrite Zmult_assoc.
+ rewrite Z.mul_assoc.
replace ((wB / 4) * 2) with (wB / 2).
ring.
pattern wB at 1; rewrite <- wB_div_4.
change 4 with (2 * 2).
- rewrite <- Zmult_assoc; rewrite (Zmult_comm 2).
+ rewrite <- Z.mul_assoc; rewrite (Z.mul_comm 2).
rewrite Z_div_mult; try ring; auto with zarith.
assert (U4 : [+|c|] <= wB -2); auto with zarith.
- apply Zle_trans with (1 := H2).
+ apply Z.le_trans with (1 := H2).
match goal with |- ?X <= ?Y =>
replace Y with (2 * (wB/ 2 - 1)); auto with zarith
end.
@@ -741,10 +739,10 @@ intros x; case x; simpl ww_is_even.
assert (U5: X < wB / 4 * wB)
end.
rewrite H1; auto with zarith.
- contradict U; apply Zlt_not_le.
- apply Zmult_lt_reg_r with wB; auto with zarith.
+ contradict U; apply Z.lt_nge.
+ apply Z.mul_lt_mono_pos_r with wB; auto with zarith.
destruct (spec_to_Z w4);auto with zarith.
- apply Zle_lt_trans with (2 := U5).
+ apply Z.le_lt_trans with (2 := U5).
unfold ww_to_Z, zn2z_to_Z.
destruct (spec_to_Z w3);auto with zarith.
generalize (@spec_w_div2s c w0 w4 U1 H2).
@@ -766,7 +764,7 @@ intros x; case x; simpl ww_is_even.
unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1.
rewrite <- Hw0.
match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
+ transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
end.
repeat rewrite Zsquare_mult.
rewrite wwB_wBwB; ring.
@@ -779,17 +777,17 @@ intros x; case x; simpl ww_is_even.
match goal with |- ?X - ?Y * ?Y <= _ =>
assert (V := Zsquare_pos Y);
rewrite Zsquare_mult in V;
- apply Zle_trans with X; auto with zarith;
+ apply Z.le_trans with X; auto with zarith;
clear V
end.
match goal with |- ?X * wB + ?Y <= 2 * (?Z * wB + ?T) =>
- apply Zle_trans with ((2 * Z - 1) * wB + wB); auto with zarith
+ apply Z.le_trans with ((2 * Z - 1) * wB + wB); auto with zarith
end.
destruct (spec_to_Z w1);auto with zarith.
match goal with |- ?X <= _ =>
replace X with (2 * [|w4|] * wB); auto with zarith
end.
- rewrite Zmult_plus_distr_r; rewrite Zmult_assoc.
+ rewrite Z.mul_add_distr_l; rewrite Z.mul_assoc.
destruct (spec_to_Z w5); auto with zarith.
ring.
intros z; replace [-[C1 z]] with (- wwB + [[z]]).
@@ -815,7 +813,7 @@ intros x; case x; simpl ww_is_even.
unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1.
rewrite <- Hw0.
match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
+ transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
end.
repeat rewrite Zsquare_mult.
rewrite wwB_wBwB; ring.
@@ -828,11 +826,11 @@ intros x; case x; simpl ww_is_even.
destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith.
assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)).
assert (0 < [[WW w4 w5]]); auto with zarith.
- apply Zlt_le_trans with (wB/ 2 * wB + 0); auto with zarith.
- autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith.
- apply Zmult_lt_reg_r with 2; auto with zarith.
+ apply Z.lt_le_trans with (wB/ 2 * wB + 0); auto with zarith.
+ autorewrite with rm10; apply Z.mul_pos_pos; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
autorewrite with rm10.
- rewrite Zmult_comm; rewrite wB_div_2; auto with zarith.
+ rewrite Z.mul_comm; rewrite wB_div_2; auto with zarith.
case (spec_to_Z w5);auto with zarith.
case (spec_to_Z w5);auto with zarith.
simpl.
@@ -840,11 +838,11 @@ intros x; case x; simpl ww_is_even.
assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith.
split; auto with zarith.
assert (wwB <= 2 * [[WW w4 w5]]); auto with zarith.
- apply Zle_trans with (2 * ([|w4|] * wB)).
- rewrite wwB_wBwB; rewrite Zpower_2.
- rewrite Zmult_assoc; apply Zmult_le_compat_r; auto with zarith.
- rewrite <- wB_div_2; auto with zarith.
+ apply Z.le_trans with (2 * ([|w4|] * wB)).
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
+ rewrite Z.mul_assoc; apply Z.mul_le_mono_nonneg_r; auto with zarith.
assert (V2 := spec_to_Z w5);auto with zarith.
+ rewrite <- wB_div_2; auto with zarith.
simpl ww_to_Z; assert (V2 := spec_to_Z w5);auto with zarith.
assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith.
intros z1; change [-[C1 z1]] with (-wwB + [[z1]]).
@@ -856,21 +854,21 @@ intros x; case x; simpl ww_is_even.
rewrite ww_add_mult_mult_2.
rename V1 into VV1.
assert (VV2: 0 < [[WW w4 w5]]); auto with zarith.
- apply Zlt_le_trans with (wB/ 2 * wB + 0); auto with zarith.
- autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith.
- apply Zmult_lt_reg_r with 2; auto with zarith.
+ apply Z.lt_le_trans with (wB/ 2 * wB + 0); auto with zarith.
+ autorewrite with rm10; apply Z.mul_pos_pos; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
autorewrite with rm10.
- rewrite Zmult_comm; rewrite wB_div_2; auto with zarith.
+ rewrite Z.mul_comm; rewrite wB_div_2; auto with zarith.
assert (VV3 := spec_to_Z w5);auto with zarith.
assert (VV3 := spec_to_Z w5);auto with zarith.
simpl.
assert (VV3 := spec_to_Z w5);auto with zarith.
assert (VV3: wwB <= 2 * [[WW w4 w5]]); auto with zarith.
- apply Zle_trans with (2 * ([|w4|] * wB)).
- rewrite wwB_wBwB; rewrite Zpower_2.
- rewrite Zmult_assoc; apply Zmult_le_compat_r; auto with zarith.
- rewrite <- wB_div_2; auto with zarith.
+ apply Z.le_trans with (2 * ([|w4|] * wB)).
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
+ rewrite Z.mul_assoc; apply Z.mul_le_mono_nonneg_r; auto with zarith.
case (spec_to_Z w5);auto with zarith.
+ rewrite <- wB_div_2; auto with zarith.
simpl ww_to_Z; assert (V4 := spec_to_Z w5);auto with zarith.
rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]);
auto with zarith.
@@ -892,7 +890,7 @@ intros x; case x; simpl ww_is_even.
rewrite <- Hw0.
split.
match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
+ transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
end.
repeat rewrite Zsquare_mult.
rewrite wwB_wBwB; ring.
@@ -905,17 +903,17 @@ intros x; case x; simpl ww_is_even.
assert (V2 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith.
assert (V3 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z1);auto with zarith.
split; auto with zarith.
- rewrite (Zplus_comm (-wwB)); rewrite <- Zplus_assoc.
+ rewrite (Z.add_comm (-wwB)); rewrite <- Z.add_assoc.
rewrite H5.
match goal with |- 0 <= ?X + (?Y - ?Z) =>
- apply Zle_trans with (X - Z); auto with zarith
+ apply Z.le_trans with (X - Z); auto with zarith
end.
2: generalize (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w6 w1)); unfold ww_to_Z; auto with zarith.
rewrite V1.
match goal with |- 0 <= ?X - 1 - ?Y =>
assert (Y < X); auto with zarith
end.
- apply Zlt_le_trans with wwB; auto with zarith.
+ apply Z.lt_le_trans with wwB; auto with zarith.
intros (H3, H4).
match goal with |- context [ww_sub_c ?y ?z] =>
generalize (spec_ww_sub_c y z); case (ww_sub_c y z)
@@ -933,7 +931,7 @@ intros x; case x; simpl ww_is_even.
unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1.
rewrite <- Hw0.
match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
+ transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
end.
repeat rewrite Zsquare_mult.
rewrite wwB_wBwB; ring.
@@ -945,27 +943,27 @@ intros x; case x; simpl ww_is_even.
simpl ww_to_Z.
rewrite H5.
simpl ww_to_Z.
- rewrite wwB_wBwB; rewrite Zpower_2.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
match goal with |- ?X * ?Y + (?Z * ?Y + ?T - ?U) <= _ =>
- apply Zle_trans with (X * Y + (Z * Y + T - 0));
+ apply Z.le_trans with (X * Y + (Z * Y + T - 0));
auto with zarith
end.
assert (V := Zsquare_pos [|w5|]);
rewrite Zsquare_mult in V; auto with zarith.
autorewrite with rm10.
match goal with |- _ <= 2 * (?U * ?V + ?W) =>
- apply Zle_trans with (2 * U * V + 0);
+ apply Z.le_trans with (2 * U * V + 0);
auto with zarith
end.
match goal with |- ?X * ?Y + (?Z * ?Y + ?T) <= _ =>
replace (X * Y + (Z * Y + T)) with ((X + Z) * Y + T);
try ring
end.
- apply Zlt_le_weak; apply beta_lex_inv; auto with zarith.
+ apply Z.lt_le_incl; apply beta_lex_inv; auto with zarith.
destruct (spec_to_Z w1);auto with zarith.
destruct (spec_to_Z w5);auto with zarith.
- rewrite Zmult_plus_distr_r; auto with zarith.
- rewrite Zmult_assoc; auto with zarith.
+ rewrite Z.mul_add_distr_l; auto with zarith.
+ rewrite Z.mul_assoc; auto with zarith.
intros z; replace [-[C1 z]] with (- wwB + [[z]]).
2: simpl; case wwB; auto with zarith.
intros H5; rewrite spec_w_square_c in H5;
@@ -984,7 +982,7 @@ intros x; case x; simpl ww_is_even.
rewrite <- Hw0.
split.
match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
+ transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
end.
repeat rewrite Zsquare_mult.
rewrite wwB_wBwB; ring.
@@ -995,40 +993,38 @@ intros x; case x; simpl ww_is_even.
repeat rewrite Zsquare_mult; ring.
rewrite V.
simpl ww_to_Z.
- rewrite wwB_wBwB; rewrite Zpower_2.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
match goal with |- (?Z * ?Y + ?T - ?U) + ?X * ?Y <= _ =>
- apply Zle_trans with ((Z * Y + T - 0) + X * Y);
+ apply Z.le_trans with ((Z * Y + T - 0) + X * Y);
auto with zarith
end.
assert (V1 := Zsquare_pos [|w5|]);
rewrite Zsquare_mult in V1; auto with zarith.
autorewrite with rm10.
match goal with |- _ <= 2 * (?U * ?V + ?W) =>
- apply Zle_trans with (2 * U * V + 0);
+ apply Z.le_trans with (2 * U * V + 0);
auto with zarith
end.
match goal with |- (?Z * ?Y + ?T) + ?X * ?Y <= _ =>
replace ((Z * Y + T) + X * Y) with ((X + Z) * Y + T);
try ring
end.
- apply Zlt_le_weak; apply beta_lex_inv; auto with zarith.
+ apply Z.lt_le_incl; apply beta_lex_inv; auto with zarith.
destruct (spec_to_Z w1);auto with zarith.
destruct (spec_to_Z w5);auto with zarith.
- rewrite Zmult_plus_distr_r; auto with zarith.
- rewrite Zmult_assoc; auto with zarith.
- case Zle_lt_or_eq with (1 := H2); clear H2; intros H2.
+ rewrite Z.mul_add_distr_l; auto with zarith.
+ rewrite Z.mul_assoc; auto with zarith.
+ Z.le_elim H2.
intros c1 (H3, H4).
- match type of H3 with ?X = ?Y =>
- absurd (X < Y)
- end.
- apply Zle_not_lt; rewrite <- H3; auto with zarith.
- rewrite Zmult_plus_distr_l.
- apply Zlt_le_trans with ((2 * [|w4|]) * wB + 0);
+ match type of H3 with ?X = ?Y => absurd (X < Y) end.
+ apply Z.le_ngt; rewrite <- H3; auto with zarith.
+ rewrite Z.mul_add_distr_r.
+ apply Z.lt_le_trans with ((2 * [|w4|]) * wB + 0);
auto with zarith.
apply beta_lex_inv; auto with zarith.
destruct (spec_to_Z w0);auto with zarith.
assert (V1 := spec_to_Z w5);auto with zarith.
- rewrite (Zmult_comm wB); auto with zarith.
+ rewrite (Z.mul_comm wB); auto with zarith.
assert (0 <= [|w5|] * (2 * [|w4|])); auto with zarith.
intros c1 (H3, H4); rewrite H2 in H3.
match type of H3 with ?X + ?Y = (?Z + ?T) * ?U + ?V =>
@@ -1038,20 +1034,19 @@ intros x; case x; simpl ww_is_even.
end.
assert (V1 := spec_to_Z w0);auto with zarith.
assert (V2 := spec_to_Z w5);auto with zarith.
- case (Zle_lt_or_eq 0 [|w5|]); auto with zarith; intros V3.
- match type of VV with ?X = ?Y =>
- absurd (X < Y)
- end.
- apply Zle_not_lt; rewrite <- VV; auto with zarith.
- apply Zlt_le_trans with wB; auto with zarith.
+ case V2; intros V3 _.
+ Z.le_elim V3; auto with zarith.
+ match type of VV with ?X = ?Y => absurd (X < Y) end.
+ apply Z.le_ngt; rewrite <- VV; auto with zarith.
+ apply Z.lt_le_trans with wB; auto with zarith.
match goal with |- _ <= ?X + _ =>
- apply Zle_trans with X; auto with zarith
+ apply Z.le_trans with X; auto with zarith
end.
match goal with |- _ <= _ * ?X =>
- apply Zle_trans with (1 * X); auto with zarith
+ apply Z.le_trans with (1 * X); auto with zarith
end.
autorewrite with rm10.
- rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith.
+ rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; auto with zarith.
rewrite <- V3 in VV; generalize VV; autorewrite with rm10;
clear VV; intros VV.
rewrite spec_ww_add_c; auto with zarith.
@@ -1067,7 +1062,7 @@ intros x; case x; simpl ww_is_even.
simpl ww_to_Z in H1; rewrite H1.
rewrite <- Hw0.
match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
- apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
+ transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
end.
repeat rewrite Zsquare_mult.
rewrite wwB_wBwB; ring.
@@ -1079,41 +1074,41 @@ intros x; case x; simpl ww_is_even.
simpl ww_to_Z; unfold ww_to_Z.
rewrite spec_w_Bm1; auto with zarith.
split.
- rewrite wwB_wBwB; rewrite Zpower_2.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
match goal with |- _ <= -?X + (2 * (?Z * ?T + ?U) + ?V) =>
assert (X <= 2 * Z * T); auto with zarith
end.
- apply Zmult_le_compat_r; auto with zarith.
- rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith.
- rewrite Zmult_plus_distr_r; auto with zarith.
- rewrite Zmult_assoc; auto with zarith.
+ apply Z.mul_le_mono_nonneg_r; auto with zarith.
+ rewrite <- wB_div_2; apply Z.mul_le_mono_nonneg_l; auto with zarith.
+ rewrite Z.mul_add_distr_l; auto with zarith.
+ rewrite Z.mul_assoc; auto with zarith.
match goal with |- _ + ?X < _ =>
replace X with ((2 * (([|w4|]) + 1) * wB) - 1); try ring
end.
assert (2 * ([|w4|] + 1) * wB <= 2 * wwB); auto with zarith.
- rewrite <- Zmult_assoc; apply Zmult_le_compat_l; auto with zarith.
- rewrite wwB_wBwB; rewrite Zpower_2.
- apply Zmult_le_compat_r; auto with zarith.
+ rewrite <- Z.mul_assoc; apply Z.mul_le_mono_nonneg_l; auto with zarith.
+ rewrite wwB_wBwB; rewrite Z.pow_2_r.
+ apply Z.mul_le_mono_nonneg_r; auto with zarith.
case (spec_to_Z w4);auto with zarith.
Qed.
Lemma spec_ww_is_zero: forall x,
if ww_is_zero x then [[x]] = 0 else 0 < [[x]].
intro x; unfold ww_is_zero.
- rewrite spec_ww_compare. case Zcompare_spec;
+ rewrite spec_ww_compare. case Z.compare_spec;
auto with zarith.
simpl ww_to_Z.
assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z x);auto with zarith.
Qed.
Lemma wwB_4_2: 2 * (wwB / 4) = wwB/ 2.
- pattern wwB at 1; rewrite wwB_wBwB; rewrite Zpower_2.
+ pattern wwB at 1; rewrite wwB_wBwB; rewrite Z.pow_2_r.
rewrite <- wB_div_2.
match goal with |- context[(2 * ?X) * (2 * ?Z)] =>
replace ((2 * X) * (2 * Z)) with ((X * Z) * 4); try ring
end.
rewrite Z_div_mult; auto with zarith.
- rewrite Zmult_assoc; rewrite wB_div_2.
+ rewrite Z.mul_assoc; rewrite wB_div_2.
rewrite wwB_div_2; ring.
Qed.
@@ -1129,10 +1124,10 @@ Qed.
intros H2.
generalize (spec_ww_head0 x H2); case (ww_head0 x); autorewrite with rm10.
intros (H3, H4); split; auto with zarith.
- apply Zle_trans with (2 := H3).
+ apply Z.le_trans with (2 := H3).
apply Zdiv_le_compat_l; auto with zarith.
intros xh xl (H3, H4); split; auto with zarith.
- apply Zle_trans with (2 := H3).
+ apply Z.le_trans with (2 := H3).
apply Zdiv_le_compat_l; auto with zarith.
intros H1.
case (spec_to_w_Z (ww_head0 x)); intros Hv1 Hv2.
@@ -1156,24 +1151,24 @@ Qed.
case (spec_ww_head0 x); auto; intros Hv3 Hv4.
assert (Hu: forall u, 0 < u -> 2 * 2 ^ (u - 1) = 2 ^u).
intros u Hu.
- pattern 2 at 1; rewrite <- Zpower_1_r.
+ pattern 2 at 1; rewrite <- Z.pow_1_r.
rewrite <- Zpower_exp; auto with zarith.
ring_simplify (1 + (u - 1)); auto with zarith.
split; auto with zarith.
- apply Zmult_le_reg_r with 2; auto with zarith.
- repeat rewrite (fun x => Zmult_comm x 2).
+ apply Z.mul_le_mono_pos_r with 2; auto with zarith.
+ repeat rewrite (fun x => Z.mul_comm x 2).
rewrite wwB_4_2.
- rewrite Zmult_assoc; rewrite Hu; auto with zarith.
- apply Zle_lt_trans with (2 * 2 ^ ([[ww_head0 x]] - 1) * [[x]]); auto with zarith;
+ rewrite Z.mul_assoc; rewrite Hu; auto with zarith.
+ apply Z.le_lt_trans with (2 * 2 ^ ([[ww_head0 x]] - 1) * [[x]]); auto with zarith;
rewrite Hu; auto with zarith.
- apply Zmult_le_compat_r; auto with zarith.
+ apply Z.mul_le_mono_nonneg_r; auto with zarith.
apply Zpower_le_monotone; auto with zarith.
Qed.
Theorem wwB_4_wB_4: wwB / 4 = wB / 4 * wB.
- apply sym_equal; apply Zdiv_unique with 0;
- auto with zarith.
- rewrite Zmult_assoc; rewrite wB_div_4; auto with zarith.
+ Proof.
+ symmetry; apply Zdiv_unique with 0; auto with zarith.
+ rewrite Z.mul_assoc; rewrite wB_div_4; auto with zarith.
rewrite wwB_wBwB; ring.
Qed.
@@ -1182,10 +1177,10 @@ Qed.
assert (U := wB_pos w_digits).
intro x; unfold ww_sqrt.
generalize (spec_ww_is_zero x); case (ww_is_zero x).
- simpl ww_to_Z; simpl Zpower; unfold Zpower_pos; simpl;
+ simpl ww_to_Z; simpl Z.pow; unfold Z.pow_pos; simpl;
auto with zarith.
intros H1.
- rewrite spec_ww_compare. case Zcompare_spec;
+ rewrite spec_ww_compare. case Z.compare_spec;
simpl ww_to_Z; autorewrite with rm10.
generalize H1; case x.
intros HH; contradict HH; simpl ww_to_Z; auto with zarith.
@@ -1203,7 +1198,7 @@ Qed.
intros w3 (H6, H7); rewrite H6.
assert (V1 := spec_to_Z w3);auto with zarith.
split; auto with zarith.
- apply Zle_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith.
+ apply Z.le_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith.
match goal with |- ?X < ?Z =>
replace Z with (X + 1); auto with zarith
end.
@@ -1211,7 +1206,7 @@ Qed.
intros w3 (H6, H7); rewrite H6.
assert (V1 := spec_to_Z w3);auto with zarith.
split; auto with zarith.
- apply Zle_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith.
+ apply Z.le_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith.
match goal with |- ?X < ?Z =>
replace Z with (X + 1); auto with zarith
end.
@@ -1221,42 +1216,42 @@ Qed.
case (spec_ww_head1 x); intros Hp1 Hp2.
generalize (Hp2 H1); clear Hp2; intros Hp2.
assert (Hv2: [[ww_head1 x]] <= Zpos (xO w_digits)).
- case (Zle_or_lt (Zpos (xO w_digits)) [[ww_head1 x]]); auto with zarith; intros HH1.
+ case (Z.le_gt_cases (Zpos (xO w_digits)) [[ww_head1 x]]); auto with zarith; intros HH1.
case Hp2; intros _ HH2; contradict HH2.
- apply Zle_not_lt; unfold base.
- apply Zle_trans with (2 ^ [[ww_head1 x]]).
+ apply Z.le_ngt; unfold base.
+ apply Z.le_trans with (2 ^ [[ww_head1 x]]).
apply Zpower_le_monotone; auto with zarith.
pattern (2 ^ [[ww_head1 x]]) at 1;
- rewrite <- (Zmult_1_r (2 ^ [[ww_head1 x]])).
- apply Zmult_le_compat_l; auto with zarith.
+ rewrite <- (Z.mul_1_r (2 ^ [[ww_head1 x]])).
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
generalize (spec_ww_add_mul_div x W0 (ww_head1 x) Hv2);
case ww_add_mul_div.
simpl ww_to_Z; autorewrite with w_rewrite rm10.
rewrite Zmod_small; auto with zarith.
- intros H2; case (Zmult_integral _ _ (sym_equal H2)); clear H2; intros H2.
- rewrite H2; unfold Zpower, Zpower_pos; simpl; auto with zarith.
+ intros H2. symmetry in H2. rewrite Z.mul_eq_0 in H2. destruct H2 as [H2|H2].
+ rewrite H2; unfold Z.pow, Z.pow_pos; simpl; auto with zarith.
match type of H2 with ?X = ?Y =>
absurd (Y < X); try (rewrite H2; auto with zarith; fail)
end.
- apply Zpower_gt_0; auto with zarith.
+ apply Z.pow_pos_nonneg; auto with zarith.
split; auto with zarith.
- case Hp2; intros _ tmp; apply Zle_lt_trans with (2 := tmp);
+ case Hp2; intros _ tmp; apply Z.le_lt_trans with (2 := tmp);
clear tmp.
- rewrite Zmult_comm; apply Zmult_le_compat_r; auto with zarith.
+ rewrite Z.mul_comm; apply Z.mul_le_mono_nonneg_r; auto with zarith.
assert (Hv0: [[ww_head1 x]] = 2 * ([[ww_head1 x]]/2)).
pattern [[ww_head1 x]] at 1; rewrite (Z_div_mod_eq [[ww_head1 x]] 2);
auto with zarith.
generalize (spec_ww_is_even (ww_head1 x)); rewrite Hp1;
- intros tmp; rewrite tmp; rewrite Zplus_0_r; auto.
+ intros tmp; rewrite tmp; rewrite Z.add_0_r; auto.
intros w0 w1; autorewrite with w_rewrite rm10.
rewrite Zmod_small; auto with zarith.
- 2: rewrite Zmult_comm; auto with zarith.
+ 2: rewrite Z.mul_comm; auto with zarith.
intros H2.
assert (V: wB/4 <= [|w0|]).
apply beta_lex with 0 [|w1|] wB; auto with zarith; autorewrite with rm10.
simpl ww_to_Z in H2; rewrite H2.
rewrite <- wwB_4_wB_4; auto with zarith.
- rewrite Zmult_comm; auto with zarith.
+ rewrite Z.mul_comm; auto with zarith.
assert (V1 := spec_to_Z w1);auto with zarith.
generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith.
case (w_sqrt2 w0 w1); intros w2 c.
@@ -1267,13 +1262,13 @@ Qed.
rewrite spec_ww_pred; rewrite spec_ww_zdigits.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
- apply Zlt_le_trans with (Zpos (xO w_digits)); auto with zarith.
+ apply Z.lt_le_trans with (Zpos (xO w_digits)); auto with zarith.
unfold base; apply Zpower2_le_lin; auto with zarith.
assert (Hv4: [[ww_head1 x]]/2 < wB).
- apply Zle_lt_trans with (Zpos w_digits).
- apply Zmult_le_reg_r with 2; auto with zarith.
- repeat rewrite (fun x => Zmult_comm x 2).
- rewrite <- Hv0; rewrite <- Zpos_xO; auto.
+ apply Z.le_lt_trans with (Zpos w_digits).
+ apply Z.mul_le_mono_pos_r with 2; auto with zarith.
+ repeat rewrite (fun x => Z.mul_comm x 2).
+ rewrite <- Hv0; rewrite <- Pos2Z.inj_xO; auto.
unfold base; apply Zpower2_lt_lin; auto with zarith.
assert (Hv5: [[(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))]]
= [[ww_head1 x]]/2).
@@ -1281,12 +1276,12 @@ Qed.
simpl ww_to_Z; autorewrite with rm10.
rewrite Hv3.
ring_simplify (Zpos (xO w_digits) - (Zpos (xO w_digits) - 1)).
- rewrite Zpower_1_r.
+ rewrite Z.pow_1_r.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
- apply Zlt_le_trans with (1 := Hv4); auto with zarith.
+ apply Z.lt_le_trans with (1 := Hv4); auto with zarith.
unfold base; apply Zpower_le_monotone; auto with zarith.
- split; unfold ww_digits; try rewrite Zpos_xO; auto with zarith.
+ split; unfold ww_digits; try rewrite Pos2Z.inj_xO; auto with zarith.
rewrite Hv3; auto with zarith.
assert (Hv6: [|low(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))|]
= [[ww_head1 x]]/2).
@@ -1302,13 +1297,13 @@ Qed.
rewrite Zmod_small.
simpl ww_to_Z in H2; rewrite H2; auto with zarith.
intros (H4, H5); split.
- apply Zmult_le_reg_r with (2 ^ [[ww_head1 x]]); auto with zarith.
+ apply Z.mul_le_mono_pos_r with (2 ^ [[ww_head1 x]]); auto with zarith.
rewrite H4.
- apply Zle_trans with ([|w2|] ^ 2); auto with zarith.
- rewrite Zmult_comm.
+ apply Z.le_trans with ([|w2|] ^ 2); auto with zarith.
+ rewrite Z.mul_comm.
pattern [[ww_head1 x]] at 1;
rewrite Hv0; auto with zarith.
- rewrite (Zmult_comm 2); rewrite Zpower_mult;
+ rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r;
auto with zarith.
assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2);
try (intros; repeat rewrite Zsquare_mult; ring);
@@ -1324,17 +1319,17 @@ Qed.
case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]] / 2))); auto with zarith.
case c; unfold interp_carry; autorewrite with rm10;
intros w3; assert (V3 := spec_to_Z w3);auto with zarith.
- apply Zmult_lt_reg_r with (2 ^ [[ww_head1 x]]); auto with zarith.
+ apply Z.mul_lt_mono_pos_r with (2 ^ [[ww_head1 x]]); auto with zarith.
rewrite H4.
- apply Zle_lt_trans with ([|w2|] ^ 2 + 2 * [|w2|]); auto with zarith.
- apply Zlt_le_trans with (([|w2|] + 1) ^ 2); auto with zarith.
+ apply Z.le_lt_trans with ([|w2|] ^ 2 + 2 * [|w2|]); auto with zarith.
+ apply Z.lt_le_trans with (([|w2|] + 1) ^ 2); auto with zarith.
match goal with |- ?X < ?Y =>
replace Y with (X + 1); auto with zarith
end.
repeat rewrite (Zsquare_mult); ring.
- rewrite Zmult_comm.
+ rewrite Z.mul_comm.
pattern [[ww_head1 x]] at 1; rewrite Hv0.
- rewrite (Zmult_comm 2); rewrite Zpower_mult;
+ rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r;
auto with zarith.
assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2);
try (intros; repeat rewrite Zsquare_mult; ring);
@@ -1343,20 +1338,20 @@ Qed.
split; auto with zarith.
pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]]/2)));
auto with zarith.
- rewrite <- Zplus_assoc; rewrite Zmult_plus_distr_r.
- autorewrite with rm10; apply Zplus_le_compat_l; auto with zarith.
+ rewrite <- Z.add_assoc; rewrite Z.mul_add_distr_l.
+ autorewrite with rm10; apply Z.add_le_mono_l; auto with zarith.
case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]]/2))); auto with zarith.
split; auto with zarith.
- apply Zle_lt_trans with ([|w2|]); auto with zarith.
+ apply Z.le_lt_trans with ([|w2|]); auto with zarith.
apply Zdiv_le_upper_bound; auto with zarith.
pattern [|w2|] at 1; replace [|w2|] with ([|w2|] * 2 ^0);
auto with zarith.
- apply Zmult_le_compat_l; auto with zarith.
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
apply Zpower_le_monotone; auto with zarith.
- rewrite Zpower_0_r; autorewrite with rm10; auto.
+ rewrite Z.pow_0_r; autorewrite with rm10; auto.
split; auto with zarith.
- rewrite Hv0 in Hv2; rewrite (Zpos_xO w_digits) in Hv2; auto with zarith.
- apply Zle_lt_trans with (Zpos w_digits); auto with zarith.
+ rewrite Hv0 in Hv2; rewrite (Pos2Z.inj_xO w_digits) in Hv2; auto with zarith.
+ apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_lt_lin; auto with zarith.
rewrite spec_w_sub; auto with zarith.
rewrite Hv6; rewrite spec_w_zdigits; auto with zarith.
@@ -1364,10 +1359,10 @@ Qed.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
assert ([[ww_head1 x]]/2 <= Zpos w_digits); auto with zarith.
- apply Zmult_le_reg_r with 2; auto with zarith.
- repeat rewrite (fun x => Zmult_comm x 2).
- rewrite <- Hv0; rewrite <- Zpos_xO; auto with zarith.
- apply Zle_lt_trans with (Zpos w_digits); auto with zarith.
+ apply Z.mul_le_mono_pos_r with 2; auto with zarith.
+ repeat rewrite (fun x => Z.mul_comm x 2).
+ rewrite <- Hv0; rewrite <- Pos2Z.inj_xO; auto with zarith.
+ apply Z.le_lt_trans with (Zpos w_digits); auto with zarith.
unfold base; apply Zpower2_lt_lin; auto with zarith.
Qed.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
index e63e20c69..46a163cde 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
@@ -195,9 +195,9 @@ Section DoubleSub.
Lemma spec_ww_opp_c : forall x, [-[ww_opp_c x]] = -[[x]].
Proof.
destruct x as [ |xh xl];simpl. reflexivity.
- rewrite Zopp_plus_distr;generalize (spec_opp_c xl);destruct (w_opp_c xl)
+ rewrite Z.opp_add_distr;generalize (spec_opp_c xl);destruct (w_opp_c xl)
as [l|l];intros H;unfold interp_carry in H;rewrite <- H;
- rewrite Zopp_mult_distr_l.
+ rewrite <- Z.mul_opp_l.
assert ([|l|] = 0).
assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega.
rewrite H0;generalize (spec_opp_c xh);destruct (w_opp_c xh)
@@ -213,13 +213,13 @@ Section DoubleSub.
Lemma spec_ww_opp : forall x, [[ww_opp x]] = (-[[x]]) mod wwB.
Proof.
destruct x as [ |xh xl];simpl. reflexivity.
- rewrite Zopp_plus_distr;rewrite Zopp_mult_distr_l.
+ rewrite Z.opp_add_distr, <- Z.mul_opp_l.
generalize (spec_opp_c xl);destruct (w_opp_c xl)
as [l|l];intros H;unfold interp_carry in H;rewrite <- H;simpl ww_to_Z.
- rewrite spec_w_0;rewrite Zplus_0_r;rewrite wwB_wBwB.
+ rewrite spec_w_0;rewrite Z.add_0_r;rewrite wwB_wBwB.
assert ([|l|] = 0).
assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega.
- rewrite H0;rewrite Zplus_0_r; rewrite Zpower_2;
+ rewrite H0;rewrite Z.add_0_r; rewrite Z.pow_2_r;
rewrite Zmult_mod_distr_r;try apply lt_0_wB.
rewrite spec_opp;trivial.
apply Zmod_unique with (q:= -1).
@@ -240,7 +240,7 @@ Section DoubleSub.
simpl ww_to_Z;replace (([|xh|]*wB+[|xl|])-1) with ([|xh|]*wB+([|xl|]-1)).
2:ring. generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l];
intros H;unfold interp_carry in H;rewrite <- H. simpl;apply spec_w_WW.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
assert ([|l|] = wB - 1).
assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega.
rewrite H0;change ([|xh|] + -1) with ([|xh|] - 1).
@@ -263,7 +263,7 @@ Section DoubleSub.
generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1;
unfold interp_carry in H1;rewrite <- H1;unfold interp_carry;
try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1).
generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h];
intros H1;unfold interp_carry in *;rewrite <- H1;simpl ww_to_Z;
@@ -274,7 +274,7 @@ Section DoubleSub.
forall x y, [-[ww_sub_carry_c x y]] = [[x]] - [[y]] - 1.
Proof.
destruct y as [ |yh yl];simpl.
- unfold Zminus;simpl;rewrite Zplus_0_r;exact (spec_ww_pred_c x).
+ unfold Z.sub;simpl;rewrite Z.add_0_r;exact (spec_ww_pred_c x).
destruct x as [ |xh xl].
unfold interp_carry;rewrite spec_w_WW;simpl ww_to_Z;rewrite wwB_wBwB;
repeat rewrite spec_opp_carry;ring.
@@ -286,7 +286,7 @@ Section DoubleSub.
generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1;
unfold interp_carry in H1;rewrite <- H1;unfold interp_carry;
try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1).
generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h];
intros H1;unfold interp_carry in *;rewrite <- H1;try rewrite spec_w_WW;
@@ -303,7 +303,7 @@ Section DoubleSub.
unfold interp_carry in H;rewrite <- H;simpl ww_to_Z.
rewrite Zmod_small. apply spec_w_WW.
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh l)).
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
change ([|xh|] + -1) with ([|xh|] - 1).
assert ([|l|] = wB - 1).
assert (H1:= spec_to_Z l);assert (H2:= spec_to_Z xl);omega.
@@ -322,7 +322,7 @@ Section DoubleSub.
unfold interp_carry in H;rewrite <- H.
rewrite spec_w_WW;rewrite (mod_wwB w_digits w_to_Z spec_to_Z).
rewrite spec_sub;trivial.
- simpl ww_to_Z;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ simpl ww_to_Z;rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial.
Qed.
@@ -341,7 +341,7 @@ Section DoubleSub.
generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl)as[l|l];
intros H;unfold interp_carry in H;rewrite <- H;rewrite spec_w_WW.
rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub;trivial.
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Z.add_assoc;rewrite <- Z.mul_add_distr_r.
rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial.
Qed.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
index a274b8395..948203822 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
@@ -13,7 +13,7 @@ Set Implicit Arguments.
Require Import ZArith.
Local Open Scope Z_scope.
-Definition base digits := Zpower 2 (Zpos digits).
+Definition base digits := Z.pow 2 (Zpos digits).
Section Carry.
diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v
index 2dd1c3eec..8ed59632e 100644
--- a/theories/Numbers/Cyclic/Int31/Cyclic31.v
+++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v
@@ -368,7 +368,7 @@ Section Basics.
(** Variant of [phi] via [recrbis] *)
Let Phi := fun b (_:int31) =>
- match b with D0 => Zdouble | D1 => Zdouble_plus_one end.
+ match b with D0 => Z.double | D1 => Z.succ_double end.
Definition phibis_aux n x := recrbis_aux n _ Z0 Phi x.
@@ -381,7 +381,7 @@ Section Basics.
(** Recursive equations satisfied by [phi] *)
Lemma phi_eqn1 : forall x, firstr x = D0 ->
- phi x = Zdouble (phi (shiftr x)).
+ phi x = Z.double (phi (shiftr x)).
Proof.
intros.
case_eq (iszero x); intros.
@@ -391,7 +391,7 @@ Section Basics.
Qed.
Lemma phi_eqn2 : forall x, firstr x = D1 ->
- phi x = Zdouble_plus_one (phi (shiftr x)).
+ phi x = Z.succ_double (phi (shiftr x)).
Proof.
intros.
case_eq (iszero x); intros.
@@ -401,7 +401,7 @@ Section Basics.
Qed.
Lemma phi_twice_firstl : forall x, firstl x = D0 ->
- phi (twice x) = Zdouble (phi x).
+ phi (twice x) = Z.double (phi x).
Proof.
intros.
rewrite phi_eqn1; auto; [ | destruct x; auto ].
@@ -410,7 +410,7 @@ Section Basics.
Qed.
Lemma phi_twice_plus_one_firstl : forall x, firstl x = D0 ->
- phi (twice_plus_one x) = Zdouble_plus_one (phi x).
+ phi (twice_plus_one x) = Z.succ_double (phi x).
Proof.
intros.
rewrite phi_eqn2; auto; [ | destruct x; auto ].
@@ -430,13 +430,13 @@ Section Basics.
unfold phibis_aux, recrbis_aux; fold recrbis_aux;
fold (phibis_aux n (shiftr x)).
destruct (firstr x).
- specialize IHn with (shiftr x); rewrite Zdouble_mult; omega.
- specialize IHn with (shiftr x); rewrite Zdouble_plus_one_mult; omega.
+ specialize IHn with (shiftr x); rewrite Z.double_spec; omega.
+ specialize IHn with (shiftr x); rewrite Z.succ_double_spec; omega.
Qed.
Lemma phibis_aux_bounded :
forall n x, n <= size ->
- (phibis_aux n (nshiftr (size-n) x) < 2 ^ (Z_of_nat n))%Z.
+ (phibis_aux n (nshiftr (size-n) x) < 2 ^ (Z.of_nat n))%Z.
Proof.
induction n.
simpl; unfold phibis_aux; simpl; auto with zarith.
@@ -450,13 +450,13 @@ Section Basics.
assert (H1 : n <= size) by omega.
specialize (IHn x H1).
set (y:=phibis_aux n (nshiftr (size - n) x)) in *.
- rewrite inj_S, Zpower_Zsucc; auto with zarith.
+ rewrite Nat2Z.inj_succ, Z.pow_succ_r; auto with zarith.
case_eq (firstr (nshiftr (size - S n) x)); intros.
- rewrite Zdouble_mult; auto with zarith.
- rewrite Zdouble_plus_one_mult; auto with zarith.
+ rewrite Z.double_spec; auto with zarith.
+ rewrite Z.succ_double_spec; auto with zarith.
Qed.
- Lemma phi_bounded : forall x, (0 <= phi x < 2 ^ (Z_of_nat size))%Z.
+ Lemma phi_bounded : forall x, (0 <= phi x < 2 ^ (Z.of_nat size))%Z.
Proof.
intros.
rewrite <- phibis_aux_equiv.
@@ -468,32 +468,32 @@ Section Basics.
Lemma phibis_aux_lowerbound :
forall n x, firstr (nshiftr n x) = D1 ->
- (2 ^ Z_of_nat n <= phibis_aux (S n) x)%Z.
+ (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z.
Proof.
induction n.
intros.
unfold nshiftr in H; simpl in *.
unfold phibis_aux, recrbis_aux.
- rewrite H, Zdouble_plus_one_mult; omega.
+ rewrite H, Z.succ_double_spec; omega.
intros.
remember (S n) as m.
unfold phibis_aux, recrbis_aux; fold recrbis_aux;
fold (phibis_aux m (shiftr x)).
subst m.
- rewrite inj_S, Zpower_Zsucc; auto with zarith.
- assert (2^(Z_of_nat n) <= phibis_aux (S n) (shiftr x))%Z.
+ rewrite Nat2Z.inj_succ, Z.pow_succ_r; auto with zarith.
+ assert (2^(Z.of_nat n) <= phibis_aux (S n) (shiftr x))%Z.
apply IHn.
rewrite <- nshiftr_S_tail; auto.
destruct (firstr x).
- change (Zdouble (phibis_aux (S n) (shiftr x))) with
+ change (Z.double (phibis_aux (S n) (shiftr x))) with
(2*(phibis_aux (S n) (shiftr x)))%Z.
omega.
- rewrite Zdouble_plus_one_mult; omega.
+ rewrite Z.succ_double_spec; omega.
Qed.
Lemma phi_lowerbound :
- forall x, firstl x = D1 -> (2^(Z_of_nat (pred size)) <= phi x)%Z.
+ forall x, firstl x = D1 -> (2^(Z.of_nat (pred size)) <= phi x)%Z.
Proof.
intros.
generalize (phibis_aux_lowerbound (pred size) x).
@@ -776,7 +776,7 @@ Section Basics.
(** First, recursive equations *)
Lemma phi_inv_double_plus_one : forall z,
- phi_inv (Zdouble_plus_one z) = twice_plus_one (phi_inv z).
+ phi_inv (Z.succ_double z) = twice_plus_one (phi_inv z).
Proof.
destruct z; simpl; auto.
induction p; simpl.
@@ -788,20 +788,20 @@ Section Basics.
Qed.
Lemma phi_inv_double : forall z,
- phi_inv (Zdouble z) = twice (phi_inv z).
+ phi_inv (Z.double z) = twice (phi_inv z).
Proof.
destruct z; simpl; auto.
rewrite incr_twice_plus_one; auto.
Qed.
Lemma phi_inv_incr : forall z,
- phi_inv (Zsucc z) = incr (phi_inv z).
+ phi_inv (Z.succ z) = incr (phi_inv z).
Proof.
destruct z.
simpl; auto.
simpl; auto.
induction p; simpl; auto.
- rewrite Pplus_one_succ_r, IHp, incr_twice_plus_one; auto.
+ rewrite <- Pos.add_1_r, IHp, incr_twice_plus_one; auto.
rewrite incr_twice; auto.
simpl; auto.
destruct p; simpl; auto.
@@ -908,15 +908,15 @@ Section Basics.
Local Open Scope Z_scope.
Lemma p2ibis_spec : forall n p, (n<=size)%nat ->
- Zpos p = (Z_of_N (fst (p2ibis n p)))*2^(Z_of_nat n) +
+ Zpos p = (Z.of_N (fst (p2ibis n p)))*2^(Z.of_nat n) +
phi (snd (p2ibis n p)).
Proof.
induction n; intros.
- simpl; rewrite Pmult_1_r; auto.
- replace (2^(Z_of_nat (S n)))%Z with (2*2^(Z_of_nat n))%Z by
- (rewrite <- Zpower_Zsucc, <- Zpos_P_of_succ_nat;
+ simpl; rewrite Pos.mul_1_r; auto.
+ replace (2^(Z.of_nat (S n)))%Z with (2*2^(Z.of_nat n))%Z by
+ (rewrite <- Z.pow_succ_r, <- Zpos_P_of_succ_nat;
auto with zarith).
- rewrite (Zmult_comm 2).
+ rewrite (Z.mul_comm 2).
assert (n<=size)%nat by omega.
destruct p; simpl; [ | | auto];
specialize (IHn p H0);
@@ -924,13 +924,13 @@ Section Basics.
destruct (p2ibis n p) as (r,i); simpl in *; intros.
change (Zpos p~1) with (2*Zpos p + 1)%Z.
- rewrite phi_twice_plus_one_firstl, Zdouble_plus_one_mult.
+ rewrite phi_twice_plus_one_firstl, Z.succ_double_spec.
rewrite IHn; ring.
apply (nshiftr_0_firstl n); auto; try omega.
change (Zpos p~0) with (2*Zpos p)%Z.
rewrite phi_twice_firstl.
- change (Zdouble (phi i)) with (2*(phi i))%Z.
+ change (Z.double (phi i)) with (2*(phi i))%Z.
rewrite IHn; ring.
apply (nshiftr_0_firstl n); auto; try omega.
Qed.
@@ -956,12 +956,12 @@ Section Basics.
for the positive case. *)
Lemma phi_phi_inv_positive : forall p,
- phi (phi_inv_positive p) = (Zpos p) mod (2^(Z_of_nat size)).
+ phi (phi_inv_positive p) = (Zpos p) mod (2^(Z.of_nat size)).
Proof.
intros.
replace (phi_inv_positive p) with (snd (p2ibis size p)).
rewrite (p2ibis_spec size p) by auto.
- rewrite Zplus_comm, Z_mod_plus.
+ rewrite Z.add_comm, Z_mod_plus.
symmetry; apply Zmod_small.
apply phi_bounded.
auto with zarith.
@@ -978,7 +978,7 @@ Section Basics.
Proof.
intros.
unfold mul31.
- rewrite <- Zdouble_mult, <- phi_twice_firstl, phi_inv_phi; auto.
+ rewrite <- Z.double_spec, <- phi_twice_firstl, phi_inv_phi; auto.
Qed.
Lemma double_twice_plus_one_firstl : forall x, firstl x = D0 ->
@@ -987,7 +987,7 @@ Section Basics.
intros.
rewrite double_twice_firstl; auto.
unfold add31.
- rewrite phi_twice_firstl, <- Zdouble_plus_one_mult,
+ rewrite phi_twice_firstl, <- Z.succ_double_spec,
<- phi_twice_plus_one_firstl, phi_inv_phi; auto.
Qed.
@@ -1016,7 +1016,7 @@ Section Basics.
Qed.
Lemma positive_to_int31_spec : forall p,
- Zpos p = (Z_of_N (fst (positive_to_int31 p)))*2^(Z_of_nat size) +
+ Zpos p = (Z.of_N (fst (positive_to_int31 p)))*2^(Z.of_nat size) +
phi (snd (positive_to_int31 p)).
Proof.
unfold positive_to_int31.
@@ -1029,43 +1029,43 @@ Section Basics.
[phi o twice] and so one. *)
Lemma phi_twice : forall x,
- phi (twice x) = (Zdouble (phi x)) mod 2^(Z_of_nat size).
+ phi (twice x) = (Z.double (phi x)) mod 2^(Z.of_nat size).
Proof.
intros.
pattern x at 1; rewrite <- (phi_inv_phi x).
rewrite <- phi_inv_double.
- assert (0 <= Zdouble (phi x)).
- rewrite Zdouble_mult; generalize (phi_bounded x); omega.
- destruct (Zdouble (phi x)).
+ assert (0 <= Z.double (phi x)).
+ rewrite Z.double_spec; generalize (phi_bounded x); omega.
+ destruct (Z.double (phi x)).
simpl; auto.
apply phi_phi_inv_positive.
compute in H; elim H; auto.
Qed.
Lemma phi_twice_plus_one : forall x,
- phi (twice_plus_one x) = (Zdouble_plus_one (phi x)) mod 2^(Z_of_nat size).
+ phi (twice_plus_one x) = (Z.succ_double (phi x)) mod 2^(Z.of_nat size).
Proof.
intros.
pattern x at 1; rewrite <- (phi_inv_phi x).
rewrite <- phi_inv_double_plus_one.
- assert (0 <= Zdouble_plus_one (phi x)).
- rewrite Zdouble_plus_one_mult; generalize (phi_bounded x); omega.
- destruct (Zdouble_plus_one (phi x)).
+ assert (0 <= Z.succ_double (phi x)).
+ rewrite Z.succ_double_spec; generalize (phi_bounded x); omega.
+ destruct (Z.succ_double (phi x)).
simpl; auto.
apply phi_phi_inv_positive.
compute in H; elim H; auto.
Qed.
Lemma phi_incr : forall x,
- phi (incr x) = (Zsucc (phi x)) mod 2^(Z_of_nat size).
+ phi (incr x) = (Z.succ (phi x)) mod 2^(Z.of_nat size).
Proof.
intros.
pattern x at 1; rewrite <- (phi_inv_phi x).
rewrite <- phi_inv_incr.
- assert (0 <= Zsucc (phi x)).
- change (Zsucc (phi x)) with ((phi x)+1)%Z;
+ assert (0 <= Z.succ (phi x)).
+ change (Z.succ (phi x)) with ((phi x)+1)%Z;
generalize (phi_bounded x); omega.
- destruct (Zsucc (phi x)).
+ destruct (Z.succ (phi x)).
simpl; auto.
apply phi_phi_inv_positive.
compute in H; elim H; auto.
@@ -1075,7 +1075,7 @@ Section Basics.
in the negative case *)
Lemma phi_phi_inv_negative :
- forall p, phi (incr (complement_negative p)) = (Zneg p) mod 2^(Z_of_nat size).
+ forall p, phi (incr (complement_negative p)) = (Zneg p) mod 2^(Z.of_nat size).
Proof.
induction p.
@@ -1083,21 +1083,21 @@ Section Basics.
rewrite phi_incr in IHp.
rewrite incr_twice, phi_twice_plus_one.
remember (phi (complement_negative p)) as q.
- rewrite Zdouble_plus_one_mult.
- replace (2*q+1) with (2*(Zsucc q)-1) by omega.
+ rewrite Z.succ_double_spec.
+ replace (2*q+1) with (2*(Z.succ q)-1) by omega.
rewrite <- Zminus_mod_idemp_l, <- Zmult_mod_idemp_r, IHp.
rewrite Zmult_mod_idemp_r, Zminus_mod_idemp_l; auto with zarith.
simpl complement_negative.
rewrite incr_twice_plus_one, phi_twice.
remember (phi (incr (complement_negative p))) as q.
- rewrite Zdouble_mult, IHp, Zmult_mod_idemp_r; auto with zarith.
+ rewrite Z.double_spec, IHp, Zmult_mod_idemp_r; auto with zarith.
simpl; auto.
Qed.
Lemma phi_phi_inv :
- forall z, phi (phi_inv z) = z mod 2 ^ (Z_of_nat size).
+ forall z, phi (phi_inv z) = z mod 2 ^ (Z.of_nat size).
Proof.
destruct z.
simpl; auto.
@@ -1167,7 +1167,7 @@ Section Int31_Specs.
Notation "[| x |]" := (phi x) (at level 0, x at level 99).
- Local Notation wB := (2 ^ (Z_of_nat size)).
+ Local Notation wB := (2 ^ (Z.of_nat size)).
Lemma wB_pos : wB > 0.
Proof.
@@ -1221,14 +1221,14 @@ Section Int31_Specs.
set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y.
assert ((X+Y) mod wB ?= X+Y <> Eq -> [+|C1 (phi_inv (X+Y))|] = X+Y).
- unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros.
+ unfold interp_carry; rewrite phi_phi_inv, Z.compare_eq_iff; intros.
destruct (Z_lt_le_dec (X+Y) wB).
contradict H1; auto using Zmod_small with zarith.
rewrite <- (Z_mod_plus_full (X+Y) (-1) wB).
rewrite Zmod_small; romega.
- generalize (Zcompare_Eq_eq ((X+Y) mod wB) (X+Y)); intros Heq.
- destruct Zcompare; intros;
+ generalize (Z.compare_eq ((X+Y) mod wB) (X+Y)); intros Heq.
+ destruct Z.compare; intros;
[ rewrite phi_phi_inv; auto | now apply H1 | now apply H1].
Qed.
@@ -1245,14 +1245,14 @@ Section Int31_Specs.
set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y.
assert ((X+Y+1) mod wB ?= X+Y+1 <> Eq -> [+|C1 (phi_inv (X+Y+1))|] = X+Y+1).
- unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros.
+ unfold interp_carry; rewrite phi_phi_inv, Z.compare_eq_iff; intros.
destruct (Z_lt_le_dec (X+Y+1) wB).
contradict H1; auto using Zmod_small with zarith.
rewrite <- (Z_mod_plus_full (X+Y+1) (-1) wB).
rewrite Zmod_small; romega.
- generalize (Zcompare_Eq_eq ((X+Y+1) mod wB) (X+Y+1)); intros Heq.
- destruct Zcompare; intros;
+ generalize (Z.compare_eq ((X+Y+1) mod wB) (X+Y+1)); intros Heq.
+ destruct Z.compare; intros;
[ rewrite phi_phi_inv; auto | now apply H1 | now apply H1].
Qed.
@@ -1284,14 +1284,14 @@ Section Int31_Specs.
set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y.
assert ((X-Y) mod wB ?= X-Y <> Eq -> [-|C1 (phi_inv (X-Y))|] = X-Y).
- unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros.
+ unfold interp_carry; rewrite phi_phi_inv, Z.compare_eq_iff; intros.
destruct (Z_lt_le_dec (X-Y) 0).
rewrite <- (Z_mod_plus_full (X-Y) 1 wB).
rewrite Zmod_small; romega.
contradict H1; apply Zmod_small; romega.
- generalize (Zcompare_Eq_eq ((X-Y) mod wB) (X-Y)); intros Heq.
- destruct Zcompare; intros;
+ generalize (Z.compare_eq ((X-Y) mod wB) (X-Y)); intros Heq.
+ destruct Z.compare; intros;
[ rewrite phi_phi_inv; auto | now apply H1 | now apply H1].
Qed.
@@ -1303,14 +1303,14 @@ Section Int31_Specs.
set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y.
assert ((X-Y-1) mod wB ?= X-Y-1 <> Eq -> [-|C1 (phi_inv (X-Y-1))|] = X-Y-1).
- unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros.
+ unfold interp_carry; rewrite phi_phi_inv, Z.compare_eq_iff; intros.
destruct (Z_lt_le_dec (X-Y-1) 0).
rewrite <- (Z_mod_plus_full (X-Y-1) 1 wB).
rewrite Zmod_small; romega.
contradict H1; apply Zmod_small; romega.
- generalize (Zcompare_Eq_eq ((X-Y-1) mod wB) (X-Y-1)); intros Heq.
- destruct Zcompare; intros;
+ generalize (Z.compare_eq ((X-Y-1) mod wB) (X-Y-1)); intros Heq.
+ destruct Z.compare; intros;
[ rewrite phi_phi_inv; auto | now apply H1 | now apply H1].
Qed.
@@ -1386,7 +1386,7 @@ Section Int31_Specs.
apply Zmod_small.
generalize (phi_bounded x)(phi_bounded y); intros.
change (wB^2) with (wB * wB).
- auto using Zmult_lt_compat with zarith.
+ auto using Z.mul_lt_mono_nonneg with zarith.
Qed.
Lemma spec_mul : forall x y, [|x*y|] = ([|x|] * [|y|]) mod wB.
@@ -1412,29 +1412,26 @@ Section Int31_Specs.
generalize (phi_bounded a1)(phi_bounded a2)(phi_bounded b); intros.
assert ([|b|]>0) by (auto with zarith).
generalize (Z_div_mod (phi2 a1 a2) [|b|] H4) (Z_div_pos (phi2 a1 a2) [|b|] H4).
- unfold Zdiv; destruct (Zdiv_eucl (phi2 a1 a2) [|b|]); simpl.
+ unfold Z.div; destruct (Z.div_eucl (phi2 a1 a2) [|b|]); simpl.
rewrite ?phi_phi_inv.
destruct 1; intros.
unfold phi2 in *.
change base with wB; change base with wB in H5.
- change (Zpower_pos 2 31) with wB; change (Zpower_pos 2 31) with wB in H.
- rewrite H5, Zmult_comm.
+ change (Z.pow_pos 2 31) with wB; change (Z.pow_pos 2 31) with wB in H.
+ rewrite H5, Z.mul_comm.
replace (z0 mod wB) with z0 by (symmetry; apply Zmod_small; omega).
replace (z mod wB) with z; auto with zarith.
symmetry; apply Zmod_small.
split.
apply H7; change base with wB; auto with zarith.
- apply Zmult_gt_0_lt_reg_r with [|b|].
- omega.
- rewrite Zmult_comm.
- apply Zle_lt_trans with ([|b|]*z+z0).
- omega.
+ apply Z.mul_lt_mono_pos_r with [|b|]; [omega| ].
+ rewrite Z.mul_comm.
+ apply Z.le_lt_trans with ([|b|]*z+z0); [omega| ].
rewrite <- H5.
- apply Zle_lt_trans with ([|a1|]*wB+(wB-1)).
- omega.
+ apply Z.le_lt_trans with ([|a1|]*wB+(wB-1)); [omega | ].
replace ([|a1|]*wB+(wB-1)) with (wB*([|a1|]+1)-1) by ring.
assert (wB*([|a1|]+1) <= wB*[|b|]); try omega.
- apply Zmult_le_compat; omega.
+ apply Z.mul_le_mono_nonneg; omega.
Qed.
Lemma spec_div : forall a b, 0 < [|b|] ->
@@ -1445,20 +1442,20 @@ Section Int31_Specs.
unfold div31; intros.
assert ([|b|]>0) by (auto with zarith).
generalize (Z_div_mod [|a|] [|b|] H0) (Z_div_pos [|a|] [|b|] H0).
- unfold Zdiv; destruct (Zdiv_eucl [|a|] [|b|]); simpl.
+ unfold Z.div; destruct (Z.div_eucl [|a|] [|b|]); simpl.
rewrite ?phi_phi_inv.
destruct 1; intros.
- rewrite H1, Zmult_comm.
+ rewrite H1, Z.mul_comm.
generalize (phi_bounded a)(phi_bounded b); intros.
replace (z0 mod wB) with z0 by (symmetry; apply Zmod_small; omega).
replace (z mod wB) with z; auto with zarith.
symmetry; apply Zmod_small.
split; auto with zarith.
- apply Zle_lt_trans with [|a|]; auto with zarith.
+ apply Z.le_lt_trans with [|a|]; auto with zarith.
rewrite H1.
- apply Zle_trans with ([|b|]*z); try omega.
- rewrite <- (Zmult_1_l z) at 1.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.le_trans with ([|b|]*z); try omega.
+ rewrite <- (Z.mul_1_l z) at 1.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
Qed.
Lemma spec_mod : forall a b, 0 < [|b|] ->
@@ -1466,9 +1463,9 @@ Section Int31_Specs.
Proof.
unfold div31; intros.
assert ([|b|]>0) by (auto with zarith).
- unfold Zmod.
+ unfold Z.modulo.
generalize (Z_div_mod [|a|] [|b|] H0).
- destruct (Zdiv_eucl [|a|] [|b|]); simpl.
+ destruct (Z.div_eucl [|a|] [|b|]); simpl.
rewrite ?phi_phi_inv.
destruct 1; intros.
generalize (phi_bounded b); intros.
@@ -1506,12 +1503,12 @@ Section Int31_Specs.
destruct [|b|].
unfold size; auto with zarith.
intros (_,H).
- cut (Psize p <= size)%nat; [ omega | rewrite <- Zpower2_Psize; auto].
+ cut (Pos.size_nat p <= size)%nat; [ omega | rewrite <- Zpower2_Psize; auto].
intros (H,_); compute in H; elim H; auto.
Qed.
Lemma iter_int31_iter_nat : forall A f i a,
- iter_int31 i A f a = iter_nat (Zabs_nat [|i|]) A f a.
+ iter_int31 i A f a = iter_nat (Z.abs_nat [|i|]) A f a.
Proof.
intros.
unfold iter_int31.
@@ -1528,15 +1525,15 @@ Section Int31_Specs.
rewrite <- iter_nat_plus.
f_equal.
- rewrite Zdouble_mult, Zmult_comm, <- Zplus_diag_eq_mult_2.
- symmetry; apply Zabs_nat_Zplus; auto with zarith.
+ rewrite Z.double_spec, <- Z.add_diag.
+ symmetry; apply Zabs2Nat.inj_add; auto with zarith.
- change (iter_nat (S (Zabs_nat z + Zabs_nat z)) A f a =
- iter_nat (Zabs_nat (Zdouble_plus_one z)) A f a); f_equal.
- rewrite Zdouble_plus_one_mult, Zmult_comm, <- Zplus_diag_eq_mult_2.
- rewrite Zabs_nat_Zplus; auto with zarith.
- rewrite Zabs_nat_Zplus; auto with zarith.
- change (Zabs_nat 1) with 1%nat; omega.
+ change (iter_nat (S (Z.abs_nat z + Z.abs_nat z)) A f a =
+ iter_nat (Z.abs_nat (Z.succ_double z)) A f a); f_equal.
+ rewrite Z.succ_double_spec, <- Z.add_diag.
+ rewrite Zabs2Nat.inj_add; auto with zarith.
+ rewrite Zabs2Nat.inj_add; auto with zarith.
+ change (Z.abs_nat 1) with 1%nat; omega.
Qed.
Fixpoint addmuldiv31_alt n i j :=
@@ -1546,12 +1543,12 @@ Section Int31_Specs.
end.
Lemma addmuldiv31_equiv : forall p x y,
- addmuldiv31 p x y = addmuldiv31_alt (Zabs_nat [|p|]) x y.
+ addmuldiv31 p x y = addmuldiv31_alt (Z.abs_nat [|p|]) x y.
Proof.
intros.
unfold addmuldiv31.
rewrite iter_int31_iter_nat.
- set (n:=Zabs_nat [|p|]); clearbody n; clear p.
+ set (n:=Z.abs_nat [|p|]); clearbody n; clear p.
revert x y; induction n.
simpl; auto.
intros.
@@ -1566,21 +1563,21 @@ Section Int31_Specs.
Proof.
intros.
rewrite addmuldiv31_equiv.
- assert ([|p|] = Z_of_nat (Zabs_nat [|p|])).
- rewrite inj_Zabs_nat; symmetry; apply Zabs_eq.
+ assert ([|p|] = Z.of_nat (Z.abs_nat [|p|])).
+ rewrite Zabs2Nat.id_abs; symmetry; apply Z.abs_eq.
destruct (phi_bounded p); auto.
- rewrite H0; rewrite H0 in H; clear H0; rewrite Zabs_nat_Z_of_nat.
- set (n := Zabs_nat [|p|]) in *; clearbody n.
+ rewrite H0; rewrite H0 in H; clear H0; rewrite Zabs2Nat.id.
+ set (n := Z.abs_nat [|p|]) in *; clearbody n.
assert (n <= 31)%nat.
- rewrite inj_le_iff; auto with zarith.
+ rewrite Nat2Z.inj_le; auto with zarith.
clear p H; revert x y.
induction n.
simpl; intros.
- change (Zpower_pos 2 31) with (2^31).
- rewrite Zmult_1_r.
+ change (Z.pow_pos 2 31) with (2^31).
+ rewrite Z.mul_1_r.
replace ([|y|] / 2^31) with 0.
- rewrite Zplus_0_r.
+ rewrite Z.add_0_r.
symmetry; apply Zmod_small; apply phi_bounded.
symmetry; apply Zdiv_small; apply phi_bounded.
@@ -1588,43 +1585,43 @@ Section Int31_Specs.
rewrite IHn; [ | omega ].
case_eq (firstl y); intros.
- rewrite phi_twice, Zdouble_mult.
+ rewrite phi_twice, Z.double_spec.
rewrite phi_twice_firstl; auto.
- change (Zdouble [|y|]) with (2*[|y|]).
- rewrite inj_S, Zpower_Zsucc; auto with zarith.
+ change (Z.double [|y|]) with (2*[|y|]).
+ rewrite Nat2Z.inj_succ, Z.pow_succ_r; auto with zarith.
rewrite Zplus_mod; rewrite Zmult_mod_idemp_l; rewrite <- Zplus_mod.
f_equal.
- apply Zplus_eq_compat.
+ f_equal.
ring.
- replace (31-Z_of_nat n) with (Zsucc(31-Zsucc(Z_of_nat n))) by ring.
- rewrite Zpower_Zsucc, <- Zdiv_Zdiv; auto with zarith.
- rewrite Zmult_comm, Z_div_mult; auto with zarith.
+ replace (31-Z.of_nat n) with (Z.succ(31-Z.succ(Z.of_nat n))) by ring.
+ rewrite Z.pow_succ_r, <- Zdiv_Zdiv; auto with zarith.
+ rewrite Z.mul_comm, Z_div_mult; auto with zarith.
- rewrite phi_twice_plus_one, Zdouble_plus_one_mult.
+ rewrite phi_twice_plus_one, Z.succ_double_spec.
rewrite phi_twice; auto.
- change (Zdouble [|y|]) with (2*[|y|]).
- rewrite inj_S, Zpower_Zsucc; auto with zarith.
+ change (Z.double [|y|]) with (2*[|y|]).
+ rewrite Nat2Z.inj_succ, Z.pow_succ_r; auto with zarith.
rewrite Zplus_mod; rewrite Zmult_mod_idemp_l; rewrite <- Zplus_mod.
- rewrite Zmult_plus_distr_l, Zmult_1_l, <- Zplus_assoc.
+ rewrite Z.mul_add_distr_r, Z.mul_1_l, <- Z.add_assoc.
+ f_equal.
f_equal.
- apply Zplus_eq_compat.
ring.
assert ((2*[|y|]) mod wB = 2*[|y|] - wB).
clear - H. symmetry. apply Zmod_unique with 1; [ | ring ].
generalize (phi_lowerbound _ H) (phi_bounded y).
- set (wB' := 2^Z_of_nat (pred size)).
+ set (wB' := 2^Z.of_nat (pred size)).
replace wB with (2*wB'); [ omega | ].
- unfold wB'. rewrite <- Zpower_Zsucc, <- inj_S by (auto with zarith).
+ unfold wB'. rewrite <- Z.pow_succ_r, <- Nat2Z.inj_succ by (auto with zarith).
f_equal.
rewrite H1.
- replace wB with (2^(Z_of_nat n)*2^(31-Z_of_nat n)) by
+ replace wB with (2^(Z.of_nat n)*2^(31-Z.of_nat n)) by
(rewrite <- Zpower_exp; auto with zarith; f_equal; unfold size; ring).
- unfold Zminus; rewrite Zopp_mult_distr_l.
+ unfold Z.sub; rewrite <- Z.mul_opp_l.
rewrite Z_div_plus; auto with zarith.
ring_simplify.
- replace (31+-Z_of_nat n) with (Zsucc(31-Zsucc(Z_of_nat n))) by ring.
- rewrite Zpower_Zsucc, <- Zdiv_Zdiv; auto with zarith.
- rewrite Zmult_comm, Z_div_mult; auto with zarith.
+ replace (31+-Z.of_nat n) with (Z.succ(31-Z.succ(Z.of_nat n))) by ring.
+ rewrite Z.pow_succ_r, <- Zdiv_Zdiv; auto with zarith.
+ rewrite Z.mul_comm, Z_div_mult; auto with zarith.
Qed.
Lemma spec_pos_mod : forall w p,
@@ -1637,25 +1634,25 @@ Section Int31_Specs.
generalize (phi_bounded w).
symmetry; apply Zmod_small.
split; auto with zarith.
- apply Zlt_le_trans with wB; auto with zarith.
+ apply Z.lt_le_trans with wB; auto with zarith.
apply Zpower_le_monotone; auto with zarith.
intros.
case_eq ([|p|] ?= 31); intros;
- [ apply H; rewrite (Zcompare_Eq_eq _ _ H0); auto with zarith | |
+ [ apply H; rewrite (Z.compare_eq _ _ H0); auto with zarith | |
apply H; change ([|p|]>31)%Z in H0; auto with zarith ].
change ([|p|]<31) in H0.
rewrite spec_add_mul_div by auto with zarith.
- change [|0|] with 0%Z; rewrite Zmult_0_l, Zplus_0_l.
+ change [|0|] with 0%Z; rewrite Z.mul_0_l, Z.add_0_l.
generalize (phi_bounded p)(phi_bounded w); intros.
assert (31-[|p|]<wB).
- apply Zle_lt_trans with 31%Z; auto with zarith.
+ apply Z.le_lt_trans with 31%Z; auto with zarith.
compute; auto.
assert ([|31-p|]=31-[|p|]).
unfold sub31; rewrite phi_phi_inv.
change [|31|] with 31%Z.
apply Zmod_small; auto with zarith.
rewrite spec_add_mul_div by (rewrite H4; auto with zarith).
- change [|0|] with 0%Z; rewrite Zdiv_0_l, Zplus_0_r.
+ change [|0|] with 0%Z; rewrite Zdiv_0_l, Z.add_0_r.
rewrite H4.
apply shift_unshift_mod_2; auto with zarith.
Qed.
@@ -1682,7 +1679,7 @@ Section Int31_Specs.
end.
Lemma head031_equiv :
- forall x, [|head031 x|] = Z_of_nat (head031_alt size x).
+ forall x, [|head031 x|] = Z.of_nat (head031_alt size x).
Proof.
intros.
case_eq (iszero x); intros.
@@ -1690,7 +1687,7 @@ Section Int31_Specs.
simpl; auto.
unfold head031, recl.
- change On with (phi_inv (Z_of_nat (31-size))).
+ change On with (phi_inv (Z.of_nat (31-size))).
replace (head031_alt size x) with
(head031_alt size x + (31 - size))%nat by auto.
assert (size <= 31)%nat by auto with arith.
@@ -1700,12 +1697,12 @@ Section Int31_Specs.
unfold recl_aux; fold recl_aux.
unfold head031_alt; fold head031_alt.
rewrite H.
- assert ([|phi_inv (Z_of_nat (31-S n))|] = Z_of_nat (31 - S n)).
+ assert ([|phi_inv (Z.of_nat (31-S n))|] = Z.of_nat (31 - S n)).
rewrite phi_phi_inv.
apply Zmod_small.
split.
- change 0 with (Z_of_nat O); apply inj_le; omega.
- apply Zle_lt_trans with (Z_of_nat 31).
+ change 0 with (Z.of_nat O); apply inj_le; omega.
+ apply Z.le_lt_trans with (Z.of_nat 31).
apply inj_le; omega.
compute; auto.
case_eq (firstl x); intros; auto.
@@ -1718,7 +1715,7 @@ Section Int31_Specs.
f_equal.
change [|In|] with 1.
replace (31-n)%nat with (S (31 - S n))%nat by omega.
- rewrite inj_S; ring.
+ rewrite Nat2Z.inj_succ; ring.
clear - H H2.
rewrite (sneakr_shiftl x) in H.
@@ -1747,16 +1744,16 @@ Section Int31_Specs.
revert x H H0.
unfold size at 2 5.
induction size.
- simpl Z_of_nat.
+ simpl Z.of_nat.
intros.
compute in H0; rewrite H0 in H; discriminate.
intros.
simpl head031_alt.
case_eq (firstl x); intros.
- rewrite (inj_S (head031_alt n (shiftl x))), Zpower_Zsucc; auto with zarith.
- rewrite <- Zmult_assoc, Zmult_comm, <- Zmult_assoc, <-(Zmult_comm 2).
- rewrite <- Zdouble_mult, <- (phi_twice_firstl _ H1).
+ rewrite (Nat2Z.inj_succ (head031_alt n (shiftl x))), Z.pow_succ_r; auto with zarith.
+ rewrite <- Z.mul_assoc, Z.mul_comm, <- Z.mul_assoc, <-(Z.mul_comm 2).
+ rewrite <- Z.double_spec, <- (phi_twice_firstl _ H1).
apply IHn.
rewrite phi_nz; rewrite phi_nz in H; contradict H.
@@ -1765,9 +1762,9 @@ Section Int31_Specs.
rewrite <- nshiftl_S_tail; auto.
- change (2^(Z_of_nat 0)) with 1; rewrite Zmult_1_l.
+ change (2^(Z.of_nat 0)) with 1; rewrite Z.mul_1_l.
generalize (phi_bounded x); unfold size; split; auto with zarith.
- change (2^(Z_of_nat 31)/2) with (2^(Z_of_nat (pred size))).
+ change (2^(Z.of_nat 31)/2) with (2^(Z.of_nat (pred size))).
apply phi_lowerbound; auto.
Qed.
@@ -1790,7 +1787,7 @@ Section Int31_Specs.
end.
Lemma tail031_equiv :
- forall x, [|tail031 x|] = Z_of_nat (tail031_alt size x).
+ forall x, [|tail031 x|] = Z.of_nat (tail031_alt size x).
Proof.
intros.
case_eq (iszero x); intros.
@@ -1798,7 +1795,7 @@ Section Int31_Specs.
simpl; auto.
unfold tail031, recr.
- change On with (phi_inv (Z_of_nat (31-size))).
+ change On with (phi_inv (Z.of_nat (31-size))).
replace (tail031_alt size x) with
(tail031_alt size x + (31 - size))%nat by auto.
assert (size <= 31)%nat by auto with arith.
@@ -1808,12 +1805,12 @@ Section Int31_Specs.
unfold recr_aux; fold recr_aux.
unfold tail031_alt; fold tail031_alt.
rewrite H.
- assert ([|phi_inv (Z_of_nat (31-S n))|] = Z_of_nat (31 - S n)).
+ assert ([|phi_inv (Z.of_nat (31-S n))|] = Z.of_nat (31 - S n)).
rewrite phi_phi_inv.
apply Zmod_small.
split.
- change 0 with (Z_of_nat O); apply inj_le; omega.
- apply Zle_lt_trans with (Z_of_nat 31).
+ change 0 with (Z.of_nat O); apply inj_le; omega.
+ apply Z.le_lt_trans with (Z.of_nat 31).
apply inj_le; omega.
compute; auto.
case_eq (firstr x); intros; auto.
@@ -1826,7 +1823,7 @@ Section Int31_Specs.
f_equal.
change [|In|] with 1.
replace (31-n)%nat with (S (31 - S n))%nat by omega.
- rewrite inj_S; ring.
+ rewrite Nat2Z.inj_succ; ring.
clear - H H2.
rewrite (sneakl_shiftr x) in H.
@@ -1844,14 +1841,14 @@ Section Int31_Specs.
apply nshiftr_size.
revert x H H0.
induction size.
- simpl Z_of_nat.
+ simpl Z.of_nat.
intros.
compute in H0; rewrite H0 in H; discriminate.
intros.
simpl tail031_alt.
case_eq (firstr x); intros.
- rewrite (inj_S (tail031_alt n (shiftr x))), Zpower_Zsucc; auto with zarith.
+ rewrite (Nat2Z.inj_succ (tail031_alt n (shiftr x))), Z.pow_succ_r; auto with zarith.
destruct (IHn (shiftr x)) as (y & Hy1 & Hy2).
rewrite phi_nz; rewrite phi_nz in H; contradict H.
@@ -1861,13 +1858,13 @@ Section Int31_Specs.
exists y; split; auto.
rewrite phi_eqn1; auto.
- rewrite Zdouble_mult, Hy2; ring.
+ rewrite Z.double_spec, Hy2; ring.
exists [|shiftr x|].
split.
generalize (phi_bounded (shiftr x)); auto with zarith.
rewrite phi_eqn2; auto.
- rewrite Zdouble_plus_one_mult; simpl; ring.
+ rewrite Z.succ_double_spec; simpl; ring.
Qed.
(* Sqrt *)
@@ -1886,30 +1883,30 @@ Section Int31_Specs.
Proof.
intros Hj; generalize Hj k; pattern j; apply natlike_ind;
auto; clear k j Hj.
- intros _ k Hk; repeat rewrite Zplus_0_l.
- apply Zmult_le_0_compat; generalize (Z_div_pos k 2); auto with zarith.
+ intros _ k Hk; repeat rewrite Z.add_0_l.
+ apply Z.mul_nonneg_nonneg; generalize (Z_div_pos k 2); auto with zarith.
intros j Hj Hrec _ k Hk; pattern k; apply natlike_ind; auto; clear k Hk.
- rewrite Zmult_0_r, Zplus_0_r, Zplus_0_l.
- generalize (sqr_pos (Zsucc j / 2)) (quotient_by_2 (Zsucc j));
- unfold Zsucc.
- rewrite Zpower_2, Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r.
+ rewrite Z.mul_0_r, Z.add_0_r, Z.add_0_l.
+ generalize (sqr_pos (Z.succ j / 2)) (quotient_by_2 (Z.succ j));
+ unfold Z.succ.
+ rewrite Z.pow_2_r, Z.mul_add_distr_r; repeat rewrite Z.mul_add_distr_l.
auto with zarith.
intros k Hk _.
- replace ((Zsucc j + Zsucc k) / 2) with ((j + k)/2 + 1).
+ replace ((Z.succ j + Z.succ k) / 2) with ((j + k)/2 + 1).
generalize (Hrec Hj k Hk) (quotient_by_2 (j + k)).
- unfold Zsucc; repeat rewrite Zpower_2;
- repeat rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r.
- repeat rewrite Zmult_1_l; repeat rewrite Zmult_1_r.
+ unfold Z.succ; repeat rewrite Z.pow_2_r;
+ repeat rewrite Z.mul_add_distr_r; repeat rewrite Z.mul_add_distr_l.
+ repeat rewrite Z.mul_1_l; repeat rewrite Z.mul_1_r.
auto with zarith.
- rewrite Zplus_comm, <- Z_div_plus_full_l; auto with zarith.
- apply f_equal2 with (f := Zdiv); auto with zarith.
+ rewrite Z.add_comm, <- Z_div_plus_full_l; auto with zarith.
+ apply f_equal2 with (f := Z.div); auto with zarith.
Qed.
Lemma sqrt_main i j: 0 <= i -> 0 < j -> i < ((j + (i/j))/2 + 1) ^ 2.
Proof.
intros Hi Hj.
assert (Hij: 0 <= i/j) by (apply Z_div_pos; auto with zarith).
- apply Zlt_le_trans with (2 := sqrt_main_trick _ _ (Zlt_le_weak _ _ Hj) Hij).
+ apply Z.lt_le_trans with (2 := sqrt_main_trick _ _ (Z.lt_le_incl _ _ Hj) Hij).
pattern i at 1; rewrite (Z_div_mod_eq i j); case (Z_mod_lt i j); auto with zarith.
Qed.
@@ -1919,48 +1916,34 @@ Section Int31_Specs.
assert (H1: 0 <= i - 2) by auto with zarith.
assert (H2: 1 <= (i / 2) ^ 2); auto with zarith.
replace i with (1* 2 + (i - 2)); auto with zarith.
- rewrite Zpower_2, Z_div_plus_full_l; auto with zarith.
+ rewrite Z.pow_2_r, Z_div_plus_full_l; auto with zarith.
generalize (sqr_pos ((i - 2)/ 2)) (Z_div_pos (i - 2) 2).
- rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r.
+ rewrite Z.mul_add_distr_r; repeat rewrite Z.mul_add_distr_l.
auto with zarith.
generalize (quotient_by_2 i).
- rewrite Zpower_2 in H2 |- *;
- repeat (rewrite Zmult_plus_distr_l ||
- rewrite Zmult_plus_distr_r ||
- rewrite Zmult_1_l || rewrite Zmult_1_r).
+ rewrite Z.pow_2_r in H2 |- *;
+ repeat (rewrite Z.mul_add_distr_r ||
+ rewrite Z.mul_add_distr_l ||
+ rewrite Z.mul_1_l || rewrite Z.mul_1_r).
auto with zarith.
Qed.
Lemma sqrt_test_true i j: 0 <= i -> 0 < j -> i/j >= j -> j ^ 2 <= i.
Proof.
- intros Hi Hj Hd; rewrite Zpower_2.
- apply Zle_trans with (j * (i/j)); auto with zarith.
+ intros Hi Hj Hd; rewrite Z.pow_2_r.
+ apply Z.le_trans with (j * (i/j)); auto with zarith.
apply Z_mult_div_ge; auto with zarith.
Qed.
Lemma sqrt_test_false i j: 0 <= i -> 0 < j -> i/j < j -> (j + (i/j))/2 < j.
Proof.
- intros Hi Hj H; case (Zle_or_lt j ((j + (i/j))/2)); auto.
- intros H1; contradict H; apply Zle_not_lt.
+ intros Hi Hj H; case (Z.le_gt_cases j ((j + (i/j))/2)); auto.
+ intros H1; contradict H; apply Z.le_ngt.
assert (2 * j <= j + (i/j)); auto with zarith.
- apply Zle_trans with (2 * ((j + (i/j))/2)); auto with zarith.
+ apply Z.le_trans with (2 * ((j + (i/j))/2)); auto with zarith.
apply Z_mult_div_ge; auto with zarith.
Qed.
- (* George's trick *)
- Inductive ZcompareSpec (i j: Z): comparison -> Prop :=
- ZcompareSpecEq: i = j -> ZcompareSpec i j Eq
- | ZcompareSpecLt: i < j -> ZcompareSpec i j Lt
- | ZcompareSpecGt: j < i -> ZcompareSpec i j Gt.
-
- Lemma Zcompare_spec i j: ZcompareSpec i j (i ?= j).
- Proof.
- case_eq (Zcompare i j); intros H.
- apply ZcompareSpecEq; apply Zcompare_Eq_eq; auto.
- apply ZcompareSpecLt; auto.
- apply ZcompareSpecGt; apply Zgt_lt; auto.
- Qed.
-
Lemma sqrt31_step_def rec i j:
sqrt31_step rec i j =
match (fst (i/j) ?= j)%int31 with
@@ -1987,65 +1970,66 @@ Section Int31_Specs.
[|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) ->
[|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2.
Proof.
- assert (Hp2: 0 < [|2|]) by exact (refl_equal Lt).
+ assert (Hp2: 0 < [|2|]) by exact (eq_refl Lt).
intros Hi Hj Hij H31 Hrec; rewrite sqrt31_step_def.
rewrite spec_compare, div31_phi; auto.
- case Zcompare_spec; auto; intros Hc;
+ case Z.compare_spec; auto; intros Hc;
try (split; auto; apply sqrt_test_true; auto with zarith; fail).
apply Hrec; repeat rewrite div31_phi; auto with zarith.
replace [|(j + fst (i / j)%int31)|] with ([|j|] + [|i|] / [|j|]).
split.
- case (Zle_lt_or_eq 1 [|j|]); auto with zarith; intros Hj1.
+ apply Z.le_succ_l in Hj. change (1 <= [|j|]) in Hj.
+ Z.le_elim Hj.
replace ([|j|] + [|i|]/[|j|]) with
(1 * 2 + (([|j|] - 2) + [|i|] / [|j|])); try ring.
rewrite Z_div_plus_full_l; auto with zarith.
assert (0 <= [|i|]/ [|j|]) by (apply Z_div_pos; auto with zarith).
assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / [|2|]) ; auto with zarith.
- rewrite <- Hj1, Zdiv_1_r.
+ rewrite <- Hj, Zdiv_1_r.
replace (1 + [|i|])%Z with (1 * 2 + ([|i|] - 1))%Z; try ring.
rewrite Z_div_plus_full_l; auto with zarith.
assert (0 <= ([|i|] - 1) /2)%Z by (apply Z_div_pos; auto with zarith).
change ([|2|]) with 2%Z; auto with zarith.
apply sqrt_test_false; auto with zarith.
rewrite spec_add, div31_phi; auto.
- apply sym_equal; apply Zmod_small.
+ symmetry; apply Zmod_small.
split; auto with zarith.
replace [|j + fst (i / j)%int31|] with ([|j|] + [|i|] / [|j|]).
apply sqrt_main; auto with zarith.
rewrite spec_add, div31_phi; auto.
- apply sym_equal; apply Zmod_small.
+ symmetry; apply Zmod_small.
split; auto with zarith.
Qed.
Lemma iter31_sqrt_correct n rec i j: 0 < [|i|] -> 0 < [|j|] ->
- [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < 2 ^ (Z_of_nat size) ->
- (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] ->
- [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < 2 ^ (Z_of_nat size) ->
+ [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < 2 ^ (Z.of_nat size) ->
+ (forall j1, 0 < [|j1|] -> 2^(Z.of_nat n) + [|j1|] <= [|j|] ->
+ [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < 2 ^ (Z.of_nat size) ->
[|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) ->
[|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2.
Proof.
revert rec i j; elim n; unfold iter31_sqrt; fold iter31_sqrt; clear n.
intros rec i j Hi Hj Hij H31 Hrec; apply sqrt31_step_correct; auto with zarith.
intros; apply Hrec; auto with zarith.
- rewrite Zpower_0_r; auto with zarith.
+ rewrite Z.pow_0_r; auto with zarith.
intros n Hrec rec i j Hi Hj Hij H31 HHrec.
apply sqrt31_step_correct; auto.
intros j1 Hj1 Hjp1; apply Hrec; auto with zarith.
intros j2 Hj2 H2j2 Hjp2 Hj31; apply Hrec; auto with zarith.
intros j3 Hj3 Hpj3.
apply HHrec; auto.
- rewrite inj_S, Zpower_Zsucc.
- apply Zle_trans with (2 ^Z_of_nat n + [|j2|]); auto with zarith.
- apply Zle_0_nat.
+ rewrite Nat2Z.inj_succ, Z.pow_succ_r.
+ apply Z.le_trans with (2 ^Z.of_nat n + [|j2|]); auto with zarith.
+ apply Nat2Z.is_nonneg.
Qed.
Lemma spec_sqrt : forall x,
[|sqrt31 x|] ^ 2 <= [|x|] < ([|sqrt31 x|] + 1) ^ 2.
Proof.
intros i; unfold sqrt31.
- rewrite spec_compare. case Zcompare_spec; change [|1|] with 1;
+ rewrite spec_compare. case Z.compare_spec; change [|1|] with 1;
intros Hi; auto with zarith.
- repeat rewrite Zpower_2; auto with zarith.
+ repeat rewrite Z.pow_2_r; auto with zarith.
apply iter31_sqrt_correct; auto with zarith.
rewrite div31_phi; change ([|2|]) with 2; auto with zarith.
replace ([|i|]) with (1 * 2 + ([|i|] - 2))%Z; try ring.
@@ -2054,18 +2038,18 @@ Section Int31_Specs.
rewrite div31_phi; change ([|2|]) with 2; auto with zarith.
apply sqrt_init; auto.
rewrite div31_phi; change ([|2|]) with 2; auto with zarith.
- apply Zle_lt_trans with ([|i|]).
+ apply Z.le_lt_trans with ([|i|]).
apply Z_mult_div_ge; auto with zarith.
case (phi_bounded i); auto.
- intros j2 H1 H2; contradict H2; apply Zlt_not_le.
+ intros j2 H1 H2; contradict H2; apply Z.lt_nge.
rewrite div31_phi; change ([|2|]) with 2; auto with zarith.
- apply Zle_lt_trans with ([|i|]); auto with zarith.
+ apply Z.le_lt_trans with ([|i|]); auto with zarith.
assert (0 <= [|i|]/2)%Z by (apply Z_div_pos; auto with zarith).
- apply Zle_trans with (2 * ([|i|]/2)); auto with zarith.
+ apply Z.le_trans with (2 * ([|i|]/2)); auto with zarith.
apply Z_mult_div_ge; auto with zarith.
case (phi_bounded i); unfold size; auto with zarith.
change [|0|] with 0; auto with zarith.
- case (phi_bounded i); repeat rewrite Zpower_2; auto with zarith.
+ case (phi_bounded i); repeat rewrite Z.pow_2_r; auto with zarith.
Qed.
Lemma sqrt312_step_def rec ih il j:
@@ -2095,10 +2079,10 @@ Section Int31_Specs.
case (phi_bounded il); intros Hbil _.
case (phi_bounded ih); intros Hbih Hbih1.
assert (([|ih|] < [|j|] + 1)%Z); auto with zarith.
- apply Zlt_square_simpl; auto with zarith.
- repeat rewrite <-Zpower_2; apply Zle_lt_trans with (2 := H1).
- apply Zle_trans with ([|ih|] * base)%Z; unfold phi2, base;
- try rewrite Zpower_2; auto with zarith.
+ apply Z.square_lt_simpl_nonneg; auto with zarith.
+ repeat rewrite <-Z.pow_2_r; apply Z.le_lt_trans with (2 := H1).
+ apply Z.le_trans with ([|ih|] * base)%Z; unfold phi2, base;
+ try rewrite Z.pow_2_r; auto with zarith.
Qed.
Lemma div312_phi ih il j: (2^30 <= [|j|] -> [|ih|] < [|j|] ->
@@ -2108,7 +2092,7 @@ Section Int31_Specs.
generalize (spec_div21 ih il j Hj Hj1).
case div3121; intros q r (Hq, Hr).
apply Zdiv_unique with (phi r); auto with zarith.
- simpl fst; apply trans_equal with (1 := Hq); ring.
+ simpl fst; apply eq_trans with (1 := Hq); ring.
Qed.
Lemma sqrt312_step_correct rec ih il j:
@@ -2118,32 +2102,33 @@ Section Int31_Specs.
[|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il
< ([|sqrt312_step rec ih il j|] + 1) ^ 2.
Proof.
- assert (Hp2: (0 < [|2|])%Z) by exact (refl_equal Lt).
+ assert (Hp2: (0 < [|2|])%Z) by exact (eq_refl Lt).
intros Hih Hj Hij Hrec; rewrite sqrt312_step_def.
assert (H1: ([|ih|] <= [|j|])%Z) by (apply sqrt312_lower_bound with il; auto).
case (phi_bounded ih); intros Hih1 _.
case (phi_bounded il); intros Hil1 _.
case (phi_bounded j); intros _ Hj1.
assert (Hp3: (0 < phi2 ih il)).
- unfold phi2; apply Zlt_le_trans with ([|ih|] * base)%Z; auto with zarith.
- apply Zmult_lt_0_compat; auto with zarith.
- apply Zlt_le_trans with (2:= Hih); auto with zarith.
- rewrite spec_compare. case Zcompare_spec; intros Hc1.
+ unfold phi2; apply Z.lt_le_trans with ([|ih|] * base)%Z; auto with zarith.
+ apply Z.mul_pos_pos; auto with zarith.
+ apply Z.lt_le_trans with (2:= Hih); auto with zarith.
+ rewrite spec_compare. case Z.compare_spec; intros Hc1.
split; auto.
apply sqrt_test_true; auto.
unfold phi2, base; auto with zarith.
unfold phi2; rewrite Hc1.
assert (0 <= [|il|]/[|j|]) by (apply Z_div_pos; auto with zarith).
- rewrite Zmult_comm, Z_div_plus_full_l; unfold base; auto with zarith.
- unfold Zpower, Zpower_pos in Hj1; simpl in Hj1; auto with zarith.
- case (Zle_or_lt (2 ^ 30) [|j|]); intros Hjj.
- rewrite spec_compare; case Zcompare_spec;
+ rewrite Z.mul_comm, Z_div_plus_full_l; unfold base; auto with zarith.
+ unfold Z.pow, Z.pow_pos in Hj1; simpl in Hj1; auto with zarith.
+ case (Z.le_gt_cases (2 ^ 30) [|j|]); intros Hjj.
+ rewrite spec_compare; case Z.compare_spec;
rewrite div312_phi; auto; intros Hc;
try (split; auto; apply sqrt_test_true; auto with zarith; fail).
apply Hrec.
assert (Hf1: 0 <= phi2 ih il/ [|j|]) by (apply Z_div_pos; auto with zarith).
- case (Zle_lt_or_eq 1 ([|j|])); auto with zarith; intros Hf2.
- 2: contradict Hc; apply Zle_not_lt; rewrite <- Hf2, Zdiv_1_r; auto with zarith.
+ apply Z.le_succ_l in Hj. change (1 <= [|j|]) in Hj.
+ Z.le_elim Hj.
+ 2: contradict Hc; apply Z.le_ngt; rewrite <- Hj, Zdiv_1_r; auto with zarith.
assert (Hf3: 0 < ([|j|] + phi2 ih il / [|j|]) / 2).
replace ([|j|] + phi2 ih il/ [|j|])%Z with
(1 * 2 + (([|j|] - 2) + phi2 ih il / [|j|])); try ring.
@@ -2157,9 +2142,9 @@ Section Int31_Specs.
rewrite div31_phi; change [|2|] with 2%Z; auto with zarith.
intros HH; rewrite HH; clear HH; auto with zarith.
rewrite spec_add, div31_phi; change [|2|] with 2%Z; auto.
- rewrite Zmult_1_l; intros HH.
- rewrite Zplus_comm, <- Z_div_plus_full_l; auto with zarith.
- change (phi v30 * 2) with (2 ^ Z_of_nat size).
+ rewrite Z.mul_1_l; intros HH.
+ rewrite Z.add_comm, <- Z_div_plus_full_l; auto with zarith.
+ change (phi v30 * 2) with (2 ^ Z.of_nat size).
rewrite HH, Zmod_small; auto with zarith.
replace (phi
match j +c fst (div3121 ih il j) with
@@ -2173,41 +2158,41 @@ Section Int31_Specs.
rewrite div31_phi; auto with zarith.
intros HH; rewrite HH; auto with zarith.
intros HH; rewrite <- HH.
- change (1 * 2 ^ Z_of_nat size) with (phi (v30) * 2).
+ change (1 * 2 ^ Z.of_nat size) with (phi (v30) * 2).
rewrite Z_div_plus_full_l; auto with zarith.
- rewrite Zplus_comm.
+ rewrite Z.add_comm.
rewrite spec_add, Zmod_small.
rewrite div31_phi; auto.
split; auto with zarith.
case (phi_bounded (fst (r/2)%int31));
case (phi_bounded v30); auto with zarith.
rewrite div31_phi; change (phi 2) with 2%Z; auto.
- change (2 ^Z_of_nat size) with (base/2 + phi v30).
+ change (2 ^Z.of_nat size) with (base/2 + phi v30).
assert (phi r / 2 < base/2); auto with zarith.
- apply Zmult_gt_0_lt_reg_r with 2; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
change (base/2 * 2) with base.
- apply Zle_lt_trans with (phi r).
- rewrite Zmult_comm; apply Z_mult_div_ge; auto with zarith.
+ apply Z.le_lt_trans with (phi r).
+ rewrite Z.mul_comm; apply Z_mult_div_ge; auto with zarith.
case (phi_bounded r); auto with zarith.
- contradict Hij; apply Zle_not_lt.
+ contradict Hij; apply Z.le_ngt.
assert ((1 + [|j|]) <= 2 ^ 30); auto with zarith.
- apply Zle_trans with ((2 ^ 30) * (2 ^ 30)); auto with zarith.
+ apply Z.le_trans with ((2 ^ 30) * (2 ^ 30)); auto with zarith.
assert (0 <= 1 + [|j|]); auto with zarith.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
change ((2 ^ 30) * (2 ^ 30)) with ((2 ^ 29) * base).
- apply Zle_trans with ([|ih|] * base); auto with zarith.
+ apply Z.le_trans with ([|ih|] * base); auto with zarith.
unfold phi2, base; auto with zarith.
split; auto.
apply sqrt_test_true; auto.
unfold phi2, base; auto with zarith.
- apply Zle_ge; apply Zle_trans with (([|j|] * base)/[|j|]).
- rewrite Zmult_comm, Z_div_mult; auto with zarith.
- apply Zge_le; apply Z_div_ge; auto with zarith.
+ apply Z.le_ge; apply Z.le_trans with (([|j|] * base)/[|j|]).
+ rewrite Z.mul_comm, Z_div_mult; auto with zarith.
+ apply Z.ge_le; apply Z_div_ge; auto with zarith.
Qed.
Lemma iter312_sqrt_correct n rec ih il j:
2^29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 ->
- (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] ->
+ (forall j1, 0 < [|j1|] -> 2^(Z.of_nat n) + [|j1|] <= [|j|] ->
phi2 ih il < ([|j1|] + 1) ^ 2 ->
[|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) ->
[|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il
@@ -2216,16 +2201,16 @@ Section Int31_Specs.
revert rec ih il j; elim n; unfold iter312_sqrt; fold iter312_sqrt; clear n.
intros rec ih il j Hi Hj Hij Hrec; apply sqrt312_step_correct; auto with zarith.
intros; apply Hrec; auto with zarith.
- rewrite Zpower_0_r; auto with zarith.
+ rewrite Z.pow_0_r; auto with zarith.
intros n Hrec rec ih il j Hi Hj Hij HHrec.
apply sqrt312_step_correct; auto.
intros j1 Hj1 Hjp1; apply Hrec; auto with zarith.
intros j2 Hj2 H2j2 Hjp2; apply Hrec; auto with zarith.
intros j3 Hj3 Hpj3.
apply HHrec; auto.
- rewrite inj_S, Zpower_Zsucc.
- apply Zle_trans with (2 ^Z_of_nat n + [|j2|])%Z; auto with zarith.
- apply Zle_0_nat.
+ rewrite Nat2Z.inj_succ, Z.pow_succ_r.
+ apply Z.le_trans with (2 ^Z.of_nat n + [|j2|])%Z; auto with zarith.
+ apply Nat2Z.is_nonneg.
Qed.
Lemma spec_sqrt2 : forall x y,
@@ -2240,30 +2225,30 @@ Section Int31_Specs.
(intros s; ring).
assert (Hb: 0 <= base) by (red; intros HH; discriminate).
assert (Hi2: phi2 ih il < (phi Tn + 1) ^ 2).
- change ((phi Tn + 1) ^ 2) with (2^62).
- apply Zle_lt_trans with ((2^31 -1) * base + (2^31 - 1)); auto with zarith.
- 2: simpl; unfold Zpower_pos; simpl; auto with zarith.
- case (phi_bounded ih); case (phi_bounded il); intros H1 H2 H3 H4.
- unfold base, Zpower, Zpower_pos in H2,H4; simpl in H2,H4.
- unfold phi2,Zpower, Zpower_pos. simpl Pos.iter; auto with zarith.
+ { change ((phi Tn + 1) ^ 2) with (2^62).
+ apply Z.le_lt_trans with ((2^31 -1) * base + (2^31 - 1)); auto with zarith.
+ 2: simpl; unfold Z.pow_pos; simpl; auto with zarith.
+ case (phi_bounded ih); case (phi_bounded il); intros H1 H2 H3 H4.
+ unfold base, Z.pow, Z.pow_pos in H2,H4; simpl in H2,H4.
+ unfold phi2,Z.pow, Z.pow_pos. simpl Pos.iter; auto with zarith. }
case (iter312_sqrt_correct 31 (fun _ _ j => j) ih il Tn); auto with zarith.
change [|Tn|] with 2147483647; auto with zarith.
intros j1 _ HH; contradict HH.
- apply Zlt_not_le.
+ apply Z.lt_nge.
change [|Tn|] with 2147483647; auto with zarith.
- change (2 ^ Z_of_nat 31) with 2147483648; auto with zarith.
+ change (2 ^ Z.of_nat 31) with 2147483648; auto with zarith.
case (phi_bounded j1); auto with zarith.
set (s := iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn).
intros Hs1 Hs2.
generalize (spec_mul_c s s); case mul31c.
simpl zn2z_to_Z; intros HH.
assert ([|s|] = 0).
- case (Zmult_integral _ _ (sym_equal HH)); auto.
- contradict Hs2; apply Zle_not_lt; rewrite H.
+ { symmetry in HH. rewrite Z.mul_eq_0 in HH. destruct HH; auto. }
+ contradict Hs2; apply Z.le_ngt; rewrite H.
change ((0 + 1) ^ 2) with 1.
- apply Zle_trans with (2 ^ Z_of_nat size / 4 * base).
+ apply Z.le_trans with (2 ^ Z.of_nat size / 4 * base).
simpl; auto with zarith.
- apply Zle_trans with ([|ih|] * base); auto with zarith.
+ apply Z.le_trans with ([|ih|] * base); auto with zarith.
unfold phi2; case (phi_bounded il); auto with zarith.
intros ih1 il1.
change [||WW ih1 il1||] with (phi2 ih1 il1).
@@ -2271,10 +2256,10 @@ Section Int31_Specs.
generalize (spec_sub_c il il1).
case sub31c; intros il2 Hil2.
simpl interp_carry in Hil2.
- rewrite spec_compare; case Zcompare_spec.
+ rewrite spec_compare; case Z.compare_spec.
unfold interp_carry.
intros H1; split.
- rewrite Zpower_2, <- Hihl1.
+ rewrite Z.pow_2_r, <- Hihl1.
unfold phi2; ring[Hil2 H1].
replace [|il2|] with (phi2 ih il - phi2 ih1 il1).
rewrite Hihl1.
@@ -2282,109 +2267,111 @@ Section Int31_Specs.
unfold phi2; rewrite H1, Hil2; ring.
unfold interp_carry.
intros H1; contradict Hs1.
- apply Zlt_not_le; rewrite Zpower_2, <-Hihl1.
+ apply Z.lt_nge; rewrite Z.pow_2_r, <-Hihl1.
unfold phi2.
case (phi_bounded il); intros _ H2.
- apply Zlt_le_trans with (([|ih|] + 1) * base + 0).
- rewrite Zmult_plus_distr_l, Zplus_0_r; auto with zarith.
+ apply Z.lt_le_trans with (([|ih|] + 1) * base + 0).
+ rewrite Z.mul_add_distr_r, Z.add_0_r; auto with zarith.
case (phi_bounded il1); intros H3 _.
- apply Zplus_le_compat; auto with zarith.
- unfold interp_carry; change (1 * 2 ^ Z_of_nat size) with base.
- rewrite Zpower_2, <- Hihl1, Hil2.
+ apply Z.add_le_mono; auto with zarith.
+ unfold interp_carry; change (1 * 2 ^ Z.of_nat size) with base.
+ rewrite Z.pow_2_r, <- Hihl1, Hil2.
intros H1.
- case (Zle_lt_or_eq ([|ih1|] + 1) ([|ih|])); auto with zarith.
- intros H2; contradict Hs2; apply Zle_not_lt.
+ rewrite <- Z.le_succ_l, <- Z.add_1_r in H1.
+ Z.le_elim H1.
+ contradict Hs2; apply Z.le_ngt.
replace (([|s|] + 1) ^ 2) with (phi2 ih1 il1 + 2 * [|s|] + 1).
unfold phi2.
case (phi_bounded il); intros Hpil _.
assert (Hl1l: [|il1|] <= [|il|]).
- case (phi_bounded il2); rewrite Hil2; auto with zarith.
+ { case (phi_bounded il2); rewrite Hil2; auto with zarith. }
assert ([|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base); auto with zarith.
- case (phi_bounded s); change (2 ^ Z_of_nat size) with base; intros _ Hps.
+ case (phi_bounded s); change (2 ^ Z.of_nat size) with base; intros _ Hps.
case (phi_bounded ih1); intros Hpih1 _; auto with zarith.
- apply Zle_trans with (([|ih1|] + 2) * base); auto with zarith.
- rewrite Zmult_plus_distr_l.
+ apply Z.le_trans with (([|ih1|] + 2) * base); auto with zarith.
+ rewrite Z.mul_add_distr_r.
assert (2 * [|s|] + 1 <= 2 * base); auto with zarith.
rewrite Hihl1, Hbin; auto.
- intros H2; split.
- unfold phi2; rewrite <- H2; ring.
+ split.
+ unfold phi2; rewrite <- H1; ring.
replace (base + ([|il|] - [|il1|])) with (phi2 ih il - ([|s|] * [|s|])).
rewrite <-Hbin in Hs2; auto with zarith.
- rewrite <- Hihl1; unfold phi2; rewrite <- H2; ring.
+ rewrite <- Hihl1; unfold phi2; rewrite <- H1; ring.
unfold interp_carry in Hil2 |- *.
- unfold interp_carry; change (1 * 2 ^ Z_of_nat size) with base.
+ unfold interp_carry; change (1 * 2 ^ Z.of_nat size) with base.
assert (Hsih: [|ih - 1|] = [|ih|] - 1).
- rewrite spec_sub, Zmod_small; auto; change [|1|] with 1.
- case (phi_bounded ih); intros H1 H2.
- generalize Hih; change (2 ^ Z_of_nat size / 4) with 536870912.
- split; auto with zarith.
- rewrite spec_compare; case Zcompare_spec.
+ { rewrite spec_sub, Zmod_small; auto; change [|1|] with 1.
+ case (phi_bounded ih); intros H1 H2.
+ generalize Hih; change (2 ^ Z.of_nat size / 4) with 536870912.
+ split; auto with zarith. }
+ rewrite spec_compare; case Z.compare_spec.
rewrite Hsih.
intros H1; split.
- rewrite Zpower_2, <- Hihl1.
+ rewrite Z.pow_2_r, <- Hihl1.
unfold phi2; rewrite <-H1.
- apply trans_equal with ([|ih|] * base + [|il1|] + ([|il|] - [|il1|])).
+ transitivity ([|ih|] * base + [|il1|] + ([|il|] - [|il1|])).
ring.
rewrite <-Hil2.
- change (2 ^ Z_of_nat size) with base; ring.
+ change (2 ^ Z.of_nat size) with base; ring.
replace [|il2|] with (phi2 ih il - phi2 ih1 il1).
rewrite Hihl1.
rewrite <-Hbin in Hs2; auto with zarith.
unfold phi2.
rewrite <-H1.
ring_simplify.
- apply trans_equal with (base + ([|il|] - [|il1|])).
+ transitivity (base + ([|il|] - [|il1|])).
ring.
rewrite <-Hil2.
- change (2 ^ Z_of_nat size) with base; ring.
+ change (2 ^ Z.of_nat size) with base; ring.
rewrite Hsih; intros H1.
assert (He: [|ih|] = [|ih1|]).
- apply Zle_antisym; auto with zarith.
- case (Zle_or_lt [|ih1|] [|ih|]); auto; intros H2.
- contradict Hs1; apply Zlt_not_le; rewrite Zpower_2, <-Hihl1.
- unfold phi2.
- case (phi_bounded il); change (2 ^ Z_of_nat size) with base;
+ { apply Z.le_antisymm; auto with zarith.
+ case (Z.le_gt_cases [|ih1|] [|ih|]); auto; intros H2.
+ contradict Hs1; apply Z.lt_nge; rewrite Z.pow_2_r, <-Hihl1.
+ unfold phi2.
+ case (phi_bounded il); change (2 ^ Z.of_nat size) with base;
intros _ Hpil1.
- apply Zlt_le_trans with (([|ih|] + 1) * base).
- rewrite Zmult_plus_distr_l, Zmult_1_l; auto with zarith.
- case (phi_bounded il1); intros Hpil2 _.
- apply Zle_trans with (([|ih1|]) * base); auto with zarith.
- rewrite Zpower_2, <-Hihl1; unfold phi2; rewrite <-He.
- contradict Hs1; apply Zlt_not_le; rewrite Zpower_2, <-Hihl1.
+ apply Z.lt_le_trans with (([|ih|] + 1) * base).
+ rewrite Z.mul_add_distr_r, Z.mul_1_l; auto with zarith.
+ case (phi_bounded il1); intros Hpil2 _.
+ apply Z.le_trans with (([|ih1|]) * base); auto with zarith. }
+ rewrite Z.pow_2_r, <-Hihl1; unfold phi2; rewrite <-He.
+ contradict Hs1; apply Z.lt_nge; rewrite Z.pow_2_r, <-Hihl1.
unfold phi2; rewrite He.
assert (phi il - phi il1 < 0); auto with zarith.
rewrite <-Hil2.
case (phi_bounded il2); auto with zarith.
intros H1.
- rewrite Zpower_2, <-Hihl1.
- case (Zle_lt_or_eq ([|ih1|] + 2) [|ih|]); auto with zarith.
- intros H2; contradict Hs2; apply Zle_not_lt.
+ rewrite Z.pow_2_r, <-Hihl1.
+ assert (H2 : [|ih1|]+2 <= [|ih|]); auto with zarith.
+ Z.le_elim H2.
+ contradict Hs2; apply Z.le_ngt.
replace (([|s|] + 1) ^ 2) with (phi2 ih1 il1 + 2 * [|s|] + 1).
unfold phi2.
assert ([|ih1|] * base + 2 * phi s + 1 <= [|ih|] * base + ([|il|] - [|il1|]));
auto with zarith.
rewrite <-Hil2.
- change (-1 * 2 ^ Z_of_nat size) with (-base).
+ change (-1 * 2 ^ Z.of_nat size) with (-base).
case (phi_bounded il2); intros Hpil2 _.
- apply Zle_trans with ([|ih|] * base + - base); auto with zarith.
- case (phi_bounded s); change (2 ^ Z_of_nat size) with base; intros _ Hps.
+ apply Z.le_trans with ([|ih|] * base + - base); auto with zarith.
+ case (phi_bounded s); change (2 ^ Z.of_nat size) with base; intros _ Hps.
assert (2 * [|s|] + 1 <= 2 * base); auto with zarith.
- apply Zle_trans with ([|ih1|] * base + 2 * base); auto with zarith.
+ apply Z.le_trans with ([|ih1|] * base + 2 * base); auto with zarith.
assert (Hi: ([|ih1|] + 3) * base <= [|ih|] * base); auto with zarith.
- rewrite Zmult_plus_distr_l in Hi; auto with zarith.
+ rewrite Z.mul_add_distr_r in Hi; auto with zarith.
rewrite Hihl1, Hbin; auto.
- intros H2; unfold phi2; rewrite <-H2.
+ unfold phi2; rewrite <-H2.
split.
replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring.
rewrite <-Hil2.
- change (-1 * 2 ^ Z_of_nat size) with (-base); ring.
+ change (-1 * 2 ^ Z.of_nat size) with (-base); ring.
replace (base + [|il2|]) with (phi2 ih il - phi2 ih1 il1).
rewrite Hihl1.
rewrite <-Hbin in Hs2; auto with zarith.
unfold phi2; rewrite <-H2.
replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring.
rewrite <-Hil2.
- change (-1 * 2 ^ Z_of_nat size) with (-base); ring.
+ change (-1 * 2 ^ Z.of_nat size) with (-base); ring.
Qed.
(** [iszero] *)
@@ -2394,7 +2381,7 @@ Qed.
clear; unfold ZnZ.eq0; simpl.
unfold compare31; simpl; intros.
change [|0|] with 0 in H.
- apply Zcompare_Eq_eq.
+ apply Z.compare_eq.
now destruct ([|x|] ?= 0).
Qed.
@@ -2412,7 +2399,7 @@ Qed.
destruct H; auto with zarith.
replace ([|x|] mod 2) with [|r|].
destruct H; auto with zarith.
- case Zcompare_spec; auto with zarith.
+ case Z.compare_spec; auto with zarith.
apply Zmod_unique with [|q|]; auto with zarith.
Qed.
diff --git a/theories/Numbers/Cyclic/Int31/Int31.v b/theories/Numbers/Cyclic/Int31/Int31.v
index 20f750f64..5415b379b 100644
--- a/theories/Numbers/Cyclic/Int31/Int31.v
+++ b/theories/Numbers/Cyclic/Int31/Int31.v
@@ -117,12 +117,12 @@ Definition iszero : int31 -> bool := Eval compute in
It seems to work, but later "unfold iszero" takes forever. *)
-(** [base] is [2^31], obtained via iterations of [Zdouble].
+(** [base] is [2^31], obtained via iterations of [Z.double].
It can also be seen as the smallest b > 0 s.t. phi_inv b = 0
(see below) *)
Definition base := Eval compute in
- iter_nat size Z Zdouble 1%Z.
+ iter_nat size Z Z.double 1%Z.
(** * Recursors *)
@@ -155,11 +155,11 @@ Definition recr := recr_aux size.
(** * Conversions *)
-(** From int31 to Z, we simply iterates [Zdouble] or [Zdouble_plus_one]. *)
+(** From int31 to Z, we simply iterates [Z.double] or [Z.succ_double]. *)
Definition phi : int31 -> Z :=
recr Z (0%Z)
- (fun b _ => match b with D0 => Zdouble | D1 => Zdouble_plus_one end).
+ (fun b _ => match b with D0 => Z.double | D1 => Z.succ_double end).
(** From positive to int31. An abstract definition could be :
[ phi_inv (2n) = 2*(phi_inv n) /\
@@ -293,13 +293,13 @@ Notation "n '*c' m" := (mul31c n m) (at level 40, no associativity) : int31_scop
(** Division of a double size word modulo [2^31] *)
Definition div3121 (nh nl m : int31) :=
- let (q,r) := Zdiv_eucl (phi2 nh nl) (phi m) in
+ let (q,r) := Z.div_eucl (phi2 nh nl) (phi m) in
(phi_inv q, phi_inv r).
(** Division modulo [2^31] *)
Definition div31 (n m : int31) :=
- let (q,r) := Zdiv_eucl (phi n) (phi m) in
+ let (q,r) := Z.div_eucl (phi n) (phi m) in
(phi_inv q, phi_inv r).
Notation "n / m" := (div31 n m) : int31_scope.
@@ -391,7 +391,7 @@ Eval lazy delta [On In Twon] in
| Lt => iter31_sqrt 31 (fun i j => j) i (fst (i/Twon))
end.
-Definition v30 := Eval compute in (addmuldiv31 (phi_inv (Z_of_nat size - 1)) In On).
+Definition v30 := Eval compute in (addmuldiv31 (phi_inv (Z.of_nat size - 1)) In On).
Definition sqrt312_step (rec: int31 -> int31 -> int31 -> int31)
(ih il j: int31) :=
@@ -452,7 +452,7 @@ Definition positive_to_int31 (p:positive) := p2i size p.
It is used as default answer for numbers of zeros
in [head0] and [tail0] *)
-Definition T31 : int31 := Eval compute in phi_inv (Z_of_nat size).
+Definition T31 : int31 := Eval compute in phi_inv (Z.of_nat size).
Definition head031 (i:int31) :=
recl _ (fun _ => T31)
diff --git a/theories/Numbers/Cyclic/Int31/Ring31.v b/theories/Numbers/Cyclic/Int31/Ring31.v
index 23e8bd338..ef3400223 100644
--- a/theories/Numbers/Cyclic/Int31/Ring31.v
+++ b/theories/Numbers/Cyclic/Int31/Ring31.v
@@ -81,7 +81,7 @@ Qed.
Lemma eqb31_eq : forall x y, eqb31 x y = true <-> x=y.
Proof.
unfold eqb31. intros x y.
-rewrite Cyclic31.spec_compare. case Zcompare_spec.
+rewrite Cyclic31.spec_compare. case Z.compare_spec.
intuition. apply Int31_canonic; auto.
intuition; subst; auto with zarith; try discriminate.
intuition; subst; auto with zarith; try discriminate.
diff --git a/theories/Numbers/Cyclic/ZModulo/ZModulo.v b/theories/Numbers/Cyclic/ZModulo/ZModulo.v
index d039fdcbf..6945d0757 100644
--- a/theories/Numbers/Cyclic/ZModulo/ZModulo.v
+++ b/theories/Numbers/Cyclic/ZModulo/ZModulo.v
@@ -76,22 +76,22 @@ Section ZModulo.
Qed.
Definition of_pos x :=
- let (q,r) := Zdiv_eucl_POS x wB in (N_of_Z q, r).
+ let (q,r) := Z.pos_div_eucl x wB in (N_of_Z q, r).
Lemma spec_of_pos : forall p,
- Zpos p = (Z_of_N (fst (of_pos p)))*wB + [|(snd (of_pos p))|].
+ Zpos p = (Z.of_N (fst (of_pos p)))*wB + [|(snd (of_pos p))|].
Proof.
intros; unfold of_pos; simpl.
generalize (Z_div_mod_POS wB wB_pos p).
- destruct (Zdiv_eucl_POS p wB); simpl; destruct 1.
+ destruct (Z.pos_div_eucl p wB); simpl; destruct 1.
unfold to_Z; rewrite Zmod_small; auto.
assert (0 <= z).
replace z with (Zpos p / wB) by
(symmetry; apply Zdiv_unique with z0; auto).
apply Z_div_pos; auto with zarith.
- replace (Z_of_N (N_of_Z z)) with z by
+ replace (Z.of_N (N_of_Z z)) with z by
(destruct z; simpl; auto; elim H1; auto).
- rewrite Zmult_comm; auto.
+ rewrite Z.mul_comm; auto.
Qed.
Lemma spec_zdigits : [|zdigits|] = Zpos digits.
@@ -118,7 +118,7 @@ Section ZModulo.
unfold to_Z, one.
apply Zmod_small; split; auto with zarith.
unfold wB, base.
- apply Zlt_trans with (Zpos digits); auto.
+ apply Z.lt_trans with (Zpos digits); auto.
apply Zpower2_lt_lin; auto with zarith.
Qed.
@@ -128,14 +128,14 @@ Section ZModulo.
apply Zmod_small; split; auto with zarith.
unfold wB, base.
cut (1 <= 2 ^ Zpos digits); auto with zarith.
- apply Zle_trans with (Zpos digits); auto with zarith.
+ apply Z.le_trans with (Zpos digits); auto with zarith.
apply Zpower2_le_lin; auto with zarith.
Qed.
- Definition compare x y := Zcompare [|x|] [|y|].
+ Definition compare x y := Z.compare [|x|] [|y|].
Lemma spec_compare : forall x y,
- compare x y = Zcompare [|x|] [|y|].
+ compare x y = Z.compare [|x|] [|y|].
Proof. reflexivity. Qed.
Definition eq0 x :=
@@ -183,7 +183,7 @@ Section ZModulo.
Qed.
Definition succ_c x :=
- let y := Zsucc x in
+ let y := Z.succ x in
if eq0 y then C1 0 else C0 y.
Definition add_c x y :=
@@ -194,29 +194,28 @@ Section ZModulo.
let z := [|x|]+[|y|]+1 in
if Z_lt_le_dec z wB then C0 z else C1 (z-wB).
- Definition succ := Zsucc.
- Definition add := Zplus.
+ Definition succ := Z.succ.
+ Definition add := Z.add.
Definition add_carry x y := x + y + 1.
Lemma Zmod_equal :
forall x y z, z>0 -> (x-y) mod z = 0 -> x mod z = y mod z.
Proof.
intros.
- generalize (Z_div_mod_eq (x-y) z H); rewrite H0, Zplus_0_r.
+ generalize (Z_div_mod_eq (x-y) z H); rewrite H0, Z.add_0_r.
remember ((x-y)/z) as k.
- intros H1; symmetry in H1; rewrite <- Zeq_plus_swap in H1.
- subst x.
- rewrite Zplus_comm, Zmult_comm, Z_mod_plus; auto.
+ rewrite Z.sub_move_r, Z.add_comm, Z.mul_comm. intros ->.
+ now apply Z_mod_plus.
Qed.
Lemma spec_succ_c : forall x, [+|succ_c x|] = [|x|] + 1.
Proof.
- intros; unfold succ_c, to_Z, Zsucc.
+ intros; unfold succ_c, to_Z, Z.succ.
case_eq (eq0 (x+1)); intros; unfold interp_carry.
- rewrite Zmult_1_l.
+ rewrite Z.mul_1_l.
replace (wB + 0 mod wB) with wB by auto with zarith.
- symmetry; rewrite Zeq_plus_swap.
+ symmetry. rewrite Z.add_move_r.
assert ((x+1) mod wB = 0) by (apply spec_eq0; auto).
replace (wB-1) with ((wB-1) mod wB) by
(apply Zmod_small; generalize wB_pos; omega).
@@ -227,7 +226,7 @@ Section ZModulo.
unfold eq0, to_Z in *; now destruct ((x+1) mod wB).
assert (x mod wB + 1 <> wB).
contradict H0.
- rewrite Zeq_plus_swap in H0; simpl in H0.
+ rewrite Z.add_move_r in H0; simpl in H0.
rewrite <- Zplus_mod_idemp_l; rewrite H0.
replace (wB-1+1) with wB; auto with zarith; apply Z_mod_same; auto.
rewrite <- Zplus_mod_idemp_l.
@@ -241,7 +240,7 @@ Section ZModulo.
destruct Z_lt_le_dec.
apply Zmod_small;
generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega.
- rewrite Zmult_1_l, Zplus_comm, Zeq_plus_swap.
+ rewrite Z.mul_1_l, Z.add_comm, Z.add_move_r.
apply Zmod_small;
generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega.
Qed.
@@ -252,14 +251,14 @@ Section ZModulo.
destruct Z_lt_le_dec.
apply Zmod_small;
generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega.
- rewrite Zmult_1_l, Zplus_comm, Zeq_plus_swap.
+ rewrite Z.mul_1_l, Z.add_comm, Z.add_move_r.
apply Zmod_small;
generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega.
Qed.
Lemma spec_succ : forall x, [|succ x|] = ([|x|] + 1) mod wB.
Proof.
- intros; unfold succ, to_Z, Zsucc.
+ intros; unfold succ, to_Z, Z.succ.
symmetry; apply Zplus_mod_idemp_l.
Qed.
@@ -288,8 +287,8 @@ Section ZModulo.
let z := [|x|]-[|y|]-1 in
if Z_lt_le_dec z 0 then C1 (wB+z) else C0 z.
- Definition pred := Zpred.
- Definition sub := Zminus.
+ Definition pred := Z.pred.
+ Definition sub := Z.sub.
Definition sub_carry x y := x - y - 1.
Lemma spec_pred_c : forall x, [-|pred_c x|] = [|x|] - 1.
@@ -337,7 +336,7 @@ Section ZModulo.
Lemma spec_pred : forall x, [|pred x|] = ([|x|] - 1) mod wB.
Proof.
- intros; unfold pred, to_Z, Zpred.
+ intros; unfold pred, to_Z, Z.pred.
rewrite <- Zplus_mod_idemp_l; auto.
Qed.
@@ -357,19 +356,19 @@ Section ZModulo.
Qed.
Definition mul_c x y :=
- let (h,l) := Zdiv_eucl ([|x|]*[|y|]) wB in
+ let (h,l) := Z.div_eucl ([|x|]*[|y|]) wB in
if eq0 h then if eq0 l then W0 else WW h l else WW h l.
- Definition mul := Zmult.
+ Definition mul := Z.mul.
Definition square_c x := mul_c x x.
Lemma spec_mul_c : forall x y, [|| mul_c x y ||] = [|x|] * [|y|].
Proof.
intros; unfold mul_c, zn2z_to_Z.
- assert (Zdiv_eucl ([|x|]*[|y|]) wB = (([|x|]*[|y|])/wB,([|x|]*[|y|]) mod wB)).
- unfold Zmod, Zdiv; destruct Zdiv_eucl; auto.
- generalize (Z_div_mod ([|x|]*[|y|]) wB wB_pos); destruct Zdiv_eucl as (h,l).
+ assert (Z.div_eucl ([|x|]*[|y|]) wB = (([|x|]*[|y|])/wB,([|x|]*[|y|]) mod wB)).
+ unfold Z.modulo, Z.div; destruct Z.div_eucl; auto.
+ generalize (Z_div_mod ([|x|]*[|y|]) wB wB_pos); destruct Z.div_eucl as (h,l).
destruct 1; injection H; clear H; intros.
rewrite H0.
assert ([|l|] = l).
@@ -380,7 +379,7 @@ Section ZModulo.
split.
apply Z_div_pos; auto with zarith.
apply Zdiv_lt_upper_bound; auto with zarith.
- apply Zmult_lt_compat; auto with zarith.
+ apply Z.mul_lt_mono_nonneg; auto with zarith.
clear H H0 H1 H2.
case_eq (eq0 h); simpl; intros.
case_eq (eq0 l); simpl; intros.
@@ -399,7 +398,7 @@ Section ZModulo.
intros x; exact (spec_mul_c x x).
Qed.
- Definition div x y := Zdiv_eucl [|x|] [|y|].
+ Definition div x y := Z.div_eucl [|x|] [|y|].
Lemma spec_div : forall a b, 0 < [|b|] ->
let (q,r) := div a b in
@@ -408,10 +407,10 @@ Section ZModulo.
Proof.
intros; unfold div.
assert ([|b|]>0) by auto with zarith.
- assert (Zdiv_eucl [|a|] [|b|] = ([|a|]/[|b|], [|a|] mod [|b|])).
- unfold Zmod, Zdiv; destruct Zdiv_eucl; auto.
+ assert (Z.div_eucl [|a|] [|b|] = ([|a|]/[|b|], [|a|] mod [|b|])).
+ unfold Z.modulo, Z.div; destruct Z.div_eucl; auto.
generalize (Z_div_mod [|a|] [|b|] H0).
- destruct Zdiv_eucl as (q,r); destruct 1; intros.
+ destruct Z.div_eucl as (q,r); destruct 1; intros.
injection H1; clear H1; intros.
assert ([|r|]=r).
apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|];
@@ -422,10 +421,10 @@ Section ZModulo.
split.
apply Z_div_pos; auto with zarith.
apply Zdiv_lt_upper_bound; auto with zarith.
- apply Zlt_le_trans with (wB*1).
- rewrite Zmult_1_r; auto with zarith.
- apply Zmult_le_compat; generalize wB_pos; auto with zarith.
- rewrite H5, H6; rewrite Zmult_comm; auto with zarith.
+ apply Z.lt_le_trans with (wB*1).
+ rewrite Z.mul_1_r; auto with zarith.
+ apply Z.mul_le_mono_nonneg; generalize wB_pos; auto with zarith.
+ rewrite H5, H6; rewrite Z.mul_comm; auto with zarith.
Qed.
Definition div_gt := div.
@@ -458,28 +457,28 @@ Section ZModulo.
intros; apply spec_modulo; auto.
Qed.
- Definition gcd x y := Zgcd [|x|] [|y|].
- Definition gcd_gt x y := Zgcd [|x|] [|y|].
+ Definition gcd x y := Z.gcd [|x|] [|y|].
+ Definition gcd_gt x y := Z.gcd [|x|] [|y|].
- Lemma Zgcd_bound : forall a b, 0<=a -> 0<=b -> Zgcd a b <= Zmax a b.
+ Lemma Zgcd_bound : forall a b, 0<=a -> 0<=b -> Z.gcd a b <= Z.max a b.
Proof.
intros.
generalize (Zgcd_is_gcd a b); inversion_clear 1.
destruct H2 as (q,H2); destruct H3 as (q',H3); clear H4.
- assert (H4:=Zgcd_is_pos a b).
- destruct (Z_eq_dec (Zgcd a b) 0).
+ assert (H4:=Z.gcd_nonneg a b).
+ destruct (Z.eq_dec (Z.gcd a b) 0).
rewrite e; generalize (Zmax_spec a b); omega.
assert (0 <= q).
- apply Zmult_le_reg_r with (Zgcd a b); auto with zarith.
- destruct (Z_eq_dec q 0).
+ apply Z.mul_le_mono_pos_r with (Z.gcd a b); auto with zarith.
+ destruct (Z.eq_dec q 0).
subst q; simpl in *; subst a; simpl; auto.
generalize (Zmax_spec 0 b) (Zabs_spec b); omega.
- apply Zle_trans with a.
+ apply Z.le_trans with a.
rewrite H2 at 2.
- rewrite <- (Zmult_1_l (Zgcd a b)) at 1.
- apply Zmult_le_compat; auto with zarith.
+ rewrite <- (Z.mul_1_l (Z.gcd a b)) at 1.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
generalize (Zmax_spec a b); omega.
Qed.
@@ -488,12 +487,12 @@ Section ZModulo.
intros; unfold gcd.
generalize (Z_mod_lt a wB wB_pos)(Z_mod_lt b wB wB_pos); intros.
fold [|a|] in *; fold [|b|] in *.
- replace ([|Zgcd [|a|] [|b|]|]) with (Zgcd [|a|] [|b|]).
+ replace ([|Z.gcd [|a|] [|b|]|]) with (Z.gcd [|a|] [|b|]).
apply Zgcd_is_gcd.
symmetry; apply Zmod_small.
split.
- apply Zgcd_is_pos.
- apply Zle_lt_trans with (Zmax [|a|] [|b|]).
+ apply Z.gcd_nonneg.
+ apply Z.le_lt_trans with (Z.max [|a|] [|b|]).
apply Zgcd_bound; auto with zarith.
generalize (Zmax_spec [|a|] [|b|]); omega.
Qed.
@@ -505,7 +504,7 @@ Section ZModulo.
Qed.
Definition div21 a1 a2 b :=
- Zdiv_eucl ([|a1|]*wB+[|a2|]) [|b|].
+ Z.div_eucl ([|a1|]*wB+[|a2|]) [|b|].
Lemma spec_div21 : forall a1 a2 b,
wB/2 <= [|b|] ->
@@ -519,10 +518,10 @@ Section ZModulo.
generalize (Z_mod_lt a2 wB wB_pos); fold [|a2|]; intros.
assert ([|b|]>0) by auto with zarith.
remember ([|a1|]*wB+[|a2|]) as a.
- assert (Zdiv_eucl a [|b|] = (a/[|b|], a mod [|b|])).
- unfold Zmod, Zdiv; destruct Zdiv_eucl; auto.
+ assert (Z.div_eucl a [|b|] = (a/[|b|], a mod [|b|])).
+ unfold Z.modulo, Z.div; destruct Z.div_eucl; auto.
generalize (Z_div_mod a [|b|] H3).
- destruct Zdiv_eucl as (q,r); destruct 1; intros.
+ destruct Z.div_eucl as (q,r); destruct 1; intros.
injection H4; clear H4; intros.
assert ([|r|]=r).
apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|];
@@ -536,8 +535,8 @@ Section ZModulo.
apply Zdiv_lt_upper_bound; auto with zarith.
subst a.
replace (wB*[|b|]) with (([|b|]-1)*wB + wB) by ring.
- apply Zlt_le_trans with ([|a1|]*wB+wB); auto with zarith.
- rewrite H8, H9; rewrite Zmult_comm; auto with zarith.
+ apply Z.lt_le_trans with ([|a1|]*wB+wB); auto with zarith.
+ rewrite H8, H9; rewrite Z.mul_comm; auto with zarith.
Qed.
Definition add_mul_div p x y :=
@@ -560,17 +559,17 @@ Section ZModulo.
generalize (Z_mod_lt [|w|] (2 ^ [|p|])); intros.
split.
destruct H; auto with zarith.
- apply Zle_lt_trans with [|w|]; auto with zarith.
+ apply Z.le_lt_trans with [|w|]; auto with zarith.
apply Zmod_le; auto with zarith.
Qed.
Definition is_even x :=
- if Z_eq_dec ([|x|] mod 2) 0 then true else false.
+ if Z.eq_dec ([|x|] mod 2) 0 then true else false.
Lemma spec_is_even : forall x,
if is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1.
Proof.
- intros; unfold is_even; destruct Z_eq_dec; auto.
+ intros; unfold is_even; destruct Z.eq_dec; auto.
generalize (Z_mod_lt [|x|] 2); omega.
Qed.
@@ -580,12 +579,12 @@ Section ZModulo.
Proof.
intros.
unfold sqrt.
- repeat rewrite Zpower_2.
+ repeat rewrite Z.pow_2_r.
replace [|Z.sqrt [|x|]|] with (Z.sqrt [|x|]).
apply Z.sqrt_spec; auto with zarith.
symmetry; apply Zmod_small.
split. apply Z.sqrt_nonneg; auto.
- apply Zle_lt_trans with [|x|]; auto.
+ apply Z.le_lt_trans with [|x|]; auto.
apply Z.sqrt_le_lin; auto.
Qed.
@@ -616,22 +615,22 @@ Section ZModulo.
destruct (Z_lt_le_dec s wB); auto.
assert (wB * wB <= Zpos p).
rewrite U.
- apply Zle_trans with (s*s); try omega.
- apply Zmult_le_compat; generalize wB_pos; auto with zarith.
+ apply Z.le_trans with (s*s); try omega.
+ apply Z.mul_le_mono_nonneg; generalize wB_pos; auto with zarith.
assert (Zpos p < wB*wB).
rewrite Heqz.
replace (wB*wB) with ((wB-1)*wB+wB) by ring.
- apply Zplus_le_lt_compat; auto with zarith.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.add_le_lt_mono; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
generalize (spec_to_Z x); auto with zarith.
generalize wB_pos; auto with zarith.
omega.
replace [|s|] with s by (symmetry; apply Zmod_small; auto with zarith).
destruct Z_lt_le_dec; unfold interp_carry.
replace [|r|] with r by (symmetry; apply Zmod_small; auto with zarith).
- rewrite Zpower_2; auto with zarith.
+ rewrite Z.pow_2_r; auto with zarith.
replace [|r-wB|] with (r-wB) by (symmetry; apply Zmod_small; auto with zarith).
- rewrite Zpower_2; omega.
+ rewrite Z.pow_2_r; omega.
assert (0<=Zneg p).
rewrite Heqz; generalize wB_pos; auto with zarith.
@@ -667,15 +666,15 @@ Section ZModulo.
cut (log_inf x < p - 1); [omega| ].
apply IHx.
change (Zpos x~1) with (2*(Zpos x)+1) in H.
- replace p with (Zsucc (p-1)) in H; auto with zarith.
- rewrite Zpower_Zsucc in H; auto with zarith.
+ replace p with (Z.succ (p-1)) in H; auto with zarith.
+ rewrite Z.pow_succ_r in H; auto with zarith.
assert (0 < p) by (destruct p; compute; auto with zarith; discriminate).
cut (log_inf x < p - 1); [omega| ].
apply IHx.
change (Zpos x~0) with (2*(Zpos x)) in H.
- replace p with (Zsucc (p-1)) in H; auto with zarith.
- rewrite Zpower_Zsucc in H; auto with zarith.
+ replace p with (Z.succ (p-1)) in H; auto with zarith.
+ rewrite Z.pow_succ_r in H; auto with zarith.
simpl; intros; destruct p; compute; auto with zarith.
Qed.
@@ -696,27 +695,27 @@ Section ZModulo.
unfold zdigits.
unfold wB, base in *.
apply log_inf_bounded; auto with zarith.
- apply Zlt_trans with zdigits.
+ apply Z.lt_trans with zdigits.
omega.
unfold zdigits, wB, base; apply Zpower2_lt_lin; auto with zarith.
unfold to_Z; rewrite (Zmod_small _ _ H3).
destruct H2.
split.
- apply Zle_trans with (2^(zdigits - log_inf p - 1)*(2^log_inf p)).
+ apply Z.le_trans with (2^(zdigits - log_inf p - 1)*(2^log_inf p)).
apply Zdiv_le_upper_bound; auto with zarith.
rewrite <- Zpower_exp; auto with zarith.
- rewrite Zmult_comm; rewrite <- Zpower_Zsucc; auto with zarith.
- replace (Zsucc (zdigits - log_inf p -1 +log_inf p)) with zdigits
+ rewrite Z.mul_comm; rewrite <- Z.pow_succ_r; auto with zarith.
+ replace (Z.succ (zdigits - log_inf p -1 +log_inf p)) with zdigits
by ring.
unfold wB, base, zdigits; auto with zarith.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
- apply Zlt_le_trans
- with (2^(zdigits - log_inf p - 1)*(2^(Zsucc (log_inf p)))).
- apply Zmult_lt_compat_l; auto with zarith.
+ apply Z.lt_le_trans
+ with (2^(zdigits - log_inf p - 1)*(2^(Z.succ (log_inf p)))).
+ apply Z.mul_lt_mono_pos_l; auto with zarith.
rewrite <- Zpower_exp; auto with zarith.
- replace (zdigits - log_inf p -1 +Zsucc (log_inf p)) with zdigits
+ replace (zdigits - log_inf p -1 +Z.succ (log_inf p)) with zdigits
by ring.
unfold wB, base, zdigits; auto with zarith.
Qed.
@@ -739,18 +738,18 @@ Section ZModulo.
assert (d <> xH).
intro; subst.
compute in H; destruct p; discriminate.
- assert (Zsucc (Zpos (Ppred d)) = Zpos d).
+ assert (Z.succ (Zpos (Pos.pred d)) = Zpos d).
simpl; f_equal.
- rewrite <- Pplus_one_succ_r.
- destruct (Psucc_pred d); auto.
+ rewrite Pos.add_1_r.
+ destruct (Pos.succ_pred_or d); auto.
rewrite H1 in H0; elim H0; auto.
- assert (Ptail p < Zpos (Ppred d)).
+ assert (Ptail p < Zpos (Pos.pred d)).
apply IHp.
- apply Zmult_lt_reg_r with 2; auto with zarith.
- rewrite (Zmult_comm (Zpos p)).
+ apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
+ rewrite (Z.mul_comm (Zpos p)).
change (2 * Zpos p) with (Zpos p~0).
- rewrite Zmult_comm.
- rewrite <- Zpower_Zsucc; auto with zarith.
+ rewrite Z.mul_comm.
+ rewrite <- Z.pow_succ_r; auto with zarith.
rewrite H1; auto.
rewrite <- H1; omega.
Qed.
@@ -779,20 +778,20 @@ Section ZModulo.
apply Zmod_small.
split; auto.
unfold wB, base in *.
- apply Zlt_trans with (Zpos digits).
+ apply Z.lt_trans with (Zpos digits).
apply Ptail_bounded; auto with zarith.
apply Zpower2_lt_lin; auto with zarith.
rewrite H1.
clear; induction p.
- exists (Zpos p); simpl; rewrite Pmult_1_r; auto with zarith.
+ exists (Zpos p); simpl; rewrite Pos.mul_1_r; auto with zarith.
destruct IHp as (y & Yp & Ye).
exists y.
split; auto.
change (Zpos p~0) with (2*Zpos p).
rewrite Ye.
- change (Ptail p~0) with (Zsucc (Ptail p)).
- rewrite Zpower_Zsucc; auto; ring.
+ change (Ptail p~0) with (Z.succ (Ptail p)).
+ rewrite Z.pow_succ_r; auto; ring.
exists 0; simpl; auto with zarith.
Qed.
diff --git a/theories/Numbers/Integer/Abstract/ZBits.v b/theories/Numbers/Integer/Abstract/ZBits.v
index 92afbcb53..7c3a99215 100644
--- a/theories/Numbers/Integer/Abstract/ZBits.v
+++ b/theories/Numbers/Integer/Abstract/ZBits.v
@@ -292,7 +292,7 @@ Proof.
Qed.
(** Hence the number of bits of [a] is [1+log2 a]
- (see [Psize] and [Psize_pos]).
+ (see [Pos.size_nat] and [Pos.size]).
*)
(** For negative numbers, things are the other ways around:
diff --git a/theories/Numbers/Integer/Abstract/ZDivFloor.v b/theories/Numbers/Integer/Abstract/ZDivFloor.v
index 14003d892..24c8dc8b0 100644
--- a/theories/Numbers/Integer/Abstract/ZDivFloor.v
+++ b/theories/Numbers/Integer/Abstract/ZDivFloor.v
@@ -16,7 +16,7 @@ Require Import ZAxioms ZMulOrder ZSgnAbs NZDiv.
[a = bq+r /\ 0 <= |r| < |b| /\ Sign(r) = Sign(b)]
- This is the convention followed historically by [Zdiv] in Coq, and
+ This is the convention followed historically by [Z.div] in Coq, and
corresponds to convention "F" in the following paper:
R. Boute, "The Euclidean definition of the functions div and mod",
diff --git a/theories/Numbers/Integer/Abstract/ZPow.v b/theories/Numbers/Integer/Abstract/ZPow.v
index 53d84dce1..c9c98b1e5 100644
--- a/theories/Numbers/Integer/Abstract/ZPow.v
+++ b/theories/Numbers/Integer/Abstract/ZPow.v
@@ -18,6 +18,17 @@ Module Type ZPowProp
Include NZPowProp A A B.
+(** A particular case of [pow_add_r], with no precondition *)
+
+Lemma pow_twice_r a b : a^(2*b) == a^b * a^b.
+Proof.
+ rewrite two_succ. nzsimpl.
+ destruct (le_gt_cases 0 b).
+ - now rewrite pow_add_r.
+ - rewrite !pow_neg_r. now nzsimpl. trivial.
+ now apply add_neg_neg.
+Qed.
+
(** Parity of power *)
Lemma even_pow : forall a b, 0<b -> even (a^b) = even a.
diff --git a/theories/Numbers/Integer/BigZ/BigZ.v b/theories/Numbers/Integer/BigZ/BigZ.v
index 443777f52..ae1701be7 100644
--- a/theories/Numbers/Integer/BigZ/BigZ.v
+++ b/theories/Numbers/Integer/BigZ/BigZ.v
@@ -76,7 +76,7 @@ Infix "÷" := BigZ.quot (at level 40, left associativity) : bigZ_scope.
(** Some additional results about [BigZ] *)
Theorem spec_to_Z: forall n : bigZ,
- BigN.to_Z (BigZ.to_N n) = ((Zsgn [n]) * [n])%Z.
+ BigN.to_Z (BigZ.to_N n) = ((Z.sgn [n]) * [n])%Z.
Proof.
intros n; case n; simpl; intros p;
generalize (BigN.spec_pos p); case (BigN.to_Z p); auto.
@@ -85,7 +85,7 @@ intros p1 H1; case H1; auto.
Qed.
Theorem spec_to_N n:
- ([n] = Zsgn [n] * (BigN.to_Z (BigZ.to_N n)))%Z.
+ ([n] = Z.sgn [n] * (BigN.to_Z (BigZ.to_N n)))%Z.
Proof.
case n; simpl; intros p;
generalize (BigN.spec_pos p); case (BigN.to_Z p); auto.
@@ -118,7 +118,7 @@ Qed.
Lemma BigZeqb_correct : forall x y, (x =? y) = true -> x==y.
Proof. now apply BigZ.eqb_eq. Qed.
-Definition BigZ_of_N n := BigZ.of_Z (Z_of_N n).
+Definition BigZ_of_N n := BigZ.of_Z (Z.of_N n).
Lemma BigZpower : power_theory 1 BigZ.mul BigZ.eq BigZ_of_N BigZ.pow.
Proof.
@@ -139,7 +139,7 @@ BigZ.zify. auto with zarith.
intros NEQ.
generalize (BigZ.spec_div_eucl a b).
generalize (Z_div_mod_full [a] [b] NEQ).
-destruct BigZ.div_eucl as (q,r), Zdiv_eucl as (q',r').
+destruct BigZ.div_eucl as (q,r), Z.div_eucl as (q',r').
intros (EQ,_). injection 1. intros EQr EQq.
BigZ.zify. rewrite EQr, EQq; auto.
Qed.
diff --git a/theories/Numbers/Integer/BigZ/ZMake.v b/theories/Numbers/Integer/BigZ/ZMake.v
index 0142b36be..d1c817f54 100644
--- a/theories/Numbers/Integer/BigZ/ZMake.v
+++ b/theories/Numbers/Integer/BigZ/ZMake.v
@@ -21,92 +21,92 @@ Open Scope Z_scope.
[NSig.NType] to a structure of integers [ZSig.ZType].
*)
-Module Make (N:NType) <: ZType.
+Module Make (NN:NType) <: ZType.
Inductive t_ :=
- | Pos : N.t -> t_
- | Neg : N.t -> t_.
+ | Pos : NN.t -> t_
+ | Neg : NN.t -> t_.
Definition t := t_.
Bind Scope abstract_scope with t t_.
- Definition zero := Pos N.zero.
- Definition one := Pos N.one.
- Definition two := Pos N.two.
- Definition minus_one := Neg N.one.
+ Definition zero := Pos NN.zero.
+ Definition one := Pos NN.one.
+ Definition two := Pos NN.two.
+ Definition minus_one := Neg NN.one.
Definition of_Z x :=
match x with
- | Zpos x => Pos (N.of_N (Npos x))
+ | Zpos x => Pos (NN.of_N (Npos x))
| Z0 => zero
- | Zneg x => Neg (N.of_N (Npos x))
+ | Zneg x => Neg (NN.of_N (Npos x))
end.
Definition to_Z x :=
match x with
- | Pos nx => N.to_Z nx
- | Neg nx => Zopp (N.to_Z nx)
+ | Pos nx => NN.to_Z nx
+ | Neg nx => Z.opp (NN.to_Z nx)
end.
Theorem spec_of_Z: forall x, to_Z (of_Z x) = x.
Proof.
intros x; case x; unfold to_Z, of_Z, zero.
- exact N.spec_0.
- intros; rewrite N.spec_of_N; auto.
- intros; rewrite N.spec_of_N; auto.
+ exact NN.spec_0.
+ intros; rewrite NN.spec_of_N; auto.
+ intros; rewrite NN.spec_of_N; auto.
Qed.
Definition eq x y := (to_Z x = to_Z y).
Theorem spec_0: to_Z zero = 0.
- exact N.spec_0.
+ exact NN.spec_0.
Qed.
Theorem spec_1: to_Z one = 1.
- exact N.spec_1.
+ exact NN.spec_1.
Qed.
Theorem spec_2: to_Z two = 2.
- exact N.spec_2.
+ exact NN.spec_2.
Qed.
Theorem spec_m1: to_Z minus_one = -1.
- simpl; rewrite N.spec_1; auto.
+ simpl; rewrite NN.spec_1; auto.
Qed.
Definition compare x y :=
match x, y with
- | Pos nx, Pos ny => N.compare nx ny
+ | Pos nx, Pos ny => NN.compare nx ny
| Pos nx, Neg ny =>
- match N.compare nx N.zero with
+ match NN.compare nx NN.zero with
| Gt => Gt
- | _ => N.compare ny N.zero
+ | _ => NN.compare ny NN.zero
end
| Neg nx, Pos ny =>
- match N.compare N.zero nx with
+ match NN.compare NN.zero nx with
| Lt => Lt
- | _ => N.compare N.zero ny
+ | _ => NN.compare NN.zero ny
end
- | Neg nx, Neg ny => N.compare ny nx
+ | Neg nx, Neg ny => NN.compare ny nx
end.
Theorem spec_compare :
- forall x y, compare x y = Zcompare (to_Z x) (to_Z y).
+ forall x y, compare x y = Z.compare (to_Z x) (to_Z y).
Proof.
unfold compare, to_Z.
destruct x as [x|x], y as [y|y];
- rewrite ?N.spec_compare, ?N.spec_0, <-?Zcompare_opp; auto;
- assert (Hx:=N.spec_pos x); assert (Hy:=N.spec_pos y);
- set (X:=N.to_Z x) in *; set (Y:=N.to_Z y) in *; clearbody X Y.
- destruct (Zcompare_spec X 0) as [EQ|LT|GT].
- rewrite EQ. rewrite <- Zopp_0 at 2. apply Zcompare_opp.
- exfalso. omega.
- symmetry. change (X > -Y). omega.
- destruct (Zcompare_spec 0 X) as [EQ|LT|GT].
- rewrite <- EQ. rewrite Zopp_0; auto.
- symmetry. change (-X < Y). omega.
- exfalso. omega.
+ rewrite ?NN.spec_compare, ?NN.spec_0, ?Z.compare_opp; auto;
+ assert (Hx:=NN.spec_pos x); assert (Hy:=NN.spec_pos y);
+ set (X:=NN.to_Z x) in *; set (Y:=NN.to_Z y) in *; clearbody X Y.
+ - destruct (Z.compare_spec X 0) as [EQ|LT|GT].
+ + rewrite <- Z.opp_0 in EQ. now rewrite EQ, Z.compare_opp.
+ + exfalso. omega.
+ + symmetry. change (X > -Y). omega.
+ - destruct (Z.compare_spec 0 X) as [EQ|LT|GT].
+ + rewrite <- EQ, Z.opp_0; auto.
+ + symmetry. change (-X < Y). omega.
+ + exfalso. omega.
Qed.
Definition eqb x y :=
@@ -155,14 +155,14 @@ Module Make (N:NType) <: ZType.
Definition min n m := match compare n m with Gt => m | _ => n end.
Definition max n m := match compare n m with Lt => m | _ => n end.
- Theorem spec_min : forall n m, to_Z (min n m) = Zmin (to_Z n) (to_Z m).
+ Theorem spec_min : forall n m, to_Z (min n m) = Z.min (to_Z n) (to_Z m).
Proof.
- unfold min, Zmin. intros. rewrite spec_compare. destruct Zcompare; auto.
+ unfold min, Z.min. intros. rewrite spec_compare. destruct Z.compare; auto.
Qed.
- Theorem spec_max : forall n m, to_Z (max n m) = Zmax (to_Z n) (to_Z m).
+ Theorem spec_max : forall n m, to_Z (max n m) = Z.max (to_Z n) (to_Z m).
Proof.
- unfold max, Zmax. intros. rewrite spec_compare. destruct Zcompare; auto.
+ unfold max, Z.max. intros. rewrite spec_compare. destruct Z.compare; auto.
Qed.
Definition to_N x :=
@@ -173,11 +173,11 @@ Module Make (N:NType) <: ZType.
Definition abs x := Pos (to_N x).
- Theorem spec_abs: forall x, to_Z (abs x) = Zabs (to_Z x).
+ Theorem spec_abs: forall x, to_Z (abs x) = Z.abs (to_Z x).
Proof.
- intros x; case x; clear x; intros x; assert (F:=N.spec_pos x).
- simpl; rewrite Zabs_eq; auto.
- simpl; rewrite Zabs_non_eq; simpl; auto with zarith.
+ intros x; case x; clear x; intros x; assert (F:=NN.spec_pos x).
+ simpl; rewrite Z.abs_eq; auto.
+ simpl; rewrite Z.abs_neq; simpl; auto with zarith.
Qed.
Definition opp x :=
@@ -193,10 +193,10 @@ Module Make (N:NType) <: ZType.
Definition succ x :=
match x with
- | Pos n => Pos (N.succ n)
+ | Pos n => Pos (NN.succ n)
| Neg n =>
- match N.compare N.zero n with
- | Lt => Neg (N.pred n)
+ match NN.compare NN.zero n with
+ | Lt => Neg (NN.pred n)
| _ => one
end
end.
@@ -204,134 +204,134 @@ Module Make (N:NType) <: ZType.
Theorem spec_succ: forall n, to_Z (succ n) = to_Z n + 1.
Proof.
intros x; case x; clear x; intros x.
- exact (N.spec_succ x).
- simpl. rewrite N.spec_compare. case Zcompare_spec; rewrite ?N.spec_0; simpl.
- intros HH; rewrite <- HH; rewrite N.spec_1; ring.
- intros HH; rewrite N.spec_pred, Zmax_r; auto with zarith.
- generalize (N.spec_pos x); auto with zarith.
+ exact (NN.spec_succ x).
+ simpl. rewrite NN.spec_compare. case Z.compare_spec; rewrite ?NN.spec_0; simpl.
+ intros HH; rewrite <- HH; rewrite NN.spec_1; ring.
+ intros HH; rewrite NN.spec_pred, Z.max_r; auto with zarith.
+ generalize (NN.spec_pos x); auto with zarith.
Qed.
Definition add x y :=
match x, y with
- | Pos nx, Pos ny => Pos (N.add nx ny)
+ | Pos nx, Pos ny => Pos (NN.add nx ny)
| Pos nx, Neg ny =>
- match N.compare nx ny with
- | Gt => Pos (N.sub nx ny)
+ match NN.compare nx ny with
+ | Gt => Pos (NN.sub nx ny)
| Eq => zero
- | Lt => Neg (N.sub ny nx)
+ | Lt => Neg (NN.sub ny nx)
end
| Neg nx, Pos ny =>
- match N.compare nx ny with
- | Gt => Neg (N.sub nx ny)
+ match NN.compare nx ny with
+ | Gt => Neg (NN.sub nx ny)
| Eq => zero
- | Lt => Pos (N.sub ny nx)
+ | Lt => Pos (NN.sub ny nx)
end
- | Neg nx, Neg ny => Neg (N.add nx ny)
+ | Neg nx, Neg ny => Neg (NN.add nx ny)
end.
Theorem spec_add: forall x y, to_Z (add x y) = to_Z x + to_Z y.
Proof.
unfold add, to_Z; intros [x | x] [y | y];
- try (rewrite N.spec_add; auto with zarith);
- rewrite N.spec_compare; case Zcompare_spec;
- unfold zero; rewrite ?N.spec_0, ?N.spec_sub; omega with *.
+ try (rewrite NN.spec_add; auto with zarith);
+ rewrite NN.spec_compare; case Z.compare_spec;
+ unfold zero; rewrite ?NN.spec_0, ?NN.spec_sub; omega with *.
Qed.
Definition pred x :=
match x with
| Pos nx =>
- match N.compare N.zero nx with
- | Lt => Pos (N.pred nx)
+ match NN.compare NN.zero nx with
+ | Lt => Pos (NN.pred nx)
| _ => minus_one
end
- | Neg nx => Neg (N.succ nx)
+ | Neg nx => Neg (NN.succ nx)
end.
Theorem spec_pred: forall x, to_Z (pred x) = to_Z x - 1.
Proof.
unfold pred, to_Z, minus_one; intros [x | x];
- try (rewrite N.spec_succ; ring).
- rewrite N.spec_compare; case Zcompare_spec;
- rewrite ?N.spec_0, ?N.spec_1, ?N.spec_pred;
- generalize (N.spec_pos x); omega with *.
+ try (rewrite NN.spec_succ; ring).
+ rewrite NN.spec_compare; case Z.compare_spec;
+ rewrite ?NN.spec_0, ?NN.spec_1, ?NN.spec_pred;
+ generalize (NN.spec_pos x); omega with *.
Qed.
Definition sub x y :=
match x, y with
| Pos nx, Pos ny =>
- match N.compare nx ny with
- | Gt => Pos (N.sub nx ny)
+ match NN.compare nx ny with
+ | Gt => Pos (NN.sub nx ny)
| Eq => zero
- | Lt => Neg (N.sub ny nx)
+ | Lt => Neg (NN.sub ny nx)
end
- | Pos nx, Neg ny => Pos (N.add nx ny)
- | Neg nx, Pos ny => Neg (N.add nx ny)
+ | Pos nx, Neg ny => Pos (NN.add nx ny)
+ | Neg nx, Pos ny => Neg (NN.add nx ny)
| Neg nx, Neg ny =>
- match N.compare nx ny with
- | Gt => Neg (N.sub nx ny)
+ match NN.compare nx ny with
+ | Gt => Neg (NN.sub nx ny)
| Eq => zero
- | Lt => Pos (N.sub ny nx)
+ | Lt => Pos (NN.sub ny nx)
end
end.
Theorem spec_sub: forall x y, to_Z (sub x y) = to_Z x - to_Z y.
Proof.
unfold sub, to_Z; intros [x | x] [y | y];
- try (rewrite N.spec_add; auto with zarith);
- rewrite N.spec_compare; case Zcompare_spec;
- unfold zero; rewrite ?N.spec_0, ?N.spec_sub; omega with *.
+ try (rewrite NN.spec_add; auto with zarith);
+ rewrite NN.spec_compare; case Z.compare_spec;
+ unfold zero; rewrite ?NN.spec_0, ?NN.spec_sub; omega with *.
Qed.
Definition mul x y :=
match x, y with
- | Pos nx, Pos ny => Pos (N.mul nx ny)
- | Pos nx, Neg ny => Neg (N.mul nx ny)
- | Neg nx, Pos ny => Neg (N.mul nx ny)
- | Neg nx, Neg ny => Pos (N.mul nx ny)
+ | Pos nx, Pos ny => Pos (NN.mul nx ny)
+ | Pos nx, Neg ny => Neg (NN.mul nx ny)
+ | Neg nx, Pos ny => Neg (NN.mul nx ny)
+ | Neg nx, Neg ny => Pos (NN.mul nx ny)
end.
Theorem spec_mul: forall x y, to_Z (mul x y) = to_Z x * to_Z y.
Proof.
- unfold mul, to_Z; intros [x | x] [y | y]; rewrite N.spec_mul; ring.
+ unfold mul, to_Z; intros [x | x] [y | y]; rewrite NN.spec_mul; ring.
Qed.
Definition square x :=
match x with
- | Pos nx => Pos (N.square nx)
- | Neg nx => Pos (N.square nx)
+ | Pos nx => Pos (NN.square nx)
+ | Neg nx => Pos (NN.square nx)
end.
Theorem spec_square: forall x, to_Z (square x) = to_Z x * to_Z x.
Proof.
- unfold square, to_Z; intros [x | x]; rewrite N.spec_square; ring.
+ unfold square, to_Z; intros [x | x]; rewrite NN.spec_square; ring.
Qed.
Definition pow_pos x p :=
match x with
- | Pos nx => Pos (N.pow_pos nx p)
+ | Pos nx => Pos (NN.pow_pos nx p)
| Neg nx =>
match p with
| xH => x
- | xO _ => Pos (N.pow_pos nx p)
- | xI _ => Neg (N.pow_pos nx p)
+ | xO _ => Pos (NN.pow_pos nx p)
+ | xI _ => Neg (NN.pow_pos nx p)
end
end.
Theorem spec_pow_pos: forall x n, to_Z (pow_pos x n) = to_Z x ^ Zpos n.
Proof.
assert (F0: forall x, (-x)^2 = x^2).
- intros x; rewrite Zpower_2; ring.
+ intros x; rewrite Z.pow_2_r; ring.
unfold pow_pos, to_Z; intros [x | x] [p | p |];
- try rewrite N.spec_pow_pos; try ring.
+ try rewrite NN.spec_pow_pos; try ring.
assert (F: 0 <= 2 * Zpos p).
assert (0 <= Zpos p); auto with zarith.
- rewrite Zpos_xI; repeat rewrite Zpower_exp; auto with zarith.
- repeat rewrite Zpower_mult; auto with zarith.
+ rewrite Pos2Z.inj_xI; repeat rewrite Zpower_exp; auto with zarith.
+ repeat rewrite Z.pow_mul_r; auto with zarith.
rewrite F0; ring.
assert (F: 0 <= 2 * Zpos p).
assert (0 <= Zpos p); auto with zarith.
- rewrite Zpos_xO; repeat rewrite Zpower_exp; auto with zarith.
- repeat rewrite Zpower_mult; auto with zarith.
+ rewrite Pos2Z.inj_xO; repeat rewrite Zpower_exp; auto with zarith.
+ repeat rewrite Z.pow_mul_r; auto with zarith.
rewrite F0; ring.
Qed.
@@ -341,9 +341,9 @@ Module Make (N:NType) <: ZType.
| Npos p => pow_pos x p
end.
- Theorem spec_pow_N: forall x n, to_Z (pow_N x n) = to_Z x ^ Z_of_N n.
+ Theorem spec_pow_N: forall x n, to_Z (pow_N x n) = to_Z x ^ Z.of_N n.
Proof.
- destruct n; simpl. apply N.spec_1.
+ destruct n; simpl. apply NN.spec_1.
apply spec_pow_pos.
Qed.
@@ -357,38 +357,38 @@ Module Make (N:NType) <: ZType.
Theorem spec_pow: forall x y, to_Z (pow x y) = to_Z x ^ to_Z y.
Proof.
intros. unfold pow. destruct (to_Z y); simpl.
- apply N.spec_1.
+ apply NN.spec_1.
apply spec_pow_pos.
- apply N.spec_0.
+ apply NN.spec_0.
Qed.
Definition log2 x :=
match x with
- | Pos nx => Pos (N.log2 nx)
+ | Pos nx => Pos (NN.log2 nx)
| Neg nx => zero
end.
Theorem spec_log2: forall x, to_Z (log2 x) = Z.log2 (to_Z x).
Proof.
- intros. destruct x as [p|p]; simpl. apply N.spec_log2.
- rewrite N.spec_0.
- destruct (Z_le_lt_eq_dec _ _ (N.spec_pos p)) as [LT|EQ].
+ intros. destruct x as [p|p]; simpl. apply NN.spec_log2.
+ rewrite NN.spec_0.
+ destruct (Z_le_lt_eq_dec _ _ (NN.spec_pos p)) as [LT|EQ].
rewrite Z.log2_nonpos; auto with zarith.
now rewrite <- EQ.
Qed.
Definition sqrt x :=
match x with
- | Pos nx => Pos (N.sqrt nx)
- | Neg nx => Neg N.zero
+ | Pos nx => Pos (NN.sqrt nx)
+ | Neg nx => Neg NN.zero
end.
Theorem spec_sqrt: forall x, to_Z (sqrt x) = Z.sqrt (to_Z x).
Proof.
destruct x as [p|p]; simpl.
- apply N.spec_sqrt.
- rewrite N.spec_0.
- destruct (Z_le_lt_eq_dec _ _ (N.spec_pos p)) as [LT|EQ].
+ apply NN.spec_sqrt.
+ rewrite NN.spec_0.
+ destruct (Z_le_lt_eq_dec _ _ (NN.spec_pos p)) as [LT|EQ].
rewrite Z.sqrt_neg; auto with zarith.
now rewrite <- EQ.
Qed.
@@ -396,68 +396,68 @@ Module Make (N:NType) <: ZType.
Definition div_eucl x y :=
match x, y with
| Pos nx, Pos ny =>
- let (q, r) := N.div_eucl nx ny in
+ let (q, r) := NN.div_eucl nx ny in
(Pos q, Pos r)
| Pos nx, Neg ny =>
- let (q, r) := N.div_eucl nx ny in
- if N.eqb N.zero r
+ let (q, r) := NN.div_eucl nx ny in
+ if NN.eqb NN.zero r
then (Neg q, zero)
- else (Neg (N.succ q), Neg (N.sub ny r))
+ else (Neg (NN.succ q), Neg (NN.sub ny r))
| Neg nx, Pos ny =>
- let (q, r) := N.div_eucl nx ny in
- if N.eqb N.zero r
+ let (q, r) := NN.div_eucl nx ny in
+ if NN.eqb NN.zero r
then (Neg q, zero)
- else (Neg (N.succ q), Pos (N.sub ny r))
+ else (Neg (NN.succ q), Pos (NN.sub ny r))
| Neg nx, Neg ny =>
- let (q, r) := N.div_eucl nx ny in
+ let (q, r) := NN.div_eucl nx ny in
(Pos q, Neg r)
end.
Ltac break_nonneg x px EQx :=
let H := fresh "H" in
- assert (H:=N.spec_pos x);
- destruct (N.to_Z x) as [|px|px]_eqn:EQx;
+ assert (H:=NN.spec_pos x);
+ destruct (NN.to_Z x) as [|px|px]_eqn:EQx;
[clear H|clear H|elim H; reflexivity].
Theorem spec_div_eucl: forall x y,
let (q,r) := div_eucl x y in
- (to_Z q, to_Z r) = Zdiv_eucl (to_Z x) (to_Z y).
+ (to_Z q, to_Z r) = Z.div_eucl (to_Z x) (to_Z y).
Proof.
unfold div_eucl, to_Z. intros [x | x] [y | y].
(* Pos Pos *)
- generalize (N.spec_div_eucl x y); destruct (N.div_eucl x y); auto.
+ generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y); auto.
(* Pos Neg *)
- generalize (N.spec_div_eucl x y); destruct (N.div_eucl x y) as (q,r).
+ generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y) as (q,r).
break_nonneg x px EQx; break_nonneg y py EQy;
- try (injection 1; intros Hr Hq; rewrite N.spec_eqb, N.spec_0, Hr;
- simpl; rewrite Hq, N.spec_0; auto).
+ try (injection 1; intros Hr Hq; rewrite NN.spec_eqb, NN.spec_0, Hr;
+ simpl; rewrite Hq, NN.spec_0; auto).
change (- Zpos py) with (Zneg py).
assert (GT : Zpos py > 0) by (compute; auto).
generalize (Z_div_mod (Zpos px) (Zpos py) GT).
- unfold Zdiv_eucl. destruct (Zdiv_eucl_POS px (Zpos py)) as (q',r').
+ unfold Z.div_eucl. destruct (Z.pos_div_eucl px (Zpos py)) as (q',r').
intros (EQ,MOD). injection 1. intros Hr' Hq'.
- rewrite N.spec_eqb, N.spec_0, Hr'.
+ rewrite NN.spec_eqb, NN.spec_0, Hr'.
break_nonneg r pr EQr.
- subst; simpl. rewrite N.spec_0; auto.
+ subst; simpl. rewrite NN.spec_0; auto.
subst. lazy iota beta delta [Z.eqb].
- rewrite N.spec_sub, N.spec_succ, EQy, EQr. f_equal. omega with *.
+ rewrite NN.spec_sub, NN.spec_succ, EQy, EQr. f_equal. omega with *.
(* Neg Pos *)
- generalize (N.spec_div_eucl x y); destruct (N.div_eucl x y) as (q,r).
+ generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y) as (q,r).
break_nonneg x px EQx; break_nonneg y py EQy;
- try (injection 1; intros Hr Hq; rewrite N.spec_eqb, N.spec_0, Hr;
- simpl; rewrite Hq, N.spec_0; auto).
+ try (injection 1; intros Hr Hq; rewrite NN.spec_eqb, NN.spec_0, Hr;
+ simpl; rewrite Hq, NN.spec_0; auto).
change (- Zpos px) with (Zneg px).
assert (GT : Zpos py > 0) by (compute; auto).
generalize (Z_div_mod (Zpos px) (Zpos py) GT).
- unfold Zdiv_eucl. destruct (Zdiv_eucl_POS px (Zpos py)) as (q',r').
+ unfold Z.div_eucl. destruct (Z.pos_div_eucl px (Zpos py)) as (q',r').
intros (EQ,MOD). injection 1. intros Hr' Hq'.
- rewrite N.spec_eqb, N.spec_0, Hr'.
+ rewrite NN.spec_eqb, NN.spec_0, Hr'.
break_nonneg r pr EQr.
- subst; simpl. rewrite N.spec_0; auto.
+ subst; simpl. rewrite NN.spec_0; auto.
subst. lazy iota beta delta [Z.eqb].
- rewrite N.spec_sub, N.spec_succ, EQy, EQr. f_equal. omega with *.
+ rewrite NN.spec_sub, NN.spec_succ, EQy, EQr. f_equal. omega with *.
(* Neg Neg *)
- generalize (N.spec_div_eucl x y); destruct (N.div_eucl x y) as (q,r).
+ generalize (NN.spec_div_eucl x y); destruct (NN.div_eucl x y) as (q,r).
break_nonneg x px EQx; break_nonneg y py EQy;
try (injection 1; intros Hr Hq; rewrite Hr, Hq; auto).
simpl. intros <-; auto.
@@ -468,8 +468,8 @@ Module Make (N:NType) <: ZType.
Definition spec_div: forall x y,
to_Z (div x y) = to_Z x / to_Z y.
Proof.
- intros x y; generalize (spec_div_eucl x y); unfold div, Zdiv.
- case div_eucl; case Zdiv_eucl; simpl; auto.
+ intros x y; generalize (spec_div_eucl x y); unfold div, Z.div.
+ case div_eucl; case Z.div_eucl; simpl; auto.
intros q r q11 r1 H; injection H; auto.
Qed.
@@ -478,38 +478,38 @@ Module Make (N:NType) <: ZType.
Theorem spec_modulo:
forall x y, to_Z (modulo x y) = to_Z x mod to_Z y.
Proof.
- intros x y; generalize (spec_div_eucl x y); unfold modulo, Zmod.
- case div_eucl; case Zdiv_eucl; simpl; auto.
+ intros x y; generalize (spec_div_eucl x y); unfold modulo, Z.modulo.
+ case div_eucl; case Z.div_eucl; simpl; auto.
intros q r q11 r1 H; injection H; auto.
Qed.
Definition quot x y :=
match x, y with
- | Pos nx, Pos ny => Pos (N.div nx ny)
- | Pos nx, Neg ny => Neg (N.div nx ny)
- | Neg nx, Pos ny => Neg (N.div nx ny)
- | Neg nx, Neg ny => Pos (N.div nx ny)
+ | Pos nx, Pos ny => Pos (NN.div nx ny)
+ | Pos nx, Neg ny => Neg (NN.div nx ny)
+ | Neg nx, Pos ny => Neg (NN.div nx ny)
+ | Neg nx, Neg ny => Pos (NN.div nx ny)
end.
Definition rem x y :=
if eqb y zero then x
else
match x, y with
- | Pos nx, Pos ny => Pos (N.modulo nx ny)
- | Pos nx, Neg ny => Pos (N.modulo nx ny)
- | Neg nx, Pos ny => Neg (N.modulo nx ny)
- | Neg nx, Neg ny => Neg (N.modulo nx ny)
+ | Pos nx, Pos ny => Pos (NN.modulo nx ny)
+ | Pos nx, Neg ny => Pos (NN.modulo nx ny)
+ | Neg nx, Pos ny => Neg (NN.modulo nx ny)
+ | Neg nx, Neg ny => Neg (NN.modulo nx ny)
end.
Lemma spec_quot : forall x y, to_Z (quot x y) = (to_Z x) ÷ (to_Z y).
Proof.
- intros [x|x] [y|y]; simpl; symmetry; rewrite N.spec_div;
+ intros [x|x] [y|y]; simpl; symmetry; rewrite NN.spec_div;
(* Nota: we rely here on [forall a b, a ÷ 0 = b / 0] *)
- destruct (Z.eq_dec (N.to_Z y) 0) as [EQ|NEQ];
- try (rewrite EQ; now destruct (N.to_Z x));
+ destruct (Z.eq_dec (NN.to_Z y) 0) as [EQ|NEQ];
+ try (rewrite EQ; now destruct (NN.to_Z x));
rewrite ?Z.quot_opp_r, ?Z.quot_opp_l, ?Z.opp_involutive, ?Z.opp_inj_wd;
trivial; apply Z.quot_div_nonneg;
- generalize (N.spec_pos x) (N.spec_pos y); Z.order.
+ generalize (NN.spec_pos x) (NN.spec_pos y); Z.order.
Qed.
Lemma spec_rem : forall x y,
@@ -521,26 +521,26 @@ Module Make (N:NType) <: ZType.
rewrite Hy. now destruct (to_Z x).
destruct x as [x|x], y as [y|y]; simpl in *; symmetry;
rewrite ?Z.eq_opp_l, ?Z.opp_0 in Hy;
- rewrite N.spec_modulo, ?Z.rem_opp_r, ?Z.rem_opp_l, ?Z.opp_involutive,
+ rewrite NN.spec_modulo, ?Z.rem_opp_r, ?Z.rem_opp_l, ?Z.opp_involutive,
?Z.opp_inj_wd;
trivial; apply Z.rem_mod_nonneg;
- generalize (N.spec_pos x) (N.spec_pos y); Z.order.
+ generalize (NN.spec_pos x) (NN.spec_pos y); Z.order.
Qed.
Definition gcd x y :=
match x, y with
- | Pos nx, Pos ny => Pos (N.gcd nx ny)
- | Pos nx, Neg ny => Pos (N.gcd nx ny)
- | Neg nx, Pos ny => Pos (N.gcd nx ny)
- | Neg nx, Neg ny => Pos (N.gcd nx ny)
+ | Pos nx, Pos ny => Pos (NN.gcd nx ny)
+ | Pos nx, Neg ny => Pos (NN.gcd nx ny)
+ | Neg nx, Pos ny => Pos (NN.gcd nx ny)
+ | Neg nx, Neg ny => Pos (NN.gcd nx ny)
end.
- Theorem spec_gcd: forall a b, to_Z (gcd a b) = Zgcd (to_Z a) (to_Z b).
+ Theorem spec_gcd: forall a b, to_Z (gcd a b) = Z.gcd (to_Z a) (to_Z b).
Proof.
- unfold gcd, Zgcd, to_Z; intros [x | x] [y | y]; rewrite N.spec_gcd; unfold Zgcd;
- auto; case N.to_Z; simpl; auto with zarith;
- try rewrite Zabs_Zopp; auto;
- case N.to_Z; simpl; auto with zarith.
+ unfold gcd, Z.gcd, to_Z; intros [x | x] [y | y]; rewrite NN.spec_gcd; unfold Z.gcd;
+ auto; case NN.to_Z; simpl; auto with zarith;
+ try rewrite Z.abs_opp; auto;
+ case NN.to_Z; simpl; auto with zarith.
Qed.
Definition sgn x :=
@@ -550,124 +550,124 @@ Module Make (N:NType) <: ZType.
| Gt => minus_one
end.
- Lemma spec_sgn : forall x, to_Z (sgn x) = Zsgn (to_Z x).
+ Lemma spec_sgn : forall x, to_Z (sgn x) = Z.sgn (to_Z x).
Proof.
- intros. unfold sgn. rewrite spec_compare. case Zcompare_spec.
+ intros. unfold sgn. rewrite spec_compare. case Z.compare_spec.
rewrite spec_0. intros <-; auto.
- rewrite spec_0, spec_1. symmetry. rewrite Zsgn_pos; auto.
- rewrite spec_0, spec_m1. symmetry. rewrite Zsgn_neg; auto with zarith.
+ rewrite spec_0, spec_1. symmetry. rewrite Z.sgn_pos_iff; auto.
+ rewrite spec_0, spec_m1. symmetry. rewrite Z.sgn_neg_iff; auto with zarith.
Qed.
Definition even z :=
match z with
- | Pos n => N.even n
- | Neg n => N.even n
+ | Pos n => NN.even n
+ | Neg n => NN.even n
end.
Definition odd z :=
match z with
- | Pos n => N.odd n
- | Neg n => N.odd n
+ | Pos n => NN.odd n
+ | Neg n => NN.odd n
end.
- Lemma spec_even : forall z, even z = Zeven_bool (to_Z z).
+ Lemma spec_even : forall z, even z = Z.even (to_Z z).
Proof.
- intros [n|n]; simpl; rewrite N.spec_even; trivial.
- destruct (N.to_Z n) as [|p|p]; now try destruct p.
+ intros [n|n]; simpl; rewrite NN.spec_even; trivial.
+ destruct (NN.to_Z n) as [|p|p]; now try destruct p.
Qed.
- Lemma spec_odd : forall z, odd z = Zodd_bool (to_Z z).
+ Lemma spec_odd : forall z, odd z = Z.odd (to_Z z).
Proof.
- intros [n|n]; simpl; rewrite N.spec_odd; trivial.
- destruct (N.to_Z n) as [|p|p]; now try destruct p.
+ intros [n|n]; simpl; rewrite NN.spec_odd; trivial.
+ destruct (NN.to_Z n) as [|p|p]; now try destruct p.
Qed.
Definition norm_pos z :=
match z with
| Pos _ => z
- | Neg n => if N.eqb n N.zero then Pos n else z
+ | Neg n => if NN.eqb n NN.zero then Pos n else z
end.
Definition testbit a n :=
match norm_pos n, norm_pos a with
- | Pos p, Pos a => N.testbit a p
- | Pos p, Neg a => negb (N.testbit (N.pred a) p)
+ | Pos p, Pos a => NN.testbit a p
+ | Pos p, Neg a => negb (NN.testbit (NN.pred a) p)
| Neg p, _ => false
end.
Definition shiftl a n :=
match norm_pos a, n with
- | Pos a, Pos n => Pos (N.shiftl a n)
- | Pos a, Neg n => Pos (N.shiftr a n)
- | Neg a, Pos n => Neg (N.shiftl a n)
- | Neg a, Neg n => Neg (N.succ (N.shiftr (N.pred a) n))
+ | Pos a, Pos n => Pos (NN.shiftl a n)
+ | Pos a, Neg n => Pos (NN.shiftr a n)
+ | Neg a, Pos n => Neg (NN.shiftl a n)
+ | Neg a, Neg n => Neg (NN.succ (NN.shiftr (NN.pred a) n))
end.
Definition shiftr a n := shiftl a (opp n).
Definition lor a b :=
match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (N.lor a b)
- | Neg a, Pos b => Neg (N.succ (N.ldiff (N.pred a) b))
- | Pos a, Neg b => Neg (N.succ (N.ldiff (N.pred b) a))
- | Neg a, Neg b => Neg (N.succ (N.land (N.pred a) (N.pred b)))
+ | Pos a, Pos b => Pos (NN.lor a b)
+ | Neg a, Pos b => Neg (NN.succ (NN.ldiff (NN.pred a) b))
+ | Pos a, Neg b => Neg (NN.succ (NN.ldiff (NN.pred b) a))
+ | Neg a, Neg b => Neg (NN.succ (NN.land (NN.pred a) (NN.pred b)))
end.
Definition land a b :=
match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (N.land a b)
- | Neg a, Pos b => Pos (N.ldiff b (N.pred a))
- | Pos a, Neg b => Pos (N.ldiff a (N.pred b))
- | Neg a, Neg b => Neg (N.succ (N.lor (N.pred a) (N.pred b)))
+ | Pos a, Pos b => Pos (NN.land a b)
+ | Neg a, Pos b => Pos (NN.ldiff b (NN.pred a))
+ | Pos a, Neg b => Pos (NN.ldiff a (NN.pred b))
+ | Neg a, Neg b => Neg (NN.succ (NN.lor (NN.pred a) (NN.pred b)))
end.
Definition ldiff a b :=
match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (N.ldiff a b)
- | Neg a, Pos b => Neg (N.succ (N.lor (N.pred a) b))
- | Pos a, Neg b => Pos (N.land a (N.pred b))
- | Neg a, Neg b => Pos (N.ldiff (N.pred b) (N.pred a))
+ | Pos a, Pos b => Pos (NN.ldiff a b)
+ | Neg a, Pos b => Neg (NN.succ (NN.lor (NN.pred a) b))
+ | Pos a, Neg b => Pos (NN.land a (NN.pred b))
+ | Neg a, Neg b => Pos (NN.ldiff (NN.pred b) (NN.pred a))
end.
Definition lxor a b :=
match norm_pos a, norm_pos b with
- | Pos a, Pos b => Pos (N.lxor a b)
- | Neg a, Pos b => Neg (N.succ (N.lxor (N.pred a) b))
- | Pos a, Neg b => Neg (N.succ (N.lxor a (N.pred b)))
- | Neg a, Neg b => Pos (N.lxor (N.pred a) (N.pred b))
+ | Pos a, Pos b => Pos (NN.lxor a b)
+ | Neg a, Pos b => Neg (NN.succ (NN.lxor (NN.pred a) b))
+ | Pos a, Neg b => Neg (NN.succ (NN.lxor a (NN.pred b)))
+ | Neg a, Neg b => Pos (NN.lxor (NN.pred a) (NN.pred b))
end.
Definition div2 x := shiftr x one.
Lemma Zlnot_alt1 : forall x, -(x+1) = Z.lnot x.
Proof.
- unfold Z.lnot, Zpred; auto with zarith.
+ unfold Z.lnot, Z.pred; auto with zarith.
Qed.
Lemma Zlnot_alt2 : forall x, Z.lnot (x-1) = -x.
Proof.
- unfold Z.lnot, Zpred; auto with zarith.
+ unfold Z.lnot, Z.pred; auto with zarith.
Qed.
Lemma Zlnot_alt3 : forall x, Z.lnot (-x) = x-1.
Proof.
- unfold Z.lnot, Zpred; auto with zarith.
+ unfold Z.lnot, Z.pred; auto with zarith.
Qed.
Lemma spec_norm_pos : forall x, to_Z (norm_pos x) = to_Z x.
Proof.
intros [x|x]; simpl; trivial.
- rewrite N.spec_eqb, N.spec_0.
+ rewrite NN.spec_eqb, NN.spec_0.
case Z.eqb_spec; simpl; auto with zarith.
Qed.
Lemma spec_norm_pos_pos : forall x y, norm_pos x = Neg y ->
- 0 < N.to_Z y.
+ 0 < NN.to_Z y.
Proof.
intros [x|x] y; simpl; try easy.
- rewrite N.spec_eqb, N.spec_0.
+ rewrite NN.spec_eqb, NN.spec_0.
case Z.eqb_spec; simpl; try easy.
- inversion 2. subst. generalize (N.spec_pos y); auto with zarith.
+ inversion 2. subst. generalize (NN.spec_pos y); auto with zarith.
Qed.
Ltac destr_norm_pos x :=
@@ -682,9 +682,9 @@ Module Make (N:NType) <: ZType.
Proof.
intros x p. unfold testbit.
destr_norm_pos p; simpl. destr_norm_pos x; simpl.
- apply N.spec_testbit.
- rewrite N.spec_testbit, N.spec_pred, Zmax_r by auto with zarith.
- symmetry. apply Z.bits_opp. apply N.spec_pos.
+ apply NN.spec_testbit.
+ rewrite NN.spec_testbit, NN.spec_pred, Z.max_r by auto with zarith.
+ symmetry. apply Z.bits_opp. apply NN.spec_pos.
symmetry. apply Z.testbit_neg_r; auto with zarith.
Qed.
@@ -692,13 +692,13 @@ Module Make (N:NType) <: ZType.
Proof.
intros x p. unfold shiftl.
destr_norm_pos x; destruct p as [p|p]; simpl;
- assert (Hp := N.spec_pos p).
- apply N.spec_shiftl.
- rewrite Z.shiftl_opp_r. apply N.spec_shiftr.
- rewrite !N.spec_shiftl.
- rewrite !Z.shiftl_mul_pow2 by apply N.spec_pos.
- apply Zopp_mult_distr_l.
- rewrite Z.shiftl_opp_r, N.spec_succ, N.spec_shiftr, N.spec_pred, Zmax_r
+ assert (Hp := NN.spec_pos p).
+ apply NN.spec_shiftl.
+ rewrite Z.shiftl_opp_r. apply NN.spec_shiftr.
+ rewrite !NN.spec_shiftl.
+ rewrite !Z.shiftl_mul_pow2 by apply NN.spec_pos.
+ symmetry. apply Z.mul_opp_l.
+ rewrite Z.shiftl_opp_r, NN.spec_succ, NN.spec_shiftr, NN.spec_pred, Z.max_r
by auto with zarith.
now rewrite Zlnot_alt1, Z.lnot_shiftr, Zlnot_alt2.
Qed.
@@ -713,8 +713,8 @@ Module Make (N:NType) <: ZType.
Proof.
intros x y. unfold land.
destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?N.spec_succ, ?N.spec_land, ?N.spec_ldiff, ?N.spec_lor,
- ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1; auto with zarith.
+ rewrite ?NN.spec_succ, ?NN.spec_land, ?NN.spec_ldiff, ?NN.spec_lor,
+ ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1; auto with zarith.
now rewrite Z.ldiff_land, Zlnot_alt2.
now rewrite Z.ldiff_land, Z.land_comm, Zlnot_alt2.
now rewrite Z.lnot_lor, !Zlnot_alt2.
@@ -724,8 +724,8 @@ Module Make (N:NType) <: ZType.
Proof.
intros x y. unfold lor.
destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?N.spec_succ, ?N.spec_land, ?N.spec_ldiff, ?N.spec_lor,
- ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1; auto with zarith.
+ rewrite ?NN.spec_succ, ?NN.spec_land, ?NN.spec_ldiff, ?NN.spec_lor,
+ ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1; auto with zarith.
now rewrite Z.lnot_ldiff, Z.lor_comm, Zlnot_alt2.
now rewrite Z.lnot_ldiff, Zlnot_alt2.
now rewrite Z.lnot_land, !Zlnot_alt2.
@@ -735,8 +735,8 @@ Module Make (N:NType) <: ZType.
Proof.
intros x y. unfold ldiff.
destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?N.spec_succ, ?N.spec_land, ?N.spec_ldiff, ?N.spec_lor,
- ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1; auto with zarith.
+ rewrite ?NN.spec_succ, ?NN.spec_land, ?NN.spec_ldiff, ?NN.spec_lor,
+ ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1; auto with zarith.
now rewrite Z.ldiff_land, Zlnot_alt3.
now rewrite Z.lnot_lor, Z.ldiff_land, <- Zlnot_alt2.
now rewrite 2 Z.ldiff_land, Zlnot_alt2, Z.land_comm, Zlnot_alt3.
@@ -746,7 +746,7 @@ Module Make (N:NType) <: ZType.
Proof.
intros x y. unfold lxor.
destr_norm_pos x; destr_norm_pos y; simpl;
- rewrite ?N.spec_succ, ?N.spec_lxor, ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1;
+ rewrite ?NN.spec_succ, ?NN.spec_lxor, ?NN.spec_pred, ?Z.max_r, ?Zlnot_alt1;
auto with zarith.
now rewrite !Z.lnot_lxor_r, Zlnot_alt2.
now rewrite !Z.lnot_lxor_l, Zlnot_alt2.
diff --git a/theories/Numbers/Integer/Binary/ZBinary.v b/theories/Numbers/Integer/Binary/ZBinary.v
index d7c0abd8a..a756eb256 100644
--- a/theories/Numbers/Integer/Binary/ZBinary.v
+++ b/theories/Numbers/Integer/Binary/ZBinary.v
@@ -36,7 +36,7 @@ End TestOrder.
(** Z forms a ring *)
-(*Lemma Zring : ring_theory 0 1 NZadd NZmul NZsub Zopp NZeq.
+(*Lemma Zring : ring_theory 0 1 NZadd NZmul NZsub Z.opp NZeq.
Proof.
constructor.
exact Zadd_0_l.
diff --git a/theories/Numbers/Integer/NatPairs/ZNatPairs.v b/theories/Numbers/Integer/NatPairs/ZNatPairs.v
index dbcc1961e..6abc78164 100644
--- a/theories/Numbers/Integer/NatPairs/ZNatPairs.v
+++ b/theories/Numbers/Integer/NatPairs/ZNatPairs.v
@@ -59,7 +59,7 @@ Definition le (n m : t) := n#1 + m#2 <= m#1 + n#2.
Definition min (n m : t) : t := (min (n#1 + m#2) (m#1 + n#2), n#2 + m#2).
Definition max (n m : t) : t := (max (n#1 + m#2) (m#1 + n#2), n#2 + m#2).
-(** NB : We do not have [Zpred (Zsucc n) = n] but only [Zpred (Zsucc n) == n].
+(** NB : We do not have [Z.pred (Z.succ n) = n] but only [Z.pred (Z.succ n) == n].
It could be possible to consider as canonical only pairs where
one of the elements is 0, and make all operations convert
canonical values into other canonical values. In that case, we
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
index bfbc063ce..5d3836eff 100644
--- a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
+++ b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
@@ -93,7 +93,7 @@ replace z with (-(-z))%Z in * by (auto with zarith).
remember (-z)%Z as z'.
pattern z'; apply natlike_ind.
apply B0.
-intros; rewrite Zopp_succ; unfold Zpred; apply BP; auto.
+intros; rewrite Z.opp_succ; unfold Z.pred; apply BP; auto.
subst z'; auto with zarith.
Qed.
@@ -364,7 +364,7 @@ Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
Theorem div_mod : forall a b, ~b==0 -> a == b*(div a b) + (modulo a b).
Proof.
-intros a b. zify. intros. apply Z_div_mod_eq_full; auto.
+intros a b. zify. intros. apply Z.div_mod; auto.
Qed.
Theorem mod_pos_bound :
diff --git a/theories/Numbers/NatInt/NZAxioms.v b/theories/Numbers/NatInt/NZAxioms.v
index fcd98787f..ef68c3bca 100644
--- a/theories/Numbers/NatInt/NZAxioms.v
+++ b/theories/Numbers/NatInt/NZAxioms.v
@@ -43,7 +43,7 @@ End IsNZDomain.
(** Axiomatization of some more constants
Simply denoting "1" for (S 0) and so on works ok when implementing
- by nat, but leaves some (Nsucc N0) when implementing by N.
+ by nat, but leaves some (N.succ N0) when implementing by N.
*)
Module Type OneTwo (Import T:Typ).
diff --git a/theories/Numbers/NatInt/NZMulOrder.v b/theories/Numbers/NatInt/NZMulOrder.v
index 97306f934..f6de98560 100644
--- a/theories/Numbers/NatInt/NZMulOrder.v
+++ b/theories/Numbers/NatInt/NZMulOrder.v
@@ -271,9 +271,9 @@ Definition mul_eq_0 := eq_mul_0.
Definition mul_eq_0_l := eq_mul_0_l.
Definition mul_eq_0_r := eq_mul_0_r.
-Theorem lt_0_mul : forall n m, 0 < n * m <-> (0 < n /\ 0 < m) \/ (m < 0 /\ n < 0).
+Theorem lt_0_mul n m : 0 < n * m <-> (0 < n /\ 0 < m) \/ (m < 0 /\ n < 0).
Proof.
-intros n m; split; [intro H | intros [[H1 H2] | [H1 H2]]].
+split; [intro H | intros [[H1 H2] | [H1 H2]]].
destruct (lt_trichotomy n 0) as [H1 | [H1 | H1]];
[| rewrite H1 in H; rewrite mul_0_l in H; false_hyp H lt_irrefl |];
(destruct (lt_trichotomy m 0) as [H2 | [H2 | H2]];
diff --git a/theories/Numbers/Natural/Abstract/NBits.v b/theories/Numbers/Natural/Abstract/NBits.v
index c66f003ec..777e52eb2 100644
--- a/theories/Numbers/Natural/Abstract/NBits.v
+++ b/theories/Numbers/Natural/Abstract/NBits.v
@@ -236,7 +236,7 @@ Proof.
Qed.
(** Hence the number of bits of [a] is [1+log2 a]
- (see [Psize] and [Psize_pos]).
+ (see [Pos.size_nat] and [Pos.size]).
*)
(** Testing bits after division or multiplication by a power of two *)
diff --git a/theories/Numbers/Natural/BigN/BigN.v b/theories/Numbers/Natural/BigN/BigN.v
index 7f205b388..10eb3c208 100644
--- a/theories/Numbers/Natural/BigN/BigN.v
+++ b/theories/Numbers/Natural/BigN/BigN.v
@@ -119,7 +119,7 @@ BigN.zify. auto with zarith.
intros NEQ.
generalize (BigN.spec_div_eucl a b).
generalize (Z_div_mod_full [a] [b] NEQ).
-destruct BigN.div_eucl as (q,r), Zdiv_eucl as (q',r').
+destruct BigN.div_eucl as (q,r), Z.div_eucl as (q',r').
intros (EQ,_). injection 1. intros EQr EQq.
BigN.zify. rewrite EQr, EQq; auto.
Qed.
diff --git a/theories/Numbers/Natural/BigN/NMake.v b/theories/Numbers/Natural/BigN/NMake.v
index 952f61833..020706290 100644
--- a/theories/Numbers/Natural/BigN/NMake.v
+++ b/theories/Numbers/Natural/BigN/NMake.v
@@ -65,8 +65,8 @@ Module Make (W0:CyclicType) <: NType.
intros.
change (Zpos (ZnZ.digits (dom_op n)) <= Zpos (ZnZ.digits (dom_op m))).
rewrite !digits_dom_op, !Pshiftl_nat_Zpower.
- apply Zmult_le_compat_l; auto with zarith.
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
+ apply Z.pow_le_mono_r; auto with zarith.
Qed.
Definition to_N (x : t) := Z.to_N (to_Z x).
@@ -186,12 +186,12 @@ Module Make (W0:CyclicType) <: NType.
exact spec_0.
Qed.
- Lemma spec_pred : forall x, [pred x] = Zmax 0 ([x]-1).
+ Lemma spec_pred x : [pred x] = Z.max 0 ([x]-1).
Proof.
- intros. destruct (Zle_lt_or_eq _ _ (spec_pos x)).
- rewrite Zmax_r; auto with zarith.
- apply spec_pred_pos; auto.
- rewrite <- H; apply spec_pred0; auto.
+ rewrite Z.max_comm.
+ destruct (Z.max_spec ([x]-1) 0) as [(H,->)|(H,->)].
+ - apply spec_pred0; generalize (spec_pos x); auto with zarith.
+ - apply spec_pred_pos; auto with zarith.
Qed.
(** * Subtraction *)
@@ -230,11 +230,11 @@ Module Make (W0:CyclicType) <: NType.
exact spec_0.
Qed.
- Lemma spec_sub : forall x y, [sub x y] = Zmax 0 ([x]-[y]).
+ Lemma spec_sub : forall x y, [sub x y] = Z.max 0 ([x]-[y]).
Proof.
- intros. destruct (Zle_or_lt [y] [x]).
- rewrite Zmax_r; auto with zarith. apply spec_sub_pos; auto.
- rewrite Zmax_l; auto with zarith. apply spec_sub0; auto.
+ intros. destruct (Z.le_gt_cases [y] [x]).
+ rewrite Z.max_r; auto with zarith. apply spec_sub_pos; auto.
+ rewrite Z.max_l; auto with zarith. apply spec_sub0; auto.
Qed.
(** * Comparison *)
@@ -249,7 +249,7 @@ Module Make (W0:CyclicType) <: NType.
Let spec_comparen_m:
forall n m (x : word (dom_t n) (S m)) (y : dom_t n),
- comparen_m n m x y = Zcompare (eval n (S m) x) (ZnZ.to_Z y).
+ comparen_m n m x y = Z.compare (eval n (S m) x) (ZnZ.to_Z y).
Proof.
intros n m x y.
unfold comparen_m, eval.
@@ -287,10 +287,8 @@ Module Make (W0:CyclicType) <: NType.
lazy beta iota delta [iter_sym dom_op dom_t comparen_m]. reflexivity.
Qed.
-(** TODO: no need for ZnZ.Spec_rect , Spec_ind, and so on... *)
-
Theorem spec_compare : forall x y,
- compare x y = Zcompare [x] [y].
+ compare x y = Z.compare [x] [y].
Proof.
intros x y. rewrite compare_fold. apply spec_iter_sym; clear x y.
intros. apply ZnZ.spec_compare.
@@ -298,7 +296,7 @@ Module Make (W0:CyclicType) <: NType.
intros n m x y; unfold comparenm.
rewrite (spec_cast_l n m x), (spec_cast_r n m y).
unfold to_Z; apply ZnZ.spec_compare.
- intros. subst. apply Zcompare_antisym.
+ intros. subst. now rewrite <- Z.compare_antisym.
Qed.
Definition eqb (x y : t) : bool :=
@@ -346,14 +344,14 @@ Module Make (W0:CyclicType) <: NType.
Definition min (n m : t) : t := match compare n m with Gt => m | _ => n end.
Definition max (n m : t) : t := match compare n m with Lt => m | _ => n end.
- Theorem spec_max : forall n m, [max n m] = Zmax [n] [m].
+ Theorem spec_max : forall n m, [max n m] = Z.max [n] [m].
Proof.
- intros. unfold max, Zmax. rewrite spec_compare; destruct Zcompare; reflexivity.
+ intros. unfold max, Z.max. rewrite spec_compare; destruct Z.compare; reflexivity.
Qed.
- Theorem spec_min : forall n m, [min n m] = Zmin [n] [m].
+ Theorem spec_min : forall n m, [min n m] = Z.min [n] [m].
Proof.
- intros. unfold min, Zmin. rewrite spec_compare; destruct Zcompare; reflexivity.
+ intros. unfold min, Z.min. rewrite spec_compare; destruct Z.compare; reflexivity.
Qed.
(** * Multiplication *)
@@ -437,7 +435,7 @@ Module Make (W0:CyclicType) <: NType.
intros; unfold wn_mul.
generalize (spec_mul_add_n1 n m x y ZnZ.zero).
case DoubleMul.double_mul_add_n1; intros q r Hqr.
- rewrite ZnZ.spec_0, Zplus_0_r in Hqr. rewrite <- Hqr.
+ rewrite ZnZ.spec_0, Z.add_0_r in Hqr. rewrite <- Hqr.
generalize (ZnZ.spec_eq0 q); case ZnZ.eq0; intros HH.
rewrite HH; auto. simpl. apply spec_mk_t_w'.
clear.
@@ -458,7 +456,7 @@ Module Make (W0:CyclicType) <: NType.
intros n m x y; unfold mulnm. rewrite spec_reduce_n.
rewrite (spec_cast_l n m x), (spec_cast_r n m y).
apply spec_muln.
- intros. rewrite Zmult_comm; auto.
+ intros. rewrite Z.mul_comm; auto.
Qed.
(** * Division by a smaller number *)
@@ -519,7 +517,7 @@ Module Make (W0:CyclicType) <: NType.
apply DoubleBase.spec_get_low.
apply spec_zeron.
exact ZnZ.spec_to_Z.
- apply Zle_lt_trans with (ZnZ.to_Z y); auto.
+ apply Z.le_lt_trans with (ZnZ.to_Z y); auto.
rewrite <- nmake_double; auto.
case (ZnZ.spec_to_Z y); auto.
Qed.
@@ -580,9 +578,9 @@ Module Make (W0:CyclicType) <: NType.
intros x y H1 H2; generalize (spec_div_gt_aux x y H1 H2); case div_gt.
intros q r (H3, H4); split.
apply (Zdiv_unique [x] [y] [q] [r]); auto.
- rewrite Zmult_comm; auto.
+ rewrite Z.mul_comm; auto.
apply (Zmod_unique [x] [y] [q] [r]); auto.
- rewrite Zmult_comm; auto.
+ rewrite Z.mul_comm; auto.
Qed.
(** * General Division *)
@@ -597,7 +595,7 @@ Module Make (W0:CyclicType) <: NType.
Theorem spec_div_eucl: forall x y,
let (q,r) := div_eucl x y in
- ([q], [r]) = Zdiv_eucl [x] [y].
+ ([q], [r]) = Z.div_eucl [x] [y].
Proof.
intros x y. unfold div_eucl.
rewrite spec_eqb, spec_compare, spec_0.
@@ -606,16 +604,16 @@ Module Make (W0:CyclicType) <: NType.
intros H'.
assert (H : 0 < [y]) by (generalize (spec_pos y); auto with zarith).
clear H'.
- case Zcompare_spec; intros Cmp;
+ case Z.compare_spec; intros Cmp;
rewrite ?spec_0, ?spec_1; intros; auto with zarith.
- rewrite Cmp; generalize (Z_div_same [y] (Zlt_gt _ _ H))
- (Z_mod_same [y] (Zlt_gt _ _ H));
- unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.
+ rewrite Cmp; generalize (Z_div_same [y] (Z.lt_gt _ _ H))
+ (Z_mod_same [y] (Z.lt_gt _ _ H));
+ unfold Z.div, Z.modulo; case Z.div_eucl; intros; subst; auto.
assert (LeLt: 0 <= [x] < [y]) by (generalize (spec_pos x); auto).
generalize (Zdiv_small _ _ LeLt) (Zmod_small _ _ LeLt);
- unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.
- generalize (spec_div_gt _ _ (Zlt_gt _ _ Cmp) H); auto.
- unfold Zdiv, Zmod; case Zdiv_eucl; case div_gt.
+ unfold Z.div, Z.modulo; case Z.div_eucl; intros; subst; auto.
+ generalize (spec_div_gt _ _ (Z.lt_gt _ _ Cmp) H); auto.
+ unfold Z.div, Z.modulo; case Z.div_eucl; case div_gt.
intros a b c d (H1, H2); subst; auto.
Qed.
@@ -626,7 +624,7 @@ Module Make (W0:CyclicType) <: NType.
Proof.
intros x y; unfold div; generalize (spec_div_eucl x y);
case div_eucl; simpl fst.
- intros xx yy; unfold Zdiv; case Zdiv_eucl; intros qq rr H;
+ intros xx yy; unfold Z.div; case Z.div_eucl; intros qq rr H;
injection H; auto.
Qed.
@@ -730,10 +728,10 @@ Module Make (W0:CyclicType) <: NType.
intro H'.
assert (H : 0 < [y]) by (generalize (spec_pos y); auto with zarith).
clear H'.
- case Zcompare_spec;
+ case Z.compare_spec;
rewrite ?spec_0, ?spec_1; intros; try split; auto with zarith.
- rewrite H0; apply sym_equal; apply Z_mod_same; auto with zarith.
- apply sym_equal; apply Zmod_small; auto with zarith.
+ rewrite H0; symmetry; apply Z_mod_same; auto with zarith.
+ symmetry; apply Zmod_small; auto with zarith.
generalize (spec_pos x); auto with zarith.
apply spec_mod_gt; auto with zarith.
Qed.
@@ -775,7 +773,7 @@ Module Make (W0:CyclicType) <: NType.
Proof.
intros x.
symmetry. apply Z.sqrt_unique.
- rewrite <- ! Zpower_2. apply spec_sqrt_aux.
+ rewrite <- ! Z.pow_2_r. apply spec_sqrt_aux.
Qed.
(** * Power *)
@@ -791,14 +789,14 @@ Module Make (W0:CyclicType) <: NType.
Proof.
intros x n; generalize x; elim n; clear n x; simpl pow_pos.
intros; rewrite spec_mul; rewrite spec_square; rewrite H.
- rewrite Zpos_xI; rewrite Zpower_exp; auto with zarith.
- rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.
- rewrite Zpower_2; rewrite Zpower_1_r; auto.
+ rewrite Pos2Z.inj_xI; rewrite Zpower_exp; auto with zarith.
+ rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r; auto with zarith.
+ rewrite Z.pow_2_r; rewrite Z.pow_1_r; auto.
intros; rewrite spec_square; rewrite H.
- rewrite Zpos_xO; auto with zarith.
- rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.
- rewrite Zpower_2; auto.
- intros; rewrite Zpower_1_r; auto.
+ rewrite Pos2Z.inj_xO; auto with zarith.
+ rewrite (Z.mul_comm 2); rewrite Z.pow_mul_r; auto with zarith.
+ rewrite Z.pow_2_r; auto.
+ intros; rewrite Z.pow_1_r; auto.
Qed.
Definition pow_N (x:t)(n:N) : t := match n with
@@ -806,7 +804,7 @@ Module Make (W0:CyclicType) <: NType.
| BinNat.Npos p => pow_pos x p
end.
- Theorem spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z_of_N n.
+ Theorem spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z.of_N n.
Proof.
destruct n; simpl. apply spec_1.
apply spec_pow_pos.
@@ -867,15 +865,15 @@ Module Make (W0:CyclicType) <: NType.
Zis_gcd [a] [b] [gcd_gt_body a b cont].
Proof.
intros a b cont p H2 H3 H4; unfold gcd_gt_body.
- rewrite ! spec_compare, spec_0. case Zcompare_spec.
+ rewrite ! spec_compare, spec_0. case Z.compare_spec.
intros ->; apply Zis_gcd_0.
intros HH; absurd (0 <= [b]); auto with zarith.
case (spec_digits b); auto with zarith.
- intros H5; case Zcompare_spec.
- intros H6; rewrite <- (Zmult_1_r [b]).
+ intros H5; case Z.compare_spec.
+ intros H6; rewrite <- (Z.mul_1_r [b]).
rewrite (Z_div_mod_eq [a] [b]); auto with zarith.
rewrite <- spec_mod_gt; auto with zarith.
- rewrite H6; rewrite Zplus_0_r.
+ rewrite H6; rewrite Z.add_0_r.
apply Zis_gcd_mult; apply Zis_gcd_1.
intros; apply False_ind.
case (spec_digits (mod_gt a b)); auto with zarith.
@@ -890,24 +888,19 @@ Module Make (W0:CyclicType) <: NType.
rewrite <- spec_mod_gt; auto with zarith.
repeat rewrite <- spec_mod_gt; auto with zarith.
apply H4; auto with zarith.
- apply Zmult_lt_reg_r with 2; auto with zarith.
- apply Zle_lt_trans with ([b] + [mod_gt a b]); auto with zarith.
- apply Zle_lt_trans with (([a]/[b]) * [b] + [mod_gt a b]); auto with zarith.
- apply Zplus_le_compat_r.
- pattern [b] at 1; rewrite <- (Zmult_1_l [b]).
- apply Zmult_le_compat_r; auto with zarith.
- case (Zle_lt_or_eq 0 ([a]/[b])); auto with zarith.
- intros HH; rewrite (Z_div_mod_eq [a] [b]) in H2;
- try rewrite <- HH in H2; auto with zarith.
- case (Z_mod_lt [a] [b]); auto with zarith.
- rewrite Zmult_comm; rewrite spec_mod_gt; auto with zarith.
- rewrite <- Z_div_mod_eq; auto with zarith.
- pattern 2 at 2; rewrite <- (Zpower_1_r 2).
- rewrite <- Zpower_exp; auto with zarith.
- ring_simplify (p - 1 + 1); auto.
- case (Zle_lt_or_eq 0 p); auto with zarith.
- generalize H3; case p; simpl Zpower; auto with zarith.
- intros HH; generalize H3; rewrite <- HH; simpl Zpower; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with 2; auto with zarith.
+ apply Z.le_lt_trans with ([b] + [mod_gt a b]); auto with zarith.
+ apply Z.le_lt_trans with (([a]/[b]) * [b] + [mod_gt a b]); auto with zarith.
+ - apply Z.add_le_mono_r.
+ rewrite <- (Z.mul_1_l [b]) at 1.
+ apply Z.mul_le_mono_nonneg_r; auto with zarith.
+ change 1 with (Z.succ 0). apply Z.le_succ_l.
+ apply Z.div_str_pos; auto with zarith.
+ - rewrite Z.mul_comm; rewrite spec_mod_gt; auto with zarith.
+ rewrite <- Z_div_mod_eq; auto with zarith.
+ rewrite Z.mul_comm, <- Z.pow_succ_r, Z.sub_1_r, Z.succ_pred; auto.
+ apply Z.le_0_sub. change 1 with (Z.succ 0). apply Z.le_succ_l.
+ destruct p; simpl in H3; auto with zarith.
Qed.
Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) : t :=
@@ -931,7 +924,7 @@ Module Make (W0:CyclicType) <: NType.
apply Hrec with (Zpos p + n); auto.
replace (Zpos p + (Zpos p + n)) with
(Zpos (xI p) + n - 1); auto.
- rewrite Zpos_xI; ring.
+ rewrite Pos2Z.inj_xI; ring.
intros a2 b2 H9 H10.
apply Hrec with n; auto.
intros p Hrec n a b cont H2 H3 H4.
@@ -940,18 +933,18 @@ Module Make (W0:CyclicType) <: NType.
apply Hrec with (Zpos p + n - 1); auto.
replace (Zpos p + (Zpos p + n - 1)) with
(Zpos (xO p) + n - 1); auto.
- rewrite Zpos_xO; ring.
+ rewrite Pos2Z.inj_xO; ring.
intros a2 b2 H9 H10.
apply Hrec with (n - 1); auto.
replace (Zpos p + (n - 1)) with
(Zpos p + n - 1); auto with zarith.
intros a3 b3 H12 H13; apply H4; auto with zarith.
- apply Zlt_le_trans with (1 := H12).
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.lt_le_trans with (1 := H12).
+ apply Z.pow_le_mono_r; auto with zarith.
intros n a b cont H H2 H3.
simpl gcd_gt_aux.
apply Zspec_gcd_gt_body with (n + 1); auto with zarith.
- rewrite Zplus_comm; auto.
+ rewrite Z.add_comm; auto.
intros a1 b1 H5 H6; apply H3; auto.
replace n with (n + 1 - 1); auto; try ring.
Qed.
@@ -965,14 +958,14 @@ Module Make (W0:CyclicType) <: NType.
Definition gcd_gt a b := gcd_gt_aux (digits a) gcd_cont a b.
Theorem spec_gcd_gt: forall a b,
- [a] > [b] -> [gcd_gt a b] = Zgcd [a] [b].
+ [a] > [b] -> [gcd_gt a b] = Z.gcd [a] [b].
Proof.
intros a b H2.
case (spec_digits (gcd_gt a b)); intros H3 H4.
case (spec_digits a); intros H5 H6.
- apply sym_equal; apply Zis_gcd_gcd; auto with zarith.
+ symmetry; apply Zis_gcd_gcd; auto with zarith.
unfold gcd_gt; apply Zspec_gcd_gt_aux with 0; auto with zarith.
- intros a1 a2; rewrite Zpower_0_r.
+ intros a1 a2; rewrite Z.pow_0_r.
case (spec_digits a2); intros H7 H8;
intros; apply False_ind; auto with zarith.
Qed.
@@ -984,18 +977,18 @@ Module Make (W0:CyclicType) <: NType.
| Gt => gcd_gt a b
end.
- Theorem spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b].
+ Theorem spec_gcd: forall a b, [gcd a b] = Z.gcd [a] [b].
Proof.
intros a b.
case (spec_digits a); intros H1 H2.
case (spec_digits b); intros H3 H4.
- unfold gcd. rewrite spec_compare. case Zcompare_spec.
- intros HH; rewrite HH; apply sym_equal; apply Zis_gcd_gcd; auto.
+ unfold gcd. rewrite spec_compare. case Z.compare_spec.
+ intros HH; rewrite HH; symmetry; apply Zis_gcd_gcd; auto.
apply Zis_gcd_refl.
- intros; apply trans_equal with (Zgcd [b] [a]).
+ intros; transitivity (Z.gcd [b] [a]).
apply spec_gcd_gt; auto with zarith.
apply Zis_gcd_gcd; auto with zarith.
- apply Zgcd_is_pos.
+ apply Z.gcd_nonneg.
apply Zis_gcd_sym; apply Zgcd_is_gcd.
intros; apply spec_gcd_gt; auto with zarith.
Qed.
@@ -1017,22 +1010,22 @@ Module Make (W0:CyclicType) <: NType.
exact (ZnZ.spec_is_even x).
Qed.
- Theorem spec_even: forall x, even x = Zeven_bool [x].
+ Theorem spec_even: forall x, even x = Z.even [x].
Proof.
intros x. assert (H := spec_even_aux x). symmetry.
- rewrite (Z_div_mod_eq_full [x] 2); auto with zarith.
- destruct (even x); rewrite H, ?Zplus_0_r.
+ rewrite (Z.div_mod [x] 2); auto with zarith.
+ destruct (even x); rewrite H, ?Z.add_0_r.
rewrite Zeven_bool_iff. apply Zeven_2p.
apply not_true_is_false. rewrite Zeven_bool_iff.
apply Zodd_not_Zeven. apply Zodd_2p_plus_1.
Qed.
- Theorem spec_odd: forall x, odd x = Zodd_bool [x].
+ Theorem spec_odd: forall x, odd x = Z.odd [x].
Proof.
intros x. unfold odd.
assert (H := spec_even_aux x). symmetry.
- rewrite (Z_div_mod_eq_full [x] 2); auto with zarith.
- destruct (even x); rewrite H, ?Zplus_0_r; simpl negb.
+ rewrite (Z.div_mod [x] 2); auto with zarith.
+ destruct (even x); rewrite H, ?Z.add_0_r; simpl negb.
apply not_true_is_false. rewrite Zodd_bool_iff.
apply Zeven_not_Zodd. apply Zeven_2p.
apply Zodd_bool_iff. apply Zodd_2p_plus_1.
@@ -1041,27 +1034,21 @@ Module Make (W0:CyclicType) <: NType.
(** * Conversion *)
Definition pheight p :=
- Peano.pred (nat_of_P (get_height (ZnZ.digits (dom_op 0)) (plength p))).
+ Peano.pred (Pos.to_nat (get_height (ZnZ.digits (dom_op 0)) (plength p))).
Theorem pheight_correct: forall p,
- Zpos p < 2 ^ (Zpos (ZnZ.digits (dom_op 0)) * 2 ^ (Z_of_nat (pheight p))).
+ Zpos p < 2 ^ (Zpos (ZnZ.digits (dom_op 0)) * 2 ^ (Z.of_nat (pheight p))).
Proof.
intros p; unfold pheight.
- assert (F1: forall x, Z_of_nat (Peano.pred (nat_of_P x)) = Zpos x - 1).
- intros x.
- assert (Zsucc (Z_of_nat (Peano.pred (nat_of_P x))) = Zpos x); auto with zarith.
- rewrite <- inj_S.
- rewrite <- (fun x => S_pred x 0); auto with zarith.
- rewrite Zpos_eq_Z_of_nat_o_nat_of_P; auto.
- apply lt_le_trans with 1%nat; auto with zarith.
- exact (le_Pmult_nat x 1).
- rewrite F1; clear F1.
+ rewrite Nat2Z.inj_pred by apply Pos2Nat.is_pos.
+ rewrite positive_nat_Z.
+ rewrite <- Z.sub_1_r.
assert (F2:= (get_height_correct (ZnZ.digits (dom_op 0)) (plength p))).
- apply Zlt_le_trans with (Zpos (Psucc p)).
- rewrite Zpos_succ_morphism; auto with zarith.
- apply Zle_trans with (1 := plength_pred_correct (Psucc p)).
- rewrite Ppred_succ.
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.lt_le_trans with (Zpos (Pos.succ p)).
+ rewrite Pos2Z.inj_succ; auto with zarith.
+ apply Z.le_trans with (1 := plength_pred_correct (Pos.succ p)).
+ rewrite Pos.pred_succ.
+ apply Z.pow_le_mono_r; auto with zarith.
Qed.
Definition of_pos (x:positive) : t :=
@@ -1076,8 +1063,8 @@ Module Make (W0:CyclicType) <: NType.
simpl.
apply ZnZ.of_pos_correct.
unfold base.
- apply Zlt_le_trans with (1 := pheight_correct x).
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.lt_le_trans with (1 := pheight_correct x).
+ apply Z.pow_le_mono_r; auto with zarith.
rewrite (digits_dom_op (_ _)), Pshiftl_nat_Zpower. auto with zarith.
Qed.
@@ -1088,7 +1075,7 @@ Module Make (W0:CyclicType) <: NType.
end.
Theorem spec_of_N: forall x,
- [of_N x] = Z_of_N x.
+ [of_N x] = Z.of_N x.
Proof.
intros x; case x.
simpl of_N. exact spec_0.
@@ -1122,7 +1109,7 @@ Module Make (W0:CyclicType) <: NType.
intros. apply Zdiv_unique with 0; auto with zarith.
change 2 with (2^1) at 2.
rewrite <- Zpower_exp; auto with zarith.
- rewrite Zplus_0_r. f_equal. auto with zarith.
+ rewrite Z.add_0_r. f_equal. auto with zarith.
Qed.
Theorem spec_head0: forall x, 0 < [x] ->
@@ -1212,9 +1199,9 @@ Module Make (W0:CyclicType) <: NType.
set (d := ZnZ.digits (dom_op n)) in *; clearbody d.
destruct (Z_lt_le_dec h (Zpos d)); auto. exfalso.
assert (1 * 2^Zpos d <= ZnZ.to_Z x * 2^h).
- apply Zmult_le_compat; auto with zarith.
- apply Zpower_le_monotone2; auto with zarith.
- rewrite Zmult_comm in H0. auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
+ apply Z.pow_le_mono_r; auto with zarith.
+ rewrite Z.mul_comm in H0. auto with zarith.
Qed.
Lemma spec_log2_pos : forall x, [x]<>0 ->
@@ -1232,13 +1219,13 @@ Module Make (W0:CyclicType) <: NType.
assert (H2 := ZnZ.spec_to_Z (ZnZ.zdigits (dom_op n))).
assert (H3 := head0_zdigits n x).
rewrite Zmod_small by auto with zarith.
+ rewrite Z.sub_simpl_r.
rewrite (Z.mul_lt_mono_pos_l (2^(ZnZ.to_Z (ZnZ.head0 x))));
auto with zarith.
rewrite (Z.mul_le_mono_pos_l _ _ (2^(ZnZ.to_Z (ZnZ.head0 x))));
auto with zarith.
rewrite <- 2 Zpower_exp; auto with zarith.
- rewrite Z.add_sub_assoc, Zplus_minus.
- rewrite Z.sub_simpl_r, Zplus_minus.
+ rewrite !Z.add_sub_assoc, !Z.add_simpl_l.
rewrite ZnZ.spec_zdigits.
rewrite pow2_pos_minus_1 by (red; auto).
apply ZnZ.spec_head0; auto with zarith.
@@ -1294,12 +1281,12 @@ Module Make (W0:CyclicType) <: NType.
Proof.
intros x y z HH HH1 HH2.
split; auto with zarith.
- apply Zle_lt_trans with (2 := HH2); auto with zarith.
+ apply Z.le_lt_trans with (2 := HH2); auto with zarith.
apply Zdiv_le_upper_bound; auto with zarith.
pattern x at 1; replace x with (x * 2 ^ 0); auto with zarith.
- apply Zmult_le_compat_l; auto.
- apply Zpower_le_monotone2; auto with zarith.
- rewrite Zpower_0_r; ring.
+ apply Z.mul_le_mono_nonneg_l; auto.
+ apply Z.pow_le_mono_r; auto with zarith.
+ rewrite Z.pow_0_r; ring.
Qed.
Theorem spec_shiftr_pow2 : forall x n,
@@ -1315,7 +1302,7 @@ Module Make (W0:CyclicType) <: NType.
rewrite spec_reduce.
rewrite ZnZ.spec_zdigits in H.
rewrite ZnZ.spec_add_mul_div by auto with zarith.
- rewrite ZnZ.spec_0, Zmult_0_l, Zplus_0_l.
+ rewrite ZnZ.spec_0, Z.mul_0_l, Z.add_0_l.
rewrite Zmod_small.
f_equal. f_equal. auto with zarith.
split. auto with zarith.
@@ -1324,8 +1311,8 @@ Module Make (W0:CyclicType) <: NType.
rewrite ZnZ.spec_0. symmetry.
apply Zdiv_small.
split; auto with zarith.
- apply Zlt_le_trans with (base (ZnZ.digits (dom_op n))); auto with zarith.
- unfold base. apply Zpower_le_monotone2; auto with zarith.
+ apply Z.lt_le_trans with (base (ZnZ.digits (dom_op n))); auto with zarith.
+ unfold base. apply Z.pow_le_mono_r; auto with zarith.
rewrite ZnZ.spec_zdigits in H.
generalize (ZnZ.spec_to_Z d); auto with zarith.
Qed.
@@ -1370,21 +1357,21 @@ Module Make (W0:CyclicType) <: NType.
destruct (ZnZ.spec_to_Z x).
destruct (ZnZ.spec_to_Z p).
rewrite ZnZ.spec_add_mul_div by (omega with *).
- rewrite ZnZ.spec_0, Zdiv_0_l, Zplus_0_r.
+ rewrite ZnZ.spec_0, Zdiv_0_l, Z.add_0_r.
apply Zmod_small. unfold base.
split; auto with zarith.
- rewrite Zmult_comm.
- apply Zlt_le_trans with (2^(ZnZ.to_Z p + K)).
+ rewrite Z.mul_comm.
+ apply Z.lt_le_trans with (2^(ZnZ.to_Z p + K)).
rewrite Zpower_exp; auto with zarith.
- apply Zmult_lt_compat_l; auto with zarith.
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.mul_lt_mono_pos_l; auto with zarith.
+ apply Z.pow_le_mono_r; auto with zarith.
Qed.
Theorem spec_unsafe_shiftl: forall x p,
[p] <= [head0 x] -> [unsafe_shiftl x p] = [x] * 2 ^ [p].
Proof.
intros.
- destruct (Z_eq_dec [x] 0) as [EQ|NEQ].
+ destruct (Z.eq_dec [x] 0) as [EQ|NEQ].
(* [x] = 0 *)
apply spec_unsafe_shiftl_aux with 0; auto with zarith.
now rewrite EQ.
@@ -1421,7 +1408,7 @@ Module Make (W0:CyclicType) <: NType.
Proof.
intros x. rewrite ! digits_level, double_size_level.
rewrite 2 digits_dom_op, 2 Pshiftl_nat_Zpower,
- inj_S, Zpower_Zsucc; auto with zarith.
+ Nat2Z.inj_succ, Z.pow_succ_r; auto with zarith.
ring.
Qed.
@@ -1438,46 +1425,47 @@ Module Make (W0:CyclicType) <: NType.
assert (F1:= spec_pos (head0 x)).
assert (F2: 0 < Zpos (digits x)).
red; auto.
- case (Zle_lt_or_eq _ _ (spec_pos x)); intros HH.
+ assert (HH := spec_pos x). Z.le_elim HH.
generalize HH; rewrite <- (spec_double_size x); intros HH1.
case (spec_head0 x HH); intros _ HH2.
case (spec_head0 _ HH1).
rewrite (spec_double_size x); rewrite (spec_double_size_digits x).
intros HH3 _.
- case (Zle_or_lt ([head0 (double_size x)]) (2 * [head0 x])); auto; intros HH4.
+ case (Z.le_gt_cases ([head0 (double_size x)]) (2 * [head0 x])); auto; intros HH4.
absurd (2 ^ (2 * [head0 x] )* [x] < 2 ^ [head0 (double_size x)] * [x]); auto.
- apply Zle_not_lt.
- apply Zmult_le_compat_r; auto with zarith.
- apply Zpower_le_monotone2; auto; auto with zarith.
+ apply Z.le_ngt.
+ apply Z.mul_le_mono_nonneg_r; auto with zarith.
+ apply Z.pow_le_mono_r; auto; auto with zarith.
assert (HH5: 2 ^[head0 x] <= 2 ^(Zpos (digits x) - 1)).
- case (Zle_lt_or_eq 1 [x]); auto with zarith; intros HH5.
- apply Zmult_le_reg_r with (2 ^ 1); auto with zarith.
- rewrite <- (fun x y z => Zpower_exp x (y - z)); auto with zarith.
- assert (tmp: forall x, x - 1 + 1 = x); [intros; ring | rewrite tmp; clear tmp].
- apply Zle_trans with (2 := Zlt_le_weak _ _ HH2).
- apply Zmult_le_compat_l; auto with zarith.
- rewrite Zpower_1_r; auto with zarith.
- apply Zpower_le_monotone2; auto with zarith.
- case (Zle_or_lt (Zpos (digits x)) [head0 x]); auto with zarith; intros HH6.
- absurd (2 ^ Zpos (digits x) <= 2 ^ [head0 x] * [x]); auto with zarith.
- rewrite <- HH5; rewrite Zmult_1_r.
- apply Zpower_le_monotone2; auto with zarith.
- rewrite (Zmult_comm 2).
- rewrite Zpower_mult; auto with zarith.
- rewrite Zpower_2.
- apply Zlt_le_trans with (2 := HH3).
- rewrite <- Zmult_assoc.
+ { apply Z.le_succ_l in HH. change (1 <= [x]) in HH.
+ Z.le_elim HH.
+ - apply Z.mul_le_mono_pos_r with (2 ^ 1); auto with zarith.
+ rewrite <- (fun x y z => Z.pow_add_r x (y - z)); auto with zarith.
+ rewrite Z.sub_add.
+ apply Z.le_trans with (2 := Z.lt_le_incl _ _ HH2).
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
+ rewrite Z.pow_1_r; auto with zarith.
+ - apply Z.pow_le_mono_r; auto with zarith.
+ case (Z.le_gt_cases (Zpos (digits x)) [head0 x]); auto with zarith; intros HH6.
+ absurd (2 ^ Zpos (digits x) <= 2 ^ [head0 x] * [x]); auto with zarith.
+ rewrite <- HH; rewrite Z.mul_1_r.
+ apply Z.pow_le_mono_r; auto with zarith. }
+ rewrite (Z.mul_comm 2).
+ rewrite Z.pow_mul_r; auto with zarith.
+ rewrite Z.pow_2_r.
+ apply Z.lt_le_trans with (2 := HH3).
+ rewrite <- Z.mul_assoc.
replace (2 * Zpos (digits x) - 1) with
((Zpos (digits x) - 1) + (Zpos (digits x))).
rewrite Zpower_exp; auto with zarith.
apply Zmult_lt_compat2; auto with zarith.
split; auto with zarith.
- apply Zmult_lt_0_compat; auto with zarith.
- rewrite Zpos_xO; ring.
- apply Zlt_le_weak; auto.
+ apply Z.mul_pos_pos; auto with zarith.
+ rewrite Pos2Z.inj_xO; ring.
+ apply Z.lt_le_incl; auto.
repeat rewrite spec_head00; auto.
rewrite spec_double_size_digits.
- rewrite Zpos_xO; auto with zarith.
+ rewrite Pos2Z.inj_xO; auto with zarith.
rewrite spec_double_size; auto.
Qed.
@@ -1485,24 +1473,26 @@ Module Make (W0:CyclicType) <: NType.
forall x, 0 < [head0 (double_size x)].
Proof.
intros x.
- assert (F: 0 < Zpos (digits x)).
- red; auto.
- case (Zle_lt_or_eq _ _ (spec_pos (head0 (double_size x)))); auto; intros F0.
- case (Zle_lt_or_eq _ _ (spec_pos (head0 x))); intros F1.
- apply Zlt_le_trans with (2 := (spec_double_size_head0 x)); auto with zarith.
- case (Zle_lt_or_eq _ _ (spec_pos x)); intros F3.
+ assert (F := Pos2Z.is_pos (digits x)).
+ assert (F0 := spec_pos (head0 (double_size x))).
+ Z.le_elim F0; auto.
+ assert (F1 := spec_pos (head0 x)).
+ Z.le_elim F1.
+ apply Z.lt_le_trans with (2 := (spec_double_size_head0 x)); auto with zarith.
+ assert (F3 := spec_pos x).
+ Z.le_elim F3.
generalize F3; rewrite <- (spec_double_size x); intros F4.
absurd (2 ^ (Zpos (xO (digits x)) - 1) < 2 ^ (Zpos (digits x))).
- apply Zle_not_lt.
- apply Zpower_le_monotone2; auto with zarith.
- rewrite Zpos_xO; auto with zarith.
+ { apply Z.le_ngt.
+ apply Z.pow_le_mono_r; auto with zarith.
+ rewrite Pos2Z.inj_xO; auto with zarith. }
case (spec_head0 x F3).
- rewrite <- F1; rewrite Zpower_0_r; rewrite Zmult_1_l; intros _ HH.
- apply Zle_lt_trans with (2 := HH).
+ rewrite <- F1; rewrite Z.pow_0_r; rewrite Z.mul_1_l; intros _ HH.
+ apply Z.le_lt_trans with (2 := HH).
case (spec_head0 _ F4).
rewrite (spec_double_size x); rewrite (spec_double_size_digits x).
- rewrite <- F0; rewrite Zpower_0_r; rewrite Zmult_1_l; auto.
- generalize F1; rewrite (spec_head00 _ (sym_equal F3)); auto with zarith.
+ rewrite <- F0; rewrite Z.pow_0_r; rewrite Z.mul_1_l; auto.
+ generalize F1; rewrite (spec_head00 _ (eq_sym F3)); auto with zarith.
Qed.
(** Finally we iterate [double_size] enough before [unsafe_shiftl]
@@ -1521,14 +1511,14 @@ Module Make (W0:CyclicType) <: NType.
[shiftl_aux_body cont x n] = [x] * 2 ^ [n].
Proof.
intros n x p cont H1 H2; unfold shiftl_aux_body.
- rewrite spec_compare; case Zcompare_spec; intros H.
+ rewrite spec_compare; case Z.compare_spec; intros H.
apply spec_unsafe_shiftl; auto with zarith.
apply spec_unsafe_shiftl; auto with zarith.
rewrite H2.
rewrite spec_double_size; auto.
- rewrite Zplus_comm; rewrite Zpower_exp; auto with zarith.
- apply Zle_trans with (2 := spec_double_size_head0 x).
- rewrite Zpower_1_r; apply Zmult_le_compat_l; auto with zarith.
+ rewrite Z.add_comm; rewrite Zpower_exp; auto with zarith.
+ apply Z.le_trans with (2 := spec_double_size_head0 x).
+ rewrite Z.pow_1_r; apply Z.mul_le_mono_nonneg_l; auto with zarith.
Qed.
Fixpoint shiftl_aux p cont x n :=
@@ -1550,27 +1540,27 @@ Module Make (W0:CyclicType) <: NType.
apply spec_shiftl_aux_body with (q); auto.
intros x1 H3; apply Hrec with (q + 1)%positive; auto.
intros x2 H4; apply Hrec with (p + q + 1)%positive; auto.
- rewrite <- Pplus_assoc.
- rewrite Zpos_plus_distr; auto.
+ rewrite <- Pos.add_assoc.
+ rewrite Pos2Z.inj_add; auto.
intros x3 H5; apply H2.
- rewrite Zpos_xI.
+ rewrite Pos2Z.inj_xI.
replace (2 * Zpos p + 1 + Zpos q) with (Zpos p + Zpos (p + q + 1));
auto.
- repeat rewrite Zpos_plus_distr; ring.
+ rewrite !Pos2Z.inj_add; ring.
intros p Hrec q n x cont H1 H2.
apply spec_shiftl_aux_body with (q); auto.
intros x1 H3; apply Hrec with (q); auto.
- apply Zle_trans with (2 := H3); auto with zarith.
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.le_trans with (2 := H3); auto with zarith.
+ apply Z.pow_le_mono_r; auto with zarith.
intros x2 H4; apply Hrec with (p + q)%positive; auto.
intros x3 H5; apply H2.
- rewrite (Zpos_xO p).
+ rewrite (Pos2Z.inj_xO p).
replace (2 * Zpos p + Zpos q) with (Zpos p + Zpos (p + q));
auto.
- repeat rewrite Zpos_plus_distr; ring.
+ rewrite Pos2Z.inj_add; ring.
intros q n x cont H1 H2.
apply spec_shiftl_aux_body with (q); auto.
- rewrite Zplus_comm; auto.
+ rewrite Z.add_comm; auto.
Qed.
Definition shiftl x n :=
@@ -1582,25 +1572,25 @@ Module Make (W0:CyclicType) <: NType.
[shiftl x n] = [x] * 2 ^ [n].
Proof.
intros x n; unfold shiftl, shiftl_aux_body.
- rewrite spec_compare; case Zcompare_spec; intros H.
+ rewrite spec_compare; case Z.compare_spec; intros H.
apply spec_unsafe_shiftl; auto with zarith.
apply spec_unsafe_shiftl; auto with zarith.
rewrite <- (spec_double_size x).
- rewrite spec_compare; case Zcompare_spec; intros H1.
+ rewrite spec_compare; case Z.compare_spec; intros H1.
apply spec_unsafe_shiftl; auto with zarith.
apply spec_unsafe_shiftl; auto with zarith.
rewrite <- (spec_double_size (double_size x)).
apply spec_shiftl_aux with 1%positive.
- apply Zle_trans with (2 := spec_double_size_head0 (double_size x)).
+ apply Z.le_trans with (2 := spec_double_size_head0 (double_size x)).
replace (2 ^ 1) with (2 * 1).
- apply Zmult_le_compat_l; auto with zarith.
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
generalize (spec_double_size_head0_pos x); auto with zarith.
- rewrite Zpower_1_r; ring.
+ rewrite Z.pow_1_r; ring.
intros x1 H2; apply spec_unsafe_shiftl.
- apply Zle_trans with (2 := H2).
- apply Zle_trans with (2 ^ Zpos (digits n)); auto with zarith.
+ apply Z.le_trans with (2 := H2).
+ apply Z.le_trans with (2 ^ Zpos (digits n)); auto with zarith.
case (spec_digits n); auto with zarith.
- apply Zpower_le_monotone2; auto with zarith.
+ apply Z.pow_le_mono_r; auto with zarith.
Qed.
Lemma spec_shiftl: forall x p, [shiftl x p] = Z.shiftl [x] [p].
diff --git a/theories/Numbers/Natural/BigN/Nbasic.v b/theories/Numbers/Natural/BigN/Nbasic.v
index 4717d0b23..756c7f9a2 100644
--- a/theories/Numbers/Natural/BigN/Nbasic.v
+++ b/theories/Numbers/Natural/BigN/Nbasic.v
@@ -32,7 +32,7 @@ Proof.
transitivity (2 * (2 ^ Z.of_nat n * Zpos p)).
rewrite <- IHn. auto.
rewrite Z.mul_assoc.
- rewrite inj_S.
+ rewrite Nat2Z.inj_succ.
rewrite <- Z.pow_succ_r; auto with zarith.
Qed.
@@ -41,39 +41,39 @@ Qed.
Fixpoint plength (p: positive) : positive :=
match p with
xH => xH
- | xO p1 => Psucc (plength p1)
- | xI p1 => Psucc (plength p1)
+ | xO p1 => Pos.succ (plength p1)
+ | xI p1 => Pos.succ (plength p1)
end.
Theorem plength_correct: forall p, (Zpos p < 2 ^ Zpos (plength p))%Z.
-assert (F: (forall p, 2 ^ (Zpos (Psucc p)) = 2 * 2 ^ Zpos p)%Z).
-intros p; replace (Zpos (Psucc p)) with (1 + Zpos p)%Z.
+assert (F: (forall p, 2 ^ (Zpos (Pos.succ p)) = 2 * 2 ^ Zpos p)%Z).
+intros p; replace (Zpos (Pos.succ p)) with (1 + Zpos p)%Z.
rewrite Zpower_exp; auto with zarith.
-rewrite Zpos_succ_morphism; unfold Zsucc; auto with zarith.
+rewrite Pos2Z.inj_succ; unfold Z.succ; auto with zarith.
intros p; elim p; simpl plength; auto.
-intros p1 Hp1; rewrite F; repeat rewrite Zpos_xI.
+intros p1 Hp1; rewrite F; repeat rewrite Pos2Z.inj_xI.
assert (tmp: (forall p, 2 * p = p + p)%Z);
try repeat rewrite tmp; auto with zarith.
-intros p1 Hp1; rewrite F; rewrite (Zpos_xO p1).
+intros p1 Hp1; rewrite F; rewrite (Pos2Z.inj_xO p1).
assert (tmp: (forall p, 2 * p = p + p)%Z);
try repeat rewrite tmp; auto with zarith.
-rewrite Zpower_1_r; auto with zarith.
+rewrite Z.pow_1_r; auto with zarith.
Qed.
-Theorem plength_pred_correct: forall p, (Zpos p <= 2 ^ Zpos (plength (Ppred p)))%Z.
-intros p; case (Psucc_pred p); intros H1.
+Theorem plength_pred_correct: forall p, (Zpos p <= 2 ^ Zpos (plength (Pos.pred p)))%Z.
+intros p; case (Pos.succ_pred_or p); intros H1.
subst; simpl plength.
-rewrite Zpower_1_r; auto with zarith.
+rewrite Z.pow_1_r; auto with zarith.
pattern p at 1; rewrite <- H1.
-rewrite Zpos_succ_morphism; unfold Zsucc; auto with zarith.
-generalize (plength_correct (Ppred p)); auto with zarith.
+rewrite Pos2Z.inj_succ; unfold Z.succ; auto with zarith.
+generalize (plength_correct (Pos.pred p)); auto with zarith.
Qed.
Definition Pdiv p q :=
- match Zdiv (Zpos p) (Zpos q) with
+ match Z.div (Zpos p) (Zpos q) with
Zpos q1 => match (Zpos p) - (Zpos q) * (Zpos q1) with
Z0 => q1
- | _ => (Psucc q1)
+ | _ => (Pos.succ q1)
end
| _ => xH
end.
@@ -85,20 +85,20 @@ unfold Pdiv.
assert (H1: Zpos q > 0); auto with zarith.
assert (H1b: Zpos p >= 0); auto with zarith.
generalize (Z_div_ge0 (Zpos p) (Zpos q) H1 H1b).
-generalize (Z_div_mod_eq (Zpos p) (Zpos q) H1); case Zdiv.
- intros HH _; rewrite HH; rewrite Zmult_0_r; rewrite Zmult_1_r; simpl.
+generalize (Z_div_mod_eq (Zpos p) (Zpos q) H1); case Z.div.
+ intros HH _; rewrite HH; rewrite Z.mul_0_r; rewrite Z.mul_1_r; simpl.
case (Z_mod_lt (Zpos p) (Zpos q) H1); auto with zarith.
intros q1 H2.
replace (Zpos p - Zpos q * Zpos q1) with (Zpos p mod Zpos q).
2: pattern (Zpos p) at 2; rewrite H2; auto with zarith.
generalize H2 (Z_mod_lt (Zpos p) (Zpos q) H1); clear H2;
- case Zmod.
+ case Z.modulo.
intros HH _; rewrite HH; auto with zarith.
- intros r1 HH (_,HH1); rewrite HH; rewrite Zpos_succ_morphism.
- unfold Zsucc; rewrite Zmult_plus_distr_r; auto with zarith.
+ intros r1 HH (_,HH1); rewrite HH; rewrite Pos2Z.inj_succ.
+ unfold Z.succ; rewrite Z.mul_add_distr_l; auto with zarith.
intros r1 _ (HH,_); case HH; auto.
intros q1 HH; rewrite HH.
-unfold Zge; simpl Zcompare; intros HH1; case HH1; auto.
+unfold Z.ge; simpl Z.compare; intros HH1; case HH1; auto.
Qed.
Definition is_one p := match p with xH => true | _ => false end.
@@ -109,7 +109,7 @@ Qed.
Definition get_height digits p :=
let r := Pdiv p digits in
- if is_one r then xH else Psucc (plength (Ppred r)).
+ if is_one r then xH else Pos.succ (plength (Pos.pred r)).
Theorem get_height_correct:
forall digits N,
@@ -119,13 +119,13 @@ unfold get_height.
assert (H1 := Pdiv_le N digits).
case_eq (is_one (Pdiv N digits)); intros H2.
rewrite (is_one_one _ H2) in H1.
-rewrite Zmult_1_r in H1.
-change (2^(1-1))%Z with 1; rewrite Zmult_1_r; auto.
+rewrite Z.mul_1_r in H1.
+change (2^(1-1))%Z with 1; rewrite Z.mul_1_r; auto.
clear H2.
-apply Zle_trans with (1 := H1).
-apply Zmult_le_compat_l; auto with zarith.
-rewrite Zpos_succ_morphism; unfold Zsucc.
-rewrite Zplus_comm; rewrite Zminus_plus.
+apply Z.le_trans with (1 := H1).
+apply Z.mul_le_mono_nonneg_l; auto with zarith.
+rewrite Pos2Z.inj_succ; unfold Z.succ.
+rewrite Z.add_comm; rewrite Z.add_simpl_l.
apply plength_pred_correct.
Qed.
@@ -152,18 +152,18 @@ Open Scope nat_scope.
Fixpoint plusnS (n m: nat) {struct n} : (n + S m = S (n + m))%nat :=
match n return (n + S m = S (n + m))%nat with
- | 0 => refl_equal (S m)
+ | 0 => eq_refl (S m)
| S n1 =>
let v := S (S n1 + m) in
- eq_ind_r (fun n => S n = v) (refl_equal v) (plusnS n1 m)
+ eq_ind_r (fun n => S n = v) (eq_refl v) (plusnS n1 m)
end.
Fixpoint plusn0 n : n + 0 = n :=
match n return (n + 0 = n) with
- | 0 => refl_equal 0
+ | 0 => eq_refl 0
| S n1 =>
let v := S n1 in
- eq_ind_r (fun n : nat => S n = v) (refl_equal v) (plusn0 n1)
+ eq_ind_r (fun n : nat => S n = v) (eq_refl v) (plusn0 n1)
end.
Fixpoint diff (m n: nat) {struct m}: nat * nat :=
@@ -177,8 +177,8 @@ Fixpoint diff_l (m n : nat) {struct m} : fst (diff m n) + n = max m n :=
match m return fst (diff m n) + n = max m n with
| 0 =>
match n return (n = max 0 n) with
- | 0 => refl_equal _
- | S n0 => refl_equal _
+ | 0 => eq_refl _
+ | S n0 => eq_refl _
end
| S m1 =>
match n return (fst (diff (S m1) n) + n = max (S m1) n)
@@ -188,7 +188,7 @@ Fixpoint diff_l (m n : nat) {struct m} : fst (diff m n) + n = max m n :=
let v := fst (diff m1 n1) + n1 in
let v1 := fst (diff m1 n1) + S n1 in
eq_ind v (fun n => v1 = S n)
- (eq_ind v1 (fun n => v1 = n) (refl_equal v1) (S v) (plusnS _ _))
+ (eq_ind v1 (fun n => v1 = n) (eq_refl v1) (S v) (plusnS _ _))
_ (diff_l _ _)
end
end.
@@ -197,17 +197,17 @@ Fixpoint diff_r (m n: nat) {struct m}: snd (diff m n) + m = max m n :=
match m return (snd (diff m n) + m = max m n) with
| 0 =>
match n return (snd (diff 0 n) + 0 = max 0 n) with
- | 0 => refl_equal _
+ | 0 => eq_refl _
| S _ => plusn0 _
end
| S m =>
match n return (snd (diff (S m) n) + S m = max (S m) n) with
- | 0 => refl_equal (snd (diff (S m) 0) + S m)
+ | 0 => eq_refl (snd (diff (S m) 0) + S m)
| S n1 =>
let v := S (max m n1) in
eq_ind_r (fun n => n = v)
(eq_ind_r (fun n => S n = v)
- (refl_equal v) (diff_r _ _)) (plusnS _ _)
+ (eq_refl v) (diff_r _ _)) (plusnS _ _)
end
end.
@@ -216,7 +216,7 @@ Fixpoint diff_r (m n: nat) {struct m}: snd (diff m n) + m = max m n :=
Definition castm (m n: nat) (H: m = n) (x: word w (S m)):
(word w (S n)) :=
match H in (_ = y) return (word w (S y)) with
- | refl_equal => x
+ | eq_refl => x
end.
Variable m: nat.
@@ -314,7 +314,7 @@ Section CompareRec.
Lemma base_xO: forall n, base (xO n) = (base n)^2.
Proof.
intros n1; unfold base.
- rewrite (Zpos_xO n1); rewrite Zmult_comm; rewrite Zpower_mult; auto with zarith.
+ rewrite (Pos2Z.inj_xO n1); rewrite Z.mul_comm; rewrite Z.pow_mul_r; auto with zarith.
Qed.
Let double_to_Z_pos: forall n x, 0 <= double_to_Z n x < double_wB n :=
@@ -332,13 +332,13 @@ Section CompareRec.
rewrite 2 Hrec.
simpl double_to_Z.
set (wB := DoubleBase.double_wB wm_base n).
- case Zcompare_spec; intros Cmp.
+ case Z.compare_spec; intros Cmp.
rewrite <- Cmp. reflexivity.
- symmetry. apply Zgt_lt, Zlt_gt. (* ;-) *)
+ symmetry. apply Z.gt_lt, Z.lt_gt. (* ;-) *)
assert (0 < wB).
unfold wB, DoubleBase.double_wB, base; auto with zarith.
- change 0 with (0 + 0); apply Zplus_lt_le_compat; auto with zarith.
- apply Zmult_lt_0_compat; auto with zarith.
+ change 0 with (0 + 0); apply Z.add_lt_le_mono; auto with zarith.
+ apply Z.mul_pos_pos; auto with zarith.
case (double_to_Z_pos n xl); auto with zarith.
case (double_to_Z_pos n xh); intros; exfalso; omega.
Qed.
@@ -358,9 +358,9 @@ Section CompareRec.
end.
Variable spec_compare: forall x y,
- compare x y = Zcompare (w_to_Z x) (w_to_Z y).
+ compare x y = Z.compare (w_to_Z x) (w_to_Z y).
Variable spec_compare_m: forall x y,
- compare_m x y = Zcompare (wm_to_Z x) (w_to_Z y).
+ compare_m x y = Z.compare (wm_to_Z x) (w_to_Z y).
Variable wm_base_lt: forall x,
0 <= w_to_Z x < base (wm_base).
@@ -369,35 +369,35 @@ Section CompareRec.
Proof.
intros n x; elim n; simpl; auto; clear n.
intros n (H0, H); split; auto.
- apply Zlt_le_trans with (1:= H).
+ apply Z.lt_le_trans with (1:= H).
unfold double_wB, DoubleBase.double_wB; simpl.
rewrite Pshiftl_nat_S, base_xO.
set (u := base (Pos.shiftl_nat wm_base n)).
assert (0 < u).
unfold u, base; auto with zarith.
replace (u^2) with (u * u); simpl; auto with zarith.
- apply Zle_trans with (1 * u); auto with zarith.
- unfold Zpower_pos; simpl; ring.
+ apply Z.le_trans with (1 * u); auto with zarith.
+ unfold Z.pow_pos; simpl; ring.
Qed.
Lemma spec_compare_mn_1: forall n x y,
- compare_mn_1 n x y = Zcompare (double_to_Z n x) (w_to_Z y).
+ compare_mn_1 n x y = Z.compare (double_to_Z n x) (w_to_Z y).
Proof.
intros n; elim n; simpl; auto; clear n.
intros n Hrec x; case x; clear x; auto.
intros y; rewrite spec_compare; rewrite w_to_Z_0. reflexivity.
intros xh xl y; simpl;
- rewrite spec_compare0_mn, Hrec. case Zcompare_spec.
+ rewrite spec_compare0_mn, Hrec. case Z.compare_spec.
intros H1b.
- rewrite <- H1b; rewrite Zmult_0_l; rewrite Zplus_0_l; auto.
- symmetry. apply Zlt_gt.
+ rewrite <- H1b; rewrite Z.mul_0_l; rewrite Z.add_0_l; auto.
+ symmetry. apply Z.lt_gt.
case (double_wB_lt n y); intros _ H0.
- apply Zlt_le_trans with (1:= H0).
+ apply Z.lt_le_trans with (1:= H0).
fold double_wB.
case (double_to_Z_pos n xl); intros H1 H2.
- apply Zle_trans with (double_to_Z n xh * double_wB n); auto with zarith.
- apply Zle_trans with (1 * double_wB n); auto with zarith.
+ apply Z.le_trans with (double_to_Z n xh * double_wB n); auto with zarith.
+ apply Z.le_trans with (1 * double_wB n); auto with zarith.
case (double_to_Z_pos n xh); intros; exfalso; omega.
Qed.
@@ -440,8 +440,8 @@ End AddS.
Proof.
intros x; elim x; clear x; [intros x1 Hrec | intros x1 Hrec | idtac];
intros y; case y; clear y; intros y1 H || intros H; simpl length_pos;
- try (rewrite (Zpos_xI x1) || rewrite (Zpos_xO x1));
- try (rewrite (Zpos_xI y1) || rewrite (Zpos_xO y1));
+ try (rewrite (Pos2Z.inj_xI x1) || rewrite (Pos2Z.inj_xO x1));
+ try (rewrite (Pos2Z.inj_xI y1) || rewrite (Pos2Z.inj_xO y1));
try (inversion H; fail);
try (assert (Zpos x1 < Zpos y1); [apply Hrec; apply lt_S_n | idtac]; auto with zarith);
assert (0 < Zpos y1); auto with zarith; red; auto.
diff --git a/theories/Numbers/Natural/Binary/NBinary.v b/theories/Numbers/Natural/Binary/NBinary.v
index 43ca67dd3..3111d327c 100644
--- a/theories/Numbers/Natural/Binary/NBinary.v
+++ b/theories/Numbers/Natural/Binary/NBinary.v
@@ -31,8 +31,8 @@ Time Eval vm_compute in (log 500000). (* 11 sec *)
Fixpoint binposlog (p : positive) : N :=
match p with
| xH => 0
-| xO p' => Nsucc (binposlog p')
-| xI p' => Nsucc (binposlog p')
+| xO p' => N.succ (binposlog p')
+| xI p' => N.succ (binposlog p')
end.
Definition binlog (n : N) : N :=
diff --git a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
index 2c7884ac4..ce92f44d9 100644
--- a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
+++ b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
@@ -318,7 +318,7 @@ Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
Theorem div_mod : forall a b, ~b==0 -> a == b*(div a b) + (modulo a b).
Proof.
-intros a b. zify. intros. apply Z_div_mod_eq_full; auto.
+intros a b. zify. intros. apply Z.div_mod; auto.
Qed.
Theorem mod_bound_pos : forall a b, 0<=a -> 0<b ->
@@ -444,7 +444,7 @@ Qed.
(** Recursion *)
Definition recursion (A : Type) (a : A) (f : NN.t -> A -> A) (n : NN.t) :=
- Nrect (fun _ => A) a (fun n a => f (NN.of_N n) a) (NN.to_N n).
+ N.peano_rect (fun _ => A) a (fun n a => f (NN.of_N n) a) (NN.to_N n).
Arguments recursion [A] a f n.
Instance recursion_wd (A : Type) (Aeq : relation A) :
@@ -457,7 +457,7 @@ unfold NN.to_N.
rewrite <- Exx'; clear x' Exx'.
induction (Z.to_N [x]) using N.peano_ind.
simpl; auto.
-rewrite 2 Nrect_step. now apply Eff'.
+rewrite 2 N.peano_rect_succ. now apply Eff'.
Qed.
Theorem recursion_0 :
@@ -474,7 +474,7 @@ Proof.
unfold eq, recursion; intros A Aeq a f EAaa f_wd n.
replace (to_N (succ n)) with (N.succ (to_N n)) by
(zify; now rewrite <- Z2N.inj_succ by apply spec_pos).
-rewrite Nrect_step.
+rewrite N.peano_rect_succ.
apply f_wd; auto.
zify. now rewrite Z2N.id by apply spec_pos.
fold (recursion a f n). apply recursion_wd; auto. red; auto.
diff --git a/theories/Numbers/Rational/BigQ/BigQ.v b/theories/Numbers/Rational/BigQ/BigQ.v
index 424db5b75..bc50c4148 100644
--- a/theories/Numbers/Rational/BigQ/BigQ.v
+++ b/theories/Numbers/Rational/BigQ/BigQ.v
@@ -26,10 +26,10 @@ Module BigN_BigZ <: NType_ZType BigN.BigN BigZ.
reflexivity.
Qed.
Definition Zabs_N := BigZ.to_N.
- Lemma spec_Zabs_N : forall z, BigN.to_Z (Zabs_N z) = Zabs (BigZ.to_Z z).
+ Lemma spec_Zabs_N : forall z, BigN.to_Z (Zabs_N z) = Z.abs (BigZ.to_Z z).
Proof.
unfold Zabs_N; intros.
- rewrite BigZ.spec_to_Z, Zmult_comm; apply Zsgn_Zabs.
+ rewrite BigZ.spec_to_Z, Z.mul_comm; apply Z.sgn_abs.
Qed.
End BigN_BigZ.
@@ -89,10 +89,10 @@ exact BigQ.div_mul_inv. exact BigQ.mul_inv_diag_l.
Qed.
Lemma BigQpowerth :
- power_theory 1 BigQ.mul BigQ.eq Z_of_N BigQ.power.
+ power_theory 1 BigQ.mul BigQ.eq Z.of_N BigQ.power.
Proof.
constructor. intros. BigQ.qify.
-replace ([r] ^ Z_of_N n)%Q with (pow_N 1 Qmult [r] n)%Q by (now destruct n).
+replace ([r] ^ Z.of_N n)%Q with (pow_N 1 Qmult [r] n)%Q by (now destruct n).
destruct n. reflexivity.
induction p; simpl; auto; rewrite ?BigQ.spec_mul, ?IHp; reflexivity.
Qed.
diff --git a/theories/Numbers/Rational/BigQ/QMake.v b/theories/Numbers/Rational/BigQ/QMake.v
index 995fbb9ee..78ad7c470 100644
--- a/theories/Numbers/Rational/BigQ/QMake.v
+++ b/theories/Numbers/Rational/BigQ/QMake.v
@@ -19,14 +19,14 @@ Require Import NSig ZSig QSig.
denominators. But first we will need some glue between [NType] and
[ZType]. *)
-Module Type NType_ZType (N:NType)(Z:ZType).
- Parameter Z_of_N : N.t -> Z.t.
- Parameter spec_Z_of_N : forall n, Z.to_Z (Z_of_N n) = N.to_Z n.
- Parameter Zabs_N : Z.t -> N.t.
- Parameter spec_Zabs_N : forall z, N.to_Z (Zabs_N z) = Zabs (Z.to_Z z).
+Module Type NType_ZType (NN:NType)(ZZ:ZType).
+ Parameter Z_of_N : NN.t -> ZZ.t.
+ Parameter spec_Z_of_N : forall n, ZZ.to_Z (Z_of_N n) = NN.to_Z n.
+ Parameter Zabs_N : ZZ.t -> NN.t.
+ Parameter spec_Zabs_N : forall z, NN.to_Z (Zabs_N z) = Z.abs (ZZ.to_Z z).
End NType_ZType.
-Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
+Module Make (NN:NType)(ZZ:ZType)(Import NZ:NType_ZType NN ZZ) <: QType.
(** The notation of a rational number is either an integer x,
interpreted as itself or a pair (x,y) of an integer x and a natural
@@ -34,8 +34,8 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
interpreted as 0. *)
Inductive t_ :=
- | Qz : Z.t -> t_
- | Qq : Z.t -> N.t -> t_.
+ | Qz : ZZ.t -> t_
+ | Qq : ZZ.t -> NN.t -> t_.
Definition t := t_.
@@ -45,41 +45,41 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Local Open Scope Q_scope.
- Definition of_Z x: t := Qz (Z.of_Z x).
+ Definition of_Z x: t := Qz (ZZ.of_Z x).
Definition of_Q (q:Q) : t :=
let (x,y) := q in
match y with
- | 1%positive => Qz (Z.of_Z x)
- | _ => Qq (Z.of_Z x) (N.of_N (Npos y))
+ | 1%positive => Qz (ZZ.of_Z x)
+ | _ => Qq (ZZ.of_Z x) (NN.of_N (Npos y))
end.
Definition to_Q (q: t) :=
match q with
- | Qz x => Z.to_Z x # 1
- | Qq x y => if N.eqb y N.zero then 0
- else Z.to_Z x # Z2P (N.to_Z y)
+ | Qz x => ZZ.to_Z x # 1
+ | Qq x y => if NN.eqb y NN.zero then 0
+ else ZZ.to_Z x # Z.to_pos (NN.to_Z y)
end.
Notation "[ x ]" := (to_Q x).
Lemma N_to_Z_pos :
- forall x, (N.to_Z x <> N.to_Z N.zero)%Z -> (0 < N.to_Z x)%Z.
+ forall x, (NN.to_Z x <> NN.to_Z NN.zero)%Z -> (0 < NN.to_Z x)%Z.
Proof.
- intros x; rewrite N.spec_0; generalize (N.spec_pos x). romega.
+ intros x; rewrite NN.spec_0; generalize (NN.spec_pos x). romega.
Qed.
Ltac destr_zcompare := case Z.compare_spec; intros ?H.
Ltac destr_eqb :=
match goal with
- | |- context [Z.eqb ?x ?y] =>
- rewrite (Z.spec_eqb x y);
- case (Z.eqb_spec (Z.to_Z x) (Z.to_Z y));
+ | |- context [ZZ.eqb ?x ?y] =>
+ rewrite (ZZ.spec_eqb x y);
+ case (Z.eqb_spec (ZZ.to_Z x) (ZZ.to_Z y));
destr_eqb
- | |- context [N.eqb ?x ?y] =>
- rewrite (N.spec_eqb x y);
- case (Z.eqb_spec (N.to_Z x) (N.to_Z y));
+ | |- context [NN.eqb ?x ?y] =>
+ rewrite (NN.spec_eqb x y);
+ case (Z.eqb_spec (NN.to_Z x) (NN.to_Z y));
[ | let H:=fresh "H" in
try (intro H;generalize (N_to_Z_pos _ H); clear H)];
destr_eqb
@@ -87,11 +87,11 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
end.
Hint Rewrite
- Zplus_0_r Zplus_0_l Zmult_0_r Zmult_0_l Zmult_1_r Zmult_1_l
- Z.spec_0 N.spec_0 Z.spec_1 N.spec_1 Z.spec_m1 Z.spec_opp
- Z.spec_compare N.spec_compare
- Z.spec_add N.spec_add Z.spec_mul N.spec_mul Z.spec_div N.spec_div
- Z.spec_gcd N.spec_gcd Zgcd_Zabs Zgcd_1
+ Z.add_0_r Z.add_0_l Z.mul_0_r Z.mul_0_l Z.mul_1_r Z.mul_1_l
+ ZZ.spec_0 NN.spec_0 ZZ.spec_1 NN.spec_1 ZZ.spec_m1 ZZ.spec_opp
+ ZZ.spec_compare NN.spec_compare
+ ZZ.spec_add NN.spec_add ZZ.spec_mul NN.spec_mul ZZ.spec_div NN.spec_div
+ ZZ.spec_gcd NN.spec_gcd Z.gcd_abs_l Z.gcd_1_r
spec_Z_of_N spec_Zabs_N
: nz.
@@ -99,13 +99,13 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Ltac qsimpl := try red; unfold to_Q; simpl; intros;
destr_eqb; simpl; nzsimpl; intros;
- rewrite ?Z2P_correct by auto;
+ rewrite ?Z2Pos.id by auto;
auto.
Theorem strong_spec_of_Q: forall q: Q, [of_Q q] = q.
Proof.
- intros(x,y); destruct y; simpl; rewrite ?Z.spec_of_Z; auto;
- destr_eqb; now rewrite ?N.spec_0, ?N.spec_of_N.
+ intros(x,y); destruct y; simpl; rewrite ?ZZ.spec_of_Z; auto;
+ destr_eqb; now rewrite ?NN.spec_0, ?NN.spec_of_N.
Qed.
Theorem spec_of_Q: forall q: Q, [of_Q q] == q.
@@ -115,9 +115,9 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Definition eq x y := [x] == [y].
- Definition zero: t := Qz Z.zero.
- Definition one: t := Qz Z.one.
- Definition minus_one: t := Qz Z.minus_one.
+ Definition zero: t := Qz ZZ.zero.
+ Definition one: t := Qz ZZ.one.
+ Definition minus_one: t := Qz ZZ.minus_one.
Lemma spec_0: [zero] == 0.
Proof.
@@ -136,20 +136,20 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Definition compare (x y: t) :=
match x, y with
- | Qz zx, Qz zy => Z.compare zx zy
+ | Qz zx, Qz zy => ZZ.compare zx zy
| Qz zx, Qq ny dy =>
- if N.eqb dy N.zero then Z.compare zx Z.zero
- else Z.compare (Z.mul zx (Z_of_N dy)) ny
+ if NN.eqb dy NN.zero then ZZ.compare zx ZZ.zero
+ else ZZ.compare (ZZ.mul zx (Z_of_N dy)) ny
| Qq nx dx, Qz zy =>
- if N.eqb dx N.zero then Z.compare Z.zero zy
- else Z.compare nx (Z.mul zy (Z_of_N dx))
+ if NN.eqb dx NN.zero then ZZ.compare ZZ.zero zy
+ else ZZ.compare nx (ZZ.mul zy (Z_of_N dx))
| Qq nx dx, Qq ny dy =>
- match N.eqb dx N.zero, N.eqb dy N.zero with
+ match NN.eqb dx NN.zero, NN.eqb dy NN.zero with
| true, true => Eq
- | true, false => Z.compare Z.zero ny
- | false, true => Z.compare nx Z.zero
- | false, false => Z.compare (Z.mul nx (Z_of_N dy))
- (Z.mul ny (Z_of_N dx))
+ | true, false => ZZ.compare ZZ.zero ny
+ | false, true => ZZ.compare nx ZZ.zero
+ | false, false => ZZ.compare (ZZ.mul nx (Z_of_N dy))
+ (ZZ.mul ny (Z_of_N dx))
end
end.
@@ -188,7 +188,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(** [check_int] : is a reduced fraction [n/d] in fact a integer ? *)
Definition check_int n d :=
- match N.compare N.one d with
+ match NN.compare NN.one d with
| Lt => Qq n d
| Eq => Qz n
| Gt => zero (* n/0 encodes 0 *)
@@ -207,9 +207,9 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(** Normalisation function *)
Definition norm n d : t :=
- let gcd := N.gcd (Zabs_N n) d in
- match N.compare N.one gcd with
- | Lt => check_int (Z.div n (Z_of_N gcd)) (N.div d gcd)
+ let gcd := NN.gcd (Zabs_N n) d in
+ match NN.compare NN.one gcd with
+ | Lt => check_int (ZZ.div n (Z_of_N gcd)) (NN.div d gcd)
| Eq => check_int n d
| Gt => zero (* gcd = 0 => both numbers are 0 *)
end.
@@ -217,8 +217,8 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_norm: forall n q, [norm n q] == [Qq n q].
Proof.
intros p q; unfold norm.
- assert (Hp := N.spec_pos (Zabs_N p)).
- assert (Hq := N.spec_pos q).
+ assert (Hp := NN.spec_pos (Zabs_N p)).
+ assert (Hq := NN.spec_pos q).
nzsimpl.
destr_zcompare.
(* Eq *)
@@ -226,15 +226,15 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(* Lt *)
rewrite strong_spec_check_int.
qsimpl.
- generalize (Zgcd_div_pos (Z.to_Z p) (N.to_Z q)). romega.
- replace (N.to_Z q) with 0%Z in * by assumption.
+ generalize (Zgcd_div_pos (ZZ.to_Z p) (NN.to_Z q)). romega.
+ replace (NN.to_Z q) with 0%Z in * by assumption.
rewrite Zdiv_0_l in *; auto with zarith.
apply Zgcd_div_swap0; romega.
(* Gt *)
qsimpl.
- assert (H' : Zgcd (Z.to_Z p) (N.to_Z q) = 0%Z).
- generalize (Zgcd_is_pos (Z.to_Z p) (N.to_Z q)); romega.
- symmetry; apply (Zgcd_inv_0_l _ _ H'); auto.
+ assert (H' : Z.gcd (ZZ.to_Z p) (NN.to_Z q) = 0%Z).
+ generalize (Z.gcd_nonneg (ZZ.to_Z p) (NN.to_Z q)); romega.
+ symmetry; apply (Z.gcd_eq_0_l _ _ H'); auto.
Qed.
Theorem strong_spec_norm : forall p q, [norm p q] = Qred [Qq p q].
@@ -244,8 +244,8 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(apply Qred_complete; apply spec_norm).
symmetry; apply Qred_identity.
unfold norm.
- assert (Hp := N.spec_pos (Zabs_N p)).
- assert (Hq := N.spec_pos q).
+ assert (Hp := NN.spec_pos (Zabs_N p)).
+ assert (Hq := NN.spec_pos q).
nzsimpl.
destr_zcompare; rewrite ?strong_spec_check_int.
(* Eq *)
@@ -253,10 +253,10 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(* Lt *)
qsimpl.
rewrite Zgcd_1_rel_prime.
- destruct (Z_lt_le_dec 0 (N.to_Z q)).
+ destruct (Z_lt_le_dec 0 (NN.to_Z q)).
apply Zis_gcd_rel_prime; auto with zarith.
apply Zgcd_is_gcd.
- replace (N.to_Z q) with 0%Z in * by romega.
+ replace (NN.to_Z q) with 0%Z in * by romega.
rewrite Zdiv_0_l in *; romega.
(* Gt *)
simpl; auto with zarith.
@@ -292,20 +292,20 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
match x with
| Qz zx =>
match y with
- | Qz zy => Qz (Z.add zx zy)
+ | Qz zy => Qz (ZZ.add zx zy)
| Qq ny dy =>
- if N.eqb dy N.zero then x
- else Qq (Z.add (Z.mul zx (Z_of_N dy)) ny) dy
+ if NN.eqb dy NN.zero then x
+ else Qq (ZZ.add (ZZ.mul zx (Z_of_N dy)) ny) dy
end
| Qq nx dx =>
- if N.eqb dx N.zero then y
+ if NN.eqb dx NN.zero then y
else match y with
- | Qz zy => Qq (Z.add nx (Z.mul zy (Z_of_N dx))) dx
+ | Qz zy => Qq (ZZ.add nx (ZZ.mul zy (Z_of_N dx))) dx
| Qq ny dy =>
- if N.eqb dy N.zero then x
+ if NN.eqb dy NN.zero then x
else
- let n := Z.add (Z.mul nx (Z_of_N dy)) (Z.mul ny (Z_of_N dx)) in
- let d := N.mul dx dy in
+ let n := ZZ.add (ZZ.mul nx (Z_of_N dy)) (ZZ.mul ny (Z_of_N dx)) in
+ let d := NN.mul dx dy in
Qq n d
end
end.
@@ -314,30 +314,30 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Proof.
intros [x | nx dx] [y | ny dy]; unfold Qplus; qsimpl;
auto with zarith.
- rewrite Pmult_1_r, Z2P_correct; auto.
- rewrite Pmult_1_r, Z2P_correct; auto.
- destruct (Zmult_integral (N.to_Z dx) (N.to_Z dy)); intuition.
- rewrite Zpos_mult_morphism, 2 Z2P_correct; auto.
+ rewrite Pos.mul_1_r, Z2Pos.id; auto.
+ rewrite Pos.mul_1_r, Z2Pos.id; auto.
+ rewrite Z.mul_eq_0 in *; intuition.
+ rewrite Pos2Z.inj_mul, 2 Z2Pos.id; auto.
Qed.
Definition add_norm (x y: t): t :=
match x with
| Qz zx =>
match y with
- | Qz zy => Qz (Z.add zx zy)
+ | Qz zy => Qz (ZZ.add zx zy)
| Qq ny dy =>
- if N.eqb dy N.zero then x
- else norm (Z.add (Z.mul zx (Z_of_N dy)) ny) dy
+ if NN.eqb dy NN.zero then x
+ else norm (ZZ.add (ZZ.mul zx (Z_of_N dy)) ny) dy
end
| Qq nx dx =>
- if N.eqb dx N.zero then y
+ if NN.eqb dx NN.zero then y
else match y with
- | Qz zy => norm (Z.add nx (Z.mul zy (Z_of_N dx))) dx
+ | Qz zy => norm (ZZ.add nx (ZZ.mul zy (Z_of_N dx))) dx
| Qq ny dy =>
- if N.eqb dy N.zero then x
+ if NN.eqb dy NN.zero then x
else
- let n := Z.add (Z.mul nx (Z_of_N dy)) (Z.mul ny (Z_of_N dx)) in
- let d := N.mul dx dy in
+ let n := ZZ.add (ZZ.mul nx (Z_of_N dy)) (ZZ.mul ny (Z_of_N dx)) in
+ let d := NN.mul dx dy in
norm n d
end
end.
@@ -363,18 +363,18 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Definition opp (x: t): t :=
match x with
- | Qz zx => Qz (Z.opp zx)
- | Qq nx dx => Qq (Z.opp nx) dx
+ | Qz zx => Qz (ZZ.opp zx)
+ | Qq nx dx => Qq (ZZ.opp nx) dx
end.
Theorem strong_spec_opp: forall q, [opp q] = -[q].
Proof.
intros [z | x y]; simpl.
- rewrite Z.spec_opp; auto.
- match goal with |- context[N.eqb ?X ?Y] =>
- generalize (N.spec_eqb X Y); case N.eqb
- end; auto; rewrite N.spec_0.
- rewrite Z.spec_opp; auto.
+ rewrite ZZ.spec_opp; auto.
+ match goal with |- context[NN.eqb ?X ?Y] =>
+ generalize (NN.spec_eqb X Y); case NN.eqb
+ end; auto; rewrite NN.spec_0.
+ rewrite ZZ.spec_opp; auto.
Qed.
Theorem spec_opp : forall q, [opp q] == -[q].
@@ -416,28 +416,28 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Definition mul (x y: t): t :=
match x, y with
- | Qz zx, Qz zy => Qz (Z.mul zx zy)
- | Qz zx, Qq ny dy => Qq (Z.mul zx ny) dy
- | Qq nx dx, Qz zy => Qq (Z.mul nx zy) dx
- | Qq nx dx, Qq ny dy => Qq (Z.mul nx ny) (N.mul dx dy)
+ | Qz zx, Qz zy => Qz (ZZ.mul zx zy)
+ | Qz zx, Qq ny dy => Qq (ZZ.mul zx ny) dy
+ | Qq nx dx, Qz zy => Qq (ZZ.mul nx zy) dx
+ | Qq nx dx, Qq ny dy => Qq (ZZ.mul nx ny) (NN.mul dx dy)
end.
Ltac nsubst :=
- match goal with E : N.to_Z _ = _ |- _ => rewrite E in * end.
+ match goal with E : NN.to_Z _ = _ |- _ => rewrite E in * end.
Theorem spec_mul : forall x y, [mul x y] == [x] * [y].
Proof.
intros [x | nx dx] [y | ny dy]; unfold Qmult; simpl; qsimpl.
- rewrite Pmult_1_r, Z2P_correct; auto.
- destruct (Zmult_integral (N.to_Z dx) (N.to_Z dy)); intuition.
+ rewrite Pos.mul_1_r, Z2Pos.id; auto.
+ rewrite Z.mul_eq_0 in *; intuition.
nsubst; auto with zarith.
nsubst; auto with zarith.
nsubst; nzsimpl; auto with zarith.
- rewrite Zpos_mult_morphism, 2 Z2P_correct; auto.
+ rewrite Pos2Z.inj_mul, 2 Z2Pos.id; auto.
Qed.
Definition norm_denum n d :=
- if N.eqb d N.one then Qz n else Qq n d.
+ if NN.eqb d NN.one then Qz n else Qq n d.
Lemma spec_norm_denum : forall n d,
[norm_denum n d] == [Qq n d].
@@ -448,40 +448,40 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Qed.
Definition irred n d :=
- let gcd := N.gcd (Zabs_N n) d in
- match N.compare gcd N.one with
- | Gt => (Z.div n (Z_of_N gcd), N.div d gcd)
+ let gcd := NN.gcd (Zabs_N n) d in
+ match NN.compare gcd NN.one with
+ | Gt => (ZZ.div n (Z_of_N gcd), NN.div d gcd)
| _ => (n, d)
end.
Lemma spec_irred : forall n d, exists g,
let (n',d') := irred n d in
- (Z.to_Z n' * g = Z.to_Z n)%Z /\ (N.to_Z d' * g = N.to_Z d)%Z.
+ (ZZ.to_Z n' * g = ZZ.to_Z n)%Z /\ (NN.to_Z d' * g = NN.to_Z d)%Z.
Proof.
intros.
unfold irred; nzsimpl; simpl.
destr_zcompare.
exists 1%Z; nzsimpl; auto.
exists 0%Z; nzsimpl.
- assert (Zgcd (Z.to_Z n) (N.to_Z d) = 0%Z).
- generalize (Zgcd_is_pos (Z.to_Z n) (N.to_Z d)); romega.
+ assert (Z.gcd (ZZ.to_Z n) (NN.to_Z d) = 0%Z).
+ generalize (Z.gcd_nonneg (ZZ.to_Z n) (NN.to_Z d)); romega.
clear H.
split.
- symmetry; apply (Zgcd_inv_0_l _ _ H0).
- symmetry; apply (Zgcd_inv_0_r _ _ H0).
- exists (Zgcd (Z.to_Z n) (N.to_Z d)).
+ symmetry; apply (Z.gcd_eq_0_l _ _ H0).
+ symmetry; apply (Z.gcd_eq_0_r _ _ H0).
+ exists (Z.gcd (ZZ.to_Z n) (NN.to_Z d)).
simpl.
split.
nzsimpl.
- destruct (Zgcd_is_gcd (Z.to_Z n) (N.to_Z d)).
- rewrite Zmult_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
+ destruct (Zgcd_is_gcd (ZZ.to_Z n) (NN.to_Z d)).
+ rewrite Z.mul_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
nzsimpl.
- destruct (Zgcd_is_gcd (Z.to_Z n) (N.to_Z d)).
- rewrite Zmult_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
+ destruct (Zgcd_is_gcd (ZZ.to_Z n) (NN.to_Z d)).
+ rewrite Z.mul_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
Qed.
Lemma spec_irred_zero : forall n d,
- (N.to_Z d = 0)%Z <-> (N.to_Z (snd (irred n d)) = 0)%Z.
+ (NN.to_Z d = 0)%Z <-> (NN.to_Z (snd (irred n d)) = 0)%Z.
Proof.
intros.
unfold irred.
@@ -494,8 +494,8 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
nzsimpl; destr_zcompare; simpl; auto.
nzsimpl.
intros.
- generalize (N.spec_pos d); intros.
- destruct (N.to_Z d); auto.
+ generalize (NN.spec_pos d); intros.
+ destruct (NN.to_Z d); auto.
assert (0 < 0)%Z.
rewrite <- H0 at 2.
apply Zgcd_div_pos; auto with zarith.
@@ -505,49 +505,49 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Qed.
Lemma strong_spec_irred : forall n d,
- (N.to_Z d <> 0%Z) ->
- let (n',d') := irred n d in Zgcd (Z.to_Z n') (N.to_Z d') = 1%Z.
+ (NN.to_Z d <> 0%Z) ->
+ let (n',d') := irred n d in Z.gcd (ZZ.to_Z n') (NN.to_Z d') = 1%Z.
Proof.
unfold irred; intros.
nzsimpl.
destr_zcompare; simpl; auto.
elim H.
- apply (Zgcd_inv_0_r (Z.to_Z n)).
- generalize (Zgcd_is_pos (Z.to_Z n) (N.to_Z d)); romega.
+ apply (Z.gcd_eq_0_r (ZZ.to_Z n)).
+ generalize (Z.gcd_nonneg (ZZ.to_Z n) (NN.to_Z d)); romega.
nzsimpl.
rewrite Zgcd_1_rel_prime.
apply Zis_gcd_rel_prime.
- generalize (N.spec_pos d); romega.
- generalize (Zgcd_is_pos (Z.to_Z n) (N.to_Z d)); romega.
+ generalize (NN.spec_pos d); romega.
+ generalize (Z.gcd_nonneg (ZZ.to_Z n) (NN.to_Z d)); romega.
apply Zgcd_is_gcd; auto.
Qed.
Definition mul_norm_Qz_Qq z n d :=
- if Z.eqb z Z.zero then zero
+ if ZZ.eqb z ZZ.zero then zero
else
- let gcd := N.gcd (Zabs_N z) d in
- match N.compare gcd N.one with
+ let gcd := NN.gcd (Zabs_N z) d in
+ match NN.compare gcd NN.one with
| Gt =>
- let z := Z.div z (Z_of_N gcd) in
- let d := N.div d gcd in
- norm_denum (Z.mul z n) d
- | _ => Qq (Z.mul z n) d
+ let z := ZZ.div z (Z_of_N gcd) in
+ let d := NN.div d gcd in
+ norm_denum (ZZ.mul z n) d
+ | _ => Qq (ZZ.mul z n) d
end.
Definition mul_norm (x y: t): t :=
match x, y with
- | Qz zx, Qz zy => Qz (Z.mul zx zy)
+ | Qz zx, Qz zy => Qz (ZZ.mul zx zy)
| Qz zx, Qq ny dy => mul_norm_Qz_Qq zx ny dy
| Qq nx dx, Qz zy => mul_norm_Qz_Qq zy nx dx
| Qq nx dx, Qq ny dy =>
let (nx, dy) := irred nx dy in
let (ny, dx) := irred ny dx in
- norm_denum (Z.mul ny nx) (N.mul dx dy)
+ norm_denum (ZZ.mul ny nx) (NN.mul dx dy)
end.
Lemma spec_mul_norm_Qz_Qq : forall z n d,
- [mul_norm_Qz_Qq z n d] == [Qq (Z.mul z n) d].
+ [mul_norm_Qz_Qq z n d] == [Qq (ZZ.mul z n) d].
Proof.
intros z n d; unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
destr_eqb; nzsimpl; intros Hz.
@@ -558,7 +558,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
qsimpl.
rewrite Zdiv_gcd_zero in GT; auto with zarith.
nsubst. rewrite Zdiv_0_l in *; discriminate.
- rewrite <- Zmult_assoc, (Zmult_comm (Z.to_Z n)), Zmult_assoc.
+ rewrite <- Z.mul_assoc, (Z.mul_comm (ZZ.to_Z n)), Z.mul_assoc.
rewrite Zgcd_div_swap0; try romega.
ring.
Qed.
@@ -582,34 +582,34 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destr_eqb; simpl; nzsimpl; auto.
nzsimpl; rewrite Hd, Zdiv_0_l; auto with zarith.
- rewrite Z2P_correct in H; auto.
+ rewrite Z2Pos.id in H; auto.
unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
destr_eqb; intros Hz; simpl; nzsimpl; simpl; auto.
destruct Z_le_gt_dec as [H'|H'].
simpl; nzsimpl.
destr_eqb; simpl; nzsimpl; auto.
intros.
- rewrite Z2P_correct; auto.
+ rewrite Z2Pos.id; auto.
apply Zgcd_mult_rel_prime; auto.
- generalize (Zgcd_inv_0_l (Z.to_Z z) (N.to_Z d))
- (Zgcd_is_pos (Z.to_Z z) (N.to_Z d)); romega.
+ generalize (Z.gcd_eq_0_l (ZZ.to_Z z) (NN.to_Z d))
+ (Z.gcd_nonneg (ZZ.to_Z z) (NN.to_Z d)); romega.
destr_eqb; simpl; nzsimpl; auto.
unfold norm_denum.
destr_eqb; nzsimpl; simpl; destr_eqb; simpl; auto.
intros; nzsimpl.
- rewrite Z2P_correct; auto.
+ rewrite Z2Pos.id; auto.
apply Zgcd_mult_rel_prime.
rewrite Zgcd_1_rel_prime.
apply Zis_gcd_rel_prime.
- generalize (N.spec_pos d); romega.
- generalize (Zgcd_is_pos (Z.to_Z z) (N.to_Z d)); romega.
+ generalize (NN.spec_pos d); romega.
+ generalize (Z.gcd_nonneg (ZZ.to_Z z) (NN.to_Z d)); romega.
apply Zgcd_is_gcd.
- destruct (Zgcd_is_gcd (Z.to_Z z) (N.to_Z d)) as [ (z0,Hz0) (d0,Hd0) Hzd].
- replace (N.to_Z d / Zgcd (Z.to_Z z) (N.to_Z d))%Z with d0.
+ destruct (Zgcd_is_gcd (ZZ.to_Z z) (NN.to_Z d)) as [ (z0,Hz0) (d0,Hd0) Hzd].
+ replace (NN.to_Z d / Z.gcd (ZZ.to_Z z) (NN.to_Z d))%Z with d0.
rewrite Zgcd_1_rel_prime in *.
apply bezout_rel_prime.
destruct (rel_prime_bezout _ _ H) as [u v Huv].
- apply Bezout_intro with u (v*(Zgcd (Z.to_Z z) (N.to_Z d)))%Z.
+ apply Bezout_intro with u (v*(Z.gcd (ZZ.to_Z z) (NN.to_Z d)))%Z.
rewrite <- Huv; rewrite Hd0 at 2; ring.
rewrite Hd0 at 1.
symmetry; apply Z_div_mult_full; auto with zarith.
@@ -634,14 +634,14 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
qsimpl.
match goal with E : (_ * _ = 0)%Z |- _ =>
- destruct (Zmult_integral _ _ E) as [Eq|Eq] end.
+ rewrite Z.mul_eq_0 in E; destruct E as [Eq|Eq] end.
rewrite Eq in *; simpl in *.
rewrite <- Hg2' in *; auto with zarith.
rewrite Eq in *; simpl in *.
rewrite <- Hg2 in *; auto with zarith.
match goal with E : (_ * _ = 0)%Z |- _ =>
- destruct (Zmult_integral _ _ E) as [Eq|Eq] end.
+ rewrite Z.mul_eq_0 in E; destruct E as [Eq|Eq] end.
rewrite Hz' in Eq; rewrite Eq in *; auto with zarith.
rewrite Hz in Eq; rewrite Eq in *; auto with zarith.
@@ -671,31 +671,31 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
unfold norm_denum; qsimpl.
- assert (NEQ : N.to_Z dy <> 0%Z) by
+ assert (NEQ : NN.to_Z dy <> 0%Z) by
(rewrite Hz; intros EQ; rewrite EQ in *; romega).
specialize (Hgc NEQ).
- assert (NEQ' : N.to_Z dx <> 0%Z) by
+ assert (NEQ' : NN.to_Z dx <> 0%Z) by
(rewrite Hz'; intro EQ; rewrite EQ in *; romega).
specialize (Hgc' NEQ').
revert H H0.
rewrite 2 strong_spec_red, 2 Qred_iff; simpl.
destr_eqb; simpl; nzsimpl; try romega; intros.
- rewrite Z2P_correct in *; auto.
+ rewrite Z2Pos.id in *; auto.
- apply Zgcd_mult_rel_prime; rewrite Zgcd_comm;
- apply Zgcd_mult_rel_prime; rewrite Zgcd_comm; auto.
+ apply Zgcd_mult_rel_prime; rewrite Z.gcd_comm;
+ apply Zgcd_mult_rel_prime; rewrite Z.gcd_comm; auto.
rewrite Zgcd_1_rel_prime in *.
apply bezout_rel_prime.
- destruct (rel_prime_bezout (Z.to_Z ny) (N.to_Z dy)) as [u v Huv]; trivial.
+ destruct (rel_prime_bezout (ZZ.to_Z ny) (NN.to_Z dy)) as [u v Huv]; trivial.
apply Bezout_intro with (u*g')%Z (v*g)%Z.
rewrite <- Huv, <- Hg1', <- Hg2. ring.
rewrite Zgcd_1_rel_prime in *.
apply bezout_rel_prime.
- destruct (rel_prime_bezout (Z.to_Z nx) (N.to_Z dx)) as [u v Huv]; trivial.
+ destruct (rel_prime_bezout (ZZ.to_Z nx) (NN.to_Z dx)) as [u v Huv]; trivial.
apply Bezout_intro with (u*g)%Z (v*g')%Z.
rewrite <- Huv, <- Hg2', <- Hg1. ring.
Qed.
@@ -703,16 +703,16 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Definition inv (x: t): t :=
match x with
| Qz z =>
- match Z.compare Z.zero z with
+ match ZZ.compare ZZ.zero z with
| Eq => zero
- | Lt => Qq Z.one (Zabs_N z)
- | Gt => Qq Z.minus_one (Zabs_N z)
+ | Lt => Qq ZZ.one (Zabs_N z)
+ | Gt => Qq ZZ.minus_one (Zabs_N z)
end
| Qq n d =>
- match Z.compare Z.zero n with
+ match ZZ.compare ZZ.zero n with
| Eq => zero
| Lt => Qq (Z_of_N d) (Zabs_N n)
- | Gt => Qq (Z.opp (Z_of_N d)) (Zabs_N n)
+ | Gt => Qq (ZZ.opp (Z_of_N d)) (Zabs_N n)
end
end.
@@ -721,29 +721,29 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destruct x as [ z | n d ].
(* Qz z *)
simpl.
- rewrite Z.spec_compare; destr_zcompare.
+ rewrite ZZ.spec_compare; destr_zcompare.
(* 0 = z *)
rewrite <- H.
simpl; nzsimpl; compute; auto.
(* 0 < z *)
simpl.
- destr_eqb; nzsimpl; [ intros; rewrite Zabs_eq in *; romega | intros _ ].
- set (z':=Z.to_Z z) in *; clearbody z'.
+ destr_eqb; nzsimpl; [ intros; rewrite Z.abs_eq in *; romega | intros _ ].
+ set (z':=ZZ.to_Z z) in *; clearbody z'.
red; simpl.
- rewrite Zabs_eq by romega.
- rewrite Z2P_correct by auto.
+ rewrite Z.abs_eq by romega.
+ rewrite Z2Pos.id by auto.
unfold Qinv; simpl; destruct z'; simpl; auto; discriminate.
(* 0 > z *)
simpl.
- destr_eqb; nzsimpl; [ intros; rewrite Zabs_non_eq in *; romega | intros _ ].
- set (z':=Z.to_Z z) in *; clearbody z'.
+ destr_eqb; nzsimpl; [ intros; rewrite Z.abs_neq in *; romega | intros _ ].
+ set (z':=ZZ.to_Z z) in *; clearbody z'.
red; simpl.
- rewrite Zabs_non_eq by romega.
- rewrite Z2P_correct by romega.
+ rewrite Z.abs_neq by romega.
+ rewrite Z2Pos.id by romega.
unfold Qinv; simpl; destruct z'; simpl; auto; discriminate.
(* Qq n d *)
simpl.
- rewrite Z.spec_compare; destr_zcompare.
+ rewrite ZZ.spec_compare; destr_zcompare.
(* 0 = n *)
rewrite <- H.
simpl; nzsimpl.
@@ -751,51 +751,51 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(* 0 < n *)
simpl.
destr_eqb; nzsimpl; intros.
- intros; rewrite Zabs_eq in *; romega.
- intros; rewrite Zabs_eq in *; romega.
+ intros; rewrite Z.abs_eq in *; romega.
+ intros; rewrite Z.abs_eq in *; romega.
nsubst; compute; auto.
- set (n':=Z.to_Z n) in *; clearbody n'.
- rewrite Zabs_eq by romega.
+ set (n':=ZZ.to_Z n) in *; clearbody n'.
+ rewrite Z.abs_eq by romega.
red; simpl.
- rewrite Z2P_correct by auto.
+ rewrite Z2Pos.id by auto.
unfold Qinv; simpl; destruct n'; simpl; auto; try discriminate.
- rewrite Zpos_mult_morphism, Z2P_correct; auto.
+ rewrite Pos2Z.inj_mul, Z2Pos.id; auto.
(* 0 > n *)
simpl.
destr_eqb; nzsimpl; intros.
- intros; rewrite Zabs_non_eq in *; romega.
- intros; rewrite Zabs_non_eq in *; romega.
+ intros; rewrite Z.abs_neq in *; romega.
+ intros; rewrite Z.abs_neq in *; romega.
nsubst; compute; auto.
- set (n':=Z.to_Z n) in *; clearbody n'.
+ set (n':=ZZ.to_Z n) in *; clearbody n'.
red; simpl; nzsimpl.
- rewrite Zabs_non_eq by romega.
- rewrite Z2P_correct by romega.
+ rewrite Z.abs_neq by romega.
+ rewrite Z2Pos.id by romega.
unfold Qinv; simpl; destruct n'; simpl; auto; try discriminate.
- assert (T : forall x, Zneg x = Zopp (Zpos x)) by auto.
- rewrite T, Zpos_mult_morphism, Z2P_correct; auto; ring.
+ assert (T : forall x, Zneg x = Z.opp (Zpos x)) by auto.
+ rewrite T, Pos2Z.inj_mul, Z2Pos.id; auto; ring.
Qed.
Definition inv_norm (x: t): t :=
match x with
| Qz z =>
- match Z.compare Z.zero z with
+ match ZZ.compare ZZ.zero z with
| Eq => zero
- | Lt => Qq Z.one (Zabs_N z)
- | Gt => Qq Z.minus_one (Zabs_N z)
+ | Lt => Qq ZZ.one (Zabs_N z)
+ | Gt => Qq ZZ.minus_one (Zabs_N z)
end
| Qq n d =>
- if N.eqb d N.zero then zero else
- match Z.compare Z.zero n with
+ if NN.eqb d NN.zero then zero else
+ match ZZ.compare ZZ.zero n with
| Eq => zero
| Lt =>
- match Z.compare n Z.one with
+ match ZZ.compare n ZZ.one with
| Gt => Qq (Z_of_N d) (Zabs_N n)
| _ => Qz (Z_of_N d)
end
| Gt =>
- match Z.compare n Z.minus_one with
- | Lt => Qq (Z.opp (Z_of_N d)) (Zabs_N n)
- | _ => Qz (Z.opp (Z_of_N d))
+ match ZZ.compare n ZZ.minus_one with
+ | Lt => Qq (ZZ.opp (Z_of_N d)) (Zabs_N n)
+ | _ => Qz (ZZ.opp (Z_of_N d))
end
end
end.
@@ -807,7 +807,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destruct x as [ z | n d ].
(* Qz z *)
simpl.
- rewrite Z.spec_compare; destr_zcompare; auto with qarith.
+ rewrite ZZ.spec_compare; destr_zcompare; auto with qarith.
(* Qq n d *)
simpl; nzsimpl; destr_eqb.
destr_zcompare; simpl; auto with qarith.
@@ -818,12 +818,12 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(* 0 < n *)
destr_zcompare; auto with qarith.
destr_zcompare; nzsimpl; simpl; auto with qarith; intros.
- destr_eqb; nzsimpl; [ intros; rewrite Zabs_eq in *; romega | intros _ ].
+ destr_eqb; nzsimpl; [ intros; rewrite Z.abs_eq in *; romega | intros _ ].
rewrite H0; auto with qarith.
romega.
(* 0 > n *)
destr_zcompare; nzsimpl; simpl; auto with qarith.
- destr_eqb; nzsimpl; [ intros; rewrite Zabs_non_eq in *; romega | intros _ ].
+ destr_eqb; nzsimpl; [ intros; rewrite Z.abs_neq in *; romega | intros _ ].
rewrite H0; auto with qarith.
romega.
Qed.
@@ -847,36 +847,36 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(* 0 < n *)
destr_zcompare; simpl; nzsimpl; auto.
destr_eqb; nzsimpl; simpl; auto.
- rewrite Zabs_eq; romega.
+ rewrite Z.abs_eq; romega.
intros _.
rewrite strong_spec_norm; simpl; nzsimpl.
destr_eqb; nzsimpl.
- rewrite Zabs_eq; romega.
+ rewrite Z.abs_eq; romega.
intros _.
rewrite Qred_iff.
simpl.
- rewrite Zabs_eq; auto with zarith.
- rewrite Z2P_correct in *; auto.
- rewrite Zgcd_comm; auto.
+ rewrite Z.abs_eq; auto with zarith.
+ rewrite Z2Pos.id in *; auto.
+ rewrite Z.gcd_comm; auto.
(* 0 > n *)
destr_eqb; nzsimpl; simpl; auto; intros.
destr_zcompare; simpl; nzsimpl; auto.
destr_eqb; nzsimpl.
- rewrite Zabs_non_eq; romega.
+ rewrite Z.abs_neq; romega.
intros _.
rewrite strong_spec_norm; simpl; nzsimpl.
destr_eqb; nzsimpl.
- rewrite Zabs_non_eq; romega.
+ rewrite Z.abs_neq; romega.
intros _.
rewrite Qred_iff.
simpl.
- rewrite Z2P_correct in *; auto.
+ rewrite Z2Pos.id in *; auto.
intros.
- rewrite Zgcd_comm, Zgcd_Zabs, Zgcd_comm.
+ rewrite Z.gcd_comm, Z.gcd_abs_l, Z.gcd_comm.
apply Zis_gcd_gcd; auto with zarith.
apply Zis_gcd_minus.
- rewrite Zopp_involutive, <- H1; apply Zgcd_is_gcd.
- rewrite Zabs_non_eq; romega.
+ rewrite Z.opp_involutive, <- H1; apply Zgcd_is_gcd.
+ rewrite Z.abs_neq; romega.
Qed.
Definition div x y := mul x (inv y).
@@ -909,31 +909,30 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Definition square (x: t): t :=
match x with
- | Qz zx => Qz (Z.square zx)
- | Qq nx dx => Qq (Z.square nx) (N.square dx)
+ | Qz zx => Qz (ZZ.square zx)
+ | Qq nx dx => Qq (ZZ.square nx) (NN.square dx)
end.
Theorem spec_square : forall x, [square x] == [x] ^ 2.
Proof.
destruct x as [ z | n d ].
- simpl; rewrite Z.spec_square; red; auto.
+ simpl; rewrite ZZ.spec_square; red; auto.
simpl.
destr_eqb; nzsimpl; intros.
apply Qeq_refl.
- rewrite N.spec_square in *; nzsimpl.
- match goal with E : (_ * _ = 0)%Z |- _ =>
- elim (Zmult_integral _ _ E); romega end.
- rewrite N.spec_square in *; nzsimpl; nsubst; romega.
- rewrite Z.spec_square, N.spec_square.
+ rewrite NN.spec_square in *; nzsimpl.
+ rewrite Z.mul_eq_0 in *; romega.
+ rewrite NN.spec_square in *; nzsimpl; nsubst; romega.
+ rewrite ZZ.spec_square, NN.spec_square.
red; simpl.
- rewrite Zpos_mult_morphism; rewrite !Z2P_correct; auto.
- apply Zmult_lt_0_compat; auto.
+ rewrite Pos2Z.inj_mul; rewrite !Z2Pos.id; auto.
+ apply Z.mul_pos_pos; auto.
Qed.
Definition power_pos (x : t) p : t :=
match x with
- | Qz zx => Qz (Z.pow_pos zx p)
- | Qq nx dx => Qq (Z.pow_pos nx p) (N.pow_pos dx p)
+ | Qz zx => Qz (ZZ.pow_pos zx p)
+ | Qq nx dx => Qq (ZZ.pow_pos nx p) (NN.pow_pos dx p)
end.
Theorem spec_power_pos : forall x p, [power_pos x p] == [x] ^ Zpos p.
@@ -941,26 +940,26 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
intros [ z | n d ] p; unfold power_pos.
(* Qz *)
simpl.
- rewrite Z.spec_pow_pos.
- rewrite Qpower_decomp.
+ rewrite ZZ.spec_pow_pos, Qpower_decomp.
red; simpl; f_equal.
- rewrite Zpower_pos_1_l; auto.
+ now rewrite Pos2Z.inj_pow, Z.pow_1_l.
(* Qq *)
simpl.
- rewrite Z.spec_pow_pos.
+ rewrite ZZ.spec_pow_pos.
destr_eqb; nzsimpl; intros.
- apply Qeq_sym; apply Qpower_positive_0.
- rewrite N.spec_pow_pos in *.
- assert (0 < N.to_Z d ^ ' p)%Z by
- (apply Zpower_gt_0; auto with zarith).
- romega.
- exfalso.
- rewrite N.spec_pow_pos in *. nsubst.
- rewrite Zpower_0_l in *; [romega|discriminate].
- rewrite Qpower_decomp.
- red; simpl; do 3 f_equal.
- rewrite Z2P_correct by (generalize (N.spec_pos d); romega).
- rewrite N.spec_pow_pos. auto.
+ - apply Qeq_sym; apply Qpower_positive_0.
+ - rewrite NN.spec_pow_pos in *.
+ assert (0 < NN.to_Z d ^ ' p)%Z by
+ (apply Z.pow_pos_nonneg; auto with zarith).
+ romega.
+ - exfalso.
+ rewrite NN.spec_pow_pos in *. nsubst.
+ rewrite Z.pow_0_l' in *; [romega|discriminate].
+ - rewrite Qpower_decomp.
+ red; simpl; do 3 f_equal.
+ apply Pos2Z.inj. rewrite Pos2Z.inj_pow.
+ rewrite 2 Z2Pos.id by (generalize (NN.spec_pos d); romega).
+ now rewrite NN.spec_pow_pos.
Qed.
Instance strong_spec_power_pos x p `(Reduced x) : Reduced (power_pos x p).
@@ -976,10 +975,10 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
unfold Reduced; rewrite strong_spec_red, Qred_iff; simpl.
destr_eqb; nzsimpl; simpl; intros.
exfalso.
- rewrite N.spec_pow_pos in *. nsubst.
- rewrite Zpower_0_l in *; [romega|discriminate].
- rewrite Z2P_correct in *; auto.
- rewrite N.spec_pow_pos, Z.spec_pow_pos; auto.
+ rewrite NN.spec_pow_pos in *. nsubst.
+ rewrite Z.pow_0_l' in *; [romega|discriminate].
+ rewrite Z2Pos.id in *; auto.
+ rewrite NN.spec_pow_pos, ZZ.spec_pow_pos; auto.
rewrite Zgcd_1_rel_prime in *.
apply rel_prime_Zpower; auto with zarith.
Qed.
@@ -1086,7 +1085,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
[[add x y]] = [[x]] + [[y]].
Proof.
unfold to_Qc.
- apply trans_equal with (!! ([x] + [y])).
+ transitivity (!! ([x] + [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_add; auto.
@@ -1100,7 +1099,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
[[add_norm x y]] = [[x]] + [[y]].
Proof.
unfold to_Qc.
- apply trans_equal with (!! ([x] + [y])).
+ transitivity (!! ([x] + [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_add_norm; auto.
@@ -1148,7 +1147,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
[[mul x y]] = [[x]] * [[y]].
Proof.
unfold to_Qc.
- apply trans_equal with (!! ([x] * [y])).
+ transitivity (!! ([x] * [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_mul; auto.
@@ -1162,7 +1161,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
[[mul_norm x y]] = [[x]] * [[y]].
Proof.
unfold to_Qc.
- apply trans_equal with (!! ([x] * [y])).
+ transitivity (!! ([x] * [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_mul_norm; auto.
@@ -1186,7 +1185,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
[[inv x]] = /[[x]].
Proof.
unfold to_Qc.
- apply trans_equal with (!! (/[x])).
+ transitivity (!! (/[x])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_inv; auto.
@@ -1200,7 +1199,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
[[inv_norm x]] = /[[x]].
Proof.
unfold to_Qc.
- apply trans_equal with (!! (/[x])).
+ transitivity (!! (/[x])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_inv_norm; auto.
@@ -1248,7 +1247,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_squarec x: [[square x]] = [[x]]^2.
Proof.
unfold to_Qc.
- apply trans_equal with (!! ([x]^2)).
+ transitivity (!! ([x]^2)).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_square; auto.
@@ -1262,24 +1261,24 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Qed.
Theorem spec_power_posc x p:
- [[power_pos x p]] = [[x]] ^ nat_of_P p.
+ [[power_pos x p]] = [[x]] ^ Pos.to_nat p.
Proof.
unfold to_Qc.
- apply trans_equal with (!! ([x]^Zpos p)).
+ transitivity (!! ([x]^Zpos p)).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete; apply spec_power_pos; auto.
- induction p using Pind.
+ induction p using Pos.peano_ind.
simpl; ring.
- rewrite Psucc_S; simpl Qcpower.
+ rewrite Pos2Nat.inj_succ; simpl Qcpower.
rewrite <- IHp; clear IHp.
unfold Qcmult, Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
apply Qred_complete.
- setoid_replace ([x] ^ ' Psucc p)%Q with ([x] * [x] ^ ' p)%Q.
+ setoid_replace ([x] ^ ' Pos.succ p)%Q with ([x] * [x] ^ ' p)%Q.
apply Qmult_comp; apply Qeq_sym; apply Qred_correct.
simpl.
- rewrite Pplus_one_succ_l.
+ rewrite <- Pos.add_1_l.
rewrite Qpower_plus_positive; simpl; apply Qeq_refl.
Qed.
diff --git a/theories/PArith/BinPos.v b/theories/PArith/BinPos.v
index e22cac7be..58fe090dd 100644
--- a/theories/PArith/BinPos.v
+++ b/theories/PArith/BinPos.v
@@ -1871,6 +1871,10 @@ Notation xI := xI (only parsing).
Notation xO := xO (only parsing).
Notation xH := xH (only parsing).
+Notation IsNul := Pos.IsNul (only parsing).
+Notation IsPos := Pos.IsPos (only parsing).
+Notation IsNeg := Pos.IsNeg (only parsing).
+
Notation Psucc := Pos.succ (compat "8.3").
Notation Pplus := Pos.add (compat "8.3").
Notation Pplus_carry := Pos.add_carry (compat "8.3").
@@ -1882,9 +1886,6 @@ Notation nat_of_P := Pos.to_nat (compat "8.3").
Notation P_of_succ_nat := Pos.of_succ_nat (compat "8.3").
Notation Pdouble_minus_one := Pos.pred_double (compat "8.3").
Notation positive_mask := Pos.mask (compat "8.3").
-Notation IsNul := Pos.IsNul (compat "8.3").
-Notation IsPos := Pos.IsPos (compat "8.3").
-Notation IsNeg := Pos.IsNeg (compat "8.3").
Notation positive_mask_rect := Pos.mask_rect (compat "8.3").
Notation positive_mask_ind := Pos.mask_ind (compat "8.3").
Notation positive_mask_rec := Pos.mask_rec (compat "8.3").
@@ -2056,24 +2057,24 @@ Lemma Peqb_true_eq x y : Pos.eqb x y = true -> x=y.
Proof. apply Pos.eqb_eq. Qed.
Lemma Pcompare_eq_Gt p q : (p ?= q) = Gt <-> p > q.
Proof. reflexivity. Qed.
-Lemma Pplus_one_succ_r p : Psucc p = p + 1.
+Lemma Pplus_one_succ_r p : Pos.succ p = p + 1.
Proof (eq_sym (Pos.add_1_r p)).
-Lemma Pplus_one_succ_l p : Psucc p = 1 + p.
+Lemma Pplus_one_succ_l p : Pos.succ p = 1 + p.
Proof (eq_sym (Pos.add_1_l p)).
-Lemma Pcompare_refl p : Pcompare p p Eq = Eq.
+Lemma Pcompare_refl p : Pos.compare_cont p p Eq = Eq.
Proof (Pos.compare_cont_refl p Eq).
-Lemma Pcompare_Eq_eq : forall p q, Pcompare p q Eq = Eq -> p = q.
+Lemma Pcompare_Eq_eq : forall p q, Pos.compare_cont p q Eq = Eq -> p = q.
Proof Pos.compare_eq.
-Lemma ZC4 p q : Pcompare p q Eq = CompOpp (Pcompare q p Eq).
+Lemma ZC4 p q : Pos.compare_cont p q Eq = CompOpp (Pos.compare_cont q p Eq).
Proof (Pos.compare_antisym q p).
-Lemma Ppred_minus p : Ppred p = p - 1.
+Lemma Ppred_minus p : Pos.pred p = p - 1.
Proof (eq_sym (Pos.sub_1_r p)).
Lemma Pminus_mask_Gt p q :
p > q ->
exists h : positive,
- Pminus_mask p q = IsPos h /\
- q + h = p /\ (h = 1 \/ Pminus_mask_carry p q = IsPos (Ppred h)).
+ Pos.sub_mask p q = IsPos h /\
+ q + h = p /\ (h = 1 \/ Pos.sub_mask_carry p q = IsPos (Pos.pred h)).
Proof.
intros H. apply Pos.gt_lt in H.
destruct (Pos.sub_mask_pos p q H) as (r & U).
diff --git a/theories/PArith/Pnat.v b/theories/PArith/Pnat.v
index 0fa6972b7..e9c77b05d 100644
--- a/theories/PArith/Pnat.v
+++ b/theories/PArith/Pnat.v
@@ -30,7 +30,7 @@ Qed.
Theorem inj_add p q : to_nat (p + q) = to_nat p + to_nat q.
Proof.
- revert q. induction p using Pind; intros q.
+ revert q. induction p using peano_ind; intros q.
now rewrite add_1_l, inj_succ.
now rewrite add_succ_l, !inj_succ, IHp.
Qed.
@@ -410,23 +410,24 @@ Notation P_of_succ_nat_o_nat_of_P_eq_succ := Pos2SuccNat.id_succ (compat "8.3").
Notation pred_o_P_of_succ_nat_o_nat_of_P_eq_id := Pos2SuccNat.pred_id (compat "8.3").
Lemma nat_of_P_minus_morphism p q :
- Pcompare p q Eq = Gt -> Pos.to_nat (p - q) = Pos.to_nat p - Pos.to_nat q.
-Proof (fun H => Pos2Nat.inj_sub p q (ZC1 _ _ H)).
+ Pos.compare_cont p q Eq = Gt ->
+ Pos.to_nat (p - q) = Pos.to_nat p - Pos.to_nat q.
+Proof (fun H => Pos2Nat.inj_sub p q (Pos.gt_lt _ _ H)).
Lemma nat_of_P_lt_Lt_compare_morphism p q :
- Pcompare p q Eq = Lt -> Pos.to_nat p < Pos.to_nat q.
+ Pos.compare_cont p q Eq = Lt -> Pos.to_nat p < Pos.to_nat q.
Proof (proj1 (Pos2Nat.inj_lt p q)).
Lemma nat_of_P_gt_Gt_compare_morphism p q :
- Pcompare p q Eq = Gt -> Pos.to_nat p > Pos.to_nat q.
+ Pos.compare_cont p q Eq = Gt -> Pos.to_nat p > Pos.to_nat q.
Proof (proj1 (Pos2Nat.inj_gt p q)).
Lemma nat_of_P_lt_Lt_compare_complement_morphism p q :
- Pos.to_nat p < Pos.to_nat q -> Pcompare p q Eq = Lt.
+ Pos.to_nat p < Pos.to_nat q -> Pos.compare_cont p q Eq = Lt.
Proof (proj2 (Pos2Nat.inj_lt p q)).
Definition nat_of_P_gt_Gt_compare_complement_morphism p q :
- Pos.to_nat p > Pos.to_nat q -> Pcompare p q Eq = Gt.
+ Pos.to_nat p > Pos.to_nat q -> Pos.compare_cont p q Eq = Gt.
Proof (proj2 (Pos2Nat.inj_gt p q)).
(** Old intermediate results about [Pmult_nat] *)
@@ -445,7 +446,7 @@ Proof.
Qed.
Lemma Pmult_nat_succ_morphism :
- forall p n, Pmult_nat (Psucc p) n = n + Pmult_nat p n.
+ forall p n, Pmult_nat (Pos.succ p) n = n + Pmult_nat p n.
Proof.
intros. now rewrite !Pmult_nat_mult, Pos2Nat.inj_succ.
Qed.
@@ -457,7 +458,7 @@ Proof.
Qed.
Theorem Pmult_nat_plus_carry_morphism :
- forall p q n, Pmult_nat (Pplus_carry p q) n = n + Pmult_nat (p + q) n.
+ forall p q n, Pmult_nat (Pos.add_carry p q) n = n + Pmult_nat (p + q) n.
Proof.
intros. now rewrite Pos.add_carry_spec, Pmult_nat_succ_morphism.
Qed.
diff --git a/theories/Program/Subset.v b/theories/Program/Subset.v
index dda0ed68d..4642021c4 100644
--- a/theories/Program/Subset.v
+++ b/theories/Program/Subset.v
@@ -106,7 +106,7 @@ Ltac rewrite_match_eq H :=
[ |- ?T ] =>
match T with
context [ match_eq ?A ?B ?t ?f ] =>
- rewrite (match_eq_rewrite A B t f (exist _ _ (sym_eq H)))
+ rewrite (match_eq_rewrite A B t f (exist _ _ (eq_sym H)))
end
end.
diff --git a/theories/QArith/QArith_base.v b/theories/QArith/QArith_base.v
index 94ea4906c..7f70d6efd 100644
--- a/theories/QArith/QArith_base.v
+++ b/theories/QArith/QArith_base.v
@@ -20,7 +20,7 @@ Delimit Scope Q_scope with Q.
Bind Scope Q_scope with Q.
Arguments Qmake _%Z _%positive.
Open Scope Q_scope.
-Ltac simpl_mult := repeat rewrite Zpos_mult_morphism.
+Ltac simpl_mult := rewrite ?Pos2Z.inj_mul.
(** [a#b] denotes the fraction [a] over [b]. *)
@@ -58,79 +58,65 @@ Qed.
Definition Qcompare (p q : Q) := (Qnum p * QDen q ?= Qnum q * QDen p)%Z.
Notation "p ?= q" := (Qcompare p q) : Q_scope.
-Lemma Qeq_alt : forall p q, (p == q) <-> (p ?= q) = Eq.
+Lemma Qeq_alt p q : (p == q) <-> (p ?= q) = Eq.
Proof.
-unfold Qeq, Qcompare; intros; split; intros.
-rewrite H; apply Zcompare_refl.
-apply Zcompare_Eq_eq; auto.
+symmetry. apply Z.compare_eq_iff.
Qed.
-Lemma Qlt_alt : forall p q, (p<q) <-> (p?=q = Lt).
+Lemma Qlt_alt p q : (p<q) <-> (p?=q = Lt).
Proof.
-unfold Qlt, Qcompare, Zlt; split; auto.
+reflexivity.
Qed.
-Lemma Qgt_alt : forall p q, (p>q) <-> (p?=q = Gt).
+Lemma Qgt_alt p q : (p>q) <-> (p?=q = Gt).
Proof.
-unfold Qlt, Qcompare, Zlt.
-intros; rewrite Zcompare_Gt_Lt_antisym; split; auto.
+symmetry. apply Z.gt_lt_iff.
Qed.
-Lemma Qle_alt : forall p q, (p<=q) <-> (p?=q <> Gt).
+Lemma Qle_alt p q : (p<=q) <-> (p?=q <> Gt).
Proof.
-unfold Qle, Qcompare, Zle; split; auto.
+reflexivity.
Qed.
-Lemma Qge_alt : forall p q, (p>=q) <-> (p?=q <> Lt).
+Lemma Qge_alt p q : (p>=q) <-> (p?=q <> Lt).
Proof.
-unfold Qle, Qcompare, Zle.
-split; intros; contradict H.
-rewrite Zcompare_Gt_Lt_antisym; auto.
-rewrite Zcompare_Gt_Lt_antisym in H; auto.
+symmetry. apply Z.ge_le_iff.
Qed.
Hint Unfold Qeq Qlt Qle : qarith.
Hint Extern 5 (?X1 <> ?X2) => intro; discriminate: qarith.
-Lemma Qcompare_antisym : forall x y, CompOpp (x ?= y) = (y ?= x).
+Lemma Qcompare_antisym x y : CompOpp (x ?= y) = (y ?= x).
Proof.
- unfold "?=". intros. apply Zcompare_antisym.
+ symmetry. apply Z.compare_antisym.
Qed.
-Lemma Qcompare_spec : forall x y, CompareSpec (x==y) (x<y) (y<x) (x ?= y).
+Lemma Qcompare_spec x y : CompareSpec (x==y) (x<y) (y<x) (x ?= y).
Proof.
- intros.
- destruct (x ?= y) as [ ]_eqn:H; constructor; auto.
- rewrite Qeq_alt; auto.
- rewrite Qlt_alt, <- Qcompare_antisym, H; auto.
+ unfold Qeq, Qlt, Qcompare. case Z.compare_spec; now constructor.
Qed.
(** * Properties of equality. *)
-Theorem Qeq_refl : forall x, x == x.
+Theorem Qeq_refl x : x == x.
Proof.
auto with qarith.
Qed.
-Theorem Qeq_sym : forall x y, x == y -> y == x.
+Theorem Qeq_sym x y : x == y -> y == x.
Proof.
auto with qarith.
Qed.
-Theorem Qeq_trans : forall x y z, x == y -> y == z -> x == z.
+Theorem Qeq_trans x y z : x == y -> y == z -> x == z.
Proof.
-unfold Qeq; intros.
-apply Zmult_reg_l with (QDen y).
-auto with qarith.
-transitivity (Qnum x * QDen y * QDen z)%Z; try ring.
-rewrite H.
-transitivity (Qnum y * QDen z * QDen x)%Z; try ring.
-rewrite H0; ring.
+unfold Qeq; intros XY YZ.
+apply Z.mul_reg_r with (QDen y); [auto with qarith|].
+now rewrite Z.mul_shuffle0, XY, Z.mul_shuffle0, YZ, Z.mul_shuffle0.
Qed.
-Hint Resolve Qeq_refl : qarith.
-Hint Resolve Qeq_sym : qarith.
-Hint Resolve Qeq_trans : qarith.
+Hint Immediate Qeq_sym : qarith.
+Hint Resolve Qeq_refl Qeq_trans : qarith.
(** In a word, [Qeq] is a setoid equality. *)
@@ -139,50 +125,48 @@ Proof. split; red; eauto with qarith. Qed.
(** Furthermore, this equality is decidable: *)
-Theorem Qeq_dec : forall x y, {x==y} + {~ x==y}.
+Theorem Qeq_dec x y : {x==y} + {~ x==y}.
Proof.
- intros; case (Z_eq_dec (Qnum x * QDen y) (Qnum y * QDen x)); auto.
+ apply Z.eq_dec.
Defined.
Definition Qeq_bool x y :=
(Zeq_bool (Qnum x * QDen y) (Qnum y * QDen x))%Z.
Definition Qle_bool x y :=
- (Zle_bool (Qnum x * QDen y) (Qnum y * QDen x))%Z.
+ (Z.leb (Qnum x * QDen y) (Qnum y * QDen x))%Z.
-Lemma Qeq_bool_iff : forall x y, Qeq_bool x y = true <-> x == y.
+Lemma Qeq_bool_iff x y : Qeq_bool x y = true <-> x == y.
Proof.
- unfold Qeq_bool, Qeq; intros.
symmetry; apply Zeq_is_eq_bool.
Qed.
-Lemma Qeq_bool_eq : forall x y, Qeq_bool x y = true -> x == y.
+Lemma Qeq_bool_eq x y : Qeq_bool x y = true -> x == y.
Proof.
- intros; rewrite <- Qeq_bool_iff; auto.
+ apply Qeq_bool_iff.
Qed.
-Lemma Qeq_eq_bool : forall x y, x == y -> Qeq_bool x y = true.
+Lemma Qeq_eq_bool x y : x == y -> Qeq_bool x y = true.
Proof.
- intros; rewrite Qeq_bool_iff; auto.
+ apply Qeq_bool_iff.
Qed.
-Lemma Qeq_bool_neq : forall x y, Qeq_bool x y = false -> ~ x == y.
+Lemma Qeq_bool_neq x y : Qeq_bool x y = false -> ~ x == y.
Proof.
- intros x y H; rewrite <- Qeq_bool_iff, H; discriminate.
+ rewrite <- Qeq_bool_iff. now intros ->.
Qed.
-Lemma Qle_bool_iff : forall x y, Qle_bool x y = true <-> x <= y.
+Lemma Qle_bool_iff x y : Qle_bool x y = true <-> x <= y.
Proof.
- unfold Qle_bool, Qle; intros.
symmetry; apply Zle_is_le_bool.
Qed.
-Lemma Qle_bool_imp_le : forall x y, Qle_bool x y = true -> x <= y.
+Lemma Qle_bool_imp_le x y : Qle_bool x y = true -> x <= y.
Proof.
- intros; rewrite <- Qle_bool_iff; auto.
+ apply Qle_bool_iff.
Qed.
-Theorem Qnot_eq_sym : forall x y, ~x == y -> ~y == x.
+Theorem Qnot_eq_sym x y : ~x == y -> ~y == x.
Proof.
auto with qarith.
Qed.
@@ -223,12 +207,9 @@ Infix "/" := Qdiv : Q_scope.
Notation " ' x " := (Zpos x) (at level 20, no associativity) : Z_scope.
-Lemma Qmake_Qdiv : forall a b, a#b==inject_Z a/inject_Z ('b).
+Lemma Qmake_Qdiv a b : a#b==inject_Z a/inject_Z ('b).
Proof.
-intros a b.
-unfold Qeq.
-simpl.
-ring.
+unfold Qeq. simpl. ring.
Qed.
(** * Setoid compatibility results *)
@@ -281,17 +262,13 @@ Instance Qinv_comp : Proper (Qeq==>Qeq) Qinv.
Proof.
unfold Qeq, Qinv; simpl.
Open Scope Z_scope.
- intros (p1, p2) (q1, q2); simpl.
- case p1; simpl.
- intros.
- assert (q1 = 0).
- elim (Zmult_integral q1 ('p2)); auto with zarith.
- intros; discriminate.
- subst; auto.
- case q1; simpl; intros; try discriminate.
- rewrite (Pmult_comm p2 p); rewrite (Pmult_comm q2 p0); auto.
- case q1; simpl; intros; try discriminate.
- rewrite (Pmult_comm p2 p); rewrite (Pmult_comm q2 p0); auto.
+ intros (p1, p2) (q1, q2) EQ; simpl in *.
+ destruct q1; simpl in *.
+ - apply Z.mul_eq_0 in EQ. destruct EQ; now subst.
+ - destruct p1; simpl in *; try discriminate.
+ now rewrite Pos.mul_comm, <- EQ, Pos.mul_comm.
+ - destruct p1; simpl in *; try discriminate.
+ now rewrite Pos.mul_comm, <- EQ, Pos.mul_comm.
Close Scope Z_scope.
Qed.
@@ -368,7 +345,7 @@ Qed.
Lemma Qplus_0_r : forall x, x+0 == x.
Proof.
intros (x1, x2); unfold Qeq, Qplus; simpl.
- rewrite Pmult_comm; simpl; ring.
+ rewrite Pos.mul_comm; simpl; ring.
Qed.
(** Commutativity of addition: *)
@@ -376,7 +353,7 @@ Qed.
Theorem Qplus_comm : forall x y, x+y == y+x.
Proof.
intros (x1, x2); unfold Qeq, Qplus; simpl.
- intros; rewrite Pmult_comm; ring.
+ intros; rewrite Pos.mul_comm; ring.
Qed.
@@ -419,7 +396,7 @@ Qed.
Theorem Qmult_assoc : forall n m p, n*(m*p)==(n*m)*p.
Proof.
- intros; red; simpl; rewrite Pmult_assoc; ring.
+ intros; red; simpl; rewrite Pos.mul_assoc; ring.
Qed.
(** multiplication and zero *)
@@ -444,15 +421,15 @@ Qed.
Theorem Qmult_1_r : forall n, n*1==n.
Proof.
intro; red; simpl.
- rewrite Zmult_1_r with (n := Qnum n).
- rewrite Pmult_comm; simpl; trivial.
+ rewrite Z.mul_1_r with (n := Qnum n).
+ rewrite Pos.mul_comm; simpl; trivial.
Qed.
(** Commutativity of multiplication *)
Theorem Qmult_comm : forall x y, x*y==y*x.
Proof.
- intros; red; simpl; rewrite Pmult_comm; ring.
+ intros; red; simpl; rewrite Pos.mul_comm; ring.
Qed.
(** Distributivity over [Qadd] *)
@@ -474,17 +451,15 @@ Qed.
Theorem Qmult_integral : forall x y, x*y==0 -> x==0 \/ y==0.
Proof.
intros (x1,x2) (y1,y2).
- unfold Qeq, Qmult; simpl; intros.
- destruct (Zmult_integral (x1*1)%Z (y1*1)%Z); auto.
- rewrite <- H; ring.
+ unfold Qeq, Qmult; simpl.
+ now rewrite <- Z.mul_eq_0, !Z.mul_1_r.
Qed.
Theorem Qmult_integral_l : forall x y, ~ x == 0 -> x*y == 0 -> y == 0.
Proof.
intros (x1, x2) (y1, y2).
- unfold Qeq, Qmult; simpl; intros.
- apply Zmult_integral_l with x1; auto with zarith.
- rewrite <- H0; ring.
+ unfold Qeq, Qmult; simpl.
+ rewrite !Z.mul_1_r, Z.mul_eq_0. intuition.
Qed.
@@ -561,12 +536,12 @@ Qed.
(** * Properties of order upon Q. *)
-Lemma Qle_refl : forall x, x<=x.
+Lemma Qle_refl x : x<=x.
Proof.
unfold Qle; auto with zarith.
Qed.
-Lemma Qle_antisym : forall x y, x<=y -> y<=x -> x==y.
+Lemma Qle_antisym x y : x<=y -> y<=x -> x==y.
Proof.
unfold Qle, Qeq; auto with zarith.
Qed.
@@ -575,52 +550,46 @@ 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 Zmult_le_reg_r with ('y2).
- red; trivial.
- apply Zle_trans with (y1 * 'x2 * 'z2).
- replace (x1 * 'z2 * 'y2) with (x1 * 'y2 * 'z2) by ring.
- apply Zmult_le_compat_r; auto with zarith.
- replace (y1 * 'x2 * 'z2) with (y1 * 'z2 * 'x2) by ring.
- replace (z1 * 'x2 * 'y2) with (z1 * 'y2 * 'x2) by ring.
- apply Zmult_le_compat_r; auto with zarith.
+ apply Z.mul_le_mono_pos_r with ('y2); [easy|].
+ apply Z.le_trans with (y1 * 'x2 * '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.
Close Scope Z_scope.
Qed.
Hint Resolve Qle_trans : qarith.
-Lemma Qlt_irrefl : forall x, ~x<x.
+Lemma Qlt_irrefl x : ~x<x.
Proof.
unfold Qlt. auto with zarith.
Qed.
-Lemma Qlt_not_eq : forall x y, x<y -> ~ x==y.
+Lemma Qlt_not_eq x y : x<y -> ~ x==y.
Proof.
unfold Qlt, Qeq; auto with zarith.
Qed.
Lemma Zle_Qle (x y: Z): (x <= y)%Z = (inject_Z x <= inject_Z y).
Proof.
- unfold Qle. intros. simpl.
- do 2 rewrite Zmult_1_r. reflexivity.
+ unfold Qle. simpl. now rewrite !Z.mul_1_r.
Qed.
Lemma Zlt_Qlt (x y: Z): (x < y)%Z = (inject_Z x < inject_Z y).
Proof.
- unfold Qlt. intros. simpl.
- do 2 rewrite Zmult_1_r. reflexivity.
+ unfold Qlt. simpl. now rewrite !Z.mul_1_r.
Qed.
(** Large = strict or equal *)
-Lemma Qle_lteq : forall x y, x<=y <-> x<y \/ x==y.
+Lemma Qle_lteq x y : x<=y <-> x<y \/ x==y.
Proof.
- intros.
rewrite Qeq_alt, Qle_alt, Qlt_alt.
destruct (x ?= y); intuition; discriminate.
Qed.
-Lemma Qlt_le_weak : forall x y, x<y -> x<=y.
+Lemma Qlt_le_weak x y : x<y -> x<=y.
Proof.
unfold Qle, Qlt; auto with zarith.
Qed.
@@ -629,15 +598,11 @@ 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 Zgt_lt.
- apply Zmult_gt_reg_r with ('y2).
- red; trivial.
- apply Zgt_le_trans with (y1 * 'x2 * 'z2).
- replace (y1 * 'x2 * 'z2) with (y1 * 'z2 * 'x2) by ring.
- replace (z1 * 'x2 * 'y2) with (z1 * 'y2 * 'x2) by ring.
- apply Zmult_gt_compat_r; auto with zarith.
- replace (x1 * 'z2 * 'y2) with (x1 * 'y2 * 'z2) by ring.
- apply Zmult_le_compat_r; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with ('y2); [easy|].
+ apply Z.le_lt_trans with (y1 * 'x2 * '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.
Close Scope Z_scope.
Qed.
@@ -645,15 +610,11 @@ 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 Zgt_lt.
- apply Zmult_gt_reg_r with ('y2).
- red; trivial.
- apply Zle_gt_trans with (y1 * 'x2 * 'z2).
- replace (y1 * 'x2 * 'z2) with (y1 * 'z2 * 'x2) by ring.
- replace (z1 * 'x2 * 'y2) with (z1 * 'y2 * 'x2) by ring.
- apply Zmult_le_compat_r; auto with zarith.
- replace (x1 * 'z2 * 'y2) with (x1 * 'y2 * 'z2) by ring.
- apply Zmult_gt_compat_r; auto with zarith.
+ apply Z.mul_lt_mono_pos_r with ('y2); [easy|].
+ apply Z.lt_le_trans with (y1 * 'x2 * '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.
Close Scope Z_scope.
Qed.
@@ -688,7 +649,7 @@ Qed.
Lemma Qle_lt_or_eq : forall x y, x<=y -> x<y \/ x==y.
Proof.
- unfold Qle, Qlt, Qeq; intros; apply Zle_lt_or_eq; auto.
+ unfold Qle, Qlt, Qeq; intros; now apply Z.lt_eq_cases.
Qed.
Hint Resolve Qle_not_lt Qlt_not_le Qnot_le_lt Qnot_lt_le
@@ -713,7 +674,7 @@ Defined.
Lemma Qopp_le_compat : forall p q, p<=q -> -q <= -p.
Proof.
intros (a1,a2) (b1,b2); unfold Qle, Qlt; simpl.
- do 2 rewrite <- Zopp_mult_distr_l; omega.
+ rewrite !Z.mul_opp_l. omega.
Qed.
Hint Resolve Qopp_le_compat : qarith.
@@ -721,15 +682,13 @@ Hint Resolve Qopp_le_compat : qarith.
Lemma Qle_minus_iff : forall p q, p <= q <-> 0 <= q+-p.
Proof.
intros (x1,x2) (y1,y2); unfold Qle; simpl.
- rewrite <- Zopp_mult_distr_l.
- split; omega.
+ rewrite Z.mul_opp_l. omega.
Qed.
Lemma Qlt_minus_iff : forall p q, p < q <-> 0 < q+-p.
Proof.
intros (x1,x2) (y1,y2); unfold Qlt; simpl.
- rewrite <- Zopp_mult_distr_l.
- split; omega.
+ rewrite Z.mul_opp_l. omega.
Qed.
Lemma Qplus_le_compat :
@@ -740,8 +699,8 @@ Proof.
Open Scope Z_scope.
intros.
match goal with |- ?a <= ?b => ring_simplify a b end.
- rewrite Zplus_comm.
- apply Zplus_le_compat.
+ rewrite Z.add_comm.
+ apply Z.add_le_mono.
match goal with |- ?a <= ?b => ring_simplify z1 t1 ('z2) ('t2) a b end.
auto with zarith.
match goal with |- ?a <= ?b => ring_simplify x1 y1 ('x2) ('y2) a b end.
@@ -757,13 +716,12 @@ Proof.
Open Scope Z_scope.
intros.
match goal with |- ?a < ?b => ring_simplify a b end.
- rewrite Zplus_comm.
- apply Zplus_le_lt_compat.
+ rewrite Z.add_comm.
+ apply Z.add_le_lt_mono.
match goal with |- ?a <= ?b => ring_simplify z1 t1 ('z2) ('t2) a b end.
auto with zarith.
match goal with |- ?a < ?b => ring_simplify x1 y1 ('x2) ('y2) a b end.
- assert (forall p, 0 < ' p) by reflexivity.
- repeat (apply Zmult_lt_compat_r; auto).
+ do 2 (apply Z.mul_lt_mono_pos_r;try easy).
Close Scope Z_scope.
Qed.
@@ -802,20 +760,20 @@ Proof.
intros (a1,a2) (b1,b2) (c1,c2); unfold Qle, Qlt; simpl.
Open Scope Z_scope.
intros; simpl_mult.
- replace (a1*c1*('b2*'c2)) with ((a1*'b2)*(c1*'c2)) by ring.
- replace (b1*c1*('a2*'c2)) with ((b1*'a2)*(c1*'c2)) by ring.
- apply Zmult_le_compat_r; auto with zarith.
+ rewrite Z.mul_shuffle1, (Z.mul_shuffle1 b1).
+ apply Z.mul_le_mono_nonneg_r; auto with zarith.
Close Scope Z_scope.
Qed.
-Lemma Qmult_lt_0_le_reg_r : forall x y z, 0 < z -> x*z <= y*z -> x <= y.
+Lemma Qmult_lt_0_le_reg_r : forall x y z, 0 < z -> x*z <= y*z -> x <= y.
Proof.
intros (a1,a2) (b1,b2) (c1,c2); unfold Qle, Qlt; simpl.
Open Scope Z_scope.
simpl_mult.
- replace (a1*c1*('b2*'c2)) with ((a1*'b2)*(c1*'c2)) by ring.
- replace (b1*c1*('a2*'c2)) with ((b1*'a2)*(c1*'c2)) by ring.
- intros; apply Zmult_le_reg_r with (c1*'c2); auto with zarith.
+ rewrite Z.mul_shuffle1, (Z.mul_shuffle1 b1).
+ intros LT LE.
+ apply Z.mul_le_mono_pos_r in LE; trivial.
+ apply Z.mul_pos_pos; [omega|easy].
Close Scope Z_scope.
Qed.
@@ -837,12 +795,9 @@ Proof.
intros (a1,a2) (b1,b2) (c1,c2); unfold Qle, Qlt; simpl.
Open Scope Z_scope.
intros; simpl_mult.
- replace (a1*c1*('b2*'c2)) with ((a1*'b2)*(c1*'c2)) by ring.
- replace (b1*c1*('a2*'c2)) with ((b1*'a2)*(c1*'c2)) by ring.
- apply Zmult_lt_compat_r; auto with zarith.
- apply Zmult_lt_0_compat.
- omega.
- compute; auto.
+ rewrite Z.mul_shuffle1, (Z.mul_shuffle1 b1).
+ apply Z.mul_lt_mono_pos_r; auto with zarith.
+ apply Z.mul_pos_pos; [omega|reflexivity].
Close Scope Z_scope.
Qed.
@@ -852,15 +807,9 @@ Proof.
intros (a1,a2) (b1,b2) (c1,c2).
unfold Qle, Qlt; simpl.
simpl_mult.
- replace (a1*c1*('b2*'c2)) with ((a1*'b2)*(c1*'c2)) by ring.
- replace (b1*c1*('a2*'c2)) with ((b1*'a2)*(c1*'c2)) by ring.
- assert (forall p, 0 < ' p) by reflexivity.
- split; intros.
- apply Zmult_lt_reg_r with (c1*'c2); auto with zarith.
- apply Zmult_lt_0_compat; auto with zarith.
- apply Zmult_lt_compat_r; auto with zarith.
- apply Zmult_lt_0_compat. omega.
- compute; auto.
+ rewrite Z.mul_shuffle1, (Z.mul_shuffle1 b1).
+ intro LT. rewrite <- Z.mul_lt_mono_pos_r. reflexivity.
+ apply Z.mul_pos_pos; [omega|reflexivity].
Close Scope Z_scope.
Qed.
diff --git a/theories/QArith/Qabs.v b/theories/QArith/Qabs.v
index 557fabc89..5ae668b05 100644
--- a/theories/QArith/Qabs.v
+++ b/theories/QArith/Qabs.v
@@ -11,7 +11,7 @@ Require Export Qreduction.
Hint Resolve Qlt_le_weak : qarith.
-Definition Qabs (x:Q) := let (n,d):=x in (Zabs n#d).
+Definition Qabs (x:Q) := let (n,d):=x in (Z.abs n#d).
Lemma Qabs_case : forall (x:Q) (P : Q -> Type), (0 <= x -> P x) -> (x <= 0 -> P (- x)) -> P (Qabs x).
Proof.
@@ -26,9 +26,9 @@ intros [xn xd] [yn yd] H.
simpl.
unfold Qeq in *.
simpl in *.
-change (' yd)%Z with (Zabs (' yd)).
-change (' xd)%Z with (Zabs (' xd)).
-repeat rewrite <- Zabs_Zmult.
+change (' yd)%Z with (Z.abs (' yd)).
+change (' xd)%Z with (Z.abs (' xd)).
+repeat rewrite <- Z.abs_mul.
congruence.
Qed.
@@ -61,7 +61,7 @@ auto.
apply (Qopp_le_compat x 0).
Qed.
-Lemma Zabs_Qabs : forall n d, (Zabs n#d)==Qabs (n#d).
+Lemma Zabs_Qabs : forall n d, (Z.abs n#d)==Qabs (n#d).
Proof.
intros [|n|n]; reflexivity.
Qed.
@@ -85,25 +85,25 @@ intros [xn xd] [yn yd].
unfold Qplus.
unfold Qle.
simpl.
-apply Zmult_le_compat_r;auto with *.
-change (' yd)%Z with (Zabs (' yd)).
-change (' xd)%Z with (Zabs (' xd)).
-repeat rewrite <- Zabs_Zmult.
-apply Zabs_triangle.
+apply Z.mul_le_mono_nonneg_r;auto with *.
+change (' yd)%Z with (Z.abs (' yd)).
+change (' xd)%Z with (Z.abs (' xd)).
+repeat rewrite <- Z.abs_mul.
+apply Z.abs_triangle.
Qed.
Lemma Qabs_Qmult : forall a b, Qabs (a*b) == (Qabs a)*(Qabs b).
Proof.
intros [an ad] [bn bd].
simpl.
-rewrite Zabs_Zmult.
+rewrite Z.abs_mul.
reflexivity.
Qed.
Lemma Qabs_Qminus x y: Qabs (x - y) = Qabs (y - x).
Proof.
unfold Qminus, Qopp. simpl.
- rewrite Pmult_comm, <- Zabs_Zopp.
+ rewrite Pos.mul_comm, <- Z.abs_opp.
do 2 f_equal. ring.
Qed.
diff --git a/theories/QArith/Qcanon.v b/theories/QArith/Qcanon.v
index fea2ba398..05a27cc43 100644
--- a/theories/QArith/Qcanon.v
+++ b/theories/QArith/Qcanon.v
@@ -22,39 +22,39 @@ Arguments Qcmake this%Q _.
Open Scope Qc_scope.
Lemma Qred_identity :
- forall q:Q, Zgcd (Qnum q) (QDen q) = 1%Z -> Qred q = q.
+ forall q:Q, Z.gcd (Qnum q) (QDen q) = 1%Z -> Qred q = q.
Proof.
unfold Qred; intros (a,b); simpl.
- generalize (Zggcd_gcd a ('b)) (Zggcd_correct_divisors a ('b)).
+ generalize (Z.ggcd_gcd a ('b)) (Z.ggcd_correct_divisors a ('b)).
intros.
rewrite H1 in H; clear H1.
- destruct (Zggcd a ('b)) as (g,(aa,bb)); simpl in *; subst.
+ destruct (Z.ggcd a ('b)) as (g,(aa,bb)); simpl in *; subst.
destruct H0.
- rewrite Zmult_1_l in H, H0.
+ rewrite Z.mul_1_l in H, H0.
subst; simpl; auto.
Qed.
Lemma Qred_identity2 :
- forall q:Q, Qred q = q -> Zgcd (Qnum q) (QDen q) = 1%Z.
+ forall q:Q, Qred q = q -> Z.gcd (Qnum q) (QDen q) = 1%Z.
Proof.
unfold Qred; intros (a,b); simpl.
- generalize (Zggcd_gcd a ('b)) (Zggcd_correct_divisors a ('b)) (Zgcd_is_pos a ('b)).
+ generalize (Z.ggcd_gcd a ('b)) (Z.ggcd_correct_divisors a ('b)) (Z.gcd_nonneg a ('b)).
intros.
rewrite <- H; rewrite <- H in H1; clear H.
- destruct (Zggcd a ('b)) as (g,(aa,bb)); simpl in *; subst.
+ destruct (Z.ggcd a ('b)) as (g,(aa,bb)); simpl in *; subst.
injection H2; intros; clear H2.
destruct H0.
clear H0 H3.
destruct g as [|g|g]; destruct bb as [|bb|bb]; simpl in *; try discriminate.
f_equal.
- apply Pmult_reg_r with bb.
+ apply Pos.mul_reg_r with bb.
injection H2; intros.
rewrite <- H0.
rewrite H; simpl; auto.
elim H1; auto.
Qed.
-Lemma Qred_iff : forall q:Q, Qred q = q <-> Zgcd (Qnum q) (QDen q) = 1%Z.
+Lemma Qred_iff : forall q:Q, Qred q = q <-> Z.gcd (Qnum q) (QDen q) = 1%Z.
Proof.
split; intros.
apply Qred_identity2; auto.
diff --git a/theories/QArith/Qfield.v b/theories/QArith/Qfield.v
index 5e27f3814..811bfb6df 100644
--- a/theories/QArith/Qfield.v
+++ b/theories/QArith/Qfield.v
@@ -38,7 +38,7 @@ Proof.
exact Hp.
Qed.
-Lemma Qpower_theory : power_theory 1 Qmult Qeq Z_of_N Qpower.
+Lemma Qpower_theory : power_theory 1 Qmult Qeq Z.of_N Qpower.
Proof.
constructor.
intros r [|n];
@@ -66,7 +66,7 @@ Ltac Qpow_tac t :=
match t with
| Z0 => N0
| Zpos ?n => Ncst (Npos n)
- | Z_of_N ?n => Ncst n
+ | Z.of_N ?n => Ncst n
| NtoZ ?n => Ncst n
| _ => NotConstant
end.
diff --git a/theories/QArith/Qpower.v b/theories/QArith/Qpower.v
index b05ee6495..700b4f3b2 100644
--- a/theories/QArith/Qpower.v
+++ b/theories/QArith/Qpower.v
@@ -101,10 +101,9 @@ Lemma Qpower_plus_positive : forall a n m, Qpower_positive a (n+m) == (Qpower_po
Proof.
intros a n m.
unfold Qpower_positive.
-apply pow_pos_Pplus.
+apply pow_pos_add.
apply Q_Setoid.
apply Qmult_comp.
-apply Qmult_comm.
apply Qmult_assoc.
Qed.
@@ -114,21 +113,18 @@ intros a [|n|n]; simpl; try reflexivity.
symmetry; apply Qinv_involutive.
Qed.
-Lemma Qpower_minus_positive : forall a (n m:positive), (Pcompare n m Eq=Gt)%positive -> Qpower_positive a (n-m)%positive == (Qpower_positive a n)/(Qpower_positive a m).
+Lemma Qpower_minus_positive : forall a (n m:positive),
+ (m < n)%positive ->
+ Qpower_positive a (n-m)%positive == (Qpower_positive a n)/(Qpower_positive a m).
Proof.
intros a n m H.
-destruct (Qeq_dec a 0).
- rewrite q.
- repeat rewrite Qpower_positive_0.
- reflexivity.
-rewrite <- (Qdiv_mult_l (Qpower_positive a (n - m)) (Qpower_positive a m)) by
- (apply Qpower_not_0_positive; assumption).
-apply Qdiv_comp;[|reflexivity].
-rewrite Qmult_comm.
-rewrite <- Qpower_plus_positive.
-rewrite Pplus_minus.
-reflexivity.
-assumption.
+destruct (Qeq_dec a 0) as [EQ|NEQ].
+- now rewrite EQ, !Qpower_positive_0.
+- rewrite <- (Qdiv_mult_l (Qpower_positive a (n - m)) (Qpower_positive a m)) by
+ (now apply Qpower_not_0_positive).
+ f_equiv.
+ rewrite <- Qpower_plus_positive.
+ now rewrite Pos.sub_add.
Qed.
Lemma Qpower_plus : forall a n m, ~a==0 -> a^(n+m) == a^n*a^m.
@@ -140,8 +136,6 @@ rewrite ?Z.pos_sub_spec;
case Pos.compare_spec; intros H0; simpl; subst;
try rewrite Qpower_minus_positive;
try (field; try split; apply Qpower_not_0_positive);
- try assumption;
- apply ZC2;
assumption.
Qed.
@@ -158,13 +152,14 @@ apply Qpower_plus.
assumption.
Qed.
-Lemma Qpower_mult_positive : forall a n m, Qpower_positive a (n*m) == Qpower_positive (Qpower_positive a n) m.
+Lemma Qpower_mult_positive : forall a n m,
+ Qpower_positive a (n*m) == Qpower_positive (Qpower_positive a n) m.
Proof.
intros a n m.
-induction n using Pind.
+induction n using Pos.peano_ind.
reflexivity.
-rewrite Pmult_Sn_m.
-rewrite Pplus_one_succ_l.
+rewrite Pos.mul_succ_l.
+rewrite <- Pos.add_1_l.
do 2 rewrite Qpower_plus_positive.
rewrite IHn.
rewrite Qmult_power_positive.
@@ -184,11 +179,11 @@ Qed.
Lemma Zpower_Qpower : forall (a n:Z), (0<=n)%Z -> inject_Z (a^n) == (inject_Z a)^n.
Proof.
intros a [|n|n] H;[reflexivity| |elim H; reflexivity].
-induction n using Pind.
+induction n using Pos.peano_ind.
replace (a^1)%Z with a by ring.
ring.
-rewrite Zpos_succ_morphism.
-unfold Zsucc.
+rewrite Pos2Z.inj_succ.
+unfold Z.succ.
rewrite Zpower_exp; auto with *; try discriminate.
rewrite Qpower_plus' by discriminate.
rewrite <- IHn by discriminate.
@@ -209,31 +204,20 @@ setoid_replace (0+ - a) with (-a) in A by ring.
apply Qmult_le_0_compat; assumption.
Qed.
-Theorem Qpower_decomp: forall p x y,
- Qpower_positive (x #y) p == x ^ Zpos p # (Z2P ((Zpos y) ^ Zpos p)).
-Proof.
-induction p; intros; unfold Qmult; simpl.
-(* xI *)
-rewrite IHp, xI_succ_xO, <-Pplus_diag, Pplus_one_succ_l.
-repeat rewrite Zpower_pos_is_exp.
-red; unfold Qmult, Qnum, Qden, Zpower.
-repeat rewrite Zpos_mult_morphism.
-repeat rewrite Z2P_correct.
-repeat rewrite Zpower_pos_1_r; ring.
-apply Zpower_pos_pos; red; auto.
-repeat apply Zmult_lt_0_compat; red; auto;
- apply Zpower_pos_pos; red; auto.
-(* xO *)
-rewrite IHp, <-Pplus_diag.
-repeat rewrite Zpower_pos_is_exp.
-red; unfold Qmult, Qnum, Qden, Zpower.
-repeat rewrite Zpos_mult_morphism.
-repeat rewrite Z2P_correct; try ring.
-apply Zpower_pos_pos; red; auto.
-repeat apply Zmult_lt_0_compat; auto;
- apply Zpower_pos_pos; red; auto.
-(* xO *)
-unfold Qmult; simpl.
-red; simpl; rewrite Zpower_pos_1_r;
- rewrite Zpos_mult_morphism; ring.
+Theorem Qpower_decomp p x y :
+ Qpower_positive (x#y) p = x ^ Zpos p # (y ^ p).
+Proof.
+induction p; intros; simpl Qpower_positive; rewrite ?IHp.
+- (* xI *)
+ unfold Qmult, Qnum, Qden. f_equal.
+ + now rewrite <- Z.pow_twice_r, <- Z.pow_succ_r.
+ + apply Pos2Z.inj; rewrite !Pos2Z.inj_mul, !Pos2Z.inj_pow.
+ now rewrite <- Z.pow_twice_r, <- Z.pow_succ_r.
+- (* xO *)
+ unfold Qmult, Qnum, Qden. f_equal.
+ + now rewrite <- Z.pow_twice_r.
+ + apply Pos2Z.inj; rewrite !Pos2Z.inj_mul, !Pos2Z.inj_pow.
+ now rewrite <- Z.pow_twice_r.
+- (* xO *)
+ now rewrite Z.pow_1_r, Pos.pow_1_r.
Qed.
diff --git a/theories/QArith/Qreduction.v b/theories/QArith/Qreduction.v
index e39eca0cb..73ebea492 100644
--- a/theories/QArith/Qreduction.v
+++ b/theories/QArith/Qreduction.v
@@ -11,46 +11,29 @@
Require Export QArith_base.
Require Import Znumtheory.
-(** First, a function that (tries to) build a positive back from a Z. *)
+Notation Z2P := Z.to_pos (compat "8.3").
+Notation Z2P_correct := Z2Pos.id (compat "8.3").
-Definition Z2P (z : Z) :=
- match z with
- | Z0 => 1%positive
- | Zpos p => p
- | Zneg p => p
- end.
-
-Lemma Z2P_correct : forall z : Z, (0 < z)%Z -> Zpos (Z2P z) = z.
-Proof.
- simple destruct z; simpl in |- *; auto; intros; discriminate.
-Qed.
-
-Lemma Z2P_correct2 : forall z : Z, 0%Z <> z -> Zpos (Z2P z) = Zabs z.
-Proof.
- simple destruct z; simpl in |- *; auto; intros; elim H; auto.
-Qed.
-
-(** Simplification of fractions using [Zgcd].
+(** 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 (Zggcd q1 ('q2))
- in r1#(Z2P r2).
+ let (r1,r2) := snd (Z.ggcd q1 ('q2))
+ in r1#(Z.to_pos r2).
Lemma Qred_correct : forall q, (Qred q) == q.
Proof.
unfold Qred, Qeq; intros (n,d); simpl.
- generalize (Zggcd_gcd n ('d)) (Zgcd_nonneg n ('d))
- (Zggcd_correct_divisors n ('d)).
- destruct (Zggcd n (Zpos d)) as (g,(nn,dd)); simpl.
+ generalize (Z.ggcd_gcd n ('d)) (Z.gcd_nonneg n ('d))
+ (Z.ggcd_correct_divisors n ('d)).
+ destruct (Z.ggcd n (Zpos d)) as (g,(nn,dd)); simpl.
Open Scope Z_scope.
intros Hg LE (Hn,Hd). rewrite Hd, Hn.
rewrite <- Hg in LE; clear Hg.
assert (0 <> g) by (intro; subst; discriminate).
- rewrite Z2P_correct. ring.
- apply Zmult_gt_0_lt_0_reg_r with g; auto with zarith.
- now rewrite Zmult_comm, <- Hd.
+ rewrite Z2Pos.id. ring.
+ rewrite <- (Z.mul_pos_cancel_l g); [now rewrite <- Hd | omega].
Close Scope Z_scope.
Qed.
@@ -59,68 +42,54 @@ Proof.
intros (a,b) (c,d).
unfold Qred, Qeq in *; simpl in *.
Open Scope Z_scope.
- generalize (Zggcd_gcd a ('b)) (Zgcd_is_gcd a ('b))
- (Zgcd_nonneg a ('b)) (Zggcd_correct_divisors a ('b)).
- destruct (Zggcd a (Zpos b)) as (g,(aa,bb)).
- generalize (Zggcd_gcd c ('d)) (Zgcd_is_gcd c ('d))
- (Zgcd_nonneg c ('d)) (Zggcd_correct_divisors c ('d)).
- destruct (Zggcd c (Zpos d)) as (g',(cc,dd)).
- simpl.
- intro H; rewrite <- H; clear H.
- intros Hg'1 Hg'2 (Hg'3,Hg'4).
- intro H; rewrite <- H; clear H.
- intros Hg1 Hg2 (Hg3,Hg4).
- intros.
- assert (g <> 0) by (intro; subst g; discriminate).
- assert (g' <> 0) by (intro; subst g'; discriminate).
+ intros H.
+ generalize (Z.ggcd_gcd a ('b)) (Zgcd_is_gcd a ('b))
+ (Z.gcd_nonneg a ('b)) (Z.ggcd_correct_divisors a ('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)).
+ 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').
elim (rel_prime_cross_prod aa bb cc dd).
- congruence.
- unfold rel_prime in |- *.
- (*rel_prime*)
- constructor.
- exists aa; auto with zarith.
- exists bb; auto with zarith.
- intros.
- inversion Hg1.
- destruct (H6 (g*x)) as (x',Hx).
- rewrite Hg3.
- destruct H2 as (xa,Hxa); exists xa; rewrite Hxa; ring.
- rewrite Hg4.
- destruct H3 as (xb,Hxb); exists xb; rewrite Hxb; ring.
- exists x'.
- apply Zmult_reg_l with g; auto. rewrite Hx at 1; ring.
- (* /rel_prime *)
- unfold rel_prime in |- *.
- (* rel_prime *)
- constructor.
- exists cc; auto with zarith.
- exists dd; auto with zarith.
- intros.
- inversion Hg'1.
- destruct (H6 (g'*x)) as (x',Hx).
- rewrite Hg'3.
- destruct H2 as (xc,Hxc); exists xc; rewrite Hxc; ring.
- rewrite Hg'4.
- destruct H3 as (xd,Hxd); exists xd; rewrite Hxd; ring.
- exists x'.
- apply Zmult_reg_l with g'; auto. rewrite Hx at 1; ring.
- (* /rel_prime *)
- assert (0<bb); [|auto with zarith].
- apply Zmult_gt_0_lt_0_reg_r with g.
- omega.
- rewrite Zmult_comm.
- rewrite <- Hg4; compute; auto.
- assert (0<dd); [|auto with zarith].
- apply Zmult_gt_0_lt_0_reg_r with g'.
- omega.
- rewrite Zmult_comm.
- rewrite <- Hg'4; compute; auto.
- apply Zmult_reg_l with (g'*g).
- intro H2; elim (Zmult_integral _ _ H2); auto.
- replace (g'*g*(aa*dd)) with ((g*aa)*(g'*dd)); [|ring].
- replace (g'*g*(bb*cc)) with ((g'*cc)*(g*bb)); [|ring].
- rewrite <- Hg3; rewrite <- Hg4; rewrite <- Hg'3; rewrite <- Hg'4; auto.
+ - congruence.
+ - (*rel_prime*)
+ constructor.
+ * exists aa; auto with zarith.
+ * exists bb; auto with zarith.
+ * intros x Ha Hb.
+ destruct Hg1 as (Hg11,Hg12,Hg13).
+ destruct (Hg13 (g*x)) as (x',Hx).
+ { rewrite Hg3.
+ destruct Ha as (xa,Hxa); exists xa; rewrite Hxa; ring. }
+ { rewrite Hg4.
+ destruct Hb as (xb,Hxb); exists xb; rewrite Hxb; ring. }
+ exists x'.
+ apply Z.mul_reg_l with g; auto. rewrite Hx at 1; ring.
+ - (* rel_prime *)
+ constructor.
+ * exists cc; auto with zarith.
+ * exists dd; auto with zarith.
+ * intros x Hc Hd.
+ inversion Hg'1 as (Hg'11,Hg'12,Hg'13).
+ destruct (Hg'13 (g'*x)) as (x',Hx).
+ { rewrite Hg'3.
+ destruct Hc as (xc,Hxc); exists xc; rewrite Hxc; ring. }
+ { rewrite Hg'4.
+ destruct Hd as (xd,Hxd); exists xd; rewrite Hxd; ring. }
+ exists x'.
+ apply Z.mul_reg_l with g'; auto. rewrite Hx at 1; ring.
+ - apply Z.lt_gt.
+ rewrite <- (Z.mul_pos_cancel_l g); [now rewrite <- Hg4 | omega].
+ - apply Z.lt_gt.
+ rewrite <- (Z.mul_pos_cancel_l g'); [now rewrite <- Hg'4 | omega].
+ - apply Z.mul_reg_l with (g*g').
+ * rewrite Z.mul_eq_0. now destruct 1.
+ * rewrite Z.mul_shuffle1, <- Hg3, <- Hg'4.
+ now rewrite Z.mul_shuffle1, <- Hg'3, <- Hg4, H, Z.mul_comm.
Close Scope Z_scope.
Qed.
@@ -137,39 +106,39 @@ Definition Qminus' x y := Qred (Qminus x y).
Lemma Qplus'_correct : forall p q : Q, (Qplus' p q)==(Qplus p q).
Proof.
- intros; unfold Qplus' in |- *; apply Qred_correct; auto.
+ intros; unfold Qplus'; apply Qred_correct; auto.
Qed.
Lemma Qmult'_correct : forall p q : Q, (Qmult' p q)==(Qmult p q).
Proof.
- intros; unfold Qmult' in |- *; apply Qred_correct; auto.
+ intros; unfold Qmult'; apply Qred_correct; auto.
Qed.
Lemma Qminus'_correct : forall p q : Q, (Qminus' p q)==(Qminus p q).
Proof.
- intros; unfold Qminus' in |- *; apply Qred_correct; auto.
+ intros; unfold Qminus'; apply Qred_correct; auto.
Qed.
Add Morphism Qplus' : Qplus'_comp.
Proof.
- intros; unfold Qplus' in |- *.
- rewrite H; rewrite H0; auto with qarith.
+ intros; unfold Qplus'.
+ rewrite H, H0; auto with qarith.
Qed.
Add Morphism Qmult' : Qmult'_comp.
- intros; unfold Qmult' in |- *.
- rewrite H; rewrite H0; auto with qarith.
+ intros; unfold Qmult'.
+ rewrite H, H0; auto with qarith.
Qed.
Add Morphism Qminus' : Qminus'_comp.
- intros; unfold Qminus' in |- *.
- rewrite H; rewrite H0; auto with qarith.
+ intros; unfold Qminus'.
+ rewrite H, H0; auto with qarith.
Qed.
Lemma Qred_opp: forall q, Qred (-q) = - (Qred q).
Proof.
intros (x, y); unfold Qred; simpl.
- rewrite Zggcd_opp; case Zggcd; intros p1 (p2, p3); simpl.
+ rewrite Z.ggcd_opp; case Z.ggcd; intros p1 (p2, p3); simpl.
unfold Qopp; auto.
Qed.
@@ -178,4 +147,3 @@ Theorem Qred_compare: forall x y,
Proof.
intros x y; apply Qcompare_comp; apply Qeq_sym; apply Qred_correct.
Qed.
-
diff --git a/theories/QArith/Qround.v b/theories/QArith/Qround.v
index ce363a33b..cfc1fb141 100644
--- a/theories/QArith/Qround.v
+++ b/theories/QArith/Qround.v
@@ -11,7 +11,7 @@ Require Import QArith.
Lemma Qopp_lt_compat: forall p q : Q, p < q -> - q < - p.
Proof.
intros (a1,a2) (b1,b2); unfold Qle, Qlt; simpl.
-do 2 rewrite <- Zopp_mult_distr_l; omega.
+rewrite !Z.mul_opp_l; omega.
Qed.
Hint Resolve Qopp_lt_compat : qarith.
@@ -20,7 +20,7 @@ Hint Resolve Qopp_lt_compat : qarith.
Coercion Local inject_Z : Z >-> Q.
-Definition Qfloor (x:Q) := let (n,d) := x in Zdiv n (Zpos d).
+Definition Qfloor (x:Q) := let (n,d) := x in Z.div n (Zpos d).
Definition Qceiling (x:Q) := (-(Qfloor (-x)))%Z.
Lemma Qfloor_Z : forall z:Z, Qfloor z = z.
@@ -46,7 +46,7 @@ simpl.
unfold Qle.
simpl.
replace (n*1)%Z with n by ring.
-rewrite Zmult_comm.
+rewrite Z.mul_comm.
apply Z_mult_div_ge.
auto with *.
Qed.
@@ -81,7 +81,7 @@ ring_simplify.
replace (n / ' d * ' d + ' d)%Z with
(('d * (n / 'd) + n mod 'd) + 'd - n mod 'd)%Z by ring.
rewrite <- Z_div_mod_eq; auto with*.
-rewrite <- Zlt_plus_swap.
+rewrite <- Z.lt_add_lt_sub_r.
destruct (Z_mod_lt n ('d)); auto with *.
Qed.
@@ -107,7 +107,7 @@ 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 (Zmult_comm ('yd) ('xd)).
+rewrite (Z.mul_comm ('yd) ('xd)).
apply Z_div_le; auto with *.
Qed.
@@ -125,7 +125,7 @@ Hint Resolve Qceiling_resp_le : qarith.
Add Morphism Qfloor with signature Qeq ==> eq as Qfloor_comp.
Proof.
intros x y H.
-apply Zle_antisym.
+apply Z.le_antisymm.
auto with *.
symmetry in H; auto with *.
Qed.
@@ -133,7 +133,7 @@ Qed.
Add Morphism Qceiling with signature Qeq ==> eq as Qceiling_comp.
Proof.
intros x y H.
-apply Zle_antisym.
+apply Z.le_antisymm.
auto with *.
symmetry in H; auto with *.
Qed.
@@ -142,9 +142,9 @@ Lemma Zdiv_Qdiv (n m: Z): (n / m)%Z = Qfloor (n / m).
Proof.
unfold Qfloor. intros. simpl.
destruct m as [?|?|p]; simpl.
- now rewrite Zdiv_0_r, Zmult_0_r.
- now rewrite Zmult_1_r.
- rewrite <- Zopp_eq_mult_neg_1.
- rewrite <- (Zopp_involutive (Zpos p)).
+ now rewrite Zdiv_0_r, Z.mul_0_r.
+ now rewrite Z.mul_1_r.
+ rewrite <- Z.opp_eq_mul_m1.
+ rewrite <- (Z.opp_involutive (Zpos p)).
now rewrite Zdiv_opp_opp.
Qed.
diff --git a/theories/Reals/AltSeries.v b/theories/Reals/AltSeries.v
index 648055d93..87f1aaf72 100644
--- a/theories/Reals/AltSeries.v
+++ b/theories/Reals/AltSeries.v
@@ -374,7 +374,7 @@ Proof.
replace (/ (2 * eps) * (INR N * (2 * eps))) with
(INR N * (2 * eps * / (2 * eps))); [ idtac | ring ].
rewrite <- Rinv_r_sym.
- rewrite Rmult_1_r; replace (INR N) with (IZR (Z_of_nat N)).
+ rewrite Rmult_1_r; replace (INR N) with (IZR (Z.of_nat N)).
rewrite <- H4.
elim H1; intros; assumption.
symmetry in |- *; apply INR_IZR_INZ.
diff --git a/theories/Reals/Exp_prop.v b/theories/Reals/Exp_prop.v
index 2a828a90a..9a38888da 100644
--- a/theories/Reals/Exp_prop.v
+++ b/theories/Reals/Exp_prop.v
@@ -770,7 +770,7 @@ Proof.
split.
unfold D_x, no_cond in |- *; split.
trivial.
- apply (sym_not_eq H6).
+ apply (not_eq_sym H6).
rewrite Rminus_0_r; apply H7.
unfold SFL in |- *.
case (cv 0); intros.
diff --git a/theories/Reals/RIneq.v b/theories/Reals/RIneq.v
index 7cf372e63..2f58201f7 100644
--- a/theories/Reals/RIneq.v
+++ b/theories/Reals/RIneq.v
@@ -58,7 +58,7 @@ Qed.
Lemma Rgt_not_eq : forall r1 r2, r1 > r2 -> r1 <> r2.
Proof.
- intros; apply sym_not_eq; apply Rlt_not_eq; auto with real.
+ intros; apply not_eq_sym; apply Rlt_not_eq; auto with real.
Qed.
(**********)
@@ -278,8 +278,8 @@ Proof. intros; red; apply Rlt_eq_compat with (r2:=r4) (r4:=r2); auto. Qed.
Lemma Rle_trans : forall r1 r2 r3, r1 <= r2 -> r2 <= r3 -> r1 <= r3.
Proof.
- generalize trans_eq Rlt_trans Rlt_eq_compat.
- unfold Rle in |- *.
+ generalize eq_trans Rlt_trans Rlt_eq_compat.
+ unfold Rle.
intuition eauto 2.
Qed.
@@ -1389,7 +1389,7 @@ Qed.
(**********)
Lemma tech_Rplus : forall r (s:R), 0 <= r -> 0 < s -> r + s <> 0.
Proof.
- intros; apply sym_not_eq; apply Rlt_not_eq.
+ intros; apply not_eq_sym; apply Rlt_not_eq.
rewrite Rplus_comm; replace 0 with (0 + 0); auto with real.
Qed.
Hint Immediate tech_Rplus: real.
@@ -1599,11 +1599,11 @@ Qed.
Hint Resolve lt_1_INR: real.
(**********)
-Lemma pos_INR_nat_of_P : forall p:positive, 0 < INR (nat_of_P p).
+Lemma pos_INR_nat_of_P : forall p:positive, 0 < INR (Pos.to_nat p).
Proof.
intro; apply lt_0_INR.
simpl in |- *; auto with real.
- apply nat_of_P_pos.
+ apply Pos2Nat.is_pos.
Qed.
Hint Resolve pos_INR_nat_of_P: real.
@@ -1666,7 +1666,7 @@ Proof.
case (le_lt_or_eq _ _ H1); intros H2.
apply Rlt_dichotomy_converse; auto with real.
exfalso; auto.
- apply sym_not_eq; apply Rlt_dichotomy_converse; auto with real.
+ apply not_eq_sym; apply Rlt_dichotomy_converse; auto with real.
Qed.
Hint Resolve not_INR: real.
@@ -1703,16 +1703,16 @@ Hint Resolve not_1_INR: real.
(**********)
-Lemma IZN : forall n:Z, (0 <= n)%Z -> exists m : nat, n = Z_of_nat m.
+Lemma IZN : forall n:Z, (0 <= n)%Z -> exists m : nat, n = Z.of_nat m.
Proof.
intros z; idtac; apply Z_of_nat_complete; assumption.
Qed.
(**********)
-Lemma INR_IZR_INZ : forall n:nat, INR n = IZR (Z_of_nat n).
+Lemma INR_IZR_INZ : forall n:nat, INR n = IZR (Z.of_nat n).
Proof.
simple induction n; auto with real.
- intros; simpl in |- *; rewrite nat_of_P_of_succ_nat;
+ intros; simpl in |- *; rewrite SuccNat2Pos.id_succ;
auto with real.
Qed.
@@ -1720,13 +1720,13 @@ Lemma plus_IZR_NEG_POS :
forall p q:positive, IZR (Zpos p + Zneg q) = IZR (Zpos p) + IZR (Zneg q).
Proof.
intros p q; simpl. rewrite Z.pos_sub_spec.
- case Pcompare_spec; intros H; simpl.
+ case Pos.compare_spec; intros H; simpl.
subst. ring.
- rewrite Pminus_minus by trivial.
- rewrite minus_INR by (now apply lt_le_weak, Plt_lt).
+ rewrite Pos2Nat.inj_sub by trivial.
+ rewrite minus_INR by (now apply lt_le_weak, Pos2Nat.inj_lt).
ring.
- rewrite Pminus_minus by trivial.
- rewrite minus_INR by (now apply lt_le_weak, Plt_lt).
+ rewrite Pos2Nat.inj_sub by trivial.
+ rewrite minus_INR by (now apply lt_le_weak, Pos2Nat.inj_lt).
ring.
Qed.
@@ -1734,10 +1734,10 @@ Qed.
Lemma plus_IZR : forall n m:Z, IZR (n + m) = IZR n + IZR m.
Proof.
intro z; destruct z; intro t; destruct t; intros; auto with real.
- simpl; intros; rewrite Pplus_plus; auto with real.
+ simpl; intros; rewrite Pos2Nat.inj_add; auto with real.
apply plus_IZR_NEG_POS.
- rewrite Zplus_comm; rewrite Rplus_comm; apply plus_IZR_NEG_POS.
- simpl; intros; rewrite Pplus_plus; rewrite plus_INR;
+ rewrite Z.add_comm; rewrite Rplus_comm; apply plus_IZR_NEG_POS.
+ simpl; intros; rewrite Pos2Nat.inj_add; rewrite plus_INR;
auto with real.
Qed.
@@ -1745,31 +1745,31 @@ Qed.
Lemma mult_IZR : forall n m:Z, IZR (n * m) = IZR n * IZR m.
Proof.
intros z t; case z; case t; simpl in |- *; auto with real.
- intros t1 z1; rewrite Pmult_mult; auto with real.
- intros t1 z1; rewrite Pmult_mult; auto with real.
+ intros t1 z1; rewrite Pos2Nat.inj_mul; auto with real.
+ intros t1 z1; rewrite Pos2Nat.inj_mul; auto with real.
rewrite Rmult_comm.
rewrite Ropp_mult_distr_l_reverse; auto with real.
apply Ropp_eq_compat; rewrite mult_comm; auto with real.
- intros t1 z1; rewrite Pmult_mult; auto with real.
+ intros t1 z1; rewrite Pos2Nat.inj_mul; auto with real.
rewrite Ropp_mult_distr_l_reverse; auto with real.
- intros t1 z1; rewrite Pmult_mult; auto with real.
+ intros t1 z1; rewrite Pos2Nat.inj_mul; auto with real.
rewrite Rmult_opp_opp; auto with real.
Qed.
-Lemma pow_IZR : forall z n, pow (IZR z) n = IZR (Zpower z (Z_of_nat n)).
+Lemma pow_IZR : forall z n, pow (IZR z) n = IZR (Z.pow z (Z.of_nat n)).
Proof.
intros z [|n];simpl;trivial.
rewrite Zpower_pos_nat.
- rewrite nat_of_P_of_succ_nat. unfold Zpower_nat;simpl.
+ rewrite SuccNat2Pos.id_succ. unfold Zpower_nat;simpl.
rewrite mult_IZR.
induction n;simpl;trivial.
rewrite mult_IZR;ring[IHn].
Qed.
(**********)
-Lemma succ_IZR : forall n:Z, IZR (Zsucc n) = IZR n + 1.
+Lemma succ_IZR : forall n:Z, IZR (Z.succ n) = IZR n + 1.
Proof.
- intro; change 1 with (IZR 1); unfold Zsucc; apply plus_IZR.
+ intro; change 1 with (IZR 1); unfold Z.succ; apply plus_IZR.
Qed.
(**********)
@@ -1782,7 +1782,7 @@ Definition Ropp_Ropp_IZR := opp_IZR.
Lemma minus_IZR : forall n m:Z, IZR (n - m) = IZR n - IZR m.
Proof.
- intros; unfold Zminus, Rminus.
+ intros; unfold Z.sub, Rminus.
rewrite <- opp_IZR.
apply plus_IZR.
Qed.
@@ -1790,7 +1790,7 @@ Qed.
(**********)
Lemma Z_R_minus : forall n m:Z, IZR n - IZR m = IZR (n - m).
Proof.
- intros z1 z2; unfold Rminus in |- *; unfold Zminus in |- *.
+ intros z1 z2; unfold Rminus in |- *; unfold Z.sub in |- *.
rewrite <- (Ropp_Ropp_IZR z2); symmetry in |- *; apply plus_IZR.
Qed.
@@ -1799,7 +1799,7 @@ Lemma lt_0_IZR : forall n:Z, 0 < IZR n -> (0 < n)%Z.
Proof.
intro z; case z; simpl in |- *; intros.
absurd (0 < 0); auto with real.
- unfold Zlt in |- *; simpl in |- *; trivial.
+ unfold Z.lt in |- *; simpl in |- *; trivial.
case Rlt_not_le with (1 := H).
replace 0 with (-0); auto with real.
Qed.
@@ -1807,7 +1807,7 @@ Qed.
(**********)
Lemma lt_IZR : forall n m:Z, IZR n < IZR m -> (n < m)%Z.
Proof.
- intros z1 z2 H; apply Zlt_0_minus_lt.
+ intros z1 z2 H; apply Z.lt_0_sub.
apply lt_0_IZR.
rewrite <- Z_R_minus.
exact (Rgt_minus (IZR z2) (IZR z1) H).
@@ -1817,8 +1817,8 @@ Qed.
Lemma eq_IZR_R0 : forall n:Z, IZR n = 0 -> n = 0%Z.
Proof.
intro z; destruct z; simpl in |- *; intros; auto with zarith.
- case (Rlt_not_eq 0 (INR (nat_of_P p))); auto with real.
- case (Rlt_not_eq (- INR (nat_of_P p)) 0); auto with real.
+ case (Rlt_not_eq 0 (INR (Pos.to_nat p))); auto with real.
+ case (Rlt_not_eq (- INR (Pos.to_nat p)) 0); auto with real.
apply Ropp_lt_gt_0_contravar. unfold Rgt in |- *; apply pos_INR_nat_of_P.
Qed.
@@ -1841,7 +1841,7 @@ Qed.
Lemma le_0_IZR : forall n:Z, 0 <= IZR n -> (0 <= n)%Z.
Proof.
unfold Rle in |- *; intros z [H| H].
- red in |- *; intro; apply (Zlt_le_weak 0 z (lt_0_IZR z H)); assumption.
+ red in |- *; intro; apply (Z.lt_le_incl 0 z (lt_0_IZR z H)); assumption.
rewrite (eq_IZR_R0 z); auto with zarith real.
Qed.
@@ -1849,7 +1849,7 @@ Qed.
Lemma le_IZR : forall n m:Z, IZR n <= IZR m -> (n <= m)%Z.
Proof.
unfold Rle in |- *; intros z1 z2 [H| H].
- apply (Zlt_le_weak z1 z2); auto with real.
+ apply (Z.lt_le_incl z1 z2); auto with real.
apply lt_IZR; trivial.
rewrite (eq_IZR z1 z2); auto with zarith real.
Qed.
@@ -1885,10 +1885,10 @@ Qed.
Lemma one_IZR_lt1 : forall n:Z, -1 < IZR n < 1 -> n = 0%Z.
Proof.
intros z [H1 H2].
- apply Zle_antisym.
- apply Zlt_succ_le; apply lt_IZR; trivial.
- replace 0%Z with (Zsucc (-1)); trivial.
- apply Zlt_le_succ; apply lt_IZR; trivial.
+ apply Z.le_antisymm.
+ apply Z.lt_succ_r; apply lt_IZR; trivial.
+ replace 0%Z with (Z.succ (-1)); trivial.
+ apply Z.le_succ_l; apply lt_IZR; trivial.
Qed.
Lemma one_IZR_r_R1 :
diff --git a/theories/Reals/R_Ifp.v b/theories/Reals/R_Ifp.v
index 9e04a7daf..c089b648d 100644
--- a/theories/Reals/R_Ifp.v
+++ b/theories/Reals/R_Ifp.v
@@ -58,7 +58,7 @@ Proof.
intros a b; rewrite b; clear a b; rewrite <- Z_R_minus;
cut (up 0 = 1%Z).
intro; rewrite H1;
- rewrite (Rminus_diag_eq (IZR 1) (IZR 1) (refl_equal (IZR 1)));
+ rewrite (Rminus_diag_eq (IZR 1) (IZR 1) (eq_refl (IZR 1)));
apply Ropp_0.
elim (archimed 0); intros; clear H2; unfold Rgt in H1;
rewrite (Rminus_0_r (IZR (up 0))) in H0; generalize (lt_O_IZR (up 0) H1);
@@ -130,16 +130,16 @@ Proof.
Qed.
(**********)
-Lemma Int_part_INR : forall n:nat, Int_part (INR n) = Z_of_nat n.
+Lemma Int_part_INR : forall n:nat, Int_part (INR n) = Z.of_nat n.
Proof.
intros n; unfold Int_part in |- *.
- cut (up (INR n) = (Z_of_nat n + Z_of_nat 1)%Z).
+ cut (up (INR n) = (Z.of_nat n + Z.of_nat 1)%Z).
intros H'; rewrite H'; simpl in |- *; ring.
- apply sym_equal; apply tech_up; auto.
- replace (Z_of_nat n + Z_of_nat 1)%Z with (Z_of_nat (S n)).
+ symmetry; apply tech_up; auto.
+ replace (Z.of_nat n + Z.of_nat 1)%Z with (Z.of_nat (S n)).
repeat rewrite <- INR_IZR_INZ.
apply lt_INR; auto.
- rewrite Zplus_comm; rewrite <- Znat.inj_plus; simpl in |- *; auto.
+ rewrite Z.add_comm; rewrite <- Znat.Nat2Z.inj_add; simpl in |- *; auto.
rewrite plus_IZR; simpl in |- *; auto with real.
repeat rewrite <- INR_IZR_INZ; auto with real.
Qed.
diff --git a/theories/Reals/Ranalysis2.v b/theories/Reals/Ranalysis2.v
index ed80ac433..732a61017 100644
--- a/theories/Reals/Ranalysis2.v
+++ b/theories/Reals/Ranalysis2.v
@@ -429,7 +429,7 @@ Proof.
rewrite Rmult_1_r in H12; rewrite <- Rinv_r_sym in H12;
[ idtac | discrR ].
cut (IZR 1 < IZR 2).
- unfold IZR in |- *; unfold INR, nat_of_P in |- *; simpl in |- *; intro;
+ unfold IZR in |- *; unfold INR, Pos.to_nat in |- *; simpl in |- *; intro;
elim (Rlt_irrefl 1 (Rlt_trans _ _ _ H13 H12)).
apply IZR_lt; omega.
unfold Rabs in |- *; case (Rcase_abs (/ 2)); intro.
diff --git a/theories/Reals/Raxioms.v b/theories/Reals/Raxioms.v
index 8f01d7d03..259e1ac04 100644
--- a/theories/Reals/Raxioms.v
+++ b/theories/Reals/Raxioms.v
@@ -122,8 +122,8 @@ Arguments INR n%nat.
Definition IZR (z:Z) : R :=
match z with
| Z0 => 0
- | Zpos n => INR (nat_of_P n)
- | Zneg n => - INR (nat_of_P n)
+ | Zpos n => INR (Pos.to_nat n)
+ | Zneg n => - INR (Pos.to_nat n)
end.
Arguments IZR z%Z.
diff --git a/theories/Reals/Rbasic_fun.v b/theories/Reals/Rbasic_fun.v
index 4bc7fd104..ee9ea921c 100644
--- a/theories/Reals/Rbasic_fun.v
+++ b/theories/Reals/Rbasic_fun.v
@@ -623,7 +623,7 @@ Proof.
apply RmaxLess1; auto.
Qed.
-Lemma Rabs_Zabs : forall z:Z, Rabs (IZR z) = IZR (Zabs z).
+Lemma Rabs_Zabs : forall z:Z, Rabs (IZR z) = IZR (Z.abs z).
Proof.
intros z; case z; simpl in |- *; auto with real.
apply Rabs_right; auto with real.
@@ -632,7 +632,7 @@ Proof.
apply Rabs_right; auto with real zarith.
Qed.
-Lemma abs_IZR : forall z, IZR (Zabs z) = Rabs (IZR z).
+Lemma abs_IZR : forall z, IZR (Z.abs z) = Rabs (IZR z).
Proof.
intros.
now rewrite Rabs_Zabs.
diff --git a/theories/Reals/Rderiv.v b/theories/Reals/Rderiv.v
index 105d8347d..8f9b99c28 100644
--- a/theories/Reals/Rderiv.v
+++ b/theories/Reals/Rderiv.v
@@ -85,7 +85,7 @@ Proof.
unfold Rgt in |- *; intro; elim (H7 H5); clear H7;
intros; clear H4 H5; generalize (H1 x1 (conj H3 H8));
clear H1; intro; unfold D_x in H3; elim H3; intros;
- generalize (sym_not_eq H5); clear H5; intro H5;
+ generalize (not_eq_sym H5); clear H5; intro H5;
generalize (Rminus_eq_contra x1 x0 H5); intro; generalize H1;
pattern (d x0) at 1 in |- *;
rewrite <- (let (H1, H2) := Rmult_ne (d x0) in H2);
@@ -177,8 +177,8 @@ Proof.
unfold D_in in |- *; intros; unfold limit1_in in |- *;
unfold limit_in in |- *; unfold Rdiv in |- *; intros;
simpl in |- *; split with eps; split; auto.
- intros; rewrite (Rminus_diag_eq y y (refl_equal y)); rewrite Rmult_0_l;
- unfold R_dist in |- *; rewrite (Rminus_diag_eq 0 0 (refl_equal 0));
+ intros; rewrite (Rminus_diag_eq y y (eq_refl y)); rewrite Rmult_0_l;
+ unfold R_dist in |- *; rewrite (Rminus_diag_eq 0 0 (eq_refl 0));
unfold Rabs in |- *; case (Rcase_abs 0); intro.
absurd (0 < 0); auto.
red in |- *; intro; apply (Rlt_irrefl 0 H1).
@@ -193,8 +193,8 @@ Proof.
unfold limit_in in |- *; intros; simpl in |- *; split with eps;
split; auto.
intros; elim H0; clear H0; intros; unfold D_x in H0; elim H0; intros;
- rewrite (Rinv_r (x - x0) (Rminus_eq_contra x x0 (sym_not_eq H3)));
- unfold R_dist in |- *; rewrite (Rminus_diag_eq 1 1 (refl_equal 1));
+ rewrite (Rinv_r (x - x0) (Rminus_eq_contra x x0 (not_eq_sym H3)));
+ unfold R_dist in |- *; rewrite (Rminus_diag_eq 1 1 (eq_refl 1));
unfold Rabs in |- *; case (Rcase_abs 0); intro.
absurd (0 < 0); auto.
red in |- *; intro; apply (Rlt_irrefl 0 r).
@@ -270,7 +270,7 @@ Proof.
ring.
unfold limit1_in in |- *; unfold limit_in in |- *; simpl in |- *; intros;
split with eps; split; auto; intros; elim (R_dist_refl (g x0) (g x0));
- intros a b; rewrite (b (refl_equal (g x0))); unfold Rgt in H;
+ intros a b; rewrite (b (eq_refl (g x0))); unfold Rgt in H;
assumption.
Qed.
@@ -391,7 +391,7 @@ Proof.
rewrite (Rminus_diag_eq (f x2) (f x0) H12) in H16;
rewrite (Rmult_0_l (/ (x2 - x0))) in H16;
rewrite (Rmult_0_l (dg (f x0))) in H16; rewrite H12;
- rewrite (Rminus_diag_eq (g (f x0)) (g (f x0)) (refl_equal (g (f x0))));
+ rewrite (Rminus_diag_eq (g (f x0)) (g (f x0)) (eq_refl (g (f x0))));
rewrite (Rmult_0_l (/ (x2 - x0))); assumption.
clear H10 H5; elim H11; clear H11; intros; elim H5; clear H5; intros;
cut
diff --git a/theories/Reals/Rfunctions.v b/theories/Reals/Rfunctions.v
index 5ba0a187c..4c4903cb2 100644
--- a/theories/Reals/Rfunctions.v
+++ b/theories/Reals/Rfunctions.v
@@ -25,8 +25,8 @@ Require Export SplitRmult.
Require Export ArithProp.
Require Import Omega.
Require Import Zpower.
-Open Local Scope nat_scope.
-Open Local Scope R_scope.
+Local Open Scope nat_scope.
+Local Open Scope R_scope.
(*******************************)
(** * Lemmas about factorial *)
@@ -221,8 +221,8 @@ Qed.
Lemma RPow_abs : forall (x:R) (n:nat), Rabs x ^ n = Rabs (x ^ n).
Proof.
intro; simple induction n; simpl.
- apply sym_eq; apply Rabs_pos_eq; apply Rlt_le; apply Rlt_0_1.
- intros; rewrite H; apply sym_eq; apply Rabs_mult.
+ symmetry; apply Rabs_pos_eq; apply Rlt_le; apply Rlt_0_1.
+ intros; rewrite H; symmetry; apply Rabs_mult.
Qed.
@@ -526,13 +526,13 @@ Qed.
(*i Due to L.Thery i*)
Ltac case_eq name :=
- generalize (refl_equal name); pattern name at -1; case name.
+ generalize (eq_refl name); pattern name at -1; case name.
Definition powerRZ (x:R) (n:Z) :=
match n with
| Z0 => 1
- | Zpos p => x ^ nat_of_P p
- | Zneg p => / x ^ nat_of_P p
+ | Zpos p => x ^ Pos.to_nat p
+ | Zneg p => / x ^ Pos.to_nat p
end.
Infix Local "^Z" := powerRZ (at level 30, right associativity) : R_scope.
@@ -548,7 +548,7 @@ Proof.
reflexivity.
Qed.
-Lemma powerRZ_1 : forall x:R, x ^Z Zsucc 0 = x.
+Lemma powerRZ_1 : forall x:R, x ^Z Z.succ 0 = x.
Proof.
simpl; auto with real.
Qed.
@@ -558,67 +558,63 @@ Proof.
destruct z; simpl; auto with real.
Qed.
+Lemma powerRZ_pos_sub (x:R) (n m:positive) : x <> 0 ->
+ x ^Z (Z.pos_sub n m) = x ^ Pos.to_nat n * / x ^ Pos.to_nat m.
+Proof.
+ intro Hx.
+ rewrite Z.pos_sub_spec.
+ case Pos.compare_spec; intro H; simpl.
+ - subst; auto with real.
+ - rewrite Pos2Nat.inj_sub by trivial.
+ rewrite Pos2Nat.inj_lt in H.
+ rewrite (pow_RN_plus x _ (Pos.to_nat n)) by auto with real.
+ rewrite plus_comm, le_plus_minus_r by auto with real.
+ rewrite Rinv_mult_distr, Rinv_involutive; auto with real.
+ - rewrite Pos2Nat.inj_sub by trivial.
+ rewrite Pos2Nat.inj_lt in H.
+ rewrite (pow_RN_plus x _ (Pos.to_nat m)) by auto with real.
+ rewrite plus_comm, le_plus_minus_r by auto with real.
+ reflexivity.
+Qed.
+
Lemma powerRZ_add :
forall (x:R) (n m:Z), x <> 0 -> x ^Z (n + m) = x ^Z n * x ^Z m.
Proof.
- intro x; destruct n as [| n1| n1]; destruct m as [| m1| m1]; simpl;
- auto with real.
-(* POS/POS *)
- rewrite Pplus_plus; auto with real.
-(* POS/NEG *)
- rewrite Z.pos_sub_spec.
- case Pcompare_spec; intros; simpl.
- subst; auto with real.
- rewrite Pminus_minus by trivial.
- rewrite (pow_RN_plus x _ (nat_of_P n1)) by auto with real.
- rewrite plus_comm, le_plus_minus_r by (now apply lt_le_weak, Plt_lt).
- rewrite Rinv_mult_distr, Rinv_involutive; auto with real.
- rewrite Pminus_minus by trivial.
- rewrite (pow_RN_plus x _ (nat_of_P m1)) by auto with real.
- rewrite plus_comm, le_plus_minus_r by (now apply lt_le_weak, Plt_lt).
- reflexivity.
-(* NEG/POS *)
- rewrite Z.pos_sub_spec.
- case Pcompare_spec; intros; simpl.
- subst; auto with real.
- rewrite Pminus_minus by trivial.
- rewrite (pow_RN_plus x _ (nat_of_P m1)) by auto with real.
- rewrite plus_comm, le_plus_minus_r by (now apply lt_le_weak, Plt_lt).
- rewrite Rinv_mult_distr, Rinv_involutive; auto with real.
- rewrite Pminus_minus by trivial.
- rewrite (pow_RN_plus x _ (nat_of_P n1)) by auto with real.
- rewrite plus_comm, le_plus_minus_r by (now apply lt_le_weak, Plt_lt).
- auto with real.
-(* NEG/NEG *)
- rewrite Pplus_plus; auto with real.
- intros H'; rewrite pow_add; auto with real.
- apply Rinv_mult_distr; auto.
- apply pow_nonzero; auto.
- apply pow_nonzero; auto.
+ intros x [|n|n] [|m|m]; simpl; intros; auto with real.
+ - (* + + *)
+ rewrite Pos2Nat.inj_add; auto with real.
+ - (* + - *)
+ now apply powerRZ_pos_sub.
+ - (* - + *)
+ rewrite Rmult_comm. now apply powerRZ_pos_sub.
+ - (* - - *)
+ rewrite Pos2Nat.inj_add; auto with real.
+ rewrite pow_add; auto with real.
+ apply Rinv_mult_distr; apply pow_nonzero; auto.
Qed.
Hint Resolve powerRZ_O powerRZ_1 powerRZ_NOR powerRZ_add: real.
Lemma Zpower_nat_powerRZ :
- forall n m:nat, IZR (Zpower_nat (Z_of_nat n) m) = INR n ^Z Z_of_nat m.
+ forall n m:nat, IZR (Zpower_nat (Z.of_nat n) m) = INR n ^Z Z.of_nat m.
Proof.
intros n m; elim m; simpl; auto with real.
- intros m1 H'; rewrite nat_of_P_of_succ_nat; simpl.
- replace (Zpower_nat (Z_of_nat n) (S m1)) with
- (Z_of_nat n * Zpower_nat (Z_of_nat n) m1)%Z.
+ intros m1 H'; rewrite SuccNat2Pos.id_succ; simpl.
+ replace (Zpower_nat (Z.of_nat n) (S m1)) with
+ (Z.of_nat n * Zpower_nat (Z.of_nat n) m1)%Z.
rewrite mult_IZR; auto with real.
repeat rewrite <- INR_IZR_INZ; simpl.
rewrite H'; simpl.
case m1; simpl; auto with real.
- intros m2; rewrite nat_of_P_of_succ_nat; auto.
+ intros m2; rewrite SuccNat2Pos.id_succ; auto.
unfold Zpower_nat; auto.
Qed.
Lemma Zpower_pos_powerRZ :
- forall n m, IZR (Zpower_pos n m) = IZR n ^Z Zpos m.
+ forall n m, IZR (Z.pow_pos n m) = IZR n ^Z Zpos m.
Proof.
intros.
rewrite Zpower_pos_nat; simpl.
- induction (nat_of_P m).
+ induction (Pos.to_nat m).
easy.
unfold Zpower_nat; simpl.
rewrite mult_IZR.
@@ -638,10 +634,10 @@ Qed.
Hint Resolve powerRZ_le: real.
Lemma Zpower_nat_powerRZ_absolu :
- forall n m:Z, (0 <= m)%Z -> IZR (Zpower_nat n (Zabs_nat m)) = IZR n ^Z m.
+ forall n m:Z, (0 <= m)%Z -> IZR (Zpower_nat n (Z.abs_nat m)) = IZR n ^Z m.
Proof.
intros n m; case m; simpl; auto with zarith.
- intros p H'; elim (nat_of_P p); simpl; auto with zarith.
+ intros p H'; elim (Pos.to_nat p); simpl; auto with zarith.
intros n0 H'0; rewrite <- H'0; simpl; auto with zarith.
rewrite <- mult_IZR; auto.
intros p H'; absurd (0 <= Zneg p)%Z; auto with zarith.
@@ -650,9 +646,9 @@ Qed.
Lemma powerRZ_R1 : forall n:Z, 1 ^Z n = 1.
Proof.
intros n; case n; simpl; auto.
- intros p; elim (nat_of_P p); simpl; auto; intros n0 H'; rewrite H';
+ intros p; elim (Pos.to_nat p); simpl; auto; intros n0 H'; rewrite H';
ring.
- intros p; elim (nat_of_P p); simpl.
+ intros p; elim (Pos.to_nat p); simpl.
exact Rinv_1.
intros n1 H'; rewrite Rinv_mult_distr; try rewrite Rinv_1; try rewrite H';
auto with real.
@@ -760,9 +756,9 @@ Qed.
Lemma R_dist_refl : forall x y:R, R_dist x y = 0 <-> x = y.
Proof.
unfold R_dist; intros; split_Rabs; split; intros.
- rewrite (Ropp_minus_distr x y) in H; apply sym_eq;
+ rewrite (Ropp_minus_distr x y) in H; symmetry;
apply (Rminus_diag_uniq y x H).
- rewrite (Ropp_minus_distr x y); generalize (sym_eq H); intro;
+ rewrite (Ropp_minus_distr x y); generalize (eq_sym H); intro;
apply (Rminus_diag_eq y x H0).
apply (Rminus_diag_uniq x y H).
apply (Rminus_diag_eq x y H).
diff --git a/theories/Reals/RiemannInt_SF.v b/theories/Reals/RiemannInt_SF.v
index 013cbdce1..18c24c184 100644
--- a/theories/Reals/RiemannInt_SF.v
+++ b/theories/Reals/RiemannInt_SF.v
@@ -21,7 +21,7 @@ Set Implicit Arguments.
Definition Nbound (I:nat -> Prop) : Prop :=
exists n : nat, (forall i:nat, I i -> (i <= n)%nat).
-Lemma IZN_var : forall z:Z, (0 <= z)%Z -> {n : nat | z = Z_of_nat n}.
+Lemma IZN_var : forall z:Z, (0 <= z)%Z -> {n : nat | z = Z.of_nat n}.
Proof.
intros; apply Z_of_nat_complete_inf; assumption.
Qed.
diff --git a/theories/Reals/Rlimit.v b/theories/Reals/Rlimit.v
index 5c864de38..80b900af7 100644
--- a/theories/Reals/Rlimit.v
+++ b/theories/Reals/Rlimit.v
@@ -196,7 +196,7 @@ Proof.
case (H0 (dist R_met (f x0) l)); auto.
intros alpha1 [H2 H3]; apply H3; auto; split; auto.
case (dist_refl R_met x0 x0); intros Hr1 Hr2; rewrite Hr2; auto.
- case (dist_refl R_met (f x0) l); intros Hr1 Hr2; apply sym_eq; auto.
+ case (dist_refl R_met (f x0) l); intros Hr1 Hr2; symmetry; auto.
Qed.
(*********)
@@ -270,7 +270,7 @@ Lemma limit_free :
Proof.
unfold limit1_in in |- *; unfold limit_in in |- *; simpl in |- *; intros;
split with eps; split; auto; intros; elim (R_dist_refl (f x) (f x));
- intros a b; rewrite (b (refl_equal (f x))); unfold Rgt in H;
+ intros a b; rewrite (b (eq_refl (f x))); unfold Rgt in H;
assumption.
Qed.
diff --git a/theories/Reals/Rpower.v b/theories/Reals/Rpower.v
index 146e97361..b5bd2b734 100644
--- a/theories/Reals/Rpower.v
+++ b/theories/Reals/Rpower.v
@@ -561,7 +561,7 @@ Proof.
apply Rminus_eq_contra.
red in |- *; intros H2; case HD2.
symmetry in |- *; apply (ln_inv _ _ HD1 Hy H2).
- apply Rminus_eq_contra; apply (sym_not_eq HD2).
+ apply Rminus_eq_contra; apply (not_eq_sym HD2).
apply Rinv_neq_0_compat; apply Rminus_eq_contra; red in |- *; intros H2;
case HD2; apply ln_inv; auto.
assumption.
@@ -588,7 +588,7 @@ Proof.
[ idtac | ring ].
apply H1.
elim H2; intros H3 _; unfold D_x in H3; elim H3; clear H3; intros _ H3;
- apply Rminus_eq_contra; apply (sym_not_eq (A:=R));
+ apply Rminus_eq_contra; apply (not_eq_sym (A:=R));
apply H3.
elim H2; clear H2; intros _ H2; apply H2.
assumption.
@@ -625,7 +625,7 @@ Proof.
replace (- h - x / 2 + (x / 2 + x / 2 + h)) with (x / 2); [ apply H7 | ring ].
apply r.
apply Rplus_lt_le_0_compat; [ assumption | apply Rge_le; apply r ].
- apply (sym_not_eq (A:=R)); apply Rminus_not_eq; replace (x + h - x) with h;
+ apply (not_eq_sym (A:=R)); apply Rminus_not_eq; replace (x + h - x) with h;
[ apply H5 | ring ].
replace (x + h - x) with h;
[ apply Rlt_le_trans with alp;
diff --git a/theories/Reals/Rseries.v b/theories/Reals/Rseries.v
index 479d381d4..9d1ba7381 100644
--- a/theories/Reals/Rseries.v
+++ b/theories/Reals/Rseries.v
@@ -182,13 +182,13 @@ Section sequence.
assert (Hs0: forall n, sum n = 0).
intros n.
- specialize (Hm1 (sum n) (ex_intro _ _ (refl_equal _))).
+ specialize (Hm1 (sum n) (ex_intro _ _ (eq_refl _))).
apply Rle_antisym with (2 := proj1 (Hsum n)).
now rewrite <- Hm.
assert (Hub: forall n, Un n <= l - eps).
intros n.
- generalize (refl_equal (sum (S n))).
+ generalize (eq_refl (sum (S n))).
simpl sum at 1.
rewrite 2!Hs0, Rplus_0_l.
unfold test.
@@ -238,7 +238,7 @@ Section sequence.
rewrite (IHN H6), Rplus_0_l.
unfold test.
destruct Rle_lt_dec.
- apply refl_equal.
+ apply eq_refl.
now elim Rlt_not_le with (1 := r).
destruct (le_or_lt N n) as [Hn|Hn].
@@ -272,13 +272,13 @@ Section sequence.
Proof.
intro; induction N as [| N HrecN].
split with (Un 0); intros; rewrite (le_n_O_eq n H);
- apply (Req_le (Un n) (Un n) (refl_equal (Un n))).
+ apply (Req_le (Un n) (Un n) (eq_refl (Un n))).
elim HrecN; clear HrecN; intros; split with (Rmax (Un (S N)) x); intros;
elim (Rmax_Rle (Un (S N)) x (Un n)); intros; clear H1;
inversion H0.
rewrite <- H1; rewrite <- H1 in H2;
apply
- (H2 (or_introl (Un n <= x) (Req_le (Un n) (Un n) (refl_equal (Un n))))).
+ (H2 (or_introl (Un n <= x) (Req_le (Un n) (Un n) (eq_refl (Un n))))).
apply (H2 (or_intror (Un n <= Un (S N)) (H n H3))).
Qed.
diff --git a/theories/Reals/Rsqrt_def.v b/theories/Reals/Rsqrt_def.v
index 7c3b4699f..223649ef5 100644
--- a/theories/Reals/Rsqrt_def.v
+++ b/theories/Reals/Rsqrt_def.v
@@ -418,7 +418,7 @@ Proof.
unfold D_x, no_cond in |- *.
split.
trivial.
- apply (sym_not_eq (A:=R)); assumption.
+ apply (not_eq_sym (A:=R)); assumption.
apply H5; assumption.
Qed.
diff --git a/theories/Reals/Rtopology.v b/theories/Reals/Rtopology.v
index f1142d245..814e336c2 100644
--- a/theories/Reals/Rtopology.v
+++ b/theories/Reals/Rtopology.v
@@ -287,7 +287,7 @@ Proof.
apply H5; split.
unfold D_x, no_cond in |- *; split.
trivial.
- apply (sym_not_eq (A:=R)); apply H7.
+ apply (not_eq_sym (A:=R)); apply H7.
unfold disc in H6; apply H6.
intros; unfold continuity_pt in |- *; unfold continue_in in |- *;
unfold limit1_in in |- *; unfold limit_in in |- *;
@@ -1581,7 +1581,7 @@ Proof.
right; elim H1; intros; elim H2; intros; exists x; exists x0; intros.
split;
[ assumption
- | split; [ assumption | apply (sym_not_eq (A:=R)); assumption ] ].
+ | split; [ assumption | apply (not_eq_sym (A:=R)); assumption ] ].
left; exists x; split.
assumption.
intros; case (Req_dec x0 x); intro.
diff --git a/theories/Reals/Rtrigo_alt.v b/theories/Reals/Rtrigo_alt.v
index 08fda178b..4a405660c 100644
--- a/theories/Reals/Rtrigo_alt.v
+++ b/theories/Reals/Rtrigo_alt.v
@@ -84,7 +84,7 @@ Proof.
left; apply pow_lt; assumption.
apply Rmult_le_reg_l with (INR (fact (S (S (2 * S n0 + 1))))).
rewrite <- H3; apply lt_INR_0; apply neq_O_lt; red in |- *; intro;
- assert (H5 := sym_eq H4); elim (fact_neq_0 _ H5).
+ assert (H5 := eq_sym H4); elim (fact_neq_0 _ H5).
rewrite <- H3; rewrite (Rmult_comm (INR (fact (2 * S (S n0) + 1))));
rewrite Rmult_assoc; rewrite <- Rinv_l_sym.
rewrite Rmult_1_r; rewrite H3; do 2 rewrite fact_simpl; do 2 rewrite mult_INR;
@@ -274,7 +274,7 @@ Proof.
apply pow_le; assumption.
apply Rmult_le_reg_l with (INR (fact (S (S (2 * S n1))))).
rewrite <- H4; apply lt_INR_0; apply neq_O_lt; red in |- *; intro;
- assert (H6 := sym_eq H5); elim (fact_neq_0 _ H6).
+ assert (H6 := eq_sym H5); elim (fact_neq_0 _ H6).
rewrite <- H4; rewrite (Rmult_comm (INR (fact (2 * S (S n1)))));
rewrite Rmult_assoc; rewrite <- Rinv_l_sym.
rewrite Rmult_1_r; rewrite H4; do 2 rewrite fact_simpl; do 2 rewrite mult_INR;
diff --git a/theories/Reals/Rtrigo_calc.v b/theories/Reals/Rtrigo_calc.v
index 40989418c..90ba205f6 100644
--- a/theories/Reals/Rtrigo_calc.v
+++ b/theories/Reals/Rtrigo_calc.v
@@ -150,7 +150,7 @@ Proof.
intro H2;
[ assumption
| absurd (0 = sqrt 2);
- [ apply (sym_not_eq (A:=R)); apply sqrt2_neq_0 | assumption ] ] ].
+ [ apply (not_eq_sym (A:=R)); apply sqrt2_neq_0 | assumption ] ] ].
Qed.
Lemma Rlt_sqrt3_0 : 0 < sqrt 3.
diff --git a/theories/Reals/Rtrigo_def.v b/theories/Reals/Rtrigo_def.v
index c64931353..9b6edab68 100644
--- a/theories/Reals/Rtrigo_def.v
+++ b/theories/Reals/Rtrigo_def.v
@@ -130,17 +130,17 @@ Proof.
intro; cut (0 <= up (/ eps))%Z.
intro; assert (H2 := IZN _ H1); elim H2; intros; exists (max x 1).
split.
- cut (0 < IZR (Z_of_nat x)).
- intro; rewrite INR_IZR_INZ; apply Rle_lt_trans with (/ IZR (Z_of_nat x)).
- apply Rmult_le_reg_l with (IZR (Z_of_nat x)).
+ cut (0 < IZR (Z.of_nat x)).
+ intro; rewrite INR_IZR_INZ; apply Rle_lt_trans with (/ IZR (Z.of_nat x)).
+ apply Rmult_le_reg_l with (IZR (Z.of_nat x)).
assumption.
rewrite <- Rinv_r_sym;
[ idtac | red in |- *; intro; rewrite H5 in H4; elim (Rlt_irrefl _ H4) ].
- apply Rmult_le_reg_l with (IZR (Z_of_nat (max x 1))).
- apply Rlt_le_trans with (IZR (Z_of_nat x)).
+ apply Rmult_le_reg_l with (IZR (Z.of_nat (max x 1))).
+ apply Rlt_le_trans with (IZR (Z.of_nat x)).
assumption.
repeat rewrite <- INR_IZR_INZ; apply le_INR; apply le_max_l.
- rewrite Rmult_1_r; rewrite (Rmult_comm (IZR (Z_of_nat (max x 1))));
+ rewrite Rmult_1_r; rewrite (Rmult_comm (IZR (Z.of_nat (max x 1))));
rewrite Rmult_assoc; rewrite <- Rinv_l_sym.
rewrite Rmult_1_r; repeat rewrite <- INR_IZR_INZ; apply le_INR;
apply le_max_l.
diff --git a/theories/Reals/Rtrigo_fun.v b/theories/Reals/Rtrigo_fun.v
index b77201411..04deb0904 100644
--- a/theories/Reals/Rtrigo_fun.v
+++ b/theories/Reals/Rtrigo_fun.v
@@ -39,7 +39,7 @@ Proof.
in H4; rewrite (let (H1, H2) := Rmult_ne (/ INR (S n)) in H1) in H4;
rewrite (Rmult_comm (/ INR (S n))) in H4;
rewrite (Rmult_assoc eps (/ INR (S n)) (INR (S n))) in H4;
- rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (sym_not_equal (O_S n)))) in H4;
+ rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (not_eq_sym (O_S n)))) in H4;
rewrite (let (H1, H2) := Rmult_ne eps in H1) in H4;
assumption.
apply Rlt_minus; unfold Rgt in a; rewrite <- Rinv_1;
@@ -72,10 +72,10 @@ Proof.
in H6; rewrite (let (H1, H2) := Rmult_ne (/ INR (S n)) in H1) in H6;
rewrite (Rmult_comm (/ INR (S n))) in H6;
rewrite (Rmult_assoc eps (/ INR (S n)) (INR (S n))) in H6;
- rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (sym_not_equal (O_S n)))) in H6;
+ rewrite (Rinv_l (INR (S n)) (not_O_INR (S n) (not_eq_sym (O_S n)))) in H6;
rewrite (let (H1, H2) := Rmult_ne eps in H1) in H6;
assumption.
- cut (IZR (up (/ eps - 1)) = IZR (Z_of_nat x));
+ cut (IZR (up (/ eps - 1)) = IZR (Z.of_nat x));
[ intro | rewrite H1; trivial ].
elim (archimed (/ eps - 1)); intros; clear H6; unfold Rgt in H5;
rewrite H4 in H5; rewrite INR_IZR_INZ; assumption.
@@ -89,11 +89,11 @@ Proof.
generalize (Rmult_lt_compat_l (/ eps) eps 1 (Rinv_0_lt_compat eps H) H0);
rewrite
(Rinv_l eps
- (sym_not_eq (Rlt_dichotomy_converse 0 eps (or_introl (0 > eps) H))))
+ (not_eq_sym (Rlt_dichotomy_converse 0 eps (or_introl (0 > eps) H))))
; rewrite (let (H1, H2) := Rmult_ne (/ eps) in H1);
intro; fold (/ eps - 1 > 0) in |- *; apply Rgt_minus;
unfold Rgt in |- *; assumption.
- right; rewrite H0; rewrite Rinv_1; apply sym_eq; apply Rminus_diag_eq; auto.
+ right; rewrite H0; rewrite Rinv_1; symmetry; apply Rminus_diag_eq; auto.
elim (archimed (/ eps - 1)); intros; clear H1; unfold Rgt in H0; apply Rlt_le;
assumption.
Qed.
diff --git a/theories/Reals/Rtrigo_reg.v b/theories/Reals/Rtrigo_reg.v
index e5befbcf4..0a05e6df4 100644
--- a/theories/Reals/Rtrigo_reg.v
+++ b/theories/Reals/Rtrigo_reg.v
@@ -214,7 +214,7 @@ Proof.
split.
unfold D_x, no_cond in |- *; split.
trivial.
- apply (sym_not_eq (A:=R)); apply H6.
+ apply (not_eq_sym (A:=R)); apply H6.
unfold Rminus in |- *; rewrite Ropp_0; rewrite Rplus_0_r; apply H7.
unfold Boule in |- *; unfold Rminus in |- *; rewrite Ropp_0;
rewrite Rplus_0_r; rewrite Rabs_R0; apply (cond_pos r).
@@ -298,7 +298,7 @@ Proof.
split.
unfold D_x, no_cond in |- *; split.
trivial.
- apply (sym_not_eq (A:=R)); unfold Rdiv in |- *; apply prod_neq_R0.
+ apply (not_eq_sym (A:=R)); unfold Rdiv in |- *; apply prod_neq_R0.
apply H7.
apply Rinv_neq_0_compat; discrR.
apply Rlt_trans with (del_c / 2).
diff --git a/theories/Reals/SeqProp.v b/theories/Reals/SeqProp.v
index 75c57401b..7b6eadc2f 100644
--- a/theories/Reals/SeqProp.v
+++ b/theories/Reals/SeqProp.v
@@ -1131,7 +1131,7 @@ Proof.
rewrite Rinv_mult_distr.
apply Rmult_le_compat_l.
left; apply Rinv_0_lt_compat; apply lt_INR_0; apply neq_O_lt; red in |- *;
- intro; assert (H10 := sym_eq H9); elim (fact_neq_0 _ H10).
+ intro; assert (H10 := eq_sym H9); elim (fact_neq_0 _ H10).
left; apply Rinv_lt_contravar.
apply Rmult_lt_0_compat; apply lt_INR_0; apply lt_O_Sn.
apply lt_INR; apply lt_n_S.
@@ -1156,7 +1156,7 @@ Proof.
replace (Rabs x ^ 1) with (Rabs x); [ idtac | simpl in |- *; ring ].
replace (M_nat + n + 1)%nat with (S (M_nat + n)).
apply Rmult_le_reg_l with (INR (fact (S (M_nat + n)))).
- apply lt_INR_0; apply neq_O_lt; red in |- *; intro; assert (H9 := sym_eq H8);
+ apply lt_INR_0; apply neq_O_lt; red; intro; assert (H9 := eq_sym H8);
elim (fact_neq_0 _ H9).
rewrite (Rmult_comm (Rabs x)); rewrite <- Rmult_assoc; rewrite <- Rinv_r_sym.
rewrite Rmult_1_l.
@@ -1173,7 +1173,7 @@ Proof.
intro; unfold Un in |- *; unfold Rdiv in |- *; apply Rmult_lt_0_compat.
apply pow_lt; assumption.
apply Rinv_0_lt_compat; apply lt_INR_0; apply neq_O_lt; red in |- *; intro;
- assert (H8 := sym_eq H7); elim (fact_neq_0 _ H8).
+ assert (H8 := eq_sym H7); elim (fact_neq_0 _ H8).
clear Un Vn; apply INR_le; simpl in |- *.
induction M_nat as [| M_nat HrecM_nat].
assert (H6 := archimed (Rabs x)); fold M in H6; elim H6; intros.
@@ -1192,7 +1192,7 @@ Proof.
unfold Rdiv in |- *; rewrite Rabs_mult; rewrite (Rabs_right (/ INR (fact n))).
rewrite RPow_abs; right; reflexivity.
apply Rle_ge; left; apply Rinv_0_lt_compat; apply lt_INR_0; apply neq_O_lt;
- red in |- *; intro; assert (H4 := sym_eq H3); elim (fact_neq_0 _ H4).
+ red; intro; assert (H4 := eq_sym H3); elim (fact_neq_0 _ H4).
apply Rle_ge; unfold Rdiv in |- *; apply Rmult_le_pos.
case (Req_dec x 0); intro.
rewrite H3; rewrite Rabs_R0.
@@ -1201,6 +1201,6 @@ Proof.
| simpl in |- *; rewrite Rmult_0_l; right; reflexivity ].
left; apply pow_lt; apply Rabs_pos_lt; assumption.
left; apply Rinv_0_lt_compat; apply lt_INR_0; apply neq_O_lt; red in |- *;
- intro; assert (H4 := sym_eq H3); elim (fact_neq_0 _ H4).
+ intro; assert (H4 := eq_sym H3); elim (fact_neq_0 _ H4).
apply H1; assumption.
Qed.
diff --git a/theories/Reals/Sqrt_reg.v b/theories/Reals/Sqrt_reg.v
index d00ed1786..4e704a274 100644
--- a/theories/Reals/Sqrt_reg.v
+++ b/theories/Reals/Sqrt_reg.v
@@ -250,7 +250,7 @@ Proof.
unfold D_x, no_cond in |- *.
split.
trivial.
- apply (sym_not_eq (A:=R)); exact H8.
+ apply (not_eq_sym (A:=R)); exact H8.
unfold Rminus in |- *; rewrite Ropp_0; rewrite Rplus_0_r;
apply Rlt_le_trans with alpha1.
exact H9.
diff --git a/theories/Relations/Relation_Operators.v b/theories/Relations/Relation_Operators.v
index abf23997d..14aebecce 100644
--- a/theories/Relations/Relation_Operators.v
+++ b/theories/Relations/Relation_Operators.v
@@ -149,13 +149,13 @@ Section Lexicographic_Product.
Variable leA : A -> A -> Prop.
Variable leB : forall x:A, B x -> B x -> Prop.
- Inductive lexprod : sigS B -> sigS B -> Prop :=
+ Inductive lexprod : sigT B -> sigT B -> Prop :=
| left_lex :
forall (x x':A) (y:B x) (y':B x'),
- leA x x' -> lexprod (existS B x y) (existS B x' y')
+ leA x x' -> lexprod (existT B x y) (existT B x' y')
| right_lex :
forall (x:A) (y y':B x),
- leB x y y' -> lexprod (existS B x y) (existS B x y').
+ leB x y y' -> lexprod (existT B x y) (existT B x y').
End Lexicographic_Product.
diff --git a/theories/Sorting/PermutSetoid.v b/theories/Sorting/PermutSetoid.v
index 6f0d76a62..b2b15c705 100644
--- a/theories/Sorting/PermutSetoid.v
+++ b/theories/Sorting/PermutSetoid.v
@@ -71,7 +71,7 @@ Qed.
Lemma permut_sym :
forall l1 l2 : list A, permutation l1 l2 -> permutation l2 l1.
Proof.
- unfold permutation, meq; intros; apply sym_eq; trivial.
+ unfold permutation, meq; intros; symmetry; trivial.
Qed.
Lemma permut_trans :
diff --git a/theories/Strings/Ascii.v b/theories/Strings/Ascii.v
index 1ed9140a5..4f6fe3f86 100644
--- a/theories/Strings/Ascii.v
+++ b/theories/Strings/Ascii.v
@@ -65,7 +65,7 @@ Definition ascii_of_N (n : N) :=
(** Same for [nat] *)
-Definition ascii_of_nat (a : nat) := ascii_of_N (N_of_nat a).
+Definition ascii_of_nat (a : nat) := ascii_of_N (N.of_nat a).
(** The opposite functions *)
@@ -81,7 +81,7 @@ Definition N_of_ascii (a : ascii) : N :=
let (a0,a1,a2,a3,a4,a5,a6,a7) := a in
N_of_digits (a0::a1::a2::a3::a4::a5::a6::a7::nil).
-Definition nat_of_ascii (a : ascii) : nat := nat_of_N (N_of_ascii a).
+Definition nat_of_ascii (a : ascii) : nat := N.to_nat (N_of_ascii a).
(** Proofs that we have indeed opposite function (below 256) *)
@@ -111,10 +111,10 @@ Theorem nat_ascii_embedding :
Proof.
intros. unfold nat_of_ascii, ascii_of_nat.
rewrite N_ascii_embedding.
- apply nat_of_N_of_nat.
- unfold Nlt.
- change 256%N with (N_of_nat 256).
- rewrite <- N_of_nat_compare.
+ apply Nat2N.id.
+ unfold N.lt.
+ change 256%N with (N.of_nat 256).
+ rewrite <- Nat2N.inj_compare.
rewrite <- Compare_dec.nat_compare_lt. auto.
Qed.
diff --git a/theories/Structures/DecidableTypeEx.v b/theories/Structures/DecidableTypeEx.v
index 2c02f8ddd..971fcd7f8 100644
--- a/theories/Structures/DecidableTypeEx.v
+++ b/theories/Structures/DecidableTypeEx.v
@@ -79,9 +79,9 @@ End PairDecidableType.
Module PairUsualDecidableType(D1 D2:UsualDecidableType) <: UsualDecidableType.
Definition t := prod D1.t D2.t.
Definition eq := @eq t.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
+ Definition eq_refl := @eq_refl t.
+ Definition eq_sym := @eq_sym t.
+ Definition eq_trans := @eq_trans t.
Definition eq_dec : forall x y, { eq x y }+{ ~eq x y }.
Proof.
intros (x1,x2) (y1,y2);
diff --git a/theories/Structures/OrderedTypeEx.v b/theories/Structures/OrderedTypeEx.v
index adeba9e48..83130deb2 100644
--- a/theories/Structures/OrderedTypeEx.v
+++ b/theories/Structures/OrderedTypeEx.v
@@ -21,9 +21,9 @@ Module Type UsualOrderedType.
Parameter Inline t : Type.
Definition eq := @eq t.
Parameter Inline lt : t -> t -> Prop.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
+ Definition eq_refl := @eq_refl t.
+ Definition eq_sym := @eq_sym t.
+ Definition eq_trans := @eq_trans t.
Axiom lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
Axiom lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
Parameter compare : forall x y : t, Compare lt eq x y.
@@ -41,9 +41,9 @@ Module Nat_as_OT <: UsualOrderedType.
Definition t := nat.
Definition eq := @eq nat.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
+ Definition eq_refl := @eq_refl t.
+ Definition eq_sym := @eq_sym t.
+ Definition eq_trans := @eq_trans t.
Definition lt := lt.
@@ -53,12 +53,12 @@ Module Nat_as_OT <: UsualOrderedType.
Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
Proof. unfold lt, eq; intros; omega. Qed.
- Definition compare : forall x y : t, Compare lt eq x y.
+ Definition compare x y : Compare lt eq x y.
Proof.
- intros x y; destruct (nat_compare x y) as [ | | ]_eqn.
- apply EQ. apply nat_compare_eq; assumption.
- apply LT. apply nat_compare_Lt_lt; assumption.
- apply GT. apply nat_compare_Gt_gt; assumption.
+ case_eq (nat_compare x y); intro.
+ - apply EQ. now apply nat_compare_eq.
+ - apply LT. now apply nat_compare_Lt_lt.
+ - apply GT. now apply nat_compare_Gt_gt.
Defined.
Definition eq_dec := eq_nat_dec.
@@ -68,15 +68,15 @@ End Nat_as_OT.
(** [Z] is an ordered type with respect to the usual order on integers. *)
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Module Z_as_OT <: UsualOrderedType.
Definition t := Z.
Definition eq := @eq Z.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
+ Definition eq_refl := @eq_refl t.
+ Definition eq_sym := @eq_sym t.
+ Definition eq_trans := @eq_trans t.
Definition lt (x y:Z) := (x<y).
@@ -86,81 +86,73 @@ Module Z_as_OT <: UsualOrderedType.
Lemma lt_not_eq : forall x y, x<y -> ~ x=y.
Proof. intros; omega. Qed.
- Definition compare : forall x y, Compare lt eq x y.
+ Definition compare x y : Compare lt eq x y.
Proof.
- intros x y; destruct (x ?= y) as [ | | ]_eqn.
- apply EQ; apply Zcompare_Eq_eq; assumption.
- apply LT; assumption.
- apply GT; apply Zgt_lt; assumption.
+ case_eq (x ?= y); intro.
+ - apply EQ. now apply Z.compare_eq.
+ - apply LT. assumption.
+ - apply GT. now apply Z.gt_lt.
Defined.
- Definition eq_dec := Z_eq_dec.
+ Definition eq_dec := Z.eq_dec.
End Z_as_OT.
(** [positive] is an ordered type with respect to the usual order on natural numbers. *)
-Open Local Scope positive_scope.
+Local Open Scope positive_scope.
Module Positive_as_OT <: UsualOrderedType.
Definition t:=positive.
Definition eq:=@eq positive.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
+ Definition eq_refl := @eq_refl t.
+ Definition eq_sym := @eq_sym t.
+ Definition eq_trans := @eq_trans t.
- Definition lt := Plt.
+ Definition lt := Pos.lt.
- Definition lt_trans := Plt_trans.
+ Definition lt_trans := Pos.lt_trans.
Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
Proof.
- intros x y H. contradict H. rewrite H. apply Plt_irrefl.
+ intros x y H. contradict H. rewrite H. apply Pos.lt_irrefl.
Qed.
- Definition compare : forall x y : t, Compare lt eq x y.
+ Definition compare x y : Compare lt eq x y.
Proof.
- intros x y. destruct (x ?= y) as [ | | ]_eqn.
- apply EQ; apply Pcompare_Eq_eq; assumption.
- apply LT; assumption.
- apply GT; apply ZC1; assumption.
+ case_eq (x ?= y); intros H.
+ - apply EQ. now apply Pos.compare_eq.
+ - apply LT; assumption.
+ - apply GT. now apply Pos.gt_lt.
Defined.
- Definition eq_dec : forall x y, { eq x y } + { ~ eq x y }.
- Proof.
- intros; unfold eq; decide equality.
- Defined.
+ Definition eq_dec := Pos.eq_dec.
End Positive_as_OT.
(** [N] is an ordered type with respect to the usual order on natural numbers. *)
-Open Local Scope positive_scope.
-
Module N_as_OT <: UsualOrderedType.
Definition t:=N.
Definition eq:=@eq N.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
+ Definition eq_refl := @eq_refl t.
+ Definition eq_sym := @eq_sym t.
+ Definition eq_trans := @eq_trans t.
- Definition lt:=Nlt.
- Definition lt_trans := Nlt_trans.
- Definition lt_not_eq := Nlt_not_eq.
+ Definition lt := N.lt.
+ Definition lt_trans := N.lt_trans.
+ Definition lt_not_eq := N.lt_neq.
- Definition compare : forall x y : t, Compare lt eq x y.
+ Definition compare x y : Compare lt eq x y.
Proof.
- intros x y. destruct (x ?= y)%N as [ | | ]_eqn.
- apply EQ; apply Ncompare_Eq_eq; assumption.
- apply LT; assumption.
- apply GT. apply Ngt_Nlt; assumption.
+ case_eq (x ?= y)%N; intro.
+ - apply EQ. now apply N.compare_eq.
+ - apply LT. assumption.
+ - apply GT. now apply N.gt_lt.
Defined.
- Definition eq_dec : forall x y, { eq x y } + { ~ eq x y }.
- Proof.
- intros. unfold eq. decide equality. apply Positive_as_OT.eq_dec.
- Defined.
+ Definition eq_dec := N.eq_dec.
End N_as_OT.
@@ -240,9 +232,9 @@ End PairOrderedType.
Module PositiveOrderedTypeBits <: UsualOrderedType.
Definition t:=positive.
Definition eq:=@eq positive.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
+ Definition eq_refl := @eq_refl t.
+ Definition eq_sym := @eq_sym t.
+ Definition eq_trans := @eq_trans t.
Fixpoint bits_lt (p q:positive) : Prop :=
match p, q with
@@ -286,38 +278,38 @@ Module PositiveOrderedTypeBits <: UsualOrderedType.
Definition compare : forall x y : t, Compare lt eq x y.
Proof.
induction x; destruct y.
- (* I I *)
- destruct (IHx y).
- apply LT; auto.
- apply EQ; rewrite e; red; auto.
- apply GT; auto.
- (* I O *)
- apply GT; simpl; auto.
- (* I H *)
- apply GT; simpl; auto.
- (* O I *)
- apply LT; simpl; auto.
- (* O O *)
- destruct (IHx y).
- apply LT; auto.
- apply EQ; rewrite e; red; auto.
- apply GT; auto.
- (* O H *)
- apply LT; simpl; auto.
- (* H I *)
- apply LT; simpl; auto.
- (* H O *)
- apply GT; simpl; auto.
- (* H H *)
- apply EQ; red; auto.
+ - (* I I *)
+ destruct (IHx y).
+ apply LT; auto.
+ apply EQ; rewrite e; red; auto.
+ apply GT; auto.
+ - (* I O *)
+ apply GT; simpl; auto.
+ - (* I H *)
+ apply GT; simpl; auto.
+ - (* O I *)
+ apply LT; simpl; auto.
+ - (* O O *)
+ destruct (IHx y).
+ apply LT; auto.
+ apply EQ; rewrite e; red; auto.
+ apply GT; auto.
+ - (* O H *)
+ apply LT; simpl; auto.
+ - (* H I *)
+ apply LT; simpl; auto.
+ - (* H O *)
+ apply GT; simpl; auto.
+ - (* H H *)
+ apply EQ; red; auto.
Qed.
Lemma eq_dec (x y: positive): {x = y} + {x <> y}.
Proof.
intros. case_eq (x ?= y); intros.
- left. apply Pcompare_Eq_eq; auto.
- right. red. intro. subst y. rewrite (Pos.compare_refl x) in H. discriminate.
- right. red. intro. subst y. rewrite (Pos.compare_refl x) in H. discriminate.
+ - left. now apply Pos.compare_eq.
+ - right. intro. subst y. now rewrite (Pos.compare_refl x) in *.
+ - right. intro. subst y. now rewrite (Pos.compare_refl x) in *.
Qed.
End PositiveOrderedTypeBits.
diff --git a/theories/Unicode/Utf8.v b/theories/Unicode/Utf8.v
index 86ab47768..db7ff08c7 100644
--- a/theories/Unicode/Utf8.v
+++ b/theories/Unicode/Utf8.v
@@ -20,5 +20,5 @@ Check ∀ x z, True -> (∃ y v, x + v ≥ y + z) ∨ x ≤ 0.
(* Integer Arithmetic *)
(* TODO: this should come after ZArith
-Notation "x ≤ y" := (Zle x y) (at level 70, no associativity).
+Notation "x ≤ y" := (Z.le x y) (at level 70, no associativity).
*)
diff --git a/theories/Wellfounded/Lexicographic_Product.v b/theories/Wellfounded/Lexicographic_Product.v
index 0e0961004..9a1d66f43 100644
--- a/theories/Wellfounded/Lexicographic_Product.v
+++ b/theories/Wellfounded/Lexicographic_Product.v
@@ -27,7 +27,7 @@ Section WfLexicographic_Product.
forall x:A,
Acc leA x ->
(forall x0:A, clos_trans A leA x0 x -> well_founded (leB x0)) ->
- forall y:B x, Acc (leB x) y -> Acc LexProd (existS B x y).
+ forall y:B x, Acc (leB x) y -> Acc LexProd (existT B x y).
Proof.
induction 1 as [x _ IHAcc]; intros H2 y.
induction 1 as [x0 H IHAcc0]; intros.
diff --git a/theories/ZArith/BinInt.v b/theories/ZArith/BinInt.v
index 0f709d686..a5c8affe7 100644
--- a/theories/ZArith/BinInt.v
+++ b/theories/ZArith/BinInt.v
@@ -82,8 +82,8 @@ Lemma pos_sub_spec p q :
pos_sub p q =
match (p ?= q)%positive with
| Eq => 0
- | Lt => Zneg (q - p)
- | Gt => Zpos (p - q)
+ | Lt => neg (q - p)
+ | Gt => pos (p - q)
end.
Proof.
revert q. induction p; destruct q; simpl; trivial;
@@ -95,6 +95,18 @@ Proof.
subst; unfold Pos.sub; simpl; now rewrite Pos.sub_mask_diag.
Qed.
+Lemma pos_sub_discr p q :
+ match pos_sub p q with
+ | Z0 => p = q
+ | pos k => p = q + k
+ | neg k => q = p + k
+ end%positive.
+Proof.
+ rewrite pos_sub_spec.
+ case Pos.compare_spec; auto; intros;
+ now rewrite Pos.add_comm, Pos.sub_add.
+Qed.
+
(** Particular cases of the previous result *)
Lemma pos_sub_diag p : pos_sub p p = 0.
@@ -102,12 +114,12 @@ Proof.
now rewrite pos_sub_spec, Pos.compare_refl.
Qed.
-Lemma pos_sub_lt p q : (p < q)%positive -> pos_sub p q = Zneg (q - p).
+Lemma pos_sub_lt p q : (p < q)%positive -> pos_sub p q = neg (q - p).
Proof.
intros H. now rewrite pos_sub_spec, H.
Qed.
-Lemma pos_sub_gt p q : (q < p)%positive -> pos_sub p q = Zpos (p - q).
+Lemma pos_sub_gt p q : (q < p)%positive -> pos_sub p q = pos (p - q).
Proof.
intros H. now rewrite pos_sub_spec, Pos.compare_antisym, H.
Qed.
@@ -120,89 +132,6 @@ Proof.
rewrite <- IHp; now destruct pos_sub.
Qed.
-(** * Results concerning [Zpos] and [Zneg] and the operators *)
-
-Lemma opp_Zneg p : - Zneg p = Zpos p.
-Proof.
- reflexivity.
-Qed.
-
-Lemma opp_Zpos p : - Zpos p = Zneg p.
-Proof.
- reflexivity.
-Qed.
-
-Lemma succ_Zpos p : succ (Zpos p) = Zpos (Pos.succ p).
-Proof.
- simpl. f_equal. apply Pos.add_1_r.
-Qed.
-
-Lemma add_Zpos p q : Zpos p + Zpos q = Zpos (p+q).
-Proof.
- reflexivity.
-Qed.
-
-Lemma add_Zneg p q : Zneg p + Zneg q = Zneg (p+q).
-Proof.
- reflexivity.
-Qed.
-
-Lemma add_Zpos_Zneg p q : Zpos p + Zneg q = pos_sub p q.
-Proof.
- reflexivity.
-Qed.
-
-Lemma add_Zneg_Zpos p q : Zneg p + Zpos q = pos_sub q p.
-Proof.
- reflexivity.
-Qed.
-
-Lemma sub_Zpos n m : (n < m)%positive -> Zpos m - Zpos n = Zpos (m-n).
-Proof.
- intros H. simpl. now apply pos_sub_gt.
-Qed.
-
-Lemma mul_Zpos (p q : positive) : Zpos p * Zpos q = Zpos (p*q).
-Proof.
- reflexivity.
-Qed.
-
-Lemma pow_Zpos p q : (Zpos p)^(Zpos q) = Zpos (p^q).
-Proof.
- unfold Pos.pow, pow, pow_pos.
- symmetry. now apply Pos.iter_swap_gen.
-Qed.
-
-Lemma inj_Zpos p q : Zpos p = Zpos q <-> p = q.
-Proof.
- split; intros H. now injection H. now f_equal.
-Qed.
-
-Lemma inj_Zneg p q : Zneg p = Zneg q <-> p = q.
-Proof.
- split; intros H. now injection H. now f_equal.
-Qed.
-
-Lemma pos_xI p : Zpos p~1 = 2 * Zpos p + 1.
-Proof.
- reflexivity.
-Qed.
-
-Lemma pos_xO p : Zpos p~0 = 2 * Zpos p.
-Proof.
- reflexivity.
-Qed.
-
-Lemma neg_xI p : Zneg p~1 = 2 * Zneg p - 1.
-Proof.
- reflexivity.
-Qed.
-
-Lemma neg_xO p : Zneg p~0 = 2 * Zneg p.
-Proof.
- reflexivity.
-Qed.
-
(** In the following module, we group results that are needed now
to prove specifications of operations, but will also be provided
later by the generic functor of properties. *)
@@ -242,7 +171,7 @@ Qed.
(** ** Addition is associative *)
Lemma pos_sub_add p q r :
- pos_sub (p + q) r = Zpos p + pos_sub q r.
+ pos_sub (p + q) r = pos p + pos_sub q r.
Proof.
simpl. rewrite !pos_sub_spec.
case (Pos.compare_spec q r); intros E0.
@@ -269,19 +198,19 @@ Qed.
Lemma add_assoc n m p : n + (m + p) = n + m + p.
Proof.
- assert (AUX : forall x y z, Zpos x + (y + z) = Zpos x + y + z).
+ assert (AUX : forall x y z, pos x + (y + z) = pos x + y + z).
{ intros x [|y|y] [|z|z]; rewrite ?add_0_r; trivial.
- simpl. now rewrite Pos.add_assoc.
- - simpl (_ + Zneg _). symmetry. apply pos_sub_add.
- - simpl (Zneg _ + _); simpl (_ + Zneg _).
- now rewrite (add_comm _ (Zpos _)), <- 2 pos_sub_add, Pos.add_comm.
- - apply opp_inj. rewrite !opp_add_distr, opp_Zpos, !opp_Zneg.
- simpl (Zneg _ + _); simpl (_ + Zneg _).
+ - simpl (_ + neg _). symmetry. apply pos_sub_add.
+ - simpl (neg _ + _); simpl (_ + neg _).
+ now rewrite (add_comm _ (pos _)), <- 2 pos_sub_add, Pos.add_comm.
+ - apply opp_inj. rewrite !opp_add_distr. simpl opp.
+ simpl (neg _ + _); simpl (_ + neg _).
rewrite add_comm, Pos.add_comm. apply pos_sub_add. }
destruct n.
- trivial.
- apply AUX.
- - apply opp_inj. rewrite !opp_add_distr, opp_Zneg. apply AUX.
+ - apply opp_inj. rewrite !opp_add_distr. simpl opp. apply AUX.
Qed.
(** ** Subtraction and successor *)
@@ -354,7 +283,7 @@ Qed.
(** ** Distributivity of multiplication over addition *)
Lemma mul_add_distr_pos (p:positive) n m :
- Zpos p * (n + m) = Zpos p * n + Zpos p * m.
+ pos p * (n + m) = pos p * n + pos p * m.
Proof.
destruct n as [|n|n], m as [|m|m]; simpl; trivial;
rewrite ?pos_sub_spec, ?Pos.mul_compare_mono_l; try case Pos.compare_spec;
@@ -365,7 +294,8 @@ Lemma mul_add_distr_l n m p : n * (m + p) = n * m + n * p.
Proof.
destruct n as [|n|n]. trivial.
apply mul_add_distr_pos.
- rewrite <- opp_Zpos, !mul_opp_l, <- opp_add_distr. f_equal.
+ change (neg n) with (- pos n).
+ rewrite !mul_opp_l, <- opp_add_distr. f_equal.
apply mul_add_distr_pos.
Qed.
@@ -374,6 +304,57 @@ Proof.
rewrite !(mul_comm _ p). apply mul_add_distr_l.
Qed.
+(** ** Basic properties of divisibility *)
+
+Lemma divide_Zpos p q : (pos p|pos q) <-> (p|q)%positive.
+Proof.
+ split.
+ intros ([ |r|r],H); simpl in *; destr_eq H. exists r; auto.
+ intros (r,H). exists (pos r); simpl; now f_equal.
+Qed.
+
+Lemma divide_Zpos_Zneg_r n p : (n|pos p) <-> (n|neg p).
+Proof.
+ split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- H.
+Qed.
+
+Lemma divide_Zpos_Zneg_l n p : (pos p|n) <-> (neg p|n).
+Proof.
+ split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- mul_opp_r.
+Qed.
+
+(** ** Conversions between [Z.testbit] and [N.testbit] *)
+
+Lemma testbit_of_N a n :
+ testbit (of_N a) (of_N n) = N.testbit a n.
+Proof.
+ destruct a as [|a], n; simpl; trivial. now destruct a.
+Qed.
+
+Lemma testbit_of_N' a n : 0<=n ->
+ testbit (of_N a) n = N.testbit a (to_N n).
+Proof.
+ intro Hn. rewrite <- testbit_of_N. f_equal.
+ destruct n; trivial; now destruct Hn.
+Qed.
+
+Lemma testbit_Zpos a n : 0<=n ->
+ testbit (pos a) n = N.testbit (N.pos a) (to_N n).
+Proof.
+ intro Hn. now rewrite <- testbit_of_N'.
+Qed.
+
+Lemma testbit_Zneg a n : 0<=n ->
+ testbit (neg a) n = negb (N.testbit (Pos.pred_N a) (to_N n)).
+Proof.
+ intro Hn.
+ rewrite <- testbit_of_N' by trivial.
+ destruct n as [ |n|n];
+ [ | simpl; now destruct (Pos.pred_N a) | now destruct Hn].
+ unfold testbit.
+ now destruct a as [|[ | | ]| ].
+Qed.
+
End Private_BootStrap.
(** * Proofs of specifications *)
@@ -454,9 +435,8 @@ Qed.
Lemma eqb_eq n m : (n =? m) = true <-> n = m.
Proof.
- destruct n, m; simpl; try (now split).
- rewrite inj_Zpos. apply Pos.eqb_eq.
- rewrite inj_Zneg. apply Pos.eqb_eq.
+ destruct n, m; simpl; try (now split); rewrite Pos.eqb_eq;
+ split; (now injection 1) || (intros; now f_equal).
Qed.
Lemma ltb_lt n m : (n <? m) = true <-> n < m.
@@ -580,7 +560,7 @@ Qed.
(** For folding back a [pow_pos] into a [pow] *)
-Lemma pow_pos_fold n p : pow_pos n p = n ^ (Zpos p).
+Lemma pow_pos_fold n p : pow_pos n p = n ^ (pos p).
Proof.
reflexivity.
Qed.
@@ -607,7 +587,7 @@ Lemma sqrt_spec n : 0<=n ->
let s := sqrt n in s*s <= n < (succ s)*(succ s).
Proof.
destruct n. now repeat split. unfold sqrt.
- rewrite succ_Zpos. intros _. apply (Pos.sqrt_spec p).
+ intros _. simpl succ. rewrite Pos.add_1_r. apply (Pos.sqrt_spec p).
now destruct 1.
Qed.
@@ -627,8 +607,10 @@ Qed.
Lemma log2_spec n : 0 < n -> 2^(log2 n) <= n < 2^(succ (log2 n)).
Proof.
+ assert (Pow : forall p q, pos (p^q) = (pos p)^(pos q)).
+ { intros. now apply Pos.iter_swap_gen. }
destruct n as [|[p|p|]|]; intros Hn; split; try easy; unfold log2;
- rewrite ?succ_Zpos, pow_Zpos.
+ simpl succ; rewrite ?Pos.add_1_r, <- Pow.
change (2^Pos.size p <= Pos.succ (p~0))%positive.
apply Pos.lt_le_incl, Pos.lt_succ_r, Pos.size_le.
apply Pos.size_gt.
@@ -678,20 +660,22 @@ Qed.
(** ** Correctness proofs for Trunc division *)
Lemma pos_div_eucl_eq a b : 0 < b ->
- let (q, r) := pos_div_eucl a b in Zpos a = q * b + r.
+ let (q, r) := pos_div_eucl a b in pos a = q * b + r.
Proof.
intros Hb.
induction a; unfold pos_div_eucl; fold pos_div_eucl.
- (* ~1 *)
destruct pos_div_eucl as (q,r).
- rewrite pos_xI, IHa, mul_add_distr_l, mul_assoc.
+ change (pos a~1) with (2*(pos a)+1).
+ rewrite IHa, mul_add_distr_l, mul_assoc.
destruct ltb.
now rewrite add_assoc.
rewrite mul_add_distr_r, mul_1_l, <- !add_assoc. f_equal.
unfold sub. now rewrite (add_comm _ (-b)), add_assoc, add_opp_diag_r.
- (* ~0 *)
destruct pos_div_eucl as (q,r).
- rewrite (pos_xO a), IHa, mul_add_distr_l, mul_assoc.
+ change (pos a~0) with (2*pos a).
+ rewrite IHa, mul_add_distr_l, mul_assoc.
destruct ltb.
trivial.
rewrite mul_add_distr_r, mul_1_l, <- !add_assoc. f_equal.
@@ -709,21 +693,23 @@ Lemma div_eucl_eq a b : b<>0 ->
Proof.
destruct a as [ |a|a], b as [ |b|b]; unfold div_eucl; trivial;
(now destruct 1) || intros _;
- generalize (pos_div_eucl_eq a (Zpos b) (eq_refl _));
- destruct pos_div_eucl as (q,r); rewrite <- ?opp_Zpos, mul_comm;
- intros ->.
- - (* Zpos Zpos *)
+ generalize (pos_div_eucl_eq a (pos b) (eq_refl _));
+ destruct pos_div_eucl as (q,r); rewrite mul_comm.
+ - (* pos pos *)
trivial.
- - (* Zpos Zneg *)
- destruct r as [ |r|r]; rewrite !mul_opp_opp; trivial;
+ - (* pos neg *)
+ intros ->.
+ destruct r as [ |r|r]; rewrite <- !mul_opp_comm; trivial;
rewrite mul_add_distr_l, mul_1_r, <- add_assoc; f_equal;
now rewrite add_assoc, add_opp_diag_r.
- - (* Zneg Zpos *)
+ - (* neg pos *)
+ change (neg a) with (- pos a). intros ->.
rewrite (opp_add_distr _ r), <- mul_opp_r.
destruct r as [ |r|r]; trivial;
rewrite opp_add_distr, mul_add_distr_l, <- add_assoc; f_equal;
unfold sub; now rewrite add_assoc, mul_opp_r, mul_1_r, add_opp_diag_l.
- - (* Zneg Zneg *)
+ - (* neg neg *)
+ change (neg a) with (- pos a). intros ->.
now rewrite opp_add_distr, <- mul_opp_l.
Qed.
@@ -735,10 +721,10 @@ Qed.
Lemma pos_div_eucl_bound a b : 0<b -> 0 <= snd (pos_div_eucl a b) < b.
Proof.
- assert (AUX : forall m p, m < Zpos (p~0) -> m - Zpos p < Zpos p).
+ assert (AUX : forall m p, m < pos (p~0) -> m - pos p < pos p).
intros m p. unfold lt.
- rewrite (compare_sub m), (compare_sub _ (Zpos _)). unfold sub.
- rewrite <- add_assoc. simpl opp; simpl (Zneg _ + _).
+ rewrite (compare_sub m), (compare_sub _ (pos _)). unfold sub.
+ rewrite <- add_assoc. simpl opp; simpl (neg _ + _).
now rewrite Pos.add_diag.
intros Hb.
destruct b as [|b|b]; discriminate Hb || clear Hb.
@@ -770,7 +756,7 @@ Proof.
destruct a as [|a|a]; unfold modulo, div_eucl.
now split.
now apply pos_div_eucl_bound.
- generalize (pos_div_eucl_bound a (Zpos b) (eq_refl _)).
+ generalize (pos_div_eucl_bound a (pos b) (eq_refl _)).
destruct pos_div_eucl as (q,r); unfold snd; intros (Hr,Hr').
destruct r as [|r|r]; (now destruct Hr) || clear Hr.
now split.
@@ -787,17 +773,17 @@ Proof.
destruct b as [|b|b]; try easy; intros _.
destruct a as [|a|a]; unfold modulo, div_eucl.
now split.
- generalize (pos_div_eucl_bound a (Zpos b) (eq_refl _)).
+ generalize (pos_div_eucl_bound a (pos b) (eq_refl _)).
destruct pos_div_eucl as (q,r); unfold snd; intros (Hr,Hr').
destruct r as [|r|r]; (now destruct Hr) || clear Hr.
now split.
split.
unfold lt in *; simpl in *. rewrite pos_sub_lt by trivial.
rewrite <- Pos.compare_antisym. now apply Pos.sub_decr.
- change (Zneg b - Zneg r <= 0). unfold le, lt in *.
+ change (neg b - neg r <= 0). unfold le, lt in *.
rewrite <- compare_sub. simpl in *.
now rewrite <- Pos.compare_antisym, Hr'.
- generalize (pos_div_eucl_bound a (Zpos b) (eq_refl _)).
+ generalize (pos_div_eucl_bound a (pos b) (eq_refl _)).
destruct pos_div_eucl as (q,r); unfold snd; intros (Hr,Hr').
split; destruct r; try easy.
red; simpl; now rewrite <- Pos.compare_antisym.
@@ -808,9 +794,10 @@ Qed.
Theorem quotrem_eq a b : let (q,r) := quotrem a b in a = q * b + r.
Proof.
destruct a as [|a|a], b as [|b|b]; simpl; trivial;
- generalize (N.pos_div_eucl_spec a (Npos b)); case N.pos_div_eucl; trivial;
- intros q r; rewrite <- ?opp_Zpos;
- change (Zpos a) with (of_N (Npos a)); intros ->; now destruct q, r.
+ generalize (N.pos_div_eucl_spec a (N.pos b)); case N.pos_div_eucl; trivial;
+ intros q r;
+ try change (neg a) with (-pos a);
+ change (pos a) with (of_N (N.pos a)); intros ->; now destruct q, r.
Qed.
Lemma quot_rem' a b : a = b*(a÷b) + rem a b.
@@ -829,7 +816,7 @@ Proof.
destruct a as [|a|a]; (now destruct Ha) || clear Ha.
compute. now split.
unfold rem, quotrem.
- assert (H := N.pos_div_eucl_remainder a (Npos b)).
+ assert (H := N.pos_div_eucl_remainder a (N.pos b)).
destruct N.pos_div_eucl as (q,[|r]); simpl; split; try easy.
now apply H.
Qed.
@@ -852,25 +839,6 @@ Proof. intros _. apply rem_opp_l'. Qed.
Lemma rem_opp_r a b : b<>0 -> rem a (-b) = rem a b.
Proof. intros _. apply rem_opp_r'. Qed.
-(** ** Basic properties of divisibility *)
-
-Lemma divide_Zpos p q : (Zpos p|Zpos q) <-> (p|q)%positive.
-Proof.
- split.
- intros ([ |r|r],H); simpl in *; destr_eq H. exists r; auto.
- intros (r,H). exists (Zpos r); simpl; now f_equal.
-Qed.
-
-Lemma divide_Zpos_Zneg_r n p : (n|Zpos p) <-> (n|Zneg p).
-Proof.
- split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- H.
-Qed.
-
-Lemma divide_Zpos_Zneg_l n p : (Zpos p|n) <-> (Zneg p|n).
-Proof.
- split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- mul_opp_r.
-Qed.
-
(** ** Correctness proofs for gcd *)
Lemma ggcd_gcd a b : fst (ggcd a b) = gcd a b.
@@ -905,7 +873,7 @@ Qed.
Lemma gcd_greatest a b c : (c|a) -> (c|b) -> (c | gcd a b).
Proof.
- assert (H : forall p q r, (r|Zpos p) -> (r|Zpos q) -> (r|Zpos (Pos.gcd p q))).
+ assert (H : forall p q r, (r|pos p) -> (r|pos q) -> (r|pos (Pos.gcd p q))).
{ intros p q [|r|r] H H'.
destruct H; now rewrite mul_comm in *.
apply divide_Zpos, Pos.gcd_greatest; now apply divide_Zpos.
@@ -930,38 +898,6 @@ Proof.
destruct (Pos.ggcd a b) as (g,(aa,bb)); auto.
Qed.
-(** ** Conversions between [Z.testbit] and [N.testbit] *)
-
-Lemma testbit_of_N a n :
- testbit (of_N a) (of_N n) = N.testbit a n.
-Proof.
- destruct a as [|a], n; simpl; trivial. now destruct a.
-Qed.
-
-Lemma testbit_of_N' a n : 0<=n ->
- testbit (of_N a) n = N.testbit a (to_N n).
-Proof.
- intro Hn. rewrite <- testbit_of_N. f_equal.
- destruct n; trivial; now destruct Hn.
-Qed.
-
-Lemma testbit_Zpos a n : 0<=n ->
- testbit (Zpos a) n = N.testbit (Npos a) (to_N n).
-Proof.
- intro Hn. now rewrite <- testbit_of_N'.
-Qed.
-
-Lemma testbit_Zneg a n : 0<=n ->
- testbit (Zneg a) n = negb (N.testbit (Pos.pred_N a) (to_N n)).
-Proof.
- intro Hn.
- rewrite <- testbit_of_N' by trivial.
- destruct n as [ |n|n];
- [ | simpl; now destruct (Ppred_N a) | now destruct Hn].
- unfold testbit.
- now destruct a as [|[ | | ]| ].
-Qed.
-
(** ** Proofs of specifications for bitwise operations *)
Lemma div2_spec a : div2 a = shiftr a 1.
@@ -994,9 +930,9 @@ Lemma testbit_odd_succ a n : 0<=n ->
Proof.
destruct n as [|n|n]; (now destruct 1) || intros _.
destruct a as [|[a|a|]|[a|a|]]; simpl; trivial. now destruct a.
- unfold testbit. rewrite succ_Zpos.
+ unfold testbit; simpl.
destruct a as [|a|[a|a|]]; simpl; trivial;
- rewrite ?Pos.pred_N_succ; now destruct n.
+ rewrite ?Pos.add_1_r, ?Pos.pred_N_succ; now destruct n.
Qed.
Lemma testbit_even_succ a n : 0<=n ->
@@ -1004,9 +940,9 @@ Lemma testbit_even_succ a n : 0<=n ->
Proof.
destruct n as [|n|n]; (now destruct 1) || intros _.
destruct a as [|[a|a|]|[a|a|]]; simpl; trivial. now destruct a.
- unfold testbit. rewrite succ_Zpos.
+ unfold testbit; simpl.
destruct a as [|a|[a|a|]]; simpl; trivial;
- rewrite ?Pos.pred_N_succ; now destruct n.
+ rewrite ?Pos.add_1_r, ?Pos.pred_N_succ; now destruct n.
Qed.
(** Correctness proofs about [Z.shiftr] and [Z.shiftl] *)
@@ -1017,9 +953,9 @@ Proof.
intros Hn Hm. unfold shiftr.
destruct n as [ |n|n]; (now destruct Hn) || clear Hn; simpl.
now rewrite add_0_r.
- assert (forall p, to_N (m + Zpos p) = (to_N m + Npos p)%N).
+ assert (forall p, to_N (m + pos p) = (to_N m + N.pos p)%N).
destruct m; trivial; now destruct Hm.
- assert (forall p, 0 <= m + Zpos p).
+ assert (forall p, 0 <= m + pos p).
destruct m; easy || now destruct Hm.
destruct a as [ |a|a].
(* a = 0 *)
@@ -1027,15 +963,15 @@ Proof.
by (apply Pos.iter_invariant; intros; subst; trivial).
now rewrite 2 testbit_0_l.
(* a > 0 *)
- change (Zpos a) with (of_N (Npos a)) at 1.
- rewrite <- (Pos.iter_swap_gen _ _ _ Ndiv2) by now intros [|[ | | ]].
+ change (pos a) with (of_N (N.pos a)) at 1.
+ rewrite <- (Pos.iter_swap_gen _ _ _ N.div2) by now intros [|[ | | ]].
rewrite testbit_Zpos, testbit_of_N', H; trivial.
- exact (N.shiftr_spec' (Npos a) (Npos n) (to_N m)).
+ exact (N.shiftr_spec' (N.pos a) (N.pos n) (to_N m)).
(* a < 0 *)
- rewrite <- (Pos.iter_swap_gen _ _ _ Pdiv2_up) by trivial.
+ rewrite <- (Pos.iter_swap_gen _ _ _ Pos.div2_up) by trivial.
rewrite 2 testbit_Zneg, H; trivial. f_equal.
- rewrite (Pos.iter_swap_gen _ _ _ _ Ndiv2) by exact N.pred_div2_up.
- exact (N.shiftr_spec' (Ppred_N a) (Npos n) (to_N m)).
+ rewrite (Pos.iter_swap_gen _ _ _ _ N.div2) by exact N.pred_div2_up.
+ exact (N.shiftr_spec' (Pos.pred_N a) (N.pos n) (to_N m)).
Qed.
Lemma shiftl_spec_low a n m : m<n ->
@@ -1052,11 +988,11 @@ Proof.
(* a > 0 *)
rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
rewrite testbit_Zpos by easy.
- exact (N.shiftl_spec_low (Npos a) (Npos n) (Npos m) H).
+ exact (N.shiftl_spec_low (N.pos a) (N.pos n) (N.pos m) H).
(* a < 0 *)
rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
rewrite testbit_Zneg by easy.
- now rewrite (N.pos_pred_shiftl_low a (Npos n)).
+ now rewrite (N.pos_pred_shiftl_low a (N.pos n)).
Qed.
Lemma shiftl_spec_high a n m : 0<=m -> n<=m ->
@@ -1066,9 +1002,9 @@ Proof.
destruct n as [ |n|n]. simpl. now rewrite sub_0_r.
(* n > 0 *)
destruct m as [ |m|m]; try (now destruct H).
- assert (0 <= Zpos m - Zpos n).
+ assert (0 <= pos m - pos n).
red. now rewrite compare_antisym, <- compare_sub, <- compare_antisym.
- assert (EQ : to_N (Zpos m - Zpos n) = (Npos m - Npos n)%N).
+ assert (EQ : to_N (pos m - pos n) = (N.pos m - N.pos n)%N).
red in H. simpl in H. simpl to_N.
rewrite pos_sub_spec, Pos.compare_antisym.
destruct (Pos.compare_spec n m) as [H'|H'|H']; try (now destruct H).
@@ -1083,16 +1019,16 @@ Proof.
(* ... a > 0 *)
rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
rewrite 2 testbit_Zpos, EQ by easy.
- exact (N.shiftl_spec_high' (Npos p) (Npos n) (Npos m) H).
+ exact (N.shiftl_spec_high' (N.pos p) (N.pos n) (N.pos m) H).
(* ... a < 0 *)
rewrite <- (Pos.iter_swap_gen _ _ _ xO) by trivial.
rewrite 2 testbit_Zneg, EQ by easy. f_equal.
simpl to_N.
rewrite <- N.shiftl_spec_high by easy.
- now apply (N.pos_pred_shiftl_high p (Npos n)).
+ now apply (N.pos_pred_shiftl_high p (N.pos n)).
(* n < 0 *)
unfold sub. simpl.
- now apply (shiftr_spec_aux a (Zpos n) m).
+ now apply (shiftr_spec_aux a (pos n) m).
Qed.
Lemma shiftr_spec a n m : 0<=m ->
@@ -1180,11 +1116,11 @@ Proof.
induction p using Pos.peano_ind.
now apply (Hs 0).
rewrite <- Pos.add_1_r.
- now apply (Hs (Zpos p)).
+ now apply (Hs (pos p)).
induction p using Pos.peano_ind.
now apply (Hp 0).
rewrite <- Pos.add_1_r.
- now apply (Hp (Zneg p)).
+ now apply (Hp (neg p)).
Qed.
Lemma bi_induction (P : Z -> Prop) :
@@ -1341,7 +1277,7 @@ Qed.
End Z.
-(** Export Notations *)
+(** Re-export Notations *)
Infix "+" := Z.add : Z_scope.
Notation "- x" := (Z.opp x) : Z_scope.
@@ -1351,27 +1287,271 @@ Infix "^" := Z.pow : Z_scope.
Infix "/" := Z.div : Z_scope.
Infix "mod" := Z.modulo (at level 40, no associativity) : Z_scope.
Infix "÷" := Z.quot (at level 40, left associativity) : Z_scope.
-
-(* TODO : transition from Zdivide *)
-Notation "( x | y )" := (Z.divide x y) (at level 0) : Z_scope.
-
Infix "?=" := Z.compare (at level 70, no associativity) : Z_scope.
-
+Infix "=?" := Z.eqb (at level 70, no associativity) : Z_scope.
+Infix "<=?" := Z.leb (at level 70, no associativity) : Z_scope.
+Infix "<?" := Z.ltb (at level 70, no associativity) : Z_scope.
+Infix ">=?" := Z.geb (at level 70, no associativity) : Z_scope.
+Infix ">?" := Z.gtb (at level 70, no associativity) : Z_scope.
+Notation "( x | y )" := (Z.divide x y) (at level 0) : Z_scope.
Infix "<=" := Z.le : Z_scope.
Infix "<" := Z.lt : Z_scope.
Infix ">=" := Z.ge : Z_scope.
Infix ">" := Z.gt : Z_scope.
-
Notation "x <= y <= z" := (x <= y /\ y <= z) : Z_scope.
Notation "x <= y < z" := (x <= y /\ y < z) : Z_scope.
Notation "x < y < z" := (x < y /\ y < z) : Z_scope.
Notation "x < y <= z" := (x < y /\ y <= z) : Z_scope.
-Infix "=?" := Z.eqb (at level 70, no associativity) : Z_scope.
-Infix "<=?" := Z.leb (at level 70, no associativity) : Z_scope.
-Infix "<?" := Z.ltb (at level 70, no associativity) : Z_scope.
-Infix ">=?" := Z.geb (at level 70, no associativity) : Z_scope.
-Infix ">?" := Z.gtb (at level 70, no associativity) : Z_scope.
+(** Conversions from / to positive numbers *)
+
+Module Pos2Z.
+
+Lemma id p : Z.to_pos (Z.pos p) = p.
+Proof. reflexivity. Qed.
+
+Lemma inj p q : Z.pos p = Z.pos q -> p = q.
+Proof. now injection 1. Qed.
+
+Lemma inj_iff p q : Z.pos p = Z.pos q <-> p = q.
+Proof. split. apply inj. intros; now f_equal. Qed.
+
+Lemma is_pos p : 0 < Z.pos p.
+Proof. reflexivity. Qed.
+
+Lemma is_nonneg p : 0 <= Z.pos p.
+Proof. easy. Qed.
+
+Lemma inj_1 : Z.pos 1 = 1.
+Proof. reflexivity. Qed.
+
+Lemma inj_xO p : Z.pos p~0 = 2 * Z.pos p.
+Proof. reflexivity. Qed.
+
+Lemma inj_xI p : Z.pos p~1 = 2 * Z.pos p + 1.
+Proof. reflexivity. Qed.
+
+Lemma inj_succ p : Z.pos (Pos.succ p) = Z.succ (Z.pos p).
+Proof. simpl. now rewrite Pos.add_1_r. Qed.
+
+Lemma inj_add p q : Z.pos (p+q) = Z.pos p + Z.pos q.
+Proof. reflexivity. Qed.
+
+Lemma inj_sub p q : (p < q)%positive ->
+ Z.pos (q-p) = Z.pos q - Z.pos p.
+Proof. intros. simpl. now rewrite Z.pos_sub_gt. Qed.
+
+Lemma inj_sub_max p q : Z.pos (p - q) = Z.max 1 (Z.pos p - Z.pos q).
+Proof.
+ simpl. rewrite Z.pos_sub_spec. case Pos.compare_spec; intros.
+ - subst; now rewrite Pos.sub_diag.
+ - now rewrite Pos.sub_lt.
+ - now destruct (p-q)%positive.
+Qed.
+
+Lemma inj_pred p : p <> 1%positive ->
+ Z.pos (Pos.pred p) = Z.pred (Z.pos p).
+Proof. destruct p; easy || now destruct 1. Qed.
+
+Lemma inj_mul p q : Z.pos (p*q) = Z.pos p * Z.pos q.
+Proof. reflexivity. Qed.
+
+Lemma inj_pow_pos p q : Z.pos (p^q) = Z.pow_pos (Z.pos p) q.
+Proof. now apply Pos.iter_swap_gen. Qed.
+
+Lemma inj_pow p q : Z.pos (p^q) = (Z.pos p)^(Z.pos q).
+Proof. apply inj_pow_pos. Qed.
+
+Lemma inj_square p : Z.pos (Pos.square p) = Z.square (Z.pos p).
+Proof. reflexivity. Qed.
+
+Lemma inj_compare p q : (p ?= q)%positive = (Z.pos p ?= Z.pos q).
+Proof. reflexivity. Qed.
+
+Lemma inj_leb p q : (p <=? q)%positive = (Z.pos p <=? Z.pos q).
+Proof. reflexivity. Qed.
+
+Lemma inj_ltb p q : (p <? q)%positive = (Z.pos p <? Z.pos q).
+Proof. reflexivity. Qed.
+
+Lemma inj_eqb p q : (p =? q)%positive = (Z.pos p =? Z.pos q).
+Proof. reflexivity. Qed.
+
+Lemma inj_max p q : Z.pos (Pos.max p q) = Z.max (Z.pos p) (Z.pos q).
+Proof.
+ unfold Z.max, Pos.max. rewrite inj_compare. now case Z.compare_spec.
+Qed.
+
+Lemma inj_min p q : Z.pos (Pos.min p q) = Z.min (Z.pos p) (Z.pos q).
+Proof.
+ unfold Z.min, Pos.min. rewrite inj_compare. now case Z.compare_spec.
+Qed.
+
+Lemma inj_sqrt p : Z.pos (Pos.sqrt p) = Z.sqrt (Z.pos p).
+Proof. reflexivity. Qed.
+
+Lemma inj_gcd p q : Z.pos (Pos.gcd p q) = Z.gcd (Z.pos p) (Z.pos q).
+Proof. reflexivity. Qed.
+
+Definition inj_divide p q : (Z.pos p|Z.pos q) <-> (p|q)%positive.
+Proof. apply Z.Private_BootStrap.divide_Zpos. Qed.
+
+Lemma inj_testbit a n : 0<=n ->
+ Z.testbit (Z.pos a) n = N.testbit (N.pos a) (Z.to_N n).
+Proof. apply Z.Private_BootStrap.testbit_Zpos. Qed.
+
+(** Some results concerning Z.neg *)
+
+Lemma inj_neg p q : Z.neg p = Z.neg q -> p = q.
+Proof. now injection 1. Qed.
+
+Lemma inj_neg_iff p q : Z.neg p = Z.neg q <-> p = q.
+Proof. split. apply inj_neg. intros; now f_equal. Qed.
+
+Lemma neg_is_neg p : Z.neg p < 0.
+Proof. reflexivity. Qed.
+
+Lemma neg_is_nonpos p : Z.neg p <= 0.
+Proof. easy. Qed.
+
+Lemma neg_xO p : Z.neg p~0 = 2 * Z.neg p.
+Proof. reflexivity. Qed.
+
+Lemma neg_xI p : Z.neg p~1 = 2 * Z.neg p - 1.
+Proof. reflexivity. Qed.
+
+Lemma opp_neg p : - Z.neg p = Z.pos p.
+Proof. reflexivity. Qed.
+
+Lemma opp_pos p : - Z.pos p = Z.neg p.
+Proof. reflexivity. Qed.
+
+Lemma add_neg_neg p q : Z.neg p + Z.neg q = Z.neg (p+q).
+Proof. reflexivity. Qed.
+
+Lemma add_pos_neg p q : Z.pos p + Z.neg q = Z.pos_sub p q.
+Proof. reflexivity. Qed.
+
+Lemma add_neg_pos p q : Z.neg p + Z.pos q = Z.pos_sub q p.
+Proof. reflexivity. Qed.
+
+Lemma divide_pos_neg_r n p : (n|Z.pos p) <-> (n|Z.neg p).
+Proof. apply Z.Private_BootStrap.divide_Zpos_Zneg_r. Qed.
+
+Lemma divide_pos_neg_l n p : (Z.pos p|n) <-> (Z.neg p|n).
+Proof. apply Z.Private_BootStrap.divide_Zpos_Zneg_l. Qed.
+
+Lemma testbit_neg a n : 0<=n ->
+ Z.testbit (Z.neg a) n = negb (N.testbit (Pos.pred_N a) (Z.to_N n)).
+Proof. apply Z.Private_BootStrap.testbit_Zneg. Qed.
+
+End Pos2Z.
+
+Module Z2Pos.
+
+Lemma id x : 0 < x -> Z.pos (Z.to_pos x) = x.
+Proof. now destruct x. Qed.
+
+Lemma inj x y : 0 < x -> 0 < y -> Z.to_pos x = Z.to_pos y -> x = y.
+Proof.
+ destruct x; simpl; try easy. intros _ H ->. now apply id.
+Qed.
+
+Lemma inj_iff x y : 0 < x -> 0 < y -> (Z.to_pos x = Z.to_pos y <-> x = y).
+Proof. split. now apply inj. intros; now f_equal. Qed.
+
+Lemma to_pos_nonpos x : x <= 0 -> Z.to_pos x = 1%positive.
+Proof. destruct x; trivial. now destruct 1. Qed.
+
+Lemma inj_1 : Z.to_pos 1 = 1%positive.
+Proof. reflexivity. Qed.
+
+Lemma inj_double x : 0 < x ->
+ Z.to_pos (Z.double x) = (Z.to_pos x)~0%positive.
+Proof. now destruct x. Qed.
+
+Lemma inj_succ_double x : 0 < x ->
+ Z.to_pos (Z.succ_double x) = (Z.to_pos x)~1%positive.
+Proof. now destruct x. Qed.
+
+Lemma inj_succ x : 0 < x -> Z.to_pos (Z.succ x) = Pos.succ (Z.to_pos x).
+Proof.
+ destruct x; try easy. simpl. now rewrite Pos.add_1_r.
+Qed.
+
+Lemma inj_add x y : 0 < x -> 0 < y ->
+ Z.to_pos (x+y) = (Z.to_pos x + Z.to_pos y)%positive.
+Proof. destruct x; easy || now destruct y. Qed.
+
+Lemma inj_sub x y : 0 < x < y ->
+ Z.to_pos (y-x) = (Z.to_pos y - Z.to_pos x)%positive.
+Proof.
+ destruct x; try easy. destruct y; try easy. simpl.
+ intros. now rewrite Z.pos_sub_gt.
+Qed.
+
+Lemma inj_pred x : 1 < x -> Z.to_pos (Z.pred x) = Pos.pred (Z.to_pos x).
+Proof. now destruct x as [|[x|x|]|]. Qed.
+
+Lemma inj_mul x y : 0 < x -> 0 < y ->
+ Z.to_pos (x*y) = (Z.to_pos x * Z.to_pos y)%positive.
+Proof. destruct x; easy || now destruct y. Qed.
+
+Lemma inj_pow x y : 0 < x -> 0 < y ->
+ Z.to_pos (x^y) = (Z.to_pos x ^ Z.to_pos y)%positive.
+Proof.
+ intros. apply Pos2Z.inj. rewrite Pos2Z.inj_pow, !id; trivial.
+ apply Z.pow_pos_nonneg. trivial. now apply Z.lt_le_incl.
+Qed.
+
+Lemma inj_pow_pos x p : 0 < x ->
+ Z.to_pos (Z.pow_pos x p) = ((Z.to_pos x)^p)%positive.
+Proof. intros. now apply (inj_pow x (Z.pos p)). Qed.
+
+Lemma inj_compare x y : 0 < x -> 0 < y ->
+ (x ?= y) = (Z.to_pos x ?= Z.to_pos y)%positive.
+Proof. destruct x; easy || now destruct y. Qed.
+
+Lemma inj_leb x y : 0 < x -> 0 < y ->
+ (x <=? y) = (Z.to_pos x <=? Z.to_pos y)%positive.
+Proof. destruct x; easy || now destruct y. Qed.
+
+Lemma inj_ltb x y : 0 < x -> 0 < y ->
+ (x <? y) = (Z.to_pos x <? Z.to_pos y)%positive.
+Proof. destruct x; easy || now destruct y. Qed.
+
+Lemma inj_eqb x y : 0 < x -> 0 < y ->
+ (x =? y) = (Z.to_pos x =? Z.to_pos y)%positive.
+Proof. destruct x; easy || now destruct y. Qed.
+
+Lemma inj_max x y :
+ Z.to_pos (Z.max x y) = Pos.max (Z.to_pos x) (Z.to_pos y).
+Proof.
+ destruct x; simpl; try rewrite Pos.max_1_l.
+ - now destruct y.
+ - destruct y; simpl; now rewrite ?Pos.max_1_r, <- ?Pos2Z.inj_max.
+ - destruct y; simpl; rewrite ?Pos.max_1_r; trivial.
+ apply to_pos_nonpos. now apply Z.max_lub.
+Qed.
+
+Lemma inj_min x y :
+ Z.to_pos (Z.min x y) = Pos.min (Z.to_pos x) (Z.to_pos y).
+Proof.
+ destruct x; simpl; try rewrite Pos.min_1_l.
+ - now destruct y.
+ - destruct y; simpl; now rewrite ?Pos.min_1_r, <- ?Pos2Z.inj_min.
+ - destruct y; simpl; rewrite ?Pos.min_1_r; trivial.
+ apply to_pos_nonpos. apply Z.min_le_iff. now left.
+Qed.
+
+Lemma inj_sqrt x : Z.to_pos (Z.sqrt x) = Pos.sqrt (Z.to_pos x).
+Proof. now destruct x. Qed.
+
+Lemma inj_gcd x y : 0 < x -> 0 < y ->
+ Z.to_pos (Z.gcd x y) = Pos.gcd (Z.to_pos x) (Z.to_pos y).
+Proof. destruct x; easy || now destruct y. Qed.
+
+End Z2Pos.
(** Compatibility Notations *)
@@ -1404,7 +1584,6 @@ Notation Z_of_N := Z.of_N (compat "8.3").
Notation Zind := Z.peano_ind (compat "8.3").
Notation Zopp_0 := Z.opp_0 (compat "8.3").
-Notation Zopp_neg := Z.opp_Zneg (compat "8.3").
Notation Zopp_involutive := Z.opp_involutive (compat "8.3").
Notation Zopp_inj := Z.opp_inj (compat "8.3").
Notation Zplus_0_l := Z.add_0_l (compat "8.3").
@@ -1452,10 +1631,18 @@ Notation Zmult_reg_l := Z.mul_reg_l (compat "8.3").
Notation Zmult_reg_r := Z.mul_reg_r (compat "8.3").
Notation Zmult_succ_l := Z.mul_succ_l (compat "8.3").
Notation Zmult_succ_r := Z.mul_succ_r (compat "8.3").
-Notation Zpos_xI := Z.pos_xI (compat "8.3").
-Notation Zpos_xO := Z.pos_xO (compat "8.3").
-Notation Zneg_xI := Z.neg_xI (compat "8.3").
-Notation Zneg_xO := Z.neg_xO (compat "8.3").
+
+Notation Zpos_xI := Pos2Z.inj_xI (compat "8.3").
+Notation Zpos_xO := Pos2Z.inj_xO (compat "8.3").
+Notation Zneg_xI := Pos2Z.neg_xI (compat "8.3").
+Notation Zneg_xO := Pos2Z.neg_xO (compat "8.3").
+Notation Zopp_neg := Pos2Z.opp_neg (compat "8.3").
+Notation Zpos_succ_morphism := Pos2Z.inj_succ (compat "8.3").
+Notation Zpos_mult_morphism := Pos2Z.inj_mul (compat "8.3").
+Notation Zpos_minus_morphism := Pos2Z.inj_sub (compat "8.3").
+Notation Zpos_eq_rev := Pos2Z.inj (compat "8.3").
+Notation Zpos_plus_distr := Pos2Z.inj_add (compat "8.3").
+Notation Zneg_plus_distr := Pos2Z.add_neg_neg (compat "8.3").
Notation Z := Z (only parsing).
Notation Z_rect := Z_rect (only parsing).
@@ -1482,8 +1669,6 @@ Lemma Zplus_0_r_reverse : forall n, n = n + 0.
Proof (SYM1 Z.add_0_r).
Lemma Zplus_eq_compat : forall n m p q, n=m -> p=q -> n+p=m+q.
Proof (f_equal2 Z.add).
-Lemma Zpos_succ_morphism : forall p, Zpos (Psucc p) = Zsucc (Zpos p).
-Proof (SYM1 Z.succ_Zpos).
Lemma Zsucc_pred : forall n, n = Z.succ (Z.pred n).
Proof (SYM1 Z.succ_pred).
Lemma Zpred_succ : forall n, n = Z.pred (Z.succ n).
@@ -1506,15 +1691,10 @@ Lemma Zminus_plus_simpl_l_reverse : forall n m p, n - m = p + n - (p + m).
Proof (SYM3 Zminus_plus_simpl_l).
Lemma Zminus_plus_simpl_r : forall n m p, n + p - (m + p) = n - m.
Proof (fun n m p => Z.add_add_simpl_r_r n p m).
-Lemma Zpos_minus_morphism : forall a b,
- Pcompare a b Eq = Lt -> Zpos (b - a) = Zpos b - Zpos a.
-Proof. intros. now rewrite Z.sub_Zpos. Qed.
Lemma Zeq_minus : forall n m, n = m -> n - m = 0.
Proof (fun n m => proj2 (Z.sub_move_0_r n m)).
Lemma Zminus_eq : forall n m, n - m = 0 -> n = m.
Proof (fun n m => proj1 (Z.sub_move_0_r n m)).
-Lemma Zpos_mult_morphism : forall p q, Zpos (p * q) = Zpos p * Zpos q.
-Proof (SYM2 Z.mul_Zpos).
Lemma Zmult_0_r_reverse : forall n, 0 = n * 0.
Proof (SYM1 Z.mul_0_r).
Lemma Zmult_assoc_reverse : forall n m p, n * m * p = n * (m * p).
@@ -1529,20 +1709,14 @@ Lemma Zopp_mult_distr_r : forall n m, - (n * m) = n * - m.
Proof (SYM2 Z.mul_opp_r).
Lemma Zmult_minus_distr_l : forall n m p, p * (n - m) = p * n - p * m.
Proof (fun n m p => Z.mul_sub_distr_l p n m).
-Lemma Zmult_succ_r_reverse : forall n m, n * m + n = n * Zsucc m.
+Lemma Zmult_succ_r_reverse : forall n m, n * m + n = n * Z.succ m.
Proof (SYM2 Z.mul_succ_r).
-Lemma Zmult_succ_l_reverse : forall n m, n * m + m = Zsucc n * m.
+Lemma Zmult_succ_l_reverse : forall n m, n * m + m = Z.succ n * m.
Proof (SYM2 Z.mul_succ_l).
-Lemma Zpos_eq : forall p q, p = q -> Zpos p = Zpos q.
-Proof (fun p q => proj2 (Z.inj_Zpos p q)).
-Lemma Zpos_eq_rev : forall p q, Zpos p = Zpos q -> p = q.
-Proof (fun p q => proj1 (Z.inj_Zpos p q)).
-Lemma Zpos_eq_iff : forall p q, p = q <-> Zpos p = Zpos q.
-Proof (fun p q => iff_sym (Z.inj_Zpos p q)).
-Lemma Zpos_plus_distr : forall p q, Zpos (p + q) = Zpos p + Zpos q.
-Proof (SYM2 Z.add_Zpos).
-Lemma Zneg_plus_distr : forall p q, Zneg (p + q) = Zneg p + Zneg q.
-Proof (SYM2 Z.add_Zneg).
+Lemma Zpos_eq : forall p q, p = q -> Z.pos p = Z.pos q.
+Proof. congruence. Qed.
+Lemma Zpos_eq_iff : forall p q, p = q <-> Z.pos p = Z.pos q.
+Proof (fun p q => iff_sym (Pos2Z.inj_iff p q)).
Hint Immediate Zsucc_pred: zarith.
diff --git a/theories/ZArith/BinIntDef.v b/theories/ZArith/BinIntDef.v
index c03b8517b..6e75e53b9 100644
--- a/theories/ZArith/BinIntDef.v
+++ b/theories/ZArith/BinIntDef.v
@@ -22,6 +22,11 @@ Module Z.
Definition t := Z.
+(** ** Nicer names [Z.pos] and [Z.neg] for contructors *)
+
+Notation pos := Zpos.
+Notation neg := Zneg.
+
(** ** Constants *)
Definition zero := 0.
@@ -33,22 +38,22 @@ Definition two := 2.
Definition double x :=
match x with
| 0 => 0
- | Zpos p => Zpos p~0
- | Zneg p => Zneg p~0
+ | pos p => pos p~0
+ | neg p => neg p~0
end.
Definition succ_double x :=
match x with
| 0 => 1
- | Zpos p => Zpos p~1
- | Zneg p => Zneg (Pos.pred_double p)
+ | pos p => pos p~1
+ | neg p => neg (Pos.pred_double p)
end.
Definition pred_double x :=
match x with
| 0 => -1
- | Zneg p => Zneg p~1
- | Zpos p => Zpos (Pos.pred_double p)
+ | neg p => neg p~1
+ | pos p => pos (Pos.pred_double p)
end.
(** ** Subtraction of positive into Z *)
@@ -57,12 +62,12 @@ Fixpoint pos_sub (x y:positive) {struct y} : Z :=
match x, y with
| p~1, q~1 => double (pos_sub p q)
| p~1, q~0 => succ_double (pos_sub p q)
- | p~1, 1 => Zpos p~0
+ | p~1, 1 => pos p~0
| p~0, q~1 => pred_double (pos_sub p q)
| p~0, q~0 => double (pos_sub p q)
- | p~0, 1 => Zpos (Pos.pred_double p)
- | 1, q~1 => Zneg q~0
- | 1, q~0 => Zneg (Pos.pred_double q)
+ | p~0, 1 => pos (Pos.pred_double p)
+ | 1, q~1 => neg q~0
+ | 1, q~0 => neg (Pos.pred_double q)
| 1, 1 => Z0
end%positive.
@@ -72,10 +77,10 @@ Definition add x y :=
match x, y with
| 0, y => y
| x, 0 => x
- | Zpos x', Zpos y' => Zpos (x' + y')
- | Zpos x', Zneg y' => pos_sub x' y'
- | Zneg x', Zpos y' => pos_sub y' x'
- | Zneg x', Zneg y' => Zneg (x' + y')
+ | pos x', pos y' => pos (x' + y')
+ | pos x', neg y' => pos_sub x' y'
+ | neg x', pos y' => pos_sub y' x'
+ | neg x', neg y' => neg (x' + y')
end.
Infix "+" := add : Z_scope.
@@ -85,8 +90,8 @@ Infix "+" := add : Z_scope.
Definition opp x :=
match x with
| 0 => 0
- | Zpos x => Zneg x
- | Zneg x => Zpos x
+ | pos x => neg x
+ | neg x => pos x
end.
Notation "- x" := (opp x) : Z_scope.
@@ -111,10 +116,10 @@ Definition mul x y :=
match x, y with
| 0, _ => 0
| _, 0 => 0
- | Zpos x', Zpos y' => Zpos (x' * y')
- | Zpos x', Zneg y' => Zneg (x' * y')
- | Zneg x', Zpos y' => Zneg (x' * y')
- | Zneg x', Zneg y' => Zpos (x' * y')
+ | pos x', pos y' => pos (x' * y')
+ | pos x', neg y' => neg (x' * y')
+ | neg x', pos y' => neg (x' * y')
+ | neg x', neg y' => pos (x' * y')
end.
Infix "*" := mul : Z_scope.
@@ -125,9 +130,9 @@ Definition pow_pos (z:Z) (n:positive) := Pos.iter n (mul z) 1.
Definition pow x y :=
match y with
- | Zpos p => pow_pos x p
+ | pos p => pow_pos x p
| 0 => 1
- | Zneg _ => 0
+ | neg _ => 0
end.
Infix "^" := pow : Z_scope.
@@ -137,8 +142,8 @@ Infix "^" := pow : Z_scope.
Definition square x :=
match x with
| 0 => 0
- | Zpos p => Zpos (Pos.square p)
- | Zneg p => Zpos (Pos.square p)
+ | pos p => pos (Pos.square p)
+ | neg p => pos (Pos.square p)
end.
(** ** Comparison *)
@@ -146,14 +151,14 @@ Definition square x :=
Definition compare x y :=
match x, y with
| 0, 0 => Eq
- | 0, Zpos y' => Lt
- | 0, Zneg y' => Gt
- | Zpos x', 0 => Gt
- | Zpos x', Zpos y' => (x' ?= y')%positive
- | Zpos x', Zneg y' => Gt
- | Zneg x', 0 => Lt
- | Zneg x', Zpos y' => Lt
- | Zneg x', Zneg y' => CompOpp ((x' ?= y')%positive)
+ | 0, pos y' => Lt
+ | 0, neg y' => Gt
+ | pos x', 0 => Gt
+ | pos x', pos y' => (x' ?= y')%positive
+ | pos x', neg y' => Gt
+ | neg x', 0 => Lt
+ | neg x', pos y' => Lt
+ | neg x', neg y' => CompOpp ((x' ?= y')%positive)
end.
Infix "?=" := compare (at level 70, no associativity) : Z_scope.
@@ -163,8 +168,8 @@ Infix "?=" := compare (at level 70, no associativity) : Z_scope.
Definition sgn z :=
match z with
| 0 => 0
- | Zpos p => 1
- | Zneg p => -1
+ | pos p => 1
+ | neg p => -1
end.
(** Boolean equality and comparisons *)
@@ -200,8 +205,8 @@ Definition gtb x y :=
Fixpoint eqb x y :=
match x, y with
| 0, 0 => true
- | Zpos p, Zpos q => Pos.eqb p q
- | Zneg p, Zneg q => Pos.eqb p q
+ | pos p, pos q => Pos.eqb p q
+ | neg p, neg q => Pos.eqb p q
| _, _ => false
end.
@@ -230,8 +235,8 @@ Definition min n m :=
Definition abs z :=
match z with
| 0 => 0
- | Zpos p => Zpos p
- | Zneg p => Zpos p
+ | pos p => pos p
+ | neg p => pos p
end.
(** ** Conversions *)
@@ -241,24 +246,24 @@ Definition abs z :=
Definition abs_nat (z:Z) : nat :=
match z with
| 0 => 0%nat
- | Zpos p => Pos.to_nat p
- | Zneg p => Pos.to_nat p
+ | pos p => Pos.to_nat p
+ | neg p => Pos.to_nat p
end.
(** From [Z] to [N] via absolute value *)
Definition abs_N (z:Z) : N :=
match z with
- | Z0 => 0%N
- | Zpos p => Npos p
- | Zneg p => Npos p
+ | 0 => 0%N
+ | pos p => N.pos p
+ | neg p => N.pos p
end.
(** From [Z] to [nat] by rounding negative numbers to 0 *)
Definition to_nat (z:Z) : nat :=
match z with
- | Zpos p => Pos.to_nat p
+ | pos p => Pos.to_nat p
| _ => O
end.
@@ -266,7 +271,7 @@ Definition to_nat (z:Z) : nat :=
Definition to_N (z:Z) : N :=
match z with
- | Zpos p => Npos p
+ | pos p => N.pos p
| _ => 0%N
end.
@@ -275,15 +280,23 @@ Definition to_N (z:Z) : N :=
Definition of_nat (n:nat) : Z :=
match n with
| O => 0
- | S n => Zpos (Pos.of_succ_nat n)
+ | S n => pos (Pos.of_succ_nat n)
end.
(** From [N] to [Z] *)
Definition of_N (n:N) : Z :=
match n with
- | N0 => 0
- | Npos p => Zpos p
+ | 0%N => 0
+ | N.pos p => pos p
+ end.
+
+(** From [Z] to [positive] by rounding nonpositive numbers to 1 *)
+
+Definition to_pos (z:Z) : positive :=
+ match z with
+ | pos p => p
+ | _ => 1%positive
end.
(** ** Iteration of a function
@@ -293,7 +306,7 @@ Definition of_N (n:N) : Z :=
Definition iter (n:Z) {A} (f:A -> A) (x:A) :=
match n with
- | Zpos p => Pos.iter p f x
+ | pos p => Pos.iter p f x
| _ => x
end.
@@ -348,17 +361,17 @@ Definition div_eucl (a b:Z) : Z * Z :=
match a, b with
| 0, _ => (0, 0)
| _, 0 => (0, 0)
- | Zpos a', Zpos _ => pos_div_eucl a' b
- | Zneg a', Zpos _ =>
+ | pos a', pos _ => pos_div_eucl a' b
+ | neg a', pos _ =>
let (q, r) := pos_div_eucl a' b in
match r with
| 0 => (- q, 0)
| _ => (- (q + 1), b - r)
end
- | Zneg a', Zneg b' =>
- let (q, r) := pos_div_eucl a' (Zpos b') in (q, - r)
- | Zpos a', Zneg b' =>
- let (q, r) := pos_div_eucl a' (Zpos b') in
+ | neg a', neg b' =>
+ let (q, r) := pos_div_eucl a' (pos b') in (q, - r)
+ | pos a', neg b' =>
+ let (q, r) := pos_div_eucl a' (pos b') in
match r with
| 0 => (- q, 0)
| _ => (- (q + 1), b + r)
@@ -392,14 +405,14 @@ Definition quotrem (a b:Z) : Z * Z :=
match a, b with
| 0, _ => (0, 0)
| _, 0 => (0, a)
- | Zpos a, Zpos b =>
- let (q, r) := N.pos_div_eucl a (Npos b) in (of_N q, of_N r)
- | Zneg a, Zpos b =>
- let (q, r) := N.pos_div_eucl a (Npos b) in (-of_N q, - of_N r)
- | Zpos a, Zneg b =>
- let (q, r) := N.pos_div_eucl a (Npos b) in (-of_N q, of_N r)
- | Zneg a, Zneg b =>
- let (q, r) := N.pos_div_eucl a (Npos b) in (of_N q, - of_N r)
+ | pos a, pos b =>
+ let (q, r) := N.pos_div_eucl a (N.pos b) in (of_N q, of_N r)
+ | neg a, pos b =>
+ let (q, r) := N.pos_div_eucl a (N.pos b) in (-of_N q, - of_N r)
+ | pos a, neg b =>
+ let (q, r) := N.pos_div_eucl a (N.pos b) in (-of_N q, of_N r)
+ | neg a, neg b =>
+ let (q, r) := N.pos_div_eucl a (N.pos b) in (of_N q, - of_N r)
end.
Definition quot a b := fst (quotrem a b).
@@ -414,16 +427,16 @@ Infix "÷" := quot (at level 40, left associativity) : Z_scope.
Definition even z :=
match z with
| 0 => true
- | Zpos (xO _) => true
- | Zneg (xO _) => true
+ | pos (xO _) => true
+ | neg (xO _) => true
| _ => false
end.
Definition odd z :=
match z with
| 0 => false
- | Zpos (xO _) => false
- | Zneg (xO _) => false
+ | pos (xO _) => false
+ | neg (xO _) => false
| _ => true
end.
@@ -437,9 +450,9 @@ Definition odd z :=
Definition div2 z :=
match z with
| 0 => 0
- | Zpos 1 => 0
- | Zpos p => Zpos (Pos.div2 p)
- | Zneg p => Zneg (Pos.div2_up p)
+ | pos 1 => 0
+ | pos p => pos (Pos.div2 p)
+ | neg p => neg (Pos.div2_up p)
end.
(** [quot2] performs rounding toward zero, it is hence a particular
@@ -449,21 +462,21 @@ Definition div2 z :=
Definition quot2 (z:Z) :=
match z with
| 0 => 0
- | Zpos 1 => 0
- | Zpos p => Zpos (Pos.div2 p)
- | Zneg 1 => 0
- | Zneg p => Zneg (Pos.div2 p)
+ | pos 1 => 0
+ | pos p => pos (Pos.div2 p)
+ | neg 1 => 0
+ | neg p => neg (Pos.div2 p)
end.
-(** NB: [Z.quot2] used to be named [Zdiv2] in Coq <= 8.3 *)
+(** NB: [Z.quot2] used to be named [Z.div2] in Coq <= 8.3 *)
(** * Base-2 logarithm *)
Definition log2 z :=
match z with
- | Zpos (p~1) => Zpos (Pos.size p)
- | Zpos (p~0) => Zpos (Pos.size p)
+ | pos (p~1) => pos (Pos.size p)
+ | pos (p~0) => pos (Pos.size p)
| _ => 0
end.
@@ -473,17 +486,17 @@ Definition log2 z :=
Definition sqrtrem n :=
match n with
| 0 => (0, 0)
- | Zpos p =>
+ | pos p =>
match Pos.sqrtrem p with
- | (s, IsPos r) => (Zpos s, Zpos r)
- | (s, _) => (Zpos s, 0)
+ | (s, IsPos r) => (pos s, pos r)
+ | (s, _) => (pos s, 0)
end
- | Zneg _ => (0,0)
+ | neg _ => (0,0)
end.
Definition sqrt n :=
match n with
- | Zpos p => Zpos (Pos.sqrt p)
+ | pos p => pos (Pos.sqrt p)
| _ => 0
end.
@@ -494,10 +507,10 @@ Definition gcd a b :=
match a,b with
| 0, _ => abs b
| _, 0 => abs a
- | Zpos a, Zpos b => Zpos (Pos.gcd a b)
- | Zpos a, Zneg b => Zpos (Pos.gcd a b)
- | Zneg a, Zpos b => Zpos (Pos.gcd a b)
- | Zneg a, Zneg b => Zpos (Pos.gcd a b)
+ | pos a, pos b => pos (Pos.gcd a b)
+ | pos a, neg b => pos (Pos.gcd a b)
+ | neg a, pos b => pos (Pos.gcd a b)
+ | neg a, neg b => pos (Pos.gcd a b)
end.
(** A generalized gcd, also computing division of a and b by gcd. *)
@@ -506,14 +519,14 @@ Definition ggcd a b : Z*(Z*Z) :=
match a,b with
| 0, _ => (abs b,(0, sgn b))
| _, 0 => (abs a,(sgn a, 0))
- | Zpos a, Zpos b =>
- let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zpos aa, Zpos bb))
- | Zpos a, Zneg b =>
- let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zpos aa, Zneg bb))
- | Zneg a, Zpos b =>
- let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zneg aa, Zpos bb))
- | Zneg a, Zneg b =>
- let '(g,(aa,bb)) := Pos.ggcd a b in (Zpos g, (Zneg aa, Zneg bb))
+ | pos a, pos b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (pos g, (pos aa, pos bb))
+ | pos a, neg b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (pos g, (pos aa, neg bb))
+ | neg a, pos b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (pos g, (neg aa, pos bb))
+ | neg a, neg b =>
+ let '(g,(aa,bb)) := Pos.ggcd a b in (pos g, (neg aa, neg bb))
end.
@@ -532,13 +545,13 @@ Definition ggcd a b : Z*(Z*Z) :=
Definition testbit a n :=
match n with
| 0 => odd a
- | Zpos p =>
+ | pos p =>
match a with
| 0 => false
- | Zpos a => Pos.testbit a (Npos p)
- | Zneg a => negb (N.testbit (Pos.pred_N a) (Npos p))
+ | pos a => Pos.testbit a (N.pos p)
+ | neg a => negb (N.testbit (Pos.pred_N a) (N.pos p))
end
- | Zneg _ => false
+ | neg _ => false
end.
(** Shifts
@@ -555,8 +568,8 @@ Definition testbit a n :=
Definition shiftl a n :=
match n with
| 0 => a
- | Zpos p => Pos.iter p (mul 2) a
- | Zneg p => Pos.iter p div2 a
+ | pos p => Pos.iter p (mul 2) a
+ | neg p => Pos.iter p div2 a
end.
Definition shiftr a n := shiftl a (-n).
@@ -567,40 +580,40 @@ Definition lor a b :=
match a, b with
| 0, _ => b
| _, 0 => a
- | Zpos a, Zpos b => Zpos (Pos.lor a b)
- | Zneg a, Zpos b => Zneg (N.succ_pos (N.ldiff (Pos.pred_N a) (Npos b)))
- | Zpos a, Zneg b => Zneg (N.succ_pos (N.ldiff (Pos.pred_N b) (Npos a)))
- | Zneg a, Zneg b => Zneg (N.succ_pos (N.land (Pos.pred_N a) (Pos.pred_N b)))
+ | pos a, pos b => pos (Pos.lor a b)
+ | neg a, pos b => neg (N.succ_pos (N.ldiff (Pos.pred_N a) (N.pos b)))
+ | pos a, neg b => neg (N.succ_pos (N.ldiff (Pos.pred_N b) (N.pos a)))
+ | neg a, neg b => neg (N.succ_pos (N.land (Pos.pred_N a) (Pos.pred_N b)))
end.
Definition land a b :=
match a, b with
| 0, _ => 0
| _, 0 => 0
- | Zpos a, Zpos b => of_N (Pos.land a b)
- | Zneg a, Zpos b => of_N (N.ldiff (Npos b) (Pos.pred_N a))
- | Zpos a, Zneg b => of_N (N.ldiff (Npos a) (Pos.pred_N b))
- | Zneg a, Zneg b => Zneg (N.succ_pos (N.lor (Pos.pred_N a) (Pos.pred_N b)))
+ | pos a, pos b => of_N (Pos.land a b)
+ | neg a, pos b => of_N (N.ldiff (N.pos b) (Pos.pred_N a))
+ | pos a, neg b => of_N (N.ldiff (N.pos a) (Pos.pred_N b))
+ | neg a, neg b => neg (N.succ_pos (N.lor (Pos.pred_N a) (Pos.pred_N b)))
end.
Definition ldiff a b :=
match a, b with
| 0, _ => 0
| _, 0 => a
- | Zpos a, Zpos b => of_N (Pos.ldiff a b)
- | Zneg a, Zpos b => Zneg (N.succ_pos (N.lor (Pos.pred_N a) (Npos b)))
- | Zpos a, Zneg b => of_N (N.land (Npos a) (Pos.pred_N b))
- | Zneg a, Zneg b => of_N (N.ldiff (Pos.pred_N b) (Pos.pred_N a))
+ | pos a, pos b => of_N (Pos.ldiff a b)
+ | neg a, pos b => neg (N.succ_pos (N.lor (Pos.pred_N a) (N.pos b)))
+ | pos a, neg b => of_N (N.land (N.pos a) (Pos.pred_N b))
+ | neg a, neg b => of_N (N.ldiff (Pos.pred_N b) (Pos.pred_N a))
end.
Definition lxor a b :=
match a, b with
| 0, _ => b
| _, 0 => a
- | Zpos a, Zpos b => of_N (Pos.lxor a b)
- | Zneg a, Zpos b => Zneg (N.succ_pos (N.lxor (Pos.pred_N a) (Npos b)))
- | Zpos a, Zneg b => Zneg (N.succ_pos (N.lxor (Npos a) (Pos.pred_N b)))
- | Zneg a, Zneg b => of_N (N.lxor (Pos.pred_N a) (Pos.pred_N b))
+ | pos a, pos b => of_N (Pos.lxor a b)
+ | neg a, pos b => neg (N.succ_pos (N.lxor (Pos.pred_N a) (N.pos b)))
+ | pos a, neg b => neg (N.succ_pos (N.lxor (N.pos a) (Pos.pred_N b)))
+ | neg a, neg b => of_N (N.lxor (Pos.pred_N a) (Pos.pred_N b))
end.
End Z. \ No newline at end of file
diff --git a/theories/ZArith/Int.v b/theories/ZArith/Int.v
index 7c840c563..384c046f3 100644
--- a/theories/ZArith/Int.v
+++ b/theories/ZArith/Int.v
@@ -250,7 +250,7 @@ Module MoreInt (Import I:Int).
| EZplus e1 e2 => ((ez2z e1)+(ez2z e2))%Z
| EZminus e1 e2 => ((ez2z e1)-(ez2z e2))%Z
| EZmult e1 e2 => ((ez2z e1)*(ez2z e2))%Z
- | EZmax e1 e2 => Zmax (ez2z e1) (ez2z e2)
+ | EZmax e1 e2 => Z.max (ez2z e1) (ez2z e2)
| EZopp e => (-(ez2z e))%Z
| EZofI e => i2z (ei2i e)
| EZraw z => z
@@ -367,14 +367,14 @@ Module Z_as_Int <: Int.
Definition _1 := 1.
Definition _2 := 2.
Definition _3 := 3.
- Definition plus := Zplus.
- Definition opp := Zopp.
- Definition minus := Zminus.
- Definition mult := Zmult.
- Definition max := Zmax.
+ Definition plus := Z.add.
+ Definition opp := Z.opp.
+ Definition minus := Z.sub.
+ Definition mult := Z.mul.
+ Definition max := Z.max.
Definition gt_le_dec := Z_gt_le_dec.
Definition ge_lt_dec := Z_ge_lt_dec.
- Definition eq_dec := Z_eq_dec.
+ Definition eq_dec := Z.eq_dec.
Definition i2z : int -> Z := fun n => n.
Lemma i2z_eq : forall n p, i2z n=i2z p -> n = p. Proof. auto. Qed.
Lemma i2z_0 : i2z _0 = 0. Proof. auto. Qed.
@@ -385,5 +385,5 @@ Module Z_as_Int <: Int.
Lemma i2z_opp n : i2z (- n) = - i2z n. Proof. auto. Qed.
Lemma i2z_minus n p : i2z (n - p) = i2z n - i2z p. Proof. auto. Qed.
Lemma i2z_mult n p : i2z (n * p) = i2z n * i2z p. Proof. auto. Qed.
- Lemma i2z_max n p : i2z (max n p) = Zmax (i2z n) (i2z p). Proof. auto. Qed.
+ Lemma i2z_max n p : i2z (max n p) = Z.max (i2z n) (i2z p). Proof. auto. Qed.
End Z_as_Int.
diff --git a/theories/ZArith/Wf_Z.v b/theories/ZArith/Wf_Z.v
index bcccc1269..bb84d0c8c 100644
--- a/theories/ZArith/Wf_Z.v
+++ b/theories/ZArith/Wf_Z.v
@@ -39,7 +39,7 @@ Proof.
Qed.
Lemma Z_of_nat_complete_inf (x : Z) :
- 0 <= x -> {n : nat | x = Z_of_nat n}.
+ 0 <= x -> {n : nat | x = Z.of_nat n}.
Proof.
intros H. exists (Z.to_nat x). symmetry. now apply Z2Nat.id.
Qed.
@@ -53,7 +53,7 @@ Qed.
Lemma Z_of_nat_set :
forall P:Z -> Set,
- (forall n:nat, P (Z_of_nat n)) -> forall x:Z, 0 <= x -> P x.
+ (forall n:nat, P (Z.of_nat n)) -> forall x:Z, 0 <= x -> P x.
Proof.
intros P H x Hx. now destruct (Z_of_nat_complete_inf x Hx) as (n,->).
Qed.
@@ -129,7 +129,7 @@ Section Efficient_Rec.
- now destruct Hz.
Qed.
- (** A more general induction principle on non-negative numbers using [Zlt]. *)
+ (** A more general induction principle on non-negative numbers using [Z.lt]. *)
Lemma Zlt_0_rec :
forall P:Z -> Type,
@@ -155,7 +155,7 @@ Section Efficient_Rec.
exact Zlt_0_rec.
Qed.
- (** Obsolete version of [Zlt] induction principle on non-negative numbers *)
+ (** Obsolete version of [Z.lt] induction principle on non-negative numbers *)
Lemma Z_lt_rec :
forall P:Z -> Type,
@@ -173,7 +173,7 @@ Section Efficient_Rec.
exact Z_lt_rec.
Qed.
- (** An even more general induction principle using [Zlt]. *)
+ (** An even more general induction principle using [Z.lt]. *)
Lemma Zlt_lower_bound_rec :
forall P:Z -> Type, forall z:Z,
diff --git a/theories/ZArith/ZArith_dec.v b/theories/ZArith/ZArith_dec.v
index ca26787bd..cac3cd4e5 100644
--- a/theories/ZArith/ZArith_dec.v
+++ b/theories/ZArith/ZArith_dec.v
@@ -21,20 +21,16 @@ Proof.
Defined.
(* end hide *)
-Lemma Zcompare_rect :
- forall (P:Type) (n m:Z),
- ((n ?= m) = Eq -> P) -> ((n ?= m) = Lt -> P) -> ((n ?= m) = Gt -> P) -> P.
+Lemma Zcompare_rect (P:Type) (n m:Z) :
+ ((n ?= m) = Eq -> P) -> ((n ?= m) = Lt -> P) -> ((n ?= m) = Gt -> P) -> P.
Proof.
- intros * H1 H2 H3.
+ intros H1 H2 H3.
destruct (n ?= m); auto.
Defined.
-Lemma Zcompare_rec :
- forall (P:Set) (n m:Z),
- ((n ?= m) = Eq -> P) -> ((n ?= m) = Lt -> P) -> ((n ?= m) = Gt -> P) -> P.
-Proof.
- intro; apply Zcompare_rect.
-Defined.
+Lemma Zcompare_rec (P:Set) (n m:Z) :
+ ((n ?= m) = Eq -> P) -> ((n ?= m) = Lt -> P) -> ((n ?= m) = Gt -> P) -> P.
+Proof. apply Zcompare_rect. Defined.
Notation Z_eq_dec := Z.eq_dec (compat "8.3").
@@ -46,38 +42,22 @@ Section decidability.
Definition Z_lt_dec : {x < y} + {~ x < y}.
Proof.
- unfold Zlt in |- *.
- apply Zcompare_rec with (n := x) (m := y); intro H.
- right. rewrite H. discriminate.
- left; assumption.
- right. rewrite H. discriminate.
+ unfold Z.lt; case Z.compare; (now left) || (now right).
Defined.
Definition Z_le_dec : {x <= y} + {~ x <= y}.
Proof.
- unfold Zle in |- *.
- apply Zcompare_rec with (n := x) (m := y); intro H.
- left. rewrite H. discriminate.
- left. rewrite H. discriminate.
- right. tauto.
+ unfold Z.le; case Z.compare; (now left) || (right; tauto).
Defined.
Definition Z_gt_dec : {x > y} + {~ x > y}.
Proof.
- unfold Zgt in |- *.
- apply Zcompare_rec with (n := x) (m := y); intro H.
- right. rewrite H. discriminate.
- right. rewrite H. discriminate.
- left; assumption.
+ unfold Z.gt; case Z.compare; (now left) || (now right).
Defined.
Definition Z_ge_dec : {x >= y} + {~ x >= y}.
Proof.
- unfold Zge in |- *.
- apply Zcompare_rec with (n := x) (m := y); intro H.
- left. rewrite H. discriminate.
- right. tauto.
- left. rewrite H. discriminate.
+ unfold Z.ge; case Z.compare; (now left) || (right; tauto).
Defined.
Definition Z_lt_ge_dec : {x < y} + {x >= y}.
@@ -87,16 +67,15 @@ Section decidability.
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.
+ * now left.
+ * right; now apply Z.ge_le.
Defined.
Definition Z_le_gt_dec : {x <= y} + {x > y}.
Proof.
elim Z_le_dec; auto with arith.
- intro. right. apply Znot_le_gt; auto with arith.
+ intro. right. Z.swap_greater. now apply Z.nle_gt.
Defined.
Definition Z_gt_le_dec : {x > y} + {x <= y}.
@@ -107,15 +86,15 @@ Section decidability.
Definition Z_ge_lt_dec : {x >= y} + {x < y}.
Proof.
elim Z_ge_dec; auto with arith.
- intro. right. apply Znot_ge_lt; auto with arith.
+ intro. right. Z.swap_greater. now apply Z.lt_nge.
Defined.
Definition Z_le_lt_eq_dec : x <= y -> {x < y} + {x = y}.
Proof.
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. right. elim (Z.compare_eq_iff x y); auto with arith.
+ intro. left. elim (Z.compare_eq_iff x y); auto with arith.
intro H1. absurd (x > y); auto with arith.
Defined.
@@ -132,8 +111,8 @@ Proof.
assumption.
intro.
right.
- apply Zle_lt_trans with (m := x).
- apply Zge_le.
+ apply Z.le_lt_trans with (m := x).
+ apply Z.ge_le.
assumption.
assumption.
Defined.
@@ -142,20 +121,16 @@ 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.
+ - now left.
+ - right.
+ apply Z.add_lt_mono_l with (p := x).
+ now rewrite Z.add_0_r.
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 ];
+ [ right; apply Z.add_lt_mono_l with (p := x); rewrite Z.add_0_r | left ];
assumption.
Defined.
@@ -167,7 +142,7 @@ Proof.
left.
assumption.
intro H0.
- generalize (Zge_le _ _ H0).
+ generalize (Z.ge_le _ _ H0).
intro.
case (Z_le_lt_eq_dec _ _ H1).
intro.
@@ -189,13 +164,13 @@ Proof.
left.
assumption.
intro H.
- generalize (Zge_le _ _ H).
+ generalize (Z.ge_le _ _ H).
intro H0.
case (Z_le_lt_eq_dec y x H0).
intro H1.
left.
right.
- apply Zlt_gt.
+ apply Z.lt_gt.
assumption.
intro.
right.
@@ -207,7 +182,7 @@ 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;
+ case (Z.eq_dec x y); intro H;
[ right; assumption | left; apply (not_Zeq_inf _ _ H) ].
Defined.
@@ -215,12 +190,12 @@ Defined.
(* To deprecate ? *)
Corollary Z_zerop : forall x:Z, {x = 0} + {x <> 0}.
Proof.
- exact (fun x:Z => Z_eq_dec x 0).
+ exact (fun x:Z => Z.eq_dec x 0).
Defined.
Corollary Z_notzerop : forall (x:Z), {x <> 0} + {x = 0}.
Proof (fun x => sumbool_not _ _ (Z_zerop x)).
Corollary Z_noteq_dec : forall (x y:Z), {x <> y} + {x = y}.
-Proof (fun x y => sumbool_not _ _ (Z_eq_dec x y)).
+Proof (fun x y => sumbool_not _ _ (Z.eq_dec x y)).
(* end hide *)
diff --git a/theories/ZArith/Zabs.v b/theories/ZArith/Zabs.v
index 2e43b3554..1eaae5df0 100644
--- a/theories/ZArith/Zabs.v
+++ b/theories/ZArith/Zabs.v
@@ -77,9 +77,9 @@ Notation Zsgn_null := Z.sgn_null_iff (compat "8.3").
(** A characterization of the sign function: *)
Lemma Zsgn_spec x :
- 0 < x /\ Zsgn x = 1 \/
- 0 = x /\ Zsgn x = 0 \/
- 0 > x /\ Zsgn x = -1.
+ 0 < x /\ Z.sgn x = 1 \/
+ 0 = x /\ Z.sgn x = 0 \/
+ 0 > x /\ Z.sgn x = -1.
Proof.
intros. Z.swap_greater. apply Z.sgn_spec.
Qed.
@@ -99,7 +99,7 @@ Proof.
intros (H,H'). apply Zabs2Nat.inj_le; trivial. now transitivity n.
Qed.
-Lemma Zabs_nat_lt n m : 0 <= n < m -> (Zabs_nat n < Zabs_nat m)%nat.
+Lemma Zabs_nat_lt n m : 0 <= n < m -> (Z.abs_nat n < Z.abs_nat m)%nat.
Proof.
intros (H,H'). apply Zabs2Nat.inj_lt; trivial.
transitivity n; trivial. now apply Z.lt_le_incl.
diff --git a/theories/ZArith/Zbool.v b/theories/ZArith/Zbool.v
index 7c1e096ed..3a86a821b 100644
--- a/theories/ZArith/Zbool.v
+++ b/theories/ZArith/Zbool.v
@@ -25,7 +25,7 @@ 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_eq_bool (x y:Z) := bool_of_sumbool (Z_eq_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).
diff --git a/theories/ZArith/Zcompare.v b/theories/ZArith/Zcompare.v
index 703a3972f..e369aa6ab 100644
--- a/theories/ZArith/Zcompare.v
+++ b/theories/ZArith/Zcompare.v
@@ -7,7 +7,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(** Binary Integers : results about Zcompare *)
+(** Binary Integers : results about Z.compare *)
(** Initial author: Pierre Crégut (CNET, Lannion, France *)
(** THIS FILE IS DEPRECATED.
@@ -85,7 +85,7 @@ Qed.
Lemma Zcompare_succ_compat n m : (Z.succ n ?= Z.succ m) = (n ?= m).
Proof.
- rewrite <- 2 Z.add_1_l. apply Zcompare_plus_compat.
+ rewrite <- 2 Z.add_1_l. apply Z.add_compare_mono_l.
Qed.
(** * Multiplication and comparison *)
@@ -106,7 +106,7 @@ Qed.
Lemma Zmult_compare_compat_r n m p :
p > 0 -> (n ?= m) = (n * p ?= m * p).
Proof.
- intros; rewrite 2 (Zmult_comm _ p); now apply Zmult_compare_compat_l.
+ intros; rewrite 2 (Z.mul_comm _ p); now apply Zmult_compare_compat_l.
Qed.
(** * Relating [x ?= y] to [=], [<=], [<], [>=] or [>] *)
diff --git a/theories/ZArith/Zcomplements.v b/theories/ZArith/Zcomplements.v
index 5a2c3cc32..dde0745eb 100644
--- a/theories/ZArith/Zcomplements.v
+++ b/theories/ZArith/Zcomplements.v
@@ -39,8 +39,8 @@ Proof. reflexivity. Qed.
Lemma floor_ok : forall p:positive, floor p <= Zpos p < 2 * floor p.
Proof.
unfold floor. induction p; simpl.
- - rewrite !Z.pos_xI, (Z.pos_xO (xO _)), Z.pos_xO. omega.
- - rewrite (Z.pos_xO (xO _)), (Z.pos_xO p), Z.pos_xO. omega.
+ - rewrite !Pos2Z.inj_xI, (Pos2Z.inj_xO (xO _)), Pos2Z.inj_xO. omega.
+ - rewrite (Pos2Z.inj_xO (xO _)), (Pos2Z.inj_xO p), Pos2Z.inj_xO. omega.
- omega.
Qed.
@@ -107,7 +107,7 @@ Require Import List.
Fixpoint Zlength_aux (acc:Z) (A:Type) (l:list A) : Z :=
match l with
| nil => acc
- | _ :: l => Zlength_aux (Zsucc acc) A l
+ | _ :: l => Zlength_aux (Z.succ acc) A l
end.
Definition Zlength := Zlength_aux 0.
diff --git a/theories/ZArith/Zdigits.v b/theories/ZArith/Zdigits.v
index ff1d96df4..a9348785a 100644
--- a/theories/ZArith/Zdigits.v
+++ b/theories/ZArith/Zdigits.v
@@ -64,7 +64,7 @@ Section ENCODING_VALUE.
(** We compute the binary value via a Horner scheme.
Computation stops at the vector length without checks.
- We define a function Zmod2 similar to Zdiv2 returning the
+ We define a function Zmod2 similar to Z.div2 returning the
quotient of division z=2q+r with 0<=r<=1.
The two's complement value is also computed via a Horner scheme
with Zmod2, the parameter is the size minus one.
@@ -88,7 +88,7 @@ Section ENCODING_VALUE.
Lemma Zmod2_twice :
- forall z:Z, z = (2 * Zmod2 z + bit_value (Zeven.Zodd_bool z))%Z.
+ forall z:Z, z = (2 * Zmod2 z + bit_value (Z.odd z))%Z.
Proof.
destruct z; simpl in |- *.
trivial.
@@ -97,7 +97,7 @@ Section ENCODING_VALUE.
destruct p; simpl in |- *.
destruct p as [p| p| ]; simpl in |- *.
- rewrite <- (Pdouble_minus_one_o_succ_eq_xI p); trivial.
+ rewrite <- (Pos.pred_double_succ p); trivial.
trivial.
@@ -113,15 +113,15 @@ Section ENCODING_VALUE.
simple induction n; intros.
exact Bnil.
- exact (Bcons (Zeven.Zodd_bool H0) n0 (H (Zeven.Zdiv2 H0))).
+ exact (Bcons (Z.odd H0) n0 (H (Z.div2 H0))).
Defined.
Lemma Z_to_two_compl : forall n:nat, Z -> Bvector (S n).
Proof.
simple induction n; intros.
- exact (Bcons (Zeven.Zodd_bool H) 0 Bnil).
+ exact (Bcons (Z.odd H) 0 Bnil).
- exact (Bcons (Zeven.Zodd_bool H0) (S n0) (H (Zmod2 H0))).
+ exact (Bcons (Z.odd H0) (S n0) (H (Zmod2 H0))).
Defined.
End ENCODING_VALUE.
@@ -175,27 +175,27 @@ Section Z_BRIC_A_BRAC.
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.
+ intros; rewrite (Pos.succ_pred_double 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)).
+ Bcons (Z.odd z) n (Z_to_binary n (Z.div2 z)).
Proof.
intros; auto.
Qed.
Lemma Z_div2_value :
forall z:Z,
- (z >= 0)%Z -> (bit_value (Zeven.Zodd_bool z) + 2 * Zeven.Zdiv2 z)%Z = z.
+ (z >= 0)%Z -> (bit_value (Z.odd z) + 2 * Z.div2 z)%Z = z.
Proof.
destruct z as [| p| p]; auto.
destruct p; auto.
intro H; elim H; trivial.
Qed.
- Lemma Pdiv2 : forall z:Z, (z >= 0)%Z -> (Zeven.Zdiv2 z >= 0)%Z.
+ Lemma Pdiv2 : forall z:Z, (z >= 0)%Z -> (Z.div2 z >= 0)%Z.
Proof.
destruct z as [| p| p].
auto.
@@ -209,10 +209,10 @@ Section Z_BRIC_A_BRAC.
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.
+ (z < two_power_nat (S n))%Z -> (Z.div2 z < two_power_nat n)%Z.
Proof.
intros.
- cut (2 * Zeven.Zdiv2 z < 2 * two_power_nat n)%Z; intros.
+ cut (2 * Z.div2 z < 2 * two_power_nat n)%Z; intros.
omega.
rewrite <- two_power_nat_S.
@@ -225,13 +225,13 @@ Section Z_BRIC_A_BRAC.
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)).
+ Bcons (Z.odd z) (S n) (Z_to_two_compl n (Zmod2 z)).
Proof.
intros; auto.
Qed.
Lemma Zeven_bit_value :
- forall z:Z, Zeven.Zeven z -> bit_value (Zeven.Zodd_bool z) = 0%Z.
+ forall z:Z, Zeven.Zeven z -> bit_value (Z.odd z) = 0%Z.
Proof.
destruct z; unfold bit_value in |- *; auto.
destruct p; tauto || (intro H; elim H).
@@ -239,7 +239,7 @@ Section Z_BRIC_A_BRAC.
Qed.
Lemma Zodd_bit_value :
- forall z:Z, Zeven.Zodd z -> bit_value (Zeven.Zodd_bool z) = 1%Z.
+ forall z:Z, Zeven.Zodd z -> bit_value (Z.odd z) = 1%Z.
Proof.
destruct z; unfold bit_value in |- *; auto.
intros; elim H.
diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v
index 057fd6d34..0546c3f6f 100644
--- a/theories/ZArith/Zdiv.v
+++ b/theories/ZArith/Zdiv.v
@@ -63,8 +63,8 @@ Definition Remainder r b := 0 <= r < b \/ b < r <= 0.
Definition Remainder_alt r b := Z.abs r < Z.abs b /\ Z.sgn r <> - Z.sgn b.
-(* In the last formulation, [ Zsgn r <> - Zsgn b ] is less nice than saying
- [ Zsgn r = Zsgn b ], but at least it works even when [r] is null. *)
+(* In the last formulation, [ Z.sgn r <> - Z.sgn b ] is less nice than saying
+ [ Z.sgn r = Z.sgn b ], but at least it works even when [r] is null. *)
Lemma Remainder_equiv : forall r b, Remainder r b <-> Remainder_alt r b.
Proof.
@@ -89,7 +89,7 @@ Proof.
now destruct Hb. left; now apply POS. right; now apply NEG.
Qed.
-(** The same results as before, stated separately in terms of Zdiv and Zmod *)
+(** The same results as before, stated separately in terms of Z.div and Z.modulo *)
Lemma Z_mod_remainder a b : b<>0 -> Remainder (a mod b) b.
Proof.
@@ -98,7 +98,7 @@ Proof.
Qed.
Lemma Z_mod_lt a b : b > 0 -> 0 <= a mod b < b.
-Proof (fun Hb => Z.mod_pos_bound a b (Zgt_lt _ _ Hb)).
+Proof (fun Hb => Z.mod_pos_bound a b (Z.gt_lt _ _ Hb)).
Lemma Z_mod_neg a b : b < 0 -> b < a mod b <= 0.
Proof (Z.mod_neg_bound a b).
@@ -220,7 +220,7 @@ Proof. intros. zero_or_not b. apply Z.mod_mul. auto. Qed.
Lemma Z_div_mult_full : forall a b:Z, b <> 0 -> (a*b)/b = a.
Proof Z.div_mul.
-(** * Order results about Zmod and Zdiv *)
+(** * Order results about Z.modulo and Z.div *)
(* Division of positive numbers is positive. *)
@@ -248,12 +248,12 @@ Proof Z.div_small.
Theorem Zmod_small: forall a n, 0 <= a < n -> a mod n = a.
Proof Z.mod_small.
-(** [Zge] is compatible with a positive division. *)
+(** [Z.ge] is compatible with a positive division. *)
Lemma Z_div_ge : forall a b c:Z, c > 0 -> a >= b -> a/c >= b/c.
-Proof. intros. apply Zle_ge. apply Z.div_le_mono; auto with zarith. Qed.
+Proof. intros. apply Z.le_ge. apply Z.div_le_mono; auto with zarith. Qed.
-(** Same, with [Zle]. *)
+(** Same, with [Z.le]. *)
Lemma Z_div_le : forall a b c:Z, c > 0 -> a <= b -> a/c <= b/c.
Proof. intros. apply Z.div_le_mono; auto with zarith. Qed.
@@ -264,7 +264,7 @@ Lemma Z_mult_div_ge : forall a b:Z, b > 0 -> b*(a/b) <= a.
Proof. intros. apply Z.mul_div_le; auto with zarith. Qed.
Lemma Z_mult_div_ge_neg : forall a b:Z, b < 0 -> b*(a/b) >= a.
-Proof. intros. apply Zle_ge. apply Z.mul_div_ge; auto with zarith. Qed.
+Proof. intros. apply Z.le_ge. apply Z.mul_div_ge; auto with zarith. Qed.
(** The previous inequalities are exact iff the modulo is zero. *)
@@ -279,7 +279,7 @@ Proof. intros; rewrite Z.div_exact; auto. Qed.
Theorem Zmod_le: forall a b, 0 < b -> 0 <= a -> a mod b <= a.
Proof. intros. apply Z.mod_le; auto. Qed.
-(** Some additionnal inequalities about Zdiv. *)
+(** Some additionnal inequalities about Z.div. *)
Theorem Zdiv_lt_upper_bound:
forall a b q, 0 < b -> a < q*b -> a/b < q.
@@ -307,7 +307,7 @@ Proof.
destruct Z.pos_div_eucl as (q,r); destruct r; omega with *.
Qed.
-(** * Relations between usual operations and Zmod and Zdiv *)
+(** * Relations between usual operations and Z.modulo and Z.div *)
Lemma Z_mod_plus_full : forall a b c:Z, (a + b * c) mod c = a mod c.
Proof. intros. zero_or_not c. apply Z.mod_add; auto. Qed.
@@ -318,9 +318,9 @@ Proof Z.div_add.
Theorem Z_div_plus_full_l: forall a b c : Z, b <> 0 -> (a * b + c) / b = a + c / b.
Proof Z.div_add_l.
-(** [Zopp] and [Zdiv], [Zmod].
+(** [Z.opp] and [Z.div], [Z.modulo].
Due to the choice of convention for our Euclidean division,
- some of the relations about [Zopp] and divisions are rather complex. *)
+ some of the relations about [Z.opp] and divisions are rather complex. *)
Lemma Zdiv_opp_opp : forall a b:Z, (-a)/(-b) = a/b.
Proof. intros. zero_or_not b. apply Z.div_opp_opp; auto. Qed.
@@ -365,22 +365,22 @@ Proof. intros. zero_or_not b. apply Z.div_mul_cancel_r; auto. Qed.
Lemma Zdiv_mult_cancel_l : forall a b c:Z,
c<>0 -> (c*a)/(c*b) = a/b.
Proof.
- intros. rewrite (Zmult_comm c b); zero_or_not b.
- rewrite (Zmult_comm b c). apply Z.div_mul_cancel_l; auto.
+ intros. rewrite (Z.mul_comm c b); zero_or_not b.
+ rewrite (Z.mul_comm b c). apply Z.div_mul_cancel_l; auto.
Qed.
Lemma Zmult_mod_distr_l: forall a b c,
(c*a) mod (c*b) = c * (a mod b).
Proof.
- intros. zero_or_not c. rewrite (Zmult_comm c b); zero_or_not b.
- rewrite (Zmult_comm b c). apply Z.mul_mod_distr_l; auto.
+ intros. zero_or_not c. rewrite (Z.mul_comm c b); zero_or_not b.
+ rewrite (Z.mul_comm b c). apply Z.mul_mod_distr_l; auto.
Qed.
Lemma Zmult_mod_distr_r: forall a b c,
(a*c) mod (b*c) = (a mod b) * c.
Proof.
- intros. zero_or_not b. rewrite (Zmult_comm b c); zero_or_not c.
- rewrite (Zmult_comm c b). apply Z.mul_mod_distr_r; auto.
+ intros. zero_or_not b. rewrite (Z.mul_comm b c); zero_or_not c.
+ rewrite (Z.mul_comm c b). apply Z.mul_mod_distr_r; auto.
Qed.
(** Operations modulo. *)
@@ -464,22 +464,22 @@ Proof.
constructor; [exact eqm_refl | exact eqm_sym | exact eqm_trans].
Qed.
-Instance Zplus_eqm : Proper (eqm ==> eqm ==> eqm) Zplus.
+Instance Zplus_eqm : Proper (eqm ==> eqm ==> eqm) Z.add.
Proof.
unfold eqm; repeat red; intros. rewrite Zplus_mod, H, H0, <- Zplus_mod; auto.
Qed.
-Instance Zminus_eqm : Proper (eqm ==> eqm ==> eqm) Zminus.
+Instance Zminus_eqm : Proper (eqm ==> eqm ==> eqm) Z.sub.
Proof.
unfold eqm; repeat red; intros. rewrite Zminus_mod, H, H0, <- Zminus_mod; auto.
Qed.
-Instance Zmult_eqm : Proper (eqm ==> eqm ==> eqm) Zmult.
+Instance Zmult_eqm : Proper (eqm ==> eqm ==> eqm) Z.mul.
Proof.
unfold eqm; repeat red; intros. rewrite Zmult_mod, H, H0, <- Zmult_mod; auto.
Qed.
-Instance Zopp_eqm : Proper (eqm ==> eqm) Zopp.
+Instance Zopp_eqm : Proper (eqm ==> eqm) Z.opp.
Proof.
intros x y H. change ((-x)==(-y)) with ((0-x)==(0-y)). now rewrite H.
Qed.
@@ -489,7 +489,7 @@ Proof.
intros; exact (Zmod_mod a N).
Qed.
-(* NB: Zmod and Zdiv are not morphisms with respect to eqm.
+(* NB: Z.modulo and Z.div are not morphisms with respect to eqm.
For instance, let (==) be (eqm 2). Then we have (3 == 1) but:
~ (3 mod 3 == 1 mod 3)
~ (1 mod 3 == 1 mod 1)
@@ -501,7 +501,7 @@ End EqualityModulo.
Lemma Zdiv_Zdiv : forall a b c, 0<=b -> 0<=c -> (a/b)/c = a/(b*c).
Proof.
- intros. zero_or_not b. rewrite Zmult_comm. zero_or_not c.
+ intros. zero_or_not b. rewrite Z.mul_comm. zero_or_not c.
rewrite Z.mul_comm. apply Z.div_div; auto with zarith.
Qed.
@@ -515,7 +515,7 @@ Theorem Zdiv_mult_le:
Proof.
intros. zero_or_not b. apply Z.div_mul_le; auto with zarith. Qed.
-(** Zmod is related to divisibility (see more in Znumtheory) *)
+(** Z.modulo is related to divisibility (see more in Znumtheory) *)
Lemma Zmod_divides : forall a b, b<>0 ->
(a mod b = 0 <-> exists c, a = b*c).
@@ -536,17 +536,17 @@ Qed.
Lemma Zmod_even : forall a, a mod 2 = if Z.even a then 0 else 1.
Proof.
- intros a. rewrite Zmod_odd, Zodd_even_bool. now destruct Zeven_bool.
+ intros a. rewrite Zmod_odd, Zodd_even_bool. now destruct Z.even.
Qed.
Lemma Zodd_mod : forall a, Z.odd a = Zeq_bool (a mod 2) 1.
Proof.
- intros a. rewrite Zmod_odd. now destruct Zodd_bool.
+ intros a. rewrite Zmod_odd. now destruct Z.odd.
Qed.
Lemma Zeven_mod : forall a, Z.even a = Zeq_bool (a mod 2) 0.
Proof.
- intros a. rewrite Zmod_even. now destruct Zeven_bool.
+ intros a. rewrite Zmod_even. now destruct Z.even.
Qed.
(** * Compatibility *)
@@ -593,7 +593,7 @@ Proof.
intros; apply Z_mod_zero_opp_full; auto with zarith.
Qed.
-(** * A direct way to compute Zmod *)
+(** * A direct way to compute Z.modulo *)
Fixpoint Zmod_POS (a : positive) (b : Z) : Z :=
match a with
@@ -675,7 +675,7 @@ Proof.
exists (- q, r).
elim Hqr; intros.
split.
- rewrite <- Zmult_opp_comm; assumption.
+ rewrite <- Z.mul_opp_comm; assumption.
rewrite Z.abs_neq; [ assumption | omega ].
Qed.
@@ -692,7 +692,7 @@ Proof.
apply (Zdiv_unique _ _ _ (Z.of_nat (n mod m))).
split. auto with zarith.
now apply inj_lt, Nat.mod_upper_bound.
- rewrite <- inj_mult, <- inj_plus.
+ rewrite <- Nat2Z.inj_mul, <- Nat2Z.inj_add.
now apply inj_eq, Nat.div_mod.
Qed.
@@ -703,6 +703,6 @@ Proof.
apply (Zmod_unique _ _ (Z.of_nat n / Z.of_nat m)).
split. auto with zarith.
now apply inj_lt, Nat.mod_upper_bound.
- rewrite <- div_Zdiv, <- inj_mult, <- inj_plus by trivial.
+ rewrite <- div_Zdiv, <- Nat2Z.inj_mul, <- Nat2Z.inj_add by trivial.
now apply inj_eq, Nat.div_mod.
Qed.
diff --git a/theories/ZArith/Zeven.v b/theories/ZArith/Zeven.v
index a032a801d..bef8ee78b 100644
--- a/theories/ZArith/Zeven.v
+++ b/theories/ZArith/Zeven.v
@@ -136,7 +136,7 @@ Notation Zodd_bool_succ := Z.odd_succ (compat "8.3").
Notation Zodd_bool_pred := Z.odd_pred (compat "8.3").
(******************************************************************)
-(** * Definition of [Zquot2], [Zdiv2] and properties wrt [Zeven]
+(** * Definition of [Z.quot2], [Z.div2] and properties wrt [Zeven]
and [Zodd] *)
Notation Zdiv2 := Z.div2 (compat "8.3").
@@ -225,7 +225,7 @@ Lemma Zsplit2 n :
{p : Z * Z | let (x1, x2) := p in n = x1 + x2 /\ (x1 = x2 \/ x2 = x1 + 1)}.
Proof.
destruct (Z_modulo_2 n) as [(y,Hy)|(y,Hy)];
- rewrite Z.mul_comm, <- Zplus_diag_eq_mult_2 in Hy.
+ rewrite <- Z.add_diag in Hy.
- exists (y, y). split. assumption. now left.
- exists (y, y + 1). split. now rewrite Z.add_assoc. now right.
Qed.
diff --git a/theories/ZArith/Zgcd_alt.v b/theories/ZArith/Zgcd_alt.v
index ebf3d0246..a1fedd536 100644
--- a/theories/ZArith/Zgcd_alt.v
+++ b/theories/ZArith/Zgcd_alt.v
@@ -6,14 +6,14 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(** * Zgcd_alt : an alternate version of Zgcd, based on Euclid's algorithm *)
+(** * Zgcd_alt : an alternate version of Z.gcd, based on Euclid's algorithm *)
(**
Author: Pierre Letouzey
*)
-(** The alternate [Zgcd_alt] given here used to be the main [Zgcd]
- function (see file [Znumtheory]), but this main [Zgcd] is now
+(** The alternate [Zgcd_alt] given here used to be the main [Z.gcd]
+ function (see file [Znumtheory]), but this main [Z.gcd] is now
based on a modern binary-efficient algorithm. This earlier
version, based on Euclid's algorithm of iterated modulo, is kept
here due to both its intrinsic interest and its use as reference
@@ -35,22 +35,22 @@ Open Scope Z_scope.
match n with
| O => 1 (* arbitrary, since n should be big enough *)
| S n => match a with
- | Z0 => Zabs b
- | Zpos _ => Zgcdn n (Zmod b a) a
- | Zneg a => Zgcdn n (Zmod b (Zpos a)) (Zpos a)
+ | Z0 => Z.abs b
+ | Zpos _ => Zgcdn n (Z.modulo b a) a
+ | Zneg a => Zgcdn n (Z.modulo b (Zpos a)) (Zpos a)
end
end.
Definition Zgcd_bound (a:Z) :=
match a with
| Z0 => S O
- | Zpos p => let n := Psize p in (n+n)%nat
- | Zneg p => let n := Psize p in (n+n)%nat
+ | Zpos p => let n := Pos.size_nat p in (n+n)%nat
+ | Zneg p => let n := Pos.size_nat p in (n+n)%nat
end.
Definition Zgcd_alt a b := Zgcdn (Zgcd_bound a) a b.
- (** A first obvious fact : [Zgcd a b] is positive. *)
+ (** A first obvious fact : [Z.gcd a b] is positive. *)
Lemma Zgcdn_pos : forall n a b,
0 <= Zgcdn n a b.
@@ -62,28 +62,28 @@ Open Scope Z_scope.
Lemma Zgcd_alt_pos : forall a b, 0 <= Zgcd_alt a b.
Proof.
- intros; unfold Zgcd; apply Zgcdn_pos; auto.
+ intros; unfold Z.gcd; apply Zgcdn_pos; auto.
Qed.
- (** We now prove that Zgcd is indeed a gcd. *)
+ (** We now prove that Z.gcd is indeed a gcd. *)
(** 1) We prove a weaker & easier bound. *)
Lemma Zgcdn_linear_bound : forall n a b,
- Zabs a < Z_of_nat n -> Zis_gcd a b (Zgcdn n a b).
+ Z.abs a < Z.of_nat n -> Zis_gcd a b (Zgcdn n a b).
Proof.
induction n.
simpl; intros.
- exfalso; generalize (Zabs_pos a); omega.
+ exfalso; generalize (Z.abs_nonneg a); omega.
destruct a; intros; simpl;
[ generalize (Zis_gcd_0_abs b); intuition | | ];
- unfold Zmod;
- generalize (Z_div_mod b (Zpos p) (refl_equal Gt));
- destruct (Zdiv_eucl b (Zpos p)) as (q,r);
+ unfold Z.modulo;
+ generalize (Z_div_mod b (Zpos p) (eq_refl Gt));
+ destruct (Z.div_eucl b (Zpos p)) as (q,r);
intros (H0,H1);
- rewrite inj_S in H; simpl Zabs in H;
- (assert (H2: Zabs r < Z_of_nat n) by
- (rewrite Zabs_eq; auto with zarith));
+ rewrite Nat2Z.inj_succ in H; simpl Z.abs in H;
+ (assert (H2: Z.abs r < Z.of_nat n) by
+ (rewrite Z.abs_eq; auto with zarith));
assert (IH:=IHn r (Zpos p) H2); clear IHn;
simpl in IH |- *;
rewrite H0.
@@ -122,7 +122,7 @@ Open Scope Z_scope.
Proof.
induction 1.
auto with zarith.
- apply Zle_trans with (fibonacci m); auto.
+ apply Z.le_trans with (fibonacci m); auto.
clear.
destruct m.
simpl; auto with zarith.
@@ -142,53 +142,38 @@ Open Scope Z_scope.
fibonacci (S (S n)) <= b.
Proof.
induction n.
- simpl; intros.
- destruct a; omega.
- intros.
- destruct a; [simpl in *; omega| | destruct H; discriminate].
- revert H1; revert H0.
- set (m:=S n) in *; (assert (m=S n) by auto); clearbody m.
- pattern m at 2; rewrite H0.
- simpl Zgcdn.
- unfold Zmod; generalize (Z_div_mod b (Zpos p) (refl_equal Gt)).
- destruct (Zdiv_eucl b (Zpos p)) as (q,r).
- intros (H1,H2).
- destruct H2.
- destruct (Zle_lt_or_eq _ _ H2).
- generalize (IHn _ _ (conj H4 H3)).
- intros H5 H6 H7.
- replace (fibonacci (S (S m))) with (fibonacci (S m) + fibonacci m) by auto.
- assert (r = Zpos p * (-q) + b) by (rewrite H1; ring).
- destruct H5; auto.
- pattern r at 1; rewrite H8.
- apply Zis_gcd_sym.
- apply Zis_gcd_for_euclid2; auto.
- apply Zis_gcd_sym; auto.
- split; auto.
- rewrite H1.
- apply Zplus_le_compat; auto.
- apply Zle_trans with (Zpos p * 1); auto.
- ring_simplify (Zpos p * 1); auto.
- apply Zmult_le_compat_l.
- destruct q.
- omega.
- assert (0 < Zpos p0) by (compute; auto).
- omega.
- assert (Zpos p * Zneg p0 < 0) by (compute; auto).
- omega.
- compute; intros; discriminate.
- (* r=0 *)
- subst r.
- simpl; rewrite H0.
- intros.
- simpl in H4.
- simpl in H5.
- destruct n.
- simpl in H5.
- simpl.
- omega.
- simpl in H5.
- elim H5; auto.
+ intros [|a|a]; intros; simpl; omega.
+ intros [|a|a] b (Ha,Ha'); [simpl; omega | | easy ].
+ remember (S n) as m.
+ rewrite Heqm at 2. simpl Zgcdn.
+ unfold Z.modulo; generalize (Z_div_mod b (Zpos a) eq_refl).
+ destruct (Z.div_eucl b (Zpos a)) as (q,r).
+ intros (EQ,(Hr,Hr')).
+ Z.le_elim Hr.
+ - (* r > 0 *)
+ replace (fibonacci (S (S m))) with (fibonacci (S m) + fibonacci m) by auto.
+ intros.
+ destruct (IHn r (Zpos a) (conj Hr Hr')); auto.
+ + assert (EQ' : r = Zpos a * (-q) + b) by (rewrite EQ; ring).
+ rewrite EQ' at 1.
+ apply Zis_gcd_sym.
+ apply Zis_gcd_for_euclid2; auto.
+ apply Zis_gcd_sym; auto.
+ + split; auto.
+ rewrite EQ.
+ apply Z.add_le_mono; auto.
+ apply Z.le_trans with (Zpos a * 1); auto.
+ now rewrite Z.mul_1_r.
+ apply Z.mul_le_mono_nonneg_l; auto with zarith.
+ change 1 with (Z.succ 0). apply Z.le_succ_l.
+ destruct q; auto with zarith.
+ assert (Zpos a * Zneg p < 0) by now compute. omega.
+ - (* r = 0 *)
+ clear IHn EQ Hr'; intros _.
+ subst r; simpl; rewrite Heqm.
+ destruct n.
+ + simpl. omega.
+ + now destruct 1.
Qed.
(** 3b) We reformulate the previous result in a more positive way. *)
@@ -199,18 +184,18 @@ Open Scope Z_scope.
Proof.
destruct a; [ destruct 1; exfalso; omega | | destruct 1; discriminate].
cut (forall k n b,
- k = (S (nat_of_P p) - n)%nat ->
+ k = (S (Pos.to_nat p) - n)%nat ->
0 < Zpos p < b -> Zpos p < fibonacci (S n) ->
Zis_gcd (Zpos p) b (Zgcdn n (Zpos p) b)).
destruct 2; eauto.
clear n; induction k.
intros.
- assert (nat_of_P p < n)%nat by omega.
+ assert (Pos.to_nat p < n)%nat by omega.
apply Zgcdn_linear_bound.
simpl.
generalize (inj_le _ _ H2).
- rewrite inj_S.
- rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P; auto.
+ rewrite Nat2Z.inj_succ.
+ rewrite positive_nat_Z; auto.
omega.
intros.
generalize (Zgcdn_worst_is_fibonacci n (Zpos p) b H0); intros.
@@ -233,77 +218,69 @@ Open Scope Z_scope.
induction p; [ | | compute; auto ];
simpl Zgcd_bound in *;
rewrite plus_comm; simpl plus;
- set (n:= (Psize p+Psize p)%nat) in *; simpl;
+ set (n:= (Pos.size_nat p+Pos.size_nat p)%nat) in *; simpl;
assert (n <> O) by (unfold n; destruct p; simpl; auto).
destruct n as [ |m]; [elim H; auto| ].
- generalize (fibonacci_pos m); rewrite Zpos_xI; omega.
+ generalize (fibonacci_pos m); rewrite Pos2Z.inj_xI; omega.
destruct n as [ |m]; [elim H; auto| ].
- generalize (fibonacci_pos m); rewrite Zpos_xO; omega.
+ generalize (fibonacci_pos m); rewrite Pos2Z.inj_xO; omega.
Qed.
(* 5) the end: we glue everything together and take care of
situations not corresponding to [0<a<b]. *)
- Lemma Zgcdn_is_gcd :
- forall n a b, (Zgcd_bound a <= n)%nat ->
- Zis_gcd a b (Zgcdn n a b).
+ Lemma Zgcd_bound_opp a : Zgcd_bound (-a) = Zgcd_bound a.
+ Proof.
+ now destruct a.
+ Qed.
+
+ Lemma Zgcdn_opp n a b : Zgcdn n (-a) b = Zgcdn n a b.
+ Proof.
+ induction n; simpl; auto.
+ destruct a; simpl; auto.
+ Qed.
+
+ Lemma Zgcdn_is_gcd_pos n a b : (Zgcd_bound (Zpos a) <= n)%nat ->
+ Zis_gcd (Zpos a) b (Zgcdn n (Zpos a) b).
+ Proof.
+ intros.
+ generalize (Zgcd_bound_fibonacci (Zpos a)).
+ simpl Zgcd_bound in *.
+ remember (Pos.size_nat a+Pos.size_nat a)%nat as m.
+ assert (1 < m)%nat.
+ { rewrite Heqm; destruct a; simpl; rewrite 1?plus_comm;
+ auto with arith. }
+ destruct m as [ |m]; [inversion H0; auto| ].
+ destruct n as [ |n]; [inversion H; auto| ].
+ simpl Zgcdn.
+ unfold Z.modulo.
+ generalize (Z_div_mod b (Zpos a) (eq_refl Gt)).
+ destruct (Z.div_eucl b (Zpos a)) as (q,r).
+ intros (->,(H1,H2)) H3.
+ apply Zis_gcd_for_euclid2.
+ Z.le_elim H1.
+ + apply Zgcdn_ok_before_fibonacci; auto.
+ apply Z.lt_le_trans with (fibonacci (S m));
+ [ omega | apply fibonacci_incr; auto].
+ + subst r; simpl.
+ destruct m as [ |m]; [exfalso; omega| ].
+ destruct n as [ |n]; [exfalso; omega| ].
+ simpl; apply Zis_gcd_sym; apply Zis_gcd_0.
+ Qed.
+
+ Lemma Zgcdn_is_gcd n a b :
+ (Zgcd_bound a <= n)%nat -> Zis_gcd a b (Zgcdn n a b).
Proof.
- destruct a; intros.
- simpl in H.
- destruct n; [exfalso; omega | ].
- simpl; generalize (Zis_gcd_0_abs b); intuition.
- (*Zpos*)
- generalize (Zgcd_bound_fibonacci (Zpos p)).
- simpl Zgcd_bound in *.
- remember (Psize p+Psize p)%nat as m.
- assert (1 < m)%nat.
- rewrite Heqm; destruct p; simpl; rewrite 1? plus_comm;
- auto with arith.
- destruct m as [ |m]; [inversion H0; auto| ].
- destruct n as [ |n]; [inversion H; auto| ].
- simpl Zgcdn.
- unfold Zmod.
- generalize (Z_div_mod b (Zpos p) (refl_equal Gt)).
- destruct (Zdiv_eucl b (Zpos p)) as (q,r).
- intros (H2,H3) H4.
- rewrite H2.
- apply Zis_gcd_for_euclid2.
- destruct H3.
- destruct (Zle_lt_or_eq _ _ H1).
- apply Zgcdn_ok_before_fibonacci; auto.
- apply Zlt_le_trans with (fibonacci (S m)); [ omega | apply fibonacci_incr; auto].
- subst r; simpl.
- destruct m as [ |m]; [exfalso; omega| ].
- destruct n as [ |n]; [exfalso; omega| ].
- simpl; apply Zis_gcd_sym; apply Zis_gcd_0.
- (*Zneg*)
- generalize (Zgcd_bound_fibonacci (Zpos p)).
- simpl Zgcd_bound in *.
- remember (Psize p+Psize p)%nat as m.
- assert (1 < m)%nat.
- rewrite Heqm; destruct p; simpl; rewrite 1? plus_comm;
- auto with arith.
- destruct m as [ |m]; [inversion H0; auto| ].
- destruct n as [ |n]; [inversion H; auto| ].
- simpl Zgcdn.
- unfold Zmod.
- generalize (Z_div_mod b (Zpos p) (refl_equal Gt)).
- destruct (Zdiv_eucl b (Zpos p)) as (q,r).
- intros (H1,H2) H3.
- rewrite H1.
- apply Zis_gcd_minus.
- apply Zis_gcd_sym.
- apply Zis_gcd_for_euclid2.
- destruct H2.
- destruct (Zle_lt_or_eq _ _ H2).
- apply Zgcdn_ok_before_fibonacci; auto.
- apply Zlt_le_trans with (fibonacci (S m)); [ omega | apply fibonacci_incr; auto].
- subst r; simpl.
- destruct m as [ |m]; [exfalso; omega| ].
- destruct n as [ |n]; [exfalso; omega| ].
- simpl; apply Zis_gcd_sym; apply Zis_gcd_0.
+ destruct a.
+ - simpl; intros.
+ destruct n; [exfalso; omega | ].
+ simpl; generalize (Zis_gcd_0_abs b); intuition.
+ - apply Zgcdn_is_gcd_pos.
+ - rewrite <- Zgcd_bound_opp, <- Zgcdn_opp.
+ intros. apply Zis_gcd_minus, Zis_gcd_sym. simpl Z.opp.
+ now apply Zgcdn_is_gcd_pos.
Qed.
Lemma Zgcd_is_gcd :
diff --git a/theories/ZArith/Zhints.v b/theories/ZArith/Zhints.v
index 6a14d6934..e50a5e445 100644
--- a/theories/ZArith/Zhints.v
+++ b/theories/ZArith/Zhints.v
@@ -43,58 +43,59 @@ Hint Resolve
(** Should clearly be declared as hints *)
(** Lemmas ending by eq *)
- Zsucc_eq_compat (* :(n,m:Z)`n = m`->`(Zs n) = (Zs m)` *)
-
- (** Lemmas ending by Zgt *)
- 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_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 *)
- 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` *)
- 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|` *)
+ Zsucc_eq_compat (* n = m -> Z.succ n = Z.succ m *)
+
+ (** Lemmas ending by Z.gt *)
+ Zsucc_gt_compat (* m > n -> Z.succ m > Z.succ n *)
+ Zgt_succ (* Z.succ n > n *)
+ Zorder.Zgt_pos_0 (* Z.pos p > 0 *)
+ Zplus_gt_compat_l (* n > m -> p+n > p+m *)
+ Zplus_gt_compat_r (* n > m -> n+p > m+p *)
+
+ (** Lemmas ending by Z.lt *)
+ Pos2Z.is_pos (* 0 < Z.pos p *)
+ Z.lt_succ_diag_r (* n < Z.succ n *)
+ Zsucc_lt_compat (* n < m -> Z.succ n < Z.succ m *)
+ Z.lt_pred_l (* Z.pred n < n *)
+ Zplus_lt_compat_l (* n < m -> p+n < p+m *)
+ Zplus_lt_compat_r (* n < m -> n+p < m+p *)
+
+ (** Lemmas ending by Z.le *)
+ Nat2Z.is_nonneg (* 0 <= Z.of_nat n *)
+ Pos2Z.is_nonneg (* 0 <= Z.pos p *)
+ Z.le_refl (* n <= n *)
+ Z.le_succ_diag_r (* n <= Z.succ n *)
+ Zsucc_le_compat (* m <= n -> Z.succ m <= Z.succ n *)
+ Z.le_pred_l (* Z.pred n <= n *)
+ Z.le_min_l (* Z.min n m <= n *)
+ Z.le_min_r (* Z.min n m <= m *)
+ Zplus_le_compat_l (* n <= m -> p+n <= p+m *)
+ Zplus_le_compat_r (* a <= b -> a+c <= b+c *)
+ Z.abs_nonneg (* 0 <= |x| *)
(** ** Irreversible simplification lemmas *)
(** Probably to be declared as hints, when no other simplification is possible *)
(** Lemmas ending by eq *)
- 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` *)
+ Z_eq_mult (* y = 0 -> y*x = 0 *)
+ Zplus_eq_compat (* n = m -> p = q -> n+p = m+q *)
- (** Lemmas ending by Zge *)
- 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` *)
+ (** Lemmas ending by Z.ge *)
+ Zorder.Zmult_ge_compat_r (* a >= b -> c >= 0 -> a*c >= b*c *)
+ Zorder.Zmult_ge_compat_l (* 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 *)
- 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 *)
- 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` *)
+ a >= c -> b >= d -> c >= 0 -> d >= 0 -> a*b >= c*d *)
+
+ (** Lemmas ending by Z.lt *)
+ Zorder.Zmult_gt_0_compat (* a > 0 -> b > 0 -> a*b > 0 *)
+ Z.lt_lt_succ_r (* n < m -> n < Z.succ m *)
+
+ (** Lemmas ending by Z.le *)
+ Z.mul_nonneg_nonneg (* 0 <= x -> 0 <= y -> 0 <= x*y *)
+ Zorder.Zmult_le_compat_r (* a <= b -> 0 <= c -> a*c <= b*c *)
+ Zorder.Zmult_le_compat_l (* a <= b -> 0 <= c -> c*a <= c*b *)
+ Z.add_nonneg_nonneg (* 0 <= x -> 0 <= y -> 0 <= x+y *)
+ Z.le_le_succ_r (* x <= y -> x <= Z.succ y *)
+ Z.add_le_mono (* n <= m -> p <= q -> n+p <= m+q *)
: zarith.
diff --git a/theories/ZArith/Zlogarithm.v b/theories/ZArith/Zlogarithm.v
index 30948ca7a..3711ea021 100644
--- a/theories/ZArith/Zlogarithm.v
+++ b/theories/ZArith/Zlogarithm.v
@@ -34,22 +34,22 @@ Section Log_pos. (* Log of positive integers *)
Fixpoint log_inf (p:positive) : Z :=
match p with
| xH => 0 (* 1 *)
- | xO q => Zsucc (log_inf q) (* 2n *)
- | xI q => Zsucc (log_inf q) (* 2n+1 *)
+ | xO q => Z.succ (log_inf q) (* 2n *)
+ | xI q => Z.succ (log_inf q) (* 2n+1 *)
end.
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 *)
+ | xO n => Z.succ (log_sup n) (* 2n *)
+ | xI n => Z.succ (Z.succ (log_inf n)) (* 2n+1 *)
end.
Hint Unfold log_inf log_sup.
Lemma Psize_log_inf : forall p, Zpos (Pos.size p) = Z.succ (log_inf p).
Proof.
- induction p; simpl; now rewrite <- ?Z.succ_Zpos, ?IHp.
+ induction p; simpl; now rewrite ?Pos2Z.inj_succ, ?IHp.
Qed.
Lemma Zlog2_log_inf : forall p, Z.log2 (Zpos p) = log_inf p.
@@ -71,24 +71,24 @@ Section Log_pos. (* Log of positive integers *)
(** Then we give the specifications of [log_inf] and [log_sup]
and prove their validity *)
- Hint Resolve Zle_trans: zarith.
+ Hint Resolve Z.le_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)).
+ 0 <= log_inf x /\ two_p (log_inf x) <= Zpos x < two_p (Z.succ (log_inf x)).
Proof.
simple induction x; intros; simpl in |- *;
[ elim H; intros Hp HR; clear H; split;
[ auto with zarith
- | rewrite two_p_S with (x := Zsucc (log_inf p)) by (apply Zle_le_succ; trivial);
+ | rewrite two_p_S with (x := Z.succ (log_inf p)) by (apply Z.le_le_succ_r; trivial);
rewrite two_p_S by trivial;
- rewrite two_p_S in HR by trivial; rewrite (BinInt.Zpos_xI p);
+ rewrite two_p_S in HR by trivial; rewrite (BinInt.Pos2Z.inj_xI p);
omega ]
| elim H; intros Hp HR; clear H; split;
[ auto with zarith
- | rewrite two_p_S with (x := Zsucc (log_inf p)) by (apply Zle_le_succ; trivial);
+ | rewrite two_p_S with (x := Z.succ (log_inf p)) by (apply Z.le_le_succ_r; trivial);
rewrite two_p_S by trivial;
- rewrite two_p_S in HR by trivial; rewrite (BinInt.Zpos_xO p);
+ rewrite two_p_S in HR by trivial; rewrite (BinInt.Pos2Z.inj_xO p);
omega ]
| unfold two_power_pos in |- *; unfold shift_pos in |- *; simpl in |- *;
omega ].
@@ -112,12 +112,12 @@ Section Log_pos. (* Log of positive integers *)
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).
+ else log_sup p = Z.succ (log_inf p).
Proof.
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
+ rewrite BinInt.Pos2Z.inj_xI; unfold Z.succ; omega
| elim H; clear H; intro Hif;
[ left; simpl in |- *;
rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0));
@@ -126,21 +126,21 @@ Section Log_pos. (* Log of positive integers *)
auto
| right; simpl in |- *;
rewrite (two_p_S (log_inf p0) (log_inf_correct1 p0));
- rewrite BinInt.Zpos_xO; unfold Zsucc in |- *;
+ rewrite BinInt.Pos2Z.inj_xO; unfold Z.succ;
omega ]
| left; auto ].
Qed.
Theorem log_sup_correct2 :
- forall x:positive, two_p (Zpred (log_sup x)) < Zpos x <= two_p (log_sup x).
+ forall x:positive, two_p (Z.pred (log_sup x)) < Zpos x <= two_p (log_sup x).
Proof.
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_refl ].
+ split; [ apply two_p_pred; apply log_sup_correct1 | apply Z.le_refl ].
intros [E1 E2]; rewrite E2.
- rewrite <- (Zpred_succ (log_inf x)).
+ rewrite (Z.pred_succ (log_inf x)).
generalize (log_inf_correct2 x); omega.
Qed.
@@ -149,7 +149,7 @@ Section Log_pos. (* Log of positive integers *)
simple induction p; simpl in |- *; intros; omega.
Qed.
- Lemma log_sup_le_Slog_inf : forall p:positive, log_sup p <= Zsucc (log_inf p).
+ Lemma log_sup_le_Slog_inf : forall p:positive, log_sup p <= Z.succ (log_inf p).
Proof.
simple induction p; simpl in |- *; intros; omega.
Qed.
@@ -161,8 +161,8 @@ Section Log_pos. (* Log of positive integers *)
| xH => 0
| xO xH => 1
| xI xH => 2
- | xO y => Zsucc (log_near y)
- | xI y => Zsucc (log_near y)
+ | xO y => Z.succ (log_near y)
+ | xI y => Z.succ (log_near y)
end.
Theorem log_near_correct1 : forall p:positive, 0 <= log_near p.
@@ -171,12 +171,10 @@ Section Log_pos. (* Log of positive integers *)
[ 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 ].
+ intros; apply Z.le_le_succ_r.
+ generalize H0; now elim p1.
+ intros; apply Z.le_le_succ_r.
+ generalize H0; now elim p1.
Qed.
Theorem log_near_correct2 :
@@ -219,19 +217,19 @@ Section divers.
Lemma ZERO_le_N_digits : forall x:Z, 0 <= N_digits x.
Proof.
simple induction x; simpl in |- *;
- [ apply Zle_refl | exact log_inf_correct1 | exact log_inf_correct1 ].
+ [ apply Z.le_refl | exact log_inf_correct1 | exact log_inf_correct1 ].
Qed.
- Lemma log_inf_shift_nat : forall n:nat, log_inf (shift_nat n 1) = Z_of_nat n.
+ Lemma log_inf_shift_nat : forall n:nat, log_inf (shift_nat n 1) = Z.of_nat n.
Proof.
simple induction n; intros;
- [ try trivial | rewrite Znat.inj_S; rewrite <- H; reflexivity ].
+ [ try trivial | rewrite Nat2Z.inj_succ; rewrite <- H; reflexivity ].
Qed.
- Lemma log_sup_shift_nat : forall n:nat, log_sup (shift_nat n 1) = Z_of_nat n.
+ Lemma log_sup_shift_nat : forall n:nat, log_sup (shift_nat n 1) = Z.of_nat n.
Proof.
simple induction n; intros;
- [ try trivial | rewrite Znat.inj_S; rewrite <- H; reflexivity ].
+ [ try trivial | rewrite Nat2Z.inj_succ; rewrite <- H; reflexivity ].
Qed.
(** [Is_power p] means that p is a power of two *)
diff --git a/theories/ZArith/Zmax.v b/theories/ZArith/Zmax.v
index a3eeb4160..86e26a605 100644
--- a/theories/ZArith/Zmax.v
+++ b/theories/ZArith/Zmax.v
@@ -12,12 +12,38 @@ Require Export BinInt Zcompare Zorder.
Local Open Scope Z_scope.
-(** Definition [Zmax] is now [BinInt.Z.max]. *)
-
-(** * Characterization of maximum on binary integer numbers *)
+(** Definition [Z.max] is now [BinInt.Z.max]. *)
+
+(** Exact compatibility *)
+
+Notation Zmax_case := Z.max_case (compat "8.3").
+Notation Zmax_case_strong := Z.max_case_strong (compat "8.3").
+Notation Zmax_right := Z.max_r (compat "8.3").
+Notation Zle_max_l := Z.le_max_l (compat "8.3").
+Notation Zle_max_r := Z.le_max_r (compat "8.3").
+Notation Zmax_lub := Z.max_lub (compat "8.3").
+Notation Zmax_lub_lt := Z.max_lub_lt (compat "8.3").
+Notation Zle_max_compat_r := Z.max_le_compat_r (compat "8.3").
+Notation Zle_max_compat_l := Z.max_le_compat_l (compat "8.3").
+Notation Zmax_idempotent := Z.max_id (compat "8.3").
+Notation Zmax_n_n := Z.max_id (compat "8.3").
+Notation Zmax_comm := Z.max_comm (compat "8.3").
+Notation Zmax_assoc := Z.max_assoc (compat "8.3").
+Notation Zmax_irreducible_dec := Z.max_dec (compat "8.3").
+Notation Zmax_le_prime := Z.max_le (compat "8.3").
+Notation Zsucc_max_distr := Z.succ_max_distr (compat "8.3").
+Notation Zmax_SS := Z.succ_max_distr (compat "8.3").
+Notation Zplus_max_distr_l := Z.add_max_distr_l (compat "8.3").
+Notation Zplus_max_distr_r := Z.add_max_distr_r (compat "8.3").
+Notation Zmax_plus := Z.add_max_distr_r (compat "8.3").
+Notation Zmax1 := Z.le_max_l (compat "8.3").
+Notation Zmax2 := Z.le_max_r (compat "8.3").
+Notation Zmax_irreducible_inf := Z.max_dec (compat "8.3").
+Notation Zmax_le_prime_inf := Z.max_le (compat "8.3").
+Notation Zpos_max := Pos2Z.inj_max (compat "8.3").
+Notation Zpos_minus := Pos2Z.inj_sub_max (compat "8.3").
-Definition Zmax_case := Z.max_case.
-Definition Zmax_case_strong := Z.max_case_strong.
+(** Slightly different lemmas *)
Lemma Zmax_spec x y :
x >= y /\ Z.max x y = x \/ x < y /\ Z.max x y = y.
@@ -26,86 +52,9 @@ Proof.
Qed.
Lemma Zmax_left n m : n>=m -> Z.max n m = n.
-Proof. Z.swap_greater. apply Zmax_l. Qed.
-
-Lemma Zmax_right : forall n m, n<=m -> Z.max n m = m. Proof Zmax_r.
-
-(** * Least upper bound properties of max *)
-
-Lemma Zle_max_l : forall n m, n <= Z.max n m. Proof Z.le_max_l.
-Lemma Zle_max_r : forall n m, m <= Z.max n m. Proof Z.le_max_r.
-
-Lemma Zmax_lub : forall n m p, n <= p -> m <= p -> Z.max n m <= p.
-Proof Z.max_lub.
-
-Lemma Zmax_lub_lt : forall n m p:Z, n < p -> m < p -> Z.max n m < p.
-Proof Z.max_lub_lt.
-
-
-(** * Compatibility with order *)
-
-Lemma Zle_max_compat_r : forall n m p, n <= m -> Z.max n p <= Z.max m p.
-Proof Z.max_le_compat_r.
-
-Lemma Zle_max_compat_l : forall n m p, n <= m -> Z.max p n <= Z.max p m.
-Proof Z.max_le_compat_l.
-
-
-(** * Semi-lattice properties of max *)
-
-Lemma Zmax_idempotent : forall n, Z.max n n = n. Proof Z.max_id.
-Lemma Zmax_comm : forall n m, Z.max n m = Z.max m n. Proof Z.max_comm.
-Lemma Zmax_assoc : forall n m p, Z.max n (Z.max m p) = Z.max (Z.max n m) p.
-Proof Z.max_assoc.
-
-(** * Additional properties of max *)
-
-Lemma Zmax_irreducible_dec : forall n m, {Z.max n m = n} + {Z.max n m = m}.
-Proof Z.max_dec.
+Proof. Z.swap_greater. apply Z.max_l. Qed.
-Lemma Zmax_le_prime : forall n m p, p <= Z.max n m -> p <= n \/ p <= m.
-Proof Z.max_le.
-
-
-(** * Operations preserving max *)
-
-Lemma Zsucc_max_distr :
- forall n m, Z.succ (Z.max n m) = Z.max (Z.succ n) (Z.succ m).
-Proof Z.succ_max_distr.
-
-Lemma Zplus_max_distr_l : forall n m p, Z.max (p + n) (p + m) = p + Z.max n m.
-Proof Z.add_max_distr_l.
-
-Lemma Zplus_max_distr_r : forall n m p, Z.max (n + p) (m + p) = Z.max n m + p.
-Proof Z.add_max_distr_r.
-
-(** * Maximum and Zpos *)
-
-Lemma Zpos_max p q : Zpos (Pos.max p q) = Z.max (Zpos p) (Zpos q).
-Proof.
- unfold Zmax, Pmax. simpl.
- case Pos.compare_spec; auto; congruence.
-Qed.
-
-Lemma Zpos_max_1 p : Z.max 1 (Zpos p) = Zpos p.
+Lemma Zpos_max_1 p : Z.max 1 (Z.pos p) = Z.pos p.
Proof.
now destruct p.
Qed.
-
-(** * Characterization of Pos.sub in term of Z.sub and Z.max *)
-
-Lemma Zpos_minus p q : Zpos (p - q) = Z.max 1 (Zpos p - Zpos q).
-Proof.
- simpl. rewrite Z.pos_sub_spec. case Pos.compare_spec; intros H.
- subst; now rewrite Pos.sub_diag.
- now rewrite Pos.sub_lt.
- symmetry. apply Zpos_max_1.
-Qed.
-
-(* begin hide *)
-(* Compatibility *)
-Notation Zmax1 := Z.le_max_l (compat "8.3").
-Notation Zmax2 := Z.le_max_r (compat "8.3").
-Notation Zmax_irreducible_inf := Z.max_dec (compat "8.3").
-Notation Zmax_le_prime_inf := Z.max_le (compat "8.3").
-(* end hide *)
diff --git a/theories/ZArith/Zmin.v b/theories/ZArith/Zmin.v
index fbb31632c..d0b29f31f 100644
--- a/theories/ZArith/Zmin.v
+++ b/theories/ZArith/Zmin.v
@@ -12,12 +12,30 @@ Require Import BinInt Zcompare Zorder.
Local Open Scope Z_scope.
-(** Definition [Zmin] is now [BinInt.Z.min]. *)
-
-(** * Characterization of the minimum on binary integer numbers *)
+(** Definition [Z.min] is now [BinInt.Z.min]. *)
+
+(** Exact compatibility *)
+
+Notation Zmin_case := Z.min_case (compat "8.3").
+Notation Zmin_case_strong := Z.min_case_strong (compat "8.3").
+Notation Zle_min_l := Z.le_min_l (compat "8.3").
+Notation Zle_min_r := Z.le_min_r (compat "8.3").
+Notation Zmin_glb := Z.min_glb (compat "8.3").
+Notation Zmin_glb_lt := Z.min_glb_lt (compat "8.3").
+Notation Zle_min_compat_r := Z.min_le_compat_r (compat "8.3").
+Notation Zle_min_compat_l := Z.min_le_compat_l (compat "8.3").
+Notation Zmin_idempotent := Z.min_id (compat "8.3").
+Notation Zmin_n_n := Z.min_id (compat "8.3").
+Notation Zmin_comm := Z.min_comm (compat "8.3").
+Notation Zmin_assoc := Z.min_assoc (compat "8.3").
+Notation Zmin_irreducible_inf := Z.min_dec (compat "8.3").
+Notation Zsucc_min_distr := Z.succ_min_distr (compat "8.3").
+Notation Zmin_SS := Z.succ_min_distr (compat "8.3").
+Notation Zplus_min_distr_r := Z.add_min_distr_r (compat "8.3").
+Notation Zmin_plus := Z.add_min_distr_r (compat "8.3").
+Notation Zpos_min := Pos2Z.inj_min (compat "8.3").
-Definition Zmin_case := Z.min_case.
-Definition Zmin_case_strong := Z.min_case_strong.
+(** Slightly different lemmas *)
Lemma Zmin_spec x y :
x <= y /\ Z.min x y = x \/ x > y /\ Z.min x y = y.
@@ -25,71 +43,15 @@ Proof.
Z.swap_greater. rewrite Z.min_comm. destruct (Z.min_spec y x); auto.
Qed.
-(** * Greatest lower bound properties of min *)
-
-Lemma Zle_min_l : forall n m, Z.min n m <= n. Proof Z.le_min_l.
-Lemma Zle_min_r : forall n m, Z.min n m <= m. Proof Z.le_min_r.
-
-Lemma Zmin_glb : forall n m p, p <= n -> p <= m -> p <= Z.min n m.
-Proof Z.min_glb.
-Lemma Zmin_glb_lt : forall n m p, p < n -> p < m -> p < Z.min n m.
-Proof Z.min_glb_lt.
-
-(** * Compatibility with order *)
-
-Lemma Zle_min_compat_r : forall n m p, n <= m -> Z.min n p <= Z.min m p.
-Proof Z.min_le_compat_r.
-Lemma Zle_min_compat_l : forall n m p, n <= m -> Z.min p n <= Z.min p m.
-Proof Z.min_le_compat_l.
-
-(** * Semi-lattice properties of min *)
-
-Lemma Zmin_idempotent : forall n, Z.min n n = n. Proof Z.min_id.
-Notation Zmin_n_n := Z.min_id (compat "8.3").
-Lemma Zmin_comm : forall n m, Z.min n m = Z.min m n. Proof Z.min_comm.
-Lemma Zmin_assoc : forall n m p, Z.min n (Z.min m p) = Z.min (Z.min n m) p.
-Proof Z.min_assoc.
-
-(** * Additional properties of min *)
-
-Lemma Zmin_irreducible_inf : forall n m, {Z.min n m = n} + {Z.min n m = m}.
-Proof Z.min_dec.
-
Lemma Zmin_irreducible n m : Z.min n m = n \/ Z.min n m = m.
Proof. destruct (Z.min_dec n m); auto. Qed.
-Notation Zmin_or := Zmin_irreducible (only parsing).
+Notation Zmin_or := Zmin_irreducible (compat "8.3").
Lemma Zmin_le_prime_inf n m p : Z.min n m <= p -> {n <= p} + {m <= p}.
-Proof. apply Zmin_case; auto. Qed.
-
-(** * Operations preserving min *)
-
-Lemma Zsucc_min_distr :
- forall n m, Z.succ (Z.min n m) = Z.min (Z.succ n) (Z.succ m).
-Proof Z.succ_min_distr.
-
-Notation Zmin_SS := Z.succ_min_distr (only parsing).
-
-Lemma Zplus_min_distr_r :
- forall n m p, Z.min (n + p) (m + p) = Z.min n m + p.
-Proof Z.add_min_distr_r.
-
-Notation Zmin_plus := Z.add_min_distr_r (only parsing).
-
-(** * Minimum and Zpos *)
-
-Lemma Zpos_min p q : Zpos (Pos.min p q) = Z.min (Zpos p) (Zpos q).
-Proof.
- unfold Z.min, Pos.min; simpl. destruct Pos.compare; auto.
-Qed.
+Proof. apply Z.min_case; auto. Qed.
Lemma Zpos_min_1 p : Z.min 1 (Zpos p) = 1.
Proof.
now destruct p.
Qed.
-
-
-
-
-
diff --git a/theories/ZArith/Zmisc.v b/theories/ZArith/Zmisc.v
index a6da9d7e7..bfd57d758 100644
--- a/theories/ZArith/Zmisc.v
+++ b/theories/ZArith/Zmisc.v
@@ -21,8 +21,9 @@ Open Local Scope Z_scope.
Notation iter := @Z.iter (compat "8.3").
Lemma iter_nat_of_Z : forall n A f x, 0 <= n ->
- iter n A f x = iter_nat (Z.abs_nat n) A f x.
+ Z.iter n f x = iter_nat (Z.abs_nat n) A f x.
+Proof.
intros n A f x; case n; auto.
-intros p _; unfold Z.iter, Z.abs_nat; apply iter_nat_of_P.
+intros p _; unfold Z.iter, Z.abs_nat; apply Pos2Nat.inj_iter.
intros p abs; case abs; trivial.
Qed.
diff --git a/theories/ZArith/Znat.v b/theories/ZArith/Znat.v
index 2c3288f6c..7cb2b7060 100644
--- a/theories/ZArith/Znat.v
+++ b/theories/ZArith/Znat.v
@@ -14,6 +14,18 @@ Require Import BinPos BinInt BinNat Pnat Nnat.
Local Open Scope Z_scope.
+(** Conversions between integers and natural numbers
+
+ Seven sections:
+ - chains of conversions (combining two conversions)
+ - module N2Z : from N to Z
+ - module Z2N : from Z to N (negative numbers rounded to 0)
+ - module Zabs2N : from Z to N (via the absolute value)
+ - module Nat2Z : from nat to Z
+ - module Z2Nat : from Z to nat (negative numbers rounded to 0)
+ - module Zabs2Nat : from Z to nat (via the absolute value)
+*)
+
(** * Chains of conversions *)
(** When combining successive conversions, we have the following
@@ -254,9 +266,13 @@ Qed.
Lemma inj_pow n m : Z.of_N (n^m) = (Z.of_N n)^(Z.of_N m).
Proof.
- symmetry. destruct n, m; trivial. now apply Z.pow_0_l. apply Z.pow_Zpos.
+ destruct n, m; trivial. now rewrite Z.pow_0_l. apply Pos2Z.inj_pow.
Qed.
+Lemma inj_testbit a n :
+ Z.testbit (Z.of_N a) (Z.of_N n) = N.testbit a n.
+Proof. apply Z.Private_BootStrap.testbit_of_N. Qed.
+
End N2Z.
Module Z2N.
@@ -408,6 +424,10 @@ Proof.
- now destruct 2.
Qed.
+Lemma inj_testbit a n : 0<=n ->
+ Z.testbit (Z.of_N a) n = N.testbit a (Z.to_N n).
+Proof. apply Z.Private_BootStrap.testbit_of_N'. Qed.
+
End Z2N.
Module Zabs2N.
@@ -526,9 +546,9 @@ Proof.
intros. rewrite abs_N_nonneg. now apply Z2N.inj_quot. now apply Z.quot_pos.
destruct n, m; trivial; simpl.
- trivial.
- - now rewrite <- Z.opp_Zpos, Z.quot_opp_r, inj_opp.
- - now rewrite <- Z.opp_Zpos, Z.quot_opp_l, inj_opp.
- - now rewrite <- 2 Z.opp_Zpos, Z.quot_opp_opp.
+ - now rewrite <- Pos2Z.opp_pos, Z.quot_opp_r, inj_opp.
+ - now rewrite <- Pos2Z.opp_pos, Z.quot_opp_l, inj_opp.
+ - now rewrite <- 2 Pos2Z.opp_pos, Z.quot_opp_opp.
Qed.
Lemma inj_rem n m : Z.abs_N (Z.rem n m) = ((Z.abs_N n) mod (Z.abs_N m))%N.
@@ -538,9 +558,9 @@ Proof.
intros. rewrite abs_N_nonneg. now apply Z2N.inj_rem. now apply Z.rem_nonneg.
destruct n, m; trivial; simpl.
- trivial.
- - now rewrite <- Z.opp_Zpos, Z.rem_opp_r.
- - now rewrite <- Z.opp_Zpos, Z.rem_opp_l, inj_opp.
- - now rewrite <- 2 Z.opp_Zpos, Z.rem_opp_opp, inj_opp.
+ - now rewrite <- Pos2Z.opp_pos, Z.rem_opp_r.
+ - now rewrite <- Pos2Z.opp_pos, Z.rem_opp_l, inj_opp.
+ - now rewrite <- 2 Pos2Z.opp_pos, Z.rem_opp_opp, inj_opp.
Qed.
Lemma inj_pow n m : 0<=m -> Z.abs_N (n^m) = ((Z.abs_N n)^(Z.abs_N m))%N.
@@ -584,7 +604,7 @@ Qed.
Lemma inj_succ n : Z.of_nat (S n) = Z.succ (Z.of_nat n).
Proof.
- destruct n. trivial. simpl. symmetry. apply Z.succ_Zpos.
+ destruct n. trivial. simpl. apply Pos2Z.inj_succ.
Qed.
(** [Z.of_N] produce non-negative integers *)
@@ -915,10 +935,10 @@ End Zabs2Nat.
Definition neq (x y:nat) := x <> y.
-Lemma inj_neq n m : neq n m -> Zne (Z_of_nat n) (Z_of_nat m).
+Lemma inj_neq n m : neq n m -> Zne (Z.of_nat n) (Z.of_nat m).
Proof. intros H H'. now apply H, Nat2Z.inj. Qed.
-Lemma Zpos_P_of_succ_nat n : Zpos (P_of_succ_nat n) = Zsucc (Z_of_nat n).
+Lemma Zpos_P_of_succ_nat n : Zpos (Pos.of_succ_nat n) = Z.succ (Z.of_nat n).
Proof (Nat2Z.inj_succ n).
(** For these one, used in omega, a Definition is necessary *)
@@ -953,7 +973,7 @@ Notation inj_max := Nat2Z.inj_max (compat "8.3").
Notation Z_of_nat_of_P := positive_nat_Z (compat "8.3").
Notation Zpos_eq_Z_of_nat_o_nat_of_P :=
- (fun p => sym_eq (positive_nat_Z p)) (compat "8.3").
+ (fun p => eq_sym (positive_nat_Z p)) (compat "8.3").
Notation Z_of_nat_of_N := N_nat_Z (compat "8.3").
Notation Z_of_N_of_nat := nat_N_Z (compat "8.3").
@@ -991,7 +1011,7 @@ Notation Zabs_N_plus := Zabs2N.inj_add (compat "8.3").
Notation Zabs_N_mult_abs := Zabs2N.inj_mul_abs (compat "8.3").
Notation Zabs_N_mult := Zabs2N.inj_mul (compat "8.3").
-Theorem inj_minus2 : forall n m:nat, (m > n)%nat -> Z_of_nat (n - m) = 0.
+Theorem inj_minus2 : forall n m:nat, (m > n)%nat -> Z.of_nat (n - m) = 0.
Proof.
intros. rewrite not_le_minus_0; auto with arith.
Qed.
diff --git a/theories/ZArith/Znumtheory.v b/theories/ZArith/Znumtheory.v
index 00019b1a3..33f4dc7f4 100644
--- a/theories/ZArith/Znumtheory.v
+++ b/theories/ZArith/Znumtheory.v
@@ -17,7 +17,7 @@ Require Import Wf_nat.
Open Scope Z_scope.
(** This file contains some notions of number theory upon Z numbers:
- - a divisibility predicate [Zdivide]
+ - a divisibility predicate [Z.divide]
- a gcd predicate [gcd]
- Euclid algorithm [euclid]
- a relatively prime predicate [rel_prime]
@@ -35,7 +35,7 @@ Notation Zgcd_greatest := Z.gcd_greatest (compat "8.3").
Notation Zgcd_nonneg := Z.gcd_nonneg (compat "8.3").
Notation Zggcd_opp := Z.ggcd_opp (compat "8.3").
-(** The former specialized inductive predicate [Zdivide] is now
+(** The former specialized inductive predicate [Z.divide] is now
a generic existential predicate. *)
Notation Zdivide := Z.divide (compat "8.3").
@@ -76,11 +76,11 @@ Proof. apply Z.divide_abs_l. Qed.
Theorem Zdivide_Zabs_inv_l a b : (a | b) -> (Z.abs a | b).
Proof. apply Z.divide_abs_l. Qed.
-Hint Resolve Zdivide_refl Zone_divide Zdivide_0: zarith.
-Hint Resolve Zmult_divide_compat_l Zmult_divide_compat_r: zarith.
-Hint Resolve Zdivide_plus_r Zdivide_opp_r Zdivide_opp_r_rev Zdivide_opp_l
- Zdivide_opp_l_rev Zdivide_minus_l Zdivide_mult_l Zdivide_mult_r
- Zdivide_factor_r Zdivide_factor_l: zarith.
+Hint Resolve Z.divide_refl Z.divide_1_l Z.divide_0_r: zarith.
+Hint Resolve Z.mul_divide_mono_l Z.mul_divide_mono_r: zarith.
+Hint Resolve Z.divide_add_r Zdivide_opp_r Zdivide_opp_r_rev Zdivide_opp_l
+ Zdivide_opp_l_rev Z.divide_sub_r Z.divide_mul_l Z.divide_mul_r
+ Z.divide_factor_l Z.divide_factor_r: zarith.
(** Auxiliary result. *)
@@ -108,7 +108,7 @@ Proof.
now apply Z.divide_pos_le.
Qed.
-(** [Zdivide] can be expressed using [Zmod]. *)
+(** [Z.divide] can be expressed using [Z.modulo]. *)
Lemma Zmod_divide : forall a b, b<>0 -> a mod b = 0 -> (b | a).
Proof.
@@ -120,7 +120,7 @@ Proof.
intros a b (c,->); apply Z_mod_mult.
Qed.
-(** [Zdivide] is hence decidable *)
+(** [Z.divide] is hence decidable *)
Lemma Zdivide_dec a b : {(a | b)} + {~ (a | b)}.
Proof.
@@ -193,14 +193,16 @@ Qed.
(** * Greatest common divisor (gcd). *)
-(** There is no unicity of the gcd; hence we define the predicate [gcd a b d]
- expressing that [d] is a gcd of [a] and [b].
- (We show later that the [gcd] is actually unique if we discard its sign.) *)
+(** There is no unicity of the gcd; hence we define the predicate
+ [Zis_gcd a b g] expressing that [g] is a gcd of [a] and [b].
+ (We show later that the [gcd] is actually unique if we discard its sign.) *)
-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.
+Inductive Zis_gcd (a b g:Z) : Prop :=
+ Zis_gcd_intro :
+ (g | a) ->
+ (g | b) ->
+ (forall x, (x | a) -> (x | b) -> (x | g)) ->
+ Zis_gcd a b g.
(** Trivial properties of [gcd] *)
@@ -246,12 +248,10 @@ Hint Resolve Zis_gcd_sym Zis_gcd_0 Zis_gcd_minus Zis_gcd_opp: zarith.
Theorem Zis_gcd_unique: forall a b c d : Z,
Zis_gcd a b c -> Zis_gcd a b d -> c = d \/ c = (- d).
Proof.
-intros a b c d H1 H2.
-inversion_clear H1 as [Hc1 Hc2 Hc3].
-inversion_clear H2 as [Hd1 Hd2 Hd3].
-assert (H3: Zdivide c d); auto.
-assert (H4: Zdivide d c); auto.
-apply Zdivide_antisym; auto.
+intros a b c d [Hc1 Hc2 Hc3] [Hd1 Hd2 Hd3].
+assert (c|d) by auto.
+assert (d|c) by auto.
+apply Z.divide_antisym; auto.
Qed.
@@ -357,7 +357,7 @@ Proof.
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).
+ exact (Z.divide_antisym d d' Hdd' Hd'd).
Qed.
(** * Bezout's coefficients *)
@@ -450,18 +450,18 @@ Lemma rel_prime_cross_prod :
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).
+ elim (Z.divide_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.
+ rewrite Z.mul_comm in H3.
+ apply Z.mul_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 Z.mul_comm.
rewrite <- H3.
auto with zarith.
red in |- *; auto with zarith.
@@ -492,12 +492,12 @@ Proof.
exists b'; auto with zarith.
intros x (xa,H5) (xb,H6).
destruct (H4 (x*g)) as (x',Hx').
- exists xa; rewrite Zmult_assoc; rewrite <- H5; auto.
- exists xb; rewrite Zmult_assoc; rewrite <- H6; auto.
+ exists xa; rewrite Z.mul_assoc; rewrite <- H5; auto.
+ exists xb; rewrite Z.mul_assoc; rewrite <- H6; auto.
replace g with (1*g) in Hx'; auto with zarith.
- do 2 rewrite Zmult_assoc in Hx'.
- apply Zmult_reg_r in Hx'; trivial.
- rewrite Zmult_1_r in Hx'.
+ do 2 rewrite Z.mul_assoc in Hx'.
+ apply Z.mul_reg_r in Hx'; trivial.
+ rewrite Z.mul_1_r in Hx'.
exists x'; auto with zarith.
Qed.
@@ -512,9 +512,9 @@ Theorem rel_prime_div: forall p q r,
Proof.
intros p q r H (u, H1); subst.
inversion_clear H as [H1 H2 H3].
- red; apply Zis_gcd_intro; try apply Zone_divide.
+ red; apply Zis_gcd_intro; try apply Z.divide_1_l.
intros x H4 H5; apply H3; auto.
- apply Zdivide_mult_r; auto.
+ apply Z.divide_mul_r; auto.
Qed.
Theorem rel_prime_1: forall n, rel_prime 1 n.
@@ -578,10 +578,10 @@ Proof.
destruct 1; intros.
assert
(a = - p \/ - p < a < -1 \/ a = -1 \/ a = 0 \/ a = 1 \/ 1 < a < p \/ a = p).
- { assert (Zabs a <= Zabs p) as H2.
+ { assert (Z.abs a <= Z.abs p) as H2.
apply Zdivide_bounds; [ assumption | omega ].
revert H2.
- pattern (Zabs a); apply Zabs_ind; pattern (Zabs p); apply Zabs_ind;
+ pattern (Z.abs a); apply Zabs_ind; pattern (Z.abs p); apply Zabs_ind;
intros; omega. }
intuition idtac.
(* -p < a < -1 *)
@@ -589,7 +589,7 @@ Proof.
inversion H2.
assert (- a | - a) by auto with zarith.
assert (- a | p) by auto with zarith.
- apply H7, Zdivide_1 in H8; intuition.
+ apply H7, Z.divide_1_r in H8; intuition.
(* a = 0 *)
- inversion H1. subst a; omega.
(* 1 < a < p *)
@@ -597,7 +597,7 @@ Proof.
inversion H2.
assert (a | a) by auto with zarith.
assert (a | p) by auto with zarith.
- apply H7, Zdivide_1 in H8; intuition.
+ apply H7, Z.divide_1_r in H8; intuition.
Qed.
(** A prime number is relatively prime with any number it does not divide *)
@@ -621,7 +621,7 @@ Proof.
intros a p Hp [H1 H2].
apply rel_prime_sym; apply prime_rel_prime; auto.
intros [q Hq]; subst a.
- case (Zle_or_lt q 0); intros Hl.
+ case (Z.le_gt_cases q 0); intros Hl.
absurd (q * p <= 0 * p); auto with zarith.
absurd (1 * p <= q * p); auto with zarith.
Qed.
@@ -651,87 +651,79 @@ Qed.
Lemma prime_2: prime 2.
Proof.
apply prime_intro; auto with zarith.
- intros n [H1 H2]; case Zle_lt_or_eq with ( 1 := H1 ); auto with zarith;
- clear H1; intros H1.
- contradict H2; auto with zarith.
- subst n; red; auto with zarith.
- apply Zis_gcd_intro; auto with zarith.
+ intros n (H,H'); Z.le_elim H; auto with zarith.
+ - contradict H'; auto with zarith.
+ - subst n. constructor; auto with zarith.
Qed.
Theorem prime_3: prime 3.
Proof.
apply prime_intro; auto with zarith.
- intros n [H1 H2]; case Zle_lt_or_eq with ( 1 := H1 ); auto with zarith;
- clear H1; intros H1.
- case (Zle_lt_or_eq 2 n); auto with zarith; clear H1; intros H1.
- contradict H2; auto with zarith.
- subst n; red; auto with zarith.
- apply Zis_gcd_intro; auto with zarith.
- intros x [q1 Hq1] [q2 Hq2].
- exists (q2 - q1).
- apply trans_equal with (3 - 2); auto with zarith.
- rewrite Hq1; rewrite Hq2; ring.
- subst n; red; auto with zarith.
- apply Zis_gcd_intro; auto with zarith.
+ intros n (H,H'); Z.le_elim H; auto with zarith.
+ - replace n with 2 by omega.
+ constructor; auto with zarith.
+ intros x (q,Hq) (q',Hq').
+ exists (q' - q). ring_simplify. now rewrite <- Hq, <- Hq'.
+ - replace n with 1 by trivial.
+ constructor; auto with zarith.
Qed.
-Theorem prime_ge_2: forall p, prime p -> 2 <= p.
+Theorem prime_ge_2 p : prime p -> 2 <= p.
Proof.
- intros p Hp; inversion Hp; auto with zarith.
+ intros (Hp,_); auto with zarith.
Qed.
Definition prime' p := 1<p /\ (forall n, 1<n<p -> ~ (n|p)).
-Theorem prime_alt:
- forall p, prime' p <-> prime p.
-Proof.
- split; destruct 1; intros.
- (* prime -> prime' *)
- constructor; auto; intros.
- red; apply Zis_gcd_intro; auto with zarith; intros.
- case (Zle_lt_or_eq 0 (Zabs x)); auto with zarith; intros H6.
- case (Zle_lt_or_eq 1 (Zabs x)); auto with zarith; intros H7.
- case (Zle_lt_or_eq (Zabs x) p); auto with zarith.
- apply Zdivide_le; auto with zarith.
- apply Zdivide_Zabs_inv_l; auto.
- intros H8; case (H0 (Zabs x)); auto.
- apply Zdivide_Zabs_inv_l; auto.
- intros H8; subst p; absurd (Zabs x <= n); auto with zarith.
- apply Zdivide_le; auto with zarith.
- apply Zdivide_Zabs_inv_l; auto.
- rewrite H7; pattern (Zabs x); apply Zabs_intro; auto with zarith.
- absurd (0%Z = p); auto with zarith.
- assert (x=0) by (destruct x; simpl in *; now auto).
- subst x; elim H3; intro q; rewrite Zmult_0_r; auto.
- (* prime' -> prime *)
- split; auto; intros.
- intros H2.
- case (Zis_gcd_unique n p n 1); auto with zarith.
- apply Zis_gcd_intro; auto with zarith.
- apply H0; auto with zarith.
+Lemma Z_0_1_more x : 0<=x -> x=0 \/ x=1 \/ 1<x.
+Proof.
+ intros H. Z.le_elim H; auto.
+ apply Z.le_succ_l in H. change (1 <= x) in H. Z.le_elim H; auto.
+Qed.
+
+Theorem prime_alt p : prime' p <-> prime p.
+Proof.
+ split; intros (Hp,H).
+ - (* prime -> prime' *)
+ constructor; trivial; intros n Hn.
+ constructor; auto with zarith; intros x Hxn Hxp.
+ rewrite <- Z.divide_abs_l in Hxn, Hxp |- *.
+ assert (Hx := Z.abs_nonneg x).
+ set (y:=Z.abs x) in *; clearbody y; clear x; rename y into x.
+ destruct (Z_0_1_more x Hx) as [->|[->|Hx']].
+ + exfalso. apply Z.divide_0_l in Hxn. omega.
+ + now exists 1.
+ + elim (H x); auto.
+ split; trivial.
+ apply Z.le_lt_trans with n; auto with zarith.
+ apply Z.divide_pos_le; auto with zarith.
+ - (* prime' -> prime *)
+ constructor; trivial. intros n Hn Hnp.
+ case (Zis_gcd_unique n p n 1); auto with zarith.
+ constructor; auto with zarith.
+ apply H; auto with zarith.
Qed.
Theorem square_not_prime: forall a, ~ prime (a * a).
Proof.
intros a Ha.
- rewrite <- (Zabs_square a) in Ha.
- assert (0 <= Zabs a) by auto with zarith.
- set (b:=Zabs a) in *; clearbody b.
- rewrite <- prime_alt in Ha; destruct Ha.
- case (Zle_lt_or_eq 0 b); auto with zarith; intros Hza1; [ | subst; omega].
- case (Zle_lt_or_eq 1 b); auto with zarith; intros Hza2; [ | subst; omega].
- assert (Hza3 := Zmult_lt_compat_r 1 b b Hza1 Hza2).
- rewrite Zmult_1_l in Hza3.
- elim (H1 _ (conj Hza2 Hza3)).
- exists b; auto.
+ rewrite <- (Z.abs_square a) in Ha.
+ assert (H:=Z.abs_nonneg a).
+ set (b:=Z.abs a) in *; clearbody b; clear a; rename b into a.
+ rewrite <- prime_alt in Ha; destruct Ha as (Ha,Ha').
+ assert (H' : 1 < a) by now apply (Z.square_lt_simpl_nonneg 1).
+ apply (Ha' a).
+ + split; trivial.
+ rewrite <- (Z.mul_1_l a) at 1. apply Z.mul_lt_mono_pos_r; omega.
+ + exists a; auto.
Qed.
Theorem prime_div_prime: forall p q,
prime p -> prime q -> (p | q) -> p = q.
Proof.
intros p q H H1 H2;
- assert (Hp: 0 < p); try apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
- assert (Hq: 0 < q); try apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+ assert (Hp: 0 < p); try apply Z.lt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+ assert (Hq: 0 < q); try apply Z.lt_le_trans with 2; try apply prime_ge_2; auto with zarith.
case prime_divisors with (2 := H2); auto.
intros H4; contradict Hp; subst; auto with zarith.
intros [H4| [H4 | H4]]; subst; auto.
@@ -768,11 +760,11 @@ Theorem Zis_gcd_gcd: forall a b c : Z,
0 <= c -> Zis_gcd a b c -> Z.gcd a b = c.
Proof.
intros a b c H1 H2.
- case (Zis_gcd_uniqueness_apart_sign a b c (Zgcd a b)); auto.
+ case (Zis_gcd_uniqueness_apart_sign a b c (Z.gcd a b)); auto.
apply Zgcd_is_gcd; auto.
Z.le_elim H1.
- generalize (Z.gcd_nonneg a b); auto with zarith.
- subst. now case (Z.gcd a b).
+ - generalize (Z.gcd_nonneg a b); auto with zarith.
+ - subst. now case (Z.gcd a b).
Qed.
Notation Zgcd_inv_0_l := Z.gcd_eq_0_l (compat "8.3").
@@ -786,8 +778,8 @@ Proof.
intros a b Hg Hb.
assert (F := Zgcd_is_gcd a b); inversion F as [F1 F2 F3].
pattern b at 2; rewrite (Zdivide_Zdiv_eq (Z.gcd a b) b); auto.
- repeat rewrite Zmult_assoc; f_equal.
- rewrite Zmult_comm.
+ repeat rewrite Z.mul_assoc; f_equal.
+ rewrite Z.mul_comm.
rewrite <- Zdivide_Zdiv_eq; auto.
Qed.
@@ -798,17 +790,17 @@ Theorem Zgcd_div_swap : forall a b c : Z,
Proof.
intros a b c Hg Hb.
assert (F := Zgcd_is_gcd a b); inversion F as [F1 F2 F3].
- pattern b at 2; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto.
- repeat rewrite Zmult_assoc; f_equal.
+ pattern b at 2; rewrite (Zdivide_Zdiv_eq (Z.gcd a b) b); auto.
+ repeat rewrite Z.mul_assoc; f_equal.
rewrite Zdivide_Zdiv_eq_2; auto.
- repeat rewrite <- Zmult_assoc; f_equal.
- rewrite Zmult_comm.
+ repeat rewrite <- Z.mul_assoc; f_equal.
+ rewrite Z.mul_comm.
rewrite <- Zdivide_Zdiv_eq; auto.
Qed.
Notation Zgcd_comm := Z.gcd_comm (compat "8.3").
-Lemma Zgcd_ass a b c : Zgcd (Zgcd a b) c = Zgcd a (Zgcd b c).
+Lemma Zgcd_ass a b c : Z.gcd (Z.gcd a b) c = Z.gcd a (Z.gcd b c).
Proof.
symmetry. apply Z.gcd_assoc.
Qed.
@@ -817,23 +809,23 @@ Notation Zgcd_Zabs := Z.gcd_abs_l (compat "8.3").
Notation Zgcd_0 := Z.gcd_0_r (compat "8.3").
Notation Zgcd_1 := Z.gcd_1_r (compat "8.3").
-Hint Resolve Zgcd_0 Zgcd_1 : zarith.
+Hint Resolve Z.gcd_0_r Z.gcd_1_r : zarith.
Theorem Zgcd_1_rel_prime : forall a b,
Z.gcd a b = 1 <-> rel_prime a b.
Proof.
unfold rel_prime; split; intro H.
rewrite <- H; apply Zgcd_is_gcd.
- case (Zis_gcd_unique a b (Zgcd a b) 1); auto.
+ case (Zis_gcd_unique a b (Z.gcd a b) 1); auto.
apply Zgcd_is_gcd.
- intros H2; absurd (0 <= Zgcd a b); auto with zarith.
- generalize (Zgcd_is_pos a b); auto with zarith.
+ intros H2; absurd (0 <= Z.gcd a b); auto with zarith.
+ generalize (Z.gcd_nonneg a b); auto with zarith.
Qed.
Definition rel_prime_dec: forall a b,
{ rel_prime a b }+{ ~ rel_prime a b }.
Proof.
- intros a b; case (Z_eq_dec (Zgcd a b) 1); intros H1.
+ intros a b; case (Z.eq_dec (Z.gcd a b) 1); intros H1.
left; apply -> Zgcd_1_rel_prime; auto.
right; contradict H1; apply <- Zgcd_1_rel_prime; auto.
Defined.
@@ -851,25 +843,24 @@ Proof.
intros x Hx IH; destruct IH as [F|E].
destruct (rel_prime_dec x p) as [Y|N].
left; intros n [HH1 HH2].
- case (Zgt_succ_gt_or_eq x n); auto with zarith.
- intros HH3; subst x; auto.
- case (Z_lt_dec 1 x); intros HH1.
- right; exists x; split; auto with zarith.
- left; intros n [HHH1 HHH2]; contradict HHH1; auto with zarith.
- right; destruct E as (n,((H0,H2),H3)); exists n; auto with zarith.
+ rewrite Z.lt_succ_r in HH2.
+ Z.le_elim HH2; subst; auto with zarith.
+ - case (Z_lt_dec 1 x); intros HH1.
+ * right; exists x; split; auto with zarith.
+ * left; intros n [HHH1 HHH2]; contradict HHH1; auto with zarith.
+ - right; destruct E as (n,((H0,H2),H3)); exists n; auto with zarith.
Defined.
Definition prime_dec: forall p, { prime p }+{ ~ prime p }.
Proof.
intros p; case (Z_lt_dec 1 p); intros H1.
- case (prime_dec_aux p p); intros H2.
- left; apply prime_intro; auto.
- intros n [Hn1 Hn2]; case Zle_lt_or_eq with ( 1 := Hn1 ); auto.
- intros HH; subst n.
- red; apply Zis_gcd_intro; auto with zarith.
- right; intros H3; inversion_clear H3 as [Hp1 Hp2].
- case H2; intros n [Hn1 Hn2]; case Hn2; auto with zarith.
- right; intros H3; inversion_clear H3 as [Hp1 Hp2]; case H1; auto.
+ + case (prime_dec_aux p p); intros H2.
+ * left; apply prime_intro; auto.
+ intros n (Hn1,Hn2). Z.le_elim Hn1; auto; subst n.
+ constructor; auto with zarith.
+ * right; intros H3; inversion_clear H3 as [Hp1 Hp2].
+ case H2; intros n [Hn1 Hn2]; case Hn2; auto with zarith.
+ + right; intros H3; inversion_clear H3 as [Hp1 Hp2]; case H1; auto.
Defined.
Theorem not_prime_divide:
@@ -877,29 +868,16 @@ Theorem not_prime_divide:
Proof.
intros p Hp Hp1.
case (prime_dec_aux p p); intros H1.
- elim Hp1; constructor; auto.
- intros n [Hn1 Hn2].
- case Zle_lt_or_eq with ( 1 := Hn1 ); auto with zarith.
- intros H2; subst n; red; apply Zis_gcd_intro; auto with zarith.
- case H1; intros n [Hn1 Hn2].
- generalize (Zgcd_is_pos n p); intros Hpos.
- case (Zle_lt_or_eq 0 (Zgcd n p)); auto with zarith; intros H3.
- case (Zle_lt_or_eq 1 (Zgcd n p)); auto with zarith; intros H4.
- exists (Zgcd n p); split; auto.
- split; auto.
- apply Zle_lt_trans with n; auto with zarith.
- generalize (Zgcd_is_gcd n p); intros tmp; inversion_clear tmp as [Hr1 Hr2 Hr3].
- case Hr1; intros q Hq.
- case (Zle_or_lt q 0); auto with zarith; intros Ht.
- absurd (n <= 0 * Zgcd n p) ; auto with zarith.
- pattern n at 1; rewrite Hq; auto with zarith.
- apply Zle_trans with (1 * Zgcd n p); auto with zarith.
- pattern n at 2; rewrite Hq; auto with zarith.
- generalize (Zgcd_is_gcd n p); intros Ht; inversion Ht; auto.
- case Hn2; red.
- rewrite H4; apply Zgcd_is_gcd.
- generalize (Zgcd_is_gcd n p); rewrite <- H3; intros tmp;
- inversion_clear tmp as [Hr1 Hr2 Hr3].
- absurd (n = 0); auto with zarith.
- case Hr1; auto with zarith.
+ - elim Hp1; constructor; auto.
+ intros n (Hn1,Hn2).
+ Z.le_elim Hn1; auto with zarith.
+ subst n; constructor; auto with zarith.
+ - case H1; intros n (Hn1,Hn2).
+ destruct (Z_0_1_more _ (Z.gcd_nonneg n p)) as [H|[H|H]].
+ + exfalso. apply Z.gcd_eq_0_l in H. omega.
+ + elim Hn2. red. rewrite <- H. apply Zgcd_is_gcd.
+ + exists (Z.gcd n p); split; [ split; auto | apply Z.gcd_divide_r ].
+ apply Z.le_lt_trans with n; auto with zarith.
+ apply Z.divide_pos_le; auto with zarith.
+ apply Z.gcd_divide_l.
Qed.
diff --git a/theories/ZArith/Zorder.v b/theories/ZArith/Zorder.v
index d22e2d57c..0d5b9feb7 100644
--- a/theories/ZArith/Zorder.v
+++ b/theories/ZArith/Zorder.v
@@ -167,8 +167,10 @@ Notation Zle_or_lt := Z.le_gt_cases (compat "8.3").
Notation Zlt_trans := Z.lt_trans (compat "8.3").
-Lemma Zgt_trans : forall n m p:Z, n > m -> m > p -> n > p.
-Proof Zcompare_Gt_trans.
+Lemma Zgt_trans n m p : n > m -> m > p -> n > p.
+Proof.
+ Z.swap_greater. intros; now transitivity m.
+Qed.
(** Mixed transitivity *)
@@ -304,7 +306,7 @@ Proof.
Qed.
Hint Resolve Z.le_succ_diag_r: zarith.
-Hint Resolve Zle_le_succ: zarith.
+Hint Resolve Z.le_le_succ_r: zarith.
(** Relating order wrt successor and order wrt predecessor *)
@@ -345,7 +347,7 @@ Proof.
easy.
Qed.
-(* weaker but useful (in [Zpower] for instance) *)
+(* weaker but useful (in [Z.pow] for instance) *)
Lemma Zle_0_pos : forall p:positive, 0 <= Zpos p.
Proof.
easy.
@@ -361,7 +363,7 @@ Proof.
induction n; simpl; intros. apply Z.le_refl. easy.
Qed.
-Hint Immediate Zeq_le: zarith.
+Hint Immediate Z.eq_le_incl: zarith.
(** Derived lemma *)
diff --git a/theories/ZArith/Zpow_alt.v b/theories/ZArith/Zpow_alt.v
index a90eedb41..880245c53 100644
--- a/theories/ZArith/Zpow_alt.v
+++ b/theories/ZArith/Zpow_alt.v
@@ -79,5 +79,5 @@ Qed.
Lemma Zpower_alt_Ppow p q : (Zpos p)^^(Zpos q) = Zpos (p^q).
Proof.
- now rewrite Zpower_equiv, Z.pow_Zpos.
+ now rewrite Zpower_equiv, Pos2Z.inj_pow.
Qed.
diff --git a/theories/ZArith/Zpow_def.v b/theories/ZArith/Zpow_def.v
index 4fe411584..2e23e982e 100644
--- a/theories/ZArith/Zpow_def.v
+++ b/theories/ZArith/Zpow_def.v
@@ -19,7 +19,7 @@ Notation Zpower := Z.pow (compat "8.3").
Notation Zpower_0_r := Z.pow_0_r (compat "8.3").
Notation Zpower_succ_r := Z.pow_succ_r (compat "8.3").
Notation Zpower_neg_r := Z.pow_neg_r (compat "8.3").
-Notation Zpower_Ppow := Z.pow_Zpos (compat "8.3").
+Notation Zpower_Ppow := Pos2Z.inj_pow (compat "8.3").
Lemma Zpower_theory : power_theory 1 Z.mul (@eq Z) Z.of_N Z.pow.
Proof.
diff --git a/theories/ZArith/Zpow_facts.v b/theories/ZArith/Zpow_facts.v
index 5d025322b..fa63a190e 100644
--- a/theories/ZArith/Zpow_facts.v
+++ b/theories/ZArith/Zpow_facts.v
@@ -85,15 +85,15 @@ Proof.
assert (Hn := Nat2Z.is_nonneg n).
destruct p; simpl Pos.size_nat.
- specialize IHn with p.
- rewrite Z.pos_xI, Nat2Z.inj_succ, Z.pow_succ_r; omega.
+ rewrite Pos2Z.inj_xI, Nat2Z.inj_succ, Z.pow_succ_r; omega.
- specialize IHn with p.
- rewrite Z.pos_xO, Nat2Z.inj_succ, Z.pow_succ_r; omega.
+ rewrite Pos2Z.inj_xO, Nat2Z.inj_succ, Z.pow_succ_r; omega.
- split; auto with zarith.
intros _. apply Z.pow_gt_1. easy.
now rewrite Nat2Z.inj_succ, Z.lt_succ_r.
Qed.
-(** * Zpower and modulo *)
+(** * Z.pow and modulo *)
Theorem Zpower_mod p q n :
0 < n -> (p^q) mod n = ((p mod n)^q) mod n.
@@ -106,7 +106,7 @@ Proof.
- rewrite !Z.pow_neg_r; auto with zarith.
Qed.
-(** A direct way to compute Zpower modulo **)
+(** A direct way to compute Z.pow modulo **)
Fixpoint Zpow_mod_pos (a: Z)(m: positive)(n : Z) : Z :=
match m with
@@ -231,7 +231,7 @@ Proof.
exists n; destruct H; rewrite Z.mul_0_r in H; auto.
Qed.
-(** * Zsquare: a direct definition of [z^2] *)
+(** * Z.square: a direct definition of [z^2] *)
Notation Psquare := Pos.square (compat "8.3").
Notation Zsquare := Z.square (compat "8.3").
diff --git a/theories/ZArith/Zpower.v b/theories/ZArith/Zpower.v
index 5052d01a0..880c4376d 100644
--- a/theories/ZArith/Zpower.v
+++ b/theories/ZArith/Zpower.v
@@ -78,7 +78,7 @@ Proof.
Qed.
Hint Immediate Zpower_nat_is_exp Zpower_pos_is_exp : zarith.
-Hint Unfold Zpower_pos Zpower_nat: zarith.
+Hint Unfold Z.pow_pos Zpower_nat: zarith.
Theorem Zpower_exp x n m :
n >= 0 -> m >= 0 -> x ^ (n + m) = x ^ n * x ^ m.
@@ -181,7 +181,7 @@ Section Powers_of_2.
Qed.
Theorem shift_pos_correct p x :
- Zpos (shift_pos p x) = Zpower_pos 2 p * Zpos x.
+ Zpos (shift_pos p x) = Z.pow_pos 2 p * Zpos x.
Proof.
now rewrite shift_pos_nat, Zpower_pos_nat, shift_nat_correct.
Qed.
@@ -266,13 +266,13 @@ Section power_div_with_rest.
apply Pos.iter_invariant; [|omega].
intros ((q,r),d) (H,H'). unfold Zdiv_rest_aux.
destruct q as [ |[q|q| ]|[q|q| ]]; try omega.
- - rewrite Z.pos_xI, Z.mul_add_distr_r in H.
+ - rewrite Pos2Z.inj_xI, Z.mul_add_distr_r in H.
rewrite Z.mul_shuffle3, Z.mul_assoc. omega.
- - rewrite Z.pos_xO in H.
+ - rewrite Pos2Z.inj_xO in H.
rewrite Z.mul_shuffle3, Z.mul_assoc. omega.
- - rewrite Z.neg_xI, Z.mul_sub_distr_r in H.
+ - rewrite Pos2Z.neg_xI, Z.mul_sub_distr_r in H.
rewrite Z.mul_sub_distr_r, Z.mul_shuffle3, Z.mul_assoc. omega.
- - rewrite Z.neg_xO in H.
+ - rewrite Pos2Z.neg_xO in H.
rewrite Z.mul_shuffle3, Z.mul_assoc. omega.
Qed.
diff --git a/theories/ZArith/Zquot.v b/theories/ZArith/Zquot.v
index 5b79fbce8..accce71f1 100644
--- a/theories/ZArith/Zquot.v
+++ b/theories/ZArith/Zquot.v
@@ -11,51 +11,95 @@ Require Import Nnat ZArith_base ROmega ZArithRing Zdiv Morphisms.
Local Open Scope Z_scope.
(** This file provides results about the Round-Toward-Zero Euclidean
- division [Zquotrem], whose projections are [Zquot] and [Zrem].
- Definition of this division can be found in file [BinIntDef].
+ division [Z.quotrem], whose projections are [Z.quot] (noted ÷)
+ and [Z.rem].
- This division and the one defined in Zdiv agree only on positive
- numbers. Otherwise, Zdiv performs Round-Toward-Bottom (a.k.a Floor).
+ This division and [Z.div] agree only on positive numbers.
+ Otherwise, [Z.div] performs Round-Toward-Bottom (a.k.a Floor).
- The current approach is compatible with the division of usual
+ This [Z.quot] is compatible with the division of usual
programming languages such as Ocaml. In addition, it has nicer
properties with respect to opposite and other usual operations.
+
+ The definition of this division is now in file [BinIntDef],
+ while most of the results about here are now in the main module
+ [BinInt.Z], thanks to the generic "Numbers" layer. Remain here:
+
+ - some compatibility notation for old names.
+
+ - some extra results with less preconditions (in particular
+ exploiting the arbitrary value of division by 0).
*)
-(** * Relation between division on N and on Z. *)
+Notation Ndiv_Zquot := N2Z.inj_quot (compat "8.3").
+Notation Nmod_Zrem := N2Z.inj_rem (compat "8.3").
+Notation Z_quot_rem_eq := Z.quot_rem' (compat "8.3").
+Notation Zrem_lt := Z.rem_bound_abs (compat "8.3").
+Notation Zquot_unique := Z.quot_unique (compat "8.3").
+Notation Zrem_unique := Z.rem_unique (compat "8.3").
+Notation Zrem_1_r := Z.rem_1_r (compat "8.3").
+Notation Zquot_1_r := Z.quot_1_r (compat "8.3").
+Notation Zrem_1_l := Z.rem_1_l (compat "8.3").
+Notation Zquot_1_l := Z.quot_1_l (compat "8.3").
+Notation Z_quot_same := Z.quot_same (compat "8.3").
+Notation Z_quot_mult := Z.quot_mul (compat "8.3").
+Notation Zquot_small := Z.quot_small (compat "8.3").
+Notation Zrem_small := Z.rem_small (compat "8.3").
+Notation Zquot2_quot := Zquot2_quot (compat "8.3").
+
+(** Particular values taken for [a÷0] and [(Z.rem a 0)].
+ We avise to not rely on these arbitrary values. *)
+
+Lemma Zquot_0_r a : a ÷ 0 = 0.
+Proof. now destruct a. Qed.
+
+Lemma Zrem_0_r a : Z.rem a 0 = a.
+Proof. now destruct a. Qed.
+
+(** The following results are expressed without the [b<>0] condition
+ whenever possible. *)
+
+Lemma Zrem_0_l a : Z.rem 0 a = 0.
+Proof. now destruct a. Qed.
+
+Lemma Zquot_0_l a : 0÷a = 0.
+Proof. now destruct a. Qed.
+
+Hint Resolve Zrem_0_l Zrem_0_r Zquot_0_l Zquot_0_r Z.quot_1_r Z.rem_1_r
+ : zarith.
-Lemma Ndiv_Zquot : forall a b:N,
- Z_of_N (a/b) = (Z_of_N a ÷ Z_of_N b).
-Proof.
- intros.
- destruct a; destruct b; simpl; auto.
- unfold N.div, Z.quot; simpl. destruct N.pos_div_eucl; auto.
-Qed.
+Ltac zero_or_not a :=
+ destruct (Z.eq_decidable a 0) as [->|?];
+ [rewrite ?Zquot_0_l, ?Zrem_0_l, ?Zquot_0_r, ?Zrem_0_r;
+ auto with zarith|].
-Lemma Nmod_Zrem : forall a b:N,
- Z.of_N (a mod b) = Z.rem (Z.of_N a) (Z.of_N b).
-Proof.
- intros.
- destruct a; destruct b; simpl; auto.
- unfold N.modulo, Z.rem; simpl; destruct N.pos_div_eucl; auto.
-Qed.
+Lemma Z_rem_same a : Z.rem a a = 0.
+Proof. zero_or_not a. now apply Z.rem_same. Qed.
-(** * Characterization of this euclidean division. *)
+Lemma Z_rem_mult a b : Z.rem (a*b) b = 0.
+Proof. zero_or_not b. now apply Z.rem_mul. Qed.
-(** First, the usual equation [a=q*b+r]. Notice that [a mod 0]
- has been chosen to be [a], so this equation holds even for [b=0].
-*)
+(** * Division and Opposite *)
-Notation Z_quot_rem_eq := Z.quot_rem' (compat "8.3").
+(* The precise equalities that are invalid with "historic" Zdiv. *)
+
+Theorem Zquot_opp_l a b : (-a)÷b = -(a÷b).
+Proof. zero_or_not b. now apply Z.quot_opp_l. Qed.
-(** Then, the inequalities constraining the remainder:
- The remainder is bounded by the divisor, in term of absolute values *)
+Theorem Zquot_opp_r a b : a÷(-b) = -(a÷b).
+Proof. zero_or_not b. now apply Z.quot_opp_r. Qed.
-Theorem Zrem_lt : forall a b:Z, b<>0 ->
- Z.abs (Z.rem a b) < Z.abs b.
-Proof.
- apply Z.rem_bound_abs.
-Qed.
+Theorem Zrem_opp_l a b : Z.rem (-a) b = -(Z.rem a b).
+Proof. zero_or_not b. now apply Z.rem_opp_l. Qed.
+
+Theorem Zrem_opp_r a b : Z.rem a (-b) = Z.rem a b.
+Proof. zero_or_not b. now apply Z.rem_opp_r. Qed.
+
+Theorem Zquot_opp_opp a b : (-a)÷(-b) = a÷b.
+Proof. zero_or_not b. now apply Z.quot_opp_opp. Qed.
+
+Theorem Zrem_opp_opp a b : Z.rem (-a) (-b) = -(Z.rem a b).
+Proof. zero_or_not b. now apply Z.rem_opp_opp. Qed.
(** The sign of the remainder is the one of [a]. Due to the possible
nullity of [a], a general result is to be stated in the following form:
@@ -63,41 +107,33 @@ Qed.
Theorem Zrem_sgn a b : 0 <= Z.sgn (Z.rem a b) * Z.sgn a.
Proof.
- destruct b as [ |b|b]; destruct a as [ |a|a]; simpl; auto with zarith;
- unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl;
- simpl; destruct n0; simpl; auto with zarith.
+ zero_or_not b.
+ - apply Z.square_nonneg.
+ - zero_or_not (Z.rem a b).
+ rewrite Z.rem_sign_nz; trivial. apply Z.square_nonneg.
Qed.
(** This can also be said in a simplier way: *)
Theorem Zrem_sgn2 a b : 0 <= (Z.rem a b) * a.
Proof.
- rewrite <-Z.sgn_nonneg, Z.sgn_mul; apply Zrem_sgn.
+ zero_or_not b.
+ - apply Z.square_nonneg.
+ - now apply Z.rem_sign_mul.
Qed.
-(** Reformulation of [Zquot_lt] and [Zrem_sgn] in 2
- then 4 particular cases. *)
+(** Reformulation of [Z.rem_bound_abs] in 2 then 4 particular cases. *)
Theorem Zrem_lt_pos a b : 0<=a -> b<>0 -> 0 <= Z.rem a b < Z.abs b.
Proof.
- intros.
- assert (0 <= Z.rem a b).
- generalize (Zrem_sgn a b).
- destruct (Zle_lt_or_eq 0 a H).
- rewrite <- Zsgn_pos in H1; rewrite H1. romega with *.
- subst a; simpl; auto.
- generalize (Zrem_lt a b H0); romega with *.
+ intros; generalize (Z.rem_nonneg a b) (Z.rem_bound_abs a b);
+ romega with *.
Qed.
Theorem Zrem_lt_neg a b : a<=0 -> b<>0 -> -Z.abs b < Z.rem a b <= 0.
Proof.
- intros.
- assert (Z.rem a b <= 0).
- generalize (Zrem_sgn a b).
- destruct (Zle_lt_or_eq a 0 H).
- rewrite <- Zsgn_neg in H1; rewrite H1; romega with *.
- subst a; simpl; auto.
- generalize (Zrem_lt a b H0); romega with *.
+ intros; generalize (Z.rem_nonpos a b) (Z.rem_bound_abs a b);
+ romega with *.
Qed.
Theorem Zrem_lt_pos_pos a b : 0<=a -> 0<b -> 0 <= Z.rem a b < b.
@@ -120,45 +156,6 @@ Proof.
intros; generalize (Zrem_lt_neg a b); romega with *.
Qed.
-(** * Division and Opposite *)
-
-(* The precise equalities that are invalid with "historic" Zdiv. *)
-
-Theorem Zquot_opp_l a b : (-a)÷b = -(a÷b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold Z.quot, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem Zquot_opp_r a b : a÷(-b) = -(a÷b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold Z.quot, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem Zrem_opp_l a b : Z.rem (-a) b = -(Z.rem a b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem Zrem_opp_r a b : Z.rem a (-b) = Z.rem a b.
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem Zquot_opp_opp a b : (-a)÷(-b) = a÷b.
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold Z.quot, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
-Qed.
-
-Theorem Zrem_opp_opp a b : Z.rem (-a) (-b) = -(Z.rem a b).
-Proof.
- destruct a; destruct b; simpl; auto;
- unfold Z.rem, Z.quotrem; destruct N.pos_div_eucl; simpl; auto with zarith.
-Qed.
(** * Unicity results *)
@@ -172,170 +169,93 @@ Lemma Remainder_equiv : forall a b r,
Remainder a b r <-> Remainder_alt a b r.
Proof.
unfold Remainder, Remainder_alt; intuition.
- romega with *.
- romega with *.
- rewrite <-(Zmult_opp_opp).
- apply Zmult_le_0_compat; romega.
- assert (0 <= Z.sgn r * Z.sgn a) by (rewrite <-Z.sgn_mul, Z.sgn_nonneg; auto).
- destruct r; simpl Z.sgn in *; romega with *.
+ - romega with *.
+ - romega with *.
+ - rewrite <-(Z.mul_opp_opp). apply Z.mul_nonneg_nonneg; romega.
+ - assert (0 <= Z.sgn r * Z.sgn a).
+ { rewrite <-Z.sgn_mul, Z.sgn_nonneg; auto. }
+ destruct r; simpl Z.sgn in *; romega with *.
Qed.
-Theorem Zquot_mod_unique_full:
- forall a b q r, Remainder a b r ->
- a = b*q + r -> q = a÷b /\ r = Z.rem a b.
+Theorem Zquot_mod_unique_full a b q r :
+ Remainder a b r -> a = b*q + r -> q = a÷b /\ r = Z.rem a b.
Proof.
destruct 1 as [(H,H0)|(H,H0)]; intros.
apply Zdiv_mod_unique with b; auto.
apply Zrem_lt_pos; auto.
romega with *.
- rewrite <- H1; apply Z_quot_rem_eq.
+ rewrite <- H1; apply Z.quot_rem'.
- rewrite <- (Zopp_involutive a).
+ rewrite <- (Z.opp_involutive a).
rewrite Zquot_opp_l, Zrem_opp_l.
generalize (Zdiv_mod_unique b (-q) (-a÷b) (-r) (Z.rem (-a) b)).
generalize (Zrem_lt_pos (-a) b).
- rewrite <-Z_quot_rem_eq, <-Zopp_mult_distr_r, <-Zopp_plus_distr, <-H1.
+ rewrite <-Z.quot_rem', Z.mul_opp_r, <-Z.opp_add_distr, <-H1.
romega with *.
Qed.
-Theorem Zquot_unique_full:
- forall a b q r, Remainder a b r ->
- a = b*q + r -> q = a÷b.
+Theorem Zquot_unique_full a b q r :
+ Remainder a b r -> a = b*q + r -> q = a÷b.
Proof.
intros; destruct (Zquot_mod_unique_full a b q r); auto.
Qed.
-Theorem Zquot_unique:
- forall a b q r, 0 <= a -> 0 <= r < b ->
- a = b*q + r -> q = a÷b.
-Proof. exact Z.quot_unique. Qed.
-
-Theorem Zrem_unique_full:
- forall a b q r, Remainder a b r ->
- a = b*q + r -> r = Z.rem a b.
+Theorem Zrem_unique_full a b q r :
+ Remainder a b r -> a = b*q + r -> r = Z.rem a b.
Proof.
intros; destruct (Zquot_mod_unique_full a b q r); auto.
Qed.
-Theorem Zrem_unique:
- forall a b q r, 0 <= a -> 0 <= r < b ->
- a = b*q + r -> r = Z.rem a b.
-Proof. exact Z.rem_unique. Qed.
-
-(** * Basic values of divisions and modulo. *)
-
-Lemma Zrem_0_l: forall a, Z.rem 0 a = 0.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma Zrem_0_r: forall a, Z.rem a 0 = a.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma Zquot_0_l: forall a, 0÷a = 0.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma Zquot_0_r: forall a, a÷0 = 0.
-Proof.
- destruct a; simpl; auto.
-Qed.
-
-Lemma Zrem_1_r: forall a, Z.rem a 1 = 0.
-Proof. exact Z.rem_1_r. Qed.
-
-Lemma Zquot_1_r: forall a, a÷1 = a.
-Proof. exact Z.quot_1_r. Qed.
-
-Hint Resolve Zrem_0_l Zrem_0_r Zquot_0_l Zquot_0_r Zquot_1_r Zrem_1_r
- : zarith.
-
-Lemma Zquot_1_l: forall a, 1 < a -> 1÷a = 0.
-Proof. exact Z.quot_1_l. Qed.
-
-Lemma Zrem_1_l: forall a, 1 < a -> Z.rem 1 a = 1.
-Proof. exact Z.rem_1_l. Qed.
-
-Lemma Z_quot_same : forall a:Z, a<>0 -> a÷a = 1.
-Proof. exact Z.quot_same. Qed.
-
-Ltac zero_or_not a :=
- destruct (Z.eq_dec a 0);
- [subst; rewrite ?Zrem_0_l, ?Zquot_0_l, ?Zrem_0_r, ?Zquot_0_r;
- auto with zarith|].
-
-Lemma Z_rem_same : forall a, Z.rem a a = 0.
-Proof. intros. zero_or_not a. apply Z.rem_same; auto. Qed.
-
-Lemma Z_rem_mult : forall a b, Z.rem (a*b) b = 0.
-Proof. intros. zero_or_not b. apply Z.rem_mul; auto. Qed.
-
-Lemma Z_quot_mult : forall a b:Z, b <> 0 -> (a*b)÷b = a.
-Proof. exact Z.quot_mul. Qed.
-
(** * Order results about Zrem and Zquot *)
(* Division of positive numbers is positive. *)
-Lemma Z_quot_pos: forall a b, 0 <= a -> 0 <= b -> 0 <= a÷b.
+Lemma Z_quot_pos a b : 0 <= a -> 0 <= b -> 0 <= a÷b.
Proof. intros. zero_or_not b. apply Z.quot_pos; auto with zarith. Qed.
(** As soon as the divisor is greater or equal than 2,
the division is strictly decreasing. *)
-Lemma Z_quot_lt : forall a b:Z, 0 < a -> 2 <= b -> a÷b < a.
+Lemma Z_quot_lt a b : 0 < a -> 2 <= b -> a÷b < a.
Proof. intros. apply Z.quot_lt; auto with zarith. Qed.
-(** A division of a small number by a bigger one yields zero. *)
+(** [<=] is compatible with a positive division. *)
-Theorem Zquot_small: forall a b, 0 <= a < b -> a÷b = 0.
-Proof. exact Z.quot_small. Qed.
-
-(** Same situation, in term of modulo: *)
-
-Theorem Zrem_small: forall a n, 0 <= a < n -> Z.rem a n = a.
-Proof. exact Z.rem_small. Qed.
-
-(** [Zge] is compatible with a positive division. *)
-
-Lemma Z_quot_monotone : forall a b c, 0<=c -> a<=b -> a÷c <= b÷c.
+Lemma Z_quot_monotone a b c : 0<=c -> a<=b -> a÷c <= b÷c.
Proof. intros. zero_or_not c. apply Z.quot_le_mono; auto with zarith. Qed.
-(** With our choice of division, rounding of (a÷b) is always done toward zero: *)
+(** With our choice of division, rounding of (a÷b) is always done toward 0: *)
-Lemma Z_mult_quot_le : forall a b:Z, 0 <= a -> 0 <= b*(a÷b) <= a.
+Lemma Z_mult_quot_le a b : 0 <= a -> 0 <= b*(a÷b) <= a.
Proof. intros. zero_or_not b. apply Z.mul_quot_le; auto with zarith. Qed.
-Lemma Z_mult_quot_ge : forall a b:Z, a <= 0 -> a <= b*(a÷b) <= 0.
+Lemma Z_mult_quot_ge a b : a <= 0 -> a <= b*(a÷b) <= 0.
Proof. intros. zero_or_not b. apply Z.mul_quot_ge; auto with zarith. Qed.
(** The previous inequalities between [b*(a÷b)] and [a] are exact
iff the modulo is zero. *)
-Lemma Z_quot_exact_full : forall a b:Z, a = b*(a÷b) <-> Z.rem a b = 0.
+Lemma Z_quot_exact_full a b : a = b*(a÷b) <-> Z.rem a b = 0.
Proof. intros. zero_or_not b. intuition. apply Z.quot_exact; auto. Qed.
(** A modulo cannot grow beyond its starting point. *)
-Theorem Zrem_le: forall a b, 0 <= a -> 0 <= b -> Z.rem a b <= a.
+Theorem Zrem_le a b : 0 <= a -> 0 <= b -> Z.rem a b <= a.
Proof. intros. zero_or_not b. apply Z.rem_le; auto with zarith. Qed.
(** Some additionnal inequalities about Zdiv. *)
Theorem Zquot_le_upper_bound:
forall a b q, 0 < b -> a <= q*b -> a÷b <= q.
-Proof. intros a b q; rewrite Zmult_comm; apply Z.quot_le_upper_bound. Qed.
+Proof. intros a b q; rewrite Z.mul_comm; apply Z.quot_le_upper_bound. Qed.
Theorem Zquot_lt_upper_bound:
forall a b q, 0 <= a -> 0 < b -> a < q*b -> a÷b < q.
-Proof. intros a b q; rewrite Zmult_comm; apply Z.quot_lt_upper_bound. Qed.
+Proof. intros a b q; rewrite Z.mul_comm; apply Z.quot_lt_upper_bound. Qed.
Theorem Zquot_le_lower_bound:
forall a b q, 0 < b -> q*b <= a -> q <= a÷b.
-Proof. intros a b q; rewrite Zmult_comm; apply Z.quot_le_lower_bound. Qed.
+Proof. intros a b q; rewrite Z.mul_comm; apply Z.quot_le_lower_bound. Qed.
Theorem Zquot_sgn: forall a b,
0 <= Z.sgn (a÷b) * Z.sgn a * Z.sgn b.
@@ -374,22 +294,22 @@ Proof. intros. zero_or_not b. apply Z.quot_mul_cancel_r; auto. Qed.
Lemma Zquot_mult_cancel_l : forall a b c:Z,
c<>0 -> (c*a)÷(c*b) = a÷b.
Proof.
- intros. rewrite (Zmult_comm c b). zero_or_not b.
- rewrite (Zmult_comm b c). apply Z.quot_mul_cancel_l; auto.
+ intros. rewrite (Z.mul_comm c b). zero_or_not b.
+ rewrite (Z.mul_comm b c). apply Z.quot_mul_cancel_l; auto.
Qed.
Lemma Zmult_rem_distr_l: forall a b c,
Z.rem (c*a) (c*b) = c * (Z.rem a b).
Proof.
- intros. zero_or_not c. rewrite (Zmult_comm c b). zero_or_not b.
- rewrite (Zmult_comm b c). apply Z.mul_rem_distr_l; auto.
+ intros. zero_or_not c. rewrite (Z.mul_comm c b). zero_or_not b.
+ rewrite (Z.mul_comm b c). apply Z.mul_rem_distr_l; auto.
Qed.
Lemma Zmult_rem_distr_r: forall a b c,
Z.rem (a*c) (b*c) = (Z.rem a b) * c.
Proof.
- intros. zero_or_not b. rewrite (Zmult_comm b c). zero_or_not c.
- rewrite (Zmult_comm c b). apply Z.mul_rem_distr_r; auto.
+ intros. zero_or_not b. rewrite (Z.mul_comm b c). zero_or_not c.
+ rewrite (Z.mul_comm c b). apply Z.mul_rem_distr_r; auto.
Qed.
(** Operations modulo. *)
@@ -424,7 +344,7 @@ Lemma Zplus_rem_idemp_r: forall a b n,
Z.rem (b + Z.rem a n) n = Z.rem (b + a) n.
Proof.
intros. zero_or_not n. apply Z.add_rem_idemp_r; auto.
- rewrite Zmult_comm; auto.
+ rewrite Z.mul_comm; auto.
Qed.
Lemma Zmult_rem_idemp_l: forall a b n, Z.rem (Z.rem a n * b) n = Z.rem (a * b) n.
@@ -437,8 +357,8 @@ Proof. intros. zero_or_not n. apply Z.mul_rem_idemp_r; auto. Qed.
Lemma Zquot_Zquot : forall a b c, (a÷b)÷c = a÷(b*c).
Proof.
- intros. zero_or_not b. rewrite Zmult_comm. zero_or_not c.
- rewrite Zmult_comm. apply Z.quot_quot; auto.
+ intros. zero_or_not b. rewrite Z.mul_comm. zero_or_not c.
+ rewrite Z.mul_comm. apply Z.quot_quot; auto.
Qed.
(** A last inequality: *)
@@ -468,28 +388,26 @@ Proof.
right. destruct p; simpl; split; now auto with zarith.
Qed.
-Notation Zquot2_quot := Zquot2_quot (only parsing).
-
Lemma Zrem_odd : forall a, Z.rem a 2 = if Z.odd a then Z.sgn a else 0.
Proof.
intros. symmetry.
- apply Zrem_unique_full with (Zquot2 a).
+ apply Zrem_unique_full with (Z.quot2 a).
apply Zquot2_odd_remainder.
apply Zquot2_odd_eqn.
Qed.
Lemma Zrem_even : forall a, Z.rem a 2 = if Z.even a then 0 else Z.sgn a.
Proof.
- intros a. rewrite Zrem_odd, Zodd_even_bool. now destruct Zeven_bool.
+ intros a. rewrite Zrem_odd, Zodd_even_bool. now destruct Z.even.
Qed.
-Lemma Zeven_rem : forall a, Z.even a = Zeq_bool (Z.rem a 2) 0.
+Lemma Zeven_rem : forall a, Z.even a = Z.eqb (Z.rem a 2) 0.
Proof.
intros a. rewrite Zrem_even.
destruct a as [ |p|p]; trivial; now destruct p.
Qed.
-Lemma Zodd_rem : forall a, Z.odd a = negb (Zeq_bool (Z.rem a 2) 0).
+Lemma Zodd_rem : forall a, Z.odd a = negb (Z.eqb (Z.rem a 2) 0).
Proof.
intros a. rewrite Zrem_odd.
destruct a as [ |p|p]; trivial; now destruct p.
@@ -505,18 +423,17 @@ Proof.
intros.
apply Zdiv_mod_unique with b.
apply Zrem_lt_pos; auto with zarith.
- rewrite Zabs_eq; auto with *; apply Z_mod_lt; auto with *.
+ rewrite Z.abs_eq; auto with *; apply Z_mod_lt; auto with *.
rewrite <- Z_div_mod_eq; auto with *.
- symmetry; apply Z_quot_rem_eq; auto with *.
+ symmetry; apply Z.quot_rem; auto with *.
Qed.
Theorem Zquot_Zdiv_pos : forall a b, 0 <= a -> 0 <= b ->
a÷b = a/b.
Proof.
- intros a b Ha Hb.
- destruct (Zle_lt_or_eq _ _ Hb).
- generalize (Zquotrem_Zdiv_eucl_pos a b Ha H); intuition.
- subst; rewrite Zquot_0_r, Zdiv_0_r; reflexivity.
+ intros a b Ha Hb. Z.le_elim Hb.
+ - generalize (Zquotrem_Zdiv_eucl_pos a b Ha Hb); intuition.
+ - subst; now rewrite Zquot_0_r, Zdiv_0_r.
Qed.
Theorem Zrem_Zmod_pos : forall a b, 0 <= a -> 0 < b ->
diff --git a/theories/ZArith/Zsqrt_compat.v b/theories/ZArith/Zsqrt_compat.v
index 4584c3f8f..54f6f2e9a 100644
--- a/theories/ZArith/Zsqrt_compat.v
+++ b/theories/ZArith/Zsqrt_compat.v
@@ -32,12 +32,12 @@ Ltac compute_POS :=
| |- context [(Zpos (xI ?X1))] =>
match constr:X1 with
| context [1%positive] => fail 1
- | _ => rewrite (BinInt.Zpos_xI X1)
+ | _ => rewrite (Pos2Z.inj_xI X1)
end
| |- context [(Zpos (xO ?X1))] =>
match constr:X1 with
| context [1%positive] => fail 1
- | _ => rewrite (BinInt.Zpos_xO X1)
+ | _ => rewrite (Pos2Z.inj_xO X1)
end
end.
@@ -115,7 +115,7 @@ Definition Zsqrt :
fun h =>
match sqrtrempos p with
| c_sqrt s r Heq Hint =>
- existS
+ existT
(fun s:Z =>
{r : Z |
Zpos p = s * s + r /\ s * s <= Zpos p < (s + 1) * (s + 1)})
@@ -131,10 +131,10 @@ Definition Zsqrt :
{s : Z &
{r : Z |
Zneg p = s * s + r /\ s * s <= Zneg p < (s + 1) * (s + 1)}}
- (h (refl_equal Datatypes.Gt))
+ (h (eq_refl Datatypes.Gt))
| Z0 =>
fun h =>
- existS
+ existT
(fun s:Z =>
{r : Z | 0 = s * s + r /\ s * s <= 0 < (s + 1) * (s + 1)}) 0
(exist
@@ -149,8 +149,8 @@ Defined.
Definition Zsqrt_plain (x:Z) : Z :=
match x with
| Zpos p =>
- match Zsqrt (Zpos p) (Zorder.Zle_0_pos p) with
- | existS s _ => s
+ match Zsqrt (Zpos p) (Pos2Z.is_nonneg p) with
+ | existT s _ => s
end
| Zneg p => 0
| Z0 => 0
@@ -164,12 +164,11 @@ Theorem Zsqrt_interval :
Zsqrt_plain n * Zsqrt_plain n <= n <
(Zsqrt_plain n + 1) * (Zsqrt_plain n + 1).
Proof.
- 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.
+ intros [|p|p] Hp.
+ - now compute.
+ - unfold Zsqrt_plain.
+ now destruct Zsqrt as (s & r & Heq & Hint).
+ - now elim Hp.
Qed.
(** Positivity *)
@@ -177,9 +176,9 @@ Qed.
Theorem Zsqrt_plain_is_pos: forall n, 0 <= n -> 0 <= Zsqrt_plain n.
Proof.
intros n m; case (Zsqrt_interval n); auto with zarith.
- intros H1 H2; case (Zle_or_lt 0 (Zsqrt_plain n)); auto.
- intros H3; contradict H2; auto; apply Zle_not_lt.
- apply Zle_trans with ( 2 := H1 ).
+ intros H1 H2; case (Z.le_gt_cases 0 (Zsqrt_plain n)); auto.
+ intros H3; contradict H2; auto; apply Z.le_ngt.
+ apply Z.le_trans with ( 2 := H1 ).
replace ((Zsqrt_plain n + 1) * (Zsqrt_plain n + 1))
with (Zsqrt_plain n * Zsqrt_plain n + (2 * Zsqrt_plain n + 1));
auto with zarith.
@@ -194,13 +193,13 @@ Proof.
generalize (Zsqrt_plain_is_pos (a * a)); auto with zarith; intros Haa.
case (Zsqrt_interval (a * a)); auto with zarith.
intros H1 H2.
- case (Zle_or_lt a (Zsqrt_plain (a * a))); intros H3; auto.
- case Zle_lt_or_eq with (1:=H3); auto; clear H3; intros H3.
- contradict H1; auto; apply Zlt_not_le; auto with zarith.
- apply Zle_lt_trans with (a * Zsqrt_plain (a * a)); auto with zarith.
- apply Zmult_lt_compat_r; auto with zarith.
- contradict H2; auto; apply Zle_not_lt; auto with zarith.
- apply Zmult_le_compat; auto with zarith.
+ case (Z.le_gt_cases a (Zsqrt_plain (a * a))); intros H3.
+ - Z.le_elim H3; auto.
+ contradict H1; auto; apply Z.lt_nge; auto with zarith.
+ apply Z.le_lt_trans with (a * Zsqrt_plain (a * a)); auto with zarith.
+ apply Z.mul_lt_mono_pos_r; auto with zarith.
+ - contradict H2; auto; apply Z.le_ngt; auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
Qed.
(** [Zsqrt_plain] is increasing *)
@@ -208,16 +207,16 @@ Qed.
Theorem Zsqrt_le:
forall p q, 0 <= p <= q -> Zsqrt_plain p <= Zsqrt_plain q.
Proof.
- intros p q [H1 H2]; case Zle_lt_or_eq with (1:=H2); clear H2; intros H2;
- [ | subst q; auto with zarith].
- case (Zle_or_lt (Zsqrt_plain p) (Zsqrt_plain q)); auto; intros H3.
+ intros p q [H1 H2].
+ Z.le_elim H2; [ | subst q; auto with zarith].
+ case (Z.le_gt_cases (Zsqrt_plain p) (Zsqrt_plain q)); auto; intros H3.
assert (Hp: (0 <= Zsqrt_plain q)).
- apply Zsqrt_plain_is_pos; auto with zarith.
+ { apply Zsqrt_plain_is_pos; auto with zarith. }
absurd (q <= p); auto with zarith.
- apply Zle_trans with ((Zsqrt_plain q + 1) * (Zsqrt_plain q + 1)).
+ apply Z.le_trans with ((Zsqrt_plain q + 1) * (Zsqrt_plain q + 1)).
case (Zsqrt_interval q); auto with zarith.
- apply Zle_trans with (Zsqrt_plain p * Zsqrt_plain p); auto with zarith.
- apply Zmult_le_compat; auto with zarith.
+ apply Z.le_trans with (Zsqrt_plain p * Zsqrt_plain p); auto with zarith.
+ apply Z.mul_le_mono_nonneg; auto with zarith.
case (Zsqrt_interval p); auto with zarith.
Qed.
diff --git a/theories/ZArith/Zwf.v b/theories/ZArith/Zwf.v
index 30802f824..efe5b6847 100644
--- a/theories/ZArith/Zwf.v
+++ b/theories/ZArith/Zwf.v
@@ -29,7 +29,7 @@ Section wf_proof.
(** The proof of well-foundness is classic: we do the proof by induction
on a measure in nat, which is here [|x-c|] *)
- Let f (z:Z) := Zabs_nat (z - c).
+ Let f (z:Z) := Z.abs_nat (z - c).
Lemma Zwf_well_founded : well_founded (Zwf c).
red in |- *; intros.
@@ -45,12 +45,12 @@ Section wf_proof.
apply Acc_intro; intros.
apply H.
unfold Zwf in H1.
- case (Zle_or_lt c y); intro; auto with zarith.
+ case (Z.le_gt_cases 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.
+ unfold f.
+ apply Zabs2Nat.inj_lt; omega.
apply (H (S (f a))); auto.
Qed.
@@ -75,18 +75,15 @@ Section wf_proof_up.
(** The proof of well-foundness is classic: we do the proof by induction
on a measure in nat, which is here [|c-x|] *)
- Let f (z:Z) := Zabs_nat (c - z).
+ Let f (z:Z) := Z.abs_nat (c - z).
Lemma Zwf_up_well_founded : well_founded (Zwf_up c).
Proof.
apply well_founded_lt_compat with (f := f).
- unfold Zwf_up, f in |- *.
+ unfold Zwf_up, f.
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.
+ apply Zabs2Nat.inj_lt; try (apply Z.le_0_sub; intuition).
+ now apply Z.sub_lt_mono_l.
Qed.
End wf_proof_up.