summaryrefslogtreecommitdiff
path: root/theories/Numbers
diff options
context:
space:
mode:
authorGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
committerGravatar Stephane Glondu <steph@glondu.net>2010-07-21 09:46:51 +0200
commit5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 (patch)
tree631ad791a7685edafeb1fb2e8faeedc8379318ae /theories/Numbers
parentda178a880e3ace820b41d38b191d3785b82991f5 (diff)
Imported Upstream snapshot 8.3~beta0+13298
Diffstat (limited to 'theories/Numbers')
-rw-r--r--theories/Numbers/BigNumPrelude.v96
-rw-r--r--theories/Numbers/Cyclic/Abstract/CyclicAxioms.v159
-rw-r--r--theories/Numbers/Cyclic/Abstract/NZCyclic.v173
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v74
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v94
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v168
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v324
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v144
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v66
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v114
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v94
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v76
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v18
-rw-r--r--theories/Numbers/Cyclic/Int31/Cyclic31.v464
-rw-r--r--theories/Numbers/Cyclic/Int31/Int31.v141
-rw-r--r--theories/Numbers/Cyclic/Int31/Ring31.v103
-rw-r--r--theories/Numbers/Cyclic/ZModulo/ZModulo.v227
-rw-r--r--theories/Numbers/Integer/Abstract/ZAdd.v318
-rw-r--r--theories/Numbers/Integer/Abstract/ZAddOrder.v337
-rw-r--r--theories/Numbers/Integer/Abstract/ZAxioms.v61
-rw-r--r--theories/Numbers/Integer/Abstract/ZBase.v69
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivEucl.v605
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivFloor.v632
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivTrunc.v532
-rw-r--r--theories/Numbers/Integer/Abstract/ZDomain.v69
-rw-r--r--theories/Numbers/Integer/Abstract/ZLt.v402
-rw-r--r--theories/Numbers/Integer/Abstract/ZMul.v105
-rw-r--r--theories/Numbers/Integer/Abstract/ZMulOrder.v356
-rw-r--r--theories/Numbers/Integer/Abstract/ZProperties.v24
-rw-r--r--theories/Numbers/Integer/Abstract/ZSgnAbs.v348
-rw-r--r--theories/Numbers/Integer/BigZ/BigZ.v173
-rw-r--r--theories/Numbers/Integer/BigZ/ZMake.v379
-rw-r--r--theories/Numbers/Integer/Binary/ZBinary.v277
-rw-r--r--theories/Numbers/Integer/NatPairs/ZNatPairs.v506
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSig.v116
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v267
-rw-r--r--theories/Numbers/NaryFunctions.v70
-rw-r--r--theories/Numbers/NatInt/NZAdd.v87
-rw-r--r--theories/Numbers/NatInt/NZAddOrder.v141
-rw-r--r--theories/Numbers/NatInt/NZAxioms.v202
-rw-r--r--theories/Numbers/NatInt/NZBase.v69
-rw-r--r--theories/Numbers/NatInt/NZDiv.v542
-rw-r--r--theories/Numbers/NatInt/NZDomain.v417
-rw-r--r--theories/Numbers/NatInt/NZMul.v74
-rw-r--r--theories/Numbers/NatInt/NZMulOrder.v325
-rw-r--r--theories/Numbers/NatInt/NZOrder.v708
-rw-r--r--theories/Numbers/NatInt/NZProperties.v20
-rw-r--r--theories/Numbers/Natural/Abstract/NAdd.v109
-rw-r--r--theories/Numbers/Natural/Abstract/NAddOrder.v88
-rw-r--r--theories/Numbers/Natural/Abstract/NAxioms.v58
-rw-r--r--theories/Numbers/Natural/Abstract/NBase.v180
-rw-r--r--theories/Numbers/Natural/Abstract/NDefOps.v477
-rw-r--r--theories/Numbers/Natural/Abstract/NDiv.v239
-rw-r--r--theories/Numbers/Natural/Abstract/NIso.v84
-rw-r--r--theories/Numbers/Natural/Abstract/NMul.v87
-rw-r--r--theories/Numbers/Natural/Abstract/NMulOrder.v101
-rw-r--r--theories/Numbers/Natural/Abstract/NOrder.v390
-rw-r--r--theories/Numbers/Natural/Abstract/NProperties.v22
-rw-r--r--theories/Numbers/Natural/Abstract/NStrongRec.v231
-rw-r--r--theories/Numbers/Natural/Abstract/NSub.v196
-rw-r--r--theories/Numbers/Natural/BigN/BigN.v192
-rw-r--r--theories/Numbers/Natural/BigN/NMake.v524
-rw-r--r--theories/Numbers/Natural/BigN/NMake_gen.ml929
-rw-r--r--theories/Numbers/Natural/BigN/Nbasic.v64
-rw-r--r--theories/Numbers/Natural/Binary/NBinDefs.v267
-rw-r--r--theories/Numbers/Natural/Binary/NBinary.v173
-rw-r--r--theories/Numbers/Natural/Peano/NPeano.v249
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSig.v119
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v292
-rw-r--r--theories/Numbers/NumPrelude.v152
-rw-r--r--theories/Numbers/Rational/BigQ/BigQ.v207
-rw-r--r--theories/Numbers/Rational/BigQ/QMake.v721
-rw-r--r--theories/Numbers/Rational/SpecViaQ/QSig.v202
-rw-r--r--theories/Numbers/vo.itarget70
74 files changed, 9726 insertions, 7463 deletions
diff --git a/theories/Numbers/BigNumPrelude.v b/theories/Numbers/BigNumPrelude.v
index 83ecd10d..dd7d9046 100644
--- a/theories/Numbers/BigNumPrelude.v
+++ b/theories/Numbers/BigNumPrelude.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: BigNumPrelude.v 11207 2008-07-04 16:50:32Z letouzey $ i*)
+(*i $Id$ i*)
(** * BigNumPrelude *)
@@ -21,6 +21,8 @@ Require Export ZArith.
Require Export Znumtheory.
Require Export Zpow_facts.
+Declare ML Module "numbers_syntax_plugin".
+
(* *** Nota Bene ***
All results that were general enough has been moved in ZArith.
Only remain here specialized lemmas and compatibility elements.
@@ -28,8 +30,8 @@ Require Export Zpow_facts.
*)
-Open Local Scope Z_scope.
-
+Local Open Scope Z_scope.
+
(* For compatibility of scripts, weaker version of some lemmas of Zdiv *)
Lemma Zlt0_not_eq : forall n, 0<n -> n<>0.
@@ -43,14 +45,14 @@ 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 (Zle _ _) =>
(match goal with
|- Zpos _ <= Zpos _ => exact (refl_equal _)
| H: _ <= ?p |- _ <= ?p => apply Zle_trans with (2 := H)
| H: _ < ?p |- _ <= ?p => apply Zlt_le_weak; apply Zle_lt_trans with (2 := H)
end).
-Hint Extern 2 (Zlt _ _) =>
+Hint Extern 2 (Zlt _ _) =>
(match goal with
|- Zpos _ < Zpos _ => exact (refl_equal _)
| H: _ <= ?p |- _ <= ?p => apply Zlt_le_trans with (2 := H)
@@ -60,13 +62,13 @@ Hint Extern 2 (Zlt _ _) =>
Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
-(**************************************
+(**************************************
Properties of order and product
**************************************)
- Theorem beta_lex: forall a b c d beta,
- a * beta + b <= c * beta + d ->
- 0 <= b < beta -> 0 <= d < beta ->
+ Theorem beta_lex: forall a b c d beta,
+ a * beta + b <= c * beta + d ->
+ 0 <= b < beta -> 0 <= d < beta ->
a <= c.
Proof.
intros a b c d beta H1 (H3, H4) (H5, H6).
@@ -78,15 +80,15 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
Theorem beta_lex_inv: forall a b c d beta,
a < c -> 0 <= b < beta ->
- 0 <= d < beta ->
- a * beta + b < c * beta + d.
+ 0 <= d < beta ->
+ 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.
Qed.
- Lemma beta_mult : forall h l beta,
+ 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.
@@ -94,7 +96,7 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
apply beta_lex_inv;auto with zarith.
Qed.
- Lemma Zmult_lt_b :
+ Lemma Zmult_lt_b :
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.
@@ -104,17 +106,17 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith.
Qed.
Lemma sum_mul_carry : forall xh xl yh yl wc cc beta,
- 1 < beta ->
+ 1 < beta ->
0 <= wc < beta ->
0 <= xh < beta ->
0 <= xl < beta ->
0 <= yh < beta ->
0 <= yl < beta ->
0 <= cc < beta^2 ->
- wc*beta^2 + cc = xh*yl + xl*yh ->
+ wc*beta^2 + cc = xh*yl + xl*yh ->
0 <= wc <= 1.
Proof.
- intros xh xl yh yl wc cc beta U H1 H2 H3 H4 H5 H6 H7.
+ intros xh xl yh yl wc cc beta U H1 H2 H3 H4 H5 H6 H7.
assert (H8 := Zmult_lt_b beta xh yl H2 H5).
assert (H9 := Zmult_lt_b beta xl yh H3 H4).
split;auto with zarith.
@@ -132,7 +134,7 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: 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);
+ repeat (rewrite Zmult_minus_distr_l || rewrite Zmult_minus_distr_r);
rewrite Zpower_2; auto with zarith.
Qed.
@@ -147,7 +149,7 @@ Hint Resolve Zlt_gt Zle_ge Z_div_pos: 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);
+ repeat (rewrite Zmult_minus_distr_l || rewrite Zmult_minus_distr_r);
rewrite Zpower_2; auto with zarith.
Qed.
@@ -199,9 +201,9 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
apply Zplus_le_lt_compat; auto with zarith.
replace b with ((b - a) + a); try ring.
rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
+ 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;
+ rewrite (Zmult_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l;
auto with zarith.
rewrite (Zmult_comm (2 ^a)); apply Zmult_le_compat_r; auto with zarith.
match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
@@ -222,22 +224,22 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
split; auto with zarith.
assert (0 <= 2 ^a * r); auto with zarith.
apply Zplus_le_0_compat; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
pattern (2 ^ b) at 2;replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); try ring.
apply Zplus_le_lt_compat; auto with zarith.
replace b with ((b - a) + a); try ring.
rewrite Zpower_exp; auto with zarith.
- pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
+ pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a));
try rewrite <- Zmult_minus_distr_r.
repeat rewrite (fun x => Zmult_comm x (2 ^ a)); rewrite Zmult_mod_distr_l;
auto with zarith.
apply Zmult_le_compat_l; auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
Qed.
- Theorem Zdiv_shift_r:
+ Theorem Zdiv_shift_r:
forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a ->
(r * 2 ^a + t) / (2 ^ b) = (r * 2 ^a) / (2 ^ b).
Proof.
@@ -251,7 +253,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
rewrite <- Zmod_shift_r; auto with zarith.
rewrite (Zmult_comm (2 ^ b)); rewrite Z_div_plus_full_l; auto with zarith.
rewrite (fun x y => @Zdiv_small (x mod y)); auto with zarith.
- match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
+ match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end;
auto with zarith.
Qed.
@@ -262,8 +264,8 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
a * 2^p = a / 2^(n - p) * 2^n + (a*2^p) mod 2^n.
Proof.
intros n p a H1 H2.
- pattern (a*2^p) at 1;replace (a*2^p) with
- (a*2^p/2^n * 2^n + a*2^p mod 2^n).
+ pattern (a*2^p) at 1;replace (a*2^p) with
+ (a*2^p/2^n * 2^n + a*2^p mod 2^n).
2:symmetry;rewrite (Zmult_comm (a*2^p/2^n));apply Z_div_mod_eq.
replace (a * 2 ^ p / 2 ^ n) with (a / 2 ^ (n - p));trivial.
replace (2^n) with (2^(n-p)*2^p).
@@ -277,8 +279,8 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Qed.
- Lemma shift_unshift_mod_2 : forall n p a, 0 <= p <= n ->
- ((a * 2 ^ (n - p)) mod (2^n) / 2 ^ (n - p)) mod (2^n) =
+ Lemma shift_unshift_mod_2 : forall n p a, 0 <= p <= n ->
+ ((a * 2 ^ (n - p)) mod (2^n) / 2 ^ (n - p)) mod (2^n) =
a mod 2 ^ p.
Proof.
intros.
@@ -310,16 +312,16 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Lemma div_le_0 : forall p x, 0 <= x -> 0 <= x / 2 ^ p.
Proof.
intros p x Hle;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_le_lower_bound;auto with zarith.
+ apply Zdiv_le_lower_bound;auto with zarith.
replace (2^p) with 0.
destruct x;compute;intro;discriminate.
destruct p;trivial;discriminate z.
Qed.
-
+
Lemma div_lt : forall p x y, 0 <= x < y -> x / 2^p < y.
Proof.
intros p x y H;destruct (Z_le_gt_dec 0 p).
- apply Zdiv_lt_upper_bound;auto with zarith.
+ apply Zdiv_lt_upper_bound;auto with zarith.
apply Zlt_le_trans with y;auto with zarith.
rewrite <- (Zmult_1_r y);apply Zmult_le_compat;auto with zarith.
assert (0 < 2^p);auto with zarith.
@@ -331,7 +333,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Theorem Zgcd_div_pos a b:
0 < b -> 0 < Zgcd a b -> 0 < b / Zgcd a b.
Proof.
- intros a b Ha Hg.
+ 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.
@@ -343,7 +345,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Theorem Zdiv_neg a b:
a < 0 -> 0 < b -> a / b < 0.
Proof.
- intros a b Ha Hb.
+ intros Ha Hb.
assert (b > 0) by omega.
generalize (Z_mult_div_ge a _ H); intros.
assert (b * (a / b) < 0)%Z.
@@ -354,22 +356,8 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
compute in H1; discriminate.
compute; auto.
Qed.
-
- Lemma Zgcd_Zabs : forall z z', Zgcd (Zabs z) z' = Zgcd z z'.
- Proof.
- destruct z; simpl; auto.
- Qed.
- Lemma Zgcd_sym : forall p q, Zgcd p q = Zgcd q p.
- Proof.
- intros.
- apply Zis_gcd_gcd.
- apply Zgcd_is_pos.
- apply Zis_gcd_sym.
- apply Zgcd_is_gcd.
- Qed.
-
- Lemma Zdiv_gcd_zero : forall a b, b / Zgcd a b = 0 -> b <> 0 ->
+ Lemma Zdiv_gcd_zero : forall a b, b / Zgcd a b = 0 -> b <> 0 ->
Zgcd a b = 0.
Proof.
intros.
@@ -381,13 +369,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
intros; subst k; simpl in *; subst b; elim H0; auto.
Qed.
- Lemma Zgcd_1 : forall z, Zgcd z 1 = 1.
- Proof.
- intros; apply Zis_gcd_gcd; auto with zarith; apply Zis_gcd_1.
- Qed.
- Hint Resolve Zgcd_1.
-
- Lemma Zgcd_mult_rel_prime : forall a b c,
+ Lemma Zgcd_mult_rel_prime : forall a b c,
Zgcd a c = 1 -> Zgcd b c = 1 -> Zgcd (a*b) c = 1.
Proof.
intros.
@@ -396,7 +378,7 @@ Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a.
Qed.
Lemma Zcompare_gt : forall (A:Type)(a a':A)(p q:Z),
- match (p?=q)%Z with Gt => a | _ => a' end =
+ match (p?=q)%Z with Gt => a | _ => a' end =
if Z_le_gt_dec p q then a' else a.
Proof.
intros.
diff --git a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
index 528d78c3..51df2fa3 100644
--- a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
+++ b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v
@@ -8,12 +8,12 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(* $Id: CyclicAxioms.v 11012 2008-05-28 16:34:43Z letouzey $ *)
+(* $Id$ *)
(** * Signature and specification of a bounded integer structure *)
-(** This file specifies how to represent [Z/nZ] when [n=2^d],
- [d] being the number of digits of these bounded integers. *)
+(** This file specifies how to represent [Z/nZ] when [n=2^d],
+ [d] being the number of digits of these bounded integers. *)
Set Implicit Arguments.
@@ -22,7 +22,7 @@ Require Import Znumtheory.
Require Import BigNumPrelude.
Require Import DoubleType.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
(** First, a description via an operator record and a spec record. *)
@@ -33,7 +33,7 @@ Section Z_nZ_Op.
Record znz_op := mk_znz_op {
(* Conversion functions with Z *)
- znz_digits : positive;
+ znz_digits : positive;
znz_zdigits: znz;
znz_to_Z : znz -> Z;
znz_of_pos : positive -> N * znz; (* Euclidean division by [2^digits] *)
@@ -78,12 +78,12 @@ Section Z_nZ_Op.
znz_div : znz -> znz -> znz * znz;
znz_mod_gt : znz -> znz -> znz; (* specialized version of [znz_mod] *)
- znz_mod : znz -> znz -> znz;
+ znz_mod : znz -> znz -> znz;
znz_gcd_gt : znz -> znz -> znz; (* specialized version of [znz_gcd] *)
- znz_gcd : znz -> znz -> znz;
+ znz_gcd : znz -> znz -> znz;
(* [znz_add_mul_div p i j] is a combination of the [(digits-p)]
- low bits of [i] above the [p] high bits of [j]:
+ low bits of [i] above the [p] high bits of [j]:
[znz_add_mul_div p i j = i*2^p+j/2^(digits-p)] *)
znz_add_mul_div : znz -> znz -> znz -> znz;
(* [znz_pos_mod p i] is [i mod 2^p] *)
@@ -135,7 +135,7 @@ Section Z_nZ_Spec.
Let w_mul_c := w_op.(znz_mul_c).
Let w_mul := w_op.(znz_mul).
Let w_square_c := w_op.(znz_square_c).
-
+
Let w_div21 := w_op.(znz_div21).
Let w_div_gt := w_op.(znz_div_gt).
Let w_div := w_op.(znz_div).
@@ -229,25 +229,25 @@ Section Z_nZ_Spec.
spec_div : forall a b, 0 < [|b|] ->
let (q,r) := w_div a b in
[|a|] = [|q|] * [|b|] + [|r|] /\
- 0 <= [|r|] < [|b|];
-
+ 0 <= [|r|] < [|b|];
+
spec_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] ->
[|w_mod_gt a b|] = [|a|] mod [|b|];
spec_mod : forall a b, 0 < [|b|] ->
[|w_mod a b|] = [|a|] mod [|b|];
-
+
spec_gcd_gt : forall a b, [|a|] > [|b|] ->
Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|];
spec_gcd : forall a b, Zis_gcd [|a|] [|b|] [|w_gcd a b|];
-
+
(* shift operations *)
spec_head00: forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits;
spec_head0 : forall x, 0 < [|x|] ->
- wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB;
+ wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB;
spec_tail00: forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits;
- spec_tail0 : forall x, 0 < [|x|] ->
- exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|w_tail0 x|]) ;
+ spec_tail0 : forall x, 0 < [|x|] ->
+ exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|w_tail0 x|]) ;
spec_add_mul_div : forall x y p,
[|p|] <= Zpos w_digits ->
[| w_add_mul_div p x y |] =
@@ -272,23 +272,23 @@ End Z_nZ_Spec.
(** Generic construction of double words *)
Section WW.
-
+
Variable w : Type.
Variable w_op : znz_op w.
Variable op_spec : znz_spec w_op.
-
+
Let wB := base w_op.(znz_digits).
Let w_to_Z := w_op.(znz_to_Z).
Let w_eq0 := w_op.(znz_eq0).
Let w_0 := w_op.(znz_0).
- Definition znz_W0 h :=
+ Definition znz_W0 h :=
if w_eq0 h then W0 else WW h w_0.
- Definition znz_0W l :=
+ Definition znz_0W l :=
if w_eq0 l then W0 else WW w_0 l.
- Definition znz_WW h l :=
+ Definition znz_WW h l :=
if w_eq0 h then znz_0W l else WW h l.
Lemma spec_W0 : forall h,
@@ -300,7 +300,7 @@ Section WW.
unfold w_0; rewrite op_spec.(spec_0); auto with zarith.
Qed.
- Lemma spec_0W : forall l,
+ Lemma spec_0W : forall l,
zn2z_to_Z wB w_to_Z (znz_0W l) = w_to_Z l.
Proof.
unfold zn2z_to_Z, znz_0W, w_to_Z; simpl; intros.
@@ -309,7 +309,7 @@ Section WW.
unfold w_0; rewrite op_spec.(spec_0); auto with zarith.
Qed.
- Lemma spec_WW : forall h l,
+ Lemma spec_WW : forall h l,
zn2z_to_Z wB w_to_Z (znz_WW h l) = (w_to_Z h)*wB + w_to_Z l.
Proof.
unfold znz_WW, w_to_Z; simpl; intros.
@@ -324,7 +324,7 @@ End WW.
(** Injecting [Z] numbers into a cyclic structure *)
Section znz_of_pos.
-
+
Variable w : Type.
Variable w_op : znz_op w.
Variable op_spec : znz_spec w_op.
@@ -349,7 +349,7 @@ Section znz_of_pos.
apply Zle_trans with X; auto with zarith
end.
match goal with |- ?X <= _ =>
- pattern X at 1; rewrite <- (Zmult_1_l);
+ pattern X at 1; rewrite <- (Zmult_1_l);
apply Zmult_le_compat_r; auto with zarith
end.
case p1; simpl; intros; red; simpl; intros; discriminate.
@@ -373,3 +373,112 @@ Module Type CyclicType.
Parameter w_op : znz_op w.
Parameter w_spec : znz_spec w_op.
End CyclicType.
+
+
+(** A Cyclic structure can be seen as a ring *)
+
+Module CyclicRing (Import Cyclic : CyclicType).
+
+Definition t := w.
+
+Local Notation "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99).
+
+Definition eq (n m : t) := [| n |] = [| m |].
+Definition zero : t := w_op.(znz_0).
+Definition one := w_op.(znz_1).
+Definition add := w_op.(znz_add).
+Definition sub := w_op.(znz_sub).
+Definition mul := w_op.(znz_mul).
+Definition opp := w_op.(znz_opp).
+
+Local Infix "==" := eq (at level 70).
+Local Notation "0" := zero.
+Local Notation "1" := one.
+Local Infix "+" := add.
+Local Infix "-" := sub.
+Local Infix "*" := mul.
+Local Notation "!!" := (base (znz_digits w_op)).
+
+Hint Rewrite
+ w_spec.(spec_0) w_spec.(spec_1)
+ w_spec.(spec_add) w_spec.(spec_mul) w_spec.(spec_opp) w_spec.(spec_sub)
+ : cyclic.
+
+Ltac zify :=
+ unfold eq, zero, one, add, sub, mul, opp in *; autorewrite with cyclic.
+
+Lemma add_0_l : forall x, 0 + x == x.
+Proof.
+intros. zify. rewrite Zplus_0_l.
+apply Zmod_small. apply w_spec.(spec_to_Z).
+Qed.
+
+Lemma add_comm : forall x y, x + y == y + x.
+Proof.
+intros. zify. now rewrite Zplus_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.
+Qed.
+
+Lemma mul_1_l : forall x, 1 * x == x.
+Proof.
+intros. zify. rewrite Zmult_1_l.
+apply Zmod_small. apply w_spec.(spec_to_Z).
+Qed.
+
+Lemma mul_comm : forall x y, x * y == y * x.
+Proof.
+intros. zify. now rewrite Zmult_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.
+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.
+Qed.
+
+Lemma add_opp_r : forall x y, x + opp y == x-y.
+Proof.
+intros. zify. rewrite <- Zminus_mod_idemp_r. unfold Zminus.
+destruct (Z_eq_dec ([|y|] mod !!) 0) as [EQ|NEQ].
+rewrite Z_mod_zero_opp_full, EQ, 2 Zplus_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.
+Qed.
+
+Lemma add_opp_diag_r : forall x, x + opp x == 0.
+Proof.
+intros. red. rewrite add_opp_r. zify. now rewrite Zminus_diag, Zmod_0_l.
+Qed.
+
+Lemma CyclicRing : ring_theory 0 1 add mul sub opp eq.
+Proof.
+constructor.
+exact add_0_l. exact add_comm. exact add_assoc.
+exact mul_1_l. exact mul_comm. exact mul_assoc.
+exact mul_add_distr_r.
+symmetry. apply add_opp_r.
+exact add_opp_diag_r.
+Qed.
+
+Definition eqb x y :=
+ match w_op.(znz_compare) x y with Eq => true | _ => false end.
+
+Lemma eqb_eq : forall x y, eqb x y = true <-> x == y.
+Proof.
+ intros. unfold eqb, eq. generalize (w_spec.(spec_compare) x y).
+ destruct (w_op.(znz_compare) x y); intuition; try discriminate.
+Qed.
+
+Lemma eqb_correct : forall x y, eqb x y = true -> x==y.
+Proof. now apply eqb_eq. Qed.
+
+End CyclicRing. \ No newline at end of file
diff --git a/theories/Numbers/Cyclic/Abstract/NZCyclic.v b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
index fb3f0cef..517e48ad 100644
--- a/theories/Numbers/Cyclic/Abstract/NZCyclic.v
+++ b/theories/Numbers/Cyclic/Abstract/NZCyclic.v
@@ -8,7 +8,7 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZCyclic.v 11238 2008-07-19 09:34:03Z herbelin $ i*)
+(*i $Id$ i*)
Require Export NZAxioms.
Require Import BigNumPrelude.
@@ -17,89 +17,79 @@ Require Import CyclicAxioms.
(** * From [CyclicType] to [NZAxiomsSig] *)
-(** A [Z/nZ] representation given by a module type [CyclicType]
- implements [NZAxiomsSig], e.g. the common properties between
- N and Z with no ordering. Notice that the [n] in [Z/nZ] is
+(** A [Z/nZ] representation given by a module type [CyclicType]
+ implements [NZAxiomsSig], e.g. the common properties between
+ N and Z with no ordering. Notice that the [n] in [Z/nZ] is
a power of 2.
*)
Module NZCyclicAxiomsMod (Import Cyclic : CyclicType) <: NZAxiomsSig.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
-Definition NZ := w.
+Definition t := w.
-Definition NZ_to_Z : NZ -> Z := znz_to_Z w_op.
-Definition Z_to_NZ : Z -> NZ := znz_of_Z w_op.
-Notation Local wB := (base w_op.(znz_digits)).
+Definition NZ_to_Z : t -> Z := znz_to_Z w_op.
+Definition Z_to_NZ : Z -> t := znz_of_Z w_op.
+Local Notation wB := (base w_op.(znz_digits)).
-Notation Local "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99).
+Local Notation "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99).
-Definition NZeq (n m : NZ) := [| n |] = [| m |].
-Definition NZ0 := w_op.(znz_0).
-Definition NZsucc := w_op.(znz_succ).
-Definition NZpred := w_op.(znz_pred).
-Definition NZadd := w_op.(znz_add).
-Definition NZsub := w_op.(znz_sub).
-Definition NZmul := w_op.(znz_mul).
+Definition eq (n m : t) := [| n |] = [| m |].
+Definition zero := w_op.(znz_0).
+Definition succ := w_op.(znz_succ).
+Definition pred := w_op.(znz_pred).
+Definition add := w_op.(znz_add).
+Definition sub := w_op.(znz_sub).
+Definition mul := w_op.(znz_mul).
-Theorem NZeq_equiv : equiv NZ NZeq.
-Proof.
-unfold equiv, reflexive, symmetric, transitive, NZeq; repeat split; intros; auto.
-now transitivity [| y |].
-Qed.
+Local Infix "==" := eq (at level 70).
+Local Notation "0" := zero.
+Local Notation S := succ.
+Local Notation P := pred.
+Local Infix "+" := add.
+Local Infix "-" := sub.
+Local Infix "*" := mul.
-Add Relation NZ NZeq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
-as NZeq_rel.
+Hint Rewrite w_spec.(spec_0) w_spec.(spec_succ) w_spec.(spec_pred)
+ w_spec.(spec_add) w_spec.(spec_mul) w_spec.(spec_sub) : w.
+Ltac wsimpl :=
+ unfold eq, zero, succ, pred, add, sub, mul; autorewrite with w.
+Ltac wcongruence := repeat red; intros; wsimpl; congruence.
-Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd.
+Instance eq_equiv : Equivalence eq.
Proof.
-unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_succ). now rewrite H.
+unfold eq. firstorder.
Qed.
-Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd.
+Instance succ_wd : Proper (eq ==> eq) succ.
Proof.
-unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_pred). now rewrite H.
+wcongruence.
Qed.
-Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd.
+Instance pred_wd : Proper (eq ==> eq) pred.
Proof.
-unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_add).
-now rewrite H1, H2.
+wcongruence.
Qed.
-Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd.
+Instance add_wd : Proper (eq ==> eq ==> eq) add.
Proof.
-unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_sub).
-now rewrite H1, H2.
+wcongruence.
Qed.
-Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd.
+Instance sub_wd : Proper (eq ==> eq ==> eq) sub.
Proof.
-unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_mul).
-now rewrite H1, H2.
+wcongruence.
Qed.
-Delimit Scope IntScope with Int.
-Bind Scope IntScope with NZ.
-Open Local Scope IntScope.
-Notation "x == y" := (NZeq x y) (at level 70) : IntScope.
-Notation "x ~= y" := (~ NZeq x y) (at level 70) : IntScope.
-Notation "0" := NZ0 : IntScope.
-Notation S x := (NZsucc x).
-Notation P x := (NZpred x).
-(*Notation "1" := (S 0) : IntScope.*)
-Notation "x + y" := (NZadd x y) : IntScope.
-Notation "x - y" := (NZsub x y) : IntScope.
-Notation "x * y" := (NZmul x y) : IntScope.
+Instance mul_wd : Proper (eq ==> eq ==> eq) mul.
+Proof.
+wcongruence.
+Qed.
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 Zlt; auto with zarith.
Qed.
Theorem gt_wB_0 : 0 < wB.
@@ -107,7 +97,7 @@ Proof.
pose proof gt_wB_1; auto with zarith.
Qed.
-Lemma NZsucc_mod_wB : forall n : Z, (n + 1) mod wB = ((n mod wB) + 1) mod wB.
+Lemma succ_mod_wB : forall n : Z, (n + 1) mod wB = ((n mod wB) + 1) mod wB.
Proof.
intro n.
pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zplus_mod.
@@ -115,7 +105,7 @@ reflexivity.
now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]].
Qed.
-Lemma NZpred_mod_wB : forall n : Z, (n - 1) mod wB = ((n mod wB) - 1) mod wB.
+Lemma pred_mod_wB : forall n : Z, (n - 1) mod wB = ((n mod wB) - 1) mod wB.
Proof.
intro n.
pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zminus_mod.
@@ -123,34 +113,32 @@ reflexivity.
now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]].
Qed.
-Lemma NZ_to_Z_mod : forall n : NZ, [| n |] mod wB = [| n |].
+Lemma NZ_to_Z_mod : forall n, [| n |] mod wB = [| n |].
Proof.
intro n; rewrite Zmod_small. reflexivity. apply w_spec.(spec_to_Z).
Qed.
-Theorem NZpred_succ : forall n : NZ, P (S n) == n.
+Theorem pred_succ : forall n, P (S n) == n.
Proof.
-intro n; unfold NZsucc, NZpred, NZeq. rewrite w_spec.(spec_pred), w_spec.(spec_succ).
-rewrite <- NZpred_mod_wB.
+intro n. wsimpl.
+rewrite <- pred_mod_wB.
replace ([| n |] + 1 - 1)%Z with [| n |] by auto with zarith. apply NZ_to_Z_mod.
Qed.
-Lemma Z_to_NZ_0 : Z_to_NZ 0%Z == 0%Int.
+Lemma Z_to_NZ_0 : Z_to_NZ 0%Z == 0.
Proof.
-unfold NZeq, NZ_to_Z, Z_to_NZ. rewrite znz_of_Z_correct.
-symmetry; apply w_spec.(spec_0).
+unfold NZ_to_Z, Z_to_NZ. wsimpl.
+rewrite znz_of_Z_correct; auto.
exact w_spec. split; [auto with zarith |apply gt_wB_0].
Qed.
Section Induction.
-Variable A : NZ -> Prop.
-Hypothesis A_wd : predicate_wd NZeq A.
+Variable A : t -> Prop.
+Hypothesis A_wd : Proper (eq ==> iff) A.
Hypothesis A0 : A 0.
-Hypothesis AS : forall n : NZ, A n <-> A (S n). (* Below, we use only -> direction *)
-
-Add Morphism A with signature NZeq ==> iff as A_morph.
-Proof. apply A_wd. Qed.
+Hypothesis AS : forall n, A n <-> A (S n).
+ (* Below, we use only -> direction *)
Let B (n : Z) := A (Z_to_NZ n).
@@ -163,8 +151,8 @@ Lemma BS : forall n : Z, 0 <= n -> n < wB - 1 -> B n -> B (n + 1).
Proof.
intros n H1 H2 H3.
unfold B in *. apply -> AS in H3.
-setoid_replace (Z_to_NZ (n + 1)) with (S (Z_to_NZ n)) using relation NZeq. assumption.
-unfold NZeq. rewrite w_spec.(spec_succ).
+setoid_replace (Z_to_NZ (n + 1)) with (S (Z_to_NZ n)). assumption.
+wsimpl.
unfold NZ_to_Z, Z_to_NZ.
do 2 (rewrite znz_of_Z_correct; [ | exact w_spec | auto with zarith]).
symmetry; apply Zmod_small; auto with zarith.
@@ -177,11 +165,11 @@ apply Zbounded_induction with wB.
apply B0. apply BS. assumption. assumption.
Qed.
-Theorem NZinduction : forall n : NZ, A n.
+Theorem bi_induction : forall n, A n.
Proof.
-intro n. setoid_replace n with (Z_to_NZ (NZ_to_Z n)) using relation NZeq.
+intro n. setoid_replace n with (Z_to_NZ (NZ_to_Z n)).
apply B_holds. apply w_spec.(spec_to_Z).
-unfold NZeq, NZ_to_Z, Z_to_NZ; rewrite znz_of_Z_correct.
+unfold eq, NZ_to_Z, Z_to_NZ; rewrite znz_of_Z_correct.
reflexivity.
exact w_spec.
apply w_spec.(spec_to_Z).
@@ -189,47 +177,40 @@ Qed.
End Induction.
-Theorem NZadd_0_l : forall n : NZ, 0 + n == n.
+Theorem add_0_l : forall n, 0 + n == n.
Proof.
-intro n; unfold NZadd, NZ0, NZeq. rewrite w_spec.(spec_add). rewrite w_spec.(spec_0).
+intro n. wsimpl.
rewrite Zplus_0_l. rewrite Zmod_small; [reflexivity | apply w_spec.(spec_to_Z)].
Qed.
-Theorem NZadd_succ_l : forall n m : NZ, (S n) + m == S (n + m).
+Theorem add_succ_l : forall n m, (S n) + m == S (n + m).
Proof.
-intros n m; unfold NZadd, NZsucc, NZeq. rewrite w_spec.(spec_add).
-do 2 rewrite w_spec.(spec_succ). rewrite w_spec.(spec_add).
-rewrite NZsucc_mod_wB. repeat rewrite Zplus_mod_idemp_l; try apply gt_wB_0.
+intros n m. wsimpl.
+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.
Qed.
-Theorem NZsub_0_r : forall n : NZ, n - 0 == n.
+Theorem sub_0_r : forall n, n - 0 == n.
Proof.
-intro n; unfold NZsub, NZ0, NZeq. rewrite w_spec.(spec_sub).
-rewrite w_spec.(spec_0). rewrite Zminus_0_r. apply NZ_to_Z_mod.
+intro n. wsimpl. rewrite Zminus_0_r. apply NZ_to_Z_mod.
Qed.
-Theorem NZsub_succ_r : forall n m : NZ, n - (S m) == P (n - m).
+Theorem sub_succ_r : forall n m, n - (S m) == P (n - m).
Proof.
-intros n m; unfold NZsub, NZsucc, NZpred, NZeq.
-rewrite w_spec.(spec_pred). do 2 rewrite w_spec.(spec_sub).
-rewrite w_spec.(spec_succ). rewrite Zminus_mod_idemp_r.
-rewrite Zminus_mod_idemp_l.
-now replace ([|n|] - ([|m|] + 1))%Z with ([|n|] - [|m|] - 1)%Z by auto with zarith.
+intros n m. wsimpl. rewrite Zminus_mod_idemp_r, Zminus_mod_idemp_l.
+now replace ([|n|] - ([|m|] + 1))%Z with ([|n|] - [|m|] - 1)%Z
+ by auto with zarith.
Qed.
-Theorem NZmul_0_l : forall n : NZ, 0 * n == 0.
+Theorem mul_0_l : forall n, 0 * n == 0.
Proof.
-intro n; unfold NZmul, NZ0, NZ, NZeq. rewrite w_spec.(spec_mul).
-rewrite w_spec.(spec_0). now rewrite Zmult_0_l.
+intro n. wsimpl. now rewrite Zmult_0_l.
Qed.
-Theorem NZmul_succ_l : forall n m : NZ, (S n) * m == n * m + m.
+Theorem mul_succ_l : forall n m, (S n) * m == n * m + m.
Proof.
-intros n m; unfold NZmul, NZsucc, NZadd, NZeq. rewrite w_spec.(spec_mul).
-rewrite w_spec.(spec_add), w_spec.(spec_mul), w_spec.(spec_succ).
-rewrite Zplus_mod_idemp_l, Zmult_mod_idemp_l.
+intros n m. wsimpl. rewrite Zplus_mod_idemp_l, Zmult_mod_idemp_l.
now rewrite Zmult_plus_distr_l, Zmult_1_l.
Qed.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
index 61d8d0fb..aa798e1c 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleAdd.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -17,7 +17,7 @@ Require Import BigNumPrelude.
Require Import DoubleType.
Require Import DoubleBase.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section DoubleAdd.
Variable w : Type.
@@ -36,10 +36,10 @@ Section DoubleAdd.
Definition ww_succ_c x :=
match x with
| W0 => C0 ww_1
- | WW xh xl =>
+ | WW xh xl =>
match w_succ_c xl with
| C0 l => C0 (WW xh l)
- | C1 l =>
+ | C1 l =>
match w_succ_c xh with
| C0 h => C0 (WW h w_0)
| C1 h => C1 W0
@@ -47,13 +47,13 @@ Section DoubleAdd.
end
end.
- Definition ww_succ x :=
+ Definition ww_succ x :=
match x with
| W0 => ww_1
| WW xh xl =>
match w_succ_c xl with
| C0 l => WW xh l
- | C1 l => w_W0 (w_succ xh)
+ | C1 l => w_W0 (w_succ xh)
end
end.
@@ -63,12 +63,12 @@ Section DoubleAdd.
| _, W0 => C0 x
| WW xh xl, WW yh yl =>
match w_add_c xl yl with
- | C0 l =>
+ | C0 l =>
match w_add_c xh yh with
| C0 h => C0 (WW h l)
| C1 h => C1 (w_WW h l)
- end
- | C1 l =>
+ end
+ | C1 l =>
match w_add_carry_c xh yh with
| C0 h => C0 (WW h l)
| C1 h => C1 (w_WW h l)
@@ -85,12 +85,12 @@ Section DoubleAdd.
| _, W0 => f0 x
| WW xh xl, WW yh yl =>
match w_add_c xl yl with
- | C0 l =>
+ | C0 l =>
match w_add_c xh yh with
| C0 h => f0 (WW h l)
| C1 h => f1 (w_WW h l)
- end
- | C1 l =>
+ end
+ | C1 l =>
match w_add_carry_c xh yh with
| C0 h => f0 (WW h l)
| C1 h => f1 (w_WW h l)
@@ -118,12 +118,12 @@ Section DoubleAdd.
| WW xh xl, W0 => ww_succ_c (WW xh xl)
| WW xh xl, WW yh yl =>
match w_add_carry_c xl yl with
- | C0 l =>
+ | C0 l =>
match w_add_c xh yh with
| C0 h => C0 (WW h l)
| C1 h => C1 (WW h l)
end
- | C1 l =>
+ | C1 l =>
match w_add_carry_c xh yh with
| C0 h => C0 (WW h l)
| C1 h => C1 (w_WW h l)
@@ -131,7 +131,7 @@ Section DoubleAdd.
end
end.
- Definition ww_add_carry x y :=
+ Definition ww_add_carry x y :=
match x, y with
| W0, W0 => ww_1
| W0, WW yh yl => ww_succ (WW yh yl)
@@ -146,7 +146,7 @@ Section DoubleAdd.
(*Section DoubleProof.*)
Variable w_digits : positive.
Variable w_to_Z : w -> Z.
-
+
Notation wB := (base w_digits).
Notation wwB := (base (ww_digits w_digits)).
@@ -157,11 +157,11 @@ Section DoubleAdd.
(interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[+[ c ]]" :=
+ (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[-[ c ]]" :=
+ (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
Variable spec_w_0 : [|w_0|] = 0.
@@ -172,7 +172,7 @@ Section DoubleAdd.
Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB.
Variable spec_w_succ_c : forall x, [+|w_succ_c x|] = [|x|] + 1.
Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
- Variable spec_w_add_carry_c :
+ Variable spec_w_add_carry_c :
forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1.
Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB.
Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB.
@@ -187,11 +187,11 @@ Section DoubleAdd.
rewrite <- Zplus_assoc;rewrite <- H;rewrite Zmult_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.
+ intro H1;unfold interp_carry in H1.
simpl;rewrite H1;rewrite spec_w_0;ring.
unfold interp_carry;simpl ww_to_Z;rewrite wwB_wBwB.
assert ([|xh|] = wB - 1). generalize (spec_to_Z xh)(spec_to_Z h);omega.
- rewrite H2;ring.
+ rewrite H2;ring.
Qed.
Lemma spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]].
@@ -222,12 +222,12 @@ Section DoubleAdd.
Proof.
destruct x as [ |xh xl];simpl;trivial.
apply spec_f0;trivial.
- destruct y as [ |yh yl];simpl.
+ destruct y as [ |yh yl];simpl.
apply spec_f0;simpl;rewrite Zplus_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 *.
+ 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.
@@ -236,12 +236,12 @@ Section DoubleAdd.
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_f1. simpl;rewrite spec_w_WW;rewrite wwB_wBwB.
- rewrite Zplus_assoc; rewrite Zpower_2; rewrite <- Zmult_plus_distr_l.
+ 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.
Qed.
-
+
End Cont.
Lemma spec_ww_add_carry_c :
@@ -251,16 +251,16 @@ Section DoubleAdd.
exact (spec_ww_succ_c y).
destruct y as [ |yh yl];simpl.
rewrite Zplus_0_r;exact (spec_ww_succ_c (WW xh xl)).
- replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1)
+ 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)
+ 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.
- 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.
+ 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.
Qed.
@@ -287,9 +287,9 @@ Section DoubleAdd.
rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
destruct y as [ |yh yl].
change [[W0]] with 0;rewrite Zplus_0_r.
- rewrite Zmod_small;trivial.
+ rewrite Zmod_small;trivial.
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)).
- simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]))
+ simpl. 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];
unfold interp_carry;intros H;simpl;rewrite <- H.
@@ -305,14 +305,14 @@ Section DoubleAdd.
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)).
- simpl;replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1)
+ 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)
+ 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(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial.
- Qed.
+ Qed.
(* End DoubleProof. *)
End DoubleAdd.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
index 952516ac..88c34915 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleBase.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -16,7 +16,7 @@ Require Import ZArith.
Require Import BigNumPrelude.
Require Import DoubleType.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section DoubleBase.
Variable w : Type.
@@ -29,8 +29,8 @@ Section DoubleBase.
Variable w_zdigits: w.
Variable w_add: w -> w -> zn2z w.
Variable w_to_Z : w -> Z.
- Variable w_compare : w -> w -> comparison.
-
+ Variable w_compare : w -> w -> comparison.
+
Definition ww_digits := xO w_digits.
Definition ww_zdigits := w_add w_zdigits w_zdigits.
@@ -46,7 +46,7 @@ Section DoubleBase.
| W0, W0 => W0
| _, _ => WW xh xl
end.
-
+
Definition ww_W0 h : zn2z (zn2z w) :=
match h with
| W0 => W0
@@ -58,10 +58,10 @@ Section DoubleBase.
| W0 => W0
| _ => WW W0 l
end.
-
- Definition double_WW (n:nat) :=
- match n return word w n -> word w n -> word w (S n) with
- | O => w_WW
+
+ Definition double_WW (n:nat) :=
+ match n return word w n -> word w n -> word w (S n) with
+ | O => w_WW
| S n =>
fun (h l : zn2z (word w n)) =>
match h, l with
@@ -70,8 +70,8 @@ Section DoubleBase.
end
end.
- Fixpoint double_digits (n:nat) : positive :=
- match n with
+ Fixpoint double_digits (n:nat) : positive :=
+ match n with
| O => w_digits
| S n => xO (double_digits n)
end.
@@ -80,7 +80,7 @@ Section DoubleBase.
Fixpoint double_to_Z (n:nat) : word w n -> Z :=
match n return word w n -> Z with
- | O => w_to_Z
+ | O => w_to_Z
| S n => zn2z_to_Z (double_wB n) (double_to_Z n)
end.
@@ -98,21 +98,21 @@ Section DoubleBase.
end.
Definition double_0 n : word w n :=
- match n return word w n with
+ match n return word w n with
| O => w_0
| S _ => W0
end.
-
+
Definition double_split (n:nat) (x:zn2z (word w n)) :=
- match x with
- | W0 =>
- match n return word w n * word w n with
+ match x with
+ | W0 =>
+ match n return word w n * word w n with
| O => (w_0,w_0)
| S _ => (W0, W0)
end
| WW h l => (h,l)
end.
-
+
Definition ww_compare x y :=
match x, y with
| W0, W0 => Eq
@@ -148,15 +148,15 @@ Section DoubleBase.
end
end.
-
+
Section DoubleProof.
Notation wB := (base w_digits).
Notation wwB := (base ww_digits).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[[ x ]]" := (ww_to_Z x) (at level 0, x at level 99).
- Notation "[+[ c ]]" :=
+ Notation "[+[ c ]]" :=
(interp_carry 1 wwB ww_to_Z c) (at level 0, x at level 99).
- Notation "[-[ c ]]" :=
+ Notation "[-[ c ]]" :=
(interp_carry (-1) wwB ww_to_Z c) (at level 0, x at level 99).
Notation "[! n | x !]" := (double_to_Z n x) (at level 0, x at level 99).
@@ -188,7 +188,7 @@ Section DoubleBase.
Proof. simpl;rewrite spec_w_Bm1;rewrite wwB_wBwB;ring. Qed.
Lemma lt_0_wB : 0 < wB.
- Proof.
+ Proof.
unfold base;apply Zpower_gt_0. unfold Zlt;reflexivity.
unfold Zle;intros H;discriminate H.
Qed.
@@ -197,25 +197,25 @@ Section DoubleBase.
Proof. rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_lt_0_compat;apply lt_0_wB. Qed.
Lemma wB_pos: 1 < wB.
- Proof.
+ 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.
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.
+
+ 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.
+ apply Zlt_le_weak;trivial.
Qed.
Theorem wB_div_2: 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
+ 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.
@@ -228,7 +228,7 @@ Section DoubleBase.
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
+ 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.
pattern wB at 1; rewrite <- wB_div_2; auto.
@@ -236,11 +236,11 @@ Section DoubleBase.
repeat (rewrite (Zmult_comm 2); rewrite Z_div_mult); auto with zarith.
Qed.
- Lemma mod_wwB : forall z x,
+ Lemma mod_wwB : forall z x,
(z*wB + [|x|]) mod wwB = (z mod wB)*wB + [|x|].
Proof.
intros z x.
- rewrite Zplus_mod.
+ rewrite Zplus_mod.
pattern wwB at 1;rewrite wwB_wBwB; rewrite Zpower_2.
rewrite Zmult_mod_distr_r;try apply lt_0_wB.
rewrite (Zmod_small [|x|]).
@@ -260,8 +260,8 @@ Section DoubleBase.
destruct (spec_to_Z x);trivial.
Qed.
- Lemma wB_div_plus : forall x y p,
- 0 <= p ->
+ Lemma wB_div_plus : forall x y p,
+ 0 <= p ->
([|x|]*wB + [|y|]) / 2^(Zpos w_digits + p) = [|x|] / 2^p.
Proof.
clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
@@ -277,7 +277,7 @@ Section DoubleBase.
assert (0 < Zpos w_digits). compute;reflexivity.
unfold ww_digits;rewrite Zpos_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.
@@ -298,7 +298,7 @@ Section DoubleBase.
Proof.
intros n;unfold double_wB;simpl.
unfold base;rewrite (Zpos_xO (double_digits n)).
- replace (2 * Zpos (double_digits n)) with
+ replace (2 * Zpos (double_digits n)) with
(Zpos (double_digits n) + Zpos (double_digits n)).
symmetry; apply Zpower_exp;intro;discriminate.
ring.
@@ -327,7 +327,7 @@ Section DoubleBase.
unfold base; auto with zarith.
Qed.
- Lemma spec_double_to_Z :
+ Lemma spec_double_to_Z :
forall n (x:word w n), 0 <= [!n | x!] < double_wB n.
Proof.
clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
@@ -347,7 +347,7 @@ Section DoubleBase.
Qed.
Lemma spec_get_low:
- forall n x,
+ forall n x,
[!n | x!] < wB -> [|get_low n x|] = [!n | x!].
Proof.
clear spec_w_1 spec_w_Bm1.
@@ -380,19 +380,19 @@ Section DoubleBase.
Qed.
Lemma spec_extend_aux : forall n x, [!S n|extend_aux n x!] = [[x]].
- Proof. induction n;simpl;trivial. Qed.
+ Proof. induction n;simpl;trivial. Qed.
Lemma spec_extend : forall n x, [!S n|extend n x!] = [|x|].
- Proof.
+ Proof.
intros n x;assert (H:= spec_w_0W x);unfold extend.
- destruct (w_0W x);simpl;trivial.
+ destruct (w_0W x);simpl;trivial.
rewrite <- H;exact (spec_extend_aux n (WW w0 w1)).
Qed.
Lemma spec_double_0 : forall n, [!n|double_0 n!] = 0.
Proof. destruct n;trivial. Qed.
- Lemma spec_double_split : forall n x,
+ Lemma spec_double_split : forall n x,
let (h,l) := double_split n x in
[!S n|x!] = [!n|h!] * double_wB n + [!n|l!].
Proof.
@@ -401,9 +401,9 @@ Section DoubleBase.
rewrite spec_w_0;trivial.
Qed.
- Lemma wB_lex_inv: forall a b c d,
- a < c ->
- a * wB + [|b|] < c * wB + [|d|].
+ Lemma wB_lex_inv: forall a b c d,
+ a < c ->
+ a * wB + [|b|] < c * wB + [|d|].
Proof.
intros a b c d H1; apply beta_lex_inv with (1 := H1); auto.
Qed.
@@ -420,7 +420,7 @@ Section DoubleBase.
intros H;rewrite spec_w_0 in H.
rewrite <- H;simpl;rewrite <- spec_w_0;apply spec_w_compare.
change 0 with (0*wB+0);pattern 0 at 2;rewrite <- spec_w_0.
- apply wB_lex_inv;trivial.
+ apply wB_lex_inv;trivial.
absurd (0 <= [|yh|]). apply Zgt_not_le;trivial.
destruct (spec_to_Z yh);trivial.
generalize (spec_w_compare xh w_0);destruct (w_compare xh w_0);
@@ -429,8 +429,8 @@ Section DoubleBase.
absurd (0 <= [|xh|]). apply Zgt_not_le;apply Zlt_gt;trivial.
destruct (spec_to_Z xh);trivial.
apply Zlt_gt;change 0 with (0*wB+0);pattern 0 at 2;rewrite <- spec_w_0.
- apply wB_lex_inv;apply Zgt_lt;trivial.
-
+ apply wB_lex_inv;apply Zgt_lt;trivial.
+
generalize (spec_w_compare xh yh);destruct (w_compare xh yh);intros H.
rewrite H;generalize (spec_w_compare xl yl);destruct (w_compare xl yl);
intros H1;[rewrite H1|apply Zplus_lt_compat_l|apply Zplus_gt_compat_l];
@@ -439,7 +439,7 @@ Section DoubleBase.
apply Zlt_gt;apply wB_lex_inv;apply Zgt_lt;trivial.
Qed.
-
+
End DoubleProof.
End DoubleBase.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
index cca32a59..eea29e7c 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleCyclic.v 11012 2008-05-28 16:34:43Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -22,10 +22,10 @@ Require Import DoubleMul.
Require Import DoubleSqrt.
Require Import DoubleLift.
Require Import DoubleDivn1.
-Require Import DoubleDiv.
+Require Import DoubleDiv.
Require Import CyclicAxioms.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section Z_2nZ.
@@ -80,7 +80,7 @@ Section Z_2nZ.
Let w_gcd_gt := w_op.(znz_gcd_gt).
Let w_gcd := w_op.(znz_gcd).
- Let w_add_mul_div := w_op.(znz_add_mul_div).
+ Let w_add_mul_div := w_op.(znz_add_mul_div).
Let w_pos_mod := w_op.(znz_pos_mod).
@@ -93,7 +93,7 @@ Section Z_2nZ.
Let wB := base w_digits.
Let w_Bm2 := w_pred w_Bm1.
-
+
Let ww_1 := ww_1 w_0 w_1.
Let ww_Bm1 := ww_Bm1 w_Bm1.
@@ -112,16 +112,16 @@ Section Z_2nZ.
Let ww_of_pos p :=
match w_of_pos p with
| (N0, l) => (N0, WW w_0 l)
- | (Npos ph,l) =>
+ | (Npos ph,l) =>
let (n,h) := w_of_pos ph in (n, w_WW h l)
end.
Let head0 :=
- Eval lazy beta delta [ww_head0] in
+ Eval lazy beta delta [ww_head0] in
ww_head0 w_0 w_0W w_compare w_head0 w_add2 w_zdigits _ww_zdigits.
Let tail0 :=
- Eval lazy beta delta [ww_tail0] in
+ Eval lazy beta delta [ww_tail0] in
ww_tail0 w_0 w_0W w_compare w_tail0 w_add2 w_zdigits _ww_zdigits.
Let ww_WW := Eval lazy beta delta [ww_WW] in (@ww_WW w).
@@ -132,7 +132,7 @@ Section Z_2nZ.
Let compare :=
Eval lazy beta delta[ww_compare] in ww_compare w_0 w_compare.
- Let eq0 (x:zn2z w) :=
+ Let eq0 (x:zn2z w) :=
match x with
| W0 => true
| _ => false
@@ -147,7 +147,7 @@ Section Z_2nZ.
Let opp_carry :=
Eval lazy beta delta [ww_opp_carry] in ww_opp_carry w_WW ww_Bm1 w_opp_carry.
-
+
(* ** Additions ** *)
Let succ_c :=
@@ -157,16 +157,16 @@ Section Z_2nZ.
Eval lazy beta delta [ww_add_c] in ww_add_c w_WW w_add_c w_add_carry_c.
Let add_carry_c :=
- Eval lazy beta iota delta [ww_add_carry_c ww_succ_c] in
+ Eval lazy beta iota delta [ww_add_carry_c ww_succ_c] in
ww_add_carry_c w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c.
- Let succ :=
+ Let succ :=
Eval lazy beta delta [ww_succ] in ww_succ w_W0 ww_1 w_succ_c w_succ.
Let add :=
Eval lazy beta delta [ww_add] in ww_add w_add_c w_add w_add_carry.
- Let add_carry :=
+ Let add_carry :=
Eval lazy beta iota delta [ww_add_carry ww_succ] in
ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ w_add w_add_carry.
@@ -174,9 +174,9 @@ Section Z_2nZ.
Let pred_c :=
Eval lazy beta delta [ww_pred_c] in ww_pred_c w_Bm1 w_WW ww_Bm1 w_pred_c.
-
+
Let sub_c :=
- Eval lazy beta iota delta [ww_sub_c ww_opp_c] in
+ Eval lazy beta iota delta [ww_sub_c ww_opp_c] in
ww_sub_c w_0 w_WW w_opp_c w_opp_carry w_sub_c w_sub_carry_c.
Let sub_carry_c :=
@@ -186,8 +186,8 @@ Section Z_2nZ.
Let pred :=
Eval lazy beta delta [ww_pred] in ww_pred w_Bm1 w_WW ww_Bm1 w_pred_c w_pred.
- Let sub :=
- Eval lazy beta iota delta [ww_sub ww_opp] in
+ Let sub :=
+ Eval lazy beta iota delta [ww_sub ww_opp] in
ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry.
Let sub_carry :=
@@ -204,7 +204,7 @@ Section Z_2nZ.
Let karatsuba_c :=
Eval lazy beta iota delta [ww_karatsuba_c double_mul_c kara_prod] in
- ww_karatsuba_c w_0 w_1 w_WW w_W0 w_compare w_add w_sub w_mul_c
+ ww_karatsuba_c w_0 w_1 w_WW w_W0 w_compare w_add w_sub w_mul_c
add_c add add_carry sub_c sub.
Let mul :=
@@ -219,7 +219,7 @@ Section Z_2nZ.
Let div32 :=
Eval lazy beta iota delta [w_div32] in
- w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c
+ w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c
w_add w_add_carry w_pred w_sub w_mul_c w_div21 sub_c.
Let div21 :=
@@ -234,40 +234,40 @@ Section Z_2nZ.
Let div_gt :=
Eval lazy beta delta [ww_div_gt] in
- ww_div_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp
+ ww_div_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp
w_opp_carry w_sub_c w_sub w_sub_carry
w_div_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits.
Let div :=
Eval lazy beta delta [ww_div] in ww_div ww_1 compare div_gt.
-
+
Let mod_gt :=
Eval lazy beta delta [ww_mod_gt] in
ww_mod_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry
w_mod_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits add_mul_div w_zdigits.
- Let mod_ :=
+ Let mod_ :=
Eval lazy beta delta [ww_mod] in ww_mod compare mod_gt.
- Let pos_mod :=
- Eval lazy beta delta [ww_pos_mod] in
+ Let pos_mod :=
+ Eval lazy beta delta [ww_pos_mod] in
ww_pos_mod w_0 w_zdigits w_WW w_pos_mod compare w_0W low sub _ww_zdigits.
- Let is_even :=
+ Let is_even :=
Eval lazy beta delta [ww_is_even] in ww_is_even w_is_even.
- Let sqrt2 :=
+ Let sqrt2 :=
Eval lazy beta delta [ww_sqrt2] in
ww_sqrt2 w_is_even w_compare w_0 w_1 w_Bm1 w_0W w_sub w_square_c
w_div21 w_add_mul_div w_zdigits w_add_c w_sqrt2 w_pred pred_c
pred add_c add sub_c add_mul_div.
- Let sqrt :=
+ Let sqrt :=
Eval lazy beta delta [ww_sqrt] in
ww_sqrt w_is_even w_0 w_sub w_add_mul_div w_zdigits
_ww_zdigits w_sqrt2 pred add_mul_div head0 compare low.
- Let gcd_gt_fix :=
+ Let gcd_gt_fix :=
Eval cbv beta delta [ww_gcd_gt_aux ww_gcd_gt_body] in
ww_gcd_gt_aux w_0 w_WW w_0W w_compare w_opp_c w_opp w_opp_carry
w_sub_c w_sub w_sub_carry w_gcd_gt
@@ -278,7 +278,7 @@ Section Z_2nZ.
Eval lazy beta delta [gcd_cont] in gcd_cont ww_1 w_1 w_compare.
Let gcd_gt :=
- Eval lazy beta delta [ww_gcd_gt] in
+ Eval lazy beta delta [ww_gcd_gt] in
ww_gcd_gt w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont.
Let gcd :=
@@ -286,18 +286,18 @@ Section Z_2nZ.
ww_gcd compare w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont.
(* ** Record of operators on 2 words *)
-
- Definition mk_zn2z_op :=
+
+ Definition mk_zn2z_op :=
mk_znz_op _ww_digits _ww_zdigits
to_Z ww_of_pos head0 tail0
W0 ww_1 ww_Bm1
compare eq0
opp_c opp opp_carry
- succ_c add_c add_carry_c
- succ add add_carry
- pred_c sub_c sub_carry_c
+ succ_c add_c add_carry_c
+ succ add add_carry
+ pred_c sub_c sub_carry_c
pred sub sub_carry
- mul_c mul square_c
+ mul_c mul square_c
div21 div_gt div
mod_gt mod_
gcd_gt gcd
@@ -307,17 +307,17 @@ Section Z_2nZ.
sqrt2
sqrt.
- Definition mk_zn2z_op_karatsuba :=
+ Definition mk_zn2z_op_karatsuba :=
mk_znz_op _ww_digits _ww_zdigits
to_Z ww_of_pos head0 tail0
W0 ww_1 ww_Bm1
compare eq0
opp_c opp opp_carry
- succ_c add_c add_carry_c
- succ add add_carry
- pred_c sub_c sub_carry_c
+ succ_c add_c add_carry_c
+ succ add add_carry
+ pred_c sub_c sub_carry_c
pred sub sub_carry
- karatsuba_c mul square_c
+ karatsuba_c mul square_c
div21 div_gt div
mod_gt mod_
gcd_gt gcd
@@ -330,7 +330,7 @@ Section Z_2nZ.
(* Proof *)
Variable op_spec : znz_spec w_op.
- Hint Resolve
+ Hint Resolve
(spec_to_Z op_spec)
(spec_of_pos op_spec)
(spec_0 op_spec)
@@ -358,13 +358,13 @@ Section Z_2nZ.
(spec_square_c op_spec)
(spec_div21 op_spec)
(spec_div_gt op_spec)
- (spec_div op_spec)
+ (spec_div op_spec)
(spec_mod_gt op_spec)
- (spec_mod op_spec)
+ (spec_mod op_spec)
(spec_gcd_gt op_spec)
- (spec_gcd op_spec)
- (spec_head0 op_spec)
- (spec_tail0 op_spec)
+ (spec_gcd op_spec)
+ (spec_head0 op_spec)
+ (spec_tail0 op_spec)
(spec_add_mul_div op_spec)
(spec_pos_mod)
(spec_is_even)
@@ -417,20 +417,20 @@ Section Z_2nZ.
Let spec_ww_Bm1 : [|ww_Bm1|] = wwB - 1.
Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed.
- Let spec_ww_compare :
+ Let spec_ww_compare :
forall x y,
match compare x y with
| Eq => [|x|] = [|y|]
| Lt => [|x|] < [|y|]
| Gt => [|x|] > [|y|]
end.
- Proof.
- refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);auto.
- exact (spec_compare op_spec).
+ Proof.
+ refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);auto.
+ exact (spec_compare op_spec).
Qed.
Let spec_ww_eq0 : forall x, eq0 x = true -> [|x|] = 0.
- Proof. destruct x;simpl;intros;trivial;discriminate. Qed.
+ Proof. destruct x;simpl;intros;trivial;discriminate. Qed.
Let spec_ww_opp_c : forall x, [-|opp_c x|] = -[|x|].
Proof.
@@ -440,7 +440,7 @@ Section Z_2nZ.
Let spec_ww_opp : forall x, [|opp x|] = (-[|x|]) mod wwB.
Proof.
- refine(spec_ww_opp w_0 w_0 W0 w_opp_c w_opp_carry w_opp
+ refine(spec_ww_opp w_0 w_0 W0 w_opp_c w_opp_carry w_opp
w_digits w_to_Z _ _ _ _ _);
auto.
Qed.
@@ -480,25 +480,25 @@ Section Z_2nZ.
Let spec_ww_add_carry : forall x y, [|add_carry x y|]=([|x|]+[|y|]+1)mod wwB.
Proof.
- refine (spec_ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ
+ refine (spec_ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ
w_add w_add_carry w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto.
Qed.
Let spec_ww_pred_c : forall x, [-|pred_c x|] = [|x|] - 1.
Proof.
- refine (spec_ww_pred_c w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_digits w_to_Z
+ refine (spec_ww_pred_c w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_digits w_to_Z
_ _ _ _ _);wwauto.
Qed.
Let spec_ww_sub_c : forall x y, [-|sub_c x y|] = [|x|] - [|y|].
Proof.
- refine (spec_ww_sub_c w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c
+ refine (spec_ww_sub_c w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c
w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _);wwauto.
Qed.
Let spec_ww_sub_carry_c : forall x y, [-|sub_carry_c x y|] = [|x|]-[|y|]-1.
Proof.
- refine (spec_ww_sub_carry_c w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c
+ refine (spec_ww_sub_carry_c w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c
w_sub_c w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto.
Qed.
@@ -533,17 +533,17 @@ Section Z_2nZ.
_ _ _ _ _ _ _ _ _ _ _ _); wwauto.
unfold w_digits; apply spec_more_than_1_digit; auto.
exact (spec_compare op_spec).
- Qed.
+ Qed.
Let spec_ww_mul : forall x y, [|mul x y|] = ([|x|] * [|y|]) mod wwB.
Proof.
refine (spec_ww_mul w_W0 w_add w_mul_c w_mul add w_digits w_to_Z _ _ _ _ _);
- wwauto.
+ wwauto.
Qed.
Let spec_ww_square_c : forall x, [[square_c x]] = [|x|] * [|x|].
Proof.
- refine (spec_ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add
+ refine (spec_ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add
add_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _);wwauto.
Qed.
@@ -574,7 +574,7 @@ Section Z_2nZ.
0 <= [|r|] < [|b|].
Proof.
refine (spec_ww_div21 w_0 w_0W div32 ww_1 compare sub w_digits w_to_Z
- _ _ _ _ _ _ _);wwauto.
+ _ _ _ _ _ _ _);wwauto.
Qed.
Let spec_add2: forall x y,
@@ -602,7 +602,7 @@ Section Z_2nZ.
unfold wB, base; auto with zarith.
Qed.
- Let spec_ww_digits:
+ Let spec_ww_digits:
[|_ww_zdigits|] = Zpos (xO w_digits).
Proof.
unfold w_to_Z, _ww_zdigits.
@@ -615,7 +615,7 @@ Section Z_2nZ.
Let spec_ww_head00 : forall x, [|x|] = 0 -> [|head0 x|] = Zpos _ww_digits.
Proof.
- refine (spec_ww_head00 w_0 w_0W
+ 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.
exact (spec_compare op_spec).
@@ -626,8 +626,8 @@ Section Z_2nZ.
Let spec_ww_head0 : forall x, 0 < [|x|] ->
wwB/ 2 <= 2 ^ [|head0 x|] * [|x|] < wwB.
Proof.
- refine (spec_ww_head0 w_0 w_0W w_compare w_head0
- w_add2 w_zdigits _ww_zdigits
+ refine (spec_ww_head0 w_0 w_0W w_compare w_head0
+ w_add2 w_zdigits _ww_zdigits
w_to_Z _ _ _ _ _ _ _);wwauto.
exact (spec_compare op_spec).
exact (spec_zdigits op_spec).
@@ -635,7 +635,7 @@ Section Z_2nZ.
Let spec_ww_tail00 : forall x, [|x|] = 0 -> [|tail0 x|] = Zpos _ww_digits.
Proof.
- refine (spec_ww_tail00 w_0 w_0W
+ 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.
exact (spec_compare op_spec).
@@ -647,7 +647,7 @@ Section Z_2nZ.
Let spec_ww_tail0 : forall x, 0 < [|x|] ->
exists y, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail0 x|].
Proof.
- refine (spec_ww_tail0 (w_digits := w_digits) w_0 w_0W w_compare w_tail0
+ refine (spec_ww_tail0 (w_digits := w_digits) w_0 w_0W w_compare w_tail0
w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ _ _ _ _);wwauto.
exact (spec_compare op_spec).
exact (spec_zdigits op_spec).
@@ -659,19 +659,19 @@ Section Z_2nZ.
([|x|] * (2 ^ [|p|]) +
[|y|] / (2 ^ ((Zpos _ww_digits) - [|p|]))) mod wwB.
Proof.
- refine (@spec_ww_add_mul_div w w_0 w_WW w_W0 w_0W compare w_add_mul_div
+ refine (@spec_ww_add_mul_div w w_0 w_WW w_W0 w_0W compare w_add_mul_div
sub w_digits w_zdigits low w_to_Z
_ _ _ _ _ _ _ _ _ _ _);wwauto.
exact (spec_zdigits op_spec).
Qed.
- Let spec_ww_div_gt : forall a b,
+ Let spec_ww_div_gt : forall a b,
[|a|] > [|b|] -> 0 < [|b|] ->
let (q,r) := div_gt a b in
[|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|].
Proof.
-refine
-(@spec_ww_div_gt w w_digits w_0 w_WW w_0W w_compare w_eq0
+refine
+(@spec_ww_div_gt w w_digits w_0 w_WW w_0W w_compare w_eq0
w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt
w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
@@ -707,14 +707,14 @@ refine
refine (spec_ww_div w_digits ww_1 compare div_gt w_to_Z _ _ _ _);auto.
Qed.
- Let spec_ww_mod_gt : forall a b,
+ Let spec_ww_mod_gt : forall a b,
[|a|] > [|b|] -> 0 < [|b|] ->
[|mod_gt a b|] = [|a|] mod [|b|].
Proof.
- refine (@spec_ww_mod_gt w w_digits w_0 w_WW w_0W w_compare w_eq0
+ refine (@spec_ww_mod_gt w w_digits w_0 w_WW w_0W w_compare w_eq0
w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt w_mod_gt
- w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div
- w_zdigits w_to_Z
+ w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div
+ w_zdigits w_to_Z
_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
exact (spec_compare op_spec).
exact (spec_div_gt op_spec).
@@ -731,12 +731,12 @@ refine
Let spec_ww_gcd_gt : forall a b, [|a|] > [|b|] ->
Zis_gcd [|a|] [|b|] [|gcd_gt a b|].
Proof.
- refine (@spec_ww_gcd_gt w w_digits W0 w_to_Z _
+ refine (@spec_ww_gcd_gt w w_digits W0 w_to_Z _
w_0 w_0 w_eq0 w_gcd_gt _ww_digits
_ gcd_gt_fix _ _ _ _ gcd_cont _);auto.
refine (@spec_ww_gcd_gt_aux w w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp
w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0
- w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
+ w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
exact (spec_compare op_spec).
exact (spec_div21 op_spec).
@@ -753,7 +753,7 @@ refine
_ww_digits _ gcd_gt_fix _ _ _ _ gcd_cont _);auto.
refine (@spec_ww_gcd_gt_aux w w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp
w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0
- w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
+ w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z
_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto.
exact (spec_compare op_spec).
exact (spec_div21 op_spec).
@@ -798,7 +798,7 @@ refine
Let spec_ww_sqrt : forall x,
[|sqrt x|] ^ 2 <= [|x|] < ([|sqrt x|] + 1) ^ 2.
Proof.
- refine (@spec_ww_sqrt w w_is_even w_0 w_1 w_Bm1
+ refine (@spec_ww_sqrt w w_is_even w_0 w_1 w_Bm1
w_sub w_add_mul_div w_digits w_zdigits _ww_zdigits
w_sqrt2 pred add_mul_div head0 compare
_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto.
@@ -814,7 +814,7 @@ refine
apply mk_znz_spec;auto.
exact spec_ww_add_mul_div.
- refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW
+ refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW
w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z
_ _ _ _ _ _ _ _ _ _ _ _);wwauto.
exact (spec_pos_mod op_spec).
@@ -828,7 +828,7 @@ refine
Proof.
apply mk_znz_spec;auto.
exact spec_ww_add_mul_div.
- refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW
+ refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW
w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z
_ _ _ _ _ _ _ _ _ _ _ _);wwauto.
exact (spec_pos_mod op_spec).
@@ -838,10 +838,10 @@ refine
rewrite <- Zpos_xO; exact spec_ww_digits.
Qed.
-End Z_2nZ.
-
+End Z_2nZ.
+
Section MulAdd.
-
+
Variable w: Type.
Variable op: znz_op w.
Variable sop: znz_spec op.
@@ -870,7 +870,7 @@ Section MulAdd.
End MulAdd.
-(** Modular versions of DoubleCyclic *)
+(** Modular versions of DoubleCyclic *)
Module DoubleCyclic (C:CyclicType) <: CyclicType.
Definition w := zn2z C.w.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
index 075aef59..9204b4e0 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleDiv.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -20,7 +20,7 @@ Require Import DoubleDivn1.
Require Import DoubleAdd.
Require Import DoubleSub.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Ltac zarith := auto with zarith.
@@ -41,13 +41,13 @@ Section POS_MOD.
Variable ww_zdigits : zn2z w.
- Definition ww_pos_mod p x :=
+ Definition ww_pos_mod p x :=
let zdigits := w_0W w_zdigits in
match x with
| W0 => W0
| WW xh xl =>
match ww_compare p zdigits with
- | Eq => w_WW w_0 xl
+ | Eq => w_WW w_0 xl
| Lt => w_WW w_0 (w_pos_mod (low p) xl)
| Gt =>
match ww_compare p ww_zdigits with
@@ -87,7 +87,7 @@ Section POS_MOD.
| Lt => [[x]] < [[y]]
| Gt => [[x]] > [[y]]
end.
- Variable spec_ww_sub: forall x y,
+ Variable spec_ww_sub: forall x y,
[[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits.
@@ -106,7 +106,7 @@ Section POS_MOD.
unfold ww_pos_mod; case w1.
simpl; rewrite Zmod_small; split; auto with zarith.
intros xh xl; generalize (spec_ww_compare p (w_0W w_zdigits));
- case ww_compare;
+ case ww_compare;
rewrite spec_w_0W; rewrite spec_zdigits; fold wB;
intros H1.
rewrite H1; simpl ww_to_Z.
@@ -135,13 +135,13 @@ Section POS_MOD.
autorewrite with w_rewrite rm10.
rewrite Zmod_mod; auto with zarith.
generalize (spec_ww_compare p ww_zdigits);
- case ww_compare; rewrite spec_ww_zdigits;
+ case ww_compare; rewrite spec_ww_zdigits;
rewrite spec_zdigits; intros H2.
replace (2^[[p]]) with wwB.
rewrite Zmod_small; auto with zarith.
unfold base; rewrite H2.
rewrite spec_ww_digits; auto.
- assert (HH0: [|low (ww_sub p (w_0W w_zdigits))|] =
+ assert (HH0: [|low (ww_sub p (w_0W w_zdigits))|] =
[[p]] - Zpos w_digits).
rewrite spec_low.
rewrite spec_ww_sub.
@@ -152,11 +152,11 @@ generalize (spec_ww_compare p ww_zdigits);
apply Zlt_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;
+ rewrite spec_ww_digits;
apply f_equal with (f := Zpower 2); rewrite Zpos_xO; auto with zarith.
simpl ww_to_Z; autorewrite with w_rewrite.
rewrite spec_pos_mod; rewrite HH0.
- pattern [|xh|] at 2;
+ 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.
@@ -196,7 +196,7 @@ generalize (spec_ww_compare p ww_zdigits);
split; auto with zarith.
rewrite Zpos_xO; auto with zarith.
Qed.
-
+
End POS_MOD.
Section DoubleDiv32.
@@ -222,24 +222,24 @@ Section DoubleDiv32.
match w_compare a1 b1 with
| Lt =>
let (q,r) := w_div21 a1 a2 b1 in
- match ww_sub_c (w_WW r a3) (w_mul_c q b2) with
+ match ww_sub_c (w_WW r a3) (w_mul_c q b2) with
| C0 r1 => (q,r1)
| C1 r1 =>
let q := w_pred q in
- ww_add_c_cont w_WW w_add_c w_add_carry_c
+ ww_add_c_cont w_WW w_add_c w_add_carry_c
(fun r2=>(w_pred q, ww_add w_add_c w_add w_add_carry r2 (WW b1 b2)))
(fun r2 => (q,r2))
r1 (WW b1 b2)
end
| Eq =>
- ww_add_c_cont w_WW w_add_c w_add_carry_c
+ ww_add_c_cont w_WW w_add_c w_add_carry_c
(fun r => (w_Bm2, ww_add w_add_c w_add w_add_carry r (WW b1 b2)))
(fun r => (w_Bm1,r))
(WW (w_sub a2 b2) a3) (WW b1 b2)
| Gt => (w_0, W0) (* cas absurde *)
end.
- (* Proof *)
+ (* Proof *)
Variable w_digits : positive.
Variable w_to_Z : w -> Z.
@@ -253,8 +253,8 @@ Section DoubleDiv32.
(interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[-[ c ]]" :=
+ (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
@@ -273,7 +273,7 @@ Section DoubleDiv32.
| Gt => [|x|] > [|y|]
end.
Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|].
- Variable spec_w_add_carry_c :
+ Variable spec_w_add_carry_c :
forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1.
Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB.
@@ -315,8 +315,8 @@ Section DoubleDiv32.
wB/2 <= [|b1|] ->
[[WW a1 a2]] < [[WW b1 b2]] ->
let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- [|a1|] * wwB + [|a2|] * wB + [|a3|] =
- [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
+ [|a1|] * wwB + [|a2|] * wB + [|a3|] =
+ [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
0 <= [[r]] < [|b1|] * wB + [|b2|].
Proof.
intros a1 a2 a3 b1 b2 Hle Hlt.
@@ -327,17 +327,17 @@ Section DoubleDiv32.
match w_compare a1 b1 with
| Lt =>
let (q,r) := w_div21 a1 a2 b1 in
- match ww_sub_c (w_WW r a3) (w_mul_c q b2) with
+ match ww_sub_c (w_WW r a3) (w_mul_c q b2) with
| C0 r1 => (q,r1)
| C1 r1 =>
let q := w_pred q in
- ww_add_c_cont w_WW w_add_c w_add_carry_c
+ ww_add_c_cont w_WW w_add_c w_add_carry_c
(fun r2=>(w_pred q, ww_add w_add_c w_add w_add_carry r2 (WW b1 b2)))
(fun r2 => (q,r2))
r1 (WW b1 b2)
end
| Eq =>
- ww_add_c_cont w_WW w_add_c w_add_carry_c
+ ww_add_c_cont w_WW w_add_c w_add_carry_c
(fun r => (w_Bm2, ww_add w_add_c w_add w_add_carry r (WW b1 b2)))
(fun r => (w_Bm1,r))
(WW (w_sub a2 b2) a3) (WW b1 b2)
@@ -360,7 +360,7 @@ Section DoubleDiv32.
[|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
0 <= [[r]] < [|b1|] * wB + [|b2|]);eauto.
rewrite H0;intros r.
- repeat
+ 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.
assert (0<= ([[r]] + ([|b1|] * wB + [|b2|])) - wwB < [|b1|] * wB + [|b2|]).
@@ -385,7 +385,7 @@ Section DoubleDiv32.
1 ([[r]] + ([|b1|] * wB + [|b2|]) - wwB));zarith;try (ring;fail).
split. rewrite H1;rewrite Hcmp;ring. trivial.
Spec_ww_to_Z (WW b1 b2). simpl in HH4;zarith.
- rewrite H0;intros r;repeat
+ rewrite H0;intros r;repeat
(rewrite spec_w_Bm1 || rewrite spec_w_Bm2);
simpl ww_to_Z;try rewrite Zmult_1_l;intros H1.
assert ([[r]]=([|a2|]-[|b2|])*wB+[|a3|]+([|b1|]*wB+[|b2|])). zarith.
@@ -409,7 +409,7 @@ Section DoubleDiv32.
as [r1|r1];repeat (rewrite spec_w_WW || rewrite spec_mul_c);
unfold interp_carry;intros H1.
rewrite H1.
- split. ring. split.
+ 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|]).
assert ( 0 <= [|q|] * [|b2|]);zarith.
@@ -418,7 +418,7 @@ Section DoubleDiv32.
rewrite <- H1;ring.
Spec_ww_to_Z r1; assert (0 <= [|r|]*wB). zarith.
assert (0 < [|q|] * [|b2|]). zarith.
- assert (0 < [|q|]).
+ assert (0 < [|q|]).
apply Zmult_lt_0_reg_r_2 with [|b2|];zarith.
eapply spec_ww_add_c_cont with (P :=
fun (x y:zn2z w) (res:w*zn2z w) =>
@@ -440,18 +440,18 @@ Section DoubleDiv32.
wwB * 1 +
([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2 * ([|b1|] * wB + [|b2|]))).
rewrite H7;rewrite H2;ring.
- assert
- ([|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|])
+ assert
+ ([|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|])
< [|b1|]*wB + [|b2|]).
Spec_ww_to_Z r2;omega.
Spec_ww_to_Z (WW b1 b2). simpl in HH5.
- assert
- (0 <= [|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|])
+ assert
+ (0 <= [|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|])
< 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 <- (Zmod_unique
+ rewrite <- (Zmod_unique
([[r2]] + ([|b1|] * wB + [|b2|]))
wwB
1
@@ -486,7 +486,7 @@ Section DoubleDiv21.
Definition ww_div21 a1 a2 b :=
match a1 with
- | W0 =>
+ | W0 =>
match ww_compare a2 b with
| Gt => (ww_1, ww_sub a2 b)
| Eq => (ww_1, W0)
@@ -529,8 +529,8 @@ Section DoubleDiv21.
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[-[ c ]]" :=
+ (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
Variable spec_w_0 : [|w_0|] = 0.
@@ -540,8 +540,8 @@ Section DoubleDiv21.
wB/2 <= [|b1|] ->
[[WW a1 a2]] < [[WW b1 b2]] ->
let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- [|a1|] * wwB + [|a2|] * wB + [|a3|] =
- [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
+ [|a1|] * wwB + [|a2|] * wB + [|a3|] =
+ [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
0 <= [[r]] < [|b1|] * wB + [|b2|].
Variable spec_ww_1 : [[ww_1]] = 1.
Variable spec_ww_compare : forall x y,
@@ -591,10 +591,10 @@ Section DoubleDiv21.
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|]);[
+ 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);
- rewrite <- wwB_div_2; trivial
+ 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;
intros (H1,H2) ]).
@@ -611,10 +611,10 @@ Section DoubleDiv21.
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 <- Zpower_2; 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|]).
+ replace (([|a1h|] * wB + [|a1l|]) * wB^2 + ([|a3|] * wB + [|a4|]))
+ with (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB+ [|a4|]).
rewrite H1;ring. rewrite wwB_wBwB;ring.
change [|a4|] with (0*wB+[|a4|]);apply beta_lex_inv;zarith.
assert (1 <= wB/2);zarith.
@@ -624,7 +624,7 @@ Section DoubleDiv21.
intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end.
split;trivial.
replace (([|a1h|] * wB + [|a1l|]) * wwB + ([|a3|] * wB + [|a4|])) with
- (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB + [|a4|]);
+ (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB + [|a4|]);
[rewrite H1 | rewrite wwB_wBwB;ring].
replace (([|q1|]*([|b1|]*wB+[|b2|])+([|r1|]*wB+[|r2|]))*wB+[|a4|]) with
(([|q1|]*([|b1|]*wB+[|b2|]))*wB+([|r1|]*wwB+[|r2|]*wB+[|a4|]));
@@ -666,22 +666,22 @@ Section DoubleDivGt.
Eval lazy beta iota delta [ww_sub ww_opp] in
let p := w_head0 bh in
match w_compare p w_0 with
- | Gt =>
+ | Gt =>
let b1 := w_add_mul_div p bh bl in
let b2 := w_add_mul_div p bl w_0 in
let a1 := w_add_mul_div p w_0 ah in
let a2 := w_add_mul_div p ah al in
let a3 := w_add_mul_div p al w_0 in
let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- (WW w_0 q, ww_add_mul_div
+ (WW w_0 q, 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 p)) W0 r)
| _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
w_opp w_sub w_sub_carry (WW ah al) (WW bh bl))
end.
- Definition ww_div_gt a b :=
- Eval lazy beta iota delta [ww_div_gt_aux double_divn1
+ Definition ww_div_gt a b :=
+ Eval lazy beta iota delta [ww_div_gt_aux double_divn1
double_divn1_p double_divn1_p_aux double_divn1_0 double_divn1_0_aux
double_split double_0 double_WW] in
match a, b with
@@ -691,11 +691,11 @@ Section DoubleDivGt.
if w_eq0 ah then
let (q,r) := w_div_gt al bl in
(WW w_0 q, w_0W r)
- else
+ else
match w_compare w_0 bh with
- | Eq =>
+ | Eq =>
let(q,r):=
- double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
+ double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
w_compare w_sub 1 a bl in
(q, w_0W r)
| Lt => ww_div_gt_aux ah al bh bl
@@ -707,7 +707,7 @@ Section DoubleDivGt.
Eval lazy beta iota delta [ww_sub ww_opp] in
let p := w_head0 bh in
match w_compare p w_0 with
- | Gt =>
+ | Gt =>
let b1 := w_add_mul_div p bh bl in
let b2 := w_add_mul_div p bl w_0 in
let a1 := w_add_mul_div p w_0 ah in
@@ -716,13 +716,13 @@ Section DoubleDivGt.
let (q,r) := w_div32 a1 a2 a3 b1 b2 in
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 p)) W0 r
- | _ =>
+ | _ =>
ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
w_opp w_sub w_sub_carry (WW ah al) (WW bh bl)
end.
- Definition ww_mod_gt a b :=
- Eval lazy beta iota delta [ww_mod_gt_aux double_modn1
+ Definition ww_mod_gt a b :=
+ Eval lazy beta iota delta [ww_mod_gt_aux double_modn1
double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux
double_split double_0 double_WW snd] in
match a, b with
@@ -730,10 +730,10 @@ Section DoubleDivGt.
| _, W0 => W0
| WW ah al, WW bh bl =>
if w_eq0 ah then w_0W (w_mod_gt al bl)
- else
+ else
match w_compare w_0 bh with
- | Eq =>
- w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
+ | Eq =>
+ w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
w_compare w_sub 1 a bl)
| Lt => ww_mod_gt_aux ah al bh bl
| Gt => W0 (* cas absurde *)
@@ -741,14 +741,14 @@ Section DoubleDivGt.
end.
Definition ww_gcd_gt_body (cont: w->w->w->w->zn2z w) (ah al bh bl: w) :=
- Eval lazy beta iota delta [ww_mod_gt_aux double_modn1
+ Eval lazy beta iota delta [ww_mod_gt_aux double_modn1
double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux
double_split double_0 double_WW snd] in
match w_compare w_0 bh with
| Eq =>
match w_compare w_0 bl with
| Eq => WW ah al (* normalement n'arrive pas si forme normale *)
- | Lt =>
+ | Lt =>
let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
w_compare w_sub 1 (WW ah al) bl in
WW w_0 (w_gcd_gt bl m)
@@ -757,14 +757,14 @@ Section DoubleDivGt.
| Lt =>
let m := ww_mod_gt_aux ah al bh bl in
match m with
- | W0 => WW bh bl
+ | W0 => WW bh bl
| WW mh ml =>
match w_compare w_0 mh with
| Eq =>
match w_compare w_0 ml with
| Eq => WW bh bl
- | _ =>
- let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
+ | _ =>
+ let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
w_compare w_sub 1 (WW bh bl) ml in
WW w_0 (w_gcd_gt ml r)
end
@@ -779,18 +779,18 @@ Section DoubleDivGt.
end
| Gt => W0 (* absurde *)
end.
-
- Fixpoint ww_gcd_gt_aux
- (p:positive) (cont: w -> w -> w -> w -> zn2z w) (ah al bh bl : w)
+
+ Fixpoint ww_gcd_gt_aux
+ (p:positive) (cont: w -> w -> w -> w -> zn2z w) (ah al bh bl : w)
{struct p} : zn2z w :=
- ww_gcd_gt_body
+ ww_gcd_gt_body
(fun mh ml rh rl => match p with
| xH => cont mh ml rh rl
| xO p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl
| xI p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl
end) ah al bh bl.
-
+
(* Proof *)
Variable w_to_Z : w -> Z.
@@ -816,7 +816,7 @@ Section DoubleDivGt.
| Gt => [|x|] > [|y|]
end.
Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0.
-
+
Variable spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|].
Variable spec_opp : forall x, [|w_opp x|] = (-[|x|]) mod wB.
Variable spec_opp_carry : forall x, [|w_opp_carry x|] = wB - [|x|] - 1.
@@ -854,8 +854,8 @@ Section DoubleDivGt.
wB/2 <= [|b1|] ->
[[WW a1 a2]] < [[WW b1 b2]] ->
let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- [|a1|] * wwB + [|a2|] * wB + [|a3|] =
- [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
+ [|a1|] * wwB + [|a2|] * wB + [|a3|] =
+ [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\
0 <= [[r]] < [|b1|] * wB + [|b2|].
Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits.
@@ -899,14 +899,14 @@ Section DoubleDivGt.
change
(let (q, r) := let p := w_head0 bh in
match w_compare p w_0 with
- | Gt =>
+ | Gt =>
let b1 := w_add_mul_div p bh bl in
let b2 := w_add_mul_div p bl w_0 in
let a1 := w_add_mul_div p w_0 ah in
let a2 := w_add_mul_div p ah al in
let a3 := w_add_mul_div p al w_0 in
let (q,r) := w_div32 a1 a2 a3 b1 b2 in
- (WW w_0 q, ww_add_mul_div
+ (WW w_0 q, 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 p)) W0 r)
| _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c
@@ -931,7 +931,7 @@ Section DoubleDivGt.
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.
- elimtype False.
+ 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.
@@ -945,11 +945,11 @@ Section DoubleDivGt.
(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.
- Spec_w_to_Z ah;Spec_w_to_Z bh.
+ 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 Zmult_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.
@@ -964,15 +964,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 Zmult_plus_distr_l.
+ rewrite <- (Zplus_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 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.
+ rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_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.
@@ -1027,7 +1027,7 @@ Section DoubleDivGt.
[[a]] = [[q]] * [[b]] + [[r]] /\
0 <= [[r]] < [[b]].
Proof.
- intros a b Hgt Hpos;unfold ww_div_gt.
+ intros a b Hgt Hpos;unfold ww_div_gt.
change (let (q,r) := match a, b with
| W0, _ => (W0,W0)
| _, W0 => (W0,W0)
@@ -1035,23 +1035,23 @@ Section DoubleDivGt.
if w_eq0 ah then
let (q,r) := w_div_gt al bl in
(WW w_0 q, w_0W r)
- else
+ else
match w_compare w_0 bh with
- | Eq =>
+ | Eq =>
let(q,r):=
- double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
+ double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
w_compare w_sub 1 a bl in
(q, w_0W r)
| Lt => ww_div_gt_aux ah al bh bl
| Gt => (W0,W0) (* cas absurde *)
end
- end in [[a]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]).
+ end in [[a]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]).
destruct a as [ |ah al]. simpl in Hgt;omega.
destruct b as [ |bh bl]. simpl in Hpos;omega.
Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl.
assert (H:=@spec_eq0 ah);destruct (w_eq0 ah).
simpl ww_to_Z;rewrite H;trivial. simpl in Hgt;rewrite H in Hgt;trivial.
- assert ([|bh|] <= 0).
+ assert ([|bh|] <= 0).
apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith.
assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;rewrite H1;simpl in Hgt.
simpl. simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos.
@@ -1066,12 +1066,12 @@ Section DoubleDivGt.
w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0
spec_add_mul_div spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hpos).
unfold double_to_Z,double_wB,double_digits in H2.
- destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
+ destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
w_compare w_sub 1
(WW ah al) bl).
rewrite spec_w_0W;unfold ww_to_Z;trivial.
apply spec_ww_div_gt_aux;trivial. rewrite spec_w_0 in Hcmp;trivial.
- rewrite spec_w_0 in Hcmp;elimtype False;omega.
+ rewrite spec_w_0 in Hcmp;exfalso;omega.
Qed.
Lemma spec_ww_mod_gt_aux_eq : forall ah al bh bl,
@@ -1104,26 +1104,26 @@ Section DoubleDivGt.
rewrite Zmult_comm in H;destruct H.
symmetry;apply Zmod_unique with [|q|];trivial.
Qed.
-
+
Lemma spec_ww_mod_gt_eq : forall a b, [[a]] > [[b]] -> 0 < [[b]] ->
[[ww_mod_gt a b]] = [[snd (ww_div_gt a b)]].
Proof.
intros a b Hgt Hpos.
- change (ww_mod_gt a b) with
+ change (ww_mod_gt a b) with
(match a, b with
| W0, _ => W0
| _, W0 => W0
| WW ah al, WW bh bl =>
if w_eq0 ah then w_0W (w_mod_gt al bl)
- else
+ else
match w_compare w_0 bh with
- | Eq =>
- w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
+ | Eq =>
+ w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
w_compare w_sub 1 a bl)
| Lt => ww_mod_gt_aux ah al bh bl
| Gt => W0 (* cas absurde *)
end end).
- change (ww_div_gt a b) with
+ change (ww_div_gt a b) with
(match a, b with
| W0, _ => (W0,W0)
| _, W0 => (W0,W0)
@@ -1131,11 +1131,11 @@ Section DoubleDivGt.
if w_eq0 ah then
let (q,r) := w_div_gt al bl in
(WW w_0 q, w_0W r)
- else
+ else
match w_compare w_0 bh with
- | Eq =>
+ | Eq =>
let(q,r):=
- double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
+ double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21
w_compare w_sub 1 a bl in
(q, w_0W r)
| Lt => ww_div_gt_aux ah al bh bl
@@ -1147,7 +1147,7 @@ Section DoubleDivGt.
Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl.
assert (H:=@spec_eq0 ah);destruct (w_eq0 ah).
simpl in Hgt;rewrite H in Hgt;trivial.
- assert ([|bh|] <= 0).
+ assert ([|bh|] <= 0).
apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith.
assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt.
simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos.
@@ -1155,7 +1155,7 @@ Section DoubleDivGt.
destruct (w_div_gt al bl);simpl;rewrite spec_w_0W;trivial.
clear H.
assert (H2 := spec_compare w_0 bh);destruct (w_compare w_0 bh).
- rewrite (@spec_double_modn1_aux w w_zdigits w_0 w_WW w_head0 w_add_mul_div
+ 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
(WW ah al) bl);simpl;trivial.
@@ -1174,7 +1174,7 @@ Section DoubleDivGt.
rewrite Zmult_comm;trivial.
Qed.
- Lemma Zis_gcd_mod : forall a b d,
+ Lemma Zis_gcd_mod : forall a b d,
0 < b -> Zis_gcd b (a mod b) d -> Zis_gcd a b d.
Proof.
intros a b d H H1; apply Zis_gcd_for_euclid with (a/b).
@@ -1182,12 +1182,12 @@ Section DoubleDivGt.
ring_simplify (b * (a / b) + a mod b - a / b * b);trivial. zarith.
Qed.
- Lemma spec_ww_gcd_gt_aux_body :
+ Lemma spec_ww_gcd_gt_aux_body :
forall ah al bh bl n cont,
- [[WW bh bl]] <= 2^n ->
+ [[WW bh bl]] <= 2^n ->
[[WW ah al]] > [[WW bh bl]] ->
- (forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^(n-1) ->
+ (forall xh xl yh yl,
+ [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^(n-1) ->
Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) ->
Zis_gcd [[WW ah al]] [[WW bh bl]] [[ww_gcd_gt_body cont ah al bh bl]].
Proof.
@@ -1196,7 +1196,7 @@ Section DoubleDivGt.
| Eq =>
match w_compare w_0 bl with
| Eq => WW ah al (* normalement n'arrive pas si forme normale *)
- | Lt =>
+ | Lt =>
let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
w_compare w_sub 1 (WW ah al) bl in
WW w_0 (w_gcd_gt bl m)
@@ -1205,14 +1205,14 @@ Section DoubleDivGt.
| Lt =>
let m := ww_mod_gt_aux ah al bh bl in
match m with
- | W0 => WW bh bl
+ | W0 => WW bh bl
| WW mh ml =>
match w_compare w_0 mh with
| Eq =>
match w_compare w_0 ml with
| Eq => WW bh bl
- | _ =>
- let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
+ | _ =>
+ let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21
w_compare w_sub 1 (WW bh bl) ml in
WW w_0 (w_gcd_gt ml r)
end
@@ -1227,10 +1227,10 @@ Section DoubleDivGt.
end
| Gt => W0 (* absurde *)
end).
- assert (Hbh := spec_compare w_0 bh);destruct (w_compare w_0 bh).
+ assert (Hbh := spec_compare w_0 bh);destruct (w_compare w_0 bh).
simpl ww_to_Z in *. rewrite spec_w_0 in Hbh;rewrite <- Hbh;
rewrite Zmult_0_l;rewrite Zplus_0_l.
- assert (Hbl := spec_compare w_0 bl); destruct (w_compare w_0 bl).
+ assert (Hbl := spec_compare w_0 bl); destruct (w_compare w_0 bl).
rewrite spec_w_0 in Hbl;rewrite <- Hbl;apply Zis_gcd_0.
simpl;rewrite spec_w_0;rewrite Zmult_0_l;rewrite Zplus_0_l.
rewrite spec_w_0 in Hbl.
@@ -1239,54 +1239,54 @@ Section DoubleDivGt.
rewrite <- (@spec_double_modn1 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 spec_add_mul_div
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 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 =>
destruct (Z_mod_lt x y);zarith end.
- rewrite spec_w_0 in Hbl;Spec_w_to_Z bl;elimtype False;omega.
+ rewrite spec_w_0 in Hbl;Spec_w_to_Z bl;exfalso;omega.
rewrite spec_w_0 in Hbh;assert (H:= spec_ww_mod_gt_aux _ _ _ Hgt Hbh).
- assert (H2 : 0 < [[WW bh bl]]).
+ 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.
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.
- assert (Hmh := spec_compare w_0 mh);destruct (w_compare w_0 mh).
+ simpl;apply Zis_gcd_0;zarith.
+ assert (Hmh := spec_compare w_0 mh);destruct (w_compare w_0 mh).
simpl;rewrite spec_w_0 in Hmh; rewrite <- Hmh;simpl.
- assert (Hml := spec_compare w_0 ml);destruct (w_compare w_0 ml).
+ assert (Hml := spec_compare w_0 ml);destruct (w_compare w_0 ml).
rewrite <- Hml;rewrite spec_w_0;simpl;apply Zis_gcd_0.
- simpl;rewrite spec_w_0;simpl.
+ simpl;rewrite spec_w_0;simpl.
rewrite spec_w_0 in Hml. apply Zis_gcd_mod;zarith.
change ([|bh|] * wB + [|bl|]) with (double_to_Z w_digits w_to_Z 1 (WW bh bl)).
rewrite <- (@spec_double_modn1 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 spec_add_mul_div
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 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 =>
destruct (Z_mod_lt x y);zarith end.
- rewrite spec_w_0 in Hml;Spec_w_to_Z ml;elimtype False;omega.
+ rewrite spec_w_0 in Hml;Spec_w_to_Z ml;exfalso;omega.
rewrite spec_w_0 in Hmh. assert ([[WW bh bl]] > [[WW mh ml]]).
- rewrite H;simpl; apply Zlt_gt;match goal with | |- ?x mod ?y < ?y =>
+ rewrite H;simpl; apply Zlt_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]]).
+ 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.
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 Zlt_gt;match goal with | |- ?x mod ?y < ?y =>
destruct (Z_mod_lt x y);zarith end.
- apply Zle_trans with (2^n/2).
- apply Zdiv_le_lower_bound;zarith.
+ apply Zle_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.
pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3'.
destruct (Zle_lt_or_eq _ _ H4').
- assert (H6' : [[WW bh bl]] mod [[WW mh ml]] =
+ 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]])).
@@ -1300,14 +1300,14 @@ Section DoubleDivGt.
rewrite Z_div_mult;zarith.
assert (2^1 <= 2^n). change (2^1) with 2;zarith.
assert (H7 := @Zpower_le_monotone_inv 2 1 n);zarith.
- rewrite spec_w_0 in Hmh;Spec_w_to_Z mh;elimtype False;zarith.
- rewrite spec_w_0 in Hbh;Spec_w_to_Z bh;elimtype False;zarith.
+ rewrite spec_w_0 in Hmh;Spec_w_to_Z mh;exfalso;zarith.
+ rewrite spec_w_0 in Hbh;Spec_w_to_Z bh;exfalso;zarith.
Qed.
- Lemma spec_ww_gcd_gt_aux :
+ Lemma spec_ww_gcd_gt_aux :
forall p cont n,
- (forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
+ (forall xh xl yh yl,
+ [[WW xh xl]] > [[WW yh yl]] ->
[[WW yh yl]] <= 2^n ->
Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) ->
forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] ->
@@ -1334,7 +1334,7 @@ Section DoubleDivGt.
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 Zpower_le_monotone2;zarith.
apply spec_ww_gcd_gt_aux_body with (n := n+1);trivial.
rewrite Zplus_comm;trivial.
ring_simplify (n + 1 - 1);trivial.
@@ -1352,16 +1352,16 @@ Section DoubleDiv.
Variable ww_div_gt : zn2z w -> zn2z w -> zn2z w * zn2z w.
Variable ww_mod_gt : zn2z w -> zn2z w -> zn2z w.
- Definition ww_div a b :=
- match ww_compare a b with
- | Gt => ww_div_gt a b
+ Definition ww_div a b :=
+ match ww_compare a b with
+ | Gt => ww_div_gt a b
| Eq => (ww_1, W0)
| Lt => (W0, a)
end.
- Definition ww_mod a b :=
- match ww_compare a b with
- | Gt => ww_mod_gt a b
+ Definition ww_mod a b :=
+ match ww_compare a b with
+ | Gt => ww_mod_gt a b
| Eq => W0
| Lt => a
end.
@@ -1401,7 +1401,7 @@ Section DoubleDiv.
Proof.
intros a b Hpos;unfold ww_div.
assert (H:=spec_ww_compare a b);destruct (ww_compare a b).
- simpl;rewrite spec_ww_1;split;zarith.
+ simpl;rewrite spec_ww_1;split;zarith.
simpl;split;[ring|Spec_ww_to_Z a;zarith].
apply spec_ww_div_gt;trivial.
Qed.
@@ -1409,7 +1409,7 @@ Section DoubleDiv.
Lemma spec_ww_mod : forall a b, 0 < [[b]] ->
[[ww_mod a b]] = [[a]] mod [[b]].
Proof.
- intros a b Hpos;unfold ww_mod.
+ intros a b Hpos;unfold ww_mod.
assert (H := spec_ww_compare a b);destruct (ww_compare a b).
simpl;apply Zmod_unique with 1;try rewrite H;zarith.
Spec_ww_to_Z a;symmetry;apply Zmod_small;zarith.
@@ -1424,8 +1424,8 @@ Section DoubleDiv.
Variable w_gcd_gt : w -> w -> w.
Variable _ww_digits : positive.
Variable spec_ww_digits_ : _ww_digits = xO w_digits.
- Variable ww_gcd_gt_fix :
- positive -> (w -> w -> w -> w -> zn2z w) ->
+ Variable ww_gcd_gt_fix :
+ positive -> (w -> w -> w -> w -> zn2z w) ->
w -> w -> w -> w -> zn2z w.
Variable spec_w_0 : [|w_0|] = 0.
@@ -1440,10 +1440,10 @@ Section DoubleDiv.
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|].
- Variable spec_gcd_gt_fix :
+ Variable spec_gcd_gt_fix :
forall p cont n,
- (forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
+ (forall xh xl yh yl,
+ [[WW xh xl]] > [[WW yh yl]] ->
[[WW yh yl]] <= 2^n ->
Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) ->
forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] ->
@@ -1451,20 +1451,20 @@ Section DoubleDiv.
Zis_gcd [[WW ah al]] [[WW bh bl]]
[[ww_gcd_gt_fix p cont ah al bh bl]].
- Definition gcd_cont (xh xl yh yl:w) :=
+ Definition gcd_cont (xh xl yh yl:w) :=
match w_compare w_1 yl with
- | Eq => ww_1
+ | Eq => ww_1
| _ => WW xh xl
end.
- Lemma spec_gcd_cont : forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
+ Lemma spec_gcd_cont : forall xh xl yh yl,
+ [[WW xh xl]] > [[WW yh yl]] ->
[[WW yh yl]] <= 1 ->
Zis_gcd [[WW xh xl]] [[WW yh yl]] [[gcd_cont xh xl yh yl]].
Proof.
intros xh xl yh yl Hgt' Hle. simpl in Hle.
assert ([|yh|] = 0).
- change 1 with (0*wB+1) in Hle.
+ change 1 with (0*wB+1) in Hle.
assert (0 <= 1 < wB). split;zarith. apply wB_pos.
assert (H1:= beta_lex _ _ _ _ _ Hle (spec_to_Z yl) H).
Spec_w_to_Z yh;zarith.
@@ -1473,20 +1473,20 @@ Section DoubleDiv.
simpl;rewrite H;simpl;destruct (w_compare w_1 yl).
rewrite spec_ww_1;rewrite <- Hcmpy;apply Zis_gcd_mod;zarith.
rewrite <- (Zmod_unique ([|xh|]*wB+[|xl|]) 1 ([|xh|]*wB+[|xl|]) 0);zarith.
- rewrite H in Hle; elimtype False;zarith.
+ rewrite H in Hle; exfalso;zarith.
assert ([|yl|] = 0). Spec_w_to_Z yl;zarith.
rewrite H0;simpl;apply Zis_gcd_0;trivial.
Qed.
-
+
Variable cont : w -> w -> w -> w -> zn2z w.
- Variable spec_cont : forall xh xl yh yl,
- [[WW xh xl]] > [[WW yh yl]] ->
+ Variable spec_cont : forall xh xl yh yl,
+ [[WW xh xl]] > [[WW yh yl]] ->
[[WW yh yl]] <= 1 ->
Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]].
-
- Definition ww_gcd_gt a b :=
- match a, b with
+
+ Definition ww_gcd_gt a b :=
+ match a, b with
| W0, _ => b
| _, W0 => a
| WW ah al, WW bh bl =>
@@ -1509,8 +1509,8 @@ Section DoubleDiv.
destruct a as [ |ah al]. simpl;apply Zis_gcd_sym;apply Zis_gcd_0.
destruct b as [ |bh bl]. simpl;apply Zis_gcd_0.
simpl in Hgt. generalize (@spec_eq0 ah);destruct (w_eq0 ah);intros.
- simpl;rewrite H in Hgt;trivial;rewrite H;trivial;rewrite spec_w_0;simpl.
- assert ([|bh|] <= 0).
+ simpl;rewrite H in Hgt;trivial;rewrite H;trivial;rewrite spec_w_0;simpl.
+ assert ([|bh|] <= 0).
apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith.
Spec_w_to_Z bh;assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt.
rewrite H1;simpl;auto. clear H.
@@ -1522,7 +1522,7 @@ Section DoubleDiv.
Lemma spec_ww_gcd : forall a b, Zis_gcd [[a]] [[b]] [[ww_gcd a b]].
Proof.
intros a b.
- change (ww_gcd a b) with
+ change (ww_gcd a b) with
(match ww_compare a b with
| Gt => ww_gcd_gt a b
| Eq => a
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
index d6f6a05f..386bbb9e 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleDivn1.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -17,7 +17,7 @@ Require Import BigNumPrelude.
Require Import DoubleType.
Require Import DoubleBase.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section GENDIVN1.
@@ -31,19 +31,19 @@ Section GENDIVN1.
Variable w_div21 : w -> w -> w -> w * w.
Variable w_compare : w -> w -> comparison.
Variable w_sub : w -> w -> w.
-
-
+
+
(* ** For proofs ** *)
Variable w_to_Z : w -> Z.
-
- Notation wB := (base w_digits).
+
+ Notation wB := (base w_digits).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
- Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x)
+ Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x)
(at level 0, x at level 99).
Notation "[[ x ]]" := (zn2z_to_Z wB w_to_Z x) (at level 0, x at level 99).
-
+
Variable spec_to_Z : forall x, 0 <= [| x |] < wB.
Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits.
Variable spec_0 : [|w_0|] = 0.
@@ -68,10 +68,10 @@ Section GENDIVN1.
| Lt => [|x|] < [|y|]
| Gt => [|x|] > [|y|]
end.
- Variable spec_sub: forall x y,
+ Variable spec_sub: forall x y,
[|w_sub x y|] = ([|x|] - [|y|]) mod wB.
-
+
Section DIVAUX.
Variable b2p : w.
@@ -85,10 +85,10 @@ Section GENDIVN1.
Fixpoint double_divn1_0 (n:nat) : w -> word w n -> word w n * w :=
match n return w -> word w n -> word w n * w with
- | O => fun r x => w_div21 r x b2p
- | S n => double_divn1_0_aux n (double_divn1_0 n)
+ | O => fun r x => w_div21 r x b2p
+ | S n => double_divn1_0_aux n (double_divn1_0 n)
end.
-
+
Lemma spec_split : forall (n : nat) (x : zn2z (word w n)),
let (h, l) := double_split w_0 n x in
[!S n | x!] = [!n | h!] * double_wB w_digits n + [!n | l!].
@@ -132,11 +132,11 @@ Section GENDIVN1.
induction n;simpl;intros;trivial.
unfold double_modn1_0_aux, double_divn1_0_aux.
destruct (double_split w_0 n x) as (hh,hl).
- rewrite (IHn r hh).
+ rewrite (IHn r hh).
destruct (double_divn1_0 n r hh) as (qh,rh);simpl.
rewrite IHn. destruct (double_divn1_0 n rh hl);trivial.
Qed.
-
+
Variable p : w.
Variable p_bounded : [|p|] <= Zpos w_digits.
@@ -148,18 +148,18 @@ Section GENDIVN1.
intros;apply spec_add_mul_div;auto.
Qed.
- Definition double_divn1_p_aux n
- (divn1 : w -> word w n -> word w n -> word w n * w) r h l :=
+ Definition double_divn1_p_aux n
+ (divn1 : w -> word w n -> word w n -> word w n * w) r h l :=
let (hh,hl) := double_split w_0 n h in
- let (lh,ll) := double_split w_0 n l in
+ let (lh,ll) := double_split w_0 n l in
let (qh,rh) := divn1 r hh hl in
let (ql,rl) := divn1 rh hl lh in
(double_WW w_WW n qh ql, rl).
Fixpoint double_divn1_p (n:nat) : w -> word w n -> word w n -> word w n * w :=
match n return w -> word w n -> word w n -> word w n * w with
- | O => fun r h l => w_div21 r (w_add_mul_div p h l) b2p
- | S n => double_divn1_p_aux n (double_divn1_p n)
+ | O => fun r h l => w_div21 r (w_add_mul_div p h l) b2p
+ | S n => double_divn1_p_aux n (double_divn1_p n)
end.
Lemma p_lt_double_digits : forall n, [|p|] <= Zpos (double_digits w_digits n).
@@ -175,8 +175,8 @@ Section GENDIVN1.
Lemma spec_double_divn1_p : forall n r h l,
[|r|] < [|b2p|] ->
let (q,r') := double_divn1_p n r h l in
- [|r|] * double_wB w_digits n +
- ([!n|h!]*2^[|p|] +
+ [|r|] * double_wB w_digits n +
+ ([!n|h!]*2^[|p|] +
[!n|l!] / (2^(Zpos(double_digits w_digits n) - [|p|])))
mod double_wB w_digits n = [!n|q!] * [|b2p|] + [|r'|] /\
0 <= [|r'|] < [|b2p|].
@@ -198,26 +198,26 @@ Section GENDIVN1.
([!n|lh!] * double_wB w_digits n + [!n|ll!]) /
2^(Zpos (double_digits w_digits (S n)) - [|p|])) mod
(double_wB w_digits n * double_wB w_digits n)) with
- (([|r|] * double_wB w_digits n + ([!n|hh!] * 2^[|p|] +
+ (([|r|] * double_wB w_digits n + ([!n|hh!] * 2^[|p|] +
[!n|hl!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod
double_wB w_digits n) * double_wB w_digits n +
- ([!n|hl!] * 2^[|p|] +
- [!n|lh!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod
+ ([!n|hl!] * 2^[|p|] +
+ [!n|lh!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod
double_wB w_digits n).
generalize (IHn r hh hl H);destruct (double_divn1_p n r hh hl) as (qh,rh);
intros (H3,H4);rewrite H3.
- assert ([|rh|] < [|b2p|]). omega.
+ assert ([|rh|] < [|b2p|]). omega.
replace (([!n|qh!] * [|b2p|] + [|rh|]) * double_wB w_digits n +
([!n|hl!] * 2 ^ [|p|] +
[!n|lh!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])) mod
- double_wB w_digits n) with
+ double_wB w_digits n) with
([!n|qh!] * [|b2p|] *double_wB w_digits n + ([|rh|]*double_wB w_digits n +
([!n|hl!] * 2 ^ [|p|] +
[!n|lh!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])) mod
double_wB w_digits n)). 2:ring.
generalize (IHn rh hl lh H0);destruct (double_divn1_p n rh hl lh) as (ql,rl);
intros (H5,H6);rewrite H5.
- split;[rewrite spec_double_WW;trivial;ring|trivial].
+ split;[rewrite spec_double_WW;trivial;ring|trivial].
assert (Uhh := spec_double_to_Z w_digits w_to_Z spec_to_Z n hh);
unfold double_wB,base in Uhh.
assert (Uhl := spec_double_to_Z w_digits w_to_Z spec_to_Z n hl);
@@ -228,37 +228,37 @@ Section GENDIVN1.
unfold double_wB,base in Ull.
unfold double_wB,base.
assert (UU:=p_lt_double_digits n).
- rewrite Zdiv_shift_r;auto with zarith.
- 2:change (Zpos (double_digits w_digits (S n)))
+ rewrite Zdiv_shift_r;auto with zarith.
+ 2:change (Zpos (double_digits w_digits (S n)))
with (2*Zpos (double_digits w_digits n));auto with zarith.
replace (2 ^ (Zpos (double_digits w_digits (S n)) - [|p|])) with
(2^(Zpos (double_digits w_digits n) - [|p|])*2^Zpos (double_digits w_digits n)).
rewrite Zdiv_mult_cancel_r;auto with zarith.
- rewrite Zmult_plus_distr_l with (p:= 2^[|p|]).
+ rewrite Zmult_plus_distr_l with (p:= 2^[|p|]).
pattern ([!n|hl!] * 2^[|p|]) at 2;
rewrite (shift_unshift_mod (Zpos(double_digits w_digits n))([|p|])([!n|hl!]));
auto with zarith.
- rewrite Zplus_assoc.
- replace
+ rewrite Zplus_assoc.
+ replace
([!n|hh!] * 2^Zpos (double_digits w_digits n)* 2^[|p|] +
([!n|hl!] / 2^(Zpos (double_digits w_digits n)-[|p|])*
2^Zpos(double_digits w_digits n)))
- with
- (([!n|hh!] *2^[|p|] + double_to_Z w_digits w_to_Z n hl /
+ with
+ (([!n|hh!] *2^[|p|] + double_to_Z w_digits w_to_Z n hl /
2^(Zpos (double_digits w_digits n)-[|p|]))
* 2^Zpos(double_digits w_digits n));try (ring;fail).
rewrite <- Zplus_assoc.
rewrite <- (Zmod_shift_r ([|p|]));auto with zarith.
- replace
+ replace
(2 ^ Zpos (double_digits w_digits n) * 2 ^ Zpos (double_digits w_digits n)) with
(2 ^ (Zpos (double_digits w_digits n) + Zpos (double_digits w_digits n))).
rewrite (Zmod_shift_r (Zpos (double_digits w_digits n)));auto with zarith.
replace (2 ^ (Zpos (double_digits w_digits n) + Zpos (double_digits w_digits n)))
- with (2^Zpos(double_digits w_digits n) *2^Zpos(double_digits w_digits n)).
+ with (2^Zpos(double_digits w_digits n) *2^Zpos(double_digits w_digits n)).
rewrite (Zmult_comm (([!n|hh!] * 2 ^ [|p|] +
[!n|hl!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])))).
rewrite Zmult_mod_distr_l;auto with zarith.
- ring.
+ ring.
rewrite Zpower_exp;auto with zarith.
assert (0 < Zpos (double_digits w_digits n)). unfold Zlt;reflexivity.
auto with zarith.
@@ -267,24 +267,24 @@ Section GENDIVN1.
split;auto with zarith.
apply Zdiv_lt_upper_bound;auto with zarith.
rewrite <- Zpower_exp;auto with zarith.
- replace ([|p|] + (Zpos (double_digits w_digits n) - [|p|])) with
+ replace ([|p|] + (Zpos (double_digits w_digits n) - [|p|])) with
(Zpos(double_digits w_digits n));auto with zarith.
rewrite <- Zpower_exp;auto with zarith.
- replace (Zpos (double_digits w_digits (S n)) - [|p|]) with
- (Zpos (double_digits w_digits n) - [|p|] +
+ replace (Zpos (double_digits w_digits (S n)) - [|p|]) with
+ (Zpos (double_digits w_digits n) - [|p|] +
Zpos (double_digits w_digits n));trivial.
- change (Zpos (double_digits w_digits (S n))) with
+ change (Zpos (double_digits w_digits (S n))) with
(2*Zpos (double_digits w_digits n)). ring.
Qed.
Definition double_modn1_p_aux n (modn1 : w -> word w n -> word w n -> w) r h l:=
let (hh,hl) := double_split w_0 n h in
- let (lh,ll) := double_split w_0 n l in
+ let (lh,ll) := double_split w_0 n l in
modn1 (modn1 r hh hl) hl lh.
Fixpoint double_modn1_p (n:nat) : w -> word w n -> word w n -> w :=
match n return w -> word w n -> word w n -> w with
- | O => fun r h l => snd (w_div21 r (w_add_mul_div p h l) b2p)
+ | O => fun r h l => snd (w_div21 r (w_add_mul_div p h l) b2p)
| S n => double_modn1_p_aux n (double_modn1_p n)
end.
@@ -302,8 +302,8 @@ Section GENDIVN1.
Fixpoint high (n:nat) : word w n -> w :=
match n return word w n -> w with
- | O => fun a => a
- | S n =>
+ | O => fun a => a
+ | S n =>
fun (a:zn2z (word w n)) =>
match a with
| W0 => w_0
@@ -314,20 +314,20 @@ Section GENDIVN1.
Lemma spec_double_digits:forall n, Zpos w_digits <= Zpos (double_digits w_digits n).
Proof.
induction n;simpl;auto with zarith.
- change (Zpos (xO (double_digits w_digits n))) with
+ change (Zpos (xO (double_digits w_digits n))) with
(2*Zpos (double_digits w_digits n)).
assert (0 < Zpos w_digits);auto with zarith.
exact (refl_equal Lt).
Qed.
- Lemma spec_high : forall n (x:word w n),
+ Lemma spec_high : forall n (x:word w n),
[|high n x|] = [!n|x!] / 2^(Zpos (double_digits w_digits n) - Zpos w_digits).
Proof.
induction n;intros.
unfold high,double_digits,double_to_Z.
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 (U2 := spec_double_digits n).
assert (U3 : 0 < Zpos w_digits). exact (refl_equal Lt).
destruct x;unfold high;fold high.
unfold double_to_Z,zn2z_to_Z;rewrite spec_0.
@@ -337,31 +337,31 @@ Section GENDIVN1.
simpl [!S n|WW w0 w1!].
unfold double_wB,base;rewrite Zdiv_shift_r;auto with zarith.
replace (2 ^ (Zpos (double_digits w_digits (S n)) - Zpos w_digits)) with
- (2^(Zpos (double_digits w_digits n) - Zpos w_digits) *
+ (2^(Zpos (double_digits w_digits n) - Zpos w_digits) *
2^Zpos (double_digits w_digits n)).
rewrite Zdiv_mult_cancel_r;auto with zarith.
rewrite <- Zpower_exp;auto with zarith.
- replace (Zpos (double_digits w_digits n) - Zpos w_digits +
+ replace (Zpos (double_digits w_digits n) - Zpos w_digits +
Zpos (double_digits w_digits n)) with
(Zpos (double_digits w_digits (S n)) - Zpos w_digits);trivial.
- change (Zpos (double_digits w_digits (S n))) with
+ change (Zpos (double_digits w_digits (S n))) with
(2*Zpos (double_digits w_digits n));ring.
- change (Zpos (double_digits w_digits (S n))) with
+ change (Zpos (double_digits w_digits (S n))) with
(2*Zpos (double_digits w_digits n)); auto with zarith.
Qed.
-
- Definition double_divn1 (n:nat) (a:word w n) (b:w) :=
+
+ Definition double_divn1 (n:nat) (a:word w n) (b:w) :=
let p := w_head0 b in
match w_compare p w_0 with
| Gt =>
let b2p := w_add_mul_div p b w_0 in
let ha := high n a in
let k := w_sub w_zdigits p in
- let lsr_n := w_add_mul_div k w_0 in
+ let lsr_n := w_add_mul_div k w_0 in
let r0 := w_add_mul_div p w_0 ha in
let (q,r) := double_divn1_p b2p p n r0 a (double_0 w_0 n) in
(q, lsr_n r)
- | _ => double_divn1_0 b n w_0 a
+ | _ => double_divn1_0 b n w_0 a
end.
Lemma spec_double_divn1 : forall n a b,
@@ -392,21 +392,21 @@ Section GENDIVN1.
apply Zmult_le_compat;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|] =
+ 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 Zmod_small; auto with zarith.
assert (H5 := spec_to_Z (high n a)).
- assert
+ assert
([|w_add_mul_div (w_head0 b) w_0 (high n a)|]
<[|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.
assert (([|high n a|]/2^(Zpos w_digits - [|w_head0 b|])) < wB).
- apply Zdiv_lt_upper_bound;auto with zarith.
+ apply Zdiv_lt_upper_bound;auto with zarith.
apply Zlt_le_trans with wB;auto with zarith.
pattern wB at 1;replace wB with (wB*1);try ring.
apply Zmult_le_compat;auto with zarith.
@@ -420,8 +420,8 @@ Section GENDIVN1.
apply Zmult_le_compat;auto with zarith.
pattern 2 at 1;rewrite <- Zpower_1_r.
apply Zpower_le_monotone;split;auto with zarith.
- rewrite <- H4 in H0.
- assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith.
+ rewrite <- H4 in H0.
+ assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith.
assert (H7:= spec_double_divn1_p H0 Hb3 n a (double_0 w_0 n) H6).
destruct (double_divn1_p (w_add_mul_div (w_head0 b) b w_0) (w_head0 b) n
(w_add_mul_div (w_head0 b) w_0 (high n a)) a
@@ -436,7 +436,7 @@ Section GENDIVN1.
rewrite Zmod_small;auto with zarith.
rewrite spec_high. rewrite Zdiv_Zdiv;auto with zarith.
rewrite <- Zpower_exp;auto with zarith.
- replace (Zpos (double_digits w_digits n) - Zpos w_digits +
+ replace (Zpos (double_digits w_digits n) - Zpos w_digits +
(Zpos w_digits - [|w_head0 b|]))
with (Zpos (double_digits w_digits n) - [|w_head0 b|]);trivial;ring.
assert (H8 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith.
@@ -448,11 +448,11 @@ Section GENDIVN1.
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|]
+ 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.
- replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|])
+ 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).
@@ -474,11 +474,11 @@ Section GENDIVN1.
split.
rewrite <- (Z_div_mult [!n|a!] (2^[|w_head0 b|]));auto with zarith.
rewrite H71;rewrite H9.
- replace ([!n|q!] * (2 ^ [|w_head0 b|] * [|b|]))
+ replace ([!n|q!] * (2 ^ [|w_head0 b|] * [|b|]))
with ([!n|q!] *[|b|] * 2^[|w_head0 b|]);
try (ring;fail).
rewrite Z_div_plus_l;auto with zarith.
- assert (H10 := spec_to_Z
+ assert (H10 := spec_to_Z
(w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r));split;
auto with zarith.
rewrite H9.
@@ -487,19 +487,19 @@ Section GENDIVN1.
exact (spec_double_to_Z w_digits w_to_Z spec_to_Z n a).
Qed.
-
- Definition double_modn1 (n:nat) (a:word w n) (b:w) :=
+
+ Definition double_modn1 (n:nat) (a:word w n) (b:w) :=
let p := w_head0 b in
match w_compare p w_0 with
| Gt =>
let b2p := w_add_mul_div p b w_0 in
let ha := high n a in
let k := w_sub w_zdigits p in
- let lsr_n := w_add_mul_div k w_0 in
+ let lsr_n := w_add_mul_div k w_0 in
let r0 := w_add_mul_div p w_0 ha in
let r := double_modn1_p b2p p n r0 a (double_0 w_0 n) in
lsr_n r
- | _ => double_modn1_0 b n w_0 a
+ | _ => double_modn1_0 b n w_0 a
end.
Lemma spec_double_modn1_aux : forall n a b,
@@ -525,4 +525,4 @@ Section GENDIVN1.
destruct H1 as (h1,h2);rewrite h1;ring.
Qed.
-End GENDIVN1.
+End GENDIVN1.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
index 50c72487..21e694e5 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleLift.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -17,7 +17,7 @@ Require Import BigNumPrelude.
Require Import DoubleType.
Require Import DoubleBase.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section DoubleLift.
Variable w : Type.
@@ -61,13 +61,13 @@ Section DoubleLift.
(* 0 < p < ww_digits *)
- Definition ww_add_mul_div p x y :=
+ Definition ww_add_mul_div p x y :=
let zdigits := w_0W w_zdigits in
match x, y with
| W0, W0 => W0
| W0, WW yh yl =>
match ww_compare p zdigits with
- | Eq => w_0W yh
+ | Eq => w_0W yh
| Lt => w_0W (w_add_mul_div (low p) w_0 yh)
| Gt =>
let n := low (ww_sub p zdigits) in
@@ -75,15 +75,15 @@ Section DoubleLift.
end
| WW xh xl, W0 =>
match ww_compare p zdigits with
- | Eq => w_W0 xl
+ | Eq => w_W0 xl
| Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl w_0)
| Gt =>
let n := low (ww_sub p zdigits) in
- w_W0 (w_add_mul_div n xl w_0)
+ w_W0 (w_add_mul_div n xl w_0)
end
| WW xh xl, WW yh yl =>
match ww_compare p zdigits with
- | Eq => w_WW xl yh
+ | Eq => w_WW xl yh
| Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl yh)
| Gt =>
let n := low (ww_sub p zdigits) in
@@ -93,7 +93,7 @@ Section DoubleLift.
Section DoubleProof.
Variable w_to_Z : w -> Z.
-
+
Notation wB := (base w_digits).
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
@@ -122,21 +122,21 @@ Section DoubleLift.
Variable spec_w_head0 : forall x, 0 < [|x|] ->
wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB.
Variable spec_w_tail00 : forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits.
- Variable spec_w_tail0 : forall x, 0 < [|x|] ->
+ Variable spec_w_tail0 : forall x, 0 < [|x|] ->
exists y, 0 <= y /\ [|x|] = (2* y + 1) * (2 ^ [|w_tail0 x|]).
Variable spec_w_add_mul_div : forall x y p,
[|p|] <= Zpos w_digits ->
[| w_add_mul_div p x y |] =
([|x|] * (2 ^ [|p|]) +
[|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB.
- Variable spec_w_add: forall x y,
+ Variable spec_w_add: forall x y,
[[w_add x y]] = [|x|] + [|y|].
- Variable spec_ww_sub: forall x y,
+ Variable spec_ww_sub: forall x y,
[[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits.
Variable spec_low: forall x, [| low x|] = [[x]] mod wB.
-
+
Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos ww_Digits.
Hint Resolve div_le_0 div_lt w_to_Z_wwB: lift.
@@ -168,7 +168,7 @@ Section DoubleLift.
rewrite spec_w_0; auto with zarith.
rewrite spec_w_0; auto with zarith.
Qed.
-
+
Lemma spec_ww_head0 : forall x, 0 < [[x]] ->
wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB.
Proof.
@@ -179,7 +179,7 @@ Section DoubleLift.
assert (H0 := spec_compare w_0 xh);rewrite spec_w_0 in H0.
destruct (w_compare w_0 xh).
rewrite <- H0. simpl Zplus. rewrite <- H0 in H;simpl in H.
- case (spec_to_Z w_zdigits);
+ 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.
@@ -209,7 +209,7 @@ Section DoubleLift.
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 <- Zpower_exp;zarith.
+ rewrite <- Zpower_exp;zarith.
rewrite Zmult_comm;ring_simplify (Zpos w_digits - p + p);fold wB;zarith.
assert (H1 := spec_to_Z xh);zarith.
Qed.
@@ -293,8 +293,8 @@ Section DoubleLift.
Qed.
Hint Rewrite Zdiv_0_l Zmult_0_l Zplus_0_l Zmult_0_r Zplus_0_r
- spec_w_W0 spec_w_0W spec_w_WW spec_w_0
- (wB_div w_digits w_to_Z spec_to_Z)
+ 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.
Ltac w_rewrite := autorewrite with w_rewrite;trivial.
@@ -303,12 +303,12 @@ Section DoubleLift.
[[p]] <= Zpos (xO w_digits) ->
[[match ww_compare p zdigits with
| Eq => w_WW xl yh
- | Lt => w_WW (w_add_mul_div (low p) xh xl)
+ | Lt => w_WW (w_add_mul_div (low p) xh xl)
(w_add_mul_div (low p) xl yh)
| Gt =>
let n := low (ww_sub p zdigits) in
w_WW (w_add_mul_div n xl yh) (w_add_mul_div n yh yl)
- end]] =
+ end]] =
([[WW xh xl]] * (2^[[p]]) +
[[WW yh yl]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB.
Proof.
@@ -317,7 +317,7 @@ Section DoubleLift.
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.
- replace (Zpos w_digits + Zpos w_digits - [[p]]) with
+ 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));
@@ -330,7 +330,7 @@ Section DoubleLift.
fold wB.
rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;rewrite <- Zplus_assoc.
rewrite <- Zpower_2.
- rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|].
+ rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|].
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xl yh)). ring.
simpl ww_to_Z; w_rewrite;zarith.
assert (HH0: [|low p|] = [[p]]).
@@ -353,7 +353,7 @@ Section DoubleLift.
rewrite Zmult_plus_distr_l.
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.
+ replace ([|xh|] * wB * 2^[[p]]) with ([|xh|] * 2^[[p]] * wB). 2:ring.
rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc.
unfold base at 5;rewrite <- Zmod_shift_r;zarith.
unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits));
@@ -387,8 +387,8 @@ Section DoubleLift.
lazy zeta; simpl ww_to_Z; w_rewrite;zarith.
repeat rewrite spec_w_add_mul_div;zarith.
rewrite HH0.
- pattern wB at 5;replace wB with
- (2^(([[p]] - Zpos w_digits)
+ 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 Z_div_plus_l;zarith.
@@ -401,28 +401,28 @@ Section DoubleLift.
repeat rewrite <- Zplus_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)
+ 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.
- replace ([|xh|] * wB * 2 ^ u) with
+ replace ([|xh|] * wB * 2 ^ u) with
([|xh|]*2^u*wB). 2:ring.
- repeat rewrite <- Zplus_assoc.
+ repeat rewrite <- Zplus_assoc.
rewrite (Zplus_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.
+ unfold u; split;zarith.
split;zarith. unfold u; apply Zdiv_lt_upper_bound;zarith.
rewrite <- Zpower_exp;zarith.
- fold u.
- ring_simplify (u + (Zpos w_digits - u)); fold
+ fold u.
+ ring_simplify (u + (Zpos w_digits - u)); fold
wB;zarith. unfold ww_digits;rewrite Zpos_xO;zarith.
unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith.
unfold u; split;zarith.
unfold u; split;zarith.
apply Zdiv_lt_upper_bound;zarith.
rewrite <- Zpower_exp;zarith.
- fold u.
+ fold u.
ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith.
unfold u;zarith.
unfold u;zarith.
@@ -446,7 +446,7 @@ 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.
- generalize (spec_ww_compare p (w_0W w_zdigits));
+ 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.
generalize H1; w_rewrite; rewrite spec_zdigits; clear H1; intros H1.
@@ -459,7 +459,7 @@ Section DoubleLift.
rewrite HH0; auto with zarith.
replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial].
intros Heq;rewrite <- Heq;clear Heq.
- generalize (spec_ww_compare p (w_0W w_zdigits));
+ 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.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
index c7d83acc..7090c76a 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleMul.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -17,7 +17,7 @@ Require Import BigNumPrelude.
Require Import DoubleType.
Require Import DoubleBase.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section DoubleMul.
Variable w : Type.
@@ -45,7 +45,7 @@ Section DoubleMul.
(* (xh*B+xl) (yh*B + yl)
xh*yh = hh = |hhh|hhl|B2
xh*yl +xl*yh = cc = |cch|ccl|B
- xl*yl = ll = |llh|lll
+ xl*yl = ll = |llh|lll
*)
Definition double_mul_c (cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w) x y :=
@@ -56,7 +56,7 @@ Section DoubleMul.
let hh := w_mul_c xh yh in
let ll := w_mul_c xl yl in
let (wc,cc) := cross xh xl yh yl hh ll in
- match cc with
+ match cc with
| W0 => WW (ww_add hh (w_W0 wc)) ll
| WW cch ccl =>
match ww_add_c (w_W0 ccl) ll with
@@ -67,8 +67,8 @@ Section DoubleMul.
end.
Definition ww_mul_c :=
- double_mul_c
- (fun xh xl yh yl hh ll=>
+ double_mul_c
+ (fun xh xl yh yl hh ll=>
match ww_add_c (w_mul_c xh yl) (w_mul_c xl yh) with
| C0 cc => (w_0, cc)
| C1 cc => (w_1, cc)
@@ -77,11 +77,11 @@ Section DoubleMul.
Definition w_2 := w_add w_1 w_1.
Definition kara_prod xh xl yh yl hh ll :=
- match ww_add_c hh ll with
+ match ww_add_c hh ll with
C0 m =>
match w_compare xl xh with
Eq => (w_0, m)
- | Lt =>
+ | Lt =>
match w_compare yl yh with
Eq => (w_0, m)
| Lt => (w_0, ww_sub m (w_mul_c (w_sub xh xl) (w_sub yh yl)))
@@ -89,7 +89,7 @@ Section DoubleMul.
C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1)
end
end
- | Gt =>
+ | Gt =>
match w_compare yl yh with
Eq => (w_0, m)
| Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with
@@ -101,17 +101,17 @@ Section DoubleMul.
| C1 m =>
match w_compare xl xh with
Eq => (w_1, m)
- | Lt =>
+ | Lt =>
match w_compare yl yh with
Eq => (w_1, m)
| Lt => match ww_sub_c m (w_mul_c (w_sub xh xl) (w_sub yh yl)) with
C0 m1 => (w_1, m1) | C1 m1 => (w_0, m1)
- end
+ end
| Gt => match ww_add_c m (w_mul_c (w_sub xh xl) (w_sub yl yh)) with
C1 m1 => (w_2, m1) | C0 m1 => (w_1, m1)
end
end
- | Gt =>
+ | Gt =>
match w_compare yl yh with
Eq => (w_1, m)
| Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with
@@ -129,8 +129,8 @@ Section DoubleMul.
Definition ww_mul x y :=
match x, y with
| W0, _ => W0
- | _, W0 => W0
- | WW xh xl, WW yh yl =>
+ | _, W0 => W0
+ | WW xh xl, WW yh yl =>
let ccl := w_add (w_mul xh yl) (w_mul xl yh) in
ww_add (w_W0 ccl) (w_mul_c xl yl)
end.
@@ -161,9 +161,9 @@ Section DoubleMul.
Variable w_mul_add : w -> w -> w -> w * w.
Fixpoint double_mul_add_n1 (n:nat) : word w n -> w -> w -> w * word w n :=
- match n return word w n -> w -> w -> w * word w n with
- | O => w_mul_add
- | S n1 =>
+ match n return word w n -> w -> w -> w * word w n with
+ | O => w_mul_add
+ | S n1 =>
let mul_add := double_mul_add_n1 n1 in
fun x y r =>
match x with
@@ -183,11 +183,11 @@ Section DoubleMul.
Variable wn_0W : wn -> zn2z wn.
Variable wn_WW : wn -> wn -> zn2z wn.
Variable w_mul_add_n1 : wn -> w -> w -> w*wn.
- Fixpoint double_mul_add_mn1 (m:nat) :
+ Fixpoint double_mul_add_mn1 (m:nat) :
word wn m -> w -> w -> w*word wn m :=
- match m return word wn m -> w -> w -> w*word wn m with
- | O => w_mul_add_n1
- | S m1 =>
+ match m return word wn m -> w -> w -> w*word wn m with
+ | O => w_mul_add_n1
+ | S m1 =>
let mul_add := double_mul_add_mn1 m1 in
fun x y r =>
match x with
@@ -207,11 +207,11 @@ Section DoubleMul.
| WW h l =>
match w_add_c l r with
| C0 lr => (h,lr)
- | C1 lr => (w_succ h, lr)
+ | C1 lr => (w_succ h, lr)
end
end.
-
+
(*Section DoubleProof. *)
Variable w_digits : positive.
Variable w_to_Z : w -> Z.
@@ -225,11 +225,11 @@ Section DoubleMul.
(interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[+[ c ]]" :=
+ (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[-[ c ]]" :=
+ (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
Notation "[|| x ||]" :=
@@ -269,8 +269,8 @@ Section DoubleMul.
forall x y, [[ww_add_carry x y]] = ([[x]] + [[y]] + 1) mod wwB.
Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB.
Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]].
-
-
+
+
Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB.
Proof. intros x;apply spec_ww_to_Z;auto. Qed.
@@ -281,21 +281,21 @@ Section DoubleMul.
Ltac zarith := auto with zarith mult.
Lemma wBwB_lex: forall a b c d,
- a * wB^2 + [[b]] <= c * wB^2 + [[d]] ->
+ a * wB^2 + [[b]] <= c * wB^2 + [[d]] ->
a <= c.
- Proof.
+ Proof.
intros a b c d H; apply beta_lex with [[b]] [[d]] (wB^2);zarith.
Qed.
- Lemma wBwB_lex_inv: forall a b c d,
- a < c ->
- a * wB^2 + [[b]] < c * wB^2 + [[d]].
+ Lemma wBwB_lex_inv: forall a b c d,
+ a < c ->
+ a * wB^2 + [[b]] < c * wB^2 + [[d]].
Proof.
intros a b c d H; apply beta_lex_inv; zarith.
Qed.
Lemma sum_mul_carry : forall xh xl yh yl wc cc,
- [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] ->
+ [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] ->
0 <= [|wc|] <= 1.
Proof.
intros.
@@ -303,14 +303,14 @@ Section DoubleMul.
apply wB_pos.
Qed.
- Theorem mult_add_ineq: forall xH yH crossH,
+ Theorem mult_add_ineq: forall xH yH crossH,
0 <= [|xH|] * [|yH|] + [|crossH|] < wwB.
Proof.
intros;rewrite wwB_wBwB;apply mult_add_ineq;zarith.
Qed.
-
+
Hint Resolve mult_add_ineq : mult.
-
+
Lemma spec_mul_aux : forall xh xl yh yl wc (cc:zn2z w) hh ll,
[[hh]] = [|xh|] * [|yh|] ->
[[ll]] = [|xl|] * [|yl|] ->
@@ -325,9 +325,9 @@ Section DoubleMul.
end||] = ([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|]).
Proof.
intros;assert (U1 := wB_pos w_digits).
- replace (([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|])) with
+ replace (([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|])) with
([|xh|]*[|yh|]*wB^2 + ([|xh|]*[|yl|] + [|xl|]*[|yh|])*wB + [|xl|]*[|yl|]).
- 2:ring. rewrite <- H1;rewrite <- H;rewrite <- H0.
+ 2:ring. rewrite <- H1;rewrite <- H;rewrite <- H0.
assert (H2 := sum_mul_carry _ _ _ _ _ _ H1).
destruct cc as [ | cch ccl]; simpl zn2z_to_Z; simpl ww_to_Z.
rewrite spec_ww_add;rewrite spec_w_W0;rewrite Zmod_small;
@@ -346,7 +346,7 @@ Section DoubleMul.
rewrite <- Zmult_plus_distr_l.
assert (((2 * wB - 4) + 2)*wB <= ([|wc|] * wB + [|cch|])*wB).
apply Zmult_le_compat;zarith.
- rewrite Zmult_plus_distr_l in H3.
+ rewrite Zmult_plus_distr_l 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;
@@ -363,8 +363,8 @@ Section DoubleMul.
(forall xh xl yh yl hh ll,
[[hh]] = [|xh|]*[|yh|] ->
[[ll]] = [|xl|]*[|yl|] ->
- let (wc,cc) := cross xh xl yh yl hh ll in
- [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]) ->
+ let (wc,cc) := cross xh xl yh yl hh ll in
+ [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]) ->
forall x y, [||double_mul_c cross x y||] = [[x]] * [[y]].
Proof.
intros cross Hcross x y;destruct x as [ |xh xl];simpl;trivial.
@@ -376,7 +376,7 @@ Section DoubleMul.
rewrite <- wwB_wBwB;trivial.
Qed.
- Lemma spec_ww_mul_c : forall x y, [||ww_mul_c x y||] = [[x]] * [[y]].
+ Lemma spec_ww_mul_c : forall x y, [||ww_mul_c x y||] = [[x]] * [[y]].
Proof.
intros x y;unfold ww_mul_c;apply spec_double_mul_c.
intros xh xl yh yl hh ll H1 H2.
@@ -402,9 +402,9 @@ Section DoubleMul.
let (wc,cc) := kara_prod xh xl yh yl hh ll in
[|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|].
Proof.
- intros xh xl yh yl hh ll H H0; rewrite <- kara_prod_aux;
+ intros xh xl yh yl hh ll H H0; rewrite <- kara_prod_aux;
rewrite <- H; rewrite <- H0; unfold kara_prod.
- assert (Hxh := (spec_to_Z xh)); assert (Hxl := (spec_to_Z xl));
+ assert (Hxh := (spec_to_Z xh)); assert (Hxl := (spec_to_Z xl));
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)).
@@ -412,7 +412,7 @@ Section DoubleMul.
try rewrite Hxlh; try rewrite spec_w_0; try (ring; fail).
generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh.
rewrite Hylh; rewrite spec_w_0; try (ring; fail).
- 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.
@@ -508,8 +508,8 @@ Section DoubleMul.
repeat rewrite Zmod_small; auto with zarith; try (ring; fail).
Qed.
- Lemma sub_carry : forall xh xl yh yl z,
- 0 <= z ->
+ Lemma sub_carry : forall xh xl yh yl z,
+ 0 <= z ->
[|xh|]*[|yl|] + [|xl|]*[|yh|] = wwB + z ->
z < wwB.
Proof.
@@ -519,7 +519,7 @@ Section DoubleMul.
generalize (Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)).
rewrite <- wwB_wBwB;intros H1 H2.
assert (H3 := wB_pos w_digits).
- assert (2*wB <= wwB).
+ assert (2*wB <= wwB).
rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat;zarith.
omega.
Qed.
@@ -528,7 +528,7 @@ Section DoubleMul.
let H:= fresh "H" in
assert (H:= spec_ww_to_Z x).
- Ltac Zmult_lt_b x y :=
+ Ltac Zmult_lt_b x y :=
let H := fresh "H" in
assert (H := Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)).
@@ -582,7 +582,7 @@ Section DoubleMul.
Variable w_mul_add : w -> w -> w -> w * w.
Variable spec_w_mul_add : forall x y r,
let (h,l):= w_mul_add x y r in
- [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|].
+ [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|].
Lemma spec_double_mul_add_n1 : forall n x y r,
let (h,l) := double_mul_add_n1 w_mul_add n x y r in
@@ -590,7 +590,7 @@ Section DoubleMul.
Proof.
induction n;intros x y r;trivial.
exact (spec_w_mul_add x y r).
- unfold double_mul_add_n1;destruct x as[ |xh xl];
+ unfold double_mul_add_n1;destruct x as[ |xh xl];
fold(double_mul_add_n1 w_mul_add).
rewrite spec_w_0;rewrite spec_extend;simpl;trivial.
assert(H:=IHn xl y r);destruct (double_mul_add_n1 w_mul_add n xl y r)as(rl,l).
@@ -599,13 +599,13 @@ Section DoubleMul.
rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc;rewrite <- H.
rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
rewrite U;ring.
- Qed.
-
+ Qed.
+
End DoubleMulAddn1Proof.
- Lemma spec_w_mul_add : forall x y r,
+ Lemma spec_w_mul_add : forall x y r,
let (h,l):= w_mul_add x y r in
- [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|].
+ [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|].
Proof.
intros x y r;unfold w_mul_add;assert (H:=spec_w_mul_c x y);
destruct (w_mul_c x y) as [ |h l];simpl;rewrite <- H.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
index 043ff351..83a2e717 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleSqrt.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -17,7 +17,7 @@ Require Import BigNumPrelude.
Require Import DoubleType.
Require Import DoubleBase.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section DoubleSqrt.
Variable w : Type.
@@ -52,7 +52,7 @@ Section DoubleSqrt.
Let wwBm1 := ww_Bm1 w_Bm1.
- Definition ww_is_even x :=
+ Definition ww_is_even x :=
match x with
| W0 => true
| WW xh xl => w_is_even xl
@@ -62,7 +62,7 @@ Section DoubleSqrt.
match w_compare x z with
| Eq =>
match w_compare y z with
- Eq => (C1 w_1, w_0)
+ Eq => (C1 w_1, w_0)
| Gt => (C1 w_1, w_sub y z)
| Lt => (C1 w_0, y)
end
@@ -120,7 +120,7 @@ Section DoubleSqrt.
let ( q, r) := w_sqrt2 x1 x2 in
let (q1, r1) := w_div2s r y1 q in
match q1 with
- C0 q1 =>
+ C0 q1 =>
let q2 := w_square_c q1 in
let a := WW q q1 in
match r1 with
@@ -132,9 +132,9 @@ Section DoubleSqrt.
| C0 r2 =>
match ww_sub_c (WW r2 y2) q2 with
C0 r3 => (a, C0 r3)
- | C1 r3 =>
+ | C1 r3 =>
let a2 := ww_add_mul_div (w_0W w_1) a W0 in
- match ww_pred_c a2 with
+ match ww_pred_c a2 with
C0 a3 =>
(ww_pred a, ww_add_c a3 r3)
| C1 a3 =>
@@ -166,20 +166,20 @@ Section DoubleSqrt.
| Gt =>
match ww_add_mul_div p x W0 with
W0 => W0
- | WW x1 x2 =>
+ | WW x1 x2 =>
let (r, _) := w_sqrt2 x1 x2 in
- WW w_0 (w_add_mul_div
- (w_sub w_zdigits
+ WW w_0 (w_add_mul_div
+ (w_sub w_zdigits
(low (ww_add_mul_div (ww_pred ww_zdigits)
W0 p))) w_0 r)
end
- | _ =>
+ | _ =>
match x with
W0 => W0
| WW x1 x2 => WW w_0 (fst (w_sqrt2 x1 x2))
end
end.
-
+
Variable w_to_Z : w -> Z.
@@ -192,11 +192,11 @@ Section DoubleSqrt.
(interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[+[ c ]]" :=
+ (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[-[ c ]]" :=
+ (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
Notation "[|| x ||]" :=
@@ -269,14 +269,12 @@ Section DoubleSqrt.
Let spec_ww_Bm1 : [[wwBm1]] = wwB - 1.
Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed.
-
- Hint Rewrite spec_w_0 spec_w_1 w_Bm1 spec_w_WW spec_w_sub
- spec_w_div21 spec_w_add_mul_div spec_ww_Bm1
- spec_w_add_c spec_w_sqrt2: w_rewrite.
+ Hint Rewrite spec_w_0 spec_w_1 spec_w_WW spec_w_sub
+ spec_w_add_mul_div spec_ww_Bm1 spec_w_add_c : w_rewrite.
Lemma spec_ww_is_even : forall x,
if ww_is_even x then [[x]] mod 2 = 0 else [[x]] mod 2 = 1.
-clear spec_more_than_1_digit.
+clear spec_more_than_1_digit.
intros x; case x; simpl ww_is_even.
simpl.
rewrite Zmod_small; auto with zarith.
@@ -379,8 +377,8 @@ intros x; case x; simpl ww_is_even.
end.
rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith.
destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
- split; auto with zarith.
- apply Zdiv_lt_upper_bound; auto with zarith.
+ split; auto with zarith.
+ apply Zdiv_lt_upper_bound; auto with zarith.
rewrite Hp; ring.
Qed.
@@ -402,7 +400,7 @@ intros x; case x; simpl ww_is_even.
rewrite Zmax_right; auto with zarith.
rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith.
destruct (spec_to_Z w1) as [H1 H2];auto with zarith.
- split; auto with zarith.
+ split; auto with zarith.
unfold base.
match goal with |- _ < _ ^ ?X =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
@@ -434,7 +432,7 @@ intros x; case x; simpl ww_is_even.
intros w1.
rewrite spec_ww_add_mul_div; auto with zarith.
autorewrite with w_rewrite rm10.
- rewrite spec_w_0W; rewrite spec_w_1.
+ rewrite spec_w_0W; rewrite spec_w_1.
rewrite Zpower_1_r; auto with zarith.
rewrite Zmult_comm; auto.
rewrite spec_w_0W; rewrite spec_w_1; auto with zarith.
@@ -458,7 +456,7 @@ intros x; case x; simpl ww_is_even.
match goal with |- 0 <= ?X - 1 =>
assert (0 < X); auto with zarith
end.
- apply Zpower_gt_0; auto with zarith.
+ apply Zpower_gt_0; auto with zarith.
match goal with |- 0 <= ?X - 1 =>
assert (0 < X); auto with zarith; red; reflexivity
end.
@@ -542,7 +540,7 @@ intros x; case x; simpl ww_is_even.
rewrite add_mult_div_2_plus_1; unfold base.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
+ rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
@@ -559,7 +557,7 @@ intros x; case x; simpl ww_is_even.
unfold base.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
+ rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
@@ -592,7 +590,7 @@ intros x; case x; simpl ww_is_even.
rewrite H1; unfold base.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
+ rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
@@ -611,7 +609,7 @@ intros x; case x; simpl ww_is_even.
rewrite H1; unfold base.
match goal with |- context[_ ^ ?X] =>
assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith;
- rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
+ rewrite <- (tmp X); clear tmp; rewrite Zpower_exp;
try rewrite Zpower_1_r; auto with zarith
end.
rewrite Zpos_minus; auto with zarith.
@@ -682,7 +680,7 @@ intros x; case x; simpl ww_is_even.
rewrite Zsquare_mult; replace (p * p) with ((- p) * (- p)); try ring.
apply Zmult_le_0_compat; auto with zarith.
Qed.
-
+
Lemma spec_split: forall x,
[|fst (split x)|] * wB + [|snd (split x)|] = [[x]].
intros x; case x; simpl; autorewrite with w_rewrite;
@@ -751,7 +749,7 @@ intros x; case x; simpl ww_is_even.
match goal with |- ?X <= ?Y =>
replace Y with (2 * (wB/ 2 - 1)); auto with zarith
end.
- pattern wB at 2; rewrite <- wB_div_2; auto with zarith.
+ pattern wB at 2; rewrite <- wB_div_2; auto with zarith.
match type of H1 with ?X = _ =>
assert (U5: X < wB / 4 * wB)
end.
@@ -764,9 +762,9 @@ intros x; case x; simpl ww_is_even.
destruct (spec_to_Z w3);auto with zarith.
generalize (@spec_w_div2s c w0 w4 U1 H2).
case (w_div2s c w0 w4).
- intros c0; case c0; intros w5;
+ intros c0; case c0; intros w5;
repeat (rewrite C0_id || rewrite C1_plus_wB).
- intros c1; case c1; intros w6;
+ intros c1; case c1; intros w6;
repeat (rewrite C0_id || rewrite C1_plus_wB).
intros (H3, H4).
match goal with |- context [ww_sub_c ?y ?z] =>
@@ -1038,7 +1036,7 @@ intros x; case x; simpl ww_is_even.
end.
apply Zle_not_lt; rewrite <- H3; auto with zarith.
rewrite Zmult_plus_distr_l.
- apply Zlt_le_trans with ((2 * [|w4|]) * wB + 0);
+ apply Zlt_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.
@@ -1119,9 +1117,9 @@ intros x; case x; simpl ww_is_even.
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.
+ Qed.
+
+ Lemma wwB_4_2: 2 * (wwB / 4) = wwB/ 2.
pattern wwB at 1; rewrite wwB_wBwB; rewrite Zpower_2.
rewrite <- wB_div_2.
match goal with |- context[(2 * ?X) * (2 * ?Z)] =>
@@ -1134,7 +1132,7 @@ intros x; case x; simpl ww_is_even.
Lemma spec_ww_head1
- : forall x : zn2z w,
+ : forall x : zn2z w,
(ww_is_even (ww_head1 x) = true) /\
(0 < [[x]] -> wwB / 4 <= 2 ^ [[ww_head1 x]] * [[x]] < wwB).
assert (U := wB_pos w_digits).
@@ -1167,7 +1165,7 @@ intros x; case x; simpl ww_is_even.
rewrite Hp.
rewrite Zminus_mod; auto with zarith.
rewrite H2; repeat rewrite Zmod_small; auto with zarith.
- intros H3; rewrite Hp.
+ intros H3; rewrite Hp.
case (spec_ww_head0 x); auto; intros Hv3 Hv4.
assert (Hu: forall u, 0 < u -> 2 * 2 ^ (u - 1) = 2 ^u).
intros u Hu.
@@ -1189,7 +1187,7 @@ intros x; case x; simpl ww_is_even.
apply sym_equal; apply Zdiv_unique with 0;
auto with zarith.
rewrite Zmult_assoc; rewrite wB_div_4; auto with zarith.
- rewrite wwB_wBwB; ring.
+ rewrite wwB_wBwB; ring.
Qed.
Lemma spec_ww_sqrt : forall x,
@@ -1198,14 +1196,14 @@ intros x; case x; simpl ww_is_even.
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;
- auto with zarith.
+ auto with zarith.
intros H1.
generalize (spec_ww_compare (ww_head1 x) W0); case ww_compare;
simpl ww_to_Z; autorewrite with rm10.
generalize H1; case x.
intros HH; contradict HH; simpl ww_to_Z; auto with zarith.
intros w0 w1; simpl ww_to_Z; autorewrite with w_rewrite rm10.
- intros H2; case (spec_ww_head1 (WW w0 w1)); intros H3 H4 H5.
+ intros H2; case (spec_ww_head1 (WW w0 w1)); intros H3 H4 H5.
generalize (H4 H2); clear H4; rewrite H5; clear H5; autorewrite with rm10.
intros (H4, H5).
assert (V: wB/4 <= [|w0|]).
@@ -1241,7 +1239,7 @@ intros x; case x; simpl ww_is_even.
apply Zle_not_lt; unfold base.
apply Zle_trans with (2 ^ [[ww_head1 x]]).
apply Zpower_le_monotone; auto with zarith.
- pattern (2 ^ [[ww_head1 x]]) at 1;
+ pattern (2 ^ [[ww_head1 x]]) at 1;
rewrite <- (Zmult_1_r (2 ^ [[ww_head1 x]])).
apply Zmult_le_compat_l; auto with zarith.
generalize (spec_ww_add_mul_div x W0 (ww_head1 x) Hv2);
@@ -1283,13 +1281,13 @@ intros x; case x; simpl ww_is_even.
rewrite Zmod_small; auto with zarith.
split; auto with zarith.
apply Zlt_le_trans with (Zpos (xO w_digits)); auto with zarith.
- unfold base; apply Zpower2_le_lin; auto with zarith.
+ 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.
- unfold base; apply Zpower2_lt_lin; auto with zarith.
+ 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).
rewrite spec_ww_add_mul_div.
@@ -1330,14 +1328,14 @@ intros x; case x; simpl ww_is_even.
rewrite tmp; clear tmp.
apply Zpower_le_monotone3; auto with zarith.
split; auto with zarith.
- pattern [|w2|] at 2;
+ pattern [|w2|] at 2;
rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]] / 2)));
auto with zarith.
match goal with |- ?X <= ?X + ?Y =>
assert (0 <= Y); auto with zarith
end.
case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]] / 2))); auto with zarith.
- case c; unfold interp_carry; autorewrite with rm10;
+ 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.
rewrite H4.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
index 269d62bb..a7e55671 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleSub.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
@@ -17,7 +17,7 @@ Require Import BigNumPrelude.
Require Import DoubleType.
Require Import DoubleBase.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section DoubleSub.
Variable w : Type.
@@ -39,7 +39,7 @@ Section DoubleSub.
Definition ww_opp_c x :=
match x with
| W0 => C0 W0
- | WW xh xl =>
+ | WW xh xl =>
match w_opp_c xl with
| C0 _ =>
match w_opp_c xh with
@@ -53,7 +53,7 @@ Section DoubleSub.
Definition ww_opp x :=
match x with
| W0 => W0
- | WW xh xl =>
+ | WW xh xl =>
match w_opp_c xl with
| C0 _ => WW (w_opp xh) w_0
| C1 l => WW (w_opp_carry xh) l
@@ -72,14 +72,14 @@ Section DoubleSub.
| WW xh xl =>
match w_pred_c xl with
| C0 l => C0 (w_WW xh l)
- | C1 _ =>
- match w_pred_c xh with
+ | C1 _ =>
+ match w_pred_c xh with
| C0 h => C0 (WW h w_Bm1)
| C1 _ => C1 ww_Bm1
end
end
end.
-
+
Definition ww_pred x :=
match x with
| W0 => ww_Bm1
@@ -89,19 +89,19 @@ Section DoubleSub.
| C1 l => WW (w_pred xh) w_Bm1
end
end.
-
+
Definition ww_sub_c x y :=
match y, x with
| W0, _ => C0 x
| WW yh yl, W0 => ww_opp_c (WW yh yl)
| WW yh yl, WW xh xl =>
match w_sub_c xl yl with
- | C0 l =>
+ | C0 l =>
match w_sub_c xh yh with
| C0 h => C0 (w_WW h l)
| C1 h => C1 (WW h l)
end
- | C1 l =>
+ | C1 l =>
match w_sub_carry_c xh yh with
| C0 h => C0 (WW h l)
| C1 h => C1 (WW h l)
@@ -109,7 +109,7 @@ Section DoubleSub.
end
end.
- Definition ww_sub x y :=
+ Definition ww_sub x y :=
match y, x with
| W0, _ => x
| WW yh yl, W0 => ww_opp (WW yh yl)
@@ -127,7 +127,7 @@ Section DoubleSub.
| WW yh yl, W0 => C1 (ww_opp_carry (WW yh yl))
| WW yh yl, WW xh xl =>
match w_sub_carry_c xl yl with
- | C0 l =>
+ | C0 l =>
match w_sub_c xh yh with
| C0 h => C0 (w_WW h l)
| C1 h => C1 (WW h l)
@@ -155,7 +155,7 @@ Section DoubleSub.
(*Section DoubleProof.*)
Variable w_digits : positive.
Variable w_to_Z : w -> Z.
-
+
Notation wB := (base w_digits).
Notation wwB := (base (ww_digits w_digits)).
@@ -166,13 +166,13 @@ Section DoubleSub.
(interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
- Notation "[+[ c ]]" :=
- (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[+[ c ]]" :=
+ (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
- Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
+ Notation "[-[ c ]]" :=
+ (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
(at level 0, x at level 99).
-
+
Variable spec_w_0 : [|w_0|] = 0.
Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1.
Variable spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1.
@@ -187,7 +187,7 @@ Section DoubleSub.
Variable spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|].
Variable spec_sub_carry_c :
forall x y, [-|w_sub_carry_c x y|] = [|x|] - [|y|] - 1.
-
+
Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB.
Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB.
Variable spec_sub_carry :
@@ -197,12 +197,12 @@ 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 Zopp_plus_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 Zopp_mult_distr_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)
+ rewrite H0;generalize (spec_opp_c xh);destruct (w_opp_c xh)
as [h|h];intros H1;unfold interp_carry in *;rewrite <- H1.
assert ([|h|] = 0).
assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega.
@@ -216,7 +216,7 @@ Section DoubleSub.
Proof.
destruct x as [ |xh xl];simpl. reflexivity.
rewrite Zopp_plus_distr;rewrite Zopp_mult_distr_l.
- generalize (spec_opp_c xl);destruct (w_opp_c xl)
+ 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.
assert ([|l|] = 0).
@@ -247,7 +247,7 @@ Section DoubleSub.
assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega.
rewrite H0;change ([|xh|] + -1) with ([|xh|] - 1).
generalize (spec_pred_c xh);destruct (w_pred_c xh) as [h|h];
- intros H1;unfold interp_carry in H1;rewrite <- H1.
+ intros H1;unfold interp_carry in H1;rewrite <- H1.
simpl;rewrite spec_w_Bm1;ring.
assert ([|h|] = wB - 1).
assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega.
@@ -258,14 +258,14 @@ Section DoubleSub.
Proof.
destruct y as [ |yh yl];simpl. ring.
destruct x as [ |xh xl];simpl. exact (spec_ww_opp_c (WW yh yl)).
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]))
+ replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]))
with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|])). 2:ring.
generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl) as [l|l];intros H;
unfold interp_carry in H;rewrite <- H.
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 Zplus_assoc;rewrite <- Zmult_plus_distr_l.
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;
@@ -275,37 +275,37 @@ Section DoubleSub.
Lemma spec_ww_sub_carry_c :
forall x y, [-[ww_sub_carry_c x y]] = [[x]] - [[y]] - 1.
Proof.
- destruct y as [ |yh yl];simpl.
+ destruct y as [ |yh yl];simpl.
unfold Zminus;simpl;rewrite Zplus_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.
simpl ww_to_Z.
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1)
+ replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1)
with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|]-1)). 2:ring.
- generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl)
+ 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.
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 Zplus_assoc;rewrite <- Zmult_plus_distr_l.
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;
simpl ww_to_Z; try rewrite wwB_wBwB;ring.
- Qed.
-
+ Qed.
+
Lemma spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB.
Proof.
- destruct x as [ |xh xl];simpl.
+ destruct x as [ |xh xl];simpl.
apply Zmod_unique with (-1). apply spec_ww_to_Z;trivial.
rewrite spec_ww_Bm1;ring.
replace ([|xh|]*wB + [|xl|] - 1) with ([|xh|]*wB + ([|xl|] - 1)). 2:ring.
generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l];intro H;
unfold interp_carry in H;rewrite <- H;simpl ww_to_Z.
- rewrite Zmod_small. apply spec_w_WW.
+ rewrite Zmod_small. apply spec_w_WW.
exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh l)).
- rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
+ rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l.
change ([|xh|] + -1) with ([|xh|] - 1).
assert ([|l|] = wB - 1).
assert (H1:= spec_to_Z l);assert (H2:= spec_to_Z xl);omega.
@@ -318,7 +318,7 @@ Section DoubleSub.
destruct y as [ |yh yl];simpl.
ring_simplify ([[x]] - 0);rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
destruct x as [ |xh xl];simpl. exact (spec_ww_opp (WW yh yl)).
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]))
+ replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]))
with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|])). 2:ring.
generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl)as[l|l];intros H;
unfold interp_carry in H;rewrite <- H.
@@ -338,7 +338,7 @@ Section DoubleSub.
apply spec_ww_to_Z;trivial.
fold (ww_opp_carry (WW yh yl)).
rewrite (spec_ww_opp_carry (WW yh yl));simpl ww_to_Z;ring.
- replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1)
+ replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1)
with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|] - 1)). 2:ring.
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.
@@ -354,4 +354,4 @@ End DoubleSub.
-
+
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
index 28d40094..88cbb484 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
@@ -8,12 +8,12 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: DoubleType.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Set Implicit Arguments.
Require Import ZArith.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Definition base digits := Zpower 2 (Zpos digits).
@@ -37,10 +37,10 @@ Section Zn2Z.
Variable znz : Type.
- (** From a type [znz] representing a cyclic structure Z/nZ,
+ (** From a type [znz] representing a cyclic structure Z/nZ,
we produce a representation of Z/2nZ by pairs of elements of [znz]
- (plus a special case for zero). High half of the new number comes
- first.
+ (plus a special case for zero). High half of the new number comes
+ first.
*)
Inductive zn2z :=
@@ -57,10 +57,10 @@ End Zn2Z.
Implicit Arguments W0 [znz].
-(** From a cyclic representation [w], we iterate the [zn2z] construct
- [n] times, gaining the type of binary trees of depth at most [n],
- whose leafs are either W0 (if depth < n) or elements of w
- (if depth = n).
+(** From a cyclic representation [w], we iterate the [zn2z] construct
+ [n] times, gaining the type of binary trees of depth at most [n],
+ whose leafs are either W0 (if depth < n) or elements of w
+ (if depth = n).
*)
Fixpoint word (w:Type) (n:nat) : Type :=
diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v
index 6da1c6ec..8addf5b9 100644
--- a/theories/Numbers/Cyclic/Int31/Cyclic31.v
+++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: Cyclic31.v 11907 2009-02-10 23:54:28Z letouzey $ i*)
+(*i $Id$ i*)
(** * Int31 numbers defines indeed a cyclic structure : Z/(2^31)Z *)
@@ -24,8 +24,8 @@ Require Import BigNumPrelude.
Require Import CyclicAxioms.
Require Import ROmega.
-Open Scope nat_scope.
-Open Scope int31_scope.
+Local Open Scope nat_scope.
+Local Open Scope int31_scope.
Section Basics.
@@ -34,9 +34,9 @@ Section Basics.
Lemma iszero_eq0 : forall x, iszero x = true -> x=0.
Proof.
destruct x; simpl; intros.
- repeat
- match goal with H:(if ?d then _ else _) = true |- _ =>
- destruct d; try discriminate
+ repeat
+ match goal with H:(if ?d then _ else _) = true |- _ =>
+ destruct d; try discriminate
end.
reflexivity.
Qed.
@@ -46,26 +46,26 @@ Section Basics.
intros x H Eq; rewrite Eq in H; simpl in *; discriminate.
Qed.
- Lemma sneakl_shiftr : forall x,
+ Lemma sneakl_shiftr : forall x,
x = sneakl (firstr x) (shiftr x).
Proof.
destruct x; simpl; auto.
Qed.
- Lemma sneakr_shiftl : forall x,
+ Lemma sneakr_shiftl : forall x,
x = sneakr (firstl x) (shiftl x).
Proof.
destruct x; simpl; auto.
Qed.
- Lemma twice_zero : forall x,
+ Lemma twice_zero : forall x,
twice x = 0 <-> twice_plus_one x = 1.
Proof.
- destruct x; simpl in *; split;
+ destruct x; simpl in *; split;
intro H; injection H; intros; subst; auto.
Qed.
- Lemma twice_or_twice_plus_one : forall x,
+ Lemma twice_or_twice_plus_one : forall x,
x = twice (shiftr x) \/ x = twice_plus_one (shiftr x).
Proof.
intros; case_eq (firstr x); intros.
@@ -79,13 +79,13 @@ Section Basics.
Definition nshiftr n x := iter_nat n _ shiftr x.
- Lemma nshiftr_S :
+ Lemma nshiftr_S :
forall n x, nshiftr (S n) x = shiftr (nshiftr n x).
Proof.
reflexivity.
Qed.
- Lemma nshiftr_S_tail :
+ Lemma nshiftr_S_tail :
forall n x, nshiftr (S n) x = nshiftr n (shiftr x).
Proof.
induction n; simpl; auto.
@@ -103,7 +103,7 @@ Section Basics.
destruct x; simpl; auto.
Qed.
- Lemma nshiftr_above_size : forall k x, size<=k ->
+ Lemma nshiftr_above_size : forall k x, size<=k ->
nshiftr k x = 0.
Proof.
intros.
@@ -117,13 +117,13 @@ Section Basics.
Definition nshiftl n x := iter_nat n _ shiftl x.
- Lemma nshiftl_S :
+ Lemma nshiftl_S :
forall n x, nshiftl (S n) x = shiftl (nshiftl n x).
Proof.
reflexivity.
Qed.
- Lemma nshiftl_S_tail :
+ Lemma nshiftl_S_tail :
forall n x, nshiftl (S n) x = nshiftl n (shiftl x).
Proof.
induction n; simpl; auto.
@@ -141,7 +141,7 @@ Section Basics.
destruct x; simpl; auto.
Qed.
- Lemma nshiftl_above_size : forall k x, size<=k ->
+ Lemma nshiftl_above_size : forall k x, size<=k ->
nshiftl k x = 0.
Proof.
intros.
@@ -151,27 +151,27 @@ Section Basics.
simpl; rewrite nshiftl_S, IHn; auto.
Qed.
- Lemma firstr_firstl :
+ Lemma firstr_firstl :
forall x, firstr x = firstl (nshiftl (pred size) x).
Proof.
destruct x; simpl; auto.
Qed.
- Lemma firstl_firstr :
+ Lemma firstl_firstr :
forall x, firstl x = firstr (nshiftr (pred size) x).
Proof.
destruct x; simpl; auto.
Qed.
-
+
(** More advanced results about [nshiftr] *)
- Lemma nshiftr_predsize_0_firstl : forall x,
+ Lemma nshiftr_predsize_0_firstl : forall x,
nshiftr (pred size) x = 0 -> firstl x = D0.
Proof.
destruct x; compute; intros H; injection H; intros; subst; auto.
Qed.
- Lemma nshiftr_0_propagates : forall n p x, n <= p ->
+ Lemma nshiftr_0_propagates : forall n p x, n <= p ->
nshiftr n x = 0 -> nshiftr p x = 0.
Proof.
intros.
@@ -181,7 +181,7 @@ Section Basics.
simpl; rewrite nshiftr_S; rewrite IHn0; auto.
Qed.
- Lemma nshiftr_0_firstl : forall n x, n < size ->
+ Lemma nshiftr_0_firstl : forall n x, n < size ->
nshiftr n x = 0 -> firstl x = D0.
Proof.
intros.
@@ -194,8 +194,8 @@ Section Basics.
(** Not used for the moment. Are they really useful ? *)
Lemma int31_ind_sneakl : forall P : int31->Prop,
- P 0 ->
- (forall x d, P x -> P (sneakl d x)) ->
+ P 0 ->
+ (forall x d, P x -> P (sneakl d x)) ->
forall x, P x.
Proof.
intros.
@@ -210,10 +210,10 @@ Section Basics.
change x with (nshiftr (size-size) x); auto.
Qed.
- Lemma int31_ind_twice : forall P : int31->Prop,
- P 0 ->
- (forall x, P x -> P (twice x)) ->
- (forall x, P x -> P (twice_plus_one x)) ->
+ Lemma int31_ind_twice : forall P : int31->Prop,
+ P 0 ->
+ (forall x, P x -> P (twice x)) ->
+ (forall x, P x -> P (twice_plus_one x)) ->
forall x, P x.
Proof.
induction x using int31_ind_sneakl; auto.
@@ -224,21 +224,21 @@ Section Basics.
(** * Some generic results about [recr] *)
Section Recr.
-
+
(** [recr] satisfies the fixpoint equation used for its definition. *)
Variable (A:Type)(case0:A)(caserec:digits->int31->A->A).
-
- Lemma recr_aux_eqn : forall n x, iszero x = false ->
- recr_aux (S n) A case0 caserec x =
+
+ Lemma recr_aux_eqn : forall n x, iszero x = false ->
+ recr_aux (S n) A case0 caserec x =
caserec (firstr x) (shiftr x) (recr_aux n A case0 caserec (shiftr x)).
Proof.
intros; simpl; rewrite H; auto.
Qed.
- Lemma recr_aux_converges :
+ Lemma recr_aux_converges :
forall n p x, n <= size -> n <= p ->
- recr_aux n A case0 caserec (nshiftr (size - n) x) =
+ recr_aux n A case0 caserec (nshiftr (size - n) x) =
recr_aux p A case0 caserec (nshiftr (size - n) x).
Proof.
induction n.
@@ -255,8 +255,8 @@ Section Basics.
apply IHn; auto with arith.
Qed.
- Lemma recr_eqn : forall x, iszero x = false ->
- recr A case0 caserec x =
+ Lemma recr_eqn : forall x, iszero x = false ->
+ recr A case0 caserec x =
caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x)).
Proof.
intros.
@@ -265,11 +265,11 @@ Section Basics.
rewrite (recr_aux_converges size (S size)); auto with arith.
rewrite recr_aux_eqn; auto.
Qed.
-
- (** [recr] is usually equivalent to a variant [recrbis]
+
+ (** [recr] is usually equivalent to a variant [recrbis]
written without [iszero] check. *)
- Fixpoint recrbis_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A)
+ Fixpoint recrbis_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A)
(i:int31) : A :=
match n with
| O => case0
@@ -277,7 +277,7 @@ Section Basics.
let si := shiftr i in
caserec (firstr i) si (recrbis_aux next A case0 caserec si)
end.
-
+
Definition recrbis := recrbis_aux size.
Hypothesis case0_caserec : caserec D0 0 case0 = case0.
@@ -291,8 +291,8 @@ Section Basics.
replace (recrbis_aux n A case0 caserec 0) with case0; auto.
clear H IHn; induction n; simpl; congruence.
Qed.
-
- Lemma recrbis_equiv : forall x,
+
+ Lemma recrbis_equiv : forall x,
recrbis A case0 caserec x = recr A case0 caserec x.
Proof.
intros; apply recrbis_aux_equiv; auto.
@@ -348,7 +348,7 @@ Section Basics.
rewrite incr_eqn1; destruct x; simpl; auto.
Qed.
- Lemma incr_twice_plus_one_firstl :
+ Lemma incr_twice_plus_one_firstl :
forall x, firstl x = D0 -> incr (twice_plus_one x) = twice (incr x).
Proof.
intros.
@@ -356,9 +356,9 @@ Section Basics.
f_equal; f_equal.
destruct x; simpl in *; rewrite H; auto.
Qed.
-
- (** The previous result is actually true even without the
- constraint on [firstl], but this is harder to prove
+
+ (** The previous result is actually true even without the
+ constraint on [firstl], but this is harder to prove
(see later). *)
End Incr.
@@ -369,9 +369,9 @@ Section Basics.
(** Variant of [phi] via [recrbis] *)
- Let Phi := fun b (_:int31) =>
+ Let Phi := fun b (_:int31) =>
match b with D0 => Zdouble | D1 => Zdouble_plus_one end.
-
+
Definition phibis_aux n x := recrbis_aux n _ Z0 Phi x.
Lemma phibis_aux_equiv : forall x, phibis_aux size x = phi x.
@@ -382,7 +382,7 @@ Section Basics.
(** Recursive equations satisfied by [phi] *)
- Lemma phi_eqn1 : forall x, firstr x = D0 ->
+ Lemma phi_eqn1 : forall x, firstr x = D0 ->
phi x = Zdouble (phi (shiftr x)).
Proof.
intros.
@@ -392,7 +392,7 @@ Section Basics.
rewrite H; auto.
Qed.
- Lemma phi_eqn2 : forall x, firstr x = D1 ->
+ Lemma phi_eqn2 : forall x, firstr x = D1 ->
phi x = Zdouble_plus_one (phi (shiftr x)).
Proof.
intros.
@@ -402,7 +402,7 @@ Section Basics.
rewrite H; auto.
Qed.
- Lemma phi_twice_firstl : forall x, firstl x = D0 ->
+ Lemma phi_twice_firstl : forall x, firstl x = D0 ->
phi (twice x) = Zdouble (phi x).
Proof.
intros.
@@ -411,7 +411,7 @@ Section Basics.
destruct x; simpl in *; rewrite H; auto.
Qed.
- Lemma phi_twice_plus_one_firstl : forall x, firstl x = D0 ->
+ Lemma phi_twice_plus_one_firstl : forall x, firstl x = D0 ->
phi (twice_plus_one x) = Zdouble_plus_one (phi x).
Proof.
intros.
@@ -427,23 +427,23 @@ Section Basics.
Lemma phibis_aux_pos : forall n x, (0 <= phibis_aux n x)%Z.
Proof.
induction n.
- simpl; unfold phibis_aux; simpl; auto with zarith.
+ simpl; unfold phibis_aux; simpl; auto with zarith.
intros.
- unfold phibis_aux, recrbis_aux; fold recrbis_aux;
+ 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.
Qed.
- Lemma phibis_aux_bounded :
- forall n x, n <= size ->
+ Lemma phibis_aux_bounded :
+ forall n x, n <= size ->
(phibis_aux n (nshiftr (size-n) x) < 2 ^ (Z_of_nat n))%Z.
Proof.
induction n.
simpl; unfold phibis_aux; simpl; auto with zarith.
intros.
- unfold phibis_aux, recrbis_aux; fold recrbis_aux;
+ unfold phibis_aux, recrbis_aux; fold recrbis_aux;
fold (phibis_aux n (shiftr (nshiftr (size - S n) x))).
assert (shiftr (nshiftr (size - S n) x) = nshiftr (size-n) x).
replace (size - n)%nat with (S (size - (S n))) by omega.
@@ -468,8 +468,8 @@ Section Basics.
apply phibis_aux_bounded; auto.
Qed.
- Lemma phibis_aux_lowerbound :
- forall n x, firstr (nshiftr n x) = D1 ->
+ Lemma phibis_aux_lowerbound :
+ forall n x, firstr (nshiftr n x) = D1 ->
(2 ^ Z_of_nat n <= phibis_aux (S n) x)%Z.
Proof.
induction n.
@@ -480,7 +480,7 @@ Section Basics.
intros.
remember (S n) as m.
- unfold phibis_aux, recrbis_aux; fold recrbis_aux;
+ unfold phibis_aux, recrbis_aux; fold recrbis_aux;
fold (phibis_aux m (shiftr x)).
subst m.
rewrite inj_S, Zpower_Zsucc; auto with zarith.
@@ -488,13 +488,13 @@ Section Basics.
apply IHn.
rewrite <- nshiftr_S_tail; auto.
destruct (firstr x).
- change (Zdouble (phibis_aux (S n) (shiftr x))) with
+ change (Zdouble (phibis_aux (S n) (shiftr x))) with
(2*(phibis_aux (S n) (shiftr x)))%Z.
omega.
rewrite Zdouble_plus_one_mult; omega.
Qed.
- Lemma phi_lowerbound :
+ Lemma phi_lowerbound :
forall x, firstl x = D1 -> (2^(Z_of_nat (pred size)) <= phi x)%Z.
Proof.
intros.
@@ -508,9 +508,9 @@ Section Basics.
Section EqShiftL.
- (** After killing [n] bits at the left, are the numbers equal ?*)
+ (** After killing [n] bits at the left, are the numbers equal ?*)
- Definition EqShiftL n x y :=
+ Definition EqShiftL n x y :=
nshiftl n x = nshiftl n y.
Lemma EqShiftL_zero : forall x y, EqShiftL O x y <-> x = y.
@@ -523,7 +523,7 @@ Section Basics.
red; intros; rewrite 2 nshiftl_above_size; auto.
Qed.
- Lemma EqShiftL_le : forall k k' x y, k <= k' ->
+ Lemma EqShiftL_le : forall k k' x y, k <= k' ->
EqShiftL k x y -> EqShiftL k' x y.
Proof.
unfold EqShiftL; intros.
@@ -534,18 +534,18 @@ Section Basics.
rewrite 2 nshiftl_S; f_equal; auto.
Qed.
- Lemma EqShiftL_firstr : forall k x y, k < size ->
+ Lemma EqShiftL_firstr : forall k x y, k < size ->
EqShiftL k x y -> firstr x = firstr y.
Proof.
intros.
rewrite 2 firstr_firstl.
f_equal.
- apply EqShiftL_le with k; auto.
+ apply EqShiftL_le with k; auto.
unfold size.
auto with arith.
Qed.
- Lemma EqShiftL_twice : forall k x y,
+ Lemma EqShiftL_twice : forall k x y,
EqShiftL k (twice x) (twice y) <-> EqShiftL (S k) x y.
Proof.
intros; unfold EqShiftL.
@@ -553,7 +553,7 @@ Section Basics.
Qed.
(** * From int31 to list of digits. *)
-
+
(** Lower (=rightmost) bits comes first. *)
Definition i2l := recrbis _ nil (fun d _ rec => d::rec).
@@ -561,10 +561,10 @@ Section Basics.
Lemma i2l_length : forall x, length (i2l x) = size.
Proof.
intros; reflexivity.
- Qed.
+ Qed.
- Fixpoint lshiftl l x :=
- match l with
+ Fixpoint lshiftl l x :=
+ match l with
| nil => x
| d::l => sneakl d (lshiftl l x)
end.
@@ -576,19 +576,19 @@ Section Basics.
destruct x; compute; auto.
Qed.
- Lemma i2l_sneakr : forall x d,
+ Lemma i2l_sneakr : forall x d,
i2l (sneakr d x) = tail (i2l x) ++ d::nil.
Proof.
destruct x; compute; auto.
Qed.
- Lemma i2l_sneakl : forall x d,
+ Lemma i2l_sneakl : forall x d,
i2l (sneakl d x) = d :: removelast (i2l x).
Proof.
destruct x; compute; auto.
Qed.
- Lemma i2l_l2i : forall l, length l = size ->
+ Lemma i2l_l2i : forall l, length l = size ->
i2l (l2i l) = l.
Proof.
repeat (destruct l as [ |? l]; [intros; discriminate | ]).
@@ -596,9 +596,9 @@ Section Basics.
intros _; compute; auto.
Qed.
- Fixpoint cstlist (A:Type)(a:A) n :=
- match n with
- | O => nil
+ Fixpoint cstlist (A:Type)(a:A) n :=
+ match n with
+ | O => nil
| S n => a::cstlist _ a n
end.
@@ -612,7 +612,7 @@ Section Basics.
induction (i2l x); simpl; f_equal; auto.
rewrite H0; clear H0.
reflexivity.
-
+
intros.
rewrite nshiftl_S.
unfold shiftl; rewrite i2l_sneakl.
@@ -657,10 +657,10 @@ Section Basics.
f_equal; auto.
Qed.
- (** This equivalence allows to prove easily the following delicate
+ (** This equivalence allows to prove easily the following delicate
result *)
- Lemma EqShiftL_twice_plus_one : forall k x y,
+ Lemma EqShiftL_twice_plus_one : forall k x y,
EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y.
Proof.
intros.
@@ -683,7 +683,7 @@ Section Basics.
subst lx n; rewrite i2l_length; omega.
Qed.
- Lemma EqShiftL_shiftr : forall k x y, EqShiftL k x y ->
+ Lemma EqShiftL_shiftr : forall k x y, EqShiftL k x y ->
EqShiftL (S k) (shiftr x) (shiftr y).
Proof.
intros.
@@ -704,41 +704,41 @@ Section Basics.
omega.
Qed.
- Lemma EqShiftL_incrbis : forall n k x y, n<=size ->
+ Lemma EqShiftL_incrbis : forall n k x y, n<=size ->
(n+k=S size)%nat ->
- EqShiftL k x y ->
+ EqShiftL k x y ->
EqShiftL k (incrbis_aux n x) (incrbis_aux n y).
Proof.
induction n; simpl; intros.
red; auto.
- destruct (eq_nat_dec k size).
+ destruct (eq_nat_dec k size).
subst k; apply EqShiftL_size; auto.
- unfold incrbis_aux; simpl;
+ unfold incrbis_aux; simpl;
fold (incrbis_aux n (shiftr x)); fold (incrbis_aux n (shiftr y)).
rewrite (EqShiftL_firstr k x y); auto; try omega.
case_eq (firstr y); intros.
rewrite EqShiftL_twice_plus_one.
apply EqShiftL_shiftr; auto.
-
+
rewrite EqShiftL_twice.
apply IHn; try omega.
apply EqShiftL_shiftr; auto.
Qed.
- Lemma EqShiftL_incr : forall x y,
+ Lemma EqShiftL_incr : forall x y,
EqShiftL 1 x y -> EqShiftL 1 (incr x) (incr y).
Proof.
intros.
rewrite <- 2 incrbis_aux_equiv.
apply EqShiftL_incrbis; auto.
Qed.
-
+
End EqShiftL.
(** * More equations about [incr] *)
- Lemma incr_twice_plus_one :
+ Lemma incr_twice_plus_one :
forall x, incr (twice_plus_one x) = twice (incr x).
Proof.
intros.
@@ -757,7 +757,7 @@ Section Basics.
destruct (incr (shiftr x)); simpl; discriminate.
Qed.
- Lemma incr_inv : forall x y,
+ Lemma incr_inv : forall x y,
incr x = twice_plus_one y -> x = twice y.
Proof.
intros.
@@ -777,7 +777,7 @@ Section Basics.
(** First, recursive equations *)
- Lemma phi_inv_double_plus_one : forall z,
+ Lemma phi_inv_double_plus_one : forall z,
phi_inv (Zdouble_plus_one z) = twice_plus_one (phi_inv z).
Proof.
destruct z; simpl; auto.
@@ -789,14 +789,14 @@ Section Basics.
auto.
Qed.
- Lemma phi_inv_double : forall z,
+ Lemma phi_inv_double : forall z,
phi_inv (Zdouble z) = twice (phi_inv z).
Proof.
destruct z; simpl; auto.
rewrite incr_twice_plus_one; auto.
Qed.
- Lemma phi_inv_incr : forall z,
+ Lemma phi_inv_incr : forall z,
phi_inv (Zsucc z) = incr (phi_inv z).
Proof.
destruct z.
@@ -816,19 +816,19 @@ Section Basics.
rewrite incr_twice_plus_one; auto.
Qed.
- (** [phi_inv o inv], the always-exact and easy-to-prove trip :
+ (** [phi_inv o inv], the always-exact and easy-to-prove trip :
from int31 to Z and then back to int31. *)
- Lemma phi_inv_phi_aux :
- forall n x, n <= size ->
- phi_inv (phibis_aux n (nshiftr (size-n) x)) =
+ Lemma phi_inv_phi_aux :
+ forall n x, n <= size ->
+ phi_inv (phibis_aux n (nshiftr (size-n) x)) =
nshiftr (size-n) x.
Proof.
induction n.
intros; simpl.
rewrite nshiftr_size; auto.
intros.
- unfold phibis_aux, recrbis_aux; fold recrbis_aux;
+ unfold phibis_aux, recrbis_aux; fold recrbis_aux;
fold (phibis_aux n (shiftr (nshiftr (size-S n) x))).
assert (shiftr (nshiftr (size - S n) x) = nshiftr (size-n) x).
replace (size - n)%nat with (S (size - (S n))); auto; omega.
@@ -863,10 +863,10 @@ Section Basics.
(** * [positive_to_int31] *)
- (** A variant of [p2i] with [twice] and [twice_plus_one] instead of
+ (** A variant of [p2i] with [twice] and [twice_plus_one] instead of
[2*i] and [2*i+1] *)
- Fixpoint p2ibis n p : (N*int31)%type :=
+ Fixpoint p2ibis n p : (N*int31)%type :=
match n with
| O => (Npos p, On)
| S n => match p with
@@ -876,7 +876,7 @@ Section Basics.
end
end.
- Lemma p2ibis_bounded : forall n p,
+ Lemma p2ibis_bounded : forall n p,
nshiftr n (snd (p2ibis n p)) = 0.
Proof.
induction n.
@@ -906,20 +906,20 @@ Section Basics.
replace (shiftr In) with 0; auto.
apply nshiftr_n_0.
Qed.
-
+
Lemma p2ibis_spec : forall n p, n<=size ->
- 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)))%Z.
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;
+ 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;
auto with zarith).
rewrite (Zmult_comm 2).
assert (n<=size) by omega.
- destruct p; simpl; [ | | auto];
- specialize (IHn p H0);
+ destruct p; simpl; [ | | auto];
+ specialize (IHn p H0);
generalize (p2ibis_bounded n p);
destruct (p2ibis n p) as (r,i); simpl in *; intros.
@@ -937,25 +937,25 @@ Section Basics.
(** We now prove that this [p2ibis] is related to [phi_inv_positive] *)
- Lemma phi_inv_positive_p2ibis : forall n p, (n<=size)%nat ->
+ Lemma phi_inv_positive_p2ibis : forall n p, (n<=size)%nat ->
EqShiftL (size-n) (phi_inv_positive p) (snd (p2ibis n p)).
Proof.
induction n.
intros.
apply EqShiftL_size; auto.
intros.
- simpl p2ibis; destruct p; [ | | red; auto];
- specialize IHn with p;
- destruct (p2ibis n p); simpl snd in *; simpl phi_inv_positive;
- rewrite ?EqShiftL_twice_plus_one, ?EqShiftL_twice;
- replace (S (size - S n))%nat with (size - n)%nat by omega;
+ simpl p2ibis; destruct p; [ | | red; auto];
+ specialize IHn with p;
+ destruct (p2ibis n p); simpl snd in *; simpl phi_inv_positive;
+ rewrite ?EqShiftL_twice_plus_one, ?EqShiftL_twice;
+ replace (S (size - S n))%nat with (size - n)%nat by omega;
apply IHn; omega.
Qed.
(** This gives the expected result about [phi o phi_inv], at least
for the positive case. *)
- Lemma phi_phi_inv_positive : forall p,
+ Lemma phi_phi_inv_positive : forall p,
phi (phi_inv_positive p) = (Zpos p) mod (2^(Z_of_nat size)).
Proof.
intros.
@@ -975,12 +975,12 @@ Section Basics.
Lemma double_twice_firstl : forall x, firstl x = D0 -> Twon*x = twice x.
Proof.
- intros.
+ intros.
unfold mul31.
rewrite <- Zdouble_mult, <- phi_twice_firstl, phi_inv_phi; auto.
Qed.
- Lemma double_twice_plus_one_firstl : forall x, firstl x = D0 ->
+ Lemma double_twice_plus_one_firstl : forall x, firstl x = D0 ->
Twon*x+In = twice_plus_one x.
Proof.
intros.
@@ -989,14 +989,14 @@ Section Basics.
rewrite phi_twice_firstl, <- Zdouble_plus_one_mult,
<- phi_twice_plus_one_firstl, phi_inv_phi; auto.
Qed.
-
- Lemma p2i_p2ibis : forall n p, (n<=size)%nat ->
+
+ Lemma p2i_p2ibis : forall n p, (n<=size)%nat ->
p2i n p = p2ibis n p.
Proof.
induction n; simpl; auto; intros.
- destruct p; auto; specialize IHn with p;
- generalize (p2ibis_bounded n p);
- rewrite IHn; try omega; destruct (p2ibis n p); simpl; intros;
+ destruct p; auto; specialize IHn with p;
+ generalize (p2ibis_bounded n p);
+ rewrite IHn; try omega; destruct (p2ibis n p); simpl; intros;
f_equal; auto.
apply double_twice_plus_one_firstl.
apply (nshiftr_0_firstl n); auto; omega.
@@ -1004,7 +1004,7 @@ Section Basics.
apply (nshiftr_0_firstl n); auto; omega.
Qed.
- Lemma positive_to_int31_phi_inv_positive : forall p,
+ Lemma positive_to_int31_phi_inv_positive : forall p,
snd (positive_to_int31 p) = phi_inv_positive p.
Proof.
intros; unfold positive_to_int31.
@@ -1014,8 +1014,8 @@ Section Basics.
apply (phi_inv_positive_p2ibis size); auto.
Qed.
- Lemma positive_to_int31_spec : forall p,
- Zpos p = ((Z_of_N (fst (positive_to_int31 p)))*2^(Z_of_nat size) +
+ Lemma positive_to_int31_spec : forall p,
+ Zpos p = ((Z_of_N (fst (positive_to_int31 p)))*2^(Z_of_nat size) +
phi (snd (positive_to_int31 p)))%Z.
Proof.
unfold positive_to_int31.
@@ -1023,11 +1023,11 @@ Section Basics.
apply p2ibis_spec; auto.
Qed.
- (** Thanks to the result about [phi o phi_inv_positive], we can
- now establish easily the most general results about
+ (** Thanks to the result about [phi o phi_inv_positive], we can
+ now establish easily the most general results about
[phi o twice] and so one. *)
-
- Lemma phi_twice : forall x,
+
+ Lemma phi_twice : forall x,
phi (twice x) = (Zdouble (phi x)) mod 2^(Z_of_nat size).
Proof.
intros.
@@ -1041,7 +1041,7 @@ Section Basics.
compute in H; elim H; auto.
Qed.
- Lemma phi_twice_plus_one : forall x,
+ Lemma phi_twice_plus_one : forall x,
phi (twice_plus_one x) = (Zdouble_plus_one (phi x)) mod 2^(Z_of_nat size).
Proof.
intros.
@@ -1055,14 +1055,14 @@ Section Basics.
compute in H; elim H; auto.
Qed.
- Lemma phi_incr : forall x,
+ Lemma phi_incr : forall x,
phi (incr x) = (Zsucc (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))%Z.
- change (Zsucc (phi x)) with ((phi x)+1)%Z;
+ change (Zsucc (phi x)) with ((phi x)+1)%Z;
generalize (phi_bounded x); omega.
destruct (Zsucc (phi x)).
simpl; auto.
@@ -1070,10 +1070,10 @@ Section Basics.
compute in H; elim H; auto.
Qed.
- (** With the previous results, we can deal with [phi o phi_inv] even
+ (** With the previous results, we can deal with [phi o phi_inv] even
in the negative case *)
- Lemma phi_phi_inv_negative :
+ Lemma phi_phi_inv_negative :
forall p, phi (incr (complement_negative p)) = (Zneg p) mod 2^(Z_of_nat size).
Proof.
induction p.
@@ -1091,11 +1091,11 @@ Section Basics.
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.
-
+
simpl; auto.
Qed.
- Lemma phi_phi_inv :
+ Lemma phi_phi_inv :
forall z, phi (phi_inv z) = z mod 2 ^ (Z_of_nat size).
Proof.
destruct z.
@@ -1120,7 +1120,7 @@ Let w_pos_mod p i :=
end.
(** Parity test *)
-Let w_iseven i :=
+Let w_iseven i :=
let (_,r) := i/2 in
match r ?= 0 with Eq => true | _ => false end.
@@ -1140,7 +1140,7 @@ Definition int31_op := (mk_znz_op
w_iszero
(* Basic arithmetic operations *)
(fun i => 0 -c i)
- (fun i => 0 - i)
+ opp31
(fun i => 0-i-1)
(fun i => i +c 1)
add31c
@@ -1181,14 +1181,14 @@ Definition int31_op := (mk_znz_op
End Int31_Op.
Section Int31_Spec.
-
- Open Local Scope Z_scope.
+
+ Local Open Scope Z_scope.
Notation "[| x |]" := (phi x) (at level 0, x at level 99).
- Notation Local wB := (2 ^ (Z_of_nat size)).
-
- Lemma wB_pos : wB > 0.
+ Local Notation wB := (2 ^ (Z_of_nat size)).
+
+ Lemma wB_pos : wB > 0.
Proof.
auto with zarith.
Qed.
@@ -1216,12 +1216,12 @@ Section Int31_Spec.
Proof.
reflexivity.
Qed.
-
+
Lemma spec_1 : [| 1 |] = 1.
Proof.
reflexivity.
Qed.
-
+
Lemma spec_Bm1 : [| Tn |] = wB - 1.
Proof.
reflexivity.
@@ -1252,16 +1252,16 @@ Section Int31_Spec.
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.
+ rewrite Zmod_small; romega.
generalize (Zcompare_Eq_eq ((X+Y) mod wB) (X+Y)); intros Heq.
- destruct Zcompare; intros;
+ destruct Zcompare; intros;
[ rewrite phi_phi_inv; auto | now apply H1 | now apply H1].
Qed.
Lemma spec_succ_c : forall x, [+|add31c x 1|] = [|x|] + 1.
Proof.
- intros; apply spec_add_c.
+ intros; apply spec_add_c.
Qed.
Lemma spec_add_carry_c : forall x y, [+|add31carryc x y|] = [|x|] + [|y|] + 1.
@@ -1279,7 +1279,7 @@ Section Int31_Spec.
rewrite Zmod_small; romega.
generalize (Zcompare_Eq_eq ((X+Y+1) mod wB) (X+Y+1)); intros Heq.
- destruct Zcompare; intros;
+ destruct Zcompare; intros;
[ rewrite phi_phi_inv; auto | now apply H1 | now apply H1].
Qed.
@@ -1304,7 +1304,7 @@ Section Int31_Spec.
(** Substraction *)
Lemma spec_sub_c : forall x y, [-|sub31c x y|] = [|x|] - [|y|].
- Proof.
+ Proof.
unfold sub31c, sub31, interp_carry; intros.
rewrite phi_phi_inv.
generalize (phi_bounded x)(phi_bounded y); intros.
@@ -1337,7 +1337,7 @@ Section Int31_Spec.
contradict H1; apply Zmod_small; romega.
generalize (Zcompare_Eq_eq ((X-Y-1) mod wB) (X-Y-1)); intros Heq.
- destruct Zcompare; intros;
+ destruct Zcompare; intros;
[ rewrite phi_phi_inv; auto | now apply H1 | now apply H1].
Qed.
@@ -1355,7 +1355,7 @@ Section Int31_Spec.
Qed.
Lemma spec_opp_c : forall x, [-|sub31c 0 x|] = -[|x|].
- Proof.
+ Proof.
intros; apply spec_sub_c.
Qed.
@@ -1402,7 +1402,7 @@ Section Int31_Spec.
change (wB*wB) with (wB^2); ring.
unfold phi_inv2.
- destruct x; unfold zn2z_to_Z; rewrite ?phi_phi_inv;
+ destruct x; unfold zn2z_to_Z; rewrite ?phi_phi_inv;
change base with wB; auto.
Qed.
@@ -1426,7 +1426,7 @@ Section Int31_Spec.
intros; apply spec_mul_c.
Qed.
- (** Division *)
+ (** Division *)
Lemma spec_div21 : forall a1 a2 b,
wB/2 <= [|b|] ->
@@ -1537,7 +1537,7 @@ Section Int31_Spec.
intros (H,_); compute in H; elim H; auto.
Qed.
- Lemma iter_int31_iter_nat : forall A f i a,
+ Lemma iter_int31_iter_nat : forall A f i a,
iter_int31 i A f a = iter_nat (Zabs_nat [|i|]) A f a.
Proof.
intros.
@@ -1548,17 +1548,17 @@ Section Int31_Spec.
revert i a; induction size.
simpl; auto.
simpl; intros.
- case_eq (firstr i); intros H; rewrite 2 IHn;
+ case_eq (firstr i); intros H; rewrite 2 IHn;
unfold phibis_aux; simpl; rewrite H; fold (phibis_aux n (shiftr i));
- generalize (phibis_aux_pos n (shiftr i)); intros;
- set (z := phibis_aux n (shiftr i)) in *; clearbody z;
+ generalize (phibis_aux_pos n (shiftr i)); intros;
+ set (z := phibis_aux n (shiftr i)) in *; clearbody z;
rewrite <- iter_nat_plus.
f_equal.
rewrite Zdouble_mult, Zmult_comm, <- Zplus_diag_eq_mult_2.
symmetry; apply Zabs_nat_Zplus; auto with zarith.
- change (iter_nat (S (Zabs_nat z + Zabs_nat z)) A f a =
+ 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.
@@ -1566,13 +1566,13 @@ Section Int31_Spec.
change (Zabs_nat 1) with 1%nat; omega.
Qed.
- Fixpoint addmuldiv31_alt n i j :=
- match n with
- | O => i
+ Fixpoint addmuldiv31_alt n i j :=
+ match n with
+ | O => i
| S n => addmuldiv31_alt n (sneakl (firstl j) i) (shiftl j)
end.
- Lemma addmuldiv31_equiv : forall p x y,
+ Lemma addmuldiv31_equiv : forall p x y,
addmuldiv31 p x y = addmuldiv31_alt (Zabs_nat [|p|]) x y.
Proof.
intros.
@@ -1588,7 +1588,7 @@ Section Int31_Spec.
Qed.
Lemma spec_add_mul_div : forall x y p, [|p|] <= Zpos 31 ->
- [| addmuldiv31 p x y |] =
+ [| addmuldiv31 p x y |] =
([|x|] * (2 ^ [|p|]) + [|y|] / (2 ^ ((Zpos 31) - [|p|]))) mod wB.
Proof.
intros.
@@ -1626,7 +1626,7 @@ Section Int31_Spec.
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.
-
+
rewrite phi_twice_plus_one, Zdouble_plus_one_mult.
rewrite phi_twice; auto.
change (Zdouble [|y|]) with (2*[|y|]).
@@ -1644,7 +1644,7 @@ Section Int31_Spec.
unfold wB'. rewrite <- Zpower_Zsucc, <- inj_S 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.
rewrite Z_div_plus; auto with zarith.
@@ -1669,8 +1669,8 @@ Section Int31_Spec.
apply Zlt_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 | |
+ case_eq ([|p|] ?= 31); intros;
+ [ apply H; rewrite (Zcompare_Eq_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.
@@ -1701,16 +1701,16 @@ Section Int31_Spec.
simpl; auto.
Qed.
- Fixpoint head031_alt n x :=
- match n with
+ Fixpoint head031_alt n x :=
+ match n with
| O => 0%nat
- | S n => match firstl x with
+ | S n => match firstl x with
| D0 => S (head031_alt n (shiftl x))
| D1 => 0%nat
end
end.
- Lemma head031_equiv :
+ Lemma head031_equiv :
forall x, [|head031 x|] = Z_of_nat (head031_alt size x).
Proof.
intros.
@@ -1720,10 +1720,10 @@ Section Int31_Spec.
unfold head031, recl.
change On with (phi_inv (Z_of_nat (31-size))).
- replace (head031_alt size x) with
+ replace (head031_alt size x) with
(head031_alt size x + (31 - size))%nat by (apply plus_0_r; auto).
assert (size <= 31)%nat by auto with arith.
-
+
revert x H; induction size; intros.
simpl; auto.
unfold recl_aux; fold recl_aux.
@@ -1748,7 +1748,7 @@ Section Int31_Spec.
change [|In|] with 1.
replace (31-n)%nat with (S (31 - S n))%nat by omega.
rewrite inj_S; ring.
-
+
clear - H H2.
rewrite (sneakr_shiftl x) in H.
rewrite H2 in H.
@@ -1793,7 +1793,7 @@ Section Int31_Spec.
rewrite (sneakr_shiftl x), H1, H; auto.
rewrite <- nshiftl_S_tail; auto.
-
+
change (2^(Z_of_nat 0)) with 1; rewrite Zmult_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))).
@@ -1809,16 +1809,16 @@ Section Int31_Spec.
simpl; auto.
Qed.
- Fixpoint tail031_alt n x :=
- match n with
+ Fixpoint tail031_alt n x :=
+ match n with
| O => 0%nat
- | S n => match firstr x with
+ | S n => match firstr x with
| D0 => S (tail031_alt n (shiftr x))
| D1 => 0%nat
end
end.
- Lemma tail031_equiv :
+ Lemma tail031_equiv :
forall x, [|tail031 x|] = Z_of_nat (tail031_alt size x).
Proof.
intros.
@@ -1828,10 +1828,10 @@ Section Int31_Spec.
unfold tail031, recr.
change On with (phi_inv (Z_of_nat (31-size))).
- replace (tail031_alt size x) with
+ replace (tail031_alt size x) with
(tail031_alt size x + (31 - size))%nat by (apply plus_0_r; auto).
assert (size <= 31)%nat by auto with arith.
-
+
revert x H; induction size; intros.
simpl; auto.
unfold recr_aux; fold recr_aux.
@@ -1856,7 +1856,7 @@ Section Int31_Spec.
change [|In|] with 1.
replace (31-n)%nat with (S (31 - S n))%nat by omega.
rewrite inj_S; ring.
-
+
clear - H H2.
rewrite (sneakl_shiftr x) in H.
rewrite H2 in H.
@@ -1864,7 +1864,7 @@ Section Int31_Spec.
rewrite (iszero_eq0 _ H0) in H; discriminate.
Qed.
- Lemma spec_tail0 : forall x, 0 < [|x|] ->
+ Lemma spec_tail0 : forall x, 0 < [|x|] ->
exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|tail031 x|]).
Proof.
intros.
@@ -1882,23 +1882,23 @@ Section Int31_Spec.
case_eq (firstr x); intros.
rewrite (inj_S (tail031_alt n (shiftr x))), Zpower_Zsucc; auto with zarith.
destruct (IHn (shiftr x)) as (y & Hy1 & Hy2).
-
+
rewrite phi_nz; rewrite phi_nz in H; contradict H.
rewrite (sneakl_shiftr x), H1, H; auto.
rewrite <- nshiftr_S_tail; auto.
-
+
exists y; split; auto.
rewrite phi_eqn1; auto.
rewrite Zdouble_mult, 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.
Qed.
-
+
(* Sqrt *)
(* Direct transcription of an old proof
@@ -1906,27 +1906,27 @@ Section Int31_Spec.
Lemma quotient_by_2 a: a - 1 <= (a/2) + (a/2).
Proof.
- intros a; case (Z_mod_lt a 2); auto with zarith.
+ case (Z_mod_lt a 2); auto with zarith.
intros H1; rewrite Zmod_eq_full; auto with zarith.
Qed.
- Lemma sqrt_main_trick j k: 0 <= j -> 0 <= k ->
+ Lemma sqrt_main_trick j k: 0 <= j -> 0 <= k ->
(j * k) + j <= ((j + k)/2 + 1) ^ 2.
Proof.
- intros j k Hj; generalize Hj k; pattern j; apply natlike_ind;
+ 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 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));
+ 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.
auto with zarith.
intros k Hk _.
replace ((Zsucc j + Zsucc k) / 2) with ((j + k)/2 + 1).
generalize (Hrec Hj k Hk) (quotient_by_2 (j + k)).
- unfold Zsucc; repeat rewrite Zpower_2;
+ 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.
auto with zarith.
@@ -1936,7 +1936,7 @@ Section Int31_Spec.
Lemma sqrt_main i j: 0 <= i -> 0 < j -> i < ((j + (i/j))/2 + 1) ^ 2.
Proof.
- intros i j Hi Hj.
+ 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).
pattern i at 1; rewrite (Z_div_mod_eq i j); case (Z_mod_lt i j); auto with zarith.
@@ -1944,7 +1944,7 @@ Section Int31_Spec.
Lemma sqrt_init i: 1 < i -> i < (i/2 + 1) ^ 2.
Proof.
- intros i Hi.
+ intros Hi.
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.
@@ -1962,14 +1962,14 @@ Section Int31_Spec.
Lemma sqrt_test_true i j: 0 <= i -> 0 < j -> i/j >= j -> j ^ 2 <= i.
Proof.
- intros i j Hi Hj Hd; rewrite Zpower_2.
+ intros Hi Hj Hd; rewrite Zpower_2.
apply Zle_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 i j Hi Hj H; case (Zle_or_lt j ((j + (i/j))/2)); auto.
+ intros Hi Hj H; case (Zle_or_lt j ((j + (i/j))/2)); auto.
intros H1; contradict H; apply Zle_not_lt.
assert (2 * j <= j + (i/j)); auto with zarith.
apply Zle_trans with (2 * ((j + (i/j))/2)); auto with zarith.
@@ -1984,32 +1984,32 @@ Section Int31_Spec.
Lemma Zcompare_spec i j: ZcompareSpec i j (i ?= j).
Proof.
- intros i j; case_eq (Zcompare i j); intros H.
+ 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 =
+ sqrt31_step rec i j =
match (fst (i/j) ?= j)%int31 with
Lt => rec i (fst ((j + fst(i/j))/2))%int31
| _ => j
end.
Proof.
- intros rec i j; unfold sqrt31_step; case div31; intros.
+ unfold sqrt31_step; case div31; intros.
simpl; case compare31; auto.
Qed.
Lemma div31_phi i j: 0 < [|j|] -> [|fst (i/j)%int31|] = [|i|]/[|j|].
- intros i j Hj; generalize (spec_div i j Hj).
+ intros Hj; generalize (spec_div i j Hj).
case div31; intros q r; simpl fst.
intros (H1,H2); apply Zdiv_unique with [|r|]; auto with zarith.
rewrite H1; ring.
Qed.
- Lemma sqrt31_step_correct rec i j:
- 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 ->
+ Lemma sqrt31_step_correct rec i j:
+ 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 ->
2 * [|j|] < wB ->
(forall j1 : int31,
0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 ->
@@ -2017,15 +2017,15 @@ Section Int31_Spec.
[|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2.
Proof.
assert (Hp2: 0 < [|2|]) by exact (refl_equal Lt).
- intros rec i j Hi Hj Hij H31 Hrec; rewrite sqrt31_step_def.
- generalize (spec_compare (fst (i/j)%int31) j); case compare31;
+ intros Hi Hj Hij H31 Hrec; rewrite sqrt31_step_def.
+ generalize (spec_compare (fst (i/j)%int31) j); case compare31;
rewrite div31_phi; 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.
- replace ([|j|] + [|i|]/[|j|]) with
+ 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).
@@ -2048,12 +2048,12 @@ Section Int31_Spec.
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|] ->
+ (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.
- intros n; elim n; unfold iter31_sqrt; fold iter31_sqrt; clear n.
+ 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.
@@ -2098,7 +2098,7 @@ Section Int31_Spec.
Qed.
Lemma sqrt312_step_def rec ih il j:
- sqrt312_step rec ih il j =
+ sqrt312_step rec ih il j =
match (ih ?= j)%int31 with
Eq => j
| Gt => j
@@ -2112,14 +2112,14 @@ Section Int31_Spec.
end
end.
Proof.
- intros rec ih il j; unfold sqrt312_step; case div3121; intros.
+ unfold sqrt312_step; case div3121; intros.
simpl; case compare31; auto.
Qed.
- Lemma sqrt312_lower_bound ih il j:
+ Lemma sqrt312_lower_bound ih il j:
phi2 ih il < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|].
Proof.
- intros ih il j H1.
+ intros H1.
case (phi_bounded j); intros Hbj _.
case (phi_bounded il); intros Hbil _.
case (phi_bounded ih); intros Hbih Hbih1.
@@ -2133,22 +2133,22 @@ Section Int31_Spec.
Lemma div312_phi ih il j: (2^30 <= [|j|] -> [|ih|] < [|j|] ->
[|fst (div3121 ih il j)|] = phi2 ih il/[|j|])%Z.
Proof.
- intros ih il j Hj Hj1.
+ intros Hj Hj1.
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.
Qed.
- Lemma sqrt312_step_correct rec ih il j:
- 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 ->
+ Lemma sqrt312_step_correct rec ih il j:
+ 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 ->
(forall j1, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 ->
[|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) ->
- [|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il
+ [|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).
- intros rec ih il j Hih Hj Hij Hrec; rewrite sqrt312_step_def.
+ 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 _.
@@ -2174,7 +2174,7 @@ Section Int31_Spec.
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.
assert (Hf3: 0 < ([|j|] + phi2 ih il / [|j|]) / 2).
- replace ([|j|] + phi2 ih il/ [|j|])%Z with
+ replace ([|j|] + phi2 ih il/ [|j|])%Z with
(1 * 2 + (([|j|] - 2) + phi2 ih il / [|j|])); try ring.
rewrite Z_div_plus_full_l; auto with zarith.
assert (0 <= ([|j|] - 2 + phi2 ih il / [|j|]) / 2) ; auto with zarith.
@@ -2213,7 +2213,7 @@ Section Int31_Spec.
rewrite div31_phi; change (phi 2) with 2%Z; auto.
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 Zmult_gt_0_lt_reg_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.
@@ -2234,15 +2234,15 @@ Section Int31_Spec.
apply Zge_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|] ->
- phi2 ih il < ([|j1|] + 1) ^ 2 ->
+ 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|] ->
+ 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
+ [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il
< ([|iter312_sqrt n rec ih il j|] + 1) ^ 2.
Proof.
- intros n; elim n; unfold iter312_sqrt; fold iter312_sqrt; clear n.
+ 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.
@@ -2265,7 +2265,7 @@ Section Int31_Spec.
Proof.
intros ih il Hih; unfold sqrt312.
change [||WW ih il||] with (phi2 ih il).
- assert (Hbin: forall s, s * s + 2* s + 1 = (s + 1) ^ 2) by
+ assert (Hbin: forall s, s * s + 2* s + 1 = (s + 1) ^ 2) by
(intros s; ring).
assert (Hb: 0 <= base) by (red; intros HH; discriminate).
assert (Hi2: phi2 ih il < (phi Tn + 1) ^ 2).
@@ -2428,9 +2428,9 @@ Section Int31_Spec.
apply Zcompare_Eq_eq.
now destruct ([|x|] ?= 0).
Qed.
-
+
(* Even *)
-
+
Let w_is_even := int31_op.(znz_is_even).
Lemma spec_is_even : forall x,
@@ -2460,13 +2460,13 @@ Section Int31_Spec.
exact spec_more_than_1_digit.
exact spec_0.
- exact spec_1.
+ exact spec_1.
exact spec_Bm1.
exact spec_compare.
exact spec_eq0.
- exact spec_opp_c.
+ exact spec_opp_c.
exact spec_opp.
exact spec_opp_carry.
@@ -2500,7 +2500,7 @@ Section Int31_Spec.
exact spec_head00.
exact spec_head0.
- exact spec_tail00.
+ exact spec_tail00.
exact spec_tail0.
exact spec_add_mul_div.
diff --git a/theories/Numbers/Cyclic/Int31/Int31.v b/theories/Numbers/Cyclic/Int31/Int31.v
index 154b436b..cc224254 100644
--- a/theories/Numbers/Cyclic/Int31/Int31.v
+++ b/theories/Numbers/Cyclic/Int31/Int31.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: Int31.v 11072 2008-06-08 16:13:37Z herbelin $ i*)
+(*i $Id$ i*)
Require Import NaryFunctions.
Require Import Wf_nat.
@@ -17,7 +17,7 @@ Require Export DoubleType.
Unset Boxed Definitions.
-(** * 31-bit integers *)
+(** * 31-bit integers *)
(** This file contains basic definitions of a 31-bit integer
arithmetic. In fact it is more general than that. The only reason
@@ -36,11 +36,13 @@ Definition size := 31%nat.
Inductive digits : Type := D0 | D1.
(** The type of 31-bit integers *)
-
-(** The type [int31] has a unique constructor [I31] that expects
+
+(** The type [int31] has a unique constructor [I31] that expects
31 arguments of type [digits]. *)
-Inductive int31 : Type := I31 : nfun digits size int31.
+Definition digits31 t := Eval compute in nfun digits size t.
+
+Inductive int31 : Type := I31 : digits31 int31.
(* spiwack: Registration of the type of integers, so that the matchs in
the functions below perform dynamic decompilation (otherwise some segfault
@@ -50,7 +52,7 @@ Register int31 as int31 type in "coq_int31" by True.
Delimit Scope int31_scope with int31.
Bind Scope int31_scope with int31.
-Open Scope int31_scope.
+Local Open Scope int31_scope.
(** * Constants *)
@@ -69,26 +71,26 @@ Definition Twon : int31 := Eval compute in (napply_cst _ _ D0 (size-2) I31) D1 D
(** * Bits manipulation *)
-(** [sneakr b x] shifts [x] to the right by one bit.
+(** [sneakr b x] shifts [x] to the right by one bit.
Rightmost digit is lost while leftmost digit becomes [b].
- Pseudo-code is
+ Pseudo-code is
[ match x with (I31 d0 ... dN) => I31 b d0 ... d(N-1) end ]
*)
Definition sneakr : digits -> int31 -> int31 := Eval compute in
fun b => int31_rect _ (napply_except_last _ _ (size-1) (I31 b)).
-(** [sneakl b x] shifts [x] to the left by one bit.
+(** [sneakl b x] shifts [x] to the left by one bit.
Leftmost digit is lost while rightmost digit becomes [b].
- Pseudo-code is
+ Pseudo-code is
[ match x with (I31 d0 ... dN) => I31 d1 ... dN b end ]
*)
-Definition sneakl : digits -> int31 -> int31 := Eval compute in
+Definition sneakl : digits -> int31 -> int31 := Eval compute in
fun b => int31_rect _ (fun _ => napply_then_last _ _ b (size-1) I31).
-(** [shiftl], [shiftr], [twice] and [twice_plus_one] are direct
+(** [shiftl], [shiftr], [twice] and [twice_plus_one] are direct
consequences of [sneakl] and [sneakr]. *)
Definition shiftl := sneakl D0.
@@ -96,31 +98,31 @@ Definition shiftr := sneakr D0.
Definition twice := sneakl D0.
Definition twice_plus_one := sneakl D1.
-(** [firstl x] returns the leftmost digit of number [x].
+(** [firstl x] returns the leftmost digit of number [x].
Pseudo-code is [ match x with (I31 d0 ... dN) => d0 end ] *)
-Definition firstl : int31 -> digits := Eval compute in
+Definition firstl : int31 -> digits := Eval compute in
int31_rect _ (fun d => napply_discard _ _ d (size-1)).
-(** [firstr x] returns the rightmost digit of number [x].
+(** [firstr x] returns the rightmost digit of number [x].
Pseudo-code is [ match x with (I31 d0 ... dN) => dN end ] *)
-Definition firstr : int31 -> digits := Eval compute in
+Definition firstr : int31 -> digits := Eval compute in
int31_rect _ (napply_discard _ _ (fun d=>d) (size-1)).
-(** [iszero x] is true iff [x = I31 D0 ... D0]. Pseudo-code is
+(** [iszero x] is true iff [x = I31 D0 ... D0]. Pseudo-code is
[ match x with (I31 D0 ... D0) => true | _ => false end ] *)
-Definition iszero : int31 -> bool := Eval compute in
- let f d b := match d with D0 => b | D1 => false end
+Definition iszero : int31 -> bool := Eval compute in
+ let f d b := match d with D0 => b | D1 => false end
in int31_rect _ (nfold_bis _ _ f true size).
-(* NB: DO NOT transform the above match in a nicer (if then else).
+(* NB: DO NOT transform the above match in a nicer (if then else).
It seems to work, but later "unfold iszero" takes forever. *)
-(** [base] is [2^31], obtained via iterations of [Zdouble].
- It can also be seen as the smallest b > 0 s.t. phi_inv b = 0
+(** [base] is [2^31], obtained via iterations of [Zdouble].
+ It can also be seen as the smallest b > 0 s.t. phi_inv b = 0
(see below) *)
Definition base := Eval compute in
@@ -140,7 +142,7 @@ Fixpoint recl_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A)
caserec (firstl i) si (recl_aux next A case0 caserec si)
end.
-Fixpoint recr_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A)
+Fixpoint recr_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A)
(i:int31) : A :=
match n with
| O => case0
@@ -159,22 +161,22 @@ Definition recr := recr_aux size.
(** From int31 to Z, we simply iterates [Zdouble] or [Zdouble_plus_one]. *)
-Definition phi : int31 -> Z :=
+Definition phi : int31 -> Z :=
recr Z (0%Z)
(fun b _ => match b with D0 => Zdouble | D1 => Zdouble_plus_one end).
-(** From positive to int31. An abstract definition could be :
- [ phi_inv (2n) = 2*(phi_inv n) /\
+(** From positive to int31. An abstract definition could be :
+ [ phi_inv (2n) = 2*(phi_inv n) /\
phi_inv 2n+1 = 2*(phi_inv n) + 1 ] *)
-Fixpoint phi_inv_positive p :=
+Fixpoint phi_inv_positive p :=
match p with
| xI q => twice_plus_one (phi_inv_positive q)
| xO q => twice (phi_inv_positive q)
| xH => In
end.
-(** The negative part : 2-complement *)
+(** The negative part : 2-complement *)
Fixpoint complement_negative p :=
match p with
@@ -186,9 +188,9 @@ Fixpoint complement_negative p :=
(** A simple incrementation function *)
Definition incr : int31 -> int31 :=
- recr int31 In
- (fun b si rec => match b with
- | D0 => sneakl D1 si
+ recr int31 In
+ (fun b si rec => match b with
+ | D0 => sneakl D1 si
| D1 => sneakl D0 rec end).
(** We can now define the conversion from Z to int31. *)
@@ -196,11 +198,11 @@ Definition incr : int31 -> int31 :=
Definition phi_inv : Z -> int31 := fun n =>
match n with
| Z0 => On
- | Zpos p => phi_inv_positive p
+ | Zpos p => phi_inv_positive p
| Zneg p => incr (complement_negative p)
end.
-(** [phi_inv2] is similar to [phi_inv] but returns a double word
+(** [phi_inv2] is similar to [phi_inv] but returns a double word
[zn2z int31] *)
Definition phi_inv2 n :=
@@ -211,7 +213,7 @@ Definition phi_inv2 n :=
(** [phi2] is similar to [phi] but takes a double word (two args) *)
-Definition phi2 nh nl :=
+Definition phi2 nh nl :=
((phi nh)*base+(phi nl))%Z.
(** * Addition *)
@@ -227,11 +229,11 @@ Notation "n + m" := (add31 n m) : int31_scope.
(* mode, (phi n)+(phi m) is computed twice*)
(* it may be considered to optimize it *)
-Definition add31c (n m : int31) :=
+Definition add31c (n m : int31) :=
let npm := n+m in
- match (phi npm ?= (phi n)+(phi m))%Z with
- | Eq => C0 npm
- | _ => C1 npm
+ match (phi npm ?= (phi n)+(phi m))%Z with
+ | Eq => C0 npm
+ | _ => C1 npm
end.
Notation "n '+c' m" := (add31c n m) (at level 50, no associativity) : int31_scope.
@@ -254,7 +256,7 @@ Notation "n - m" := (sub31 n m) : int31_scope.
(** Subtraction with carry (thus exact) *)
-Definition sub31c (n m : int31) :=
+Definition sub31c (n m : int31) :=
let nmm := n-m in
match (phi nmm ?= (phi n)-(phi m))%Z with
| Eq => C0 nmm
@@ -272,6 +274,10 @@ Definition sub31carryc (n m : int31) :=
| _ => C1 nmmmone
end.
+(** Opposite *)
+
+Definition opp31 x := On - x.
+Notation "- x" := (opp31 x) : int31_scope.
(** Multiplication *)
@@ -290,13 +296,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) :=
+Definition div3121 (nh nl m : int31) :=
let (q,r) := Zdiv_eucl (phi2 nh nl) (phi m) in
(phi_inv q, phi_inv r).
(** Division modulo [2^31] *)
-Definition div31 (n m : int31) :=
+Definition div31 (n m : int31) :=
let (q,r) := Zdiv_eucl (phi n) (phi m) in
(phi_inv q, phi_inv r).
Notation "n / m" := (div31 n m) : int31_scope.
@@ -307,13 +313,16 @@ Notation "n / m" := (div31 n m) : int31_scope.
Definition compare31 (n m : int31) := ((phi n)?=(phi m))%Z.
Notation "n ?= m" := (compare31 n m) (at level 70, no associativity) : int31_scope.
+Definition eqb31 (n m : int31) :=
+ match n ?= m with Eq => true | _ => false end.
+
-(** Computing the [i]-th iterate of a function:
+(** Computing the [i]-th iterate of a function:
[iter_int31 i A f = f^i] *)
Definition iter_int31 i A f :=
- recr (A->A) (fun x => x)
- (fun b si rec => match b with
+ recr (A->A) (fun x => x)
+ (fun b si rec => match b with
| D0 => fun x => rec (rec x)
| D1 => fun x => f (rec (rec x))
end)
@@ -322,9 +331,9 @@ Definition iter_int31 i A f :=
(** Combining the [(31-p)] low bits of [i] above the [p] high bits of [j]:
[addmuldiv31 p i j = i*2^p+j/2^(31-p)] (modulo [2^31]) *)
-Definition addmuldiv31 p i j :=
- let (res, _ ) :=
- iter_int31 p (int31*int31)
+Definition addmuldiv31 p i j :=
+ let (res, _ ) :=
+ iter_int31 p (int31*int31)
(fun ij => let (i,j) := ij in (sneakl (firstl j) i, shiftl j))
(i,j)
in
@@ -346,7 +355,7 @@ Register addmuldiv31 as int31 addmuldiv in "coq_int31" by True.
Definition gcd31 (i j:int31) :=
(fix euler (guard:nat) (i j:int31) {struct guard} :=
- match guard with
+ match guard with
| O => In
| S p => match j ?= On with
| Eq => i
@@ -370,17 +379,17 @@ Eval lazy delta [Twon] in
| _ => j
end.
-Fixpoint iter31_sqrt (n: nat) (rec: int31 -> int31 -> int31)
+Fixpoint iter31_sqrt (n: nat) (rec: int31 -> int31 -> int31)
(i j: int31) {struct n} : int31 :=
- sqrt31_step
+ sqrt31_step
(match n with
O => rec
| S n => (iter31_sqrt n (iter31_sqrt n rec))
end) i j.
-Definition sqrt31 i :=
+Definition sqrt31 i :=
Eval lazy delta [On In Twon] in
- match compare31 In i with
+ match compare31 In i with
Gt => On
| Eq => In
| Lt => iter31_sqrt 31 (fun i j => j) i (fst (i/Twon))
@@ -388,7 +397,7 @@ Eval lazy delta [On In Twon] in
Definition v30 := Eval compute in (addmuldiv31 (phi_inv (Z_of_nat size - 1)) In On).
-Definition sqrt312_step (rec: int31 -> int31 -> int31 -> int31)
+Definition sqrt312_step (rec: int31 -> int31 -> int31 -> int31)
(ih il j: int31) :=
Eval lazy delta [Twon v30] in
match ih ?= j with Eq => j | Gt => j | _ =>
@@ -401,28 +410,28 @@ Eval lazy delta [Twon v30] in
| _ => j
end end.
-Fixpoint iter312_sqrt (n: nat)
- (rec: int31 -> int31 -> int31 -> int31)
+Fixpoint iter312_sqrt (n: nat)
+ (rec: int31 -> int31 -> int31 -> int31)
(ih il j: int31) {struct n} : int31 :=
- sqrt312_step
+ sqrt312_step
(match n with
O => rec
| S n => (iter312_sqrt n (iter312_sqrt n rec))
end) ih il j.
-Definition sqrt312 ih il :=
+Definition sqrt312 ih il :=
Eval lazy delta [On In] in
let s := iter312_sqrt 31 (fun ih il j => j) ih il Tn in
match s *c s with
W0 => (On, C0 On) (* impossible *)
| WW ih1 il1 =>
match il -c il1 with
- C0 il2 =>
+ C0 il2 =>
match ih ?= ih1 with
Gt => (s, C1 il2)
| _ => (s, C0 il2)
end
- | C1 il2 =>
+ | C1 il2 =>
match (ih - In) ?= ih1 with (* we could parametrize ih - 1 *)
Gt => (s, C1 il2)
| _ => (s, C0 il2)
@@ -431,7 +440,7 @@ Eval lazy delta [On In] in
end.
-Fixpoint p2i n p : (N*int31)%type :=
+Fixpoint p2i n p : (N*int31)%type :=
match n with
| O => (Npos p, On)
| S n => match p with
@@ -444,26 +453,26 @@ Fixpoint p2i n p : (N*int31)%type :=
Definition positive_to_int31 (p:positive) := p2i size p.
(** Constant 31 converted into type int31.
- It is used as default answer for numbers of zeros
+ 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 head031 (i:int31) :=
- recl _ (fun _ => T31)
- (fun b si rec n => match b with
+ recl _ (fun _ => T31)
+ (fun b si rec n => match b with
| D0 => rec (add31 n In)
| D1 => n
end)
i On.
Definition tail031 (i:int31) :=
- recr _ (fun _ => T31)
- (fun b si rec n => match b with
+ recr _ (fun _ => T31)
+ (fun b si rec n => match b with
| D0 => rec (add31 n In)
| D1 => n
end)
i On.
Register head031 as int31 head0 in "coq_int31" by True.
-Register tail031 as int31 tail0 in "coq_int31" by True.
+Register tail031 as int31 tail0 in "coq_int31" by True.
diff --git a/theories/Numbers/Cyclic/Int31/Ring31.v b/theories/Numbers/Cyclic/Int31/Ring31.v
new file mode 100644
index 00000000..2ec406b0
--- /dev/null
+++ b/theories/Numbers/Cyclic/Int31/Ring31.v
@@ -0,0 +1,103 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+(** * Int31 numbers defines Z/(2^31)Z, and can hence be equipped
+ with a ring structure and a ring tactic *)
+
+Require Import Int31 Cyclic31 CyclicAxioms.
+
+Local Open Scope int31_scope.
+
+(** Detection of constants *)
+
+Local Open Scope list_scope.
+
+Ltac isInt31cst_lst l :=
+ match l with
+ | nil => constr:true
+ | ?t::?l => match t with
+ | D1 => isInt31cst_lst l
+ | D0 => isInt31cst_lst l
+ | _ => constr:false
+ end
+ | _ => constr:false
+ end.
+
+Ltac isInt31cst t :=
+ match t with
+ | I31 ?i0 ?i1 ?i2 ?i3 ?i4 ?i5 ?i6 ?i7 ?i8 ?i9 ?i10
+ ?i11 ?i12 ?i13 ?i14 ?i15 ?i16 ?i17 ?i18 ?i19 ?i20
+ ?i21 ?i22 ?i23 ?i24 ?i25 ?i26 ?i27 ?i28 ?i29 ?i30 =>
+ let l :=
+ constr:(i0::i1::i2::i3::i4::i5::i6::i7::i8::i9::i10
+ ::i11::i12::i13::i14::i15::i16::i17::i18::i19::i20
+ ::i21::i22::i23::i24::i25::i26::i27::i28::i29::i30::nil)
+ in isInt31cst_lst l
+ | Int31.On => constr:true
+ | Int31.In => constr:true
+ | Int31.Tn => constr:true
+ | Int31.Twon => constr:true
+ | _ => constr:false
+ end.
+
+Ltac Int31cst t :=
+ match isInt31cst t with
+ | true => constr:t
+ | false => constr:NotConstant
+ end.
+
+(** The generic ring structure inferred from the Cyclic structure *)
+
+Module Int31ring := CyclicRing Int31Cyclic.
+
+(** Unlike in the generic [CyclicRing], we can use Leibniz here. *)
+
+Lemma Int31_canonic : forall x y, phi x = phi y -> x = y.
+Proof.
+ intros x y EQ.
+ now rewrite <- (phi_inv_phi x), <- (phi_inv_phi y), EQ.
+Qed.
+
+Lemma ring_theory_switch_eq :
+ forall A (R R':A->A->Prop) zero one add mul sub opp,
+ (forall x y : A, R x y -> R' x y) ->
+ ring_theory zero one add mul sub opp R ->
+ ring_theory zero one add mul sub opp R'.
+Proof.
+intros A R R' zero one add mul sub opp Impl Ring.
+constructor; intros; apply Impl; apply Ring.
+Qed.
+
+Lemma Int31Ring : ring_theory 0 1 add31 mul31 sub31 opp31 Logic.eq.
+Proof.
+exact (ring_theory_switch_eq _ _ _ _ _ _ _ _ _ Int31_canonic Int31ring.CyclicRing).
+Qed.
+
+Lemma eqb31_eq : forall x y, eqb31 x y = true <-> x=y.
+Proof.
+unfold eqb31. intros x y.
+generalize (Cyclic31.spec_compare x y).
+destruct (x ?= y); intuition; subst; auto with zarith; try discriminate.
+apply Int31_canonic; auto.
+Qed.
+
+Lemma eqb31_correct : forall x y, eqb31 x y = true -> x=y.
+Proof. now apply eqb31_eq. Qed.
+
+Add Ring Int31Ring : Int31Ring
+ (decidable eqb31_correct,
+ constants [Int31cst]).
+
+Section TestRing.
+Let test : forall x y, 1 + x*y + x*x + 1 = 1*1 + 1 + y*x + 1*x*x.
+intros. ring.
+Qed.
+End TestRing.
+
diff --git a/theories/Numbers/Cyclic/ZModulo/ZModulo.v b/theories/Numbers/Cyclic/ZModulo/ZModulo.v
index 7c770e97..4f0f6c7c 100644
--- a/theories/Numbers/Cyclic/ZModulo/ZModulo.v
+++ b/theories/Numbers/Cyclic/ZModulo/ZModulo.v
@@ -6,13 +6,13 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: ZModulo.v 11033 2008-06-01 22:56:50Z letouzey $ *)
+(* $Id$ *)
-(** * Type [Z] viewed modulo a particular constant corresponds to [Z/nZ]
+(** * Type [Z] viewed modulo a particular constant corresponds to [Z/nZ]
as defined abstractly in CyclicAxioms. *)
-(** Even if the construction provided here is not reused for building
- the efficient arbitrary precision numbers, it provides a simple
+(** Even if the construction provided here is not reused for building
+ the efficient arbitrary precision numbers, it provides a simple
implementation of CyclicAxioms, hence ensuring its coherence. *)
Set Implicit Arguments.
@@ -24,7 +24,7 @@ Require Import BigNumPrelude.
Require Import DoubleType.
Require Import CyclicAxioms.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Section ZModulo.
@@ -56,9 +56,9 @@ Section ZModulo.
destruct 1; auto.
Qed.
Let digits_gt_1 := spec_more_than_1_digit.
-
+
Lemma wB_pos : wB > 0.
- Proof.
+ Proof.
unfold wB, base; auto with zarith.
Qed.
Hint Resolve wB_pos.
@@ -79,7 +79,7 @@ Section ZModulo.
auto.
Qed.
- Definition znz_of_pos x :=
+ Definition znz_of_pos x :=
let (q,r) := Zdiv_eucl_POS x wB in (N_of_Z q, r).
Lemma spec_of_pos : forall p,
@@ -90,10 +90,10 @@ Section ZModulo.
destruct (Zdiv_eucl_POS p wB); simpl; destruct 1.
unfold znz_to_Z; rewrite Zmod_small; auto.
assert (0 <= z).
- replace z with (Zpos p / wB) by
+ 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.
Qed.
@@ -110,7 +110,7 @@ Section ZModulo.
Definition znz_0 := 0.
Definition znz_1 := 1.
Definition znz_Bm1 := wB - 1.
-
+
Lemma spec_0 : [|znz_0|] = 0.
Proof.
unfold znz_to_Z, znz_0.
@@ -121,7 +121,7 @@ Section ZModulo.
Proof.
unfold znz_to_Z, znz_1.
apply Zmod_small; split; auto with zarith.
- unfold wB, base.
+ unfold wB, base.
apply Zlt_trans with (Zpos digits); auto.
apply Zpower2_lt_lin; auto with zarith.
Qed.
@@ -138,7 +138,7 @@ Section ZModulo.
Definition znz_compare x y := Zcompare [|x|] [|y|].
- Lemma spec_compare : forall x y,
+ Lemma spec_compare : forall x y,
match znz_compare x y with
| Eq => [|x|] = [|y|]
| Lt => [|x|] < [|y|]
@@ -150,19 +150,19 @@ Section ZModulo.
intros; apply Zcompare_Eq_eq; auto.
Qed.
- Definition znz_eq0 x :=
+ Definition znz_eq0 x :=
match [|x|] with Z0 => true | _ => false end.
-
+
Lemma spec_eq0 : forall x, znz_eq0 x = true -> [|x|] = 0.
Proof.
unfold znz_eq0; intros; now destruct [|x|].
Qed.
- Definition znz_opp_c x :=
+ Definition znz_opp_c x :=
if znz_eq0 x then C0 0 else C1 (- x).
Definition znz_opp x := - x.
Definition znz_opp_carry x := - x - 1.
-
+
Lemma spec_opp_c : forall x, [-|znz_opp_c x|] = -[|x|].
Proof.
intros; unfold znz_opp_c, znz_to_Z; auto.
@@ -180,7 +180,7 @@ Section ZModulo.
change ((- x) mod wB = (0 - (x mod wB)) mod wB).
rewrite Zminus_mod_idemp_r; simpl; auto.
Qed.
-
+
Lemma spec_opp_carry : forall x, [|znz_opp_carry x|] = wB - [|x|] - 1.
Proof.
intros; unfold znz_opp_carry, znz_to_Z; auto.
@@ -194,15 +194,15 @@ Section ZModulo.
generalize (Z_mod_lt x wB wB_pos); omega.
Qed.
- Definition znz_succ_c x :=
- let y := Zsucc x in
+ Definition znz_succ_c x :=
+ let y := Zsucc x in
if znz_eq0 y then C1 0 else C0 y.
- Definition znz_add_c x y :=
- let z := [|x|] + [|y|] in
+ Definition znz_add_c x y :=
+ let z := [|x|] + [|y|] in
if Z_lt_le_dec z wB then C0 z else C1 (z-wB).
- Definition znz_add_carry_c x y :=
+ Definition znz_add_carry_c x y :=
let z := [|x|]+[|y|]+1 in
if Z_lt_le_dec z wB then C0 z else C1 (z-wB).
@@ -210,7 +210,7 @@ Section ZModulo.
Definition znz_add := Zplus.
Definition znz_add_carry x y := x + y + 1.
- Lemma Zmod_equal :
+ Lemma Zmod_equal :
forall x y z, z>0 -> (x-y) mod z = 0 -> x mod z = y mod z.
Proof.
intros.
@@ -225,12 +225,12 @@ Section ZModulo.
Proof.
intros; unfold znz_succ_c, znz_to_Z, Zsucc.
case_eq (znz_eq0 (x+1)); intros; unfold interp_carry.
-
+
rewrite Zmult_1_l.
replace (wB + 0 mod wB) with wB by auto with zarith.
symmetry; rewrite Zeq_plus_swap.
assert ((x+1) mod wB = 0) by (apply spec_eq0; auto).
- replace (wB-1) with ((wB-1) mod wB) by
+ replace (wB-1) with ((wB-1) mod wB) by
(apply Zmod_small; generalize wB_pos; omega).
rewrite <- Zminus_mod_idemp_l; rewrite Z_mod_same; simpl; auto.
apply Zmod_equal; auto.
@@ -289,15 +289,15 @@ Section ZModulo.
rewrite Zplus_mod_idemp_l; auto.
Qed.
- Definition znz_pred_c x :=
+ Definition znz_pred_c x :=
if znz_eq0 x then C1 (wB-1) else C0 (x-1).
- Definition znz_sub_c x y :=
- let z := [|x|]-[|y|] in
+ Definition znz_sub_c x y :=
+ let z := [|x|]-[|y|] in
if Z_lt_le_dec z 0 then C1 (wB+z) else C0 z.
- Definition znz_sub_carry_c x y :=
- let z := [|x|]-[|y|]-1 in
+ Definition znz_sub_carry_c x y :=
+ let z := [|x|]-[|y|]-1 in
if Z_lt_le_dec z 0 then C1 (wB+z) else C0 z.
Definition znz_pred := Zpred.
@@ -323,7 +323,7 @@ Section ZModulo.
Proof.
intros; unfold znz_sub_c, znz_to_Z, interp_carry.
destruct Z_lt_le_dec.
- replace ((wB + (x mod wB - y mod wB)) mod wB) with
+ replace ((wB + (x mod wB - y mod wB)) mod wB) with
(wB + (x mod wB - y mod wB)).
omega.
symmetry; apply Zmod_small.
@@ -337,7 +337,7 @@ Section ZModulo.
Proof.
intros; unfold znz_sub_carry_c, znz_to_Z, interp_carry.
destruct Z_lt_le_dec.
- replace ((wB + (x mod wB - y mod wB - 1)) mod wB) with
+ replace ((wB + (x mod wB - y mod wB - 1)) mod wB) with
(wB + (x mod wB - y mod wB -1)).
omega.
symmetry; apply Zmod_small.
@@ -358,7 +358,7 @@ Section ZModulo.
intros; unfold znz_sub, znz_to_Z; apply Zminus_mod.
Qed.
- Lemma spec_sub_carry :
+ Lemma spec_sub_carry :
forall x y, [|znz_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB.
Proof.
intros; unfold znz_sub_carry, znz_to_Z.
@@ -367,15 +367,15 @@ Section ZModulo.
rewrite Zminus_mod_idemp_l.
auto.
Qed.
-
- Definition znz_mul_c x y :=
+
+ Definition znz_mul_c x y :=
let (h,l) := Zdiv_eucl ([|x|]*[|y|]) wB in
if znz_eq0 h then if znz_eq0 l then W0 else WW h l else WW h l.
Definition znz_mul := Zmult.
Definition znz_square_c x := znz_mul_c x x.
-
+
Lemma spec_mul_c : forall x y, [|| znz_mul_c x y ||] = [|x|] * [|y|].
Proof.
intros; unfold znz_mul_c, zn2z_to_Z.
@@ -426,7 +426,7 @@ Section ZModulo.
destruct Zdiv_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|];
+ apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|];
auto with zarith.
assert ([|q|]=q).
apply Zmod_small.
@@ -453,7 +453,7 @@ Section ZModulo.
Definition znz_mod x y := [|x|] mod [|y|].
Definition znz_mod_gt x y := [|x|] mod [|y|].
-
+
Lemma spec_mod : forall a b, 0 < [|b|] ->
[|znz_mod a b|] = [|a|] mod [|b|].
Proof.
@@ -469,7 +469,7 @@ Section ZModulo.
Proof.
intros; apply spec_mod; auto.
Qed.
-
+
Definition znz_gcd x y := Zgcd [|x|] [|y|].
Definition znz_gcd_gt x y := Zgcd [|x|] [|y|].
@@ -516,7 +516,7 @@ Section ZModulo.
intros. apply spec_gcd; auto.
Qed.
- Definition znz_div21 a1 a2 b :=
+ Definition znz_div21 a1 a2 b :=
Zdiv_eucl ([|a1|]*wB+[|a2|]) [|b|].
Lemma spec_div21 : forall a1 a2 b,
@@ -537,7 +537,7 @@ Section ZModulo.
destruct Zdiv_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|];
+ apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|];
auto with zarith.
assert ([|q|]=q).
apply Zmod_small.
@@ -546,7 +546,6 @@ Section ZModulo.
apply Z_div_pos; auto with zarith.
subst a; auto with zarith.
apply Zdiv_lt_upper_bound; auto with zarith.
- subst a; 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.
@@ -577,7 +576,7 @@ Section ZModulo.
apply Zmod_le; auto with zarith.
Qed.
- Definition znz_is_even x :=
+ Definition znz_is_even x :=
if Z_eq_dec ([|x|] mod 2) 0 then true else false.
Lemma spec_is_even : forall x,
@@ -587,7 +586,7 @@ Section ZModulo.
generalize (Z_mod_lt [|x|] 2); omega.
Qed.
- Definition znz_sqrt x := Zsqrt_plain [|x|].
+ Definition znz_sqrt x := Zsqrt_plain [|x|].
Lemma spec_sqrt : forall x,
[|znz_sqrt x|] ^ 2 <= [|x|] < ([|znz_sqrt x|] + 1) ^ 2.
Proof.
@@ -610,12 +609,12 @@ Section ZModulo.
generalize wB_pos; auto with zarith.
Qed.
- Definition znz_sqrt2 x y :=
- let z := [|x|]*wB+[|y|] in
- match z with
+ Definition znz_sqrt2 x y :=
+ let z := [|x|]*wB+[|y|] in
+ match z with
| Z0 => (0, C0 0)
- | Zpos p =>
- let (s,r,_,_) := sqrtrempos p in
+ | Zpos p =>
+ let (s,r,_,_) := sqrtrempos p in
(s, if Z_lt_le_dec r wB then C0 r else C1 (r-wB))
| Zneg _ => (0, C0 0)
end.
@@ -652,7 +651,7 @@ Section ZModulo.
rewrite Zpower_2; auto with zarith.
replace [|r-wB|] with (r-wB) by (symmetry; apply Zmod_small; auto with zarith).
rewrite Zpower_2; omega.
-
+
assert (0<=Zneg p).
rewrite Heqz; generalize wB_pos; auto with zarith.
compute in H0; elim H0; auto.
@@ -666,8 +665,8 @@ Section ZModulo.
apply two_power_pos_correct.
Qed.
- Definition znz_head0 x := match [|x|] with
- | Z0 => znz_zdigits
+ Definition znz_head0 x := match [|x|] with
+ | Z0 => znz_zdigits
| Zpos p => znz_zdigits - log_inf p - 1
| _ => 0
end.
@@ -696,7 +695,7 @@ Section ZModulo.
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.
-
+
simpl; intros; destruct p; compute; auto with zarith.
Qed.
@@ -731,8 +730,8 @@ Section ZModulo.
by ring.
unfold wB, base, znz_zdigits; auto with zarith.
apply Zmult_le_compat; auto with zarith.
-
- apply Zlt_le_trans
+
+ apply Zlt_le_trans
with (2^(znz_zdigits - log_inf p - 1)*(2^(Zsucc (log_inf p)))).
apply Zmult_lt_compat_l; auto with zarith.
rewrite <- Zpower_exp; auto with zarith.
@@ -741,17 +740,17 @@ Section ZModulo.
unfold wB, base, znz_zdigits; auto with zarith.
Qed.
- Fixpoint Ptail p := match p with
+ Fixpoint Ptail p := match p with
| xO p => (Ptail p)+1
| _ => 0
- end.
+ end.
Lemma Ptail_pos : forall p, 0 <= Ptail p.
Proof.
induction p; simpl; auto with zarith.
Qed.
Hint Resolve Ptail_pos.
-
+
Lemma Ptail_bounded : forall p d, Zpos p < 2^(Zpos d) -> Ptail p < Zpos d.
Proof.
induction p; try (compute; auto; fail).
@@ -776,7 +775,7 @@ Section ZModulo.
Qed.
Definition znz_tail0 x :=
- match [|x|] with
+ match [|x|] with
| Z0 => znz_zdigits
| Zpos p => Ptail p
| Zneg _ => 0
@@ -789,7 +788,7 @@ Section ZModulo.
apply spec_zdigits.
Qed.
- Lemma spec_tail0 : forall x, 0 < [|x|] ->
+ Lemma spec_tail0 : forall x, 0 < [|x|] ->
exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|znz_tail0 x|]).
Proof.
intros; unfold znz_tail0.
@@ -819,7 +818,7 @@ Section ZModulo.
(** Let's now group everything in two records *)
- Definition zmod_op := mk_znz_op
+ Definition zmod_op := mk_znz_op
(znz_digits : positive)
(znz_zdigits: znz)
(znz_to_Z : znz -> Z)
@@ -860,11 +859,11 @@ Section ZModulo.
(znz_div_gt : znz -> znz -> znz * znz)
(znz_div : znz -> znz -> znz * znz)
- (znz_mod_gt : znz -> znz -> znz)
- (znz_mod : znz -> znz -> znz)
+ (znz_mod_gt : znz -> znz -> znz)
+ (znz_mod : znz -> znz -> znz)
(znz_gcd_gt : znz -> znz -> znz)
- (znz_gcd : znz -> znz -> znz)
+ (znz_gcd : znz -> znz -> znz)
(znz_add_mul_div : znz -> znz -> znz -> znz)
(znz_pos_mod : znz -> znz -> znz)
@@ -879,54 +878,54 @@ Section ZModulo.
spec_more_than_1_digit
spec_0
- spec_1
- spec_Bm1
-
- spec_compare
- spec_eq0
-
- spec_opp_c
- spec_opp
- spec_opp_carry
-
- spec_succ_c
- spec_add_c
- spec_add_carry_c
- spec_succ
- spec_add
- spec_add_carry
-
- spec_pred_c
- spec_sub_c
- spec_sub_carry_c
- spec_pred
- spec_sub
- spec_sub_carry
-
- spec_mul_c
- spec_mul
- spec_square_c
-
- spec_div21
- spec_div_gt
- spec_div
-
- spec_mod_gt
- spec_mod
-
- spec_gcd_gt
- spec_gcd
-
- spec_head00
- spec_head0
- spec_tail00
- spec_tail0
-
- spec_add_mul_div
- spec_pos_mod
-
- spec_is_even
- spec_sqrt2
+ spec_1
+ spec_Bm1
+
+ spec_compare
+ spec_eq0
+
+ spec_opp_c
+ spec_opp
+ spec_opp_carry
+
+ spec_succ_c
+ spec_add_c
+ spec_add_carry_c
+ spec_succ
+ spec_add
+ spec_add_carry
+
+ spec_pred_c
+ spec_sub_c
+ spec_sub_carry_c
+ spec_pred
+ spec_sub
+ spec_sub_carry
+
+ spec_mul_c
+ spec_mul
+ spec_square_c
+
+ spec_div21
+ spec_div_gt
+ spec_div
+
+ spec_mod_gt
+ spec_mod
+
+ spec_gcd_gt
+ spec_gcd
+
+ spec_head00
+ spec_head0
+ spec_tail00
+ spec_tail0
+
+ spec_add_mul_div
+ spec_pos_mod
+
+ spec_is_even
+ spec_sqrt2
spec_sqrt.
End ZModulo.
@@ -935,7 +934,7 @@ End ZModulo.
Module Type PositiveNotOne.
Parameter p : positive.
- Axiom not_one : p<> 1%positive.
+ Axiom not_one : p<> 1%positive.
End PositiveNotOne.
Module ZModuloCyclicType (P:PositiveNotOne) <: CyclicType.
diff --git a/theories/Numbers/Integer/Abstract/ZAdd.v b/theories/Numbers/Integer/Abstract/ZAdd.v
index df941d90..5663408d 100644
--- a/theories/Numbers/Integer/Abstract/ZAdd.v
+++ b/theories/Numbers/Integer/Abstract/ZAdd.v
@@ -8,338 +8,286 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZAdd.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export ZBase.
-Module ZAddPropFunct (Import ZAxiomsMod : ZAxiomsSig).
-Module Export ZBasePropMod := ZBasePropFunct ZAxiomsMod.
-Open Local Scope IntScope.
+Module ZAddPropFunct (Import Z : ZAxiomsSig').
+Include ZBasePropFunct Z.
-Theorem Zadd_wd :
- forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> n1 + m1 == n2 + m2.
-Proof NZadd_wd.
+(** Theorems that are either not valid on N or have different proofs
+ on N and Z *)
-Theorem Zadd_0_l : forall n : Z, 0 + n == n.
-Proof NZadd_0_l.
-
-Theorem Zadd_succ_l : forall n m : Z, (S n) + m == S (n + m).
-Proof NZadd_succ_l.
-
-Theorem Zsub_0_r : forall n : Z, n - 0 == n.
-Proof NZsub_0_r.
-
-Theorem Zsub_succ_r : forall n m : Z, n - (S m) == P (n - m).
-Proof NZsub_succ_r.
-
-Theorem Zopp_0 : - 0 == 0.
-Proof Zopp_0.
-
-Theorem Zopp_succ : forall n : Z, - (S n) == P (- n).
-Proof Zopp_succ.
-
-(* Theorems that are valid for both natural numbers and integers *)
-
-Theorem Zadd_0_r : forall n : Z, n + 0 == n.
-Proof NZadd_0_r.
-
-Theorem Zadd_succ_r : forall n m : Z, n + S m == S (n + m).
-Proof NZadd_succ_r.
-
-Theorem Zadd_comm : forall n m : Z, n + m == m + n.
-Proof NZadd_comm.
-
-Theorem Zadd_assoc : forall n m p : Z, n + (m + p) == (n + m) + p.
-Proof NZadd_assoc.
-
-Theorem Zadd_shuffle1 : forall n m p q : Z, (n + m) + (p + q) == (n + p) + (m + q).
-Proof NZadd_shuffle1.
-
-Theorem Zadd_shuffle2 : forall n m p q : Z, (n + m) + (p + q) == (n + q) + (m + p).
-Proof NZadd_shuffle2.
-
-Theorem Zadd_1_l : forall n : Z, 1 + n == S n.
-Proof NZadd_1_l.
-
-Theorem Zadd_1_r : forall n : Z, n + 1 == S n.
-Proof NZadd_1_r.
-
-Theorem Zadd_cancel_l : forall n m p : Z, p + n == p + m <-> n == m.
-Proof NZadd_cancel_l.
-
-Theorem Zadd_cancel_r : forall n m p : Z, n + p == m + p <-> n == m.
-Proof NZadd_cancel_r.
-
-(* Theorems that are either not valid on N or have different proofs on N and Z *)
-
-Theorem Zadd_pred_l : forall n m : Z, P n + m == P (n + m).
+Theorem add_pred_l : forall n m, P n + m == P (n + m).
Proof.
intros n m.
-rewrite <- (Zsucc_pred n) at 2.
-rewrite Zadd_succ_l. now rewrite Zpred_succ.
+rewrite <- (succ_pred n) at 2.
+rewrite add_succ_l. now rewrite pred_succ.
Qed.
-Theorem Zadd_pred_r : forall n m : Z, n + P m == P (n + m).
+Theorem add_pred_r : forall n m, n + P m == P (n + m).
Proof.
-intros n m; rewrite (Zadd_comm n (P m)), (Zadd_comm n m);
-apply Zadd_pred_l.
+intros n m; rewrite (add_comm n (P m)), (add_comm n m);
+apply add_pred_l.
Qed.
-Theorem Zadd_opp_r : forall n m : Z, n + (- m) == n - m.
+Theorem add_opp_r : forall n m, n + (- m) == n - m.
Proof.
-NZinduct m.
-rewrite Zopp_0; rewrite Zsub_0_r; now rewrite Zadd_0_r.
-intro m. rewrite Zopp_succ, Zsub_succ_r, Zadd_pred_r; now rewrite Zpred_inj_wd.
+nzinduct m.
+rewrite opp_0; rewrite sub_0_r; now rewrite add_0_r.
+intro m. rewrite opp_succ, sub_succ_r, add_pred_r; now rewrite pred_inj_wd.
Qed.
-Theorem Zsub_0_l : forall n : Z, 0 - n == - n.
+Theorem sub_0_l : forall n, 0 - n == - n.
Proof.
-intro n; rewrite <- Zadd_opp_r; now rewrite Zadd_0_l.
+intro n; rewrite <- add_opp_r; now rewrite add_0_l.
Qed.
-Theorem Zsub_succ_l : forall n m : Z, S n - m == S (n - m).
+Theorem sub_succ_l : forall n m, S n - m == S (n - m).
Proof.
-intros n m; do 2 rewrite <- Zadd_opp_r; now rewrite Zadd_succ_l.
+intros n m; do 2 rewrite <- add_opp_r; now rewrite add_succ_l.
Qed.
-Theorem Zsub_pred_l : forall n m : Z, P n - m == P (n - m).
+Theorem sub_pred_l : forall n m, P n - m == P (n - m).
Proof.
-intros n m. rewrite <- (Zsucc_pred n) at 2.
-rewrite Zsub_succ_l; now rewrite Zpred_succ.
+intros n m. rewrite <- (succ_pred n) at 2.
+rewrite sub_succ_l; now rewrite pred_succ.
Qed.
-Theorem Zsub_pred_r : forall n m : Z, n - (P m) == S (n - m).
+Theorem sub_pred_r : forall n m, n - (P m) == S (n - m).
Proof.
-intros n m. rewrite <- (Zsucc_pred m) at 2.
-rewrite Zsub_succ_r; now rewrite Zsucc_pred.
+intros n m. rewrite <- (succ_pred m) at 2.
+rewrite sub_succ_r; now rewrite succ_pred.
Qed.
-Theorem Zopp_pred : forall n : Z, - (P n) == S (- n).
+Theorem opp_pred : forall n, - (P n) == S (- n).
Proof.
-intro n. rewrite <- (Zsucc_pred n) at 2.
-rewrite Zopp_succ. now rewrite Zsucc_pred.
+intro n. rewrite <- (succ_pred n) at 2.
+rewrite opp_succ. now rewrite succ_pred.
Qed.
-Theorem Zsub_diag : forall n : Z, n - n == 0.
+Theorem sub_diag : forall n, n - n == 0.
Proof.
-NZinduct n.
-now rewrite Zsub_0_r.
-intro n. rewrite Zsub_succ_r, Zsub_succ_l; now rewrite Zpred_succ.
+nzinduct n.
+now rewrite sub_0_r.
+intro n. rewrite sub_succ_r, sub_succ_l; now rewrite pred_succ.
Qed.
-Theorem Zadd_opp_diag_l : forall n : Z, - n + n == 0.
+Theorem add_opp_diag_l : forall n, - n + n == 0.
Proof.
-intro n; now rewrite Zadd_comm, Zadd_opp_r, Zsub_diag.
+intro n; now rewrite add_comm, add_opp_r, sub_diag.
Qed.
-Theorem Zadd_opp_diag_r : forall n : Z, n + (- n) == 0.
+Theorem add_opp_diag_r : forall n, n + (- n) == 0.
Proof.
-intro n; rewrite Zadd_comm; apply Zadd_opp_diag_l.
+intro n; rewrite add_comm; apply add_opp_diag_l.
Qed.
-Theorem Zadd_opp_l : forall n m : Z, - m + n == n - m.
+Theorem add_opp_l : forall n m, - m + n == n - m.
Proof.
-intros n m; rewrite <- Zadd_opp_r; now rewrite Zadd_comm.
+intros n m; rewrite <- add_opp_r; now rewrite add_comm.
Qed.
-Theorem Zadd_sub_assoc : forall n m p : Z, n + (m - p) == (n + m) - p.
+Theorem add_sub_assoc : forall n m p, n + (m - p) == (n + m) - p.
Proof.
-intros n m p; do 2 rewrite <- Zadd_opp_r; now rewrite Zadd_assoc.
+intros n m p; do 2 rewrite <- add_opp_r; now rewrite add_assoc.
Qed.
-Theorem Zopp_involutive : forall n : Z, - (- n) == n.
+Theorem opp_involutive : forall n, - (- n) == n.
Proof.
-NZinduct n.
-now do 2 rewrite Zopp_0.
-intro n. rewrite Zopp_succ, Zopp_pred; now rewrite Zsucc_inj_wd.
+nzinduct n.
+now do 2 rewrite opp_0.
+intro n. rewrite opp_succ, opp_pred; now rewrite succ_inj_wd.
Qed.
-Theorem Zopp_add_distr : forall n m : Z, - (n + m) == - n + (- m).
+Theorem opp_add_distr : forall n m, - (n + m) == - n + (- m).
Proof.
-intros n m; NZinduct n.
-rewrite Zopp_0; now do 2 rewrite Zadd_0_l.
-intro n. rewrite Zadd_succ_l; do 2 rewrite Zopp_succ; rewrite Zadd_pred_l.
-now rewrite Zpred_inj_wd.
+intros n m; nzinduct n.
+rewrite opp_0; now do 2 rewrite add_0_l.
+intro n. rewrite add_succ_l; do 2 rewrite opp_succ; rewrite add_pred_l.
+now rewrite pred_inj_wd.
Qed.
-Theorem Zopp_sub_distr : forall n m : Z, - (n - m) == - n + m.
+Theorem opp_sub_distr : forall n m, - (n - m) == - n + m.
Proof.
-intros n m; rewrite <- Zadd_opp_r, Zopp_add_distr.
-now rewrite Zopp_involutive.
+intros n m; rewrite <- add_opp_r, opp_add_distr.
+now rewrite opp_involutive.
Qed.
-Theorem Zopp_inj : forall n m : Z, - n == - m -> n == m.
+Theorem opp_inj : forall n m, - n == - m -> n == m.
Proof.
-intros n m H. apply Zopp_wd in H. now do 2 rewrite Zopp_involutive in H.
+intros n m H. apply opp_wd in H. now do 2 rewrite opp_involutive in H.
Qed.
-Theorem Zopp_inj_wd : forall n m : Z, - n == - m <-> n == m.
+Theorem opp_inj_wd : forall n m, - n == - m <-> n == m.
Proof.
-intros n m; split; [apply Zopp_inj | apply Zopp_wd].
+intros n m; split; [apply opp_inj | apply opp_wd].
Qed.
-Theorem Zeq_opp_l : forall n m : Z, - n == m <-> n == - m.
+Theorem eq_opp_l : forall n m, - n == m <-> n == - m.
Proof.
-intros n m. now rewrite <- (Zopp_inj_wd (- n) m), Zopp_involutive.
+intros n m. now rewrite <- (opp_inj_wd (- n) m), opp_involutive.
Qed.
-Theorem Zeq_opp_r : forall n m : Z, n == - m <-> - n == m.
+Theorem eq_opp_r : forall n m, n == - m <-> - n == m.
Proof.
-symmetry; apply Zeq_opp_l.
+symmetry; apply eq_opp_l.
Qed.
-Theorem Zsub_add_distr : forall n m p : Z, n - (m + p) == (n - m) - p.
+Theorem sub_add_distr : forall n m p, n - (m + p) == (n - m) - p.
Proof.
-intros n m p; rewrite <- Zadd_opp_r, Zopp_add_distr, Zadd_assoc.
-now do 2 rewrite Zadd_opp_r.
+intros n m p; rewrite <- add_opp_r, opp_add_distr, add_assoc.
+now do 2 rewrite add_opp_r.
Qed.
-Theorem Zsub_sub_distr : forall n m p : Z, n - (m - p) == (n - m) + p.
+Theorem sub_sub_distr : forall n m p, n - (m - p) == (n - m) + p.
Proof.
-intros n m p; rewrite <- Zadd_opp_r, Zopp_sub_distr, Zadd_assoc.
-now rewrite Zadd_opp_r.
+intros n m p; rewrite <- add_opp_r, opp_sub_distr, add_assoc.
+now rewrite add_opp_r.
Qed.
-Theorem sub_opp_l : forall n m : Z, - n - m == - m - n.
+Theorem sub_opp_l : forall n m, - n - m == - m - n.
Proof.
-intros n m. do 2 rewrite <- Zadd_opp_r. now rewrite Zadd_comm.
+intros n m. do 2 rewrite <- add_opp_r. now rewrite add_comm.
Qed.
-Theorem Zsub_opp_r : forall n m : Z, n - (- m) == n + m.
+Theorem sub_opp_r : forall n m, n - (- m) == n + m.
Proof.
-intros n m; rewrite <- Zadd_opp_r; now rewrite Zopp_involutive.
+intros n m; rewrite <- add_opp_r; now rewrite opp_involutive.
Qed.
-Theorem Zadd_sub_swap : forall n m p : Z, n + m - p == n - p + m.
+Theorem add_sub_swap : forall n m p, n + m - p == n - p + m.
Proof.
-intros n m p. rewrite <- Zadd_sub_assoc, <- (Zadd_opp_r n p), <- Zadd_assoc.
-now rewrite Zadd_opp_l.
+intros n m p. rewrite <- add_sub_assoc, <- (add_opp_r n p), <- add_assoc.
+now rewrite add_opp_l.
Qed.
-Theorem Zsub_cancel_l : forall n m p : Z, n - m == n - p <-> m == p.
+Theorem sub_cancel_l : forall n m p, n - m == n - p <-> m == p.
Proof.
-intros n m p. rewrite <- (Zadd_cancel_l (n - m) (n - p) (- n)).
-do 2 rewrite Zadd_sub_assoc. rewrite Zadd_opp_diag_l; do 2 rewrite Zsub_0_l.
-apply Zopp_inj_wd.
+intros n m p. rewrite <- (add_cancel_l (n - m) (n - p) (- n)).
+do 2 rewrite add_sub_assoc. rewrite add_opp_diag_l; do 2 rewrite sub_0_l.
+apply opp_inj_wd.
Qed.
-Theorem Zsub_cancel_r : forall n m p : Z, n - p == m - p <-> n == m.
+Theorem sub_cancel_r : forall n m p, n - p == m - p <-> n == m.
Proof.
intros n m p.
-stepl (n - p + p == m - p + p) by apply Zadd_cancel_r.
-now do 2 rewrite <- Zsub_sub_distr, Zsub_diag, Zsub_0_r.
+stepl (n - p + p == m - p + p) by apply add_cancel_r.
+now do 2 rewrite <- sub_sub_distr, sub_diag, sub_0_r.
Qed.
-(* The next several theorems are devoted to moving terms from one side of
-an equation to the other. The name contains the operation in the original
-equation (add or sub) and the indication whether the left or right term
-is moved. *)
+(** The next several theorems are devoted to moving terms from one
+ side of an equation to the other. The name contains the operation
+ in the original equation ([add] or [sub]) and the indication
+ whether the left or right term is moved. *)
-Theorem Zadd_move_l : forall n m p : Z, n + m == p <-> m == p - n.
+Theorem add_move_l : forall n m p, n + m == p <-> m == p - n.
Proof.
intros n m p.
-stepl (n + m - n == p - n) by apply Zsub_cancel_r.
-now rewrite Zadd_comm, <- Zadd_sub_assoc, Zsub_diag, Zadd_0_r.
+stepl (n + m - n == p - n) by apply sub_cancel_r.
+now rewrite add_comm, <- add_sub_assoc, sub_diag, add_0_r.
Qed.
-Theorem Zadd_move_r : forall n m p : Z, n + m == p <-> n == p - m.
+Theorem add_move_r : forall n m p, n + m == p <-> n == p - m.
Proof.
-intros n m p; rewrite Zadd_comm; now apply Zadd_move_l.
+intros n m p; rewrite add_comm; now apply add_move_l.
Qed.
-(* The two theorems above do not allow rewriting subformulas of the form
-n - m == p to n == p + m since subtraction is in the right-hand side of
-the equation. Hence the following two theorems. *)
+(** The two theorems above do not allow rewriting subformulas of the
+ form [n - m == p] to [n == p + m] since subtraction is in the
+ right-hand side of the equation. Hence the following two
+ theorems. *)
-Theorem Zsub_move_l : forall n m p : Z, n - m == p <-> - m == p - n.
+Theorem sub_move_l : forall n m p, n - m == p <-> - m == p - n.
Proof.
-intros n m p; rewrite <- (Zadd_opp_r n m); apply Zadd_move_l.
+intros n m p; rewrite <- (add_opp_r n m); apply add_move_l.
Qed.
-Theorem Zsub_move_r : forall n m p : Z, n - m == p <-> n == p + m.
+Theorem sub_move_r : forall n m p, n - m == p <-> n == p + m.
Proof.
-intros n m p; rewrite <- (Zadd_opp_r n m). now rewrite Zadd_move_r, Zsub_opp_r.
+intros n m p; rewrite <- (add_opp_r n m). now rewrite add_move_r, sub_opp_r.
Qed.
-Theorem Zadd_move_0_l : forall n m : Z, n + m == 0 <-> m == - n.
+Theorem add_move_0_l : forall n m, n + m == 0 <-> m == - n.
Proof.
-intros n m; now rewrite Zadd_move_l, Zsub_0_l.
+intros n m; now rewrite add_move_l, sub_0_l.
Qed.
-Theorem Zadd_move_0_r : forall n m : Z, n + m == 0 <-> n == - m.
+Theorem add_move_0_r : forall n m, n + m == 0 <-> n == - m.
Proof.
-intros n m; now rewrite Zadd_move_r, Zsub_0_l.
+intros n m; now rewrite add_move_r, sub_0_l.
Qed.
-Theorem Zsub_move_0_l : forall n m : Z, n - m == 0 <-> - m == - n.
+Theorem sub_move_0_l : forall n m, n - m == 0 <-> - m == - n.
Proof.
-intros n m. now rewrite Zsub_move_l, Zsub_0_l.
+intros n m. now rewrite sub_move_l, sub_0_l.
Qed.
-Theorem Zsub_move_0_r : forall n m : Z, n - m == 0 <-> n == m.
+Theorem sub_move_0_r : forall n m, n - m == 0 <-> n == m.
Proof.
-intros n m. now rewrite Zsub_move_r, Zadd_0_l.
+intros n m. now rewrite sub_move_r, add_0_l.
Qed.
-(* The following section is devoted to cancellation of like terms. The name
-includes the first operator and the position of the term being canceled. *)
+(** The following section is devoted to cancellation of like
+ terms. The name includes the first operator and the position of
+ the term being canceled. *)
-Theorem Zadd_simpl_l : forall n m : Z, n + m - n == m.
+Theorem add_simpl_l : forall n m, n + m - n == m.
Proof.
-intros; now rewrite Zadd_sub_swap, Zsub_diag, Zadd_0_l.
+intros; now rewrite add_sub_swap, sub_diag, add_0_l.
Qed.
-Theorem Zadd_simpl_r : forall n m : Z, n + m - m == n.
+Theorem add_simpl_r : forall n m, n + m - m == n.
Proof.
-intros; now rewrite <- Zadd_sub_assoc, Zsub_diag, Zadd_0_r.
+intros; now rewrite <- add_sub_assoc, sub_diag, add_0_r.
Qed.
-Theorem Zsub_simpl_l : forall n m : Z, - n - m + n == - m.
+Theorem sub_simpl_l : forall n m, - n - m + n == - m.
Proof.
-intros; now rewrite <- Zadd_sub_swap, Zadd_opp_diag_l, Zsub_0_l.
+intros; now rewrite <- add_sub_swap, add_opp_diag_l, sub_0_l.
Qed.
-Theorem Zsub_simpl_r : forall n m : Z, n - m + m == n.
+Theorem sub_simpl_r : forall n m, n - m + m == n.
Proof.
-intros; now rewrite <- Zsub_sub_distr, Zsub_diag, Zsub_0_r.
+intros; now rewrite <- sub_sub_distr, sub_diag, sub_0_r.
Qed.
-(* Now we have two sums or differences; the name includes the two operators
-and the position of the terms being canceled *)
+(** Now we have two sums or differences; the name includes the two
+ operators and the position of the terms being canceled *)
-Theorem Zadd_add_simpl_l_l : forall n m p : Z, (n + m) - (n + p) == m - p.
+Theorem add_add_simpl_l_l : forall n m p, (n + m) - (n + p) == m - p.
Proof.
-intros n m p. now rewrite (Zadd_comm n m), <- Zadd_sub_assoc,
-Zsub_add_distr, Zsub_diag, Zsub_0_l, Zadd_opp_r.
+intros n m p. now rewrite (add_comm n m), <- add_sub_assoc,
+sub_add_distr, sub_diag, sub_0_l, add_opp_r.
Qed.
-Theorem Zadd_add_simpl_l_r : forall n m p : Z, (n + m) - (p + n) == m - p.
+Theorem add_add_simpl_l_r : forall n m p, (n + m) - (p + n) == m - p.
Proof.
-intros n m p. rewrite (Zadd_comm p n); apply Zadd_add_simpl_l_l.
+intros n m p. rewrite (add_comm p n); apply add_add_simpl_l_l.
Qed.
-Theorem Zadd_add_simpl_r_l : forall n m p : Z, (n + m) - (m + p) == n - p.
+Theorem add_add_simpl_r_l : forall n m p, (n + m) - (m + p) == n - p.
Proof.
-intros n m p. rewrite (Zadd_comm n m); apply Zadd_add_simpl_l_l.
+intros n m p. rewrite (add_comm n m); apply add_add_simpl_l_l.
Qed.
-Theorem Zadd_add_simpl_r_r : forall n m p : Z, (n + m) - (p + m) == n - p.
+Theorem add_add_simpl_r_r : forall n m p, (n + m) - (p + m) == n - p.
Proof.
-intros n m p. rewrite (Zadd_comm p m); apply Zadd_add_simpl_r_l.
+intros n m p. rewrite (add_comm p m); apply add_add_simpl_r_l.
Qed.
-Theorem Zsub_add_simpl_r_l : forall n m p : Z, (n - m) + (m + p) == n + p.
+Theorem sub_add_simpl_r_l : forall n m p, (n - m) + (m + p) == n + p.
Proof.
-intros n m p. now rewrite <- Zsub_sub_distr, Zsub_add_distr, Zsub_diag,
-Zsub_0_l, Zsub_opp_r.
+intros n m p. now rewrite <- sub_sub_distr, sub_add_distr, sub_diag,
+sub_0_l, sub_opp_r.
Qed.
-Theorem Zsub_add_simpl_r_r : forall n m p : Z, (n - m) + (p + m) == n + p.
+Theorem sub_add_simpl_r_r : forall n m p, (n - m) + (p + m) == n + p.
Proof.
-intros n m p. rewrite (Zadd_comm p m); apply Zsub_add_simpl_r_l.
+intros n m p. rewrite (add_comm p m); apply sub_add_simpl_r_l.
Qed.
-(* Of course, there are many other variants *)
+(** Of course, there are many other variants *)
End ZAddPropFunct.
diff --git a/theories/Numbers/Integer/Abstract/ZAddOrder.v b/theories/Numbers/Integer/Abstract/ZAddOrder.v
index 101ea634..de12993f 100644
--- a/theories/Numbers/Integer/Abstract/ZAddOrder.v
+++ b/theories/Numbers/Integer/Abstract/ZAddOrder.v
@@ -8,365 +8,292 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZAddOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export ZLt.
-Module ZAddOrderPropFunct (Import ZAxiomsMod : ZAxiomsSig).
-Module Export ZOrderPropMod := ZOrderPropFunct ZAxiomsMod.
-Open Local Scope IntScope.
+Module ZAddOrderPropFunct (Import Z : ZAxiomsSig').
+Include ZOrderPropFunct Z.
-(* Theorems that are true on both natural numbers and integers *)
+(** Theorems that are either not valid on N or have different proofs
+ on N and Z *)
-Theorem Zadd_lt_mono_l : forall n m p : Z, n < m <-> p + n < p + m.
-Proof NZadd_lt_mono_l.
-
-Theorem Zadd_lt_mono_r : forall n m p : Z, n < m <-> n + p < m + p.
-Proof NZadd_lt_mono_r.
-
-Theorem Zadd_lt_mono : forall n m p q : Z, n < m -> p < q -> n + p < m + q.
-Proof NZadd_lt_mono.
-
-Theorem Zadd_le_mono_l : forall n m p : Z, n <= m <-> p + n <= p + m.
-Proof NZadd_le_mono_l.
-
-Theorem Zadd_le_mono_r : forall n m p : Z, n <= m <-> n + p <= m + p.
-Proof NZadd_le_mono_r.
-
-Theorem Zadd_le_mono : forall n m p q : Z, n <= m -> p <= q -> n + p <= m + q.
-Proof NZadd_le_mono.
-
-Theorem Zadd_lt_le_mono : forall n m p q : Z, n < m -> p <= q -> n + p < m + q.
-Proof NZadd_lt_le_mono.
-
-Theorem Zadd_le_lt_mono : forall n m p q : Z, n <= m -> p < q -> n + p < m + q.
-Proof NZadd_le_lt_mono.
-
-Theorem Zadd_pos_pos : forall n m : Z, 0 < n -> 0 < m -> 0 < n + m.
-Proof NZadd_pos_pos.
-
-Theorem Zadd_pos_nonneg : forall n m : Z, 0 < n -> 0 <= m -> 0 < n + m.
-Proof NZadd_pos_nonneg.
-
-Theorem Zadd_nonneg_pos : forall n m : Z, 0 <= n -> 0 < m -> 0 < n + m.
-Proof NZadd_nonneg_pos.
-
-Theorem Zadd_nonneg_nonneg : forall n m : Z, 0 <= n -> 0 <= m -> 0 <= n + m.
-Proof NZadd_nonneg_nonneg.
-
-Theorem Zlt_add_pos_l : forall n m : Z, 0 < n -> m < n + m.
-Proof NZlt_add_pos_l.
-
-Theorem Zlt_add_pos_r : forall n m : Z, 0 < n -> m < m + n.
-Proof NZlt_add_pos_r.
-
-Theorem Zle_lt_add_lt : forall n m p q : Z, n <= m -> p + m < q + n -> p < q.
-Proof NZle_lt_add_lt.
-
-Theorem Zlt_le_add_lt : forall n m p q : Z, n < m -> p + m <= q + n -> p < q.
-Proof NZlt_le_add_lt.
-
-Theorem Zle_le_add_le : forall n m p q : Z, n <= m -> p + m <= q + n -> p <= q.
-Proof NZle_le_add_le.
-
-Theorem Zadd_lt_cases : forall n m p q : Z, n + m < p + q -> n < p \/ m < q.
-Proof NZadd_lt_cases.
-
-Theorem Zadd_le_cases : forall n m p q : Z, n + m <= p + q -> n <= p \/ m <= q.
-Proof NZadd_le_cases.
-
-Theorem Zadd_neg_cases : forall n m : Z, n + m < 0 -> n < 0 \/ m < 0.
-Proof NZadd_neg_cases.
-
-Theorem Zadd_pos_cases : forall n m : Z, 0 < n + m -> 0 < n \/ 0 < m.
-Proof NZadd_pos_cases.
-
-Theorem Zadd_nonpos_cases : forall n m : Z, n + m <= 0 -> n <= 0 \/ m <= 0.
-Proof NZadd_nonpos_cases.
-
-Theorem Zadd_nonneg_cases : forall n m : Z, 0 <= n + m -> 0 <= n \/ 0 <= m.
-Proof NZadd_nonneg_cases.
-
-(* Theorems that are either not valid on N or have different proofs on N and Z *)
-
-Theorem Zadd_neg_neg : forall n m : Z, n < 0 -> m < 0 -> n + m < 0.
+Theorem add_neg_neg : forall n m, n < 0 -> m < 0 -> n + m < 0.
Proof.
-intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_lt_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_lt_mono.
Qed.
-Theorem Zadd_neg_nonpos : forall n m : Z, n < 0 -> m <= 0 -> n + m < 0.
+Theorem add_neg_nonpos : forall n m, n < 0 -> m <= 0 -> n + m < 0.
Proof.
-intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_lt_le_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_lt_le_mono.
Qed.
-Theorem Zadd_nonpos_neg : forall n m : Z, n <= 0 -> m < 0 -> n + m < 0.
+Theorem add_nonpos_neg : forall n m, n <= 0 -> m < 0 -> n + m < 0.
Proof.
-intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_le_lt_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_le_lt_mono.
Qed.
-Theorem Zadd_nonpos_nonpos : forall n m : Z, n <= 0 -> m <= 0 -> n + m <= 0.
+Theorem add_nonpos_nonpos : forall n m, n <= 0 -> m <= 0 -> n + m <= 0.
Proof.
-intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_le_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_le_mono.
Qed.
(** Sub and order *)
-Theorem Zlt_0_sub : forall n m : Z, 0 < m - n <-> n < m.
+Theorem lt_0_sub : forall n m, 0 < m - n <-> n < m.
Proof.
-intros n m. stepl (0 + n < m - n + n) by symmetry; apply Zadd_lt_mono_r.
-rewrite Zadd_0_l; now rewrite Zsub_simpl_r.
+intros n m. stepl (0 + n < m - n + n) by symmetry; apply add_lt_mono_r.
+rewrite add_0_l; now rewrite sub_simpl_r.
Qed.
-Notation Zsub_pos := Zlt_0_sub (only parsing).
+Notation sub_pos := lt_0_sub (only parsing).
-Theorem Zle_0_sub : forall n m : Z, 0 <= m - n <-> n <= m.
+Theorem le_0_sub : forall n m, 0 <= m - n <-> n <= m.
Proof.
-intros n m; stepl (0 + n <= m - n + n) by symmetry; apply Zadd_le_mono_r.
-rewrite Zadd_0_l; now rewrite Zsub_simpl_r.
+intros n m; stepl (0 + n <= m - n + n) by symmetry; apply add_le_mono_r.
+rewrite add_0_l; now rewrite sub_simpl_r.
Qed.
-Notation Zsub_nonneg := Zle_0_sub (only parsing).
+Notation sub_nonneg := le_0_sub (only parsing).
-Theorem Zlt_sub_0 : forall n m : Z, n - m < 0 <-> n < m.
+Theorem lt_sub_0 : forall n m, n - m < 0 <-> n < m.
Proof.
-intros n m. stepl (n - m + m < 0 + m) by symmetry; apply Zadd_lt_mono_r.
-rewrite Zadd_0_l; now rewrite Zsub_simpl_r.
+intros n m. stepl (n - m + m < 0 + m) by symmetry; apply add_lt_mono_r.
+rewrite add_0_l; now rewrite sub_simpl_r.
Qed.
-Notation Zsub_neg := Zlt_sub_0 (only parsing).
+Notation sub_neg := lt_sub_0 (only parsing).
-Theorem Zle_sub_0 : forall n m : Z, n - m <= 0 <-> n <= m.
+Theorem le_sub_0 : forall n m, n - m <= 0 <-> n <= m.
Proof.
-intros n m. stepl (n - m + m <= 0 + m) by symmetry; apply Zadd_le_mono_r.
-rewrite Zadd_0_l; now rewrite Zsub_simpl_r.
+intros n m. stepl (n - m + m <= 0 + m) by symmetry; apply add_le_mono_r.
+rewrite add_0_l; now rewrite sub_simpl_r.
Qed.
-Notation Zsub_nonpos := Zle_sub_0 (only parsing).
+Notation sub_nonpos := le_sub_0 (only parsing).
-Theorem Zopp_lt_mono : forall n m : Z, n < m <-> - m < - n.
+Theorem opp_lt_mono : forall n m, n < m <-> - m < - n.
Proof.
-intros n m. stepr (m + - m < m + - n) by symmetry; apply Zadd_lt_mono_l.
-do 2 rewrite Zadd_opp_r. rewrite Zsub_diag. symmetry; apply Zlt_0_sub.
+intros n m. stepr (m + - m < m + - n) by symmetry; apply add_lt_mono_l.
+do 2 rewrite add_opp_r. rewrite sub_diag. symmetry; apply lt_0_sub.
Qed.
-Theorem Zopp_le_mono : forall n m : Z, n <= m <-> - m <= - n.
+Theorem opp_le_mono : forall n m, n <= m <-> - m <= - n.
Proof.
-intros n m. stepr (m + - m <= m + - n) by symmetry; apply Zadd_le_mono_l.
-do 2 rewrite Zadd_opp_r. rewrite Zsub_diag. symmetry; apply Zle_0_sub.
+intros n m. stepr (m + - m <= m + - n) by symmetry; apply add_le_mono_l.
+do 2 rewrite add_opp_r. rewrite sub_diag. symmetry; apply le_0_sub.
Qed.
-Theorem Zopp_pos_neg : forall n : Z, 0 < - n <-> n < 0.
+Theorem opp_pos_neg : forall n, 0 < - n <-> n < 0.
Proof.
-intro n; rewrite (Zopp_lt_mono n 0); now rewrite Zopp_0.
+intro n; rewrite (opp_lt_mono n 0); now rewrite opp_0.
Qed.
-Theorem Zopp_neg_pos : forall n : Z, - n < 0 <-> 0 < n.
+Theorem opp_neg_pos : forall n, - n < 0 <-> 0 < n.
Proof.
-intro n. rewrite (Zopp_lt_mono 0 n). now rewrite Zopp_0.
+intro n. rewrite (opp_lt_mono 0 n). now rewrite opp_0.
Qed.
-Theorem Zopp_nonneg_nonpos : forall n : Z, 0 <= - n <-> n <= 0.
+Theorem opp_nonneg_nonpos : forall n, 0 <= - n <-> n <= 0.
Proof.
-intro n; rewrite (Zopp_le_mono n 0); now rewrite Zopp_0.
+intro n; rewrite (opp_le_mono n 0); now rewrite opp_0.
Qed.
-Theorem Zopp_nonpos_nonneg : forall n : Z, - n <= 0 <-> 0 <= n.
+Theorem opp_nonpos_nonneg : forall n, - n <= 0 <-> 0 <= n.
Proof.
-intro n. rewrite (Zopp_le_mono 0 n). now rewrite Zopp_0.
+intro n. rewrite (opp_le_mono 0 n). now rewrite opp_0.
Qed.
-Theorem Zsub_lt_mono_l : forall n m p : Z, n < m <-> p - m < p - n.
+Theorem sub_lt_mono_l : forall n m p, n < m <-> p - m < p - n.
Proof.
-intros n m p. do 2 rewrite <- Zadd_opp_r. rewrite <- Zadd_lt_mono_l.
-apply Zopp_lt_mono.
+intros n m p. do 2 rewrite <- add_opp_r. rewrite <- add_lt_mono_l.
+apply opp_lt_mono.
Qed.
-Theorem Zsub_lt_mono_r : forall n m p : Z, n < m <-> n - p < m - p.
+Theorem sub_lt_mono_r : forall n m p, n < m <-> n - p < m - p.
Proof.
-intros n m p; do 2 rewrite <- Zadd_opp_r; apply Zadd_lt_mono_r.
+intros n m p; do 2 rewrite <- add_opp_r; apply add_lt_mono_r.
Qed.
-Theorem Zsub_lt_mono : forall n m p q : Z, n < m -> q < p -> n - p < m - q.
+Theorem sub_lt_mono : forall n m p q, n < m -> q < p -> n - p < m - q.
Proof.
intros n m p q H1 H2.
-apply NZlt_trans with (m - p);
-[now apply -> Zsub_lt_mono_r | now apply -> Zsub_lt_mono_l].
+apply lt_trans with (m - p);
+[now apply -> sub_lt_mono_r | now apply -> sub_lt_mono_l].
Qed.
-Theorem Zsub_le_mono_l : forall n m p : Z, n <= m <-> p - m <= p - n.
+Theorem sub_le_mono_l : forall n m p, n <= m <-> p - m <= p - n.
Proof.
-intros n m p; do 2 rewrite <- Zadd_opp_r; rewrite <- Zadd_le_mono_l;
-apply Zopp_le_mono.
+intros n m p; do 2 rewrite <- add_opp_r; rewrite <- add_le_mono_l;
+apply opp_le_mono.
Qed.
-Theorem Zsub_le_mono_r : forall n m p : Z, n <= m <-> n - p <= m - p.
+Theorem sub_le_mono_r : forall n m p, n <= m <-> n - p <= m - p.
Proof.
-intros n m p; do 2 rewrite <- Zadd_opp_r; apply Zadd_le_mono_r.
+intros n m p; do 2 rewrite <- add_opp_r; apply add_le_mono_r.
Qed.
-Theorem Zsub_le_mono : forall n m p q : Z, n <= m -> q <= p -> n - p <= m - q.
+Theorem sub_le_mono : forall n m p q, n <= m -> q <= p -> n - p <= m - q.
Proof.
intros n m p q H1 H2.
-apply NZle_trans with (m - p);
-[now apply -> Zsub_le_mono_r | now apply -> Zsub_le_mono_l].
+apply le_trans with (m - p);
+[now apply -> sub_le_mono_r | now apply -> sub_le_mono_l].
Qed.
-Theorem Zsub_lt_le_mono : forall n m p q : Z, n < m -> q <= p -> n - p < m - q.
+Theorem sub_lt_le_mono : forall n m p q, n < m -> q <= p -> n - p < m - q.
Proof.
intros n m p q H1 H2.
-apply NZlt_le_trans with (m - p);
-[now apply -> Zsub_lt_mono_r | now apply -> Zsub_le_mono_l].
+apply lt_le_trans with (m - p);
+[now apply -> sub_lt_mono_r | now apply -> sub_le_mono_l].
Qed.
-Theorem Zsub_le_lt_mono : forall n m p q : Z, n <= m -> q < p -> n - p < m - q.
+Theorem sub_le_lt_mono : forall n m p q, n <= m -> q < p -> n - p < m - q.
Proof.
intros n m p q H1 H2.
-apply NZle_lt_trans with (m - p);
-[now apply -> Zsub_le_mono_r | now apply -> Zsub_lt_mono_l].
+apply le_lt_trans with (m - p);
+[now apply -> sub_le_mono_r | now apply -> sub_lt_mono_l].
Qed.
-Theorem Zle_lt_sub_lt : forall n m p q : Z, n <= m -> p - n < q - m -> p < q.
+Theorem le_lt_sub_lt : forall n m p q, n <= m -> p - n < q - m -> p < q.
Proof.
-intros n m p q H1 H2. apply (Zle_lt_add_lt (- m) (- n));
-[now apply -> Zopp_le_mono | now do 2 rewrite Zadd_opp_r].
+intros n m p q H1 H2. apply (le_lt_add_lt (- m) (- n));
+[now apply -> opp_le_mono | now do 2 rewrite add_opp_r].
Qed.
-Theorem Zlt_le_sub_lt : forall n m p q : Z, n < m -> p - n <= q - m -> p < q.
+Theorem lt_le_sub_lt : forall n m p q, n < m -> p - n <= q - m -> p < q.
Proof.
-intros n m p q H1 H2. apply (Zlt_le_add_lt (- m) (- n));
-[now apply -> Zopp_lt_mono | now do 2 rewrite Zadd_opp_r].
+intros n m p q H1 H2. apply (lt_le_add_lt (- m) (- n));
+[now apply -> opp_lt_mono | now do 2 rewrite add_opp_r].
Qed.
-Theorem Zle_le_sub_lt : forall n m p q : Z, n <= m -> p - n <= q - m -> p <= q.
+Theorem le_le_sub_lt : forall n m p q, n <= m -> p - n <= q - m -> p <= q.
Proof.
-intros n m p q H1 H2. apply (Zle_le_add_le (- m) (- n));
-[now apply -> Zopp_le_mono | now do 2 rewrite Zadd_opp_r].
+intros n m p q H1 H2. apply (le_le_add_le (- m) (- n));
+[now apply -> opp_le_mono | now do 2 rewrite add_opp_r].
Qed.
-Theorem Zlt_add_lt_sub_r : forall n m p : Z, n + p < m <-> n < m - p.
+Theorem lt_add_lt_sub_r : forall n m p, n + p < m <-> n < m - p.
Proof.
-intros n m p. stepl (n + p - p < m - p) by symmetry; apply Zsub_lt_mono_r.
-now rewrite Zadd_simpl_r.
+intros n m p. stepl (n + p - p < m - p) by symmetry; apply sub_lt_mono_r.
+now rewrite add_simpl_r.
Qed.
-Theorem Zle_add_le_sub_r : forall n m p : Z, n + p <= m <-> n <= m - p.
+Theorem le_add_le_sub_r : forall n m p, n + p <= m <-> n <= m - p.
Proof.
-intros n m p. stepl (n + p - p <= m - p) by symmetry; apply Zsub_le_mono_r.
-now rewrite Zadd_simpl_r.
+intros n m p. stepl (n + p - p <= m - p) by symmetry; apply sub_le_mono_r.
+now rewrite add_simpl_r.
Qed.
-Theorem Zlt_add_lt_sub_l : forall n m p : Z, n + p < m <-> p < m - n.
+Theorem lt_add_lt_sub_l : forall n m p, n + p < m <-> p < m - n.
Proof.
-intros n m p. rewrite Zadd_comm; apply Zlt_add_lt_sub_r.
+intros n m p. rewrite add_comm; apply lt_add_lt_sub_r.
Qed.
-Theorem Zle_add_le_sub_l : forall n m p : Z, n + p <= m <-> p <= m - n.
+Theorem le_add_le_sub_l : forall n m p, n + p <= m <-> p <= m - n.
Proof.
-intros n m p. rewrite Zadd_comm; apply Zle_add_le_sub_r.
+intros n m p. rewrite add_comm; apply le_add_le_sub_r.
Qed.
-Theorem Zlt_sub_lt_add_r : forall n m p : Z, n - p < m <-> n < m + p.
+Theorem lt_sub_lt_add_r : forall n m p, n - p < m <-> n < m + p.
Proof.
-intros n m p. stepl (n - p + p < m + p) by symmetry; apply Zadd_lt_mono_r.
-now rewrite Zsub_simpl_r.
+intros n m p. stepl (n - p + p < m + p) by symmetry; apply add_lt_mono_r.
+now rewrite sub_simpl_r.
Qed.
-Theorem Zle_sub_le_add_r : forall n m p : Z, n - p <= m <-> n <= m + p.
+Theorem le_sub_le_add_r : forall n m p, n - p <= m <-> n <= m + p.
Proof.
-intros n m p. stepl (n - p + p <= m + p) by symmetry; apply Zadd_le_mono_r.
-now rewrite Zsub_simpl_r.
+intros n m p. stepl (n - p + p <= m + p) by symmetry; apply add_le_mono_r.
+now rewrite sub_simpl_r.
Qed.
-Theorem Zlt_sub_lt_add_l : forall n m p : Z, n - m < p <-> n < m + p.
+Theorem lt_sub_lt_add_l : forall n m p, n - m < p <-> n < m + p.
Proof.
-intros n m p. rewrite Zadd_comm; apply Zlt_sub_lt_add_r.
+intros n m p. rewrite add_comm; apply lt_sub_lt_add_r.
Qed.
-Theorem Zle_sub_le_add_l : forall n m p : Z, n - m <= p <-> n <= m + p.
+Theorem le_sub_le_add_l : forall n m p, n - m <= p <-> n <= m + p.
Proof.
-intros n m p. rewrite Zadd_comm; apply Zle_sub_le_add_r.
+intros n m p. rewrite add_comm; apply le_sub_le_add_r.
Qed.
-Theorem Zlt_sub_lt_add : forall n m p q : Z, n - m < p - q <-> n + q < m + p.
+Theorem lt_sub_lt_add : forall n m p q, n - m < p - q <-> n + q < m + p.
Proof.
-intros n m p q. rewrite Zlt_sub_lt_add_l. rewrite Zadd_sub_assoc.
-now rewrite <- Zlt_add_lt_sub_r.
+intros n m p q. rewrite lt_sub_lt_add_l. rewrite add_sub_assoc.
+now rewrite <- lt_add_lt_sub_r.
Qed.
-Theorem Zle_sub_le_add : forall n m p q : Z, n - m <= p - q <-> n + q <= m + p.
+Theorem le_sub_le_add : forall n m p q, n - m <= p - q <-> n + q <= m + p.
Proof.
-intros n m p q. rewrite Zle_sub_le_add_l. rewrite Zadd_sub_assoc.
-now rewrite <- Zle_add_le_sub_r.
+intros n m p q. rewrite le_sub_le_add_l. rewrite add_sub_assoc.
+now rewrite <- le_add_le_sub_r.
Qed.
-Theorem Zlt_sub_pos : forall n m : Z, 0 < m <-> n - m < n.
+Theorem lt_sub_pos : forall n m, 0 < m <-> n - m < n.
Proof.
-intros n m. stepr (n - m < n - 0) by now rewrite Zsub_0_r. apply Zsub_lt_mono_l.
+intros n m. stepr (n - m < n - 0) by now rewrite sub_0_r. apply sub_lt_mono_l.
Qed.
-Theorem Zle_sub_nonneg : forall n m : Z, 0 <= m <-> n - m <= n.
+Theorem le_sub_nonneg : forall n m, 0 <= m <-> n - m <= n.
Proof.
-intros n m. stepr (n - m <= n - 0) by now rewrite Zsub_0_r. apply Zsub_le_mono_l.
+intros n m. stepr (n - m <= n - 0) by now rewrite sub_0_r. apply sub_le_mono_l.
Qed.
-Theorem Zsub_lt_cases : forall n m p q : Z, n - m < p - q -> n < m \/ q < p.
+Theorem sub_lt_cases : forall n m p q, n - m < p - q -> n < m \/ q < p.
Proof.
-intros n m p q H. rewrite Zlt_sub_lt_add in H. now apply Zadd_lt_cases.
+intros n m p q H. rewrite lt_sub_lt_add in H. now apply add_lt_cases.
Qed.
-Theorem Zsub_le_cases : forall n m p q : Z, n - m <= p - q -> n <= m \/ q <= p.
+Theorem sub_le_cases : forall n m p q, n - m <= p - q -> n <= m \/ q <= p.
Proof.
-intros n m p q H. rewrite Zle_sub_le_add in H. now apply Zadd_le_cases.
+intros n m p q H. rewrite le_sub_le_add in H. now apply add_le_cases.
Qed.
-Theorem Zsub_neg_cases : forall n m : Z, n - m < 0 -> n < 0 \/ 0 < m.
+Theorem sub_neg_cases : forall n m, n - m < 0 -> n < 0 \/ 0 < m.
Proof.
-intros n m H; rewrite <- Zadd_opp_r in H.
-setoid_replace (0 < m) with (- m < 0) using relation iff by (symmetry; apply Zopp_neg_pos).
-now apply Zadd_neg_cases.
+intros n m H; rewrite <- add_opp_r in H.
+setoid_replace (0 < m) with (- m < 0) using relation iff by (symmetry; apply opp_neg_pos).
+now apply add_neg_cases.
Qed.
-Theorem Zsub_pos_cases : forall n m : Z, 0 < n - m -> 0 < n \/ m < 0.
+Theorem sub_pos_cases : forall n m, 0 < n - m -> 0 < n \/ m < 0.
Proof.
-intros n m H; rewrite <- Zadd_opp_r in H.
-setoid_replace (m < 0) with (0 < - m) using relation iff by (symmetry; apply Zopp_pos_neg).
-now apply Zadd_pos_cases.
+intros n m H; rewrite <- add_opp_r in H.
+setoid_replace (m < 0) with (0 < - m) using relation iff by (symmetry; apply opp_pos_neg).
+now apply add_pos_cases.
Qed.
-Theorem Zsub_nonpos_cases : forall n m : Z, n - m <= 0 -> n <= 0 \/ 0 <= m.
+Theorem sub_nonpos_cases : forall n m, n - m <= 0 -> n <= 0 \/ 0 <= m.
Proof.
-intros n m H; rewrite <- Zadd_opp_r in H.
-setoid_replace (0 <= m) with (- m <= 0) using relation iff by (symmetry; apply Zopp_nonpos_nonneg).
-now apply Zadd_nonpos_cases.
+intros n m H; rewrite <- add_opp_r in H.
+setoid_replace (0 <= m) with (- m <= 0) using relation iff by (symmetry; apply opp_nonpos_nonneg).
+now apply add_nonpos_cases.
Qed.
-Theorem Zsub_nonneg_cases : forall n m : Z, 0 <= n - m -> 0 <= n \/ m <= 0.
+Theorem sub_nonneg_cases : forall n m, 0 <= n - m -> 0 <= n \/ m <= 0.
Proof.
-intros n m H; rewrite <- Zadd_opp_r in H.
-setoid_replace (m <= 0) with (0 <= - m) using relation iff by (symmetry; apply Zopp_nonneg_nonpos).
-now apply Zadd_nonneg_cases.
+intros n m H; rewrite <- add_opp_r in H.
+setoid_replace (m <= 0) with (0 <= - m) using relation iff by (symmetry; apply opp_nonneg_nonpos).
+now apply add_nonneg_cases.
Qed.
Section PosNeg.
-Variable P : Z -> Prop.
-Hypothesis P_wd : predicate_wd Zeq P.
-
-Add Morphism P with signature Zeq ==> iff as P_morph. Proof. exact P_wd. Qed.
+Variable P : Z.t -> Prop.
+Hypothesis P_wd : Proper (Z.eq ==> iff) P.
-Theorem Z0_pos_neg :
- P 0 -> (forall n : Z, 0 < n -> P n /\ P (- n)) -> forall n : Z, P n.
+Theorem zero_pos_neg :
+ P 0 -> (forall n, 0 < n -> P n /\ P (- n)) -> forall n, P n.
Proof.
-intros H1 H2 n. destruct (Zlt_trichotomy n 0) as [H3 | [H3 | H3]].
-apply <- Zopp_pos_neg in H3. apply H2 in H3. destruct H3 as [_ H3].
-now rewrite Zopp_involutive in H3.
+intros H1 H2 n. destruct (lt_trichotomy n 0) as [H3 | [H3 | H3]].
+apply <- opp_pos_neg in H3. apply H2 in H3. destruct H3 as [_ H3].
+now rewrite opp_involutive in H3.
now rewrite H3.
apply H2 in H3; now destruct H3.
Qed.
End PosNeg.
-Ltac Z0_pos_neg n := induction_maker n ltac:(apply Z0_pos_neg).
+Ltac zero_pos_neg n := induction_maker n ltac:(apply zero_pos_neg).
End ZAddOrderPropFunct.
diff --git a/theories/Numbers/Integer/Abstract/ZAxioms.v b/theories/Numbers/Integer/Abstract/ZAxioms.v
index c4a4b6b8..9158a214 100644
--- a/theories/Numbers/Integer/Abstract/ZAxioms.v
+++ b/theories/Numbers/Integer/Abstract/ZAxioms.v
@@ -8,58 +8,31 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export NZAxioms.
Set Implicit Arguments.
-Module Type ZAxiomsSig.
-Declare Module Export NZOrdAxiomsMod : NZOrdAxiomsSig.
+Module Type Opp (Import T:Typ).
+ Parameter Inline opp : t -> t.
+End Opp.
-Delimit Scope IntScope with Int.
-Notation Z := NZ.
-Notation Zeq := NZeq.
-Notation Z0 := NZ0.
-Notation Z1 := (NZsucc NZ0).
-Notation S := NZsucc.
-Notation P := NZpred.
-Notation Zadd := NZadd.
-Notation Zmul := NZmul.
-Notation Zsub := NZsub.
-Notation Zlt := NZlt.
-Notation Zle := NZle.
-Notation Zmin := NZmin.
-Notation Zmax := NZmax.
-Notation "x == y" := (NZeq x y) (at level 70) : IntScope.
-Notation "x ~= y" := (~ NZeq x y) (at level 70) : IntScope.
-Notation "0" := NZ0 : IntScope.
-Notation "1" := (NZsucc NZ0) : IntScope.
-Notation "x + y" := (NZadd x y) : IntScope.
-Notation "x - y" := (NZsub x y) : IntScope.
-Notation "x * y" := (NZmul x y) : IntScope.
-Notation "x < y" := (NZlt x y) : IntScope.
-Notation "x <= y" := (NZle x y) : IntScope.
-Notation "x > y" := (NZlt y x) (only parsing) : IntScope.
-Notation "x >= y" := (NZle y x) (only parsing) : IntScope.
+Module Type OppNotation (T:Typ)(Import O : Opp T).
+ Notation "- x" := (opp x) (at level 35, right associativity).
+End OppNotation.
-Parameter Zopp : Z -> Z.
+Module Type Opp' (T:Typ) := Opp T <+ OppNotation T.
-(*Notation "- 1" := (Zopp 1) : IntScope.
-Check (-1).*)
+(** We obtain integers by postulating that every number has a predecessor. *)
-Add Morphism Zopp with signature Zeq ==> Zeq as Zopp_wd.
+Module Type IsOpp (Import Z : NZAxiomsSig')(Import O : Opp' Z).
+ Declare Instance opp_wd : Proper (eq==>eq) opp.
+ Axiom succ_pred : forall n, S (P n) == n.
+ Axiom opp_0 : - 0 == 0.
+ Axiom opp_succ : forall n, - (S n) == P (- n).
+End IsOpp.
-Notation "- x" := (Zopp x) (at level 35, right associativity) : IntScope.
-Notation "- 1" := (Zopp (NZsucc NZ0)) : IntScope.
-
-Open Local Scope IntScope.
-
-(* Integers are obtained by postulating that every number has a predecessor *)
-Axiom Zsucc_pred : forall n : Z, S (P n) == n.
-
-Axiom Zopp_0 : - 0 == 0.
-Axiom Zopp_succ : forall n : Z, - (S n) == P (- n).
-
-End ZAxiomsSig.
+Module Type ZAxiomsSig := NZOrdAxiomsSig <+ Opp <+ IsOpp.
+Module Type ZAxiomsSig' := NZOrdAxiomsSig' <+ Opp' <+ IsOpp.
diff --git a/theories/Numbers/Integer/Abstract/ZBase.v b/theories/Numbers/Integer/Abstract/ZBase.v
index 0f71f2cc..44bb02ec 100644
--- a/theories/Numbers/Integer/Abstract/ZBase.v
+++ b/theories/Numbers/Integer/Abstract/ZBase.v
@@ -8,78 +8,25 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZBase.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
Require Export Decidable.
Require Export ZAxioms.
-Require Import NZMulOrder.
+Require Import NZProperties.
-Module ZBasePropFunct (Import ZAxiomsMod : ZAxiomsSig).
-
-(* Note: writing "Export" instead of "Import" on the previous line leads to
-some warnings about hiding repeated declarations and results in the loss of
-notations in Zadd and later *)
-
-Open Local Scope IntScope.
-
-Module Export NZMulOrderMod := NZMulOrderPropFunct NZOrdAxiomsMod.
-
-Theorem Zsucc_wd : forall n1 n2 : Z, n1 == n2 -> S n1 == S n2.
-Proof NZsucc_wd.
-
-Theorem Zpred_wd : forall n1 n2 : Z, n1 == n2 -> P n1 == P n2.
-Proof NZpred_wd.
-
-Theorem Zpred_succ : forall n : Z, P (S n) == n.
-Proof NZpred_succ.
-
-Theorem Zeq_refl : forall n : Z, n == n.
-Proof (proj1 NZeq_equiv).
-
-Theorem Zeq_sym : forall n m : Z, n == m -> m == n.
-Proof (proj2 (proj2 NZeq_equiv)).
-
-Theorem Zeq_trans : forall n m p : Z, n == m -> m == p -> n == p.
-Proof (proj1 (proj2 NZeq_equiv)).
-
-Theorem Zneq_sym : forall n m : Z, n ~= m -> m ~= n.
-Proof NZneq_sym.
-
-Theorem Zsucc_inj : forall n1 n2 : Z, S n1 == S n2 -> n1 == n2.
-Proof NZsucc_inj.
-
-Theorem Zsucc_inj_wd : forall n1 n2 : Z, S n1 == S n2 <-> n1 == n2.
-Proof NZsucc_inj_wd.
-
-Theorem Zsucc_inj_wd_neg : forall n m : Z, S n ~= S m <-> n ~= m.
-Proof NZsucc_inj_wd_neg.
-
-(* Decidability and stability of equality was proved only in NZOrder, but
-since it does not mention order, we'll put it here *)
-
-Theorem Zeq_dec : forall n m : Z, decidable (n == m).
-Proof NZeq_dec.
-
-Theorem Zeq_dne : forall n m : Z, ~ ~ n == m <-> n == m.
-Proof NZeq_dne.
-
-Theorem Zcentral_induction :
-forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z, A z ->
- (forall n : Z, A n <-> A (S n)) ->
- forall n : Z, A n.
-Proof NZcentral_induction.
+Module ZBasePropFunct (Import Z : ZAxiomsSig').
+Include NZPropFunct Z.
(* Theorems that are true for integers but not for natural numbers *)
-Theorem Zpred_inj : forall n m : Z, P n == P m -> n == m.
+Theorem pred_inj : forall n m, P n == P m -> n == m.
Proof.
-intros n m H. apply NZsucc_wd in H. now do 2 rewrite Zsucc_pred in H.
+intros n m H. apply succ_wd in H. now do 2 rewrite succ_pred in H.
Qed.
-Theorem Zpred_inj_wd : forall n1 n2 : Z, P n1 == P n2 <-> n1 == n2.
+Theorem pred_inj_wd : forall n1 n2, P n1 == P n2 <-> n1 == n2.
Proof.
-intros n1 n2; split; [apply Zpred_inj | apply NZpred_wd].
+intros n1 n2; split; [apply pred_inj | apply pred_wd].
Qed.
End ZBasePropFunct.
diff --git a/theories/Numbers/Integer/Abstract/ZDivEucl.v b/theories/Numbers/Integer/Abstract/ZDivEucl.v
new file mode 100644
index 00000000..bcd16fec
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZDivEucl.v
@@ -0,0 +1,605 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * Euclidean Division for integers, Euclid convention
+
+ We use here the "usual" formulation of the Euclid Theorem
+ [forall a b, b<>0 -> exists b q, a = b*q+r /\ 0 < r < |b| ]
+
+ The outcome of the modulo function is hence always positive.
+ This corresponds to convention "E" in the following paper:
+
+ R. Boute, "The Euclidean definition of the functions div and mod",
+ ACM Transactions on Programming Languages and Systems,
+ Vol. 14, No.2, pp. 127-144, April 1992.
+
+ See files [ZDivTrunc] and [ZDivFloor] for others conventions.
+*)
+
+Require Import ZAxioms ZProperties NZDiv.
+
+Module Type ZDivSpecific (Import Z : ZAxiomsExtSig')(Import DM : DivMod' Z).
+ Axiom mod_always_pos : forall a b, 0 <= a mod b < abs b.
+End ZDivSpecific.
+
+Module Type ZDiv (Z:ZAxiomsExtSig)
+ := DivMod Z <+ NZDivCommon Z <+ ZDivSpecific Z.
+
+Module Type ZDivSig := ZAxiomsExtSig <+ ZDiv.
+Module Type ZDivSig' := ZAxiomsExtSig' <+ ZDiv <+ DivModNotation.
+
+Module ZDivPropFunct (Import Z : ZDivSig')(Import ZP : ZPropSig Z).
+
+(** We benefit from what already exists for NZ *)
+
+ Module ZD <: NZDiv Z.
+ Definition div := div.
+ Definition modulo := modulo.
+ Definition div_wd := div_wd.
+ Definition mod_wd := mod_wd.
+ Definition div_mod := div_mod.
+ Lemma mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+ Proof.
+ intros. rewrite <- (abs_eq b) at 3 by now apply lt_le_incl.
+ apply mod_always_pos.
+ Qed.
+ End ZD.
+ Module Import NZDivP := NZDivPropFunct Z ZP ZD.
+
+(** Another formulation of the main equation *)
+
+Lemma mod_eq :
+ forall a b, b~=0 -> a mod b == a - b*(a/b).
+Proof.
+intros.
+rewrite <- add_move_l.
+symmetry. now apply div_mod.
+Qed.
+
+Ltac pos_or_neg a :=
+ let LT := fresh "LT" in
+ let LE := fresh "LE" in
+ destruct (le_gt_cases 0 a) as [LE|LT]; [|rewrite <- opp_pos_neg in LT].
+
+(** Uniqueness theorems *)
+
+Theorem div_mod_unique : forall b q1 q2 r1 r2 : t,
+ 0<=r1<abs b -> 0<=r2<abs b ->
+ b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.
+Proof.
+intros b q1 q2 r1 r2 Hr1 Hr2 EQ.
+pos_or_neg b.
+rewrite abs_eq in * by trivial.
+apply div_mod_unique with b; trivial.
+rewrite abs_neq' in * by auto using lt_le_incl.
+rewrite eq_sym_iff. apply div_mod_unique with (-b); trivial.
+rewrite 2 mul_opp_l.
+rewrite add_move_l, sub_opp_r.
+rewrite <-add_assoc.
+symmetry. rewrite add_move_l, sub_opp_r.
+now rewrite (add_comm r2), (add_comm r1).
+Qed.
+
+Theorem div_unique:
+ forall a b q r, 0<=r<abs b -> a == b*q + r -> q == a/b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0).
+ pos_or_neg b.
+ rewrite abs_eq in Hr; intuition; order.
+ rewrite <- opp_0, eq_opp_r. rewrite abs_neq' in Hr; intuition; order.
+destruct (div_mod_unique b q (a/b) r (a mod b)); trivial.
+now apply mod_always_pos.
+now rewrite <- div_mod.
+Qed.
+
+Theorem mod_unique:
+ forall a b q r, 0<=r<abs b -> a == b*q + r -> r == a mod b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0).
+ pos_or_neg b.
+ rewrite abs_eq in Hr; intuition; order.
+ rewrite <- opp_0, eq_opp_r. rewrite abs_neq' in Hr; intuition; order.
+destruct (div_mod_unique b q (a/b) r (a mod b)); trivial.
+now apply mod_always_pos.
+now rewrite <- div_mod.
+Qed.
+
+(** Sign rules *)
+
+Lemma div_opp_r : forall a b, b~=0 -> a/(-b) == -(a/b).
+Proof.
+intros. symmetry.
+apply div_unique with (a mod b).
+rewrite abs_opp; apply mod_always_pos.
+rewrite mul_opp_opp; now apply div_mod.
+Qed.
+
+Lemma mod_opp_r : forall a b, b~=0 -> a mod (-b) == a mod b.
+Proof.
+intros. symmetry.
+apply mod_unique with (-(a/b)).
+rewrite abs_opp; apply mod_always_pos.
+rewrite mul_opp_opp; now apply div_mod.
+Qed.
+
+Lemma div_opp_l_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a)/b == -(a/b).
+Proof.
+intros a b Hb Hab. symmetry.
+apply div_unique with (-(a mod b)).
+rewrite Hab, opp_0. split; [order|].
+pos_or_neg b; [rewrite abs_eq | rewrite abs_neq']; order.
+now rewrite mul_opp_r, <-opp_add_distr, <-div_mod.
+Qed.
+
+Lemma div_opp_l_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a)/b == -(a/b)-sgn b.
+Proof.
+intros a b Hb Hab. symmetry.
+apply div_unique with (abs b -(a mod b)).
+rewrite lt_sub_lt_add_l.
+rewrite <- le_add_le_sub_l. nzsimpl.
+rewrite <- (add_0_l (abs b)) at 2.
+rewrite <- add_lt_mono_r.
+destruct (mod_always_pos a b); intuition order.
+rewrite <- 2 add_opp_r, mul_add_distr_l, 2 mul_opp_r.
+rewrite sgn_abs.
+rewrite add_shuffle2, add_opp_diag_l; nzsimpl.
+rewrite <-opp_add_distr, <-div_mod; order.
+Qed.
+
+Lemma mod_opp_l_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a) mod b == 0.
+Proof.
+intros a b Hb Hab. symmetry.
+apply mod_unique with (-(a/b)).
+split; [order|now rewrite abs_pos].
+now rewrite <-opp_0, <-Hab, mul_opp_r, <-opp_add_distr, <-div_mod.
+Qed.
+
+Lemma mod_opp_l_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a) mod b == abs b - (a mod b).
+Proof.
+intros a b Hb Hab. symmetry.
+apply mod_unique with (-(a/b)-sgn b).
+rewrite lt_sub_lt_add_l.
+rewrite <- le_add_le_sub_l. nzsimpl.
+rewrite <- (add_0_l (abs b)) at 2.
+rewrite <- add_lt_mono_r.
+destruct (mod_always_pos a b); intuition order.
+rewrite <- 2 add_opp_r, mul_add_distr_l, 2 mul_opp_r.
+rewrite sgn_abs.
+rewrite add_shuffle2, add_opp_diag_l; nzsimpl.
+rewrite <-opp_add_distr, <-div_mod; order.
+Qed.
+
+Lemma div_opp_opp_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a)/(-b) == a/b.
+Proof.
+intros. now rewrite div_opp_r, div_opp_l_z, opp_involutive.
+Qed.
+
+Lemma div_opp_opp_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a)/(-b) == a/b + sgn(b).
+Proof.
+intros. rewrite div_opp_r, div_opp_l_nz by trivial.
+now rewrite opp_sub_distr, opp_involutive.
+Qed.
+
+Lemma mod_opp_opp_z : forall a b, b~=0 -> a mod b == 0 ->
+ (-a) mod (-b) == 0.
+Proof.
+intros. now rewrite mod_opp_r, mod_opp_l_z.
+Qed.
+
+Lemma mod_opp_opp_nz : forall a b, b~=0 -> a mod b ~= 0 ->
+ (-a) mod (-b) == abs b - a mod b.
+Proof.
+intros. now rewrite mod_opp_r, mod_opp_l_nz.
+Qed.
+
+(** A division by itself returns 1 *)
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof.
+intros. symmetry. apply div_unique with 0.
+split; [order|now rewrite abs_pos].
+now nzsimpl.
+Qed.
+
+Lemma mod_same : forall a, a~=0 -> a mod a == 0.
+Proof.
+intros.
+rewrite mod_eq, div_same by trivial. nzsimpl. apply sub_diag.
+Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, 0<=a<b -> a/b == 0.
+Proof. exact div_small. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.
+Proof. exact mod_small. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, a~=0 -> 0/a == 0.
+Proof.
+intros. pos_or_neg a. apply div_0_l; order.
+apply opp_inj. rewrite <- div_opp_r, opp_0 by trivial. now apply div_0_l.
+Qed.
+
+Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.
+Proof.
+intros; rewrite mod_eq, div_0_l; now nzsimpl.
+Qed.
+
+Lemma div_1_r: forall a, a/1 == a.
+Proof.
+intros. symmetry. apply div_unique with 0.
+assert (H:=lt_0_1); rewrite abs_pos; intuition; order.
+now nzsimpl.
+Qed.
+
+Lemma mod_1_r: forall a, a mod 1 == 0.
+Proof.
+intros. rewrite mod_eq, div_1_r; nzsimpl; auto using sub_diag.
+apply neq_sym, lt_neq; apply lt_0_1.
+Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof. exact div_1_l. Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof. exact mod_1_l. Qed.
+
+Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.
+Proof.
+intros. symmetry. apply div_unique with 0.
+split; [order|now rewrite abs_pos].
+nzsimpl; apply mul_comm.
+Qed.
+
+Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.
+Proof.
+intros. rewrite mod_eq, div_mul by trivial. rewrite mul_comm; apply sub_diag.
+Qed.
+
+(** * Order results about mod and div *)
+
+(** A modulo cannot grow beyond its starting point. *)
+
+Theorem mod_le: forall a b, 0<=a -> b~=0 -> a mod b <= a.
+Proof.
+intros. pos_or_neg b. apply mod_le; order.
+rewrite <- mod_opp_r by trivial. apply mod_le; order.
+Qed.
+
+Theorem div_pos : forall a b, 0<=a -> 0<b -> 0<= a/b.
+Proof. exact div_pos. Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof. exact div_str_pos. Qed.
+
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> 0<=a<abs b).
+Proof.
+intros a b Hb.
+split.
+intros EQ.
+rewrite (div_mod a b Hb), EQ; nzsimpl.
+apply mod_always_pos.
+intros. pos_or_neg b.
+apply div_small.
+now rewrite <- (abs_eq b).
+apply opp_inj; rewrite opp_0, <- div_opp_r by trivial.
+apply div_small.
+rewrite <- (abs_neq' b) by order. trivial.
+Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> 0<=a<abs b).
+Proof.
+intros.
+rewrite <- div_small_iff, mod_eq by trivial.
+rewrite sub_move_r, <- (add_0_r a) at 1. rewrite add_cancel_l.
+rewrite eq_sym_iff, eq_mul_0. tauto.
+Qed.
+
+(** As soon as the divisor is strictly greater than 1,
+ the division is strictly decreasing. *)
+
+Lemma div_lt : forall a b, 0<a -> 1<b -> a/b < a.
+Proof. exact div_lt. Qed.
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono : forall a b c, 0<c -> a<=b -> a/c <= b/c.
+Proof.
+intros a b c Hc Hab.
+rewrite lt_eq_cases in Hab. destruct Hab as [LT|EQ];
+ [|rewrite EQ; order].
+rewrite <- lt_succ_r.
+rewrite (mul_lt_mono_pos_l c) by order.
+nzsimpl.
+rewrite (add_lt_mono_r _ _ (a mod c)).
+rewrite <- div_mod by order.
+apply lt_le_trans with b; trivial.
+rewrite (div_mod b c) at 1 by order.
+rewrite <- add_assoc, <- add_le_mono_l.
+apply le_trans with (c+0).
+nzsimpl; destruct (mod_always_pos b c); try order.
+rewrite abs_eq in *; order.
+rewrite <- add_le_mono_l. destruct (mod_always_pos a c); order.
+Qed.
+
+(** In this convention, [div] performs Rounding-Toward-Bottom
+ when divisor is positive, and Rounding-Toward-Top otherwise.
+
+ Since we cannot speak of rational values here, we express this
+ fact by multiplying back by [b], and this leads to a nice
+ unique statement.
+*)
+
+Lemma mul_div_le : forall a b, b~=0 -> b*(a/b) <= a.
+Proof.
+intros.
+rewrite (div_mod a b) at 2; trivial.
+rewrite <- (add_0_r (b*(a/b))) at 1.
+rewrite <- add_le_mono_l.
+now destruct (mod_always_pos a b).
+Qed.
+
+(** Giving a reversed bound is slightly more complex *)
+
+Lemma mul_succ_div_gt: forall a b, 0<b -> a < b*(S (a/b)).
+Proof.
+intros.
+nzsimpl.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- add_lt_mono_l.
+destruct (mod_always_pos a b).
+rewrite abs_eq in *; order.
+Qed.
+
+Lemma mul_pred_div_gt: forall a b, b<0 -> a < b*(P (a/b)).
+Proof.
+intros a b Hb.
+rewrite mul_pred_r, <- add_opp_r.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- add_lt_mono_l.
+destruct (mod_always_pos a b).
+rewrite <- opp_pos_neg in Hb. rewrite abs_neq' in *; order.
+Qed.
+
+(** NB: The three previous properties could be used as
+ specifications for [div]. *)
+
+(** Inequality [mul_div_le] is exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).
+Proof.
+intros.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- (add_0_r (b*(a/b))) at 2.
+apply add_cancel_l.
+Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, 0<b -> a < b*q -> a/b < q.
+Proof.
+intros.
+rewrite (mul_lt_mono_pos_l b) by trivial.
+apply le_lt_trans with a; trivial.
+apply mul_div_le; order.
+Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, 0<b -> a <= b*q -> a/b <= q.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; trivial. now rewrite mul_comm.
+Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, 0<b -> b*q <= a -> q <= a/b.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; trivial. now rewrite mul_comm.
+Qed.
+
+(** A division respects opposite monotonicity for the divisor *)
+
+Lemma div_le_compat_l: forall p q r, 0<=p -> 0<q<=r -> p/r <= p/q.
+Proof. exact div_le_compat_l. Qed.
+
+(** * Relations between usual operations and mod and div *)
+
+Lemma mod_add : forall a b c, c~=0 ->
+ (a + b * c) mod c == a mod c.
+Proof.
+intros.
+symmetry.
+apply mod_unique with (a/c+b); trivial.
+now apply mod_always_pos.
+rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+now rewrite mul_comm.
+Qed.
+
+Lemma div_add : forall a b c, c~=0 ->
+ (a + b * c) / c == a / c + b.
+Proof.
+intros.
+apply (mul_cancel_l _ _ c); try order.
+apply (add_cancel_r _ _ ((a+b*c) mod c)).
+rewrite <- div_mod, mod_add by order.
+rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+now rewrite mul_comm.
+Qed.
+
+Lemma div_add_l: forall a b c, b~=0 ->
+ (a * b + c) / b == a + c / b.
+Proof.
+ intros a b c. rewrite (add_comm _ c), (add_comm a).
+ now apply div_add.
+Qed.
+
+(** Cancellations. *)
+
+(** With the current convention, the following isn't always true
+ when [c<0]: [-3*-1 / -2*-1 = 3/2 = 1] while [-3/-2 = 2] *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> 0<c ->
+ (a*c)/(b*c) == a/b.
+Proof.
+intros.
+symmetry.
+apply div_unique with ((a mod b)*c).
+(* ineqs *)
+rewrite abs_mul, (abs_eq c) by order.
+rewrite <-(mul_0_l c), <-mul_lt_mono_pos_r, <-mul_le_mono_pos_r by trivial.
+apply mod_always_pos.
+(* equation *)
+rewrite (div_mod a b) at 1 by order.
+rewrite mul_add_distr_r.
+rewrite add_cancel_r.
+rewrite <- 2 mul_assoc. now rewrite (mul_comm c).
+Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> 0<c ->
+ (c*a)/(c*b) == a/b.
+Proof.
+intros. rewrite !(mul_comm c); now apply div_mul_cancel_r.
+Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> 0<c ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof.
+intros.
+rewrite <- (add_cancel_l _ _ ((c*b)* ((c*a)/(c*b)))).
+rewrite <- div_mod.
+rewrite div_mul_cancel_l by trivial.
+rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+apply div_mod; order.
+rewrite <- neq_mul_0; intuition; order.
+Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> 0<c ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof.
+ intros. rewrite !(mul_comm _ c); now rewrite mul_mod_distr_l.
+Qed.
+
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof.
+intros. rewrite mod_small_iff by trivial.
+now apply mod_always_pos.
+Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof.
+ intros a b n Hn. symmetry.
+ rewrite (div_mod a n) at 1 by order.
+ rewrite add_comm, (mul_comm n), (mul_comm _ b).
+ rewrite mul_add_distr_l, mul_assoc.
+ rewrite mod_add by trivial.
+ now rewrite mul_comm.
+Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof.
+ intros. rewrite !(mul_comm a). now apply mul_mod_idemp_l.
+Qed.
+
+Theorem mul_mod: forall a b n, n~=0 ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof.
+ intros. now rewrite mul_mod_idemp_l, mul_mod_idemp_r.
+Qed.
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof.
+ intros a b n Hn. symmetry.
+ rewrite (div_mod a n) at 1 by order.
+ rewrite <- add_assoc, add_comm, mul_comm.
+ now rewrite mod_add.
+Qed.
+
+Lemma add_mod_idemp_r : forall a b n, n~=0 ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof.
+ intros. rewrite !(add_comm a). now apply add_mod_idemp_l.
+Qed.
+
+Theorem add_mod: forall a b n, n~=0 ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof.
+ intros. now rewrite add_mod_idemp_l, add_mod_idemp_r.
+Qed.
+
+(** With the current convention, the following result isn't always
+ true for negative divisors. For instance
+ [ 3/(-2)/(-2) = 1 <> 0 = 3 / (-2*-2) ]. *)
+
+Lemma div_div : forall a b c, 0<b -> 0<c ->
+ (a/b)/c == a/(b*c).
+Proof.
+ intros a b c Hb Hc.
+ apply div_unique with (b*((a/b) mod c) + a mod b).
+ (* begin 0<= ... <abs(b*c) *)
+ rewrite abs_mul.
+ destruct (mod_always_pos (a/b) c), (mod_always_pos a b).
+ split.
+ apply add_nonneg_nonneg; trivial.
+ apply mul_nonneg_nonneg; order.
+ apply lt_le_trans with (b*((a/b) mod c) + abs b).
+ now rewrite <- add_lt_mono_l.
+ rewrite (abs_eq b) by order.
+ now rewrite <- mul_succ_r, <- mul_le_mono_pos_l, le_succ_l.
+ (* end 0<= ... < abs(b*c) *)
+ rewrite (div_mod a b) at 1 by order.
+ rewrite add_assoc, add_cancel_r.
+ rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+ apply div_mod; order.
+Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, 0<=a -> 0<b -> 0<=c -> c*(a/b) <= (c*a)/b.
+Proof. exact div_mul_le. Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, b~=0 ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof.
+intros a b Hb. split.
+intros Hab. exists (a/b). rewrite (div_mod a b Hb) at 1.
+ rewrite Hab; now nzsimpl.
+intros (c,Hc).
+rewrite Hc, mul_comm.
+now apply mod_mul.
+Qed.
+
+
+End ZDivPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZDivFloor.v b/theories/Numbers/Integer/Abstract/ZDivFloor.v
new file mode 100644
index 00000000..1e7624ba
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZDivFloor.v
@@ -0,0 +1,632 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * Euclidean Division for integers (Floor convention)
+
+ We use here the convention known as Floor, or Round-Toward-Bottom,
+ where [a/b] is the closest integer below the exact fraction.
+ It can be summarized by:
+
+ [a = bq+r /\ 0 <= |r| < |b| /\ Sign(r) = Sign(b)]
+
+ This is the convention followed historically by [Zdiv] in Coq, and
+ corresponds to convention "F" in the following paper:
+
+ R. Boute, "The Euclidean definition of the functions div and mod",
+ ACM Transactions on Programming Languages and Systems,
+ Vol. 14, No.2, pp. 127-144, April 1992.
+
+ See files [ZDivTrunc] and [ZDivEucl] for others conventions.
+*)
+
+Require Import ZAxioms ZProperties NZDiv.
+
+Module Type ZDivSpecific (Import Z:ZAxiomsSig')(Import DM : DivMod' Z).
+ Axiom mod_pos_bound : forall a b, 0 < b -> 0 <= a mod b < b.
+ Axiom mod_neg_bound : forall a b, b < 0 -> b < a mod b <= 0.
+End ZDivSpecific.
+
+Module Type ZDiv (Z:ZAxiomsSig)
+ := DivMod Z <+ NZDivCommon Z <+ ZDivSpecific Z.
+
+Module Type ZDivSig := ZAxiomsExtSig <+ ZDiv.
+Module Type ZDivSig' := ZAxiomsExtSig' <+ ZDiv <+ DivModNotation.
+
+Module ZDivPropFunct (Import Z : ZDivSig')(Import ZP : ZPropSig Z).
+
+(** We benefit from what already exists for NZ *)
+
+ Module ZD <: NZDiv Z.
+ Definition div := div.
+ Definition modulo := modulo.
+ Definition div_wd := div_wd.
+ Definition mod_wd := mod_wd.
+ Definition div_mod := div_mod.
+ Lemma mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+ Proof. intros. now apply mod_pos_bound. Qed.
+ End ZD.
+ Module Import NZDivP := NZDivPropFunct Z ZP ZD.
+
+(** Another formulation of the main equation *)
+
+Lemma mod_eq :
+ forall a b, b~=0 -> a mod b == a - b*(a/b).
+Proof.
+intros.
+rewrite <- add_move_l.
+symmetry. now apply div_mod.
+Qed.
+
+(** Uniqueness theorems *)
+
+Theorem div_mod_unique : forall b q1 q2 r1 r2 : t,
+ (0<=r1<b \/ b<r1<=0) -> (0<=r2<b \/ b<r2<=0) ->
+ b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.
+Proof.
+intros b q1 q2 r1 r2 Hr1 Hr2 EQ.
+destruct Hr1; destruct Hr2; try (intuition; order).
+apply div_mod_unique with b; trivial.
+rewrite <- (opp_inj_wd r1 r2).
+apply div_mod_unique with (-b); trivial.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; tauto.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; tauto.
+now rewrite 2 mul_opp_l, <- 2 opp_add_distr, opp_inj_wd.
+Qed.
+
+Theorem div_unique:
+ forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> q == a/b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0) by (destruct Hr; intuition; order).
+destruct (div_mod_unique b q (a/b) r (a mod b)); trivial.
+destruct Hr; [left; apply mod_pos_bound|right; apply mod_neg_bound];
+ intuition order.
+now rewrite <- div_mod.
+Qed.
+
+Theorem div_unique_pos:
+ forall a b q r, 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto. Qed.
+
+Theorem div_unique_neg:
+ forall a b q r, 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto. Qed.
+
+Theorem mod_unique:
+ forall a b q r, (0<=r<b \/ b<r<=0) -> a == b*q + r -> r == a mod b.
+Proof.
+intros a b q r Hr EQ.
+assert (Hb : b~=0) by (destruct Hr; intuition; order).
+destruct (div_mod_unique b q (a/b) r (a mod b)); trivial.
+destruct Hr; [left; apply mod_pos_bound|right; apply mod_neg_bound];
+ intuition order.
+now rewrite <- div_mod.
+Qed.
+
+Theorem mod_unique_pos:
+ forall a b q r, 0<=r<b -> a == b*q + r -> r == a mod b.
+Proof. intros; apply mod_unique with q; auto. Qed.
+
+Theorem mod_unique_neg:
+ forall a b q r, b<r<=0 -> a == b*q + r -> r == a mod b.
+Proof. intros; apply mod_unique with q; auto. Qed.
+
+(** Sign rules *)
+
+Ltac pos_or_neg a :=
+ let LT := fresh "LT" in
+ let LE := fresh "LE" in
+ destruct (le_gt_cases 0 a) as [LE|LT]; [|rewrite <- opp_pos_neg in LT].
+
+Fact mod_bound_or : forall a b, b~=0 -> 0<=a mod b<b \/ b<a mod b<=0.
+Proof.
+intros.
+destruct (lt_ge_cases 0 b); [left|right].
+ apply mod_pos_bound; trivial. apply mod_neg_bound; order.
+Qed.
+
+Fact opp_mod_bound_or : forall a b, b~=0 ->
+ 0 <= -(a mod b) < -b \/ -b < -(a mod b) <= 0.
+Proof.
+intros.
+destruct (lt_ge_cases 0 b); [right|left].
+rewrite <- opp_lt_mono, opp_nonpos_nonneg.
+ destruct (mod_pos_bound a b); intuition; order.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos.
+ destruct (mod_neg_bound a b); intuition; order.
+Qed.
+
+Lemma div_opp_opp : forall a b, b~=0 -> -a/-b == a/b.
+Proof.
+intros. symmetry. apply div_unique with (- (a mod b)).
+now apply opp_mod_bound_or.
+rewrite mul_opp_l, <- opp_add_distr, <- div_mod; order.
+Qed.
+
+Lemma mod_opp_opp : forall a b, b~=0 -> (-a) mod (-b) == - (a mod b).
+Proof.
+intros. symmetry. apply mod_unique with (a/b).
+now apply opp_mod_bound_or.
+rewrite mul_opp_l, <- opp_add_distr, <- div_mod; order.
+Qed.
+
+(** With the current conventions, the other sign rules are rather complex. *)
+
+Lemma div_opp_l_z :
+ forall a b, b~=0 -> a mod b == 0 -> (-a)/b == -(a/b).
+Proof.
+intros a b Hb H. symmetry. apply div_unique with 0.
+destruct (lt_ge_cases 0 b); [left|right]; intuition; order.
+rewrite <- opp_0, <- H.
+rewrite mul_opp_r, <- opp_add_distr, <- div_mod; order.
+Qed.
+
+Lemma div_opp_l_nz :
+ forall a b, b~=0 -> a mod b ~= 0 -> (-a)/b == -(a/b)-1.
+Proof.
+intros a b Hb H. symmetry. apply div_unique with (b - a mod b).
+destruct (lt_ge_cases 0 b); [left|right].
+rewrite le_0_sub. rewrite <- (sub_0_r b) at 5. rewrite <- sub_lt_mono_l.
+destruct (mod_pos_bound a b); intuition; order.
+rewrite le_sub_0. rewrite <- (sub_0_r b) at 1. rewrite <- sub_lt_mono_l.
+destruct (mod_neg_bound a b); intuition; order.
+rewrite <- (add_opp_r b), mul_sub_distr_l, mul_1_r, sub_add_simpl_r_l.
+rewrite mul_opp_r, <-opp_add_distr, <-div_mod; order.
+Qed.
+
+Lemma mod_opp_l_z :
+ forall a b, b~=0 -> a mod b == 0 -> (-a) mod b == 0.
+Proof.
+intros a b Hb H. symmetry. apply mod_unique with (-(a/b)).
+destruct (lt_ge_cases 0 b); [left|right]; intuition; order.
+rewrite <- opp_0, <- H.
+rewrite mul_opp_r, <- opp_add_distr, <- div_mod; order.
+Qed.
+
+Lemma mod_opp_l_nz :
+ forall a b, b~=0 -> a mod b ~= 0 -> (-a) mod b == b - a mod b.
+Proof.
+intros a b Hb H. symmetry. apply mod_unique with (-(a/b)-1).
+destruct (lt_ge_cases 0 b); [left|right].
+rewrite le_0_sub. rewrite <- (sub_0_r b) at 5. rewrite <- sub_lt_mono_l.
+destruct (mod_pos_bound a b); intuition; order.
+rewrite le_sub_0. rewrite <- (sub_0_r b) at 1. rewrite <- sub_lt_mono_l.
+destruct (mod_neg_bound a b); intuition; order.
+rewrite <- (add_opp_r b), mul_sub_distr_l, mul_1_r, sub_add_simpl_r_l.
+rewrite mul_opp_r, <-opp_add_distr, <-div_mod; order.
+Qed.
+
+Lemma div_opp_r_z :
+ forall a b, b~=0 -> a mod b == 0 -> a/(-b) == -(a/b).
+Proof.
+intros. rewrite <- (opp_involutive a) at 1.
+rewrite div_opp_opp; auto using div_opp_l_z.
+Qed.
+
+Lemma div_opp_r_nz :
+ forall a b, b~=0 -> a mod b ~= 0 -> a/(-b) == -(a/b)-1.
+Proof.
+intros. rewrite <- (opp_involutive a) at 1.
+rewrite div_opp_opp; auto using div_opp_l_nz.
+Qed.
+
+Lemma mod_opp_r_z :
+ forall a b, b~=0 -> a mod b == 0 -> a mod (-b) == 0.
+Proof.
+intros. rewrite <- (opp_involutive a) at 1.
+now rewrite mod_opp_opp, mod_opp_l_z, opp_0.
+Qed.
+
+Lemma mod_opp_r_nz :
+ forall a b, b~=0 -> a mod b ~= 0 -> a mod (-b) == (a mod b) - b.
+Proof.
+intros. rewrite <- (opp_involutive a) at 1.
+rewrite mod_opp_opp, mod_opp_l_nz by trivial.
+now rewrite opp_sub_distr, add_comm, add_opp_r.
+Qed.
+
+(** The sign of [a mod b] is the one of [b] *)
+
+(* TODO: a proper sgn function and theory *)
+
+Lemma mod_sign : forall a b, b~=0 -> (0 <= (a mod b) * b).
+Proof.
+intros. destruct (lt_ge_cases 0 b).
+apply mul_nonneg_nonneg; destruct (mod_pos_bound a b); order.
+apply mul_nonpos_nonpos; destruct (mod_neg_bound a b); order.
+Qed.
+
+(** A division by itself returns 1 *)
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof.
+intros. pos_or_neg a. apply div_same; order.
+rewrite <- div_opp_opp by trivial. now apply div_same.
+Qed.
+
+Lemma mod_same : forall a, a~=0 -> a mod a == 0.
+Proof.
+intros. rewrite mod_eq, div_same by trivial. nzsimpl. apply sub_diag.
+Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, 0<=a<b -> a/b == 0.
+Proof. exact div_small. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.
+Proof. exact mod_small. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, a~=0 -> 0/a == 0.
+Proof.
+intros. pos_or_neg a. apply div_0_l; order.
+rewrite <- div_opp_opp, opp_0 by trivial. now apply div_0_l.
+Qed.
+
+Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.
+Proof.
+intros; rewrite mod_eq, div_0_l; now nzsimpl.
+Qed.
+
+Lemma div_1_r: forall a, a/1 == a.
+Proof.
+intros. symmetry. apply div_unique with 0. left. split; order || apply lt_0_1.
+now nzsimpl.
+Qed.
+
+Lemma mod_1_r: forall a, a mod 1 == 0.
+Proof.
+intros. rewrite mod_eq, div_1_r; nzsimpl; auto using sub_diag.
+intro EQ; symmetry in EQ; revert EQ; apply lt_neq; apply lt_0_1.
+Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof. exact div_1_l. Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof. exact mod_1_l. Qed.
+
+Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.
+Proof.
+intros. symmetry. apply div_unique with 0.
+destruct (lt_ge_cases 0 b); [left|right]; split; order.
+nzsimpl; apply mul_comm.
+Qed.
+
+Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.
+Proof.
+intros. rewrite mod_eq, div_mul by trivial. rewrite mul_comm; apply sub_diag.
+Qed.
+
+(** * Order results about mod and div *)
+
+(** A modulo cannot grow beyond its starting point. *)
+
+Theorem mod_le: forall a b, 0<=a -> 0<b -> a mod b <= a.
+Proof. exact mod_le. Qed.
+
+Theorem div_pos : forall a b, 0<=a -> 0<b -> 0<= a/b.
+Proof. exact div_pos. Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof. exact div_str_pos. Qed.
+
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> 0<=a<b \/ b<a<=0).
+Proof.
+intros a b Hb.
+split.
+intros EQ.
+rewrite (div_mod a b Hb), EQ; nzsimpl.
+now apply mod_bound_or.
+destruct 1. now apply div_small.
+rewrite <- div_opp_opp by trivial. apply div_small; trivial.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; tauto.
+Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> 0<=a<b \/ b<a<=0).
+Proof.
+intros.
+rewrite <- div_small_iff, mod_eq by trivial.
+rewrite sub_move_r, <- (add_0_r a) at 1. rewrite add_cancel_l.
+rewrite eq_sym_iff, eq_mul_0. tauto.
+Qed.
+
+(** As soon as the divisor is strictly greater than 1,
+ the division is strictly decreasing. *)
+
+Lemma div_lt : forall a b, 0<a -> 1<b -> a/b < a.
+Proof. exact div_lt. Qed.
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono : forall a b c, 0<c -> a<=b -> a/c <= b/c.
+Proof.
+intros a b c Hc Hab.
+rewrite lt_eq_cases in Hab. destruct Hab as [LT|EQ];
+ [|rewrite EQ; order].
+rewrite <- lt_succ_r.
+rewrite (mul_lt_mono_pos_l c) by order.
+nzsimpl.
+rewrite (add_lt_mono_r _ _ (a mod c)).
+rewrite <- div_mod by order.
+apply lt_le_trans with b; trivial.
+rewrite (div_mod b c) at 1 by order.
+rewrite <- add_assoc, <- add_le_mono_l.
+apply le_trans with (c+0).
+nzsimpl; destruct (mod_pos_bound b c); order.
+rewrite <- add_le_mono_l. destruct (mod_pos_bound a c); order.
+Qed.
+
+(** In this convention, [div] performs Rounding-Toward-Bottom.
+
+ Since we cannot speak of rational values here, we express this
+ fact by multiplying back by [b], and this leads to separates
+ statements according to the sign of [b].
+
+ First, [a/b] is below the exact fraction ...
+*)
+
+Lemma mul_div_le : forall a b, 0<b -> b*(a/b) <= a.
+Proof.
+intros.
+rewrite (div_mod a b) at 2; try order.
+rewrite <- (add_0_r (b*(a/b))) at 1.
+rewrite <- add_le_mono_l.
+now destruct (mod_pos_bound a b).
+Qed.
+
+Lemma mul_div_ge : forall a b, b<0 -> a <= b*(a/b).
+Proof.
+intros. rewrite <- div_opp_opp, opp_le_mono, <-mul_opp_l by order.
+apply mul_div_le. now rewrite opp_pos_neg.
+Qed.
+
+(** ... and moreover it is the larger such integer, since [S(a/b)]
+ is strictly above the exact fraction.
+*)
+
+Lemma mul_succ_div_gt: forall a b, 0<b -> a < b*(S (a/b)).
+Proof.
+intros.
+nzsimpl.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- add_lt_mono_l.
+destruct (mod_pos_bound a b); order.
+Qed.
+
+Lemma mul_succ_div_lt: forall a b, b<0 -> b*(S (a/b)) < a.
+Proof.
+intros. rewrite <- div_opp_opp, opp_lt_mono, <-mul_opp_l by order.
+apply mul_succ_div_gt. now rewrite opp_pos_neg.
+Qed.
+
+(** NB: The four previous properties could be used as
+ specifications for [div]. *)
+
+(** Inequality [mul_div_le] is exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).
+Proof.
+intros.
+rewrite (div_mod a b) at 1; try order.
+rewrite <- (add_0_r (b*(a/b))) at 2.
+apply add_cancel_l.
+Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, 0<b -> a < b*q -> a/b < q.
+Proof.
+intros.
+rewrite (mul_lt_mono_pos_l b) by trivial.
+apply le_lt_trans with a; trivial.
+now apply mul_div_le.
+Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, 0<b -> a <= b*q -> a/b <= q.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; trivial. now rewrite mul_comm.
+Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, 0<b -> b*q <= a -> q <= a/b.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; trivial. now rewrite mul_comm.
+Qed.
+
+(** A division respects opposite monotonicity for the divisor *)
+
+Lemma div_le_compat_l: forall p q r, 0<=p -> 0<q<=r -> p/r <= p/q.
+Proof. exact div_le_compat_l. Qed.
+
+(** * Relations between usual operations and mod and div *)
+
+Lemma mod_add : forall a b c, c~=0 ->
+ (a + b * c) mod c == a mod c.
+Proof.
+intros.
+symmetry.
+apply mod_unique with (a/c+b); trivial.
+now apply mod_bound_or.
+rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+now rewrite mul_comm.
+Qed.
+
+Lemma div_add : forall a b c, c~=0 ->
+ (a + b * c) / c == a / c + b.
+Proof.
+intros.
+apply (mul_cancel_l _ _ c); try order.
+apply (add_cancel_r _ _ ((a+b*c) mod c)).
+rewrite <- div_mod, mod_add by order.
+rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+now rewrite mul_comm.
+Qed.
+
+Lemma div_add_l: forall a b c, b~=0 ->
+ (a * b + c) / b == a + c / b.
+Proof.
+ intros a b c. rewrite (add_comm _ c), (add_comm a).
+ now apply div_add.
+Qed.
+
+(** Cancellations. *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> c~=0 ->
+ (a*c)/(b*c) == a/b.
+Proof.
+intros.
+symmetry.
+apply div_unique with ((a mod b)*c).
+(* ineqs *)
+destruct (lt_ge_cases 0 c).
+rewrite <-(mul_0_l c), <-2mul_lt_mono_pos_r, <-2mul_le_mono_pos_r by trivial.
+now apply mod_bound_or.
+rewrite <-(mul_0_l c), <-2mul_lt_mono_neg_r, <-2mul_le_mono_neg_r by order.
+destruct (mod_bound_or a b); tauto.
+(* equation *)
+rewrite (div_mod a b) at 1 by order.
+rewrite mul_add_distr_r.
+rewrite add_cancel_r.
+rewrite <- 2 mul_assoc. now rewrite (mul_comm c).
+Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> c~=0 ->
+ (c*a)/(c*b) == a/b.
+Proof.
+intros. rewrite !(mul_comm c); now apply div_mul_cancel_r.
+Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> c~=0 ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof.
+intros.
+rewrite <- (add_cancel_l _ _ ((c*b)* ((c*a)/(c*b)))).
+rewrite <- div_mod.
+rewrite div_mul_cancel_l by trivial.
+rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+apply div_mod; order.
+rewrite <- neq_mul_0; auto.
+Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> c~=0 ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof.
+ intros. rewrite !(mul_comm _ c); now rewrite mul_mod_distr_l.
+Qed.
+
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof.
+intros. rewrite mod_small_iff by trivial.
+now apply mod_bound_or.
+Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof.
+ intros a b n Hn. symmetry.
+ rewrite (div_mod a n) at 1 by order.
+ rewrite add_comm, (mul_comm n), (mul_comm _ b).
+ rewrite mul_add_distr_l, mul_assoc.
+ intros. rewrite mod_add by trivial.
+ now rewrite mul_comm.
+Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof.
+ intros. rewrite !(mul_comm a). now apply mul_mod_idemp_l.
+Qed.
+
+Theorem mul_mod: forall a b n, n~=0 ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof.
+ intros. now rewrite mul_mod_idemp_l, mul_mod_idemp_r.
+Qed.
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof.
+ intros a b n Hn. symmetry.
+ rewrite (div_mod a n) at 1 by order.
+ rewrite <- add_assoc, add_comm, mul_comm.
+ intros. now rewrite mod_add.
+Qed.
+
+Lemma add_mod_idemp_r : forall a b n, n~=0 ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof.
+ intros. rewrite !(add_comm a). now apply add_mod_idemp_l.
+Qed.
+
+Theorem add_mod: forall a b n, n~=0 ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof.
+ intros. now rewrite add_mod_idemp_l, add_mod_idemp_r.
+Qed.
+
+(** With the current convention, the following result isn't always
+ true for negative divisors. For instance
+ [ 3/(-2)/(-2) = 1 <> 0 = 3 / (-2*-2) ]. *)
+
+Lemma div_div : forall a b c, 0<b -> 0<c ->
+ (a/b)/c == a/(b*c).
+Proof.
+ intros a b c Hb Hc.
+ apply div_unique with (b*((a/b) mod c) + a mod b).
+ (* begin 0<= ... <b*c \/ ... *)
+ left.
+ destruct (mod_pos_bound (a/b) c), (mod_pos_bound a b); trivial.
+ split.
+ apply add_nonneg_nonneg; trivial.
+ apply mul_nonneg_nonneg; order.
+ apply lt_le_trans with (b*((a/b) mod c) + b).
+ now rewrite <- add_lt_mono_l.
+ now rewrite <- mul_succ_r, <- mul_le_mono_pos_l, le_succ_l.
+ (* end 0<= ... < b*c \/ ... *)
+ rewrite (div_mod a b) at 1 by order.
+ rewrite add_assoc, add_cancel_r.
+ rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+ apply div_mod; order.
+Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, 0<=a -> 0<b -> 0<=c -> c*(a/b) <= (c*a)/b.
+Proof. exact div_mul_le. Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, b~=0 ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof.
+intros a b Hb. split.
+intros Hab. exists (a/b). rewrite (div_mod a b Hb) at 1.
+ rewrite Hab. now nzsimpl.
+intros (c,Hc).
+rewrite Hc, mul_comm.
+now apply mod_mul.
+Qed.
+
+End ZDivPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZDivTrunc.v b/theories/Numbers/Integer/Abstract/ZDivTrunc.v
new file mode 100644
index 00000000..3200ba2a
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZDivTrunc.v
@@ -0,0 +1,532 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * Euclidean Division for integers (Trunc convention)
+
+ We use here the convention known as Trunc, or Round-Toward-Zero,
+ where [a/b] is the integer with the largest absolute value to
+ be between zero and the exact fraction. It can be summarized by:
+
+ [a = bq+r /\ 0 <= |r| < |b| /\ Sign(r) = Sign(a)]
+
+ This is the convention of Ocaml and many other systems (C, ASM, ...).
+ This convention is named "T" in the following paper:
+
+ R. Boute, "The Euclidean definition of the functions div and mod",
+ ACM Transactions on Programming Languages and Systems,
+ Vol. 14, No.2, pp. 127-144, April 1992.
+
+ See files [ZDivFloor] and [ZDivEucl] for others conventions.
+*)
+
+Require Import ZAxioms ZProperties NZDiv.
+
+Module Type ZDivSpecific (Import Z:ZAxiomsSig')(Import DM : DivMod' Z).
+ Axiom mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+ Axiom mod_opp_l : forall a b, b ~= 0 -> (-a) mod b == - (a mod b).
+ Axiom mod_opp_r : forall a b, b ~= 0 -> a mod (-b) == a mod b.
+End ZDivSpecific.
+
+Module Type ZDiv (Z:ZAxiomsSig)
+ := DivMod Z <+ NZDivCommon Z <+ ZDivSpecific Z.
+
+Module Type ZDivSig := ZAxiomsExtSig <+ ZDiv.
+Module Type ZDivSig' := ZAxiomsExtSig' <+ ZDiv <+ DivModNotation.
+
+Module ZDivPropFunct (Import Z : ZDivSig')(Import ZP : ZPropSig Z).
+
+(** We benefit from what already exists for NZ *)
+
+ Module Import NZDivP := NZDivPropFunct Z ZP Z.
+
+Ltac pos_or_neg a :=
+ let LT := fresh "LT" in
+ let LE := fresh "LE" in
+ destruct (le_gt_cases 0 a) as [LE|LT]; [|rewrite <- opp_pos_neg in LT].
+
+(** Another formulation of the main equation *)
+
+Lemma mod_eq :
+ forall a b, b~=0 -> a mod b == a - b*(a/b).
+Proof.
+intros.
+rewrite <- add_move_l.
+symmetry. now apply div_mod.
+Qed.
+
+(** A few sign rules (simple ones) *)
+
+Lemma mod_opp_opp : forall a b, b ~= 0 -> (-a) mod (-b) == - (a mod b).
+Proof. intros. now rewrite mod_opp_r, mod_opp_l. Qed.
+
+Lemma div_opp_l : forall a b, b ~= 0 -> (-a)/b == -(a/b).
+Proof.
+intros.
+rewrite <- (mul_cancel_l _ _ b) by trivial.
+rewrite <- (add_cancel_r _ _ ((-a) mod b)).
+now rewrite <- div_mod, mod_opp_l, mul_opp_r, <- opp_add_distr, <- div_mod.
+Qed.
+
+Lemma div_opp_r : forall a b, b ~= 0 -> a/(-b) == -(a/b).
+Proof.
+intros.
+assert (-b ~= 0) by (now rewrite eq_opp_l, opp_0).
+rewrite <- (mul_cancel_l _ _ (-b)) by trivial.
+rewrite <- (add_cancel_r _ _ (a mod (-b))).
+now rewrite <- div_mod, mod_opp_r, mul_opp_opp, <- div_mod.
+Qed.
+
+Lemma div_opp_opp : forall a b, b ~= 0 -> (-a)/(-b) == a/b.
+Proof. intros. now rewrite div_opp_r, div_opp_l, opp_involutive. Qed.
+
+(** The sign of [a mod b] is the one of [a] *)
+
+(* TODO: a proper sgn function and theory *)
+
+Lemma mod_sign : forall a b, b~=0 -> 0 <= (a mod b) * a.
+Proof.
+assert (Aux : forall a b, 0<b -> 0 <= (a mod b) * a).
+ intros. pos_or_neg a.
+ apply mul_nonneg_nonneg; trivial. now destruct (mod_bound a b).
+ rewrite <- mul_opp_opp, <- mod_opp_l by order.
+ apply mul_nonneg_nonneg; try order. destruct (mod_bound (-a) b); order.
+intros. pos_or_neg b. apply Aux; order.
+rewrite <- mod_opp_r by order. apply Aux; order.
+Qed.
+
+
+(** Uniqueness theorems *)
+
+Theorem div_mod_unique : forall b q1 q2 r1 r2 : t,
+ (0<=r1<b \/ b<r1<=0) -> (0<=r2<b \/ b<r2<=0) ->
+ b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.
+Proof.
+intros b q1 q2 r1 r2 Hr1 Hr2 EQ.
+destruct Hr1; destruct Hr2; try (intuition; order).
+apply div_mod_unique with b; trivial.
+rewrite <- (opp_inj_wd r1 r2).
+apply div_mod_unique with (-b); trivial.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; tauto.
+rewrite <- opp_lt_mono, opp_nonneg_nonpos; tauto.
+now rewrite 2 mul_opp_l, <- 2 opp_add_distr, opp_inj_wd.
+Qed.
+
+Theorem div_unique:
+ forall a b q r, 0<=a -> 0<=r<b -> a == b*q + r -> q == a/b.
+Proof. intros; now apply div_unique with r. Qed.
+
+Theorem mod_unique:
+ forall a b q r, 0<=a -> 0<=r<b -> a == b*q + r -> r == a mod b.
+Proof. intros; now apply mod_unique with q. Qed.
+
+(** A division by itself returns 1 *)
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof.
+intros. pos_or_neg a. apply div_same; order.
+rewrite <- div_opp_opp by trivial. now apply div_same.
+Qed.
+
+Lemma mod_same : forall a, a~=0 -> a mod a == 0.
+Proof.
+intros. rewrite mod_eq, div_same by trivial. nzsimpl. apply sub_diag.
+Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, 0<=a<b -> a/b == 0.
+Proof. exact div_small. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.
+Proof. exact mod_small. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, a~=0 -> 0/a == 0.
+Proof.
+intros. pos_or_neg a. apply div_0_l; order.
+rewrite <- div_opp_opp, opp_0 by trivial. now apply div_0_l.
+Qed.
+
+Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.
+Proof.
+intros; rewrite mod_eq, div_0_l; now nzsimpl.
+Qed.
+
+Lemma div_1_r: forall a, a/1 == a.
+Proof.
+intros. pos_or_neg a. now apply div_1_r.
+apply opp_inj. rewrite <- div_opp_l. apply div_1_r; order.
+intro EQ; symmetry in EQ; revert EQ; apply lt_neq, lt_0_1.
+Qed.
+
+Lemma mod_1_r: forall a, a mod 1 == 0.
+Proof.
+intros. rewrite mod_eq, div_1_r; nzsimpl; auto using sub_diag.
+intro EQ; symmetry in EQ; revert EQ; apply lt_neq; apply lt_0_1.
+Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof. exact div_1_l. Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof. exact mod_1_l. Qed.
+
+Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.
+Proof.
+intros. pos_or_neg a; pos_or_neg b. apply div_mul; order.
+rewrite <- div_opp_opp, <- mul_opp_r by order. apply div_mul; order.
+rewrite <- opp_inj_wd, <- div_opp_l, <- mul_opp_l by order. apply div_mul; order.
+rewrite <- opp_inj_wd, <- div_opp_r, <- mul_opp_opp by order. apply div_mul; order.
+Qed.
+
+Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.
+Proof.
+intros. rewrite mod_eq, div_mul by trivial. rewrite mul_comm; apply sub_diag.
+Qed.
+
+(** * Order results about mod and div *)
+
+(** A modulo cannot grow beyond its starting point. *)
+
+Theorem mod_le: forall a b, 0<=a -> 0<b -> a mod b <= a.
+Proof. exact mod_le. Qed.
+
+Theorem div_pos : forall a b, 0<=a -> 0<b -> 0<= a/b.
+Proof. exact div_pos. Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof. exact div_str_pos. Qed.
+
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> abs a < abs b).
+Proof.
+intros. pos_or_neg a; pos_or_neg b.
+rewrite div_small_iff; try order. rewrite 2 abs_eq; intuition; order.
+rewrite <- opp_inj_wd, opp_0, <- div_opp_r, div_small_iff by order.
+ rewrite (abs_eq a), (abs_neq' b); intuition; order.
+rewrite <- opp_inj_wd, opp_0, <- div_opp_l, div_small_iff by order.
+ rewrite (abs_neq' a), (abs_eq b); intuition; order.
+rewrite <- div_opp_opp, div_small_iff by order.
+ rewrite (abs_neq' a), (abs_neq' b); intuition; order.
+Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> abs a < abs b).
+Proof.
+intros. rewrite mod_eq, <- div_small_iff by order.
+rewrite sub_move_r, <- (add_0_r a) at 1. rewrite add_cancel_l.
+rewrite eq_sym_iff, eq_mul_0. tauto.
+Qed.
+
+(** As soon as the divisor is strictly greater than 1,
+ the division is strictly decreasing. *)
+
+Lemma div_lt : forall a b, 0<a -> 1<b -> a/b < a.
+Proof. exact div_lt. Qed.
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono : forall a b c, 0<c -> a<=b -> a/c <= b/c.
+Proof.
+intros. pos_or_neg a. apply div_le_mono; auto.
+pos_or_neg b. apply le_trans with 0.
+ rewrite <- opp_nonneg_nonpos, <- div_opp_l by order.
+ apply div_pos; order.
+ apply div_pos; order.
+rewrite opp_le_mono in *. rewrite <- 2 div_opp_l by order.
+ apply div_le_mono; intuition; order.
+Qed.
+
+(** With this choice of division,
+ rounding of div is always done toward zero: *)
+
+Lemma mul_div_le : forall a b, 0<=a -> b~=0 -> 0 <= b*(a/b) <= a.
+Proof.
+intros. pos_or_neg b.
+split.
+apply mul_nonneg_nonneg; [|apply div_pos]; order.
+apply mul_div_le; order.
+rewrite <- mul_opp_opp, <- div_opp_r by order.
+split.
+apply mul_nonneg_nonneg; [|apply div_pos]; order.
+apply mul_div_le; order.
+Qed.
+
+Lemma mul_div_ge : forall a b, a<=0 -> b~=0 -> a <= b*(a/b) <= 0.
+Proof.
+intros.
+rewrite <- opp_nonneg_nonpos, opp_le_mono, <-mul_opp_r, <-div_opp_l by order.
+rewrite <- opp_nonneg_nonpos in *.
+destruct (mul_div_le (-a) b); tauto.
+Qed.
+
+(** For positive numbers, considering [S (a/b)] leads to an upper bound for [a] *)
+
+Lemma mul_succ_div_gt: forall a b, 0<=a -> 0<b -> a < b*(S (a/b)).
+Proof. exact mul_succ_div_gt. Qed.
+
+(** Similar results with negative numbers *)
+
+Lemma mul_pred_div_lt: forall a b, a<=0 -> 0<b -> b*(P (a/b)) < a.
+Proof.
+intros.
+rewrite opp_lt_mono, <- mul_opp_r, opp_pred, <- div_opp_l by order.
+rewrite <- opp_nonneg_nonpos in *.
+now apply mul_succ_div_gt.
+Qed.
+
+Lemma mul_pred_div_gt: forall a b, 0<=a -> b<0 -> a < b*(P (a/b)).
+Proof.
+intros.
+rewrite <- mul_opp_opp, opp_pred, <- div_opp_r by order.
+rewrite <- opp_pos_neg in *.
+now apply mul_succ_div_gt.
+Qed.
+
+Lemma mul_succ_div_lt: forall a b, a<=0 -> b<0 -> b*(S (a/b)) < a.
+Proof.
+intros.
+rewrite opp_lt_mono, <- mul_opp_l, <- div_opp_opp by order.
+rewrite <- opp_nonneg_nonpos, <- opp_pos_neg in *.
+now apply mul_succ_div_gt.
+Qed.
+
+(** Inequality [mul_div_le] is exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).
+Proof.
+intros. rewrite mod_eq by order. rewrite sub_move_r; nzsimpl; tauto.
+Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, 0<=a -> 0<b -> a < b*q -> a/b < q.
+Proof. exact div_lt_upper_bound. Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, 0<b -> a <= b*q -> a/b <= q.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; trivial. now rewrite mul_comm.
+Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, 0<b -> b*q <= a -> q <= a/b.
+Proof.
+intros.
+rewrite <- (div_mul q b) by order.
+apply div_le_mono; trivial. now rewrite mul_comm.
+Qed.
+
+(** A division respects opposite monotonicity for the divisor *)
+
+Lemma div_le_compat_l: forall p q r, 0<=p -> 0<q<=r -> p/r <= p/q.
+Proof. exact div_le_compat_l. Qed.
+
+(** * Relations between usual operations and mod and div *)
+
+(** Unlike with other division conventions, some results here aren't
+ always valid, and need to be restricted. For instance
+ [(a+b*c) mod c <> a mod c] for [a=9,b=-5,c=2] *)
+
+Lemma mod_add : forall a b c, c~=0 -> 0 <= (a+b*c)*a ->
+ (a + b * c) mod c == a mod c.
+Proof.
+assert (forall a b c, c~=0 -> 0<=a -> 0<=a+b*c -> (a+b*c) mod c == a mod c).
+ intros. pos_or_neg c. apply mod_add; order.
+ rewrite <- (mod_opp_r a), <- (mod_opp_r (a+b*c)) by order.
+ rewrite <- mul_opp_opp in *.
+ apply mod_add; order.
+intros a b c Hc Habc.
+destruct (le_0_mul _ _ Habc) as [(Habc',Ha)|(Habc',Ha)]. auto.
+apply opp_inj. revert Ha Habc'.
+rewrite <- 2 opp_nonneg_nonpos.
+rewrite <- 2 mod_opp_l, opp_add_distr, <- mul_opp_l by order. auto.
+Qed.
+
+Lemma div_add : forall a b c, c~=0 -> 0 <= (a+b*c)*a ->
+ (a + b * c) / c == a / c + b.
+Proof.
+intros.
+rewrite <- (mul_cancel_l _ _ c) by trivial.
+rewrite <- (add_cancel_r _ _ ((a+b*c) mod c)).
+rewrite <- div_mod, mod_add by trivial.
+now rewrite mul_add_distr_l, add_shuffle0, <-div_mod, mul_comm.
+Qed.
+
+Lemma div_add_l: forall a b c, b~=0 -> 0 <= (a*b+c)*c ->
+ (a * b + c) / b == a + c / b.
+Proof.
+ intros a b c. rewrite add_comm, (add_comm a). now apply div_add.
+Qed.
+
+(** Cancellations. *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> c~=0 ->
+ (a*c)/(b*c) == a/b.
+Proof.
+assert (Aux1 : forall a b c, 0<=a -> 0<b -> c~=0 -> (a*c)/(b*c) == a/b).
+ intros. pos_or_neg c. apply div_mul_cancel_r; order.
+ rewrite <- div_opp_opp, <- 2 mul_opp_r. apply div_mul_cancel_r; order.
+ rewrite <- neq_mul_0; intuition order.
+assert (Aux2 : forall a b c, 0<=a -> b~=0 -> c~=0 -> (a*c)/(b*c) == a/b).
+ intros. pos_or_neg b. apply Aux1; order.
+ apply opp_inj. rewrite <- 2 div_opp_r, <- mul_opp_l; try order. apply Aux1; order.
+ rewrite <- neq_mul_0; intuition order.
+intros. pos_or_neg a. apply Aux2; order.
+apply opp_inj. rewrite <- 2 div_opp_l, <- mul_opp_l; try order. apply Aux2; order.
+rewrite <- neq_mul_0; intuition order.
+Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> c~=0 ->
+ (c*a)/(c*b) == a/b.
+Proof.
+intros. rewrite !(mul_comm c); now apply div_mul_cancel_r.
+Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> c~=0 ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof.
+intros.
+assert (b*c ~= 0) by (rewrite <- neq_mul_0; tauto).
+rewrite ! mod_eq by trivial.
+rewrite div_mul_cancel_r by order.
+now rewrite mul_sub_distr_r, <- !mul_assoc, (mul_comm (a/b) c).
+Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> c~=0 ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof.
+intros; rewrite !(mul_comm c); now apply mul_mod_distr_r.
+Qed.
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof.
+intros. pos_or_neg a; pos_or_neg n. apply mod_mod; order.
+rewrite <- ! (mod_opp_r _ n) by trivial. apply mod_mod; order.
+apply opp_inj. rewrite <- !mod_opp_l by order. apply mod_mod; order.
+apply opp_inj. rewrite <- !mod_opp_opp by order. apply mod_mod; order.
+Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof.
+assert (Aux1 : forall a b n, 0<=a -> 0<=b -> n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n).
+ intros. pos_or_neg n. apply mul_mod_idemp_l; order.
+ rewrite <- ! (mod_opp_r _ n) by order. apply mul_mod_idemp_l; order.
+assert (Aux2 : forall a b n, 0<=a -> n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n).
+ intros. pos_or_neg b. now apply Aux1.
+ apply opp_inj. rewrite <-2 mod_opp_l, <-2 mul_opp_r by order.
+ apply Aux1; order.
+intros a b n Hn. pos_or_neg a. now apply Aux2.
+apply opp_inj. rewrite <-2 mod_opp_l, <-2 mul_opp_l, <-mod_opp_l by order.
+apply Aux2; order.
+Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof.
+intros. rewrite !(mul_comm a). now apply mul_mod_idemp_l.
+Qed.
+
+Theorem mul_mod: forall a b n, n~=0 ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof.
+intros. now rewrite mul_mod_idemp_l, mul_mod_idemp_r.
+Qed.
+
+(** addition and modulo
+
+ Generally speaking, unlike with other conventions, we don't have
+ [(a+b) mod n = (a mod n + b mod n) mod n]
+ for any a and b.
+ For instance, take (8 + (-10)) mod 3 = -2 whereas
+ (8 mod 3 + (-10 mod 3)) mod 3 = 1.
+*)
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 -> 0 <= a*b ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof.
+assert (Aux : forall a b n, 0<=a -> 0<=b -> n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n).
+ intros. pos_or_neg n. apply add_mod_idemp_l; order.
+ rewrite <- ! (mod_opp_r _ n) by order. apply add_mod_idemp_l; order.
+intros a b n Hn Hab. destruct (le_0_mul _ _ Hab) as [(Ha,Hb)|(Ha,Hb)].
+now apply Aux.
+apply opp_inj. rewrite <-2 mod_opp_l, 2 opp_add_distr, <-mod_opp_l by order.
+rewrite <- opp_nonneg_nonpos in *.
+now apply Aux.
+Qed.
+
+Lemma add_mod_idemp_r : forall a b n, n~=0 -> 0 <= a*b ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof.
+intros. rewrite !(add_comm a). apply add_mod_idemp_l; trivial.
+now rewrite mul_comm.
+Qed.
+
+Theorem add_mod: forall a b n, n~=0 -> 0 <= a*b ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof.
+intros a b n Hn Hab. rewrite add_mod_idemp_l, add_mod_idemp_r; trivial.
+reflexivity.
+destruct (le_0_mul _ _ Hab) as [(Ha,Hb)|(Ha,Hb)];
+ destruct (le_0_mul _ _ (mod_sign b n Hn)) as [(Hb',Hm)|(Hb',Hm)];
+ auto using mul_nonneg_nonneg, mul_nonpos_nonpos.
+ setoid_replace b with 0 by order. rewrite mod_0_l by order. nzsimpl; order.
+ setoid_replace b with 0 by order. rewrite mod_0_l by order. nzsimpl; order.
+Qed.
+
+
+(** Conversely, the following result needs less restrictions here. *)
+
+Lemma div_div : forall a b c, b~=0 -> c~=0 ->
+ (a/b)/c == a/(b*c).
+Proof.
+assert (Aux1 : forall a b c, 0<=a -> 0<b -> c~=0 -> (a/b)/c == a/(b*c)).
+ intros. pos_or_neg c. apply div_div; order.
+ apply opp_inj. rewrite <- 2 div_opp_r, <- mul_opp_r; trivial.
+ apply div_div; order.
+ rewrite <- neq_mul_0; intuition order.
+assert (Aux2 : forall a b c, 0<=a -> b~=0 -> c~=0 -> (a/b)/c == a/(b*c)).
+ intros. pos_or_neg b. apply Aux1; order.
+ apply opp_inj. rewrite <- div_opp_l, <- 2 div_opp_r, <- mul_opp_l; trivial.
+ apply Aux1; trivial.
+ rewrite <- neq_mul_0; intuition order.
+intros. pos_or_neg a. apply Aux2; order.
+apply opp_inj. rewrite <- 3 div_opp_l; try order. apply Aux2; order.
+rewrite <- neq_mul_0. tauto.
+Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, 0<=a -> 0<b -> 0<=c -> c*(a/b) <= (c*a)/b.
+Proof. exact div_mul_le. Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, b~=0 ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof.
+ intros a b Hb. split.
+ intros Hab. exists (a/b). rewrite (div_mod a b Hb) at 1.
+ rewrite Hab; now nzsimpl.
+ intros (c,Hc). rewrite Hc, mul_comm. now apply mod_mul.
+Qed.
+
+End ZDivPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZDomain.v b/theories/Numbers/Integer/Abstract/ZDomain.v
deleted file mode 100644
index 9a17e151..00000000
--- a/theories/Numbers/Integer/Abstract/ZDomain.v
+++ /dev/null
@@ -1,69 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Evgeny Makarov, INRIA, 2007 *)
-(************************************************************************)
-
-(*i $Id: ZDomain.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
-
-Require Export NumPrelude.
-
-Module Type ZDomainSignature.
-
-Parameter Inline Z : Set.
-Parameter Inline Zeq : Z -> Z -> Prop.
-Parameter Inline e : Z -> Z -> bool.
-
-Axiom eq_equiv_e : forall x y : Z, Zeq x y <-> e x y.
-Axiom eq_equiv : equiv Z Zeq.
-
-Add Relation Z Zeq
- reflexivity proved by (proj1 eq_equiv)
- symmetry proved by (proj2 (proj2 eq_equiv))
- transitivity proved by (proj1 (proj2 eq_equiv))
-as eq_rel.
-
-Delimit Scope IntScope with Int.
-Bind Scope IntScope with Z.
-Notation "x == y" := (Zeq x y) (at level 70) : IntScope.
-Notation "x # y" := (~ Zeq x y) (at level 70) : IntScope.
-
-End ZDomainSignature.
-
-Module ZDomainProperties (Import ZDomainModule : ZDomainSignature).
-Open Local Scope IntScope.
-
-Add Morphism e with signature Zeq ==> Zeq ==> eq_bool as e_wd.
-Proof.
-intros x x' Exx' y y' Eyy'.
-case_eq (e x y); case_eq (e x' y'); intros H1 H2; trivial.
-assert (x == y); [apply <- eq_equiv_e; now rewrite H2 |
-assert (x' == y'); [rewrite <- Exx'; now rewrite <- Eyy' |
-rewrite <- H1; assert (H3 : e x' y'); [now apply -> eq_equiv_e | now inversion H3]]].
-assert (x' == y'); [apply <- eq_equiv_e; now rewrite H1 |
-assert (x == y); [rewrite Exx'; now rewrite Eyy' |
-rewrite <- H2; assert (H3 : e x y); [now apply -> eq_equiv_e | now inversion H3]]].
-Qed.
-
-Theorem neq_sym : forall n m, n # m -> m # n.
-Proof.
-intros n m H1 H2; symmetry in H2; false_hyp H2 H1.
-Qed.
-
-Theorem ZE_stepl : forall x y z : Z, x == y -> x == z -> z == y.
-Proof.
-intros x y z H1 H2; now rewrite <- H1.
-Qed.
-
-Declare Left Step ZE_stepl.
-
-(* The right step lemma is just transitivity of Zeq *)
-Declare Right Step (proj1 (proj2 eq_equiv)).
-
-End ZDomainProperties.
-
-
diff --git a/theories/Numbers/Integer/Abstract/ZLt.v b/theories/Numbers/Integer/Abstract/ZLt.v
index 2a88a535..849bf6b4 100644
--- a/theories/Numbers/Integer/Abstract/ZLt.v
+++ b/theories/Numbers/Integer/Abstract/ZLt.v
@@ -8,424 +8,126 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZLt.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export ZMul.
-Module ZOrderPropFunct (Import ZAxiomsMod : ZAxiomsSig).
-Module Export ZMulPropMod := ZMulPropFunct ZAxiomsMod.
-Open Local Scope IntScope.
+Module ZOrderPropFunct (Import Z : ZAxiomsSig').
+Include ZMulPropFunct Z.
-(* Axioms *)
+(** Instances of earlier theorems for m == 0 *)
-Theorem Zlt_wd :
- forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> (n1 < m1 <-> n2 < m2).
-Proof NZlt_wd.
-
-Theorem Zle_wd :
- forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> (n1 <= m1 <-> n2 <= m2).
-Proof NZle_wd.
-
-Theorem Zmin_wd :
- forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> Zmin n1 m1 == Zmin n2 m2.
-Proof NZmin_wd.
-
-Theorem Zmax_wd :
- forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> Zmax n1 m1 == Zmax n2 m2.
-Proof NZmax_wd.
-
-Theorem Zlt_eq_cases : forall n m : Z, n <= m <-> n < m \/ n == m.
-Proof NZlt_eq_cases.
-
-Theorem Zlt_irrefl : forall n : Z, ~ n < n.
-Proof NZlt_irrefl.
-
-Theorem Zlt_succ_r : forall n m : Z, n < S m <-> n <= m.
-Proof NZlt_succ_r.
-
-Theorem Zmin_l : forall n m : Z, n <= m -> Zmin n m == n.
-Proof NZmin_l.
-
-Theorem Zmin_r : forall n m : Z, m <= n -> Zmin n m == m.
-Proof NZmin_r.
-
-Theorem Zmax_l : forall n m : Z, m <= n -> Zmax n m == n.
-Proof NZmax_l.
-
-Theorem Zmax_r : forall n m : Z, n <= m -> Zmax n m == m.
-Proof NZmax_r.
-
-(* Renaming theorems from NZOrder.v *)
-
-Theorem Zlt_le_incl : forall n m : Z, n < m -> n <= m.
-Proof NZlt_le_incl.
-
-Theorem Zlt_neq : forall n m : Z, n < m -> n ~= m.
-Proof NZlt_neq.
-
-Theorem Zle_neq : forall n m : Z, n < m <-> n <= m /\ n ~= m.
-Proof NZle_neq.
-
-Theorem Zle_refl : forall n : Z, n <= n.
-Proof NZle_refl.
-
-Theorem Zlt_succ_diag_r : forall n : Z, n < S n.
-Proof NZlt_succ_diag_r.
-
-Theorem Zle_succ_diag_r : forall n : Z, n <= S n.
-Proof NZle_succ_diag_r.
-
-Theorem Zlt_0_1 : 0 < 1.
-Proof NZlt_0_1.
-
-Theorem Zle_0_1 : 0 <= 1.
-Proof NZle_0_1.
-
-Theorem Zlt_lt_succ_r : forall n m : Z, n < m -> n < S m.
-Proof NZlt_lt_succ_r.
-
-Theorem Zle_le_succ_r : forall n m : Z, n <= m -> n <= S m.
-Proof NZle_le_succ_r.
-
-Theorem Zle_succ_r : forall n m : Z, n <= S m <-> n <= m \/ n == S m.
-Proof NZle_succ_r.
-
-Theorem Zneq_succ_diag_l : forall n : Z, S n ~= n.
-Proof NZneq_succ_diag_l.
-
-Theorem Zneq_succ_diag_r : forall n : Z, n ~= S n.
-Proof NZneq_succ_diag_r.
-
-Theorem Znlt_succ_diag_l : forall n : Z, ~ S n < n.
-Proof NZnlt_succ_diag_l.
-
-Theorem Znle_succ_diag_l : forall n : Z, ~ S n <= n.
-Proof NZnle_succ_diag_l.
-
-Theorem Zle_succ_l : forall n m : Z, S n <= m <-> n < m.
-Proof NZle_succ_l.
-
-Theorem Zlt_succ_l : forall n m : Z, S n < m -> n < m.
-Proof NZlt_succ_l.
-
-Theorem Zsucc_lt_mono : forall n m : Z, n < m <-> S n < S m.
-Proof NZsucc_lt_mono.
-
-Theorem Zsucc_le_mono : forall n m : Z, n <= m <-> S n <= S m.
-Proof NZsucc_le_mono.
-
-Theorem Zlt_asymm : forall n m, n < m -> ~ m < n.
-Proof NZlt_asymm.
-
-Notation Zlt_ngt := Zlt_asymm (only parsing).
-
-Theorem Zlt_trans : forall n m p : Z, n < m -> m < p -> n < p.
-Proof NZlt_trans.
-
-Theorem Zle_trans : forall n m p : Z, n <= m -> m <= p -> n <= p.
-Proof NZle_trans.
-
-Theorem Zle_lt_trans : forall n m p : Z, n <= m -> m < p -> n < p.
-Proof NZle_lt_trans.
-
-Theorem Zlt_le_trans : forall n m p : Z, n < m -> m <= p -> n < p.
-Proof NZlt_le_trans.
-
-Theorem Zle_antisymm : forall n m : Z, n <= m -> m <= n -> n == m.
-Proof NZle_antisymm.
-
-Theorem Zlt_1_l : forall n m : Z, 0 < n -> n < m -> 1 < m.
-Proof NZlt_1_l.
-
-(** Trichotomy, decidability, and double negation elimination *)
-
-Theorem Zlt_trichotomy : forall n m : Z, n < m \/ n == m \/ m < n.
-Proof NZlt_trichotomy.
-
-Notation Zlt_eq_gt_cases := Zlt_trichotomy (only parsing).
-
-Theorem Zlt_gt_cases : forall n m : Z, n ~= m <-> n < m \/ n > m.
-Proof NZlt_gt_cases.
-
-Theorem Zle_gt_cases : forall n m : Z, n <= m \/ n > m.
-Proof NZle_gt_cases.
-
-Theorem Zlt_ge_cases : forall n m : Z, n < m \/ n >= m.
-Proof NZlt_ge_cases.
-
-Theorem Zle_ge_cases : forall n m : Z, n <= m \/ n >= m.
-Proof NZle_ge_cases.
-
-(** Instances of the previous theorems for m == 0 *)
-
-Theorem Zneg_pos_cases : forall n : Z, n ~= 0 <-> n < 0 \/ n > 0.
+Theorem neg_pos_cases : forall n, n ~= 0 <-> n < 0 \/ n > 0.
Proof.
-intro; apply Zlt_gt_cases.
+intro; apply lt_gt_cases.
Qed.
-Theorem Znonpos_pos_cases : forall n : Z, n <= 0 \/ n > 0.
+Theorem nonpos_pos_cases : forall n, n <= 0 \/ n > 0.
Proof.
-intro; apply Zle_gt_cases.
+intro; apply le_gt_cases.
Qed.
-Theorem Zneg_nonneg_cases : forall n : Z, n < 0 \/ n >= 0.
+Theorem neg_nonneg_cases : forall n, n < 0 \/ n >= 0.
Proof.
-intro; apply Zlt_ge_cases.
+intro; apply lt_ge_cases.
Qed.
-Theorem Znonpos_nonneg_cases : forall n : Z, n <= 0 \/ n >= 0.
+Theorem nonpos_nonneg_cases : forall n, n <= 0 \/ n >= 0.
Proof.
-intro; apply Zle_ge_cases.
+intro; apply le_ge_cases.
Qed.
-Theorem Zle_ngt : forall n m : Z, n <= m <-> ~ n > m.
-Proof NZle_ngt.
-
-Theorem Znlt_ge : forall n m : Z, ~ n < m <-> n >= m.
-Proof NZnlt_ge.
-
-Theorem Zlt_dec : forall n m : Z, decidable (n < m).
-Proof NZlt_dec.
-
-Theorem Zlt_dne : forall n m, ~ ~ n < m <-> n < m.
-Proof NZlt_dne.
-
-Theorem Znle_gt : forall n m : Z, ~ n <= m <-> n > m.
-Proof NZnle_gt.
-
-Theorem Zlt_nge : forall n m : Z, n < m <-> ~ n >= m.
-Proof NZlt_nge.
-
-Theorem Zle_dec : forall n m : Z, decidable (n <= m).
-Proof NZle_dec.
-
-Theorem Zle_dne : forall n m : Z, ~ ~ n <= m <-> n <= m.
-Proof NZle_dne.
-
-Theorem Znlt_succ_r : forall n m : Z, ~ m < S n <-> n < m.
-Proof NZnlt_succ_r.
-
-Theorem Zlt_exists_pred :
- forall z n : Z, z < n -> exists k : Z, n == S k /\ z <= k.
-Proof NZlt_exists_pred.
-
-Theorem Zlt_succ_iter_r :
- forall (n : nat) (m : Z), m < NZsucc_iter (Datatypes.S n) m.
-Proof NZlt_succ_iter_r.
-
-Theorem Zneq_succ_iter_l :
- forall (n : nat) (m : Z), NZsucc_iter (Datatypes.S n) m ~= m.
-Proof NZneq_succ_iter_l.
-
-(** Stronger variant of induction with assumptions n >= 0 (n < 0)
-in the induction step *)
-
-Theorem Zright_induction :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z, A z ->
- (forall n : Z, z <= n -> A n -> A (S n)) ->
- forall n : Z, z <= n -> A n.
-Proof NZright_induction.
-
-Theorem Zleft_induction :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z, A z ->
- (forall n : Z, n < z -> A (S n) -> A n) ->
- forall n : Z, n <= z -> A n.
-Proof NZleft_induction.
-
-Theorem Zright_induction' :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z,
- (forall n : Z, n <= z -> A n) ->
- (forall n : Z, z <= n -> A n -> A (S n)) ->
- forall n : Z, A n.
-Proof NZright_induction'.
-
-Theorem Zleft_induction' :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z,
- (forall n : Z, z <= n -> A n) ->
- (forall n : Z, n < z -> A (S n) -> A n) ->
- forall n : Z, A n.
-Proof NZleft_induction'.
-
-Theorem Zstrong_right_induction :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z,
- (forall n : Z, z <= n -> (forall m : Z, z <= m -> m < n -> A m) -> A n) ->
- forall n : Z, z <= n -> A n.
-Proof NZstrong_right_induction.
-
-Theorem Zstrong_left_induction :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z,
- (forall n : Z, n <= z -> (forall m : Z, m <= z -> S n <= m -> A m) -> A n) ->
- forall n : Z, n <= z -> A n.
-Proof NZstrong_left_induction.
-
-Theorem Zstrong_right_induction' :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z,
- (forall n : Z, n <= z -> A n) ->
- (forall n : Z, z <= n -> (forall m : Z, z <= m -> m < n -> A m) -> A n) ->
- forall n : Z, A n.
-Proof NZstrong_right_induction'.
-
-Theorem Zstrong_left_induction' :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z,
- (forall n : Z, z <= n -> A n) ->
- (forall n : Z, n <= z -> (forall m : Z, m <= z -> S n <= m -> A m) -> A n) ->
- forall n : Z, A n.
-Proof NZstrong_left_induction'.
-
-Theorem Zorder_induction :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z, A z ->
- (forall n : Z, z <= n -> A n -> A (S n)) ->
- (forall n : Z, n < z -> A (S n) -> A n) ->
- forall n : Z, A n.
-Proof NZorder_induction.
-
-Theorem Zorder_induction' :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall z : Z, A z ->
- (forall n : Z, z <= n -> A n -> A (S n)) ->
- (forall n : Z, n <= z -> A n -> A (P n)) ->
- forall n : Z, A n.
-Proof NZorder_induction'.
-
-Theorem Zorder_induction_0 :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- A 0 ->
- (forall n : Z, 0 <= n -> A n -> A (S n)) ->
- (forall n : Z, n < 0 -> A (S n) -> A n) ->
- forall n : Z, A n.
-Proof NZorder_induction_0.
-
-Theorem Zorder_induction'_0 :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- A 0 ->
- (forall n : Z, 0 <= n -> A n -> A (S n)) ->
- (forall n : Z, n <= 0 -> A n -> A (P n)) ->
- forall n : Z, A n.
-Proof NZorder_induction'_0.
-
-Ltac Zinduct n := induction_maker n ltac:(apply Zorder_induction_0).
-
-(** Elimintation principle for < *)
-
-Theorem Zlt_ind :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall n : Z, A (S n) ->
- (forall m : Z, n < m -> A m -> A (S m)) -> forall m : Z, n < m -> A m.
-Proof NZlt_ind.
-
-(** Elimintation principle for <= *)
-
-Theorem Zle_ind :
- forall A : Z -> Prop, predicate_wd Zeq A ->
- forall n : Z, A n ->
- (forall m : Z, n <= m -> A m -> A (S m)) -> forall m : Z, n <= m -> A m.
-Proof NZle_ind.
-
-(** Well-founded relations *)
-
-Theorem Zlt_wf : forall z : Z, well_founded (fun n m : Z => z <= n /\ n < m).
-Proof NZlt_wf.
-
-Theorem Zgt_wf : forall z : Z, well_founded (fun n m : Z => m < n /\ n <= z).
-Proof NZgt_wf.
+Ltac zinduct n := induction_maker n ltac:(apply order_induction_0).
-(* Theorems that are either not valid on N or have different proofs on N and Z *)
+(** Theorems that are either not valid on N or have different proofs
+ on N and Z *)
-Theorem Zlt_pred_l : forall n : Z, P n < n.
+Theorem lt_pred_l : forall n, P n < n.
Proof.
-intro n; rewrite <- (Zsucc_pred n) at 2; apply Zlt_succ_diag_r.
+intro n; rewrite <- (succ_pred n) at 2; apply lt_succ_diag_r.
Qed.
-Theorem Zle_pred_l : forall n : Z, P n <= n.
+Theorem le_pred_l : forall n, P n <= n.
Proof.
-intro; apply Zlt_le_incl; apply Zlt_pred_l.
+intro; apply lt_le_incl; apply lt_pred_l.
Qed.
-Theorem Zlt_le_pred : forall n m : Z, n < m <-> n <= P m.
+Theorem lt_le_pred : forall n m, n < m <-> n <= P m.
Proof.
-intros n m; rewrite <- (Zsucc_pred m); rewrite Zpred_succ. apply Zlt_succ_r.
+intros n m; rewrite <- (succ_pred m); rewrite pred_succ. apply lt_succ_r.
Qed.
-Theorem Znle_pred_r : forall n : Z, ~ n <= P n.
+Theorem nle_pred_r : forall n, ~ n <= P n.
Proof.
-intro; rewrite <- Zlt_le_pred; apply Zlt_irrefl.
+intro; rewrite <- lt_le_pred; apply lt_irrefl.
Qed.
-Theorem Zlt_pred_le : forall n m : Z, P n < m <-> n <= m.
+Theorem lt_pred_le : forall n m, P n < m <-> n <= m.
Proof.
-intros n m; rewrite <- (Zsucc_pred n) at 2.
-symmetry; apply Zle_succ_l.
+intros n m; rewrite <- (succ_pred n) at 2.
+symmetry; apply le_succ_l.
Qed.
-Theorem Zlt_lt_pred : forall n m : Z, n < m -> P n < m.
+Theorem lt_lt_pred : forall n m, n < m -> P n < m.
Proof.
-intros; apply <- Zlt_pred_le; now apply Zlt_le_incl.
+intros; apply <- lt_pred_le; now apply lt_le_incl.
Qed.
-Theorem Zle_le_pred : forall n m : Z, n <= m -> P n <= m.
+Theorem le_le_pred : forall n m, n <= m -> P n <= m.
Proof.
-intros; apply Zlt_le_incl; now apply <- Zlt_pred_le.
+intros; apply lt_le_incl; now apply <- lt_pred_le.
Qed.
-Theorem Zlt_pred_lt : forall n m : Z, n < P m -> n < m.
+Theorem lt_pred_lt : forall n m, n < P m -> n < m.
Proof.
-intros n m H; apply Zlt_trans with (P m); [assumption | apply Zlt_pred_l].
+intros n m H; apply lt_trans with (P m); [assumption | apply lt_pred_l].
Qed.
-Theorem Zle_pred_lt : forall n m : Z, n <= P m -> n <= m.
+Theorem le_pred_lt : forall n m, n <= P m -> n <= m.
Proof.
-intros; apply Zlt_le_incl; now apply <- Zlt_le_pred.
+intros; apply lt_le_incl; now apply <- lt_le_pred.
Qed.
-Theorem Zpred_lt_mono : forall n m : Z, n < m <-> P n < P m.
+Theorem pred_lt_mono : forall n m, n < m <-> P n < P m.
Proof.
-intros; rewrite Zlt_le_pred; symmetry; apply Zlt_pred_le.
+intros; rewrite lt_le_pred; symmetry; apply lt_pred_le.
Qed.
-Theorem Zpred_le_mono : forall n m : Z, n <= m <-> P n <= P m.
+Theorem pred_le_mono : forall n m, n <= m <-> P n <= P m.
Proof.
-intros; rewrite <- Zlt_pred_le; now rewrite Zlt_le_pred.
+intros; rewrite <- lt_pred_le; now rewrite lt_le_pred.
Qed.
-Theorem Zlt_succ_lt_pred : forall n m : Z, S n < m <-> n < P m.
+Theorem lt_succ_lt_pred : forall n m, S n < m <-> n < P m.
Proof.
-intros n m; now rewrite (Zpred_lt_mono (S n) m), Zpred_succ.
+intros n m; now rewrite (pred_lt_mono (S n) m), pred_succ.
Qed.
-Theorem Zle_succ_le_pred : forall n m : Z, S n <= m <-> n <= P m.
+Theorem le_succ_le_pred : forall n m, S n <= m <-> n <= P m.
Proof.
-intros n m; now rewrite (Zpred_le_mono (S n) m), Zpred_succ.
+intros n m; now rewrite (pred_le_mono (S n) m), pred_succ.
Qed.
-Theorem Zlt_pred_lt_succ : forall n m : Z, P n < m <-> n < S m.
+Theorem lt_pred_lt_succ : forall n m, P n < m <-> n < S m.
Proof.
-intros; rewrite Zlt_pred_le; symmetry; apply Zlt_succ_r.
+intros; rewrite lt_pred_le; symmetry; apply lt_succ_r.
Qed.
-Theorem Zle_pred_lt_succ : forall n m : Z, P n <= m <-> n <= S m.
+Theorem le_pred_lt_succ : forall n m, P n <= m <-> n <= S m.
Proof.
-intros n m; now rewrite (Zpred_le_mono n (S m)), Zpred_succ.
+intros n m; now rewrite (pred_le_mono n (S m)), pred_succ.
Qed.
-Theorem Zneq_pred_l : forall n : Z, P n ~= n.
+Theorem neq_pred_l : forall n, P n ~= n.
Proof.
-intro; apply Zlt_neq; apply Zlt_pred_l.
+intro; apply lt_neq; apply lt_pred_l.
Qed.
-Theorem Zlt_n1_r : forall n m : Z, n < m -> m < 0 -> n < -1.
+Theorem lt_n1_r : forall n m, n < m -> m < 0 -> n < -(1).
Proof.
-intros n m H1 H2. apply -> Zlt_le_pred in H2.
-setoid_replace (P 0) with (-1) in H2. now apply NZlt_le_trans with m.
-apply <- Zeq_opp_r. now rewrite Zopp_pred, Zopp_0.
+intros n m H1 H2. apply -> lt_le_pred in H2.
+setoid_replace (P 0) with (-(1)) in H2. now apply lt_le_trans with m.
+apply <- eq_opp_r. now rewrite opp_pred, opp_0.
Qed.
End ZOrderPropFunct.
diff --git a/theories/Numbers/Integer/Abstract/ZMul.v b/theories/Numbers/Integer/Abstract/ZMul.v
index c48d1b4c..84d840ad 100644
--- a/theories/Numbers/Integer/Abstract/ZMul.v
+++ b/theories/Numbers/Integer/Abstract/ZMul.v
@@ -8,106 +8,63 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZMul.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export ZAdd.
-Module ZMulPropFunct (Import ZAxiomsMod : ZAxiomsSig).
-Module Export ZAddPropMod := ZAddPropFunct ZAxiomsMod.
-Open Local Scope IntScope.
+Module ZMulPropFunct (Import Z : ZAxiomsSig').
+Include ZAddPropFunct Z.
-Theorem Zmul_wd :
- forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> n1 * m1 == n2 * m2.
-Proof NZmul_wd.
+(** A note on naming: right (correspondingly, left) distributivity
+ happens when the sum is multiplied by a number on the right
+ (left), not when the sum itself is the right (left) factor in the
+ product (see planetmath.org and mathworld.wolfram.com). In the old
+ library BinInt, distributivity over subtraction was named
+ correctly, but distributivity over addition was named
+ incorrectly. The names in Isabelle/HOL library are also
+ incorrect. *)
-Theorem Zmul_0_l : forall n : Z, 0 * n == 0.
-Proof NZmul_0_l.
+(** Theorems that are either not valid on N or have different proofs
+ on N and Z *)
-Theorem Zmul_succ_l : forall n m : Z, (S n) * m == n * m + m.
-Proof NZmul_succ_l.
-
-(* Theorems that are valid for both natural numbers and integers *)
-
-Theorem Zmul_0_r : forall n : Z, n * 0 == 0.
-Proof NZmul_0_r.
-
-Theorem Zmul_succ_r : forall n m : Z, n * (S m) == n * m + n.
-Proof NZmul_succ_r.
-
-Theorem Zmul_comm : forall n m : Z, n * m == m * n.
-Proof NZmul_comm.
-
-Theorem Zmul_add_distr_r : forall n m p : Z, (n + m) * p == n * p + m * p.
-Proof NZmul_add_distr_r.
-
-Theorem Zmul_add_distr_l : forall n m p : Z, n * (m + p) == n * m + n * p.
-Proof NZmul_add_distr_l.
-
-(* A note on naming: right (correspondingly, left) distributivity happens
-when the sum is multiplied by a number on the right (left), not when the
-sum itself is the right (left) factor in the product (see planetmath.org
-and mathworld.wolfram.com). In the old library BinInt, distributivity over
-subtraction was named correctly, but distributivity over addition was named
-incorrectly. The names in Isabelle/HOL library are also incorrect. *)
-
-Theorem Zmul_assoc : forall n m p : Z, n * (m * p) == (n * m) * p.
-Proof NZmul_assoc.
-
-Theorem Zmul_1_l : forall n : Z, 1 * n == n.
-Proof NZmul_1_l.
-
-Theorem Zmul_1_r : forall n : Z, n * 1 == n.
-Proof NZmul_1_r.
-
-(* The following two theorems are true in an ordered ring,
-but since they don't mention order, we'll put them here *)
-
-Theorem Zeq_mul_0 : forall n m : Z, n * m == 0 <-> n == 0 \/ m == 0.
-Proof NZeq_mul_0.
-
-Theorem Zneq_mul_0 : forall n m : Z, n ~= 0 /\ m ~= 0 <-> n * m ~= 0.
-Proof NZneq_mul_0.
-
-(* Theorems that are either not valid on N or have different proofs on N and Z *)
-
-Theorem Zmul_pred_r : forall n m : Z, n * (P m) == n * m - n.
+Theorem mul_pred_r : forall n m, n * (P m) == n * m - n.
Proof.
intros n m.
-rewrite <- (Zsucc_pred m) at 2.
-now rewrite Zmul_succ_r, <- Zadd_sub_assoc, Zsub_diag, Zadd_0_r.
+rewrite <- (succ_pred m) at 2.
+now rewrite mul_succ_r, <- add_sub_assoc, sub_diag, add_0_r.
Qed.
-Theorem Zmul_pred_l : forall n m : Z, (P n) * m == n * m - m.
+Theorem mul_pred_l : forall n m, (P n) * m == n * m - m.
Proof.
-intros n m; rewrite (Zmul_comm (P n) m), (Zmul_comm n m). apply Zmul_pred_r.
+intros n m; rewrite (mul_comm (P n) m), (mul_comm n m). apply mul_pred_r.
Qed.
-Theorem Zmul_opp_l : forall n m : Z, (- n) * m == - (n * m).
+Theorem mul_opp_l : forall n m, (- n) * m == - (n * m).
Proof.
-intros n m. apply -> Zadd_move_0_r.
-now rewrite <- Zmul_add_distr_r, Zadd_opp_diag_l, Zmul_0_l.
+intros n m. apply -> add_move_0_r.
+now rewrite <- mul_add_distr_r, add_opp_diag_l, mul_0_l.
Qed.
-Theorem Zmul_opp_r : forall n m : Z, n * (- m) == - (n * m).
+Theorem mul_opp_r : forall n m, n * (- m) == - (n * m).
Proof.
-intros n m; rewrite (Zmul_comm n (- m)), (Zmul_comm n m); apply Zmul_opp_l.
+intros n m; rewrite (mul_comm n (- m)), (mul_comm n m); apply mul_opp_l.
Qed.
-Theorem Zmul_opp_opp : forall n m : Z, (- n) * (- m) == n * m.
+Theorem mul_opp_opp : forall n m, (- n) * (- m) == n * m.
Proof.
-intros n m; now rewrite Zmul_opp_l, Zmul_opp_r, Zopp_involutive.
+intros n m; now rewrite mul_opp_l, mul_opp_r, opp_involutive.
Qed.
-Theorem Zmul_sub_distr_l : forall n m p : Z, n * (m - p) == n * m - n * p.
+Theorem mul_sub_distr_l : forall n m p, n * (m - p) == n * m - n * p.
Proof.
-intros n m p. do 2 rewrite <- Zadd_opp_r. rewrite Zmul_add_distr_l.
-now rewrite Zmul_opp_r.
+intros n m p. do 2 rewrite <- add_opp_r. rewrite mul_add_distr_l.
+now rewrite mul_opp_r.
Qed.
-Theorem Zmul_sub_distr_r : forall n m p : Z, (n - m) * p == n * p - m * p.
+Theorem mul_sub_distr_r : forall n m p, (n - m) * p == n * p - m * p.
Proof.
-intros n m p; rewrite (Zmul_comm (n - m) p), (Zmul_comm n p), (Zmul_comm m p);
-now apply Zmul_sub_distr_l.
+intros n m p; rewrite (mul_comm (n - m) p), (mul_comm n p), (mul_comm m p);
+now apply mul_sub_distr_l.
Qed.
End ZMulPropFunct.
diff --git a/theories/Numbers/Integer/Abstract/ZMulOrder.v b/theories/Numbers/Integer/Abstract/ZMulOrder.v
index c7996ffd..99be58eb 100644
--- a/theories/Numbers/Integer/Abstract/ZMulOrder.v
+++ b/theories/Numbers/Integer/Abstract/ZMulOrder.v
@@ -8,335 +8,225 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZMulOrder.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
Require Export ZAddOrder.
-Module ZMulOrderPropFunct (Import ZAxiomsMod : ZAxiomsSig).
-Module Export ZAddOrderPropMod := ZAddOrderPropFunct ZAxiomsMod.
-Open Local Scope IntScope.
+Module Type ZMulOrderPropFunct (Import Z : ZAxiomsSig').
+Include ZAddOrderPropFunct Z.
-Theorem Zmul_lt_pred :
- forall p q n m : Z, S p == q -> (p * n < p * m <-> q * n + m < q * m + n).
-Proof NZmul_lt_pred.
+Local Notation "- 1" := (-(1)).
-Theorem Zmul_lt_mono_pos_l : forall p n m : Z, 0 < p -> (n < m <-> p * n < p * m).
-Proof NZmul_lt_mono_pos_l.
-
-Theorem Zmul_lt_mono_pos_r : forall p n m : Z, 0 < p -> (n < m <-> n * p < m * p).
-Proof NZmul_lt_mono_pos_r.
-
-Theorem Zmul_lt_mono_neg_l : forall p n m : Z, p < 0 -> (n < m <-> p * m < p * n).
-Proof NZmul_lt_mono_neg_l.
-
-Theorem Zmul_lt_mono_neg_r : forall p n m : Z, p < 0 -> (n < m <-> m * p < n * p).
-Proof NZmul_lt_mono_neg_r.
-
-Theorem Zmul_le_mono_nonneg_l : forall n m p : Z, 0 <= p -> n <= m -> p * n <= p * m.
-Proof NZmul_le_mono_nonneg_l.
-
-Theorem Zmul_le_mono_nonpos_l : forall n m p : Z, p <= 0 -> n <= m -> p * m <= p * n.
-Proof NZmul_le_mono_nonpos_l.
-
-Theorem Zmul_le_mono_nonneg_r : forall n m p : Z, 0 <= p -> n <= m -> n * p <= m * p.
-Proof NZmul_le_mono_nonneg_r.
-
-Theorem Zmul_le_mono_nonpos_r : forall n m p : Z, p <= 0 -> n <= m -> m * p <= n * p.
-Proof NZmul_le_mono_nonpos_r.
-
-Theorem Zmul_cancel_l : forall n m p : Z, p ~= 0 -> (p * n == p * m <-> n == m).
-Proof NZmul_cancel_l.
-
-Theorem Zmul_cancel_r : forall n m p : Z, p ~= 0 -> (n * p == m * p <-> n == m).
-Proof NZmul_cancel_r.
-
-Theorem Zmul_id_l : forall n m : Z, m ~= 0 -> (n * m == m <-> n == 1).
-Proof NZmul_id_l.
-
-Theorem Zmul_id_r : forall n m : Z, n ~= 0 -> (n * m == n <-> m == 1).
-Proof NZmul_id_r.
-
-Theorem Zmul_le_mono_pos_l : forall n m p : Z, 0 < p -> (n <= m <-> p * n <= p * m).
-Proof NZmul_le_mono_pos_l.
-
-Theorem Zmul_le_mono_pos_r : forall n m p : Z, 0 < p -> (n <= m <-> n * p <= m * p).
-Proof NZmul_le_mono_pos_r.
-
-Theorem Zmul_le_mono_neg_l : forall n m p : Z, p < 0 -> (n <= m <-> p * m <= p * n).
-Proof NZmul_le_mono_neg_l.
-
-Theorem Zmul_le_mono_neg_r : forall n m p : Z, p < 0 -> (n <= m <-> m * p <= n * p).
-Proof NZmul_le_mono_neg_r.
-
-Theorem Zmul_lt_mono_nonneg :
- forall n m p q : Z, 0 <= n -> n < m -> 0 <= p -> p < q -> n * p < m * q.
-Proof NZmul_lt_mono_nonneg.
-
-Theorem Zmul_lt_mono_nonpos :
- forall n m p q : Z, m <= 0 -> n < m -> q <= 0 -> p < q -> m * q < n * p.
+Theorem mul_lt_mono_nonpos :
+ forall n m p q, m <= 0 -> n < m -> q <= 0 -> p < q -> m * q < n * p.
Proof.
intros n m p q H1 H2 H3 H4.
-apply Zle_lt_trans with (m * p).
-apply Zmul_le_mono_nonpos_l; [assumption | now apply Zlt_le_incl].
-apply -> Zmul_lt_mono_neg_r; [assumption | now apply Zlt_le_trans with q].
+apply le_lt_trans with (m * p).
+apply mul_le_mono_nonpos_l; [assumption | now apply lt_le_incl].
+apply -> mul_lt_mono_neg_r; [assumption | now apply lt_le_trans with q].
Qed.
-Theorem Zmul_le_mono_nonneg :
- forall n m p q : Z, 0 <= n -> n <= m -> 0 <= p -> p <= q -> n * p <= m * q.
-Proof NZmul_le_mono_nonneg.
-
-Theorem Zmul_le_mono_nonpos :
- forall n m p q : Z, m <= 0 -> n <= m -> q <= 0 -> p <= q -> m * q <= n * p.
+Theorem mul_le_mono_nonpos :
+ forall n m p q, m <= 0 -> n <= m -> q <= 0 -> p <= q -> m * q <= n * p.
Proof.
intros n m p q H1 H2 H3 H4.
-apply Zle_trans with (m * p).
-now apply Zmul_le_mono_nonpos_l.
-apply Zmul_le_mono_nonpos_r; [now apply Zle_trans with q | assumption].
-Qed.
-
-Theorem Zmul_pos_pos : forall n m : Z, 0 < n -> 0 < m -> 0 < n * m.
-Proof NZmul_pos_pos.
-
-Theorem Zmul_neg_neg : forall n m : Z, n < 0 -> m < 0 -> 0 < n * m.
-Proof NZmul_neg_neg.
-
-Theorem Zmul_pos_neg : forall n m : Z, 0 < n -> m < 0 -> n * m < 0.
-Proof NZmul_pos_neg.
-
-Theorem Zmul_neg_pos : forall n m : Z, n < 0 -> 0 < m -> n * m < 0.
-Proof NZmul_neg_pos.
-
-Theorem Zmul_nonneg_nonneg : forall n m : Z, 0 <= n -> 0 <= m -> 0 <= n * m.
-Proof.
-intros n m H1 H2.
-rewrite <- (Zmul_0_l m). now apply Zmul_le_mono_nonneg_r.
+apply le_trans with (m * p).
+now apply mul_le_mono_nonpos_l.
+apply mul_le_mono_nonpos_r; [now apply le_trans with q | assumption].
Qed.
-Theorem Zmul_nonpos_nonpos : forall n m : Z, n <= 0 -> m <= 0 -> 0 <= n * m.
+Theorem mul_nonpos_nonpos : forall n m, n <= 0 -> m <= 0 -> 0 <= n * m.
Proof.
intros n m H1 H2.
-rewrite <- (Zmul_0_l m). now apply Zmul_le_mono_nonpos_r.
+rewrite <- (mul_0_l m). now apply mul_le_mono_nonpos_r.
Qed.
-Theorem Zmul_nonneg_nonpos : forall n m : Z, 0 <= n -> m <= 0 -> n * m <= 0.
+Theorem mul_nonneg_nonpos : forall n m, 0 <= n -> m <= 0 -> n * m <= 0.
Proof.
intros n m H1 H2.
-rewrite <- (Zmul_0_l m). now apply Zmul_le_mono_nonpos_r.
+rewrite <- (mul_0_l m). now apply mul_le_mono_nonpos_r.
Qed.
-Theorem Zmul_nonpos_nonneg : forall n m : Z, n <= 0 -> 0 <= m -> n * m <= 0.
+Theorem mul_nonpos_nonneg : forall n m, n <= 0 -> 0 <= m -> n * m <= 0.
Proof.
-intros; rewrite Zmul_comm; now apply Zmul_nonneg_nonpos.
+intros; rewrite mul_comm; now apply mul_nonneg_nonpos.
Qed.
-Theorem Zlt_1_mul_pos : forall n m : Z, 1 < n -> 0 < m -> 1 < n * m.
-Proof NZlt_1_mul_pos.
-
-Theorem Zeq_mul_0 : forall n m : Z, n * m == 0 <-> n == 0 \/ m == 0.
-Proof NZeq_mul_0.
-
-Theorem Zneq_mul_0 : forall n m : Z, n ~= 0 /\ m ~= 0 <-> n * m ~= 0.
-Proof NZneq_mul_0.
-
-Theorem Zeq_square_0 : forall n : Z, n * n == 0 <-> n == 0.
-Proof NZeq_square_0.
+Notation mul_pos := lt_0_mul (only parsing).
-Theorem Zeq_mul_0_l : forall n m : Z, n * m == 0 -> m ~= 0 -> n == 0.
-Proof NZeq_mul_0_l.
-
-Theorem Zeq_mul_0_r : forall n m : Z, n * m == 0 -> n ~= 0 -> m == 0.
-Proof NZeq_mul_0_r.
-
-Theorem Zlt_0_mul : forall n m : Z, 0 < n * m <-> 0 < n /\ 0 < m \/ m < 0 /\ n < 0.
-Proof NZlt_0_mul.
-
-Notation Zmul_pos := Zlt_0_mul (only parsing).
-
-Theorem Zlt_mul_0 :
- forall n m : Z, n * m < 0 <-> n < 0 /\ m > 0 \/ n > 0 /\ m < 0.
+Theorem lt_mul_0 :
+ forall n m, n * m < 0 <-> n < 0 /\ m > 0 \/ n > 0 /\ m < 0.
Proof.
intros n m; split; [intro H | intros [[H1 H2] | [H1 H2]]].
-destruct (Zlt_trichotomy n 0) as [H1 | [H1 | H1]];
-[| rewrite H1 in H; rewrite Zmul_0_l in H; false_hyp H Zlt_irrefl |];
-(destruct (Zlt_trichotomy m 0) as [H2 | [H2 | H2]];
-[| rewrite H2 in H; rewrite Zmul_0_r in H; false_hyp H Zlt_irrefl |]);
+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]];
+[| rewrite H2 in H; rewrite mul_0_r in H; false_hyp H lt_irrefl |]);
try (left; now split); try (right; now split).
-assert (H3 : n * m > 0) by now apply Zmul_neg_neg.
-elimtype False; now apply (Zlt_asymm (n * m) 0).
-assert (H3 : n * m > 0) by now apply Zmul_pos_pos.
-elimtype False; now apply (Zlt_asymm (n * m) 0).
-now apply Zmul_neg_pos. now apply Zmul_pos_neg.
+assert (H3 : n * m > 0) by now apply mul_neg_neg.
+exfalso; now apply (lt_asymm (n * m) 0).
+assert (H3 : n * m > 0) by now apply mul_pos_pos.
+exfalso; now apply (lt_asymm (n * m) 0).
+now apply mul_neg_pos. now apply mul_pos_neg.
Qed.
-Notation Zmul_neg := Zlt_mul_0 (only parsing).
+Notation mul_neg := lt_mul_0 (only parsing).
-Theorem Zle_0_mul :
- forall n m : Z, 0 <= n * m -> 0 <= n /\ 0 <= m \/ n <= 0 /\ m <= 0.
+Theorem le_0_mul :
+ forall n m, 0 <= n * m -> 0 <= n /\ 0 <= m \/ n <= 0 /\ m <= 0.
Proof.
-assert (R : forall n : Z, 0 == n <-> n == 0) by (intros; split; apply Zeq_sym).
-intros n m. repeat rewrite Zlt_eq_cases. repeat rewrite R.
-rewrite Zlt_0_mul, Zeq_mul_0.
-pose proof (Zlt_trichotomy n 0); pose proof (Zlt_trichotomy m 0). tauto.
+assert (R : forall n, 0 == n <-> n == 0) by (intros; split; apply eq_sym).
+intros n m. repeat rewrite lt_eq_cases. repeat rewrite R.
+rewrite lt_0_mul, eq_mul_0.
+pose proof (lt_trichotomy n 0); pose proof (lt_trichotomy m 0). tauto.
Qed.
-Notation Zmul_nonneg := Zle_0_mul (only parsing).
+Notation mul_nonneg := le_0_mul (only parsing).
-Theorem Zle_mul_0 :
- forall n m : Z, n * m <= 0 -> 0 <= n /\ m <= 0 \/ n <= 0 /\ 0 <= m.
+Theorem le_mul_0 :
+ forall n m, n * m <= 0 -> 0 <= n /\ m <= 0 \/ n <= 0 /\ 0 <= m.
Proof.
-assert (R : forall n : Z, 0 == n <-> n == 0) by (intros; split; apply Zeq_sym).
-intros n m. repeat rewrite Zlt_eq_cases. repeat rewrite R.
-rewrite Zlt_mul_0, Zeq_mul_0.
-pose proof (Zlt_trichotomy n 0); pose proof (Zlt_trichotomy m 0). tauto.
+assert (R : forall n, 0 == n <-> n == 0) by (intros; split; apply eq_sym).
+intros n m. repeat rewrite lt_eq_cases. repeat rewrite R.
+rewrite lt_mul_0, eq_mul_0.
+pose proof (lt_trichotomy n 0); pose proof (lt_trichotomy m 0). tauto.
Qed.
-Notation Zmul_nonpos := Zle_mul_0 (only parsing).
+Notation mul_nonpos := le_mul_0 (only parsing).
-Theorem Zle_0_square : forall n : Z, 0 <= n * n.
+Theorem le_0_square : forall n, 0 <= n * n.
Proof.
-intro n; destruct (Zneg_nonneg_cases n).
-apply Zlt_le_incl; now apply Zmul_neg_neg.
-now apply Zmul_nonneg_nonneg.
+intro n; destruct (neg_nonneg_cases n).
+apply lt_le_incl; now apply mul_neg_neg.
+now apply mul_nonneg_nonneg.
Qed.
-Notation Zsquare_nonneg := Zle_0_square (only parsing).
+Notation square_nonneg := le_0_square (only parsing).
-Theorem Znlt_square_0 : forall n : Z, ~ n * n < 0.
+Theorem nlt_square_0 : forall n, ~ n * n < 0.
Proof.
-intros n H. apply -> Zlt_nge in H. apply H. apply Zsquare_nonneg.
+intros n H. apply -> lt_nge in H. apply H. apply square_nonneg.
Qed.
-Theorem Zsquare_lt_mono_nonneg : forall n m : Z, 0 <= n -> n < m -> n * n < m * m.
-Proof NZsquare_lt_mono_nonneg.
-
-Theorem Zsquare_lt_mono_nonpos : forall n m : Z, n <= 0 -> m < n -> n * n < m * m.
+Theorem square_lt_mono_nonpos : forall n m, n <= 0 -> m < n -> n * n < m * m.
Proof.
-intros n m H1 H2. now apply Zmul_lt_mono_nonpos.
+intros n m H1 H2. now apply mul_lt_mono_nonpos.
Qed.
-Theorem Zsquare_le_mono_nonneg : forall n m : Z, 0 <= n -> n <= m -> n * n <= m * m.
-Proof NZsquare_le_mono_nonneg.
-
-Theorem Zsquare_le_mono_nonpos : forall n m : Z, n <= 0 -> m <= n -> n * n <= m * m.
+Theorem square_le_mono_nonpos : forall n m, n <= 0 -> m <= n -> n * n <= m * m.
Proof.
-intros n m H1 H2. now apply Zmul_le_mono_nonpos.
+intros n m H1 H2. now apply mul_le_mono_nonpos.
Qed.
-Theorem Zsquare_lt_simpl_nonneg : forall n m : Z, 0 <= m -> n * n < m * m -> n < m.
-Proof NZsquare_lt_simpl_nonneg.
-
-Theorem Zsquare_le_simpl_nonneg : forall n m : Z, 0 <= m -> n * n <= m * m -> n <= m.
-Proof NZsquare_le_simpl_nonneg.
-
-Theorem Zsquare_lt_simpl_nonpos : forall n m : Z, m <= 0 -> n * n < m * m -> m < n.
+Theorem square_lt_simpl_nonpos : forall n m, m <= 0 -> n * n < m * m -> m < n.
Proof.
-intros n m H1 H2. destruct (Zle_gt_cases n 0).
-destruct (NZlt_ge_cases m n).
-assumption. assert (F : m * m <= n * n) by now apply Zsquare_le_mono_nonpos.
-apply -> NZle_ngt in F. false_hyp H2 F.
-now apply Zle_lt_trans with 0.
+intros n m H1 H2. destruct (le_gt_cases n 0).
+destruct (lt_ge_cases m n).
+assumption. assert (F : m * m <= n * n) by now apply square_le_mono_nonpos.
+apply -> le_ngt in F. false_hyp H2 F.
+now apply le_lt_trans with 0.
Qed.
-Theorem Zsquare_le_simpl_nonpos : forall n m : NZ, m <= 0 -> n * n <= m * m -> m <= n.
+Theorem square_le_simpl_nonpos : forall n m, m <= 0 -> n * n <= m * m -> m <= n.
Proof.
-intros n m H1 H2. destruct (NZle_gt_cases n 0).
-destruct (NZle_gt_cases m n).
-assumption. assert (F : m * m < n * n) by now apply Zsquare_lt_mono_nonpos.
-apply -> NZlt_nge in F. false_hyp H2 F.
-apply Zlt_le_incl; now apply NZle_lt_trans with 0.
+intros n m H1 H2. destruct (le_gt_cases n 0).
+destruct (le_gt_cases m n).
+assumption. assert (F : m * m < n * n) by now apply square_lt_mono_nonpos.
+apply -> lt_nge in F. false_hyp H2 F.
+apply lt_le_incl; now apply le_lt_trans with 0.
Qed.
-Theorem Zmul_2_mono_l : forall n m : Z, n < m -> 1 + (1 + 1) * n < (1 + 1) * m.
-Proof NZmul_2_mono_l.
-
-Theorem Zlt_1_mul_neg : forall n m : Z, n < -1 -> m < 0 -> 1 < n * m.
+Theorem lt_1_mul_neg : forall n m, n < -1 -> m < 0 -> 1 < n * m.
Proof.
-intros n m H1 H2. apply -> (NZmul_lt_mono_neg_r m) in H1.
-apply <- Zopp_pos_neg in H2. rewrite Zmul_opp_l, Zmul_1_l in H1.
-now apply Zlt_1_l with (- m).
+intros n m H1 H2. apply -> (mul_lt_mono_neg_r m) in H1.
+apply <- opp_pos_neg in H2. rewrite mul_opp_l, mul_1_l in H1.
+now apply lt_1_l with (- m).
assumption.
Qed.
-Theorem Zlt_mul_n1_neg : forall n m : Z, 1 < n -> m < 0 -> n * m < -1.
+Theorem lt_mul_n1_neg : forall n m, 1 < n -> m < 0 -> n * m < -1.
Proof.
-intros n m H1 H2. apply -> (NZmul_lt_mono_neg_r m) in H1.
-rewrite Zmul_1_l in H1. now apply Zlt_n1_r with m.
+intros n m H1 H2. apply -> (mul_lt_mono_neg_r m) in H1.
+rewrite mul_1_l in H1. now apply lt_n1_r with m.
assumption.
Qed.
-Theorem Zlt_mul_n1_pos : forall n m : Z, n < -1 -> 0 < m -> n * m < -1.
+Theorem lt_mul_n1_pos : forall n m, n < -1 -> 0 < m -> n * m < -1.
Proof.
-intros n m H1 H2. apply -> (NZmul_lt_mono_pos_r m) in H1.
-rewrite Zmul_opp_l, Zmul_1_l in H1.
-apply <- Zopp_neg_pos in H2. now apply Zlt_n1_r with (- m).
+intros n m H1 H2. apply -> (mul_lt_mono_pos_r m) in H1.
+rewrite mul_opp_l, mul_1_l in H1.
+apply <- opp_neg_pos in H2. now apply lt_n1_r with (- m).
assumption.
Qed.
-Theorem Zlt_1_mul_l : forall n m : Z, 1 < n -> n * m < -1 \/ n * m == 0 \/ 1 < n * m.
+Theorem lt_1_mul_l : forall n m, 1 < n ->
+ n * m < -1 \/ n * m == 0 \/ 1 < n * m.
Proof.
-intros n m H; destruct (Zlt_trichotomy m 0) as [H1 | [H1 | H1]].
-left. now apply Zlt_mul_n1_neg.
-right; left; now rewrite H1, Zmul_0_r.
-right; right; now apply Zlt_1_mul_pos.
+intros n m H; destruct (lt_trichotomy m 0) as [H1 | [H1 | H1]].
+left. now apply lt_mul_n1_neg.
+right; left; now rewrite H1, mul_0_r.
+right; right; now apply lt_1_mul_pos.
Qed.
-Theorem Zlt_n1_mul_r : forall n m : Z, n < -1 -> n * m < -1 \/ n * m == 0 \/ 1 < n * m.
+Theorem lt_n1_mul_r : forall n m, n < -1 ->
+ n * m < -1 \/ n * m == 0 \/ 1 < n * m.
Proof.
-intros n m H; destruct (Zlt_trichotomy m 0) as [H1 | [H1 | H1]].
-right; right. now apply Zlt_1_mul_neg.
-right; left; now rewrite H1, Zmul_0_r.
-left. now apply Zlt_mul_n1_pos.
+intros n m H; destruct (lt_trichotomy m 0) as [H1 | [H1 | H1]].
+right; right. now apply lt_1_mul_neg.
+right; left; now rewrite H1, mul_0_r.
+left. now apply lt_mul_n1_pos.
Qed.
-Theorem Zeq_mul_1 : forall n m : Z, n * m == 1 -> n == 1 \/ n == -1.
+Theorem eq_mul_1 : forall n m, n * m == 1 -> n == 1 \/ n == -1.
Proof.
assert (F : ~ 1 < -1).
intro H.
-assert (H1 : -1 < 0). apply <- Zopp_neg_pos. apply Zlt_succ_diag_r.
-assert (H2 : 1 < 0) by now apply Zlt_trans with (-1). false_hyp H2 Znlt_succ_diag_l.
-Z0_pos_neg n.
-intros m H; rewrite Zmul_0_l in H; false_hyp H Zneq_succ_diag_r.
-intros n H; split; apply <- Zle_succ_l in H; le_elim H.
-intros m H1; apply (Zlt_1_mul_l n m) in H.
+assert (H1 : -1 < 0). apply <- opp_neg_pos. apply lt_succ_diag_r.
+assert (H2 : 1 < 0) by now apply lt_trans with (-1).
+false_hyp H2 nlt_succ_diag_l.
+zero_pos_neg n.
+intros m H; rewrite mul_0_l in H; false_hyp H neq_succ_diag_r.
+intros n H; split; apply <- le_succ_l in H; le_elim H.
+intros m H1; apply (lt_1_mul_l n m) in H.
rewrite H1 in H; destruct H as [H | [H | H]].
-false_hyp H F. false_hyp H Zneq_succ_diag_l. false_hyp H Zlt_irrefl.
+false_hyp H F. false_hyp H neq_succ_diag_l. false_hyp H lt_irrefl.
intros; now left.
-intros m H1; apply (Zlt_1_mul_l n m) in H. rewrite Zmul_opp_l in H1;
-apply -> Zeq_opp_l in H1. rewrite H1 in H; destruct H as [H | [H | H]].
-false_hyp H Zlt_irrefl. apply -> Zeq_opp_l in H. rewrite Zopp_0 in H.
-false_hyp H Zneq_succ_diag_l. false_hyp H F.
-intros; right; symmetry; now apply Zopp_wd.
+intros m H1; apply (lt_1_mul_l n m) in H. rewrite mul_opp_l in H1;
+apply -> eq_opp_l in H1. rewrite H1 in H; destruct H as [H | [H | H]].
+false_hyp H lt_irrefl. apply -> eq_opp_l in H. rewrite opp_0 in H.
+false_hyp H neq_succ_diag_l. false_hyp H F.
+intros; right; symmetry; now apply opp_wd.
Qed.
-Theorem Zlt_mul_diag_l : forall n m : Z, n < 0 -> (1 < m <-> n * m < n).
+Theorem lt_mul_diag_l : forall n m, n < 0 -> (1 < m <-> n * m < n).
Proof.
-intros n m H. stepr (n * m < n * 1) by now rewrite Zmul_1_r.
-now apply Zmul_lt_mono_neg_l.
+intros n m H. stepr (n * m < n * 1) by now rewrite mul_1_r.
+now apply mul_lt_mono_neg_l.
Qed.
-Theorem Zlt_mul_diag_r : forall n m : Z, 0 < n -> (1 < m <-> n < n * m).
+Theorem lt_mul_diag_r : forall n m, 0 < n -> (1 < m <-> n < n * m).
Proof.
-intros n m H. stepr (n * 1 < n * m) by now rewrite Zmul_1_r.
-now apply Zmul_lt_mono_pos_l.
+intros n m H. stepr (n * 1 < n * m) by now rewrite mul_1_r.
+now apply mul_lt_mono_pos_l.
Qed.
-Theorem Zle_mul_diag_l : forall n m : Z, n < 0 -> (1 <= m <-> n * m <= n).
+Theorem le_mul_diag_l : forall n m, n < 0 -> (1 <= m <-> n * m <= n).
Proof.
-intros n m H. stepr (n * m <= n * 1) by now rewrite Zmul_1_r.
-now apply Zmul_le_mono_neg_l.
+intros n m H. stepr (n * m <= n * 1) by now rewrite mul_1_r.
+now apply mul_le_mono_neg_l.
Qed.
-Theorem Zle_mul_diag_r : forall n m : Z, 0 < n -> (1 <= m <-> n <= n * m).
+Theorem le_mul_diag_r : forall n m, 0 < n -> (1 <= m <-> n <= n * m).
Proof.
-intros n m H. stepr (n * 1 <= n * m) by now rewrite Zmul_1_r.
-now apply Zmul_le_mono_pos_l.
+intros n m H. stepr (n * 1 <= n * m) by now rewrite mul_1_r.
+now apply mul_le_mono_pos_l.
Qed.
-Theorem Zlt_mul_r : forall n m p : Z, 0 < n -> 1 < p -> n < m -> n < m * p.
+Theorem lt_mul_r : forall n m p, 0 < n -> 1 < p -> n < m -> n < m * p.
Proof.
-intros. stepl (n * 1) by now rewrite Zmul_1_r.
-apply Zmul_lt_mono_nonneg.
-now apply Zlt_le_incl. assumption. apply Zle_0_1. assumption.
+intros. stepl (n * 1) by now rewrite mul_1_r.
+apply mul_lt_mono_nonneg.
+now apply lt_le_incl. assumption. apply le_0_1. assumption.
Qed.
End ZMulOrderPropFunct.
diff --git a/theories/Numbers/Integer/Abstract/ZProperties.v b/theories/Numbers/Integer/Abstract/ZProperties.v
new file mode 100644
index 00000000..dc46edda
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZProperties.v
@@ -0,0 +1,24 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+Require Export ZAxioms ZMulOrder ZSgnAbs.
+
+(** This functor summarizes all known facts about Z.
+ For the moment it is only an alias to [ZMulOrderPropFunct], which
+ subsumes all others, plus properties of [sgn] and [abs].
+*)
+
+Module Type ZPropSig (Z:ZAxiomsExtSig) :=
+ ZMulOrderPropFunct Z <+ ZSgnAbsPropSig Z.
+
+Module ZPropFunct (Z:ZAxiomsExtSig) <: ZPropSig Z.
+ Include ZPropSig Z.
+End ZPropFunct.
+
diff --git a/theories/Numbers/Integer/Abstract/ZSgnAbs.v b/theories/Numbers/Integer/Abstract/ZSgnAbs.v
new file mode 100644
index 00000000..8b191613
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZSgnAbs.v
@@ -0,0 +1,348 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Export ZMulOrder.
+
+(** An axiomatization of [abs]. *)
+
+Module Type HasAbs(Import Z : ZAxiomsSig').
+ Parameter Inline abs : t -> t.
+ Axiom abs_eq : forall n, 0<=n -> abs n == n.
+ Axiom abs_neq : forall n, n<=0 -> abs n == -n.
+End HasAbs.
+
+(** Since we already have [max], we could have defined [abs]. *)
+
+Module GenericAbs (Import Z : ZAxiomsSig')
+ (Import ZP : ZMulOrderPropFunct Z) <: HasAbs Z.
+ Definition abs n := max n (-n).
+ Lemma abs_eq : forall n, 0<=n -> abs n == n.
+ Proof.
+ intros. unfold abs. apply max_l.
+ apply le_trans with 0; auto.
+ rewrite opp_nonpos_nonneg; auto.
+ Qed.
+ Lemma abs_neq : forall n, n<=0 -> abs n == -n.
+ Proof.
+ intros. unfold abs. apply max_r.
+ apply le_trans with 0; auto.
+ rewrite opp_nonneg_nonpos; auto.
+ Qed.
+End GenericAbs.
+
+(** An Axiomatization of [sgn]. *)
+
+Module Type HasSgn (Import Z : ZAxiomsSig').
+ Parameter Inline sgn : t -> t.
+ Axiom sgn_null : forall n, n==0 -> sgn n == 0.
+ Axiom sgn_pos : forall n, 0<n -> sgn n == 1.
+ Axiom sgn_neg : forall n, n<0 -> sgn n == -(1).
+End HasSgn.
+
+(** We can deduce a [sgn] function from a [compare] function *)
+
+Module Type ZDecAxiomsSig := ZAxiomsSig <+ HasCompare.
+Module Type ZDecAxiomsSig' := ZAxiomsSig' <+ HasCompare.
+
+Module Type GenericSgn (Import Z : ZDecAxiomsSig')
+ (Import ZP : ZMulOrderPropFunct Z) <: HasSgn Z.
+ Definition sgn n :=
+ match compare 0 n with Eq => 0 | Lt => 1 | Gt => -(1) end.
+ Lemma sgn_null : forall n, n==0 -> sgn n == 0.
+ Proof. unfold sgn; intros. destruct (compare_spec 0 n); order. Qed.
+ Lemma sgn_pos : forall n, 0<n -> sgn n == 1.
+ Proof. unfold sgn; intros. destruct (compare_spec 0 n); order. Qed.
+ Lemma sgn_neg : forall n, n<0 -> sgn n == -(1).
+ Proof. unfold sgn; intros. destruct (compare_spec 0 n); order. Qed.
+End GenericSgn.
+
+Module Type ZAxiomsExtSig := ZAxiomsSig <+ HasAbs <+ HasSgn.
+Module Type ZAxiomsExtSig' := ZAxiomsSig' <+ HasAbs <+ HasSgn.
+
+Module Type ZSgnAbsPropSig (Import Z : ZAxiomsExtSig')
+ (Import ZP : ZMulOrderPropFunct Z).
+
+Ltac destruct_max n :=
+ destruct (le_ge_cases 0 n);
+ [rewrite (abs_eq n) by auto | rewrite (abs_neq n) by auto].
+
+Instance abs_wd : Proper (eq==>eq) abs.
+Proof.
+ intros x y EQ. destruct_max x.
+ rewrite abs_eq; trivial. now rewrite <- EQ.
+ rewrite abs_neq; try order. now rewrite opp_inj_wd.
+Qed.
+
+Lemma abs_max : forall n, abs n == max n (-n).
+Proof.
+ intros n. destruct_max n.
+ rewrite max_l; auto with relations.
+ apply le_trans with 0; auto.
+ rewrite opp_nonpos_nonneg; auto.
+ rewrite max_r; auto with relations.
+ apply le_trans with 0; auto.
+ rewrite opp_nonneg_nonpos; auto.
+Qed.
+
+Lemma abs_neq' : forall n, 0<=-n -> abs n == -n.
+Proof.
+ intros. apply abs_neq. now rewrite <- opp_nonneg_nonpos.
+Qed.
+
+Lemma abs_nonneg : forall n, 0 <= abs n.
+Proof.
+ intros n. destruct_max n; auto.
+ now rewrite opp_nonneg_nonpos.
+Qed.
+
+Lemma abs_eq_iff : forall n, abs n == n <-> 0<=n.
+Proof.
+ split; try apply abs_eq. intros EQ.
+ rewrite <- EQ. apply abs_nonneg.
+Qed.
+
+Lemma abs_neq_iff : forall n, abs n == -n <-> n<=0.
+Proof.
+ split; try apply abs_neq. intros EQ.
+ rewrite <- opp_nonneg_nonpos, <- EQ. apply abs_nonneg.
+Qed.
+
+Lemma abs_opp : forall n, abs (-n) == abs n.
+Proof.
+ intros. destruct_max n.
+ rewrite (abs_neq (-n)), opp_involutive. reflexivity.
+ now rewrite opp_nonpos_nonneg.
+ rewrite (abs_eq (-n)). reflexivity.
+ now rewrite opp_nonneg_nonpos.
+Qed.
+
+Lemma abs_0 : abs 0 == 0.
+Proof.
+ apply abs_eq. apply le_refl.
+Qed.
+
+Lemma abs_0_iff : forall n, abs n == 0 <-> n==0.
+Proof.
+ split. destruct_max n; auto.
+ now rewrite eq_opp_l, opp_0.
+ intros EQ; rewrite EQ. rewrite abs_eq; auto using eq_refl, le_refl.
+Qed.
+
+Lemma abs_pos : forall n, 0 < abs n <-> n~=0.
+Proof.
+ intros. rewrite <- abs_0_iff. split; [intros LT| intros NEQ].
+ intro EQ. rewrite EQ in LT. now elim (lt_irrefl 0).
+ assert (LE : 0 <= abs n) by apply abs_nonneg.
+ rewrite lt_eq_cases in LE; destruct LE; auto.
+ elim NEQ; auto with relations.
+Qed.
+
+Lemma abs_eq_or_opp : forall n, abs n == n \/ abs n == -n.
+Proof.
+ intros. destruct_max n; auto with relations.
+Qed.
+
+Lemma abs_or_opp_abs : forall n, n == abs n \/ n == - abs n.
+Proof.
+ intros. destruct_max n; rewrite ? opp_involutive; auto with relations.
+Qed.
+
+Lemma abs_involutive : forall n, abs (abs n) == abs n.
+Proof.
+ intros. apply abs_eq. apply abs_nonneg.
+Qed.
+
+Lemma abs_spec : forall n,
+ (0 <= n /\ abs n == n) \/ (n < 0 /\ abs n == -n).
+Proof.
+ intros. destruct (le_gt_cases 0 n).
+ left; split; auto. now apply abs_eq.
+ right; split; auto. apply abs_neq. now apply lt_le_incl.
+Qed.
+
+Lemma abs_case_strong :
+ forall (P:t->Prop) n, Proper (eq==>iff) P ->
+ (0<=n -> P n) -> (n<=0 -> P (-n)) -> P (abs n).
+Proof.
+ intros. destruct_max n; auto.
+Qed.
+
+Lemma abs_case : forall (P:t->Prop) n, Proper (eq==>iff) P ->
+ P n -> P (-n) -> P (abs n).
+Proof. intros. now apply abs_case_strong. Qed.
+
+Lemma abs_eq_cases : forall n m, abs n == abs m -> n == m \/ n == - m.
+Proof.
+ intros n m EQ. destruct (abs_or_opp_abs n) as [EQn|EQn].
+ rewrite EQn, EQ. apply abs_eq_or_opp.
+ rewrite EQn, EQ, opp_inj_wd, eq_opp_l, or_comm. apply abs_eq_or_opp.
+Qed.
+
+(** Triangular inequality *)
+
+Lemma abs_triangle : forall n m, abs (n + m) <= abs n + abs m.
+Proof.
+ intros. destruct_max n; destruct_max m.
+ rewrite abs_eq. apply le_refl. now apply add_nonneg_nonneg.
+ destruct_max (n+m); try rewrite opp_add_distr;
+ apply add_le_mono_l || apply add_le_mono_r.
+ apply le_trans with 0; auto. now rewrite opp_nonneg_nonpos.
+ apply le_trans with 0; auto. now rewrite opp_nonpos_nonneg.
+ destruct_max (n+m); try rewrite opp_add_distr;
+ apply add_le_mono_l || apply add_le_mono_r.
+ apply le_trans with 0; auto. now rewrite opp_nonneg_nonpos.
+ apply le_trans with 0; auto. now rewrite opp_nonpos_nonneg.
+ rewrite abs_neq, opp_add_distr. apply le_refl.
+ now apply add_nonpos_nonpos.
+Qed.
+
+Lemma abs_sub_triangle : forall n m, abs n - abs m <= abs (n-m).
+Proof.
+ intros.
+ rewrite le_sub_le_add_l, add_comm.
+ rewrite <- (sub_simpl_r n m) at 1.
+ apply abs_triangle.
+Qed.
+
+(** Absolute value and multiplication *)
+
+Lemma abs_mul : forall n m, abs (n * m) == abs n * abs m.
+Proof.
+ assert (H : forall n m, 0<=n -> abs (n*m) == n * abs m).
+ intros. destruct_max m.
+ rewrite abs_eq. apply eq_refl. now apply mul_nonneg_nonneg.
+ rewrite abs_neq, mul_opp_r. reflexivity. now apply mul_nonneg_nonpos .
+ intros. destruct_max n. now apply H.
+ rewrite <- mul_opp_opp, H, abs_opp. reflexivity.
+ now apply opp_nonneg_nonpos.
+Qed.
+
+Lemma abs_square : forall n, abs n * abs n == n * n.
+Proof.
+ intros. rewrite <- abs_mul. apply abs_eq. apply le_0_square.
+Qed.
+
+(** Some results about the sign function. *)
+
+Ltac destruct_sgn n :=
+ let LT := fresh "LT" in
+ let EQ := fresh "EQ" in
+ let GT := fresh "GT" in
+ destruct (lt_trichotomy 0 n) as [LT|[EQ|GT]];
+ [rewrite (sgn_pos n) by auto|
+ rewrite (sgn_null n) by auto with relations|
+ rewrite (sgn_neg n) by auto].
+
+Instance sgn_wd : Proper (eq==>eq) sgn.
+Proof.
+ intros x y Hxy. destruct_sgn x.
+ rewrite sgn_pos; auto with relations. rewrite <- Hxy; auto.
+ rewrite sgn_null; auto with relations. rewrite <- Hxy; auto with relations.
+ rewrite sgn_neg; auto with relations. rewrite <- Hxy; auto.
+Qed.
+
+Lemma sgn_spec : forall n,
+ 0 < n /\ sgn n == 1 \/
+ 0 == n /\ sgn n == 0 \/
+ 0 > n /\ sgn n == -(1).
+Proof.
+ intros n.
+ destruct_sgn n; [left|right;left|right;right]; auto with relations.
+Qed.
+
+Lemma sgn_0 : sgn 0 == 0.
+Proof.
+ now apply sgn_null.
+Qed.
+
+Lemma sgn_pos_iff : forall n, sgn n == 1 <-> 0<n.
+Proof.
+ split; try apply sgn_pos. destruct_sgn n; auto.
+ intros. elim (lt_neq 0 1); auto. apply lt_0_1.
+ intros. elim (lt_neq (-(1)) 1); auto.
+ apply lt_trans with 0. rewrite opp_neg_pos. apply lt_0_1. apply lt_0_1.
+Qed.
+
+Lemma sgn_null_iff : forall n, sgn n == 0 <-> n==0.
+Proof.
+ split; try apply sgn_null. destruct_sgn n; auto with relations.
+ intros. elim (lt_neq 0 1); auto with relations. apply lt_0_1.
+ intros. elim (lt_neq (-(1)) 0); auto.
+ rewrite opp_neg_pos. apply lt_0_1.
+Qed.
+
+Lemma sgn_neg_iff : forall n, sgn n == -(1) <-> n<0.
+Proof.
+ split; try apply sgn_neg. destruct_sgn n; auto with relations.
+ intros. elim (lt_neq (-(1)) 1); auto with relations.
+ apply lt_trans with 0. rewrite opp_neg_pos. apply lt_0_1. apply lt_0_1.
+ intros. elim (lt_neq (-(1)) 0); auto with relations.
+ rewrite opp_neg_pos. apply lt_0_1.
+Qed.
+
+Lemma sgn_opp : forall n, sgn (-n) == - sgn n.
+Proof.
+ intros. destruct_sgn n.
+ apply sgn_neg. now rewrite opp_neg_pos.
+ setoid_replace n with 0 by auto with relations.
+ rewrite opp_0. apply sgn_0.
+ rewrite opp_involutive. apply sgn_pos. now rewrite opp_pos_neg.
+Qed.
+
+Lemma sgn_nonneg : forall n, 0 <= sgn n <-> 0 <= n.
+Proof.
+ split.
+ destruct_sgn n; intros.
+ now apply lt_le_incl.
+ order.
+ elim (lt_irrefl 0). apply lt_le_trans with 1; auto using lt_0_1.
+ now rewrite <- opp_nonneg_nonpos.
+ rewrite lt_eq_cases; destruct 1.
+ rewrite sgn_pos by auto. apply lt_le_incl, lt_0_1.
+ rewrite sgn_null by auto with relations. apply le_refl.
+Qed.
+
+Lemma sgn_nonpos : forall n, sgn n <= 0 <-> n <= 0.
+Proof.
+ intros. rewrite <- 2 opp_nonneg_nonpos, <- sgn_opp. apply sgn_nonneg.
+Qed.
+
+Lemma sgn_mul : forall n m, sgn (n*m) == sgn n * sgn m.
+Proof.
+ intros. destruct_sgn n; nzsimpl.
+ destruct_sgn m.
+ apply sgn_pos. now apply mul_pos_pos.
+ apply sgn_null. rewrite eq_mul_0; auto with relations.
+ apply sgn_neg. now apply mul_pos_neg.
+ apply sgn_null. rewrite eq_mul_0; auto with relations.
+ destruct_sgn m; try rewrite mul_opp_opp; nzsimpl.
+ apply sgn_neg. now apply mul_neg_pos.
+ apply sgn_null. rewrite eq_mul_0; auto with relations.
+ apply sgn_pos. now apply mul_neg_neg.
+Qed.
+
+Lemma sgn_abs : forall n, n * sgn n == abs n.
+Proof.
+ intros. symmetry.
+ destruct_sgn n; try rewrite mul_opp_r; nzsimpl.
+ apply abs_eq. now apply lt_le_incl.
+ rewrite abs_0_iff; auto with relations.
+ apply abs_neq. now apply lt_le_incl.
+Qed.
+
+Lemma abs_sgn : forall n, abs n * sgn n == n.
+Proof.
+ intros.
+ destruct_sgn n; try rewrite mul_opp_r; nzsimpl; auto.
+ apply abs_eq. now apply lt_le_incl.
+ rewrite eq_opp_l. apply abs_neq. now apply lt_le_incl.
+Qed.
+
+End ZSgnAbsPropSig.
+
+
diff --git a/theories/Numbers/Integer/BigZ/BigZ.v b/theories/Numbers/Integer/BigZ/BigZ.v
index e5e950ac..4e024c02 100644
--- a/theories/Numbers/Integer/BigZ/BigZ.v
+++ b/theories/Numbers/Integer/BigZ/BigZ.v
@@ -8,20 +8,31 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: BigZ.v 11576 2008-11-10 19:13:15Z msozeau $ i*)
+(*i $Id$ i*)
Require Export BigN.
-Require Import ZMulOrder.
-Require Import ZSig.
-Require Import ZSigZAxioms.
-Require Import ZMake.
+Require Import ZProperties ZDivFloor ZSig ZSigZAxioms ZMake.
-Module BigZ <: ZType := ZMake.Make BigN.
+(** * [BigZ] : arbitrary large efficient integers.
-(** Module [BigZ] implements [ZAxiomsSig] *)
+ The following [BigZ] module regroups both the operations and
+ all the abstract properties:
-Module Export BigZAxiomsMod := ZSig_ZAxioms BigZ.
-Module Export BigZMulOrderPropMod := ZMulOrderPropFunct BigZAxiomsMod.
+ - [ZMake.Make BigN] provides the operations and basic specs w.r.t. ZArith
+ - [ZTypeIsZAxioms] shows (mainly) that these operations implement
+ the interface [ZAxioms]
+ - [ZPropSig] adds all generic properties derived from [ZAxioms]
+ - [ZDivPropFunct] provides generic properties of [div] and [mod]
+ ("Floor" variant)
+ - [MinMax*Properties] provides properties of [min] and [max]
+
+*)
+
+
+Module BigZ <: ZType <: OrderedTypeFull <: TotalOrder :=
+ ZMake.Make BigN <+ ZTypeIsZAxioms
+ <+ !ZPropSig <+ !ZDivPropFunct <+ HasEqBool2Dec
+ <+ !MinMaxLogicalProperties <+ !MinMaxDecProperties.
(** Notations about [BigZ] *)
@@ -31,26 +42,60 @@ Delimit Scope bigZ_scope with bigZ.
Bind Scope bigZ_scope with bigZ.
Bind Scope bigZ_scope with BigZ.t.
Bind Scope bigZ_scope with BigZ.t_.
-
-Notation Local "0" := BigZ.zero : bigZ_scope.
+(* Bind Scope has no retroactive effect, let's declare scopes by hand. *)
+Arguments Scope BigZ.Pos [bigN_scope].
+Arguments Scope BigZ.Neg [bigN_scope].
+Arguments Scope BigZ.to_Z [bigZ_scope].
+Arguments Scope BigZ.succ [bigZ_scope].
+Arguments Scope BigZ.pred [bigZ_scope].
+Arguments Scope BigZ.opp [bigZ_scope].
+Arguments Scope BigZ.square [bigZ_scope].
+Arguments Scope BigZ.add [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.sub [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.mul [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.div [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.eq [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.lt [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.le [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.eq [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.compare [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.min [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.max [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.eq_bool [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.power_pos [bigZ_scope positive_scope].
+Arguments Scope BigZ.power [bigZ_scope N_scope].
+Arguments Scope BigZ.sqrt [bigZ_scope].
+Arguments Scope BigZ.div_eucl [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.modulo [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.gcd [bigZ_scope bigZ_scope].
+
+Local Notation "0" := BigZ.zero : bigZ_scope.
+Local Notation "1" := BigZ.one : bigZ_scope.
Infix "+" := BigZ.add : bigZ_scope.
Infix "-" := BigZ.sub : bigZ_scope.
Notation "- x" := (BigZ.opp x) : bigZ_scope.
Infix "*" := BigZ.mul : bigZ_scope.
Infix "/" := BigZ.div : bigZ_scope.
+Infix "^" := BigZ.power : bigZ_scope.
Infix "?=" := BigZ.compare : bigZ_scope.
Infix "==" := BigZ.eq (at level 70, no associativity) : bigZ_scope.
+Notation "x != y" := (~x==y)%bigZ (at level 70, no associativity) : bigZ_scope.
Infix "<" := BigZ.lt : bigZ_scope.
Infix "<=" := BigZ.le : bigZ_scope.
Notation "x > y" := (BigZ.lt y x)(only parsing) : bigZ_scope.
Notation "x >= y" := (BigZ.le y x)(only parsing) : bigZ_scope.
+Notation "x < y < z" := (x<y /\ y<z)%bigZ : bigZ_scope.
+Notation "x < y <= z" := (x<y /\ y<=z)%bigZ : bigZ_scope.
+Notation "x <= y < z" := (x<=y /\ y<z)%bigZ : bigZ_scope.
+Notation "x <= y <= z" := (x<=y /\ y<=z)%bigZ : bigZ_scope.
Notation "[ i ]" := (BigZ.to_Z i) : bigZ_scope.
+Infix "mod" := BigZ.modulo (at level 40, no associativity) : bigN_scope.
-Open Scope bigZ_scope.
+Local Open Scope bigZ_scope.
(** Some additional results about [BigZ] *)
-Theorem spec_to_Z: forall n:bigZ,
+Theorem spec_to_Z: forall n : bigZ,
BigN.to_Z (BigZ.to_N n) = ((Zsgn [n]) * [n])%Z.
Proof.
intros n; case n; simpl; intros p;
@@ -62,7 +107,7 @@ Qed.
Theorem spec_to_N n:
([n] = Zsgn [n] * (BigN.to_Z (BigZ.to_N n)))%Z.
Proof.
-intros n; case n; simpl; intros p;
+case n; simpl; intros p;
generalize (BigN.spec_pos p); case (BigN.to_Z p); auto.
intros p1 H1; case H1; auto.
intros p1 H1; case H1; auto.
@@ -77,35 +122,97 @@ intros p1 _ H1; case H1; auto.
intros p1 H1; case H1; auto.
Qed.
-Lemma sub_opp : forall x y : bigZ, x - y == x + (- y).
+(** [BigZ] is a ring *)
+
+Lemma BigZring :
+ ring_theory 0 1 BigZ.add BigZ.mul BigZ.sub BigZ.opp BigZ.eq.
Proof.
-red; intros; zsimpl; auto.
+constructor.
+exact BigZ.add_0_l. exact BigZ.add_comm. exact BigZ.add_assoc.
+exact BigZ.mul_1_l. exact BigZ.mul_comm. exact BigZ.mul_assoc.
+exact BigZ.mul_add_distr_r.
+symmetry. apply BigZ.add_opp_r.
+exact BigZ.add_opp_diag_r.
Qed.
-Lemma add_opp : forall x : bigZ, x + (- x) == 0.
+Lemma BigZeqb_correct : forall x y, BigZ.eq_bool x y = true -> x==y.
+Proof. now apply BigZ.eqb_eq. Qed.
+
+Lemma BigZpower : power_theory 1 BigZ.mul BigZ.eq (@id N) BigZ.power.
Proof.
-red; intros; zsimpl; auto with zarith.
+constructor.
+intros. red. rewrite BigZ.spec_power. unfold id.
+destruct Zpower_theory as [EQ]. rewrite EQ.
+destruct n; simpl. reflexivity.
+induction p; simpl; intros; BigZ.zify; rewrite ?IHp; auto.
Qed.
-(** [BigZ] is a ring *)
+Lemma BigZdiv : div_theory BigZ.eq BigZ.add BigZ.mul (@id _)
+ (fun a b => if BigZ.eq_bool b 0 then (0,a) else BigZ.div_eucl a b).
+Proof.
+constructor. unfold id. intros a b.
+BigZ.zify.
+generalize (Zeq_bool_if [b] 0); destruct (Zeq_bool [b] 0).
+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').
+intros (EQ,_). injection 1. intros EQr EQq.
+BigZ.zify. rewrite EQr, EQq; auto.
+Qed.
-Lemma BigZring :
- ring_theory BigZ.zero BigZ.one BigZ.add BigZ.mul BigZ.sub BigZ.opp BigZ.eq.
+(** Detection of constants *)
+
+Ltac isBigZcst t :=
+ match t with
+ | BigZ.Pos ?t => isBigNcst t
+ | BigZ.Neg ?t => isBigNcst t
+ | BigZ.zero => constr:true
+ | BigZ.one => constr:true
+ | BigZ.minus_one => constr:true
+ | _ => constr:false
+ end.
+
+Ltac BigZcst t :=
+ match isBigZcst t with
+ | true => constr:t
+ | false => constr:NotConstant
+ end.
+
+(** Registration for the "ring" tactic *)
+
+Add Ring BigZr : BigZring
+ (decidable BigZeqb_correct,
+ constants [BigZcst],
+ power_tac BigZpower [Ncst],
+ div BigZdiv).
+
+Section TestRing.
+Let test : forall x y, 1 + x*y + x^2 + 1 == 1*1 + 1 + y*x + 1*x*x.
Proof.
-constructor.
-exact Zadd_0_l.
-exact Zadd_comm.
-exact Zadd_assoc.
-exact Zmul_1_l.
-exact Zmul_comm.
-exact Zmul_assoc.
-exact Zmul_add_distr_r.
-exact sub_opp.
-exact add_opp.
+intros. ring_simplify. reflexivity.
Qed.
+Let test' : forall x y, 1 + x*y + x^2 - 1*1 - y*x + 1*(-x)*x == 0.
+Proof.
+intros. ring_simplify. reflexivity.
+Qed.
+End TestRing.
+
+(** [BigZ] also benefits from an "order" tactic *)
+
+Ltac bigZ_order := BigZ.order.
+
+Section TestOrder.
+Let test : forall x y : bigZ, x<=y -> y<=x -> x==y.
+Proof. bigZ_order. Qed.
+End TestOrder.
-Add Ring BigZr : BigZring.
+(** We can use at least a bit of (r)omega by translating to [Z]. *)
-(** Todo: tactic translating from [BigZ] to [Z] + omega *)
+Section TestOmega.
+Let test : forall x y : bigZ, x<=y -> y<=x -> x==y.
+Proof. intros x y. BigZ.zify. omega. Qed.
+End TestOmega.
(** Todo: micromega *)
diff --git a/theories/Numbers/Integer/BigZ/ZMake.v b/theories/Numbers/Integer/BigZ/ZMake.v
index 98ad4c64..3196f11e 100644
--- a/theories/Numbers/Integer/BigZ/ZMake.v
+++ b/theories/Numbers/Integer/BigZ/ZMake.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZMake.v 11576 2008-11-10 19:13:15Z msozeau $ i*)
+(*i $Id$ i*)
Require Import ZArith.
Require Import BigNumPrelude.
@@ -17,31 +17,31 @@ Require Import ZSig.
Open Scope Z_scope.
-(** * ZMake
-
- A generic transformation from a structure of natural numbers
+(** * ZMake
+
+ A generic transformation from a structure of natural numbers
[NSig.NType] to a structure of integers [ZSig.ZType].
*)
Module Make (N:NType) <: ZType.
-
- Inductive t_ :=
+
+ Inductive t_ :=
| Pos : N.t -> t_
| Neg : N.t -> t_.
-
+
Definition t := t_.
Definition zero := Pos N.zero.
Definition one := Pos N.one.
Definition minus_one := Neg N.one.
- Definition of_Z x :=
+ Definition of_Z x :=
match x with
| Zpos x => Pos (N.of_N (Npos x))
| Z0 => zero
| Zneg x => Neg (N.of_N (Npos x))
end.
-
+
Definition to_Z x :=
match x with
| Pos nx => N.to_Z nx
@@ -49,6 +49,7 @@ Module Make (N:NType) <: ZType.
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.
@@ -85,72 +86,52 @@ Module Make (N:NType) <: ZType.
| Neg nx, Neg ny => N.compare ny nx
end.
- Definition lt n m := compare n m = Lt.
- Definition le n m := compare n m <> Gt.
- 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_compare :
+ forall x y, compare x y = Zcompare (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.
+ Qed.
- Theorem spec_compare: forall x y,
- match compare x y with
- Eq => to_Z x = to_Z y
- | Lt => to_Z x < to_Z y
- | Gt => to_Z x > to_Z y
- end.
- unfold compare, to_Z; intros x y; case x; case y; clear x y;
- intros x y; auto; generalize (N.spec_pos x) (N.spec_pos y).
- generalize (N.spec_compare y x); case N.compare; auto with zarith.
- generalize (N.spec_compare y N.zero); case N.compare;
- try rewrite N.spec_0; auto with zarith.
- generalize (N.spec_compare x N.zero); case N.compare;
- rewrite N.spec_0; auto with zarith.
- generalize (N.spec_compare x N.zero); case N.compare;
- rewrite N.spec_0; auto with zarith.
- generalize (N.spec_compare N.zero y); case N.compare;
- try rewrite N.spec_0; auto with zarith.
- generalize (N.spec_compare N.zero x); case N.compare;
- rewrite N.spec_0; auto with zarith.
- generalize (N.spec_compare N.zero x); case N.compare;
- rewrite N.spec_0; auto with zarith.
- generalize (N.spec_compare x y); case N.compare; auto with zarith.
- Qed.
-
- Definition eq_bool x y :=
+ Definition eq_bool x y :=
match compare x y with
| Eq => true
| _ => false
end.
- Theorem spec_eq_bool: forall x y,
- if eq_bool x y then to_Z x = to_Z y else to_Z x <> to_Z y.
- intros x y; unfold eq_bool;
- generalize (spec_compare x y); case compare; auto with zarith.
+ Theorem spec_eq_bool:
+ forall x y, eq_bool x y = Zeq_bool (to_Z x) (to_Z y).
+ Proof.
+ unfold eq_bool, Zeq_bool; intros; rewrite spec_compare; reflexivity.
Qed.
- Definition cmp_sign x y :=
- match x, y with
- | Pos nx, Neg ny =>
- if N.eq_bool ny N.zero then Eq else Gt
- | Neg nx, Pos ny =>
- if N.eq_bool nx N.zero then Eq else Lt
- | _, _ => Eq
- end.
+ Definition lt n m := to_Z n < to_Z m.
+ Definition le n m := to_Z n <= to_Z m.
+
+ 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).
+ Proof.
+ unfold min, Zmin. intros. rewrite spec_compare. destruct Zcompare; auto.
+ Qed.
+
+ Theorem spec_max : forall n m, to_Z (max n m) = Zmax (to_Z n) (to_Z m).
+ Proof.
+ unfold max, Zmax. intros. rewrite spec_compare. destruct Zcompare; auto.
+ Qed.
- Theorem spec_cmp_sign: forall x y,
- match cmp_sign x y with
- | Gt => 0 <= to_Z x /\ to_Z y < 0
- | Lt => to_Z x < 0 /\ 0 <= to_Z y
- | Eq => True
- end.
- Proof.
- intros [x | x] [y | y]; unfold cmp_sign; auto.
- generalize (N.spec_eq_bool y N.zero); case N.eq_bool; auto.
- rewrite N.spec_0; unfold to_Z.
- generalize (N.spec_pos x) (N.spec_pos y); auto with zarith.
- generalize (N.spec_eq_bool x N.zero); case N.eq_bool; auto.
- rewrite N.spec_0; unfold to_Z.
- generalize (N.spec_pos x) (N.spec_pos y); auto with zarith.
- Qed.
-
Definition to_N x :=
match x with
| Pos nx => nx
@@ -160,21 +141,23 @@ 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).
+ 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.
Qed.
-
- Definition opp x :=
- match x with
+
+ Definition opp x :=
+ match x with
| Pos nx => Neg nx
| Neg nx => Pos nx
end.
Theorem spec_opp: forall x, to_Z (opp x) = - to_Z x.
+ Proof.
intros x; case x; simpl; auto with zarith.
Qed.
-
+
Definition succ x :=
match x with
| Pos n => Pos (N.succ n)
@@ -186,12 +169,12 @@ Module Make (N:NType) <: ZType.
end.
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; generalize (N.spec_compare N.zero x); case N.compare;
- rewrite N.spec_0; simpl.
+ 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; auto with zarith.
+ intros HH; rewrite N.spec_pred, Zmax_r; auto with zarith.
generalize (N.spec_pos x); auto with zarith.
Qed.
@@ -212,19 +195,13 @@ Module Make (N:NType) <: ZType.
end
| Neg nx, Neg ny => Neg (N.add nx ny)
end.
-
+
Theorem spec_add: forall x y, to_Z (add x y) = to_Z x + to_Z y.
- unfold add, to_Z; intros [x | x] [y | y].
- exact (N.spec_add x y).
- unfold zero; generalize (N.spec_compare x y); case N.compare.
- rewrite N.spec_0; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
- unfold zero; generalize (N.spec_compare x y); case N.compare.
- rewrite N.spec_0; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
- intros; rewrite N.spec_add; try ring; auto with zarith.
+ 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 *.
Qed.
Definition pred x :=
@@ -238,17 +215,17 @@ Module Make (N:NType) <: ZType.
end.
Theorem spec_pred: forall x, to_Z (pred x) = to_Z x - 1.
- unfold pred, to_Z, minus_one; intros [x | x].
- generalize (N.spec_compare N.zero x); case N.compare;
- rewrite N.spec_0; try rewrite N.spec_1; auto with zarith.
- intros H; exact (N.spec_pred _ H).
- generalize (N.spec_pos x); auto with zarith.
- rewrite N.spec_succ; ring.
+ 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 *.
Qed.
Definition sub x y :=
match x, y with
- | Pos nx, Pos ny =>
+ | Pos nx, Pos ny =>
match N.compare nx ny with
| Gt => Pos (N.sub nx ny)
| Eq => zero
@@ -256,7 +233,7 @@ Module Make (N:NType) <: ZType.
end
| Pos nx, Neg ny => Pos (N.add nx ny)
| Neg nx, Pos ny => Neg (N.add nx ny)
- | Neg nx, Neg ny =>
+ | Neg nx, Neg ny =>
match N.compare nx ny with
| Gt => Neg (N.sub nx ny)
| Eq => zero
@@ -265,20 +242,14 @@ Module Make (N:NType) <: ZType.
end.
Theorem spec_sub: forall x y, to_Z (sub x y) = to_Z x - to_Z y.
- unfold sub, to_Z; intros [x | x] [y | y].
- unfold zero; generalize (N.spec_compare x y); case N.compare.
- rewrite N.spec_0; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
- rewrite N.spec_add; ring.
- rewrite N.spec_add; ring.
- unfold zero; generalize (N.spec_compare x y); case N.compare.
- rewrite N.spec_0; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
- intros; rewrite N.spec_sub; try ring; auto with zarith.
+ 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 *.
Qed.
- Definition mul x y :=
+ 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)
@@ -286,25 +257,26 @@ Module Make (N:NType) <: ZType.
| Neg nx, Neg ny => Pos (N.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.
Qed.
- Definition square x :=
+ Definition square x :=
match x with
| Pos nx => Pos (N.square nx)
| Neg nx => Pos (N.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.
Qed.
Definition power_pos x p :=
match x with
| Pos nx => Pos (N.power_pos nx p)
- | Neg nx =>
+ | Neg nx =>
match p with
| xH => x
| xO _ => Pos (N.power_pos nx p)
@@ -313,9 +285,10 @@ Module Make (N:NType) <: ZType.
end.
Theorem spec_power_pos: forall x n, to_Z (power_pos x n) = to_Z x ^ Zpos n.
+ Proof.
assert (F0: forall x, (-x)^2 = x^2).
intros x; rewrite Zpower_2; ring.
- unfold power_pos, to_Z; intros [x | x] [p | p |];
+ unfold power_pos, to_Z; intros [x | x] [p | p |];
try rewrite N.spec_power_pos; try ring.
assert (F: 0 <= 2 * Zpos p).
assert (0 <= Zpos p); auto with zarith.
@@ -329,15 +302,28 @@ Module Make (N:NType) <: ZType.
rewrite F0; ring.
Qed.
+ Definition power x n :=
+ match n with
+ | N0 => one
+ | Npos p => power_pos x p
+ end.
+
+ Theorem spec_power: forall x n, to_Z (power x n) = to_Z x ^ Z_of_N n.
+ Proof.
+ destruct n; simpl. rewrite N.spec_1; reflexivity.
+ apply spec_power_pos.
+ Qed.
+
+
Definition sqrt x :=
match x with
| Pos nx => Pos (N.sqrt nx)
| Neg nx => Neg N.zero
end.
-
- Theorem spec_sqrt: forall x, 0 <= to_Z x ->
+ Theorem spec_sqrt: forall x, 0 <= to_Z x ->
to_Z (sqrt x) ^ 2 <= to_Z x < (to_Z (sqrt x) + 1) ^ 2.
+ Proof.
unfold to_Z, sqrt; intros [x | x] H.
exact (N.spec_sqrt x).
replace (N.to_Z x) with 0.
@@ -353,113 +339,75 @@ Module Make (N:NType) <: ZType.
(Pos q, Pos r)
| Pos nx, Neg ny =>
let (q, r) := N.div_eucl nx ny in
- match N.compare N.zero r with
- | Eq => (Neg q, zero)
- | _ => (Neg (N.succ q), Neg (N.sub ny r))
- end
+ if N.eq_bool N.zero r
+ then (Neg q, zero)
+ else (Neg (N.succ q), Neg (N.sub ny r))
| Neg nx, Pos ny =>
let (q, r) := N.div_eucl nx ny in
- match N.compare N.zero r with
- | Eq => (Neg q, zero)
- | _ => (Neg (N.succ q), Pos (N.sub ny r))
- end
+ if N.eq_bool N.zero r
+ then (Neg q, zero)
+ else (Neg (N.succ q), Pos (N.sub ny r))
| Neg nx, Neg ny =>
let (q, r) := N.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;
+ [clear H|clear H|elim H; reflexivity].
Theorem spec_div_eucl: forall x y,
- to_Z y <> 0 ->
- let (q,r) := div_eucl x y in
- (to_Z q, to_Z r) = Zdiv_eucl (to_Z x) (to_Z y).
- unfold div_eucl, to_Z; intros [x | x] [y | y] H.
- assert (H1: 0 < N.to_Z y).
- generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y H1); case N.div_eucl; auto.
- assert (HH: 0 < N.to_Z y).
- generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y HH); case N.div_eucl; auto.
- intros q r; generalize (N.spec_pos x) HH; unfold Zdiv_eucl;
- case_eq (N.to_Z x); case_eq (N.to_Z y);
- try (intros; apply False_ind; auto with zarith; fail).
- intros p He1 He2 _ _ H1; injection H1; intros H2 H3.
- generalize (N.spec_compare N.zero r); case N.compare;
- unfold zero; rewrite N.spec_0; try rewrite H3; auto.
- rewrite H2; intros; apply False_ind; auto with zarith.
- rewrite H2; intros; apply False_ind; auto with zarith.
- intros p _ _ _ H1; discriminate H1.
- intros p He p1 He1 H1 _.
- generalize (N.spec_compare N.zero r); case N.compare.
- change (- Zpos p) with (Zneg p).
- unfold zero; lazy zeta.
- rewrite N.spec_0; intros H2; rewrite <- H2.
- intros H3; rewrite <- H3; auto.
- rewrite N.spec_0; intros H2.
- change (- Zpos p) with (Zneg p); lazy iota beta.
- intros H3; rewrite <- H3; auto.
- rewrite N.spec_succ; rewrite N.spec_sub.
- generalize H2; case (N.to_Z r).
- intros; apply False_ind; auto with zarith.
- intros p2 _; rewrite He; auto with zarith.
- change (Zneg p) with (- (Zpos p)); apply f_equal2 with (f := @pair Z Z); ring.
- intros p2 H4; discriminate H4.
- assert (N.to_Z r = (Zpos p1 mod (Zpos p))).
- unfold Zmod, Zdiv_eucl; rewrite <- H3; auto.
- case (Z_mod_lt (Zpos p1) (Zpos p)); auto with zarith.
- rewrite N.spec_0; intros H2; generalize (N.spec_pos r);
- intros; apply False_ind; auto with zarith.
- assert (HH: 0 < N.to_Z y).
- generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y HH); case N.div_eucl; auto.
- intros q r; generalize (N.spec_pos x) HH; unfold Zdiv_eucl;
- case_eq (N.to_Z x); case_eq (N.to_Z y);
- try (intros; apply False_ind; auto with zarith; fail).
- intros p He1 He2 _ _ H1; injection H1; intros H2 H3.
- generalize (N.spec_compare N.zero r); case N.compare;
- unfold zero; rewrite N.spec_0; try rewrite H3; auto.
- rewrite H2; intros; apply False_ind; auto with zarith.
- rewrite H2; intros; apply False_ind; auto with zarith.
- intros p _ _ _ H1; discriminate H1.
- intros p He p1 He1 H1 _.
- generalize (N.spec_compare N.zero r); case N.compare.
- change (- Zpos p1) with (Zneg p1).
- unfold zero; lazy zeta.
- rewrite N.spec_0; intros H2; rewrite <- H2.
- intros H3; rewrite <- H3; auto.
- rewrite N.spec_0; intros H2.
- change (- Zpos p1) with (Zneg p1); lazy iota beta.
- intros H3; rewrite <- H3; auto.
- rewrite N.spec_succ; rewrite N.spec_sub.
- generalize H2; case (N.to_Z r).
- intros; apply False_ind; auto with zarith.
- intros p2 _; rewrite He; auto with zarith.
- intros p2 H4; discriminate H4.
- assert (N.to_Z r = (Zpos p1 mod (Zpos p))).
- unfold Zmod, Zdiv_eucl; rewrite <- H3; auto.
- case (Z_mod_lt (Zpos p1) (Zpos p)); auto with zarith.
- rewrite N.spec_0; generalize (N.spec_pos r); intros; apply False_ind; auto with zarith.
- assert (H1: 0 < N.to_Z y).
- generalize (N.spec_pos y); auto with zarith.
- generalize (N.spec_div_eucl x y H1); case N.div_eucl; auto.
- intros q r; generalize (N.spec_pos x) H1; unfold Zdiv_eucl;
- case_eq (N.to_Z x); case_eq (N.to_Z y);
- try (intros; apply False_ind; auto with zarith; fail).
- change (-0) with 0; lazy iota beta; auto.
- intros p _ _ _ _ H2; injection H2.
- intros H3 H4; rewrite H3; rewrite H4; auto.
- intros p _ _ _ H2; discriminate H2.
- intros p He p1 He1 _ _ H2.
- change (- Zpos p1) with (Zneg p1); lazy iota beta.
- change (- Zpos p) with (Zneg p); lazy iota beta.
- rewrite <- H2; auto.
+ let (q,r) := div_eucl x y in
+ (to_Z q, to_Z r) = Zdiv_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.
+ (* Pos Neg *)
+ generalize (N.spec_div_eucl x y); destruct (N.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_eq_bool, N.spec_0, Hr;
+ simpl; rewrite Hq, N.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').
+ intros (EQ,MOD). injection 1. intros Hr' Hq'.
+ rewrite N.spec_eq_bool, N.spec_0, Hr'.
+ break_nonneg r pr EQr.
+ subst; simpl. rewrite N.spec_0; auto.
+ subst. lazy iota beta delta [Zeq_bool Zcompare].
+ rewrite N.spec_sub, N.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).
+ break_nonneg x px EQx; break_nonneg y py EQy;
+ try (injection 1; intros Hr Hq; rewrite N.spec_eq_bool, N.spec_0, Hr;
+ simpl; rewrite Hq, N.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').
+ intros (EQ,MOD). injection 1. intros Hr' Hq'.
+ rewrite N.spec_eq_bool, N.spec_0, Hr'.
+ break_nonneg r pr EQr.
+ subst; simpl. rewrite N.spec_0; auto.
+ subst. lazy iota beta delta [Zeq_bool Zcompare].
+ rewrite N.spec_sub, N.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).
+ break_nonneg x px EQx; break_nonneg y py EQy;
+ try (injection 1; intros Hr Hq; rewrite Hr, Hq; auto).
+ simpl. intros <-; auto.
Qed.
Definition div x y := fst (div_eucl x y).
Definition spec_div: forall x y,
- to_Z y <> 0 -> to_Z (div x y) = to_Z x / to_Z y.
- intros x y H1; generalize (spec_div_eucl x y H1); unfold div, Zdiv.
+ 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 q r q11 r1 H; injection H; auto.
Qed.
@@ -467,8 +415,9 @@ Module Make (N:NType) <: ZType.
Definition modulo x y := snd (div_eucl x y).
Theorem spec_modulo:
- forall x y, to_Z y <> 0 -> to_Z (modulo x y) = to_Z x mod to_Z y.
- intros x y H1; generalize (spec_div_eucl x y H1); unfold modulo, Zmod.
+ 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 q r q11 r1 H; injection H; auto.
Qed.
@@ -478,14 +427,30 @@ Module Make (N:NType) <: ZType.
| 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)
+ | Neg nx, Neg ny => Pos (N.gcd nx ny)
end.
Theorem spec_gcd: forall a b, to_Z (gcd a b) = Zgcd (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.
Qed.
+ Definition sgn x :=
+ match compare zero x with
+ | Lt => one
+ | Eq => zero
+ | Gt => minus_one
+ end.
+
+ Lemma spec_sgn : forall x, to_Z (sgn x) = Zsgn (to_Z x).
+ Proof.
+ intros. unfold sgn. rewrite spec_compare. case Zcompare_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.
+ Qed.
+
End Make.
diff --git a/theories/Numbers/Integer/Binary/ZBinary.v b/theories/Numbers/Integer/Binary/ZBinary.v
index 66d2a96a..835f7958 100644
--- a/theories/Numbers/Integer/Binary/ZBinary.v
+++ b/theories/Numbers/Integer/Binary/ZBinary.v
@@ -8,212 +8,103 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZBinary.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import ZMulOrder.
-Require Import ZArith.
-Open Local Scope Z_scope.
+Require Import ZAxioms ZProperties.
+Require Import ZArith_base.
-Module ZBinAxiomsMod <: ZAxiomsSig.
-Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig.
-Module Export NZAxiomsMod <: NZAxiomsSig.
+Local Open Scope Z_scope.
-Definition NZ := Z.
-Definition NZeq := (@eq Z).
-Definition NZ0 := 0.
-Definition NZsucc := Zsucc'.
-Definition NZpred := Zpred'.
-Definition NZadd := Zplus.
-Definition NZsub := Zminus.
-Definition NZmul := Zmult.
+(** * Implementation of [ZAxiomsSig] by [BinInt.Z] *)
-Theorem NZeq_equiv : equiv Z NZeq.
-Proof.
-exact (@eq_equiv Z).
-Qed.
-
-Add Relation Z NZeq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
-as NZeq_rel.
-
-Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd.
-Proof.
-congruence.
-Qed.
+Module ZBinAxiomsMod <: ZAxiomsExtSig.
-Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd.
-Proof.
-congruence.
-Qed.
+(** Bi-directional induction. *)
-Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd.
-Proof.
-congruence.
-Qed.
-
-Theorem NZpred_succ : forall n : Z, NZpred (NZsucc n) = n.
-Proof.
-exact Zpred'_succ'.
-Qed.
-
-Theorem NZinduction :
- forall A : Z -> Prop, predicate_wd NZeq A ->
- A 0 -> (forall n : Z, A n <-> A (NZsucc n)) -> forall n : Z, A n.
+Theorem bi_induction :
+ forall A : Z -> Prop, Proper (eq ==> iff) A ->
+ A 0 -> (forall n : Z, A n <-> A (Zsucc n)) -> forall n : Z, A n.
Proof.
intros A A_wd A0 AS n; apply Zind; clear n.
assumption.
-intros; now apply -> AS.
-intros n H. rewrite <- (Zsucc'_pred' n) in H. now apply <- AS.
-Qed.
-
-Theorem NZadd_0_l : forall n : Z, 0 + n = n.
-Proof.
-exact Zplus_0_l.
-Qed.
-
-Theorem NZadd_succ_l : forall n m : Z, (NZsucc n) + m = NZsucc (n + m).
-Proof.
-intros; do 2 rewrite <- Zsucc_succ'; apply Zplus_succ_l.
-Qed.
-
-Theorem NZsub_0_r : forall n : Z, n - 0 = n.
-Proof.
-exact Zminus_0_r.
-Qed.
-
-Theorem NZsub_succ_r : forall n m : Z, n - (NZsucc m) = NZpred (n - m).
-Proof.
-intros; rewrite <- Zsucc_succ'; rewrite <- Zpred_pred';
-apply Zminus_succ_r.
-Qed.
-
-Theorem NZmul_0_l : forall n : Z, 0 * n = 0.
-Proof.
-reflexivity.
-Qed.
-
-Theorem NZmul_succ_l : forall n m : Z, (NZsucc n) * m = n * m + m.
-Proof.
-intros; rewrite <- Zsucc_succ'; apply Zmult_succ_l.
-Qed.
-
-End NZAxiomsMod.
-
-Definition NZlt := Zlt.
-Definition NZle := Zle.
-Definition NZmin := Zmin.
-Definition NZmax := Zmax.
-
-Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd.
-Proof.
-unfold NZeq. intros n1 n2 H1 m1 m2 H2; rewrite H1; now rewrite H2.
-Qed.
-
-Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd.
-Proof.
-unfold NZeq. intros n1 n2 H1 m1 m2 H2; rewrite H1; now rewrite H2.
-Qed.
-
-Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd.
-Proof.
-congruence.
-Qed.
-
-Theorem NZlt_eq_cases : forall n m : Z, n <= m <-> n < m \/ n = m.
-Proof.
-intros n m; split. apply Zle_lt_or_eq.
-intro H; destruct H as [H | H]. now apply Zlt_le_weak. rewrite H; apply Zle_refl.
-Qed.
-
-Theorem NZlt_irrefl : forall n : Z, ~ n < n.
-Proof.
-exact Zlt_irrefl.
-Qed.
-
-Theorem NZlt_succ_r : forall n m : Z, n < (NZsucc m) <-> n <= m.
-Proof.
-intros; unfold NZsucc; rewrite <- Zsucc_succ'; split;
-[apply Zlt_succ_le | apply Zle_lt_succ].
-Qed.
-
-Theorem NZmin_l : forall n m : NZ, n <= m -> NZmin n m = n.
-Proof.
-unfold NZmin, Zmin, Zle; intros n m H.
-destruct (n ?= m); try reflexivity. now elim H.
-Qed.
-
-Theorem NZmin_r : forall n m : NZ, m <= n -> NZmin n m = m.
-Proof.
-unfold NZmin, Zmin, Zle; intros n m H.
-case_eq (n ?= m); intro H1; try reflexivity.
-now apply Zcompare_Eq_eq.
-apply <- Zcompare_Gt_Lt_antisym in H1. now elim H.
-Qed.
-
-Theorem NZmax_l : forall n m : NZ, m <= n -> NZmax n m = n.
-Proof.
-unfold NZmax, Zmax, Zle; intros n m H.
-case_eq (n ?= m); intro H1; try reflexivity.
-apply <- Zcompare_Gt_Lt_antisym in H1. now elim H.
-Qed.
-
-Theorem NZmax_r : forall n m : NZ, n <= m -> NZmax n m = m.
-Proof.
-unfold NZmax, Zmax, Zle; intros n m H.
-case_eq (n ?= m); intro H1.
-now apply Zcompare_Eq_eq. reflexivity. now elim H.
-Qed.
-
-End NZOrdAxiomsMod.
-
-Definition Zopp (x : Z) :=
-match x with
-| Z0 => Z0
-| Zpos x => Zneg x
-| Zneg x => Zpos x
-end.
-
-Add Morphism Zopp with signature NZeq ==> NZeq as Zopp_wd.
-Proof.
-congruence.
-Qed.
-
-Theorem Zsucc_pred : forall n : Z, NZsucc (NZpred n) = n.
-Proof.
-exact Zsucc'_pred'.
-Qed.
-
-Theorem Zopp_0 : - 0 = 0.
-Proof.
-reflexivity.
-Qed.
-
-Theorem Zopp_succ : forall n : Z, - (NZsucc n) = NZpred (- n).
-Proof.
-intro; rewrite <- Zsucc_succ'; rewrite <- Zpred_pred'. apply Zopp_succ.
-Qed.
+intros; rewrite <- Zsucc_succ'. now apply -> AS.
+intros n H. rewrite <- Zpred_pred'. rewrite Zsucc_pred in H. now apply <- AS.
+Qed.
+
+(** Basic operations. *)
+
+Definition eq_equiv : Equivalence (@eq Z) := eq_equivalence.
+Local Obligation Tactic := simpl_relation.
+Program Instance succ_wd : Proper (eq==>eq) Zsucc.
+Program Instance pred_wd : Proper (eq==>eq) Zpred.
+Program Instance add_wd : Proper (eq==>eq==>eq) Zplus.
+Program Instance sub_wd : Proper (eq==>eq==>eq) Zminus.
+Program Instance mul_wd : Proper (eq==>eq==>eq) Zmult.
+
+Definition pred_succ n := eq_sym (Zpred_succ n).
+Definition add_0_l := Zplus_0_l.
+Definition add_succ_l := Zplus_succ_l.
+Definition sub_0_r := Zminus_0_r.
+Definition sub_succ_r := Zminus_succ_r.
+Definition mul_0_l := Zmult_0_l.
+Definition mul_succ_l := Zmult_succ_l.
+
+(** Order *)
+
+Program Instance lt_wd : Proper (eq==>eq==>iff) Zlt.
+
+Definition lt_eq_cases := Zle_lt_or_eq_iff.
+Definition lt_irrefl := Zlt_irrefl.
+Definition lt_succ_r := Zlt_succ_r.
+
+Definition min_l := Zmin_l.
+Definition min_r := Zmin_r.
+Definition max_l := Zmax_l.
+Definition max_r := Zmax_r.
+
+(** Properties specific to integers, not natural numbers. *)
+
+Program Instance opp_wd : Proper (eq==>eq) Zopp.
+
+Definition succ_pred n := eq_sym (Zsucc_pred n).
+Definition opp_0 := Zopp_0.
+Definition opp_succ := Zopp_succ.
+
+(** Absolute value and sign *)
+
+Definition abs_eq := Zabs_eq.
+Definition abs_neq := Zabs_non_eq.
+
+Lemma sgn_null : forall x, x = 0 -> Zsgn x = 0.
+Proof. intros. apply <- Zsgn_null; auto. Qed.
+Lemma sgn_pos : forall x, 0 < x -> Zsgn x = 1.
+Proof. intros. apply <- Zsgn_pos; auto. Qed.
+Lemma sgn_neg : forall x, x < 0 -> Zsgn x = -1.
+Proof. intros. apply <- Zsgn_neg; auto. Qed.
+
+(** The instantiation of operations.
+ Placing them at the very end avoids having indirections in above lemmas. *)
+
+Definition t := Z.
+Definition eq := (@eq Z).
+Definition zero := 0.
+Definition succ := Zsucc.
+Definition pred := Zpred.
+Definition add := Zplus.
+Definition sub := Zminus.
+Definition mul := Zmult.
+Definition lt := Zlt.
+Definition le := Zle.
+Definition min := Zmin.
+Definition max := Zmax.
+Definition opp := Zopp.
+Definition abs := Zabs.
+Definition sgn := Zsgn.
End ZBinAxiomsMod.
-Module Export ZBinMulOrderPropMod := ZMulOrderPropFunct ZBinAxiomsMod.
+Module Export ZBinPropMod := ZPropFunct ZBinAxiomsMod.
(** Z forms a ring *)
diff --git a/theories/Numbers/Integer/NatPairs/ZNatPairs.v b/theories/Numbers/Integer/NatPairs/ZNatPairs.v
index 9427b37b..8b5624cd 100644
--- a/theories/Numbers/Integer/NatPairs/ZNatPairs.v
+++ b/theories/Numbers/Integer/NatPairs/ZNatPairs.v
@@ -8,400 +8,306 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZNatPairs.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import NSub. (* The most complete file for natural numbers *)
-Require Export ZMulOrder. (* The most complete file for integers *)
+Require Import NProperties. (* The most complete file for N *)
+Require Export ZProperties. (* The most complete file for Z *)
Require Export Ring.
-Module ZPairsAxiomsMod (Import NAxiomsMod : NAxiomsSig) <: ZAxiomsSig.
-Module Import NPropMod := NSubPropFunct NAxiomsMod. (* Get all properties of natural numbers *)
-
-(* We do not declare ring in Natural/Abstract for two reasons. First, some
-of the properties proved in NAdd and NMul are used in the new BinNat,
-and it is in turn used in Ring. Using ring in Natural/Abstract would be
-circular. It is possible, however, not to make BinNat dependent on
-Numbers/Natural and prove the properties necessary for ring from scratch
-(this is, of course, how it used to be). In addition, if we define semiring
-structures in the implementation subdirectories of Natural, we are able to
-specify binary natural numbers as the type of coefficients. For these
-reasons we define an abstract semiring here. *)
-
-Open Local Scope NatScope.
-
-Lemma Nsemi_ring : semi_ring_theory 0 1 add mul Neq.
-Proof.
-constructor.
-exact add_0_l.
-exact add_comm.
-exact add_assoc.
-exact mul_1_l.
-exact mul_0_l.
-exact mul_comm.
-exact mul_assoc.
-exact mul_add_distr_r.
-Qed.
-
-Add Ring NSR : Nsemi_ring.
-
-(* The definitios of functions (NZadd, NZmul, etc.) will be unfolded by
-the properties functor. Since we don't want Zadd_comm to refer to unfolded
-definitions of equality: fun p1 p2 : NZ => (fst p1 + snd p2) = (fst p2 + snd p1),
-we will provide an extra layer of definitions. *)
-
-Definition Z := (N * N)%type.
-Definition Z0 : Z := (0, 0).
-Definition Zeq (p1 p2 : Z) := ((fst p1) + (snd p2) == (fst p2) + (snd p1)).
-Definition Zsucc (n : Z) : Z := (S (fst n), snd n).
-Definition Zpred (n : Z) : Z := (fst n, S (snd n)).
-
-(* We do not have Zpred (Zsucc n) = n but only Zpred (Zsucc 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 could get rid of setoids and arrive at
-integers as signed natural numbers. *)
-
-Definition Zadd (n m : Z) : Z := ((fst n) + (fst m), (snd n) + (snd m)).
-Definition Zsub (n m : Z) : Z := ((fst n) + (snd m), (snd n) + (fst m)).
-
-(* Unfortunately, the elements of the pair keep increasing, even during
-subtraction *)
-
-Definition Zmul (n m : Z) : Z :=
- ((fst n) * (fst m) + (snd n) * (snd m), (fst n) * (snd m) + (snd n) * (fst m)).
-Definition Zlt (n m : Z) := (fst n) + (snd m) < (fst m) + (snd n).
-Definition Zle (n m : Z) := (fst n) + (snd m) <= (fst m) + (snd n).
-Definition Zmin (n m : Z) := (min ((fst n) + (snd m)) ((fst m) + (snd n)), (snd n) + (snd m)).
-Definition Zmax (n m : Z) := (max ((fst n) + (snd m)) ((fst m) + (snd n)), (snd n) + (snd m)).
-
-Delimit Scope IntScope with Int.
-Bind Scope IntScope with Z.
-Notation "x == y" := (Zeq x y) (at level 70) : IntScope.
-Notation "x ~= y" := (~ Zeq x y) (at level 70) : IntScope.
-Notation "0" := Z0 : IntScope.
-Notation "1" := (Zsucc Z0) : IntScope.
-Notation "x + y" := (Zadd x y) : IntScope.
-Notation "x - y" := (Zsub x y) : IntScope.
-Notation "x * y" := (Zmul x y) : IntScope.
-Notation "x < y" := (Zlt x y) : IntScope.
-Notation "x <= y" := (Zle x y) : IntScope.
-Notation "x > y" := (Zlt y x) (only parsing) : IntScope.
-Notation "x >= y" := (Zle y x) (only parsing) : IntScope.
-
-Notation Local N := NZ.
-(* To remember N without having to use a long qualifying name. since NZ will be redefined *)
-Notation Local NE := NZeq (only parsing).
-Notation Local add_wd := NZadd_wd (only parsing).
-
-Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig.
-Module Export NZAxiomsMod <: NZAxiomsSig.
-
-Definition NZ : Type := Z.
-Definition NZeq := Zeq.
-Definition NZ0 := Z0.
-Definition NZsucc := Zsucc.
-Definition NZpred := Zpred.
-Definition NZadd := Zadd.
-Definition NZsub := Zsub.
-Definition NZmul := Zmul.
-
-Theorem ZE_refl : reflexive Z Zeq.
-Proof.
-unfold reflexive, Zeq. reflexivity.
-Qed.
-
-Theorem ZE_sym : symmetric Z Zeq.
-Proof.
-unfold symmetric, Zeq; now symmetry.
-Qed.
-
-Theorem ZE_trans : transitive Z Zeq.
-Proof.
-unfold transitive, Zeq. intros n m p H1 H2.
-assert (H3 : (fst n + snd m) + (fst m + snd p) == (fst m + snd n) + (fst p + snd m))
-by now apply add_wd.
-stepl ((fst n + snd p) + (fst m + snd m)) in H3 by ring.
-stepr ((fst p + snd n) + (fst m + snd m)) in H3 by ring.
-now apply -> add_cancel_r in H3.
+Notation "s #1" := (fst s) (at level 9, format "s '#1'") : pair_scope.
+Notation "s #2" := (snd s) (at level 9, format "s '#2'") : pair_scope.
+Open Local Scope pair_scope.
+
+Module ZPairsAxiomsMod (Import N : NAxiomsSig) <: ZAxiomsSig.
+Module Import NPropMod := NPropFunct N. (* Get all properties of N *)
+
+Delimit Scope NScope with N.
+Bind Scope NScope with N.t.
+Infix "==" := N.eq (at level 70) : NScope.
+Notation "x ~= y" := (~ N.eq x y) (at level 70) : NScope.
+Notation "0" := N.zero : NScope.
+Notation "1" := (N.succ N.zero) : NScope.
+Infix "+" := N.add : NScope.
+Infix "-" := N.sub : NScope.
+Infix "*" := N.mul : NScope.
+Infix "<" := N.lt : NScope.
+Infix "<=" := N.le : NScope.
+Local Open Scope NScope.
+
+(** The definitions of functions ([add], [mul], etc.) will be unfolded
+ by the properties functor. Since we don't want [add_comm] to refer
+ to unfolded definitions of equality: [fun p1 p2 => (fst p1 +
+ snd p2) = (fst p2 + snd p1)], we will provide an extra layer of
+ definitions. *)
+
+Module Z.
+
+Definition t := (N.t * N.t)%type.
+Definition zero : t := (0, 0).
+Definition eq (p q : t) := (p#1 + q#2 == q#1 + p#2).
+Definition succ (n : t) : t := (N.succ n#1, n#2).
+Definition pred (n : t) : t := (n#1, N.succ n#2).
+Definition opp (n : t) : t := (n#2, n#1).
+Definition add (n m : t) : t := (n#1 + m#1, n#2 + m#2).
+Definition sub (n m : t) : t := (n#1 + m#2, n#2 + m#1).
+Definition mul (n m : t) : t :=
+ (n#1 * m#1 + n#2 * m#2, n#1 * m#2 + n#2 * m#1).
+Definition lt (n m : t) := n#1 + m#2 < m#1 + n#2.
+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].
+ 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
+ could get rid of setoids and arrive at integers as signed natural
+ numbers. *)
+
+(** NB : Unfortunately, the elements of the pair keep increasing during
+ many operations, even during subtraction. *)
+
+End Z.
+
+Delimit Scope ZScope with Z.
+Bind Scope ZScope with Z.t.
+Infix "==" := Z.eq (at level 70) : ZScope.
+Notation "x ~= y" := (~ Z.eq x y) (at level 70) : ZScope.
+Notation "0" := Z.zero : ZScope.
+Notation "1" := (Z.succ Z.zero) : ZScope.
+Infix "+" := Z.add : ZScope.
+Infix "-" := Z.sub : ZScope.
+Infix "*" := Z.mul : ZScope.
+Notation "- x" := (Z.opp x) : ZScope.
+Infix "<" := Z.lt : ZScope.
+Infix "<=" := Z.le : ZScope.
+Local Open Scope ZScope.
+
+Lemma sub_add_opp : forall n m, Z.sub n m = Z.add n (Z.opp m).
+Proof. reflexivity. Qed.
+
+Instance eq_equiv : Equivalence Z.eq.
+Proof.
+split.
+unfold Reflexive, Z.eq. reflexivity.
+unfold Symmetric, Z.eq; now symmetry.
+unfold Transitive, Z.eq. intros (n1,n2) (m1,m2) (p1,p2) H1 H2; simpl in *.
+apply (add_cancel_r _ _ (m1+m2)%N).
+rewrite add_shuffle2, H1, add_shuffle1, H2.
+now rewrite add_shuffle1, (add_comm m1).
+Qed.
+
+Instance pair_wd : Proper (N.eq==>N.eq==>Z.eq) (@pair N.t N.t).
+Proof.
+intros n1 n2 H1 m1 m2 H2; unfold Z.eq; simpl; now rewrite H1, H2.
+Qed.
+
+Instance succ_wd : Proper (Z.eq ==> Z.eq) Z.succ.
+Proof.
+unfold Z.succ, Z.eq; intros n m H; simpl.
+do 2 rewrite add_succ_l; now rewrite H.
Qed.
-Theorem NZeq_equiv : equiv Z Zeq.
+Instance pred_wd : Proper (Z.eq ==> Z.eq) Z.pred.
Proof.
-unfold equiv; repeat split; [apply ZE_refl | apply ZE_trans | apply ZE_sym].
-Qed.
-
-Add Relation Z Zeq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
-as NZeq_rel.
-
-Add Morphism (@pair N N) with signature NE ==> NE ==> Zeq as Zpair_wd.
-Proof.
-intros n1 n2 H1 m1 m2 H2; unfold Zeq; simpl; rewrite H1; now rewrite H2.
+unfold Z.pred, Z.eq; intros n m H; simpl.
+do 2 rewrite add_succ_r; now rewrite H.
Qed.
-Add Morphism NZsucc with signature Zeq ==> Zeq as NZsucc_wd.
+Instance add_wd : Proper (Z.eq ==> Z.eq ==> Z.eq) Z.add.
Proof.
-unfold NZsucc, Zeq; intros n m H; simpl.
-do 2 rewrite add_succ_l; now rewrite H.
+unfold Z.eq, Z.add; intros n1 m1 H1 n2 m2 H2; simpl.
+now rewrite add_shuffle1, H1, H2, add_shuffle1.
Qed.
-Add Morphism NZpred with signature Zeq ==> Zeq as NZpred_wd.
+Instance opp_wd : Proper (Z.eq ==> Z.eq) Z.opp.
Proof.
-unfold NZpred, Zeq; intros n m H; simpl.
-do 2 rewrite add_succ_r; now rewrite H.
+unfold Z.eq, Z.opp; intros (n1,n2) (m1,m2) H; simpl in *.
+now rewrite (add_comm n2), (add_comm m2).
Qed.
-Add Morphism NZadd with signature Zeq ==> Zeq ==> Zeq as NZadd_wd.
+Instance sub_wd : Proper (Z.eq ==> Z.eq ==> Z.eq) Z.sub.
Proof.
-unfold Zeq, NZadd; intros n1 m1 H1 n2 m2 H2; simpl.
-assert (H3 : (fst n1 + snd m1) + (fst n2 + snd m2) == (fst m1 + snd n1) + (fst m2 + snd n2))
-by now apply add_wd.
-stepl (fst n1 + snd m1 + (fst n2 + snd m2)) by ring.
-now stepr (fst m1 + snd n1 + (fst m2 + snd n2)) by ring.
+intros n1 m1 H1 n2 m2 H2. rewrite 2 sub_add_opp.
+apply add_wd, opp_wd; auto.
Qed.
-Add Morphism NZsub with signature Zeq ==> Zeq ==> Zeq as NZsub_wd.
+Lemma mul_comm : forall n m, n*m == m*n.
Proof.
-unfold Zeq, NZsub; intros n1 m1 H1 n2 m2 H2; simpl.
-symmetry in H2.
-assert (H3 : (fst n1 + snd m1) + (fst m2 + snd n2) == (fst m1 + snd n1) + (fst n2 + snd m2))
-by now apply add_wd.
-stepl (fst n1 + snd m1 + (fst m2 + snd n2)) by ring.
-now stepr (fst m1 + snd n1 + (fst n2 + snd m2)) by ring.
+intros (n1,n2) (m1,m2); compute.
+rewrite (add_comm (m1*n2)%N).
+apply N.add_wd; apply N.add_wd; apply mul_comm.
Qed.
-Add Morphism NZmul with signature Zeq ==> Zeq ==> Zeq as NZmul_wd.
+Instance mul_wd : Proper (Z.eq ==> Z.eq ==> Z.eq) Z.mul.
Proof.
-unfold NZmul, Zeq; intros n1 m1 H1 n2 m2 H2; simpl.
-stepl (fst n1 * fst n2 + (snd n1 * snd n2 + fst m1 * snd m2 + snd m1 * fst m2)) by ring.
-stepr (fst n1 * snd n2 + (fst m1 * fst m2 + snd m1 * snd m2 + snd n1 * fst n2)) by ring.
-apply add_mul_repl_pair with (n := fst m2) (m := snd m2); [| now idtac].
-stepl (snd n1 * snd n2 + (fst n1 * fst m2 + fst m1 * snd m2 + snd m1 * fst m2)) by ring.
-stepr (snd n1 * fst n2 + (fst n1 * snd m2 + fst m1 * fst m2 + snd m1 * snd m2)) by ring.
-apply add_mul_repl_pair with (n := snd m2) (m := fst m2);
-[| (stepl (fst n2 + snd m2) by ring); now stepr (fst m2 + snd n2) by ring].
-stepl (snd m2 * snd n1 + (fst n1 * fst m2 + fst m1 * snd m2 + snd m1 * fst m2)) by ring.
-stepr (snd m2 * fst n1 + (snd n1 * fst m2 + fst m1 * fst m2 + snd m1 * snd m2)) by ring.
-apply add_mul_repl_pair with (n := snd m1) (m := fst m1);
-[ | (stepl (fst n1 + snd m1) by ring); now stepr (fst m1 + snd n1) by ring].
-stepl (fst m2 * fst n1 + (snd m2 * snd m1 + fst m1 * snd m2 + snd m1 * fst m2)) by ring.
-stepr (fst m2 * snd n1 + (snd m2 * fst m1 + fst m1 * fst m2 + snd m1 * snd m2)) by ring.
-apply add_mul_repl_pair with (n := fst m1) (m := snd m1); [| now idtac].
-ring.
+assert (forall n, Proper (Z.eq ==> Z.eq) (Z.mul n)).
+ unfold Z.mul, Z.eq. intros (n1,n2) (p1,p2) (q1,q2) H; simpl in *.
+ rewrite add_shuffle1, (add_comm (n1*p1)%N).
+ symmetry. rewrite add_shuffle1.
+ rewrite <- ! mul_add_distr_l.
+ rewrite (add_comm p2), (add_comm q2), H.
+ reflexivity.
+intros n n' Hn m m' Hm.
+rewrite Hm, (mul_comm n), (mul_comm n'), Hn.
+reflexivity.
Qed.
Section Induction.
-Open Scope NatScope. (* automatically closes at the end of the section *)
-Variable A : Z -> Prop.
-Hypothesis A_wd : predicate_wd Zeq A.
+Variable A : Z.t -> Prop.
+Hypothesis A_wd : Proper (Z.eq==>iff) A.
-Add Morphism A with signature Zeq ==> iff as A_morph.
+Theorem bi_induction :
+ A 0 -> (forall n, A n <-> A (Z.succ n)) -> forall n, A n.
Proof.
-exact A_wd.
-Qed.
-
-Theorem NZinduction :
- A 0 -> (forall n : Z, A n <-> A (Zsucc n)) -> forall n : Z, A n. (* 0 is interpreted as in Z due to "Bind" directive *)
-Proof.
-intros A0 AS n; unfold NZ0, Zsucc, predicate_wd, fun_wd, Zeq in *.
+intros A0 AS n; unfold Z.zero, Z.succ, Z.eq in *.
destruct n as [n m].
-cut (forall p : N, A (p, 0)); [intro H1 |].
-cut (forall p : N, A (0, p)); [intro H2 |].
+cut (forall p, A (p, 0%N)); [intro H1 |].
+cut (forall p, A (0%N, p)); [intro H2 |].
destruct (add_dichotomy n m) as [[p H] | [p H]].
-rewrite (A_wd (n, m) (0, p)) by (rewrite add_0_l; now rewrite add_comm).
+rewrite (A_wd (n, m) (0%N, p)) by (rewrite add_0_l; now rewrite add_comm).
apply H2.
-rewrite (A_wd (n, m) (p, 0)) by now rewrite add_0_r. apply H1.
+rewrite (A_wd (n, m) (p, 0%N)) by now rewrite add_0_r. apply H1.
induct p. assumption. intros p IH.
-apply -> (A_wd (0, p) (1, S p)) in IH; [| now rewrite add_0_l, add_1_l].
+apply -> (A_wd (0%N, p) (1%N, N.succ p)) in IH; [| now rewrite add_0_l, add_1_l].
now apply <- AS.
induct p. assumption. intros p IH.
-replace 0 with (snd (p, 0)); [| reflexivity].
-replace (S p) with (S (fst (p, 0))); [| reflexivity]. now apply -> AS.
+replace 0%N with (snd (p, 0%N)); [| reflexivity].
+replace (N.succ p) with (N.succ (fst (p, 0%N))); [| reflexivity]. now apply -> AS.
Qed.
End Induction.
(* Time to prove theorems in the language of Z *)
-Open Local Scope IntScope.
-
-Theorem NZpred_succ : forall n : Z, Zpred (Zsucc n) == n.
+Theorem pred_succ : forall n, Z.pred (Z.succ n) == n.
Proof.
-unfold NZpred, NZsucc, Zeq; intro n; simpl.
-rewrite add_succ_l; now rewrite add_succ_r.
+unfold Z.pred, Z.succ, Z.eq; intro n; simpl; now nzsimpl.
Qed.
-Theorem NZadd_0_l : forall n : Z, 0 + n == n.
+Theorem succ_pred : forall n, Z.succ (Z.pred n) == n.
Proof.
-intro n; unfold NZadd, Zeq; simpl. now do 2 rewrite add_0_l.
+intro n; unfold Z.succ, Z.pred, Z.eq; simpl; now nzsimpl.
Qed.
-Theorem NZadd_succ_l : forall n m : Z, (Zsucc n) + m == Zsucc (n + m).
+Theorem opp_0 : - 0 == 0.
Proof.
-intros n m; unfold NZadd, Zeq; simpl. now do 2 rewrite add_succ_l.
+unfold Z.opp, Z.eq; simpl. now nzsimpl.
Qed.
-Theorem NZsub_0_r : forall n : Z, n - 0 == n.
+Theorem opp_succ : forall n, - (Z.succ n) == Z.pred (- n).
Proof.
-intro n; unfold NZsub, Zeq; simpl. now do 2 rewrite add_0_r.
-Qed.
-
-Theorem NZsub_succ_r : forall n m : Z, n - (Zsucc m) == Zpred (n - m).
-Proof.
-intros n m; unfold NZsub, Zeq; simpl. symmetry; now rewrite add_succ_r.
+reflexivity.
Qed.
-Theorem NZmul_0_l : forall n : Z, 0 * n == 0.
+Theorem add_0_l : forall n, 0 + n == n.
Proof.
-intro n; unfold NZmul, Zeq; simpl.
-repeat rewrite mul_0_l. now rewrite add_assoc.
+intro n; unfold Z.add, Z.eq; simpl. now nzsimpl.
Qed.
-Theorem NZmul_succ_l : forall n m : Z, (Zsucc n) * m == n * m + m.
+Theorem add_succ_l : forall n m, (Z.succ n) + m == Z.succ (n + m).
Proof.
-intros n m; unfold NZmul, NZsucc, Zeq; simpl.
-do 2 rewrite mul_succ_l. ring.
+intros n m; unfold Z.add, Z.eq; simpl. now nzsimpl.
Qed.
-End NZAxiomsMod.
-
-Definition NZlt := Zlt.
-Definition NZle := Zle.
-Definition NZmin := Zmin.
-Definition NZmax := Zmax.
-
-Add Morphism NZlt with signature Zeq ==> Zeq ==> iff as NZlt_wd.
+Theorem sub_0_r : forall n, n - 0 == n.
Proof.
-unfold NZlt, Zlt, Zeq; intros n1 m1 H1 n2 m2 H2; simpl. split; intro H.
-stepr (snd m1 + fst m2) by apply add_comm.
-apply (add_lt_repl_pair (fst n1) (snd n1)); [| assumption].
-stepl (snd m2 + fst n1) by apply add_comm.
-stepr (fst m2 + snd n1) by apply add_comm.
-apply (add_lt_repl_pair (snd n2) (fst n2)).
-now stepl (fst n1 + snd n2) by apply add_comm.
-stepl (fst m2 + snd n2) by apply add_comm. now stepr (fst n2 + snd m2) by apply add_comm.
-stepr (snd n1 + fst n2) by apply add_comm.
-apply (add_lt_repl_pair (fst m1) (snd m1)); [| now symmetry].
-stepl (snd n2 + fst m1) by apply add_comm.
-stepr (fst n2 + snd m1) by apply add_comm.
-apply (add_lt_repl_pair (snd m2) (fst m2)).
-now stepl (fst m1 + snd m2) by apply add_comm.
-stepl (fst n2 + snd m2) by apply add_comm. now stepr (fst m2 + snd n2) by apply add_comm.
+intro n; unfold Z.sub, Z.eq; simpl. now nzsimpl.
Qed.
-Add Morphism NZle with signature Zeq ==> Zeq ==> iff as NZle_wd.
+Theorem sub_succ_r : forall n m, n - (Z.succ m) == Z.pred (n - m).
Proof.
-unfold NZle, Zle, Zeq; intros n1 m1 H1 n2 m2 H2; simpl.
-do 2 rewrite lt_eq_cases. rewrite (NZlt_wd n1 m1 H1 n2 m2 H2). fold (m1 < m2)%Int.
-fold (n1 == n2)%Int (m1 == m2)%Int; fold (n1 == m1)%Int in H1; fold (n2 == m2)%Int in H2.
-now rewrite H1, H2.
+intros n m; unfold Z.sub, Z.eq; simpl. symmetry; now rewrite add_succ_r.
Qed.
-Add Morphism NZmin with signature Zeq ==> Zeq ==> Zeq as NZmin_wd.
+Theorem mul_0_l : forall n, 0 * n == 0.
Proof.
-intros n1 m1 H1 n2 m2 H2; unfold NZmin, Zeq; simpl.
-destruct (le_ge_cases (fst n1 + snd n2) (fst n2 + snd n1)) as [H | H].
-rewrite (min_l (fst n1 + snd n2) (fst n2 + snd n1)) by assumption.
-rewrite (min_l (fst m1 + snd m2) (fst m2 + snd m1)) by
-now apply -> (NZle_wd n1 m1 H1 n2 m2 H2).
-stepl ((fst n1 + snd m1) + (snd n2 + snd m2)) by ring.
-unfold Zeq in H1. rewrite H1. ring.
-rewrite (min_r (fst n1 + snd n2) (fst n2 + snd n1)) by assumption.
-rewrite (min_r (fst m1 + snd m2) (fst m2 + snd m1)) by
-now apply -> (NZle_wd n2 m2 H2 n1 m1 H1).
-stepl ((fst n2 + snd m2) + (snd n1 + snd m1)) by ring.
-unfold Zeq in H2. rewrite H2. ring.
+intros (n1,n2); unfold Z.mul, Z.eq; simpl; now nzsimpl.
Qed.
-Add Morphism NZmax with signature Zeq ==> Zeq ==> Zeq as NZmax_wd.
+Theorem mul_succ_l : forall n m, (Z.succ n) * m == n * m + m.
Proof.
-intros n1 m1 H1 n2 m2 H2; unfold NZmax, Zeq; simpl.
-destruct (le_ge_cases (fst n1 + snd n2) (fst n2 + snd n1)) as [H | H].
-rewrite (max_r (fst n1 + snd n2) (fst n2 + snd n1)) by assumption.
-rewrite (max_r (fst m1 + snd m2) (fst m2 + snd m1)) by
-now apply -> (NZle_wd n1 m1 H1 n2 m2 H2).
-stepl ((fst n2 + snd m2) + (snd n1 + snd m1)) by ring.
-unfold Zeq in H2. rewrite H2. ring.
-rewrite (max_l (fst n1 + snd n2) (fst n2 + snd n1)) by assumption.
-rewrite (max_l (fst m1 + snd m2) (fst m2 + snd m1)) by
-now apply -> (NZle_wd n2 m2 H2 n1 m1 H1).
-stepl ((fst n1 + snd m1) + (snd n2 + snd m2)) by ring.
-unfold Zeq in H1. rewrite H1. ring.
+intros (n1,n2) (m1,m2); unfold Z.mul, Z.succ, Z.eq; simpl; nzsimpl.
+rewrite <- (add_assoc _ m1), (add_comm m1), (add_assoc _ _ m1).
+now rewrite <- (add_assoc _ m2), (add_comm m2), (add_assoc _ (n2*m1)%N m2).
Qed.
-Open Local Scope IntScope.
+(** Order *)
-Theorem NZlt_eq_cases : forall n m : Z, n <= m <-> n < m \/ n == m.
+Lemma lt_eq_cases : forall n m, n<=m <-> n<m \/ n==m.
Proof.
-intros n m; unfold Zlt, Zle, Zeq; simpl. apply lt_eq_cases.
+intros; apply N.lt_eq_cases.
Qed.
-Theorem NZlt_irrefl : forall n : Z, ~ (n < n).
+Theorem lt_irrefl : forall n, ~ (n < n).
Proof.
-intros n; unfold Zlt, Zeq; simpl. apply lt_irrefl.
+intros; apply N.lt_irrefl.
Qed.
-Theorem NZlt_succ_r : forall n m : Z, n < (Zsucc m) <-> n <= m.
+Theorem lt_succ_r : forall n m, n < (Z.succ m) <-> n <= m.
Proof.
-intros n m; unfold Zlt, Zle, Zeq; simpl. rewrite add_succ_l; apply lt_succ_r.
+intros n m; unfold Z.lt, Z.le, Z.eq; simpl; nzsimpl. apply lt_succ_r.
Qed.
-Theorem NZmin_l : forall n m : Z, n <= m -> Zmin n m == n.
+Theorem min_l : forall n m, n <= m -> Z.min n m == n.
Proof.
-unfold Zmin, Zle, Zeq; simpl; intros n m H.
-rewrite min_l by assumption. ring.
+unfold Z.min, Z.le, Z.eq; simpl; intros (n1,n2) (m1,m2) H; simpl in *.
+rewrite min_l by assumption.
+now rewrite <- add_assoc, (add_comm m2).
Qed.
-Theorem NZmin_r : forall n m : Z, m <= n -> Zmin n m == m.
+Theorem min_r : forall n m, m <= n -> Z.min n m == m.
Proof.
-unfold Zmin, Zle, Zeq; simpl; intros n m H.
-rewrite min_r by assumption. ring.
+unfold Z.min, Z.le, Z.eq; simpl; intros (n1,n2) (m1,m2) H; simpl in *.
+rewrite min_r by assumption.
+now rewrite add_assoc.
Qed.
-Theorem NZmax_l : forall n m : Z, m <= n -> Zmax n m == n.
+Theorem max_l : forall n m, m <= n -> Z.max n m == n.
Proof.
-unfold Zmax, Zle, Zeq; simpl; intros n m H.
-rewrite max_l by assumption. ring.
+unfold Z.max, Z.le, Z.eq; simpl; intros (n1,n2) (m1,m2) H; simpl in *.
+rewrite max_l by assumption.
+now rewrite <- add_assoc, (add_comm m2).
Qed.
-Theorem NZmax_r : forall n m : Z, n <= m -> Zmax n m == m.
+Theorem max_r : forall n m, n <= m -> Z.max n m == m.
Proof.
-unfold Zmax, Zle, Zeq; simpl; intros n m H.
-rewrite max_r by assumption. ring.
+unfold Z.max, Z.le, Z.eq; simpl; intros n m H.
+rewrite max_r by assumption.
+now rewrite add_assoc.
Qed.
-End NZOrdAxiomsMod.
-
-Definition Zopp (n : Z) : Z := (snd n, fst n).
-
-Notation "- x" := (Zopp x) : IntScope.
-
-Add Morphism Zopp with signature Zeq ==> Zeq as Zopp_wd.
-Proof.
-unfold Zeq; intros n m H; simpl. symmetry.
-stepl (fst n + snd m) by apply add_comm.
-now stepr (fst m + snd n) by apply add_comm.
-Qed.
-
-Open Local Scope IntScope.
-
-Theorem Zsucc_pred : forall n : Z, Zsucc (Zpred n) == n.
+Theorem lt_nge : forall n m, n < m <-> ~(m<=n).
Proof.
-intro n; unfold Zsucc, Zpred, Zeq; simpl.
-rewrite add_succ_l; now rewrite add_succ_r.
+intros. apply lt_nge.
Qed.
-Theorem Zopp_0 : - 0 == 0.
+Instance lt_wd : Proper (Z.eq ==> Z.eq ==> iff) Z.lt.
Proof.
-unfold Zopp, Zeq; simpl. now rewrite add_0_l.
+assert (forall n, Proper (Z.eq==>iff) (Z.lt n)).
+ intros (n1,n2). apply proper_sym_impl_iff; auto with *.
+ unfold Z.lt, Z.eq; intros (r1,r2) (s1,s2) Eq H; simpl in *.
+ apply le_lt_add_lt with (r1+r2)%N (r1+r2)%N; [apply le_refl; auto with *|].
+ rewrite add_shuffle2, (add_comm s2), Eq.
+ rewrite (add_comm s1 n2), (add_shuffle1 n2), (add_comm n2 r1).
+ now rewrite <- add_lt_mono_r.
+intros n n' Hn m m' Hm.
+rewrite Hm. rewrite 2 lt_nge, 2 lt_eq_cases, Hn; auto with *.
Qed.
-Theorem Zopp_succ : forall n, - (Zsucc n) == Zpred (- n).
-Proof.
-reflexivity.
-Qed.
+Definition t := Z.t.
+Definition eq := Z.eq.
+Definition zero := Z.zero.
+Definition succ := Z.succ.
+Definition pred := Z.pred.
+Definition add := Z.add.
+Definition sub := Z.sub.
+Definition mul := Z.mul.
+Definition opp := Z.opp.
+Definition lt := Z.lt.
+Definition le := Z.le.
+Definition min := Z.min.
+Definition max := Z.max.
End ZPairsAxiomsMod.
@@ -413,9 +319,7 @@ and get their properties *)
Require Import NPeano.
Module Export ZPairsPeanoAxiomsMod := ZPairsAxiomsMod NPeanoAxiomsMod.
-Module Export ZPairsMulOrderPropMod := ZMulOrderPropFunct ZPairsPeanoAxiomsMod.
-
-Open Local Scope IntScope.
+Module Export ZPairsPropMod := ZPropFunct ZPairsPeanoAxiomsMod.
Eval compute in (3, 5) * (4, 6).
*)
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSig.v b/theories/Numbers/Integer/SpecViaZ/ZSig.v
index 0af98c74..ffa91706 100644
--- a/theories/Numbers/Integer/SpecViaZ/ZSig.v
+++ b/theories/Numbers/Integer/SpecViaZ/ZSig.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: ZSig.v 11027 2008-06-01 13:28:59Z letouzey $ i*)
+(*i $Id$ i*)
Require Import ZArith Znumtheory.
@@ -25,93 +25,77 @@ Module Type ZType.
Parameter t : Type.
Parameter to_Z : t -> Z.
- Notation "[ x ]" := (to_Z x).
+ Local Notation "[ x ]" := (to_Z x).
- Definition eq x y := ([x] = [y]).
+ Definition eq x y := [x] = [y].
+ Definition lt x y := [x] < [y].
+ Definition le x y := [x] <= [y].
Parameter of_Z : Z -> t.
Parameter spec_of_Z: forall x, to_Z (of_Z x) = x.
+ Parameter compare : t -> t -> comparison.
+ Parameter eq_bool : t -> t -> bool.
+ Parameter min : t -> t -> t.
+ Parameter max : t -> t -> t.
Parameter zero : t.
Parameter one : t.
Parameter minus_one : t.
-
- Parameter spec_0: [zero] = 0.
- Parameter spec_1: [one] = 1.
- Parameter spec_m1: [minus_one] = -1.
-
- Parameter compare : t -> t -> comparison.
-
- Parameter spec_compare: forall x y,
- match compare x y with
- | Eq => [x] = [y]
- | Lt => [x] < [y]
- | Gt => [x] > [y]
- end.
-
- Definition lt n m := compare n m = Lt.
- Definition le n m := compare n m <> Gt.
- 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.
-
- Parameter eq_bool : t -> t -> bool.
-
- Parameter spec_eq_bool: forall x y,
- if eq_bool x y then [x] = [y] else [x] <> [y].
-
Parameter succ : t -> t.
-
- Parameter spec_succ: forall n, [succ n] = [n] + 1.
-
Parameter add : t -> t -> t.
-
- Parameter spec_add: forall x y, [add x y] = [x] + [y].
-
Parameter pred : t -> t.
-
- Parameter spec_pred: forall x, [pred x] = [x] - 1.
-
Parameter sub : t -> t -> t.
-
- Parameter spec_sub: forall x y, [sub x y] = [x] - [y].
-
Parameter opp : t -> t.
-
- Parameter spec_opp: forall x, [opp x] = - [x].
-
Parameter mul : t -> t -> t.
-
- Parameter spec_mul: forall x y, [mul x y] = [x] * [y].
-
Parameter square : t -> t.
-
- Parameter spec_square: forall x, [square x] = [x] * [x].
-
Parameter power_pos : t -> positive -> t.
-
- Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
-
+ Parameter power : t -> N -> t.
Parameter sqrt : t -> t.
-
- Parameter spec_sqrt: forall x, 0 <= [x] ->
- [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
-
Parameter div_eucl : t -> t -> t * t.
-
- Parameter spec_div_eucl: forall x y, [y] <> 0 ->
- let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y].
-
Parameter div : t -> t -> t.
-
- Parameter spec_div: forall x y, [y] <> 0 -> [div x y] = [x] / [y].
-
Parameter modulo : t -> t -> t.
-
- Parameter spec_modulo: forall x y, [y] <> 0 ->
- [modulo x y] = [x] mod [y].
-
Parameter gcd : t -> t -> t.
+ Parameter sgn : t -> t.
+ Parameter abs : t -> t.
+ Parameter spec_compare: forall x y, compare x y = Zcompare [x] [y].
+ Parameter spec_eq_bool: forall x y, eq_bool x y = Zeq_bool [x] [y].
+ Parameter spec_min : forall x y, [min x y] = Zmin [x] [y].
+ Parameter spec_max : forall x y, [max x y] = Zmax [x] [y].
+ Parameter spec_0: [zero] = 0.
+ Parameter spec_1: [one] = 1.
+ Parameter spec_m1: [minus_one] = -1.
+ Parameter spec_succ: forall n, [succ n] = [n] + 1.
+ Parameter spec_add: forall x y, [add x y] = [x] + [y].
+ Parameter spec_pred: forall x, [pred x] = [x] - 1.
+ Parameter spec_sub: forall x y, [sub x y] = [x] - [y].
+ Parameter spec_opp: forall x, [opp x] = - [x].
+ Parameter spec_mul: forall x y, [mul x y] = [x] * [y].
+ Parameter spec_square: forall x, [square x] = [x] * [x].
+ Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
+ Parameter spec_power: forall x n, [power x n] = [x] ^ Z_of_N n.
+ Parameter spec_sqrt: forall x, 0 <= [x] ->
+ [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
+ Parameter spec_div_eucl: forall x y,
+ let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y].
+ Parameter spec_div: forall x y, [div x y] = [x] / [y].
+ Parameter spec_modulo: forall x y, [modulo x y] = [x] mod [y].
Parameter spec_gcd: forall a b, [gcd a b] = Zgcd (to_Z a) (to_Z b).
+ Parameter spec_sgn : forall x, [sgn x] = Zsgn [x].
+ Parameter spec_abs : forall x, [abs x] = Zabs [x].
End ZType.
+
+Module Type ZType_Notation (Import Z:ZType).
+ Notation "[ x ]" := (to_Z x).
+ Infix "==" := eq (at level 70).
+ Notation "0" := zero.
+ Infix "+" := add.
+ Infix "-" := sub.
+ Infix "*" := mul.
+ Notation "- x" := (opp x).
+ Infix "<=" := le.
+ Infix "<" := lt.
+End ZType_Notation.
+
+Module Type ZType' := ZType <+ ZType_Notation. \ No newline at end of file
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
index aceb8984..bcecb6a8 100644
--- a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
+++ b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
@@ -6,119 +6,74 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: ZSigZAxioms.v 11282 2008-07-28 11:51:53Z msozeau $ i*)
+(*i $Id$ i*)
-Require Import ZArith.
-Require Import ZAxioms.
-Require Import ZSig.
+Require Import ZArith ZAxioms ZDivFloor ZSig.
-(** * The interface [ZSig.ZType] implies the interface [ZAxiomsSig] *)
+(** * The interface [ZSig.ZType] implies the interface [ZAxiomsSig]
-Module ZSig_ZAxioms (Z:ZType) <: ZAxiomsSig.
+ It also provides [sgn], [abs], [div], [mod]
+*)
-Delimit Scope IntScope with Int.
-Bind Scope IntScope with Z.t.
-Open Local Scope IntScope.
-Notation "[ x ]" := (Z.to_Z x) : IntScope.
-Infix "==" := Z.eq (at level 70) : IntScope.
-Notation "0" := Z.zero : IntScope.
-Infix "+" := Z.add : IntScope.
-Infix "-" := Z.sub : IntScope.
-Infix "*" := Z.mul : IntScope.
-Notation "- x" := (Z.opp x) : IntScope.
-Hint Rewrite
- Z.spec_0 Z.spec_1 Z.spec_add Z.spec_sub Z.spec_pred Z.spec_succ
- Z.spec_mul Z.spec_opp Z.spec_of_Z : Zspec.
+Module ZTypeIsZAxioms (Import Z : ZType').
-Ltac zsimpl := unfold Z.eq in *; autorewrite with Zspec.
+Hint Rewrite
+ spec_0 spec_1 spec_add spec_sub spec_pred spec_succ
+ spec_mul spec_opp spec_of_Z spec_div spec_modulo
+ spec_compare spec_eq_bool spec_max spec_min spec_abs spec_sgn
+ : zsimpl.
-Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig.
-Module Export NZAxiomsMod <: NZAxiomsSig.
+Ltac zsimpl := autorewrite with zsimpl.
+Ltac zcongruence := repeat red; intros; zsimpl; congruence.
+Ltac zify := unfold eq, lt, le in *; zsimpl.
-Definition NZ := Z.t.
-Definition NZeq := Z.eq.
-Definition NZ0 := Z.zero.
-Definition NZsucc := Z.succ.
-Definition NZpred := Z.pred.
-Definition NZadd := Z.add.
-Definition NZsub := Z.sub.
-Definition NZmul := Z.mul.
+Instance eq_equiv : Equivalence eq.
+Proof. unfold eq. firstorder. Qed.
-Theorem NZeq_equiv : equiv Z.t Z.eq.
-Proof.
-repeat split; repeat red; intros; auto; congruence.
-Qed.
-
-Add Relation Z.t Z.eq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
- as NZeq_rel.
-
-Add Morphism NZsucc with signature Z.eq ==> Z.eq as NZsucc_wd.
-Proof.
-intros; zsimpl; f_equal; assumption.
-Qed.
-
-Add Morphism NZpred with signature Z.eq ==> Z.eq as NZpred_wd.
-Proof.
-intros; zsimpl; f_equal; assumption.
-Qed.
+Local Obligation Tactic := zcongruence.
-Add Morphism NZadd with signature Z.eq ==> Z.eq ==> Z.eq as NZadd_wd.
-Proof.
-intros; zsimpl; f_equal; assumption.
-Qed.
-
-Add Morphism NZsub with signature Z.eq ==> Z.eq ==> Z.eq as NZsub_wd.
-Proof.
-intros; zsimpl; f_equal; assumption.
-Qed.
+Program Instance succ_wd : Proper (eq ==> eq) succ.
+Program Instance pred_wd : Proper (eq ==> eq) pred.
+Program Instance add_wd : Proper (eq ==> eq ==> eq) add.
+Program Instance sub_wd : Proper (eq ==> eq ==> eq) sub.
+Program Instance mul_wd : Proper (eq ==> eq ==> eq) mul.
-Add Morphism NZmul with signature Z.eq ==> Z.eq ==> Z.eq as NZmul_wd.
+Theorem pred_succ : forall n, pred (succ n) == n.
Proof.
-intros; zsimpl; f_equal; assumption.
-Qed.
-
-Theorem NZpred_succ : forall n, Z.pred (Z.succ n) == n.
-Proof.
-intros; zsimpl; auto with zarith.
+intros. zify. auto with zarith.
Qed.
Section Induction.
Variable A : Z.t -> Prop.
-Hypothesis A_wd : predicate_wd Z.eq A.
+Hypothesis A_wd : Proper (eq==>iff) A.
Hypothesis A0 : A 0.
-Hypothesis AS : forall n, A n <-> A (Z.succ n).
-
-Add Morphism A with signature Z.eq ==> iff as A_morph.
-Proof. apply A_wd. Qed.
+Hypothesis AS : forall n, A n <-> A (succ n).
-Let B (z : Z) := A (Z.of_Z z).
+Let B (z : Z) := A (of_Z z).
Lemma B0 : B 0.
Proof.
unfold B; simpl.
rewrite <- (A_wd 0); auto.
-zsimpl; auto.
+zify. auto.
Qed.
Lemma BS : forall z : Z, B z -> B (z + 1).
Proof.
intros z H.
unfold B in *. apply -> AS in H.
-setoid_replace (Z.of_Z (z + 1)) with (Z.succ (Z.of_Z z)); auto.
-zsimpl; auto.
+setoid_replace (of_Z (z + 1)) with (succ (of_Z z)); auto.
+zify. auto.
Qed.
Lemma BP : forall z : Z, B z -> B (z - 1).
Proof.
intros z H.
unfold B in *. rewrite AS.
-setoid_replace (Z.succ (Z.of_Z (z - 1))) with (Z.of_Z z); auto.
-zsimpl; auto with zarith.
+setoid_replace (succ (of_Z (z - 1))) with (of_Z z); auto.
+zify. auto with zarith.
Qed.
Lemma B_holds : forall z : Z, B z.
@@ -135,172 +90,170 @@ intros; rewrite Zopp_succ; unfold Zpred; apply BP; auto.
subst z'; auto with zarith.
Qed.
-Theorem NZinduction : forall n, A n.
+Theorem bi_induction : forall n, A n.
Proof.
-intro n. setoid_replace n with (Z.of_Z (Z.to_Z n)).
+intro n. setoid_replace n with (of_Z (to_Z n)).
apply B_holds.
-zsimpl; auto.
+zify. auto.
Qed.
End Induction.
-Theorem NZadd_0_l : forall n, 0 + n == n.
+Theorem add_0_l : forall n, 0 + n == n.
Proof.
-intros; zsimpl; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZadd_succ_l : forall n m, (Z.succ n) + m == Z.succ (n + m).
+Theorem add_succ_l : forall n m, (succ n) + m == succ (n + m).
Proof.
-intros; zsimpl; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZsub_0_r : forall n, n - 0 == n.
+Theorem sub_0_r : forall n, n - 0 == n.
Proof.
-intros; zsimpl; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZsub_succ_r : forall n m, n - (Z.succ m) == Z.pred (n - m).
+Theorem sub_succ_r : forall n m, n - (succ m) == pred (n - m).
Proof.
-intros; zsimpl; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZmul_0_l : forall n, 0 * n == 0.
+Theorem mul_0_l : forall n, 0 * n == 0.
Proof.
-intros; zsimpl; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZmul_succ_l : forall n m, (Z.succ n) * m == n * m + m.
+Theorem mul_succ_l : forall n m, (succ n) * m == n * m + m.
Proof.
-intros; zsimpl; ring.
+intros. zify. ring.
Qed.
-End NZAxiomsMod.
+(** Order *)
-Definition NZlt := Z.lt.
-Definition NZle := Z.le.
-Definition NZmin := Z.min.
-Definition NZmax := Z.max.
+Lemma compare_spec : forall x y, CompSpec eq lt x y (compare x y).
+Proof.
+ intros. zify. destruct (Zcompare_spec [x] [y]); auto.
+Qed.
-Infix "<=" := Z.le : IntScope.
-Infix "<" := Z.lt : IntScope.
+Definition eqb := eq_bool.
-Lemma spec_compare_alt : forall x y, Z.compare x y = ([x] ?= [y])%Z.
+Lemma eqb_eq : forall x y, eq_bool x y = true <-> x == y.
Proof.
- intros; generalize (Z.spec_compare x y).
- destruct (Z.compare x y); auto.
- intros H; rewrite H; symmetry; apply Zcompare_refl.
+ intros. zify. symmetry. apply Zeq_is_eq_bool.
Qed.
-Lemma spec_lt : forall x y, (x<y) <-> ([x]<[y])%Z.
+Instance compare_wd : Proper (eq ==> eq ==> Logic.eq) compare.
Proof.
- intros; unfold Z.lt, Zlt; rewrite spec_compare_alt; intuition.
+intros x x' Hx y y' Hy. rewrite 2 spec_compare, Hx, Hy; intuition.
Qed.
-Lemma spec_le : forall x y, (x<=y) <-> ([x]<=[y])%Z.
+Instance lt_wd : Proper (eq ==> eq ==> iff) lt.
Proof.
- intros; unfold Z.le, Zle; rewrite spec_compare_alt; intuition.
+intros x x' Hx y y' Hy; unfold lt; rewrite Hx, Hy; intuition.
Qed.
-Lemma spec_min : forall x y, [Z.min x y] = Zmin [x] [y].
+Theorem lt_eq_cases : forall n m, n <= m <-> n < m \/ n == m.
Proof.
- intros; unfold Z.min, Zmin.
- rewrite spec_compare_alt; destruct Zcompare; auto.
+intros. zify. omega.
Qed.
-Lemma spec_max : forall x y, [Z.max x y] = Zmax [x] [y].
+Theorem lt_irrefl : forall n, ~ n < n.
Proof.
- intros; unfold Z.max, Zmax.
- rewrite spec_compare_alt; destruct Zcompare; auto.
+intros. zify. omega.
Qed.
-Add Morphism Z.compare with signature Z.eq ==> Z.eq ==> (@eq comparison) as compare_wd.
-Proof.
-intros x x' Hx y y' Hy.
-rewrite 2 spec_compare_alt; unfold Z.eq in *; rewrite Hx, Hy; intuition.
+Theorem lt_succ_r : forall n m, n < (succ m) <-> n <= m.
+Proof.
+intros. zify. omega.
Qed.
-Add Morphism Z.lt with signature Z.eq ==> Z.eq ==> iff as NZlt_wd.
+Theorem min_l : forall n m, n <= m -> min n m == n.
Proof.
-intros x x' Hx y y' Hy; unfold Z.lt; rewrite Hx, Hy; intuition.
+intros n m. zify. omega with *.
Qed.
-Add Morphism Z.le with signature Z.eq ==> Z.eq ==> iff as NZle_wd.
+Theorem min_r : forall n m, m <= n -> min n m == m.
Proof.
-intros x x' Hx y y' Hy; unfold Z.le; rewrite Hx, Hy; intuition.
+intros n m. zify. omega with *.
Qed.
-Add Morphism Z.min with signature Z.eq ==> Z.eq ==> Z.eq as NZmin_wd.
+Theorem max_l : forall n m, m <= n -> max n m == n.
Proof.
-intros; red; rewrite 2 spec_min; congruence.
+intros n m. zify. omega with *.
Qed.
-Add Morphism Z.max with signature Z.eq ==> Z.eq ==> Z.eq as NZmax_wd.
+Theorem max_r : forall n m, n <= m -> max n m == m.
Proof.
-intros; red; rewrite 2 spec_max; congruence.
+intros n m. zify. omega with *.
Qed.
-Theorem NZlt_eq_cases : forall n m, n <= m <-> n < m \/ n == m.
+(** Part specific to integers, not natural numbers *)
+
+Program Instance opp_wd : Proper (eq ==> eq) opp.
+
+Theorem succ_pred : forall n, succ (pred n) == n.
Proof.
-intros.
-unfold Z.eq; rewrite spec_lt, spec_le; omega.
+intros. zify. auto with zarith.
Qed.
-Theorem NZlt_irrefl : forall n, ~ n < n.
+Theorem opp_0 : - 0 == 0.
Proof.
-intros; rewrite spec_lt; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZlt_succ_r : forall n m, n < (Z.succ m) <-> n <= m.
+Theorem opp_succ : forall n, - (succ n) == pred (- n).
Proof.
-intros; rewrite spec_lt, spec_le, Z.spec_succ; omega.
+intros. zify. auto with zarith.
Qed.
-Theorem NZmin_l : forall n m, n <= m -> Z.min n m == n.
+Theorem abs_eq : forall n, 0 <= n -> abs n == n.
Proof.
-intros n m; unfold Z.eq; rewrite spec_le, spec_min.
-generalize (Zmin_spec [n] [m]); omega.
+intros n. zify. omega with *.
Qed.
-Theorem NZmin_r : forall n m, m <= n -> Z.min n m == m.
+Theorem abs_neq : forall n, n <= 0 -> abs n == -n.
Proof.
-intros n m; unfold Z.eq; rewrite spec_le, spec_min.
-generalize (Zmin_spec [n] [m]); omega.
+intros n. zify. omega with *.
Qed.
-Theorem NZmax_l : forall n m, m <= n -> Z.max n m == n.
+Theorem sgn_null : forall n, n==0 -> sgn n == 0.
Proof.
-intros n m; unfold Z.eq; rewrite spec_le, spec_max.
-generalize (Zmax_spec [n] [m]); omega.
+intros n. zify. omega with *.
Qed.
-Theorem NZmax_r : forall n m, n <= m -> Z.max n m == m.
+Theorem sgn_pos : forall n, 0<n -> sgn n == succ 0.
Proof.
-intros n m; unfold Z.eq; rewrite spec_le, spec_max.
-generalize (Zmax_spec [n] [m]); omega.
+intros n. zify. omega with *.
Qed.
-End NZOrdAxiomsMod.
-
-Definition Zopp := Z.opp.
-
-Add Morphism Z.opp with signature Z.eq ==> Z.eq as Zopp_wd.
+Theorem sgn_neg : forall n, n<0 -> sgn n == opp (succ 0).
Proof.
-intros; zsimpl; auto with zarith.
+intros n. zify. omega with *.
Qed.
-Theorem Zsucc_pred : forall n, Z.succ (Z.pred n) == n.
+Program Instance div_wd : Proper (eq==>eq==>eq) div.
+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.
-red; intros; zsimpl; auto with zarith.
+intros a b. zify. intros. apply Z_div_mod_eq_full; auto.
Qed.
-Theorem Zopp_0 : - 0 == 0.
+Theorem mod_pos_bound :
+ forall a b, 0 < b -> 0 <= modulo a b /\ modulo a b < b.
Proof.
-red; intros; zsimpl; auto with zarith.
+intros a b. zify. intros. apply Z_mod_lt; auto with zarith.
Qed.
-Theorem Zopp_succ : forall n, - (Z.succ n) == Z.pred (- n).
+Theorem mod_neg_bound :
+ forall a b, b < 0 -> b < modulo a b /\ modulo a b <= 0.
Proof.
-intros; zsimpl; auto with zarith.
+intros a b. zify. intros. apply Z_mod_neg; auto with zarith.
Qed.
-End ZSig_ZAxioms.
+End ZTypeIsZAxioms.
+
+Module ZType_ZAxioms (Z : ZType)
+ <: ZAxiomsSig <: ZDivSig <: HasCompare Z <: HasEqBool Z <: HasMinMax Z
+ := Z <+ ZTypeIsZAxioms.
diff --git a/theories/Numbers/NaryFunctions.v b/theories/Numbers/NaryFunctions.v
index 04a48d51..417463eb 100644
--- a/theories/Numbers/NaryFunctions.v
+++ b/theories/Numbers/NaryFunctions.v
@@ -8,27 +8,27 @@
(* Pierre Letouzey, Jerome Vouillon, PPS, Paris 7, 2008 *)
(************************************************************************)
-(*i $Id: NaryFunctions.v 10967 2008-05-22 12:59:38Z letouzey $ i*)
+(*i $Id$ i*)
-Open Local Scope type_scope.
+Local Open Scope type_scope.
Require Import List.
(** * Generic dependently-typed operators about [n]-ary functions *)
-(** The type of [n]-ary function: [nfun A n B] is
+(** The type of [n]-ary function: [nfun A n B] is
[A -> ... -> A -> B] with [n] occurences of [A] in this type. *)
-Fixpoint nfun A n B :=
+Fixpoint nfun A n B :=
match n with
- | O => B
+ | O => B
| S n => A -> (nfun A n B)
- end.
+ end.
Notation " A ^^ n --> B " := (nfun A n B)
(at level 50, n at next level) : type_scope.
-(** [napply_cst _ _ a n f] iterates [n] times the application of a
+(** [napply_cst _ _ a n f] iterates [n] times the application of a
particular constant [a] to the [n]-ary function [f]. *)
Fixpoint napply_cst (A B:Type)(a:A) n : (A^^n-->B) -> B :=
@@ -40,47 +40,47 @@ Fixpoint napply_cst (A B:Type)(a:A) n : (A^^n-->B) -> B :=
(** A generic transformation from an n-ary function to another one.*)
-Fixpoint nfun_to_nfun (A B C:Type)(f:B -> C) n :
+Fixpoint nfun_to_nfun (A B C:Type)(f:B -> C) n :
(A^^n-->B) -> (A^^n-->C) :=
- match n return (A^^n-->B) -> (A^^n-->C) with
+ match n return (A^^n-->B) -> (A^^n-->C) with
| O => f
| S n => fun g a => nfun_to_nfun _ _ _ f n (g a)
end.
-(** [napply_except_last _ _ n f] expects [n] arguments of type [A],
- applies [n-1] of them to [f] and discard the last one. *)
+(** [napply_except_last _ _ n f] expects [n] arguments of type [A],
+ applies [n-1] of them to [f] and discard the last one. *)
-Definition napply_except_last (A B:Type) :=
+Definition napply_except_last (A B:Type) :=
nfun_to_nfun A B (A->B) (fun b a => b).
-(** [napply_then_last _ _ a n f] expects [n] arguments of type [A],
- applies them to [f] and then apply [a] to the result. *)
+(** [napply_then_last _ _ a n f] expects [n] arguments of type [A],
+ applies them to [f] and then apply [a] to the result. *)
-Definition napply_then_last (A B:Type)(a:A) :=
+Definition napply_then_last (A B:Type)(a:A) :=
nfun_to_nfun A (A->B) B (fun fab => fab a).
-(** [napply_discard _ b n] expects [n] arguments, discards then,
+(** [napply_discard _ b n] expects [n] arguments, discards then,
and returns [b]. *)
Fixpoint napply_discard (A B:Type)(b:B) n : A^^n-->B :=
- match n return A^^n-->B with
+ match n return A^^n-->B with
| O => b
| S n => fun _ => napply_discard _ _ b n
end.
(** A fold function *)
-Fixpoint nfold A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
- match n return (A^^n-->B) with
+Fixpoint nfold A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
+ match n return (A^^n-->B) with
| O => b
| S n => fun a => (nfold _ _ f (f a b) n)
end.
-(** [n]-ary products : [nprod A n] is [A*...*A*unit],
+(** [n]-ary products : [nprod A n] is [A*...*A*unit],
with [n] occurrences of [A] in this type. *)
-Fixpoint nprod A n : Type := match n with
+Fixpoint nprod A n : Type := match n with
| O => unit
| S n => (A * nprod A n)%type
end.
@@ -89,54 +89,54 @@ Notation "A ^ n" := (nprod A n) : type_scope.
(** [n]-ary curryfication / uncurryfication *)
-Fixpoint ncurry (A B:Type) n : (A^n -> B) -> (A^^n-->B) :=
- match n return (A^n -> B) -> (A^^n-->B) with
+Fixpoint ncurry (A B:Type) n : (A^n -> B) -> (A^^n-->B) :=
+ match n return (A^n -> B) -> (A^^n-->B) with
| O => fun x => x tt
| S n => fun f a => ncurry _ _ n (fun p => f (a,p))
end.
-Fixpoint nuncurry (A B:Type) n : (A^^n-->B) -> (A^n -> B) :=
+Fixpoint nuncurry (A B:Type) n : (A^^n-->B) -> (A^n -> B) :=
match n return (A^^n-->B) -> (A^n -> B) with
| O => fun x _ => x
| S n => fun f p => let (x,p) := p in nuncurry _ _ n (f x) p
end.
-(** Earlier functions can also be defined via [ncurry/nuncurry].
+(** Earlier functions can also be defined via [ncurry/nuncurry].
For instance : *)
Definition nfun_to_nfun_bis A B C (f:B->C) n :
- (A^^n-->B) -> (A^^n-->C) :=
+ (A^^n-->B) -> (A^^n-->C) :=
fun anb => ncurry _ _ n (fun an => f ((nuncurry _ _ n anb) an)).
-(** We can also us it to obtain another [fold] function,
+(** We can also us it to obtain another [fold] function,
equivalent to the previous one, but with a nicer expansion
(see for instance Int31.iszero). *)
-Fixpoint nfold_bis A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
- match n return (A^^n-->B) with
+Fixpoint nfold_bis A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
+ match n return (A^^n-->B) with
| O => b
- | S n => fun a =>
+ | S n => fun a =>
nfun_to_nfun_bis _ _ _ (f a) n (nfold_bis _ _ f b n)
end.
(** From [nprod] to [list] *)
-Fixpoint nprod_to_list (A:Type) n : A^n -> list A :=
- match n with
+Fixpoint nprod_to_list (A:Type) n : A^n -> list A :=
+ match n with
| O => fun _ => nil
| S n => fun p => let (x,p) := p in x::(nprod_to_list _ n p)
end.
(** From [list] to [nprod] *)
-Fixpoint nprod_of_list (A:Type)(l:list A) : A^(length l) :=
- match l return A^(length l) with
+Fixpoint nprod_of_list (A:Type)(l:list A) : A^(length l) :=
+ match l return A^(length l) with
| nil => tt
| x::l => (x, nprod_of_list _ l)
end.
(** This gives an additional way to write the fold *)
-Definition nfold_list (A B:Type)(f:A->B->B)(b:B) n : (A^^n-->B) :=
+Definition nfold_list (A B:Type)(f:A->B->B)(b:B) n : (A^^n-->B) :=
ncurry _ _ n (fun p => fold_right f b (nprod_to_list _ _ p)).
diff --git a/theories/Numbers/NatInt/NZAdd.v b/theories/Numbers/NatInt/NZAdd.v
index c9bb5c95..9535cfdb 100644
--- a/theories/Numbers/NatInt/NZAdd.v
+++ b/theories/Numbers/NatInt/NZAdd.v
@@ -8,84 +8,83 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZAdd.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import NZAxioms.
-Require Import NZBase.
+Require Import NZAxioms NZBase.
-Module NZAddPropFunct (Import NZAxiomsMod : NZAxiomsSig).
-Module Export NZBasePropMod := NZBasePropFunct NZAxiomsMod.
-Open Local Scope NatIntScope.
+Module Type NZAddPropSig
+ (Import NZ : NZAxiomsSig')(Import NZBase : NZBasePropSig NZ).
-Theorem NZadd_0_r : forall n : NZ, n + 0 == n.
+Hint Rewrite
+ pred_succ add_0_l add_succ_l mul_0_l mul_succ_l sub_0_r sub_succ_r : nz.
+Ltac nzsimpl := autorewrite with nz.
+
+Theorem add_0_r : forall n, n + 0 == n.
Proof.
-NZinduct n. now rewrite NZadd_0_l.
-intro. rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd.
+nzinduct n. now nzsimpl.
+intro. nzsimpl. now rewrite succ_inj_wd.
Qed.
-Theorem NZadd_succ_r : forall n m : NZ, n + S m == S (n + m).
+Theorem add_succ_r : forall n m, n + S m == S (n + m).
Proof.
-intros n m; NZinduct n.
-now do 2 rewrite NZadd_0_l.
-intro. repeat rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd.
+intros n m; nzinduct n. now nzsimpl.
+intro. nzsimpl. now rewrite succ_inj_wd.
Qed.
-Theorem NZadd_comm : forall n m : NZ, n + m == m + n.
+Hint Rewrite add_0_r add_succ_r : nz.
+
+Theorem add_comm : forall n m, n + m == m + n.
Proof.
-intros n m; NZinduct n.
-rewrite NZadd_0_l; now rewrite NZadd_0_r.
-intros n. rewrite NZadd_succ_l; rewrite NZadd_succ_r. now rewrite NZsucc_inj_wd.
+intros n m; nzinduct n. now nzsimpl.
+intro. nzsimpl. now rewrite succ_inj_wd.
Qed.
-Theorem NZadd_1_l : forall n : NZ, 1 + n == S n.
+Theorem add_1_l : forall n, 1 + n == S n.
Proof.
-intro n; rewrite NZadd_succ_l; now rewrite NZadd_0_l.
+intro n; now nzsimpl.
Qed.
-Theorem NZadd_1_r : forall n : NZ, n + 1 == S n.
+Theorem add_1_r : forall n, n + 1 == S n.
Proof.
-intro n; rewrite NZadd_comm; apply NZadd_1_l.
+intro n; now nzsimpl.
Qed.
-Theorem NZadd_assoc : forall n m p : NZ, n + (m + p) == (n + m) + p.
+Theorem add_assoc : forall n m p, n + (m + p) == (n + m) + p.
Proof.
-intros n m p; NZinduct n.
-now do 2 rewrite NZadd_0_l.
-intro. do 3 rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd.
+intros n m p; nzinduct n. now nzsimpl.
+intro. nzsimpl. now rewrite succ_inj_wd.
Qed.
-Theorem NZadd_shuffle1 : forall n m p q : NZ, (n + m) + (p + q) == (n + p) + (m + q).
+Theorem add_cancel_l : forall n m p, p + n == p + m <-> n == m.
Proof.
-intros n m p q.
-rewrite <- (NZadd_assoc n m (p + q)). rewrite (NZadd_comm m (p + q)).
-rewrite <- (NZadd_assoc p q m). rewrite (NZadd_assoc n p (q + m)).
-now rewrite (NZadd_comm q m).
+intros n m p; nzinduct p. now nzsimpl.
+intro p. nzsimpl. now rewrite succ_inj_wd.
Qed.
-Theorem NZadd_shuffle2 : forall n m p q : NZ, (n + m) + (p + q) == (n + q) + (m + p).
+Theorem add_cancel_r : forall n m p, n + p == m + p <-> n == m.
Proof.
-intros n m p q.
-rewrite <- (NZadd_assoc n m (p + q)). rewrite (NZadd_assoc m p q).
-rewrite (NZadd_comm (m + p) q). now rewrite <- (NZadd_assoc n q (m + p)).
+intros n m p. rewrite (add_comm n p), (add_comm m p). apply add_cancel_l.
Qed.
-Theorem NZadd_cancel_l : forall n m p : NZ, p + n == p + m <-> n == m.
+Theorem add_shuffle0 : forall n m p, n+m+p == n+p+m.
Proof.
-intros n m p; NZinduct p.
-now do 2 rewrite NZadd_0_l.
-intro p. do 2 rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd.
+intros n m p. rewrite <- 2 add_assoc, add_cancel_l. apply add_comm.
Qed.
-Theorem NZadd_cancel_r : forall n m p : NZ, n + p == m + p <-> n == m.
+Theorem add_shuffle1 : forall n m p q, (n + m) + (p + q) == (n + p) + (m + q).
Proof.
-intros n m p. rewrite (NZadd_comm n p); rewrite (NZadd_comm m p).
-apply NZadd_cancel_l.
+intros n m p q. rewrite 2 add_assoc, add_cancel_r. apply add_shuffle0.
Qed.
-Theorem NZsub_1_r : forall n : NZ, n - 1 == P n.
+Theorem add_shuffle2 : forall n m p q, (n + m) + (p + q) == (n + q) + (m + p).
Proof.
-intro n; rewrite NZsub_succ_r; now rewrite NZsub_0_r.
+intros n m p q.
+rewrite 2 add_assoc, add_shuffle0, add_cancel_r. apply add_shuffle0.
Qed.
-End NZAddPropFunct.
+Theorem sub_1_r : forall n, n - 1 == P n.
+Proof.
+intro n; now nzsimpl.
+Qed.
+End NZAddPropSig.
diff --git a/theories/Numbers/NatInt/NZAddOrder.v b/theories/Numbers/NatInt/NZAddOrder.v
index 50d1c42f..97c12202 100644
--- a/theories/Numbers/NatInt/NZAddOrder.v
+++ b/theories/Numbers/NatInt/NZAddOrder.v
@@ -8,159 +8,146 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZAddOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import NZAxioms.
-Require Import NZOrder.
+Require Import NZAxioms NZBase NZMul NZOrder.
-Module NZAddOrderPropFunct (Import NZOrdAxiomsMod : NZOrdAxiomsSig).
-Module Export NZOrderPropMod := NZOrderPropFunct NZOrdAxiomsMod.
-Open Local Scope NatIntScope.
+Module Type NZAddOrderPropSig (Import NZ : NZOrdAxiomsSig').
+Include NZBasePropSig NZ <+ NZMulPropSig NZ <+ NZOrderPropSig NZ.
-Theorem NZadd_lt_mono_l : forall n m p : NZ, n < m <-> p + n < p + m.
+Theorem add_lt_mono_l : forall n m p, n < m <-> p + n < p + m.
Proof.
-intros n m p; NZinduct p.
-now do 2 rewrite NZadd_0_l.
-intro p. do 2 rewrite NZadd_succ_l. now rewrite <- NZsucc_lt_mono.
+intros n m p; nzinduct p. now nzsimpl.
+intro p. nzsimpl. now rewrite <- succ_lt_mono.
Qed.
-Theorem NZadd_lt_mono_r : forall n m p : NZ, n < m <-> n + p < m + p.
+Theorem add_lt_mono_r : forall n m p, n < m <-> n + p < m + p.
Proof.
-intros n m p.
-rewrite (NZadd_comm n p); rewrite (NZadd_comm m p); apply NZadd_lt_mono_l.
+intros n m p. rewrite (add_comm n p), (add_comm m p); apply add_lt_mono_l.
Qed.
-Theorem NZadd_lt_mono : forall n m p q : NZ, n < m -> p < q -> n + p < m + q.
+Theorem add_lt_mono : forall n m p q, n < m -> p < q -> n + p < m + q.
Proof.
intros n m p q H1 H2.
-apply NZlt_trans with (m + p);
-[now apply -> NZadd_lt_mono_r | now apply -> NZadd_lt_mono_l].
+apply lt_trans with (m + p);
+[now apply -> add_lt_mono_r | now apply -> add_lt_mono_l].
Qed.
-Theorem NZadd_le_mono_l : forall n m p : NZ, n <= m <-> p + n <= p + m.
+Theorem add_le_mono_l : forall n m p, n <= m <-> p + n <= p + m.
Proof.
-intros n m p; NZinduct p.
-now do 2 rewrite NZadd_0_l.
-intro p. do 2 rewrite NZadd_succ_l. now rewrite <- NZsucc_le_mono.
+intros n m p; nzinduct p. now nzsimpl.
+intro p. nzsimpl. now rewrite <- succ_le_mono.
Qed.
-Theorem NZadd_le_mono_r : forall n m p : NZ, n <= m <-> n + p <= m + p.
+Theorem add_le_mono_r : forall n m p, n <= m <-> n + p <= m + p.
Proof.
-intros n m p.
-rewrite (NZadd_comm n p); rewrite (NZadd_comm m p); apply NZadd_le_mono_l.
+intros n m p. rewrite (add_comm n p), (add_comm m p); apply add_le_mono_l.
Qed.
-Theorem NZadd_le_mono : forall n m p q : NZ, n <= m -> p <= q -> n + p <= m + q.
+Theorem add_le_mono : forall n m p q, n <= m -> p <= q -> n + p <= m + q.
Proof.
intros n m p q H1 H2.
-apply NZle_trans with (m + p);
-[now apply -> NZadd_le_mono_r | now apply -> NZadd_le_mono_l].
+apply le_trans with (m + p);
+[now apply -> add_le_mono_r | now apply -> add_le_mono_l].
Qed.
-Theorem NZadd_lt_le_mono : forall n m p q : NZ, n < m -> p <= q -> n + p < m + q.
+Theorem add_lt_le_mono : forall n m p q, n < m -> p <= q -> n + p < m + q.
Proof.
intros n m p q H1 H2.
-apply NZlt_le_trans with (m + p);
-[now apply -> NZadd_lt_mono_r | now apply -> NZadd_le_mono_l].
+apply lt_le_trans with (m + p);
+[now apply -> add_lt_mono_r | now apply -> add_le_mono_l].
Qed.
-Theorem NZadd_le_lt_mono : forall n m p q : NZ, n <= m -> p < q -> n + p < m + q.
+Theorem add_le_lt_mono : forall n m p q, n <= m -> p < q -> n + p < m + q.
Proof.
intros n m p q H1 H2.
-apply NZle_lt_trans with (m + p);
-[now apply -> NZadd_le_mono_r | now apply -> NZadd_lt_mono_l].
+apply le_lt_trans with (m + p);
+[now apply -> add_le_mono_r | now apply -> add_lt_mono_l].
Qed.
-Theorem NZadd_pos_pos : forall n m : NZ, 0 < n -> 0 < m -> 0 < n + m.
+Theorem add_pos_pos : forall n m, 0 < n -> 0 < m -> 0 < n + m.
Proof.
-intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_lt_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_lt_mono.
Qed.
-Theorem NZadd_pos_nonneg : forall n m : NZ, 0 < n -> 0 <= m -> 0 < n + m.
+Theorem add_pos_nonneg : forall n m, 0 < n -> 0 <= m -> 0 < n + m.
Proof.
-intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_lt_le_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_lt_le_mono.
Qed.
-Theorem NZadd_nonneg_pos : forall n m : NZ, 0 <= n -> 0 < m -> 0 < n + m.
+Theorem add_nonneg_pos : forall n m, 0 <= n -> 0 < m -> 0 < n + m.
Proof.
-intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_le_lt_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_le_lt_mono.
Qed.
-Theorem NZadd_nonneg_nonneg : forall n m : NZ, 0 <= n -> 0 <= m -> 0 <= n + m.
+Theorem add_nonneg_nonneg : forall n m, 0 <= n -> 0 <= m -> 0 <= n + m.
Proof.
-intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_le_mono.
+intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_le_mono.
Qed.
-Theorem NZlt_add_pos_l : forall n m : NZ, 0 < n -> m < n + m.
+Theorem lt_add_pos_l : forall n m, 0 < n -> m < n + m.
Proof.
-intros n m H. apply -> (NZadd_lt_mono_r 0 n m) in H.
-now rewrite NZadd_0_l in H.
+intros n m. rewrite (add_lt_mono_r 0 n m). now nzsimpl.
Qed.
-Theorem NZlt_add_pos_r : forall n m : NZ, 0 < n -> m < m + n.
+Theorem lt_add_pos_r : forall n m, 0 < n -> m < m + n.
Proof.
-intros; rewrite NZadd_comm; now apply NZlt_add_pos_l.
+intros; rewrite add_comm; now apply lt_add_pos_l.
Qed.
-Theorem NZle_lt_add_lt : forall n m p q : NZ, n <= m -> p + m < q + n -> p < q.
+Theorem le_lt_add_lt : forall n m p q, n <= m -> p + m < q + n -> p < q.
Proof.
-intros n m p q H1 H2. destruct (NZle_gt_cases q p); [| assumption].
-pose proof (NZadd_le_mono q p n m H H1) as H3. apply <- NZnle_gt in H2.
-false_hyp H3 H2.
+intros n m p q H1 H2. destruct (le_gt_cases q p); [| assumption].
+contradict H2. rewrite nlt_ge. now apply add_le_mono.
Qed.
-Theorem NZlt_le_add_lt : forall n m p q : NZ, n < m -> p + m <= q + n -> p < q.
+Theorem lt_le_add_lt : forall n m p q, n < m -> p + m <= q + n -> p < q.
Proof.
-intros n m p q H1 H2. destruct (NZle_gt_cases q p); [| assumption].
-pose proof (NZadd_le_lt_mono q p n m H H1) as H3. apply <- NZnle_gt in H3.
-false_hyp H2 H3.
+intros n m p q H1 H2. destruct (le_gt_cases q p); [| assumption].
+contradict H2. rewrite nle_gt. now apply add_le_lt_mono.
Qed.
-Theorem NZle_le_add_le : forall n m p q : NZ, n <= m -> p + m <= q + n -> p <= q.
+Theorem le_le_add_le : forall n m p q, n <= m -> p + m <= q + n -> p <= q.
Proof.
-intros n m p q H1 H2. destruct (NZle_gt_cases p q); [assumption |].
-pose proof (NZadd_lt_le_mono q p n m H H1) as H3. apply <- NZnle_gt in H3.
-false_hyp H2 H3.
+intros n m p q H1 H2. destruct (le_gt_cases p q); [assumption |].
+contradict H2. rewrite nle_gt. now apply add_lt_le_mono.
Qed.
-Theorem NZadd_lt_cases : forall n m p q : NZ, n + m < p + q -> n < p \/ m < q.
+Theorem add_lt_cases : forall n m p q, n + m < p + q -> n < p \/ m < q.
Proof.
intros n m p q H;
-destruct (NZle_gt_cases p n) as [H1 | H1].
-destruct (NZle_gt_cases q m) as [H2 | H2].
-pose proof (NZadd_le_mono p n q m H1 H2) as H3. apply -> NZle_ngt in H3.
-false_hyp H H3.
-now right. now left.
+destruct (le_gt_cases p n) as [H1 | H1]; [| now left].
+destruct (le_gt_cases q m) as [H2 | H2]; [| now right].
+contradict H; rewrite nlt_ge. now apply add_le_mono.
Qed.
-Theorem NZadd_le_cases : forall n m p q : NZ, n + m <= p + q -> n <= p \/ m <= q.
+Theorem add_le_cases : forall n m p q, n + m <= p + q -> n <= p \/ m <= q.
Proof.
intros n m p q H.
-destruct (NZle_gt_cases n p) as [H1 | H1]. now left.
-destruct (NZle_gt_cases m q) as [H2 | H2]. now right.
-assert (H3 : p + q < n + m) by now apply NZadd_lt_mono.
-apply -> NZle_ngt in H. false_hyp H3 H.
+destruct (le_gt_cases n p) as [H1 | H1]. now left.
+destruct (le_gt_cases m q) as [H2 | H2]. now right.
+contradict H; rewrite nle_gt. now apply add_lt_mono.
Qed.
-Theorem NZadd_neg_cases : forall n m : NZ, n + m < 0 -> n < 0 \/ m < 0.
+Theorem add_neg_cases : forall n m, n + m < 0 -> n < 0 \/ m < 0.
Proof.
-intros n m H; apply NZadd_lt_cases; now rewrite NZadd_0_l.
+intros n m H; apply add_lt_cases; now nzsimpl.
Qed.
-Theorem NZadd_pos_cases : forall n m : NZ, 0 < n + m -> 0 < n \/ 0 < m.
+Theorem add_pos_cases : forall n m, 0 < n + m -> 0 < n \/ 0 < m.
Proof.
-intros n m H; apply NZadd_lt_cases; now rewrite NZadd_0_l.
+intros n m H; apply add_lt_cases; now nzsimpl.
Qed.
-Theorem NZadd_nonpos_cases : forall n m : NZ, n + m <= 0 -> n <= 0 \/ m <= 0.
+Theorem add_nonpos_cases : forall n m, n + m <= 0 -> n <= 0 \/ m <= 0.
Proof.
-intros n m H; apply NZadd_le_cases; now rewrite NZadd_0_l.
+intros n m H; apply add_le_cases; now nzsimpl.
Qed.
-Theorem NZadd_nonneg_cases : forall n m : NZ, 0 <= n + m -> 0 <= n \/ 0 <= m.
+Theorem add_nonneg_cases : forall n m, 0 <= n + m -> 0 <= n \/ 0 <= m.
Proof.
-intros n m H; apply NZadd_le_cases; now rewrite NZadd_0_l.
+intros n m H; apply add_le_cases; now nzsimpl.
Qed.
-End NZAddOrderPropFunct.
+End NZAddOrderPropSig.
diff --git a/theories/Numbers/NatInt/NZAxioms.v b/theories/Numbers/NatInt/NZAxioms.v
index 26933646..ee7ee159 100644
--- a/theories/Numbers/NatInt/NZAxioms.v
+++ b/theories/Numbers/NatInt/NZAxioms.v
@@ -5,95 +5,115 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* Evgeny Makarov, INRIA, 2007 *)
-(************************************************************************)
-(*i $Id: NZAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
-
-Require Export NumPrelude.
-
-Module Type NZAxiomsSig.
-
-Parameter Inline NZ : Type.
-Parameter Inline NZeq : NZ -> NZ -> Prop.
-Parameter Inline NZ0 : NZ.
-Parameter Inline NZsucc : NZ -> NZ.
-Parameter Inline NZpred : NZ -> NZ.
-Parameter Inline NZadd : NZ -> NZ -> NZ.
-Parameter Inline NZsub : NZ -> NZ -> NZ.
-Parameter Inline NZmul : NZ -> NZ -> NZ.
-
-(* Unary subtraction (opp) is not defined on natural numbers, so we have
- it for integers only *)
-
-Axiom NZeq_equiv : equiv NZ NZeq.
-Add Relation NZ NZeq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
-as NZeq_rel.
-
-Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd.
-Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd.
-Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd.
-Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd.
-Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd.
-
-Delimit Scope NatIntScope with NatInt.
-Open Local Scope NatIntScope.
-Notation "x == y" := (NZeq x y) (at level 70) : NatIntScope.
-Notation "x ~= y" := (~ NZeq x y) (at level 70) : NatIntScope.
-Notation "0" := NZ0 : NatIntScope.
-Notation S := NZsucc.
-Notation P := NZpred.
-Notation "1" := (S 0) : NatIntScope.
-Notation "x + y" := (NZadd x y) : NatIntScope.
-Notation "x - y" := (NZsub x y) : NatIntScope.
-Notation "x * y" := (NZmul x y) : NatIntScope.
-
-Axiom NZpred_succ : forall n : NZ, P (S n) == n.
-
-Axiom NZinduction :
- forall A : NZ -> Prop, predicate_wd NZeq A ->
- A 0 -> (forall n : NZ, A n <-> A (S n)) -> forall n : NZ, A n.
-
-Axiom NZadd_0_l : forall n : NZ, 0 + n == n.
-Axiom NZadd_succ_l : forall n m : NZ, (S n) + m == S (n + m).
-
-Axiom NZsub_0_r : forall n : NZ, n - 0 == n.
-Axiom NZsub_succ_r : forall n m : NZ, n - (S m) == P (n - m).
-
-Axiom NZmul_0_l : forall n : NZ, 0 * n == 0.
-Axiom NZmul_succ_l : forall n m : NZ, S n * m == n * m + m.
-
-End NZAxiomsSig.
-
-Module Type NZOrdAxiomsSig.
-Declare Module Export NZAxiomsMod : NZAxiomsSig.
-Open Local Scope NatIntScope.
-
-Parameter Inline NZlt : NZ -> NZ -> Prop.
-Parameter Inline NZle : NZ -> NZ -> Prop.
-Parameter Inline NZmin : NZ -> NZ -> NZ.
-Parameter Inline NZmax : NZ -> NZ -> NZ.
-
-Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd.
-Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd.
-Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd.
-Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd.
-
-Notation "x < y" := (NZlt x y) : NatIntScope.
-Notation "x <= y" := (NZle x y) : NatIntScope.
-Notation "x > y" := (NZlt y x) (only parsing) : NatIntScope.
-Notation "x >= y" := (NZle y x) (only parsing) : NatIntScope.
-
-Axiom NZlt_eq_cases : forall n m : NZ, n <= m <-> n < m \/ n == m.
-Axiom NZlt_irrefl : forall n : NZ, ~ (n < n).
-Axiom NZlt_succ_r : forall n m : NZ, n < S m <-> n <= m.
-
-Axiom NZmin_l : forall n m : NZ, n <= m -> NZmin n m == n.
-Axiom NZmin_r : forall n m : NZ, m <= n -> NZmin n m == m.
-Axiom NZmax_l : forall n m : NZ, m <= n -> NZmax n m == n.
-Axiom NZmax_r : forall n m : NZ, n <= m -> NZmax n m == m.
-
-End NZOrdAxiomsSig.
+(** Initial Author : Evgeny Makarov, INRIA, 2007 *)
+
+(*i $Id$ i*)
+
+Require Export Equalities Orders NumPrelude GenericMinMax.
+
+(** Axiomatization of a domain with zero, successor, predecessor,
+ and a bi-directional induction principle. We require [P (S n) = n]
+ but not the other way around, since this domain is meant
+ to be either N or Z. In fact it can be a few other things,
+ for instance [Z/nZ] (See file [NZDomain] for a study of that).
+*)
+
+Module Type ZeroSuccPred (Import T:Typ).
+ Parameter Inline zero : t.
+ Parameters Inline succ pred : t -> t.
+End ZeroSuccPred.
+
+Module Type ZeroSuccPredNotation (T:Typ)(Import NZ:ZeroSuccPred T).
+ Notation "0" := zero.
+ Notation S := succ.
+ Notation P := pred.
+ Notation "1" := (S 0).
+ Notation "2" := (S 1).
+End ZeroSuccPredNotation.
+
+Module Type ZeroSuccPred' (T:Typ) :=
+ ZeroSuccPred T <+ ZeroSuccPredNotation T.
+
+Module Type IsNZDomain (Import E:Eq')(Import NZ:ZeroSuccPred' E).
+ Declare Instance succ_wd : Proper (eq ==> eq) S.
+ Declare Instance pred_wd : Proper (eq ==> eq) P.
+ Axiom pred_succ : forall n, P (S n) == n.
+ Axiom bi_induction :
+ forall A : t -> Prop, Proper (eq==>iff) A ->
+ A 0 -> (forall n, A n <-> A (S n)) -> forall n, A n.
+End IsNZDomain.
+
+Module Type NZDomainSig := EqualityType <+ ZeroSuccPred <+ IsNZDomain.
+Module Type NZDomainSig' := EqualityType' <+ ZeroSuccPred' <+ IsNZDomain.
+
+
+(** Axiomatization of basic operations : [+] [-] [*] *)
+
+Module Type AddSubMul (Import T:Typ).
+ Parameters Inline add sub mul : t -> t -> t.
+End AddSubMul.
+
+Module Type AddSubMulNotation (T:Typ)(Import NZ:AddSubMul T).
+ Notation "x + y" := (add x y).
+ Notation "x - y" := (sub x y).
+ Notation "x * y" := (mul x y).
+End AddSubMulNotation.
+
+Module Type AddSubMul' (T:Typ) := AddSubMul T <+ AddSubMulNotation T.
+
+Module Type IsAddSubMul (Import E:NZDomainSig')(Import NZ:AddSubMul' E).
+ Declare Instance add_wd : Proper (eq ==> eq ==> eq) add.
+ Declare Instance sub_wd : Proper (eq ==> eq ==> eq) sub.
+ Declare Instance mul_wd : Proper (eq ==> eq ==> eq) mul.
+ Axiom add_0_l : forall n, 0 + n == n.
+ Axiom add_succ_l : forall n m, (S n) + m == S (n + m).
+ Axiom sub_0_r : forall n, n - 0 == n.
+ Axiom sub_succ_r : forall n m, n - (S m) == P (n - m).
+ Axiom mul_0_l : forall n, 0 * n == 0.
+ Axiom mul_succ_l : forall n m, S n * m == n * m + m.
+End IsAddSubMul.
+
+Module Type NZBasicFunsSig := NZDomainSig <+ AddSubMul <+ IsAddSubMul.
+Module Type NZBasicFunsSig' := NZDomainSig' <+ AddSubMul' <+IsAddSubMul.
+
+(** Old name for the same interface: *)
+
+Module Type NZAxiomsSig := NZBasicFunsSig.
+Module Type NZAxiomsSig' := NZBasicFunsSig'.
+
+(** Axiomatization of order *)
+
+Module Type NZOrd := NZDomainSig <+ HasLt <+ HasLe.
+Module Type NZOrd' := NZDomainSig' <+ HasLt <+ HasLe <+
+ LtNotation <+ LeNotation <+ LtLeNotation.
+
+Module Type IsNZOrd (Import NZ : NZOrd').
+ Declare Instance lt_wd : Proper (eq ==> eq ==> iff) lt.
+ Axiom lt_eq_cases : forall n m, n <= m <-> n < m \/ n == m.
+ Axiom lt_irrefl : forall n, ~ (n < n).
+ Axiom lt_succ_r : forall n m, n < S m <-> n <= m.
+End IsNZOrd.
+
+(** NB: the compatibility of [le] can be proved later from [lt_wd]
+ and [lt_eq_cases] *)
+
+Module Type NZOrdSig := NZOrd <+ IsNZOrd.
+Module Type NZOrdSig' := NZOrd' <+ IsNZOrd.
+
+(** Everything together : *)
+
+Module Type NZOrdAxiomsSig <: NZBasicFunsSig <: NZOrdSig
+ := NZOrdSig <+ AddSubMul <+ IsAddSubMul <+ HasMinMax.
+Module Type NZOrdAxiomsSig' <: NZOrdAxiomsSig
+ := NZOrdSig' <+ AddSubMul' <+ IsAddSubMul <+ HasMinMax.
+
+
+(** Same, plus a comparison function. *)
+
+Module Type NZDecOrdSig := NZOrdSig <+ HasCompare.
+Module Type NZDecOrdSig' := NZOrdSig' <+ HasCompare.
+
+Module Type NZDecOrdAxiomsSig := NZOrdAxiomsSig <+ HasCompare.
+Module Type NZDecOrdAxiomsSig' := NZOrdAxiomsSig' <+ HasCompare.
+
diff --git a/theories/Numbers/NatInt/NZBase.v b/theories/Numbers/NatInt/NZBase.v
index bd4d6232..18e3b9b9 100644
--- a/theories/Numbers/NatInt/NZBase.v
+++ b/theories/Numbers/NatInt/NZBase.v
@@ -8,45 +8,54 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZBase.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
Require Import NZAxioms.
-Module NZBasePropFunct (Import NZAxiomsMod : NZAxiomsSig).
-Open Local Scope NatIntScope.
+Module Type NZBasePropSig (Import NZ : NZDomainSig').
-Theorem NZneq_sym : forall n m : NZ, n ~= m -> m ~= n.
+Include BackportEq NZ NZ. (** eq_refl, eq_sym, eq_trans *)
+
+Lemma eq_sym_iff : forall x y, x==y <-> y==x.
+Proof.
+intros; split; symmetry; auto.
+Qed.
+
+(* TODO: how register ~= (which is just a notation) as a Symmetric relation,
+ hence allowing "symmetry" tac ? *)
+
+Theorem neq_sym : forall n m, n ~= m -> m ~= n.
Proof.
intros n m H1 H2; symmetry in H2; false_hyp H2 H1.
Qed.
-Theorem NZE_stepl : forall x y z : NZ, x == y -> x == z -> z == y.
+Theorem eq_stepl : forall x y z, x == y -> x == z -> z == y.
Proof.
intros x y z H1 H2; now rewrite <- H1.
Qed.
-Declare Left Step NZE_stepl.
-(* The right step lemma is just the transitivity of NZeq *)
-Declare Right Step (proj1 (proj2 NZeq_equiv)).
+Declare Left Step eq_stepl.
+(* The right step lemma is just the transitivity of eq *)
+Declare Right Step (@Equivalence_Transitive _ _ eq_equiv).
-Theorem NZsucc_inj : forall n1 n2 : NZ, S n1 == S n2 -> n1 == n2.
+Theorem succ_inj : forall n1 n2, S n1 == S n2 -> n1 == n2.
Proof.
intros n1 n2 H.
-apply NZpred_wd in H. now do 2 rewrite NZpred_succ in H.
+apply pred_wd in H. now do 2 rewrite pred_succ in H.
Qed.
(* The following theorem is useful as an equivalence for proving
bidirectional induction steps *)
-Theorem NZsucc_inj_wd : forall n1 n2 : NZ, S n1 == S n2 <-> n1 == n2.
+Theorem succ_inj_wd : forall n1 n2, S n1 == S n2 <-> n1 == n2.
Proof.
intros; split.
-apply NZsucc_inj.
-apply NZsucc_wd.
+apply succ_inj.
+apply succ_wd.
Qed.
-Theorem NZsucc_inj_wd_neg : forall n m : NZ, S n ~= S m <-> n ~= m.
+Theorem succ_inj_wd_neg : forall n m, S n ~= S m <-> n ~= m.
Proof.
-intros; now rewrite NZsucc_inj_wd.
+intros; now rewrite succ_inj_wd.
Qed.
(* We cannot prove that the predecessor is injective, nor that it is
@@ -54,31 +63,27 @@ left-inverse to the successor at this point *)
Section CentralInduction.
-Variable A : predicate NZ.
-
-Hypothesis A_wd : predicate_wd NZeq A.
-
-Add Morphism A with signature NZeq ==> iff as A_morph.
-Proof. apply A_wd. Qed.
+Variable A : predicate t.
+Hypothesis A_wd : Proper (eq==>iff) A.
-Theorem NZcentral_induction :
- forall z : NZ, A z ->
- (forall n : NZ, A n <-> A (S n)) ->
- forall n : NZ, A n.
+Theorem central_induction :
+ forall z, A z ->
+ (forall n, A n <-> A (S n)) ->
+ forall n, A n.
Proof.
-intros z Base Step; revert Base; pattern z; apply NZinduction.
+intros z Base Step; revert Base; pattern z; apply bi_induction.
solve_predicate_wd.
-intro; now apply NZinduction.
+intro; now apply bi_induction.
intro; pose proof (Step n); tauto.
Qed.
End CentralInduction.
-Tactic Notation "NZinduct" ident(n) :=
- induction_maker n ltac:(apply NZinduction).
+Tactic Notation "nzinduct" ident(n) :=
+ induction_maker n ltac:(apply bi_induction).
-Tactic Notation "NZinduct" ident(n) constr(u) :=
- induction_maker n ltac:(apply NZcentral_induction with (z := u)).
+Tactic Notation "nzinduct" ident(n) constr(u) :=
+ induction_maker n ltac:(apply central_induction with (z := u)).
-End NZBasePropFunct.
+End NZBasePropSig.
diff --git a/theories/Numbers/NatInt/NZDiv.v b/theories/Numbers/NatInt/NZDiv.v
new file mode 100644
index 00000000..1f6c615b
--- /dev/null
+++ b/theories/Numbers/NatInt/NZDiv.v
@@ -0,0 +1,542 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Euclidean Division *)
+
+Require Import NZAxioms NZMulOrder.
+
+(** The first signatures will be common to all divisions over NZ, N and Z *)
+
+Module Type DivMod (Import T:Typ).
+ Parameters Inline div modulo : t -> t -> t.
+End DivMod.
+
+Module Type DivModNotation (T:Typ)(Import NZ:DivMod T).
+ Infix "/" := div.
+ Infix "mod" := modulo (at level 40, no associativity).
+End DivModNotation.
+
+Module Type DivMod' (T:Typ) := DivMod T <+ DivModNotation T.
+
+Module Type NZDivCommon (Import NZ : NZAxiomsSig')(Import DM : DivMod' NZ).
+ Declare Instance div_wd : Proper (eq==>eq==>eq) div.
+ Declare Instance mod_wd : Proper (eq==>eq==>eq) modulo.
+ Axiom div_mod : forall a b, b ~= 0 -> a == b*(a/b) + (a mod b).
+End NZDivCommon.
+
+(** The different divisions will only differ in the conditions
+ they impose on [modulo]. For NZ, we only describe behavior
+ on positive numbers.
+
+ NB: This axiom would also be true for N and Z, but redundant.
+*)
+
+Module Type NZDivSpecific (Import NZ : NZOrdAxiomsSig')(Import DM : DivMod' NZ).
+ Axiom mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+End NZDivSpecific.
+
+Module Type NZDiv (NZ:NZOrdAxiomsSig)
+ := DivMod NZ <+ NZDivCommon NZ <+ NZDivSpecific NZ.
+
+Module Type NZDiv' (NZ:NZOrdAxiomsSig) := NZDiv NZ <+ DivModNotation NZ.
+
+Module NZDivPropFunct
+ (Import NZ : NZOrdAxiomsSig')
+ (Import NZP : NZMulOrderPropSig NZ)
+ (Import NZD : NZDiv' NZ)
+.
+
+(** Uniqueness theorems *)
+
+Theorem div_mod_unique :
+ forall b q1 q2 r1 r2, 0<=r1<b -> 0<=r2<b ->
+ b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.
+Proof.
+intros b.
+assert (U : forall q1 q2 r1 r2,
+ b*q1+r1 == b*q2+r2 -> 0<=r1<b -> 0<=r2 -> q1<q2 -> False).
+ intros q1 q2 r1 r2 EQ LT Hr1 Hr2.
+ contradict EQ.
+ apply lt_neq.
+ apply lt_le_trans with (b*q1+b).
+ rewrite <- add_lt_mono_l. tauto.
+ apply le_trans with (b*q2).
+ rewrite mul_comm, <- mul_succ_l, mul_comm.
+ apply mul_le_mono_nonneg_l; intuition; try order.
+ rewrite le_succ_l; auto.
+ rewrite <- (add_0_r (b*q2)) at 1.
+ rewrite <- add_le_mono_l. tauto.
+
+intros q1 q2 r1 r2 Hr1 Hr2 EQ; destruct (lt_trichotomy q1 q2) as [LT|[EQ'|GT]].
+elim (U q1 q2 r1 r2); intuition.
+split; auto. rewrite EQ' in EQ. rewrite add_cancel_l in EQ; auto.
+elim (U q2 q1 r2 r1); intuition.
+Qed.
+
+Theorem div_unique:
+ forall a b q r, 0<=a -> 0<=r<b ->
+ a == b*q + r -> q == a/b.
+Proof.
+intros a b q r Ha (Hb,Hr) EQ.
+destruct (div_mod_unique b q (a/b) r (a mod b)); auto.
+apply mod_bound; order.
+rewrite <- div_mod; order.
+Qed.
+
+Theorem mod_unique:
+ forall a b q r, 0<=a -> 0<=r<b ->
+ a == b*q + r -> r == a mod b.
+Proof.
+intros a b q r Ha (Hb,Hr) EQ.
+destruct (div_mod_unique b q (a/b) r (a mod b)); auto.
+apply mod_bound; order.
+rewrite <- div_mod; order.
+Qed.
+
+
+(** A division by itself returns 1 *)
+
+Lemma div_same : forall a, 0<a -> a/a == 1.
+Proof.
+intros. symmetry.
+apply div_unique with 0; intuition; try order.
+now nzsimpl.
+Qed.
+
+Lemma mod_same : forall a, 0<a -> a mod a == 0.
+Proof.
+intros. symmetry.
+apply mod_unique with 1; intuition; try order.
+now nzsimpl.
+Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, 0<=a<b -> a/b == 0.
+Proof.
+intros. symmetry.
+apply div_unique with a; intuition; try order.
+now nzsimpl.
+Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, 0<=a<b -> a mod b == a.
+Proof.
+intros. symmetry.
+apply mod_unique with 0; intuition; try order.
+now nzsimpl.
+Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, 0<a -> 0/a == 0.
+Proof.
+intros; apply div_small; split; order.
+Qed.
+
+Lemma mod_0_l: forall a, 0<a -> 0 mod a == 0.
+Proof.
+intros; apply mod_small; split; order.
+Qed.
+
+Lemma div_1_r: forall a, 0<=a -> a/1 == a.
+Proof.
+intros. symmetry.
+apply div_unique with 0; try split; try order; try apply lt_0_1.
+now nzsimpl.
+Qed.
+
+Lemma mod_1_r: forall a, 0<=a -> a mod 1 == 0.
+Proof.
+intros. symmetry.
+apply mod_unique with a; try split; try order; try apply lt_0_1.
+now nzsimpl.
+Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof.
+intros; apply div_small; split; auto. apply le_succ_diag_r.
+Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof.
+intros; apply mod_small; split; auto. apply le_succ_diag_r.
+Qed.
+
+Lemma div_mul : forall a b, 0<=a -> 0<b -> (a*b)/b == a.
+Proof.
+intros; symmetry.
+apply div_unique with 0; try split; try order.
+apply mul_nonneg_nonneg; order.
+nzsimpl; apply mul_comm.
+Qed.
+
+Lemma mod_mul : forall a b, 0<=a -> 0<b -> (a*b) mod b == 0.
+Proof.
+intros; symmetry.
+apply mod_unique with a; try split; try order.
+apply mul_nonneg_nonneg; order.
+nzsimpl; apply mul_comm.
+Qed.
+
+
+(** * Order results about mod and div *)
+
+(** A modulo cannot grow beyond its starting point. *)
+
+Theorem mod_le: forall a b, 0<=a -> 0<b -> a mod b <= a.
+Proof.
+intros. destruct (le_gt_cases b a).
+apply le_trans with b; auto.
+apply lt_le_incl. destruct (mod_bound a b); auto.
+rewrite lt_eq_cases; right.
+apply mod_small; auto.
+Qed.
+
+
+(* Division of positive numbers is positive. *)
+
+Lemma div_pos: forall a b, 0<=a -> 0<b -> 0 <= a/b.
+Proof.
+intros.
+rewrite (mul_le_mono_pos_l _ _ b); auto; nzsimpl.
+rewrite (add_le_mono_r _ _ (a mod b)).
+rewrite <- div_mod by order.
+nzsimpl.
+apply mod_le; auto.
+Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof.
+intros a b (Hb,Hab).
+assert (LE : 0 <= a/b) by (apply div_pos; order).
+assert (MOD : a mod b < b) by (destruct (mod_bound a b); order).
+rewrite lt_eq_cases in LE; destruct LE as [LT|EQ]; auto.
+exfalso; revert Hab.
+rewrite (div_mod a b), <-EQ; nzsimpl; order.
+Qed.
+
+Lemma div_small_iff : forall a b, 0<=a -> 0<b -> (a/b==0 <-> a<b).
+Proof.
+intros a b Ha Hb; split; intros Hab.
+destruct (lt_ge_cases a b); auto.
+symmetry in Hab. contradict Hab. apply lt_neq, div_str_pos; auto.
+apply div_small; auto.
+Qed.
+
+Lemma mod_small_iff : forall a b, 0<=a -> 0<b -> (a mod b == a <-> a<b).
+Proof.
+intros a b Ha Hb. split; intros H; auto using mod_small.
+rewrite <- div_small_iff; auto.
+rewrite <- (mul_cancel_l _ _ b) by order.
+rewrite <- (add_cancel_r _ _ (a mod b)).
+rewrite <- div_mod, H by order. now nzsimpl.
+Qed.
+
+Lemma div_str_pos_iff : forall a b, 0<=a -> 0<b -> (0<a/b <-> b<=a).
+Proof.
+intros a b Ha Hb; split; intros Hab.
+destruct (lt_ge_cases a b) as [LT|LE]; auto.
+rewrite <- div_small_iff in LT; order.
+apply div_str_pos; auto.
+Qed.
+
+
+(** As soon as the divisor is strictly greater than 1,
+ the division is strictly decreasing. *)
+
+Lemma div_lt : forall a b, 0<a -> 1<b -> a/b < a.
+Proof.
+intros.
+assert (0 < b) by (apply lt_trans with 1; auto using lt_0_1).
+destruct (lt_ge_cases a b).
+rewrite div_small; try split; order.
+rewrite (div_mod a b) at 2 by order.
+apply lt_le_trans with (b*(a/b)).
+rewrite <- (mul_1_l (a/b)) at 1.
+rewrite <- mul_lt_mono_pos_r; auto.
+apply div_str_pos; auto.
+rewrite <- (add_0_r (b*(a/b))) at 1.
+rewrite <- add_le_mono_l. destruct (mod_bound a b); order.
+Qed.
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono : forall a b c, 0<c -> 0<=a<=b -> a/c <= b/c.
+Proof.
+intros a b c Hc (Ha,Hab).
+rewrite lt_eq_cases in Hab. destruct Hab as [LT|EQ];
+ [|rewrite EQ; order].
+rewrite <- lt_succ_r.
+rewrite (mul_lt_mono_pos_l c) by order.
+nzsimpl.
+rewrite (add_lt_mono_r _ _ (a mod c)).
+rewrite <- div_mod by order.
+apply lt_le_trans with b; auto.
+rewrite (div_mod b c) at 1 by order.
+rewrite <- add_assoc, <- add_le_mono_l.
+apply le_trans with (c+0).
+nzsimpl; destruct (mod_bound b c); order.
+rewrite <- add_le_mono_l. destruct (mod_bound a c); order.
+Qed.
+
+(** The following two properties could be used as specification of div *)
+
+Lemma mul_div_le : forall a b, 0<=a -> 0<b -> b*(a/b) <= a.
+Proof.
+intros.
+rewrite (add_le_mono_r _ _ (a mod b)), <- div_mod by order.
+rewrite <- (add_0_r a) at 1.
+rewrite <- add_le_mono_l. destruct (mod_bound a b); order.
+Qed.
+
+Lemma mul_succ_div_gt : forall a b, 0<=a -> 0<b -> a < b*(S (a/b)).
+Proof.
+intros.
+rewrite (div_mod a b) at 1 by order.
+rewrite (mul_succ_r).
+rewrite <- add_lt_mono_l.
+destruct (mod_bound a b); auto.
+Qed.
+
+
+(** The previous inequality is exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, 0<=a -> 0<b -> (a == b*(a/b) <-> a mod b == 0).
+Proof.
+intros. rewrite (div_mod a b) at 1 by order.
+rewrite <- (add_0_r (b*(a/b))) at 2.
+apply add_cancel_l.
+Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, 0<=a -> 0<b -> a < b*q -> a/b < q.
+Proof.
+intros.
+rewrite (mul_lt_mono_pos_l b) by order.
+apply le_lt_trans with a; auto.
+apply mul_div_le; auto.
+Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, 0<=a -> 0<b -> a <= b*q -> a/b <= q.
+Proof.
+intros.
+rewrite (mul_le_mono_pos_l _ _ b) by order.
+apply le_trans with a; auto.
+apply mul_div_le; auto.
+Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, 0<=a -> 0<b -> b*q <= a -> q <= a/b.
+Proof.
+intros a b q Ha Hb H.
+destruct (lt_ge_cases 0 q).
+rewrite <- (div_mul q b); try order.
+apply div_le_mono; auto.
+rewrite mul_comm; split; auto.
+apply lt_le_incl, mul_pos_pos; auto.
+apply le_trans with 0; auto; apply div_pos; auto.
+Qed.
+
+(** A division respects opposite monotonicity for the divisor *)
+
+Lemma div_le_compat_l: forall p q r, 0<=p -> 0<q<=r ->
+ p/r <= p/q.
+Proof.
+ intros p q r Hp (Hq,Hqr).
+ apply div_le_lower_bound; auto.
+ rewrite (div_mod p r) at 2 by order.
+ apply le_trans with (r*(p/r)).
+ apply mul_le_mono_nonneg_r; try order.
+ apply div_pos; order.
+ rewrite <- (add_0_r (r*(p/r))) at 1.
+ rewrite <- add_le_mono_l. destruct (mod_bound p r); order.
+Qed.
+
+
+(** * Relations between usual operations and mod and div *)
+
+Lemma mod_add : forall a b c, 0<=a -> 0<=a+b*c -> 0<c ->
+ (a + b * c) mod c == a mod c.
+Proof.
+ intros.
+ symmetry.
+ apply mod_unique with (a/c+b); auto.
+ apply mod_bound; auto.
+ rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+ now rewrite mul_comm.
+Qed.
+
+Lemma div_add : forall a b c, 0<=a -> 0<=a+b*c -> 0<c ->
+ (a + b * c) / c == a / c + b.
+Proof.
+ intros.
+ apply (mul_cancel_l _ _ c); try order.
+ apply (add_cancel_r _ _ ((a+b*c) mod c)).
+ rewrite <- div_mod, mod_add by order.
+ rewrite mul_add_distr_l, add_shuffle0, <- div_mod by order.
+ now rewrite mul_comm.
+Qed.
+
+Lemma div_add_l: forall a b c, 0<=c -> 0<=a*b+c -> 0<b ->
+ (a * b + c) / b == a + c / b.
+Proof.
+ intros a b c. rewrite (add_comm _ c), (add_comm a).
+ intros. apply div_add; auto.
+Qed.
+
+(** Cancellations. *)
+
+Lemma div_mul_cancel_r : forall a b c, 0<=a -> 0<b -> 0<c ->
+ (a*c)/(b*c) == a/b.
+Proof.
+ intros.
+ symmetry.
+ apply div_unique with ((a mod b)*c).
+ apply mul_nonneg_nonneg; order.
+ split.
+ apply mul_nonneg_nonneg; destruct (mod_bound a b); order.
+ rewrite <- mul_lt_mono_pos_r; auto. destruct (mod_bound a b); auto.
+ rewrite (div_mod a b) at 1 by order.
+ rewrite mul_add_distr_r.
+ rewrite add_cancel_r.
+ rewrite <- 2 mul_assoc. now rewrite (mul_comm c).
+Qed.
+
+Lemma div_mul_cancel_l : forall a b c, 0<=a -> 0<b -> 0<c ->
+ (c*a)/(c*b) == a/b.
+Proof.
+ intros. rewrite !(mul_comm c); apply div_mul_cancel_r; auto.
+Qed.
+
+Lemma mul_mod_distr_l: forall a b c, 0<=a -> 0<b -> 0<c ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof.
+ intros.
+ rewrite <- (add_cancel_l _ _ ((c*b)* ((c*a)/(c*b)))).
+ rewrite <- div_mod.
+ rewrite div_mul_cancel_l; auto.
+ rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+ apply div_mod; order.
+ rewrite <- neq_mul_0; intuition; order.
+Qed.
+
+Lemma mul_mod_distr_r: forall a b c, 0<=a -> 0<b -> 0<c ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof.
+ intros. rewrite !(mul_comm _ c); now rewrite mul_mod_distr_l.
+Qed.
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, 0<=a -> 0<n ->
+ (a mod n) mod n == a mod n.
+Proof.
+ intros. destruct (mod_bound a n); auto. now rewrite mod_small_iff.
+Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, 0<=a -> 0<=b -> 0<n ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof.
+ intros a b n Ha Hb Hn. symmetry.
+ generalize (mul_nonneg_nonneg _ _ Ha Hb).
+ rewrite (div_mod a n) at 1 2 by order.
+ rewrite add_comm, (mul_comm n), (mul_comm _ b).
+ rewrite mul_add_distr_l, mul_assoc.
+ intros. rewrite mod_add; auto.
+ now rewrite mul_comm.
+ apply mul_nonneg_nonneg; destruct (mod_bound a n); auto.
+Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, 0<=a -> 0<=b -> 0<n ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof.
+ intros. rewrite !(mul_comm a). apply mul_mod_idemp_l; auto.
+Qed.
+
+Theorem mul_mod: forall a b n, 0<=a -> 0<=b -> 0<n ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof.
+ intros. rewrite mul_mod_idemp_l, mul_mod_idemp_r; trivial. reflexivity.
+ now destruct (mod_bound b n).
+Qed.
+
+Lemma add_mod_idemp_l : forall a b n, 0<=a -> 0<=b -> 0<n ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof.
+ intros a b n Ha Hb Hn. symmetry.
+ generalize (add_nonneg_nonneg _ _ Ha Hb).
+ rewrite (div_mod a n) at 1 2 by order.
+ rewrite <- add_assoc, add_comm, mul_comm.
+ intros. rewrite mod_add; trivial. reflexivity.
+ apply add_nonneg_nonneg; auto. destruct (mod_bound a n); auto.
+Qed.
+
+Lemma add_mod_idemp_r : forall a b n, 0<=a -> 0<=b -> 0<n ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof.
+ intros. rewrite !(add_comm a). apply add_mod_idemp_l; auto.
+Qed.
+
+Theorem add_mod: forall a b n, 0<=a -> 0<=b -> 0<n ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof.
+ intros. rewrite add_mod_idemp_l, add_mod_idemp_r; trivial. reflexivity.
+ now destruct (mod_bound b n).
+Qed.
+
+Lemma div_div : forall a b c, 0<=a -> 0<b -> 0<c ->
+ (a/b)/c == a/(b*c).
+Proof.
+ intros a b c Ha Hb Hc.
+ apply div_unique with (b*((a/b) mod c) + a mod b); trivial.
+ (* begin 0<= ... <b*c *)
+ destruct (mod_bound (a/b) c), (mod_bound a b); auto using div_pos.
+ split.
+ apply add_nonneg_nonneg; auto.
+ apply mul_nonneg_nonneg; order.
+ apply lt_le_trans with (b*((a/b) mod c) + b).
+ rewrite <- add_lt_mono_l; auto.
+ rewrite <- mul_succ_r, <- mul_le_mono_pos_l, le_succ_l; auto.
+ (* end 0<= ... < b*c *)
+ rewrite (div_mod a b) at 1 by order.
+ rewrite add_assoc, add_cancel_r.
+ rewrite <- mul_assoc, <- mul_add_distr_l, mul_cancel_l by order.
+ apply div_mod; order.
+Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, 0<=a -> 0<b -> 0<=c -> c*(a/b) <= (c*a)/b.
+Proof.
+ intros.
+ apply div_le_lower_bound; auto.
+ apply mul_nonneg_nonneg; auto.
+ rewrite mul_assoc, (mul_comm b c), <- mul_assoc.
+ apply mul_le_mono_nonneg_l; auto.
+ apply mul_div_le; auto.
+Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, 0<=a -> 0<b ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof.
+ split.
+ intros. exists (a/b). rewrite div_exact; auto.
+ intros (c,Hc). rewrite Hc, mul_comm. apply mod_mul; auto.
+ rewrite (mul_le_mono_pos_l _ _ b); auto. nzsimpl. order.
+Qed.
+
+End NZDivPropFunct.
+
diff --git a/theories/Numbers/NatInt/NZDomain.v b/theories/Numbers/NatInt/NZDomain.v
new file mode 100644
index 00000000..8c3c7937
--- /dev/null
+++ b/theories/Numbers/NatInt/NZDomain.v
@@ -0,0 +1,417 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+Require Export NumPrelude NZAxioms.
+Require Import NZBase NZOrder NZAddOrder Plus Minus.
+
+(** In this file, we investigate the shape of domains satisfying
+ the [NZDomainSig] interface. In particular, we define a
+ translation from Peano numbers [nat] into NZ.
+*)
+
+(** First, a section about iterating a function. *)
+
+Section Iter.
+Variable A : Type.
+Fixpoint iter (f:A->A)(n:nat) : A -> A := fun a =>
+ match n with
+ | O => a
+ | S n => f (iter f n a)
+ end.
+Infix "^" := iter.
+
+Lemma iter_alt : forall f n m, (f^(Datatypes.S n)) m = (f^n) (f m).
+Proof.
+induction n; simpl; auto.
+intros; rewrite <- IHn; auto.
+Qed.
+
+Lemma iter_plus : forall f n n' m, (f^(n+n')) m = (f^n) ((f^n') m).
+Proof.
+induction n; simpl; auto.
+intros; rewrite IHn; auto.
+Qed.
+
+Lemma iter_plus_bis : forall f n n' m, (f^(n+n')) m = (f^n') ((f^n) m).
+Proof.
+induction n; simpl; auto.
+intros. rewrite <- iter_alt, IHn; auto.
+Qed.
+
+Global Instance iter_wd (R:relation A) : Proper ((R==>R)==>eq==>R==>R) iter.
+Proof.
+intros f f' Hf n n' Hn; subst n'. induction n; simpl; red; auto.
+Qed.
+
+End Iter.
+Implicit Arguments iter [A].
+Local Infix "^" := iter.
+
+
+Module NZDomainProp (Import NZ:NZDomainSig').
+
+(** * Relationship between points thanks to [succ] and [pred]. *)
+
+(** We prove that any points in NZ have a common descendant by [succ] *)
+
+Definition common_descendant n m := exists k, exists l, (S^k) n == (S^l) m.
+
+Instance common_descendant_wd : Proper (eq==>eq==>iff) common_descendant.
+Proof.
+unfold common_descendant. intros n n' Hn m m' Hm.
+setoid_rewrite Hn. setoid_rewrite Hm. auto with *.
+Qed.
+
+Instance common_descendant_equiv : Equivalence common_descendant.
+Proof.
+split; red.
+intros x. exists O; exists O. simpl; auto with *.
+intros x y (p & q & H); exists q; exists p; auto with *.
+intros x y z (p & q & Hpq) (r & s & Hrs).
+exists (r+p)%nat. exists (q+s)%nat.
+rewrite !iter_plus. rewrite Hpq, <-Hrs, <-iter_plus, <- iter_plus_bis.
+auto with *.
+Qed.
+
+Lemma common_descendant_with_0 : forall n, common_descendant n 0.
+Proof.
+apply bi_induction.
+intros n n' Hn. rewrite Hn; auto with *.
+reflexivity.
+split; intros (p & q & H).
+exists p; exists (Datatypes.S q). rewrite <- iter_alt; simpl.
+ apply succ_wd; auto.
+exists (Datatypes.S p); exists q. rewrite iter_alt; auto.
+Qed.
+
+Lemma common_descendant_always : forall n m, common_descendant n m.
+Proof.
+intros. transitivity 0; [|symmetry]; apply common_descendant_with_0.
+Qed.
+
+(** Thanks to [succ] being injective, we can then deduce that for any two
+ points, one is an iterated successor of the other. *)
+
+Lemma itersucc_or_itersucc : forall n m, exists k, n == (S^k) m \/ m == (S^k) n.
+Proof.
+intros n m. destruct (common_descendant_always n m) as (k & l & H).
+revert l H. induction k.
+simpl. intros; exists l; left; auto with *.
+intros. destruct l.
+simpl in *. exists (Datatypes.S k); right; auto with *.
+simpl in *. apply pred_wd in H; rewrite !pred_succ in H. eauto.
+Qed.
+
+(** Generalized version of [pred_succ] when iterating *)
+
+Lemma succ_swap_pred : forall k n m, n == (S^k) m -> m == (P^k) n.
+Proof.
+induction k.
+simpl; auto with *.
+simpl; intros. apply pred_wd in H. rewrite pred_succ in H. apply IHk in H; auto.
+rewrite <- iter_alt in H; auto.
+Qed.
+
+(** From a given point, all others are iterated successors
+ or iterated predecessors. *)
+
+Lemma itersucc_or_iterpred : forall n m, exists k, n == (S^k) m \/ n == (P^k) m.
+Proof.
+intros n m. destruct (itersucc_or_itersucc n m) as (k,[H|H]).
+exists k; left; auto.
+exists k; right. apply succ_swap_pred; auto.
+Qed.
+
+(** In particular, all points are either iterated successors of [0]
+ or iterated predecessors of [0] (or both). *)
+
+Lemma itersucc0_or_iterpred0 :
+ forall n, exists p:nat, n == (S^p) 0 \/ n == (P^p) 0.
+Proof.
+ intros n. exact (itersucc_or_iterpred n 0).
+Qed.
+
+(** * Study of initial point w.r.t. [succ] (if any). *)
+
+Definition initial n := forall m, n ~= S m.
+
+Lemma initial_alt : forall n, initial n <-> S (P n) ~= n.
+Proof.
+split. intros Bn EQ. symmetry in EQ. destruct (Bn _ EQ).
+intros NEQ m EQ. apply NEQ. rewrite EQ, pred_succ; auto with *.
+Qed.
+
+Lemma initial_alt2 : forall n, initial n <-> ~exists m, n == S m.
+Proof. firstorder. Qed.
+
+(** First case: let's assume such an initial point exists
+ (i.e. [S] isn't surjective)... *)
+
+Section InitialExists.
+Hypothesis init : t.
+Hypothesis Initial : initial init.
+
+(** ... then we have unicity of this initial point. *)
+
+Lemma initial_unique : forall m, initial m -> m == init.
+Proof.
+intros m Im. destruct (itersucc_or_itersucc init m) as (p,[H|H]).
+destruct p. now simpl in *. destruct (Initial _ H).
+destruct p. now simpl in *. destruct (Im _ H).
+Qed.
+
+(** ... then all other points are descendant of it. *)
+
+Lemma initial_ancestor : forall m, exists p, m == (S^p) init.
+Proof.
+intros m. destruct (itersucc_or_itersucc init m) as (p,[H|H]).
+destruct p; simpl in *; auto. exists O; auto with *. destruct (Initial _ H).
+exists p; auto.
+Qed.
+
+(** NB : We would like to have [pred n == n] for the initial element,
+ but nothing forces that. For instance we can have -3 as initial point,
+ and P(-3) = 2. A bit odd indeed, but legal according to [NZDomainSig].
+ We can hence have [n == (P^k) m] without [exists k', m == (S^k') n].
+*)
+
+(** We need decidability of [eq] (or classical reasoning) for this: *)
+
+Section SuccPred.
+Hypothesis eq_decidable : forall n m, n==m \/ n~=m.
+Lemma succ_pred_approx : forall n, ~initial n -> S (P n) == n.
+Proof.
+intros n NB. rewrite initial_alt in NB.
+destruct (eq_decidable (S (P n)) n); auto.
+elim NB; auto.
+Qed.
+End SuccPred.
+End InitialExists.
+
+(** Second case : let's suppose now [S] surjective, i.e. no initial point. *)
+
+Section InitialDontExists.
+
+Hypothesis succ_onto : forall n, exists m, n == S m.
+
+Lemma succ_onto_gives_succ_pred : forall n, S (P n) == n.
+Proof.
+intros n. destruct (succ_onto n) as (m,H). rewrite H, pred_succ; auto with *.
+Qed.
+
+Lemma succ_onto_pred_injective : forall n m, P n == P m -> n == m.
+Proof.
+intros n m. intros H; apply succ_wd in H.
+rewrite !succ_onto_gives_succ_pred in H; auto.
+Qed.
+
+End InitialDontExists.
+
+
+(** To summarize:
+
+ S is always injective, P is always surjective (thanks to [pred_succ]).
+
+ I) If S is not surjective, we have an initial point, which is unique.
+ This bottom is below zero: we have N shifted (or not) to the left.
+ P cannot be injective: P init = P (S (P init)).
+ (P init) can be arbitrary.
+
+ II) If S is surjective, we have [forall n, S (P n) = n], S and P are
+ bijective and reciprocal.
+
+ IIa) if [exists k<>O, 0 == S^k 0], then we have a cyclic structure Z/nZ
+ IIb) otherwise, we have Z
+*)
+
+
+(** * An alternative induction principle using [S] and [P]. *)
+
+(** It is weaker than [bi_induction]. For instance it cannot prove that
+ we can go from one point by many [S] _or_ many [P], but only by many
+ [S] mixed with many [P]. Think of a model with two copies of N:
+
+ 0, 1=S 0, 2=S 1, ...
+ 0', 1'=S 0', 2'=S 1', ...
+
+ and P 0 = 0' and P 0' = 0.
+*)
+
+Lemma bi_induction_pred :
+ forall A : t -> Prop, Proper (eq==>iff) A ->
+ A 0 -> (forall n, A n -> A (S n)) -> (forall n, A n -> A (P n)) ->
+ forall n, A n.
+Proof.
+intros. apply bi_induction; auto.
+clear n. intros n; split; auto.
+intros G; apply H2 in G. rewrite pred_succ in G; auto.
+Qed.
+
+Lemma central_induction_pred :
+ forall A : t -> Prop, Proper (eq==>iff) A -> forall n0,
+ A n0 -> (forall n, A n -> A (S n)) -> (forall n, A n -> A (P n)) ->
+ forall n, A n.
+Proof.
+intros.
+assert (A 0).
+destruct (itersucc_or_iterpred 0 n0) as (k,[Hk|Hk]); rewrite Hk; clear Hk.
+ clear H2. induction k; simpl in *; auto.
+ clear H1. induction k; simpl in *; auto.
+apply bi_induction_pred; auto.
+Qed.
+
+End NZDomainProp.
+
+(** We now focus on the translation from [nat] into [NZ].
+ First, relationship with [0], [succ], [pred].
+*)
+
+Module NZOfNat (Import NZ:NZDomainSig').
+
+Definition ofnat (n : nat) : t := (S^n) 0.
+Notation "[ n ]" := (ofnat n) (at level 7) : ofnat.
+Local Open Scope ofnat.
+
+Lemma ofnat_zero : [O] == 0.
+Proof.
+reflexivity.
+Qed.
+
+Lemma ofnat_succ : forall n, [Datatypes.S n] == succ [n].
+Proof.
+ now unfold ofnat.
+Qed.
+
+Lemma ofnat_pred : forall n, n<>O -> [Peano.pred n] == P [n].
+Proof.
+ unfold ofnat. destruct n. destruct 1; auto.
+ intros _. simpl. symmetry. apply pred_succ.
+Qed.
+
+(** Since [P 0] can be anything in NZ (either [-1], [0], or even other
+ numbers, we cannot state previous lemma for [n=O]. *)
+
+End NZOfNat.
+
+
+(** If we require in addition a strict order on NZ, we can prove that
+ [ofnat] is injective, and hence that NZ is infinite
+ (i.e. we ban Z/nZ models) *)
+
+Module NZOfNatOrd (Import NZ:NZOrdSig').
+Include NZOfNat NZ.
+Include NZOrderPropFunct NZ.
+Local Open Scope ofnat.
+
+Theorem ofnat_S_gt_0 :
+ forall n : nat, 0 < [Datatypes.S n].
+Proof.
+unfold ofnat.
+intros n; induction n as [| n IH]; simpl in *.
+apply lt_0_1.
+apply lt_trans with 1. apply lt_0_1. now rewrite <- succ_lt_mono.
+Qed.
+
+Theorem ofnat_S_neq_0 :
+ forall n : nat, 0 ~= [Datatypes.S n].
+Proof.
+intros. apply lt_neq, ofnat_S_gt_0.
+Qed.
+
+Lemma ofnat_injective : forall n m, [n]==[m] -> n = m.
+Proof.
+induction n as [|n IH]; destruct m; auto.
+intros H; elim (ofnat_S_neq_0 _ H).
+intros H; symmetry in H; elim (ofnat_S_neq_0 _ H).
+intros. f_equal. apply IH. now rewrite <- succ_inj_wd.
+Qed.
+
+Lemma ofnat_eq : forall n m, [n]==[m] <-> n = m.
+Proof.
+split. apply ofnat_injective. intros; now subst.
+Qed.
+
+(* In addition, we can prove that [ofnat] preserves order. *)
+
+Lemma ofnat_lt : forall n m : nat, [n]<[m] <-> (n<m)%nat.
+Proof.
+induction n as [|n IH]; destruct m; repeat rewrite ofnat_zero; split.
+intro H; elim (lt_irrefl _ H).
+inversion 1.
+auto with arith.
+intros; apply ofnat_S_gt_0.
+intro H; elim (lt_asymm _ _ H); apply ofnat_S_gt_0.
+inversion 1.
+rewrite !ofnat_succ, <- succ_lt_mono, IH; auto with arith.
+rewrite !ofnat_succ, <- succ_lt_mono, IH; auto with arith.
+Qed.
+
+Lemma ofnat_le : forall n m : nat, [n]<=[m] <-> (n<=m)%nat.
+Proof.
+intros. rewrite lt_eq_cases, ofnat_lt, ofnat_eq.
+split.
+destruct 1; subst; auto with arith.
+apply Lt.le_lt_or_eq.
+Qed.
+
+End NZOfNatOrd.
+
+
+(** For basic operations, we can prove correspondance with
+ their counterpart in [nat]. *)
+
+Module NZOfNatOps (Import NZ:NZAxiomsSig').
+Include NZOfNat NZ.
+Local Open Scope ofnat.
+
+Lemma ofnat_add_l : forall n m, [n]+m == (S^n) m.
+Proof.
+ induction n; intros.
+ apply add_0_l.
+ rewrite ofnat_succ, add_succ_l. simpl; apply succ_wd; auto.
+Qed.
+
+Lemma ofnat_add : forall n m, [n+m] == [n]+[m].
+Proof.
+ intros. rewrite ofnat_add_l.
+ induction n; simpl. reflexivity.
+ rewrite ofnat_succ. now apply succ_wd.
+Qed.
+
+Lemma ofnat_mul : forall n m, [n*m] == [n]*[m].
+Proof.
+ induction n; simpl; intros.
+ symmetry. apply mul_0_l.
+ rewrite plus_comm.
+ rewrite ofnat_succ, ofnat_add, mul_succ_l.
+ now apply add_wd.
+Qed.
+
+Lemma ofnat_sub_r : forall n m, n-[m] == (P^m) n.
+Proof.
+ induction m; simpl; intros.
+ rewrite ofnat_zero. apply sub_0_r.
+ rewrite ofnat_succ, sub_succ_r. now apply pred_wd.
+Qed.
+
+Lemma ofnat_sub : forall n m, m<=n -> [n-m] == [n]-[m].
+Proof.
+ intros n m H. rewrite ofnat_sub_r.
+ revert n H. induction m. intros.
+ rewrite <- minus_n_O. now simpl.
+ intros.
+ destruct n.
+ inversion H.
+ rewrite iter_alt.
+ simpl.
+ rewrite ofnat_succ, pred_succ; auto with arith.
+Qed.
+
+End NZOfNatOps.
diff --git a/theories/Numbers/NatInt/NZMul.v b/theories/Numbers/NatInt/NZMul.v
index fda8b7a3..296bd095 100644
--- a/theories/Numbers/NatInt/NZMul.v
+++ b/theories/Numbers/NatInt/NZMul.v
@@ -8,73 +8,63 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZMul.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import NZAxioms.
-Require Import NZAdd.
+Require Import NZAxioms NZBase NZAdd.
-Module NZMulPropFunct (Import NZAxiomsMod : NZAxiomsSig).
-Module Export NZAddPropMod := NZAddPropFunct NZAxiomsMod.
-Open Local Scope NatIntScope.
+Module Type NZMulPropSig
+ (Import NZ : NZAxiomsSig')(Import NZBase : NZBasePropSig NZ).
+Include NZAddPropSig NZ NZBase.
-Theorem NZmul_0_r : forall n : NZ, n * 0 == 0.
+Theorem mul_0_r : forall n, n * 0 == 0.
Proof.
-NZinduct n.
-now rewrite NZmul_0_l.
-intro. rewrite NZmul_succ_l. now rewrite NZadd_0_r.
+nzinduct n; intros; now nzsimpl.
Qed.
-Theorem NZmul_succ_r : forall n m : NZ, n * (S m) == n * m + n.
+Theorem mul_succ_r : forall n m, n * (S m) == n * m + n.
Proof.
-intros n m; NZinduct n.
-do 2 rewrite NZmul_0_l; now rewrite NZadd_0_l.
-intro n. do 2 rewrite NZmul_succ_l. do 2 rewrite NZadd_succ_r.
-rewrite NZsucc_inj_wd. rewrite <- (NZadd_assoc (n * m) m n).
-rewrite (NZadd_comm m n). rewrite NZadd_assoc.
-now rewrite NZadd_cancel_r.
+intros n m; nzinduct n. now nzsimpl.
+intro n. nzsimpl. rewrite succ_inj_wd, <- add_assoc, (add_comm m n), add_assoc.
+now rewrite add_cancel_r.
Qed.
-Theorem NZmul_comm : forall n m : NZ, n * m == m * n.
+Hint Rewrite mul_0_r mul_succ_r : nz.
+
+Theorem mul_comm : forall n m, n * m == m * n.
Proof.
-intros n m; NZinduct n.
-rewrite NZmul_0_l; now rewrite NZmul_0_r.
-intro. rewrite NZmul_succ_l; rewrite NZmul_succ_r. now rewrite NZadd_cancel_r.
+intros n m; nzinduct n. now nzsimpl.
+intro. nzsimpl. now rewrite add_cancel_r.
Qed.
-Theorem NZmul_add_distr_r : forall n m p : NZ, (n + m) * p == n * p + m * p.
+Theorem mul_add_distr_r : forall n m p, (n + m) * p == n * p + m * p.
Proof.
-intros n m p; NZinduct n.
-rewrite NZmul_0_l. now do 2 rewrite NZadd_0_l.
-intro n. rewrite NZadd_succ_l. do 2 rewrite NZmul_succ_l.
-rewrite <- (NZadd_assoc (n * p) p (m * p)).
-rewrite (NZadd_comm p (m * p)). rewrite (NZadd_assoc (n * p) (m * p) p).
-now rewrite NZadd_cancel_r.
+intros n m p; nzinduct n. now nzsimpl.
+intro n. nzsimpl. rewrite <- add_assoc, (add_comm p (m*p)), add_assoc.
+now rewrite add_cancel_r.
Qed.
-Theorem NZmul_add_distr_l : forall n m p : NZ, n * (m + p) == n * m + n * p.
+Theorem mul_add_distr_l : forall n m p, n * (m + p) == n * m + n * p.
Proof.
intros n m p.
-rewrite (NZmul_comm n (m + p)). rewrite (NZmul_comm n m).
-rewrite (NZmul_comm n p). apply NZmul_add_distr_r.
+rewrite (mul_comm n (m + p)), (mul_comm n m), (mul_comm n p).
+apply mul_add_distr_r.
Qed.
-Theorem NZmul_assoc : forall n m p : NZ, n * (m * p) == (n * m) * p.
+Theorem mul_assoc : forall n m p, n * (m * p) == (n * m) * p.
Proof.
-intros n m p; NZinduct n.
-now do 3 rewrite NZmul_0_l.
-intro n. do 2 rewrite NZmul_succ_l. rewrite NZmul_add_distr_r.
-now rewrite NZadd_cancel_r.
+intros n m p; nzinduct n. now nzsimpl.
+intro n. nzsimpl. rewrite mul_add_distr_r.
+now rewrite add_cancel_r.
Qed.
-Theorem NZmul_1_l : forall n : NZ, 1 * n == n.
+Theorem mul_1_l : forall n, 1 * n == n.
Proof.
-intro n. rewrite NZmul_succ_l; rewrite NZmul_0_l. now rewrite NZadd_0_l.
+intro n. now nzsimpl.
Qed.
-Theorem NZmul_1_r : forall n : NZ, n * 1 == n.
+Theorem mul_1_r : forall n, n * 1 == n.
Proof.
-intro n; rewrite NZmul_comm; apply NZmul_1_l.
+intro n. now nzsimpl.
Qed.
-End NZMulPropFunct.
-
+End NZMulPropSig.
diff --git a/theories/Numbers/NatInt/NZMulOrder.v b/theories/Numbers/NatInt/NZMulOrder.v
index c707bf73..7b64a698 100644
--- a/theories/Numbers/NatInt/NZMulOrder.v
+++ b/theories/Numbers/NatInt/NZMulOrder.v
@@ -8,303 +8,300 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZMulOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Import NZAxioms.
Require Import NZAddOrder.
-Module NZMulOrderPropFunct (Import NZOrdAxiomsMod : NZOrdAxiomsSig).
-Module Export NZAddOrderPropMod := NZAddOrderPropFunct NZOrdAxiomsMod.
-Open Local Scope NatIntScope.
+Module Type NZMulOrderPropSig (Import NZ : NZOrdAxiomsSig').
+Include NZAddOrderPropSig NZ.
-Theorem NZmul_lt_pred :
- forall p q n m : NZ, S p == q -> (p * n < p * m <-> q * n + m < q * m + n).
+Theorem mul_lt_pred :
+ forall p q n m, S p == q -> (p * n < p * m <-> q * n + m < q * m + n).
Proof.
-intros p q n m H. rewrite <- H. do 2 rewrite NZmul_succ_l.
-rewrite <- (NZadd_assoc (p * n) n m).
-rewrite <- (NZadd_assoc (p * m) m n).
-rewrite (NZadd_comm n m). now rewrite <- NZadd_lt_mono_r.
+intros p q n m H. rewrite <- H. nzsimpl.
+rewrite <- ! add_assoc, (add_comm n m).
+now rewrite <- add_lt_mono_r.
Qed.
-Theorem NZmul_lt_mono_pos_l : forall p n m : NZ, 0 < p -> (n < m <-> p * n < p * m).
+Theorem mul_lt_mono_pos_l : forall p n m, 0 < p -> (n < m <-> p * n < p * m).
Proof.
-NZord_induct p.
-intros n m H; false_hyp H NZlt_irrefl.
-intros p H IH n m H1. do 2 rewrite NZmul_succ_l.
-le_elim H. assert (LR : forall n m : NZ, n < m -> p * n + n < p * m + m).
-intros n1 m1 H2. apply NZadd_lt_mono; [now apply -> IH | assumption].
-split; [apply LR |]. intro H2. apply -> NZlt_dne; intro H3.
-apply <- NZle_ngt in H3. le_elim H3.
-apply NZlt_asymm in H2. apply H2. now apply LR.
-rewrite H3 in H2; false_hyp H2 NZlt_irrefl.
-rewrite <- H; do 2 rewrite NZmul_0_l; now do 2 rewrite NZadd_0_l.
-intros p H1 _ n m H2. apply NZlt_asymm in H1. false_hyp H2 H1.
+nzord_induct p.
+intros n m H; false_hyp H lt_irrefl.
+intros p H IH n m H1. nzsimpl.
+le_elim H. assert (LR : forall n m, n < m -> p * n + n < p * m + m).
+intros n1 m1 H2. apply add_lt_mono; [now apply -> IH | assumption].
+split; [apply LR |]. intro H2. apply -> lt_dne; intro H3.
+apply <- le_ngt in H3. le_elim H3.
+apply lt_asymm in H2. apply H2. now apply LR.
+rewrite H3 in H2; false_hyp H2 lt_irrefl.
+rewrite <- H; now nzsimpl.
+intros p H1 _ n m H2. destruct (lt_asymm _ _ H1 H2).
Qed.
-Theorem NZmul_lt_mono_pos_r : forall p n m : NZ, 0 < p -> (n < m <-> n * p < m * p).
+Theorem mul_lt_mono_pos_r : forall p n m, 0 < p -> (n < m <-> n * p < m * p).
Proof.
intros p n m.
-rewrite (NZmul_comm n p); rewrite (NZmul_comm m p). now apply NZmul_lt_mono_pos_l.
+rewrite (mul_comm n p), (mul_comm m p). now apply mul_lt_mono_pos_l.
Qed.
-Theorem NZmul_lt_mono_neg_l : forall p n m : NZ, p < 0 -> (n < m <-> p * m < p * n).
+Theorem mul_lt_mono_neg_l : forall p n m, p < 0 -> (n < m <-> p * m < p * n).
Proof.
-NZord_induct p.
-intros n m H; false_hyp H NZlt_irrefl.
-intros p H1 _ n m H2. apply NZlt_succ_l in H2. apply <- NZnle_gt in H2. false_hyp H1 H2.
-intros p H IH n m H1. apply <- NZle_succ_l in H.
-le_elim H. assert (LR : forall n m : NZ, n < m -> p * m < p * n).
-intros n1 m1 H2. apply (NZle_lt_add_lt n1 m1).
-now apply NZlt_le_incl. do 2 rewrite <- NZmul_succ_l. now apply -> IH.
-split; [apply LR |]. intro H2. apply -> NZlt_dne; intro H3.
-apply <- NZle_ngt in H3. le_elim H3.
-apply NZlt_asymm in H2. apply H2. now apply LR.
-rewrite H3 in H2; false_hyp H2 NZlt_irrefl.
-rewrite (NZmul_lt_pred p (S p)) by reflexivity.
-rewrite H; do 2 rewrite NZmul_0_l; now do 2 rewrite NZadd_0_l.
+nzord_induct p.
+intros n m H; false_hyp H lt_irrefl.
+intros p H1 _ n m H2. apply lt_succ_l in H2. apply <- nle_gt in H2.
+false_hyp H1 H2.
+intros p H IH n m H1. apply <- le_succ_l in H.
+le_elim H. assert (LR : forall n m, n < m -> p * m < p * n).
+intros n1 m1 H2. apply (le_lt_add_lt n1 m1).
+now apply lt_le_incl. rewrite <- 2 mul_succ_l. now apply -> IH.
+split; [apply LR |]. intro H2. apply -> lt_dne; intro H3.
+apply <- le_ngt in H3. le_elim H3.
+apply lt_asymm in H2. apply H2. now apply LR.
+rewrite H3 in H2; false_hyp H2 lt_irrefl.
+rewrite (mul_lt_pred p (S p)) by reflexivity.
+rewrite H; do 2 rewrite mul_0_l; now do 2 rewrite add_0_l.
Qed.
-Theorem NZmul_lt_mono_neg_r : forall p n m : NZ, p < 0 -> (n < m <-> m * p < n * p).
+Theorem mul_lt_mono_neg_r : forall p n m, p < 0 -> (n < m <-> m * p < n * p).
Proof.
intros p n m.
-rewrite (NZmul_comm n p); rewrite (NZmul_comm m p). now apply NZmul_lt_mono_neg_l.
+rewrite (mul_comm n p), (mul_comm m p). now apply mul_lt_mono_neg_l.
Qed.
-Theorem NZmul_le_mono_nonneg_l : forall n m p : NZ, 0 <= p -> n <= m -> p * n <= p * m.
+Theorem mul_le_mono_nonneg_l : forall n m p, 0 <= p -> n <= m -> p * n <= p * m.
Proof.
intros n m p H1 H2. le_elim H1.
-le_elim H2. apply NZlt_le_incl. now apply -> NZmul_lt_mono_pos_l.
-apply NZeq_le_incl; now rewrite H2.
-apply NZeq_le_incl; rewrite <- H1; now do 2 rewrite NZmul_0_l.
+le_elim H2. apply lt_le_incl. now apply -> mul_lt_mono_pos_l.
+apply eq_le_incl; now rewrite H2.
+apply eq_le_incl; rewrite <- H1; now do 2 rewrite mul_0_l.
Qed.
-Theorem NZmul_le_mono_nonpos_l : forall n m p : NZ, p <= 0 -> n <= m -> p * m <= p * n.
+Theorem mul_le_mono_nonpos_l : forall n m p, p <= 0 -> n <= m -> p * m <= p * n.
Proof.
intros n m p H1 H2. le_elim H1.
-le_elim H2. apply NZlt_le_incl. now apply -> NZmul_lt_mono_neg_l.
-apply NZeq_le_incl; now rewrite H2.
-apply NZeq_le_incl; rewrite H1; now do 2 rewrite NZmul_0_l.
+le_elim H2. apply lt_le_incl. now apply -> mul_lt_mono_neg_l.
+apply eq_le_incl; now rewrite H2.
+apply eq_le_incl; rewrite H1; now do 2 rewrite mul_0_l.
Qed.
-Theorem NZmul_le_mono_nonneg_r : forall n m p : NZ, 0 <= p -> n <= m -> n * p <= m * p.
+Theorem mul_le_mono_nonneg_r : forall n m p, 0 <= p -> n <= m -> n * p <= m * p.
Proof.
-intros n m p H1 H2; rewrite (NZmul_comm n p); rewrite (NZmul_comm m p);
-now apply NZmul_le_mono_nonneg_l.
+intros n m p H1 H2;
+rewrite (mul_comm n p), (mul_comm m p); now apply mul_le_mono_nonneg_l.
Qed.
-Theorem NZmul_le_mono_nonpos_r : forall n m p : NZ, p <= 0 -> n <= m -> m * p <= n * p.
+Theorem mul_le_mono_nonpos_r : forall n m p, p <= 0 -> n <= m -> m * p <= n * p.
Proof.
-intros n m p H1 H2; rewrite (NZmul_comm n p); rewrite (NZmul_comm m p);
-now apply NZmul_le_mono_nonpos_l.
+intros n m p H1 H2;
+rewrite (mul_comm n p), (mul_comm m p); now apply mul_le_mono_nonpos_l.
Qed.
-Theorem NZmul_cancel_l : forall n m p : NZ, p ~= 0 -> (p * n == p * m <-> n == m).
+Theorem mul_cancel_l : forall n m p, p ~= 0 -> (p * n == p * m <-> n == m).
Proof.
intros n m p H; split; intro H1.
-destruct (NZlt_trichotomy p 0) as [H2 | [H2 | H2]].
-apply -> NZeq_dne; intro H3. apply -> NZlt_gt_cases in H3. destruct H3 as [H3 | H3].
-assert (H4 : p * m < p * n); [now apply -> NZmul_lt_mono_neg_l |].
-rewrite H1 in H4; false_hyp H4 NZlt_irrefl.
-assert (H4 : p * n < p * m); [now apply -> NZmul_lt_mono_neg_l |].
-rewrite H1 in H4; false_hyp H4 NZlt_irrefl.
+destruct (lt_trichotomy p 0) as [H2 | [H2 | H2]].
+apply -> eq_dne; intro H3. apply -> lt_gt_cases in H3. destruct H3 as [H3 | H3].
+assert (H4 : p * m < p * n); [now apply -> mul_lt_mono_neg_l |].
+rewrite H1 in H4; false_hyp H4 lt_irrefl.
+assert (H4 : p * n < p * m); [now apply -> mul_lt_mono_neg_l |].
+rewrite H1 in H4; false_hyp H4 lt_irrefl.
false_hyp H2 H.
-apply -> NZeq_dne; intro H3. apply -> NZlt_gt_cases in H3. destruct H3 as [H3 | H3].
-assert (H4 : p * n < p * m) by (now apply -> NZmul_lt_mono_pos_l).
-rewrite H1 in H4; false_hyp H4 NZlt_irrefl.
-assert (H4 : p * m < p * n) by (now apply -> NZmul_lt_mono_pos_l).
-rewrite H1 in H4; false_hyp H4 NZlt_irrefl.
+apply -> eq_dne; intro H3. apply -> lt_gt_cases in H3. destruct H3 as [H3 | H3].
+assert (H4 : p * n < p * m) by (now apply -> mul_lt_mono_pos_l).
+rewrite H1 in H4; false_hyp H4 lt_irrefl.
+assert (H4 : p * m < p * n) by (now apply -> mul_lt_mono_pos_l).
+rewrite H1 in H4; false_hyp H4 lt_irrefl.
now rewrite H1.
Qed.
-Theorem NZmul_cancel_r : forall n m p : NZ, p ~= 0 -> (n * p == m * p <-> n == m).
+Theorem mul_cancel_r : forall n m p, p ~= 0 -> (n * p == m * p <-> n == m).
Proof.
-intros n m p. rewrite (NZmul_comm n p), (NZmul_comm m p); apply NZmul_cancel_l.
+intros n m p. rewrite (mul_comm n p), (mul_comm m p); apply mul_cancel_l.
Qed.
-Theorem NZmul_id_l : forall n m : NZ, m ~= 0 -> (n * m == m <-> n == 1).
+Theorem mul_id_l : forall n m, m ~= 0 -> (n * m == m <-> n == 1).
Proof.
intros n m H.
-stepl (n * m == 1 * m) by now rewrite NZmul_1_l. now apply NZmul_cancel_r.
+stepl (n * m == 1 * m) by now rewrite mul_1_l. now apply mul_cancel_r.
Qed.
-Theorem NZmul_id_r : forall n m : NZ, n ~= 0 -> (n * m == n <-> m == 1).
+Theorem mul_id_r : forall n m, n ~= 0 -> (n * m == n <-> m == 1).
Proof.
-intros n m; rewrite NZmul_comm; apply NZmul_id_l.
+intros n m; rewrite mul_comm; apply mul_id_l.
Qed.
-Theorem NZmul_le_mono_pos_l : forall n m p : NZ, 0 < p -> (n <= m <-> p * n <= p * m).
+Theorem mul_le_mono_pos_l : forall n m p, 0 < p -> (n <= m <-> p * n <= p * m).
Proof.
-intros n m p H; do 2 rewrite NZlt_eq_cases.
-rewrite (NZmul_lt_mono_pos_l p n m) by assumption.
-now rewrite -> (NZmul_cancel_l n m p) by
-(intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl).
+intros n m p H; do 2 rewrite lt_eq_cases.
+rewrite (mul_lt_mono_pos_l p n m) by assumption.
+now rewrite -> (mul_cancel_l n m p) by
+(intro H1; rewrite H1 in H; false_hyp H lt_irrefl).
Qed.
-Theorem NZmul_le_mono_pos_r : forall n m p : NZ, 0 < p -> (n <= m <-> n * p <= m * p).
+Theorem mul_le_mono_pos_r : forall n m p, 0 < p -> (n <= m <-> n * p <= m * p).
Proof.
-intros n m p. rewrite (NZmul_comm n p); rewrite (NZmul_comm m p);
-apply NZmul_le_mono_pos_l.
+intros n m p. rewrite (mul_comm n p), (mul_comm m p); apply mul_le_mono_pos_l.
Qed.
-Theorem NZmul_le_mono_neg_l : forall n m p : NZ, p < 0 -> (n <= m <-> p * m <= p * n).
+Theorem mul_le_mono_neg_l : forall n m p, p < 0 -> (n <= m <-> p * m <= p * n).
Proof.
-intros n m p H; do 2 rewrite NZlt_eq_cases.
-rewrite (NZmul_lt_mono_neg_l p n m); [| assumption].
-rewrite -> (NZmul_cancel_l m n p) by (intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl).
-now setoid_replace (n == m) with (m == n) using relation iff by (split; now intro).
+intros n m p H; do 2 rewrite lt_eq_cases.
+rewrite (mul_lt_mono_neg_l p n m); [| assumption].
+rewrite -> (mul_cancel_l m n p)
+ by (intro H1; rewrite H1 in H; false_hyp H lt_irrefl).
+now setoid_replace (n == m) with (m == n) by (split; now intro).
Qed.
-Theorem NZmul_le_mono_neg_r : forall n m p : NZ, p < 0 -> (n <= m <-> m * p <= n * p).
+Theorem mul_le_mono_neg_r : forall n m p, p < 0 -> (n <= m <-> m * p <= n * p).
Proof.
-intros n m p. rewrite (NZmul_comm n p); rewrite (NZmul_comm m p);
-apply NZmul_le_mono_neg_l.
+intros n m p. rewrite (mul_comm n p), (mul_comm m p); apply mul_le_mono_neg_l.
Qed.
-Theorem NZmul_lt_mono_nonneg :
- forall n m p q : NZ, 0 <= n -> n < m -> 0 <= p -> p < q -> n * p < m * q.
+Theorem mul_lt_mono_nonneg :
+ forall n m p q, 0 <= n -> n < m -> 0 <= p -> p < q -> n * p < m * q.
Proof.
intros n m p q H1 H2 H3 H4.
-apply NZle_lt_trans with (m * p).
-apply NZmul_le_mono_nonneg_r; [assumption | now apply NZlt_le_incl].
-apply -> NZmul_lt_mono_pos_l; [assumption | now apply NZle_lt_trans with n].
+apply le_lt_trans with (m * p).
+apply mul_le_mono_nonneg_r; [assumption | now apply lt_le_incl].
+apply -> mul_lt_mono_pos_l; [assumption | now apply le_lt_trans with n].
Qed.
(* There are still many variants of the theorem above. One can assume 0 < n
or 0 < p or n <= m or p <= q. *)
-Theorem NZmul_le_mono_nonneg :
- forall n m p q : NZ, 0 <= n -> n <= m -> 0 <= p -> p <= q -> n * p <= m * q.
+Theorem mul_le_mono_nonneg :
+ forall n m p q, 0 <= n -> n <= m -> 0 <= p -> p <= q -> n * p <= m * q.
Proof.
intros n m p q H1 H2 H3 H4.
le_elim H2; le_elim H4.
-apply NZlt_le_incl; now apply NZmul_lt_mono_nonneg.
-rewrite <- H4; apply NZmul_le_mono_nonneg_r; [assumption | now apply NZlt_le_incl].
-rewrite <- H2; apply NZmul_le_mono_nonneg_l; [assumption | now apply NZlt_le_incl].
-rewrite H2; rewrite H4; now apply NZeq_le_incl.
+apply lt_le_incl; now apply mul_lt_mono_nonneg.
+rewrite <- H4; apply mul_le_mono_nonneg_r; [assumption | now apply lt_le_incl].
+rewrite <- H2; apply mul_le_mono_nonneg_l; [assumption | now apply lt_le_incl].
+rewrite H2; rewrite H4; now apply eq_le_incl.
Qed.
-Theorem NZmul_pos_pos : forall n m : NZ, 0 < n -> 0 < m -> 0 < n * m.
+Theorem mul_pos_pos : forall n m, 0 < n -> 0 < m -> 0 < n * m.
Proof.
-intros n m H1 H2.
-rewrite <- (NZmul_0_l m). now apply -> NZmul_lt_mono_pos_r.
+intros n m H1 H2. rewrite <- (mul_0_l m). now apply -> mul_lt_mono_pos_r.
Qed.
-Theorem NZmul_neg_neg : forall n m : NZ, n < 0 -> m < 0 -> 0 < n * m.
+Theorem mul_neg_neg : forall n m, n < 0 -> m < 0 -> 0 < n * m.
Proof.
-intros n m H1 H2.
-rewrite <- (NZmul_0_l m). now apply -> NZmul_lt_mono_neg_r.
+intros n m H1 H2. rewrite <- (mul_0_l m). now apply -> mul_lt_mono_neg_r.
Qed.
-Theorem NZmul_pos_neg : forall n m : NZ, 0 < n -> m < 0 -> n * m < 0.
+Theorem mul_pos_neg : forall n m, 0 < n -> m < 0 -> n * m < 0.
Proof.
-intros n m H1 H2.
-rewrite <- (NZmul_0_l m). now apply -> NZmul_lt_mono_neg_r.
+intros n m H1 H2. rewrite <- (mul_0_l m). now apply -> mul_lt_mono_neg_r.
Qed.
-Theorem NZmul_neg_pos : forall n m : NZ, n < 0 -> 0 < m -> n * m < 0.
+Theorem mul_neg_pos : forall n m, n < 0 -> 0 < m -> n * m < 0.
Proof.
-intros; rewrite NZmul_comm; now apply NZmul_pos_neg.
+intros; rewrite mul_comm; now apply mul_pos_neg.
Qed.
-Theorem NZlt_1_mul_pos : forall n m : NZ, 1 < n -> 0 < m -> 1 < n * m.
+Theorem mul_nonneg_nonneg : forall n m, 0 <= n -> 0 <= m -> 0 <= n*m.
Proof.
-intros n m H1 H2. apply -> (NZmul_lt_mono_pos_r m) in H1.
-rewrite NZmul_1_l in H1. now apply NZlt_1_l with m.
+intros. rewrite <- (mul_0_l m). apply mul_le_mono_nonneg; order.
+Qed.
+
+Theorem lt_1_mul_pos : forall n m, 1 < n -> 0 < m -> 1 < n * m.
+Proof.
+intros n m H1 H2. apply -> (mul_lt_mono_pos_r m) in H1.
+rewrite mul_1_l in H1. now apply lt_1_l with m.
assumption.
Qed.
-Theorem NZeq_mul_0 : forall n m : NZ, n * m == 0 <-> n == 0 \/ m == 0.
+Theorem eq_mul_0 : forall n m, n * m == 0 <-> n == 0 \/ m == 0.
Proof.
intros n m; split.
-intro H; destruct (NZlt_trichotomy n 0) as [H1 | [H1 | H1]];
-destruct (NZlt_trichotomy m 0) as [H2 | [H2 | H2]];
+intro H; destruct (lt_trichotomy n 0) as [H1 | [H1 | H1]];
+destruct (lt_trichotomy m 0) as [H2 | [H2 | H2]];
try (now right); try (now left).
-elimtype False; now apply (NZlt_neq 0 (n * m)); [apply NZmul_neg_neg |].
-elimtype False; now apply (NZlt_neq (n * m) 0); [apply NZmul_neg_pos |].
-elimtype False; now apply (NZlt_neq (n * m) 0); [apply NZmul_pos_neg |].
-elimtype False; now apply (NZlt_neq 0 (n * m)); [apply NZmul_pos_pos |].
-intros [H | H]. now rewrite H, NZmul_0_l. now rewrite H, NZmul_0_r.
+exfalso; now apply (lt_neq 0 (n * m)); [apply mul_neg_neg |].
+exfalso; now apply (lt_neq (n * m) 0); [apply mul_neg_pos |].
+exfalso; now apply (lt_neq (n * m) 0); [apply mul_pos_neg |].
+exfalso; now apply (lt_neq 0 (n * m)); [apply mul_pos_pos |].
+intros [H | H]. now rewrite H, mul_0_l. now rewrite H, mul_0_r.
Qed.
-Theorem NZneq_mul_0 : forall n m : NZ, n ~= 0 /\ m ~= 0 <-> n * m ~= 0.
+Theorem neq_mul_0 : forall n m, n ~= 0 /\ m ~= 0 <-> n * m ~= 0.
Proof.
intros n m; split; intro H.
-intro H1; apply -> NZeq_mul_0 in H1. tauto.
+intro H1; apply -> eq_mul_0 in H1. tauto.
split; intro H1; rewrite H1 in H;
-(rewrite NZmul_0_l in H || rewrite NZmul_0_r in H); now apply H.
+(rewrite mul_0_l in H || rewrite mul_0_r in H); now apply H.
Qed.
-Theorem NZeq_square_0 : forall n : NZ, n * n == 0 <-> n == 0.
+Theorem eq_square_0 : forall n, n * n == 0 <-> n == 0.
Proof.
-intro n; rewrite NZeq_mul_0; tauto.
+intro n; rewrite eq_mul_0; tauto.
Qed.
-Theorem NZeq_mul_0_l : forall n m : NZ, n * m == 0 -> m ~= 0 -> n == 0.
+Theorem eq_mul_0_l : forall n m, n * m == 0 -> m ~= 0 -> n == 0.
Proof.
-intros n m H1 H2. apply -> NZeq_mul_0 in H1. destruct H1 as [H1 | H1].
+intros n m H1 H2. apply -> eq_mul_0 in H1. destruct H1 as [H1 | H1].
assumption. false_hyp H1 H2.
Qed.
-Theorem NZeq_mul_0_r : forall n m : NZ, n * m == 0 -> n ~= 0 -> m == 0.
+Theorem eq_mul_0_r : forall n m, n * m == 0 -> n ~= 0 -> m == 0.
Proof.
-intros n m H1 H2; apply -> NZeq_mul_0 in H1. destruct H1 as [H1 | H1].
+intros n m H1 H2; apply -> eq_mul_0 in H1. destruct H1 as [H1 | H1].
false_hyp H1 H2. assumption.
Qed.
-Theorem NZlt_0_mul : forall n m : NZ, 0 < n * m <-> (0 < n /\ 0 < m) \/ (m < 0 /\ n < 0).
+Theorem lt_0_mul : forall 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]]].
-destruct (NZlt_trichotomy n 0) as [H1 | [H1 | H1]];
-[| rewrite H1 in H; rewrite NZmul_0_l in H; false_hyp H NZlt_irrefl |];
-(destruct (NZlt_trichotomy m 0) as [H2 | [H2 | H2]];
-[| rewrite H2 in H; rewrite NZmul_0_r in H; false_hyp H NZlt_irrefl |]);
+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]];
+[| rewrite H2 in H; rewrite mul_0_r in H; false_hyp H lt_irrefl |]);
try (left; now split); try (right; now split).
-assert (H3 : n * m < 0) by now apply NZmul_neg_pos.
-elimtype False; now apply (NZlt_asymm (n * m) 0).
-assert (H3 : n * m < 0) by now apply NZmul_pos_neg.
-elimtype False; now apply (NZlt_asymm (n * m) 0).
-now apply NZmul_pos_pos. now apply NZmul_neg_neg.
+assert (H3 : n * m < 0) by now apply mul_neg_pos.
+exfalso; now apply (lt_asymm (n * m) 0).
+assert (H3 : n * m < 0) by now apply mul_pos_neg.
+exfalso; now apply (lt_asymm (n * m) 0).
+now apply mul_pos_pos. now apply mul_neg_neg.
Qed.
-Theorem NZsquare_lt_mono_nonneg : forall n m : NZ, 0 <= n -> n < m -> n * n < m * m.
+Theorem square_lt_mono_nonneg : forall n m, 0 <= n -> n < m -> n * n < m * m.
Proof.
-intros n m H1 H2. now apply NZmul_lt_mono_nonneg.
+intros n m H1 H2. now apply mul_lt_mono_nonneg.
Qed.
-Theorem NZsquare_le_mono_nonneg : forall n m : NZ, 0 <= n -> n <= m -> n * n <= m * m.
+Theorem square_le_mono_nonneg : forall n m, 0 <= n -> n <= m -> n * n <= m * m.
Proof.
-intros n m H1 H2. now apply NZmul_le_mono_nonneg.
+intros n m H1 H2. now apply mul_le_mono_nonneg.
Qed.
(* The converse theorems require nonnegativity (or nonpositivity) of the
other variable *)
-Theorem NZsquare_lt_simpl_nonneg : forall n m : NZ, 0 <= m -> n * n < m * m -> n < m.
+Theorem square_lt_simpl_nonneg : forall n m, 0 <= m -> n * n < m * m -> n < m.
Proof.
-intros n m H1 H2. destruct (NZlt_ge_cases n 0).
-now apply NZlt_le_trans with 0.
-destruct (NZlt_ge_cases n m).
-assumption. assert (F : m * m <= n * n) by now apply NZsquare_le_mono_nonneg.
-apply -> NZle_ngt in F. false_hyp H2 F.
+intros n m H1 H2. destruct (lt_ge_cases n 0).
+now apply lt_le_trans with 0.
+destruct (lt_ge_cases n m).
+assumption. assert (F : m * m <= n * n) by now apply square_le_mono_nonneg.
+apply -> le_ngt in F. false_hyp H2 F.
Qed.
-Theorem NZsquare_le_simpl_nonneg : forall n m : NZ, 0 <= m -> n * n <= m * m -> n <= m.
+Theorem square_le_simpl_nonneg : forall n m, 0 <= m -> n * n <= m * m -> n <= m.
Proof.
-intros n m H1 H2. destruct (NZlt_ge_cases n 0).
-apply NZlt_le_incl; now apply NZlt_le_trans with 0.
-destruct (NZle_gt_cases n m).
-assumption. assert (F : m * m < n * n) by now apply NZsquare_lt_mono_nonneg.
-apply -> NZlt_nge in F. false_hyp H2 F.
+intros n m H1 H2. destruct (lt_ge_cases n 0).
+apply lt_le_incl; now apply lt_le_trans with 0.
+destruct (le_gt_cases n m).
+assumption. assert (F : m * m < n * n) by now apply square_lt_mono_nonneg.
+apply -> lt_nge in F. false_hyp H2 F.
Qed.
-Theorem NZmul_2_mono_l : forall n m : NZ, n < m -> 1 + (1 + 1) * n < (1 + 1) * m.
+Theorem mul_2_mono_l : forall n m, n < m -> 1 + (1 + 1) * n < (1 + 1) * m.
Proof.
-intros n m H. apply <- NZle_succ_l in H.
-apply -> (NZmul_le_mono_pos_l (S n) m (1 + 1)) in H.
-repeat rewrite NZmul_add_distr_r in *; repeat rewrite NZmul_1_l in *.
-repeat rewrite NZadd_succ_r in *. repeat rewrite NZadd_succ_l in *. rewrite NZadd_0_l.
-now apply -> NZle_succ_l.
-apply NZadd_pos_pos; now apply NZlt_succ_diag_r.
+intros n m. rewrite <- le_succ_l, (mul_le_mono_pos_l (S n) m (1 + 1)).
+rewrite !mul_add_distr_r; nzsimpl; now rewrite le_succ_l.
+apply add_pos_pos; now apply lt_0_1.
Qed.
-End NZMulOrderPropFunct.
+End NZMulOrderPropSig.
diff --git a/theories/Numbers/NatInt/NZOrder.v b/theories/Numbers/NatInt/NZOrder.v
index d0e2faf8..14fa0bfd 100644
--- a/theories/Numbers/NatInt/NZOrder.v
+++ b/theories/Numbers/NatInt/NZOrder.v
@@ -8,659 +8,637 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NZOrder.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import NZAxioms.
-Require Import NZMul.
-Require Import Decidable.
+Require Import NZAxioms NZBase Decidable OrdersTac.
-Module NZOrderPropFunct (Import NZOrdAxiomsMod : NZOrdAxiomsSig).
-Module Export NZMulPropMod := NZMulPropFunct NZAxiomsMod.
-Open Local Scope NatIntScope.
+Module Type NZOrderPropSig
+ (Import NZ : NZOrdSig')(Import NZBase : NZBasePropSig NZ).
-Ltac le_elim H := rewrite NZlt_eq_cases in H; destruct H as [H | H].
-
-Theorem NZlt_le_incl : forall n m : NZ, n < m -> n <= m.
+Instance le_wd : Proper (eq==>eq==>iff) le.
Proof.
-intros; apply <- NZlt_eq_cases; now left.
+intros n n' Hn m m' Hm. rewrite !lt_eq_cases, !Hn, !Hm; auto with *.
Qed.
-Theorem NZeq_le_incl : forall n m : NZ, n == m -> n <= m.
-Proof.
-intros; apply <- NZlt_eq_cases; now right.
-Qed.
+Ltac le_elim H := rewrite lt_eq_cases in H; destruct H as [H | H].
-Lemma NZlt_stepl : forall x y z : NZ, x < y -> x == z -> z < y.
+Theorem lt_le_incl : forall n m, n < m -> n <= m.
Proof.
-intros x y z H1 H2; now rewrite <- H2.
+intros; apply <- lt_eq_cases; now left.
Qed.
-Lemma NZlt_stepr : forall x y z : NZ, x < y -> y == z -> x < z.
+Theorem le_refl : forall n, n <= n.
Proof.
-intros x y z H1 H2; now rewrite <- H2.
+intro; apply <- lt_eq_cases; now right.
Qed.
-Lemma NZle_stepl : forall x y z : NZ, x <= y -> x == z -> z <= y.
+Theorem lt_succ_diag_r : forall n, n < S n.
Proof.
-intros x y z H1 H2; now rewrite <- H2.
+intro n. rewrite lt_succ_r. apply le_refl.
Qed.
-Lemma NZle_stepr : forall x y z : NZ, x <= y -> y == z -> x <= z.
+Theorem le_succ_diag_r : forall n, n <= S n.
Proof.
-intros x y z H1 H2; now rewrite <- H2.
+intro; apply lt_le_incl; apply lt_succ_diag_r.
Qed.
-Declare Left Step NZlt_stepl.
-Declare Right Step NZlt_stepr.
-Declare Left Step NZle_stepl.
-Declare Right Step NZle_stepr.
-
-Theorem NZlt_neq : forall n m : NZ, n < m -> n ~= m.
+Theorem neq_succ_diag_l : forall n, S n ~= n.
Proof.
-intros n m H1 H2; rewrite H2 in H1; false_hyp H1 NZlt_irrefl.
+intros n H. apply (lt_irrefl n). rewrite <- H at 2. apply lt_succ_diag_r.
Qed.
-Theorem NZle_neq : forall n m : NZ, n < m <-> n <= m /\ n ~= m.
+Theorem neq_succ_diag_r : forall n, n ~= S n.
Proof.
-intros n m; split; [intro H | intros [H1 H2]].
-split. now apply NZlt_le_incl. now apply NZlt_neq.
-le_elim H1. assumption. false_hyp H1 H2.
+intro n; apply neq_sym, neq_succ_diag_l.
Qed.
-Theorem NZle_refl : forall n : NZ, n <= n.
+Theorem nlt_succ_diag_l : forall n, ~ S n < n.
Proof.
-intro; now apply NZeq_le_incl.
+intros n H. apply (lt_irrefl (S n)). rewrite lt_succ_r. now apply lt_le_incl.
Qed.
-Theorem NZlt_succ_diag_r : forall n : NZ, n < S n.
+Theorem nle_succ_diag_l : forall n, ~ S n <= n.
Proof.
-intro n. rewrite NZlt_succ_r. now apply NZeq_le_incl.
+intros n H; le_elim H.
+false_hyp H nlt_succ_diag_l. false_hyp H neq_succ_diag_l.
Qed.
-Theorem NZle_succ_diag_r : forall n : NZ, n <= S n.
+Theorem le_succ_l : forall n m, S n <= m <-> n < m.
Proof.
-intro; apply NZlt_le_incl; apply NZlt_succ_diag_r.
+intro n; nzinduct m n.
+split; intro H. false_hyp H nle_succ_diag_l. false_hyp H lt_irrefl.
+intro m.
+rewrite (lt_eq_cases (S n) (S m)), !lt_succ_r, (lt_eq_cases n m), succ_inj_wd.
+rewrite or_cancel_r.
+reflexivity.
+intros LE EQ; rewrite EQ in LE; false_hyp LE nle_succ_diag_l.
+intros LT EQ; rewrite EQ in LT; false_hyp LT lt_irrefl.
Qed.
-Theorem NZlt_0_1 : 0 < 1.
-Proof.
-apply NZlt_succ_diag_r.
-Qed.
+(** Trichotomy *)
-Theorem NZle_0_1 : 0 <= 1.
+Theorem le_gt_cases : forall n m, n <= m \/ n > m.
Proof.
-apply NZle_succ_diag_r.
+intros n m; nzinduct n m.
+left; apply le_refl.
+intro n. rewrite lt_succ_r, le_succ_l, !lt_eq_cases. intuition.
Qed.
-Theorem NZlt_lt_succ_r : forall n m : NZ, n < m -> n < S m.
+Theorem lt_trichotomy : forall n m, n < m \/ n == m \/ m < n.
Proof.
-intros. rewrite NZlt_succ_r. now apply NZlt_le_incl.
+intros n m. generalize (le_gt_cases n m); rewrite lt_eq_cases; tauto.
Qed.
-Theorem NZle_le_succ_r : forall n m : NZ, n <= m -> n <= S m.
-Proof.
-intros n m H. rewrite <- NZlt_succ_r in H. now apply NZlt_le_incl.
-Qed.
+Notation lt_eq_gt_cases := lt_trichotomy (only parsing).
-Theorem NZle_succ_r : forall n m : NZ, n <= S m <-> n <= m \/ n == S m.
+(** Asymmetry and transitivity. *)
+
+Theorem lt_asymm : forall n m, n < m -> ~ m < n.
Proof.
-intros n m; rewrite NZlt_eq_cases. now rewrite NZlt_succ_r.
+intros n m; nzinduct n m.
+intros H; false_hyp H lt_irrefl.
+intro n; split; intros H H1 H2.
+apply lt_succ_r in H2. le_elim H2.
+apply H; auto. apply -> le_succ_l. now apply lt_le_incl.
+rewrite H2 in H1. false_hyp H1 nlt_succ_diag_l.
+apply le_succ_l in H1. le_elim H1.
+apply H; auto. rewrite lt_succ_r. now apply lt_le_incl.
+rewrite <- H1 in H2. false_hyp H2 nlt_succ_diag_l.
Qed.
-(* The following theorem is a special case of neq_succ_iter_l below,
-but we prove it separately *)
+Notation lt_ngt := lt_asymm (only parsing).
-Theorem NZneq_succ_diag_l : forall n : NZ, S n ~= n.
+Theorem lt_trans : forall n m p, n < m -> m < p -> n < p.
Proof.
-intros n H. pose proof (NZlt_succ_diag_r n) as H1. rewrite H in H1.
-false_hyp H1 NZlt_irrefl.
+intros n m p; nzinduct p m.
+intros _ H; false_hyp H lt_irrefl.
+intro p. rewrite 2 lt_succ_r.
+split; intros H H1 H2.
+apply lt_le_incl; le_elim H2; [now apply H | now rewrite H2 in H1].
+assert (n <= p) as H3 by (auto using lt_le_incl).
+le_elim H3. assumption. rewrite <- H3 in H2.
+elim (lt_asymm n m); auto.
Qed.
-Theorem NZneq_succ_diag_r : forall n : NZ, n ~= S n.
+Theorem le_trans : forall n m p, n <= m -> m <= p -> n <= p.
Proof.
-intro n; apply NZneq_sym; apply NZneq_succ_diag_l.
+intros n m p. rewrite 3 lt_eq_cases.
+intros [LT|EQ] [LT'|EQ']; try rewrite EQ; try rewrite <- EQ';
+ generalize (lt_trans n m p); auto with relations.
Qed.
-Theorem NZnlt_succ_diag_l : forall n : NZ, ~ S n < n.
-Proof.
-intros n H; apply NZlt_lt_succ_r in H. false_hyp H NZlt_irrefl.
-Qed.
+(** Some type classes about order *)
-Theorem NZnle_succ_diag_l : forall n : NZ, ~ S n <= n.
+Instance lt_strorder : StrictOrder lt.
+Proof. split. exact lt_irrefl. exact lt_trans. Qed.
+
+Instance le_preorder : PreOrder le.
+Proof. split. exact le_refl. exact le_trans. Qed.
+
+Instance le_partialorder : PartialOrder _ le.
Proof.
-intros n H; le_elim H.
-false_hyp H NZnlt_succ_diag_l. false_hyp H NZneq_succ_diag_l.
+intros x y. compute. split.
+intro EQ; now rewrite EQ.
+rewrite 2 lt_eq_cases. intuition. elim (lt_irrefl x). now transitivity y.
Qed.
-Theorem NZle_succ_l : forall n m : NZ, S n <= m <-> n < m.
+(** We know enough now to benefit from the generic [order] tactic. *)
+
+Definition lt_compat := lt_wd.
+Definition lt_total := lt_trichotomy.
+Definition le_lteq := lt_eq_cases.
+
+Module OrderElts <: TotalOrder.
+ Definition t := t.
+ Definition eq := eq.
+ Definition lt := lt.
+ Definition le := le.
+ Definition eq_equiv := eq_equiv.
+ Definition lt_strorder := lt_strorder.
+ Definition lt_compat := lt_compat.
+ Definition lt_total := lt_total.
+ Definition le_lteq := le_lteq.
+End OrderElts.
+Module OrderTac := !MakeOrderTac OrderElts.
+Ltac order := OrderTac.order.
+
+(** Some direct consequences of [order]. *)
+
+Theorem lt_neq : forall n m, n < m -> n ~= m.
+Proof. order. Qed.
+
+Theorem le_neq : forall n m, n < m <-> n <= m /\ n ~= m.
+Proof. intuition order. Qed.
+
+Theorem eq_le_incl : forall n m, n == m -> n <= m.
+Proof. order. Qed.
+
+Lemma lt_stepl : forall x y z, x < y -> x == z -> z < y.
+Proof. order. Qed.
+
+Lemma lt_stepr : forall x y z, x < y -> y == z -> x < z.
+Proof. order. Qed.
+
+Lemma le_stepl : forall x y z, x <= y -> x == z -> z <= y.
+Proof. order. Qed.
+
+Lemma le_stepr : forall x y z, x <= y -> y == z -> x <= z.
+Proof. order. Qed.
+
+Declare Left Step lt_stepl.
+Declare Right Step lt_stepr.
+Declare Left Step le_stepl.
+Declare Right Step le_stepr.
+
+Theorem le_lt_trans : forall n m p, n <= m -> m < p -> n < p.
+Proof. order. Qed.
+
+Theorem lt_le_trans : forall n m p, n < m -> m <= p -> n < p.
+Proof. order. Qed.
+
+Theorem le_antisymm : forall n m, n <= m -> m <= n -> n == m.
+Proof. order. Qed.
+
+(** More properties of [<] and [<=] with respect to [S] and [0]. *)
+
+Theorem le_succ_r : forall n m, n <= S m <-> n <= m \/ n == S m.
Proof.
-intro n; NZinduct m n.
-setoid_replace (n < n) with False using relation iff by
- (apply -> neg_false; apply NZlt_irrefl).
-now setoid_replace (S n <= n) with False using relation iff by
- (apply -> neg_false; apply NZnle_succ_diag_l).
-intro m. rewrite NZlt_succ_r. rewrite NZle_succ_r.
-rewrite NZsucc_inj_wd.
-rewrite (NZlt_eq_cases n m).
-rewrite or_cancel_r.
-reflexivity.
-intros H1 H2; rewrite H2 in H1; false_hyp H1 NZnle_succ_diag_l.
-apply NZlt_neq.
+intros n m; rewrite lt_eq_cases. now rewrite lt_succ_r.
Qed.
-Theorem NZlt_succ_l : forall n m : NZ, S n < m -> n < m.
+Theorem lt_succ_l : forall n m, S n < m -> n < m.
Proof.
-intros n m H; apply -> NZle_succ_l; now apply NZlt_le_incl.
+intros n m H; apply -> le_succ_l; order.
Qed.
-Theorem NZsucc_lt_mono : forall n m : NZ, n < m <-> S n < S m.
+Theorem le_le_succ_r : forall n m, n <= m -> n <= S m.
Proof.
-intros n m. rewrite <- NZle_succ_l. symmetry. apply NZlt_succ_r.
+intros n m LE. rewrite <- lt_succ_r in LE. order.
Qed.
-Theorem NZsucc_le_mono : forall n m : NZ, n <= m <-> S n <= S m.
+Theorem lt_lt_succ_r : forall n m, n < m -> n < S m.
Proof.
-intros n m. do 2 rewrite NZlt_eq_cases.
-rewrite <- NZsucc_lt_mono; now rewrite NZsucc_inj_wd.
+intros. rewrite lt_succ_r. order.
Qed.
-Theorem NZlt_asymm : forall n m, n < m -> ~ m < n.
+Theorem succ_lt_mono : forall n m, n < m <-> S n < S m.
Proof.
-intros n m; NZinduct n m.
-intros H _; false_hyp H NZlt_irrefl.
-intro n; split; intros H H1 H2.
-apply NZlt_succ_l in H1. apply -> NZlt_succ_r in H2. le_elim H2.
-now apply H. rewrite H2 in H1; false_hyp H1 NZlt_irrefl.
-apply NZlt_lt_succ_r in H2. apply <- NZle_succ_l in H1. le_elim H1.
-now apply H. rewrite H1 in H2; false_hyp H2 NZlt_irrefl.
+intros n m. rewrite <- le_succ_l. symmetry. apply lt_succ_r.
Qed.
-Theorem NZlt_trans : forall n m p : NZ, n < m -> m < p -> n < p.
+Theorem succ_le_mono : forall n m, n <= m <-> S n <= S m.
Proof.
-intros n m p; NZinduct p m.
-intros _ H; false_hyp H NZlt_irrefl.
-intro p. do 2 rewrite NZlt_succ_r.
-split; intros H H1 H2.
-apply NZlt_le_incl; le_elim H2; [now apply H | now rewrite H2 in H1].
-assert (n <= p) as H3. apply H. assumption. now apply NZlt_le_incl.
-le_elim H3. assumption. rewrite <- H3 in H2.
-elimtype False; now apply (NZlt_asymm n m).
+intros n m. now rewrite 2 lt_eq_cases, <- succ_lt_mono, succ_inj_wd.
Qed.
-Theorem NZle_trans : forall n m p : NZ, n <= m -> m <= p -> n <= p.
+Theorem lt_0_1 : 0 < 1.
Proof.
-intros n m p H1 H2; le_elim H1.
-le_elim H2. apply NZlt_le_incl; now apply NZlt_trans with (m := m).
-apply NZlt_le_incl; now rewrite <- H2. now rewrite H1.
+apply lt_succ_diag_r.
Qed.
-Theorem NZle_lt_trans : forall n m p : NZ, n <= m -> m < p -> n < p.
+Theorem le_0_1 : 0 <= 1.
Proof.
-intros n m p H1 H2; le_elim H1.
-now apply NZlt_trans with (m := m). now rewrite H1.
+apply le_succ_diag_r.
Qed.
-Theorem NZlt_le_trans : forall n m p : NZ, n < m -> m <= p -> n < p.
+Theorem lt_1_l : forall n m, 0 < n -> n < m -> 1 < m.
Proof.
-intros n m p H1 H2; le_elim H2.
-now apply NZlt_trans with (m := m). now rewrite <- H2.
+intros n m H1 H2. apply <- le_succ_l in H1. order.
Qed.
-Theorem NZle_antisymm : forall n m : NZ, n <= m -> m <= n -> n == m.
+
+(** More Trichotomy, decidability and double negation elimination. *)
+
+(** The following theorem is cleary redundant, but helps not to
+remember whether one has to say le_gt_cases or lt_ge_cases *)
+
+Theorem lt_ge_cases : forall n m, n < m \/ n >= m.
Proof.
-intros n m H1 H2; now (le_elim H1; le_elim H2);
-[elimtype False; apply (NZlt_asymm n m) | | |].
+intros n m; destruct (le_gt_cases m n); intuition order.
Qed.
-Theorem NZlt_1_l : forall n m : NZ, 0 < n -> n < m -> 1 < m.
+Theorem le_ge_cases : forall n m, n <= m \/ n >= m.
Proof.
-intros n m H1 H2. apply <- NZle_succ_l in H1. now apply NZle_lt_trans with n.
+intros n m; destruct (le_gt_cases n m); intuition order.
Qed.
-(** Trichotomy, decidability, and double negation elimination *)
-
-Theorem NZlt_trichotomy : forall n m : NZ, n < m \/ n == m \/ m < n.
+Theorem lt_gt_cases : forall n m, n ~= m <-> n < m \/ n > m.
Proof.
-intros n m; NZinduct n m.
-right; now left.
-intro n; rewrite NZlt_succ_r. stepr ((S n < m \/ S n == m) \/ m <= n) by tauto.
-rewrite <- (NZlt_eq_cases (S n) m).
-setoid_replace (n == m) with (m == n) using relation iff by now split.
-stepl (n < m \/ m < n \/ m == n) by tauto. rewrite <- NZlt_eq_cases.
-apply or_iff_compat_r. symmetry; apply NZle_succ_l.
+intros n m; destruct (lt_trichotomy n m); intuition order.
Qed.
-(* Decidability of equality, even though true in each finite ring, does not
+(** Decidability of equality, even though true in each finite ring, does not
have a uniform proof. Otherwise, the proof for two fixed numbers would
reduce to a normal form that will say if the numbers are equal or not,
which cannot be true in all finite rings. Therefore, we prove decidability
in the presence of order. *)
-Theorem NZeq_dec : forall n m : NZ, decidable (n == m).
+Theorem eq_decidable : forall n m, decidable (n == m).
Proof.
-intros n m; destruct (NZlt_trichotomy n m) as [H | [H | H]].
-right; intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl.
-now left.
-right; intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl.
+intros n m; destruct (lt_trichotomy n m) as [ | [ | ]];
+ (right; order) || (left; order).
Qed.
-(* DNE stands for double-negation elimination *)
+(** DNE stands for double-negation elimination *)
-Theorem NZeq_dne : forall n m, ~ ~ n == m <-> n == m.
+Theorem eq_dne : forall n m, ~ ~ n == m <-> n == m.
Proof.
intros n m; split; intro H.
-destruct (NZeq_dec n m) as [H1 | H1].
+destruct (eq_decidable n m) as [H1 | H1].
assumption. false_hyp H1 H.
intro H1; now apply H1.
Qed.
-Theorem NZlt_gt_cases : forall n m : NZ, n ~= m <-> n < m \/ n > m.
-Proof.
-intros n m; split.
-pose proof (NZlt_trichotomy n m); tauto.
-intros H H1; destruct H as [H | H]; rewrite H1 in H; false_hyp H NZlt_irrefl.
-Qed.
+Theorem le_ngt : forall n m, n <= m <-> ~ n > m.
+Proof. intuition order. Qed.
-Theorem NZle_gt_cases : forall n m : NZ, n <= m \/ n > m.
-Proof.
-intros n m; destruct (NZlt_trichotomy n m) as [H | [H | H]].
-left; now apply NZlt_le_incl. left; now apply NZeq_le_incl. now right.
-Qed.
-
-(* The following theorem is cleary redundant, but helps not to
-remember whether one has to say le_gt_cases or lt_ge_cases *)
+(** Redundant but useful *)
-Theorem NZlt_ge_cases : forall n m : NZ, n < m \/ n >= m.
-Proof.
-intros n m; destruct (NZle_gt_cases m n); try (now left); try (now right).
-Qed.
-
-Theorem NZle_ge_cases : forall n m : NZ, n <= m \/ n >= m.
-Proof.
-intros n m; destruct (NZle_gt_cases n m) as [H | H].
-now left. right; now apply NZlt_le_incl.
-Qed.
-
-Theorem NZle_ngt : forall n m : NZ, n <= m <-> ~ n > m.
-Proof.
-intros n m. split; intro H; [intro H1 |].
-eapply NZle_lt_trans in H; [| eassumption ..]. false_hyp H NZlt_irrefl.
-destruct (NZle_gt_cases n m) as [H1 | H1].
-assumption. false_hyp H1 H.
-Qed.
-
-(* Redundant but useful *)
-
-Theorem NZnlt_ge : forall n m : NZ, ~ n < m <-> n >= m.
-Proof.
-intros n m; symmetry; apply NZle_ngt.
-Qed.
+Theorem nlt_ge : forall n m, ~ n < m <-> n >= m.
+Proof. intuition order. Qed.
-Theorem NZlt_dec : forall n m : NZ, decidable (n < m).
+Theorem lt_decidable : forall n m, decidable (n < m).
Proof.
-intros n m; destruct (NZle_gt_cases m n);
-[right; now apply -> NZle_ngt | now left].
+intros n m; destruct (le_gt_cases m n); [right|left]; order.
Qed.
-Theorem NZlt_dne : forall n m, ~ ~ n < m <-> n < m.
+Theorem lt_dne : forall n m, ~ ~ n < m <-> n < m.
Proof.
-intros n m; split; intro H;
-[destruct (NZlt_dec n m) as [H1 | H1]; [assumption | false_hyp H1 H] |
-intro H1; false_hyp H H1].
+intros n m; split; intro H.
+destruct (lt_decidable n m) as [H1 | H1]; [assumption | false_hyp H1 H].
+intro H1; false_hyp H H1.
Qed.
-Theorem NZnle_gt : forall n m : NZ, ~ n <= m <-> n > m.
-Proof.
-intros n m. rewrite NZle_ngt. apply NZlt_dne.
-Qed.
+Theorem nle_gt : forall n m, ~ n <= m <-> n > m.
+Proof. intuition order. Qed.
-(* Redundant but useful *)
+(** Redundant but useful *)
-Theorem NZlt_nge : forall n m : NZ, n < m <-> ~ n >= m.
-Proof.
-intros n m; symmetry; apply NZnle_gt.
-Qed.
+Theorem lt_nge : forall n m, n < m <-> ~ n >= m.
+Proof. intuition order. Qed.
-Theorem NZle_dec : forall n m : NZ, decidable (n <= m).
+Theorem le_decidable : forall n m, decidable (n <= m).
Proof.
-intros n m; destruct (NZle_gt_cases n m);
-[now left | right; now apply <- NZnle_gt].
+intros n m; destruct (le_gt_cases n m); [left|right]; order.
Qed.
-Theorem NZle_dne : forall n m : NZ, ~ ~ n <= m <-> n <= m.
+Theorem le_dne : forall n m, ~ ~ n <= m <-> n <= m.
Proof.
-intros n m; split; intro H;
-[destruct (NZle_dec n m) as [H1 | H1]; [assumption | false_hyp H1 H] |
-intro H1; false_hyp H H1].
+intros n m; split; intro H.
+destruct (le_decidable n m) as [H1 | H1]; [assumption | false_hyp H1 H].
+intro H1; false_hyp H H1.
Qed.
-Theorem NZnlt_succ_r : forall n m : NZ, ~ m < S n <-> n < m.
+Theorem nlt_succ_r : forall n m, ~ m < S n <-> n < m.
Proof.
-intros n m; rewrite NZlt_succ_r; apply NZnle_gt.
+intros n m; rewrite lt_succ_r. intuition order.
Qed.
-(* The difference between integers and natural numbers is that for
+(** The difference between integers and natural numbers is that for
every integer there is a predecessor, which is not true for natural
numbers. However, for both classes, every number that is bigger than
some other number has a predecessor. The proof of this fact by regular
induction does not go through, so we need to use strong
(course-of-value) induction. *)
-Lemma NZlt_exists_pred_strong :
- forall z n m : NZ, z < m -> m <= n -> exists k : NZ, m == S k /\ z <= k.
+Lemma lt_exists_pred_strong :
+ forall z n m, z < m -> m <= n -> exists k, m == S k /\ z <= k.
Proof.
-intro z; NZinduct n z.
-intros m H1 H2; apply <- NZnle_gt in H1; false_hyp H2 H1.
+intro z; nzinduct n z.
+order.
intro n; split; intros IH m H1 H2.
-apply -> NZle_succ_r in H2; destruct H2 as [H2 | H2].
-now apply IH. exists n. now split; [| rewrite <- NZlt_succ_r; rewrite <- H2].
-apply IH. assumption. now apply NZle_le_succ_r.
+apply -> le_succ_r in H2. destruct H2 as [H2 | H2].
+now apply IH. exists n. now split; [| rewrite <- lt_succ_r; rewrite <- H2].
+apply IH. assumption. now apply le_le_succ_r.
Qed.
-Theorem NZlt_exists_pred :
- forall z n : NZ, z < n -> exists k : NZ, n == S k /\ z <= k.
+Theorem lt_exists_pred :
+ forall z n, z < n -> exists k, n == S k /\ z <= k.
Proof.
-intros z n H; apply NZlt_exists_pred_strong with (z := z) (n := n).
-assumption. apply NZle_refl.
+intros z n H; apply lt_exists_pred_strong with (z := z) (n := n).
+assumption. apply le_refl.
Qed.
-(** A corollary of having an order is that NZ is infinite *)
-
-(* This section about infinity of NZ relies on the type nat and can be
-safely removed *)
-
-Definition NZsucc_iter (n : nat) (m : NZ) :=
- nat_rect (fun _ => NZ) m (fun _ l => S l) n.
-
-Theorem NZlt_succ_iter_r :
- forall (n : nat) (m : NZ), m < NZsucc_iter (Datatypes.S n) m.
-Proof.
-intros n m; induction n as [| n IH]; simpl in *.
-apply NZlt_succ_diag_r. now apply NZlt_lt_succ_r.
-Qed.
-
-Theorem NZneq_succ_iter_l :
- forall (n : nat) (m : NZ), NZsucc_iter (Datatypes.S n) m ~= m.
-Proof.
-intros n m H. pose proof (NZlt_succ_iter_r n m) as H1. rewrite H in H1.
-false_hyp H1 NZlt_irrefl.
-Qed.
-
-(* End of the section about the infinity of NZ *)
-
(** Stronger variant of induction with assumptions n >= 0 (n < 0)
in the induction step *)
Section Induction.
-Variable A : NZ -> Prop.
-Hypothesis A_wd : predicate_wd NZeq A.
-
-Add Morphism A with signature NZeq ==> iff as A_morph.
-Proof. apply A_wd. Qed.
+Variable A : t -> Prop.
+Hypothesis A_wd : Proper (eq==>iff) A.
Section Center.
-Variable z : NZ. (* A z is the basis of induction *)
+Variable z : t. (* A z is the basis of induction *)
Section RightInduction.
-Let A' (n : NZ) := forall m : NZ, z <= m -> m < n -> A m.
-Let right_step := forall n : NZ, z <= n -> A n -> A (S n).
-Let right_step' := forall n : NZ, z <= n -> A' n -> A n.
-Let right_step'' := forall n : NZ, A' n <-> A' (S n).
+Let A' (n : t) := forall m, z <= m -> m < n -> A m.
+Let right_step := forall n, z <= n -> A n -> A (S n).
+Let right_step' := forall n, z <= n -> A' n -> A n.
+Let right_step'' := forall n, A' n <-> A' (S n).
-Lemma NZrs_rs' : A z -> right_step -> right_step'.
+Lemma rs_rs' : A z -> right_step -> right_step'.
Proof.
intros Az RS n H1 H2.
-le_elim H1. apply NZlt_exists_pred in H1. destruct H1 as [k [H3 H4]].
-rewrite H3. apply RS; [assumption | apply H2; [assumption | rewrite H3; apply NZlt_succ_diag_r]].
+le_elim H1. apply lt_exists_pred in H1. destruct H1 as [k [H3 H4]].
+rewrite H3. apply RS; trivial. apply H2; trivial.
+rewrite H3; apply lt_succ_diag_r.
rewrite <- H1; apply Az.
Qed.
-Lemma NZrs'_rs'' : right_step' -> right_step''.
+Lemma rs'_rs'' : right_step' -> right_step''.
Proof.
intros RS' n; split; intros H1 m H2 H3.
-apply -> NZlt_succ_r in H3; le_elim H3;
+apply -> lt_succ_r in H3; le_elim H3;
[now apply H1 | rewrite H3 in *; now apply RS'].
-apply H1; [assumption | now apply NZlt_lt_succ_r].
+apply H1; [assumption | now apply lt_lt_succ_r].
Qed.
-Lemma NZrbase : A' z.
+Lemma rbase : A' z.
Proof.
-intros m H1 H2. apply -> NZle_ngt in H1. false_hyp H2 H1.
+intros m H1 H2. apply -> le_ngt in H1. false_hyp H2 H1.
Qed.
-Lemma NZA'A_right : (forall n : NZ, A' n) -> forall n : NZ, z <= n -> A n.
+Lemma A'A_right : (forall n, A' n) -> forall n, z <= n -> A n.
Proof.
-intros H1 n H2. apply H1 with (n := S n); [assumption | apply NZlt_succ_diag_r].
+intros H1 n H2. apply H1 with (n := S n); [assumption | apply lt_succ_diag_r].
Qed.
-Theorem NZstrong_right_induction: right_step' -> forall n : NZ, z <= n -> A n.
+Theorem strong_right_induction: right_step' -> forall n, z <= n -> A n.
Proof.
-intro RS'; apply NZA'A_right; unfold A'; NZinduct n z;
-[apply NZrbase | apply NZrs'_rs''; apply RS'].
+intro RS'; apply A'A_right; unfold A'; nzinduct n z;
+[apply rbase | apply rs'_rs''; apply RS'].
Qed.
-Theorem NZright_induction : A z -> right_step -> forall n : NZ, z <= n -> A n.
+Theorem right_induction : A z -> right_step -> forall n, z <= n -> A n.
Proof.
-intros Az RS; apply NZstrong_right_induction; now apply NZrs_rs'.
+intros Az RS; apply strong_right_induction; now apply rs_rs'.
Qed.
-Theorem NZright_induction' :
- (forall n : NZ, n <= z -> A n) -> right_step -> forall n : NZ, A n.
+Theorem right_induction' :
+ (forall n, n <= z -> A n) -> right_step -> forall n, A n.
Proof.
intros L R n.
-destruct (NZlt_trichotomy n z) as [H | [H | H]].
-apply L; now apply NZlt_le_incl.
-apply L; now apply NZeq_le_incl.
-apply NZright_induction. apply L; now apply NZeq_le_incl. assumption. now apply NZlt_le_incl.
+destruct (lt_trichotomy n z) as [H | [H | H]].
+apply L; now apply lt_le_incl.
+apply L; now apply eq_le_incl.
+apply right_induction. apply L; now apply eq_le_incl. assumption.
+now apply lt_le_incl.
Qed.
-Theorem NZstrong_right_induction' :
- (forall n : NZ, n <= z -> A n) -> right_step' -> forall n : NZ, A n.
+Theorem strong_right_induction' :
+ (forall n, n <= z -> A n) -> right_step' -> forall n, A n.
Proof.
intros L R n.
-destruct (NZlt_trichotomy n z) as [H | [H | H]].
-apply L; now apply NZlt_le_incl.
-apply L; now apply NZeq_le_incl.
-apply NZstrong_right_induction. assumption. now apply NZlt_le_incl.
+destruct (lt_trichotomy n z) as [H | [H | H]].
+apply L; now apply lt_le_incl.
+apply L; now apply eq_le_incl.
+apply strong_right_induction. assumption. now apply lt_le_incl.
Qed.
End RightInduction.
Section LeftInduction.
-Let A' (n : NZ) := forall m : NZ, m <= z -> n <= m -> A m.
-Let left_step := forall n : NZ, n < z -> A (S n) -> A n.
-Let left_step' := forall n : NZ, n <= z -> A' (S n) -> A n.
-Let left_step'' := forall n : NZ, A' n <-> A' (S n).
+Let A' (n : t) := forall m, m <= z -> n <= m -> A m.
+Let left_step := forall n, n < z -> A (S n) -> A n.
+Let left_step' := forall n, n <= z -> A' (S n) -> A n.
+Let left_step'' := forall n, A' n <-> A' (S n).
-Lemma NZls_ls' : A z -> left_step -> left_step'.
+Lemma ls_ls' : A z -> left_step -> left_step'.
Proof.
intros Az LS n H1 H2. le_elim H1.
-apply LS; [assumption | apply H2; [now apply <- NZle_succ_l | now apply NZeq_le_incl]].
+apply LS; trivial. apply H2; [now apply <- le_succ_l | now apply eq_le_incl].
rewrite H1; apply Az.
Qed.
-Lemma NZls'_ls'' : left_step' -> left_step''.
+Lemma ls'_ls'' : left_step' -> left_step''.
Proof.
intros LS' n; split; intros H1 m H2 H3.
-apply -> NZle_succ_l in H3. apply NZlt_le_incl in H3. now apply H1.
+apply -> le_succ_l in H3. apply lt_le_incl in H3. now apply H1.
le_elim H3.
-apply <- NZle_succ_l in H3. now apply H1.
+apply <- le_succ_l in H3. now apply H1.
rewrite <- H3 in *; now apply LS'.
Qed.
-Lemma NZlbase : A' (S z).
+Lemma lbase : A' (S z).
Proof.
-intros m H1 H2. apply -> NZle_succ_l in H2.
-apply -> NZle_ngt in H1. false_hyp H2 H1.
+intros m H1 H2. apply -> le_succ_l in H2.
+apply -> le_ngt in H1. false_hyp H2 H1.
Qed.
-Lemma NZA'A_left : (forall n : NZ, A' n) -> forall n : NZ, n <= z -> A n.
+Lemma A'A_left : (forall n, A' n) -> forall n, n <= z -> A n.
Proof.
-intros H1 n H2. apply H1 with (n := n); [assumption | now apply NZeq_le_incl].
+intros H1 n H2. apply H1 with (n := n); [assumption | now apply eq_le_incl].
Qed.
-Theorem NZstrong_left_induction: left_step' -> forall n : NZ, n <= z -> A n.
+Theorem strong_left_induction: left_step' -> forall n, n <= z -> A n.
Proof.
-intro LS'; apply NZA'A_left; unfold A'; NZinduct n (S z);
-[apply NZlbase | apply NZls'_ls''; apply LS'].
+intro LS'; apply A'A_left; unfold A'; nzinduct n (S z);
+[apply lbase | apply ls'_ls''; apply LS'].
Qed.
-Theorem NZleft_induction : A z -> left_step -> forall n : NZ, n <= z -> A n.
+Theorem left_induction : A z -> left_step -> forall n, n <= z -> A n.
Proof.
-intros Az LS; apply NZstrong_left_induction; now apply NZls_ls'.
+intros Az LS; apply strong_left_induction; now apply ls_ls'.
Qed.
-Theorem NZleft_induction' :
- (forall n : NZ, z <= n -> A n) -> left_step -> forall n : NZ, A n.
+Theorem left_induction' :
+ (forall n, z <= n -> A n) -> left_step -> forall n, A n.
Proof.
intros R L n.
-destruct (NZlt_trichotomy n z) as [H | [H | H]].
-apply NZleft_induction. apply R. now apply NZeq_le_incl. assumption. now apply NZlt_le_incl.
-rewrite H; apply R; now apply NZeq_le_incl.
-apply R; now apply NZlt_le_incl.
+destruct (lt_trichotomy n z) as [H | [H | H]].
+apply left_induction. apply R. now apply eq_le_incl. assumption.
+now apply lt_le_incl.
+rewrite H; apply R; now apply eq_le_incl.
+apply R; now apply lt_le_incl.
Qed.
-Theorem NZstrong_left_induction' :
- (forall n : NZ, z <= n -> A n) -> left_step' -> forall n : NZ, A n.
+Theorem strong_left_induction' :
+ (forall n, z <= n -> A n) -> left_step' -> forall n, A n.
Proof.
intros R L n.
-destruct (NZlt_trichotomy n z) as [H | [H | H]].
-apply NZstrong_left_induction; auto. now apply NZlt_le_incl.
-rewrite H; apply R; now apply NZeq_le_incl.
-apply R; now apply NZlt_le_incl.
+destruct (lt_trichotomy n z) as [H | [H | H]].
+apply strong_left_induction; auto. now apply lt_le_incl.
+rewrite H; apply R; now apply eq_le_incl.
+apply R; now apply lt_le_incl.
Qed.
End LeftInduction.
-Theorem NZorder_induction :
+Theorem order_induction :
A z ->
- (forall n : NZ, z <= n -> A n -> A (S n)) ->
- (forall n : NZ, n < z -> A (S n) -> A n) ->
- forall n : NZ, A n.
+ (forall n, z <= n -> A n -> A (S n)) ->
+ (forall n, n < z -> A (S n) -> A n) ->
+ forall n, A n.
Proof.
intros Az RS LS n.
-destruct (NZlt_trichotomy n z) as [H | [H | H]].
-now apply NZleft_induction; [| | apply NZlt_le_incl].
+destruct (lt_trichotomy n z) as [H | [H | H]].
+now apply left_induction; [| | apply lt_le_incl].
now rewrite H.
-now apply NZright_induction; [| | apply NZlt_le_incl].
+now apply right_induction; [| | apply lt_le_incl].
Qed.
-Theorem NZorder_induction' :
+Theorem order_induction' :
A z ->
- (forall n : NZ, z <= n -> A n -> A (S n)) ->
- (forall n : NZ, n <= z -> A n -> A (P n)) ->
- forall n : NZ, A n.
+ (forall n, z <= n -> A n -> A (S n)) ->
+ (forall n, n <= z -> A n -> A (P n)) ->
+ forall n, A n.
Proof.
-intros Az AS AP n; apply NZorder_induction; try assumption.
-intros m H1 H2. apply AP in H2; [| now apply <- NZle_succ_l].
-unfold predicate_wd, fun_wd in A_wd; apply -> (A_wd (P (S m)) m);
-[assumption | apply NZpred_succ].
+intros Az AS AP n; apply order_induction; try assumption.
+intros m H1 H2. apply AP in H2; [| now apply <- le_succ_l].
+apply -> (A_wd (P (S m)) m); [assumption | apply pred_succ].
Qed.
End Center.
-Theorem NZorder_induction_0 :
+Theorem order_induction_0 :
A 0 ->
- (forall n : NZ, 0 <= n -> A n -> A (S n)) ->
- (forall n : NZ, n < 0 -> A (S n) -> A n) ->
- forall n : NZ, A n.
-Proof (NZorder_induction 0).
+ (forall n, 0 <= n -> A n -> A (S n)) ->
+ (forall n, n < 0 -> A (S n) -> A n) ->
+ forall n, A n.
+Proof (order_induction 0).
-Theorem NZorder_induction'_0 :
+Theorem order_induction'_0 :
A 0 ->
- (forall n : NZ, 0 <= n -> A n -> A (S n)) ->
- (forall n : NZ, n <= 0 -> A n -> A (P n)) ->
- forall n : NZ, A n.
-Proof (NZorder_induction' 0).
+ (forall n, 0 <= n -> A n -> A (S n)) ->
+ (forall n, n <= 0 -> A n -> A (P n)) ->
+ forall n, A n.
+Proof (order_induction' 0).
(** Elimintation principle for < *)
-Theorem NZlt_ind : forall (n : NZ),
+Theorem lt_ind : forall (n : t),
A (S n) ->
- (forall m : NZ, n < m -> A m -> A (S m)) ->
- forall m : NZ, n < m -> A m.
+ (forall m, n < m -> A m -> A (S m)) ->
+ forall m, n < m -> A m.
Proof.
intros n H1 H2 m H3.
-apply NZright_induction with (S n); [assumption | | now apply <- NZle_succ_l].
-intros; apply H2; try assumption. now apply -> NZle_succ_l.
+apply right_induction with (S n); [assumption | | now apply <- le_succ_l].
+intros; apply H2; try assumption. now apply -> le_succ_l.
Qed.
(** Elimintation principle for <= *)
-Theorem NZle_ind : forall (n : NZ),
+Theorem le_ind : forall (n : t),
A n ->
- (forall m : NZ, n <= m -> A m -> A (S m)) ->
- forall m : NZ, n <= m -> A m.
+ (forall m, n <= m -> A m -> A (S m)) ->
+ forall m, n <= m -> A m.
Proof.
intros n H1 H2 m H3.
-now apply NZright_induction with n.
+now apply right_induction with n.
Qed.
End Induction.
-Tactic Notation "NZord_induct" ident(n) :=
- induction_maker n ltac:(apply NZorder_induction_0).
+Tactic Notation "nzord_induct" ident(n) :=
+ induction_maker n ltac:(apply order_induction_0).
-Tactic Notation "NZord_induct" ident(n) constr(z) :=
- induction_maker n ltac:(apply NZorder_induction with z).
+Tactic Notation "nzord_induct" ident(n) constr(z) :=
+ induction_maker n ltac:(apply order_induction with z).
Section WF.
-Variable z : NZ.
+Variable z : t.
-Let Rlt (n m : NZ) := z <= n /\ n < m.
-Let Rgt (n m : NZ) := m < n /\ n <= z.
+Let Rlt (n m : t) := z <= n /\ n < m.
+Let Rgt (n m : t) := m < n /\ n <= z.
-Add Morphism Rlt with signature NZeq ==> NZeq ==> iff as Rlt_wd.
+Instance Rlt_wd : Proper (eq ==> eq ==> iff) Rlt.
Proof.
-intros x1 x2 H1 x3 x4 H2; unfold Rlt; rewrite H1; now rewrite H2.
+intros x1 x2 H1 x3 x4 H2; unfold Rlt. rewrite H1; now rewrite H2.
Qed.
-Add Morphism Rgt with signature NZeq ==> NZeq ==> iff as Rgt_wd.
+Instance Rgt_wd : Proper (eq ==> eq ==> iff) Rgt.
Proof.
intros x1 x2 H1 x3 x4 H2; unfold Rgt; rewrite H1; now rewrite H2.
Qed.
-Lemma NZAcc_lt_wd : predicate_wd NZeq (Acc Rlt).
+Instance Acc_lt_wd : Proper (eq==>iff) (Acc Rlt).
Proof.
-unfold predicate_wd, fun_wd.
intros x1 x2 H; split; intro H1; destruct H1 as [H2];
constructor; intros; apply H2; now (rewrite H || rewrite <- H).
Qed.
-Lemma NZAcc_gt_wd : predicate_wd NZeq (Acc Rgt).
+Instance Acc_gt_wd : Proper (eq==>iff) (Acc Rgt).
Proof.
-unfold predicate_wd, fun_wd.
intros x1 x2 H; split; intro H1; destruct H1 as [H2];
constructor; intros; apply H2; now (rewrite H || rewrite <- H).
Qed.
-Theorem NZlt_wf : well_founded Rlt.
+Theorem lt_wf : well_founded Rlt.
Proof.
unfold well_founded.
-apply NZstrong_right_induction' with (z := z).
-apply NZAcc_lt_wd.
+apply strong_right_induction' with (z := z).
+apply Acc_lt_wd.
intros n H; constructor; intros y [H1 H2].
-apply <- NZnle_gt in H2. elim H2. now apply NZle_trans with z.
+apply <- nle_gt in H2. elim H2. now apply le_trans with z.
intros n H1 H2; constructor; intros m [H3 H4]. now apply H2.
Qed.
-Theorem NZgt_wf : well_founded Rgt.
+Theorem gt_wf : well_founded Rgt.
Proof.
unfold well_founded.
-apply NZstrong_left_induction' with (z := z).
-apply NZAcc_gt_wd.
+apply strong_left_induction' with (z := z).
+apply Acc_gt_wd.
intros n H; constructor; intros y [H1 H2].
-apply <- NZnle_gt in H2. elim H2. now apply NZle_lt_trans with n.
+apply <- nle_gt in H2. elim H2. now apply le_lt_trans with n.
intros n H1 H2; constructor; intros m [H3 H4].
-apply H2. assumption. now apply <- NZle_succ_l.
+apply H2. assumption. now apply <- le_succ_l.
Qed.
End WF.
-End NZOrderPropFunct.
+End NZOrderPropSig.
+
+Module NZOrderPropFunct (NZ : NZOrdSig) :=
+ NZBasePropSig NZ <+ NZOrderPropSig NZ.
+
+(** If we have moreover a [compare] function, we can build
+ an [OrderedType] structure. *)
+
+Module NZOrderedTypeFunct (NZ : NZDecOrdSig')
+ <: DecidableTypeFull <: OrderedTypeFull :=
+ NZ <+ NZOrderPropFunct <+ Compare2EqBool <+ HasEqBool2Dec.
diff --git a/theories/Numbers/NatInt/NZProperties.v b/theories/Numbers/NatInt/NZProperties.v
new file mode 100644
index 00000000..125b4f62
--- /dev/null
+++ b/theories/Numbers/NatInt/NZProperties.v
@@ -0,0 +1,20 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+(* Evgeny Makarov, INRIA, 2007 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+Require Export NZAxioms NZMulOrder.
+
+(** This functor summarizes all known facts about NZ.
+ For the moment it is only an alias to [NZMulOrderPropFunct], which
+ subsumes all others.
+*)
+
+Module Type NZPropFunct := NZMulOrderPropSig.
diff --git a/theories/Numbers/Natural/Abstract/NAdd.v b/theories/Numbers/Natural/Abstract/NAdd.v
index 91ae5b70..9f0b54a6 100644
--- a/theories/Numbers/Natural/Abstract/NAdd.v
+++ b/theories/Numbers/Natural/Abstract/NAdd.v
@@ -8,74 +8,30 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NAdd.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
Require Export NBase.
-Module NAddPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NBasePropMod := NBasePropFunct NAxiomsMod.
+Module NAddPropFunct (Import N : NAxiomsSig').
+Include NBasePropFunct N.
-Open Local Scope NatScope.
+(** For theorems about [add] that are both valid for [N] and [Z], see [NZAdd] *)
+(** Now comes theorems valid for natural numbers but not for Z *)
-Theorem add_wd :
- forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> n1 + m1 == n2 + m2.
-Proof NZadd_wd.
-
-Theorem add_0_l : forall n : N, 0 + n == n.
-Proof NZadd_0_l.
-
-Theorem add_succ_l : forall n m : N, (S n) + m == S (n + m).
-Proof NZadd_succ_l.
-
-(** Theorems that are valid for both natural numbers and integers *)
-
-Theorem add_0_r : forall n : N, n + 0 == n.
-Proof NZadd_0_r.
-
-Theorem add_succ_r : forall n m : N, n + S m == S (n + m).
-Proof NZadd_succ_r.
-
-Theorem add_comm : forall n m : N, n + m == m + n.
-Proof NZadd_comm.
-
-Theorem add_assoc : forall n m p : N, n + (m + p) == (n + m) + p.
-Proof NZadd_assoc.
-
-Theorem add_shuffle1 : forall n m p q : N, (n + m) + (p + q) == (n + p) + (m + q).
-Proof NZadd_shuffle1.
-
-Theorem add_shuffle2 : forall n m p q : N, (n + m) + (p + q) == (n + q) + (m + p).
-Proof NZadd_shuffle2.
-
-Theorem add_1_l : forall n : N, 1 + n == S n.
-Proof NZadd_1_l.
-
-Theorem add_1_r : forall n : N, n + 1 == S n.
-Proof NZadd_1_r.
-
-Theorem add_cancel_l : forall n m p : N, p + n == p + m <-> n == m.
-Proof NZadd_cancel_l.
-
-Theorem add_cancel_r : forall n m p : N, n + p == m + p <-> n == m.
-Proof NZadd_cancel_r.
-
-(* Theorems that are valid for natural numbers but cannot be proved for Z *)
-
-Theorem eq_add_0 : forall n m : N, n + m == 0 <-> n == 0 /\ m == 0.
+Theorem eq_add_0 : forall n m, n + m == 0 <-> n == 0 /\ m == 0.
Proof.
intros n m; induct n.
-(* The next command does not work with the axiom add_0_l from NAddSig *)
-rewrite add_0_l. intuition reflexivity.
-intros n IH. rewrite add_succ_l.
-setoid_replace (S (n + m) == 0) with False using relation iff by
+nzsimpl; intuition.
+intros n IH. nzsimpl.
+setoid_replace (S (n + m) == 0) with False by
(apply -> neg_false; apply neq_succ_0).
-setoid_replace (S n == 0) with False using relation iff by
+setoid_replace (S n == 0) with False by
(apply -> neg_false; apply neq_succ_0). tauto.
Qed.
Theorem eq_add_succ :
- forall n m : N, (exists p : N, n + m == S p) <->
- (exists n' : N, n == S n') \/ (exists m' : N, m == S m').
+ forall n m, (exists p, n + m == S p) <->
+ (exists n', n == S n') \/ (exists m', m == S m').
Proof.
intros n m; cases n.
split; intro H.
@@ -88,11 +44,11 @@ left; now exists n.
exists (n + m); now rewrite add_succ_l.
Qed.
-Theorem eq_add_1 : forall n m : N,
+Theorem eq_add_1 : forall n m,
n + m == 1 -> n == 1 /\ m == 0 \/ n == 0 /\ m == 1.
Proof.
intros n m H.
-assert (H1 : exists p : N, n + m == S p) by now exists 0.
+assert (H1 : exists p, n + m == S p) by now exists 0.
apply -> eq_add_succ in H1. destruct H1 as [[n' H1] | [m' H1]].
left. rewrite H1 in H; rewrite add_succ_l in H; apply succ_inj in H.
apply -> eq_add_0 in H. destruct H as [H2 H3]; rewrite H2 in H1; now split.
@@ -100,7 +56,7 @@ right. rewrite H1 in H; rewrite add_succ_r in H; apply succ_inj in H.
apply -> eq_add_0 in H. destruct H as [H2 H3]; rewrite H3 in H1; now split.
Qed.
-Theorem succ_add_discr : forall n m : N, m ~= S (n + m).
+Theorem succ_add_discr : forall n m, m ~= S (n + m).
Proof.
intro n; induct m.
apply neq_sym. apply neq_succ_0.
@@ -108,49 +64,18 @@ intros m IH H. apply succ_inj in H. rewrite add_succ_r in H.
unfold not in IH; now apply IH.
Qed.
-Theorem add_pred_l : forall n m : N, n ~= 0 -> P n + m == P (n + m).
+Theorem add_pred_l : forall n m, n ~= 0 -> P n + m == P (n + m).
Proof.
intros n m; cases n.
intro H; now elim H.
intros n IH; rewrite add_succ_l; now do 2 rewrite pred_succ.
Qed.
-Theorem add_pred_r : forall n m : N, m ~= 0 -> n + P m == P (n + m).
+Theorem add_pred_r : forall n m, m ~= 0 -> n + P m == P (n + m).
Proof.
intros n m H; rewrite (add_comm n (P m));
rewrite (add_comm n m); now apply add_pred_l.
Qed.
-(* One could define n <= m as exists p : N, p + n == m. Then we have
-dichotomy:
-
-forall n m : N, n <= m \/ m <= n,
-
-i.e.,
-
-forall n m : N, (exists p : N, p + n == m) \/ (exists p : N, p + m == n) (1)
-
-We will need (1) in the proof of induction principle for integers
-constructed as pairs of natural numbers. The formula (1) can be proved
-using properties of order and truncated subtraction. Thus, p would be
-m - n or n - m and (1) would hold by theorem sub_add from Sub.v
-depending on whether n <= m or m <= n. However, in proving induction
-for integers constructed from natural numbers we do not need to
-require implementations of order and sub; it is enough to prove (1)
-here. *)
-
-Theorem add_dichotomy :
- forall n m : N, (exists p : N, p + n == m) \/ (exists p : N, p + m == n).
-Proof.
-intros n m; induct n.
-left; exists m; apply add_0_r.
-intros n IH.
-destruct IH as [[p H] | [p H]].
-destruct (zero_or_succ p) as [H1 | [p' H1]]; rewrite H1 in H.
-rewrite add_0_l in H. right; exists (S 0); rewrite H; rewrite add_succ_l; now rewrite add_0_l.
-left; exists p'; rewrite add_succ_r; now rewrite add_succ_l in H.
-right; exists (S p). rewrite add_succ_l; now rewrite H.
-Qed.
-
End NAddPropFunct.
diff --git a/theories/Numbers/Natural/Abstract/NAddOrder.v b/theories/Numbers/Natural/Abstract/NAddOrder.v
index 7024fd00..0ce04e54 100644
--- a/theories/Numbers/Natural/Abstract/NAddOrder.v
+++ b/theories/Numbers/Natural/Abstract/NAddOrder.v
@@ -8,107 +8,41 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NAddOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export NOrder.
-Module NAddOrderPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NOrderPropMod := NOrderPropFunct NAxiomsMod.
-Open Local Scope NatScope.
+Module NAddOrderPropFunct (Import N : NAxiomsSig').
+Include NOrderPropFunct N.
-Theorem add_lt_mono_l : forall n m p : N, n < m <-> p + n < p + m.
-Proof NZadd_lt_mono_l.
+(** Theorems true for natural numbers, not for integers *)
-Theorem add_lt_mono_r : forall n m p : N, n < m <-> n + p < m + p.
-Proof NZadd_lt_mono_r.
-
-Theorem add_lt_mono : forall n m p q : N, n < m -> p < q -> n + p < m + q.
-Proof NZadd_lt_mono.
-
-Theorem add_le_mono_l : forall n m p : N, n <= m <-> p + n <= p + m.
-Proof NZadd_le_mono_l.
-
-Theorem add_le_mono_r : forall n m p : N, n <= m <-> n + p <= m + p.
-Proof NZadd_le_mono_r.
-
-Theorem add_le_mono : forall n m p q : N, n <= m -> p <= q -> n + p <= m + q.
-Proof NZadd_le_mono.
-
-Theorem add_lt_le_mono : forall n m p q : N, n < m -> p <= q -> n + p < m + q.
-Proof NZadd_lt_le_mono.
-
-Theorem add_le_lt_mono : forall n m p q : N, n <= m -> p < q -> n + p < m + q.
-Proof NZadd_le_lt_mono.
-
-Theorem add_pos_pos : forall n m : N, 0 < n -> 0 < m -> 0 < n + m.
-Proof NZadd_pos_pos.
-
-Theorem lt_add_pos_l : forall n m : N, 0 < n -> m < n + m.
-Proof NZlt_add_pos_l.
-
-Theorem lt_add_pos_r : forall n m : N, 0 < n -> m < m + n.
-Proof NZlt_add_pos_r.
-
-Theorem le_lt_add_lt : forall n m p q : N, n <= m -> p + m < q + n -> p < q.
-Proof NZle_lt_add_lt.
-
-Theorem lt_le_add_lt : forall n m p q : N, n < m -> p + m <= q + n -> p < q.
-Proof NZlt_le_add_lt.
-
-Theorem le_le_add_le : forall n m p q : N, n <= m -> p + m <= q + n -> p <= q.
-Proof NZle_le_add_le.
-
-Theorem add_lt_cases : forall n m p q : N, n + m < p + q -> n < p \/ m < q.
-Proof NZadd_lt_cases.
-
-Theorem add_le_cases : forall n m p q : N, n + m <= p + q -> n <= p \/ m <= q.
-Proof NZadd_le_cases.
-
-Theorem add_pos_cases : forall n m : N, 0 < n + m -> 0 < n \/ 0 < m.
-Proof NZadd_pos_cases.
-
-(* Theorems true for natural numbers *)
-
-Theorem le_add_r : forall n m : N, n <= n + m.
+Theorem le_add_r : forall n m, n <= n + m.
Proof.
intro n; induct m.
rewrite add_0_r; now apply eq_le_incl.
intros m IH. rewrite add_succ_r; now apply le_le_succ_r.
Qed.
-Theorem lt_lt_add_r : forall n m p : N, n < m -> n < m + p.
+Theorem lt_lt_add_r : forall n m p, n < m -> n < m + p.
Proof.
intros n m p H; rewrite <- (add_0_r n).
apply add_lt_le_mono; [assumption | apply le_0_l].
Qed.
-Theorem lt_lt_add_l : forall n m p : N, n < m -> n < p + m.
+Theorem lt_lt_add_l : forall n m p, n < m -> n < p + m.
Proof.
intros n m p; rewrite add_comm; apply lt_lt_add_r.
Qed.
-Theorem add_pos_l : forall n m : N, 0 < n -> 0 < n + m.
+Theorem add_pos_l : forall n m, 0 < n -> 0 < n + m.
Proof.
-intros; apply NZadd_pos_nonneg. assumption. apply le_0_l.
+intros; apply add_pos_nonneg. assumption. apply le_0_l.
Qed.
-Theorem add_pos_r : forall n m : N, 0 < m -> 0 < n + m.
-Proof.
-intros; apply NZadd_nonneg_pos. apply le_0_l. assumption.
-Qed.
-
-(* The following property is used to prove the correctness of the
-definition of order on integers constructed from pairs of natural numbers *)
-
-Theorem add_lt_repl_pair : forall n m n' m' u v : N,
- n + u < m + v -> n + m' == n' + m -> n' + u < m' + v.
+Theorem add_pos_r : forall n m, 0 < m -> 0 < n + m.
Proof.
-intros n m n' m' u v H1 H2.
-symmetry in H2. assert (H3 : n' + m <= n + m') by now apply eq_le_incl.
-pose proof (add_lt_le_mono _ _ _ _ H1 H3) as H4.
-rewrite (add_shuffle2 n u), (add_shuffle1 m v), (add_comm m n) in H4.
-do 2 rewrite <- add_assoc in H4. do 2 apply <- add_lt_mono_l in H4.
-now rewrite (add_comm n' u), (add_comm m' v).
+intros; apply add_nonneg_pos. apply le_0_l. assumption.
Qed.
End NAddOrderPropFunct.
diff --git a/theories/Numbers/Natural/Abstract/NAxioms.v b/theories/Numbers/Natural/Abstract/NAxioms.v
index 750cc977..42016ab1 100644
--- a/theories/Numbers/Natural/Abstract/NAxioms.v
+++ b/theories/Numbers/Natural/Abstract/NAxioms.v
@@ -8,64 +8,32 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export NZAxioms.
Set Implicit Arguments.
-Module Type NAxiomsSig.
-Declare Module Export NZOrdAxiomsMod : NZOrdAxiomsSig.
+Module Type NAxioms (Import NZ : NZDomainSig').
-Delimit Scope NatScope with Nat.
-Notation N := NZ.
-Notation Neq := NZeq.
-Notation N0 := NZ0.
-Notation N1 := (NZsucc NZ0).
-Notation S := NZsucc.
-Notation P := NZpred.
-Notation add := NZadd.
-Notation mul := NZmul.
-Notation sub := NZsub.
-Notation lt := NZlt.
-Notation le := NZle.
-Notation min := NZmin.
-Notation max := NZmax.
-Notation "x == y" := (Neq x y) (at level 70) : NatScope.
-Notation "x ~= y" := (~ Neq x y) (at level 70) : NatScope.
-Notation "0" := NZ0 : NatScope.
-Notation "1" := (NZsucc NZ0) : NatScope.
-Notation "x + y" := (NZadd x y) : NatScope.
-Notation "x - y" := (NZsub x y) : NatScope.
-Notation "x * y" := (NZmul x y) : NatScope.
-Notation "x < y" := (NZlt x y) : NatScope.
-Notation "x <= y" := (NZle x y) : NatScope.
-Notation "x > y" := (NZlt y x) (only parsing) : NatScope.
-Notation "x >= y" := (NZle y x) (only parsing) : NatScope.
-
-Open Local Scope NatScope.
+Axiom pred_0 : P 0 == 0.
-Parameter Inline recursion : forall A : Type, A -> (N -> A -> A) -> N -> A.
+Parameter Inline recursion : forall A : Type, A -> (t -> A -> A) -> t -> A.
Implicit Arguments recursion [A].
-Axiom pred_0 : P 0 == 0.
-
-Axiom recursion_wd : forall (A : Type) (Aeq : relation A),
- forall a a' : A, Aeq a a' ->
- forall f f' : N -> A -> A, fun2_eq Neq Aeq Aeq f f' ->
- forall x x' : N, x == x' ->
- Aeq (recursion a f x) (recursion a' f' x').
+Declare Instance recursion_wd (A : Type) (Aeq : relation A) :
+ Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) (@recursion A).
Axiom recursion_0 :
- forall (A : Type) (a : A) (f : N -> A -> A), recursion a f 0 = a.
+ forall (A : Type) (a : A) (f : t -> A -> A), recursion a f 0 = a.
Axiom recursion_succ :
- forall (A : Type) (Aeq : relation A) (a : A) (f : N -> A -> A),
- Aeq a a -> fun2_wd Neq Aeq Aeq f ->
- forall n : N, Aeq (recursion a f (S n)) (f n (recursion a f n)).
+ forall (A : Type) (Aeq : relation A) (a : A) (f : t -> A -> A),
+ Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
+ forall n, Aeq (recursion a f (S n)) (f n (recursion a f n)).
-(*Axiom dep_rec :
- forall A : N -> Type, A 0 -> (forall n : N, A n -> A (S n)) -> forall n : N, A n.*)
+End NAxioms.
-End NAxiomsSig.
+Module Type NAxiomsSig := NZOrdAxiomsSig <+ NAxioms.
+Module Type NAxiomsSig' := NZOrdAxiomsSig' <+ NAxioms.
diff --git a/theories/Numbers/Natural/Abstract/NBase.v b/theories/Numbers/Natural/Abstract/NBase.v
index 85e2c2ab..842f4bcf 100644
--- a/theories/Numbers/Natural/Abstract/NBase.v
+++ b/theories/Numbers/Natural/Abstract/NBase.v
@@ -8,135 +8,78 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NBase.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
Require Export Decidable.
Require Export NAxioms.
-Require Import NZMulOrder. (* The last property functor on NZ, which subsumes all others *)
+Require Import NZProperties.
-Module NBasePropFunct (Import NAxiomsMod : NAxiomsSig).
+Module NBasePropFunct (Import N : NAxiomsSig').
+(** First, we import all known facts about both natural numbers and integers. *)
+Include NZPropFunct N.
-Open Local Scope NatScope.
-
-(* We call the last property functor on NZ, which includes all the previous
-ones, to get all properties of NZ at once. This way we will include them
-only one time. *)
-
-Module Export NZMulOrderMod := NZMulOrderPropFunct NZOrdAxiomsMod.
-
-(* Here we probably need to re-prove all axioms declared in NAxioms.v to
-make sure that the definitions like N, S and add are unfolded in them,
-since unfolding is done only inside a functor. In fact, we'll do it in the
-files that prove the corresponding properties. In those files, we will also
-rename properties proved in NZ files by removing NZ from their names. In
-this way, one only has to consult, for example, NAdd.v to see all
-available properties for add, i.e., one does not have to go to NAxioms.v
-for axioms and NZAdd.v for theorems. *)
-
-Theorem succ_wd : forall n1 n2 : N, n1 == n2 -> S n1 == S n2.
-Proof NZsucc_wd.
-
-Theorem pred_wd : forall n1 n2 : N, n1 == n2 -> P n1 == P n2.
-Proof NZpred_wd.
-
-Theorem pred_succ : forall n : N, P (S n) == n.
-Proof NZpred_succ.
-
-Theorem pred_0 : P 0 == 0.
-Proof pred_0.
-
-Theorem Neq_refl : forall n : N, n == n.
-Proof (proj1 NZeq_equiv).
-
-Theorem Neq_sym : forall n m : N, n == m -> m == n.
-Proof (proj2 (proj2 NZeq_equiv)).
-
-Theorem Neq_trans : forall n m p : N, n == m -> m == p -> n == p.
-Proof (proj1 (proj2 NZeq_equiv)).
-
-Theorem neq_sym : forall n m : N, n ~= m -> m ~= n.
-Proof NZneq_sym.
-
-Theorem succ_inj : forall n1 n2 : N, S n1 == S n2 -> n1 == n2.
-Proof NZsucc_inj.
-
-Theorem succ_inj_wd : forall n1 n2 : N, S n1 == S n2 <-> n1 == n2.
-Proof NZsucc_inj_wd.
-
-Theorem succ_inj_wd_neg : forall n m : N, S n ~= S m <-> n ~= m.
-Proof NZsucc_inj_wd_neg.
-
-(* Decidability and stability of equality was proved only in NZOrder, but
-since it does not mention order, we'll put it here *)
-
-Theorem eq_dec : forall n m : N, decidable (n == m).
-Proof NZeq_dec.
-
-Theorem eq_dne : forall n m : N, ~ ~ n == m <-> n == m.
-Proof NZeq_dne.
-
-(* Now we prove that the successor of a number is not zero by defining a
+(** We prove that the successor of a number is not zero by defining a
function (by recursion) that maps 0 to false and the successor to true *)
-Definition if_zero (A : Set) (a b : A) (n : N) : A :=
+Definition if_zero (A : Type) (a b : A) (n : N.t) : A :=
recursion a (fun _ _ => b) n.
-Add Parametric Morphism (A : Set) : (if_zero A) with signature (@eq _ ==> @eq _ ==> Neq ==> @eq _) as if_zero_wd.
+Implicit Arguments if_zero [A].
+
+Instance if_zero_wd (A : Type) :
+ Proper (Logic.eq ==> Logic.eq ==> N.eq ==> Logic.eq) (@if_zero A).
Proof.
-intros; unfold if_zero. apply recursion_wd with (Aeq := (@eq A)).
-reflexivity. unfold fun2_eq; now intros. assumption.
+intros; unfold if_zero.
+repeat red; intros. apply recursion_wd; auto. repeat red; auto.
Qed.
-Theorem if_zero_0 : forall (A : Set) (a b : A), if_zero A a b 0 = a.
+Theorem if_zero_0 : forall (A : Type) (a b : A), if_zero a b 0 = a.
Proof.
unfold if_zero; intros; now rewrite recursion_0.
Qed.
-Theorem if_zero_succ : forall (A : Set) (a b : A) (n : N), if_zero A a b (S n) = b.
+Theorem if_zero_succ :
+ forall (A : Type) (a b : A) (n : N.t), if_zero a b (S n) = b.
Proof.
intros; unfold if_zero.
-now rewrite (@recursion_succ A (@eq A)); [| | unfold fun2_wd; now intros].
+now rewrite recursion_succ.
Qed.
-Implicit Arguments if_zero [A].
-
-Theorem neq_succ_0 : forall n : N, S n ~= 0.
+Theorem neq_succ_0 : forall n, S n ~= 0.
Proof.
intros n H.
-assert (true = false); [| discriminate].
-replace true with (if_zero false true (S n)) by apply if_zero_succ.
-pattern false at 2; replace false with (if_zero false true 0) by apply if_zero_0.
-now rewrite H.
+generalize (Logic.eq_refl (if_zero false true 0)).
+rewrite <- H at 1. rewrite if_zero_0, if_zero_succ; discriminate.
Qed.
-Theorem neq_0_succ : forall n : N, 0 ~= S n.
+Theorem neq_0_succ : forall n, 0 ~= S n.
Proof.
intro n; apply neq_sym; apply neq_succ_0.
Qed.
-(* Next, we show that all numbers are nonnegative and recover regular induction
-from the bidirectional induction on NZ *)
+(** Next, we show that all numbers are nonnegative and recover regular
+ induction from the bidirectional induction on NZ *)
-Theorem le_0_l : forall n : N, 0 <= n.
+Theorem le_0_l : forall n, 0 <= n.
Proof.
-NZinduct n.
-now apply NZeq_le_incl.
+nzinduct n.
+now apply eq_le_incl.
intro n; split.
-apply NZle_le_succ_r.
-intro H; apply -> NZle_succ_r in H; destruct H as [H | H].
+apply le_le_succ_r.
+intro H; apply -> le_succ_r in H; destruct H as [H | H].
assumption.
symmetry in H; false_hyp H neq_succ_0.
Qed.
Theorem induction :
- forall A : N -> Prop, predicate_wd Neq A ->
- A 0 -> (forall n : N, A n -> A (S n)) -> forall n : N, A n.
+ forall A : N.t -> Prop, Proper (N.eq==>iff) A ->
+ A 0 -> (forall n, A n -> A (S n)) -> forall n, A n.
Proof.
-intros A A_wd A0 AS n; apply NZright_induction with 0; try assumption.
+intros A A_wd A0 AS n; apply right_induction with 0; try assumption.
intros; auto; apply le_0_l. apply le_0_l.
Qed.
-(* The theorems NZinduction, NZcentral_induction and the tactic NZinduct
+(** The theorems [bi_induction], [central_induction] and the tactic [nzinduct]
refer to bidirectional induction, which is not useful on natural
numbers. Therefore, we define a new induction tactic for natural numbers.
We do not have to call "Declare Left Step" and "Declare Right Step"
@@ -146,8 +89,8 @@ from NZ. *)
Ltac induct n := induction_maker n ltac:(apply induction).
Theorem case_analysis :
- forall A : N -> Prop, predicate_wd Neq A ->
- A 0 -> (forall n : N, A (S n)) -> forall n : N, A n.
+ forall A : N.t -> Prop, Proper (N.eq==>iff) A ->
+ A 0 -> (forall n, A (S n)) -> forall n, A n.
Proof.
intros; apply induction; auto.
Qed.
@@ -173,7 +116,7 @@ now left.
intro n; right; now exists n.
Qed.
-Theorem eq_pred_0 : forall n : N, P n == 0 <-> n == 0 \/ n == 1.
+Theorem eq_pred_0 : forall n, P n == 0 <-> n == 0 \/ n == 1.
Proof.
cases n.
rewrite pred_0. setoid_replace (0 == 1) with False using relation iff. tauto.
@@ -184,34 +127,29 @@ setoid_replace (S n == 0) with False using relation iff by
rewrite succ_inj_wd. tauto.
Qed.
-Theorem succ_pred : forall n : N, n ~= 0 -> S (P n) == n.
+Theorem succ_pred : forall n, n ~= 0 -> S (P n) == n.
Proof.
cases n.
-intro H; elimtype False; now apply H.
+intro H; exfalso; now apply H.
intros; now rewrite pred_succ.
Qed.
-Theorem pred_inj : forall n m : N, n ~= 0 -> m ~= 0 -> P n == P m -> n == m.
+Theorem pred_inj : forall n m, n ~= 0 -> m ~= 0 -> P n == P m -> n == m.
Proof.
intros n m; cases n.
-intros H; elimtype False; now apply H.
+intros H; exfalso; now apply H.
intros n _; cases m.
-intros H; elimtype False; now apply H.
+intros H; exfalso; now apply H.
intros m H2 H3. do 2 rewrite pred_succ in H3. now rewrite H3.
Qed.
-(* The following induction principle is useful for reasoning about, e.g.,
+(** The following induction principle is useful for reasoning about, e.g.,
Fibonacci numbers *)
Section PairInduction.
-Variable A : N -> Prop.
-Hypothesis A_wd : predicate_wd Neq A.
-
-Add Morphism A with signature Neq ==> iff as A_morph.
-Proof.
-exact A_wd.
-Qed.
+Variable A : N.t -> Prop.
+Hypothesis A_wd : Proper (N.eq==>iff) A.
Theorem pair_induction :
A 0 -> A 1 ->
@@ -224,18 +162,12 @@ Qed.
End PairInduction.
-(*Ltac pair_induct n := induction_maker n ltac:(apply pair_induction).*)
+(** The following is useful for reasoning about, e.g., Ackermann function *)
-(* The following is useful for reasoning about, e.g., Ackermann function *)
Section TwoDimensionalInduction.
-Variable R : N -> N -> Prop.
-Hypothesis R_wd : relation_wd Neq Neq R.
-
-Add Morphism R with signature Neq ==> Neq ==> iff as R_morph.
-Proof.
-exact R_wd.
-Qed.
+Variable R : N.t -> N.t -> Prop.
+Hypothesis R_wd : Proper (N.eq==>N.eq==>iff) R.
Theorem two_dim_induction :
R 0 0 ->
@@ -251,26 +183,16 @@ Qed.
End TwoDimensionalInduction.
-(*Ltac two_dim_induct n m :=
- try intros until n;
- try intros until m;
- pattern n, m; apply two_dim_induction; clear n m;
- [solve_relation_wd | | | ].*)
Section DoubleInduction.
-Variable R : N -> N -> Prop.
-Hypothesis R_wd : relation_wd Neq Neq R.
-
-Add Morphism R with signature Neq ==> Neq ==> iff as R_morph1.
-Proof.
-exact R_wd.
-Qed.
+Variable R : N.t -> N.t -> Prop.
+Hypothesis R_wd : Proper (N.eq==>N.eq==>iff) R.
Theorem double_induction :
- (forall m : N, R 0 m) ->
- (forall n : N, R (S n) 0) ->
- (forall n m : N, R n m -> R (S n) (S m)) -> forall n m : N, R n m.
+ (forall m, R 0 m) ->
+ (forall n, R (S n) 0) ->
+ (forall n m, R n m -> R (S n) (S m)) -> forall n m, R n m.
Proof.
intros H1 H2 H3; induct n; auto.
intros n H; cases m; auto.
diff --git a/theories/Numbers/Natural/Abstract/NDefOps.v b/theories/Numbers/Natural/Abstract/NDefOps.v
index 0a8f5f1e..22eb2cb3 100644
--- a/theories/Numbers/Natural/Abstract/NDefOps.v
+++ b/theories/Numbers/Natural/Abstract/NDefOps.v
@@ -8,45 +8,47 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NDefOps.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
Require Import Bool. (* To get the orb and negb function *)
+Require Import RelationPairs.
Require Export NStrongRec.
-Module NdefOpsPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NStrongRecPropMod := NStrongRecPropFunct NAxiomsMod.
-Open Local Scope NatScope.
+Module NdefOpsPropFunct (Import N : NAxiomsSig').
+Include NStrongRecPropFunct N.
(*****************************************************)
(** Addition *)
-Definition def_add (x y : N) := recursion y (fun _ p => S p) x.
+Definition def_add (x y : N.t) := recursion y (fun _ => S) x.
-Infix Local "++" := def_add (at level 50, left associativity).
+Local Infix "+++" := def_add (at level 50, left associativity).
-Add Morphism def_add with signature Neq ==> Neq ==> Neq as def_add_wd.
+Instance def_add_prewd : Proper (N.eq==>N.eq==>N.eq) (fun _ => S).
Proof.
-unfold def_add.
-intros x x' Exx' y y' Eyy'.
-apply recursion_wd with (Aeq := Neq).
-assumption.
-unfold fun2_eq; intros _ _ _ p p' Epp'; now rewrite Epp'.
-assumption.
+intros _ _ _ p p' Epp'; now rewrite Epp'.
+Qed.
+
+Instance def_add_wd : Proper (N.eq ==> N.eq ==> N.eq) def_add.
+Proof.
+intros x x' Exx' y y' Eyy'. unfold def_add.
+(* TODO: why rewrite Exx' don't work here (or verrrry slowly) ? *)
+apply recursion_wd with (Aeq := N.eq); auto with *.
+apply def_add_prewd.
Qed.
-Theorem def_add_0_l : forall y : N, 0 ++ y == y.
+Theorem def_add_0_l : forall y, 0 +++ y == y.
Proof.
intro y. unfold def_add. now rewrite recursion_0.
Qed.
-Theorem def_add_succ_l : forall x y : N, S x ++ y == S (x ++ y).
+Theorem def_add_succ_l : forall x y, S x +++ y == S (x +++ y).
Proof.
intros x y; unfold def_add.
-rewrite (@recursion_succ N Neq); try reflexivity.
-unfold fun2_wd. intros _ _ _ m1 m2 H2. now rewrite H2.
+rewrite recursion_succ; auto with *.
Qed.
-Theorem def_add_add : forall n m : N, n ++ m == n + m.
+Theorem def_add_add : forall n m, n +++ m == n + m.
Proof.
intros n m; induct n.
now rewrite def_add_0_l, add_0_l.
@@ -56,42 +58,37 @@ Qed.
(*****************************************************)
(** Multiplication *)
-Definition def_mul (x y : N) := recursion 0 (fun _ p => p ++ x) y.
+Definition def_mul (x y : N.t) := recursion 0 (fun _ p => p +++ x) y.
-Infix Local "**" := def_mul (at level 40, left associativity).
+Local Infix "**" := def_mul (at level 40, left associativity).
-Lemma def_mul_step_wd : forall x : N, fun2_wd Neq Neq Neq (fun _ p => def_add p x).
+Instance def_mul_prewd :
+ Proper (N.eq==>N.eq==>N.eq==>N.eq) (fun x _ p => p +++ x).
Proof.
-unfold fun2_wd. intros. now apply def_add_wd.
+repeat red; intros; now apply def_add_wd.
Qed.
-Lemma def_mul_step_equal :
- forall x x' : N, x == x' ->
- fun2_eq Neq Neq Neq (fun _ p => def_add p x) (fun x p => def_add p x').
-Proof.
-unfold fun2_eq; intros; apply def_add_wd; assumption.
-Qed.
-
-Add Morphism def_mul with signature Neq ==> Neq ==> Neq as def_mul_wd.
+Instance def_mul_wd : Proper (N.eq ==> N.eq ==> N.eq) def_mul.
Proof.
unfold def_mul.
intros x x' Exx' y y' Eyy'.
-apply recursion_wd with (Aeq := Neq).
-reflexivity. apply def_mul_step_equal. assumption. assumption.
+apply recursion_wd; auto with *.
+now apply def_mul_prewd.
Qed.
-Theorem def_mul_0_r : forall x : N, x ** 0 == 0.
+Theorem def_mul_0_r : forall x, x ** 0 == 0.
Proof.
intro. unfold def_mul. now rewrite recursion_0.
Qed.
-Theorem def_mul_succ_r : forall x y : N, x ** S y == x ** y ++ x.
+Theorem def_mul_succ_r : forall x y, x ** S y == x ** y +++ x.
Proof.
intros x y; unfold def_mul.
-now rewrite (@recursion_succ N Neq); [| apply def_mul_step_wd |].
+rewrite recursion_succ; auto with *.
+now apply def_mul_prewd.
Qed.
-Theorem def_mul_mul : forall n m : N, n ** m == n * m.
+Theorem def_mul_mul : forall n m, n ** m == n * m.
Proof.
intros n m; induct m.
now rewrite def_mul_0_r, mul_0_r.
@@ -101,120 +98,99 @@ Qed.
(*****************************************************)
(** Order *)
-Definition def_ltb (m : N) : N -> bool :=
+Definition ltb (m : N.t) : N.t -> bool :=
recursion
(if_zero false true)
- (fun _ f => fun n => recursion false (fun n' _ => f n') n)
+ (fun _ f n => recursion false (fun n' _ => f n') n)
m.
-Infix Local "<<" := def_ltb (at level 70, no associativity).
-
-Lemma lt_base_wd : fun_wd Neq (@eq bool) (if_zero false true).
-unfold fun_wd; intros; now apply if_zero_wd.
-Qed.
+Local Infix "<<" := ltb (at level 70, no associativity).
-Lemma lt_step_wd :
-fun2_wd Neq (fun_eq Neq (@eq bool)) (fun_eq Neq (@eq bool))
- (fun _ f => fun n => recursion false (fun n' _ => f n') n).
+Instance ltb_prewd1 : Proper (N.eq==>Logic.eq) (if_zero false true).
Proof.
-unfold fun2_wd, fun_eq.
-intros x x' Exx' f f' Eff' y y' Eyy'.
-apply recursion_wd with (Aeq := @eq bool).
-reflexivity.
-unfold fun2_eq; intros; now apply Eff'.
-assumption.
+red; intros; apply if_zero_wd; auto.
Qed.
-Lemma lt_curry_wd :
- forall m m' : N, m == m' -> fun_eq Neq (@eq bool) (def_ltb m) (def_ltb m').
+Instance ltb_prewd2 : Proper (N.eq==>(N.eq==>Logic.eq)==>N.eq==>Logic.eq)
+ (fun _ f n => recursion false (fun n' _ => f n') n).
Proof.
-unfold def_ltb.
-intros m m' Emm'.
-apply recursion_wd with (Aeq := fun_eq Neq (@eq bool)).
-apply lt_base_wd.
-apply lt_step_wd.
-assumption.
+repeat red; intros; simpl.
+apply recursion_wd; auto with *.
+repeat red; auto.
Qed.
-Add Morphism def_ltb with signature Neq ==> Neq ==> (@eq bool) as def_ltb_wd.
+Instance ltb_wd : Proper (N.eq ==> N.eq ==> Logic.eq) ltb.
Proof.
-intros; now apply lt_curry_wd.
+unfold ltb.
+intros n n' Hn m m' Hm.
+apply f_equiv; auto with *.
+apply recursion_wd; auto; [ apply ltb_prewd1 | apply ltb_prewd2 ].
Qed.
-Theorem def_ltb_base : forall n : N, 0 << n = if_zero false true n.
+Theorem ltb_base : forall n, 0 << n = if_zero false true n.
Proof.
-intro n; unfold def_ltb; now rewrite recursion_0.
+intro n; unfold ltb; now rewrite recursion_0.
Qed.
-Theorem def_ltb_step :
- forall m n : N, S m << n = recursion false (fun n' _ => m << n') n.
+Theorem ltb_step :
+ forall m n, S m << n = recursion false (fun n' _ => m << n') n.
Proof.
-intros m n; unfold def_ltb.
-pose proof
- (@recursion_succ
- (N -> bool)
- (fun_eq Neq (@eq bool))
- (if_zero false true)
- (fun _ f => fun n => recursion false (fun n' _ => f n') n)
- lt_base_wd
- lt_step_wd
- m n n) as H.
-now rewrite H.
+intros m n; unfold ltb at 1.
+apply f_equiv; auto with *.
+rewrite recursion_succ by (apply ltb_prewd1||apply ltb_prewd2).
+fold (ltb m).
+repeat red; intros. apply recursion_wd; auto.
+repeat red; intros; now apply ltb_wd.
Qed.
(* Above, we rewrite applications of function. Is it possible to rewrite
functions themselves, i.e., rewrite (recursion lt_base lt_step (S n)) to
lt_step n (recursion lt_base lt_step n)? *)
-Theorem def_ltb_0 : forall n : N, n << 0 = false.
+Theorem ltb_0 : forall n, n << 0 = false.
Proof.
cases n.
-rewrite def_ltb_base; now rewrite if_zero_0.
-intro n; rewrite def_ltb_step. now rewrite recursion_0.
+rewrite ltb_base; now rewrite if_zero_0.
+intro n; rewrite ltb_step. now rewrite recursion_0.
Qed.
-Theorem def_ltb_0_succ : forall n : N, 0 << S n = true.
+Theorem ltb_0_succ : forall n, 0 << S n = true.
Proof.
-intro n; rewrite def_ltb_base; now rewrite if_zero_succ.
+intro n; rewrite ltb_base; now rewrite if_zero_succ.
Qed.
-Theorem succ_def_ltb_mono : forall n m : N, (S n << S m) = (n << m).
+Theorem succ_ltb_mono : forall n m, (S n << S m) = (n << m).
Proof.
intros n m.
-rewrite def_ltb_step. rewrite (@recursion_succ bool (@eq bool)); try reflexivity.
-unfold fun2_wd; intros; now apply def_ltb_wd.
+rewrite ltb_step. rewrite recursion_succ; try reflexivity.
+repeat red; intros; now apply ltb_wd.
Qed.
-Theorem def_ltb_lt : forall n m : N, n << m = true <-> n < m.
+Theorem ltb_lt : forall n m, n << m = true <-> n < m.
Proof.
double_induct n m.
cases m.
-rewrite def_ltb_0. split; intro H; [discriminate H | false_hyp H nlt_0_r].
-intro n. rewrite def_ltb_0_succ. split; intro; [apply lt_0_succ | reflexivity].
-intro n. rewrite def_ltb_0. split; intro H; [discriminate | false_hyp H nlt_0_r].
-intros n m. rewrite succ_def_ltb_mono. now rewrite <- succ_lt_mono.
+rewrite ltb_0. split; intro H; [discriminate H | false_hyp H nlt_0_r].
+intro n. rewrite ltb_0_succ. split; intro; [apply lt_0_succ | reflexivity].
+intro n. rewrite ltb_0. split; intro H; [discriminate | false_hyp H nlt_0_r].
+intros n m. rewrite succ_ltb_mono. now rewrite <- succ_lt_mono.
+Qed.
+
+Theorem ltb_ge : forall n m, n << m = false <-> n >= m.
+Proof.
+intros. rewrite <- not_true_iff_false, ltb_lt. apply nlt_ge.
Qed.
-(*
(*****************************************************)
(** Even *)
-Definition even (x : N) := recursion true (fun _ p => negb p) x.
-
-Lemma even_step_wd : fun2_wd Neq (@eq bool) (@eq bool) (fun x p => if p then false else true).
-Proof.
-unfold fun2_wd.
-intros x x' Exx' b b' Ebb'.
-unfold eq_bool; destruct b; destruct b'; now simpl.
-Qed.
+Definition even (x : N.t) := recursion true (fun _ p => negb p) x.
-Add Morphism even with signature Neq ==> (@eq bool) as even_wd.
+Instance even_wd : Proper (N.eq==>Logic.eq) even.
Proof.
-unfold even; intros.
-apply recursion_wd with (A := bool) (Aeq := (@eq bool)).
-now unfold eq_bool.
-unfold fun2_eq. intros _ _ _ b b' Ebb'. unfold eq_bool; destruct b; destruct b'; now simpl.
-assumption.
+intros n n' Hn. unfold even.
+apply recursion_wd; auto.
+congruence.
Qed.
Theorem even_0 : even 0 = true.
@@ -223,76 +199,281 @@ unfold even.
now rewrite recursion_0.
Qed.
-Theorem even_succ : forall x : N, even (S x) = negb (even x).
+Theorem even_succ : forall x, even (S x) = negb (even x).
Proof.
unfold even.
-intro x; rewrite (recursion_succ (@eq bool)); try reflexivity.
-unfold fun2_wd.
-intros _ _ _ b b' Ebb'. destruct b; destruct b'; now simpl.
+intro x; rewrite recursion_succ; try reflexivity.
+congruence.
Qed.
(*****************************************************)
(** Division by 2 *)
-Definition half_aux (x : N) : N * N :=
- recursion (0, 0) (fun _ p => let (x1, x2) := p in ((S x2, x1))) x.
+Local Notation "a <= b <= c" := (a<=b /\ b<=c).
+Local Notation "a <= b < c" := (a<=b /\ b<c).
+Local Notation "a < b <= c" := (a<b /\ b<=c).
+Local Notation "a < b < c" := (a<b /\ b<c).
+Local Notation "2" := (S 1).
-Definition half (x : N) := snd (half_aux x).
+Definition half_aux (x : N.t) : N.t * N.t :=
+ recursion (0, 0) (fun _ p => let (x1, x2) := p in (S x2, x1)) x.
-Definition E2 := prod_rel Neq Neq.
+Definition half (x : N.t) := snd (half_aux x).
-Add Relation (prod N N) E2
-reflexivity proved by (prod_rel_refl N N Neq Neq E_equiv E_equiv)
-symmetry proved by (prod_rel_sym N N Neq Neq E_equiv E_equiv)
-transitivity proved by (prod_rel_trans N N Neq Neq E_equiv E_equiv)
-as E2_rel.
+Instance half_aux_wd : Proper (N.eq ==> N.eq*N.eq) half_aux.
+Proof.
+intros x x' Hx. unfold half_aux.
+apply recursion_wd; auto with *.
+intros y y' Hy (u,v) (u',v') (Hu,Hv). compute in *.
+rewrite Hu, Hv; auto with *.
+Qed.
-Lemma half_step_wd: fun2_wd Neq E2 E2 (fun _ p => let (x1, x2) := p in ((S x2, x1))).
+Instance half_wd : Proper (N.eq==>N.eq) half.
Proof.
-unfold fun2_wd, E2, prod_rel.
-intros _ _ _ p1 p2 [H1 H2].
-destruct p1; destruct p2; simpl in *.
-now split; [rewrite H2 |].
+intros x x' Hx. unfold half. rewrite Hx; auto with *.
Qed.
-Add Morphism half with signature Neq ==> Neq as half_wd.
+Lemma half_aux_0 : half_aux 0 = (0,0).
Proof.
-unfold half.
-assert (H: forall x y, x == y -> E2 (half_aux x) (half_aux y)).
-intros x y Exy; unfold half_aux; apply recursion_wd with (Aeq := E2); unfold E2.
-unfold E2.
-unfold prod_rel; simpl; now split.
-unfold fun2_eq, prod_rel; simpl.
-intros _ _ _ p1 p2; destruct p1; destruct p2; simpl.
-intros [H1 H2]; split; [rewrite H2 | assumption]. reflexivity. assumption.
-unfold E2, prod_rel in H. intros x y Exy; apply H in Exy.
-exact (proj2 Exy).
+unfold half_aux. rewrite recursion_0; auto.
Qed.
+Lemma half_aux_succ : forall x,
+ half_aux (S x) = (S (snd (half_aux x)), fst (half_aux x)).
+Proof.
+intros.
+remember (half_aux x) as h.
+destruct h as (f,s); simpl in *.
+unfold half_aux in *.
+rewrite recursion_succ, <- Heqh; simpl; auto.
+repeat red; intros; subst; auto.
+Qed.
+
+Theorem half_aux_spec : forall n,
+ n == fst (half_aux n) + snd (half_aux n).
+Proof.
+apply induction.
+intros x x' Hx. setoid_rewrite Hx; auto with *.
+rewrite half_aux_0; simpl; rewrite add_0_l; auto with *.
+intros.
+rewrite half_aux_succ. simpl.
+rewrite add_succ_l, add_comm; auto.
+apply succ_wd; auto.
+Qed.
+
+Theorem half_aux_spec2 : forall n,
+ fst (half_aux n) == snd (half_aux n) \/
+ fst (half_aux n) == S (snd (half_aux n)).
+Proof.
+apply induction.
+intros x x' Hx. setoid_rewrite Hx; auto with *.
+rewrite half_aux_0; simpl. auto with *.
+intros.
+rewrite half_aux_succ; simpl.
+destruct H; auto with *.
+right; apply succ_wd; auto with *.
+Qed.
+
+Theorem half_0 : half 0 == 0.
+Proof.
+unfold half. rewrite half_aux_0; simpl; auto with *.
+Qed.
+
+Theorem half_1 : half 1 == 0.
+Proof.
+unfold half. rewrite half_aux_succ, half_aux_0; simpl; auto with *.
+Qed.
+
+Theorem half_double : forall n,
+ n == 2 * half n \/ n == 1 + 2 * half n.
+Proof.
+intros. unfold half.
+nzsimpl.
+destruct (half_aux_spec2 n) as [H|H]; [left|right].
+rewrite <- H at 1. apply half_aux_spec.
+rewrite <- add_succ_l. rewrite <- H at 1. apply half_aux_spec.
+Qed.
+
+Theorem half_upper_bound : forall n, 2 * half n <= n.
+Proof.
+intros.
+destruct (half_double n) as [E|E]; rewrite E at 2.
+apply le_refl.
+nzsimpl.
+apply le_le_succ_r, le_refl.
+Qed.
+
+Theorem half_lower_bound : forall n, n <= 1 + 2 * half n.
+Proof.
+intros.
+destruct (half_double n) as [E|E]; rewrite E at 1.
+nzsimpl.
+apply le_le_succ_r, le_refl.
+apply le_refl.
+Qed.
+
+Theorem half_nz : forall n, 1 < n -> 0 < half n.
+Proof.
+intros n LT.
+assert (LE : 0 <= half n) by apply le_0_l.
+le_elim LE; auto.
+destruct (half_double n) as [E|E];
+ rewrite <- LE, mul_0_r, ?add_0_r in E; rewrite E in LT.
+destruct (nlt_0_r _ LT).
+rewrite <- succ_lt_mono in LT.
+destruct (nlt_0_r _ LT).
+Qed.
+
+Theorem half_decrease : forall n, 0 < n -> half n < n.
+Proof.
+intros n LT.
+destruct (half_double n) as [E|E]; rewrite E at 2;
+ rewrite ?mul_succ_l, ?mul_0_l, ?add_0_l, ?add_assoc.
+rewrite <- add_0_l at 1.
+rewrite <- add_lt_mono_r.
+assert (LE : 0 <= half n) by apply le_0_l.
+le_elim LE; auto.
+rewrite <- LE, mul_0_r in E. rewrite E in LT. destruct (nlt_0_r _ LT).
+rewrite <- add_0_l at 1.
+rewrite <- add_lt_mono_r.
+rewrite add_succ_l. apply lt_0_succ.
+Qed.
+
+
+(*****************************************************)
+(** Power *)
+
+Definition pow (n m : N.t) := recursion 1 (fun _ r => n*r) m.
+
+Local Infix "^^" := pow (at level 30, right associativity).
+
+Instance pow_prewd :
+ Proper (N.eq==>N.eq==>N.eq==>N.eq) (fun n _ r => n*r).
+Proof.
+intros n n' Hn x x' Hx y y' Hy. rewrite Hn, Hy; auto with *.
+Qed.
+
+Instance pow_wd : Proper (N.eq==>N.eq==>N.eq) pow.
+Proof.
+intros n n' Hn m m' Hm. unfold pow.
+apply recursion_wd; auto with *.
+now apply pow_prewd.
+Qed.
+
+Lemma pow_0 : forall n, n^^0 == 1.
+Proof.
+intros. unfold pow. rewrite recursion_0. auto with *.
+Qed.
+
+Lemma pow_succ : forall n m, n^^(S m) == n*(n^^m).
+Proof.
+intros. unfold pow. rewrite recursion_succ; auto with *.
+now apply pow_prewd.
+Qed.
+
+
(*****************************************************)
(** Logarithm for the base 2 *)
-Definition log (x : N) : N :=
+Definition log (x : N.t) : N.t :=
strong_rec 0
- (fun x g =>
- if (e x 0) then 0
- else if (e x 1) then 0
+ (fun g x =>
+ if x << 2 then 0
else S (g (half x)))
x.
-Add Morphism log with signature Neq ==> Neq as log_wd.
+Instance log_prewd :
+ Proper ((N.eq==>N.eq)==>N.eq==>N.eq)
+ (fun g x => if x<<2 then 0 else S (g (half x))).
+Proof.
+intros g g' Hg n n' Hn.
+rewrite Hn.
+destruct (n' << 2); auto with *.
+apply succ_wd.
+apply Hg. rewrite Hn; auto with *.
+Qed.
+
+Instance log_wd : Proper (N.eq==>N.eq) log.
Proof.
intros x x' Exx'. unfold log.
-apply strong_rec_wd with (Aeq := Neq); try (reflexivity || assumption).
-unfold fun2_eq. intros y y' Eyy' g g' Egg'.
-assert (H : e y 0 = e y' 0); [now apply e_wd|].
-rewrite <- H; clear H.
-assert (H : e y 1 = e y' 1); [now apply e_wd|].
-rewrite <- H; clear H.
-assert (H : S (g (half y)) == S (g' (half y')));
-[apply succ_wd; apply Egg'; now apply half_wd|].
-now destruct (e y 0); destruct (e y 1).
+apply strong_rec_wd; auto with *.
+apply log_prewd.
Qed.
+
+Lemma log_good_step : forall n h1 h2,
+ (forall m, m < n -> h1 m == h2 m) ->
+ (if n << 2 then 0 else S (h1 (half n))) ==
+ (if n << 2 then 0 else S (h2 (half n))).
+Proof.
+intros n h1 h2 E.
+destruct (n<<2) as [ ]_eqn:H.
+auto with *.
+apply succ_wd, E, half_decrease.
+rewrite <- not_true_iff_false, ltb_lt, nlt_ge, le_succ_l in H.
+apply lt_succ_l; auto.
+Qed.
+Hint Resolve log_good_step.
+
+Theorem log_init : forall n, n < 2 -> log n == 0.
+Proof.
+intros n Hn. unfold log. rewrite strong_rec_fixpoint; auto with *.
+replace (n << 2) with true; auto with *.
+symmetry. now rewrite ltb_lt.
+Qed.
+
+Theorem log_step : forall n, 2 <= n -> log n == S (log (half n)).
+Proof.
+intros n Hn. unfold log. rewrite strong_rec_fixpoint; auto with *.
+replace (n << 2) with false; auto with *.
+symmetry. rewrite <- not_true_iff_false, ltb_lt, nlt_ge; auto.
+Qed.
+
+Theorem pow2_log : forall n, 0 < n -> half n < 2^^(log n) <= n.
+Proof.
+intro n; generalize (le_refl n). set (k:=n) at -2. clearbody k.
+revert k. pattern n. apply induction; clear n.
+intros n n' Hn; setoid_rewrite Hn; auto with *.
+intros k Hk1 Hk2.
+ le_elim Hk1. destruct (nlt_0_r _ Hk1).
+ rewrite Hk1 in Hk2. destruct (nlt_0_r _ Hk2).
+
+intros n IH k Hk1 Hk2.
+destruct (lt_ge_cases k 2) as [LT|LE].
+(* base *)
+rewrite log_init, pow_0 by auto.
+rewrite <- le_succ_l in Hk2.
+le_elim Hk2.
+rewrite <- nle_gt, le_succ_l in LT. destruct LT; auto.
+rewrite <- Hk2.
+rewrite half_1; auto using lt_0_1, le_refl.
+(* step *)
+rewrite log_step, pow_succ by auto.
+rewrite le_succ_l in LE.
+destruct (IH (half k)) as (IH1,IH2).
+ rewrite <- lt_succ_r. apply lt_le_trans with k; auto.
+ now apply half_decrease.
+ apply half_nz; auto.
+set (K:=2^^log (half k)) in *; clearbody K.
+split.
+rewrite <- le_succ_l in IH1.
+apply mul_le_mono_l with (p:=2) in IH1.
+eapply lt_le_trans; eauto.
+nzsimpl.
+rewrite lt_succ_r.
+eapply le_trans; [ eapply half_lower_bound | ].
+nzsimpl; apply le_refl.
+eapply le_trans; [ | eapply half_upper_bound ].
+apply mul_le_mono_l; auto.
+Qed.
+
+(** Later:
+
+Theorem log_mul : forall n m, 0 < n -> 0 < m ->
+ log (n*m) == log n + log m.
+
+Theorem log_pow2 : forall n, log (2^^n) = n.
+
*)
+
End NdefOpsPropFunct.
diff --git a/theories/Numbers/Natural/Abstract/NDiv.v b/theories/Numbers/Natural/Abstract/NDiv.v
new file mode 100644
index 00000000..0cb5665a
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NDiv.v
@@ -0,0 +1,239 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Euclidean Division *)
+
+Require Import NAxioms NProperties NZDiv.
+
+Module Type NDivSpecific (Import N : NAxiomsSig')(Import DM : DivMod' N).
+ Axiom mod_upper_bound : forall a b, b ~= 0 -> a mod b < b.
+End NDivSpecific.
+
+Module Type NDivSig := NAxiomsSig <+ DivMod <+ NZDivCommon <+ NDivSpecific.
+Module Type NDivSig' := NAxiomsSig' <+ DivMod' <+ NZDivCommon <+ NDivSpecific.
+
+Module NDivPropFunct (Import N : NDivSig')(Import NP : NPropSig N).
+
+(** We benefit from what already exists for NZ *)
+
+ Module ND <: NZDiv N.
+ Definition div := div.
+ Definition modulo := modulo.
+ Definition div_wd := div_wd.
+ Definition mod_wd := mod_wd.
+ Definition div_mod := div_mod.
+ Lemma mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
+ Proof. split. apply le_0_l. apply mod_upper_bound. order. Qed.
+ End ND.
+ Module Import NZDivP := NZDivPropFunct N NP ND.
+
+ Ltac auto' := try rewrite <- neq_0_lt_0; auto using le_0_l.
+
+(** Let's now state again theorems, but without useless hypothesis. *)
+
+(** Uniqueness theorems *)
+
+Theorem div_mod_unique :
+ forall b q1 q2 r1 r2, r1<b -> r2<b ->
+ b*q1+r1 == b*q2+r2 -> q1 == q2 /\ r1 == r2.
+Proof. intros. apply div_mod_unique with b; auto'. Qed.
+
+Theorem div_unique:
+ forall a b q r, r<b -> a == b*q + r -> q == a/b.
+Proof. intros; apply div_unique with r; auto'. Qed.
+
+Theorem mod_unique:
+ forall a b q r, r<b -> a == b*q + r -> r == a mod b.
+Proof. intros. apply mod_unique with q; auto'. Qed.
+
+(** A division by itself returns 1 *)
+
+Lemma div_same : forall a, a~=0 -> a/a == 1.
+Proof. intros. apply div_same; auto'. Qed.
+
+Lemma mod_same : forall a, a~=0 -> a mod a == 0.
+Proof. intros. apply mod_same; auto'. Qed.
+
+(** A division of a small number by a bigger one yields zero. *)
+
+Theorem div_small: forall a b, a<b -> a/b == 0.
+Proof. intros. apply div_small; auto'. Qed.
+
+(** Same situation, in term of modulo: *)
+
+Theorem mod_small: forall a b, a<b -> a mod b == a.
+Proof. intros. apply mod_small; auto'. Qed.
+
+(** * Basic values of divisions and modulo. *)
+
+Lemma div_0_l: forall a, a~=0 -> 0/a == 0.
+Proof. intros. apply div_0_l; auto'. Qed.
+
+Lemma mod_0_l: forall a, a~=0 -> 0 mod a == 0.
+Proof. intros. apply mod_0_l; auto'. Qed.
+
+Lemma div_1_r: forall a, a/1 == a.
+Proof. intros. apply div_1_r; auto'. Qed.
+
+Lemma mod_1_r: forall a, a mod 1 == 0.
+Proof. intros. apply mod_1_r; auto'. Qed.
+
+Lemma div_1_l: forall a, 1<a -> 1/a == 0.
+Proof. exact div_1_l. Qed.
+
+Lemma mod_1_l: forall a, 1<a -> 1 mod a == 1.
+Proof. exact mod_1_l. Qed.
+
+Lemma div_mul : forall a b, b~=0 -> (a*b)/b == a.
+Proof. intros. apply div_mul; auto'. Qed.
+
+Lemma mod_mul : forall a b, b~=0 -> (a*b) mod b == 0.
+Proof. intros. apply mod_mul; auto'. Qed.
+
+
+(** * Order results about mod and div *)
+
+(** A modulo cannot grow beyond its starting point. *)
+
+Theorem mod_le: forall a b, b~=0 -> a mod b <= a.
+Proof. intros. apply mod_le; auto'. Qed.
+
+Lemma div_str_pos : forall a b, 0<b<=a -> 0 < a/b.
+Proof. exact div_str_pos. Qed.
+
+Lemma div_small_iff : forall a b, b~=0 -> (a/b==0 <-> a<b).
+Proof. intros. apply div_small_iff; auto'. Qed.
+
+Lemma mod_small_iff : forall a b, b~=0 -> (a mod b == a <-> a<b).
+Proof. intros. apply mod_small_iff; auto'. Qed.
+
+Lemma div_str_pos_iff : forall a b, b~=0 -> (0<a/b <-> b<=a).
+Proof. intros. apply div_str_pos_iff; auto'. Qed.
+
+
+(** As soon as the divisor is strictly greater than 1,
+ the division is strictly decreasing. *)
+
+Lemma div_lt : forall a b, 0<a -> 1<b -> a/b < a.
+Proof. exact div_lt. Qed.
+
+(** [le] is compatible with a positive division. *)
+
+Lemma div_le_mono : forall a b c, c~=0 -> a<=b -> a/c <= b/c.
+Proof. intros. apply div_le_mono; auto'. Qed.
+
+Lemma mul_div_le : forall a b, b~=0 -> b*(a/b) <= a.
+Proof. intros. apply mul_div_le; auto'. Qed.
+
+Lemma mul_succ_div_gt: forall a b, b~=0 -> a < b*(S (a/b)).
+Proof. intros; apply mul_succ_div_gt; auto'. Qed.
+
+(** The previous inequality is exact iff the modulo is zero. *)
+
+Lemma div_exact : forall a b, b~=0 -> (a == b*(a/b) <-> a mod b == 0).
+Proof. intros. apply div_exact; auto'. Qed.
+
+(** Some additionnal inequalities about div. *)
+
+Theorem div_lt_upper_bound:
+ forall a b q, b~=0 -> a < b*q -> a/b < q.
+Proof. intros. apply div_lt_upper_bound; auto'. Qed.
+
+Theorem div_le_upper_bound:
+ forall a b q, b~=0 -> a <= b*q -> a/b <= q.
+Proof. intros; apply div_le_upper_bound; auto'. Qed.
+
+Theorem div_le_lower_bound:
+ forall a b q, b~=0 -> b*q <= a -> q <= a/b.
+Proof. intros; apply div_le_lower_bound; auto'. Qed.
+
+(** A division respects opposite monotonicity for the divisor *)
+
+Lemma div_le_compat_l: forall p q r, 0<q<=r -> p/r <= p/q.
+Proof. intros. apply div_le_compat_l. auto'. auto. Qed.
+
+(** * Relations between usual operations and mod and div *)
+
+Lemma mod_add : forall a b c, c~=0 ->
+ (a + b * c) mod c == a mod c.
+Proof. intros. apply mod_add; auto'. Qed.
+
+Lemma div_add : forall a b c, c~=0 ->
+ (a + b * c) / c == a / c + b.
+Proof. intros. apply div_add; auto'. Qed.
+
+Lemma div_add_l: forall a b c, b~=0 ->
+ (a * b + c) / b == a + c / b.
+Proof. intros. apply div_add_l; auto'. Qed.
+
+(** Cancellations. *)
+
+Lemma div_mul_cancel_r : forall a b c, b~=0 -> c~=0 ->
+ (a*c)/(b*c) == a/b.
+Proof. intros. apply div_mul_cancel_r; auto'. Qed.
+
+Lemma div_mul_cancel_l : forall a b c, b~=0 -> c~=0 ->
+ (c*a)/(c*b) == a/b.
+Proof. intros. apply div_mul_cancel_l; auto'. Qed.
+
+Lemma mul_mod_distr_r: forall a b c, b~=0 -> c~=0 ->
+ (a*c) mod (b*c) == (a mod b) * c.
+Proof. intros. apply mul_mod_distr_r; auto'. Qed.
+
+Lemma mul_mod_distr_l: forall a b c, b~=0 -> c~=0 ->
+ (c*a) mod (c*b) == c * (a mod b).
+Proof. intros. apply mul_mod_distr_l; auto'. Qed.
+
+(** Operations modulo. *)
+
+Theorem mod_mod: forall a n, n~=0 ->
+ (a mod n) mod n == a mod n.
+Proof. intros. apply mod_mod; auto'. Qed.
+
+Lemma mul_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)*b) mod n == (a*b) mod n.
+Proof. intros. apply mul_mod_idemp_l; auto'. Qed.
+
+Lemma mul_mod_idemp_r : forall a b n, n~=0 ->
+ (a*(b mod n)) mod n == (a*b) mod n.
+Proof. intros. apply mul_mod_idemp_r; auto'. Qed.
+
+Theorem mul_mod: forall a b n, n~=0 ->
+ (a * b) mod n == ((a mod n) * (b mod n)) mod n.
+Proof. intros. apply mul_mod; auto'. Qed.
+
+Lemma add_mod_idemp_l : forall a b n, n~=0 ->
+ ((a mod n)+b) mod n == (a+b) mod n.
+Proof. intros. apply add_mod_idemp_l; auto'. Qed.
+
+Lemma add_mod_idemp_r : forall a b n, n~=0 ->
+ (a+(b mod n)) mod n == (a+b) mod n.
+Proof. intros. apply add_mod_idemp_r; auto'. Qed.
+
+Theorem add_mod: forall a b n, n~=0 ->
+ (a+b) mod n == (a mod n + b mod n) mod n.
+Proof. intros. apply add_mod; auto'. Qed.
+
+Lemma div_div : forall a b c, b~=0 -> c~=0 ->
+ (a/b)/c == a/(b*c).
+Proof. intros. apply div_div; auto'. Qed.
+
+(** A last inequality: *)
+
+Theorem div_mul_le:
+ forall a b c, b~=0 -> c*(a/b) <= (c*a)/b.
+Proof. intros. apply div_mul_le; auto'. Qed.
+
+(** mod is related to divisibility *)
+
+Lemma mod_divides : forall a b, b~=0 ->
+ (a mod b == 0 <-> exists c, a == b*c).
+Proof. intros. apply mod_divides; auto'. Qed.
+
+End NDivPropFunct.
+
diff --git a/theories/Numbers/Natural/Abstract/NIso.v b/theories/Numbers/Natural/Abstract/NIso.v
index f6ccf3db..47bf38cb 100644
--- a/theories/Numbers/Natural/Abstract/NIso.v
+++ b/theories/Numbers/Natural/Abstract/NIso.v
@@ -8,51 +8,41 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NIso.v 10934 2008-05-15 21:58:20Z letouzey $ i*)
+(*i $Id$ i*)
Require Import NBase.
-Module Homomorphism (NAxiomsMod1 NAxiomsMod2 : NAxiomsSig).
+Module Homomorphism (N1 N2 : NAxiomsSig).
-Module NBasePropMod2 := NBasePropFunct NAxiomsMod2.
+Local Notation "n == m" := (N2.eq n m) (at level 70, no associativity).
-Notation Local N1 := NAxiomsMod1.N.
-Notation Local N2 := NAxiomsMod2.N.
-Notation Local Eq1 := NAxiomsMod1.Neq.
-Notation Local Eq2 := NAxiomsMod2.Neq.
-Notation Local O1 := NAxiomsMod1.N0.
-Notation Local O2 := NAxiomsMod2.N0.
-Notation Local S1 := NAxiomsMod1.S.
-Notation Local S2 := NAxiomsMod2.S.
-Notation Local "n == m" := (Eq2 n m) (at level 70, no associativity).
+Definition homomorphism (f : N1.t -> N2.t) : Prop :=
+ f N1.zero == N2.zero /\ forall n, f (N1.succ n) == N2.succ (f n).
-Definition homomorphism (f : N1 -> N2) : Prop :=
- f O1 == O2 /\ forall n : N1, f (S1 n) == S2 (f n).
+Definition natural_isomorphism : N1.t -> N2.t :=
+ N1.recursion N2.zero (fun (n : N1.t) (p : N2.t) => N2.succ p).
-Definition natural_isomorphism : N1 -> N2 :=
- NAxiomsMod1.recursion O2 (fun (n : N1) (p : N2) => S2 p).
-
-Add Morphism natural_isomorphism with signature Eq1 ==> Eq2 as natural_isomorphism_wd.
+Instance natural_isomorphism_wd : Proper (N1.eq ==> N2.eq) natural_isomorphism.
Proof.
unfold natural_isomorphism.
intros n m Eqxy.
-apply NAxiomsMod1.recursion_wd with (Aeq := Eq2).
+apply N1.recursion_wd.
reflexivity.
-unfold fun2_eq. intros _ _ _ y' y'' H. now apply NBasePropMod2.succ_wd.
+intros _ _ _ y' y'' H. now apply N2.succ_wd.
assumption.
Qed.
-Theorem natural_isomorphism_0 : natural_isomorphism O1 == O2.
+Theorem natural_isomorphism_0 : natural_isomorphism N1.zero == N2.zero.
Proof.
-unfold natural_isomorphism; now rewrite NAxiomsMod1.recursion_0.
+unfold natural_isomorphism; now rewrite N1.recursion_0.
Qed.
Theorem natural_isomorphism_succ :
- forall n : N1, natural_isomorphism (S1 n) == S2 (natural_isomorphism n).
+ forall n : N1.t, natural_isomorphism (N1.succ n) == N2.succ (natural_isomorphism n).
Proof.
unfold natural_isomorphism.
-intro n. now rewrite (@NAxiomsMod1.recursion_succ N2 NAxiomsMod2.Neq) ;
-[ | | unfold fun2_wd; intros; apply NBasePropMod2.succ_wd].
+intro n. rewrite N1.recursion_succ; auto with *.
+repeat red; intros. apply N2.succ_wd; auto.
Qed.
Theorem hom_nat_iso : homomorphism natural_isomorphism.
@@ -63,23 +53,20 @@ Qed.
End Homomorphism.
-Module Inverse (NAxiomsMod1 NAxiomsMod2 : NAxiomsSig).
+Module Inverse (N1 N2 : NAxiomsSig).
-Module Import NBasePropMod1 := NBasePropFunct NAxiomsMod1.
+Module Import NBasePropMod1 := NBasePropFunct N1.
(* This makes the tactic induct available. Since it is taken from
(NBasePropFunct NAxiomsMod1), it refers to induction on N1. *)
-Module Hom12 := Homomorphism NAxiomsMod1 NAxiomsMod2.
-Module Hom21 := Homomorphism NAxiomsMod2 NAxiomsMod1.
-
-Notation Local N1 := NAxiomsMod1.N.
-Notation Local N2 := NAxiomsMod2.N.
-Notation Local h12 := Hom12.natural_isomorphism.
-Notation Local h21 := Hom21.natural_isomorphism.
+Module Hom12 := Homomorphism N1 N2.
+Module Hom21 := Homomorphism N2 N1.
-Notation Local "n == m" := (NAxiomsMod1.Neq n m) (at level 70, no associativity).
+Local Notation h12 := Hom12.natural_isomorphism.
+Local Notation h21 := Hom21.natural_isomorphism.
+Local Notation "n == m" := (N1.eq n m) (at level 70, no associativity).
-Lemma inverse_nat_iso : forall n : N1, h21 (h12 n) == n.
+Lemma inverse_nat_iso : forall n : N1.t, h21 (h12 n) == n.
Proof.
induct n.
now rewrite Hom12.natural_isomorphism_0, Hom21.natural_isomorphism_0.
@@ -89,25 +76,20 @@ Qed.
End Inverse.
-Module Isomorphism (NAxiomsMod1 NAxiomsMod2 : NAxiomsSig).
-
-Module Hom12 := Homomorphism NAxiomsMod1 NAxiomsMod2.
-Module Hom21 := Homomorphism NAxiomsMod2 NAxiomsMod1.
+Module Isomorphism (N1 N2 : NAxiomsSig).
-Module Inverse12 := Inverse NAxiomsMod1 NAxiomsMod2.
-Module Inverse21 := Inverse NAxiomsMod2 NAxiomsMod1.
+Module Hom12 := Homomorphism N1 N2.
+Module Hom21 := Homomorphism N2 N1.
+Module Inverse12 := Inverse N1 N2.
+Module Inverse21 := Inverse N2 N1.
-Notation Local N1 := NAxiomsMod1.N.
-Notation Local N2 := NAxiomsMod2.N.
-Notation Local Eq1 := NAxiomsMod1.Neq.
-Notation Local Eq2 := NAxiomsMod2.Neq.
-Notation Local h12 := Hom12.natural_isomorphism.
-Notation Local h21 := Hom21.natural_isomorphism.
+Local Notation h12 := Hom12.natural_isomorphism.
+Local Notation h21 := Hom21.natural_isomorphism.
-Definition isomorphism (f1 : N1 -> N2) (f2 : N2 -> N1) : Prop :=
+Definition isomorphism (f1 : N1.t -> N2.t) (f2 : N2.t -> N1.t) : Prop :=
Hom12.homomorphism f1 /\ Hom21.homomorphism f2 /\
- forall n : N1, Eq1 (f2 (f1 n)) n /\
- forall n : N2, Eq2 (f1 (f2 n)) n.
+ forall n, N1.eq (f2 (f1 n)) n /\
+ forall n, N2.eq (f1 (f2 n)) n.
Theorem iso_nat_iso : isomorphism h12 h21.
Proof.
diff --git a/theories/Numbers/Natural/Abstract/NMul.v b/theories/Numbers/Natural/Abstract/NMul.v
deleted file mode 100644
index 0b00f689..00000000
--- a/theories/Numbers/Natural/Abstract/NMul.v
+++ /dev/null
@@ -1,87 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Evgeny Makarov, INRIA, 2007 *)
-(************************************************************************)
-
-(*i $Id: NMul.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
-
-Require Export NAdd.
-
-Module NMulPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NAddPropMod := NAddPropFunct NAxiomsMod.
-Open Local Scope NatScope.
-
-Theorem mul_wd :
- forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> n1 * m1 == n2 * m2.
-Proof NZmul_wd.
-
-Theorem mul_0_l : forall n : N, 0 * n == 0.
-Proof NZmul_0_l.
-
-Theorem mul_succ_l : forall n m : N, (S n) * m == n * m + m.
-Proof NZmul_succ_l.
-
-(** Theorems that are valid for both natural numbers and integers *)
-
-Theorem mul_0_r : forall n, n * 0 == 0.
-Proof NZmul_0_r.
-
-Theorem mul_succ_r : forall n m, n * (S m) == n * m + n.
-Proof NZmul_succ_r.
-
-Theorem mul_comm : forall n m : N, n * m == m * n.
-Proof NZmul_comm.
-
-Theorem mul_add_distr_r : forall n m p : N, (n + m) * p == n * p + m * p.
-Proof NZmul_add_distr_r.
-
-Theorem mul_add_distr_l : forall n m p : N, n * (m + p) == n * m + n * p.
-Proof NZmul_add_distr_l.
-
-Theorem mul_assoc : forall n m p : N, n * (m * p) == (n * m) * p.
-Proof NZmul_assoc.
-
-Theorem mul_1_l : forall n : N, 1 * n == n.
-Proof NZmul_1_l.
-
-Theorem mul_1_r : forall n : N, n * 1 == n.
-Proof NZmul_1_r.
-
-(* Theorems that cannot be proved in NZMul *)
-
-(* In proving the correctness of the definition of multiplication on
-integers constructed from pairs of natural numbers, we'll need the
-following fact about natural numbers:
-
-a * n + u == a * m + v -> n + m' == n' + m -> a * n' + u = a * m' + v
-
-Here n + m' == n' + m expresses equality of integers (n, m) and (n', m'),
-since a pair (a, b) of natural numbers represents the integer a - b. On
-integers, the formula above could be proved by moving a * m to the left,
-factoring out a and replacing n - m by n' - m'. However, the formula is
-required in the process of constructing integers, so it has to be proved
-for natural numbers, where terms cannot be moved from one side of an
-equation to the other. The proof uses the cancellation laws add_cancel_l
-and add_cancel_r. *)
-
-Theorem add_mul_repl_pair : forall a n m n' m' u v : N,
- a * n + u == a * m + v -> n + m' == n' + m -> a * n' + u == a * m' + v.
-Proof.
-intros a n m n' m' u v H1 H2.
-apply (@NZmul_wd a a) in H2; [| reflexivity].
-do 2 rewrite mul_add_distr_l in H2. symmetry in H2.
-pose proof (NZadd_wd _ _ H1 _ _ H2) as H3.
-rewrite (add_shuffle1 (a * m)), (add_comm (a * m) (a * n)) in H3.
-do 2 rewrite <- add_assoc in H3. apply -> add_cancel_l in H3.
-rewrite (add_assoc u), (add_comm (a * m)) in H3.
-apply -> add_cancel_r in H3.
-now rewrite (add_comm (a * n') u), (add_comm (a * m') v).
-Qed.
-
-End NMulPropFunct.
-
diff --git a/theories/Numbers/Natural/Abstract/NMulOrder.v b/theories/Numbers/Natural/Abstract/NMulOrder.v
index aa21fb50..a2162b13 100644
--- a/theories/Numbers/Natural/Abstract/NMulOrder.v
+++ b/theories/Numbers/Natural/Abstract/NMulOrder.v
@@ -8,122 +8,71 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NMulOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export NAddOrder.
-Module NMulOrderPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NAddOrderPropMod := NAddOrderPropFunct NAxiomsMod.
-Open Local Scope NatScope.
+Module NMulOrderPropFunct (Import N : NAxiomsSig').
+Include NAddOrderPropFunct N.
-Theorem mul_lt_pred :
- forall p q n m : N, S p == q -> (p * n < p * m <-> q * n + m < q * m + n).
-Proof NZmul_lt_pred.
+(** Theorems that are either not valid on Z or have different proofs
+ on N and Z *)
-Theorem mul_lt_mono_pos_l : forall p n m : N, 0 < p -> (n < m <-> p * n < p * m).
-Proof NZmul_lt_mono_pos_l.
-
-Theorem mul_lt_mono_pos_r : forall p n m : N, 0 < p -> (n < m <-> n * p < m * p).
-Proof NZmul_lt_mono_pos_r.
-
-Theorem mul_cancel_l : forall n m p : N, p ~= 0 -> (p * n == p * m <-> n == m).
-Proof NZmul_cancel_l.
-
-Theorem mul_cancel_r : forall n m p : N, p ~= 0 -> (n * p == m * p <-> n == m).
-Proof NZmul_cancel_r.
-
-Theorem mul_id_l : forall n m : N, m ~= 0 -> (n * m == m <-> n == 1).
-Proof NZmul_id_l.
-
-Theorem mul_id_r : forall n m : N, n ~= 0 -> (n * m == n <-> m == 1).
-Proof NZmul_id_r.
-
-Theorem mul_le_mono_pos_l : forall n m p : N, 0 < p -> (n <= m <-> p * n <= p * m).
-Proof NZmul_le_mono_pos_l.
-
-Theorem mul_le_mono_pos_r : forall n m p : N, 0 < p -> (n <= m <-> n * p <= m * p).
-Proof NZmul_le_mono_pos_r.
-
-Theorem mul_pos_pos : forall n m : N, 0 < n -> 0 < m -> 0 < n * m.
-Proof NZmul_pos_pos.
-
-Theorem lt_1_mul_pos : forall n m : N, 1 < n -> 0 < m -> 1 < n * m.
-Proof NZlt_1_mul_pos.
-
-Theorem eq_mul_0 : forall n m : N, n * m == 0 <-> n == 0 \/ m == 0.
-Proof NZeq_mul_0.
-
-Theorem neq_mul_0 : forall n m : N, n ~= 0 /\ m ~= 0 <-> n * m ~= 0.
-Proof NZneq_mul_0.
-
-Theorem eq_square_0 : forall n : N, n * n == 0 <-> n == 0.
-Proof NZeq_square_0.
-
-Theorem eq_mul_0_l : forall n m : N, n * m == 0 -> m ~= 0 -> n == 0.
-Proof NZeq_mul_0_l.
-
-Theorem eq_mul_0_r : forall n m : N, n * m == 0 -> n ~= 0 -> m == 0.
-Proof NZeq_mul_0_r.
-
-Theorem square_lt_mono : forall n m : N, n < m <-> n * n < m * m.
+Theorem square_lt_mono : forall n m, n < m <-> n * n < m * m.
Proof.
intros n m; split; intro;
-[apply NZsquare_lt_mono_nonneg | apply NZsquare_lt_simpl_nonneg];
+[apply square_lt_mono_nonneg | apply square_lt_simpl_nonneg];
try assumption; apply le_0_l.
Qed.
-Theorem square_le_mono : forall n m : N, n <= m <-> n * n <= m * m.
+Theorem square_le_mono : forall n m, n <= m <-> n * n <= m * m.
Proof.
intros n m; split; intro;
-[apply NZsquare_le_mono_nonneg | apply NZsquare_le_simpl_nonneg];
+[apply square_le_mono_nonneg | apply square_le_simpl_nonneg];
try assumption; apply le_0_l.
Qed.
-Theorem mul_2_mono_l : forall n m : N, n < m -> 1 + (1 + 1) * n < (1 + 1) * m.
-Proof NZmul_2_mono_l.
-
-(* Theorems that are either not valid on Z or have different proofs on N and Z *)
-
-Theorem mul_le_mono_l : forall n m p : N, n <= m -> p * n <= p * m.
+Theorem mul_le_mono_l : forall n m p, n <= m -> p * n <= p * m.
Proof.
-intros; apply NZmul_le_mono_nonneg_l. apply le_0_l. assumption.
+intros; apply mul_le_mono_nonneg_l. apply le_0_l. assumption.
Qed.
-Theorem mul_le_mono_r : forall n m p : N, n <= m -> n * p <= m * p.
+Theorem mul_le_mono_r : forall n m p, n <= m -> n * p <= m * p.
Proof.
-intros; apply NZmul_le_mono_nonneg_r. apply le_0_l. assumption.
+intros; apply mul_le_mono_nonneg_r. apply le_0_l. assumption.
Qed.
-Theorem mul_lt_mono : forall n m p q : N, n < m -> p < q -> n * p < m * q.
+Theorem mul_lt_mono : forall n m p q, n < m -> p < q -> n * p < m * q.
Proof.
-intros; apply NZmul_lt_mono_nonneg; try assumption; apply le_0_l.
+intros; apply mul_lt_mono_nonneg; try assumption; apply le_0_l.
Qed.
-Theorem mul_le_mono : forall n m p q : N, n <= m -> p <= q -> n * p <= m * q.
+Theorem mul_le_mono : forall n m p q, n <= m -> p <= q -> n * p <= m * q.
Proof.
-intros; apply NZmul_le_mono_nonneg; try assumption; apply le_0_l.
+intros; apply mul_le_mono_nonneg; try assumption; apply le_0_l.
Qed.
-Theorem lt_0_mul : forall n m : N, n * m > 0 <-> n > 0 /\ m > 0.
+Theorem lt_0_mul' : forall n m, n * m > 0 <-> n > 0 /\ m > 0.
Proof.
intros n m; split; [intro H | intros [H1 H2]].
-apply -> NZlt_0_mul in H. destruct H as [[H1 H2] | [H1 H2]]. now split. false_hyp H1 nlt_0_r.
-now apply NZmul_pos_pos.
+apply -> lt_0_mul in H. destruct H as [[H1 H2] | [H1 H2]]. now split.
+ false_hyp H1 nlt_0_r.
+now apply mul_pos_pos.
Qed.
-Notation mul_pos := lt_0_mul (only parsing).
+Notation mul_pos := lt_0_mul' (only parsing).
-Theorem eq_mul_1 : forall n m : N, n * m == 1 <-> n == 1 /\ m == 1.
+Theorem eq_mul_1 : forall n m, n * m == 1 <-> n == 1 /\ m == 1.
Proof.
intros n m.
split; [| intros [H1 H2]; now rewrite H1, H2, mul_1_l].
-intro H; destruct (NZlt_trichotomy n 1) as [H1 | [H1 | H1]].
+intro H; destruct (lt_trichotomy n 1) as [H1 | [H1 | H1]].
apply -> lt_1_r in H1. rewrite H1, mul_0_l in H. false_hyp H neq_0_succ.
rewrite H1, mul_1_l in H; now split.
destruct (eq_0_gt_0_cases m) as [H2 | H2].
rewrite H2, mul_0_r in H; false_hyp H neq_0_succ.
apply -> (mul_lt_mono_pos_r m) in H1; [| assumption]. rewrite mul_1_l in H1.
-assert (H3 : 1 < n * m) by now apply (lt_1_l 0 m).
+assert (H3 : 1 < n * m) by now apply (lt_1_l m).
rewrite H in H3; false_hyp H3 lt_irrefl.
Qed.
diff --git a/theories/Numbers/Natural/Abstract/NOrder.v b/theories/Numbers/Natural/Abstract/NOrder.v
index 15aed7ab..090c02ec 100644
--- a/theories/Numbers/Natural/Abstract/NOrder.v
+++ b/theories/Numbers/Natural/Abstract/NOrder.v
@@ -8,355 +8,62 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NOrder.v 11282 2008-07-28 11:51:53Z msozeau $ i*)
+(*i $Id$ i*)
-Require Export NMul.
+Require Export NAdd.
-Module NOrderPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NMulPropMod := NMulPropFunct NAxiomsMod.
-Open Local Scope NatScope.
+Module NOrderPropFunct (Import N : NAxiomsSig').
+Include NAddPropFunct N.
-(* The tactics le_less, le_equal and le_elim are inherited from NZOrder.v *)
-
-(* Axioms *)
-
-Theorem lt_wd :
- forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> (n1 < m1 <-> n2 < m2).
-Proof NZlt_wd.
-
-Theorem le_wd :
- forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> (n1 <= m1 <-> n2 <= m2).
-Proof NZle_wd.
-
-Theorem min_wd :
- forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> min n1 m1 == min n2 m2.
-Proof NZmin_wd.
-
-Theorem max_wd :
- forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> max n1 m1 == max n2 m2.
-Proof NZmax_wd.
-
-Theorem lt_eq_cases : forall n m : N, n <= m <-> n < m \/ n == m.
-Proof NZlt_eq_cases.
-
-Theorem lt_irrefl : forall n : N, ~ n < n.
-Proof NZlt_irrefl.
-
-Theorem lt_succ_r : forall n m : N, n < S m <-> n <= m.
-Proof NZlt_succ_r.
-
-Theorem min_l : forall n m : N, n <= m -> min n m == n.
-Proof NZmin_l.
-
-Theorem min_r : forall n m : N, m <= n -> min n m == m.
-Proof NZmin_r.
-
-Theorem max_l : forall n m : N, m <= n -> max n m == n.
-Proof NZmax_l.
-
-Theorem max_r : forall n m : N, n <= m -> max n m == m.
-Proof NZmax_r.
-
-(* Renaming theorems from NZOrder.v *)
-
-Theorem lt_le_incl : forall n m : N, n < m -> n <= m.
-Proof NZlt_le_incl.
-
-Theorem eq_le_incl : forall n m : N, n == m -> n <= m.
-Proof NZeq_le_incl.
-
-Theorem lt_neq : forall n m : N, n < m -> n ~= m.
-Proof NZlt_neq.
-
-Theorem le_neq : forall n m : N, n < m <-> n <= m /\ n ~= m.
-Proof NZle_neq.
-
-Theorem le_refl : forall n : N, n <= n.
-Proof NZle_refl.
-
-Theorem lt_succ_diag_r : forall n : N, n < S n.
-Proof NZlt_succ_diag_r.
-
-Theorem le_succ_diag_r : forall n : N, n <= S n.
-Proof NZle_succ_diag_r.
-
-Theorem lt_0_1 : 0 < 1.
-Proof NZlt_0_1.
-
-Theorem le_0_1 : 0 <= 1.
-Proof NZle_0_1.
-
-Theorem lt_lt_succ_r : forall n m : N, n < m -> n < S m.
-Proof NZlt_lt_succ_r.
-
-Theorem le_le_succ_r : forall n m : N, n <= m -> n <= S m.
-Proof NZle_le_succ_r.
-
-Theorem le_succ_r : forall n m : N, n <= S m <-> n <= m \/ n == S m.
-Proof NZle_succ_r.
-
-Theorem neq_succ_diag_l : forall n : N, S n ~= n.
-Proof NZneq_succ_diag_l.
-
-Theorem neq_succ_diag_r : forall n : N, n ~= S n.
-Proof NZneq_succ_diag_r.
-
-Theorem nlt_succ_diag_l : forall n : N, ~ S n < n.
-Proof NZnlt_succ_diag_l.
-
-Theorem nle_succ_diag_l : forall n : N, ~ S n <= n.
-Proof NZnle_succ_diag_l.
-
-Theorem le_succ_l : forall n m : N, S n <= m <-> n < m.
-Proof NZle_succ_l.
-
-Theorem lt_succ_l : forall n m : N, S n < m -> n < m.
-Proof NZlt_succ_l.
-
-Theorem succ_lt_mono : forall n m : N, n < m <-> S n < S m.
-Proof NZsucc_lt_mono.
-
-Theorem succ_le_mono : forall n m : N, n <= m <-> S n <= S m.
-Proof NZsucc_le_mono.
-
-Theorem lt_asymm : forall n m : N, n < m -> ~ m < n.
-Proof NZlt_asymm.
-
-Notation lt_ngt := lt_asymm (only parsing).
-
-Theorem lt_trans : forall n m p : N, n < m -> m < p -> n < p.
-Proof NZlt_trans.
-
-Theorem le_trans : forall n m p : N, n <= m -> m <= p -> n <= p.
-Proof NZle_trans.
-
-Theorem le_lt_trans : forall n m p : N, n <= m -> m < p -> n < p.
-Proof NZle_lt_trans.
-
-Theorem lt_le_trans : forall n m p : N, n < m -> m <= p -> n < p.
-Proof NZlt_le_trans.
-
-Theorem le_antisymm : forall n m : N, n <= m -> m <= n -> n == m.
-Proof NZle_antisymm.
-
-(** Trichotomy, decidability, and double negation elimination *)
-
-Theorem lt_trichotomy : forall n m : N, n < m \/ n == m \/ m < n.
-Proof NZlt_trichotomy.
-
-Notation lt_eq_gt_cases := lt_trichotomy (only parsing).
-
-Theorem lt_gt_cases : forall n m : N, n ~= m <-> n < m \/ n > m.
-Proof NZlt_gt_cases.
-
-Theorem le_gt_cases : forall n m : N, n <= m \/ n > m.
-Proof NZle_gt_cases.
-
-Theorem lt_ge_cases : forall n m : N, n < m \/ n >= m.
-Proof NZlt_ge_cases.
-
-Theorem le_ge_cases : forall n m : N, n <= m \/ n >= m.
-Proof NZle_ge_cases.
-
-Theorem le_ngt : forall n m : N, n <= m <-> ~ n > m.
-Proof NZle_ngt.
-
-Theorem nlt_ge : forall n m : N, ~ n < m <-> n >= m.
-Proof NZnlt_ge.
-
-Theorem lt_dec : forall n m : N, decidable (n < m).
-Proof NZlt_dec.
-
-Theorem lt_dne : forall n m : N, ~ ~ n < m <-> n < m.
-Proof NZlt_dne.
-
-Theorem nle_gt : forall n m : N, ~ n <= m <-> n > m.
-Proof NZnle_gt.
-
-Theorem lt_nge : forall n m : N, n < m <-> ~ n >= m.
-Proof NZlt_nge.
-
-Theorem le_dec : forall n m : N, decidable (n <= m).
-Proof NZle_dec.
-
-Theorem le_dne : forall n m : N, ~ ~ n <= m <-> n <= m.
-Proof NZle_dne.
-
-Theorem nlt_succ_r : forall n m : N, ~ m < S n <-> n < m.
-Proof NZnlt_succ_r.
-
-Theorem lt_exists_pred :
- forall z n : N, z < n -> exists k : N, n == S k /\ z <= k.
-Proof NZlt_exists_pred.
-
-Theorem lt_succ_iter_r :
- forall (n : nat) (m : N), m < NZsucc_iter (Datatypes.S n) m.
-Proof NZlt_succ_iter_r.
-
-Theorem neq_succ_iter_l :
- forall (n : nat) (m : N), NZsucc_iter (Datatypes.S n) m ~= m.
-Proof NZneq_succ_iter_l.
-
-(** Stronger variant of induction with assumptions n >= 0 (n < 0)
-in the induction step *)
-
-Theorem right_induction :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N, A z ->
- (forall n : N, z <= n -> A n -> A (S n)) ->
- forall n : N, z <= n -> A n.
-Proof NZright_induction.
-
-Theorem left_induction :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N, A z ->
- (forall n : N, n < z -> A (S n) -> A n) ->
- forall n : N, n <= z -> A n.
-Proof NZleft_induction.
-
-Theorem right_induction' :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N,
- (forall n : N, n <= z -> A n) ->
- (forall n : N, z <= n -> A n -> A (S n)) ->
- forall n : N, A n.
-Proof NZright_induction'.
-
-Theorem left_induction' :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N,
- (forall n : N, z <= n -> A n) ->
- (forall n : N, n < z -> A (S n) -> A n) ->
- forall n : N, A n.
-Proof NZleft_induction'.
-
-Theorem strong_right_induction :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N,
- (forall n : N, z <= n -> (forall m : N, z <= m -> m < n -> A m) -> A n) ->
- forall n : N, z <= n -> A n.
-Proof NZstrong_right_induction.
-
-Theorem strong_left_induction :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N,
- (forall n : N, n <= z -> (forall m : N, m <= z -> S n <= m -> A m) -> A n) ->
- forall n : N, n <= z -> A n.
-Proof NZstrong_left_induction.
-
-Theorem strong_right_induction' :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N,
- (forall n : N, n <= z -> A n) ->
- (forall n : N, z <= n -> (forall m : N, z <= m -> m < n -> A m) -> A n) ->
- forall n : N, A n.
-Proof NZstrong_right_induction'.
-
-Theorem strong_left_induction' :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N,
- (forall n : N, z <= n -> A n) ->
- (forall n : N, n <= z -> (forall m : N, m <= z -> S n <= m -> A m) -> A n) ->
- forall n : N, A n.
-Proof NZstrong_left_induction'.
-
-Theorem order_induction :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N, A z ->
- (forall n : N, z <= n -> A n -> A (S n)) ->
- (forall n : N, n < z -> A (S n) -> A n) ->
- forall n : N, A n.
-Proof NZorder_induction.
-
-Theorem order_induction' :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall z : N, A z ->
- (forall n : N, z <= n -> A n -> A (S n)) ->
- (forall n : N, n <= z -> A n -> A (P n)) ->
- forall n : N, A n.
-Proof NZorder_induction'.
-
-(* We don't need order_induction_0 and order_induction'_0 (see NZOrder and
-ZOrder) since they boil down to regular induction *)
-
-(** Elimintation principle for < *)
-
-Theorem lt_ind :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall n : N,
- A (S n) ->
- (forall m : N, n < m -> A m -> A (S m)) ->
- forall m : N, n < m -> A m.
-Proof NZlt_ind.
-
-(** Elimintation principle for <= *)
-
-Theorem le_ind :
- forall A : N -> Prop, predicate_wd Neq A ->
- forall n : N,
- A n ->
- (forall m : N, n <= m -> A m -> A (S m)) ->
- forall m : N, n <= m -> A m.
-Proof NZle_ind.
-
-(** Well-founded relations *)
-
-Theorem lt_wf : forall z : N, well_founded (fun n m : N => z <= n /\ n < m).
-Proof NZlt_wf.
-
-Theorem gt_wf : forall z : N, well_founded (fun n m : N => m < n /\ n <= z).
-Proof NZgt_wf.
+(* Theorems that are true for natural numbers but not for integers *)
Theorem lt_wf_0 : well_founded lt.
Proof.
-setoid_replace lt with (fun n m : N => 0 <= n /\ n < m)
- using relation (@relations_eq N N).
+setoid_replace lt with (fun n m => 0 <= n /\ n < m).
apply lt_wf.
intros x y; split.
intro H; split; [apply le_0_l | assumption]. now intros [_ H].
Defined.
-(* Theorems that are true for natural numbers but not for integers *)
-
(* "le_0_l : forall n : N, 0 <= n" was proved in NBase.v *)
-Theorem nlt_0_r : forall n : N, ~ n < 0.
+Theorem nlt_0_r : forall n, ~ n < 0.
Proof.
intro n; apply -> le_ngt. apply le_0_l.
Qed.
-Theorem nle_succ_0 : forall n : N, ~ (S n <= 0).
+Theorem nle_succ_0 : forall n, ~ (S n <= 0).
Proof.
intros n H; apply -> le_succ_l in H; false_hyp H nlt_0_r.
Qed.
-Theorem le_0_r : forall n : N, n <= 0 <-> n == 0.
+Theorem le_0_r : forall n, n <= 0 <-> n == 0.
Proof.
intros n; split; intro H.
le_elim H; [false_hyp H nlt_0_r | assumption].
now apply eq_le_incl.
Qed.
-Theorem lt_0_succ : forall n : N, 0 < S n.
+Theorem lt_0_succ : forall n, 0 < S n.
Proof.
induct n; [apply lt_succ_diag_r | intros n H; now apply lt_lt_succ_r].
Qed.
-Theorem neq_0_lt_0 : forall n : N, n ~= 0 <-> 0 < n.
+Theorem neq_0_lt_0 : forall n, n ~= 0 <-> 0 < n.
Proof.
cases n.
split; intro H; [now elim H | intro; now apply lt_irrefl with 0].
intro n; split; intro H; [apply lt_0_succ | apply neq_succ_0].
Qed.
-Theorem eq_0_gt_0_cases : forall n : N, n == 0 \/ 0 < n.
+Theorem eq_0_gt_0_cases : forall n, n == 0 \/ 0 < n.
Proof.
cases n.
now left.
intro; right; apply lt_0_succ.
Qed.
-Theorem zero_one : forall n : N, n == 0 \/ n == 1 \/ 1 < n.
+Theorem zero_one : forall n, n == 0 \/ n == 1 \/ 1 < n.
Proof.
induct n. now left.
cases n. intros; right; now left.
@@ -366,7 +73,7 @@ right; right. rewrite H. apply lt_succ_diag_r.
right; right. now apply lt_lt_succ_r.
Qed.
-Theorem lt_1_r : forall n : N, n < 1 <-> n == 0.
+Theorem lt_1_r : forall n, n < 1 <-> n == 0.
Proof.
cases n.
split; intro; [reflexivity | apply lt_succ_diag_r].
@@ -374,7 +81,7 @@ intros n. rewrite <- succ_lt_mono.
split; intro H; [false_hyp H nlt_0_r | false_hyp H neq_succ_0].
Qed.
-Theorem le_1_r : forall n : N, n <= 1 <-> n == 0 \/ n == 1.
+Theorem le_1_r : forall n, n <= 1 <-> n == 0 \/ n == 1.
Proof.
cases n.
split; intro; [now left | apply le_succ_diag_r].
@@ -382,36 +89,30 @@ intro n. rewrite <- succ_le_mono, le_0_r, succ_inj_wd.
split; [intro; now right | intros [H | H]; [false_hyp H neq_succ_0 | assumption]].
Qed.
-Theorem lt_lt_0 : forall n m : N, n < m -> 0 < m.
+Theorem lt_lt_0 : forall n m, n < m -> 0 < m.
Proof.
intros n m; induct n.
trivial.
intros n IH H. apply IH; now apply lt_succ_l.
Qed.
-Theorem lt_1_l : forall n m p : N, n < m -> m < p -> 1 < p.
+Theorem lt_1_l' : forall n m p, n < m -> m < p -> 1 < p.
Proof.
-intros n m p H1 H2.
-apply le_lt_trans with m. apply <- le_succ_l. apply le_lt_trans with n.
-apply le_0_l. assumption. assumption.
+intros. apply lt_1_l with m; auto.
+apply le_lt_trans with n; auto. now apply le_0_l.
Qed.
(** Elimination principlies for < and <= for relations *)
Section RelElim.
-(* FIXME: Variable R : relation N. -- does not work *)
-
-Variable R : N -> N -> Prop.
-Hypothesis R_wd : relation_wd Neq Neq R.
-
-Add Morphism R with signature Neq ==> Neq ==> iff as R_morph2.
-Proof. apply R_wd. Qed.
+Variable R : relation N.t.
+Hypothesis R_wd : Proper (N.eq==>N.eq==>iff) R.
Theorem le_ind_rel :
- (forall m : N, R 0 m) ->
- (forall n m : N, n <= m -> R n m -> R (S n) (S m)) ->
- forall n m : N, n <= m -> R n m.
+ (forall m, R 0 m) ->
+ (forall n m, n <= m -> R n m -> R (S n) (S m)) ->
+ forall n m, n <= m -> R n m.
Proof.
intros Base Step; induct n.
intros; apply Base.
@@ -422,9 +123,9 @@ intros k H1 H2. apply -> le_succ_l in H1. apply lt_le_incl in H1. auto.
Qed.
Theorem lt_ind_rel :
- (forall m : N, R 0 (S m)) ->
- (forall n m : N, n < m -> R n m -> R (S n) (S m)) ->
- forall n m : N, n < m -> R n m.
+ (forall m, R 0 (S m)) ->
+ (forall n m, n < m -> R n m -> R (S n) (S m)) ->
+ forall n m, n < m -> R n m.
Proof.
intros Base Step; induct n.
intros m H. apply lt_exists_pred in H; destruct H as [m' [H _]].
@@ -439,61 +140,64 @@ End RelElim.
(** Predecessor and order *)
-Theorem succ_pred_pos : forall n : N, 0 < n -> S (P n) == n.
+Theorem succ_pred_pos : forall n, 0 < n -> S (P n) == n.
Proof.
intros n H; apply succ_pred; intro H1; rewrite H1 in H.
false_hyp H lt_irrefl.
Qed.
-Theorem le_pred_l : forall n : N, P n <= n.
+Theorem le_pred_l : forall n, P n <= n.
Proof.
cases n.
rewrite pred_0; now apply eq_le_incl.
intros; rewrite pred_succ; apply le_succ_diag_r.
Qed.
-Theorem lt_pred_l : forall n : N, n ~= 0 -> P n < n.
+Theorem lt_pred_l : forall n, n ~= 0 -> P n < n.
Proof.
cases n.
-intro H; elimtype False; now apply H.
+intro H; exfalso; now apply H.
intros; rewrite pred_succ; apply lt_succ_diag_r.
Qed.
-Theorem le_le_pred : forall n m : N, n <= m -> P n <= m.
+Theorem le_le_pred : forall n m, n <= m -> P n <= m.
Proof.
intros n m H; apply le_trans with n. apply le_pred_l. assumption.
Qed.
-Theorem lt_lt_pred : forall n m : N, n < m -> P n < m.
+Theorem lt_lt_pred : forall n m, n < m -> P n < m.
Proof.
intros n m H; apply le_lt_trans with n. apply le_pred_l. assumption.
Qed.
-Theorem lt_le_pred : forall n m : N, n < m -> n <= P m. (* Converse is false for n == m == 0 *)
+Theorem lt_le_pred : forall n m, n < m -> n <= P m.
+ (* Converse is false for n == m == 0 *)
Proof.
intro n; cases m.
intro H; false_hyp H nlt_0_r.
intros m IH. rewrite pred_succ; now apply -> lt_succ_r.
Qed.
-Theorem lt_pred_le : forall n m : N, P n < m -> n <= m. (* Converse is false for n == m == 0 *)
+Theorem lt_pred_le : forall n m, P n < m -> n <= m.
+ (* Converse is false for n == m == 0 *)
Proof.
intros n m; cases n.
rewrite pred_0; intro H; now apply lt_le_incl.
intros n IH. rewrite pred_succ in IH. now apply <- le_succ_l.
Qed.
-Theorem lt_pred_lt : forall n m : N, n < P m -> n < m.
+Theorem lt_pred_lt : forall n m, n < P m -> n < m.
Proof.
intros n m H; apply lt_le_trans with (P m); [assumption | apply le_pred_l].
Qed.
-Theorem le_pred_le : forall n m : N, n <= P m -> n <= m.
+Theorem le_pred_le : forall n m, n <= P m -> n <= m.
Proof.
intros n m H; apply le_trans with (P m); [assumption | apply le_pred_l].
Qed.
-Theorem pred_le_mono : forall n m : N, n <= m -> P n <= P m. (* Converse is false for n == 1, m == 0 *)
+Theorem pred_le_mono : forall n m, n <= m -> P n <= P m.
+ (* Converse is false for n == 1, m == 0 *)
Proof.
intros n m H; elim H using le_ind_rel.
solve_relation_wd.
@@ -501,7 +205,7 @@ intro; rewrite pred_0; apply le_0_l.
intros p q H1 _; now do 2 rewrite pred_succ.
Qed.
-Theorem pred_lt_mono : forall n m : N, n ~= 0 -> (n < m <-> P n < P m).
+Theorem pred_lt_mono : forall n m, n ~= 0 -> (n < m <-> P n < P m).
Proof.
intros n m H1; split; intro H2.
assert (m ~= 0). apply <- neq_0_lt_0. now apply lt_lt_0 with n.
@@ -512,22 +216,24 @@ apply lt_le_trans with (P m). assumption. apply le_pred_l.
apply -> succ_lt_mono in H2. now do 2 rewrite succ_pred in H2.
Qed.
-Theorem lt_succ_lt_pred : forall n m : N, S n < m <-> n < P m.
+Theorem lt_succ_lt_pred : forall n m, S n < m <-> n < P m.
Proof.
intros n m. rewrite pred_lt_mono by apply neq_succ_0. now rewrite pred_succ.
Qed.
-Theorem le_succ_le_pred : forall n m : N, S n <= m -> n <= P m. (* Converse is false for n == m == 0 *)
+Theorem le_succ_le_pred : forall n m, S n <= m -> n <= P m.
+ (* Converse is false for n == m == 0 *)
Proof.
intros n m H. apply lt_le_pred. now apply -> le_succ_l.
Qed.
-Theorem lt_pred_lt_succ : forall n m : N, P n < m -> n < S m. (* Converse is false for n == m == 0 *)
+Theorem lt_pred_lt_succ : forall n m, P n < m -> n < S m.
+ (* Converse is false for n == m == 0 *)
Proof.
intros n m H. apply <- lt_succ_r. now apply lt_pred_le.
Qed.
-Theorem le_pred_le_succ : forall n m : N, P n <= m <-> n <= S m.
+Theorem le_pred_le_succ : forall n m, P n <= m <-> n <= S m.
Proof.
intros n m; cases n.
rewrite pred_0. split; intro H; apply le_0_l.
diff --git a/theories/Numbers/Natural/Abstract/NProperties.v b/theories/Numbers/Natural/Abstract/NProperties.v
new file mode 100644
index 00000000..30262bd9
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NProperties.v
@@ -0,0 +1,22 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id$ i*)
+
+Require Export NAxioms NSub.
+
+(** This functor summarizes all known facts about N.
+ For the moment it is only an alias to [NSubPropFunct], which
+ subsumes all others.
+*)
+
+Module Type NPropSig := NSubPropFunct.
+
+Module NPropFunct (N:NAxiomsSig) <: NPropSig N.
+ Include NPropSig N.
+End NPropFunct.
diff --git a/theories/Numbers/Natural/Abstract/NStrongRec.v b/theories/Numbers/Natural/Abstract/NStrongRec.v
index c6a6da48..cbbcdbff 100644
--- a/theories/Numbers/Natural/Abstract/NStrongRec.v
+++ b/theories/Numbers/Natural/Abstract/NStrongRec.v
@@ -8,123 +8,200 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NStrongRec.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
(** This file defined the strong (course-of-value, well-founded) recursion
and proves its properties *)
Require Export NSub.
-Module NStrongRecPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NSubPropMod := NSubPropFunct NAxiomsMod.
-Open Local Scope NatScope.
+Module NStrongRecPropFunct (Import N : NAxiomsSig').
+Include NSubPropFunct N.
Section StrongRecursion.
-Variable A : Set.
+Variable A : Type.
Variable Aeq : relation A.
+Variable Aeq_equiv : Equivalence Aeq.
+
+(** [strong_rec] allows to define a recursive function [phi] given by
+ an equation [phi(n) = F(phi)(n)] where recursive calls to [phi]
+ in [F] are made on strictly lower numbers than [n].
+
+ For [strong_rec a F n]:
+ - Parameter [a:A] is a default value used internally, it has no
+ effect on the final result.
+ - Parameter [F:(N->A)->N->A] is the step function:
+ [F f n] should return [phi(n)] when [f] is a function
+ that coincide with [phi] for numbers strictly less than [n].
+*)
-Notation Local "x ==A y" := (Aeq x y) (at level 70, no associativity).
+Definition strong_rec (a : A) (f : (N.t -> A) -> N.t -> A) (n : N.t) : A :=
+ recursion (fun _ => a) (fun _ => f) (S n) n.
-Hypothesis Aeq_equiv : equiv A Aeq.
+(** For convenience, we use in proofs an intermediate definition
+ between [recursion] and [strong_rec]. *)
-Add Relation A Aeq
- reflexivity proved by (proj1 Aeq_equiv)
- symmetry proved by (proj2 (proj2 Aeq_equiv))
- transitivity proved by (proj1 (proj2 Aeq_equiv))
-as Aeq_rel.
+Definition strong_rec0 (a : A) (f : (N.t -> A) -> N.t -> A) : N.t -> N.t -> A :=
+ recursion (fun _ => a) (fun _ => f).
-Definition strong_rec (a : A) (f : N -> (N -> A) -> A) (n : N) : A :=
-recursion
- (fun _ : N => a)
- (fun (m : N) (p : N -> A) (k : N) => f k p)
- (S n)
- n.
+Lemma strong_rec_alt : forall a f n,
+ strong_rec a f n = strong_rec0 a f (S n) n.
+Proof.
+reflexivity.
+Qed.
-Theorem strong_rec_wd :
-forall a a' : A, a ==A a' ->
- forall f f', fun2_eq Neq (fun_eq Neq Aeq) Aeq f f' ->
- forall n n', n == n' ->
- strong_rec a f n ==A strong_rec a' f' n'.
+(** We need a result similar to [f_equal], but for setoid equalities. *)
+Lemma f_equiv : forall f g x y,
+ (N.eq==>Aeq)%signature f g -> N.eq x y -> Aeq (f x) (g y).
+Proof.
+auto.
+Qed.
+
+Instance strong_rec0_wd :
+ Proper (Aeq ==> ((N.eq ==> Aeq) ==> N.eq ==> Aeq) ==> N.eq ==> N.eq ==> Aeq)
+ strong_rec0.
+Proof.
+unfold strong_rec0.
+repeat red; intros.
+apply f_equiv; auto.
+apply recursion_wd; try red; auto.
+Qed.
+
+Instance strong_rec_wd :
+ Proper (Aeq ==> ((N.eq ==> Aeq) ==> N.eq ==> Aeq) ==> N.eq ==> Aeq) strong_rec.
Proof.
intros a a' Eaa' f f' Eff' n n' Enn'.
-(* First we prove that recursion (which is on type N -> A) returns
-extensionally equal functions, and then we use the fact that n == n' *)
-assert (H : fun_eq Neq Aeq
- (recursion
- (fun _ : N => a)
- (fun (m : N) (p : N -> A) (k : N) => f k p)
- (S n))
- (recursion
- (fun _ : N => a')
- (fun (m : N) (p : N -> A) (k : N) => f' k p)
- (S n'))).
-apply recursion_wd with (Aeq := fun_eq Neq Aeq).
-unfold fun_eq; now intros.
-unfold fun2_eq. intros y y' Eyy' p p' Epp'. unfold fun_eq. auto.
+rewrite !strong_rec_alt.
+apply strong_rec0_wd; auto.
now rewrite Enn'.
-unfold strong_rec.
-now apply H.
Qed.
-(*Section FixPoint.
-
-Variable a : A.
-Variable f : N -> (N -> A) -> A.
+Section FixPoint.
-Hypothesis f_wd : fun2_wd Neq (fun_eq Neq Aeq) Aeq f.
+Variable f : (N.t -> A) -> N.t -> A.
+Variable f_wd : Proper ((N.eq==>Aeq)==>N.eq==>Aeq) f.
-Let g (n : N) : A := strong_rec a f n.
+Lemma strong_rec0_0 : forall a m,
+ (strong_rec0 a f 0 m) = a.
+Proof.
+intros. unfold strong_rec0. rewrite recursion_0; auto.
+Qed.
-Add Morphism g with signature Neq ==> Aeq as g_wd.
+Lemma strong_rec0_succ : forall a n m,
+ Aeq (strong_rec0 a f (S n) m) (f (strong_rec0 a f n) m).
Proof.
-intros n1 n2 H. unfold g. now apply strong_rec_wd.
+intros. unfold strong_rec0.
+apply f_equiv; auto with *.
+rewrite recursion_succ; try (repeat red; auto with *; fail).
+apply f_wd.
+apply recursion_wd; try red; auto with *.
Qed.
-Theorem NtoA_eq_sym : symmetric (N -> A) (fun_eq Neq Aeq).
+Lemma strong_rec_0 : forall a,
+ Aeq (strong_rec a f 0) (f (fun _ => a) 0).
Proof.
-apply fun_eq_sym.
-exact (proj2 (proj2 NZeq_equiv)).
-exact (proj2 (proj2 Aeq_equiv)).
+intros. rewrite strong_rec_alt, strong_rec0_succ.
+apply f_wd; auto with *.
+red; intros; rewrite strong_rec0_0; auto with *.
Qed.
-Theorem NtoA_eq_trans : transitive (N -> A) (fun_eq Neq Aeq).
+(* We need an assumption saying that for every n, the step function (f h n)
+calls h only on the segment [0 ... n - 1]. This means that if h1 and h2
+coincide on values < n, then (f h1 n) coincides with (f h2 n) *)
+
+Hypothesis step_good :
+ forall (n : N.t) (h1 h2 : N.t -> A),
+ (forall m : N.t, m < n -> Aeq (h1 m) (h2 m)) -> Aeq (f h1 n) (f h2 n).
+
+Lemma strong_rec0_more_steps : forall a k n m, m < n ->
+ Aeq (strong_rec0 a f n m) (strong_rec0 a f (n+k) m).
Proof.
-apply fun_eq_trans.
-exact (proj1 NZeq_equiv).
-exact (proj1 (proj2 NZeq_equiv)).
-exact (proj1 (proj2 Aeq_equiv)).
+ intros a k n. pattern n.
+ apply induction; clear n.
+
+ intros n n' Hn; setoid_rewrite Hn; auto with *.
+
+ intros m Hm. destruct (nlt_0_r _ Hm).
+
+ intros n IH m Hm.
+ rewrite lt_succ_r in Hm.
+ rewrite add_succ_l.
+ rewrite 2 strong_rec0_succ.
+ apply step_good.
+ intros m' Hm'.
+ apply IH.
+ apply lt_le_trans with m; auto.
Qed.
-Add Relation (N -> A) (fun_eq Neq Aeq)
- symmetry proved by NtoA_eq_sym
- transitivity proved by NtoA_eq_trans
-as NtoA_eq_rel.
+Lemma strong_rec0_fixpoint : forall (a : A) (n : N.t),
+ Aeq (strong_rec0 a f (S n) n) (f (fun n => strong_rec0 a f (S n) n) n).
+Proof.
+intros.
+rewrite strong_rec0_succ.
+apply step_good.
+intros m Hm.
+symmetry.
+setoid_replace n with (S m + (n - S m)).
+apply strong_rec0_more_steps.
+apply lt_succ_diag_r.
+rewrite add_comm.
+symmetry.
+apply sub_add.
+rewrite le_succ_l; auto.
+Qed.
-Add Morphism f with signature Neq ==> (fun_eq Neq Aeq) ==> Aeq as f_morph.
+Theorem strong_rec_fixpoint : forall (a : A) (n : N.t),
+ Aeq (strong_rec a f n) (f (strong_rec a f) n).
Proof.
-apply f_wd.
+intros.
+transitivity (f (fun n => strong_rec0 a f (S n) n) n).
+rewrite strong_rec_alt.
+apply strong_rec0_fixpoint.
+apply f_wd; auto with *.
+intros x x' Hx; rewrite strong_rec_alt, Hx; auto with *.
Qed.
-(* We need an assumption saying that for every n, the step function (f n h)
-calls h only on the segment [0 ... n - 1]. This means that if h1 and h2
-coincide on values < n, then (f n h1) coincides with (f n h2) *)
+(** NB: without the [step_good] hypothesis, we have proved that
+ [strong_rec a f 0] is [f (fun _ => a) 0]. Now we can prove
+ that the first argument of [f] is arbitrary in this case...
+*)
-Hypothesis step_good :
- forall (n : N) (h1 h2 : N -> A),
- (forall m : N, m < n -> Aeq (h1 m) (h2 m)) -> Aeq (f n h1) (f n h2).
+Theorem strong_rec_0_any : forall (a : A)(any : N.t->A),
+ Aeq (strong_rec a f 0) (f any 0).
+Proof.
+intros.
+rewrite strong_rec_fixpoint.
+apply step_good.
+intros m Hm. destruct (nlt_0_r _ Hm).
+Qed.
-(* Todo:
-Theorem strong_rec_fixpoint : forall n : N, Aeq (g n) (f n g).
+(** ... and that first argument of [strong_rec] is always arbitrary. *)
+
+Lemma strong_rec_any_fst_arg : forall a a' n,
+ Aeq (strong_rec a f n) (strong_rec a' f n).
Proof.
-apply induction.
-unfold predicate_wd, fun_wd.
-intros x y H. rewrite H. unfold fun_eq; apply g_wd.
-reflexivity.
-unfold g, strong_rec.
-*)
+intros a a' n.
+generalize (le_refl n).
+set (k:=n) at -2. clearbody k. revert k. pattern n.
+apply induction; clear n.
+(* compat *)
+intros n n' Hn. setoid_rewrite Hn; auto with *.
+(* 0 *)
+intros k Hk. rewrite le_0_r in Hk.
+rewrite Hk, strong_rec_0. symmetry. apply strong_rec_0_any.
+(* S *)
+intros n IH k Hk.
+rewrite 2 strong_rec_fixpoint.
+apply step_good.
+intros m Hm.
+apply IH.
+rewrite succ_le_mono.
+apply le_trans with k; auto.
+rewrite le_succ_l; auto.
+Qed.
-End FixPoint.*)
+End FixPoint.
End StrongRecursion.
Implicit Arguments strong_rec [A].
diff --git a/theories/Numbers/Natural/Abstract/NSub.v b/theories/Numbers/Natural/Abstract/NSub.v
index f67689dd..35d3b8aa 100644
--- a/theories/Numbers/Natural/Abstract/NSub.v
+++ b/theories/Numbers/Natural/Abstract/NSub.v
@@ -8,49 +8,33 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NSub.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
Require Export NMulOrder.
-Module NSubPropFunct (Import NAxiomsMod : NAxiomsSig).
-Module Export NMulOrderPropMod := NMulOrderPropFunct NAxiomsMod.
-Open Local Scope NatScope.
+Module Type NSubPropFunct (Import N : NAxiomsSig').
+Include NMulOrderPropFunct N.
-Theorem sub_wd :
- forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> n1 - m1 == n2 - m2.
-Proof NZsub_wd.
-
-Theorem sub_0_r : forall n : N, n - 0 == n.
-Proof NZsub_0_r.
-
-Theorem sub_succ_r : forall n m : N, n - (S m) == P (n - m).
-Proof NZsub_succ_r.
-
-Theorem sub_1_r : forall n : N, n - 1 == P n.
-Proof.
-intro n; rewrite sub_succ_r; now rewrite sub_0_r.
-Qed.
-
-Theorem sub_0_l : forall n : N, 0 - n == 0.
+Theorem sub_0_l : forall n, 0 - n == 0.
Proof.
induct n.
apply sub_0_r.
intros n IH; rewrite sub_succ_r; rewrite IH. now apply pred_0.
Qed.
-Theorem sub_succ : forall n m : N, S n - S m == n - m.
+Theorem sub_succ : forall n m, S n - S m == n - m.
Proof.
intro n; induct m.
rewrite sub_succ_r. do 2 rewrite sub_0_r. now rewrite pred_succ.
intros m IH. rewrite sub_succ_r. rewrite IH. now rewrite sub_succ_r.
Qed.
-Theorem sub_diag : forall n : N, n - n == 0.
+Theorem sub_diag : forall n, n - n == 0.
Proof.
induct n. apply sub_0_r. intros n IH; rewrite sub_succ; now rewrite IH.
Qed.
-Theorem sub_gt : forall n m : N, n > m -> n - m ~= 0.
+Theorem sub_gt : forall n m, n > m -> n - m ~= 0.
Proof.
intros n m H; elim H using lt_ind_rel; clear n m H.
solve_relation_wd.
@@ -58,7 +42,7 @@ intro; rewrite sub_0_r; apply neq_succ_0.
intros; now rewrite sub_succ.
Qed.
-Theorem add_sub_assoc : forall n m p : N, p <= m -> n + (m - p) == (n + m) - p.
+Theorem add_sub_assoc : forall n m p, p <= m -> n + (m - p) == (n + m) - p.
Proof.
intros n m p; induct p.
intro; now do 2 rewrite sub_0_r.
@@ -68,32 +52,32 @@ rewrite add_pred_r by (apply sub_gt; now apply -> le_succ_l).
reflexivity.
Qed.
-Theorem sub_succ_l : forall n m : N, n <= m -> S m - n == S (m - n).
+Theorem sub_succ_l : forall n m, n <= m -> S m - n == S (m - n).
Proof.
intros n m H. rewrite <- (add_1_l m). rewrite <- (add_1_l (m - n)).
symmetry; now apply add_sub_assoc.
Qed.
-Theorem add_sub : forall n m : N, (n + m) - m == n.
+Theorem add_sub : forall n m, (n + m) - m == n.
Proof.
intros n m. rewrite <- add_sub_assoc by (apply le_refl).
rewrite sub_diag; now rewrite add_0_r.
Qed.
-Theorem sub_add : forall n m : N, n <= m -> (m - n) + n == m.
+Theorem sub_add : forall n m, n <= m -> (m - n) + n == m.
Proof.
intros n m H. rewrite add_comm. rewrite add_sub_assoc by assumption.
rewrite add_comm. apply add_sub.
Qed.
-Theorem add_sub_eq_l : forall n m p : N, m + p == n -> n - m == p.
+Theorem add_sub_eq_l : forall n m p, m + p == n -> n - m == p.
Proof.
intros n m p H. symmetry.
assert (H1 : m + p - m == n - m) by now rewrite H.
rewrite add_comm in H1. now rewrite add_sub in H1.
Qed.
-Theorem add_sub_eq_r : forall n m p : N, m + p == n -> n - p == m.
+Theorem add_sub_eq_r : forall n m p, m + p == n -> n - p == m.
Proof.
intros n m p H; rewrite add_comm in H; now apply add_sub_eq_l.
Qed.
@@ -101,7 +85,7 @@ Qed.
(* This could be proved by adding m to both sides. Then the proof would
use add_sub_assoc and sub_0_le, which is proven below. *)
-Theorem add_sub_eq_nz : forall n m p : N, p ~= 0 -> n - m == p -> m + p == n.
+Theorem add_sub_eq_nz : forall n m p, p ~= 0 -> n - m == p -> m + p == n.
Proof.
intros n m p H; double_induct n m.
intros m H1; rewrite sub_0_l in H1. symmetry in H1; false_hyp H1 H.
@@ -110,14 +94,14 @@ intros n m IH H1. rewrite sub_succ in H1. apply IH in H1.
rewrite add_succ_l; now rewrite H1.
Qed.
-Theorem sub_add_distr : forall n m p : N, n - (m + p) == (n - m) - p.
+Theorem sub_add_distr : forall n m p, n - (m + p) == (n - m) - p.
Proof.
intros n m; induct p.
rewrite add_0_r; now rewrite sub_0_r.
intros p IH. rewrite add_succ_r; do 2 rewrite sub_succ_r. now rewrite IH.
Qed.
-Theorem add_sub_swap : forall n m p : N, p <= n -> n + m - p == n - p + m.
+Theorem add_sub_swap : forall n m p, p <= n -> n + m - p == n - p + m.
Proof.
intros n m p H.
rewrite (add_comm n m).
@@ -127,7 +111,7 @@ Qed.
(** Sub and order *)
-Theorem le_sub_l : forall n m : N, n - m <= n.
+Theorem le_sub_l : forall n m, n - m <= n.
Proof.
intro n; induct m.
rewrite sub_0_r; now apply eq_le_incl.
@@ -135,7 +119,7 @@ intros m IH. rewrite sub_succ_r.
apply le_trans with (n - m); [apply le_pred_l | assumption].
Qed.
-Theorem sub_0_le : forall n m : N, n - m == 0 <-> n <= m.
+Theorem sub_0_le : forall n m, n - m == 0 <-> n <= m.
Proof.
double_induct n m.
intro m; split; intro; [apply le_0_l | apply sub_0_l].
@@ -144,9 +128,86 @@ intro m; rewrite sub_0_r; split; intro H;
intros n m H. rewrite <- succ_le_mono. now rewrite sub_succ.
Qed.
+Theorem sub_add_le : forall n m, n <= n - m + m.
+Proof.
+intros.
+destruct (le_ge_cases n m) as [LE|GE].
+rewrite <- sub_0_le in LE. rewrite LE; nzsimpl.
+now rewrite <- sub_0_le.
+rewrite sub_add by assumption. apply le_refl.
+Qed.
+
+Theorem le_sub_le_add_r : forall n m p,
+ n - p <= m <-> n <= m + p.
+Proof.
+intros n m p.
+split; intros LE.
+rewrite (add_le_mono_r _ _ p) in LE.
+apply le_trans with (n-p+p); auto using sub_add_le.
+destruct (le_ge_cases n p) as [LE'|GE].
+rewrite <- sub_0_le in LE'. rewrite LE'. apply le_0_l.
+rewrite (add_le_mono_r _ _ p). now rewrite sub_add.
+Qed.
+
+Theorem le_sub_le_add_l : forall n m p, n - m <= p <-> n <= m + p.
+Proof.
+intros n m p. rewrite add_comm; apply le_sub_le_add_r.
+Qed.
+
+Theorem lt_sub_lt_add_r : forall n m p,
+ n - p < m -> n < m + p.
+Proof.
+intros n m p LT.
+rewrite (add_lt_mono_r _ _ p) in LT.
+apply le_lt_trans with (n-p+p); auto using sub_add_le.
+Qed.
+
+(** Unfortunately, we do not have [n < m + p -> n - p < m].
+ For instance [1<0+2] but not [1-2<0]. *)
+
+Theorem lt_sub_lt_add_l : forall n m p, n - m < p -> n < m + p.
+Proof.
+intros n m p. rewrite add_comm; apply lt_sub_lt_add_r.
+Qed.
+
+Theorem le_add_le_sub_r : forall n m p, n + p <= m -> n <= m - p.
+Proof.
+intros n m p LE.
+apply (add_le_mono_r _ _ p).
+rewrite sub_add. assumption.
+apply le_trans with (n+p); trivial.
+rewrite <- (add_0_l p) at 1. rewrite <- add_le_mono_r. apply le_0_l.
+Qed.
+
+(** Unfortunately, we do not have [n <= m - p -> n + p <= m].
+ For instance [0<=1-2] but not [2+0<=1]. *)
+
+Theorem le_add_le_sub_l : forall n m p, n + p <= m -> p <= m - n.
+Proof.
+intros n m p. rewrite add_comm; apply le_add_le_sub_r.
+Qed.
+
+Theorem lt_add_lt_sub_r : forall n m p, n + p < m <-> n < m - p.
+Proof.
+intros n m p.
+destruct (le_ge_cases p m) as [LE|GE].
+rewrite <- (sub_add p m) at 1 by assumption.
+now rewrite <- add_lt_mono_r.
+assert (GE' := GE). rewrite <- sub_0_le in GE'; rewrite GE'.
+split; intros LT.
+elim (lt_irrefl m). apply le_lt_trans with (n+p); trivial.
+ rewrite <- (add_0_l m). apply add_le_mono. apply le_0_l. assumption.
+now elim (nlt_0_r n).
+Qed.
+
+Theorem lt_add_lt_sub_l : forall n m p, n + p < m <-> p < m - n.
+Proof.
+intros n m p. rewrite add_comm; apply lt_add_lt_sub_r.
+Qed.
+
(** Sub and mul *)
-Theorem mul_pred_r : forall n m : N, n * (P m) == n * m - n.
+Theorem mul_pred_r : forall n m, n * (P m) == n * m - n.
Proof.
intros n m; cases m.
now rewrite pred_0, mul_0_r, sub_0_l.
@@ -155,7 +216,7 @@ now rewrite sub_diag, add_0_r.
now apply eq_le_incl.
Qed.
-Theorem mul_sub_distr_r : forall n m p : N, (n - m) * p == n * p - m * p.
+Theorem mul_sub_distr_r : forall n m p, (n - m) * p == n * p - m * p.
Proof.
intros n m p; induct n.
now rewrite sub_0_l, mul_0_l, sub_0_l.
@@ -170,11 +231,72 @@ setoid_replace ((S n * p) - m * p) with 0 by (apply <- sub_0_le; now apply mul_l
apply mul_0_l.
Qed.
-Theorem mul_sub_distr_l : forall n m p : N, p * (n - m) == p * n - p * m.
+Theorem mul_sub_distr_l : forall n m p, p * (n - m) == p * n - p * m.
Proof.
intros n m p; rewrite (mul_comm p (n - m)), (mul_comm p n), (mul_comm p m).
apply mul_sub_distr_r.
Qed.
+(** Alternative definitions of [<=] and [<] based on [+] *)
+
+Definition le_alt n m := exists p, p + n == m.
+Definition lt_alt n m := exists p, S p + n == m.
+
+Lemma le_equiv : forall n m, le_alt n m <-> n <= m.
+Proof.
+split.
+intros (p,H). rewrite <- H, add_comm. apply le_add_r.
+intro H. exists (m-n). now apply sub_add.
+Qed.
+
+Lemma lt_equiv : forall n m, lt_alt n m <-> n < m.
+Proof.
+split.
+intros (p,H). rewrite <- H, add_succ_l, lt_succ_r, add_comm. apply le_add_r.
+intro H. exists (m-S n). rewrite add_succ_l, <- add_succ_r.
+apply sub_add. now rewrite le_succ_l.
+Qed.
+
+Instance le_alt_wd : Proper (eq==>eq==>iff) le_alt.
+Proof.
+ intros x x' Hx y y' Hy; unfold le_alt.
+ setoid_rewrite Hx. setoid_rewrite Hy. auto with *.
+Qed.
+
+Instance lt_alt_wd : Proper (eq==>eq==>iff) lt_alt.
+Proof.
+ intros x x' Hx y y' Hy; unfold lt_alt.
+ setoid_rewrite Hx. setoid_rewrite Hy. auto with *.
+Qed.
+
+(** With these alternative definition, the dichotomy:
+
+[forall n m, n <= m \/ m <= n]
+
+becomes:
+
+[forall n m, (exists p, p + n == m) \/ (exists p, p + m == n)]
+
+We will need this in the proof of induction principle for integers
+constructed as pairs of natural numbers. This formula can be proved
+from know properties of [<=]. However, it can also be done directly. *)
+
+Theorem le_alt_dichotomy : forall n m, le_alt n m \/ le_alt m n.
+Proof.
+intros n m; induct n.
+left; exists m; apply add_0_r.
+intros n IH.
+destruct IH as [[p H] | [p H]].
+destruct (zero_or_succ p) as [H1 | [p' H1]]; rewrite H1 in H.
+rewrite add_0_l in H. right; exists (S 0); rewrite H, add_succ_l;
+ now rewrite add_0_l.
+left; exists p'; rewrite add_succ_r; now rewrite add_succ_l in H.
+right; exists (S p). rewrite add_succ_l; now rewrite H.
+Qed.
+
+Theorem add_dichotomy :
+ forall n m, (exists p, p + n == m) \/ (exists p, p + m == n).
+Proof. exact le_alt_dichotomy. Qed.
+
End NSubPropFunct.
diff --git a/theories/Numbers/Natural/BigN/BigN.v b/theories/Numbers/Natural/BigN/BigN.v
index 16007656..cab4b154 100644
--- a/theories/Numbers/Natural/BigN/BigN.v
+++ b/theories/Numbers/Natural/BigN/BigN.v
@@ -6,28 +6,32 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: BigN.v 11576 2008-11-10 19:13:15Z msozeau $ i*)
+(** * Efficient arbitrary large natural numbers in base 2^31 *)
-(** * Natural numbers in base 2^31 *)
-
-(**
-Author: Arnaud Spiwack
-*)
+(** Initial Author: Arnaud Spiwack *)
Require Export Int31.
-Require Import CyclicAxioms.
-Require Import Cyclic31.
-Require Import NSig.
-Require Import NSigNAxioms.
-Require Import NMake.
-Require Import NSub.
+Require Import CyclicAxioms Cyclic31 Ring31 NSig NSigNAxioms NMake
+ NProperties NDiv GenericMinMax.
+
+(** The following [BigN] module regroups both the operations and
+ all the abstract properties:
-Module BigN <: NType := NMake.Make Int31Cyclic.
+ - [NMake.Make Int31Cyclic] provides the operations and basic specs
+ w.r.t. ZArith
+ - [NTypeIsNAxioms] shows (mainly) that these operations implement
+ the interface [NAxioms]
+ - [NPropSig] adds all generic properties derived from [NAxioms]
+ - [NDivPropFunct] provides generic properties of [div] and [mod].
+ - [MinMax*Properties] provides properties of [min] and [max].
+
+*)
-(** Module [BigN] implements [NAxiomsSig] *)
+Module BigN <: NType <: OrderedTypeFull <: TotalOrder :=
+ NMake.Make Int31Cyclic <+ NTypeIsNAxioms
+ <+ !NPropSig <+ !NDivPropFunct <+ HasEqBool2Dec
+ <+ !MinMaxLogicalProperties <+ !MinMaxDecProperties.
-Module Export BigNAxiomsMod := NSig_NAxioms BigN.
-Module Export BigNSubPropMod := NSubPropFunct BigNAxiomsMod.
(** Notations about [BigN] *)
@@ -37,49 +41,171 @@ Delimit Scope bigN_scope with bigN.
Bind Scope bigN_scope with bigN.
Bind Scope bigN_scope with BigN.t.
Bind Scope bigN_scope with BigN.t_.
-
-Notation Local "0" := BigN.zero : bigN_scope. (* temporary notation *)
+(* Bind Scope has no retroactive effect, let's declare scopes by hand. *)
+Arguments Scope BigN.to_Z [bigN_scope].
+Arguments Scope BigN.succ [bigN_scope].
+Arguments Scope BigN.pred [bigN_scope].
+Arguments Scope BigN.square [bigN_scope].
+Arguments Scope BigN.add [bigN_scope bigN_scope].
+Arguments Scope BigN.sub [bigN_scope bigN_scope].
+Arguments Scope BigN.mul [bigN_scope bigN_scope].
+Arguments Scope BigN.div [bigN_scope bigN_scope].
+Arguments Scope BigN.eq [bigN_scope bigN_scope].
+Arguments Scope BigN.lt [bigN_scope bigN_scope].
+Arguments Scope BigN.le [bigN_scope bigN_scope].
+Arguments Scope BigN.eq [bigN_scope bigN_scope].
+Arguments Scope BigN.compare [bigN_scope bigN_scope].
+Arguments Scope BigN.min [bigN_scope bigN_scope].
+Arguments Scope BigN.max [bigN_scope bigN_scope].
+Arguments Scope BigN.eq_bool [bigN_scope bigN_scope].
+Arguments Scope BigN.power_pos [bigN_scope positive_scope].
+Arguments Scope BigN.power [bigN_scope N_scope].
+Arguments Scope BigN.sqrt [bigN_scope].
+Arguments Scope BigN.div_eucl [bigN_scope bigN_scope].
+Arguments Scope BigN.modulo [bigN_scope bigN_scope].
+Arguments Scope BigN.gcd [bigN_scope bigN_scope].
+
+Local Notation "0" := BigN.zero : bigN_scope. (* temporary notation *)
+Local Notation "1" := BigN.one : bigN_scope. (* temporary notation *)
Infix "+" := BigN.add : bigN_scope.
Infix "-" := BigN.sub : bigN_scope.
Infix "*" := BigN.mul : bigN_scope.
Infix "/" := BigN.div : bigN_scope.
+Infix "^" := BigN.power : bigN_scope.
Infix "?=" := BigN.compare : bigN_scope.
Infix "==" := BigN.eq (at level 70, no associativity) : bigN_scope.
+Notation "x != y" := (~x==y)%bigN (at level 70, no associativity) : bigN_scope.
Infix "<" := BigN.lt : bigN_scope.
Infix "<=" := BigN.le : bigN_scope.
Notation "x > y" := (BigN.lt y x)(only parsing) : bigN_scope.
Notation "x >= y" := (BigN.le y x)(only parsing) : bigN_scope.
+Notation "x < y < z" := (x<y /\ y<z)%bigN : bigN_scope.
+Notation "x < y <= z" := (x<y /\ y<=z)%bigN : bigN_scope.
+Notation "x <= y < z" := (x<=y /\ y<z)%bigN : bigN_scope.
+Notation "x <= y <= z" := (x<=y /\ y<=z)%bigN : bigN_scope.
Notation "[ i ]" := (BigN.to_Z i) : bigN_scope.
+Infix "mod" := BigN.modulo (at level 40, no associativity) : bigN_scope.
-Open Scope bigN_scope.
+Local Open Scope bigN_scope.
(** Example of reasoning about [BigN] *)
-Theorem succ_pred: forall q:bigN,
+Theorem succ_pred: forall q : bigN,
0 < q -> BigN.succ (BigN.pred q) == q.
Proof.
-intros; apply succ_pred.
+intros; apply BigN.succ_pred.
intro H'; rewrite H' in H; discriminate.
Qed.
(** [BigN] is a semi-ring *)
-Lemma BigNring :
- semi_ring_theory BigN.zero BigN.one BigN.add BigN.mul BigN.eq.
+Lemma BigNring : semi_ring_theory 0 1 BigN.add BigN.mul BigN.eq.
+Proof.
+constructor.
+exact BigN.add_0_l. exact BigN.add_comm. exact BigN.add_assoc.
+exact BigN.mul_1_l. exact BigN.mul_0_l. exact BigN.mul_comm.
+exact BigN.mul_assoc. exact BigN.mul_add_distr_r.
+Qed.
+
+Lemma BigNeqb_correct : forall x y, BigN.eq_bool x y = true -> x==y.
+Proof. now apply BigN.eqb_eq. Qed.
+
+Lemma BigNpower : power_theory 1 BigN.mul BigN.eq (@id N) BigN.power.
Proof.
constructor.
-exact add_0_l.
-exact add_comm.
-exact add_assoc.
-exact mul_1_l.
-exact mul_0_l.
-exact mul_comm.
-exact mul_assoc.
-exact mul_add_distr_r.
+intros. red. rewrite BigN.spec_power. unfold id.
+destruct Zpower_theory as [EQ]. rewrite EQ.
+destruct n; simpl. reflexivity.
+induction p; simpl; intros; BigN.zify; rewrite ?IHp; auto.
+Qed.
+
+Lemma BigNdiv : div_theory BigN.eq BigN.add BigN.mul (@id _)
+ (fun a b => if BigN.eq_bool b 0 then (0,a) else BigN.div_eucl a b).
+Proof.
+constructor. unfold id. intros a b.
+BigN.zify.
+generalize (Zeq_bool_if [b] 0); destruct (Zeq_bool [b] 0).
+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').
+intros (EQ,_). injection 1. intros EQr EQq.
+BigN.zify. rewrite EQr, EQq; auto.
+Qed.
+
+
+(** Detection of constants *)
+
+Ltac isStaticWordCst t :=
+ match t with
+ | W0 => constr:true
+ | WW ?t1 ?t2 =>
+ match isStaticWordCst t1 with
+ | false => constr:false
+ | true => isStaticWordCst t2
+ end
+ | _ => isInt31cst t
+ end.
+
+Ltac isBigNcst t :=
+ match t with
+ | BigN.N0 ?t => isStaticWordCst t
+ | BigN.N1 ?t => isStaticWordCst t
+ | BigN.N2 ?t => isStaticWordCst t
+ | BigN.N3 ?t => isStaticWordCst t
+ | BigN.N4 ?t => isStaticWordCst t
+ | BigN.N5 ?t => isStaticWordCst t
+ | BigN.N6 ?t => isStaticWordCst t
+ | BigN.Nn ?n ?t => match isnatcst n with
+ | true => isStaticWordCst t
+ | false => constr:false
+ end
+ | BigN.zero => constr:true
+ | BigN.one => constr:true
+ | _ => constr:false
+ end.
+
+Ltac BigNcst t :=
+ match isBigNcst t with
+ | true => constr:t
+ | false => constr:NotConstant
+ end.
+
+Ltac Ncst t :=
+ match isNcst t with
+ | true => constr:t
+ | false => constr:NotConstant
+ end.
+
+(** Registration for the "ring" tactic *)
+
+Add Ring BigNr : BigNring
+ (decidable BigNeqb_correct,
+ constants [BigNcst],
+ power_tac BigNpower [Ncst],
+ div BigNdiv).
+
+Section TestRing.
+Let test : forall x y, 1 + x*y + x^2 + 1 == 1*1 + 1 + y*x + 1*x*x.
+intros. ring_simplify. reflexivity.
Qed.
+End TestRing.
+
+(** We benefit also from an "order" tactic *)
+
+Ltac bigN_order := BigN.order.
+
+Section TestOrder.
+Let test : forall x y : bigN, x<=y -> y<=x -> x==y.
+Proof. bigN_order. Qed.
+End TestOrder.
-Add Ring BigNr : BigNring.
+(** We can use at least a bit of (r)omega by translating to [Z]. *)
-(** Todo: tactic translating from [BigN] to [Z] + omega *)
+Section TestOmega.
+Let test : forall x y : bigN, x<=y -> y<=x -> x==y.
+Proof. intros x y. BigN.zify. omega. Qed.
+End TestOmega.
(** Todo: micromega *)
diff --git a/theories/Numbers/Natural/BigN/NMake.v b/theories/Numbers/Natural/BigN/NMake.v
new file mode 100644
index 00000000..925b0535
--- /dev/null
+++ b/theories/Numbers/Natural/BigN/NMake.v
@@ -0,0 +1,524 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
+(************************************************************************)
+
+(** * NMake *)
+
+(** From a cyclic Z/nZ representation to arbitrary precision natural numbers.*)
+
+(** NB: This file contain the part which is independent from the underlying
+ representation. The representation-dependent (and macro-generated) part
+ is now in [NMake_gen]. *)
+
+Require Import BigNumPrelude ZArith CyclicAxioms.
+Require Import Nbasic Wf_nat StreamMemo NSig NMake_gen.
+
+Module Make (Import W0:CyclicType) <: NType.
+
+ (** Macro-generated part *)
+
+ Include NMake_gen.Make W0.
+
+
+ (** * Predecessor *)
+
+ Lemma spec_pred : forall x, [pred x] = Zmax 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.
+ Qed.
+
+
+ (** * Subtraction *)
+
+ Lemma spec_sub : forall x y, [sub x y] = Zmax 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.
+ Qed.
+
+ (** * Comparison *)
+
+ Theorem spec_compare : forall x y, compare x y = Zcompare [x] [y].
+ Proof.
+ intros x y. generalize (spec_compare_aux x y); destruct compare;
+ intros; symmetry; try rewrite Zcompare_Eq_iff_eq; assumption.
+ Qed.
+
+ Definition eq_bool x y :=
+ match compare x y with
+ | Eq => true
+ | _ => false
+ end.
+
+ Theorem spec_eq_bool : forall x y, eq_bool x y = Zeq_bool [x] [y].
+ Proof.
+ intros. unfold eq_bool, Zeq_bool. rewrite spec_compare; reflexivity.
+ Qed.
+
+ Theorem spec_eq_bool_aux: forall x y,
+ if eq_bool x y then [x] = [y] else [x] <> [y].
+ Proof.
+ intros x y; unfold eq_bool.
+ generalize (spec_compare_aux x y); case compare; auto with zarith.
+ Qed.
+
+ Definition lt n m := [n] < [m].
+ Definition le n m := [n] <= [m].
+
+ 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_max : forall n m, [max n m] = Zmax [n] [m].
+ Proof.
+ intros. unfold max, Zmax. rewrite spec_compare; destruct Zcompare; reflexivity.
+ Qed.
+
+ Theorem spec_min : forall n m, [min n m] = Zmin [n] [m].
+ Proof.
+ intros. unfold min, Zmin. rewrite spec_compare; destruct Zcompare; reflexivity.
+ Qed.
+
+
+ (** * Power *)
+
+ Fixpoint power_pos (x:t) (p:positive) {struct p} : t :=
+ match p with
+ | xH => x
+ | xO p => square (power_pos x p)
+ | xI p => mul (square (power_pos x p)) x
+ end.
+
+ Theorem spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
+ Proof.
+ intros x n; generalize x; elim n; clear n x; simpl power_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.
+ 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.
+ Qed.
+
+ Definition power x (n:N) := match n with
+ | BinNat.N0 => one
+ | BinNat.Npos p => power_pos x p
+ end.
+
+ Theorem spec_power: forall x n, [power x n] = [x] ^ Z_of_N n.
+ Proof.
+ destruct n; simpl. apply (spec_1 w0_spec).
+ apply spec_power_pos.
+ Qed.
+
+
+ (** * Div *)
+
+ Definition div_eucl x y :=
+ if eq_bool y zero then (zero,zero) else
+ match compare x y with
+ | Eq => (one, zero)
+ | Lt => (zero, x)
+ | Gt => div_gt x y
+ end.
+
+ Theorem spec_div_eucl: forall x y,
+ let (q,r) := div_eucl x y in
+ ([q], [r]) = Zdiv_eucl [x] [y].
+ Proof.
+ assert (F0: [zero] = 0).
+ exact (spec_0 w0_spec).
+ assert (F1: [one] = 1).
+ exact (spec_1 w0_spec).
+ intros x y. unfold div_eucl.
+ generalize (spec_eq_bool_aux y zero). destruct eq_bool; rewrite F0.
+ intro H. rewrite H. destruct [x]; auto.
+ intro H'.
+ assert (0 < [y]) by (generalize (spec_pos y); auto with zarith).
+ clear H'.
+ generalize (spec_compare_aux x y); case compare; try rewrite F0;
+ try rewrite F1; intros; auto with zarith.
+ rewrite H0; generalize (Z_div_same [y] (Zlt_gt _ _ H))
+ (Z_mod_same [y] (Zlt_gt _ _ H));
+ unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.
+ assert (F2: 0 <= [x] < [y]).
+ generalize (spec_pos x); auto.
+ generalize (Zdiv_small _ _ F2)
+ (Zmod_small _ _ F2);
+ unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.
+ generalize (spec_div_gt _ _ H0 H); auto.
+ unfold Zdiv, Zmod; case Zdiv_eucl; case div_gt.
+ intros a b c d (H1, H2); subst; auto.
+ Qed.
+
+ Definition div x y := fst (div_eucl x y).
+
+ Theorem spec_div:
+ forall x y, [div x y] = [x] / [y].
+ 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;
+ injection H; auto.
+ Qed.
+
+
+ (** * Modulo *)
+
+ Definition modulo x y :=
+ if eq_bool y zero then zero else
+ match compare x y with
+ | Eq => zero
+ | Lt => x
+ | Gt => mod_gt x y
+ end.
+
+ Theorem spec_modulo:
+ forall x y, [modulo x y] = [x] mod [y].
+ Proof.
+ assert (F0: [zero] = 0).
+ exact (spec_0 w0_spec).
+ assert (F1: [one] = 1).
+ exact (spec_1 w0_spec).
+ intros x y. unfold modulo.
+ generalize (spec_eq_bool_aux y zero). destruct eq_bool; rewrite F0.
+ intro H; rewrite H. destruct [x]; auto.
+ intro H'.
+ assert (H : 0 < [y]) by (generalize (spec_pos y); auto with zarith).
+ clear H'.
+ generalize (spec_compare_aux x y); case compare; try rewrite F0;
+ try rewrite F1; 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.
+ generalize (spec_pos x); auto with zarith.
+ apply spec_mod_gt; auto.
+ Qed.
+
+
+ (** * Gcd *)
+
+ Definition gcd_gt_body a b cont :=
+ match compare b zero with
+ | Gt =>
+ let r := mod_gt a b in
+ match compare r zero with
+ | Gt => cont r (mod_gt b r)
+ | _ => b
+ end
+ | _ => a
+ end.
+
+ Theorem Zspec_gcd_gt_body: forall a b cont p,
+ [a] > [b] -> [a] < 2 ^ p ->
+ (forall a1 b1, [a1] < 2 ^ (p - 1) -> [a1] > [b1] ->
+ Zis_gcd [a1] [b1] [cont a1 b1]) ->
+ Zis_gcd [a] [b] [gcd_gt_body a b cont].
+ Proof.
+ assert (F1: [zero] = 0).
+ unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.
+ intros a b cont p H2 H3 H4; unfold gcd_gt_body.
+ generalize (spec_compare_aux b zero); case compare; try rewrite F1.
+ intros HH; rewrite HH; apply Zis_gcd_0.
+ intros HH; absurd (0 <= [b]); auto with zarith.
+ case (spec_digits b); auto with zarith.
+ intros H5; generalize (spec_compare_aux (mod_gt a b) zero);
+ case compare; try rewrite F1.
+ intros H6; rewrite <- (Zmult_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.
+ apply Zis_gcd_mult; apply Zis_gcd_1.
+ intros; apply False_ind.
+ case (spec_digits (mod_gt a b)); auto with zarith.
+ intros H6; apply DoubleDiv.Zis_gcd_mod; auto with zarith.
+ apply DoubleDiv.Zis_gcd_mod; auto with zarith.
+ rewrite <- spec_mod_gt; auto with zarith.
+ assert (F2: [b] > [mod_gt a b]).
+ case (Z_mod_lt [a] [b]); auto with zarith.
+ repeat rewrite <- spec_mod_gt; auto with zarith.
+ assert (F3: [mod_gt a b] > [mod_gt b (mod_gt a b)]).
+ case (Z_mod_lt [b] [mod_gt a b]); auto with zarith.
+ 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.
+ Qed.
+
+ Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) {struct p} : t :=
+ gcd_gt_body a b
+ (fun a b =>
+ match p with
+ | xH => cont a b
+ | xO p => gcd_gt_aux p (gcd_gt_aux p cont) a b
+ | xI p => gcd_gt_aux p (gcd_gt_aux p cont) a b
+ end).
+
+ Theorem Zspec_gcd_gt_aux: forall p n a b cont,
+ [a] > [b] -> [a] < 2 ^ (Zpos p + n) ->
+ (forall a1 b1, [a1] < 2 ^ n -> [a1] > [b1] ->
+ Zis_gcd [a1] [b1] [cont a1 b1]) ->
+ Zis_gcd [a] [b] [gcd_gt_aux p cont a b].
+ intros p; elim p; clear p.
+ intros p Hrec n a b cont H2 H3 H4.
+ unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xI p) + n); auto.
+ intros a1 b1 H6 H7.
+ apply Hrec with (Zpos p + n); auto.
+ replace (Zpos p + (Zpos p + n)) with
+ (Zpos (xI p) + n - 1); auto.
+ rewrite Zpos_xI; ring.
+ intros a2 b2 H9 H10.
+ apply Hrec with n; auto.
+ intros p Hrec n a b cont H2 H3 H4.
+ unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xO p) + n); auto.
+ intros a1 b1 H6 H7.
+ 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.
+ 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).
+ case (Zle_or_lt 1 n); intros HH.
+ apply Zpower_le_monotone; auto with zarith.
+ apply Zle_trans with 0; auto with zarith.
+ assert (HH1: n - 1 < 0); auto with zarith.
+ generalize HH1; case (n - 1); auto with zarith.
+ intros p1 HH2; discriminate.
+ 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.
+ intros a1 b1 H5 H6; apply H3; auto.
+ replace n with (n + 1 - 1); auto; try ring.
+ Qed.
+
+ Definition gcd_cont a b :=
+ match compare one b with
+ | Eq => one
+ | _ => a
+ end.
+
+ 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].
+ 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.
+ unfold gcd_gt; apply Zspec_gcd_gt_aux with 0; auto with zarith.
+ intros a1 a2; rewrite Zpower_0_r.
+ case (spec_digits a2); intros H7 H8;
+ intros; apply False_ind; auto with zarith.
+ Qed.
+
+ Definition gcd a b :=
+ match compare a b with
+ | Eq => a
+ | Lt => gcd_gt b a
+ | Gt => gcd_gt a b
+ end.
+
+ Theorem spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b].
+ Proof.
+ intros a b.
+ case (spec_digits a); intros H1 H2.
+ case (spec_digits b); intros H3 H4.
+ unfold gcd; generalize (spec_compare_aux a b); case compare.
+ intros HH; rewrite HH; apply sym_equal; apply Zis_gcd_gcd; auto.
+ apply Zis_gcd_refl.
+ intros; apply trans_equal with (Zgcd [b] [a]).
+ apply spec_gcd_gt; auto with zarith.
+ apply Zis_gcd_gcd; auto with zarith.
+ apply Zgcd_is_pos.
+ apply Zis_gcd_sym; apply Zgcd_is_gcd.
+ intros; apply spec_gcd_gt; auto.
+ Qed.
+
+
+ (** * Conversion *)
+
+ Definition of_N x :=
+ match x with
+ | BinNat.N0 => zero
+ | Npos p => of_pos p
+ end.
+
+ Theorem spec_of_N: forall x,
+ [of_N x] = Z_of_N x.
+ Proof.
+ intros x; case x.
+ simpl of_N.
+ unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.
+ intros p; exact (spec_of_pos p).
+ Qed.
+
+
+ (** * Shift *)
+
+ Definition shiftr n x :=
+ match compare n (Ndigits x) with
+ | Lt => unsafe_shiftr n x
+ | _ => N0 w_0
+ end.
+
+ Theorem spec_shiftr: forall n x,
+ [shiftr n x] = [x] / 2 ^ [n].
+ Proof.
+ intros n x; unfold shiftr;
+ generalize (spec_compare_aux n (Ndigits x)); case compare; intros H.
+ apply trans_equal with (1 := spec_0 w0_spec).
+ apply sym_equal; apply Zdiv_small; rewrite H.
+ rewrite spec_Ndigits; exact (spec_digits x).
+ rewrite <- spec_unsafe_shiftr; auto with zarith.
+ apply trans_equal with (1 := spec_0 w0_spec).
+ apply sym_equal; apply Zdiv_small.
+ rewrite spec_Ndigits in H; case (spec_digits x); intros H1 H2.
+ split; auto.
+ apply Zlt_le_trans with (1 := H2).
+ apply Zpower_le_monotone; auto with zarith.
+ Qed.
+
+ Definition shiftl_aux_body cont n x :=
+ match compare n (head0 x) with
+ Gt => cont n (double_size x)
+ | _ => unsafe_shiftl n x
+ end.
+
+ Theorem spec_shiftl_aux_body: forall n p x cont,
+ 2^ Zpos p <= [head0 x] ->
+ (forall x, 2 ^ (Zpos p + 1) <= [head0 x]->
+ [cont n x] = [x] * 2 ^ [n]) ->
+ [shiftl_aux_body cont n x] = [x] * 2 ^ [n].
+ Proof.
+ intros n p x cont H1 H2; unfold shiftl_aux_body.
+ generalize (spec_compare_aux n (head0 x)); case compare; 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.
+ Qed.
+
+ Fixpoint shiftl_aux p cont n x {struct p} :=
+ shiftl_aux_body
+ (fun n x => match p with
+ | xH => cont n x
+ | xO p => shiftl_aux p (shiftl_aux p cont) n x
+ | xI p => shiftl_aux p (shiftl_aux p cont) n x
+ end) n x.
+
+ Theorem spec_shiftl_aux: forall p q n x cont,
+ 2 ^ (Zpos q) <= [head0 x] ->
+ (forall x, 2 ^ (Zpos p + Zpos q) <= [head0 x] ->
+ [cont n x] = [x] * 2 ^ [n]) ->
+ [shiftl_aux p cont n x] = [x] * 2 ^ [n].
+ Proof.
+ intros p; elim p; unfold shiftl_aux; fold shiftl_aux; clear p.
+ intros p Hrec q n x cont H1 H2.
+ 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.
+ intros x3 H5; apply H2.
+ rewrite Zpos_xI.
+ replace (2 * Zpos p + 1 + Zpos q) with (Zpos p + Zpos (p + q + 1));
+ auto.
+ repeat rewrite Zpos_plus_distr; 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_monotone; auto with zarith.
+ intros x2 H4; apply Hrec with (p + q)%positive; auto.
+ intros x3 H5; apply H2.
+ rewrite (Zpos_xO p).
+ replace (2 * Zpos p + Zpos q) with (Zpos p + Zpos (p + q));
+ auto.
+ repeat rewrite Zpos_plus_distr; ring.
+ intros q n x cont H1 H2.
+ apply spec_shiftl_aux_body with (q); auto.
+ rewrite Zplus_comm; auto.
+ Qed.
+
+ Definition shiftl n x :=
+ shiftl_aux_body
+ (shiftl_aux_body
+ (shiftl_aux (digits n) unsafe_shiftl)) n x.
+
+ Theorem spec_shiftl: forall n x,
+ [shiftl n x] = [x] * 2 ^ [n].
+ Proof.
+ intros n x; unfold shiftl, shiftl_aux_body.
+ generalize (spec_compare_aux n (head0 x)); case compare; intros H.
+ apply spec_unsafe_shiftl; auto with zarith.
+ apply spec_unsafe_shiftl; auto with zarith.
+ rewrite <- (spec_double_size x).
+ generalize (spec_compare_aux n (head0 (double_size x))); case compare; 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)).
+ replace (2 ^ 1) with (2 * 1).
+ apply Zmult_le_compat_l; auto with zarith.
+ generalize (spec_double_size_head0_pos x); auto with zarith.
+ rewrite Zpower_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.
+ case (spec_digits n); auto with zarith.
+ apply Zpower_le_monotone; auto with zarith.
+ Qed.
+
+
+ (** * Zero and One *)
+
+ Theorem spec_0: [zero] = 0.
+ Proof.
+ exact (spec_0 w0_spec).
+ Qed.
+
+ Theorem spec_1: [one] = 1.
+ Proof.
+ exact (spec_1 w0_spec).
+ Qed.
+
+
+End Make.
diff --git a/theories/Numbers/Natural/BigN/NMake_gen.ml b/theories/Numbers/Natural/BigN/NMake_gen.ml
index 04c7b96d..b8552a39 100644
--- a/theories/Numbers/Natural/BigN/NMake_gen.ml
+++ b/theories/Numbers/Natural/BigN/NMake_gen.ml
@@ -8,14 +8,14 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NMake_gen.ml 11576 2008-11-10 19:13:15Z msozeau $ i*)
+(*i $Id$ i*)
(*S NMake_gen.ml : this file generates NMake.v *)
(*s The two parameters that control the generation: *)
-let size = 6 (* how many times should we repeat the Z/nZ --> Z/2nZ
+let size = 6 (* how many times should we repeat the Z/nZ --> Z/2nZ
process before relying on a generic construct *)
let gen_proof = true (* should we generate proofs ? *)
@@ -27,18 +27,18 @@ let c = "N"
let pz n = if n == 0 then "w_0" else "W0"
let rec gen2 n = if n == 0 then "1" else if n == 1 then "2"
else "2 * " ^ (gen2 (n - 1))
-let rec genxO n s =
+let rec genxO n s =
if n == 0 then s else " (xO" ^ (genxO (n - 1) s) ^ ")"
-(* NB: in ocaml >= 3.10, we could use Printf.ifprintf for printing to
- /dev/null, but for being compatible with earlier ocaml and not
- relying on system-dependent stuff like open_out "/dev/null",
+(* NB: in ocaml >= 3.10, we could use Printf.ifprintf for printing to
+ /dev/null, but for being compatible with earlier ocaml and not
+ relying on system-dependent stuff like open_out "/dev/null",
let's use instead a magical hack *)
(* Standard printer, with a final newline *)
let pr s = Printf.printf (s^^"\n")
(* Printing to /dev/null *)
-let pn = (fun s -> Obj.magic (fun _ _ _ _ _ _ _ _ _ _ _ _ _ _ -> ())
+let pn = (fun s -> Obj.magic (fun _ _ _ _ _ _ _ _ _ _ _ _ _ _ -> ())
: ('a, out_channel, unit) format -> 'a)
(* Proof printer : prints iff gen_proof is true *)
let pp = if gen_proof then pr else pn
@@ -51,7 +51,7 @@ let pp0 = if gen_proof then pr0 else pn
(*s The actual printing *)
-let _ =
+let _ =
pr "(************************************************************************)";
pr "(* v * The Coq Proof Assistant / The Coq Development Team *)";
@@ -67,21 +67,13 @@ let _ =
pr "";
pr "(** From a cyclic Z/nZ representation to arbitrary precision natural numbers.*)";
pr "";
- pr "(** Remark: File automatically generated by NMake_gen.ml, DO NOT EDIT ! *)";
+ pr "(** Remark: File automatically generated by NMake_gen.ml, DO NOT EDIT ! *)";
pr "";
- pr "Require Import BigNumPrelude.";
- pr "Require Import ZArith.";
- pr "Require Import CyclicAxioms.";
- pr "Require Import DoubleType.";
- pr "Require Import DoubleMul.";
- pr "Require Import DoubleDivn1.";
- pr "Require Import DoubleCyclic.";
- pr "Require Import Nbasic.";
- pr "Require Import Wf_nat.";
- pr "Require Import StreamMemo.";
- pr "Require Import NSig.";
+ pr "Require Import BigNumPrelude ZArith CyclicAxioms";
+ pr " DoubleType DoubleMul DoubleDivn1 DoubleCyclic Nbasic";
+ pr " Wf_nat StreamMemo.";
pr "";
- pr "Module Make (Import W0:CyclicType) <: NType.";
+ pr "Module Make (Import W0:CyclicType).";
pr "";
pr " Definition w0 := W0.w.";
@@ -132,7 +124,7 @@ let _ =
pr "";
pr " Inductive %s_ :=" t;
- for i = 0 to size do
+ for i = 0 to size do
pr " | %s%i : w%i -> %s_" c i i t
done;
pr " | %sn : forall n, word w%i (S n) -> %s_." c size t;
@@ -167,20 +159,20 @@ let _ =
pr " Definition to_N x := Zabs_N (to_Z x).";
pr "";
-
+
pr " Definition eq x y := (to_Z x = to_Z y).";
pr "";
pp " (* Regular make op (no karatsuba) *)";
- pp " Fixpoint nmake_op (ww:Type) (ww_op: znz_op ww) (n: nat) : ";
+ pp " Fixpoint nmake_op (ww:Type) (ww_op: znz_op ww) (n: nat) :";
pp " znz_op (word ww n) :=";
- pp " match n return znz_op (word ww n) with ";
+ pp " match n return znz_op (word ww n) with";
pp " O => ww_op";
- pp " | S n1 => mk_zn2z_op (nmake_op ww ww_op n1) ";
+ pp " | S n1 => mk_zn2z_op (nmake_op ww ww_op n1)";
pp " end.";
pp "";
pp " (* Simplification by rewriting for nmake_op *)";
- pp " Theorem nmake_op_S: forall ww (w_op: znz_op ww) x, ";
+ pp " Theorem nmake_op_S: forall ww (w_op: znz_op ww) x,";
pp " nmake_op _ w_op (S x) = mk_zn2z_op (nmake_op _ w_op x).";
pp " auto.";
pp " Qed.";
@@ -191,7 +183,7 @@ let _ =
for i = 0 to size do
pp " Let nmake_op%i := nmake_op _ w%i_op." i i;
pp " Let eval%in n := znz_to_Z (nmake_op%i n)." i i;
- if i == 0 then
+ if i == 0 then
pr " Let extend%i := DoubleBase.extend (WW w_0)." i
else
pr " Let extend%i := DoubleBase.extend (WW (W0: w%i))." i i;
@@ -199,8 +191,8 @@ let _ =
pr "";
- pp " Theorem digits_doubled:forall n ww (w_op: znz_op ww), ";
- pp " znz_digits (nmake_op _ w_op n) = ";
+ pp " Theorem digits_doubled:forall n ww (w_op: znz_op ww),";
+ pp " znz_digits (nmake_op _ w_op n) =";
pp " DoubleBase.double_digits (znz_digits w_op) n.";
pp " Proof.";
pp " intros n; elim n; auto; clear n.";
@@ -208,7 +200,7 @@ let _ =
pp " rewrite <- Hrec; auto.";
pp " Qed.";
pp "";
- pp " Theorem nmake_double: forall n ww (w_op: znz_op ww), ";
+ pp " Theorem nmake_double: forall n ww (w_op: znz_op ww),";
pp " znz_to_Z (nmake_op _ w_op n) =";
pp " @DoubleBase.double_to_Z _ (znz_digits w_op) (znz_to_Z w_op) n.";
pp " Proof.";
@@ -220,8 +212,8 @@ let _ =
pp "";
- pp " Theorem digits_nmake:forall n ww (w_op: znz_op ww), ";
- pp " znz_digits (nmake_op _ w_op (S n)) = ";
+ pp " Theorem digits_nmake:forall n ww (w_op: znz_op ww),";
+ pp " znz_digits (nmake_op _ w_op (S n)) =";
pp " xO (znz_digits (nmake_op _ w_op n)).";
pp " Proof.";
pp " auto.";
@@ -257,30 +249,30 @@ let _ =
pp " (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (omake_op n)))).";
pp " rewrite Hrec; auto with arith.";
pp " Qed.";
- pp " ";
+ pp "";
for i = 1 to size + 2 do
pp " Let znz_to_Z_%i: forall x y," i;
- pp " znz_to_Z w%i_op (WW x y) = " i;
+ pp " znz_to_Z w%i_op (WW x y) =" i;
pp " znz_to_Z w%i_op x * base (znz_digits w%i_op) + znz_to_Z w%i_op y." (i-1) (i-1) (i-1);
pp " Proof.";
pp " auto.";
- pp " Qed. ";
+ pp " Qed.";
pp "";
done;
pp " Let znz_to_Z_n: forall n x y,";
- pp " znz_to_Z (make_op (S n)) (WW x y) = ";
+ pp " znz_to_Z (make_op (S n)) (WW x y) =";
pp " znz_to_Z (make_op n) x * base (znz_digits (make_op n)) + znz_to_Z (make_op n) y.";
pp " Proof.";
pp " intros n x y; rewrite make_op_S; auto.";
- pp " Qed. ";
+ pp " Qed.";
pp "";
pp " Let w0_spec: znz_spec w0_op := W0.w_spec.";
for i = 1 to 3 do
- pp " Let w%i_spec: znz_spec w%i_op := mk_znz2_spec w%i_spec." i i (i-1)
+ pp " Let w%i_spec: znz_spec w%i_op := mk_znz2_spec w%i_spec." i i (i-1)
done;
for i = 4 to size + 3 do
pp " Let w%i_spec : znz_spec w%i_op := mk_znz2_karatsuba_spec w%i_spec." i i (i-1)
@@ -309,14 +301,14 @@ let _ =
for i = 0 to size do
- pp " Theorem digits_w%i: znz_digits w%i_op = znz_digits (nmake_op _ w0_op %i)." i i i;
+ pp " Theorem digits_w%i: znz_digits w%i_op = znz_digits (nmake_op _ w0_op %i)." i i i;
if i == 0 then
pp " auto."
else
pp " rewrite digits_nmake; rewrite <- digits_w%i; auto." (i - 1);
pp " Qed.";
pp "";
- pp " Let spec_double_eval%in: forall n, eval%in n = DoubleBase.double_to_Z (znz_digits w%i_op) (znz_to_Z w%i_op) n." i i i i;
+ pp " Let spec_double_eval%in: forall n, eval%in n = DoubleBase.double_to_Z (znz_digits w%i_op) (znz_to_Z w%i_op) n." i i i i;
pp " Proof.";
pp " intros n; exact (nmake_double n w%i w%i_op)." i i;
pp " Qed.";
@@ -325,7 +317,7 @@ let _ =
for i = 0 to size do
for j = 0 to (size - i) do
- pp " Theorem digits_w%in%i: znz_digits w%i_op = znz_digits (nmake_op _ w%i_op %i)." i j (i + j) i j;
+ pp " Theorem digits_w%in%i: znz_digits w%i_op = znz_digits (nmake_op _ w%i_op %i)." i j (i + j) i j;
pp " Proof.";
if j == 0 then
if i == 0 then
@@ -346,7 +338,7 @@ let _ =
end;
pp " Qed.";
pp "";
- pp " Let spec_eval%in%i: forall x, [%s%i x] = eval%in %i x." i j c (i + j) i j;
+ pp " Let spec_eval%in%i: forall x, [%s%i x] = eval%in %i x." i j c (i + j) i j;
pp " Proof.";
if j == 0 then
pp " intros x; rewrite spec_double_eval%in; unfold DoubleBase.double_to_Z, to_Z; auto." i
@@ -363,7 +355,7 @@ let _ =
pp " Qed.";
if i + j <> size then
begin
- pp " Let spec_extend%in%i: forall x, [%s%i x] = [%s%i (extend%i %i x)]." i (i + j + 1) c i c (i + j + 1) i j;
+ pp " Let spec_extend%in%i: forall x, [%s%i x] = [%s%i (extend%i %i x)]." i (i + j + 1) c i c (i + j + 1) i j;
if j == 0 then
begin
pp " intros x; change (extend%i 0 x) with (WW (znz_0 w%i_op) x)." i (i + j);
@@ -393,7 +385,7 @@ let _ =
pp " Qed.";
pp "";
- pp " Let spec_eval%in%i: forall x, [%sn 0 x] = eval%in %i x." i (size - i + 1) c i (size - i + 1);
+ pp " Let spec_eval%in%i: forall x, [%sn 0 x] = eval%in %i x." i (size - i + 1) c i (size - i + 1);
pp " Proof.";
pp " intros x; case x.";
pp " auto.";
@@ -405,7 +397,7 @@ let _ =
pp " Qed.";
pp "";
- pp " Let spec_eval%in%i: forall x, [%sn 1 x] = eval%in %i x." i (size - i + 2) c i (size - i + 2);
+ pp " Let spec_eval%in%i: forall x, [%sn 1 x] = eval%in %i x." i (size - i + 2) c i (size - i + 2);
pp " intros x; case x.";
pp " auto.";
pp " intros xh xl; unfold to_Z; rewrite znz_to_Z_%i." (size + 2);
@@ -430,7 +422,7 @@ let _ =
pp " Qed.";
pp "";
- pp " Let spec_eval%in: forall n x, [%sn n x] = eval%in (S n) x." size c size;
+ pp " Let spec_eval%in: forall n x, [%sn n x] = eval%in (S n) x." size c size;
pp " intros n; elim n; clear n.";
pp " exact spec_eval%in1." size;
pp " intros n Hrec x; case x; clear x.";
@@ -446,7 +438,7 @@ let _ =
pp " Qed.";
pp "";
- pp " Let spec_extend%in: forall n x, [%s%i x] = [%sn n (extend%i n x)]." size c size c size ;
+ pp " Let spec_extend%in: forall n x, [%s%i x] = [%sn n (extend%i n x)]." size c size c size ;
pp " intros n; elim n; clear n.";
pp " intros x; change (extend%i 0 x) with (WW (znz_0 w%i_op) x)." size size;
pp " unfold to_Z.";
@@ -478,7 +470,6 @@ let _ =
pp " unfold to_Z.";
pp " case n1; auto; intros n2; repeat rewrite make_op_S; auto.";
pp " Qed.";
- pp " Hint Rewrite spec_extendn_0: extr.";
pp "";
pp " Let spec_extendn0_0: forall n wx, [%sn (S n) (WW W0 wx)] = [%sn n wx]." c c;
pp " Proof.";
@@ -489,7 +480,6 @@ let _ =
pp " case n; auto.";
pp " intros n1; rewrite make_op_S; auto.";
pp " Qed.";
- pp " Hint Rewrite spec_extendn_0: extr.";
pp "";
pp " Let spec_extend_tr: forall m n (w: word _ (S n)),";
pp " [%sn (m + n) (extend_tr w m)] = [%sn n w]." c c;
@@ -498,7 +488,6 @@ let _ =
pp " intros n x; simpl extend_tr.";
pp " simpl plus; rewrite spec_extendn0_0; auto.";
pp " Qed.";
- pp " Hint Rewrite spec_extend_tr: extr.";
pp "";
pp " Let spec_cast_l: forall n m x1,";
pp " [%sn (Max.max n m)" c;
@@ -508,7 +497,6 @@ let _ =
pp " intros n m x1; case (diff_r n m); simpl castm.";
pp " rewrite spec_extend_tr; auto.";
pp " Qed.";
- pp " Hint Rewrite spec_cast_l: extr.";
pp "";
pp " Let spec_cast_r: forall n m x1,";
pp " [%sn (Max.max n m)" c;
@@ -518,7 +506,6 @@ let _ =
pp " intros n m x1; case (diff_l n m); simpl castm.";
pp " rewrite spec_extend_tr; auto.";
pp " Qed.";
- pp " Hint Rewrite spec_cast_r: extr.";
pp "";
@@ -578,14 +565,14 @@ let _ =
pr " | %s%i wx, %s%i wy => f%i (extend%i %i wx) wy" c i c j j i (j - i - 1);
done;
if i == size then
- pr " | %s%i wx, %sn m wy => fnn m (extend%i m wx) wy" c size c size
- else
+ pr " | %s%i wx, %sn m wy => fnn m (extend%i m wx) wy" c size c size
+ else
pr " | %s%i wx, %sn m wy => fnn m (extend%i m (extend%i %i wx)) wy" c i c size i (size - i - 1);
done;
for i = 0 to size do
if i == size then
- pr " | %sn n wx, %s%i wy => fnn n wx (extend%i n wy)" c c size size
- else
+ pr " | %sn n wx, %s%i wy => fnn n wx (extend%i n wy)" c c size size
+ else
pr " | %sn n wx, %s%i wy => fnn n wx (extend%i n (extend%i %i wy))" c c i size i (size - i - 1);
done;
pr " | %sn n wx, Nn m wy =>" c;
@@ -611,17 +598,17 @@ let _ =
done;
if i == size then
pp " intros m y; rewrite (spec_extend%in m); apply Pfnn." size
- else
+ else
pp " intros m y; rewrite spec_extend%in%i; rewrite (spec_extend%in m); apply Pfnn." i size size;
done;
pp " intros n x y; case y; clear y.";
for i = 0 to size do
if i == size then
pp " intros y; rewrite (spec_extend%in n); apply Pfnn." size
- else
+ else
pp " intros y; rewrite spec_extend%in%i; rewrite (spec_extend%in n); apply Pfnn." i size size;
done;
- pp " intros m y; rewrite <- (spec_cast_l n m x); ";
+ pp " intros m y; rewrite <- (spec_cast_l n m x);";
pp " rewrite <- (spec_cast_r n m y); apply Pfnn.";
pp " Qed.";
pp "";
@@ -644,7 +631,7 @@ let _ =
pr " match y with";
for j = 0 to i - 1 do
pr " | %s%i wy =>" c j;
- if j == 0 then
+ if j == 0 then
pr " if w0_eq0 wy then ft0 x else";
pr " f%i wx (extend%i %i wy)" i j (i - j -1);
done;
@@ -653,8 +640,8 @@ let _ =
pr " | %s%i wy => f%i (extend%i %i wx) wy" c j j i (j - i - 1);
done;
if i == size then
- pr " | %sn m wy => fnn m (extend%i m wx) wy" c size
- else
+ pr " | %sn m wy => fnn m (extend%i m wx) wy" c size
+ else
pr " | %sn m wy => fnn m (extend%i m (extend%i %i wx)) wy" c size i (size - i - 1);
pr" end";
done;
@@ -665,8 +652,8 @@ let _ =
if i == 0 then
pr " if w0_eq0 wy then ft0 x else";
if i == size then
- pr " fnn n wx (extend%i n wy)" size
- else
+ pr " fnn n wx (extend%i n wy)" size
+ else
pr " fnn n wx (extend%i n (extend%i %i wy))" size i (size - i - 1);
done;
pr " | %sn m wy =>" c;
@@ -707,7 +694,7 @@ let _ =
done;
if i == size then
pp " intros m y; rewrite (spec_extend%in m); apply Pfnn." size
- else
+ else
pp " intros m y; rewrite spec_extend%in%i; rewrite (spec_extend%in m); apply Pfnn." i size size;
done;
pp " intros n x y; case y; clear y.";
@@ -721,16 +708,16 @@ let _ =
end;
if i == size then
pp " rewrite (spec_extend%in n); apply Pfnn." size
- else
+ else
pp " rewrite spec_extend%in%i; rewrite (spec_extend%in n); apply Pfnn." i size size;
done;
- pp " intros m y; rewrite <- (spec_cast_l n m x); ";
+ pp " intros m y; rewrite <- (spec_cast_l n m x);";
pp " rewrite <- (spec_cast_r n m y); apply Pfnn.";
pp " Qed.";
pp "";
pr " (* We iter the smaller argument with the bigger *)";
- pr " Definition iter (x y: t_): res := ";
+ pr " Definition iter (x y: t_): res :=";
pr0 " Eval lazy zeta beta iota delta [";
for i = 0 to size do
pr0 "extend%i " i;
@@ -748,14 +735,14 @@ let _ =
pr " | %s%i wx, %s%i wy => f%in %i wx wy" c i c j i (j - i - 1);
done;
if i == size then
- pr " | %s%i wx, %sn m wy => f%in m wx wy" c size c size
- else
+ pr " | %s%i wx, %sn m wy => f%in m wx wy" c size c size
+ else
pr " | %s%i wx, %sn m wy => f%in m (extend%i %i wx) wy" c i c size i (size - i - 1);
done;
for i = 0 to size do
if i == size then
- pr " | %sn n wx, %s%i wy => fn%i n wx wy" c c size size
- else
+ pr " | %sn n wx, %s%i wy => fn%i n wx wy" c c size size
+ else
pr " | %sn n wx, %s%i wy => fn%i n wx (extend%i %i wy)" c c i size i (size - i - 1);
done;
pr " | %sn n wx, %sn m wy => fnm n m wx wy" c c;
@@ -765,6 +752,7 @@ let _ =
pp " Ltac zg_tac := try";
pp " (red; simpl Zcompare; auto;";
pp " let t := fresh \"H\" in (intros t; discriminate t)).";
+ pp "";
pp " Lemma spec_iter: forall x y, P [x] [y] (iter x y).";
pp " Proof.";
pp " intros x; case x; clear x; unfold iter.";
@@ -779,14 +767,14 @@ let _ =
done;
if i == size then
pp " intros m y; rewrite spec_eval%in; apply Pf%in." size size
- else
+ else
pp " intros m y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pf%in." i size size size;
done;
pp " intros n x y; case y; clear y.";
for i = 0 to size do
if i == size then
pp " intros y; rewrite spec_eval%in; apply Pfn%i." size size
- else
+ else
pp " intros y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pfn%i." i size size size;
done;
pp " intros m y; apply Pfnm.";
@@ -820,8 +808,8 @@ let _ =
pr " | %s%i wy => f%in %i wx wy" c j i (j - i - 1);
done;
if i == size then
- pr " | %sn m wy => f%in m wx wy" c size
- else
+ pr " | %sn m wy => f%in m wx wy" c size
+ else
pr " | %sn m wy => f%in m (extend%i %i wx) wy" c size i (size - i - 1);
pr " end";
done;
@@ -832,8 +820,8 @@ let _ =
if i == 0 then
pr " if w0_eq0 wy then ft0 x else";
if i == size then
- pr " fn%i n wx wy" size
- else
+ pr " fn%i n wx wy" size
+ else
pr " fn%i n wx (extend%i %i wy)" size i (size - i - 1);
done;
pr " | %sn m wy => fnm n m wx wy" c;
@@ -869,7 +857,7 @@ let _ =
done;
if i == size then
pp " intros m y; rewrite spec_eval%in; apply Pf%in." size size
- else
+ else
pp " intros m y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pf%in." i size size size;
done;
pp " intros n x y; case y; clear y.";
@@ -883,7 +871,7 @@ let _ =
end;
if i == size then
pp " rewrite spec_eval%in; apply Pfn%i." size size
- else
+ else
pp " rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pfn%i." i size size size;
done;
pp " intros m y; apply Pfnm.";
@@ -897,27 +885,27 @@ let _ =
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Reduction *)";
+ pr " (** * Reduction *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
- pr " Definition reduce_0 (x:w) := %s0 x." c;
+ pr " Definition reduce_0 (x:w) := %s0 x." c;
pr " Definition reduce_1 :=";
pr " Eval lazy beta iota delta[reduce_n1] in";
pr " reduce_n1 _ _ zero w0_eq0 %s0 %s1." c c;
for i = 2 to size do
pr " Definition reduce_%i :=" i;
pr " Eval lazy beta iota delta[reduce_n1] in";
- pr " reduce_n1 _ _ zero w%i_eq0 reduce_%i %s%i."
+ pr " reduce_n1 _ _ zero w%i_eq0 reduce_%i %s%i."
(i-1) (i-1) c i
done;
pr " Definition reduce_%i :=" (size+1);
pr " Eval lazy beta iota delta[reduce_n1] in";
- pr " reduce_n1 _ _ zero w%i_eq0 reduce_%i (%sn 0)."
- size size c;
+ pr " reduce_n1 _ _ zero w%i_eq0 reduce_%i (%sn 0)."
+ size size c;
- pr " Definition reduce_n n := ";
+ pr " Definition reduce_n n :=";
pr " Eval lazy beta iota delta[reduce_n] in";
pr " reduce_n _ _ zero reduce_%i %sn n." (size + 1) c;
pr "";
@@ -927,7 +915,7 @@ let _ =
pp " intros x; unfold to_Z, reduce_0.";
pp " auto.";
pp " Qed.";
- pp " ";
+ pp "";
for i = 1 to size + 1 do
if i == size + 1 then
@@ -938,14 +926,14 @@ let _ =
pp " intros x; case x; unfold reduce_%i." i;
pp " exact (spec_0 w0_spec).";
pp " intros x1 y1.";
- pp " generalize (spec_w%i_eq0 x1); " (i - 1);
+ pp " generalize (spec_w%i_eq0 x1);" (i - 1);
pp " case w%i_eq0; intros H1; auto." (i - 1);
- if i <> 1 then
+ if i <> 1 then
pp " rewrite spec_reduce_%i." (i - 1);
pp " unfold to_Z; rewrite znz_to_Z_%i." i;
pp " unfold to_Z in H1; rewrite H1; auto.";
pp " Qed.";
- pp " ";
+ pp "";
done;
pp " Let spec_reduce_n: forall n x, [reduce_n n x] = [%sn n x]." c;
@@ -959,11 +947,11 @@ let _ =
pp " rewrite Hrec.";
pp " rewrite spec_extendn0_0; auto.";
pp " Qed.";
- pp " ";
+ pp "";
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Successor *)";
+ pr " (** * Successor *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
@@ -983,19 +971,19 @@ let _ =
for i = 0 to size-1 do
pr " | %s%i wx =>" c i;
pr " match w%i_succ_c wx with" i;
- pr " | C0 r => %s%i r" c i;
+ pr " | C0 r => %s%i r" c i;
pr " | C1 r => %s%i (WW one%i r)" c (i+1) i;
pr " end";
done;
pr " | %s%i wx =>" c size;
pr " match w%i_succ_c wx with" size;
- pr " | C0 r => %s%i r" c size;
+ pr " | C0 r => %s%i r" c size;
pr " | C1 r => %sn 0 (WW one%i r)" c size ;
pr " end";
pr " | %sn n wx =>" c;
pr " let op := make_op n in";
pr " match op.(znz_succ_c) wx with";
- pr " | C0 r => %sn n r" c;
+ pr " | C0 r => %sn n r" c;
pr " | C1 r => %sn (S n) (WW op.(znz_1) r)" c;
pr " end";
pr " end.";
@@ -1027,13 +1015,13 @@ let _ =
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Adddition *)";
+ pr " (** * Adddition *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
for i = 0 to size do
- pr " Definition w%i_add_c := znz_add_c w%i_op." i i;
+ pr " Definition w%i_add_c := znz_add_c w%i_op." i i;
pr " Definition w%i_add x y :=" i;
pr " match w%i_add_c x y with" i;
pr " | C0 r => %s%i r" c i;
@@ -1057,26 +1045,24 @@ let _ =
pp " Proof.";
pp " intros n m; unfold to_Z, w%i_add, w%i_add_c." i i;
pp " generalize (spec_add_c w%i_spec n m); case znz_add_c; auto." i;
- pp " intros ww H; rewrite <- H.";
+ pp " intros ww H; rewrite <- H.";
pp " rewrite znz_to_Z_%i; unfold interp_carry;" (i + 1);
pp " apply f_equal2 with (f := Zplus); auto;";
pp " apply f_equal2 with (f := Zmult); auto;";
pp " exact (spec_1 w%i_spec)." i;
pp " Qed.";
- pp " Hint Rewrite spec_w%i_add: addr." i;
pp "";
done;
pp " Let spec_wn_add: forall n x y, [addn n x y] = [%sn n x] + [%sn n y]." c c;
pp " Proof.";
pp " intros k n m; unfold to_Z, addn.";
pp " generalize (spec_add_c (wn_spec k) n m); case znz_add_c; auto.";
- pp " intros ww H; rewrite <- H.";
+ pp " intros ww H; rewrite <- H.";
pp " rewrite (znz_to_Z_n k); unfold interp_carry;";
pp " apply f_equal2 with (f := Zplus); auto;";
pp " apply f_equal2 with (f := Zmult); auto;";
pp " exact (spec_1 (wn_spec k)).";
pp " Qed.";
- pp " Hint Rewrite spec_wn_add: addr.";
pr " Definition add := Eval lazy beta delta [same_level] in";
pr0 " (same_level t_ ";
@@ -1101,7 +1087,7 @@ let _ =
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Predecessor *)";
+ pr " (** * Predecessor *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
@@ -1116,25 +1102,25 @@ let _ =
for i = 0 to size do
pr " | %s%i wx =>" c i;
pr " match w%i_pred_c wx with" i;
- pr " | C0 r => reduce_%i r" i;
+ pr " | C0 r => reduce_%i r" i;
pr " | C1 r => zero";
pr " end";
done;
pr " | %sn n wx =>" c;
pr " let op := make_op n in";
pr " match op.(znz_pred_c) wx with";
- pr " | C0 r => reduce_n n r";
+ pr " | C0 r => reduce_n n r";
pr " | C1 r => zero";
pr " end";
pr " end.";
pr "";
- pr " Theorem spec_pred: forall x, 0 < [x] -> [pred x] = [x] - 1.";
+ pr " Theorem spec_pred_pos : forall x, 0 < [x] -> [pred x] = [x] - 1.";
pa " Admitted.";
pp " Proof.";
pp " intros x; case x; unfold pred.";
for i = 0 to size do
- pp " intros x1 H1; unfold w%i_pred_c; " i;
+ pp " intros x1 H1; unfold w%i_pred_c;" i;
pp " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1." i;
pp " rewrite spec_reduce_%i; auto." i;
pp " unfold interp_carry; unfold to_Z.";
@@ -1143,7 +1129,7 @@ let _ =
pp " assert (znz_to_Z w%i_op x1 - 1 < 0); auto with zarith." i;
pp " unfold to_Z in H1; auto with zarith.";
done;
- pp " intros n x1 H1; ";
+ pp " intros n x1 H1;";
pp " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1.";
pp " rewrite spec_reduce_n; auto.";
pp " unfold interp_carry; unfold to_Z.";
@@ -1152,32 +1138,31 @@ let _ =
pp " assert (znz_to_Z (make_op n) x1 - 1 < 0); auto with zarith.";
pp " unfold to_Z in H1; auto with zarith.";
pp " Qed.";
- pp " ";
-
+ pp "";
+
pp " Let spec_pred0: forall x, [x] = 0 -> [pred x] = 0.";
pp " Proof.";
pp " intros x; case x; unfold pred.";
for i = 0 to size do
- pp " intros x1 H1; unfold w%i_pred_c; " i;
+ pp " intros x1 H1; unfold w%i_pred_c;" i;
pp " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1." i;
pp " unfold interp_carry; unfold to_Z.";
pp " unfold to_Z in H1; auto with zarith.";
pp " case (spec_to_Z w%i_spec y1); intros HH3 HH4; auto with zarith." i;
pp " intros; exact (spec_0 w0_spec).";
done;
- pp " intros n x1 H1; ";
+ pp " intros n x1 H1;";
pp " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1.";
pp " unfold interp_carry; unfold to_Z.";
pp " unfold to_Z in H1; auto with zarith.";
pp " case (spec_to_Z (wn_spec n) y1); intros HH3 HH4; auto with zarith.";
pp " intros; exact (spec_0 w0_spec).";
pp " Qed.";
- pr " ";
-
+ pr "";
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Subtraction *)";
+ pr " (** * Subtraction *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
@@ -1187,7 +1172,7 @@ let _ =
done;
pr "";
- for i = 0 to size do
+ for i = 0 to size do
pr " Definition w%i_sub x y :=" i;
pr " match w%i_sub_c x y with" i;
pr " | C0 r => reduce_%i r" i;
@@ -1208,8 +1193,8 @@ let _ =
pp " Let spec_w%i_sub: forall x y, [%s%i y] <= [%s%i x] -> [w%i_sub x y] = [%s%i x] - [%s%i y]." i c i c i i c i c i;
pp " Proof.";
pp " intros n m; unfold w%i_sub, w%i_sub_c." i i;
- pp " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; " i;
- if i == 0 then
+ pp " generalize (spec_sub_c w%i_spec n m); case znz_sub_c;" i;
+ if i == 0 then
pp " intros x; auto."
else
pp " intros x; try rewrite spec_reduce_%i; auto." i;
@@ -1219,11 +1204,11 @@ let _ =
pp " Qed.";
pp "";
done;
-
+
pp " Let spec_wn_sub: forall n x y, [%sn n y] <= [%sn n x] -> [subn n x y] = [%sn n x] - [%sn n y]." c c c c;
pp " Proof.";
pp " intros k n m; unfold subn.";
- pp " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; ";
+ pp " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c;";
pp " intros x; auto.";
pp " unfold interp_carry, to_Z.";
pp " case (spec_to_Z (wn_spec k) x); intros; auto with zarith.";
@@ -1238,7 +1223,7 @@ let _ =
pr "subn).";
pr "";
- pr " Theorem spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y].";
+ pr " Theorem spec_sub_pos : forall x y, [y] <= [x] -> [sub x y] = [x] - [y].";
pa " Admitted.";
pp " Proof.";
pp " unfold sub.";
@@ -1255,7 +1240,7 @@ let _ =
pp " Let spec_w%i_sub0: forall x y, [%s%i x] < [%s%i y] -> [w%i_sub x y] = 0." i c i c i i;
pp " Proof.";
pp " intros n m; unfold w%i_sub, w%i_sub_c." i i;
- pp " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; " i;
+ pp " generalize (spec_sub_c w%i_spec n m); case znz_sub_c;" i;
pp " intros x; unfold interp_carry.";
pp " unfold to_Z; case (spec_to_Z w%i_spec x); intros; auto with zarith." i;
pp " intros; unfold to_Z, zero, w_0; rewrite (spec_0 w0_spec); auto.";
@@ -1266,7 +1251,7 @@ let _ =
pp " Let spec_wn_sub0: forall n x y, [%sn n x] < [%sn n y] -> [subn n x y] = 0." c c;
pp " Proof.";
pp " intros k n m; unfold subn.";
- pp " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; ";
+ pp " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c;";
pp " intros x; unfold interp_carry.";
pp " unfold to_Z; case (spec_to_Z (wn_spec k) x); intros; auto with zarith.";
pp " intros; unfold to_Z, w_0; rewrite (spec_0 (w0_spec)); auto.";
@@ -1289,7 +1274,7 @@ let _ =
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Comparison *)";
+ pr " (** * Comparison *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
@@ -1299,7 +1284,7 @@ let _ =
pr " Definition comparen_%i :=" i;
pr " compare_mn_1 w%i w%i %s compare_%i (compare_%i %s) compare_%i." i i (pz i) i i (pz i) i
done;
- pr "";
+ pr "";
pr " Definition comparenm n m wx wy :=";
pr " let mn := Max.max n m in";
@@ -1310,8 +1295,8 @@ let _ =
pr " (castm (diff_l n m) (extend_tr wy (fst d))).";
pr "";
- pr " Definition compare := Eval lazy beta delta [iter] in ";
- pr " (iter _ ";
+ pr " Definition compare := Eval lazy beta delta [iter] in";
+ pr " (iter _";
for i = 0 to size do
pr " compare_%i" i;
pr " (fun n x y => opp_compare (comparen_%i (S n) y x))" i;
@@ -1320,15 +1305,9 @@ let _ =
pr " comparenm).";
pr "";
- pr " Definition lt n m := compare n m = Lt.";
- pr " Definition le n m := compare n m <> Gt.";
- pr " Definition min n m := match compare n m with Gt => m | _ => n end.";
- pr " Definition max n m := match compare n m with Lt => m | _ => n end.";
- pr "";
-
for i = 0 to size do
pp " Let spec_compare_%i: forall x y," i;
- pp " match compare_%i x y with " i;
+ pp " match compare_%i x y with" i;
pp " Eq => [%s%i x] = [%s%i y]" c i c i;
pp " | Lt => [%s%i x] < [%s%i y]" c i c i;
pp " | Gt => [%s%i x] > [%s%i y]" c i c i;
@@ -1337,7 +1316,7 @@ let _ =
pp " unfold compare_%i, to_Z; exact (spec_compare w%i_spec)." i i;
pp " Qed.";
pp "";
-
+
pp " Let spec_comparen_%i:" i;
pp " forall (n : nat) (x : word w%i n) (y : w%i)," i i;
pp " match comparen_%i n x y with" i;
@@ -1367,16 +1346,16 @@ let _ =
pp "";
- pr " Theorem spec_compare: forall x y,";
- pr " match compare x y with ";
+ pr " Theorem spec_compare_aux: forall x y,";
+ pr " match compare x y with";
pr " Eq => [x] = [y]";
pr " | Lt => [x] < [y]";
pr " | Gt => [x] > [y]";
pr " end.";
pa " Admitted.";
pp " Proof.";
- pp " refine (spec_iter _ (fun x y res => ";
- pp " match res with ";
+ pp " refine (spec_iter _ (fun x y res =>";
+ pp " match res with";
pp " Eq => x = y";
pp " | Lt => x < y";
pp " | Gt => x > y";
@@ -1387,12 +1366,12 @@ let _ =
pp " (fun n => comparen_%i (S n)) _ _ _" i;
done;
pp " comparenm _).";
-
+
for i = 0 to size - 1 do
pp " exact spec_compare_%i." i;
pp " intros n x y H;apply spec_opp_compare; apply spec_comparen_%i." i;
pp " intros n x y H; exact (spec_comparen_%i (S n) x y)." i;
- done;
+ done;
pp " exact spec_compare_%i." size;
pp " intros n x y;apply spec_opp_compare; apply spec_comparen_%i." size;
pp " intros n; exact (spec_comparen_%i (S n))." size;
@@ -1402,28 +1381,9 @@ let _ =
pp " Qed.";
pr "";
- pr " Definition eq_bool x y :=";
- pr " match compare x y with";
- pr " | Eq => true";
- pr " | _ => false";
- pr " end.";
- pr "";
-
-
- pr " Theorem spec_eq_bool: forall x y,";
- pr " if eq_bool x y then [x] = [y] else [x] <> [y].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros x y; unfold eq_bool.";
- pp " generalize (spec_compare x y); case compare; auto with zarith.";
- pp " Qed.";
- pr "";
-
-
-
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Multiplication *)";
+ pr " (** * Multiplication *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
@@ -1461,7 +1421,7 @@ let _ =
pr " match n return word w%i (S n) -> t_ with" i;
for j = 0 to size - i do
if (i + j) == size then
- begin
+ begin
pr " | %i%s => fun x => %sn 0 x" j "%nat" c;
pr " | %i%s => fun x => %sn 1 x" (j + 1) "%nat" c
end
@@ -1471,7 +1431,7 @@ let _ =
pr " | _ => fun _ => N0 w_0";
pr " end.";
pr "";
- done;
+ done;
for i = 0 to size - 1 do
@@ -1486,7 +1446,7 @@ let _ =
pp " repeat rewrite inj_S; unfold Zsucc; auto with zarith.";
pp " Qed.";
pp "";
- done;
+ done;
for i = 0 to size do
@@ -1497,8 +1457,8 @@ let _ =
pr " if w%i_eq0 w then %sn n r" i c;
pr " else %sn (S n) (WW (extend%i n w) r)." c i;
end
- else
- begin
+ else
+ begin
pr " if w%i_eq0 w then to_Z%i n r" i i;
pr " else to_Z%i (S n) (WW (extend%i n w) r)." i i;
end;
@@ -1514,10 +1474,10 @@ let _ =
pr " (castm (diff_l n m) (extend_tr y (fst d)))).";
pr "";
- pr " Definition mul := Eval lazy beta delta [iter0] in ";
- pr " (iter0 t_ ";
+ pr " Definition mul := Eval lazy beta delta [iter0] in";
+ pr " (iter0 t_";
for i = 0 to size do
- pr " (fun x y => reduce_%i (w%i_mul_c x y)) " (i + 1) i;
+ pr " (fun x y => reduce_%i (w%i_mul_c x y))" (i + 1) i;
pr " (fun n x y => w%i_mul n y x)" i;
pr " w%i_mul" i;
done;
@@ -1556,7 +1516,7 @@ let _ =
pp " Qed.";
pp "";
done;
-
+
pp " Lemma nmake_op_WW: forall ww ww1 n x y,";
pp " znz_to_Z (nmake_op ww ww1 (S n)) (WW x y) =";
pp " znz_to_Z (nmake_op ww ww1 n) x * base (znz_digits (nmake_op ww ww1 n)) +";
@@ -1564,21 +1524,21 @@ let _ =
pp " auto.";
pp " Qed.";
pp "";
-
+
for i = 0 to size do
pp " Lemma extend%in_spec: forall n x1," i;
- pp " znz_to_Z (nmake_op _ w%i_op (S n)) (extend%i n x1) = " i i;
+ pp " znz_to_Z (nmake_op _ w%i_op (S n)) (extend%i n x1) =" i i;
pp " znz_to_Z w%i_op x1." i;
pp " Proof.";
pp " intros n1 x2; rewrite nmake_double.";
pp " unfold extend%i." i;
pp " rewrite DoubleBase.spec_extend; auto.";
- if i == 0 then
+ if i == 0 then
pp " intros l; simpl; unfold w_0; rewrite (spec_0 w0_spec); ring.";
pp " Qed.";
pp "";
done;
-
+
pp " Lemma spec_muln:";
pp " forall n (x: word _ (S n)) y,";
pp " [%sn (S n) (znz_mul_c (make_op n) x y)] = [%sn n x] * [%sn n y]." c c c;
@@ -1588,12 +1548,13 @@ let _ =
pp " rewrite make_op_S.";
pp " case znz_mul_c; auto.";
pp " Qed.";
+ pr "";
pr " Theorem spec_mul: forall x y, [mul x y] = [x] * [y].";
pa " Admitted.";
pp " Proof.";
for i = 0 to size do
- pp " assert(F%i: " i;
+ pp " assert(F%i:" i;
pp " forall n x y,";
if i <> size then
pp0 " Z_of_nat n <= %i -> " (size - i);
@@ -1614,7 +1575,7 @@ let _ =
pp " generalize (spec_w%i_eq0 x1); case w%i_eq0; intros HH." i i;
pp " unfold to_Z in HH; rewrite HH.";
if i == size then
- begin
+ begin
pp " rewrite spec_eval%in; unfold eval%in, nmake_op%i; auto." i i i;
pp " rewrite spec_eval%in; unfold eval%in, nmake_op%i." i i i
end
@@ -1627,7 +1588,7 @@ let _ =
done;
pp " refine (spec_iter0 t_ (fun x y res => [res] = x * y)";
for i = 0 to size do
- pp " (fun x y => reduce_%i (w%i_mul_c x y)) " (i + 1) i;
+ pp " (fun x y => reduce_%i (w%i_mul_c x y))" (i + 1) i;
pp " (fun n x y => w%i_mul n y x)" i;
pp " w%i_mul _ _ _" i;
done;
@@ -1643,12 +1604,12 @@ let _ =
if i == size then
begin
pp " intros n x y; rewrite F%i; auto with zarith." i;
- pp " intros n x y; rewrite F%i; auto with zarith. " i;
+ pp " intros n x y; rewrite F%i; auto with zarith." i;
end
else
begin
pp " intros n x y H; rewrite F%i; auto with zarith." i;
- pp " intros n x y H; rewrite F%i; auto with zarith. " i;
+ pp " intros n x y H; rewrite F%i; auto with zarith." i;
end;
done;
pp " intros n m x y; unfold mulnm.";
@@ -1663,7 +1624,7 @@ let _ =
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Square *)";
+ pr " (** * Square *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
@@ -1702,42 +1663,9 @@ let _ =
pp "Qed.";
pr "";
-
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Power *)";
- pr " (* *)";
- pr " (***************************************************************)";
- pr "";
-
- pr " Fixpoint power_pos (x:%s) (p:positive) {struct p} : %s :=" t t;
- pr " match p with";
- pr " | xH => x";
- pr " | xO p => square (power_pos x p)";
- pr " | xI p => mul (square (power_pos x p)) x";
- pr " end.";
- pr "";
-
- pr " Theorem spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.";
- pa " Admitted.";
- pp " Proof.";
- pp " intros x n; generalize x; elim n; clear n x; simpl power_pos.";
- pp " intros; rewrite spec_mul; rewrite spec_square; rewrite H.";
- pp " rewrite Zpos_xI; rewrite Zpower_exp; auto with zarith.";
- pp " rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.";
- pp " rewrite Zpower_2; rewrite Zpower_1_r; auto.";
- pp " intros; rewrite spec_square; rewrite H.";
- pp " rewrite Zpos_xO; auto with zarith.";
- pp " rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.";
- pp " rewrite Zpower_2; auto.";
- pp " intros; rewrite Zpower_1_r; auto.";
- pp " Qed.";
- pp "";
- pr "";
-
- pr " (***************************************************************)";
- pr " (* *)";
- pr " (* Square root *)";
+ pr " (** * Square root *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
@@ -1772,26 +1700,26 @@ let _ =
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Division *)";
+ pr " (** * Division *)";
pr " (* *)";
pr " (***************************************************************)";
- pr "";
+ pr "";
for i = 0 to size do
pr " Definition w%i_div_gt := w%i_op.(znz_div_gt)." i i
done;
pr "";
- pp " Let spec_divn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) := ";
- pp " (spec_double_divn1 ";
+ pp " Let spec_divn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) :=";
+ pp " (spec_double_divn1";
pp " ww_op.(znz_zdigits) ww_op.(znz_0)";
pp " (znz_WW ww_op) ww_op.(znz_head0)";
pp " ww_op.(znz_add_mul_div) ww_op.(znz_div21)";
pp " ww_op.(znz_compare) ww_op.(znz_sub) (znz_to_Z ww_op)";
- pp " (spec_to_Z ww_spec) ";
+ pp " (spec_to_Z ww_spec)";
pp " (spec_zdigits ww_spec)";
pp " (spec_0 ww_spec) (spec_WW ww_spec) (spec_head0 ww_spec)";
- pp " (spec_add_mul_div ww_spec) (spec_div21 ww_spec) ";
+ pp " (spec_add_mul_div ww_spec) (spec_div21 ww_spec)";
pp " (CyclicAxioms.spec_compare ww_spec) (CyclicAxioms.spec_sub ww_spec)).";
pp "";
@@ -1811,7 +1739,7 @@ let _ =
for i = 0 to size do
pp " Lemma spec_get_end%i: forall n x y," i;
- pp " eval%in n x <= [%s%i y] -> " i c i;
+ pp " eval%in n x <= [%s%i y] ->" i c i;
pp " [%s%i (DoubleBase.get_low %s n x)] = eval%in n x." c i (pz i) i;
pp " Proof.";
pp " intros n x y H.";
@@ -1843,8 +1771,8 @@ let _ =
pr "";
pr " Definition div_gt := Eval lazy beta delta [iter] in";
- pr " (iter _ ";
- for i = 0 to size do
+ pr " (iter _";
+ for i = 0 to size do
pr " div_gt%i" i;
pr " (fun n x y => div_gt%i x (DoubleBase.get_low %s (S n) y))" i (pz i);
pr " w%i_divn1" i;
@@ -1862,10 +1790,10 @@ let _ =
pp " forall x y, [x] > [y] -> 0 < [y] ->";
pp " let (q,r) := div_gt x y in";
pp " [x] = [q] * [y] + [r] /\\ 0 <= [r] < [y]).";
- pp " refine (spec_iter (t_*t_) (fun x y res => x > y -> 0 < y ->";
+ pp " refine (spec_iter (t_*t_) (fun x y res => x > y -> 0 < y ->";
pp " let (q,r) := res in";
pp " x = [q] * y + [r] /\\ 0 <= [r] < y)";
- for i = 0 to size do
+ for i = 0 to size do
pp " div_gt%i" i;
pp " (fun n x y => div_gt%i x (DoubleBase.get_low %s (S n) y))" i (pz i);
pp " w%i_divn1 _ _ _" i;
@@ -1879,11 +1807,11 @@ let _ =
pp " intros n x y H2 H3; unfold div_gt%i, w%i_div_gt." i i
else
pp " intros n x y H1 H2 H3; unfold div_gt%i, w%i_div_gt." i i;
- pp " generalize (spec_div_gt w%i_spec x " i;
+ pp " generalize (spec_div_gt w%i_spec x" i;
pp " (DoubleBase.get_low %s (S n) y))." (pz i);
- pp0 " ";
+ pp0 "";
for j = 0 to i do
- pp0 "unfold w%i; " (i-j);
+ pp0 "unfold w%i; " (i-j);
done;
pp "case znz_div_gt.";
pp " intros xx yy H4; repeat rewrite spec_reduce_%i." i;
@@ -1897,7 +1825,7 @@ let _ =
pp " (spec_divn1 w%i w%i_op w%i_spec (S n) x y H3)." i i i;
pp0 " unfold w%i_divn1; " i;
for j = 0 to i do
- pp0 "unfold w%i; " (i-j);
+ pp0 "unfold w%i; " (i-j);
done;
pp "case double_divn1.";
pp " intros xx yy H4.";
@@ -1936,61 +1864,12 @@ let _ =
pp " Qed.";
pr "";
- pr " Definition div_eucl x y :=";
- pr " match compare x y with";
- pr " | Eq => (one, zero)";
- pr " | Lt => (zero, x)";
- pr " | Gt => div_gt x y";
- pr " end.";
- pr "";
-
- pr " Theorem spec_div_eucl: forall x y,";
- pr " 0 < [y] ->";
- pr " let (q,r) := div_eucl x y in";
- pr " ([q], [r]) = Zdiv_eucl [x] [y].";
- pa " Admitted.";
- pp " Proof.";
- pp " assert (F0: [zero] = 0).";
- pp " exact (spec_0 w0_spec).";
- pp " assert (F1: [one] = 1).";
- pp " exact (spec_1 w0_spec).";
- pp " intros x y H; generalize (spec_compare x y);";
- pp " unfold div_eucl; case compare; try rewrite F0;";
- pp " try rewrite F1; intros; auto with zarith.";
- pp " rewrite H0; generalize (Z_div_same [y] (Zlt_gt _ _ H))";
- pp " (Z_mod_same [y] (Zlt_gt _ _ H));";
- pp " unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.";
- pp " assert (F2: 0 <= [x] < [y]).";
- pp " generalize (spec_pos x); auto.";
- pp " generalize (Zdiv_small _ _ F2)";
- pp " (Zmod_small _ _ F2);";
- pp " unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.";
- pp " generalize (spec_div_gt _ _ H0 H); auto.";
- pp " unfold Zdiv, Zmod; case Zdiv_eucl; case div_gt.";
- pp " intros a b c d (H1, H2); subst; auto.";
- pp " Qed.";
- pr "";
-
- pr " Definition div x y := fst (div_eucl x y).";
- pr "";
-
- pr " Theorem spec_div:";
- pr " forall x y, 0 < [y] -> [div x y] = [x] / [y].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros x y H1; unfold div; generalize (spec_div_eucl x y H1);";
- pp " case div_eucl; simpl fst.";
- pp " intros xx yy; unfold Zdiv; case Zdiv_eucl; intros qq rr H; ";
- pp " injection H; auto.";
- pp " Qed.";
- pr "";
-
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Modulo *)";
+ pr " (** * Modulo *)";
pr " (* *)";
pr " (***************************************************************)";
- pr "";
+ pr "";
for i = 0 to size do
pr " Definition w%i_mod_gt := w%i_op.(znz_mod_gt)." i i
@@ -2015,7 +1894,7 @@ let _ =
pr "";
pr " Definition mod_gt := Eval lazy beta delta[iter] in";
- pr " (iter _ ";
+ pr " (iter _";
for i = 0 to size do
pr " (fun x y => reduce_%i (w%i_mod_gt x y))" i i;
pr " (fun n x y => reduce_%i (w%i_mod_gt x (DoubleBase.get_low %s (S n) y)))" i i (pz i);
@@ -2024,16 +1903,16 @@ let _ =
pr " mod_gtnm).";
pr "";
- pp " Let spec_modn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) := ";
- pp " (spec_double_modn1 ";
+ pp " Let spec_modn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) :=";
+ pp " (spec_double_modn1";
pp " ww_op.(znz_zdigits) ww_op.(znz_0)";
pp " (znz_WW ww_op) ww_op.(znz_head0)";
pp " ww_op.(znz_add_mul_div) ww_op.(znz_div21)";
pp " ww_op.(znz_compare) ww_op.(znz_sub) (znz_to_Z ww_op)";
- pp " (spec_to_Z ww_spec) ";
+ pp " (spec_to_Z ww_spec)";
pp " (spec_zdigits ww_spec)";
pp " (spec_0 ww_spec) (spec_WW ww_spec) (spec_head0 ww_spec)";
- pp " (spec_add_mul_div ww_spec) (spec_div21 ww_spec) ";
+ pp " (spec_add_mul_div ww_spec) (spec_div21 ww_spec)";
pp " (CyclicAxioms.spec_compare ww_spec) (CyclicAxioms.spec_sub ww_spec)).";
pp "";
@@ -2063,7 +1942,7 @@ let _ =
pp " rewrite <- (spec_get_end%i (S n) y x) in H3; auto with zarith." i;
if i == size then
pp " intros n x y H2 H3; rewrite spec_reduce_%i." i
- else
+ else
pp " intros n x y H1 H2 H3; rewrite spec_reduce_%i." i;
pp " unfold w%i_modn1, to_Z; rewrite spec_double_eval%in." i i;
pp " apply (spec_modn1 _ _ w%i_spec); auto." i;
@@ -2079,39 +1958,9 @@ let _ =
pp " Qed.";
pr "";
- pr " Definition modulo x y := ";
- pr " match compare x y with";
- pr " | Eq => zero";
- pr " | Lt => x";
- pr " | Gt => mod_gt x y";
- pr " end.";
+ pr " (** digits: a measure for gcd *)";
pr "";
- pr " Theorem spec_modulo:";
- pr " forall x y, 0 < [y] -> [modulo x y] = [x] mod [y].";
- pa " Admitted.";
- pp " Proof.";
- pp " assert (F0: [zero] = 0).";
- pp " exact (spec_0 w0_spec).";
- pp " assert (F1: [one] = 1).";
- pp " exact (spec_1 w0_spec).";
- pp " intros x y H; generalize (spec_compare x y);";
- pp " unfold modulo; case compare; try rewrite F0;";
- pp " try rewrite F1; intros; try split; auto with zarith.";
- pp " rewrite H0; apply sym_equal; apply Z_mod_same; auto with zarith.";
- pp " apply sym_equal; apply Zmod_small; auto with zarith.";
- pp " generalize (spec_pos x); auto with zarith.";
- pp " apply spec_mod_gt; auto.";
- pp " Qed.";
- pr "";
-
- pr " (***************************************************************)";
- pr " (* *)";
- pr " (* Gcd *)";
- pr " (* *)";
- pr " (***************************************************************)";
- pr "";
-
pr " Definition digits x :=";
pr " match x with";
for i = 0 to size do
@@ -2134,189 +1983,18 @@ let _ =
pp " Qed.";
pr "";
- pr " Definition gcd_gt_body a b cont :=";
- pr " match compare b zero with";
- pr " | Gt =>";
- pr " let r := mod_gt a b in";
- pr " match compare r zero with";
- pr " | Gt => cont r (mod_gt b r)";
- pr " | _ => b";
- pr " end";
- pr " | _ => a";
- pr " end.";
- pr "";
-
- pp " Theorem Zspec_gcd_gt_body: forall a b cont p,";
- pp " [a] > [b] -> [a] < 2 ^ p ->";
- pp " (forall a1 b1, [a1] < 2 ^ (p - 1) -> [a1] > [b1] ->";
- pp " Zis_gcd [a1] [b1] [cont a1 b1]) -> ";
- pp " Zis_gcd [a] [b] [gcd_gt_body a b cont].";
- pp " Proof.";
- pp " assert (F1: [zero] = 0).";
- pp " unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.";
- pp " intros a b cont p H2 H3 H4; unfold gcd_gt_body.";
- pp " generalize (spec_compare b zero); case compare; try rewrite F1.";
- pp " intros HH; rewrite HH; apply Zis_gcd_0.";
- pp " intros HH; absurd (0 <= [b]); auto with zarith.";
- pp " case (spec_digits b); auto with zarith.";
- pp " intros H5; generalize (spec_compare (mod_gt a b) zero); ";
- pp " case compare; try rewrite F1.";
- pp " intros H6; rewrite <- (Zmult_1_r [b]).";
- pp " rewrite (Z_div_mod_eq [a] [b]); auto with zarith.";
- pp " rewrite <- spec_mod_gt; auto with zarith.";
- pp " rewrite H6; rewrite Zplus_0_r.";
- pp " apply Zis_gcd_mult; apply Zis_gcd_1.";
- pp " intros; apply False_ind.";
- pp " case (spec_digits (mod_gt a b)); auto with zarith.";
- pp " intros H6; apply DoubleDiv.Zis_gcd_mod; auto with zarith.";
- pp " apply DoubleDiv.Zis_gcd_mod; auto with zarith.";
- pp " rewrite <- spec_mod_gt; auto with zarith.";
- pp " assert (F2: [b] > [mod_gt a b]).";
- pp " case (Z_mod_lt [a] [b]); auto with zarith.";
- pp " repeat rewrite <- spec_mod_gt; auto with zarith.";
- pp " assert (F3: [mod_gt a b] > [mod_gt b (mod_gt a b)]).";
- pp " case (Z_mod_lt [b] [mod_gt a b]); auto with zarith.";
- pp " rewrite <- spec_mod_gt; auto with zarith.";
- pp " repeat rewrite <- spec_mod_gt; auto with zarith.";
- pp " apply H4; auto with zarith.";
- pp " apply Zmult_lt_reg_r with 2; auto with zarith.";
- pp " apply Zle_lt_trans with ([b] + [mod_gt a b]); auto with zarith.";
- pp " apply Zle_lt_trans with (([a]/[b]) * [b] + [mod_gt a b]); auto with zarith.";
- pp " apply Zplus_le_compat_r.";
- pp " pattern [b] at 1; rewrite <- (Zmult_1_l [b]).";
- pp " apply Zmult_le_compat_r; auto with zarith.";
- pp " case (Zle_lt_or_eq 0 ([a]/[b])); auto with zarith.";
- pp " intros HH; rewrite (Z_div_mod_eq [a] [b]) in H2;";
- pp " try rewrite <- HH in H2; auto with zarith.";
- pp " case (Z_mod_lt [a] [b]); auto with zarith.";
- pp " rewrite Zmult_comm; rewrite spec_mod_gt; auto with zarith.";
- pp " rewrite <- Z_div_mod_eq; auto with zarith.";
- pp " pattern 2 at 2; rewrite <- (Zpower_1_r 2).";
- pp " rewrite <- Zpower_exp; auto with zarith.";
- pp " ring_simplify (p - 1 + 1); auto.";
- pp " case (Zle_lt_or_eq 0 p); auto with zarith.";
- pp " generalize H3; case p; simpl Zpower; auto with zarith.";
- pp " intros HH; generalize H3; rewrite <- HH; simpl Zpower; auto with zarith.";
- pp " Qed.";
- pp "";
-
- pr " Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) {struct p} : t :=";
- pr " gcd_gt_body a b";
- pr " (fun a b =>";
- pr " match p with";
- pr " | xH => cont a b";
- pr " | xO p => gcd_gt_aux p (gcd_gt_aux p cont) a b";
- pr " | xI p => gcd_gt_aux p (gcd_gt_aux p cont) a b";
- pr " end).";
- pr "";
-
- pp " Theorem Zspec_gcd_gt_aux: forall p n a b cont,";
- pp " [a] > [b] -> [a] < 2 ^ (Zpos p + n) ->";
- pp " (forall a1 b1, [a1] < 2 ^ n -> [a1] > [b1] ->";
- pp " Zis_gcd [a1] [b1] [cont a1 b1]) ->";
- pp " Zis_gcd [a] [b] [gcd_gt_aux p cont a b].";
- pp " intros p; elim p; clear p.";
- pp " intros p Hrec n a b cont H2 H3 H4.";
- pp " unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xI p) + n); auto.";
- pp " intros a1 b1 H6 H7.";
- pp " apply Hrec with (Zpos p + n); auto.";
- pp " replace (Zpos p + (Zpos p + n)) with";
- pp " (Zpos (xI p) + n - 1); auto.";
- pp " rewrite Zpos_xI; ring.";
- pp " intros a2 b2 H9 H10.";
- pp " apply Hrec with n; auto.";
- pp " intros p Hrec n a b cont H2 H3 H4.";
- pp " unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xO p) + n); auto.";
- pp " intros a1 b1 H6 H7.";
- pp " apply Hrec with (Zpos p + n - 1); auto.";
- pp " replace (Zpos p + (Zpos p + n - 1)) with";
- pp " (Zpos (xO p) + n - 1); auto.";
- pp " rewrite Zpos_xO; ring.";
- pp " intros a2 b2 H9 H10.";
- pp " apply Hrec with (n - 1); auto.";
- pp " replace (Zpos p + (n - 1)) with";
- pp " (Zpos p + n - 1); auto with zarith.";
- pp " intros a3 b3 H12 H13; apply H4; auto with zarith.";
- pp " apply Zlt_le_trans with (1 := H12).";
- pp " case (Zle_or_lt 1 n); intros HH.";
- pp " apply Zpower_le_monotone; auto with zarith.";
- pp " apply Zle_trans with 0; auto with zarith.";
- pp " assert (HH1: n - 1 < 0); auto with zarith.";
- pp " generalize HH1; case (n - 1); auto with zarith.";
- pp " intros p1 HH2; discriminate.";
- pp " intros n a b cont H H2 H3.";
- pp " simpl gcd_gt_aux.";
- pp " apply Zspec_gcd_gt_body with (n + 1); auto with zarith.";
- pp " rewrite Zplus_comm; auto.";
- pp " intros a1 b1 H5 H6; apply H3; auto.";
- pp " replace n with (n + 1 - 1); auto; try ring.";
- pp " Qed.";
- pp "";
-
- pr " Definition gcd_cont a b :=";
- pr " match compare one b with";
- pr " | Eq => one";
- pr " | _ => a";
- pr " end.";
- pr "";
-
- pr " Definition gcd_gt a b := gcd_gt_aux (digits a) gcd_cont a b.";
- pr "";
-
- pr " Theorem spec_gcd_gt: forall a b,";
- pr " [a] > [b] -> [gcd_gt a b] = Zgcd [a] [b].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros a b H2.";
- pp " case (spec_digits (gcd_gt a b)); intros H3 H4.";
- pp " case (spec_digits a); intros H5 H6.";
- pp " apply sym_equal; apply Zis_gcd_gcd; auto with zarith.";
- pp " unfold gcd_gt; apply Zspec_gcd_gt_aux with 0; auto with zarith.";
- pp " intros a1 a2; rewrite Zpower_0_r.";
- pp " case (spec_digits a2); intros H7 H8;";
- pp " intros; apply False_ind; auto with zarith.";
- pp " Qed.";
- pr "";
-
- pr " Definition gcd a b :=";
- pr " match compare a b with";
- pr " | Eq => a";
- pr " | Lt => gcd_gt b a";
- pr " | Gt => gcd_gt a b";
- pr " end.";
- pr "";
-
- pr " Theorem spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros a b.";
- pp " case (spec_digits a); intros H1 H2.";
- pp " case (spec_digits b); intros H3 H4.";
- pp " unfold gcd; generalize (spec_compare a b); case compare.";
- pp " intros HH; rewrite HH; apply sym_equal; apply Zis_gcd_gcd; auto.";
- pp " apply Zis_gcd_refl.";
- pp " intros; apply trans_equal with (Zgcd [b] [a]).";
- pp " apply spec_gcd_gt; auto with zarith.";
- pp " apply Zis_gcd_gcd; auto with zarith.";
- pp " apply Zgcd_is_pos.";
- pp " apply Zis_gcd_sym; apply Zgcd_is_gcd.";
- pp " intros; apply spec_gcd_gt; auto.";
- pp " Qed.";
- pr "";
-
-
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Conversion *)";
+ pr " (** * Conversion *)";
pr " (* *)";
pr " (***************************************************************)";
pr "";
- pr " Definition pheight p := ";
+ pr " Definition pheight p :=";
pr " Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (plength p))).";
pr "";
- pr " Theorem pheight_correct: forall p, ";
+ pr " Theorem pheight_correct: forall p,";
pr " Zpos p < 2 ^ (Zpos (znz_digits w0_op) * 2 ^ (Z_of_nat (pheight p))).";
pr " Proof.";
pr " intros p; unfold pheight.";
@@ -2400,30 +2078,12 @@ let _ =
pp " Qed.";
pr "";
- pr " Definition of_N x :=";
- pr " match x with";
- pr " | BinNat.N0 => zero";
- pr " | Npos p => of_pos p";
- pr " end.";
- pr "";
-
- pr " Theorem spec_of_N: forall x,";
- pr " [of_N x] = Z_of_N x.";
- pa " Admitted.";
- pp " Proof.";
- pp " intros x; case x.";
- pp " simpl of_N.";
- pp " unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.";
- pp " intros p; exact (spec_of_pos p).";
- pp " Qed.";
- pr "";
-
pr " (***************************************************************)";
pr " (* *)";
- pr " (* Shift *)";
+ pr " (** * Shift *)";
pr " (* *)";
pr " (***************************************************************)";
- pr "";
+ pr "";
(* Head0 *)
pr " Definition head0 w := match w with";
@@ -2443,21 +2103,21 @@ let _ =
done;
pp " intros n x; rewrite spec_reduce_n; exact (spec_head00 (wn_spec n) x).";
pp " Qed.";
- pr " ";
+ pr "";
pr " Theorem spec_head0: forall x, 0 < [x] ->";
pr " 2 ^ (Zpos (digits x) - 1) <= 2 ^ [head0 x] * [x] < 2 ^ Zpos (digits x).";
pa " Admitted.";
pp " Proof.";
pp " assert (F0: forall x, (x - 1) + 1 = x).";
- pp " intros; ring. ";
+ pp " intros; ring.";
pp " intros x; case x; unfold digits, head0; clear x.";
for i = 0 to size do
pp " intros x Hx; rewrite spec_reduce_%i." i;
pp " assert (F1:= spec_more_than_1_digit w%i_spec)." i;
pp " generalize (spec_head0 w%i_spec x Hx)." i;
pp " unfold base.";
- pp " pattern (Zpos (znz_digits w%i_op)) at 1; " i;
+ pp " pattern (Zpos (znz_digits w%i_op)) at 1;" i;
pp " rewrite <- (fun x => (F0 (Zpos x))).";
pp " rewrite Zpower_exp; auto with zarith.";
pp " rewrite Zpower_1_r; rewrite Z_div_mult; auto with zarith.";
@@ -2466,7 +2126,7 @@ let _ =
pp " assert (F1:= spec_more_than_1_digit (wn_spec n)).";
pp " generalize (spec_head0 (wn_spec n) x Hx).";
pp " unfold base.";
- pp " pattern (Zpos (znz_digits (make_op n))) at 1; ";
+ pp " pattern (Zpos (znz_digits (make_op n))) at 1;";
pp " rewrite <- (fun x => (F0 (Zpos x))).";
pp " rewrite Zpower_exp; auto with zarith.";
pp " rewrite Zpower_1_r; rewrite Z_div_mult; auto with zarith.";
@@ -2493,7 +2153,7 @@ let _ =
done;
pp " intros n x; rewrite spec_reduce_n; exact (spec_tail00 (wn_spec n) x).";
pp " Qed.";
- pr " ";
+ pr "";
pr " Theorem spec_tail0: forall x,";
@@ -2513,7 +2173,7 @@ let _ =
pr " Definition %sdigits x :=" c;
pr " match x with";
pr " | %s0 _ => %s0 w0_op.(znz_zdigits)" c c;
- for i = 1 to size do
+ for i = 1 to size do
pr " | %s%i _ => reduce_%i w%i_op.(znz_zdigits)" c i i i;
done;
pr " | %sn n _ => reduce_n n (make_op n).(znz_zdigits)" c;
@@ -2534,22 +2194,22 @@ let _ =
(* Shiftr *)
for i = 0 to size do
- pr " Definition shiftr%i n x := w%i_op.(znz_add_mul_div) (w%i_op.(znz_sub) w%i_op.(znz_zdigits) n) w%i_op.(znz_0) x." i i i i i;
+ pr " Definition unsafe_shiftr%i n x := w%i_op.(znz_add_mul_div) (w%i_op.(znz_sub) w%i_op.(znz_zdigits) n) w%i_op.(znz_0) x." i i i i i;
done;
- pr " Definition shiftrn n p x := (make_op n).(znz_add_mul_div) ((make_op n).(znz_sub) (make_op n).(znz_zdigits) p) (make_op n).(znz_0) x.";
+ pr " Definition unsafe_shiftrn n p x := (make_op n).(znz_add_mul_div) ((make_op n).(znz_sub) (make_op n).(znz_zdigits) p) (make_op n).(znz_0) x.";
pr "";
- pr " Definition shiftr := Eval lazy beta delta [same_level] in ";
- pr " same_level _ (fun n x => %s0 (shiftr0 n x))" c;
+ pr " Definition unsafe_shiftr := Eval lazy beta delta [same_level] in";
+ pr " same_level _ (fun n x => %s0 (unsafe_shiftr0 n x))" c;
for i = 1 to size do
- pr " (fun n x => reduce_%i (shiftr%i n x))" i i;
+ pr " (fun n x => reduce_%i (unsafe_shiftr%i n x))" i i;
done;
- pr " (fun n p x => reduce_n n (shiftrn n p x)).";
+ pr " (fun n p x => reduce_n n (unsafe_shiftrn n p x)).";
pr "";
- pr " Theorem spec_shiftr: forall n x,";
- pr " [n] <= [Ndigits x] -> [shiftr n x] = [x] / 2 ^ [n].";
+ pr " Theorem spec_unsafe_shiftr: forall n x,";
+ pr " [n] <= [Ndigits x] -> [unsafe_shiftr n x] = [x] / 2 ^ [n].";
pa " Admitted.";
pp " Proof.";
pp " assert (F0: forall x y, x - (x - y) = y).";
@@ -2568,7 +2228,7 @@ let _ =
pp " split; auto with zarith.";
pp " apply Zle_lt_trans with xx; auto with zarith.";
pp " apply Zpower2_lt_lin; auto with zarith.";
- pp " assert (F4: forall ww ww1 ww2 ";
+ pp " assert (F4: forall ww ww1 ww2";
pp " (ww_op: znz_op ww) (ww1_op: znz_op ww1) (ww2_op: znz_op ww2)";
pp " xx yy xx1 yy1,";
pp " znz_to_Z ww2_op yy <= znz_to_Z ww1_op (znz_zdigits ww1_op) ->";
@@ -2586,7 +2246,7 @@ let _ =
pp " rewrite <- Hy.";
pp " generalize (spec_add_mul_div Hw";
pp " (znz_0 ww_op) xx1";
- pp " (znz_sub ww_op (znz_zdigits ww_op) ";
+ pp " (znz_sub ww_op (znz_zdigits ww_op)";
pp " yy1)";
pp " ).";
pp " rewrite (spec_0 Hw).";
@@ -2612,11 +2272,11 @@ let _ =
pp " rewrite Zpos_xO.";
pp " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith." size;
pp " apply F5; auto with arith.";
- pp " intros x; case x; clear x; unfold shiftr, same_level.";
+ pp " intros x; case x; clear x; unfold unsafe_shiftr, same_level.";
for i = 0 to size do
pp " intros x y; case y; clear y.";
for j = 0 to i - 1 do
- pp " intros y; unfold shiftr%i, Ndigits." i;
+ pp " intros y; unfold unsafe_shiftr%i, Ndigits." i;
pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i j i;
pp " rewrite (spec_zdigits w%i_spec)." i;
@@ -2628,25 +2288,25 @@ let _ =
pp " try (apply sym_equal; exact (spec_extend%in%i y))." j i;
done;
- pp " intros y; unfold shiftr%i, Ndigits." i;
+ pp " intros y; unfold unsafe_shiftr%i, Ndigits." i;
pp " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i i i;
for j = i + 1 to size do
- pp " intros y; unfold shiftr%i, Ndigits." j;
+ pp " intros y; unfold unsafe_shiftr%i, Ndigits." j;
pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." j j i;
pp " try (apply sym_equal; exact (spec_extend%in%i x))." i j;
done;
if i == size then
begin
- pp " intros m y; unfold shiftrn, Ndigits.";
+ pp " intros m y; unfold unsafe_shiftrn, Ndigits.";
pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." size;
pp " try (apply sym_equal; exact (spec_extend%in m x))." size;
end
- else
+ else
begin
- pp " intros m y; unfold shiftrn, Ndigits.";
+ pp " intros m y; unfold unsafe_shiftrn, Ndigits.";
pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." i;
pp " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x])." size i (size - i - 1) i;
@@ -2654,7 +2314,7 @@ let _ =
end
done;
pp " intros n x y; case y; clear y;";
- pp " intros y; unfold shiftrn, Ndigits; try rewrite spec_reduce_n.";
+ pp " intros y; unfold unsafe_shiftrn, Ndigits; try rewrite spec_reduce_n.";
for i = 0 to size do
pp " try rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
pp " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith." i;
@@ -2684,52 +2344,23 @@ let _ =
pp " Qed.";
pr "";
- pr " Definition safe_shiftr n x := ";
- pr " match compare n (Ndigits x) with";
- pr " | Lt => shiftr n x ";
- pr " | _ => %s0 w_0" c;
- pr " end.";
- pr "";
-
-
- pr " Theorem spec_safe_shiftr: forall n x,";
- pr " [safe_shiftr n x] = [x] / 2 ^ [n].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros n x; unfold safe_shiftr;";
- pp " generalize (spec_compare n (Ndigits x)); case compare; intros H.";
- pp " apply trans_equal with (1 := spec_0 w0_spec).";
- pp " apply sym_equal; apply Zdiv_small; rewrite H.";
- pp " rewrite spec_Ndigits; exact (spec_digits x).";
- pp " rewrite <- spec_shiftr; auto with zarith.";
- pp " apply trans_equal with (1 := spec_0 w0_spec).";
- pp " apply sym_equal; apply Zdiv_small.";
- pp " rewrite spec_Ndigits in H; case (spec_digits x); intros H1 H2.";
- pp " split; auto.";
- pp " apply Zlt_le_trans with (1 := H2).";
- pp " apply Zpower_le_monotone; auto with zarith.";
- pp " Qed.";
- pr "";
-
- pr "";
-
- (* Shiftl *)
+ (* Unsafe_Shiftl *)
for i = 0 to size do
- pr " Definition shiftl%i n x := w%i_op.(znz_add_mul_div) n x w%i_op.(znz_0)." i i i
+ pr " Definition unsafe_shiftl%i n x := w%i_op.(znz_add_mul_div) n x w%i_op.(znz_0)." i i i
done;
- pr " Definition shiftln n p x := (make_op n).(znz_add_mul_div) p x (make_op n).(znz_0).";
- pr " Definition shiftl := Eval lazy beta delta [same_level] in";
- pr " same_level _ (fun n x => %s0 (shiftl0 n x))" c;
+ pr " Definition unsafe_shiftln n p x := (make_op n).(znz_add_mul_div) p x (make_op n).(znz_0).";
+ pr " Definition unsafe_shiftl := Eval lazy beta delta [same_level] in";
+ pr " same_level _ (fun n x => %s0 (unsafe_shiftl0 n x))" c;
for i = 1 to size do
- pr " (fun n x => reduce_%i (shiftl%i n x))" i i;
+ pr " (fun n x => reduce_%i (unsafe_shiftl%i n x))" i i;
done;
- pr " (fun n p x => reduce_n n (shiftln n p x)).";
+ pr " (fun n p x => reduce_n n (unsafe_shiftln n p x)).";
pr "";
pr "";
- pr " Theorem spec_shiftl: forall n x,";
- pr " [n] <= [head0 x] -> [shiftl n x] = [x] * 2 ^ [n].";
+ pr " Theorem spec_unsafe_shiftl: forall n x,";
+ pr " [n] <= [head0 x] -> [unsafe_shiftl n x] = [x] * 2 ^ [n].";
pa " Admitted.";
pp " Proof.";
pp " assert (F0: forall x y, x - (x - y) = y).";
@@ -2748,7 +2379,7 @@ let _ =
pp " split; auto with zarith.";
pp " apply Zle_lt_trans with xx; auto with zarith.";
pp " apply Zpower2_lt_lin; auto with zarith.";
- pp " assert (F4: forall ww ww1 ww2 ";
+ pp " assert (F4: forall ww ww1 ww2";
pp " (ww_op: znz_op ww) (ww1_op: znz_op ww1) (ww2_op: znz_op ww2)";
pp " xx yy xx1 yy1,";
pp " znz_to_Z ww2_op yy <= znz_to_Z ww1_op (znz_head0 ww1_op xx) ->";
@@ -2788,7 +2419,7 @@ let _ =
pp " rewrite Zmod_small; auto with zarith.";
pp " intros HH; apply HH.";
pp " rewrite Hy; apply Zle_trans with (1:= Hl).";
- pp " rewrite <- (spec_zdigits Hw). ";
+ pp " rewrite <- (spec_zdigits Hw).";
pp " apply Zle_trans with (2 := Hl1); auto.";
pp " rewrite (spec_zdigits Hw1); auto with zarith.";
pp " split; auto with zarith .";
@@ -2826,11 +2457,11 @@ let _ =
pp " rewrite Zpos_xO.";
pp " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith." size;
pp " apply F5; auto with arith.";
- pp " intros x; case x; clear x; unfold shiftl, same_level.";
+ pp " intros x; case x; clear x; unfold unsafe_shiftl, same_level.";
for i = 0 to size do
pp " intros x y; case y; clear y.";
for j = 0 to i - 1 do
- pp " intros y; unfold shiftl%i, head0." i;
+ pp " intros y; unfold unsafe_shiftl%i, head0." i;
pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i j i;
pp " rewrite (spec_zdigits w%i_spec)." i;
@@ -2841,25 +2472,25 @@ let _ =
pp " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith." j;
pp " try (apply sym_equal; exact (spec_extend%in%i y))." j i;
done;
- pp " intros y; unfold shiftl%i, head0." i;
+ pp " intros y; unfold unsafe_shiftl%i, head0." i;
pp " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i i i;
for j = i + 1 to size do
- pp " intros y; unfold shiftl%i, head0." j;
+ pp " intros y; unfold unsafe_shiftl%i, head0." j;
pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." j j i;
pp " try (apply sym_equal; exact (spec_extend%in%i x))." i j;
done;
if i == size then
begin
- pp " intros m y; unfold shiftln, head0.";
+ pp " intros m y; unfold unsafe_shiftln, head0.";
pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." size;
pp " try (apply sym_equal; exact (spec_extend%in m x))." size;
end
- else
+ else
begin
- pp " intros m y; unfold shiftln, head0.";
+ pp " intros m y; unfold unsafe_shiftln, head0.";
pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." i;
pp " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x])." size i (size - i - 1) i;
@@ -2867,7 +2498,7 @@ let _ =
end
done;
pp " intros n x y; case y; clear y;";
- pp " intros y; unfold shiftln, head0; try rewrite spec_reduce_n.";
+ pp " intros y; unfold unsafe_shiftln, head0; try rewrite spec_reduce_n.";
for i = 0 to size do
pp " try rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
pp " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith." i;
@@ -2907,7 +2538,7 @@ let _ =
pr " end.";
pr "";
- pr " Theorem spec_double_size_digits: ";
+ pr " Theorem spec_double_size_digits:";
pr " forall x, digits (double_size x) = xO (digits x).";
pa " Admitted.";
pp " Proof.";
@@ -2922,7 +2553,7 @@ let _ =
pp " Proof.";
pp " intros x; case x; unfold double_size; clear x.";
for i = 0 to size do
- pp " intros x; unfold to_Z, make_op; ";
+ pp " intros x; unfold to_Z, make_op;";
pp " rewrite znz_to_Z_%i; rewrite (spec_0 w%i_spec); auto with zarith." (i + 1) i;
done;
pp " intros n x; unfold to_Z;";
@@ -2934,7 +2565,7 @@ let _ =
pr "";
- pr " Theorem spec_double_size_head0: ";
+ pr " Theorem spec_double_size_head0:";
pr " forall x, 2 * [head0 x] <= [head0 (double_size x)].";
pa " Admitted.";
pp " Proof.";
@@ -2963,7 +2594,7 @@ let _ =
pp " apply Zmult_le_compat_l; auto with zarith.";
pp " rewrite Zpower_1_r; auto with zarith.";
pp " apply Zpower_le_monotone; auto with zarith.";
- pp " split; auto with zarith. ";
+ pp " split; auto with zarith.";
pp " case (Zle_or_lt (Zpos (digits x)) [head0 x]); auto with zarith; intros HH6.";
pp " absurd (2 ^ Zpos (digits x) <= 2 ^ [head0 x] * [x]); auto with zarith.";
pp " rewrite <- HH5; rewrite Zmult_1_r.";
@@ -2988,7 +2619,7 @@ let _ =
pp " Qed.";
pr "";
- pr " Theorem spec_double_size_head0_pos: ";
+ pr " Theorem spec_double_size_head0_pos:";
pr " forall x, 0 < [head0 (double_size x)].";
pa " Admitted.";
pp " Proof.";
@@ -3015,114 +2646,6 @@ let _ =
pp " Qed.";
pr "";
-
- (* Safe shiftl *)
-
- pr " Definition safe_shiftl_aux_body cont n x :=";
- pr " match compare n (head0 x) with";
- pr " Gt => cont n (double_size x)";
- pr " | _ => shiftl n x";
- pr " end.";
- pr "";
-
- pr " Theorem spec_safe_shift_aux_body: forall n p x cont,";
- pr " 2^ Zpos p <= [head0 x] ->";
- pr " (forall x, 2 ^ (Zpos p + 1) <= [head0 x]->";
- pr " [cont n x] = [x] * 2 ^ [n]) ->";
- pr " [safe_shiftl_aux_body cont n x] = [x] * 2 ^ [n].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros n p x cont H1 H2; unfold safe_shiftl_aux_body.";
- pp " generalize (spec_compare n (head0 x)); case compare; intros H.";
- pp " apply spec_shiftl; auto with zarith.";
- pp " apply spec_shiftl; auto with zarith.";
- pp " rewrite H2.";
- pp " rewrite spec_double_size; auto.";
- pp " rewrite Zplus_comm; rewrite Zpower_exp; auto with zarith.";
- pp " apply Zle_trans with (2 := spec_double_size_head0 x).";
- pp " rewrite Zpower_1_r; apply Zmult_le_compat_l; auto with zarith.";
- pp " Qed.";
- pr "";
-
- pr " Fixpoint safe_shiftl_aux p cont n x {struct p} :=";
- pr " safe_shiftl_aux_body ";
- pr " (fun n x => match p with";
- pr " | xH => cont n x";
- pr " | xO p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x";
- pr " | xI p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x";
- pr " end) n x.";
- pr "";
-
- pr " Theorem spec_safe_shift_aux: forall p q n x cont,";
- pr " 2 ^ (Zpos q) <= [head0 x] ->";
- pr " (forall x, 2 ^ (Zpos p + Zpos q) <= [head0 x] ->";
- pr " [cont n x] = [x] * 2 ^ [n]) -> ";
- pr " [safe_shiftl_aux p cont n x] = [x] * 2 ^ [n].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros p; elim p; unfold safe_shiftl_aux; fold safe_shiftl_aux; clear p.";
- pp " intros p Hrec q n x cont H1 H2.";
- pp " apply spec_safe_shift_aux_body with (q); auto.";
- pp " intros x1 H3; apply Hrec with (q + 1)%spositive; auto." "%";
- pp " intros x2 H4; apply Hrec with (p + q + 1)%spositive; auto." "%";
- pp " rewrite <- Pplus_assoc.";
- pp " rewrite Zpos_plus_distr; auto.";
- pp " intros x3 H5; apply H2.";
- pp " rewrite Zpos_xI.";
- pp " replace (2 * Zpos p + 1 + Zpos q) with (Zpos p + Zpos (p + q + 1));";
- pp " auto.";
- pp " repeat rewrite Zpos_plus_distr; ring.";
- pp " intros p Hrec q n x cont H1 H2.";
- pp " apply spec_safe_shift_aux_body with (q); auto.";
- pp " intros x1 H3; apply Hrec with (q); auto.";
- pp " apply Zle_trans with (2 := H3); auto with zarith.";
- pp " apply Zpower_le_monotone; auto with zarith.";
- pp " intros x2 H4; apply Hrec with (p + q)%spositive; auto." "%";
- pp " intros x3 H5; apply H2.";
- pp " rewrite (Zpos_xO p).";
- pp " replace (2 * Zpos p + Zpos q) with (Zpos p + Zpos (p + q));";
- pp " auto.";
- pp " repeat rewrite Zpos_plus_distr; ring.";
- pp " intros q n x cont H1 H2.";
- pp " apply spec_safe_shift_aux_body with (q); auto.";
- pp " rewrite Zplus_comm; auto.";
- pp " Qed.";
- pr "";
-
-
- pr " Definition safe_shiftl n x :=";
- pr " safe_shiftl_aux_body";
- pr " (safe_shiftl_aux_body";
- pr " (safe_shiftl_aux (digits n) shiftl)) n x.";
- pr "";
-
- pr " Theorem spec_safe_shift: forall n x,";
- pr " [safe_shiftl n x] = [x] * 2 ^ [n].";
- pa " Admitted.";
- pp " Proof.";
- pp " intros n x; unfold safe_shiftl, safe_shiftl_aux_body.";
- pp " generalize (spec_compare n (head0 x)); case compare; intros H.";
- pp " apply spec_shiftl; auto with zarith.";
- pp " apply spec_shiftl; auto with zarith.";
- pp " rewrite <- (spec_double_size x).";
- pp " generalize (spec_compare n (head0 (double_size x))); case compare; intros H1.";
- pp " apply spec_shiftl; auto with zarith.";
- pp " apply spec_shiftl; auto with zarith.";
- pp " rewrite <- (spec_double_size (double_size x)).";
- pp " apply spec_safe_shift_aux with 1%spositive." "%";
- pp " apply Zle_trans with (2 := spec_double_size_head0 (double_size x)).";
- pp " replace (2 ^ 1) with (2 * 1).";
- pp " apply Zmult_le_compat_l; auto with zarith.";
- pp " generalize (spec_double_size_head0_pos x); auto with zarith.";
- pp " rewrite Zpower_1_r; ring.";
- pp " intros x1 H2; apply spec_shiftl.";
- pp " apply Zle_trans with (2 := H2).";
- pp " apply Zle_trans with (2 ^ Zpos (digits n)); auto with zarith.";
- pp " case (spec_digits n); auto with zarith.";
- pp " apply Zpower_le_monotone; auto with zarith.";
- pp " Qed.";
- pr "";
-
(* even *)
pr " Definition is_even x :=";
pr " match x with";
@@ -3146,20 +2669,6 @@ let _ =
pp " Qed.";
pr "";
- pr " Theorem spec_0: [zero] = 0.";
- pa " Admitted.";
- pp " Proof.";
- pp " exact (spec_0 w0_spec).";
- pp " Qed.";
- pr "";
-
- pr " Theorem spec_1: [one] = 1.";
- pa " Admitted.";
- pp " Proof.";
- pp " exact (spec_1 w0_spec).";
- pp " Qed.";
- pr "";
-
pr "End Make.";
pr "";
diff --git a/theories/Numbers/Natural/BigN/Nbasic.v b/theories/Numbers/Natural/BigN/Nbasic.v
index ae2cfd30..d42db97d 100644
--- a/theories/Numbers/Natural/BigN/Nbasic.v
+++ b/theories/Numbers/Natural/BigN/Nbasic.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: Nbasic.v 10964 2008-05-22 11:08:13Z letouzey $ i*)
+(*i $Id$ i*)
Require Import ZArith.
Require Import BigNumPrelude.
@@ -21,7 +21,7 @@ Require Import DoubleCyclic.
(* To compute the necessary height *)
Fixpoint plength (p: positive) : positive :=
- match p with
+ match p with
xH => xH
| xO p1 => Psucc (plength p1)
| xI p1 => Psucc (plength p1)
@@ -34,10 +34,10 @@ rewrite Zpower_exp; auto with zarith.
rewrite Zpos_succ_morphism; unfold Zsucc; auto with zarith.
intros p; elim p; simpl plength; auto.
intros p1 Hp1; rewrite F; repeat rewrite Zpos_xI.
-assert (tmp: (forall p, 2 * p = p + p)%Z);
+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).
-assert (tmp: (forall p, 2 * p = p + p)%Z);
+assert (tmp: (forall p, 2 * p = p + p)%Z);
try repeat rewrite tmp; auto with zarith.
rewrite Zpower_1_r; auto with zarith.
Qed.
@@ -73,7 +73,7 @@ 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;
+generalize H2 (Z_mod_lt (Zpos p) (Zpos q) H1); clear H2;
case Zmod.
intros HH _; rewrite HH; auto with zarith.
intros r1 HH (_,HH1); rewrite HH; rewrite Zpos_succ_morphism.
@@ -121,9 +121,9 @@ Definition zn2z_word_comm : forall w n, zn2z (word w n) = word (zn2z w) n.
Defined.
Fixpoint extend (n:nat) {struct n} : forall w:Type, zn2z w -> word w (S n) :=
- match n return forall w:Type, zn2z w -> word w (S n) with
+ match n return forall w:Type, zn2z w -> word w (S n) with
| O => fun w x => x
- | S m =>
+ | S m =>
let aux := extend m in
fun w x => WW W0 (aux w x)
end.
@@ -169,7 +169,7 @@ Fixpoint diff_l (m n : nat) {struct m} : fst (diff m n) + n = max m n :=
| S n1 =>
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 v (fun n => v1 = S n)
(eq_ind v1 (fun n => v1 = n) (refl_equal v1) (S v) (plusnS _ _))
_ (diff_l _ _)
end
@@ -182,7 +182,7 @@ Fixpoint diff_r (m n: nat) {struct m}: snd (diff m n) + m = max m n :=
| 0 => refl_equal _
| S _ => plusn0 _
end
- | S m =>
+ | 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)
| S n1 =>
@@ -253,9 +253,9 @@ Section ReduceRec.
| WW xh xl =>
match xh with
| W0 => @reduce_n m xl
- | _ => @c (S m) x
+ | _ => @c (S m) x
end
- end
+ end
end.
End ReduceRec.
@@ -276,14 +276,14 @@ Section CompareRec.
Variable compare_m : wm -> w -> comparison.
Fixpoint compare0_mn (n:nat) : word wm n -> comparison :=
- match n return word wm n -> comparison with
- | O => compare0_m
+ match n return word wm n -> comparison with
+ | O => compare0_m
| S m => fun x =>
match x with
| W0 => Eq
- | WW xh xl =>
+ | WW xh xl =>
match compare0_mn m xh with
- | Eq => compare0_mn m xl
+ | Eq => compare0_mn m xl
| r => Lt
end
end
@@ -296,7 +296,7 @@ Section CompareRec.
Variable spec_compare0_m: forall x,
match compare0_m x with
Eq => w_to_Z w_0 = wm_to_Z x
- | Lt => w_to_Z w_0 < wm_to_Z x
+ | Lt => w_to_Z w_0 < wm_to_Z x
| Gt => w_to_Z w_0 > wm_to_Z x
end.
Variable wm_to_Z_pos: forall x, 0 <= wm_to_Z x < base wm_base.
@@ -341,14 +341,14 @@ Section CompareRec.
Qed.
Fixpoint compare_mn_1 (n:nat) : word wm n -> w -> comparison :=
- match n return word wm n -> w -> comparison with
- | O => compare_m
- | S m => fun x y =>
+ match n return word wm n -> w -> comparison with
+ | O => compare_m
+ | S m => fun x y =>
match x with
| W0 => compare w_0 y
- | WW xh xl =>
+ | WW xh xl =>
match compare0_mn m xh with
- | Eq => compare_mn_1 m xl y
+ | Eq => compare_mn_1 m xl y
| r => Gt
end
end
@@ -366,7 +366,7 @@ Section CompareRec.
| Lt => wm_to_Z x < w_to_Z y
| Gt => wm_to_Z x > w_to_Z y
end.
- Variable wm_base_lt: forall x,
+ Variable wm_base_lt: forall x,
0 <= w_to_Z x < base (wm_base).
Let double_wB_lt: forall n x,
@@ -385,7 +385,7 @@ Section CompareRec.
unfold Zpower_pos; simpl; ring.
Qed.
-
+
Lemma spec_compare_mn_1: forall n x y,
match compare_mn_1 n x y with
Eq => double_to_Z n x = w_to_Z y
@@ -434,7 +434,7 @@ Section AddS.
| C1 z => match incr hy with
C0 z1 => C0 (WW z1 z)
| C1 z1 => C1 (WW z1 z)
- end
+ end
end
end.
@@ -458,12 +458,12 @@ End AddS.
Fixpoint length_pos x :=
match x with xH => O | xO x1 => S (length_pos x1) | xI x1 => S (length_pos x1) end.
-
+
Theorem length_pos_lt: forall x y,
(length_pos x < length_pos y)%nat -> Zpos x < Zpos y.
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;
+ 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 (inversion H; fail);
@@ -492,20 +492,20 @@ End AddS.
Qed.
Theorem make_zop: forall w (x: znz_op w),
- znz_to_Z (mk_zn2z_op x) =
- fun z => match z with
+ znz_to_Z (mk_zn2z_op x) =
+ fun z => match z with
W0 => 0
- | WW xh xl => znz_to_Z x xh * base (znz_digits x)
+ | WW xh xl => znz_to_Z x xh * base (znz_digits x)
+ znz_to_Z x xl
end.
intros ww x; auto.
Qed.
Theorem make_kzop: forall w (x: znz_op w),
- znz_to_Z (mk_zn2z_op_karatsuba x) =
- fun z => match z with
+ znz_to_Z (mk_zn2z_op_karatsuba x) =
+ fun z => match z with
W0 => 0
- | WW xh xl => znz_to_Z x xh * base (znz_digits x)
+ | WW xh xl => znz_to_Z x xh * base (znz_digits x)
+ znz_to_Z x xl
end.
intros ww x; auto.
diff --git a/theories/Numbers/Natural/Binary/NBinDefs.v b/theories/Numbers/Natural/Binary/NBinDefs.v
deleted file mode 100644
index fc2bd2df..00000000
--- a/theories/Numbers/Natural/Binary/NBinDefs.v
+++ /dev/null
@@ -1,267 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-(* Evgeny Makarov, INRIA, 2007 *)
-(************************************************************************)
-
-(*i $Id: NBinDefs.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
-
-Require Import BinPos.
-Require Export BinNat.
-Require Import NSub.
-
-Open Local Scope N_scope.
-
-(** Implementation of [NAxiomsSig] module type via [BinNat.N] *)
-
-Module NBinaryAxiomsMod <: NAxiomsSig.
-Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig.
-Module Export NZAxiomsMod <: NZAxiomsSig.
-
-Definition NZ := N.
-Definition NZeq := @eq N.
-Definition NZ0 := N0.
-Definition NZsucc := Nsucc.
-Definition NZpred := Npred.
-Definition NZadd := Nplus.
-Definition NZsub := Nminus.
-Definition NZmul := Nmult.
-
-Theorem NZeq_equiv : equiv N NZeq.
-Proof (eq_equiv N).
-
-Add Relation N NZeq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
-as NZeq_rel.
-
-Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd.
-Proof.
-congruence.
-Qed.
-
-Theorem NZinduction :
- forall A : NZ -> Prop, predicate_wd NZeq A ->
- A N0 -> (forall n, A n <-> A (NZsucc n)) -> forall n : NZ, A n.
-Proof.
-intros A A_wd A0 AS. apply Nrect. assumption. intros; now apply -> AS.
-Qed.
-
-Theorem NZpred_succ : forall n : NZ, NZpred (NZsucc n) = n.
-Proof.
-destruct n as [| p]; simpl. reflexivity.
-case_eq (Psucc p); try (intros q H; rewrite <- H; now rewrite Ppred_succ).
-intro H; false_hyp H Psucc_not_one.
-Qed.
-
-Theorem NZadd_0_l : forall n : NZ, N0 + n = n.
-Proof.
-reflexivity.
-Qed.
-
-Theorem NZadd_succ_l : forall n m : NZ, (NZsucc n) + m = NZsucc (n + m).
-Proof.
-destruct n; destruct m.
-simpl in |- *; reflexivity.
-unfold NZsucc, NZadd, Nsucc, Nplus. rewrite <- Pplus_one_succ_l; reflexivity.
-simpl in |- *; reflexivity.
-simpl in |- *; rewrite Pplus_succ_permute_l; reflexivity.
-Qed.
-
-Theorem NZsub_0_r : forall n : NZ, n - N0 = n.
-Proof.
-now destruct n.
-Qed.
-
-Theorem NZsub_succ_r : forall n m : NZ, n - (NZsucc m) = NZpred (n - m).
-Proof.
-destruct n as [| p]; destruct m as [| q]; try reflexivity.
-now destruct p.
-simpl. rewrite Pminus_mask_succ_r, Pminus_mask_carry_spec.
-now destruct (Pminus_mask p q) as [| r |]; [| destruct r |].
-Qed.
-
-Theorem NZmul_0_l : forall n : NZ, N0 * n = N0.
-Proof.
-destruct n; reflexivity.
-Qed.
-
-Theorem NZmul_succ_l : forall n m : NZ, (NZsucc n) * m = n * m + m.
-Proof.
-destruct n as [| n]; destruct m as [| m]; simpl; try reflexivity.
-now rewrite Pmult_Sn_m, Pplus_comm.
-Qed.
-
-End NZAxiomsMod.
-
-Definition NZlt := Nlt.
-Definition NZle := Nle.
-Definition NZmin := Nmin.
-Definition NZmax := Nmax.
-
-Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd.
-Proof.
-unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2.
-Qed.
-
-Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd.
-Proof.
-unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2.
-Qed.
-
-Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd.
-Proof.
-congruence.
-Qed.
-
-Theorem NZlt_eq_cases : forall n m : N, n <= m <-> n < m \/ n = m.
-Proof.
-intros n m. unfold Nle, Nlt. rewrite <- Ncompare_eq_correct.
-destruct (n ?= m); split; intro H1; (try discriminate); try (now left); try now right.
-now elim H1. destruct H1; discriminate.
-Qed.
-
-Theorem NZlt_irrefl : forall n : NZ, ~ n < n.
-Proof.
-intro n; unfold Nlt; now rewrite Ncompare_refl.
-Qed.
-
-Theorem NZlt_succ_r : forall n m : NZ, n < (NZsucc m) <-> n <= m.
-Proof.
-intros n m; unfold Nlt, Nle; destruct n as [| p]; destruct m as [| q]; simpl;
-split; intro H; try reflexivity; try discriminate.
-destruct p; simpl; intros; discriminate. elimtype False; now apply H.
-apply -> Pcompare_p_Sq in H. destruct H as [H | H].
-now rewrite H. now rewrite H, Pcompare_refl.
-apply <- Pcompare_p_Sq. case_eq ((p ?= q)%positive Eq); intro H1.
-right; now apply Pcompare_Eq_eq. now left. elimtype False; now apply H.
-Qed.
-
-Theorem NZmin_l : forall n m : N, n <= m -> NZmin n m = n.
-Proof.
-unfold NZmin, Nmin, Nle; intros n m H.
-destruct (n ?= m); try reflexivity. now elim H.
-Qed.
-
-Theorem NZmin_r : forall n m : N, m <= n -> NZmin n m = m.
-Proof.
-unfold NZmin, Nmin, Nle; intros n m H.
-case_eq (n ?= m); intro H1; try reflexivity.
-now apply -> Ncompare_eq_correct.
-rewrite <- Ncompare_antisym, H1 in H; elim H; auto.
-Qed.
-
-Theorem NZmax_l : forall n m : N, m <= n -> NZmax n m = n.
-Proof.
-unfold NZmax, Nmax, Nle; intros n m H.
-case_eq (n ?= m); intro H1; try reflexivity.
-symmetry; now apply -> Ncompare_eq_correct.
-rewrite <- Ncompare_antisym, H1 in H; elim H; auto.
-Qed.
-
-Theorem NZmax_r : forall n m : N, n <= m -> NZmax n m = m.
-Proof.
-unfold NZmax, Nmax, Nle; intros n m H.
-destruct (n ?= m); try reflexivity. now elim H.
-Qed.
-
-End NZOrdAxiomsMod.
-
-Definition recursion (A : Type) (a : A) (f : N -> A -> A) (n : N) :=
- Nrect (fun _ => A) a f n.
-Implicit Arguments recursion [A].
-
-Theorem pred_0 : Npred N0 = N0.
-Proof.
-reflexivity.
-Qed.
-
-Theorem recursion_wd :
-forall (A : Type) (Aeq : relation A),
- forall a a' : A, Aeq a a' ->
- forall f f' : N -> A -> A, fun2_eq NZeq Aeq Aeq f f' ->
- forall x x' : N, x = x' ->
- Aeq (recursion a f x) (recursion a' f' x').
-Proof.
-unfold fun2_wd, NZeq, fun2_eq.
-intros A Aeq a a' Eaa' f f' Eff'.
-intro x; pattern x; apply Nrect.
-intros x' H; now rewrite <- H.
-clear x.
-intros x IH x' H; rewrite <- H.
-unfold recursion in *. do 2 rewrite Nrect_step.
-now apply Eff'; [| apply IH].
-Qed.
-
-Theorem recursion_0 :
- forall (A : Type) (a : A) (f : N -> A -> A), recursion a f N0 = a.
-Proof.
-intros A a f; unfold recursion; now rewrite Nrect_base.
-Qed.
-
-Theorem recursion_succ :
- forall (A : Type) (Aeq : relation A) (a : A) (f : N -> A -> A),
- Aeq a a -> fun2_wd NZeq Aeq Aeq f ->
- forall n : N, Aeq (recursion a f (Nsucc n)) (f n (recursion a f n)).
-Proof.
-unfold NZeq, recursion, fun2_wd; intros A Aeq a f EAaa f_wd n; pattern n; apply Nrect.
-rewrite Nrect_step; rewrite Nrect_base; now apply f_wd.
-clear n; intro n; do 2 rewrite Nrect_step; intro IH. apply f_wd; [reflexivity|].
-now rewrite Nrect_step.
-Qed.
-
-End NBinaryAxiomsMod.
-
-Module Export NBinarySubPropMod := NSubPropFunct NBinaryAxiomsMod.
-
-(* Some fun comparing the efficiency of the generic log defined
-by strong (course-of-value) recursion and the log defined by recursion
-on notation *)
-(* Time Eval compute in (log 100000). *) (* 98 sec *)
-
-(*
-Fixpoint binposlog (p : positive) : N :=
-match p with
-| xH => 0
-| xO p' => Nsucc (binposlog p')
-| xI p' => Nsucc (binposlog p')
-end.
-
-Definition binlog (n : N) : N :=
-match n with
-| 0 => 0
-| Npos p => binposlog p
-end.
-*)
-(* Eval compute in (binlog 1000000000000000000). *) (* Works very fast *)
-
diff --git a/theories/Numbers/Natural/Binary/NBinary.v b/theories/Numbers/Natural/Binary/NBinary.v
index 2c99128d..e593f4a5 100644
--- a/theories/Numbers/Natural/Binary/NBinary.v
+++ b/theories/Numbers/Natural/Binary/NBinary.v
@@ -8,8 +8,175 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NBinary.v 10934 2008-05-15 21:58:20Z letouzey $ i*)
+(*i $Id$ i*)
-Require Export NBinDefs.
-Require Export NArithRing.
+Require Import BinPos.
+Require Export BinNat.
+Require Import NAxioms NProperties.
+Local Open Scope N_scope.
+
+(** * Implementation of [NAxiomsSig] module type via [BinNat.N] *)
+
+Module NBinaryAxiomsMod <: NAxiomsSig.
+
+(** Bi-directional induction. *)
+
+Theorem bi_induction :
+ forall A : N -> Prop, Proper (eq==>iff) A ->
+ A N0 -> (forall n, A n <-> A (Nsucc n)) -> forall n : N, A n.
+Proof.
+intros A A_wd A0 AS. apply Nrect. assumption. intros; now apply -> AS.
+Qed.
+
+(** Basic operations. *)
+
+Definition eq_equiv : Equivalence (@eq N) := eq_equivalence.
+Local Obligation Tactic := simpl_relation.
+Program Instance succ_wd : Proper (eq==>eq) Nsucc.
+Program Instance pred_wd : Proper (eq==>eq) Npred.
+Program Instance add_wd : Proper (eq==>eq==>eq) Nplus.
+Program Instance sub_wd : Proper (eq==>eq==>eq) Nminus.
+Program Instance mul_wd : Proper (eq==>eq==>eq) Nmult.
+
+Definition pred_succ := Npred_succ.
+Definition add_0_l := Nplus_0_l.
+Definition add_succ_l := Nplus_succ.
+Definition sub_0_r := Nminus_0_r.
+Definition sub_succ_r := Nminus_succ_r.
+Definition mul_0_l := Nmult_0_l.
+Definition mul_succ_l n m := eq_trans (Nmult_Sn_m n m) (Nplus_comm _ _).
+
+(** Order *)
+
+Program Instance lt_wd : Proper (eq==>eq==>iff) Nlt.
+
+Definition lt_eq_cases := Nle_lteq.
+Definition lt_irrefl := Nlt_irrefl.
+
+Theorem lt_succ_r : forall n m, n < (Nsucc m) <-> n <= m.
+Proof.
+intros n m; unfold Nlt, Nle; destruct n as [| p]; destruct m as [| q]; simpl;
+split; intro H; try reflexivity; try discriminate.
+destruct p; simpl; intros; discriminate. exfalso; now apply H.
+apply -> Pcompare_p_Sq in H. destruct H as [H | H].
+now rewrite H. now rewrite H, Pcompare_refl.
+apply <- Pcompare_p_Sq. case_eq ((p ?= q)%positive Eq); intro H1.
+right; now apply Pcompare_Eq_eq. now left. exfalso; now apply H.
+Qed.
+
+Theorem min_l : forall n m, n <= m -> Nmin n m = n.
+Proof.
+unfold Nmin, Nle; intros n m H.
+destruct (n ?= m); try reflexivity. now elim H.
+Qed.
+
+Theorem min_r : forall n m, m <= n -> Nmin n m = m.
+Proof.
+unfold Nmin, Nle; intros n m H.
+case_eq (n ?= m); intro H1; try reflexivity.
+now apply -> Ncompare_eq_correct.
+rewrite <- Ncompare_antisym, H1 in H; elim H; auto.
+Qed.
+
+Theorem max_l : forall n m, m <= n -> Nmax n m = n.
+Proof.
+unfold Nmax, Nle; intros n m H.
+case_eq (n ?= m); intro H1; try reflexivity.
+symmetry; now apply -> Ncompare_eq_correct.
+rewrite <- Ncompare_antisym, H1 in H; elim H; auto.
+Qed.
+
+Theorem max_r : forall n m : N, n <= m -> Nmax n m = m.
+Proof.
+unfold Nmax, Nle; intros n m H.
+destruct (n ?= m); try reflexivity. now elim H.
+Qed.
+
+(** Part specific to natural numbers, not integers. *)
+
+Theorem pred_0 : Npred 0 = 0.
+Proof.
+reflexivity.
+Qed.
+
+Definition recursion (A : Type) : A -> (N -> A -> A) -> N -> A :=
+ Nrect (fun _ => A).
+Implicit Arguments recursion [A].
+
+Instance recursion_wd A (Aeq : relation A) :
+ Proper (Aeq==>(eq==>Aeq==>Aeq)==>eq==>Aeq) (@recursion A).
+Proof.
+intros a a' Eaa' f f' Eff'.
+intro x; pattern x; apply Nrect.
+intros x' H; now rewrite <- H.
+clear x.
+intros x IH x' H; rewrite <- H.
+unfold recursion in *. do 2 rewrite Nrect_step.
+now apply Eff'; [| apply IH].
+Qed.
+
+Theorem recursion_0 :
+ forall (A : Type) (a : A) (f : N -> A -> A), recursion a f N0 = a.
+Proof.
+intros A a f; unfold recursion; now rewrite Nrect_base.
+Qed.
+
+Theorem recursion_succ :
+ forall (A : Type) (Aeq : relation A) (a : A) (f : N -> A -> A),
+ Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
+ forall n : N, Aeq (recursion a f (Nsucc n)) (f n (recursion a f n)).
+Proof.
+unfold recursion; intros A Aeq a f EAaa f_wd n; pattern n; apply Nrect.
+rewrite Nrect_step; rewrite Nrect_base; now apply f_wd.
+clear n; intro n; do 2 rewrite Nrect_step; intro IH. apply f_wd; [reflexivity|].
+now rewrite Nrect_step.
+Qed.
+
+(** The instantiation of operations.
+ Placing them at the very end avoids having indirections in above lemmas. *)
+
+Definition t := N.
+Definition eq := @eq N.
+Definition zero := N0.
+Definition succ := Nsucc.
+Definition pred := Npred.
+Definition add := Nplus.
+Definition sub := Nminus.
+Definition mul := Nmult.
+Definition lt := Nlt.
+Definition le := Nle.
+Definition min := Nmin.
+Definition max := Nmax.
+
+End NBinaryAxiomsMod.
+
+Module Export NBinaryPropMod := NPropFunct NBinaryAxiomsMod.
+
+(*
+Require Import NDefOps.
+Module Import NBinaryDefOpsMod := NdefOpsPropFunct NBinaryAxiomsMod.
+
+(* Some fun comparing the efficiency of the generic log defined
+by strong (course-of-value) recursion and the log defined by recursion
+on notation *)
+
+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')
+end.
+
+Definition binlog (n : N) : N :=
+match n with
+| 0 => 0
+| Npos p => binposlog p
+end.
+
+Time Eval vm_compute in (binlog 500000). (* 0 sec *)
+Time Eval vm_compute in (binlog 1000000000000000000000000000000). (* 0 sec *)
+
+*)
diff --git a/theories/Numbers/Natural/Peano/NPeano.v b/theories/Numbers/Natural/Peano/NPeano.v
index 1c83da45..becbd243 100644
--- a/theories/Numbers/Natural/Peano/NPeano.v
+++ b/theories/Numbers/Natural/Peano/NPeano.v
@@ -8,134 +8,73 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NPeano.v 11040 2008-06-03 00:04:16Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import Arith.
-Require Import Min.
-Require Import Max.
-Require Import NSub.
+Require Import Arith MinMax NAxioms NProperties.
-Module NPeanoAxiomsMod <: NAxiomsSig.
-Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig.
-Module Export NZAxiomsMod <: NZAxiomsSig.
-
-Definition NZ := nat.
-Definition NZeq := (@eq nat).
-Definition NZ0 := 0.
-Definition NZsucc := S.
-Definition NZpred := pred.
-Definition NZadd := plus.
-Definition NZsub := minus.
-Definition NZmul := mult.
-
-Theorem NZeq_equiv : equiv nat NZeq.
-Proof (eq_equiv nat).
-
-Add Relation nat NZeq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
-as NZeq_rel.
-
-(* If we say "Add Relation nat (@eq nat)" instead of "Add Relation nat NZeq"
-then the theorem generated for succ_wd below is forall x, succ x = succ x,
-which does not match the axioms in NAxiomsSig *)
-
-Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd.
-Proof.
-congruence.
-Qed.
+(** * Implementation of [NAxiomsSig] by [nat] *)
-Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd.
-Proof.
-congruence.
-Qed.
+Module NPeanoAxiomsMod <: NAxiomsSig.
-Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd.
-Proof.
-congruence.
-Qed.
+(** Bi-directional induction. *)
-Theorem NZinduction :
- forall A : nat -> Prop, predicate_wd (@eq nat) A ->
+Theorem bi_induction :
+ forall A : nat -> Prop, Proper (eq==>iff) A ->
A 0 -> (forall n : nat, A n <-> A (S n)) -> forall n : nat, A n.
Proof.
intros A A_wd A0 AS. apply nat_ind. assumption. intros; now apply -> AS.
Qed.
-Theorem NZpred_succ : forall n : nat, pred (S n) = n.
+(** Basic operations. *)
+
+Definition eq_equiv : Equivalence (@eq nat) := eq_equivalence.
+Local Obligation Tactic := simpl_relation.
+Program Instance succ_wd : Proper (eq==>eq) S.
+Program Instance pred_wd : Proper (eq==>eq) pred.
+Program Instance add_wd : Proper (eq==>eq==>eq) plus.
+Program Instance sub_wd : Proper (eq==>eq==>eq) minus.
+Program Instance mul_wd : Proper (eq==>eq==>eq) mult.
+
+Theorem pred_succ : forall n : nat, pred (S n) = n.
Proof.
reflexivity.
Qed.
-Theorem NZadd_0_l : forall n : nat, 0 + n = n.
+Theorem add_0_l : forall n : nat, 0 + n = n.
Proof.
reflexivity.
Qed.
-Theorem NZadd_succ_l : forall n m : nat, (S n) + m = S (n + m).
+Theorem add_succ_l : forall n m : nat, (S n) + m = S (n + m).
Proof.
reflexivity.
Qed.
-Theorem NZsub_0_r : forall n : nat, n - 0 = n.
+Theorem sub_0_r : forall n : nat, n - 0 = n.
Proof.
intro n; now destruct n.
Qed.
-Theorem NZsub_succ_r : forall n m : nat, n - (S m) = pred (n - m).
+Theorem sub_succ_r : forall n m : nat, n - (S m) = pred (n - m).
Proof.
-intros n m; induction n m using nat_double_ind; simpl; auto. apply NZsub_0_r.
+intros n m; induction n m using nat_double_ind; simpl; auto. apply sub_0_r.
Qed.
-Theorem NZmul_0_l : forall n : nat, 0 * n = 0.
+Theorem mul_0_l : forall n : nat, 0 * n = 0.
Proof.
reflexivity.
Qed.
-Theorem NZmul_succ_l : forall n m : nat, S n * m = n * m + m.
+Theorem mul_succ_l : forall n m : nat, S n * m = n * m + m.
Proof.
intros n m; now rewrite plus_comm.
Qed.
-End NZAxiomsMod.
+(** Order on natural numbers *)
-Definition NZlt := lt.
-Definition NZle := le.
-Definition NZmin := min.
-Definition NZmax := max.
+Program Instance lt_wd : Proper (eq==>eq==>iff) lt.
-Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd.
-Proof.
-unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2.
-Qed.
-
-Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd.
-Proof.
-unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2.
-Qed.
-
-Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd.
-Proof.
-congruence.
-Qed.
-
-Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd.
-Proof.
-congruence.
-Qed.
-
-Theorem NZlt_eq_cases : forall n m : nat, n <= m <-> n < m \/ n = m.
+Theorem lt_eq_cases : forall n m : nat, n <= m <-> n < m \/ n = m.
Proof.
intros n m; split.
apply le_lt_or_eq.
@@ -143,59 +82,52 @@ intro H; destruct H as [H | H].
now apply lt_le_weak. rewrite H; apply le_refl.
Qed.
-Theorem NZlt_irrefl : forall n : nat, ~ (n < n).
+Theorem lt_irrefl : forall n : nat, ~ (n < n).
Proof.
exact lt_irrefl.
Qed.
-Theorem NZlt_succ_r : forall n m : nat, n < S m <-> n <= m.
+Theorem lt_succ_r : forall n m : nat, n < S m <-> n <= m.
Proof.
intros n m; split; [apply lt_n_Sm_le | apply le_lt_n_Sm].
Qed.
-Theorem NZmin_l : forall n m : nat, n <= m -> NZmin n m = n.
+Theorem min_l : forall n m : nat, n <= m -> min n m = n.
Proof.
exact min_l.
Qed.
-Theorem NZmin_r : forall n m : nat, m <= n -> NZmin n m = m.
+Theorem min_r : forall n m : nat, m <= n -> min n m = m.
Proof.
exact min_r.
Qed.
-Theorem NZmax_l : forall n m : nat, m <= n -> NZmax n m = n.
+Theorem max_l : forall n m : nat, m <= n -> max n m = n.
Proof.
exact max_l.
Qed.
-Theorem NZmax_r : forall n m : nat, n <= m -> NZmax n m = m.
+Theorem max_r : forall n m : nat, n <= m -> max n m = m.
Proof.
exact max_r.
Qed.
-End NZOrdAxiomsMod.
-
-Definition recursion : forall A : Type, A -> (nat -> A -> A) -> nat -> A :=
- fun A : Type => nat_rect (fun _ => A).
-Implicit Arguments recursion [A].
-
-Theorem succ_neq_0 : forall n : nat, S n <> 0.
-Proof.
-intros; discriminate.
-Qed.
+(** Facts specific to natural numbers, not integers. *)
Theorem pred_0 : pred 0 = 0.
Proof.
reflexivity.
Qed.
-Theorem recursion_wd : forall (A : Type) (Aeq : relation A),
- forall a a' : A, Aeq a a' ->
- forall f f' : nat -> A -> A, fun2_eq (@eq nat) Aeq Aeq f f' ->
- forall n n' : nat, n = n' ->
- Aeq (recursion a f n) (recursion a' f' n').
+Definition recursion (A : Type) : A -> (nat -> A -> A) -> nat -> A :=
+ nat_rect (fun _ => A).
+Implicit Arguments recursion [A].
+
+Instance recursion_wd (A : Type) (Aeq : relation A) :
+ Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) (@recursion A).
Proof.
-unfold fun2_eq; induction n; intros n' Enn'; rewrite <- Enn' in *; simpl; auto.
+intros a a' Ha f f' Hf n n' Hn. subst n'.
+induction n; simpl; auto. apply Hf; auto.
Qed.
Theorem recursion_0 :
@@ -206,15 +138,100 @@ Qed.
Theorem recursion_succ :
forall (A : Type) (Aeq : relation A) (a : A) (f : nat -> A -> A),
- Aeq a a -> fun2_wd (@eq nat) Aeq Aeq f ->
+ Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
forall n : nat, Aeq (recursion a f (S n)) (f n (recursion a f n)).
Proof.
-induction n; simpl; auto.
+unfold Proper, respectful in *; induction n; simpl; auto.
Qed.
-End NPeanoAxiomsMod.
+(** The instantiation of operations.
+ Placing them at the very end avoids having indirections in above lemmas. *)
-(* Now we apply the largest property functor *)
+Definition t := nat.
+Definition eq := @eq nat.
+Definition zero := 0.
+Definition succ := S.
+Definition pred := pred.
+Definition add := plus.
+Definition sub := minus.
+Definition mul := mult.
+Definition lt := lt.
+Definition le := le.
+Definition min := min.
+Definition max := max.
-Module Export NPeanoSubPropMod := NSubPropFunct NPeanoAxiomsMod.
+End NPeanoAxiomsMod.
+(** Now we apply the largest property functor *)
+
+Module Export NPeanoPropMod := NPropFunct NPeanoAxiomsMod.
+
+
+
+(** Euclidean Division *)
+
+Definition divF div x y := if leb y x then S (div (x-y) y) else 0.
+Definition modF mod x y := if leb y x then mod (x-y) y else x.
+Definition initF (_ _ : nat) := 0.
+
+Fixpoint loop {A} (F:A->A)(i:A) (n:nat) : A :=
+ match n with
+ | 0 => i
+ | S n => F (loop F i n)
+ end.
+
+Definition div x y := loop divF initF x x y.
+Definition modulo x y := loop modF initF x x y.
+Infix "/" := div : nat_scope.
+Infix "mod" := modulo (at level 40, no associativity) : nat_scope.
+
+Lemma div_mod : forall x y, y<>0 -> x = y*(x/y) + x mod y.
+Proof.
+ cut (forall n x y, y<>0 -> x<=n ->
+ x = y*(loop divF initF n x y) + (loop modF initF n x y)).
+ intros H x y Hy. apply H; auto.
+ induction n.
+ simpl; unfold initF; simpl. intros. nzsimpl. auto with arith.
+ simpl; unfold divF at 1, modF at 1.
+ intros.
+ destruct (leb y x) as [ ]_eqn:L;
+ [apply leb_complete in L | apply leb_complete_conv in L].
+ rewrite mul_succ_r, <- add_assoc, (add_comm y), add_assoc.
+ rewrite <- IHn; auto.
+ symmetry; apply sub_add; auto.
+ rewrite <- NPeanoAxiomsMod.lt_succ_r.
+ apply lt_le_trans with x; auto.
+ apply lt_minus; auto. rewrite <- neq_0_lt_0; auto.
+ nzsimpl; auto.
+Qed.
+
+Lemma mod_upper_bound : forall x y, y<>0 -> x mod y < y.
+Proof.
+ cut (forall n x y, y<>0 -> x<=n -> loop modF initF n x y < y).
+ intros H x y Hy. apply H; auto.
+ induction n.
+ simpl; unfold initF. intros. rewrite <- neq_0_lt_0; auto.
+ simpl; unfold modF at 1.
+ intros.
+ destruct (leb y x) as [ ]_eqn:L;
+ [apply leb_complete in L | apply leb_complete_conv in L]; auto.
+ apply IHn; auto.
+ rewrite <- NPeanoAxiomsMod.lt_succ_r.
+ apply lt_le_trans with x; auto.
+ apply lt_minus; auto. rewrite <- neq_0_lt_0; auto.
+Qed.
+
+Require Import NDiv.
+
+Module NDivMod <: NDivSig.
+ Include NPeanoAxiomsMod.
+ Definition div := div.
+ Definition modulo := modulo.
+ Definition div_mod := div_mod.
+ Definition mod_upper_bound := mod_upper_bound.
+ Local Obligation Tactic := simpl_relation.
+ Program Instance div_wd : Proper (eq==>eq==>eq) div.
+ Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
+End NDivMod.
+
+Module Export NDivPropMod := NDivPropFunct NDivMod NPeanoPropMod.
diff --git a/theories/Numbers/Natural/SpecViaZ/NSig.v b/theories/Numbers/Natural/SpecViaZ/NSig.v
index 0275d1e1..85639aa6 100644
--- a/theories/Numbers/Natural/SpecViaZ/NSig.v
+++ b/theories/Numbers/Natural/SpecViaZ/NSig.v
@@ -8,7 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NSig.v 11027 2008-06-01 13:28:59Z letouzey $ i*)
+(*i $Id$ i*)
Require Import ZArith Znumtheory.
@@ -25,91 +25,76 @@ Module Type NType.
Parameter t : Type.
Parameter to_Z : t -> Z.
- Notation "[ x ]" := (to_Z x).
+ Local Notation "[ x ]" := (to_Z x).
Parameter spec_pos: forall x, 0 <= [x].
Parameter of_N : N -> t.
Parameter spec_of_N: forall x, to_Z (of_N x) = Z_of_N x.
Definition to_N n := Zabs_N (to_Z n).
- Definition eq n m := ([n] = [m]).
-
- Parameter zero : t.
- Parameter one : t.
-
- Parameter spec_0: [zero] = 0.
- Parameter spec_1: [one] = 1.
+ Definition eq n m := [n] = [m].
+ Definition lt n m := [n] < [m].
+ Definition le n m := [n] <= [m].
Parameter compare : t -> t -> comparison.
-
- Parameter spec_compare: forall x y,
- match compare x y with
- | Eq => [x] = [y]
- | Lt => [x] < [y]
- | Gt => [x] > [y]
- end.
-
- Definition lt n m := compare n m = Lt.
- Definition le n m := compare n m <> Gt.
- 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.
-
Parameter eq_bool : t -> t -> bool.
-
- Parameter spec_eq_bool: forall x y,
- if eq_bool x y then [x] = [y] else [x] <> [y].
-
+ Parameter max : t -> t -> t.
+ Parameter min : t -> t -> t.
+ Parameter zero : t.
+ Parameter one : t.
Parameter succ : t -> t.
-
- Parameter spec_succ: forall n, [succ n] = [n] + 1.
-
- Parameter add : t -> t -> t.
-
- Parameter spec_add: forall x y, [add x y] = [x] + [y].
-
Parameter pred : t -> t.
-
- Parameter spec_pred: forall x, 0 < [x] -> [pred x] = [x] - 1.
- Parameter spec_pred0: forall x, [x] = 0 -> [pred x] = 0.
-
+ Parameter add : t -> t -> t.
Parameter sub : t -> t -> t.
-
- Parameter spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y].
- Parameter spec_sub0: forall x y, [x] < [y]-> [sub x y] = 0.
-
Parameter mul : t -> t -> t.
-
- Parameter spec_mul: forall x y, [mul x y] = [x] * [y].
-
Parameter square : t -> t.
-
- Parameter spec_square: forall x, [square x] = [x] * [x].
-
Parameter power_pos : t -> positive -> t.
-
- Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
-
+ Parameter power : t -> N -> t.
Parameter sqrt : t -> t.
-
- Parameter spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
-
Parameter div_eucl : t -> t -> t * t.
-
- Parameter spec_div_eucl: forall x y,
- 0 < [y] ->
- let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y].
-
Parameter div : t -> t -> t.
-
- Parameter spec_div: forall x y, 0 < [y] -> [div x y] = [x] / [y].
-
Parameter modulo : t -> t -> t.
-
- Parameter spec_modulo:
- forall x y, 0 < [y] -> [modulo x y] = [x] mod [y].
-
Parameter gcd : t -> t -> t.
-
- Parameter spec_gcd: forall a b, [gcd a b] = Zgcd (to_Z a) (to_Z b).
+ Parameter shiftr : t -> t -> t.
+ Parameter shiftl : t -> t -> t.
+ Parameter is_even : t -> bool.
+
+ Parameter spec_compare: forall x y, compare x y = Zcompare [x] [y].
+ Parameter spec_eq_bool: forall x y, eq_bool x y = Zeq_bool [x] [y].
+ Parameter spec_max : forall x y, [max x y] = Zmax [x] [y].
+ Parameter spec_min : forall x y, [min x y] = Zmin [x] [y].
+ Parameter spec_0: [zero] = 0.
+ Parameter spec_1: [one] = 1.
+ Parameter spec_succ: forall n, [succ n] = [n] + 1.
+ Parameter spec_add: forall x y, [add x y] = [x] + [y].
+ Parameter spec_pred: forall x, [pred x] = Zmax 0 ([x] - 1).
+ Parameter spec_sub: forall x y, [sub x y] = Zmax 0 ([x] - [y]).
+ Parameter spec_mul: forall x y, [mul x y] = [x] * [y].
+ Parameter spec_square: forall x, [square x] = [x] * [x].
+ Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
+ Parameter spec_power: forall x n, [power x n] = [x] ^ Z_of_N n.
+ Parameter spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
+ Parameter spec_div_eucl: forall x y,
+ let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y].
+ Parameter spec_div: forall x y, [div x y] = [x] / [y].
+ Parameter spec_modulo: forall x y, [modulo x y] = [x] mod [y].
+ Parameter spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b].
+ Parameter spec_shiftr: forall p x, [shiftr p x] = [x] / 2^[p].
+ Parameter spec_shiftl: forall p x, [shiftl p x] = [x] * 2^[p].
+ Parameter spec_is_even: forall x,
+ if is_even x then [x] mod 2 = 0 else [x] mod 2 = 1.
End NType.
+
+Module Type NType_Notation (Import N:NType).
+ Notation "[ x ]" := (to_Z x).
+ Infix "==" := eq (at level 70).
+ Notation "0" := zero.
+ Infix "+" := add.
+ Infix "-" := sub.
+ Infix "*" := mul.
+ Infix "<=" := le.
+ Infix "<" := lt.
+End NType_Notation.
+
+Module Type NType' := NType <+ NType_Notation.
diff --git a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
index 84836268..ab749bd1 100644
--- a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
+++ b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
@@ -6,101 +6,47 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: NSigNAxioms.v 11282 2008-07-28 11:51:53Z msozeau $ i*)
+(*i $Id$ i*)
-Require Import ZArith.
-Require Import Nnat.
-Require Import NAxioms.
-Require Import NSig.
+Require Import ZArith Nnat NAxioms NDiv NSig.
(** * The interface [NSig.NType] implies the interface [NAxiomsSig] *)
-Module NSig_NAxioms (N:NType) <: NAxiomsSig.
-
-Delimit Scope IntScope with Int.
-Bind Scope IntScope with N.t.
-Open Local Scope IntScope.
-Notation "[ x ]" := (N.to_Z x) : IntScope.
-Infix "==" := N.eq (at level 70) : IntScope.
-Notation "0" := N.zero : IntScope.
-Infix "+" := N.add : IntScope.
-Infix "-" := N.sub : IntScope.
-Infix "*" := N.mul : IntScope.
-
-Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig.
-Module Export NZAxiomsMod <: NZAxiomsSig.
-
-Definition NZ := N.t.
-Definition NZeq := N.eq.
-Definition NZ0 := N.zero.
-Definition NZsucc := N.succ.
-Definition NZpred := N.pred.
-Definition NZadd := N.add.
-Definition NZsub := N.sub.
-Definition NZmul := N.mul.
-
-Theorem NZeq_equiv : equiv N.t N.eq.
-Proof.
-repeat split; repeat red; intros; auto; congruence.
-Qed.
+Module NTypeIsNAxioms (Import N : NType').
-Add Relation N.t N.eq
- reflexivity proved by (proj1 NZeq_equiv)
- symmetry proved by (proj2 (proj2 NZeq_equiv))
- transitivity proved by (proj1 (proj2 NZeq_equiv))
- as NZeq_rel.
+Hint Rewrite
+ spec_0 spec_succ spec_add spec_mul spec_pred spec_sub
+ spec_div spec_modulo spec_gcd spec_compare spec_eq_bool
+ spec_max spec_min spec_power_pos spec_power
+ : nsimpl.
+Ltac nsimpl := autorewrite with nsimpl.
+Ltac ncongruence := unfold eq; repeat red; intros; nsimpl; congruence.
+Ltac zify := unfold eq, lt, le in *; nsimpl.
-Add Morphism NZsucc with signature N.eq ==> N.eq as NZsucc_wd.
-Proof.
-unfold N.eq; intros; rewrite 2 N.spec_succ; f_equal; auto.
-Qed.
+Local Obligation Tactic := ncongruence.
-Add Morphism NZpred with signature N.eq ==> N.eq as NZpred_wd.
-Proof.
-unfold N.eq; intros.
-generalize (N.spec_pos y) (N.spec_pos x) (N.spec_eq_bool x 0).
-destruct N.eq_bool; rewrite N.spec_0; intros.
-rewrite 2 N.spec_pred0; congruence.
-rewrite 2 N.spec_pred; f_equal; auto; try omega.
-Qed.
+Instance eq_equiv : Equivalence eq.
+Proof. unfold eq. firstorder. Qed.
-Add Morphism NZadd with signature N.eq ==> N.eq ==> N.eq as NZadd_wd.
-Proof.
-unfold N.eq; intros; rewrite 2 N.spec_add; f_equal; auto.
-Qed.
+Program Instance succ_wd : Proper (eq==>eq) succ.
+Program Instance pred_wd : Proper (eq==>eq) pred.
+Program Instance add_wd : Proper (eq==>eq==>eq) add.
+Program Instance sub_wd : Proper (eq==>eq==>eq) sub.
+Program Instance mul_wd : Proper (eq==>eq==>eq) mul.
-Add Morphism NZsub with signature N.eq ==> N.eq ==> N.eq as NZsub_wd.
+Theorem pred_succ : forall n, pred (succ n) == n.
Proof.
-unfold N.eq; intros x x' Hx y y' Hy.
-destruct (Z_lt_le_dec [x] [y]).
-rewrite 2 N.spec_sub0; f_equal; congruence.
-rewrite 2 N.spec_sub; f_equal; congruence.
+intros. zify. generalize (spec_pos n); omega with *.
Qed.
-Add Morphism NZmul with signature N.eq ==> N.eq ==> N.eq as NZmul_wd.
-Proof.
-unfold N.eq; intros; rewrite 2 N.spec_mul; f_equal; auto.
-Qed.
-
-Theorem NZpred_succ : forall n, N.pred (N.succ n) == n.
-Proof.
-unfold N.eq; intros.
-rewrite N.spec_pred; rewrite N.spec_succ.
-omega.
-generalize (N.spec_pos n); omega.
-Qed.
-
-Definition N_of_Z z := N.of_N (Zabs_N z).
+Definition N_of_Z z := of_N (Zabs_N z).
Section Induction.
Variable A : N.t -> Prop.
-Hypothesis A_wd : predicate_wd N.eq A.
+Hypothesis A_wd : Proper (eq==>iff) A.
Hypothesis A0 : A 0.
-Hypothesis AS : forall n, A n <-> A (N.succ n).
-
-Add Morphism A with signature N.eq ==> iff as A_morph.
-Proof. apply A_wd. Qed.
+Hypothesis AS : forall n, A n <-> A (succ n).
Let B (z : Z) := A (N_of_Z z).
@@ -108,17 +54,17 @@ Lemma B0 : B 0.
Proof.
unfold B, N_of_Z; simpl.
rewrite <- (A_wd 0); auto.
-red; rewrite N.spec_0, N.spec_of_N; auto.
+red; rewrite spec_0, spec_of_N; auto.
Qed.
Lemma BS : forall z : Z, (0 <= z)%Z -> B z -> B (z + 1).
Proof.
intros z H1 H2.
unfold B in *. apply -> AS in H2.
-setoid_replace (N_of_Z (z + 1)) with (N.succ (N_of_Z z)); auto.
-unfold N.eq. rewrite N.spec_succ.
+setoid_replace (N_of_Z (z + 1)) with (succ (N_of_Z z)); auto.
+unfold eq. rewrite spec_succ.
unfold N_of_Z.
-rewrite 2 N.spec_of_N, 2 Z_of_N_abs, 2 Zabs_eq; auto with zarith.
+rewrite 2 spec_of_N, 2 Z_of_N_abs, 2 Zabs_eq; auto with zarith.
Qed.
Lemma B_holds : forall z : Z, (0 <= z)%Z -> B z.
@@ -126,193 +72,144 @@ Proof.
exact (natlike_ind B B0 BS).
Qed.
-Theorem NZinduction : forall n, A n.
+Theorem bi_induction : forall n, A n.
Proof.
-intro n. setoid_replace n with (N_of_Z (N.to_Z n)).
-apply B_holds. apply N.spec_pos.
+intro n. setoid_replace n with (N_of_Z (to_Z n)).
+apply B_holds. apply spec_pos.
red; unfold N_of_Z.
-rewrite N.spec_of_N, Z_of_N_abs, Zabs_eq; auto.
-apply N.spec_pos.
+rewrite spec_of_N, Z_of_N_abs, Zabs_eq; auto.
+apply spec_pos.
Qed.
End Induction.
-Theorem NZadd_0_l : forall n, 0 + n == n.
+Theorem add_0_l : forall n, 0 + n == n.
Proof.
-intros; red; rewrite N.spec_add, N.spec_0; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZadd_succ_l : forall n m, (N.succ n) + m == N.succ (n + m).
+Theorem add_succ_l : forall n m, (succ n) + m == succ (n + m).
Proof.
-intros; red; rewrite N.spec_add, 2 N.spec_succ, N.spec_add; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZsub_0_r : forall n, n - 0 == n.
+Theorem sub_0_r : forall n, n - 0 == n.
Proof.
-intros; red; rewrite N.spec_sub; rewrite N.spec_0; auto with zarith.
-apply N.spec_pos.
+intros. zify. generalize (spec_pos n); omega with *.
Qed.
-Theorem NZsub_succ_r : forall n m, n - (N.succ m) == N.pred (n - m).
+Theorem sub_succ_r : forall n m, n - (succ m) == pred (n - m).
Proof.
-intros; red.
-destruct (Z_lt_le_dec [n] [N.succ m]) as [H|H].
-rewrite N.spec_sub0; auto.
-rewrite N.spec_succ in H.
-rewrite N.spec_pred0; auto.
-destruct (Z_eq_dec [n] [m]).
-rewrite N.spec_sub; auto with zarith.
-rewrite N.spec_sub0; auto with zarith.
-
-rewrite N.spec_sub, N.spec_succ in *; auto.
-rewrite N.spec_pred, N.spec_sub; auto with zarith.
-rewrite N.spec_sub; auto with zarith.
+intros. zify. omega with *.
Qed.
-Theorem NZmul_0_l : forall n, 0 * n == 0.
+Theorem mul_0_l : forall n, 0 * n == 0.
Proof.
-intros; red.
-rewrite N.spec_mul, N.spec_0; auto with zarith.
+intros. zify. auto with zarith.
Qed.
-Theorem NZmul_succ_l : forall n m, (N.succ n) * m == n * m + m.
+Theorem mul_succ_l : forall n m, (succ n) * m == n * m + m.
Proof.
-intros; red.
-rewrite N.spec_add, 2 N.spec_mul, N.spec_succ; ring.
+intros. zify. ring.
Qed.
-End NZAxiomsMod.
-
-Definition NZlt := N.lt.
-Definition NZle := N.le.
-Definition NZmin := N.min.
-Definition NZmax := N.max.
+(** Order *)
-Infix "<=" := N.le : IntScope.
-Infix "<" := N.lt : IntScope.
-
-Lemma spec_compare_alt : forall x y, N.compare x y = ([x] ?= [y])%Z.
+Lemma compare_spec : forall x y, CompSpec eq lt x y (compare x y).
Proof.
- intros; generalize (N.spec_compare x y).
- destruct (N.compare x y); auto.
- intros H; rewrite H; symmetry; apply Zcompare_refl.
+ intros. zify. destruct (Zcompare_spec [x] [y]); auto.
Qed.
-Lemma spec_lt : forall x y, (x<y) <-> ([x]<[y])%Z.
-Proof.
- intros; unfold N.lt, Zlt; rewrite spec_compare_alt; intuition.
-Qed.
+Definition eqb := eq_bool.
-Lemma spec_le : forall x y, (x<=y) <-> ([x]<=[y])%Z.
+Lemma eqb_eq : forall x y, eq_bool x y = true <-> x == y.
Proof.
- intros; unfold N.le, Zle; rewrite spec_compare_alt; intuition.
+ intros. zify. symmetry. apply Zeq_is_eq_bool.
Qed.
-Lemma spec_min : forall x y, [N.min x y] = Zmin [x] [y].
+Instance compare_wd : Proper (eq ==> eq ==> Logic.eq) compare.
Proof.
- intros; unfold N.min, Zmin.
- rewrite spec_compare_alt; destruct Zcompare; auto.
+intros x x' Hx y y' Hy. rewrite 2 spec_compare, Hx, Hy; intuition.
Qed.
-Lemma spec_max : forall x y, [N.max x y] = Zmax [x] [y].
+Instance lt_wd : Proper (eq ==> eq ==> iff) lt.
Proof.
- intros; unfold N.max, Zmax.
- rewrite spec_compare_alt; destruct Zcompare; auto.
-Qed.
-
-Add Morphism N.compare with signature N.eq ==> N.eq ==> (@eq comparison) as compare_wd.
-Proof.
-intros x x' Hx y y' Hy.
-rewrite 2 spec_compare_alt. unfold N.eq in *. rewrite Hx, Hy; intuition.
+intros x x' Hx y y' Hy; unfold lt; rewrite Hx, Hy; intuition.
Qed.
-Add Morphism N.lt with signature N.eq ==> N.eq ==> iff as NZlt_wd.
+Theorem lt_eq_cases : forall n m, n <= m <-> n < m \/ n == m.
Proof.
-intros x x' Hx y y' Hy; unfold N.lt; rewrite Hx, Hy; intuition.
+intros. zify. omega.
Qed.
-Add Morphism N.le with signature N.eq ==> N.eq ==> iff as NZle_wd.
+Theorem lt_irrefl : forall n, ~ n < n.
Proof.
-intros x x' Hx y y' Hy; unfold N.le; rewrite Hx, Hy; intuition.
+intros. zify. omega.
Qed.
-Add Morphism N.min with signature N.eq ==> N.eq ==> N.eq as NZmin_wd.
+Theorem lt_succ_r : forall n m, n < (succ m) <-> n <= m.
Proof.
-intros; red; rewrite 2 spec_min; congruence.
+intros. zify. omega.
Qed.
-Add Morphism N.max with signature N.eq ==> N.eq ==> N.eq as NZmax_wd.
+Theorem min_l : forall n m, n <= m -> min n m == n.
Proof.
-intros; red; rewrite 2 spec_max; congruence.
+intros n m. zify. omega with *.
Qed.
-Theorem NZlt_eq_cases : forall n m, n <= m <-> n < m \/ n == m.
+Theorem min_r : forall n m, m <= n -> min n m == m.
Proof.
-intros.
-unfold N.eq; rewrite spec_lt, spec_le; omega.
+intros n m. zify. omega with *.
Qed.
-Theorem NZlt_irrefl : forall n, ~ n < n.
+Theorem max_l : forall n m, m <= n -> max n m == n.
Proof.
-intros; rewrite spec_lt; auto with zarith.
+intros n m. zify. omega with *.
Qed.
-Theorem NZlt_succ_r : forall n m, n < (N.succ m) <-> n <= m.
+Theorem max_r : forall n m, n <= m -> max n m == m.
Proof.
-intros; rewrite spec_lt, spec_le, N.spec_succ; omega.
+intros n m. zify. omega with *.
Qed.
-Theorem NZmin_l : forall n m, n <= m -> N.min n m == n.
-Proof.
-intros n m; unfold N.eq; rewrite spec_le, spec_min.
-generalize (Zmin_spec [n] [m]); omega.
-Qed.
+(** Properties specific to natural numbers, not integers. *)
-Theorem NZmin_r : forall n m, m <= n -> N.min n m == m.
+Theorem pred_0 : pred 0 == 0.
Proof.
-intros n m; unfold N.eq; rewrite spec_le, spec_min.
-generalize (Zmin_spec [n] [m]); omega.
+zify. auto.
Qed.
-Theorem NZmax_l : forall n m, m <= n -> N.max n m == n.
-Proof.
-intros n m; unfold N.eq; rewrite spec_le, spec_max.
-generalize (Zmax_spec [n] [m]); omega.
-Qed.
+Program Instance div_wd : Proper (eq==>eq==>eq) div.
+Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
-Theorem NZmax_r : forall n m, n <= m -> N.max n m == m.
+Theorem div_mod : forall a b, ~b==0 -> a == b*(div a b) + (modulo a b).
Proof.
-intros n m; unfold N.eq; rewrite spec_le, spec_max.
-generalize (Zmax_spec [n] [m]); omega.
+intros a b. zify. intros. apply Z_div_mod_eq_full; auto.
Qed.
-End NZOrdAxiomsMod.
-
-Theorem pred_0 : N.pred 0 == 0.
+Theorem mod_upper_bound : forall a b, ~b==0 -> modulo a b < b.
Proof.
-red; rewrite N.spec_pred0; rewrite N.spec_0; auto.
+intros a b. zify. intros.
+destruct (Z_mod_lt [a] [b]); auto.
+generalize (spec_pos b); auto with zarith.
Qed.
Definition recursion (A : Type) (a : A) (f : N.t -> A -> A) (n : N.t) :=
Nrect (fun _ => A) a (fun n a => f (N.of_N n) a) (N.to_N n).
Implicit Arguments recursion [A].
-Theorem recursion_wd :
-forall (A : Type) (Aeq : relation A),
- forall a a' : A, Aeq a a' ->
- forall f f' : N.t -> A -> A, fun2_eq N.eq Aeq Aeq f f' ->
- forall x x' : N.t, x == x' ->
- Aeq (recursion a f x) (recursion a' f' x').
+Instance recursion_wd (A : Type) (Aeq : relation A) :
+ Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) (@recursion A).
Proof.
-unfold fun2_wd, N.eq, fun2_eq.
-intros A Aeq a a' Eaa' f f' Eff' x x' Exx'.
+unfold eq.
+intros a a' Eaa' f f' Eff' x x' Exx'.
unfold recursion.
unfold N.to_N.
rewrite <- Exx'; clear x' Exx'.
replace (Zabs_N [x]) with (N_of_nat (Zabs_nat [x])).
induction (Zabs_nat [x]).
simpl; auto.
-rewrite N_of_S, 2 Nrect_step; auto.
+rewrite N_of_S, 2 Nrect_step; auto. apply Eff'; auto.
destruct [x]; simpl; auto.
change (nat_of_P p) with (nat_of_N (Npos p)); apply N_of_nat_of_N.
change (nat_of_P p) with (nat_of_N (Npos p)); apply N_of_nat_of_N.
@@ -326,11 +223,11 @@ Qed.
Theorem recursion_succ :
forall (A : Type) (Aeq : relation A) (a : A) (f : N.t -> A -> A),
- Aeq a a -> fun2_wd N.eq Aeq Aeq f ->
- forall n, Aeq (recursion a f (N.succ n)) (f n (recursion a f n)).
+ Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
+ forall n, Aeq (recursion a f (succ n)) (f n (recursion a f n)).
Proof.
-unfold N.eq, recursion, fun2_wd; intros A Aeq a f EAaa f_wd n.
-replace (N.to_N (N.succ n)) with (Nsucc (N.to_N n)).
+unfold N.eq, recursion; intros A Aeq a f EAaa f_wd n.
+replace (N.to_N (succ n)) with (Nsucc (N.to_N n)).
rewrite Nrect_step.
apply f_wd; auto.
unfold N.to_N.
@@ -340,7 +237,6 @@ rewrite N.spec_of_N, Z_of_N_abs, Zabs_eq; auto.
fold (recursion a f n).
apply recursion_wd; auto.
red; auto.
-red; auto.
unfold N.to_N.
rewrite N.spec_succ.
@@ -349,8 +245,12 @@ apply Z_of_N_eq_rev.
rewrite Z_of_N_succ.
rewrite 2 Z_of_N_abs.
rewrite 2 Zabs_eq; auto.
-generalize (N.spec_pos n); auto with zarith.
-apply N.spec_pos; auto.
+generalize (spec_pos n); auto with zarith.
+apply spec_pos; auto.
Qed.
-End NSig_NAxioms.
+End NTypeIsNAxioms.
+
+Module NType_NAxioms (N : NType)
+ <: NAxiomsSig <: NDivSig <: HasCompare N <: HasEqBool N <: HasMinMax N
+ := N <+ NTypeIsNAxioms.
diff --git a/theories/Numbers/NumPrelude.v b/theories/Numbers/NumPrelude.v
index 95d8b366..468b0613 100644
--- a/theories/Numbers/NumPrelude.v
+++ b/theories/Numbers/NumPrelude.v
@@ -8,9 +8,9 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NumPrelude.v 11674 2008-12-12 19:48:40Z letouzey $ i*)
+(*i $Id$ i*)
-Require Export Setoid.
+Require Export Setoid Morphisms.
Set Implicit Arguments.
(*
@@ -91,85 +91,31 @@ end.
Tactic Notation "stepr" constr(t2') "in" hyp(H) "by" tactic(r) := stepr t2' in H; [| r].
-(** Extentional properties of predicates, relations and functions *)
+(** Predicates, relations, functions *)
Definition predicate (A : Type) := A -> Prop.
-Section ExtensionalProperties.
-
-Variables A B C : Type.
-Variable Aeq : relation A.
-Variable Beq : relation B.
-Variable Ceq : relation C.
-
-(* "wd" stands for "well-defined" *)
-
-Definition fun_wd (f : A -> B) := forall x y : A, Aeq x y -> Beq (f x) (f y).
-
-Definition fun2_wd (f : A -> B -> C) :=
- forall x x' : A, Aeq x x' -> forall y y' : B, Beq y y' -> Ceq (f x y) (f x' y').
-
-Definition fun_eq : relation (A -> B) :=
- fun f f' => forall x x' : A, Aeq x x' -> Beq (f x) (f' x').
-
-(* Note that reflexivity of fun_eq means that every function
-is well-defined w.r.t. Aeq and Beq, i.e.,
-forall x x' : A, Aeq x x' -> Beq (f x) (f x') *)
-
-Definition fun2_eq (f f' : A -> B -> C) :=
- forall x x' : A, Aeq x x' -> forall y y' : B, Beq y y' -> Ceq (f x y) (f' x' y').
-
-End ExtensionalProperties.
-
-(* The following definitions instantiate Beq or Ceq to iff; therefore, they
-have to be outside the ExtensionalProperties section *)
-
-Definition predicate_wd (A : Type) (Aeq : relation A) := fun_wd Aeq iff.
-
-Definition relation_wd (A B : Type) (Aeq : relation A) (Beq : relation B) :=
- fun2_wd Aeq Beq iff.
-
-Definition relations_eq (A B : Type) (R1 R2 : A -> B -> Prop) :=
- forall (x : A) (y : B), R1 x y <-> R2 x y.
-
-Theorem relations_eq_equiv :
- forall (A B : Type), equiv (A -> B -> Prop) (@relations_eq A B).
-Proof.
-intros A B; unfold equiv. split; [| split];
-unfold reflexive, symmetric, transitive, relations_eq.
-reflexivity.
-intros R1 R2 R3 H1 H2 x y; rewrite H1; apply H2.
-now symmetry.
-Qed.
-
-Add Parametric Relation (A B : Type) : (A -> B -> Prop) (@relations_eq A B)
- reflexivity proved by (proj1 (relations_eq_equiv A B))
- symmetry proved by (proj2 (proj2 (relations_eq_equiv A B)))
- transitivity proved by (proj1 (proj2 (relations_eq_equiv A B)))
-as relations_eq_rel.
-
-Add Parametric Morphism (A : Type) : (@well_founded A) with signature (@relations_eq A A) ==> iff as well_founded_wd.
+Instance well_founded_wd A :
+ Proper (@relation_equivalence A ==> iff) (@well_founded A).
Proof.
-unfold relations_eq, well_founded; intros R1 R2 H;
-split; intros H1 a; induction (H1 a) as [x H2 H3]; constructor;
-intros y H4; apply H3; [now apply <- H | now apply -> H].
+intros R1 R2 H.
+split; intros WF a; induction (WF a) as [x _ WF']; constructor;
+intros y Ryx; apply WF'; destruct (H y x); auto.
Qed.
-(* solve_predicate_wd solves the goal [predicate_wd P] for P consisting of
-morhisms and quatifiers *)
+(** [solve_predicate_wd] solves the goal [Proper (?==>iff) P]
+ for P consisting of morphisms and quantifiers *)
Ltac solve_predicate_wd :=
-unfold predicate_wd;
let x := fresh "x" in
let y := fresh "y" in
let H := fresh "H" in
intros x y H; setoid_rewrite H; reflexivity.
-(* solve_relation_wd solves the goal [relation_wd R] for R consisting of
-morhisms and quatifiers *)
+(** [solve_relation_wd] solves the goal [Proper (?==>?==>iff) R]
+ for R consisting of morphisms and quantifiers *)
Ltac solve_relation_wd :=
-unfold relation_wd, fun2_wd;
let x1 := fresh "x" in
let y1 := fresh "y" in
let H1 := fresh "H" in
@@ -191,77 +137,3 @@ Ltac induction_maker n t :=
pattern n; t; clear n;
[solve_predicate_wd | ..].
-(** Relations on cartesian product. Used in MiscFunct for defining
-functions whose domain is a product of sets by primitive recursion *)
-
-Section RelationOnProduct.
-
-Variables A B : Set.
-Variable Aeq : relation A.
-Variable Beq : relation B.
-
-Hypothesis EA_equiv : equiv A Aeq.
-Hypothesis EB_equiv : equiv B Beq.
-
-Definition prod_rel : relation (A * B) :=
- fun p1 p2 => Aeq (fst p1) (fst p2) /\ Beq (snd p1) (snd p2).
-
-Lemma prod_rel_refl : reflexive (A * B) prod_rel.
-Proof.
-unfold reflexive, prod_rel.
-destruct x; split; [apply (proj1 EA_equiv) | apply (proj1 EB_equiv)]; simpl.
-Qed.
-
-Lemma prod_rel_sym : symmetric (A * B) prod_rel.
-Proof.
-unfold symmetric, prod_rel.
-destruct x; destruct y;
-split; [apply (proj2 (proj2 EA_equiv)) | apply (proj2 (proj2 EB_equiv))]; simpl in *; tauto.
-Qed.
-
-Lemma prod_rel_trans : transitive (A * B) prod_rel.
-Proof.
-unfold transitive, prod_rel.
-destruct x; destruct y; destruct z; simpl.
-intros; split; [apply (proj1 (proj2 EA_equiv)) with (y := a0) |
-apply (proj1 (proj2 EB_equiv)) with (y := b0)]; tauto.
-Qed.
-
-Theorem prod_rel_equiv : equiv (A * B) prod_rel.
-Proof.
-unfold equiv; split; [exact prod_rel_refl | split; [exact prod_rel_trans | exact prod_rel_sym]].
-Qed.
-
-End RelationOnProduct.
-
-Implicit Arguments prod_rel [A B].
-Implicit Arguments prod_rel_equiv [A B].
-
-(** Miscellaneous *)
-
-(*Definition comp_bool (x y : comparison) : bool :=
-match x, y with
-| Lt, Lt => true
-| Eq, Eq => true
-| Gt, Gt => true
-| _, _ => false
-end.
-
-Theorem comp_bool_correct : forall x y : comparison,
- comp_bool x y <-> x = y.
-Proof.
-destruct x; destruct y; simpl; split; now intro.
-Qed.*)
-
-Lemma eq_equiv : forall A : Set, equiv A (@eq A).
-Proof.
-intro A; unfold equiv, reflexive, symmetric, transitive.
-repeat split; [exact (@trans_eq A) | exact (@sym_eq A)].
-(* It is interesting how the tactic split proves reflexivity *)
-Qed.
-
-(*Add Relation (fun A : Set => A) LE_Set
- reflexivity proved by (fun A : Set => (proj1 (eq_equiv A)))
- symmetry proved by (fun A : Set => (proj2 (proj2 (eq_equiv A))))
- transitivity proved by (fun A : Set => (proj1 (proj2 (eq_equiv A))))
-as EA_rel.*)
diff --git a/theories/Numbers/Rational/BigQ/BigQ.v b/theories/Numbers/Rational/BigQ/BigQ.v
index f01cbbc5..0bc71166 100644
--- a/theories/Numbers/Rational/BigQ/BigQ.v
+++ b/theories/Numbers/Rational/BigQ/BigQ.v
@@ -5,12 +5,13 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-(*i $Id: BigQ.v 12509 2009-11-12 15:52:50Z letouzey $ i*)
+(** * BigQ: an efficient implementation of rational numbers *)
+
+(** Initial authors: Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-Require Import Field Qfield BigN BigZ QSig QMake.
+Require Export BigZ.
+Require Import Field Qfield QSig QMake Orders GenericMinMax.
(** We choose for BigQ an implemention with
multiple representation of 0: 0, 1/0, 2/0 etc.
@@ -34,7 +35,9 @@ End BigN_BigZ.
(** This allows to build [BigQ] out of [BigN] and [BigQ] via [QMake] *)
-Module BigQ <: QSig.QType := QMake.Make BigN BigZ BigN_BigZ.
+Module BigQ <: QType <: OrderedTypeFull <: TotalOrder :=
+ QMake.Make BigN BigZ BigN_BigZ <+ !QProperties <+ HasEqBool2Dec
+ <+ !MinMaxLogicalProperties <+ !MinMaxDecProperties.
(** Notations about [BigQ] *)
@@ -43,12 +46,40 @@ Notation bigQ := BigQ.t.
Delimit Scope bigQ_scope with bigQ.
Bind Scope bigQ_scope with bigQ.
Bind Scope bigQ_scope with BigQ.t.
-
-(* Allow nice printing of rational numerals, either as (Qz 1234)
- or as (Qq 1234 5678) *)
+Bind Scope bigQ_scope with BigQ.t_.
+(* Bind Scope has no retroactive effect, let's declare scopes by hand. *)
Arguments Scope BigQ.Qz [bigZ_scope].
-Arguments Scope BigQ.Qq [bigZ_scope bigN_scope].
-
+Arguments Scope BigQ.Qq [bigZ_scope bigN_scope].
+Arguments Scope BigQ.to_Q [bigQ_scope].
+Arguments Scope BigQ.red [bigQ_scope].
+Arguments Scope BigQ.opp [bigQ_scope].
+Arguments Scope BigQ.inv [bigQ_scope].
+Arguments Scope BigQ.square [bigQ_scope].
+Arguments Scope BigQ.add [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.sub [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.mul [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.div [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.eq [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.lt [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.le [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.eq [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.compare [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.min [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.max [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.eq_bool [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.power_pos [bigQ_scope positive_scope].
+Arguments Scope BigQ.power [bigQ_scope Z_scope].
+Arguments Scope BigQ.inv_norm [bigQ_scope].
+Arguments Scope BigQ.add_norm [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.sub_norm [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.mul_norm [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.div_norm [bigQ_scope bigQ_scope].
+Arguments Scope BigQ.power_norm [bigQ_scope bigQ_scope].
+
+(** As in QArith, we use [#] to denote fractions *)
+Notation "p # q" := (BigQ.Qq p q) (at level 55, no associativity) : bigQ_scope.
+Local Notation "0" := BigQ.zero : bigQ_scope.
+Local Notation "1" := BigQ.one : bigQ_scope.
Infix "+" := BigQ.add : bigQ_scope.
Infix "-" := BigQ.sub : bigQ_scope.
Notation "- x" := (BigQ.opp x) : bigQ_scope.
@@ -57,142 +88,102 @@ Infix "/" := BigQ.div : bigQ_scope.
Infix "^" := BigQ.power : bigQ_scope.
Infix "?=" := BigQ.compare : bigQ_scope.
Infix "==" := BigQ.eq : bigQ_scope.
+Notation "x != y" := (~x==y)%bigQ (at level 70, no associativity) : bigQ_scope.
Infix "<" := BigQ.lt : bigQ_scope.
Infix "<=" := BigQ.le : bigQ_scope.
Notation "x > y" := (BigQ.lt y x)(only parsing) : bigQ_scope.
Notation "x >= y" := (BigQ.le y x)(only parsing) : bigQ_scope.
+Notation "x < y < z" := (x<y /\ y<z)%bigQ : bigQ_scope.
+Notation "x < y <= z" := (x<y /\ y<=z)%bigQ : bigQ_scope.
+Notation "x <= y < z" := (x<=y /\ y<z)%bigQ : bigQ_scope.
+Notation "x <= y <= z" := (x<=y /\ y<=z)%bigQ : bigQ_scope.
Notation "[ q ]" := (BigQ.to_Q q) : bigQ_scope.
-Open Scope bigQ_scope.
-
-(** [BigQ] is a setoid *)
-
-Add Relation BigQ.t BigQ.eq
- reflexivity proved by (fun x => Qeq_refl [x])
- symmetry proved by (fun x y => Qeq_sym [x] [y])
- transitivity proved by (fun x y z => Qeq_trans [x] [y] [z])
-as BigQeq_rel.
-
-Add Morphism BigQ.add with signature BigQ.eq ==> BigQ.eq ==> BigQ.eq as BigQadd_wd.
-Proof.
- unfold BigQ.eq; intros; rewrite !BigQ.spec_add; rewrite H, H0; apply Qeq_refl.
-Qed.
-
-Add Morphism BigQ.opp with signature BigQ.eq ==> BigQ.eq as BigQopp_wd.
-Proof.
- unfold BigQ.eq; intros; rewrite !BigQ.spec_opp; rewrite H; apply Qeq_refl.
-Qed.
-
-Add Morphism BigQ.sub with signature BigQ.eq ==> BigQ.eq ==> BigQ.eq as BigQsub_wd.
-Proof.
- unfold BigQ.eq; intros; rewrite !BigQ.spec_sub; rewrite H, H0; apply Qeq_refl.
-Qed.
-
-Add Morphism BigQ.mul with signature BigQ.eq ==> BigQ.eq ==> BigQ.eq as BigQmul_wd.
-Proof.
- unfold BigQ.eq; intros; rewrite !BigQ.spec_mul; rewrite H, H0; apply Qeq_refl.
-Qed.
-
-Add Morphism BigQ.inv with signature BigQ.eq ==> BigQ.eq as BigQinv_wd.
-Proof.
- unfold BigQ.eq; intros; rewrite !BigQ.spec_inv; rewrite H; apply Qeq_refl.
-Qed.
-
-Add Morphism BigQ.div with signature BigQ.eq ==> BigQ.eq ==> BigQ.eq as BigQdiv_wd.
-Proof.
- unfold BigQ.eq; intros; rewrite !BigQ.spec_div; rewrite H, H0; apply Qeq_refl.
-Qed.
-
-(* TODO : fix this. For the moment it's useless (horribly slow)
-Hint Rewrite
- BigQ.spec_0 BigQ.spec_1 BigQ.spec_m1 BigQ.spec_compare
- BigQ.spec_red BigQ.spec_add BigQ.spec_sub BigQ.spec_opp
- BigQ.spec_mul BigQ.spec_inv BigQ.spec_div BigQ.spec_power_pos
- BigQ.spec_square : bigq. *)
-
+Local Open Scope bigQ_scope.
(** [BigQ] is a field *)
Lemma BigQfieldth :
- field_theory BigQ.zero BigQ.one BigQ.add BigQ.mul BigQ.sub BigQ.opp BigQ.div BigQ.inv BigQ.eq.
+ field_theory 0 1 BigQ.add BigQ.mul BigQ.sub BigQ.opp
+ BigQ.div BigQ.inv BigQ.eq.
Proof.
constructor.
-constructor; intros; red.
-rewrite BigQ.spec_add, BigQ.spec_0; ring.
-rewrite ! BigQ.spec_add; ring.
-rewrite ! BigQ.spec_add; ring.
-rewrite BigQ.spec_mul, BigQ.spec_1; ring.
-rewrite ! BigQ.spec_mul; ring.
-rewrite ! BigQ.spec_mul; ring.
-rewrite BigQ.spec_add, ! BigQ.spec_mul, BigQ.spec_add; ring.
-unfold BigQ.sub; apply Qeq_refl.
-rewrite BigQ.spec_add, BigQ.spec_0, BigQ.spec_opp; ring.
-compute; discriminate.
-intros; red.
-unfold BigQ.div; apply Qeq_refl.
-intros; red.
-rewrite BigQ.spec_mul, BigQ.spec_inv, BigQ.spec_1; field.
-rewrite <- BigQ.spec_0; auto.
-Qed.
-
-Lemma BigQpowerth :
- power_theory BigQ.one BigQ.mul BigQ.eq Z_of_N BigQ.power.
-Proof.
constructor.
-intros; red.
-rewrite BigQ.spec_power.
-replace ([r] ^ Z_of_N n)%Q with (pow_N 1 Qmult [r] n)%Q.
-destruct n.
-simpl; compute; auto.
-induction p; simpl; auto; try rewrite !BigQ.spec_mul, !IHp; apply Qeq_refl.
-destruct n; reflexivity.
-Qed.
-
-Lemma BigQ_eq_bool_correct :
- forall x y, BigQ.eq_bool x y = true -> x==y.
-Proof.
-intros; generalize (BigQ.spec_eq_bool x y); rewrite H; auto.
+exact BigQ.add_0_l. exact BigQ.add_comm. exact BigQ.add_assoc.
+exact BigQ.mul_1_l. exact BigQ.mul_comm. exact BigQ.mul_assoc.
+exact BigQ.mul_add_distr_r. exact BigQ.sub_add_opp.
+exact BigQ.add_opp_diag_r. exact BigQ.neq_1_0.
+exact BigQ.div_mul_inv. exact BigQ.mul_inv_diag_l.
Qed.
-Lemma BigQ_eq_bool_complete :
- forall x y, x==y -> BigQ.eq_bool x y = true.
+Lemma BigQpowerth :
+ power_theory 1 BigQ.mul BigQ.eq Z_of_N BigQ.power.
Proof.
-intros; generalize (BigQ.spec_eq_bool x y).
-destruct BigQ.eq_bool; auto.
+constructor. intros. BigQ.qify.
+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.
-(* TODO : improve later the detection of constants ... *)
+Ltac isBigQcst t :=
+ match t with
+ | BigQ.Qz ?t => isBigZcst t
+ | BigQ.Qq ?n ?d => match isBigZcst n with
+ | true => isBigNcst d
+ | false => constr:false
+ end
+ | BigQ.zero => constr:true
+ | BigQ.one => constr:true
+ | BigQ.minus_one => constr:true
+ | _ => constr:false
+ end.
Ltac BigQcst t :=
- match t with
- | BigQ.zero => BigQ.zero
- | BigQ.one => BigQ.one
- | BigQ.minus_one => BigQ.minus_one
- | _ => NotConstant
+ match isBigQcst t with
+ | true => constr:t
+ | false => constr:NotConstant
end.
Add Field BigQfield : BigQfieldth
- (decidable BigQ_eq_bool_correct,
- completeness BigQ_eq_bool_complete,
+ (decidable BigQ.eqb_correct,
+ completeness BigQ.eqb_complete,
constants [BigQcst],
power_tac BigQpowerth [Qpow_tac]).
-Section Examples.
+Section TestField.
Let ex1 : forall x y z, (x+y)*z == (x*z)+(y*z).
intros.
ring.
Qed.
-Let ex8 : forall x, x ^ 1 == x.
+Let ex8 : forall x, x ^ 2 == x*x.
intro.
ring.
Qed.
-Let ex10 : forall x y, ~(y==BigQ.zero) -> (x/y)*y == x.
+Let ex10 : forall x y, y!=0 -> (x/y)*y == x.
intros.
field.
auto.
Qed.
-End Examples. \ No newline at end of file
+End TestField.
+
+(** [BigQ] can also benefit from an "order" tactic *)
+
+Module BigQ_Order := !OrdersTac.MakeOrderTac BigQ.
+Ltac bigQ_order := BigQ_Order.order.
+
+Section TestOrder.
+Let test : forall x y : bigQ, x<=y -> y<=x -> x==y.
+Proof. bigQ_order. Qed.
+End TestOrder.
+
+(** We can also reason by switching to QArith thanks to tactic
+ BigQ.qify. *)
+
+Section TestQify.
+Let test : forall x : bigQ, 0+x == 1*x.
+Proof. intro x. BigQ.qify. ring. Qed.
+End TestQify.
diff --git a/theories/Numbers/Rational/BigQ/QMake.v b/theories/Numbers/Rational/BigQ/QMake.v
index 494420bd..407f7b90 100644
--- a/theories/Numbers/Rational/BigQ/QMake.v
+++ b/theories/Numbers/Rational/BigQ/QMake.v
@@ -5,15 +5,20 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
-(************************************************************************)
-(*i $Id: QMake.v 11208 2008-07-04 16:57:46Z letouzey $ i*)
+(** * QMake : a generic efficient implementation of rational numbers *)
+
+(** Initial authors : Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
Require Import BigNumPrelude ROmega.
-Require Import QArith Qcanon Qpower.
+Require Import QArith Qcanon Qpower Qminmax.
Require Import NSig ZSig QSig.
+(** We will build rationals out of an implementation of integers [ZType]
+ for numerators and an implementation of natural numbers [NType] for
+ 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.
@@ -28,27 +33,27 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
number y interpreted as x/y. The pairs (x,0) and (0,y) are all
interpreted as 0. *)
- Inductive t_ :=
+ Inductive t_ :=
| Qz : Z.t -> t_
| Qq : Z.t -> N.t -> t_.
Definition t := t_.
- (** Specification with respect to [QArith] *)
+ (** Specification with respect to [QArith] *)
- Open Local Scope Q_scope.
+ Local Open Scope Q_scope.
Definition of_Z x: t := Qz (Z.of_Z x).
- Definition of_Q (q:Q) : t :=
- let (x,y) := q in
- match y with
+ 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))
end.
- Definition to_Q (q: t) :=
- match q with
+ Definition to_Q (q: t) :=
+ match q with
| Qz x => Z.to_Z x # 1
| Qq x y => if N.eq_bool y N.zero then 0
else Z.to_Z x # Z2P (N.to_Z y)
@@ -56,17 +61,56 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
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.
+ Proof.
+ intros x; rewrite N.spec_0; generalize (N.spec_pos x). romega.
+ Qed.
+(*
+ Lemma if_fun_commut : forall A B (f:A->B)(b:bool) a a',
+ f (if b then a else a') = if b then f a else f a'.
+ Proof. now destruct b. Qed.
+
+ Lemma if_fun_commut' : forall A B C D (f:A->B)(b:{C}+{D}) a a',
+ f (if b then a else a') = if b then f a else f a'.
+ Proof. now destruct b. Qed.
+*)
+ Ltac destr_eqb :=
+ match goal with
+ | |- context [Z.eq_bool ?x ?y] =>
+ rewrite (Z.spec_eq_bool x y);
+ generalize (Zeq_bool_if (Z.to_Z x) (Z.to_Z y));
+ case (Zeq_bool (Z.to_Z x) (Z.to_Z y));
+ destr_eqb
+ | |- context [N.eq_bool ?x ?y] =>
+ rewrite (N.spec_eq_bool x y);
+ generalize (Zeq_bool_if (N.to_Z x) (N.to_Z y));
+ case (Zeq_bool (N.to_Z x) (N.to_Z y));
+ [ | let H:=fresh "H" in
+ try (intro H;generalize (N_to_Z_pos _ H); clear H)];
+ destr_eqb
+ | _ => idtac
+ 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
+ spec_Z_of_N spec_Zabs_N
+ : nz.
+ Ltac nzsimpl := autorewrite with nz in *.
+
+ Ltac qsimpl := try red; unfold to_Q; simpl; intros;
+ destr_eqb; simpl; nzsimpl; intros;
+ rewrite ?Z2P_correct 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.
- generalize (N.spec_eq_bool (N.of_N (Npos y~1)) N.zero);
- case N.eq_bool; auto; rewrite N.spec_0.
- rewrite N.spec_of_N; intros; discriminate.
- rewrite N.spec_of_N; auto.
- generalize (N.spec_eq_bool (N.of_N (Npos y~0)) N.zero);
- case N.eq_bool; auto; rewrite N.spec_0.
- rewrite N.spec_of_N; intros; discriminate.
- rewrite N.spec_of_N; auto.
+ intros(x,y); destruct y; simpl; rewrite ?Z.spec_of_Z; auto;
+ destr_eqb; now rewrite ?N.spec_0, ?N.spec_of_N.
Qed.
Theorem spec_of_Q: forall q: Q, [of_Q q] == q.
@@ -82,131 +126,96 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Lemma spec_0: [zero] == 0.
Proof.
- simpl; rewrite Z.spec_0; reflexivity.
+ simpl. nzsimpl. reflexivity.
Qed.
Lemma spec_1: [one] == 1.
Proof.
- simpl; rewrite Z.spec_1; reflexivity.
+ simpl. nzsimpl. reflexivity.
Qed.
Lemma spec_m1: [minus_one] == -(1).
Proof.
- simpl; rewrite Z.spec_m1; reflexivity.
+ simpl. nzsimpl. reflexivity.
Qed.
Definition compare (x y: t) :=
match x, y with
| Qz zx, Qz zy => Z.compare zx zy
- | Qz zx, Qq ny dy =>
+ | Qz zx, Qq ny dy =>
if N.eq_bool dy N.zero then Z.compare zx Z.zero
else Z.compare (Z.mul zx (Z_of_N dy)) ny
- | Qq nx dx, Qz zy =>
- if N.eq_bool dx N.zero then Z.compare Z.zero zy
+ | Qq nx dx, Qz zy =>
+ if N.eq_bool dx N.zero then Z.compare Z.zero zy
else Z.compare nx (Z.mul zy (Z_of_N dx))
| Qq nx dx, Qq ny dy =>
match N.eq_bool dx N.zero, N.eq_bool dy N.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))
+ | false, false => Z.compare (Z.mul nx (Z_of_N dy))
(Z.mul ny (Z_of_N dx))
end
end.
- Lemma Zcompare_spec_alt :
- forall z z', Z.compare z z' = (Z.to_Z z ?= Z.to_Z z')%Z.
+ Theorem spec_compare: forall q1 q2, (compare q1 q2) = ([q1] ?= [q2]).
Proof.
- intros; generalize (Z.spec_compare z z'); destruct Z.compare; auto.
- intro H; rewrite H; symmetry; apply Zcompare_refl.
+ intros [z1 | x1 y1] [z2 | x2 y2];
+ unfold Qcompare, compare; qsimpl.
Qed.
-
- Lemma Ncompare_spec_alt :
- forall n n', N.compare n n' = (N.to_Z n ?= N.to_Z n')%Z.
+
+ Definition lt n m := [n] < [m].
+ Definition le n m := [n] <= [m].
+
+ 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.
+
+ Lemma spec_min : forall n m, [min n m] == Qmin [n] [m].
Proof.
- intros; generalize (N.spec_compare n n'); destruct N.compare; auto.
- intro H; rewrite H; symmetry; apply Zcompare_refl.
+ unfold min, Qmin, GenericMinMax.gmin. intros.
+ rewrite spec_compare; destruct Qcompare; auto with qarith.
Qed.
- Lemma N_to_Z2P : forall n, N.to_Z n <> 0%Z ->
- Zpos (Z2P (N.to_Z n)) = N.to_Z n.
+ Lemma spec_max : forall n m, [max n m] == Qmax [n] [m].
Proof.
- intros; apply Z2P_correct.
- generalize (N.spec_pos n); romega.
+ unfold max, Qmax, GenericMinMax.gmax. intros.
+ rewrite spec_compare; destruct Qcompare; auto with qarith.
Qed.
- 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
- Zcompare_spec_alt Ncompare_spec_alt
- Z.spec_add N.spec_add Z.spec_mul N.spec_mul
- Z.spec_gcd N.spec_gcd Zgcd_Zabs
- spec_Z_of_N spec_Zabs_N
- : nz.
- Ltac nzsimpl := autorewrite with nz in *.
-
- Ltac destr_neq_bool := repeat
- (match goal with |- context [N.eq_bool ?x ?y] =>
- generalize (N.spec_eq_bool x y); case N.eq_bool
- end).
-
- Ltac destr_zeq_bool := repeat
- (match goal with |- context [Z.eq_bool ?x ?y] =>
- generalize (Z.spec_eq_bool x y); case Z.eq_bool
- end).
-
- Ltac simpl_ndiv := rewrite N.spec_div by (nzsimpl; romega).
- Tactic Notation "simpl_ndiv" "in" "*" :=
- rewrite N.spec_div in * by (nzsimpl; romega).
-
- Ltac simpl_zdiv := rewrite Z.spec_div by (nzsimpl; romega).
- Tactic Notation "simpl_zdiv" "in" "*" :=
- rewrite Z.spec_div in * by (nzsimpl; romega).
-
- Ltac qsimpl := try red; unfold to_Q; simpl; intros;
- destr_neq_bool; destr_zeq_bool; simpl; nzsimpl; auto; intros.
+ Definition eq_bool n m :=
+ match compare n m with Eq => true | _ => false end.
- Theorem spec_compare: forall q1 q2, (compare q1 q2) = ([q1] ?= [q2]).
+ Theorem spec_eq_bool: forall x y, eq_bool x y = Qeq_bool [x] [y].
Proof.
- intros [z1 | x1 y1] [z2 | x2 y2];
- unfold Qcompare, compare; qsimpl; rewrite ! N_to_Z2P; auto.
+ intros. unfold eq_bool. rewrite spec_compare. reflexivity.
Qed.
- Definition lt n m := compare n m = Lt.
- Definition le n m := compare n m <> Gt.
- 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.
+ (** [check_int] : is a reduced fraction [n/d] in fact a integer ? *)
- Definition eq_bool n m :=
- match compare n m with Eq => true | _ => false end.
+ Definition check_int n d :=
+ match N.compare N.one d with
+ | Lt => Qq n d
+ | Eq => Qz n
+ | Gt => zero (* n/0 encodes 0 *)
+ end.
- Theorem spec_eq_bool: forall x y,
- if eq_bool x y then [x] == [y] else ~([x] == [y]).
+ Theorem strong_spec_check_int : forall n d, [check_int n d] = [Qq n d].
Proof.
- intros.
- unfold eq_bool.
- rewrite spec_compare.
- generalize (Qeq_alt [x] [y]).
- destruct Qcompare.
- intros H; rewrite H; auto.
- intros H H'; rewrite H in H'; discriminate.
- intros H H'; rewrite H in H'; discriminate.
+ intros; unfold check_int.
+ nzsimpl.
+ destr_zcompare.
+ simpl. rewrite <- H; qsimpl. congruence.
+ reflexivity.
+ qsimpl. exfalso; romega.
Qed.
(** Normalisation function *)
Definition norm n d : t :=
- let gcd := N.gcd (Zabs_N n) d in
+ let gcd := N.gcd (Zabs_N n) d in
match N.compare N.one gcd with
- | Lt =>
- let n := Z.div n (Z_of_N gcd) in
- let d := N.div d gcd in
- match N.compare d N.one with
- | Gt => Qq n d
- | Eq => Qz n
- | Lt => zero
- end
- | Eq => Qq n d
+ | Lt => check_int (Z.div n (Z_of_N gcd)) (N.div d gcd)
+ | Eq => check_int n d
| Gt => zero (* gcd = 0 => both numbers are 0 *)
end.
@@ -217,29 +226,16 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
assert (Hq := N.spec_pos q).
nzsimpl.
destr_zcompare.
+ (* Eq *)
+ rewrite strong_spec_check_int; reflexivity.
+ (* Lt *)
+ rewrite strong_spec_check_int.
qsimpl.
-
- simpl_ndiv.
- destr_zcompare.
- qsimpl.
- rewrite H1 in *; rewrite Zdiv_0_l in H0; discriminate.
- rewrite N_to_Z2P; auto.
- simpl_zdiv; nzsimpl.
- rewrite Zgcd_div_swap0, H0; romega.
-
- qsimpl.
- assert (0 < N.to_Z q / Zgcd (Z.to_Z p) (N.to_Z q))%Z.
- apply Zgcd_div_pos; romega.
- romega.
-
- qsimpl.
- simpl_ndiv in *; nzsimpl; romega.
- simpl_ndiv in *.
- rewrite H1, Zdiv_0_l in H2; elim H2; auto.
- rewrite 2 N_to_Z2P; auto.
- simpl_ndiv; simpl_zdiv; nzsimpl.
+ generalize (Zgcd_div_pos (Z.to_Z p) (N.to_Z q)). romega.
+ replace (N.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.
@@ -249,48 +245,37 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem strong_spec_norm : forall p q, [norm p q] = Qred [Qq p q].
Proof.
intros.
- replace (Qred [Qq p q]) with (Qred [norm p q]) by
+ replace (Qred [Qq p q]) with (Qred [norm p q]) by
(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).
nzsimpl.
- destr_zcompare.
+ destr_zcompare; rewrite ?strong_spec_check_int.
(* Eq *)
- simpl.
- destr_neq_bool; nzsimpl; simpl; auto.
- intros.
- rewrite N_to_Z2P; auto.
- (* Lt *)
- simpl_ndiv.
- destr_zcompare.
- qsimpl; auto.
qsimpl.
+ (* Lt *)
qsimpl.
- simpl_zdiv; nzsimpl.
- rewrite N_to_Z2P; auto.
- clear H1.
- simpl_ndiv; nzsimpl.
rewrite Zgcd_1_rel_prime.
destruct (Z_lt_le_dec 0 (N.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.
- rewrite Zdiv_0_l in H0; discriminate.
+ rewrite Zdiv_0_l in *; romega.
(* Gt *)
- simpl; auto.
+ simpl; auto with zarith.
Qed.
- (** Reduction function : producing irreducible fractions *)
+ (** Reduction function : producing irreducible fractions *)
- Definition red (x : t) : t :=
- match x with
+ Definition red (x : t) : t :=
+ match x with
| Qz z => x
| Qq n d => norm n d
end.
- Definition Reduced x := [red x] = [x].
+ Class Reduced x := is_reduced : [red x] = [x].
Theorem spec_red : forall x, [red x] == [x].
Proof.
@@ -304,21 +289,21 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Proof.
intros [ z | n d ].
unfold red.
- symmetry; apply Qred_identity; simpl; auto.
+ symmetry; apply Qred_identity; simpl; auto with zarith.
unfold red; apply strong_spec_norm.
Qed.
-
+
Definition add (x y: t): t :=
match x with
| Qz zx =>
match y with
| Qz zy => Qz (Z.add zx zy)
- | Qq ny dy =>
- if N.eq_bool dy N.zero then x
+ | Qq ny dy =>
+ if N.eq_bool dy N.zero then x
else Qq (Z.add (Z.mul zx (Z_of_N dy)) ny) dy
end
| Qq nx dx =>
- if N.eq_bool dx N.zero then y
+ if N.eq_bool dx N.zero then y
else match y with
| Qz zy => Qq (Z.add nx (Z.mul zy (Z_of_N dx))) dx
| Qq ny dy =>
@@ -332,19 +317,12 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_add : forall x y, [add x y] == [x] + [y].
Proof.
- intros [x | nx dx] [y | ny dy]; unfold Qplus; qsimpl.
- intuition.
- rewrite N_to_Z2P; auto.
- intuition.
- rewrite Pmult_1_r, N_to_Z2P; auto.
- intuition.
- rewrite Pmult_1_r, N_to_Z2P; auto.
- destruct (Zmult_integral _ _ H); intuition.
- rewrite Zpos_mult_morphism, 2 N_to_Z2P; auto.
- rewrite (Z2P_correct (N.to_Z dx * N.to_Z dy)); auto.
- apply Zmult_lt_0_compat.
- generalize (N.spec_pos dx); romega.
- generalize (N.spec_pos dy); romega.
+ 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.
Qed.
Definition add_norm (x y: t): t :=
@@ -352,12 +330,12 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
| Qz zx =>
match y with
| Qz zy => Qz (Z.add zx zy)
- | Qq ny dy =>
- if N.eq_bool dy N.zero then x
+ | Qq ny dy =>
+ if N.eq_bool dy N.zero then x
else norm (Z.add (Z.mul zx (Z_of_N dy)) ny) dy
end
| Qq nx dx =>
- if N.eq_bool dx N.zero then y
+ if N.eq_bool dx N.zero then y
else match y with
| Qz zy => norm (Z.add nx (Z.mul zy (Z_of_N dx))) dx
| Qq ny dy =>
@@ -372,26 +350,20 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_add_norm : forall x y, [add_norm x y] == [x] + [y].
Proof.
intros x y; rewrite <- spec_add.
- destruct x; destruct y; unfold add_norm, add;
- destr_neq_bool; auto using Qeq_refl, spec_norm.
+ destruct x; destruct y; unfold add_norm, add;
+ destr_eqb; auto using Qeq_refl, spec_norm.
Qed.
- Theorem strong_spec_add_norm : forall x y : t,
- Reduced x -> Reduced y -> Reduced (add_norm x y).
+ Instance strong_spec_add_norm x y
+ `(Reduced x, Reduced y) : Reduced (add_norm x y).
Proof.
unfold Reduced; intros.
rewrite strong_spec_red.
- rewrite <- (Qred_complete [add x y]);
+ rewrite <- (Qred_complete [add x y]);
[ | rewrite spec_add, spec_add_norm; apply Qeq_refl ].
rewrite <- strong_spec_red.
- destruct x as [zx|nx dx]; destruct y as [zy|ny dy].
- simpl in *; auto.
- simpl; intros.
- destr_neq_bool; nzsimpl; simpl; auto.
- simpl; intros.
- destr_neq_bool; nzsimpl; simpl; auto.
- simpl; intros.
- destr_neq_bool; nzsimpl; simpl; auto.
+ destruct x as [zx|nx dx]; destruct y as [zy|ny dy];
+ simpl; destr_eqb; nzsimpl; simpl; auto.
Qed.
Definition opp (x: t): t :=
@@ -404,7 +376,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Proof.
intros [z | x y]; simpl.
rewrite Z.spec_opp; auto.
- match goal with |- context[N.eq_bool ?X ?Y] =>
+ match goal with |- context[N.eq_bool ?X ?Y] =>
generalize (N.spec_eq_bool X Y); case N.eq_bool
end; auto; rewrite N.spec_0.
rewrite Z.spec_opp; auto.
@@ -415,7 +387,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
intros; rewrite strong_spec_opp; red; auto.
Qed.
- Theorem strong_spec_opp_norm : forall q, Reduced q -> Reduced (opp q).
+ Instance strong_spec_opp_norm q `(Reduced q) : Reduced (opp q).
Proof.
unfold Reduced; intros.
rewrite strong_spec_opp, <- H, !strong_spec_red, <- Qred_opp.
@@ -438,8 +410,8 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
rewrite spec_opp; ring.
Qed.
- Theorem strong_spec_sub_norm : forall x y,
- Reduced x -> Reduced y -> Reduced (sub_norm x y).
+ Instance strong_spec_sub_norm x y
+ `(Reduced x, Reduced y) : Reduced (sub_norm x y).
Proof.
intros.
unfold sub_norm.
@@ -458,35 +430,34 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
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, N_to_Z2P; auto.
- destruct (Zmult_integral _ _ H1); intuition.
- rewrite H0 in H1; elim H1; auto.
- rewrite H0 in H1; elim H1; auto.
- rewrite H in H1; nzsimpl; elim H1; auto.
- rewrite Zpos_mult_morphism, 2 N_to_Z2P; auto.
- rewrite (Z2P_correct (N.to_Z dx * N.to_Z dy)); auto.
- apply Zmult_lt_0_compat.
- generalize (N.spec_pos dx); omega.
- generalize (N.spec_pos dy); omega.
+ rewrite Pmult_1_r, Z2P_correct; auto.
+ destruct (Zmult_integral (N.to_Z dx) (N.to_Z dy)); intuition.
+ rewrite H0 in H1; auto with zarith.
+ rewrite H0 in H1; auto with zarith.
+ rewrite H in H1; nzsimpl; auto with zarith.
+ rewrite Zpos_mult_morphism, 2 Z2P_correct; auto.
Qed.
- Lemma norm_denum : forall n d,
- [if N.eq_bool d N.one then Qz n else Qq n d] == [Qq n d].
+ Definition norm_denum n d :=
+ if N.eq_bool d N.one then Qz n else Qq n d.
+
+ Lemma spec_norm_denum : forall n d,
+ [norm_denum n d] == [Qq n d].
Proof.
- intros; simpl; qsimpl.
- rewrite H0 in H; discriminate.
- rewrite N_to_Z2P, H0; auto with zarith.
+ unfold norm_denum; intros; simpl; qsimpl.
+ congruence.
+ rewrite H0 in *; auto with zarith.
Qed.
- Definition irred n d :=
+ Definition irred n d :=
let gcd := N.gcd (Zabs_N n) d in
- match N.compare gcd N.one with
+ match N.compare gcd N.one with
| Gt => (Z.div n (Z_of_N gcd), N.div d gcd)
| _ => (n, d)
end.
- Lemma spec_irred : forall n d, exists g,
- let (n',d') := irred n d in
+ 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.
Proof.
intros.
@@ -503,15 +474,15 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
exists (Zgcd (Z.to_Z n) (N.to_Z d)).
simpl.
split.
- simpl_zdiv; nzsimpl.
+ nzsimpl.
destruct (Zgcd_is_gcd (Z.to_Z n) (N.to_Z d)).
rewrite Zmult_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
- simpl_ndiv; nzsimpl.
+ nzsimpl.
destruct (Zgcd_is_gcd (Z.to_Z n) (N.to_Z d)).
rewrite Zmult_comm; symmetry; apply Zdivide_Zdiv_eq; auto with zarith.
Qed.
- Lemma spec_irred_zero : forall n d,
+ Lemma spec_irred_zero : forall n d,
(N.to_Z d = 0)%Z <-> (N.to_Z (snd (irred n d)) = 0)%Z.
Proof.
intros.
@@ -520,10 +491,10 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
nzsimpl; intros.
destr_zcompare; auto.
simpl.
- simpl_ndiv; nzsimpl.
+ nzsimpl.
rewrite H, Zdiv_0_l; auto.
nzsimpl; destr_zcompare; simpl; auto.
- simpl_ndiv; nzsimpl.
+ nzsimpl.
intros.
generalize (N.spec_pos d); intros.
destruct (N.to_Z d); auto.
@@ -535,8 +506,8 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
compute in H1; elim H1; auto.
Qed.
- Lemma strong_spec_irred : forall n d,
- (N.to_Z d <> 0%Z) ->
+ 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.
Proof.
unfold irred; intros.
@@ -546,7 +517,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
apply (Zgcd_inv_0_r (Z.to_Z n)).
generalize (Zgcd_is_pos (Z.to_Z n) (N.to_Z d)); romega.
- simpl_ndiv; simpl_zdiv; nzsimpl.
+ nzsimpl.
rewrite Zgcd_1_rel_prime.
apply Zis_gcd_rel_prime.
generalize (N.spec_pos d); romega.
@@ -554,89 +525,81 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
apply Zgcd_is_gcd; auto.
Qed.
- Definition mul_norm_Qz_Qq z n d :=
- if Z.eq_bool z Z.zero then zero
+ Definition mul_norm_Qz_Qq z n d :=
+ if Z.eq_bool z Z.zero then zero
else
let gcd := N.gcd (Zabs_N z) d in
match N.compare gcd N.one with
- | Gt =>
+ | Gt =>
let z := Z.div z (Z_of_N gcd) in
let d := N.div d gcd in
- if N.eq_bool d N.one then Qz (Z.mul z n) else Qq (Z.mul z n) d
+ norm_denum (Z.mul z n) d
| _ => Qq (Z.mul z n) d
end.
- Definition mul_norm (x y: t): t :=
+ Definition mul_norm (x y: t): t :=
match x, y with
| Qz zx, Qz zy => Qz (Z.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
- let d := N.mul dx dy in
- if N.eq_bool d N.one then Qz (Z.mul ny nx) else Qq (Z.mul ny nx) d
+ | 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)
end.
- Lemma spec_mul_norm_Qz_Qq : forall z n d,
+ Lemma spec_mul_norm_Qz_Qq : forall z n d,
[mul_norm_Qz_Qq z n d] == [Qq (Z.mul z n) d].
Proof.
intros z n d; unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
- destr_zeq_bool; intros Hz; nzsimpl.
+ destr_eqb; nzsimpl; intros Hz.
qsimpl; rewrite Hz; auto.
- assert (Hd := N.spec_pos d).
- destruct Z_le_gt_dec.
+ destruct Z_le_gt_dec; intros.
qsimpl.
- rewrite norm_denum.
+ rewrite spec_norm_denum.
qsimpl.
- simpl_ndiv in *; nzsimpl.
- rewrite (Zdiv_gcd_zero _ _ H0 H) in z0; discriminate.
- simpl_ndiv in *; nzsimpl.
- rewrite H, Zdiv_0_l in H0; elim H0; auto.
- rewrite 2 N_to_Z2P; auto.
- simpl_ndiv; simpl_zdiv; nzsimpl.
- rewrite (Zmult_comm (Z.to_Z z)), <- 2 Zmult_assoc.
- rewrite <- Zgcd_div_swap0; auto with zarith; ring.
+ rewrite Zdiv_gcd_zero in z0; auto with zarith.
+ rewrite H in *. rewrite Zdiv_0_l in *; discriminate.
+ rewrite <- Zmult_assoc, (Zmult_comm (Z.to_Z n)), Zmult_assoc.
+ rewrite Zgcd_div_swap0; try romega.
+ ring.
Qed.
- Lemma strong_spec_mul_norm_Qz_Qq : forall z n d,
- Reduced (Qq n d) -> Reduced (mul_norm_Qz_Qq z n d).
+ Instance strong_spec_mul_norm_Qz_Qq z n d :
+ forall `(Reduced (Qq n d)), Reduced (mul_norm_Qz_Qq z n d).
Proof.
- unfold Reduced; intros z n d.
+ unfold Reduced.
rewrite 2 strong_spec_red, 2 Qred_iff.
simpl; nzsimpl.
- destr_neq_bool; intros Hd H; simpl in *; nzsimpl.
-
+ destr_eqb; intros Hd H; simpl in *; nzsimpl.
+
unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
- destr_zeq_bool; intros Hz; simpl; nzsimpl; simpl; auto.
+ destr_eqb; intros Hz; simpl; nzsimpl; simpl; auto.
destruct Z_le_gt_dec.
simpl; nzsimpl.
- destr_neq_bool; simpl; nzsimpl; auto.
- intros H'; elim H'; auto.
- destr_neq_bool; simpl; nzsimpl.
- simpl_ndiv; nzsimpl; rewrite Hd, Zdiv_0_l; intros; discriminate.
+ destr_eqb; simpl; nzsimpl; auto with zarith.
+ unfold norm_denum. destr_eqb; simpl; nzsimpl.
+ rewrite Hd, Zdiv_0_l; discriminate.
intros _.
- destr_neq_bool; simpl; nzsimpl; auto.
- simpl_ndiv; nzsimpl; rewrite Hd, Zdiv_0_l; intro H'; elim H'; auto.
+ destr_eqb; simpl; nzsimpl; auto.
+ nzsimpl; rewrite Hd, Zdiv_0_l; auto with zarith.
- rewrite N_to_Z2P in H; auto.
+ rewrite Z2P_correct in H; auto.
unfold mul_norm_Qz_Qq; nzsimpl; rewrite Zcompare_gt.
- destr_zeq_bool; intros Hz; simpl; nzsimpl; simpl; auto.
+ destr_eqb; intros Hz; simpl; nzsimpl; simpl; auto.
destruct Z_le_gt_dec as [H'|H'].
simpl; nzsimpl.
- destr_neq_bool; simpl; nzsimpl; auto.
+ destr_eqb; simpl; nzsimpl; auto.
intros.
- rewrite N_to_Z2P; auto.
+ rewrite Z2P_correct; 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.
- destr_neq_bool; simpl; nzsimpl; auto.
- simpl_ndiv; simpl_zdiv; nzsimpl.
- intros.
- destr_neq_bool; simpl; nzsimpl; auto.
- simpl_ndiv in *; nzsimpl.
- intros.
- rewrite Z2P_correct.
+ destr_eqb; simpl; nzsimpl; auto.
+ unfold norm_denum.
+ destr_eqb; nzsimpl; simpl; destr_eqb; simpl; auto.
+ intros; nzsimpl.
+ rewrite Z2P_correct; auto.
apply Zgcd_mult_rel_prime.
rewrite Zgcd_1_rel_prime.
apply Zis_gcd_rel_prime.
@@ -652,9 +615,6 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
rewrite <- Huv; rewrite Hd0 at 2; ring.
rewrite Hd0 at 1.
symmetry; apply Z_div_mult_full; auto with zarith.
- apply Zgcd_div_pos.
- generalize (N.spec_pos d); romega.
- generalize (Zgcd_is_pos (Z.to_Z z) (N.to_Z d)); romega.
Qed.
Theorem spec_mul_norm : forall x y, [mul_norm x y] == [x] * [y].
@@ -670,37 +630,31 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destruct (spec_irred ny dx) as (g' & Hg').
assert (Hz := spec_irred_zero nx dy).
assert (Hz':= spec_irred_zero ny dx).
- destruct irred as (n1,d1); destruct irred as (n2,d2).
+ destruct irred as (n1,d1); destruct irred as (n2,d2).
simpl snd in *; destruct Hg as [Hg1 Hg2]; destruct Hg' as [Hg1' Hg2'].
- rewrite norm_denum.
+ rewrite spec_norm_denum.
qsimpl.
- elim H; destruct (Zmult_integral _ _ H0) as [Eq|Eq].
- rewrite <- Hz' in Eq; rewrite Eq; simpl; auto.
- rewrite <- Hz in Eq; rewrite Eq; nzsimpl; auto.
+ destruct (Zmult_integral _ _ H0) as [Eq|Eq].
+ rewrite Eq in *; simpl in *.
+ rewrite <- Hg2' in *; auto with zarith.
+ rewrite Eq in *; simpl in *.
+ rewrite <- Hg2 in *; auto with zarith.
- elim H0; destruct (Zmult_integral _ _ H) as [Eq|Eq].
- rewrite Hz' in Eq; rewrite Eq; simpl; auto.
- rewrite Hz in Eq; rewrite Eq; nzsimpl; auto.
+ destruct (Zmult_integral _ _ H) as [Eq|Eq].
+ rewrite Hz' in Eq; rewrite Eq in *; auto with zarith.
+ rewrite Hz in Eq; rewrite Eq in *; auto with zarith.
- rewrite 2 Z2P_correct.
rewrite <- Hg1, <- Hg2, <- Hg1', <- Hg2'; ring.
-
- assert (0 <= N.to_Z d2 * N.to_Z d1)%Z
- by (apply Zmult_le_0_compat; apply N.spec_pos).
- romega.
- assert (0 <= N.to_Z dx * N.to_Z dy)%Z
- by (apply Zmult_le_0_compat; apply N.spec_pos).
- romega.
Qed.
- Theorem strong_spec_mul_norm : forall x y,
- Reduced x -> Reduced y -> Reduced (mul_norm x y).
+ Instance strong_spec_mul_norm x y :
+ forall `(Reduced x, Reduced y), Reduced (mul_norm x y).
Proof.
unfold Reduced; intros.
rewrite strong_spec_red, Qred_iff.
destruct x as [zx|nx dx]; destruct y as [zy|ny dy].
- simpl in *; auto.
+ simpl in *; auto with zarith.
simpl.
rewrite <- Qred_iff, <- strong_spec_red, strong_spec_mul_norm_Qz_Qq; auto.
simpl.
@@ -712,26 +666,27 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
assert (Hz':= spec_irred_zero ny dx).
assert (Hgc := strong_spec_irred nx dy).
assert (Hgc' := strong_spec_irred ny dx).
- destruct irred as (n1,d1); destruct irred as (n2,d2).
+ destruct irred as (n1,d1); destruct irred as (n2,d2).
simpl snd in *; destruct Hg as [Hg1 Hg2]; destruct Hg' as [Hg1' Hg2'].
- destr_neq_bool; simpl; nzsimpl; intros.
- apply Zis_gcd_gcd; auto with zarith; apply Zis_gcd_1.
- destr_neq_bool; simpl; nzsimpl; intros.
- auto.
+
+ unfold norm_denum; qsimpl.
+
+ assert (NEQ : N.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
+ (rewrite Hz'; intro EQ; rewrite EQ in *; romega).
+ specialize (Hgc' NEQ').
revert H H0.
rewrite 2 strong_spec_red, 2 Qred_iff; simpl.
- destr_neq_bool; simpl; nzsimpl; intros.
- rewrite Hz in H; rewrite H in H2; nzsimpl; elim H2; auto.
- rewrite Hz' in H0; rewrite H0 in H2; nzsimpl; elim H2; auto.
- rewrite Hz in H; rewrite H in H2; nzsimpl; elim H2; auto.
+ destr_eqb; simpl; nzsimpl; try romega; intros.
+ rewrite Z2P_correct in *; auto.
- rewrite N_to_Z2P in *; auto.
- rewrite Z2P_correct.
+ apply Zgcd_mult_rel_prime; rewrite Zgcd_comm;
+ apply Zgcd_mult_rel_prime; rewrite Zgcd_comm; auto.
- apply Zgcd_mult_rel_prime; rewrite Zgcd_sym;
- apply Zgcd_mult_rel_prime; rewrite Zgcd_sym; auto.
-
rewrite Zgcd_1_rel_prime in *.
apply bezout_rel_prime.
destruct (rel_prime_bezout _ _ H4) as [u v Huv].
@@ -743,21 +698,17 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destruct (rel_prime_bezout _ _ H3) as [u v Huv].
apply Bezout_intro with (u*g)%Z (v*g')%Z.
rewrite <- Huv, <- Hg2', <- Hg1. ring.
-
- assert (0 <= N.to_Z d2 * N.to_Z d1)%Z.
- apply Zmult_le_0_compat; apply N.spec_pos.
- romega.
Qed.
- Definition inv (x: t): t :=
+ Definition inv (x: t): t :=
match x with
- | Qz z =>
- match Z.compare Z.zero z with
+ | Qz z =>
+ match Z.compare Z.zero z with
| Eq => zero
| Lt => Qq Z.one (Zabs_N z)
| Gt => Qq Z.minus_one (Zabs_N z)
end
- | Qq n d =>
+ | Qq n d =>
match Z.compare Z.zero n with
| Eq => zero
| Lt => Qq (Z_of_N d) (Zabs_N n)
@@ -770,13 +721,13 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destruct x as [ z | n d ].
(* Qz z *)
simpl.
- rewrite Zcompare_spec_alt; destr_zcompare.
+ rewrite Z.spec_compare; destr_zcompare.
(* 0 = z *)
rewrite <- H.
simpl; nzsimpl; compute; auto.
(* 0 < z *)
simpl.
- destr_neq_bool; nzsimpl; [ intros; rewrite Zabs_eq in *; romega | intros _ ].
+ destr_eqb; nzsimpl; [ intros; rewrite Zabs_eq in *; romega | intros _ ].
set (z':=Z.to_Z z) in *; clearbody z'.
red; simpl.
rewrite Zabs_eq by romega.
@@ -784,7 +735,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
unfold Qinv; simpl; destruct z'; simpl; auto; discriminate.
(* 0 > z *)
simpl.
- destr_neq_bool; nzsimpl; [ intros; rewrite Zabs_non_eq in *; romega | intros _ ].
+ destr_eqb; nzsimpl; [ intros; rewrite Zabs_non_eq in *; romega | intros _ ].
set (z':=Z.to_Z z) in *; clearbody z'.
red; simpl.
rewrite Zabs_non_eq by romega.
@@ -792,14 +743,14 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
unfold Qinv; simpl; destruct z'; simpl; auto; discriminate.
(* Qq n d *)
simpl.
- rewrite Zcompare_spec_alt; destr_zcompare.
+ rewrite Z.spec_compare; destr_zcompare.
(* 0 = n *)
rewrite <- H.
simpl; nzsimpl.
- destr_neq_bool; intros; compute; auto.
+ destr_eqb; intros; compute; auto.
(* 0 < n *)
simpl.
- destr_neq_bool; nzsimpl; intros.
+ destr_eqb; nzsimpl; intros.
intros; rewrite Zabs_eq in *; romega.
intros; rewrite Zabs_eq in *; romega.
clear H1.
@@ -811,10 +762,10 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
red; simpl.
rewrite Z2P_correct by auto.
unfold Qinv; simpl; destruct n'; simpl; auto; try discriminate.
- rewrite Zpos_mult_morphism, N_to_Z2P; auto.
+ rewrite Zpos_mult_morphism, Z2P_correct; auto.
(* 0 > n *)
simpl.
- destr_neq_bool; nzsimpl; intros.
+ destr_eqb; nzsimpl; intros.
intros; rewrite Zabs_non_eq in *; romega.
intros; rewrite Zabs_non_eq in *; romega.
clear H1.
@@ -826,28 +777,28 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
rewrite Z2P_correct 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, N_to_Z2P; auto; ring.
+ rewrite T, Zpos_mult_morphism, Z2P_correct; auto; ring.
Qed.
- Definition inv_norm (x: t): t :=
+ Definition inv_norm (x: t): t :=
match x with
- | Qz z =>
- match Z.compare Z.zero z with
+ | Qz z =>
+ match Z.compare Z.zero z with
| Eq => zero
| Lt => Qq Z.one (Zabs_N z)
| Gt => Qq Z.minus_one (Zabs_N z)
end
- | Qq n d =>
- if N.eq_bool d N.zero then zero else
- match Z.compare Z.zero n with
+ | Qq n d =>
+ if N.eq_bool d N.zero then zero else
+ match Z.compare Z.zero n with
| Eq => zero
- | Lt =>
- match Z.compare n Z.one with
+ | Lt =>
+ match Z.compare n Z.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
+ | 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))
end
@@ -861,74 +812,72 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destruct x as [ z | n d ].
(* Qz z *)
simpl.
- rewrite Zcompare_spec_alt; destr_zcompare; auto with qarith.
+ rewrite Z.spec_compare; destr_zcompare; auto with qarith.
(* Qq n d *)
- simpl; nzsimpl; destr_neq_bool.
+ simpl; nzsimpl; destr_eqb.
destr_zcompare; simpl; auto with qarith.
- destr_neq_bool; nzsimpl; auto with qarith.
+ destr_eqb; nzsimpl; auto with qarith.
intros _ Hd; rewrite Hd; auto with qarith.
- destr_neq_bool; nzsimpl; auto with qarith.
+ destr_eqb; nzsimpl; auto with qarith.
intros _ Hd; rewrite Hd; auto with qarith.
(* 0 < n *)
destr_zcompare; auto with qarith.
destr_zcompare; nzsimpl; simpl; auto with qarith; intros.
- destr_neq_bool; nzsimpl; [ intros; rewrite Zabs_eq in *; romega | intros _ ].
+ destr_eqb; nzsimpl; [ intros; rewrite Zabs_eq in *; romega | intros _ ].
rewrite H0; auto with qarith.
romega.
(* 0 > n *)
destr_zcompare; nzsimpl; simpl; auto with qarith.
- destr_neq_bool; nzsimpl; [ intros; rewrite Zabs_non_eq in *; romega | intros _ ].
+ destr_eqb; nzsimpl; [ intros; rewrite Zabs_non_eq in *; romega | intros _ ].
rewrite H0; auto with qarith.
romega.
Qed.
- Theorem strong_spec_inv_norm : forall x, Reduced x -> Reduced (inv_norm x).
+ Instance strong_spec_inv_norm x : Reduced x -> Reduced (inv_norm x).
Proof.
- unfold Reduced.
+ unfold Reduced.
intros.
destruct x as [ z | n d ].
(* Qz *)
simpl; nzsimpl.
rewrite strong_spec_red, Qred_iff.
destr_zcompare; simpl; nzsimpl; auto.
- destr_neq_bool; nzsimpl; simpl; auto.
- destr_neq_bool; nzsimpl; simpl; auto.
+ destr_eqb; nzsimpl; simpl; auto.
+ destr_eqb; nzsimpl; simpl; auto.
(* Qq n d *)
rewrite strong_spec_red, Qred_iff in H; revert H.
simpl; nzsimpl.
- destr_neq_bool; nzsimpl; auto with qarith.
+ destr_eqb; nzsimpl; auto with qarith.
destr_zcompare; simpl; nzsimpl; auto; intros.
(* 0 < n *)
destr_zcompare; simpl; nzsimpl; auto.
- destr_neq_bool; nzsimpl; simpl; auto.
+ destr_eqb; nzsimpl; simpl; auto.
rewrite Zabs_eq; romega.
intros _.
rewrite strong_spec_norm; simpl; nzsimpl.
- destr_neq_bool; nzsimpl.
+ destr_eqb; nzsimpl.
rewrite Zabs_eq; romega.
intros _.
rewrite Qred_iff.
simpl.
rewrite Zabs_eq; auto with zarith.
- rewrite N_to_Z2P in *; auto.
- rewrite Z2P_correct; auto with zarith.
- rewrite Zgcd_sym; auto.
+ rewrite Z2P_correct in *; auto.
+ rewrite Zgcd_comm; auto.
(* 0 > n *)
- destr_neq_bool; nzsimpl; simpl; auto; intros.
+ destr_eqb; nzsimpl; simpl; auto; intros.
destr_zcompare; simpl; nzsimpl; auto.
- destr_neq_bool; nzsimpl.
+ destr_eqb; nzsimpl.
rewrite Zabs_non_eq; romega.
intros _.
rewrite strong_spec_norm; simpl; nzsimpl.
- destr_neq_bool; nzsimpl.
+ destr_eqb; nzsimpl.
rewrite Zabs_non_eq; romega.
intros _.
rewrite Qred_iff.
simpl.
- rewrite N_to_Z2P in *; auto.
- rewrite Z2P_correct; auto with zarith.
+ rewrite Z2P_correct in *; auto.
intros.
- rewrite Zgcd_sym, Zgcd_Zabs, Zgcd_sym.
+ rewrite Zgcd_comm, Zgcd_Zabs, Zgcd_comm.
apply Zis_gcd_gcd; auto with zarith.
apply Zis_gcd_minus.
rewrite Zopp_involutive, <- H1; apply Zgcd_is_gcd.
@@ -939,7 +888,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_div x y: [div x y] == [x] / [y].
Proof.
- intros x y; unfold div; rewrite spec_mul; auto.
+ unfold div; rewrite spec_mul; auto.
unfold Qdiv; apply Qmult_comp.
apply Qeq_refl.
apply spec_inv; auto.
@@ -949,14 +898,14 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_div_norm x y: [div_norm x y] == [x] / [y].
Proof.
- intros x y; unfold div_norm; rewrite spec_mul_norm; auto.
+ unfold div_norm; rewrite spec_mul_norm; auto.
unfold Qdiv; apply Qmult_comp.
apply Qeq_refl.
apply spec_inv_norm; auto.
Qed.
-
- Theorem strong_spec_div_norm : forall x y,
- Reduced x -> Reduced y -> Reduced (div_norm x y).
+
+ Instance strong_spec_div_norm x y
+ `(Reduced x, Reduced y) : Reduced (div_norm x y).
Proof.
intros; unfold div_norm.
apply strong_spec_mul_norm; auto.
@@ -974,15 +923,13 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
destruct x as [ z | n d ].
simpl; rewrite Z.spec_square; red; auto.
simpl.
- destr_neq_bool; nzsimpl; intros.
+ destr_eqb; nzsimpl; intros.
apply Qeq_refl.
rewrite N.spec_square in *; nzsimpl.
- contradict H; elim (Zmult_integral _ _ H0); auto.
+ elim (Zmult_integral _ _ H0); romega.
rewrite N.spec_square in *; nzsimpl.
- rewrite H in H0; simpl in H0; elim H0; auto.
- assert (0 < N.to_Z d)%Z by (generalize (N.spec_pos d); romega).
- clear H H0.
- rewrite Z.spec_square, N.spec_square.
+ rewrite H in H0; romega.
+ rewrite Z.spec_square, N.spec_square.
red; simpl.
rewrite Zpos_mult_morphism; rewrite !Z2P_correct; auto.
apply Zmult_lt_0_compat; auto.
@@ -993,7 +940,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
| Qz zx => Qz (Z.power_pos zx p)
| Qq nx dx => Qq (Z.power_pos nx p) (N.power_pos dx p)
end.
-
+
Theorem spec_power_pos : forall x p, [power_pos x p] == [x] ^ Zpos p.
Proof.
intros [ z | n d ] p; unfold power_pos.
@@ -1006,44 +953,42 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(* Qq *)
simpl.
rewrite Z.spec_power_pos.
- destr_neq_bool; nzsimpl; intros.
+ destr_eqb; nzsimpl; intros.
apply Qeq_sym; apply Qpower_positive_0.
rewrite N.spec_power_pos in *.
- assert (0 < N.to_Z d ^ ' p)%Z.
- apply Zpower_gt_0; auto with zarith.
- generalize (N.spec_pos d); romega.
+ assert (0 < N.to_Z d ^ ' p)%Z by
+ (apply Zpower_gt_0; auto with zarith).
romega.
rewrite N.spec_power_pos, H in *.
- rewrite Zpower_0_l in H0; [ elim H0; auto | discriminate ].
+ rewrite Zpower_0_l in H0; [romega|discriminate].
rewrite Qpower_decomp.
red; simpl; do 3 f_equal.
rewrite Z2P_correct by (generalize (N.spec_pos d); romega).
rewrite N.spec_power_pos. auto.
Qed.
- Theorem strong_spec_power_pos : forall x p,
- Reduced x -> Reduced (power_pos x p).
+ Instance strong_spec_power_pos x p `(Reduced x) : Reduced (power_pos x p).
Proof.
destruct x as [z | n d]; simpl; intros.
red; simpl; auto.
red; simpl; intros.
rewrite strong_spec_norm; simpl.
- destr_neq_bool; nzsimpl; intros.
+ destr_eqb; nzsimpl; intros.
simpl; auto.
rewrite Qred_iff.
revert H.
unfold Reduced; rewrite strong_spec_red, Qred_iff; simpl.
- destr_neq_bool; nzsimpl; simpl; intros.
+ destr_eqb; nzsimpl; simpl; intros.
rewrite N.spec_power_pos in H0.
- elim H0; rewrite H; rewrite Zpower_0_l; auto; discriminate.
- rewrite N_to_Z2P in *; auto.
+ rewrite H, Zpower_0_l in *; [romega|discriminate].
+ rewrite Z2P_correct in *; auto.
rewrite N.spec_power_pos, Z.spec_power_pos; auto.
rewrite Zgcd_1_rel_prime in *.
apply rel_prime_Zpower; auto with zarith.
Qed.
- Definition power (x : t) (z : Z) : t :=
- match z with
+ Definition power (x : t) (z : Z) : t :=
+ match z with
| Z0 => one
| Zpos p => power_pos x p
| Zneg p => inv (power_pos x p)
@@ -1058,8 +1003,8 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
rewrite spec_inv, spec_power_pos; apply Qeq_refl.
Qed.
- Definition power_norm (x : t) (z : Z) : t :=
- match z with
+ Definition power_norm (x : t) (z : Z) : t :=
+ match z with
| Z0 => one
| Zpos p => power_pos x p
| Zneg p => inv_norm (power_pos x p)
@@ -1074,7 +1019,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
rewrite spec_inv_norm, spec_power_pos; apply Qeq_refl.
Qed.
- Theorem strong_spec_power_norm : forall x z,
+ Instance strong_spec_power_norm x z :
Reduced x -> Reduced (power_norm x z).
Proof.
destruct z; simpl.
@@ -1087,7 +1032,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
(** Interaction with [Qcanon.Qc] *)
-
+
Open Scope Qc_scope.
Definition of_Qc q := of_Q (this q).
@@ -1102,7 +1047,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
unfold of_Qc; rewrite strong_spec_of_Q; auto.
Qed.
- Lemma strong_spec_of_Qc_bis : forall q, Reduced (of_Qc q).
+ Instance strong_spec_of_Qc_bis q : Reduced (of_Qc q).
Proof.
intros; red; rewrite strong_spec_red, strong_spec_of_Qc.
destruct q; simpl; auto.
@@ -1143,7 +1088,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_addc x y:
[[add x y]] = [[x]] + [[y]].
Proof.
- intros x y; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! ([x] + [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
@@ -1157,7 +1102,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_add_normc x y:
[[add_norm x y]] = [[x]] + [[y]].
Proof.
- intros x y; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! ([x] + [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
@@ -1168,7 +1113,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
apply Qplus_comp; apply Qeq_sym; apply Qred_correct.
Qed.
- Theorem spec_add_normc_bis : forall x y : Qc,
+ Theorem spec_add_normc_bis : forall x y : Qc,
[add_norm (of_Qc x) (of_Qc y)] = x+y.
Proof.
intros.
@@ -1180,18 +1125,18 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_subc x y: [[sub x y]] = [[x]] - [[y]].
Proof.
- intros x y; unfold sub; rewrite spec_addc; auto.
+ unfold sub; rewrite spec_addc; auto.
rewrite spec_oppc; ring.
Qed.
Theorem spec_sub_normc x y:
[[sub_norm x y]] = [[x]] - [[y]].
Proof.
- intros x y; unfold sub_norm; rewrite spec_add_normc; auto.
+ unfold sub_norm; rewrite spec_add_normc; auto.
rewrite spec_oppc; ring.
Qed.
- Theorem spec_sub_normc_bis : forall x y : Qc,
+ Theorem spec_sub_normc_bis : forall x y : Qc,
[sub_norm (of_Qc x) (of_Qc y)] = x-y.
Proof.
intros.
@@ -1199,13 +1144,13 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
rewrite strong_spec_red.
symmetry; apply (Qred_complete (x+(-y)%Qc)%Q).
rewrite spec_sub_norm, ! strong_spec_of_Qc.
- unfold Qcopp, Q2Qc; rewrite Qred_correct; auto with qarith.
+ unfold Qcopp, Q2Qc, this. rewrite Qred_correct ; auto with qarith.
Qed.
Theorem spec_mulc x y:
[[mul x y]] = [[x]] * [[y]].
Proof.
- intros x y; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! ([x] * [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
@@ -1219,7 +1164,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_mul_normc x y:
[[mul_norm x y]] = [[x]] * [[y]].
Proof.
- intros x y; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! ([x] * [y])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
@@ -1230,7 +1175,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
apply Qmult_comp; apply Qeq_sym; apply Qred_correct.
Qed.
- Theorem spec_mul_normc_bis : forall x y : Qc,
+ Theorem spec_mul_normc_bis : forall x y : Qc,
[mul_norm (of_Qc x) (of_Qc y)] = x*y.
Proof.
intros.
@@ -1243,7 +1188,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_invc x:
[[inv x]] = /[[x]].
Proof.
- intros x; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! (/[x])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
@@ -1257,7 +1202,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_inv_normc x:
[[inv_norm x]] = /[[x]].
Proof.
- intros x; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! (/[x])).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
@@ -1268,7 +1213,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
apply Qinv_comp; apply Qeq_sym; apply Qred_correct.
Qed.
- Theorem spec_inv_normc_bis : forall x : Qc,
+ Theorem spec_inv_normc_bis : forall x : Qc,
[inv_norm (of_Qc x)] = /x.
Proof.
intros.
@@ -1280,19 +1225,19 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_divc x y: [[div x y]] = [[x]] / [[y]].
Proof.
- intros x y; unfold div; rewrite spec_mulc; auto.
+ unfold div; rewrite spec_mulc; auto.
unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto.
- apply spec_invc; auto.
+ apply spec_invc; auto.
Qed.
Theorem spec_div_normc x y: [[div_norm x y]] = [[x]] / [[y]].
Proof.
- intros x y; unfold div_norm; rewrite spec_mul_normc; auto.
+ unfold div_norm; rewrite spec_mul_normc; auto.
unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto.
apply spec_inv_normc; auto.
Qed.
- Theorem spec_div_normc_bis : forall x y : Qc,
+ Theorem spec_div_normc_bis : forall x y : Qc,
[div_norm (of_Qc x) (of_Qc y)] = x/y.
Proof.
intros.
@@ -1300,12 +1245,12 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
rewrite strong_spec_red.
symmetry; apply (Qred_complete (x*(/y)%Qc)%Q).
rewrite spec_div_norm, ! strong_spec_of_Qc.
- unfold Qcinv, Q2Qc; rewrite Qred_correct; auto with qarith.
+ unfold Qcinv, Q2Qc, this; rewrite Qred_correct; auto with qarith.
Qed.
- Theorem spec_squarec x: [[square x]] = [[x]]^2.
+ Theorem spec_squarec x: [[square x]] = [[x]]^2.
Proof.
- intros x; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! ([x]^2)).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
@@ -1322,7 +1267,7 @@ Module Make (N:NType)(Z:ZType)(Import NZ:NType_ZType N Z) <: QType.
Theorem spec_power_posc x p:
[[power_pos x p]] = [[x]] ^ nat_of_P p.
Proof.
- intros x p; unfold to_Qc.
+ unfold to_Qc.
apply trans_equal with (!! ([x]^Zpos p)).
unfold Q2Qc.
apply Qc_decomp; intros _ _; unfold this.
diff --git a/theories/Numbers/Rational/SpecViaQ/QSig.v b/theories/Numbers/Rational/SpecViaQ/QSig.v
index be9b2d4e..10d0189a 100644
--- a/theories/Numbers/Rational/SpecViaQ/QSig.v
+++ b/theories/Numbers/Rational/SpecViaQ/QSig.v
@@ -6,9 +6,9 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: QSig.v 11207 2008-07-04 16:50:32Z letouzey $ i*)
+(*i $Id$ i*)
-Require Import QArith Qpower.
+Require Import QArith Qpower Qminmax Orders RelationPairs GenericMinMax.
Open Scope Q_scope.
@@ -23,75 +23,203 @@ Module Type QType.
Parameter t : Type.
Parameter to_Q : t -> Q.
- Notation "[ x ]" := (to_Q x).
+ Local Notation "[ x ]" := (to_Q x).
Definition eq x y := [x] == [y].
+ Definition lt x y := [x] < [y].
+ Definition le x y := [x] <= [y].
Parameter of_Q : Q -> t.
Parameter spec_of_Q: forall x, to_Q (of_Q x) == x.
+ Parameter red : t -> t.
+ Parameter compare : t -> t -> comparison.
+ Parameter eq_bool : t -> t -> bool.
+ Parameter max : t -> t -> t.
+ Parameter min : t -> t -> t.
Parameter zero : t.
Parameter one : t.
Parameter minus_one : t.
+ Parameter add : t -> t -> t.
+ Parameter sub : t -> t -> t.
+ Parameter opp : t -> t.
+ Parameter mul : t -> t -> t.
+ Parameter square : t -> t.
+ Parameter inv : t -> t.
+ Parameter div : t -> t -> t.
+ Parameter power : t -> Z -> t.
+ Parameter spec_red : forall x, [red x] == [x].
+ Parameter strong_spec_red : forall x, [red x] = Qred [x].
+ Parameter spec_compare : forall x y, compare x y = ([x] ?= [y]).
+ Parameter spec_eq_bool : forall x y, eq_bool x y = Qeq_bool [x] [y].
+ Parameter spec_max : forall x y, [max x y] == Qmax [x] [y].
+ Parameter spec_min : forall x y, [min x y] == Qmin [x] [y].
Parameter spec_0: [zero] == 0.
Parameter spec_1: [one] == 1.
Parameter spec_m1: [minus_one] == -(1).
+ Parameter spec_add: forall x y, [add x y] == [x] + [y].
+ Parameter spec_sub: forall x y, [sub x y] == [x] - [y].
+ Parameter spec_opp: forall x, [opp x] == - [x].
+ Parameter spec_mul: forall x y, [mul x y] == [x] * [y].
+ Parameter spec_square: forall x, [square x] == [x] ^ 2.
+ Parameter spec_inv : forall x, [inv x] == / [x].
+ Parameter spec_div: forall x y, [div x y] == [x] / [y].
+ Parameter spec_power: forall x z, [power x z] == [x] ^ z.
- Parameter compare : t -> t -> comparison.
+End QType.
- Parameter spec_compare : forall x y, compare x y = ([x] ?= [y]).
+(** NB: several of the above functions come with [..._norm] variants
+ that expect reduced arguments and return reduced results. *)
- Definition lt n m := compare n m = Lt.
- Definition le n m := compare n m <> Gt.
- 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.
+(** TODO : also speak of specifications via Qcanon ... *)
- Parameter eq_bool : t -> t -> bool.
-
- Parameter spec_eq_bool : forall x y,
- if eq_bool x y then [x]==[y] else ~([x]==[y]).
+Module Type QType_Notation (Import Q : QType).
+ Notation "[ x ]" := (to_Q x).
+ Infix "==" := eq (at level 70).
+ Notation "x != y" := (~x==y) (at level 70).
+ Infix "<=" := le.
+ Infix "<" := lt.
+ Notation "0" := zero.
+ Notation "1" := one.
+ Infix "+" := add.
+ Infix "-" := sub.
+ Infix "*" := mul.
+ Notation "- x" := (opp x).
+ Infix "/" := div.
+ Notation "/ x" := (inv x).
+ Infix "^" := power.
+End QType_Notation.
- Parameter red : t -> t.
-
- Parameter spec_red : forall x, [red x] == [x].
- Parameter strong_spec_red : forall x, [red x] = Qred [x].
+Module Type QType' := QType <+ QType_Notation.
- Parameter add : t -> t -> t.
- Parameter spec_add: forall x y, [add x y] == [x] + [y].
+Module QProperties (Import Q : QType').
- Parameter sub : t -> t -> t.
+(** Conversion to Q *)
- Parameter spec_sub: forall x y, [sub x y] == [x] - [y].
+Hint Rewrite
+ spec_red spec_compare spec_eq_bool spec_min spec_max
+ spec_add spec_sub spec_opp spec_mul spec_square spec_inv spec_div
+ spec_power : qsimpl.
+Ltac qify := unfold eq, lt, le in *; autorewrite with qsimpl;
+ try rewrite spec_0 in *; try rewrite spec_1 in *; try rewrite spec_m1 in *.
- Parameter opp : t -> t.
+(** NB: do not add [spec_0] in the autorewrite database. Otherwise,
+ after instanciation in BigQ, this lemma become convertible to 0=0,
+ and autorewrite loops. Idem for [spec_1] and [spec_m1] *)
- Parameter spec_opp: forall x, [opp x] == - [x].
+(** Morphisms *)
- Parameter mul : t -> t -> t.
+Ltac solve_wd1 := intros x x' Hx; qify; now rewrite Hx.
+Ltac solve_wd2 := intros x x' Hx y y' Hy; qify; now rewrite Hx, Hy.
- Parameter spec_mul: forall x y, [mul x y] == [x] * [y].
+Local Obligation Tactic := solve_wd2 || solve_wd1.
- Parameter square : t -> t.
+Instance : Measure to_Q.
+Instance eq_equiv : Equivalence eq.
- Parameter spec_square: forall x, [square x] == [x] ^ 2.
+Program Instance lt_wd : Proper (eq==>eq==>iff) lt.
+Program Instance le_wd : Proper (eq==>eq==>iff) le.
+Program Instance red_wd : Proper (eq==>eq) red.
+Program Instance compare_wd : Proper (eq==>eq==>Logic.eq) compare.
+Program Instance eq_bool_wd : Proper (eq==>eq==>Logic.eq) eq_bool.
+Program Instance min_wd : Proper (eq==>eq==>eq) min.
+Program Instance max_wd : Proper (eq==>eq==>eq) max.
+Program Instance add_wd : Proper (eq==>eq==>eq) add.
+Program Instance sub_wd : Proper (eq==>eq==>eq) sub.
+Program Instance opp_wd : Proper (eq==>eq) opp.
+Program Instance mul_wd : Proper (eq==>eq==>eq) mul.
+Program Instance square_wd : Proper (eq==>eq) square.
+Program Instance inv_wd : Proper (eq==>eq) inv.
+Program Instance div_wd : Proper (eq==>eq==>eq) div.
+Program Instance power_wd : Proper (eq==>Logic.eq==>eq) power.
- Parameter inv : t -> t.
+(** Let's implement [HasCompare] *)
- Parameter spec_inv : forall x, [inv x] == / [x].
+Lemma compare_spec : forall x y, CompSpec eq lt x y (compare x y).
+Proof. intros. qify. destruct (Qcompare_spec [x] [y]); auto. Qed.
- Parameter div : t -> t -> t.
+(** Let's implement [TotalOrder] *)
- Parameter spec_div: forall x y, [div x y] == [x] / [y].
+Definition lt_compat := lt_wd.
+Instance lt_strorder : StrictOrder lt.
- Parameter power : t -> Z -> t.
+Lemma le_lteq : forall x y, x<=y <-> x<y \/ x==y.
+Proof. intros. qify. apply Qle_lteq. Qed.
- Parameter spec_power: forall x z, [power x z] == [x] ^ z.
+Lemma lt_total : forall x y, x<y \/ x==y \/ y<x.
+Proof. intros. destruct (compare_spec x y); auto. Qed.
-End QType.
+(** Let's implement [HasEqBool] *)
-(** NB: several of the above functions come with [..._norm] variants
- that expect reduced arguments and return reduced results. *)
+Definition eqb := eq_bool.
-(** TODO : also speak of specifications via Qcanon ... *)
+Lemma eqb_eq : forall x y, eq_bool x y = true <-> x == y.
+Proof. intros. qify. apply Qeq_bool_iff. Qed.
+
+Lemma eqb_correct : forall x y, eq_bool x y = true -> x == y.
+Proof. now apply eqb_eq. Qed.
+
+Lemma eqb_complete : forall x y, x == y -> eq_bool x y = true.
+Proof. now apply eqb_eq. Qed.
+
+(** Let's implement [HasMinMax] *)
+
+Lemma max_l : forall x y, y<=x -> max x y == x.
+Proof. intros x y. qify. apply Qminmax.Q.max_l. Qed.
+
+Lemma max_r : forall x y, x<=y -> max x y == y.
+Proof. intros x y. qify. apply Qminmax.Q.max_r. Qed.
+
+Lemma min_l : forall x y, x<=y -> min x y == x.
+Proof. intros x y. qify. apply Qminmax.Q.min_l. Qed.
+
+Lemma min_r : forall x y, y<=x -> min x y == y.
+Proof. intros x y. qify. apply Qminmax.Q.min_r. Qed.
+
+(** Q is a ring *)
+
+Lemma add_0_l : forall x, 0+x == x.
+Proof. intros. qify. apply Qplus_0_l. Qed.
+
+Lemma add_comm : forall x y, x+y == y+x.
+Proof. intros. qify. apply Qplus_comm. Qed.
+
+Lemma add_assoc : forall x y z, x+(y+z) == x+y+z.
+Proof. intros. qify. apply Qplus_assoc. Qed.
+
+Lemma mul_1_l : forall x, 1*x == x.
+Proof. intros. qify. apply Qmult_1_l. Qed.
+
+Lemma mul_comm : forall x y, x*y == y*x.
+Proof. intros. qify. apply Qmult_comm. Qed.
+
+Lemma mul_assoc : forall x y z, x*(y*z) == x*y*z.
+Proof. intros. qify. apply Qmult_assoc. Qed.
+
+Lemma mul_add_distr_r : forall x y z, (x+y)*z == x*z + y*z.
+Proof. intros. qify. apply Qmult_plus_distr_l. Qed.
+
+Lemma sub_add_opp : forall x y, x-y == x+(-y).
+Proof. intros. qify. now unfold Qminus. Qed.
+
+Lemma add_opp_diag_r : forall x, x+(-x) == 0.
+Proof. intros. qify. apply Qplus_opp_r. Qed.
+
+(** Q is a field *)
+
+Lemma neq_1_0 : 1!=0.
+Proof. intros. qify. apply Q_apart_0_1. Qed.
+
+Lemma div_mul_inv : forall x y, x/y == x*(/y).
+Proof. intros. qify. now unfold Qdiv. Qed.
+
+Lemma mul_inv_diag_l : forall x, x!=0 -> /x * x == 1.
+Proof. intros x. qify. rewrite Qmult_comm. apply Qmult_inv_r. Qed.
+
+End QProperties.
+
+Module QTypeExt (Q : QType)
+ <: QType <: TotalOrder <: HasCompare Q <: HasMinMax Q <: HasEqBool Q
+ := Q <+ QProperties. \ No newline at end of file
diff --git a/theories/Numbers/vo.itarget b/theories/Numbers/vo.itarget
new file mode 100644
index 00000000..175a15e9
--- /dev/null
+++ b/theories/Numbers/vo.itarget
@@ -0,0 +1,70 @@
+BigNumPrelude.vo
+Cyclic/Abstract/CyclicAxioms.vo
+Cyclic/Abstract/NZCyclic.vo
+Cyclic/DoubleCyclic/DoubleAdd.vo
+Cyclic/DoubleCyclic/DoubleBase.vo
+Cyclic/DoubleCyclic/DoubleCyclic.vo
+Cyclic/DoubleCyclic/DoubleDivn1.vo
+Cyclic/DoubleCyclic/DoubleDiv.vo
+Cyclic/DoubleCyclic/DoubleLift.vo
+Cyclic/DoubleCyclic/DoubleMul.vo
+Cyclic/DoubleCyclic/DoubleSqrt.vo
+Cyclic/DoubleCyclic/DoubleSub.vo
+Cyclic/DoubleCyclic/DoubleType.vo
+Cyclic/Int31/Int31.vo
+Cyclic/Int31/Cyclic31.vo
+Cyclic/Int31/Ring31.vo
+Cyclic/ZModulo/ZModulo.vo
+Integer/Abstract/ZAddOrder.vo
+Integer/Abstract/ZAdd.vo
+Integer/Abstract/ZAxioms.vo
+Integer/Abstract/ZBase.vo
+Integer/Abstract/ZLt.vo
+Integer/Abstract/ZMulOrder.vo
+Integer/Abstract/ZMul.vo
+Integer/Abstract/ZSgnAbs.vo
+Integer/Abstract/ZProperties.vo
+Integer/Abstract/ZDivFloor.vo
+Integer/Abstract/ZDivTrunc.vo
+Integer/Abstract/ZDivEucl.vo
+Integer/BigZ/BigZ.vo
+Integer/BigZ/ZMake.vo
+Integer/Binary/ZBinary.vo
+Integer/NatPairs/ZNatPairs.vo
+Integer/SpecViaZ/ZSig.vo
+Integer/SpecViaZ/ZSigZAxioms.vo
+NaryFunctions.vo
+NatInt/NZAddOrder.vo
+NatInt/NZAdd.vo
+NatInt/NZAxioms.vo
+NatInt/NZBase.vo
+NatInt/NZMulOrder.vo
+NatInt/NZMul.vo
+NatInt/NZOrder.vo
+NatInt/NZProperties.vo
+NatInt/NZDomain.vo
+NatInt/NZDiv.vo
+Natural/Abstract/NAddOrder.vo
+Natural/Abstract/NAdd.vo
+Natural/Abstract/NAxioms.vo
+Natural/Abstract/NBase.vo
+Natural/Abstract/NDefOps.vo
+Natural/Abstract/NIso.vo
+Natural/Abstract/NMulOrder.vo
+Natural/Abstract/NOrder.vo
+Natural/Abstract/NStrongRec.vo
+Natural/Abstract/NSub.vo
+Natural/Abstract/NProperties.vo
+Natural/Abstract/NDiv.vo
+Natural/BigN/BigN.vo
+Natural/BigN/Nbasic.vo
+Natural/BigN/NMake_gen.vo
+Natural/BigN/NMake.vo
+Natural/Binary/NBinary.vo
+Natural/Peano/NPeano.vo
+Natural/SpecViaZ/NSigNAxioms.vo
+Natural/SpecViaZ/NSig.vo
+NumPrelude.vo
+Rational/BigQ/BigQ.vo
+Rational/BigQ/QMake.vo
+Rational/SpecViaQ/QSig.vo