summaryrefslogtreecommitdiff
path: root/theories/Numbers/Cyclic/DoubleCyclic
diff options
context:
space:
mode:
Diffstat (limited to 'theories/Numbers/Cyclic/DoubleCyclic')
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v37
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v8
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v164
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v44
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v5
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v2
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v12
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v20
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v11
-rw-r--r--theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v5
10 files changed, 201 insertions, 107 deletions
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
index 1b035948..a7c28862 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -150,17 +150,17 @@ Section DoubleAdd.
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wB w_to_Z c) (at level 0, c 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)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Notation "[-[ c ]]" :=
(interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Variable spec_w_0 : [|w_0|] = 0.
Variable spec_w_1 : [|w_1|] = 1.
@@ -194,9 +194,9 @@ Section DoubleAdd.
Lemma spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]].
Proof.
- destruct x as [ |xh xl];simpl;trivial.
- destruct y as [ |yh yl];simpl. rewrite Z.add_0_r;trivial.
- replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]))
+ destruct x as [ |xh xl];trivial.
+ destruct y as [ |yh yl]. rewrite Z.add_0_r;trivial.
+ 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];
intros H;unfold interp_carry in H;rewrite <- H.
@@ -218,10 +218,11 @@ Section DoubleAdd.
Lemma spec_ww_add_c_cont : P x y (ww_add_c_cont x y).
Proof.
- destruct x as [ |xh xl];simpl;trivial.
+ destruct x as [ |xh xl];trivial.
apply spec_f0;trivial.
- destruct y as [ |yh yl];simpl.
- apply spec_f0;simpl;rewrite Z.add_0_r;trivial.
+ destruct y as [ |yh yl].
+ apply spec_f0;rewrite Z.add_0_r;trivial.
+ simpl.
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];
@@ -234,10 +235,10 @@ Section DoubleAdd.
as [h|h]; intros H1;unfold interp_carry in *.
apply spec_f0;simpl;rewrite H1. rewrite Z.mul_add_distr_r.
rewrite <- Z.add_assoc;rewrite H;ring.
- apply spec_f1. simpl;rewrite spec_w_WW;rewrite wwB_wBwB.
+ apply spec_f1. rewrite spec_w_WW;rewrite wwB_wBwB.
rewrite Z.add_assoc; rewrite Z.pow_2_r; rewrite <- Z.mul_add_distr_r.
rewrite Z.mul_1_l in H1;rewrite H1. rewrite Z.mul_add_distr_r.
- rewrite <- Z.add_assoc;rewrite H;ring.
+ rewrite <- Z.add_assoc;rewrite H; simpl; ring.
Qed.
End Cont.
@@ -245,11 +246,11 @@ Section DoubleAdd.
Lemma spec_ww_add_carry_c :
forall x y, [+[ww_add_carry_c x y]] = [[x]] + [[y]] + 1.
Proof.
- destruct x as [ |xh xl];intro y;simpl.
+ destruct x as [ |xh xl];intro y.
exact (spec_ww_succ_c y).
- destruct y as [ |yh yl];simpl.
+ destruct y as [ |yh yl].
rewrite Z.add_0_r;exact (spec_ww_succ_c (WW xh xl)).
- 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)
as [l|l];intros H;unfold interp_carry in H;rewrite <- H.
@@ -281,7 +282,7 @@ Section DoubleAdd.
Lemma spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB.
Proof.
- destruct x as [ |xh xl];intros y;simpl.
+ destruct x as [ |xh xl];intros y.
rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial.
destruct y as [ |yh yl].
change [[W0]] with 0;rewrite Z.add_0_r.
@@ -299,7 +300,7 @@ Section DoubleAdd.
Lemma spec_ww_add_carry :
forall x y, [[ww_add_carry x y]] = ([[x]] + [[y]] + 1) mod wwB.
Proof.
- destruct x as [ |xh xl];intros y;simpl.
+ destruct x as [ |xh xl];intros y.
exact (spec_ww_succ y).
destruct y as [ |yh yl].
change [[W0]] with 0;rewrite Z.add_0_r. exact (spec_ww_succ (WW xh xl)).
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
index 41a1d8ba..e68cd033 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -149,9 +149,9 @@ Section DoubleBase.
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 ]]" :=
- (interp_carry 1 wwB ww_to_Z c) (at level 0, x at level 99).
+ (interp_carry 1 wwB ww_to_Z c) (at level 0, c at level 99).
Notation "[-[ c ]]" :=
- (interp_carry (-1) wwB ww_to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wwB ww_to_Z c) (at level 0, c at level 99).
Notation "[! n | x !]" := (double_to_Z n x) (at level 0, x at level 99).
Variable spec_w_0 : [|w_0|] = 0.
@@ -287,7 +287,7 @@ Section DoubleBase.
Lemma double_wB_wwB : forall n, double_wB n * double_wB n = double_wB (S n).
Proof.
intros n;unfold double_wB;simpl.
- unfold base. rewrite Pshiftl_nat_S, (Pos2Z.inj_xO (_ << _)).
+ unfold base. rewrite (Pos2Z.inj_xO (_ << _)).
replace (2 * Zpos (w_digits << n)) with
(Zpos (w_digits << n) + Zpos (w_digits << n)) by ring.
symmetry; apply Zpower_exp;intro;discriminate.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
index e207d7eb..e137349e 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -283,6 +283,27 @@ Section Z_2nZ.
Eval lazy beta delta [ww_gcd] in
ww_gcd compare w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont.
+ Definition lor (x y : zn2z t) :=
+ match x, y with
+ | W0, _ => y
+ | _, W0 => x
+ | WW hx lx, WW hy ly => WW (ZnZ.lor hx hy) (ZnZ.lor lx ly)
+ end.
+
+ Definition land (x y : zn2z t) :=
+ match x, y with
+ | W0, _ => W0
+ | _, W0 => W0
+ | WW hx lx, WW hy ly => WW (ZnZ.land hx hy) (ZnZ.land lx ly)
+ end.
+
+ Definition lxor (x y : zn2z t) :=
+ match x, y with
+ | W0, _ => y
+ | _, W0 => x
+ | WW hx lx, WW hy ly => WW (ZnZ.lxor hx hy) (ZnZ.lxor lx ly)
+ end.
+
(* ** Record of operators on 2 words *)
Global Instance mk_zn2z_ops : ZnZ.Ops (zn2z t) | 1 :=
@@ -303,7 +324,10 @@ Section Z_2nZ.
pos_mod
is_even
sqrt2
- sqrt.
+ sqrt
+ lor
+ land
+ lxor.
Global Instance mk_zn2z_ops_karatsuba : ZnZ.Ops (zn2z t) | 2 :=
ZnZ.MkOps _ww_digits _ww_zdigits
@@ -323,10 +347,15 @@ Section Z_2nZ.
pos_mod
is_even
sqrt2
- sqrt.
+ sqrt
+ lor
+ land
+ lxor.
(* Proof *)
Context {specs : ZnZ.Specs ops}.
+
+ Create HintDb ZnZ.
Hint Resolve
ZnZ.spec_to_Z
@@ -370,24 +399,24 @@ Section Z_2nZ.
ZnZ.spec_sqrt
ZnZ.spec_WO
ZnZ.spec_OW
- ZnZ.spec_WW.
-
- Ltac wwauto := unfold ww_to_Z; auto.
+ ZnZ.spec_WW : ZnZ.
+
+ Ltac wwauto := unfold ww_to_Z; eauto with ZnZ.
Let wwB := base _ww_digits.
Notation "[| x |]" := (to_Z x) (at level 0, x at level 99).
Notation "[+| c |]" :=
- (interp_carry 1 wwB to_Z c) (at level 0, x at level 99).
+ (interp_carry 1 wwB to_Z c) (at level 0, c at level 99).
Notation "[-| c |]" :=
- (interp_carry (-1) wwB to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wwB to_Z c) (at level 0, c at level 99).
Notation "[[ x ]]" := (zn2z_to_Z wwB to_Z x) (at level 0, x at level 99).
Let spec_ww_to_Z : forall x, 0 <= [| x |] < wwB.
- Proof. refine (spec_ww_to_Z w_digits w_to_Z _);auto. Qed.
+ Proof. refine (spec_ww_to_Z w_digits w_to_Z _); wwauto. Qed.
Let spec_ww_of_pos : forall p,
Zpos p = (Z.of_N (fst (ww_of_pos p)))*wwB + [|(snd (ww_of_pos p))|].
@@ -411,15 +440,15 @@ Section Z_2nZ.
Proof. reflexivity. Qed.
Let spec_ww_1 : [|ww_1|] = 1.
- Proof. refine (spec_ww_1 w_0 w_1 w_digits w_to_Z _ _);auto. Qed.
+ Proof. refine (spec_ww_1 w_0 w_1 w_digits w_to_Z _ _);wwauto. Qed.
Let spec_ww_Bm1 : [|ww_Bm1|] = wwB - 1.
- Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed.
+ Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);wwauto. Qed.
Let spec_ww_compare :
forall x y, compare x y = Z.compare [|x|] [|y|].
Proof.
- refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);auto.
+ refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);wwauto.
Qed.
Let spec_ww_eq0 : forall x, eq0 x = true -> [|x|] = 0.
@@ -428,14 +457,14 @@ Section Z_2nZ.
Let spec_ww_opp_c : forall x, [-|opp_c x|] = -[|x|].
Proof.
refine(spec_ww_opp_c w_0 w_0 W0 w_opp_c w_opp_carry w_digits w_to_Z _ _ _ _);
- auto.
+ wwauto.
Qed.
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
w_digits w_to_Z _ _ _ _ _);
- auto.
+ wwauto.
Qed.
Let spec_ww_opp_carry : forall x, [|opp_carry x|] = wwB - [|x|] - 1.
@@ -446,7 +475,7 @@ Section Z_2nZ.
Let spec_ww_succ_c : forall x, [+|succ_c x|] = [|x|] + 1.
Proof.
- refine (spec_ww_succ_c w_0 w_0 ww_1 w_succ_c w_digits w_to_Z _ _ _ _);auto.
+ refine (spec_ww_succ_c w_0 w_0 ww_1 w_succ_c w_digits w_to_Z _ _ _ _);wwauto.
Qed.
Let spec_ww_add_c : forall x y, [+|add_c x y|] = [|x|] + [|y|].
@@ -468,7 +497,7 @@ Section Z_2nZ.
Let spec_ww_add : forall x y, [|add x y|] = ([|x|] + [|y|]) mod wwB.
Proof.
- refine (spec_ww_add w_add_c w_add w_add_carry w_digits w_to_Z _ _ _ _);auto.
+ refine (spec_ww_add w_add_c w_add w_add_carry w_digits w_to_Z _ _ _ _);wwauto.
Qed.
Let spec_ww_add_carry : forall x y, [|add_carry x y|]=([|x|]+[|y|]+1)mod wwB.
@@ -565,7 +594,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,
@@ -581,13 +610,14 @@ Section Z_2nZ.
Qed.
Let spec_low: forall x,
- w_to_Z (low x) = [|x|] mod wB.
+ w_to_Z (low x) = [|x|] mod wB.
intros x; case x; simpl low.
- unfold ww_to_Z, w_to_Z, w_0; rewrite ZnZ.spec_0; simpl; auto.
+ unfold ww_to_Z, w_to_Z, w_0; rewrite ZnZ.spec_0; simpl; wwauto.
intros xh xl; simpl.
rewrite Z.add_comm; rewrite Z_mod_plus; auto with zarith.
rewrite Zmod_small; auto with zarith.
- unfold wB, base; auto with zarith.
+ unfold wB, base; eauto with ZnZ zarith.
+ unfold wB, base; eauto with ZnZ zarith.
Qed.
Let spec_ww_digits:
@@ -605,7 +635,7 @@ Section Z_2nZ.
Proof.
refine (spec_ww_head00 w_0 w_0W
w_compare w_head0 w_add2 w_zdigits _ww_zdigits
- w_to_Z _ _ _ (eq_refl _ww_digits) _ _ _ _); auto.
+ w_to_Z _ _ _ (eq_refl _ww_digits) _ _ _ _); wwauto.
exact ZnZ.spec_head00.
exact ZnZ.spec_zdigits.
Qed.
@@ -688,7 +718,7 @@ refine
[|a|] = [|q|] * [|b|] + [|r|] /\
0 <= [|r|] < [|b|].
Proof.
- refine (spec_ww_div w_digits ww_1 compare div_gt w_to_Z _ _ _ _);auto.
+ refine (spec_ww_div w_digits ww_1 compare div_gt w_to_Z _ _ _ _);wwauto.
Qed.
Let spec_ww_mod_gt : forall a b,
@@ -708,7 +738,7 @@ refine
Let spec_ww_mod : forall a b, 0 < [|b|] -> [|mod_ a b|] = [|a|] mod [|b|].
Proof.
- refine (spec_ww_mod w_digits W0 compare mod_gt w_to_Z _ _ _);auto.
+ refine (spec_ww_mod w_digits W0 compare mod_gt w_to_Z _ _ _);wwauto.
Qed.
Let spec_ww_gcd_gt : forall a b, [|a|] > [|b|] ->
@@ -716,7 +746,7 @@ refine
Proof.
refine (@spec_ww_gcd_gt t w_digits W0 w_to_Z _
w_0 w_0 w_eq0 w_gcd_gt _ww_digits
- _ gcd_gt_fix _ _ _ _ gcd_cont _);auto.
+ _ gcd_gt_fix _ _ _ _ gcd_cont _);wwauto.
refine (@spec_ww_gcd_gt_aux t 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
@@ -725,13 +755,13 @@ refine
exact ZnZ.spec_zdigits.
exact spec_ww_add_mul_div.
refine (@spec_gcd_cont t w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare
- _ _);auto.
+ _ _);wwauto.
Qed.
Let spec_ww_gcd : forall a b, Zis_gcd [|a|] [|b|] [|gcd a b|].
Proof.
refine (@spec_ww_gcd t w_digits W0 compare w_to_Z _ _ w_0 w_0 w_eq0 w_gcd_gt
- _ww_digits _ gcd_gt_fix _ _ _ _ gcd_cont _);auto.
+ _ww_digits _ gcd_gt_fix _ _ _ _ gcd_cont _);wwauto.
refine (@spec_ww_gcd_gt_aux t 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
@@ -740,7 +770,7 @@ refine
exact ZnZ.spec_zdigits.
exact spec_ww_add_mul_div.
refine (@spec_gcd_cont t w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare
- _ _);auto.
+ _ _);wwauto.
Qed.
Let spec_ww_is_even : forall x,
@@ -779,7 +809,7 @@ refine
refine (@spec_ww_sqrt t 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.
+ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto.
exact ZnZ.spec_zdigits.
exact ZnZ.spec_more_than_1_digit.
exact ZnZ.spec_is_even.
@@ -787,6 +817,83 @@ refine
exact ZnZ.spec_sqrt2.
Qed.
+ Let wB_pos : 0 < wB.
+ Proof.
+ unfold wB, base; apply Z.pow_pos_nonneg; auto with zarith.
+ Qed.
+
+ Hint Transparent ww_to_Z.
+
+ Let ww_testbit_high n x y : Z.pos w_digits <= n ->
+ Z.testbit [|WW x y|] n =
+ Z.testbit (ZnZ.to_Z x) (n - Z.pos w_digits).
+ Proof.
+ intros Hn.
+ assert (E : ZnZ.to_Z x = [|WW x y|] / wB).
+ { simpl.
+ rewrite Z.div_add_l; eauto with ZnZ zarith.
+ now rewrite Z.div_small, Z.add_0_r; wwauto. }
+ rewrite E.
+ unfold wB, base. rewrite Z.div_pow2_bits.
+ - f_equal; auto with zarith.
+ - easy.
+ - auto with zarith.
+ Qed.
+
+ Let ww_testbit_low n x y : 0 <= n < Z.pos w_digits ->
+ Z.testbit [|WW x y|] n = Z.testbit (ZnZ.to_Z y) n.
+ Proof.
+ intros (Hn,Hn').
+ assert (E : ZnZ.to_Z y = [|WW x y|] mod wB).
+ { simpl; symmetry.
+ rewrite Z.add_comm, Z.mod_add; auto with zarith nocore.
+ apply Z.mod_small; eauto with ZnZ zarith. }
+ rewrite E.
+ unfold wB, base. symmetry. apply Z.mod_pow2_bits_low; auto.
+ Qed.
+
+ Let spec_lor x y : [|lor x y|] = Z.lor [|x|] [|y|].
+ Proof.
+ destruct x as [ |hx lx]. trivial.
+ destruct y as [ |hy ly]. now rewrite Z.lor_comm.
+ change ([|WW (ZnZ.lor hx hy) (ZnZ.lor lx ly)|] =
+ Z.lor [|WW hx lx|] [|WW hy ly|]).
+ apply Z.bits_inj'; intros n Hn.
+ rewrite Z.lor_spec.
+ destruct (Z.le_gt_cases (Z.pos w_digits) n) as [LE|GT].
+ - now rewrite !ww_testbit_high, ZnZ.spec_lor, Z.lor_spec.
+ - rewrite !ww_testbit_low; auto.
+ now rewrite ZnZ.spec_lor, Z.lor_spec.
+ Qed.
+
+ Let spec_land x y : [|land x y|] = Z.land [|x|] [|y|].
+ Proof.
+ destruct x as [ |hx lx]. trivial.
+ destruct y as [ |hy ly]. now rewrite Z.land_comm.
+ change ([|WW (ZnZ.land hx hy) (ZnZ.land lx ly)|] =
+ Z.land [|WW hx lx|] [|WW hy ly|]).
+ apply Z.bits_inj'; intros n Hn.
+ rewrite Z.land_spec.
+ destruct (Z.le_gt_cases (Z.pos w_digits) n) as [LE|GT].
+ - now rewrite !ww_testbit_high, ZnZ.spec_land, Z.land_spec.
+ - rewrite !ww_testbit_low; auto.
+ now rewrite ZnZ.spec_land, Z.land_spec.
+ Qed.
+
+ Let spec_lxor x y : [|lxor x y|] = Z.lxor [|x|] [|y|].
+ Proof.
+ destruct x as [ |hx lx]. trivial.
+ destruct y as [ |hy ly]. now rewrite Z.lxor_comm.
+ change ([|WW (ZnZ.lxor hx hy) (ZnZ.lxor lx ly)|] =
+ Z.lxor [|WW hx lx|] [|WW hy ly|]).
+ apply Z.bits_inj'; intros n Hn.
+ rewrite Z.lxor_spec.
+ destruct (Z.le_gt_cases (Z.pos w_digits) n) as [LE|GT].
+ - now rewrite !ww_testbit_high, ZnZ.spec_lxor, Z.lxor_spec.
+ - rewrite !ww_testbit_low; auto.
+ now rewrite ZnZ.spec_lxor, Z.lxor_spec.
+ Qed.
+
Global Instance mk_zn2z_specs : ZnZ.Specs mk_zn2z_ops.
Proof.
apply ZnZ.MkSpecs; auto.
@@ -816,6 +923,7 @@ refine
End Z_2nZ.
+
Section MulAdd.
Context {t : Type}{ops : ZnZ.Ops t}{specs : ZnZ.Specs ops}.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
index 08f05bbf..cd55f9d8 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -97,8 +97,7 @@ Section POS_MOD.
assert (HHHHH:= lt_0_wB w_digits).
assert (F0: forall x y, x - y + y = x); auto with zarith.
intros w1 p; case (spec_to_w_Z p); intros HH1 HH2.
- unfold ww_pos_mod; case w1.
- simpl; rewrite Zmod_small; split; auto with zarith.
+ unfold ww_pos_mod; case w1. reflexivity.
intros xh xl; rewrite spec_ww_compare.
case Z.compare_spec;
rewrite spec_w_0W; rewrite spec_zdigits; fold wB;
@@ -211,8 +210,7 @@ Section DoubleDiv32.
Variable w_div21 : w -> w -> w -> w*w.
Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w).
- Definition w_div32 a1 a2 a3 b1 b2 :=
- Eval lazy beta iota delta [ww_add_c_cont ww_add] in
+ Definition w_div32_body a1 a2 a3 b1 b2 :=
match w_compare a1 b1 with
| Lt =>
let (q,r) := w_div21 a1 a2 b1 in
@@ -233,6 +231,10 @@ Section DoubleDiv32.
| Gt => (w_0, W0) (* cas absurde *)
end.
+ Definition w_div32 a1 a2 a3 b1 b2 :=
+ Eval lazy beta iota delta [ww_add_c_cont ww_add w_div32_body] in
+ w_div32_body a1 a2 a3 b1 b2.
+
(* Proof *)
Variable w_digits : positive.
@@ -242,14 +244,14 @@ Section DoubleDiv32.
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wB w_to_Z c) (at level 0, c 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)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Variable spec_w_0 : [|w_0|] = 0.
@@ -312,26 +314,8 @@ Section DoubleDiv32.
assert (U:= lt_0_wB w_digits); assert (U1:= lt_0_wwB w_digits).
Spec_w_to_Z a1;Spec_w_to_Z a2;Spec_w_to_Z a3;Spec_w_to_Z b1;Spec_w_to_Z b2.
rewrite wwB_wBwB; rewrite Z.pow_2_r; rewrite Z.mul_assoc;rewrite <- Z.mul_add_distr_r.
- change (w_div32 a1 a2 a3 b1 b2) with
- match w_compare a1 b1 with
- | Lt =>
- let (q,r) := w_div21 a1 a2 b1 in
- 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
- (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
- (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.
+ change (w_div32 a1 a2 a3 b1 b2) with (w_div32_body a1 a2 a3 b1 b2).
+ unfold w_div32_body.
rewrite spec_compare. case Z.compare_spec; intro Hcmp.
simpl in Hlt.
rewrite Hcmp in Hlt;assert ([|a2|] < [|b2|]). omega.
@@ -520,7 +504,7 @@ Section DoubleDiv21.
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)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Variable spec_w_0 : [|w_0|] = 0.
Variable spec_to_Z : forall x, 0 <= [|x|] < wB.
@@ -782,7 +766,7 @@ Section DoubleDivGt.
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wB w_to_Z c) (at level 0, c at level 99).
Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99).
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
index 8e179ef6..6a1d741e 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -160,7 +160,7 @@ Section GENDIVN1.
Lemma p_lt_double_digits : forall n, [|p|] <= Zpos (w_digits << n).
Proof.
induction n;simpl. trivial.
- case (spec_to_Z p); rewrite Pshiftl_nat_S, Pos2Z.inj_xO;auto with zarith.
+ case (spec_to_Z p); rewrite Pos2Z.inj_xO;auto with zarith.
Qed.
Lemma spec_double_divn1_p : forall n r h l,
@@ -305,7 +305,6 @@ Section GENDIVN1.
Lemma spec_double_digits:forall n, Zpos w_digits <= Zpos (w_digits << n).
Proof.
induction n;simpl;auto with zarith.
- rewrite Pshiftl_nat_S.
change (Zpos (xO (w_digits << n))) with
(2*Zpos (w_digits << n)).
assert (0 < Zpos w_digits) by reflexivity.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
index 2d0cc0fb..ff9f50a5 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
index 1c0fc68a..537f557d 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -218,17 +218,17 @@ Section DoubleMul.
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wB w_to_Z c) (at level 0, c 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)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Notation "[-[ c ]]" :=
(interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Notation "[|| x ||]" :=
(zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99).
@@ -328,7 +328,7 @@ Section DoubleMul.
rewrite <- (Z.add_0_r ([|wc|]*wB));rewrite H;apply mult_add_ineq3;zarith.
simpl ww_to_Z in H1. assert (U:=spec_to_Z cch).
assert ([|wc|]*wB + [|cch|] <= 2*wB - 3).
- destruct (Z_le_gt_dec ([|wc|]*wB + [|cch|]) (2*wB - 3));trivial.
+ destruct (Z_le_gt_dec ([|wc|]*wB + [|cch|]) (2*wB - 3)) as [Hle|Hgt];trivial.
assert ([|xh|] * [|yl|] + [|xl|] * [|yh|] <= (2*wB - 4)*wB + 2).
ring_simplify ((2*wB - 4)*wB + 2).
assert (H4 := Zmult_lt_b _ _ _ (spec_to_Z xh) (spec_to_Z yl)).
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
index 149682f8..ab8c8617 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -185,17 +185,17 @@ Section DoubleSqrt.
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wB w_to_Z c) (at level 0, c 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)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Notation "[-[ c ]]" :=
(interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Notation "[|| x ||]" :=
(zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99).
@@ -266,8 +266,8 @@ Section DoubleSqrt.
if ww_is_even x then [[x]] mod 2 = 0 else [[x]] mod 2 = 1.
clear spec_more_than_1_digit.
intros x; case x; simpl ww_is_even.
+ reflexivity.
simpl.
- rewrite Zmod_small; auto with zarith.
intros w1 w2; simpl.
unfold base.
rewrite Zplus_mod; auto with zarith.
@@ -692,7 +692,7 @@ intros x; case x; simpl ww_is_even.
intros x y H; unfold ww_sqrt2.
repeat match goal with |- context[split ?x] =>
generalize (spec_split x); case (split x)
- end; simpl fst; simpl snd.
+ end; simpl @fst; simpl @snd.
intros w0 w1 Hw0 w2 w3 Hw1.
assert (U: wB/4 <= [|w2|]).
case (Z.le_gt_cases (wB / 4) [|w2|]); auto; intros H1.
@@ -761,7 +761,7 @@ intros x; case x; simpl ww_is_even.
auto.
split.
unfold zn2z_to_Z; rewrite <- Hw1.
- unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1.
+ unfold ww_to_Z, zn2z_to_Z in H1. rewrite H1.
rewrite <- Hw0.
match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U =>
transitivity ((X * wB) ^ 2 + (Y * wB + Z) * wB + T)
@@ -1193,7 +1193,7 @@ Qed.
rewrite <- wwB_4_wB_4; auto.
generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith.
case (w_sqrt2 w0 w1); intros w2 c.
- simpl ww_to_Z; simpl fst.
+ simpl ww_to_Z; simpl @fst.
case c; unfold interp_carry; autorewrite with rm10.
intros w3 (H6, H7); rewrite H6.
assert (V1 := spec_to_Z w3);auto with zarith.
@@ -1256,7 +1256,7 @@ Qed.
generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith.
case (w_sqrt2 w0 w1); intros w2 c.
case (spec_to_Z w2); intros HH1 HH2.
- simpl ww_to_Z; simpl fst.
+ simpl ww_to_Z; simpl @fst.
assert (Hv3: [[ww_pred ww_zdigits]]
= Zpos (xO w_digits) - 1).
rewrite spec_ww_pred; rewrite spec_ww_zdigits.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
index aaa93a21..a2df2600 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v
@@ -1,6 +1,7 @@
+
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -159,17 +160,17 @@ Section DoubleSub.
Notation wwB := (base (ww_digits w_digits)).
Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99).
Notation "[+| c |]" :=
- (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry 1 wB w_to_Z c) (at level 0, c at level 99).
Notation "[-| c |]" :=
- (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99).
+ (interp_carry (-1) wB w_to_Z c) (at level 0, c 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)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Notation "[-[ c ]]" :=
(interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c)
- (at level 0, x at level 99).
+ (at level 0, c at level 99).
Variable spec_w_0 : [|w_0|] = 0.
Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1.
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
index 1ab75307..c1f314e9 100644
--- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
+++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -14,6 +14,7 @@ Require Import ZArith.
Local Open Scope Z_scope.
Definition base digits := Z.pow 2 (Zpos digits).
+Arguments base digits: simpl never.
Section Carry.
@@ -53,7 +54,7 @@ Section Zn2Z.
End Zn2Z.
-Arguments W0 [znz].
+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],