diff options
Diffstat (limited to 'theories/Numbers/Cyclic/DoubleCyclic')
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v | 318 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v | 446 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v | 885 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v | 1540 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v | 528 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v | 487 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v | 628 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v | 1389 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v | 357 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v | 71 |
10 files changed, 6649 insertions, 0 deletions
diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v new file mode 100644 index 00000000..61d8d0fb --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v @@ -0,0 +1,318 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleAdd.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleAdd. + Variable w : Type. + Variable w_0 : w. + Variable w_1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable ww_1 : zn2z w. + Variable w_succ_c : w -> carry w. + Variable w_add_c : w -> w -> carry w. + Variable w_add_carry_c : w -> w -> carry w. + Variable w_succ : w -> w. + Variable w_add : w -> w -> w. + Variable w_add_carry : w -> w -> w. + + Definition ww_succ_c x := + match x with + | W0 => C0 ww_1 + | WW xh xl => + match w_succ_c xl with + | C0 l => C0 (WW xh l) + | C1 l => + match w_succ_c xh with + | C0 h => C0 (WW h w_0) + | C1 h => C1 W0 + end + end + end. + + 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) + end + end. + + Definition ww_add_c x y := + match x, y with + | W0, _ => C0 y + | _, W0 => C0 x + | WW xh xl, WW yh yl => + match w_add_c xl yl with + | 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 => + match w_add_carry_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (w_WW h l) + end + end + end. + + Variable R : Type. + Variable f0 f1 : zn2z w -> R. + + Definition ww_add_c_cont x y := + match x, y with + | W0, _ => f0 y + | _, W0 => f0 x + | WW xh xl, WW yh yl => + match w_add_c xl yl with + | 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 => + match w_add_carry_c xh yh with + | C0 h => f0 (WW h l) + | C1 h => f1 (w_WW h l) + end + end + end. + + (* ww_add et ww_add_carry conserve la forme normale s'il n'y a pas + de debordement *) + Definition ww_add x y := + match x, y with + | W0, _ => y + | _, W0 => x + | WW xh xl, WW yh yl => + match w_add_c xl yl with + | C0 l => WW (w_add xh yh) l + | C1 l => WW (w_add_carry xh yh) l + end + end. + + Definition ww_add_carry_c x y := + match x, y with + | W0, W0 => C0 ww_1 + | W0, WW yh yl => ww_succ_c (WW yh yl) + | 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 => + match w_add_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (WW h 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) + end + end + end. + + Definition ww_add_carry x y := + match x, y with + | W0, W0 => ww_1 + | W0, WW yh yl => ww_succ (WW yh yl) + | WW xh xl, W0 => ww_succ (WW xh xl) + | WW xh xl, WW yh yl => + match w_add_carry_c xl yl with + | C0 l => WW (w_add xh yh) l + | C1 l => WW (w_add_carry xh yh) l + end + end. + + (*Section DoubleProof.*) + Variable w_digits : positive. + 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). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (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) + (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). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + 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 : + 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. + Variable spec_w_add_carry : + forall x y, [|w_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB. + + Lemma spec_ww_succ_c : forall x, [+[ww_succ_c x]] = [[x]] + 1. + Proof. + destruct x as [ |xh xl];simpl. apply spec_ww_1. + generalize (spec_w_succ_c xl);destruct (w_succ_c xl) as [l|l]; + intro H;unfold interp_carry in H. simpl;rewrite H;ring. + rewrite <- Zplus_assoc;rewrite <- H;rewrite Zmult_1_l. + 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. + 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. + Qed. + + Lemma spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]]. + Proof. + destruct x as [ |xh xl];simpl;trivial. + destruct y as [ |yh yl];simpl. rewrite Zplus_0_r;trivial. + replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|])) + with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring. + generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l]; + intros H;unfold interp_carry in H;rewrite <- H. + generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; + intros H1;unfold interp_carry in *;rewrite <- H1. trivial. + repeat rewrite Zmult_1_l;rewrite spec_w_WW;rewrite wwB_wBwB; ring. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh) + as [h|h]; intros H1;unfold interp_carry in *;rewrite <- H1. + simpl;ring. + repeat rewrite Zmult_1_l;rewrite wwB_wBwB;rewrite spec_w_WW;ring. + Qed. + + Section Cont. + Variable P : zn2z w -> zn2z w -> R -> Prop. + Variable x y : zn2z w. + Variable spec_f0 : forall r, [[r]] = [[x]] + [[y]] -> P x y (f0 r). + Variable spec_f1 : forall r, wwB + [[r]] = [[x]] + [[y]] -> P x y (f1 r). + + Lemma spec_ww_add_c_cont : P x y (ww_add_c_cont x y). + Proof. + destruct x as [ |xh xl];simpl;trivial. + apply spec_f0;trivial. + destruct y as [ |yh yl];simpl. + apply spec_f0;simpl;rewrite Zplus_0_r;trivial. + generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l]; + intros H;unfold interp_carry in H. + generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; + intros H1;unfold interp_carry in *. + apply spec_f0. simpl;rewrite H;rewrite H1;ring. + apply spec_f1. simpl;rewrite spec_w_WW;rewrite H. + rewrite Zplus_assoc;rewrite wwB_wBwB. rewrite Zpower_2; rewrite <- Zmult_plus_distr_l. + rewrite Zmult_1_l in H1;rewrite H1;ring. + generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh) + as [h|h]; intros H1;unfold interp_carry in *. + apply spec_f0;simpl;rewrite H1. rewrite Zmult_plus_distr_l. + rewrite <- Zplus_assoc;rewrite H;ring. + apply spec_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 : + forall x y, [+[ww_add_carry_c x y]] = [[x]] + [[y]] + 1. + Proof. + destruct x as [ |xh xl];intro y;simpl. + exact (spec_ww_succ_c y). + destruct y as [ |yh yl];simpl. + rewrite Zplus_0_r;exact (spec_ww_succ_c (WW xh xl)). + replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1) + with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring. + generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) + as [l|l];intros H;unfold interp_carry in H;rewrite <- H. + generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; + intros H1;unfold interp_carry in H1;rewrite <- H1. trivial. + unfold interp_carry;repeat rewrite Zmult_1_l;simpl;rewrite wwB_wBwB;ring. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + 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. + + Lemma spec_ww_succ : forall x, [[ww_succ x]] = ([[x]] + 1) mod wwB. + Proof. + destruct x as [ |xh xl];simpl. + rewrite spec_ww_1;rewrite Zmod_small;trivial. + split;[intro;discriminate|apply wwB_pos]. + rewrite <- Zplus_assoc;generalize (spec_w_succ_c xl); + destruct (w_succ_c xl) as[l|l];intro H;unfold interp_carry in H;rewrite <-H. + rewrite Zmod_small;trivial. + rewrite wwB_wBwB;apply beta_mult;apply spec_to_Z. + assert ([|l|] = 0). clear spec_ww_1 spec_w_1 spec_w_0. + assert (H1:= spec_to_Z l); assert (H2:= spec_to_Z xl); omega. + rewrite H0;rewrite Zplus_0_r;rewrite <- Zmult_plus_distr_l;rewrite wwB_wBwB. + rewrite Zpower_2; rewrite Zmult_mod_distr_r;try apply lt_0_wB. + rewrite spec_w_W0;rewrite spec_w_succ;trivial. + Qed. + + Lemma spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB. + Proof. + destruct x as [ |xh xl];intros y;simpl. + rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial. + destruct y as [ |yh yl]. + change [[W0]] with 0;rewrite Zplus_0_r. + 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|])) + 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. + 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. + + 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. + 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) + with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring. + generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) + as [l|l];unfold interp_carry;intros H;rewrite <- H;simpl ww_to_Z. + rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial. + Qed. + +(* End DoubleProof. *) +End DoubleAdd. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v new file mode 100644 index 00000000..952516ac --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v @@ -0,0 +1,446 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleBase.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. + +Open Local Scope Z_scope. + +Section DoubleBase. + Variable w : Type. + Variable w_0 : w. + Variable w_1 : w. + Variable w_Bm1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_digits : positive. + Variable w_zdigits: w. + Variable w_add: w -> w -> zn2z w. + Variable w_to_Z : w -> Z. + Variable w_compare : w -> w -> comparison. + + Definition ww_digits := xO w_digits. + + Definition ww_zdigits := w_add w_zdigits w_zdigits. + + Definition ww_to_Z := zn2z_to_Z (base w_digits) w_to_Z. + + Definition ww_1 := WW w_0 w_1. + + Definition ww_Bm1 := WW w_Bm1 w_Bm1. + + Definition ww_WW xh xl : zn2z (zn2z w) := + match xh, xl with + | W0, W0 => W0 + | _, _ => WW xh xl + end. + + Definition ww_W0 h : zn2z (zn2z w) := + match h with + | W0 => W0 + | _ => WW h W0 + end. + + Definition ww_0W l : zn2z (zn2z w) := + match l with + | 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 + | S n => + fun (h l : zn2z (word w n)) => + match h, l with + | W0, W0 => W0 + | _, _ => WW h l + end + end. + + Fixpoint double_digits (n:nat) : positive := + match n with + | O => w_digits + | S n => xO (double_digits n) + end. + + Definition double_wB n := base (double_digits n). + + Fixpoint double_to_Z (n:nat) : word w n -> Z := + match n return word w n -> Z with + | O => w_to_Z + | S n => zn2z_to_Z (double_wB n) (double_to_Z n) + end. + + Fixpoint extend_aux (n:nat) (x:zn2z w) {struct n}: word w (S n) := + match n return word w (S n) with + | O => x + | S n1 => WW W0 (extend_aux n1 x) + end. + + Definition extend (n:nat) (x:w) : word w (S n) := + let r := w_0W x in + match r with + | W0 => W0 + | _ => extend_aux n r + end. + + Definition double_0 n : word w n := + 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 + | 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 + | W0, WW yh yl => + match w_compare w_0 yh with + | Eq => w_compare w_0 yl + | _ => Lt + end + | WW xh xl, W0 => + match w_compare xh w_0 with + | Eq => w_compare xl w_0 + | _ => Gt + end + | WW xh xl, WW yh yl => + match w_compare xh yh with + | Eq => w_compare xl yl + | Lt => Lt + | Gt => Gt + end + end. + + + (* Return the low part of the composed word*) + Fixpoint get_low (n : nat) {struct n}: + word w n -> w := + match n return (word w n -> w) with + | 0%nat => fun x => x + | S n1 => + fun x => + match x with + | W0 => w_0 + | WW _ x1 => get_low n1 x1 + 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 ]]" := + (interp_carry 1 wwB ww_to_Z c) (at level 0, x at level 99). + 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). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_compare : forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + + Lemma wwB_wBwB : wwB = wB^2. + Proof. + unfold base, ww_digits;rewrite Zpower_2; rewrite (Zpos_xO w_digits). + replace (2 * Zpos w_digits) with (Zpos w_digits + Zpos w_digits). + apply Zpower_exp; unfold Zge;simpl;intros;discriminate. + ring. + Qed. + + Lemma spec_ww_1 : [[ww_1]] = 1. + Proof. simpl;rewrite spec_w_0;rewrite spec_w_1;ring. Qed. + + Lemma spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1. + Proof. simpl;rewrite spec_w_Bm1;rewrite wwB_wBwB;ring. Qed. + + Lemma lt_0_wB : 0 < wB. + Proof. + unfold base;apply Zpower_gt_0. unfold Zlt;reflexivity. + unfold Zle;intros H;discriminate H. + Qed. + + Lemma lt_0_wwB : 0 < wwB. + Proof. rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_lt_0_compat;apply lt_0_wB. Qed. + + Lemma wB_pos: 1 < wB. + Proof. + unfold base;apply Zlt_le_trans with (2^1). unfold Zlt;reflexivity. + apply Zpower_le_monotone. unfold Zlt;reflexivity. + split;unfold Zle;intros H. discriminate H. + clear spec_w_0W w_0W spec_w_Bm1 spec_to_Z spec_w_WW w_WW. + destruct w_digits; discriminate H. + Qed. + + Lemma wwB_pos: 1 < wwB. + Proof. + assert (H:= wB_pos);rewrite wwB_wBwB;rewrite <-(Zmult_1_r 1). + rewrite Zpower_2. + apply Zmult_lt_compat2;(split;[unfold Zlt;reflexivity|trivial]). + apply Zlt_le_weak;trivial. + 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 + spec_to_Z;unfold base. + assert (2 ^ Zpos w_digits = 2 * (2 ^ (Zpos w_digits - 1))). + pattern 2 at 2; rewrite <- Zpower_1_r. + rewrite <- Zpower_exp; auto with zarith. + f_equal; auto with zarith. + case w_digits; compute; intros; discriminate. + rewrite H; f_equal; auto with zarith. + rewrite Zmult_comm; apply Z_div_mult; auto with zarith. + Qed. + + Theorem wwB_div_2 : wwB / 2 = wB / 2 * wB. + Proof. + clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W + spec_to_Z. + rewrite wwB_wBwB; rewrite Zpower_2. + pattern wB at 1; rewrite <- wB_div_2; auto. + rewrite <- Zmult_assoc. + repeat (rewrite (Zmult_comm 2); rewrite Z_div_mult); auto with zarith. + Qed. + + Lemma mod_wwB : forall z x, + (z*wB + [|x|]) mod wwB = (z mod wB)*wB + [|x|]. + Proof. + intros z x. + 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|]). + apply Zmod_small;rewrite wwB_wBwB;apply beta_mult;try apply spec_to_Z. + apply Z_mod_lt;apply Zlt_gt;apply lt_0_wB. + destruct (spec_to_Z x);split;trivial. + change [|x|] with (0*wB+[|x|]). rewrite wwB_wBwB. + rewrite Zpower_2;rewrite <- (Zplus_0_r (wB*wB));apply beta_lex_inv. + apply lt_0_wB. apply spec_to_Z. split;[apply Zle_refl | apply lt_0_wB]. + Qed. + + Lemma wB_div : forall x y, ([|x|] * wB + [|y|]) / wB = [|x|]. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + intros x y;unfold base;rewrite Zdiv_shift_r;auto with zarith. + rewrite Z_div_mult;auto with zarith. + destruct (spec_to_Z x);trivial. + Qed. + + 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. + intros x y p Hp;rewrite Zpower_exp;auto with zarith. + rewrite <- Zdiv_Zdiv;auto with zarith. + rewrite wB_div;trivial. + Qed. + + Lemma lt_wB_wwB : wB < wwB. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + unfold base;apply Zpower_lt_monotone;auto with zarith. + assert (0 < Zpos w_digits). compute;reflexivity. + unfold ww_digits;rewrite Zpos_xO;auto with zarith. + 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. + Qed. + + Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + destruct x as [ |h l];simpl. + split;[apply Zle_refl|apply lt_0_wwB]. + assert (H:=spec_to_Z h);assert (L:=spec_to_Z l);split. + apply Zplus_le_0_compat;auto with zarith. + rewrite <- (Zplus_0_r wwB);rewrite wwB_wBwB; rewrite Zpower_2; + apply beta_lex_inv;auto with zarith. + Qed. + + Lemma double_wB_wwB : forall n, double_wB n * double_wB n = double_wB (S n). + Proof. + intros n;unfold double_wB;simpl. + unfold base;rewrite (Zpos_xO (double_digits n)). + replace (2 * Zpos (double_digits n)) with + (Zpos (double_digits n) + Zpos (double_digits n)). + symmetry; apply Zpower_exp;intro;discriminate. + ring. + Qed. + + Lemma double_wB_pos: + forall n, 0 <= double_wB n. + Proof. + intros n; unfold double_wB, base; auto with zarith. + Qed. + + Lemma double_wB_more_digits: + forall n, wB <= double_wB n. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + intros n; elim n; clear n; auto. + unfold double_wB, double_digits; auto with zarith. + intros n H1; rewrite <- double_wB_wwB. + apply Zle_trans with (wB * 1). + rewrite Zmult_1_r; apply Zle_refl. + apply Zmult_le_compat; auto with zarith. + apply Zle_trans with wB; auto with zarith. + unfold base. + rewrite <- (Zpower_0_r 2). + apply Zpower_le_monotone2; auto with zarith. + unfold base; auto with zarith. + Qed. + + 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. + induction n;intros. exact (spec_to_Z x). + unfold double_to_Z;fold double_to_Z. + destruct x;unfold zn2z_to_Z. + unfold double_wB,base;split;auto with zarith. + assert (U0:= IHn w0);assert (U1:= IHn w1). + split;auto with zarith. + apply Zlt_le_trans with ((double_wB n - 1) * double_wB n + double_wB n). + assert (double_to_Z n w0*double_wB n <= (double_wB n - 1)*double_wB n). + apply Zmult_le_compat_r;auto with zarith. + auto with zarith. + rewrite <- double_wB_wwB. + replace ((double_wB n - 1) * double_wB n + double_wB n) with (double_wB n * double_wB n); + [auto with zarith | ring]. + Qed. + + Lemma spec_get_low: + forall n x, + [!n | x!] < wB -> [|get_low n x|] = [!n | x!]. + Proof. + clear spec_w_1 spec_w_Bm1. + intros n; elim n; auto; clear n. + intros n Hrec x; case x; clear x; auto. + intros xx yy H1; simpl in H1. + assert (F1: [!n | xx!] = 0). + case (Zle_lt_or_eq 0 ([!n | xx!])); auto. + case (spec_double_to_Z n xx); auto. + intros F2. + assert (F3 := double_wB_more_digits n). + assert (F4: 0 <= [!n | yy!]). + case (spec_double_to_Z n yy); auto. + assert (F5: 1 * wB <= [!n | xx!] * double_wB n); + auto with zarith. + apply Zmult_le_compat; auto with zarith. + unfold base; auto with zarith. + simpl get_low; simpl double_to_Z. + generalize H1; clear H1. + rewrite F1; rewrite Zmult_0_l; rewrite Zplus_0_l. + intros H1; apply Hrec; auto. + Qed. + + Lemma spec_double_WW : forall n (h l : word w n), + [!S n|double_WW n h l!] = [!n|h!] * double_wB n + [!n|l!]. + Proof. + induction n;simpl;intros;trivial. + destruct h;auto. + destruct l;auto. + Qed. + + Lemma spec_extend_aux : forall n x, [!S n|extend_aux n x!] = [[x]]. + Proof. induction n;simpl;trivial. Qed. + + Lemma spec_extend : forall n x, [!S n|extend n x!] = [|x|]. + Proof. + intros n x;assert (H:= spec_w_0W x);unfold extend. + 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, + let (h,l) := double_split n x in + [!S n|x!] = [!n|h!] * double_wB n + [!n|l!]. + Proof. + destruct x;simpl;auto. + destruct n;simpl;trivial. + rewrite spec_w_0;trivial. + Qed. + + 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. + + Lemma spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Proof. + destruct x as [ |xh xl];destruct y as [ |yh yl];simpl;trivial. + generalize (spec_w_compare w_0 yh);destruct (w_compare w_0 yh); + 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. + 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); + intros H;rewrite spec_w_0 in H. + rewrite H;simpl;rewrite <- spec_w_0;apply spec_w_compare. + 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. + + 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]; + trivial. + apply wB_lex_inv;trivial. + 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 new file mode 100644 index 00000000..cca32a59 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v @@ -0,0 +1,885 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleCyclic.v 11012 2008-05-28 16:34:43Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. +Require Import DoubleAdd. +Require Import DoubleSub. +Require Import DoubleMul. +Require Import DoubleSqrt. +Require Import DoubleLift. +Require Import DoubleDivn1. +Require Import DoubleDiv. +Require Import CyclicAxioms. + +Open Local Scope Z_scope. + + +Section Z_2nZ. + + Variable w : Type. + Variable w_op : znz_op w. + Let w_digits := w_op.(znz_digits). + Let w_zdigits := w_op.(znz_zdigits). + + Let w_to_Z := w_op.(znz_to_Z). + Let w_of_pos := w_op.(znz_of_pos). + Let w_head0 := w_op.(znz_head0). + Let w_tail0 := w_op.(znz_tail0). + + Let w_0 := w_op.(znz_0). + Let w_1 := w_op.(znz_1). + Let w_Bm1 := w_op.(znz_Bm1). + + Let w_compare := w_op.(znz_compare). + Let w_eq0 := w_op.(znz_eq0). + + Let w_opp_c := w_op.(znz_opp_c). + Let w_opp := w_op.(znz_opp). + Let w_opp_carry := w_op.(znz_opp_carry). + + Let w_succ_c := w_op.(znz_succ_c). + Let w_add_c := w_op.(znz_add_c). + Let w_add_carry_c := w_op.(znz_add_carry_c). + Let w_succ := w_op.(znz_succ). + Let w_add := w_op.(znz_add). + Let w_add_carry := w_op.(znz_add_carry). + + Let w_pred_c := w_op.(znz_pred_c). + Let w_sub_c := w_op.(znz_sub_c). + Let w_sub_carry_c := w_op.(znz_sub_carry_c). + Let w_pred := w_op.(znz_pred). + Let w_sub := w_op.(znz_sub). + Let w_sub_carry := w_op.(znz_sub_carry). + + + 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). + + Let w_mod_gt := w_op.(znz_mod_gt). + Let w_mod := w_op.(znz_mod). + + 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_pos_mod := w_op.(znz_pos_mod). + + Let w_is_even := w_op.(znz_is_even). + Let w_sqrt2 := w_op.(znz_sqrt2). + Let w_sqrt := w_op.(znz_sqrt). + + Let _zn2z := zn2z w. + + 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. + + Let w_add2 a b := match w_add_c a b with C0 p => WW w_0 p | C1 p => WW w_1 p end. + + Let _ww_digits := xO w_digits. + + Let _ww_zdigits := w_add2 w_zdigits w_zdigits. + + Let to_Z := zn2z_to_Z wB w_to_Z. + + Let w_W0 := znz_W0 w_op. + Let w_0W := znz_0W w_op. + Let w_WW := znz_WW w_op. + + Let ww_of_pos p := + match w_of_pos p with + | (N0, l) => (N0, WW w_0 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 + 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 + 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). + Let ww_0W := Eval lazy beta delta [ww_0W] in (@ww_0W w). + Let ww_W0 := Eval lazy beta delta [ww_W0] in (@ww_W0 w). + + (* ** Comparison ** *) + Let compare := + Eval lazy beta delta[ww_compare] in ww_compare w_0 w_compare. + + Let eq0 (x:zn2z w) := + match x with + | W0 => true + | _ => false + end. + + (* ** Opposites ** *) + Let opp_c := + Eval lazy beta delta [ww_opp_c] in ww_opp_c w_0 w_opp_c w_opp_carry. + + Let opp := + Eval lazy beta delta [ww_opp] in ww_opp w_0 w_opp_c w_opp_carry w_opp. + + 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 := + Eval lazy beta delta [ww_succ_c] in ww_succ_c w_0 ww_1 w_succ_c. + + Let add_c := + 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 + ww_add_carry_c w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c. + + 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 := + 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. + + (* ** Subtractions ** *) + + 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 + ww_sub_c w_0 w_WW w_opp_c w_opp_carry w_sub_c w_sub_carry_c. + + Let sub_carry_c := + Eval lazy beta iota delta [ww_sub_carry_c ww_pred_c ww_opp_carry] in + ww_sub_carry_c w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c w_sub_c w_sub_carry_c. + + 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 + 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 := + Eval lazy beta iota delta [ww_sub_carry ww_pred ww_opp_carry] in + ww_sub_carry w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c w_sub_carry_c w_pred + w_sub w_sub_carry. + + + (* ** Multiplication ** *) + + Let mul_c := + Eval lazy beta iota delta [ww_mul_c double_mul_c] in + ww_mul_c w_0 w_1 w_WW w_W0 w_mul_c add_c add add_carry. + + 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 + add_c add add_carry sub_c sub. + + Let mul := + Eval lazy beta delta [ww_mul] in + ww_mul w_W0 w_add w_mul_c w_mul add. + + Let square_c := + Eval lazy beta delta [ww_square_c] in + ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add add_carry. + + (* Division operation *) + + 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_add w_add_carry w_pred w_sub w_mul_c w_div21 sub_c. + + Let div21 := + Eval lazy beta iota delta [ww_div21] in + ww_div21 w_0 w_0W div32 ww_1 compare sub. + + Let low (p: zn2z w) := match p with WW _ p1 => p1 | _ => w_0 end. + + Let add_mul_div := + Eval lazy beta delta [ww_add_mul_div] in + ww_add_mul_div w_0 w_WW w_W0 w_0W compare w_add_mul_div sub w_zdigits low. + + 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 + 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_ := + Eval lazy beta delta [ww_mod] in ww_mod compare mod_gt. + + 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 := + Eval lazy beta delta [ww_is_even] in ww_is_even w_is_even. + + 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 := + 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 := + 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 + w_add_mul_div w_head0 w_div21 div32 _ww_zdigits add_mul_div + w_zdigits. + + Let gcd_cont := + 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 + ww_gcd_gt w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont. + + Let gcd := + Eval lazy beta delta [ww_gcd] in + 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 := + 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 + pred sub sub_carry + mul_c mul square_c + div21 div_gt div + mod_gt mod_ + gcd_gt gcd + add_mul_div + pos_mod + is_even + sqrt2 + sqrt. + + 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 + pred sub sub_carry + karatsuba_c mul square_c + div21 div_gt div + mod_gt mod_ + gcd_gt gcd + add_mul_div + pos_mod + is_even + sqrt2 + sqrt. + + (* Proof *) + Variable op_spec : znz_spec w_op. + + Hint Resolve + (spec_to_Z op_spec) + (spec_of_pos op_spec) + (spec_0 op_spec) + (spec_1 op_spec) + (spec_Bm1 op_spec) + (spec_compare op_spec) + (spec_eq0 op_spec) + (spec_opp_c op_spec) + (spec_opp op_spec) + (spec_opp_carry op_spec) + (spec_succ_c op_spec) + (spec_add_c op_spec) + (spec_add_carry_c op_spec) + (spec_succ op_spec) + (spec_add op_spec) + (spec_add_carry op_spec) + (spec_pred_c op_spec) + (spec_sub_c op_spec) + (spec_sub_carry_c op_spec) + (spec_pred op_spec) + (spec_sub op_spec) + (spec_sub_carry op_spec) + (spec_mul_c op_spec) + (spec_mul op_spec) + (spec_square_c op_spec) + (spec_div21 op_spec) + (spec_div_gt op_spec) + (spec_div op_spec) + (spec_mod_gt op_spec) + (spec_mod op_spec) + (spec_gcd_gt 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) + (spec_sqrt2) + (spec_sqrt) + (spec_W0 op_spec) + (spec_0W op_spec) + (spec_WW op_spec). + + Ltac wwauto := unfold ww_to_Z; auto. + + 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). + + Notation "[-| c |]" := + (interp_carry (-1) wwB to_Z c) (at level 0, x 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. + + Let spec_ww_of_pos : forall p, + Zpos p = (Z_of_N (fst (ww_of_pos p)))*wwB + [|(snd (ww_of_pos p))|]. + Proof. + unfold ww_of_pos;intros. + assert (H:= spec_of_pos op_spec p);unfold w_of_pos; + destruct (znz_of_pos w_op p). simpl in H. + rewrite H;clear H;destruct n;simpl to_Z. + simpl;unfold w_to_Z,w_0;rewrite (spec_0 op_spec);trivial. + unfold Z_of_N; assert (H:= spec_of_pos op_spec p0); + destruct (znz_of_pos w_op p0). simpl in H. + rewrite H;unfold fst, snd,Z_of_N, to_Z. + rewrite (spec_WW op_spec). + replace wwB with (wB*wB). + unfold wB,w_to_Z,w_digits;clear H;destruct n;ring. + symmetry. rewrite <- Zpower_2; exact (wwB_wBwB w_digits). + Qed. + + Let spec_ww_0 : [|W0|] = 0. + 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. + + 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 : + 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). + Qed. + + Let spec_ww_eq0 : forall x, eq0 x = true -> [|x|] = 0. + Proof. destruct x;simpl;intros;trivial;discriminate. Qed. + + 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. + 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. + Qed. + + Let spec_ww_opp_carry : forall x, [|opp_carry x|] = wwB - [|x|] - 1. + Proof. + refine (spec_ww_opp_carry w_WW ww_Bm1 w_opp_carry w_digits w_to_Z _ _ _); + wwauto. + Qed. + + 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. + Qed. + + Let spec_ww_add_c : forall x y, [+|add_c x y|] = [|x|] + [|y|]. + Proof. + refine (spec_ww_add_c w_WW w_add_c w_add_carry_c w_digits w_to_Z _ _ _);wwauto. + Qed. + + Let spec_ww_add_carry_c : forall x y, [+|add_carry_c x y|] = [|x|]+[|y|]+1. + Proof. + refine (spec_ww_add_carry_c w_0 w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c + w_digits w_to_Z _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_succ : forall x, [|succ x|] = ([|x|] + 1) mod wwB. + Proof. + refine (spec_ww_succ w_W0 ww_1 w_succ_c w_succ w_digits w_to_Z _ _ _ _ _); + wwauto. + Qed. + + 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. + Qed. + + 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 + 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 + _ _ _ _ _);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 + 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 + w_sub_c w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_pred : forall x, [|pred x|] = ([|x|] - 1) mod wwB. + Proof. + refine (spec_ww_pred w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_pred w_digits w_to_Z + _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_sub : forall x y, [|sub x y|] = ([|x|] - [|y|]) mod wwB. + Proof. + refine (spec_ww_sub w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c w_opp + w_sub w_sub_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_sub_carry : forall x y, [|sub_carry x y|]=([|x|]-[|y|]-1) mod wwB. + Proof. + refine (spec_ww_sub_carry w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c + w_sub_carry_c w_pred w_sub w_sub_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _); + wwauto. + Qed. + + Let spec_ww_mul_c : forall x y, [[mul_c x y ]] = [|x|] * [|y|]. + Proof. + refine (spec_ww_mul_c w_0 w_1 w_WW w_W0 w_mul_c add_c add add_carry w_digits + w_to_Z _ _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_karatsuba_c : forall x y, [[karatsuba_c x y ]] = [|x|] * [|y|]. + Proof. + refine (spec_ww_karatsuba_c _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + _ _ _ _ _ _ _ _ _ _ _ _); wwauto. + unfold w_digits; apply spec_more_than_1_digit; auto. + exact (spec_compare op_spec). + 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. + 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 + add_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_w_div32 : forall a1 a2 a3 b1 b2, + wB / 2 <= (w_to_Z b1) -> + [|WW a1 a2|] < [|WW b1 b2|] -> + let (q, r) := div32 a1 a2 a3 b1 b2 in + (w_to_Z a1) * wwB + (w_to_Z a2) * wB + (w_to_Z a3) = + (w_to_Z q) * ((w_to_Z b1)*wB + (w_to_Z b2)) + [|r|] /\ + 0 <= [|r|] < (w_to_Z b1)*wB + w_to_Z b2. + Proof. + refine (spec_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 w_digits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + unfold w_Bm2, w_to_Z, w_pred, w_Bm1. + rewrite (spec_pred op_spec);rewrite (spec_Bm1 op_spec). + unfold w_digits;rewrite Zmod_small. ring. + assert (H:= wB_pos(znz_digits w_op)). omega. + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + Qed. + + Let spec_ww_div21 : forall a1 a2 b, + wwB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := div21 a1 a2 b in + [|a1|] *wwB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + refine (spec_ww_div21 w_0 w_0W div32 ww_1 compare sub w_digits w_to_Z + _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_add2: forall x y, + [|w_add2 x y|] = w_to_Z x + w_to_Z y. + unfold w_add2. + intros xh xl; generalize (spec_add_c op_spec xh xl). + unfold w_add_c; case znz_add_c; unfold interp_carry; simpl ww_to_Z. + intros w0 Hw0; simpl; unfold w_to_Z; rewrite Hw0. + unfold w_0; rewrite spec_0; simpl; auto with zarith. + intros w0; rewrite Zmult_1_l; simpl. + unfold w_to_Z, w_1; rewrite spec_1; auto with zarith. + rewrite Zmult_1_l; auto. + Qed. + + Let spec_low: forall x, + w_to_Z (low x) = [|x|] mod wB. + intros x; case x; simpl low. + unfold ww_to_Z, w_to_Z, w_0; rewrite (spec_0 op_spec); simpl. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + unfold wB, base; auto with zarith. + intros xh xl; simpl. + rewrite Zplus_comm; rewrite Z_mod_plus; auto with zarith. + rewrite Zmod_small; auto with zarith. + unfold wB, base; auto with zarith. + Qed. + + Let spec_ww_digits: + [|_ww_zdigits|] = Zpos (xO w_digits). + Proof. + unfold w_to_Z, _ww_zdigits. + rewrite spec_add2. + unfold w_to_Z, w_zdigits, w_digits. + rewrite spec_zdigits; auto. + rewrite Zpos_xO; auto with zarith. + Qed. + + + Let spec_ww_head00 : forall x, [|x|] = 0 -> [|head0 x|] = Zpos _ww_digits. + Proof. + refine (spec_ww_head00 w_0 w_0W + w_compare w_head0 w_add2 w_zdigits _ww_zdigits + w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); auto. + exact (spec_compare op_spec). + exact (spec_head00 op_spec). + exact (spec_zdigits op_spec). + Qed. + + 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 + w_to_Z _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_zdigits op_spec). + Qed. + + Let spec_ww_tail00 : forall x, [|x|] = 0 -> [|tail0 x|] = Zpos _ww_digits. + Proof. + refine (spec_ww_tail00 w_0 w_0W + w_compare w_tail0 w_add2 w_zdigits _ww_zdigits + w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); wwauto. + exact (spec_compare op_spec). + exact (spec_tail00 op_spec). + exact (spec_zdigits op_spec). + Qed. + + + 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 + w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_zdigits op_spec). + Qed. + + Lemma spec_ww_add_mul_div : forall x y p, + [|p|] <= Zpos _ww_digits -> + [| add_mul_div p x y |] = + ([|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 + sub w_digits w_zdigits low w_to_Z + _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_zdigits op_spec). + Qed. + + 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 + 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 + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ +). + exact (spec_0 op_spec). + exact (spec_to_Z op_spec). + wwauto. + wwauto. + exact (spec_compare op_spec). + exact (spec_eq0 op_spec). + exact (spec_opp_c op_spec). + exact (spec_opp op_spec). + exact (spec_opp_carry op_spec). + exact (spec_sub_c op_spec). + exact (spec_sub op_spec). + exact (spec_sub_carry op_spec). + exact (spec_div_gt op_spec). + exact (spec_add_mul_div op_spec). + exact (spec_head0 op_spec). + exact (spec_div21 op_spec). + exact spec_w_div32. + exact (spec_zdigits op_spec). + exact spec_ww_digits. + exact spec_ww_1. + exact spec_ww_add_mul_div. + Qed. + + Let spec_ww_div : forall a b, 0 < [|b|] -> + let (q,r) := div a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + refine (spec_ww_div w_digits ww_1 compare div_gt w_to_Z _ _ _ _);auto. + Qed. + + 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 + 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 + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_div_gt op_spec). + exact (spec_div21 op_spec). + exact (spec_zdigits op_spec). + exact spec_ww_add_mul_div. + Qed. + + 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. + Qed. + + 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 _ + 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 + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + exact (spec_zdigits op_spec). + exact spec_ww_add_mul_div. + refine (@spec_gcd_cont w w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare + _ _);auto. + exact (spec_compare op_spec). + Qed. + + Let spec_ww_gcd : forall a b, Zis_gcd [|a|] [|b|] [|gcd a b|]. + Proof. + refine (@spec_ww_gcd w w_digits W0 compare 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 + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + exact (spec_zdigits op_spec). + exact spec_ww_add_mul_div. + refine (@spec_gcd_cont w w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare + _ _);auto. + exact (spec_compare op_spec). + Qed. + + Let spec_ww_is_even : forall x, + match is_even x with + true => [|x|] mod 2 = 0 + | false => [|x|] mod 2 = 1 + end. + Proof. + refine (@spec_ww_is_even w w_is_even w_0 w_1 w_Bm1 w_digits _ _ _ _ _); auto. + exact (spec_is_even op_spec). + Qed. + + Let spec_ww_sqrt2 : forall x y, + wwB/ 4 <= [|x|] -> + let (s,r) := sqrt2 x y in + [[WW x y]] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]. + Proof. + intros x y H. + refine (@spec_ww_sqrt2 w 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_digits w_zdigits + _ww_zdigits + w_add_c w_sqrt2 w_pred pred_c pred add_c add sub_c add_mul_div + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto. + exact (spec_zdigits op_spec). + exact (spec_more_than_1_digit op_spec). + exact (spec_is_even op_spec). + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + exact (spec_ww_add_mul_div). + exact (spec_sqrt2 op_spec). + Qed. + + 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 + w_sub w_add_mul_div w_digits w_zdigits _ww_zdigits + w_sqrt2 pred add_mul_div head0 compare + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto. + exact (spec_zdigits op_spec). + exact (spec_more_than_1_digit op_spec). + exact (spec_is_even op_spec). + exact (spec_ww_add_mul_div). + exact (spec_sqrt2 op_spec). + Qed. + + Lemma mk_znz2_spec : znz_spec mk_zn2z_op. + 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 + w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_pos_mod op_spec). + exact (spec_zdigits op_spec). + unfold w_to_Z, w_zdigits. + rewrite (spec_zdigits op_spec). + rewrite <- Zpos_xO; exact spec_ww_digits. + Qed. + + Lemma mk_znz2_karatsuba_spec : znz_spec mk_zn2z_op_karatsuba. + 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 + w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_pos_mod op_spec). + exact (spec_zdigits op_spec). + unfold w_to_Z, w_zdigits. + rewrite (spec_zdigits op_spec). + rewrite <- Zpos_xO; exact spec_ww_digits. + Qed. + +End Z_2nZ. + +Section MulAdd. + + Variable w: Type. + Variable op: znz_op w. + Variable sop: znz_spec op. + + Definition mul_add:= w_mul_add (znz_0 op) (znz_succ op) (znz_add_c op) (znz_mul_c op). + + Notation "[| x |]" := (znz_to_Z op x) (at level 0, x at level 99). + + Notation "[|| x ||]" := + (zn2z_to_Z (base (znz_digits op)) (znz_to_Z op) x) (at level 0, x at level 99). + + + Lemma spec_mul_add: forall x y z, + let (zh, zl) := mul_add x y z in + [||WW zh zl||] = [|x|] * [|y|] + [|z|]. + Proof. + intros x y z. + refine (spec_w_mul_add _ _ _ _ _ _ _ _ _ _ _ _ x y z); auto. + exact (spec_0 sop). + exact (spec_to_Z sop). + exact (spec_succ sop). + exact (spec_add_c sop). + exact (spec_mul_c sop). + Qed. + +End MulAdd. + + +(** Modular versions of DoubleCyclic *) + +Module DoubleCyclic (C:CyclicType) <: CyclicType. + Definition w := zn2z C.w. + Definition w_op := mk_zn2z_op C.w_op. + Definition w_spec := mk_znz2_spec C.w_spec. +End DoubleCyclic. + +Module DoubleCyclicKaratsuba (C:CyclicType) <: CyclicType. + Definition w := zn2z C.w. + Definition w_op := mk_zn2z_op_karatsuba C.w_op. + Definition w_spec := mk_znz2_karatsuba_spec C.w_spec. +End DoubleCyclicKaratsuba. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v new file mode 100644 index 00000000..075aef59 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v @@ -0,0 +1,1540 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleDiv.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. +Require Import DoubleDivn1. +Require Import DoubleAdd. +Require Import DoubleSub. + +Open Local Scope Z_scope. + +Ltac zarith := auto with zarith. + + +Section POS_MOD. + + Variable w:Type. + Variable w_0 : w. + Variable w_digits : positive. + Variable w_zdigits : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_pos_mod : w -> w -> w. + Variable w_compare : w -> w -> comparison. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable w_0W : w -> zn2z w. + Variable low: zn2z w -> w. + Variable ww_sub: zn2z w -> zn2z w -> zn2z w. + Variable ww_zdigits : zn2z w. + + + 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 + | Lt => w_WW w_0 (w_pos_mod (low p) xl) + | Gt => + match ww_compare p ww_zdigits with + | Lt => + let n := low (ww_sub p zdigits) in + w_WW (w_pos_mod n xh) xl + | _ => x + end + end + end. + + + 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). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + + + Variable spec_w_0 : [|w_0|] = 0. + + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + + Variable spec_pos_mod : forall w p, + [|w_pos_mod p w|] = [|w|] mod (2 ^ [|p|]). + + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + 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]] = 2 * [|w_zdigits|]. + Variable spec_ww_digits : ww_digits w_digits = xO w_digits. + + + Hint Rewrite spec_w_0 spec_w_WW : w_rewrite. + + Lemma spec_ww_pos_mod : forall w p, + [[ww_pos_mod p w]] = [[w]] mod (2 ^ [[p]]). + 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. + intros xh xl; generalize (spec_ww_compare p (w_0W w_zdigits)); + case ww_compare; + rewrite spec_w_0W; rewrite spec_zdigits; fold wB; + intros H1. + rewrite H1; simpl ww_to_Z. + autorewrite with w_rewrite rm10. + rewrite Zplus_mod; auto with zarith. + rewrite Z_mod_mult; auto with zarith. + autorewrite with rm10. + rewrite Zmod_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + autorewrite with w_rewrite rm10. + simpl ww_to_Z. + rewrite spec_pos_mod. + assert (HH0: [|low p|] = [[p]]). + rewrite spec_low. + apply Zmod_small; auto with zarith. + case (spec_to_w_Z p); intros HHH1 HHH2; split; auto with zarith. + apply Zlt_le_trans with (1 := H1). + unfold base; apply Zpower2_le_lin; auto with zarith. + rewrite HH0. + rewrite Zplus_mod; auto with zarith. + unfold base. + rewrite <- (F0 (Zpos w_digits) [[p]]). + rewrite Zpower_exp; auto with zarith. + rewrite Zmult_assoc. + rewrite Z_mod_mult; auto with zarith. + 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; + 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))|] = + [[p]] - Zpos w_digits). + rewrite spec_low. + rewrite spec_ww_sub. + rewrite spec_w_0W; rewrite spec_zdigits. + rewrite <- Zmod_div_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. + exists wB; unfold base; rewrite <- Zpower_exp; auto with zarith. + rewrite spec_ww_digits; + apply f_equal with (f := Zpower 2); rewrite Zpos_xO; auto with zarith. + simpl ww_to_Z; autorewrite with w_rewrite. + rewrite spec_pos_mod; rewrite HH0. + pattern [|xh|] at 2; + rewrite Z_div_mod_eq with (b := 2 ^ ([[p]] - Zpos w_digits)); + auto with zarith. + rewrite (fun x => (Zmult_comm (2 ^ x))); rewrite Zmult_plus_distr_l. + unfold base; rewrite <- Zmult_assoc; rewrite <- Zpower_exp; + auto with zarith. + rewrite F0; auto with zarith. + rewrite <- Zplus_assoc; rewrite Zplus_mod; auto with zarith. + rewrite Z_mod_mult; auto with zarith. + autorewrite with rm10. + rewrite Zmod_mod; auto with zarith. + apply sym_equal; apply Zmod_small; auto with zarith. + case (spec_to_Z xh); intros U1 U2. + case (spec_to_Z xl); intros U3 U4. + split; auto with zarith. + apply Zplus_le_0_compat; auto with zarith. + apply Zmult_le_0_compat; auto with zarith. + match goal with |- 0 <= ?X mod ?Y => + case (Z_mod_lt X Y); auto with zarith + end. + match goal with |- ?X mod ?Y * ?U + ?Z < ?T => + apply Zle_lt_trans with ((Y - 1) * U + Z ); + [case (Z_mod_lt X Y); auto with zarith | idtac] + end. + match goal with |- ?X * ?U + ?Y < ?Z => + apply Zle_lt_trans with (X * U + (U - 1)) + end. + apply Zplus_le_compat_l; auto with zarith. + case (spec_to_Z xl); unfold base; auto with zarith. + rewrite Zmult_minus_distr_r; rewrite <- Zpower_exp; auto with zarith. + rewrite F0; auto with zarith. + rewrite Zmod_small; auto with zarith. + case (spec_to_w_Z (WW xh xl)); intros U1 U2. + split; auto with zarith. + apply Zlt_le_trans with (1:= U2). + unfold base; rewrite spec_ww_digits. + apply Zpower_le_monotone; auto with zarith. + split; auto with zarith. + rewrite Zpos_xO; auto with zarith. + Qed. + +End POS_MOD. + +Section DoubleDiv32. + + Variable w : Type. + Variable w_0 : w. + Variable w_Bm1 : w. + Variable w_Bm2 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable w_add_c : w -> w -> carry w. + Variable w_add_carry_c : w -> w -> carry w. + Variable w_add : w -> w -> w. + Variable w_add_carry : w -> w -> w. + Variable w_pred : w -> w. + Variable w_sub : w -> w -> w. + Variable w_mul_c : w -> w -> zn2z w. + 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 + 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. + + (* Proof *) + + Variable w_digits : positive. + 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). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (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) + (at level 0, x at level 99). + + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_w_Bm2 : [|w_Bm2|] = wB - 2. + + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | 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 : + 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. + Variable spec_w_add_carry : + forall x y, [|w_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB. + + 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_mul_c : forall x y, [[ w_mul_c x y ]] = [|x|] * [|y|]. + Variable spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + + Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Theorem wB_div2: forall x, wB/2 <= x -> wB <= 2 * x. + intros x H; rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith. + Qed. + + Lemma Zmult_lt_0_reg_r_2 : forall n m : Z, 0 <= n -> 0 < m * n -> 0 < m. + Proof. + intros n m H1 H2;apply Zmult_lt_0_reg_r with n;trivial. + destruct (Zle_lt_or_eq _ _ H1);trivial. + subst;rewrite Zmult_0_r in H2;discriminate H2. + Qed. + + Theorem spec_w_div32 : forall a1 a2 a3 b1 b2, + 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]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]. + Proof. + intros a1 a2 a3 b1 b2 Hle Hlt. + assert (U:= lt_0_wB w_digits); assert (U1:= lt_0_wwB w_digits). + Spec_w_to_Z a1;Spec_w_to_Z a2;Spec_w_to_Z a3;Spec_w_to_Z b1;Spec_w_to_Z b2. + rewrite wwB_wBwB; rewrite Zpower_2; rewrite Zmult_assoc;rewrite <- Zmult_plus_distr_l. + 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. + assert (Hcmp:=spec_compare a1 b1);destruct (w_compare a1 b1). + simpl in Hlt. + rewrite Hcmp in Hlt;assert ([|a2|] < [|b2|]). omega. + assert ([[WW (w_sub a2 b2) a3]] = ([|a2|]-[|b2|])*wB + [|a3|] + wwB). + simpl;rewrite spec_sub. + assert ([|a2|] - [|b2|] = wB*(-1) + ([|a2|] - [|b2|] + wB)). ring. + assert (0 <= [|a2|] - [|b2|] + wB < wB). omega. + rewrite <-(Zmod_unique ([|a2|]-[|b2|]) wB (-1) ([|a2|]-[|b2|]+wB) H1 H0). + rewrite wwB_wBwB;ring. + assert (U2 := wB_pos w_digits). + eapply spec_ww_add_c_cont with (P := + fun (x y:zn2z w) (res:w*zn2z w) => + let (q, r) := res in + ([|a1|] * wB + [|a2|]) * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]);eauto. + rewrite H0;intros r. + repeat + (rewrite spec_ww_add;eauto || rewrite spec_w_Bm1 || rewrite spec_w_Bm2); + simpl ww_to_Z;try rewrite Zmult_1_l;intros H1. + assert (0<= ([[r]] + ([|b1|] * wB + [|b2|])) - wwB < [|b1|] * wB + [|b2|]). + Spec_ww_to_Z r;split;zarith. + rewrite H1. + assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB). + rewrite wwB_wBwB; rewrite Zpower_2; zarith. + assert (-wwB < ([|a2|] - [|b2|]) * wB + [|a3|] < 0). + split. apply Zlt_le_trans with (([|a2|] - [|b2|]) * wB);zarith. + rewrite wwB_wBwB;replace (-(wB^2)) with (-wB*wB);[zarith | ring]. + apply Zmult_lt_compat_r;zarith. + apply Zle_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith. + replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with + (([|a2|] - [|b2|] + 1) * wB + - 1);[zarith | ring]. + assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith. + replace 0 with (0*wB);zarith. + replace (([|a2|] - [|b2|]) * wB + [|a3|] + wwB + ([|b1|] * wB + [|b2|]) + + ([|b1|] * wB + [|b2|]) - wwB) with + (([|a2|] - [|b2|]) * wB + [|a3|] + 2*[|b1|] * wB + 2*[|b2|]); + [zarith | ring]. + rewrite <- (Zmod_unique ([[r]] + ([|b1|] * wB + [|b2|])) wwB + 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 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. + split. rewrite H2;rewrite Hcmp;ring. + split. Spec_ww_to_Z r;zarith. + rewrite H2. + assert (([|a2|] - [|b2|]) * wB + [|a3|] < 0);zarith. + apply Zle_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith. + replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with + (([|a2|] - [|b2|] + 1) * wB + - 1);[zarith|ring]. + assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith. + replace 0 with (0*wB);zarith. + (* Cas Lt *) + assert (Hdiv21 := spec_div21 a2 Hle Hcmp); + destruct (w_div21 a1 a2 b1) as (q, r);destruct Hdiv21. + rewrite H. + assert (Hq := spec_to_Z q). + generalize + (spec_ww_sub_c (w_WW r a3) (w_mul_c q b2)); + destruct (ww_sub_c (w_WW r a3) (w_mul_c q b2)) + as [r1|r1];repeat (rewrite spec_w_WW || rewrite spec_mul_c); + unfold interp_carry;intros H1. + rewrite H1. + split. ring. split. + rewrite <- H1;destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z r1);trivial. + apply Zle_lt_trans with ([|r|] * wB + [|a3|]). + assert ( 0 <= [|q|] * [|b2|]);zarith. + apply beta_lex_inv;zarith. + assert ([[r1]] = [|r|] * wB + [|a3|] - [|q|] * [|b2|] + wwB). + rewrite <- H1;ring. + Spec_ww_to_Z r1; assert (0 <= [|r|]*wB). zarith. + assert (0 < [|q|] * [|b2|]). zarith. + 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) => + let (q0, r0) := res in + ([|q|] * [|b1|] + [|r|]) * wB + [|a3|] = + [|q0|] * ([|b1|] * wB + [|b2|]) + [[r0]] /\ + 0 <= [[r0]] < [|b1|] * wB + [|b2|]);eauto. + intros r2;repeat (rewrite spec_pred || rewrite spec_ww_add;eauto); + simpl ww_to_Z;intros H7. + assert (0 < [|q|] - 1). + assert (1 <= [|q|]). zarith. + destruct (Zle_lt_or_eq _ _ H6);zarith. + rewrite <- H8 in H2;rewrite H2 in H7. + assert (0 < [|b1|]*wB). apply Zmult_lt_0_compat;zarith. + Spec_ww_to_Z r2. zarith. + rewrite (Zmod_small ([|q|] -1));zarith. + rewrite (Zmod_small ([|q|] -1 -1));zarith. + assert ([[r2]] + ([|b1|] * wB + [|b2|]) = + 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|]) + < [|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|]) + < 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 + ([[r2]] + ([|b1|] * wB + [|b2|])) + wwB + 1 + ([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2*([|b1|] * wB + [|b2|])) + H10 H8). + split. ring. zarith. + intros r2;repeat (rewrite spec_pred);simpl ww_to_Z;intros H7. + rewrite (Zmod_small ([|q|] -1));zarith. + split. + replace [[r2]] with ([[r1]] + ([|b1|] * wB + [|b2|]) -wwB). + rewrite H2; ring. rewrite <- H7; ring. + Spec_ww_to_Z r2;Spec_ww_to_Z r1. omega. + simpl in Hlt. + assert ([|a1|] * wB + [|a2|] <= [|b1|] * wB + [|b2|]). zarith. + assert (H1 := beta_lex _ _ _ _ _ H HH0 HH3). rewrite spec_w_0;simpl;zarith. + Qed. + + +End DoubleDiv32. + +Section DoubleDiv21. + Variable w : Type. + Variable w_0 : w. + + Variable w_0W : w -> zn2z w. + Variable w_div32 : w -> w -> w -> w -> w -> w * zn2z w. + + Variable ww_1 : zn2z w. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable ww_sub : zn2z w -> zn2z w -> zn2z w. + + + Definition ww_div21 a1 a2 b := + match a1 with + | W0 => + match ww_compare a2 b with + | Gt => (ww_1, ww_sub a2 b) + | Eq => (ww_1, W0) + | Lt => (W0, a2) + end + | WW a1h a1l => + match a2 with + | W0 => + match b with + | W0 => (W0,W0) (* cas absurde *) + | WW b1 b2 => + let (q1, r) := w_div32 a1h a1l w_0 b1 b2 in + match r with + | W0 => (WW q1 w_0, W0) + | WW r1 r2 => + let (q2, s) := w_div32 r1 r2 w_0 b1 b2 in + (WW q1 q2, s) + end + end + | WW a2h a2l => + match b with + | W0 => (W0,W0) (* cas absurde *) + | WW b1 b2 => + let (q1, r) := w_div32 a1h a1l a2h b1 b2 in + match r with + | W0 => (WW q1 w_0, w_0W a2l) + | WW r1 r2 => + let (q2, s) := w_div32 r1 r2 a2l b1 b2 in + (WW q1 q2, s) + end + end + end + end. + + (* Proof *) + + Variable w_digits : positive. + 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). + 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). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_w_div32 : forall a1 a2 a3 b1 b2, + 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]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]. + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. + + Theorem wwB_div: wwB = 2 * (wwB / 2). + Proof. + rewrite wwB_div_2; rewrite Zmult_assoc; rewrite wB_div_2; auto. + rewrite <- Zpower_2; apply wwB_wBwB. + Qed. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Theorem spec_ww_div21 : forall a1 a2 b, + wwB/2 <= [[b]] -> + [[a1]] < [[b]] -> + let (q,r) := ww_div21 a1 a2 b in + [[a1]] *wwB+[[a2]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]. + Proof. + assert (U:= lt_0_wB w_digits). + assert (U1:= lt_0_wwB w_digits). + intros a1 a2 b H Hlt; unfold ww_div21. + Spec_ww_to_Z b; assert (Eq: 0 < [[b]]). Spec_ww_to_Z a1;omega. + generalize Hlt H ;clear Hlt H;case a1. + intros H1 H2;simpl in H1;Spec_ww_to_Z a2; + match goal with |-context [ww_compare ?Y ?Z] => + generalize (spec_ww_compare Y Z); case (ww_compare Y Z) + end; simpl;try rewrite spec_ww_1;autorewrite with rm10; intros;zarith. + rewrite spec_ww_sub;simpl. rewrite Zmod_small;zarith. + split. ring. + assert (wwB <= 2*[[b]]);zarith. + rewrite wwB_div;zarith. + intros a1h a1l. Spec_w_to_Z a1h;Spec_w_to_Z a1l. Spec_ww_to_Z a2. + destruct a2 as [ |a3 a4]; + (destruct b as [ |b1 b2];[unfold Zle in Eq;discriminate Eq|idtac]); + try (Spec_w_to_Z a3; Spec_w_to_Z a4); Spec_w_to_Z b1; Spec_w_to_Z b2; + intros Hlt H; match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] => + generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U); + intros q1 r H0 + end; (assert (Eq1: wB / 2 <= [|b1|]);[ + apply (@beta_lex (wB / 2) 0 [|b1|] [|b2|] wB); auto with zarith; + autorewrite with rm10;repeat rewrite (Zmult_comm wB); + 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) ]). + split;[rewrite wwB_wBwB; rewrite Zpower_2 | trivial]. + rewrite Zmult_assoc;rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc; + rewrite <- Zpower_2; rewrite <- wwB_wBwB;rewrite H1;ring. + destruct H2 as (H2,H3);match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] => + generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U); + intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end. + split;[rewrite wwB_wBwB | trivial]. + rewrite Zpower_2. + rewrite Zmult_assoc;rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc; + rewrite <- Zpower_2. + rewrite <- 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. + split;[rewrite wwB_wBwB | split;zarith]. + 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. + assert (H_:= wB_pos w_digits);apply Zdiv_le_lower_bound;zarith. + destruct H2 as (H2,H3);match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] => + generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U); + intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end. + split;trivial. + replace (([|a1h|] * wB + [|a1l|]) * wwB + ([|a3|] * wB + [|a4|])) with + (([|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|])); + [rewrite H4;simpl|rewrite wwB_wBwB];ring. + Qed. + +End DoubleDiv21. + +Section DoubleDivGt. + Variable w : Type. + Variable w_digits : positive. + Variable w_0 : w. + + Variable w_WW : w -> w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable w_eq0 : w -> bool. + Variable w_opp_c : w -> carry w. + Variable w_opp w_opp_carry : w -> w. + Variable w_sub_c : w -> w -> carry w. + Variable w_sub w_sub_carry : w -> w -> w. + + Variable w_div_gt : w -> w -> w*w. + Variable w_mod_gt : w -> w -> w. + Variable w_gcd_gt : w -> w -> w. + Variable w_add_mul_div : w -> w -> w -> w. + Variable w_head0 : w -> w. + Variable w_div21 : w -> w -> w -> w * w. + Variable w_div32 : w -> w -> w -> w -> w -> w * zn2z w. + + + Variable _ww_zdigits : zn2z w. + Variable ww_1 : zn2z w. + Variable ww_add_mul_div : zn2z w -> zn2z w -> zn2z w -> zn2z w. + + Variable w_zdigits : w. + + Definition ww_div_gt_aux ah al bh bl := + Eval lazy beta iota delta [ww_sub ww_opp] in + let p := w_head0 bh in + match w_compare p w_0 with + | 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_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 + 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 + | W0, _ => (W0,W0) + | _, W0 => (W0,W0) + | WW ah al, WW bh bl => + if w_eq0 ah then + let (q,r) := w_div_gt al bl in + (WW w_0 q, w_0W r) + else + match w_compare w_0 bh with + | Eq => + let(q,r):= + 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. + + Definition ww_mod_gt_aux ah al bh bl := + Eval lazy beta iota delta [ww_sub ww_opp] in + let p := w_head0 bh in + match w_compare p w_0 with + | 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_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 + 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 + | W0, _ => W0 + | _, W0 => W0 + | WW ah al, WW bh bl => + if w_eq0 ah then w_0W (w_mod_gt al bl) + 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 + w_compare w_sub 1 a bl) + | Lt => ww_mod_gt_aux ah al bh bl + | Gt => W0 (* cas absurde *) + end + 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 + 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 => + 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) + | Gt => W0 (* absurde *) + end + | Lt => + let m := ww_mod_gt_aux ah al bh bl in + match m with + | 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 + w_compare w_sub 1 (WW bh bl) ml in + WW w_0 (w_gcd_gt ml r) + end + | Lt => + let r := ww_mod_gt_aux bh bl mh ml in + match r with + | W0 => m + | WW rh rl => cont mh ml rh rl + end + | Gt => W0 (* absurde *) + end + end + | Gt => W0 (* absurde *) + end. + + 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 + (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. + 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). + Notation "[-| c |]" := + (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). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | 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. + + Variable spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|]. + Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + Variable spec_sub_carry : + forall x y, [|w_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB. + + Variable spec_div_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> + let (q,r) := w_div_gt a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Variable spec_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> + [|w_mod_gt a b|] = [|a|] mod [|b|]. + Variable spec_gcd_gt : forall a b, [|a|] > [|b|] -> + Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|]. + + Variable spec_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_head0 : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ [|w_head0 x|] * [|x|] < wB. + + Variable spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + + Variable spec_w_div32 : forall a1 a2 a3 b1 b2, + 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]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]. + + Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits. + + Variable spec_ww_digits_ : [[_ww_zdigits]] = Zpos (xO w_digits). + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_ww_add_mul_div : forall x y p, + [[p]] <= Zpos (xO w_digits) -> + [[ ww_add_mul_div p x y ]] = + ([[x]] * (2^[[p]]) + + [[y]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Lemma to_Z_div_minus_p : forall x p, + 0 < [|p|] < Zpos w_digits -> + 0 <= [|x|] / 2 ^ (Zpos w_digits - [|p|]) < 2 ^ [|p|]. + Proof. + intros x p H;Spec_w_to_Z x. + split. apply Zdiv_le_lower_bound;zarith. + apply Zdiv_lt_upper_bound;zarith. + rewrite <- Zpower_exp;zarith. + ring_simplify ([|p|] + (Zpos w_digits - [|p|])); unfold base in HH;zarith. + Qed. + Hint Resolve to_Z_div_minus_p : zarith. + + Lemma spec_ww_div_gt_aux : forall ah al bh bl, + [[WW ah al]] > [[WW bh bl]] -> + 0 < [|bh|] -> + let (q,r) := ww_div_gt_aux ah al bh bl in + [[WW ah al]] = [[q]] * [[WW bh bl]] + [[r]] /\ + 0 <= [[r]] < [[WW bh bl]]. + Proof. + intros ah al bh bl Hgt Hpos;unfold ww_div_gt_aux. + change + (let (q, r) := let p := w_head0 bh in + match w_compare p w_0 with + | 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_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 in [[WW ah al]]=[[q]]*[[WW bh bl]]+[[r]] /\ 0 <=[[r]]< [[WW bh bl]]). + assert (Hh := spec_head0 Hpos). + lazy zeta. + generalize (spec_compare (w_head0 bh) w_0); case w_compare; + rewrite spec_w_0; intros HH. + generalize Hh; rewrite HH; simpl Zpower; + rewrite Zmult_1_l; intros (HH1, HH2); clear HH. + assert (wwB <= 2*[[WW bh bl]]). + apply Zle_trans with (2*[|bh|]*wB). + rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat_r; zarith. + rewrite <- wB_div_2; apply Zmult_le_compat_l; zarith. + simpl ww_to_Z;rewrite Zmult_plus_distr_r;rewrite Zmult_assoc. + Spec_w_to_Z bl;zarith. + Spec_ww_to_Z (WW ah al). + rewrite spec_ww_sub;eauto. + simpl;rewrite spec_ww_1;rewrite Zmult_1_l;simpl. + simpl ww_to_Z in Hgt, H, HH;rewrite Zmod_small;split;zarith. + case (spec_to_Z (w_head0 bh)); auto with zarith. + assert ([|w_head0 bh|] < Zpos w_digits). + destruct (Z_lt_ge_dec [|w_head0 bh|] (Zpos w_digits));trivial. + elimtype False. + 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. + unfold base;apply Zpower_le_monotone;zarith. + assert (HHHH : 0 < [|w_head0 bh|] < Zpos w_digits); auto with zarith. + assert (Hb:= Zlt_le_weak _ _ H). + generalize (spec_add_mul_div w_0 ah Hb) + (spec_add_mul_div ah al Hb) + (spec_add_mul_div al w_0 Hb) + (spec_add_mul_div bh bl Hb) + (spec_add_mul_div bl w_0 Hb); + rewrite spec_w_0; repeat rewrite Zmult_0_l;repeat rewrite Zplus_0_l; + rewrite Zdiv_0_l;repeat rewrite Zplus_0_r. + 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. + 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. + intros U1 U2 U3 V1 V2. + generalize (@spec_w_div32 (w_add_mul_div (w_head0 bh) w_0 ah) + (w_add_mul_div (w_head0 bh) ah al) + (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)). + destruct (w_div32 (w_add_mul_div (w_head0 bh) w_0 ah) + (w_add_mul_div (w_head0 bh) ah al) + (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)). + unfold base;rewrite <- shift_unshift_mod;zarith. fold wB. + replace ([|bh|] * 2 ^ [|w_head0 bh|] * wB + [|bl|] * 2 ^ [|w_head0 bh|]) with + ([[WW bh bl]] * 2^[|w_head0 bh|]). 2:simpl;ring. + fold wwB. rewrite wwB_wBwB. rewrite Zpower_2. rewrite U1;rewrite U2;rewrite U3. + rewrite Zmult_assoc. rewrite Zmult_plus_distr_l. + rewrite (Zplus_assoc ([|ah|] / 2^(Zpos(w_digits) - [|w_head0 bh|])*wB * wB)). + rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc. + 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. + intros Hd;destruct Hd;zarith. + simpl. apply beta_lex_inv;zarith. rewrite U1;rewrite V1. + assert ([|ah|] / 2 ^ (Zpos (w_digits) - [|w_head0 bh|]) < wB/2);zarith. + apply Zdiv_lt_upper_bound;zarith. + unfold base. + replace (2^Zpos (w_digits)) with (2^(Zpos (w_digits) - 1)*2). + rewrite Z_div_mult;zarith. rewrite <- Zpower_exp;zarith. + apply Zlt_le_trans with wB;zarith. + unfold base;apply Zpower_le_monotone;zarith. + pattern 2 at 2;replace 2 with (2^1);trivial. + rewrite <- Zpower_exp;zarith. ring_simplify (Zpos (w_digits) - 1 + 1);trivial. + change [[WW w_0 q]] with ([|w_0|]*wB+[|q|]);rewrite spec_w_0;rewrite + Zmult_0_l;rewrite Zplus_0_l. + replace [[ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry + _ww_zdigits (w_0W (w_head0 bh))) W0 r]] with ([[r]]/2^[|w_head0 bh|]). + assert (0 < 2^[|w_head0 bh|]). apply Zpower_gt_0;zarith. + split. + rewrite <- (Z_div_mult [[WW ah al]] (2^[|w_head0 bh|]));zarith. + rewrite H1;rewrite Zmult_assoc;apply Z_div_plus_l;trivial. + split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith. + rewrite spec_ww_add_mul_div. + rewrite spec_ww_sub; auto with zarith. + rewrite spec_ww_digits_. + change (Zpos (xO (w_digits))) with (2*Zpos (w_digits));zarith. + simpl ww_to_Z;rewrite Zmult_0_l;rewrite Zplus_0_l. + rewrite spec_w_0W. + rewrite (fun x y => Zmod_small (x-y)); auto with zarith. + ring_simplify (2 * Zpos w_digits - (2 * Zpos w_digits - [|w_head0 bh|])). + rewrite Zmod_small;zarith. + split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith. + Spec_ww_to_Z r. + apply Zlt_le_trans with wwB;zarith. + rewrite <- (Zmult_1_r wwB);apply Zmult_le_compat;zarith. + split; auto with zarith. + apply Zle_lt_trans with (2 * Zpos w_digits); auto with zarith. + unfold base, ww_digits; rewrite (Zpos_xO w_digits). + apply Zpower2_lt_lin; auto with zarith. + rewrite spec_ww_sub; auto with zarith. + rewrite spec_ww_digits_; rewrite spec_w_0W. + rewrite Zmod_small;zarith. + rewrite Zpos_xO; split; auto with zarith. + apply Zle_lt_trans with (2 * Zpos w_digits); auto with zarith. + unfold base, ww_digits; rewrite (Zpos_xO w_digits). + apply Zpower2_lt_lin; auto with zarith. + Qed. + + Lemma spec_ww_div_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + let (q,r) := ww_div_gt a b in + [[a]] = [[q]] * [[b]] + [[r]] /\ + 0 <= [[r]] < [[b]]. + Proof. + intros a b Hgt Hpos;unfold ww_div_gt. + change (let (q,r) := match a, b with + | W0, _ => (W0,W0) + | _, W0 => (W0,W0) + | WW ah al, WW bh bl => + if w_eq0 ah then + let (q,r) := w_div_gt al bl in + (WW w_0 q, w_0W r) + else + match w_compare w_0 bh with + | Eq => + let(q,r):= + 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]]). + 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). + 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. + assert (H2:=spec_div_gt Hgt Hpos);destruct (w_div_gt al bl). + repeat rewrite spec_w_0W;simpl;rewrite spec_w_0;simpl;trivial. + clear H. + assert (Hcmp := spec_compare w_0 bh); destruct (w_compare w_0 bh). + rewrite spec_w_0 in Hcmp. change [[WW bh bl]] with ([|bh|]*wB+[|bl|]). + rewrite <- Hcmp;rewrite Zmult_0_l;rewrite Zplus_0_l. + simpl in Hpos;rewrite <- Hcmp in Hpos;simpl in Hpos. + assert (H2:= @spec_double_divn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div + w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 + 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 + 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. + Qed. + + Lemma spec_ww_mod_gt_aux_eq : forall ah al bh bl, + ww_mod_gt_aux ah al bh bl = snd (ww_div_gt_aux ah al bh bl). + Proof. + intros ah al bh bl. unfold ww_mod_gt_aux, ww_div_gt_aux. + case w_compare; auto. + case w_div32; auto. + Qed. + + Lemma spec_ww_mod_gt_aux : forall ah al bh bl, + [[WW ah al]] > [[WW bh bl]] -> + 0 < [|bh|] -> + [[ww_mod_gt_aux ah al bh bl]] = [[WW ah al]] mod [[WW bh bl]]. + Proof. + intros. rewrite spec_ww_mod_gt_aux_eq;trivial. + assert (H3 := spec_ww_div_gt_aux ah al bl H H0). + destruct (ww_div_gt_aux ah al bh bl) as (q,r);simpl. simpl in H,H3. + destruct H3;apply Zmod_unique with [[q]];zarith. + rewrite H1;ring. + Qed. + + Lemma spec_w_mod_gt_eq : forall a b, [|a|] > [|b|] -> 0 <[|b|] -> + [|w_mod_gt a b|] = [|snd (w_div_gt a b)|]. + Proof. + intros a b Hgt Hpos. + rewrite spec_mod_gt;trivial. + assert (H:=spec_div_gt Hgt Hpos). + destruct (w_div_gt a b) as (q,r);simpl. + rewrite Zmult_comm in H;destruct H. + 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 + (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 + match w_compare w_0 bh with + | 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 + (match a, b with + | W0, _ => (W0,W0) + | _, W0 => (W0,W0) + | WW ah al, WW bh bl => + if w_eq0 ah then + let (q,r) := w_div_gt al bl in + (WW w_0 q, w_0W r) + else + match w_compare w_0 bh with + | Eq => + let(q,r):= + 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). + destruct a as [ |ah al];trivial. + destruct b as [ |bh bl];trivial. + 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). + 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. + rewrite spec_w_0W;rewrite spec_w_mod_gt_eq;trivial. + destruct (w_div_gt al bl);simpl;rewrite spec_w_0W;trivial. + clear H. + 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 + 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. + rewrite spec_ww_mod_gt_aux_eq;trivial;symmetry;trivial. + trivial. + Qed. + + Lemma spec_ww_mod_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + [[ww_mod_gt a b]] = [[a]] mod [[b]]. + Proof. + intros a b Hgt Hpos. + assert (H:= spec_ww_div_gt a b Hgt Hpos). + rewrite (spec_ww_mod_gt_eq a b Hgt Hpos). + destruct (ww_div_gt a b)as(q,r);destruct H. + apply Zmod_unique with[[q]];simpl;trivial. + rewrite Zmult_comm;trivial. + Qed. + + 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). + pattern a at 1;rewrite (Z_div_mod_eq a b). + ring_simplify (b * (a / b) + a mod b - a / b * b);trivial. zarith. + Qed. + + Lemma spec_ww_gcd_gt_aux_body : + forall ah al bh bl n cont, + [[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) -> + 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. + intros ah al bh bl n cont Hlog Hgt Hcont. + change (ww_gcd_gt_body cont ah al bh bl) with (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 => + 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) + | Gt => W0 (* absurde *) + end + | Lt => + let m := ww_mod_gt_aux ah al bh bl in + match m with + | 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 + w_compare w_sub 1 (WW bh bl) ml in + WW w_0 (w_gcd_gt ml r) + end + | Lt => + let r := ww_mod_gt_aux bh bl mh ml in + match r with + | W0 => m + | WW rh rl => cont mh ml rh rl + end + | Gt => W0 (* absurde *) + end + end + | Gt => W0 (* absurde *) + end). + 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). + 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. + apply Zis_gcd_mod;zarith. + change ([|ah|] * wB + [|al|]) with (double_to_Z w_digits w_to_Z 1 (WW ah al)). + rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div + 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 => + 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 Hbh;assert (H:= spec_ww_mod_gt_aux _ _ _ Hgt Hbh). + assert (H2 : 0 < [[WW bh bl]]). + simpl;Spec_w_to_Z bl. apply Zlt_le_trans with ([|bh|]*wB);zarith. + apply Zmult_lt_0_compat;zarith. + 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;rewrite spec_w_0 in Hmh; rewrite <- Hmh;simpl. + 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. + 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 => + 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 Hmh. assert ([[WW bh bl]] > [[WW mh ml]]). + 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]]). + 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 => + 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 ([|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]] = + [[WW bh bl]] - [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])). + simpl;pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3';ring. simpl in H6'. + assert ([[WW mh ml]] <= [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])). + simpl;pattern ([|mh|]*wB+[|ml|]) at 1;rewrite <- Zmult_1_r;zarith. + simpl in *;assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in H8; + zarith. + assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in *;zarith. + rewrite <- H4 in H3';rewrite Zmult_0_r in H3';simpl in H3';zarith. + pattern n at 1;replace n with (n-1+1);try ring. + rewrite Zpower_exp;zarith. change (2^1) with 2. + rewrite Z_div_mult;zarith. + 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. + Qed. + + Lemma spec_ww_gcd_gt_aux : + forall p cont n, + (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]] -> + [[WW bh bl]] <= 2^(Zpos p + n) -> + Zis_gcd [[WW ah al]] [[WW bh bl]] + [[ww_gcd_gt_aux p cont ah al bh bl]]. + Proof. + induction p;intros cont n Hcont ah al bh bl Hgt Hs;simpl ww_gcd_gt_aux. + assert (0 < Zpos p). unfold Zlt;reflexivity. + apply spec_ww_gcd_gt_aux_body with (n := Zpos (xI p) + n); + trivial;rewrite Zpos_xI. + intros. apply IHp with (n := Zpos p + n);zarith. + intros. apply IHp with (n := n );zarith. + apply Zle_trans with (2 ^ (2* Zpos p + 1+ n -1));zarith. + apply Zpower_le_monotone2;zarith. + assert (0 < Zpos p). unfold Zlt;reflexivity. + apply spec_ww_gcd_gt_aux_body with (n := Zpos (xO p) + n );trivial. + rewrite (Zpos_xO p). + intros. apply IHp with (n := Zpos p + n - 1);zarith. + intros. apply IHp with (n := n -1 );zarith. + intros;apply Hcont;zarith. + apply Zle_trans with (2^(n-1));zarith. + apply Zpower_le_monotone2;zarith. + apply Zle_trans with (2 ^ (Zpos p + n -1));zarith. + apply Zpower_le_monotone2;zarith. + apply Zle_trans with (2 ^ (2*Zpos p + n -1));zarith. + apply Zpower_le_monotone2;zarith. + apply spec_ww_gcd_gt_aux_body with (n := n+1);trivial. + rewrite Zplus_comm;trivial. + ring_simplify (n + 1 - 1);trivial. + Qed. + +End DoubleDivGt. + +Section DoubleDiv. + + Variable w : Type. + Variable w_digits : positive. + Variable ww_1 : zn2z w. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + + 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 + | 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 + | Eq => W0 + | Lt => a + end. + + 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). + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_div_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + let (q,r) := ww_div_gt a b in + [[a]] = [[q]] * [[b]] + [[r]] /\ + 0 <= [[r]] < [[b]]. + Variable spec_ww_mod_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + [[ww_mod_gt a b]] = [[a]] mod [[b]]. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Lemma spec_ww_div : forall a b, 0 < [[b]] -> + let (q,r) := ww_div a b in + [[a]] = [[q]] * [[b]] + [[r]] /\ + 0 <= [[r]] < [[b]]. + 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;split;[ring|Spec_ww_to_Z a;zarith]. + apply spec_ww_div_gt;trivial. + Qed. + + Lemma spec_ww_mod : forall a b, 0 < [[b]] -> + [[ww_mod a b]] = [[a]] mod [[b]]. + Proof. + 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. + apply spec_ww_mod_gt;trivial. + Qed. + + + Variable w_0 : w. + Variable w_1 : w. + Variable w_compare : w -> w -> comparison. + Variable w_eq0 : w -> bool. + 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) -> + w -> w -> w -> w -> zn2z w. + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + 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 : + forall p cont n, + (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]] -> + [[WW bh bl]] <= 2^(Zpos p + n) -> + 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) := + match w_compare w_1 yl with + | Eq => ww_1 + | _ => WW xh xl + end. + + 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. + 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. + unfold gcd_cont;assert (Hcmpy:=spec_compare w_1 yl); + rewrite spec_w_1 in Hcmpy. + 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. + 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]] -> + [[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 + | W0, _ => b + | _, W0 => a + | WW ah al, WW bh bl => + if w_eq0 ah then (WW w_0 (w_gcd_gt al bl)) + else ww_gcd_gt_fix _ww_digits cont ah al bh bl + end. + + Definition ww_gcd a b := + Eval lazy beta delta [ww_gcd_gt] in + match ww_compare a b with + | Gt => ww_gcd_gt a b + | Eq => a + | Lt => ww_gcd_gt b a + end. + + Lemma spec_ww_gcd_gt : forall a b, [[a]] > [[b]] -> + Zis_gcd [[a]] [[b]] [[ww_gcd_gt a b]]. + Proof. + intros a b Hgt;unfold ww_gcd_gt. + 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). + 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. + apply spec_gcd_gt_fix with (n:= 0);trivial. + rewrite Zplus_0_r;rewrite spec_ww_digits_. + change (2 ^ Zpos (xO w_digits)) with wwB. Spec_ww_to_Z (WW bh bl);zarith. + Qed. + + 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 + (match ww_compare a b with + | Gt => ww_gcd_gt a b + | Eq => a + | Lt => ww_gcd_gt b a + end). + assert (Hcmp := spec_ww_compare a b);destruct (ww_compare a b). + Spec_ww_to_Z b;rewrite Hcmp. + apply Zis_gcd_for_euclid with 1;zarith. + ring_simplify ([[b]] - 1 * [[b]]). apply Zis_gcd_0;zarith. + apply Zis_gcd_sym;apply spec_ww_gcd_gt;zarith. + apply spec_ww_gcd_gt;zarith. + Qed. + +End DoubleDiv. + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v new file mode 100644 index 00000000..d6f6a05f --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v @@ -0,0 +1,528 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleDivn1.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section GENDIVN1. + + Variable w : Type. + Variable w_digits : positive. + Variable w_zdigits : w. + Variable w_0 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_head0 : w -> w. + Variable w_add_mul_div : w -> w -> w -> w. + 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 "[| 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) + (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. + Variable spec_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_head0 : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ [|w_head0 x|] * [|x|] < wB. + Variable spec_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_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_sub: forall x y, + [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + + + + Section DIVAUX. + Variable b2p : w. + Variable b2p_le : wB/2 <= [|b2p|]. + + Definition double_divn1_0_aux n (divn1: w -> word w n -> word w n * w) r h := + let (hh,hl) := double_split w_0 n h in + let (qh,rh) := divn1 r hh in + let (ql,rl) := divn1 rh hl in + (double_WW w_WW n qh ql, rl). + + 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) + 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!]. + Proof (spec_double_split w_0 w_digits w_to_Z spec_0). + + Lemma spec_double_divn1_0 : forall n r a, + [|r|] < [|b2p|] -> + let (q,r') := double_divn1_0 n r a in + [|r|] * double_wB w_digits n + [!n|a!] = [!n|q!] * [|b2p|] + [|r'|] /\ + 0 <= [|r'|] < [|b2p|]. + Proof. + induction n;intros. + exact (spec_div21 a b2p_le H). + simpl (double_divn1_0 (S n) r a); unfold double_divn1_0_aux. + assert (H1 := spec_split n a);destruct (double_split w_0 n a) as (hh,hl). + rewrite H1. + assert (H2 := IHn r hh H);destruct (double_divn1_0 n r hh) as (qh,rh). + destruct H2. + assert ([|rh|] < [|b2p|]). omega. + assert (H4 := IHn rh hl H3);destruct (double_divn1_0 n rh hl) as (ql,rl). + destruct H4;split;trivial. + rewrite spec_double_WW;trivial. + rewrite <- double_wB_wwB. + rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite H0;rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc. + rewrite H4;ring. + Qed. + + Definition double_modn1_0_aux n (modn1:w -> word w n -> w) r h := + let (hh,hl) := double_split w_0 n h in modn1 (modn1 r hh) hl. + + Fixpoint double_modn1_0 (n:nat) : w -> word w n -> w := + match n return w -> word w n -> w with + | O => fun r x => snd (w_div21 r x b2p) + | S n => double_modn1_0_aux n (double_modn1_0 n) + end. + + Lemma spec_double_modn1_0 : forall n r x, + double_modn1_0 n r x = snd (double_divn1_0 n r x). + Proof. + 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). + 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. + + Lemma spec_add_mul_divp : forall x y, + [| w_add_mul_div p x y |] = + ([|x|] * (2 ^ [|p|]) + + [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB. + Proof. + 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 := + let (hh,hl) := double_split w_0 n h 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) + end. + + Lemma p_lt_double_digits : forall n, [|p|] <= Zpos (double_digits w_digits n). + Proof. +(* + induction n;simpl. destruct p_bounded;trivial. + case (spec_to_Z p); rewrite Zpos_xO;auto with zarith. +*) + induction n;simpl. trivial. + case (spec_to_Z p); rewrite Zpos_xO;auto with zarith. + Qed. + + 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|] + + [!n|l!] / (2^(Zpos(double_digits w_digits n) - [|p|]))) + mod double_wB w_digits n = [!n|q!] * [|b2p|] + [|r'|] /\ + 0 <= [|r'|] < [|b2p|]. + Proof. + case (spec_to_Z p); intros HH0 HH1. + induction n;intros. + simpl (double_divn1_p 0 r h l). + unfold double_to_Z, double_wB, double_digits. + rewrite <- spec_add_mul_divp. + exact (spec_div21 (w_add_mul_div p h l) b2p_le H). + simpl (double_divn1_p (S n) r h l). + unfold double_divn1_p_aux. + assert (H1 := spec_split n h);destruct (double_split w_0 n h) as (hh,hl). + rewrite H1. rewrite <- double_wB_wwB. + assert (H2 := spec_split n l);destruct (double_split w_0 n l) as (lh,ll). + rewrite H2. + replace ([|r|] * (double_wB w_digits n * double_wB w_digits n) + + (([!n|hh!] * double_wB w_digits n + [!n|hl!]) * 2 ^ [|p|] + + ([!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|] + + [!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 + 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. + 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 + ([!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]. + 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); + unfold double_wB,base in Uhl. + assert (Ulh := spec_double_to_Z w_digits w_to_Z spec_to_Z n lh); + unfold double_wB,base in Ulh. + assert (Ull := spec_double_to_Z w_digits w_to_Z spec_to_Z n ll); + 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))) + 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|]). + 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 + ([!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 / + 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 + (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)). + 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. + rewrite Zpower_exp;auto with zarith. + assert (0 < Zpos (double_digits w_digits n)). unfold Zlt;reflexivity. + auto with zarith. + apply Z_mod_lt;auto with zarith. + rewrite Zpower_exp;auto with zarith. + 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 + (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|] + + Zpos (double_digits w_digits n));trivial. + 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 + 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) + | S n => double_modn1_p_aux n (double_modn1_p n) + end. + + Lemma spec_double_modn1_p : forall n r h l , + double_modn1_p n r h l = snd (double_divn1_p n r h l). + Proof. + induction n;simpl;intros;trivial. + unfold double_modn1_p_aux, double_divn1_p_aux. + destruct(double_split w_0 n h)as(hh,hl);destruct(double_split w_0 n l) as (lh,ll). + rewrite (IHn r hh hl);destruct (double_divn1_p n r hh hl) as (qh,rh). + rewrite IHn;simpl;destruct (double_divn1_p n rh hl lh);trivial. + Qed. + + End DIVAUX. + + Fixpoint high (n:nat) : word w n -> w := + match n return word w n -> w with + | O => fun a => a + | S n => + fun (a:zn2z (word w n)) => + match a with + | W0 => w_0 + | WW h l => high n h + end + end. + + 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 + (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), + [|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 (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. + rewrite Zdiv_0_l;trivial. + assert (U0 := spec_double_to_Z w_digits w_to_Z spec_to_Z n w0); + assert (U1 := spec_double_to_Z w_digits w_to_Z spec_to_Z n w1). + 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)). + rewrite Zdiv_mult_cancel_r;auto with zarith. + rewrite <- Zpower_exp;auto with zarith. + 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 + (2*Zpos (double_digits w_digits n));ring. + 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) := + 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 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 + end. + + Lemma spec_double_divn1 : forall n a b, + 0 < [|b|] -> + let (q,r) := double_divn1 n a b in + [!n|a!] = [!n|q!] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + intros n a b H. unfold double_divn1. + case (spec_head0 H); intros H0 H1. + case (spec_to_Z (w_head0 b)); intros HH1 HH2. + generalize (spec_compare (w_head0 b) w_0); case w_compare; + rewrite spec_0; intros H2; auto with zarith. + assert (Hv1: wB/2 <= [|b|]). + generalize H0; rewrite H2; rewrite Zpower_0_r; + rewrite Zmult_1_l; auto. + assert (Hv2: [|w_0|] < [|b|]). + rewrite spec_0; auto. + generalize (spec_double_divn1_0 Hv1 n a Hv2). + rewrite spec_0;rewrite Zmult_0_l; rewrite Zplus_0_l; auto. + contradict H2; auto with zarith. + assert (HHHH : 0 < [|w_head0 b|]); auto with zarith. + assert ([|w_head0 b|] < Zpos w_digits). + case (Zle_or_lt (Zpos w_digits) [|w_head0 b|]); auto; intros HH. + assert (2 ^ [|w_head0 b|] < wB). + apply Zle_lt_trans with (2 ^ [|w_head0 b|] * [|b|]);auto with zarith. + replace (2 ^ [|w_head0 b|]) with (2^[|w_head0 b|] * 1);try (ring;fail). + apply Zmult_le_compat;auto with zarith. + assert (wB <= 2^[|w_head0 b|]). + unfold base;apply Zpower_le_monotone;auto with zarith. omega. + assert ([|w_add_mul_div (w_head0 b) b w_0|] = + 2 ^ [|w_head0 b|] * [|b|]). + rewrite (spec_add_mul_div b w_0); auto with zarith. + rewrite spec_0;rewrite Zdiv_0_l; try omega. + rewrite Zplus_0_r; rewrite Zmult_comm. + rewrite Zmod_small; auto with zarith. + assert (H5 := spec_to_Z (high n a)). + 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 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. + assert (H6 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|])); + auto with zarith. + rewrite Zmod_small;auto with zarith. + apply Zdiv_lt_upper_bound;auto with zarith. + apply Zlt_le_trans with wB;auto with zarith. + apply Zle_trans with (2 ^ [|w_head0 b|] * [|b|] * 2). + rewrite <- wB_div_2; try omega. + 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. + 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 + (double_0 w_0 n)) as (q,r). + assert (U:= spec_double_digits n). + rewrite spec_double_0 in H7;trivial;rewrite Zdiv_0_l in H7. + rewrite Zplus_0_r in H7. + rewrite spec_add_mul_div in H7;auto with zarith. + rewrite spec_0 in H7;rewrite Zmult_0_l in H7;rewrite Zplus_0_l in H7. + assert (([|high n a|] / 2 ^ (Zpos w_digits - [|w_head0 b|])) mod wB + = [!n|a!] / 2^(Zpos (double_digits w_digits n) - [|w_head0 b|])). + 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 + + (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. + split;auto with zarith. + apply Zle_lt_trans with ([|high n a|]);auto with zarith. + apply Zdiv_le_upper_bound;auto with zarith. + pattern ([|high n a|]) at 1;rewrite <- Zmult_1_r. + apply Zmult_le_compat;auto with zarith. + rewrite H8 in H7;unfold double_wB,base in H7. + rewrite <- shift_unshift_mod in H7;auto with zarith. + rewrite H4 in H7. + assert ([|w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r|] + = [|r|]/2^[|w_head0 b|]). + rewrite spec_add_mul_div. + rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l. + replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|]) + with ([|w_head0 b|]). + rewrite Zmod_small;auto with zarith. + assert (H9 := spec_to_Z r). + split;auto with zarith. + apply Zle_lt_trans with ([|r|]);auto with zarith. + apply Zdiv_le_upper_bound;auto with zarith. + pattern ([|r|]) at 1;rewrite <- Zmult_1_r. + apply Zmult_le_compat;auto with zarith. + assert (H10 := Zpower_gt_0 2 ([|w_head0 b|]));auto with zarith. + rewrite spec_sub. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + case (spec_to_Z w_zdigits); auto with zarith. + rewrite spec_sub. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + case (spec_to_Z w_zdigits); auto with zarith. + case H7; intros H71 H72. + 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|])) + with ([!n|q!] *[|b|] * 2^[|w_head0 b|]); + try (ring;fail). + rewrite Z_div_plus_l;auto with zarith. + 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. + apply Zdiv_lt_upper_bound;auto with zarith. + rewrite Zmult_comm;auto with zarith. + 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) := + 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 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 + end. + + Lemma spec_double_modn1_aux : forall n a b, + double_modn1 n a b = snd (double_divn1 n a b). + Proof. + intros n a b;unfold double_divn1,double_modn1. + generalize (spec_compare (w_head0 b) w_0); case w_compare; + rewrite spec_0; intros H2; auto with zarith. + apply spec_double_modn1_0. + apply spec_double_modn1_0. + rewrite spec_double_modn1_p. + 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 (double_0 w_0 n));simpl;trivial. + Qed. + + Lemma spec_double_modn1 : forall n a b, 0 < [|b|] -> + [|double_modn1 n a b|] = [!n|a!] mod [|b|]. + Proof. + intros n a b H;assert (H1 := spec_double_divn1 n a H). + assert (H2 := spec_double_modn1_aux n a b). + rewrite H2;destruct (double_divn1 n a b) as (q,r). + simpl;apply Zmod_unique with (double_to_Z w_digits w_to_Z n q);auto with zarith. + destruct H1 as (h1,h2);rewrite h1;ring. + Qed. + +End GENDIVN1. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v new file mode 100644 index 00000000..50c72487 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v @@ -0,0 +1,487 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleLift.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleLift. + Variable w : Type. + Variable w_0 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable w_head0 : w -> w. + Variable w_tail0 : w -> w. + Variable w_add: w -> w -> zn2z w. + Variable w_add_mul_div : w -> w -> w -> w. + Variable ww_sub: zn2z w -> zn2z w -> zn2z w. + Variable w_digits : positive. + Variable ww_Digits : positive. + Variable w_zdigits : w. + Variable ww_zdigits : zn2z w. + Variable low: zn2z w -> w. + + Definition ww_head0 x := + match x with + | W0 => ww_zdigits + | WW xh xl => + match w_compare w_0 xh with + | Eq => w_add w_zdigits (w_head0 xl) + | _ => w_0W (w_head0 xh) + end + end. + + + Definition ww_tail0 x := + match x with + | W0 => ww_zdigits + | WW xh xl => + match w_compare w_0 xl with + | Eq => w_add w_zdigits (w_tail0 xh) + | _ => w_0W (w_tail0 xl) + end + end. + + + (* 0 < p < ww_digits *) + 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 + | Lt => w_0W (w_add_mul_div (low p) w_0 yh) + | Gt => + let n := low (ww_sub p zdigits) in + w_WW (w_add_mul_div n w_0 yh) (w_add_mul_div n yh yl) + end + | WW xh xl, W0 => + match ww_compare p zdigits with + | 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) + end + | WW xh xl, WW yh yl => + match ww_compare p zdigits with + | 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 + w_WW (w_add_mul_div n xl yh) (w_add_mul_div n yh yl) + end + end. + + 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). + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_compare : forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_digits : ww_Digits = xO w_digits. + Variable spec_w_head00 : forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits. + Variable spec_w_head0 : forall x, 0 < [|x|] -> + 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|] -> + 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, + [[w_add x y]] = [|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. + Ltac zarith := auto with zarith lift. + + Lemma spec_ww_head00 : forall x, [[x]] = 0 -> [[ww_head0 x]] = Zpos ww_Digits. + Proof. + intros x; case x; unfold ww_head0. + intros HH; rewrite spec_ww_zdigits; auto. + intros xh xl; simpl; intros Hx. + case (spec_to_Z xh); intros Hx1 Hx2. + case (spec_to_Z xl); intros Hy1 Hy2. + assert (F1: [|xh|] = 0). + case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + apply Zlt_le_trans with (1 := Hy3); auto with zarith. + pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]). + apply Zplus_le_compat_r; auto with zarith. + case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith. + apply Zmult_lt_0_compat; auto with zarith. + generalize (spec_compare w_0 xh); case w_compare. + intros H; simpl. + rewrite spec_w_add; rewrite spec_w_head00. + rewrite spec_zdigits; rewrite spec_ww_digits. + rewrite Zpos_xO; auto with zarith. + rewrite F1 in Hx; auto with zarith. + 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. + clear spec_ww_zdigits. + rewrite wwB_div_2;rewrite Zmult_comm;rewrite wwB_wBwB. + assert (U:= lt_0_wB w_digits); destruct x as [ |xh xl];simpl ww_to_Z;intros H. + unfold Zlt in H;discriminate H. + 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_head0 xl)); intros HH1 HH2 HH3 HH4. + rewrite spec_w_add. + rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith. + case (spec_w_head0 H); intros H1 H2. + rewrite Zpower_2; fold wB; rewrite <- Zmult_assoc; split. + apply Zmult_le_compat_l; auto with zarith. + apply Zmult_lt_compat_l; auto with zarith. + assert (H1 := spec_w_head0 H0). + rewrite spec_w_0W. + split. + rewrite Zmult_plus_distr_r;rewrite Zmult_assoc. + apply Zle_trans with (2 ^ [|w_head0 xh|] * [|xh|] * wB). + rewrite Zmult_comm; zarith. + assert (0 <= 2 ^ [|w_head0 xh|] * [|xl|]);zarith. + assert (H2:=spec_to_Z xl);apply Zmult_le_0_compat;zarith. + case (spec_to_Z (w_head0 xh)); intros H2 _. + generalize ([|w_head0 xh|]) H1 H2;clear H1 H2; + intros p H1 H2. + assert (Eq1 : 2^p < wB). + rewrite <- (Zmult_1_r (2^p));apply Zle_lt_trans with (2^p*[|xh|]);zarith. + assert (Eq2: p < Zpos w_digits). + destruct (Zle_or_lt (Zpos w_digits) p);trivial;contradict Eq1. + apply Zle_not_lt;unfold base;apply Zpower_le_monotone;zarith. + assert (Zpos w_digits = p + (Zpos w_digits - p)). ring. + rewrite Zpower_2. + unfold base at 2;rewrite H3;rewrite Zpower_exp;zarith. + rewrite <- Zmult_assoc; apply Zmult_lt_compat_l; zarith. + rewrite <- (Zplus_0_r (2^(Zpos w_digits - p)*wB));apply beta_lex_inv;zarith. + apply Zmult_lt_reg_r with (2 ^ p); zarith. + rewrite <- Zpower_exp;zarith. + rewrite Zmult_comm;ring_simplify (Zpos w_digits - p + p);fold wB;zarith. + assert (H1 := spec_to_Z xh);zarith. + Qed. + + Lemma spec_ww_tail00 : forall x, [[x]] = 0 -> [[ww_tail0 x]] = Zpos ww_Digits. + Proof. + intros x; case x; unfold ww_tail0. + intros HH; rewrite spec_ww_zdigits; auto. + intros xh xl; simpl; intros Hx. + case (spec_to_Z xh); intros Hx1 Hx2. + case (spec_to_Z xl); intros Hy1 Hy2. + assert (F1: [|xh|] = 0). + case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + apply Zlt_le_trans with (1 := Hy3); auto with zarith. + pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]). + apply Zplus_le_compat_r; auto with zarith. + case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith. + apply Zmult_lt_0_compat; auto with zarith. + assert (F2: [|xl|] = 0). + rewrite F1 in Hx; auto with zarith. + generalize (spec_compare w_0 xl); case w_compare. + intros H; simpl. + rewrite spec_w_add; rewrite spec_w_tail00; auto. + rewrite spec_zdigits; rewrite spec_ww_digits. + rewrite Zpos_xO; auto with zarith. + rewrite spec_w_0; auto with zarith. + rewrite spec_w_0; auto with zarith. + Qed. + + Lemma spec_ww_tail0 : forall x, 0 < [[x]] -> + exists y, 0 <= y /\ [[x]] = (2 * y + 1) * 2 ^ [[ww_tail0 x]]. + Proof. + clear spec_ww_zdigits. + destruct x as [ |xh xl];simpl ww_to_Z;intros H. + unfold Zlt in H;discriminate H. + assert (H0 := spec_compare w_0 xl);rewrite spec_w_0 in H0. + destruct (w_compare w_0 xl). + rewrite <- H0; rewrite Zplus_0_r. + case (spec_to_Z (w_tail0 xh)); intros HH1 HH2. + generalize H; rewrite <- H0; rewrite Zplus_0_r; clear H; intros H. + case (@spec_w_tail0 xh). + apply Zmult_lt_reg_r with wB; auto with zarith. + unfold base; auto with zarith. + intros z (Hz1, Hz2); exists z; split; auto. + rewrite spec_w_add; rewrite (fun x => Zplus_comm [|x|]). + rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith. + rewrite Zmult_assoc; rewrite <- Hz2; auto. + + case (spec_to_Z (w_tail0 xh)); intros HH1 HH2. + case (spec_w_tail0 H0); intros z (Hz1, Hz2). + assert (Hp: [|w_tail0 xl|] < Zpos w_digits). + case (Zle_or_lt (Zpos w_digits) [|w_tail0 xl|]); auto; intros H1. + absurd (2 ^ (Zpos w_digits) <= 2 ^ [|w_tail0 xl|]). + apply Zlt_not_le. + case (spec_to_Z xl); intros HH3 HH4. + apply Zle_lt_trans with (2 := HH4). + apply Zle_trans with (1 * 2 ^ [|w_tail0 xl|]); auto with zarith. + rewrite Hz2. + apply Zmult_le_compat_r; auto with zarith. + apply Zpower_le_monotone; auto with zarith. + exists ([|xh|] * (2 ^ ((Zpos w_digits - [|w_tail0 xl|]) - 1)) + z); split. + apply Zplus_le_0_compat; auto. + apply Zmult_le_0_compat; auto with zarith. + case (spec_to_Z xh); auto. + rewrite spec_w_0W. + rewrite (Zmult_plus_distr_r 2); rewrite <- Zplus_assoc. + rewrite Zmult_plus_distr_l; rewrite <- Hz2. + apply f_equal2 with (f := Zplus); auto. + rewrite (Zmult_comm 2). + repeat rewrite <- Zmult_assoc. + apply f_equal2 with (f := Zmult); auto. + case (spec_to_Z (w_tail0 xl)); intros HH3 HH4. + pattern 2 at 2; rewrite <- Zpower_1_r. + lazy beta; repeat rewrite <- Zpower_exp; auto with zarith. + unfold base; apply f_equal with (f := Zpower 2); auto with zarith. + + contradict H0; case (spec_to_Z xl); auto with zarith. + Qed. + + Hint Rewrite Zdiv_0_l Zmult_0_l Zplus_0_l Zmult_0_r Zplus_0_r + 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. + + Lemma spec_ww_add_mul_div_aux : forall xh xl yh yl p, + let zdigits := w_0W w_zdigits in + [[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) + (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]] = + ([[WW xh xl]] * (2^[[p]]) + + [[WW yh yl]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB. + Proof. + clear spec_ww_zdigits. + intros xh xl yh yl p zdigits;assert (HwwB := wwB_pos w_digits). + case (spec_to_w_Z p); intros Hv1 Hv2. + replace (Zpos (xO w_digits)) with (Zpos w_digits + Zpos w_digits). + 2 : rewrite Zpos_xO;ring. + replace (Zpos w_digits + Zpos w_digits - [[p]]) with + (Zpos w_digits + (Zpos w_digits - [[p]])). 2:ring. + intros Hp; assert (Hxh := spec_to_Z xh);assert (Hxl:=spec_to_Z xl); + assert (Hx := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)); + simpl in Hx;assert (Hyh := spec_to_Z yh);assert (Hyl:=spec_to_Z yl); + assert (Hy:=spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW yh yl));simpl in Hy. + generalize (spec_ww_compare p zdigits); case ww_compare; intros H1. + rewrite H1; unfold zdigits; rewrite spec_w_0W. + rewrite spec_zdigits; rewrite Zminus_diag; rewrite Zplus_0_r. + simpl ww_to_Z; w_rewrite;zarith. + fold wB. + rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;rewrite <- Zplus_assoc. + rewrite <- Zpower_2. + rewrite <- 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]]). + rewrite spec_low. + apply Zmod_small. + case (spec_to_w_Z p); intros HH1 HH2; split; auto. + generalize H1; unfold zdigits; rewrite spec_w_0W; + rewrite spec_zdigits; intros tmp. + apply Zlt_le_trans with (1 := tmp). + unfold base. + apply Zpower2_le_lin; auto with zarith. + 2: generalize H1; unfold zdigits; rewrite spec_w_0W; + rewrite spec_zdigits; auto with zarith. + generalize H1; unfold zdigits; rewrite spec_w_0W; + rewrite spec_zdigits; auto; clear H1; intros H1. + assert (HH: [|low p|] <= Zpos w_digits). + rewrite HH0; auto with zarith. + repeat rewrite spec_w_add_mul_div with (1 := HH). + rewrite HH0. + 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. + 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)); + fold wB;fold wwB;zarith. + rewrite wwB_wBwB;rewrite Zpower_2; rewrite Zmult_mod_distr_r;zarith. + unfold ww_digits;rewrite Zpos_xO;zarith. apply Z_mod_lt;zarith. + split;zarith. apply Zdiv_lt_upper_bound;zarith. + rewrite <- Zpower_exp;zarith. + ring_simplify ([[p]] + (Zpos w_digits - [[p]]));fold wB;zarith. + assert (Hv: [[p]] > Zpos w_digits). + generalize H1; clear H1. + unfold zdigits; rewrite spec_w_0W; rewrite spec_zdigits; auto. + clear H1. + assert (HH0: [|low (ww_sub p zdigits)|] = [[p]] - Zpos w_digits). + rewrite spec_low. + rewrite spec_ww_sub. + unfold zdigits; rewrite spec_w_0W; rewrite spec_zdigits. + rewrite <- Zmod_div_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + exists wB; unfold base. + unfold ww_digits; rewrite (Zpos_xO w_digits). + rewrite <- Zpower_exp; auto with zarith. + apply f_equal with (f := fun x => 2 ^ x); auto with zarith. + assert (HH: [|low (ww_sub p zdigits)|] <= Zpos w_digits). + rewrite HH0; auto with zarith. + replace (Zpos w_digits + (Zpos w_digits - [[p]])) with + (Zpos w_digits - ([[p]] - Zpos w_digits)); zarith. + 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) + + (Zpos w_digits - ([[p]] - Zpos w_digits)))). + rewrite Zpower_exp;zarith. rewrite Zmult_assoc. + rewrite Z_div_plus_l;zarith. + rewrite shift_unshift_mod with (a:= [|yh|]) (p:= [[p]] - Zpos w_digits) + (n := Zpos w_digits);zarith. fold wB. + set (u := [[p]] - Zpos w_digits). + replace [[p]] with (u + Zpos w_digits);zarith. + rewrite Zpower_exp;zarith. rewrite Zmult_assoc. fold wB. + repeat rewrite Zplus_assoc. rewrite <- Zmult_plus_distr_l. + repeat rewrite <- Zplus_assoc. + unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits)); + fold wB;fold wwB;zarith. + unfold base;rewrite Zmod_shift_r with (a:= Zpos w_digits) + (b:= Zpos w_digits);fold wB;fold wwB;zarith. + rewrite wwB_wBwB; rewrite Zpower_2; rewrite Zmult_mod_distr_r;zarith. + rewrite Zmult_plus_distr_l. + replace ([|xh|] * wB * 2 ^ u) with + ([|xh|]*2^u*wB). 2:ring. + 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. + split;zarith. unfold u; apply Zdiv_lt_upper_bound;zarith. + rewrite <- Zpower_exp;zarith. + 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. + ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith. + unfold u;zarith. + unfold u;zarith. + set (u := [[p]] - Zpos w_digits). + ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith. + Qed. + + Lemma spec_ww_add_mul_div : forall x y p, + [[p]] <= Zpos (xO w_digits) -> + [[ ww_add_mul_div p x y ]] = + ([[x]] * (2^[[p]]) + + [[y]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB. + Proof. + clear spec_ww_zdigits. + intros x y p H. + destruct x as [ |xh xl]; + [assert (H1 := @spec_ww_add_mul_div_aux w_0 w_0) + |assert (H1 := @spec_ww_add_mul_div_aux xh xl)]; + (destruct y as [ |yh yl]; + [generalize (H1 w_0 w_0 p H) | generalize (H1 yh yl p H)]; + 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)); + 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. + assert (HH0: [|low p|] = [[p]]). + rewrite spec_low. + apply Zmod_small. + case (spec_to_w_Z p); intros HH1 HH2; split; auto. + apply Zlt_le_trans with (1 := H1). + unfold base; apply Zpower2_le_lin; auto with zarith. + rewrite HH0; auto with zarith. + replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial]. + intros Heq;rewrite <- Heq;clear Heq. + 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. + assert (HH: [|low (ww_sub p (w_0W w_zdigits)) |] = [[p]] - Zpos w_digits). + generalize H1; clear H1. + rewrite spec_low. + rewrite spec_ww_sub; w_rewrite; intros H1. + rewrite <- Zmod_div_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + unfold base; auto with zarith. + unfold base; auto with zarith. + exists wB; unfold base. + unfold ww_digits; rewrite (Zpos_xO w_digits). + rewrite <- Zpower_exp; auto with zarith. + apply f_equal with (f := fun x => 2 ^ x); auto with zarith. + case (spec_to_Z xh); auto with zarith. + Qed. + + End DoubleProof. + +End DoubleLift. + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v new file mode 100644 index 00000000..c7d83acc --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v @@ -0,0 +1,628 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleMul.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleMul. + Variable w : Type. + Variable w_0 : w. + Variable w_1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable w_succ : w -> w. + Variable w_add_c : w -> w -> carry w. + Variable w_add : w -> w -> w. + Variable w_sub: w -> w -> w. + Variable w_mul_c : w -> w -> zn2z w. + Variable w_mul : w -> w -> w. + Variable w_square_c : w -> zn2z w. + Variable ww_add_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_add : zn2z w -> zn2z w -> zn2z w. + Variable ww_add_carry : zn2z w -> zn2z w -> zn2z w. + Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_sub : zn2z w -> zn2z w -> zn2z w. + + (* ** Multiplication ** *) + + (* (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 + *) + + Definition double_mul_c (cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w) x y := + match x, y with + | W0, _ => W0 + | _, W0 => W0 + | WW xh xl, WW yh yl => + 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 + | W0 => WW (ww_add hh (w_W0 wc)) ll + | WW cch ccl => + match ww_add_c (w_W0 ccl) ll with + | C0 l => WW (ww_add hh (w_WW wc cch)) l + | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l + end + end + end. + + Definition ww_mul_c := + 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) + end). + + 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 + C0 m => + match w_compare xl xh with + Eq => (w_0, m) + | 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))) + | Gt => match ww_add_c m (w_mul_c (w_sub xh xl) (w_sub yl yh)) with + C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1) + end + end + | 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 + C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1) + end + | Gt => (w_0, ww_sub m (w_mul_c (w_sub xl xh) (w_sub yl yh))) + end + end + | C1 m => + match w_compare xl xh with + Eq => (w_1, m) + | 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 + | 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 => + 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 + C1 m1 => (w_2, m1) | C0 m1 => (w_1, m1) + end + | Gt => match ww_sub_c m (w_mul_c (w_sub xl xh) (w_sub yl yh)) with + C1 m1 => (w_0, m1) | C0 m1 => (w_1, m1) + end + end + end + end. + + Definition ww_karatsuba_c := double_mul_c kara_prod. + + Definition ww_mul x y := + match x, y with + | W0, _ => W0 + | _, 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. + + Definition ww_square_c x := + match x with + | W0 => W0 + | WW xh xl => + let hh := w_square_c xh in + let ll := w_square_c xl in + let xhxl := w_mul_c xh xl in + let (wc,cc) := + match ww_add_c xhxl xhxl with + | C0 cc => (w_0, cc) + | C1 cc => (w_1, cc) + end in + match cc with + | W0 => WW (ww_add hh (w_W0 wc)) ll + | WW cch ccl => + match ww_add_c (w_W0 ccl) ll with + | C0 l => WW (ww_add hh (w_WW wc cch)) l + | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l + end + end + end. + + Section DoubleMulAddn1. + 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 => + let mul_add := double_mul_add_n1 n1 in + fun x y r => + match x with + | W0 => (w_0,extend w_0W n1 r) + | WW xh xl => + let (rl,l) := mul_add xl y r in + let (rh,h) := mul_add xh y rl in + (rh, double_WW w_WW n1 h l) + end + end. + + End DoubleMulAddn1. + + Section DoubleMulAddmn1. + Variable wn: Type. + Variable extend_n : w -> wn. + 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) : + 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 => + let mul_add := double_mul_add_mn1 m1 in + fun x y r => + match x with + | W0 => (w_0,extend wn_0W m1 (extend_n r)) + | WW xh xl => + let (rl,l) := mul_add xl y r in + let (rh,h) := mul_add xh y rl in + (rh, double_WW wn_WW m1 h l) + end + end. + + End DoubleMulAddmn1. + + Definition w_mul_add x y r := + match w_mul_c x y with + | W0 => (w_0, r) + | WW h l => + match w_add_c l r with + | C0 lr => (h,lr) + | C1 lr => (w_succ h, lr) + end + end. + + + (*Section DoubleProof. *) + Variable w_digits : positive. + 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). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (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) + (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). + + Notation "[|| x ||]" := + (zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99). + + Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x) + (at level 0, x at level 99). + + Variable spec_more_than_1_digit: 1 < Zpos w_digits. + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_w_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB. + Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. + Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB. + Variable spec_w_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + + Variable spec_w_mul_c : forall x y, [[ w_mul_c x y ]] = [|x|] * [|y|]. + Variable spec_w_mul : forall x y, [|w_mul x y|] = ([|x|] * [|y|]) mod wB. + Variable spec_w_square_c : forall x, [[ w_square_c x]] = [|x|] * [|x|]. + + Variable spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]]. + Variable spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB. + Variable spec_ww_add_carry : + 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. + + Lemma spec_ww_to_Z_wBwB : forall x, 0 <= [[x]] < wB^2. + Proof. rewrite <- wwB_wBwB;apply spec_ww_to_Z. Qed. + + Hint Resolve spec_ww_to_Z spec_ww_to_Z_wBwB : mult. + Ltac zarith := auto with zarith mult. + + Lemma wBwB_lex: forall a b c d, + a * wB^2 + [[b]] <= c * wB^2 + [[d]] -> + a <= c. + 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]]. + 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|] -> + 0 <= [|wc|] <= 1. + Proof. + intros. + apply (sum_mul_carry [|xh|] [|xl|] [|yh|] [|yl|] [|wc|][[cc]] wB);zarith. + apply wB_pos. + Qed. + + 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|] -> + [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] -> + [||match cc with + | W0 => WW (ww_add hh (w_W0 wc)) ll + | WW cch ccl => + match ww_add_c (w_W0 ccl) ll with + | C0 l => WW (ww_add hh (w_WW wc cch)) l + | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l + end + end||] = ([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|]). + Proof. + intros;assert (U1 := wB_pos w_digits). + 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. + 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; + rewrite wwB_wBwB. ring. + rewrite <- (Zplus_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. + 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)). + assert (H5 := Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)). + omega. + generalize H3;clear H3;rewrite <- H1. + rewrite Zplus_assoc; rewrite Zpower_2; rewrite Zmult_assoc; + rewrite <- Zmult_plus_distr_l. + assert (((2 * wB - 4) + 2)*wB <= ([|wc|] * wB + [|cch|])*wB). + apply Zmult_le_compat;zarith. + 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; + simpl zn2z_to_Z; + try rewrite spec_ww_add;try rewrite spec_ww_add_carry;rewrite spec_w_WW; + rewrite Zmod_small;rewrite wwB_wBwB;intros. + rewrite H4;ring. rewrite H;apply mult_add_ineq2;zarith. + rewrite Zplus_assoc;rewrite Zmult_plus_distr_l. + rewrite Zmult_1_l;rewrite <- Zplus_assoc;rewrite H4;ring. + repeat rewrite <- Zplus_assoc;rewrite H;apply mult_add_ineq2;zarith. + Qed. + + Lemma spec_double_mul_c : forall cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w, + (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|]) -> + forall x y, [||double_mul_c cross x y||] = [[x]] * [[y]]. + Proof. + intros cross Hcross x y;destruct x as [ |xh xl];simpl;trivial. + destruct y as [ |yh yl];simpl. rewrite Zmult_0_r;trivial. + assert (H1:= spec_w_mul_c xh yh);assert (H2:= spec_w_mul_c xl yl). + generalize (Hcross _ _ _ _ _ _ H1 H2). + destruct (cross xh xl yh yl (w_mul_c xh yh) (w_mul_c xl yl)) as (wc,cc). + intros;apply spec_mul_aux;trivial. + rewrite <- wwB_wBwB;trivial. + Qed. + + 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. + generalize (spec_ww_add_c (w_mul_c xh yl) (w_mul_c xl yh)); + destruct (ww_add_c (w_mul_c xh yl) (w_mul_c xl yh)) as [c|c]; + unfold interp_carry;repeat rewrite spec_w_mul_c;intros H; + (rewrite spec_w_0 || rewrite spec_w_1);rewrite H;ring. + Qed. + + Lemma spec_w_2: [|w_2|] = 2. + unfold w_2; rewrite spec_w_add; rewrite spec_w_1; simpl. + apply Zmod_small; split; auto with zarith. + rewrite <- (Zpower_1_r 2); unfold base; apply Zpower_lt_monotone; auto with zarith. + Qed. + + Lemma kara_prod_aux : forall xh xl yh yl, + xh*yh + xl*yl - (xh-xl)*(yh-yl) = xh*yl + xl*yh. + Proof. intros;ring. Qed. + + Lemma spec_kara_prod : forall xh xl yh yl hh ll, + [[hh]] = [|xh|]*[|yh|] -> + [[ll]] = [|xl|]*[|yl|] -> + 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; + rewrite <- H; rewrite <- H0; unfold kara_prod. + 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)). + generalize (spec_w_compare xl xh); case (w_compare xl xh); intros Hxlh; + 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). + repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + split; auto with zarith. + simpl in Hz; rewrite Hz; rewrite H; rewrite H0. + rewrite kara_prod_aux; apply Zplus_le_0_compat; apply Zmult_le_0_compat; auto with zarith. + apply Zle_lt_trans with ([[z]]-0); auto with zarith. + unfold Zminus; apply Zplus_le_compat_l; apply Zle_left_rev; simpl; rewrite Zopp_involutive. + apply Zmult_le_0_compat; auto with zarith. + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2; + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh. + rewrite Hylh; rewrite spec_w_0; try (ring; fail). + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2; + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_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. + match goal with |- context[(?x - ?y) * (?z - ?t)] => + replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring] + end. + simpl in Hz; rewrite Hz; rewrite H; rewrite H0. + rewrite kara_prod_aux; apply Zplus_le_0_compat; apply Zmult_le_0_compat; auto with zarith. + apply Zle_lt_trans with ([[z]]-0); auto with zarith. + unfold Zminus; apply Zplus_le_compat_l; apply Zle_left_rev; simpl; rewrite Zopp_involutive. + apply Zmult_le_0_compat; auto with zarith. + (** there is a carry in hh + ll **) + rewrite Zmult_1_l. + generalize (spec_w_compare xl xh); case (w_compare xl xh); intros Hxlh; + try rewrite Hxlh; try rewrite spec_w_1; try (ring; fail). + generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh; + try rewrite Hylh; try rewrite spec_w_1; try (ring; fail). + match goal with |- context[ww_sub_c ?x ?y] => + generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l. + generalize Hz2; clear Hz2; unfold interp_carry. + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_2; unfold interp_carry in Hz2. + apply trans_equal with (wwB + (1 * wwB + [[z1]])). + ring. + rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh; + try rewrite Hylh; try rewrite spec_w_1; try (ring; fail). + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_2; unfold interp_carry in Hz2. + apply trans_equal with (wwB + (1 * wwB + [[z1]])). + ring. + rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + match goal with |- context[ww_sub_c ?x ?y] => + generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l. + match goal with |- context[(?x - ?y) * (?z - ?t)] => + replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring] + end. + generalize Hz2; clear Hz2; unfold interp_carry. + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + Qed. + + Lemma sub_carry : forall xh xl yh yl z, + 0 <= z -> + [|xh|]*[|yl|] + [|xl|]*[|yh|] = wwB + z -> + z < wwB. + Proof. + intros xh xl yh yl z Hle Heq. + destruct (Z_le_gt_dec wwB z);auto with zarith. + generalize (Zmult_lt_b _ _ _ (spec_to_Z xh) (spec_to_Z yl)). + 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). + rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat;zarith. + omega. + Qed. + + Ltac Spec_ww_to_Z x := + let H:= fresh "H" in + assert (H:= spec_ww_to_Z x). + + Ltac Zmult_lt_b x y := + let H := fresh "H" in + assert (H := Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)). + + Lemma spec_ww_karatsuba_c : forall x y, [||ww_karatsuba_c x y||]=[[x]]*[[y]]. + Proof. + intros x y; unfold ww_karatsuba_c;apply spec_double_mul_c. + intros; apply spec_kara_prod; auto. + Qed. + + Lemma spec_ww_mul : forall x y, [[ww_mul x y]] = [[x]]*[[y]] mod wwB. + Proof. + assert (U:= lt_0_wB w_digits). + assert (U1:= lt_0_wwB w_digits). + intros x y; case x; auto; intros xh xl. + case y; auto. + simpl; rewrite Zmult_0_r; rewrite Zmod_small; auto with zarith. + intros yh yl;simpl. + repeat (rewrite spec_ww_add || rewrite spec_w_W0 || rewrite spec_w_mul_c + || rewrite spec_w_add || rewrite spec_w_mul). + rewrite <- Zplus_mod; auto with zarith. + repeat (rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r). + rewrite <- Zmult_mod_distr_r; auto with zarith. + rewrite <- Zpower_2; rewrite <- wwB_wBwB; auto with zarith. + rewrite Zplus_mod; auto with zarith. + rewrite Zmod_mod; auto with zarith. + rewrite <- Zplus_mod; auto with zarith. + match goal with |- ?X mod _ = _ => + rewrite <- Z_mod_plus with (a := X) (b := [|xh|] * [|yh|]) + end; auto with zarith. + f_equal; auto; rewrite wwB_wBwB; ring. + Qed. + + Lemma spec_ww_square_c : forall x, [||ww_square_c x||] = [[x]]*[[x]]. + Proof. + destruct x as [ |xh xl];simpl;trivial. + case_eq match ww_add_c (w_mul_c xh xl) (w_mul_c xh xl) with + | C0 cc => (w_0, cc) + | C1 cc => (w_1, cc) + end;intros wc cc Heq. + apply (spec_mul_aux xh xl xh xl wc cc);trivial. + generalize Heq (spec_ww_add_c (w_mul_c xh xl) (w_mul_c xh xl));clear Heq. + rewrite spec_w_mul_c;destruct (ww_add_c (w_mul_c xh xl) (w_mul_c xh xl)); + unfold interp_carry;try rewrite Zmult_1_l;intros Heq Heq';inversion Heq; + rewrite (Zmult_comm [|xl|]);subst. + rewrite spec_w_0;rewrite Zmult_0_l;rewrite Zplus_0_l;trivial. + rewrite spec_w_1;rewrite Zmult_1_l;rewrite <- wwB_wBwB;trivial. + Qed. + + Section DoubleMulAddn1Proof. + + 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|]. + + 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 + [|h|]*double_wB w_digits n + [!n|l!] = [!n|x!]*[|y|]+[|r|]. + 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]; + 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). + assert(U:=IHn xh y rl);destruct(double_mul_add_n1 w_mul_add n xh y rl)as(rh,h). + rewrite <- double_wB_wwB. rewrite spec_double_WW;simpl;trivial. + rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc;rewrite <- H. + rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite U;ring. + Qed. + + End DoubleMulAddn1Proof. + + 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|]. + 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. + rewrite spec_w_0;trivial. + assert (U:=spec_w_add_c l r);destruct (w_add_c l r) as [lr|lr];unfold + interp_carry in U;try rewrite Zmult_1_l in H;simpl. + rewrite U;ring. rewrite spec_w_succ. rewrite Zmod_small. + rewrite <- Zplus_assoc;rewrite <- U;ring. + simpl in H;assert (H1:= Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)). + rewrite <- H in H1. + assert (H2:=spec_to_Z h);split;zarith. + case H1;clear H1;intro H1;clear H1. + replace (wB ^ 2 - 2 * wB) with ((wB - 2)*wB). 2:ring. + intros H0;assert (U1:= wB_pos w_digits). + assert (H1 := beta_lex _ _ _ _ _ H0 (spec_to_Z l));zarith. + Qed. + +(* End DoubleProof. *) + +End DoubleMul. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v new file mode 100644 index 00000000..043ff351 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v @@ -0,0 +1,1389 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleSqrt.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleSqrt. + Variable w : Type. + Variable w_is_even : w -> bool. + Variable w_compare : w -> w -> comparison. + Variable w_0 : w. + Variable w_1 : w. + Variable w_Bm1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_sub : w -> w -> w. + Variable w_sub_c : w -> w -> carry w. + Variable w_square_c : w -> zn2z w. + Variable w_div21 : w -> w -> w -> w * w. + Variable w_add_mul_div : w -> w -> w -> w. + Variable w_digits : positive. + Variable w_zdigits : w. + Variable ww_zdigits : zn2z w. + Variable w_add_c : w -> w -> carry w. + Variable w_sqrt2 : w -> w -> w * carry w. + Variable w_pred : w -> w. + Variable ww_pred_c : zn2z w -> carry (zn2z w). + Variable ww_pred : zn2z w -> zn2z w. + Variable ww_add_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_add : zn2z w -> zn2z w -> zn2z w. + Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_add_mul_div : zn2z w -> zn2z w -> zn2z w -> zn2z w. + Variable ww_head0 : zn2z w -> zn2z w. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable low : zn2z w -> w. + + Let wwBm1 := ww_Bm1 w_Bm1. + + Definition ww_is_even x := + match x with + | W0 => true + | WW xh xl => w_is_even xl + end. + + Let w_div21c x y z := + match w_compare x z with + | Eq => + match w_compare y z with + Eq => (C1 w_1, w_0) + | Gt => (C1 w_1, w_sub y z) + | Lt => (C1 w_0, y) + end + | Gt => + let x1 := w_sub x z in + let (q, r) := w_div21 x1 y z in + (C1 q, r) + | Lt => + let (q, r) := w_div21 x y z in + (C0 q, r) + end. + + Let w_div2s x y s := + match x with + C1 x1 => + let x2 := w_sub x1 s in + let (q, r) := w_div21c x2 y s in + match q with + C0 q1 => + if w_is_even q1 then + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), C0 r) + else + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), w_add_c r s) + | C1 q1 => + if w_is_even q1 then + (C1 (w_add_mul_div (w_pred w_zdigits) w_0 q1), C0 r) + else + (C1 (w_add_mul_div (w_pred w_zdigits) w_0 q1), w_add_c r s) + end + | C0 x1 => + let (q, r) := w_div21c x1 y s in + match q with + C0 q1 => + if w_is_even q1 then + (C0 (w_add_mul_div (w_pred w_zdigits) w_0 q1), C0 r) + else + (C0 (w_add_mul_div (w_pred w_zdigits) w_0 q1), w_add_c r s) + | C1 q1 => + if w_is_even q1 then + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), C0 r) + else + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), w_add_c r s) + end + end. + + Definition split x := + match x with + | W0 => (w_0,w_0) + | WW h l => (h,l) + end. + + Definition ww_sqrt2 x y := + let (x1, x2) := split x in + let (y1, y2) := split y in + let ( q, r) := w_sqrt2 x1 x2 in + let (q1, r1) := w_div2s r y1 q in + match q1 with + C0 q1 => + let q2 := w_square_c q1 in + let a := WW q q1 in + match r1 with + C1 r2 => + match ww_sub_c (WW r2 y2) q2 with + C0 r3 => (a, C1 r3) + | C1 r3 => (a, C0 r3) + end + | C0 r2 => + match ww_sub_c (WW r2 y2) q2 with + C0 r3 => (a, C0 r3) + | C1 r3 => + let a2 := ww_add_mul_div (w_0W w_1) a W0 in + match ww_pred_c a2 with + C0 a3 => + (ww_pred a, ww_add_c a3 r3) + | C1 a3 => + (ww_pred a, C0 (ww_add a3 r3)) + end + end + end + | C1 q1 => + let a1 := WW q w_Bm1 in + let a2 := ww_add_mul_div (w_0W w_1) a1 wwBm1 in + (a1, ww_add_c a2 y) + end. + + Definition ww_is_zero x := + match ww_compare W0 x with + Eq => true + | _ => false + end. + + Definition ww_head1 x := + let p := ww_head0 x in + if (ww_is_even p) then p else ww_pred p. + + Definition ww_sqrt x := + if (ww_is_zero x) then W0 + else + let p := ww_head1 x in + match ww_compare p W0 with + | Gt => + match ww_add_mul_div p x W0 with + W0 => W0 + | WW x1 x2 => + let (r, _) := w_sqrt2 x1 x2 in + 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. + + 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). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (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) + (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). + + Notation "[|| x ||]" := + (zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99). + + Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x) + (at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_w_zdigits : [|w_zdigits|] = Zpos w_digits. + Variable spec_more_than_1_digit: 1 < Zpos w_digits. + + Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos (xO w_digits). + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_w_is_even : forall x, + if w_is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1. + Variable spec_w_compare : forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_w_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + Variable spec_w_square_c : forall x, [[ w_square_c x]] = [|x|] * [|x|]. + Variable spec_w_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + 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|] / (Zpower 2 ((Zpos w_digits) - [|p|]))) mod wB. + Variable spec_ww_add_mul_div : forall x y p, + [[p]] <= Zpos (xO w_digits) -> + [[ ww_add_mul_div p x y ]] = + ([[x]] * (2^ [[p]]) + + [[y]] / (2^ (Zpos (xO w_digits) - [[p]]))) mod wwB. + Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. + Variable spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB. + Variable spec_w_sqrt2 : forall x y, + wB/ 4 <= [|x|] -> + let (s,r) := w_sqrt2 x y in + [[WW x y]] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]. + Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. + Variable spec_ww_pred_c : forall x, [-[ww_pred_c x]] = [[x]] - 1. + Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB. + Variable spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB. + Variable spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]]. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_head0 : forall x, 0 < [[x]] -> + wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB. + Variable spec_low: forall x, [|low x|] = [[x]] mod wB. + + 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. + + 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. +intros x; case x; simpl ww_is_even. + simpl. + rewrite Zmod_small; auto with zarith. + intros w1 w2; simpl. + unfold base. + rewrite Zplus_mod; auto with zarith. + rewrite (fun x y => (Zdivide_mod (x * y))); auto with zarith. + rewrite Zplus_0_l; rewrite Zmod_mod; auto with zarith. + apply spec_w_is_even; auto with zarith. + apply Zdivide_mult_r; apply Zpower_divide; auto with zarith. + red; simpl; auto. + Qed. + + + Theorem spec_w_div21c : forall a1 a2 b, + wB/2 <= [|b|] -> + let (q,r) := w_div21c a1 a2 b in + [|a1|] * wB + [|a2|] = [+|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]. + intros a1 a2 b Hb; unfold w_div21c. + assert (H: 0 < [|b|]); auto with zarith. + assert (U := wB_pos w_digits). + apply Zlt_le_trans with (2 := Hb); auto with zarith. + apply Zlt_le_trans with 1; auto with zarith. + apply Zdiv_le_lower_bound; auto with zarith. + repeat match goal with |- context[w_compare ?y ?z] => + generalize (spec_w_compare y z); + case (w_compare y z) + end. + intros H1 H2; split. + unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith. + rewrite H1; rewrite H2; ring. + autorewrite with w_rewrite; auto with zarith. + intros H1 H2; split. + unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith. + rewrite H2; ring. + destruct (spec_to_Z a2);auto with zarith. + intros H1 H2; split. + unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith. + rewrite H2; rewrite Zmod_small; auto with zarith. + ring. + destruct (spec_to_Z a2);auto with zarith. + rewrite spec_w_sub; auto with zarith. + destruct (spec_to_Z a2) as [H3 H4];auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + assert ([|a2|] < 2 * [|b|]); auto with zarith. + apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith. + rewrite wB_div_2; auto. + intros H1. + match goal with |- context[w_div21 ?y ?z ?t] => + generalize (@spec_w_div21 y z t Hb H1); + case (w_div21 y z t); simpl; autorewrite with w_rewrite; + auto + end. + intros H1. + assert (H2: [|w_sub a1 b|] < [|b|]). + rewrite spec_w_sub; auto with zarith. + rewrite Zmod_small; auto with zarith. + assert ([|a1|] < 2 * [|b|]); auto with zarith. + apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith. + rewrite wB_div_2; auto. + destruct (spec_to_Z a1);auto with zarith. + destruct (spec_to_Z a1);auto with zarith. + match goal with |- context[w_div21 ?y ?z ?t] => + generalize (@spec_w_div21 y z t Hb H2); + case (w_div21 y z t); autorewrite with w_rewrite; + auto + end. + intros w0 w1; replace [+|C1 w0|] with (wB + [|w0|]). + rewrite Zmod_small; auto with zarith. + intros (H3, H4); split; auto. + rewrite Zmult_plus_distr_l. + rewrite <- Zplus_assoc; rewrite <- H3; ring. + split; auto with zarith. + assert ([|a1|] < 2 * [|b|]); auto with zarith. + apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith. + rewrite wB_div_2; auto. + destruct (spec_to_Z a1);auto with zarith. + destruct (spec_to_Z a1);auto with zarith. + simpl; case wB; auto. + Qed. + + Theorem C0_id: forall p, [+|C0 p|] = [|p|]. + intros p; simpl; auto. + Qed. + + Theorem add_mult_div_2: forall w, + [|w_add_mul_div (w_pred w_zdigits) w_0 w|] = [|w|] / 2. + intros w1. + assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1). + rewrite spec_pred; rewrite spec_w_zdigits. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. + rewrite spec_w_add_mul_div; auto with zarith. + autorewrite with w_rewrite rm10. + match goal with |- context[?X - ?Y] => + replace (X - Y) with 1 + end. + rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith. + destruct (spec_to_Z w1) as [H1 H2];auto with zarith. + split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + rewrite Hp; ring. + Qed. + + Theorem add_mult_div_2_plus_1: forall w, + [|w_add_mul_div (w_pred w_zdigits) w_1 w|] = + [|w|] / 2 + 2 ^ Zpos (w_digits - 1). + intros w1. + assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1). + rewrite spec_pred; rewrite spec_w_zdigits. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. + autorewrite with w_rewrite rm10; auto with zarith. + match goal with |- context[?X - ?Y] => + replace (X - Y) with 1 + end; rewrite Hp; try ring. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith. + destruct (spec_to_Z w1) as [H1 H2];auto with zarith. + split; auto with zarith. + unfold base. + match goal with |- _ < _ ^ ?X => + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite <- (tmp X); clear tmp + end. + rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith. + assert (tmp: forall p, 1 + (p -1) - 1 = p - 1); auto with zarith; + rewrite tmp; clear tmp; auto with zarith. + match goal with |- ?X + ?Y < _ => + assert (Y < X); auto with zarith + end. + apply Zdiv_lt_upper_bound; auto with zarith. + pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp; + auto with zarith. + assert (tmp: forall p, (p - 1) + 1 = p); auto with zarith; + rewrite tmp; clear tmp; auto with zarith. + Qed. + + Theorem add_mult_mult_2: forall w, + [|w_add_mul_div w_1 w w_0|] = 2 * [|w|] mod wB. + intros w1. + autorewrite with w_rewrite rm10; auto with zarith. + rewrite Zpower_1_r; auto with zarith. + rewrite Zmult_comm; auto. + Qed. + + Theorem ww_add_mult_mult_2: forall w, + [[ww_add_mul_div (w_0W w_1) w W0]] = 2 * [[w]] mod wwB. + 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 Zpower_1_r; auto with zarith. + rewrite Zmult_comm; auto. + rewrite spec_w_0W; rewrite spec_w_1; auto with zarith. + red; simpl; intros; discriminate. + Qed. + + Theorem ww_add_mult_mult_2_plus_1: forall w, + [[ww_add_mul_div (w_0W w_1) w wwBm1]] = + (2 * [[w]] + 1) mod wwB. + intros w1. + rewrite spec_ww_add_mul_div; auto with zarith. + rewrite spec_w_0W; rewrite spec_w_1; auto with zarith. + rewrite Zpower_1_r; auto with zarith. + f_equal; auto. + rewrite Zmult_comm; f_equal; auto. + autorewrite with w_rewrite rm10. + unfold ww_digits, base. + apply sym_equal; apply Zdiv_unique with (r := 2 ^ (Zpos (ww_digits w_digits) - 1) -1); + auto with zarith. + unfold ww_digits; split; auto with zarith. + match goal with |- 0 <= ?X - 1 => + assert (0 < X); auto with zarith + end. + apply Zpower_gt_0; auto with zarith. + match goal with |- 0 <= ?X - 1 => + assert (0 < X); auto with zarith; red; reflexivity + end. + unfold ww_digits; autorewrite with rm10. + assert (tmp: forall p q r, p + (q - r) = p + q - r); auto with zarith; + rewrite tmp; clear tmp. + assert (tmp: forall p, p + p = 2 * p); auto with zarith; + rewrite tmp; clear tmp. + f_equal; auto. + pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp; + auto with zarith. + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite tmp; clear tmp; auto. + match goal with |- ?X - 1 >= 0 => + assert (0 < X); auto with zarith; red; reflexivity + end. + rewrite spec_w_0W; rewrite spec_w_1; auto with zarith. + red; simpl; intros; discriminate. + Qed. + + Theorem Zplus_mod_one: forall a1 b1, 0 < b1 -> (a1 + b1) mod b1 = a1 mod b1. + intros a1 b1 H; rewrite Zplus_mod; auto with zarith. + rewrite Z_mod_same; try rewrite Zplus_0_r; auto with zarith. + apply Zmod_mod; auto. + Qed. + + Lemma C1_plus_wB: forall x, [+|C1 x|] = wB + [|x|]. + unfold interp_carry; auto with zarith. + Qed. + + Theorem spec_w_div2s : forall a1 a2 b, + wB/2 <= [|b|] -> [+|a1|] <= 2 * [|b|] -> + let (q,r) := w_div2s a1 a2 b in + [+|a1|] * wB + [|a2|] = [+|q|] * (2 * [|b|]) + [+|r|] /\ 0 <= [+|r|] < 2 * [|b|]. + intros a1 a2 b H. + assert (HH: 0 < [|b|]); auto with zarith. + assert (U := wB_pos w_digits). + apply Zlt_le_trans with (2 := H); auto with zarith. + apply Zlt_le_trans with 1; auto with zarith. + apply Zdiv_le_lower_bound; auto with zarith. + unfold w_div2s; case a1; intros w0 H0. + match goal with |- context[w_div21c ?y ?z ?t] => + generalize (@spec_w_div21c y z t H); + case (w_div21c y z t); autorewrite with w_rewrite; + auto + end. + intros c w1; case c. + simpl interp_carry; intros w2 (Hw1, Hw2). + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat rewrite C0_id. + rewrite add_mult_div_2. + intros H1; split; auto with zarith. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + repeat rewrite C0_id. + rewrite add_mult_div_2. + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + intros w2; rewrite C1_plus_wB. + intros (Hw1, Hw2). + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat rewrite C0_id. + intros H1; split; auto with zarith. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1. + repeat rewrite C0_id. + 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; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + repeat rewrite C0_id. + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + rewrite add_mult_div_2_plus_1. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + 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; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + repeat rewrite C1_plus_wB in H0. + rewrite C1_plus_wB. + match goal with |- context[w_div21c ?y ?z ?t] => + generalize (@spec_w_div21c y z t H); + case (w_div21c y z t); autorewrite with w_rewrite; + auto + end. + intros c w1; case c. + intros w2 (Hw1, Hw2); rewrite C0_id in Hw1. + rewrite <- Zplus_mod_one in Hw1; auto with zarith. + rewrite Zmod_small in Hw1; auto with zarith. + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat rewrite C0_id. + intros H1; split; auto with zarith. + rewrite add_mult_div_2_plus_1. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + 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; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + repeat rewrite C0_id. + rewrite add_mult_div_2_plus_1. + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + 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; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + split; auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z w0);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + intros w2; rewrite C1_plus_wB. + rewrite <- Zplus_mod_one; auto with zarith. + rewrite Zmod_small; auto with zarith. + intros (Hw1, Hw2). + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat (rewrite C0_id || rewrite C1_plus_wB). + intros H1; split; auto with zarith. + rewrite add_mult_div_2. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + repeat (rewrite C0_id || rewrite C1_plus_wB). + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + rewrite add_mult_div_2. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + split; auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z w0);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + Qed. + + Theorem wB_div_4: 4 * (wB / 4) = wB. + Proof. + unfold base. + assert (2 ^ Zpos w_digits = + 4 * (2 ^ (Zpos w_digits - 2))). + change 4 with (2 ^ 2). + rewrite <- Zpower_exp; auto with zarith. + f_equal; auto with zarith. + rewrite H. + rewrite (fun x => (Zmult_comm 4 (2 ^x))). + rewrite Z_div_mult; auto with zarith. + Qed. + + Theorem Zsquare_mult: forall p, p ^ 2 = p * p. + intros p; change 2 with (1 + 1); rewrite Zpower_exp; + try rewrite Zpower_1_r; auto with zarith. + Qed. + + Theorem Zsquare_pos: forall p, 0 <= p ^ 2. + intros p; case (Zle_or_lt 0 p); intros H1. + rewrite Zsquare_mult; apply Zmult_le_0_compat; auto with zarith. + 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; + auto with zarith. + Qed. + + Theorem mult_wwB: forall x y, [|x|] * [|y|] < wwB. + Proof. + intros x y; rewrite wwB_wBwB; rewrite Zpower_2. + generalize (spec_to_Z x); intros U. + generalize (spec_to_Z y); intros U1. + apply Zle_lt_trans with ((wB -1 ) * (wB - 1)); auto with zarith. + apply Zmult_le_compat; auto with zarith. + repeat (rewrite Zmult_minus_distr_r || rewrite Zmult_minus_distr_l); + auto with zarith. + Qed. + Hint Resolve mult_wwB. + + Lemma spec_ww_sqrt2 : forall x y, + wwB/ 4 <= [[x]] -> + let (s,r) := ww_sqrt2 x y in + [||WW x y||] = [[s]] ^ 2 + [+[r]] /\ + [+[r]] <= 2 * [[s]]. + 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. + intros w0 w1 Hw0 w2 w3 Hw1. + assert (U: wB/4 <= [|w2|]). + case (Zle_or_lt (wB / 4) [|w2|]); auto; intros H1. + contradict H; apply Zlt_not_le. + rewrite wwB_wBwB; rewrite Zpower_2. + pattern wB at 1; rewrite <- wB_div_4; rewrite <- Zmult_assoc; + rewrite Zmult_comm. + rewrite Z_div_mult; auto with zarith. + rewrite <- Hw1. + match goal with |- _ < ?X => + pattern X; rewrite <- Zplus_0_r; apply beta_lex_inv; + auto with zarith + end. + destruct (spec_to_Z w3);auto with zarith. + generalize (@spec_w_sqrt2 w2 w3 U); case (w_sqrt2 w2 w3). + intros w4 c (H1, H2). + assert (U1: wB/2 <= [|w4|]). + case (Zle_or_lt (wB/2) [|w4|]); auto with zarith. + intros U1. + assert (U2 : [|w4|] <= wB/2 -1); auto with zarith. + assert (U3 : [|w4|] ^ 2 <= wB/4 * wB - wB + 1); auto with zarith. + match goal with |- ?X ^ 2 <= ?Y => + rewrite Zsquare_mult; + replace Y with ((wB/2 - 1) * (wB/2 -1)) + end. + apply Zmult_le_compat; auto with zarith. + destruct (spec_to_Z w4);auto with zarith. + destruct (spec_to_Z w4);auto with zarith. + pattern wB at 4 5; rewrite <- wB_div_2. + rewrite Zmult_assoc. + replace ((wB / 4) * 2) with (wB / 2). + ring. + pattern wB at 1; rewrite <- wB_div_4. + change 4 with (2 * 2). + rewrite <- Zmult_assoc; rewrite (Zmult_comm 2). + rewrite Z_div_mult; try ring; auto with zarith. + assert (U4 : [+|c|] <= wB -2); auto with zarith. + apply Zle_trans with (1 := H2). + 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. + match type of H1 with ?X = _ => + assert (U5: X < wB / 4 * wB) + end. + rewrite H1; auto with zarith. + contradict U; apply Zlt_not_le. + apply Zmult_lt_reg_r with wB; auto with zarith. + destruct (spec_to_Z w4);auto with zarith. + apply Zle_lt_trans with (2 := U5). + unfold ww_to_Z, zn2z_to_Z. + destruct (spec_to_Z w3);auto with zarith. + generalize (@spec_w_div2s c w0 w4 U1 H2). + case (w_div2s c w0 w4). + intros c0; case c0; intros w5; + repeat (rewrite C0_id || rewrite C1_plus_wB). + 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] => + generalize (spec_ww_sub_c y z); case (ww_sub_c y z) + end. + intros z; change [-[C0 z]] with ([[z]]). + change [+[C0 z]] with ([[z]]). + intros H5; rewrite spec_w_square_c in H5; + auto. + split. + unfold zn2z_to_Z; rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite H5. + unfold ww_to_Z, zn2z_to_Z. + repeat rewrite Zsquare_mult; ring. + rewrite H5. + unfold ww_to_Z, zn2z_to_Z. + match goal with |- ?X - ?Y * ?Y <= _ => + assert (V := Zsquare_pos Y); + rewrite Zsquare_mult in V; + apply Zle_trans with X; auto with zarith; + clear V + end. + match goal with |- ?X * wB + ?Y <= 2 * (?Z * wB + ?T) => + apply Zle_trans with ((2 * Z - 1) * wB + wB); auto with zarith + end. + destruct (spec_to_Z w1);auto with zarith. + match goal with |- ?X <= _ => + replace X with (2 * [|w4|] * wB); auto with zarith + end. + rewrite Zmult_plus_distr_r; rewrite Zmult_assoc. + destruct (spec_to_Z w5); auto with zarith. + ring. + intros z; replace [-[C1 z]] with (- wwB + [[z]]). + 2: simpl; case wwB; auto with zarith. + intros H5; rewrite spec_w_square_c in H5; + auto. + match goal with |- context [ww_pred_c ?y] => + generalize (spec_ww_pred_c y); case (ww_pred_c y) + end. + intros z1; change [-[C0 z1]] with ([[z1]]). + rewrite ww_add_mult_mult_2. + rewrite spec_ww_add_c. + rewrite spec_ww_pred. + rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]); + auto with zarith. + intros Hz1; rewrite Zmod_small; auto with zarith. + match type of H5 with -?X + ?Y = ?Z => + assert (V: Y = Z + X); + try (rewrite <- H5; ring) + end. + split. + unfold zn2z_to_Z; rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite V. + rewrite Hz1. + unfold ww_to_Z; simpl zn2z_to_Z. + repeat rewrite Zsquare_mult; ring. + rewrite Hz1. + destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith. + assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)). + assert (0 < [[WW w4 w5]]); auto with zarith. + apply Zlt_le_trans with (wB/ 2 * wB + 0); auto with zarith. + autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith. + apply Zmult_lt_reg_r with 2; auto with zarith. + autorewrite with rm10. + rewrite Zmult_comm; rewrite wB_div_2; auto with zarith. + case (spec_to_Z w5);auto with zarith. + case (spec_to_Z w5);auto with zarith. + simpl. + assert (V2 := spec_to_Z w5);auto with zarith. + assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith. + split; auto with zarith. + assert (wwB <= 2 * [[WW w4 w5]]); auto with zarith. + apply Zle_trans with (2 * ([|w4|] * wB)). + rewrite wwB_wBwB; rewrite Zpower_2. + rewrite Zmult_assoc; apply Zmult_le_compat_r; auto with zarith. + rewrite <- wB_div_2; auto with zarith. + assert (V2 := spec_to_Z w5);auto with zarith. + simpl ww_to_Z; assert (V2 := spec_to_Z w5);auto with zarith. + assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith. + intros z1; change [-[C1 z1]] with (-wwB + [[z1]]). + match goal with |- context[([+[C0 ?z]])] => + change [+[C0 z]] with ([[z]]) + end. + rewrite spec_ww_add; auto with zarith. + rewrite spec_ww_pred; auto with zarith. + rewrite ww_add_mult_mult_2. + rename V1 into VV1. + assert (VV2: 0 < [[WW w4 w5]]); auto with zarith. + apply Zlt_le_trans with (wB/ 2 * wB + 0); auto with zarith. + autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith. + apply Zmult_lt_reg_r with 2; auto with zarith. + autorewrite with rm10. + rewrite Zmult_comm; rewrite wB_div_2; auto with zarith. + assert (VV3 := spec_to_Z w5);auto with zarith. + assert (VV3 := spec_to_Z w5);auto with zarith. + simpl. + assert (VV3 := spec_to_Z w5);auto with zarith. + assert (VV3: wwB <= 2 * [[WW w4 w5]]); auto with zarith. + apply Zle_trans with (2 * ([|w4|] * wB)). + rewrite wwB_wBwB; rewrite Zpower_2. + rewrite Zmult_assoc; apply Zmult_le_compat_r; auto with zarith. + rewrite <- wB_div_2; auto with zarith. + case (spec_to_Z w5);auto with zarith. + simpl ww_to_Z; assert (V4 := spec_to_Z w5);auto with zarith. + rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]); + auto with zarith. + intros Hz1; rewrite Zmod_small; auto with zarith. + match type of H5 with -?X + ?Y = ?Z => + assert (V: Y = Z + X); + try (rewrite <- H5; ring) + end. + match type of Hz1 with -?X + ?Y = -?X + ?Z - 1 => + assert (V1: Y = Z - 1); + [replace (Z - 1) with (X + (-X + Z -1)); + [rewrite <- Hz1 | idtac]; ring + | idtac] + end. + rewrite <- Zmod_unique with (q := 1) (r := -wwB + [[z1]] + [[z]]); + auto with zarith. + unfold zn2z_to_Z; rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + split. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite V. + rewrite Hz1. + unfold ww_to_Z; simpl zn2z_to_Z. + repeat rewrite Zsquare_mult; ring. + assert (V2 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith. + assert (V2 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith. + assert (V3 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z1);auto with zarith. + split; auto with zarith. + rewrite (Zplus_comm (-wwB)); rewrite <- Zplus_assoc. + rewrite H5. + match goal with |- 0 <= ?X + (?Y - ?Z) => + apply Zle_trans with (X - Z); auto with zarith + end. + 2: generalize (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w6 w1)); unfold ww_to_Z; auto with zarith. + rewrite V1. + match goal with |- 0 <= ?X - 1 - ?Y => + assert (Y < X); auto with zarith + end. + apply Zlt_le_trans with wwB; auto with zarith. + intros (H3, H4). + match goal with |- context [ww_sub_c ?y ?z] => + generalize (spec_ww_sub_c y z); case (ww_sub_c y z) + end. + intros z; change [-[C0 z]] with ([[z]]). + match goal with |- context[([+[C1 ?z]])] => + replace [+[C1 z]] with (wwB + [[z]]) + end. + 2: simpl; case wwB; auto. + intros H5; rewrite spec_w_square_c in H5; + auto. + split. + change ([||WW x y||]) with ([[x]] * wwB + [[y]]). + rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite H5. + unfold ww_to_Z; simpl zn2z_to_Z. + rewrite wwB_wBwB. + repeat rewrite Zsquare_mult; ring. + simpl ww_to_Z. + rewrite H5. + simpl ww_to_Z. + rewrite wwB_wBwB; rewrite Zpower_2. + match goal with |- ?X * ?Y + (?Z * ?Y + ?T - ?U) <= _ => + apply Zle_trans with (X * Y + (Z * Y + T - 0)); + auto with zarith + end. + assert (V := Zsquare_pos [|w5|]); + rewrite Zsquare_mult in V; auto with zarith. + autorewrite with rm10. + match goal with |- _ <= 2 * (?U * ?V + ?W) => + apply Zle_trans with (2 * U * V + 0); + auto with zarith + end. + match goal with |- ?X * ?Y + (?Z * ?Y + ?T) <= _ => + replace (X * Y + (Z * Y + T)) with ((X + Z) * Y + T); + try ring + end. + apply Zlt_le_weak; apply beta_lex_inv; auto with zarith. + destruct (spec_to_Z w1);auto with zarith. + destruct (spec_to_Z w5);auto with zarith. + rewrite Zmult_plus_distr_r; auto with zarith. + rewrite Zmult_assoc; auto with zarith. + intros z; replace [-[C1 z]] with (- wwB + [[z]]). + 2: simpl; case wwB; auto with zarith. + intros H5; rewrite spec_w_square_c in H5; + auto. + match goal with |- context[([+[C0 ?z]])] => + change [+[C0 z]] with ([[z]]) + end. + match type of H5 with -?X + ?Y = ?Z => + assert (V: Y = Z + X); + try (rewrite <- H5; ring) + end. + change ([||WW x y||]) with ([[x]] * wwB + [[y]]). + simpl ww_to_Z. + rewrite <- Hw1. + simpl ww_to_Z in H1; rewrite H1. + rewrite <- Hw0. + split. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite V. + simpl ww_to_Z. + rewrite wwB_wBwB. + repeat rewrite Zsquare_mult; ring. + rewrite V. + simpl ww_to_Z. + rewrite wwB_wBwB; rewrite Zpower_2. + match goal with |- (?Z * ?Y + ?T - ?U) + ?X * ?Y <= _ => + apply Zle_trans with ((Z * Y + T - 0) + X * Y); + auto with zarith + end. + assert (V1 := Zsquare_pos [|w5|]); + rewrite Zsquare_mult in V1; auto with zarith. + autorewrite with rm10. + match goal with |- _ <= 2 * (?U * ?V + ?W) => + apply Zle_trans with (2 * U * V + 0); + auto with zarith + end. + match goal with |- (?Z * ?Y + ?T) + ?X * ?Y <= _ => + replace ((Z * Y + T) + X * Y) with ((X + Z) * Y + T); + try ring + end. + apply Zlt_le_weak; apply beta_lex_inv; auto with zarith. + destruct (spec_to_Z w1);auto with zarith. + destruct (spec_to_Z w5);auto with zarith. + rewrite Zmult_plus_distr_r; auto with zarith. + rewrite Zmult_assoc; auto with zarith. + case Zle_lt_or_eq with (1 := H2); clear H2; intros H2. + intros c1 (H3, H4). + match type of H3 with ?X = ?Y => + absurd (X < Y) + end. + apply Zle_not_lt; rewrite <- H3; auto with zarith. + rewrite Zmult_plus_distr_l. + apply Zlt_le_trans with ((2 * [|w4|]) * wB + 0); + auto with zarith. + apply beta_lex_inv; auto with zarith. + destruct (spec_to_Z w0);auto with zarith. + assert (V1 := spec_to_Z w5);auto with zarith. + rewrite (Zmult_comm wB); auto with zarith. + assert (0 <= [|w5|] * (2 * [|w4|])); auto with zarith. + intros c1 (H3, H4); rewrite H2 in H3. + match type of H3 with ?X + ?Y = (?Z + ?T) * ?U + ?V => + assert (VV: (Y = (T * U) + V)); + [replace Y with ((X + Y) - X); + [rewrite H3; ring | ring] | idtac] + end. + assert (V1 := spec_to_Z w0);auto with zarith. + assert (V2 := spec_to_Z w5);auto with zarith. + case (Zle_lt_or_eq 0 [|w5|]); auto with zarith; intros V3. + match type of VV with ?X = ?Y => + absurd (X < Y) + end. + apply Zle_not_lt; rewrite <- VV; auto with zarith. + apply Zlt_le_trans with wB; auto with zarith. + match goal with |- _ <= ?X + _ => + apply Zle_trans with X; auto with zarith + end. + match goal with |- _ <= _ * ?X => + apply Zle_trans with (1 * X); auto with zarith + end. + autorewrite with rm10. + rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith. + rewrite <- V3 in VV; generalize VV; autorewrite with rm10; + clear VV; intros VV. + rewrite spec_ww_add_c; auto with zarith. + rewrite ww_add_mult_mult_2_plus_1. + match goal with |- context[?X mod wwB] => + rewrite <- Zmod_unique with (q := 1) (r := -wwB + X) + end; auto with zarith. + simpl ww_to_Z. + rewrite spec_w_Bm1; auto with zarith. + split. + change ([||WW x y||]) with ([[x]] * wwB + [[y]]). + rewrite <- Hw1. + simpl ww_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H2. + rewrite wwB_wBwB. + repeat rewrite Zsquare_mult; ring. + assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z y);auto with zarith. + assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z y);auto with zarith. + simpl ww_to_Z; unfold ww_to_Z. + rewrite spec_w_Bm1; auto with zarith. + split. + rewrite wwB_wBwB; rewrite Zpower_2. + match goal with |- _ <= -?X + (2 * (?Z * ?T + ?U) + ?V) => + assert (X <= 2 * Z * T); auto with zarith + end. + apply Zmult_le_compat_r; auto with zarith. + rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith. + rewrite Zmult_plus_distr_r; auto with zarith. + rewrite Zmult_assoc; auto with zarith. + match goal with |- _ + ?X < _ => + replace X with ((2 * (([|w4|]) + 1) * wB) - 1); try ring + end. + assert (2 * ([|w4|] + 1) * wB <= 2 * wwB); auto with zarith. + rewrite <- Zmult_assoc; apply Zmult_le_compat_l; auto with zarith. + rewrite wwB_wBwB; rewrite Zpower_2. + apply Zmult_le_compat_r; auto with zarith. + case (spec_to_Z w4);auto with zarith. + Qed. + + Lemma spec_ww_is_zero: forall x, + if ww_is_zero x then [[x]] = 0 else 0 < [[x]]. + intro x; unfold ww_is_zero. + generalize (spec_ww_compare W0 x); case (ww_compare W0 x); + auto with zarith. + simpl ww_to_Z. + assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z x);auto with zarith. + Qed. + + Lemma wwB_4_2: 2 * (wwB / 4) = wwB/ 2. + pattern wwB at 1; rewrite wwB_wBwB; rewrite Zpower_2. + rewrite <- wB_div_2. + match goal with |- context[(2 * ?X) * (2 * ?Z)] => + replace ((2 * X) * (2 * Z)) with ((X * Z) * 4); try ring + end. + rewrite Z_div_mult; auto with zarith. + rewrite Zmult_assoc; rewrite wB_div_2. + rewrite wwB_div_2; ring. + Qed. + + + Lemma spec_ww_head1 + : 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). + intros x; unfold ww_head1. + generalize (spec_ww_is_even (ww_head0 x)); case_eq (ww_is_even (ww_head0 x)). + intros HH H1; rewrite HH; split; auto. + intros H2. + generalize (spec_ww_head0 x H2); case (ww_head0 x); autorewrite with rm10. + intros (H3, H4); split; auto with zarith. + apply Zle_trans with (2 := H3). + apply Zdiv_le_compat_l; auto with zarith. + intros xh xl (H3, H4); split; auto with zarith. + apply Zle_trans with (2 := H3). + apply Zdiv_le_compat_l; auto with zarith. + intros H1. + case (spec_to_w_Z (ww_head0 x)); intros Hv1 Hv2. + assert (Hp0: 0 < [[ww_head0 x]]). + generalize (spec_ww_is_even (ww_head0 x)); rewrite H1. + generalize Hv1; case [[ww_head0 x]]. + rewrite Zmod_small; auto with zarith. + intros; assert (0 < Zpos p); auto with zarith. + red; simpl; auto. + intros p H2; case H2; auto. + assert (Hp: [[ww_pred (ww_head0 x)]] = [[ww_head0 x]] - 1). + rewrite spec_ww_pred. + rewrite Zmod_small; auto with zarith. + intros H2; split. + generalize (spec_ww_is_even (ww_pred (ww_head0 x))); + case ww_is_even; auto. + rewrite Hp. + rewrite Zminus_mod; auto with zarith. + rewrite H2; repeat rewrite Zmod_small; auto with zarith. + 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. + pattern 2 at 1; rewrite <- Zpower_1_r. + rewrite <- Zpower_exp; auto with zarith. + ring_simplify (1 + (u - 1)); auto with zarith. + split; auto with zarith. + apply Zmult_le_reg_r with 2; auto with zarith. + repeat rewrite (fun x => Zmult_comm x 2). + rewrite wwB_4_2. + rewrite Zmult_assoc; rewrite Hu; auto with zarith. + apply Zle_lt_trans with (2 * 2 ^ ([[ww_head0 x]] - 1) * [[x]]); auto with zarith; + rewrite Hu; auto with zarith. + apply Zmult_le_compat_r; auto with zarith. + apply Zpower_le_monotone; auto with zarith. + Qed. + + Theorem wwB_4_wB_4: wwB / 4 = wB / 4 * wB. + apply sym_equal; apply Zdiv_unique with 0; + auto with zarith. + rewrite Zmult_assoc; rewrite wB_div_4; auto with zarith. + rewrite wwB_wBwB; ring. + Qed. + + Lemma spec_ww_sqrt : forall x, + [[ww_sqrt x]] ^ 2 <= [[x]] < ([[ww_sqrt x]] + 1) ^ 2. + assert (U := wB_pos w_digits). + intro x; unfold ww_sqrt. + generalize (spec_ww_is_zero x); case (ww_is_zero x). + simpl ww_to_Z; simpl Zpower; unfold Zpower_pos; simpl; + 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. + generalize (H4 H2); clear H4; rewrite H5; clear H5; autorewrite with rm10. + intros (H4, H5). + assert (V: wB/4 <= [|w0|]). + apply beta_lex with 0 [|w1|] wB; auto with zarith; autorewrite with rm10. + 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. + case c; unfold interp_carry; autorewrite with rm10. + intros w3 (H6, H7); rewrite H6. + assert (V1 := spec_to_Z w3);auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith. + match goal with |- ?X < ?Z => + replace Z with (X + 1); auto with zarith + end. + repeat rewrite Zsquare_mult; ring. + intros w3 (H6, H7); rewrite H6. + assert (V1 := spec_to_Z w3);auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith. + match goal with |- ?X < ?Z => + replace Z with (X + 1); auto with zarith + end. + repeat rewrite Zsquare_mult; ring. + intros HH; case (spec_to_w_Z (ww_head1 x)); auto with zarith. + intros Hv1. + case (spec_ww_head1 x); intros Hp1 Hp2. + generalize (Hp2 H1); clear Hp2; intros Hp2. + assert (Hv2: [[ww_head1 x]] <= Zpos (xO w_digits)). + case (Zle_or_lt (Zpos (xO w_digits)) [[ww_head1 x]]); auto with zarith; intros HH1. + case Hp2; intros _ HH2; contradict HH2. + 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; + 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); + case ww_add_mul_div. + simpl ww_to_Z; autorewrite with w_rewrite rm10. + rewrite Zmod_small; auto with zarith. + intros H2; case (Zmult_integral _ _ (sym_equal H2)); clear H2; intros H2. + rewrite H2; unfold Zpower, Zpower_pos; simpl; auto with zarith. + match type of H2 with ?X = ?Y => + absurd (Y < X); try (rewrite H2; auto with zarith; fail) + end. + apply Zpower_gt_0; auto with zarith. + split; auto with zarith. + case Hp2; intros _ tmp; apply Zle_lt_trans with (2 := tmp); + clear tmp. + rewrite Zmult_comm; apply Zmult_le_compat_r; auto with zarith. + assert (Hv0: [[ww_head1 x]] = 2 * ([[ww_head1 x]]/2)). + pattern [[ww_head1 x]] at 1; rewrite (Z_div_mod_eq [[ww_head1 x]] 2); + auto with zarith. + generalize (spec_ww_is_even (ww_head1 x)); rewrite Hp1; + intros tmp; rewrite tmp; rewrite Zplus_0_r; auto. + intros w0 w1; autorewrite with w_rewrite rm10. + rewrite Zmod_small; auto with zarith. + 2: rewrite Zmult_comm; auto with zarith. + intros H2. + assert (V: wB/4 <= [|w0|]). + apply beta_lex with 0 [|w1|] wB; auto with zarith; autorewrite with rm10. + simpl ww_to_Z in H2; rewrite H2. + rewrite <- wwB_4_wB_4; auto with zarith. + rewrite Zmult_comm; auto with zarith. + assert (V1 := spec_to_Z w1);auto with zarith. + generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith. + case (w_sqrt2 w0 w1); intros w2 c. + case (spec_to_Z w2); intros HH1 HH2. + 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. + 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. + 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. + assert (Hv5: [[(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))]] + = [[ww_head1 x]]/2). + rewrite spec_ww_add_mul_div. + simpl ww_to_Z; autorewrite with rm10. + rewrite Hv3. + ring_simplify (Zpos (xO w_digits) - (Zpos (xO w_digits) - 1)). + rewrite Zpower_1_r. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (1 := Hv4); auto with zarith. + unfold base; apply Zpower_le_monotone; auto with zarith. + split; unfold ww_digits; try rewrite Zpos_xO; auto with zarith. + rewrite Hv3; auto with zarith. + assert (Hv6: [|low(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))|] + = [[ww_head1 x]]/2). + rewrite spec_low. + rewrite Hv5; rewrite Zmod_small; auto with zarith. + rewrite spec_w_add_mul_div; auto with zarith. + rewrite spec_w_sub; auto with zarith. + rewrite spec_w_0. + simpl ww_to_Z; autorewrite with rm10. + rewrite Hv6; rewrite spec_w_zdigits. + rewrite (fun x y => Zmod_small (x - y)). + ring_simplify (Zpos w_digits - (Zpos w_digits - [[ww_head1 x]] / 2)). + rewrite Zmod_small. + simpl ww_to_Z in H2; rewrite H2; auto with zarith. + intros (H4, H5); split. + apply Zmult_le_reg_r with (2 ^ [[ww_head1 x]]); auto with zarith. + rewrite H4. + apply Zle_trans with ([|w2|] ^ 2); auto with zarith. + rewrite Zmult_comm. + pattern [[ww_head1 x]] at 1; + rewrite Hv0; auto with zarith. + rewrite (Zmult_comm 2); rewrite Zpower_mult; + auto with zarith. + assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2); + try (intros; repeat rewrite Zsquare_mult; ring); + rewrite tmp; clear tmp. + apply Zpower_le_monotone3; auto with zarith. + split; auto with zarith. + 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; + 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. + apply Zle_lt_trans with ([|w2|] ^ 2 + 2 * [|w2|]); auto with zarith. + apply Zlt_le_trans with (([|w2|] + 1) ^ 2); auto with zarith. + match goal with |- ?X < ?Y => + replace Y with (X + 1); auto with zarith + end. + repeat rewrite (Zsquare_mult); ring. + rewrite Zmult_comm. + pattern [[ww_head1 x]] at 1; rewrite Hv0. + rewrite (Zmult_comm 2); rewrite Zpower_mult; + auto with zarith. + assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2); + try (intros; repeat rewrite Zsquare_mult; ring); + rewrite tmp; clear tmp. + apply Zpower_le_monotone3; auto with zarith. + split; auto with zarith. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]]/2))); + auto with zarith. + rewrite <- Zplus_assoc; rewrite Zmult_plus_distr_r. + autorewrite with rm10; apply Zplus_le_compat_l; auto with zarith. + case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]]/2))); auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with ([|w2|]); auto with zarith. + apply Zdiv_le_upper_bound; auto with zarith. + pattern [|w2|] at 1; replace [|w2|] with ([|w2|] * 2 ^0); + auto with zarith. + apply Zmult_le_compat_l; auto with zarith. + apply Zpower_le_monotone; auto with zarith. + rewrite Zpower_0_r; autorewrite with rm10; auto. + split; auto with zarith. + rewrite Hv0 in Hv2; rewrite (Zpos_xO w_digits) in Hv2; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + rewrite spec_w_sub; auto with zarith. + rewrite Hv6; rewrite spec_w_zdigits; auto with zarith. + assert (Hv7: 0 < [[ww_head1 x]]/2); auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + assert ([[ww_head1 x]]/2 <= Zpos w_digits); auto with zarith. + apply Zmult_le_reg_r with 2; auto with zarith. + repeat rewrite (fun x => Zmult_comm x 2). + rewrite <- Hv0; rewrite <- Zpos_xO; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + Qed. + +End DoubleSqrt. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v new file mode 100644 index 00000000..269d62bb --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v @@ -0,0 +1,357 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleSub.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleSub. + Variable w : Type. + Variable w_0 : w. + Variable w_Bm1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable ww_Bm1 : zn2z w. + Variable w_opp_c : w -> carry w. + Variable w_opp_carry : w -> w. + Variable w_pred_c : w -> carry w. + Variable w_sub_c : w -> w -> carry w. + Variable w_sub_carry_c : w -> w -> carry w. + Variable w_opp : w -> w. + Variable w_pred : w -> w. + Variable w_sub : w -> w -> w. + Variable w_sub_carry : w -> w -> w. + + (* ** Opposites ** *) + Definition ww_opp_c x := + match x with + | W0 => C0 W0 + | WW xh xl => + match w_opp_c xl with + | C0 _ => + match w_opp_c xh with + | C0 h => C0 W0 + | C1 h => C1 (WW h w_0) + end + | C1 l => C1 (WW (w_opp_carry xh) l) + end + end. + + Definition ww_opp x := + match x with + | W0 => W0 + | WW xh xl => + match w_opp_c xl with + | C0 _ => WW (w_opp xh) w_0 + | C1 l => WW (w_opp_carry xh) l + end + end. + + Definition ww_opp_carry x := + match x with + | W0 => ww_Bm1 + | WW xh xl => w_WW (w_opp_carry xh) (w_opp_carry xl) + end. + + Definition ww_pred_c x := + match x with + | W0 => C1 ww_Bm1 + | WW xh xl => + match w_pred_c xl with + | C0 l => C0 (w_WW xh l) + | 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 + | WW xh xl => + match w_pred_c xl with + | C0 l => w_WW xh l + | 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 => + match w_sub_c xh yh with + | C0 h => C0 (w_WW h l) + | C1 h => C1 (WW h l) + end + | C1 l => + match w_sub_carry_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (WW h l) + end + end + end. + + Definition ww_sub x y := + match y, x with + | W0, _ => x + | WW yh yl, W0 => ww_opp (WW yh yl) + | WW yh yl, WW xh xl => + match w_sub_c xl yl with + | C0 l => w_WW (w_sub xh yh) l + | C1 l => WW (w_sub_carry xh yh) l + end + end. + + Definition ww_sub_carry_c x y := + match y, x with + | W0, W0 => C1 ww_Bm1 + | W0, WW xh xl => ww_pred_c (WW xh xl) + | 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 => + match w_sub_c xh yh with + | C0 h => C0 (w_WW h l) + | C1 h => C1 (WW h l) + end + | C1 l => + match w_sub_carry_c xh yh with + | C0 h => C0 (w_WW h l) + | C1 h => C1 (w_WW h l) + end + end + end. + + Definition ww_sub_carry x y := + match y, x with + | W0, W0 => ww_Bm1 + | W0, WW xh xl => ww_pred (WW xh xl) + | WW yh yl, W0 => ww_opp_carry (WW yh yl) + | WW yh yl, WW xh xl => + match w_sub_carry_c xl yl with + | C0 l => w_WW (w_sub xh yh) l + | C1 l => w_WW (w_sub_carry xh yh) l + end + end. + + (*Section DoubleProof.*) + Variable w_digits : positive. + 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). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (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) + (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). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + + 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. + + Variable spec_pred_c : forall x, [-|w_pred_c x|] = [|x|] - 1. + 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 : + forall x y, [|w_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB. + + + 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) + as [l|l];intros H;unfold interp_carry in H;rewrite <- H; + 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) + 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. + rewrite H2;reflexivity. + simpl ww_to_Z;rewrite wwB_wBwB;rewrite spec_w_0;ring. + unfold interp_carry;simpl ww_to_Z;rewrite wwB_wBwB;rewrite spec_opp_carry; + ring. + Qed. + + Lemma spec_ww_opp : forall x, [[ww_opp x]] = (-[[x]]) mod wwB. + Proof. + destruct x as [ |xh xl];simpl. reflexivity. + rewrite Zopp_plus_distr;rewrite Zopp_mult_distr_l. + 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). + assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega. + rewrite H0;rewrite Zplus_0_r; rewrite Zpower_2; + rewrite Zmult_mod_distr_r;try apply lt_0_wB. + rewrite spec_opp;trivial. + apply Zmod_unique with (q:= -1). + exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW (w_opp_carry xh) l)). + rewrite spec_opp_carry;rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_opp_carry : forall x, [[ww_opp_carry x]] = wwB - [[x]] - 1. + Proof. + destruct x as [ |xh xl];simpl. rewrite spec_ww_Bm1;ring. + rewrite spec_w_WW;simpl;repeat rewrite spec_opp_carry;rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_pred_c : forall x, [-[ww_pred_c x]] = [[x]] - 1. + Proof. + destruct x as [ |xh xl];unfold ww_pred_c. + unfold interp_carry;rewrite spec_ww_Bm1;simpl ww_to_Z;ring. + simpl ww_to_Z;replace (([|xh|]*wB+[|xl|])-1) with ([|xh|]*wB+([|xl|]-1)). + 2:ring. generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l]; + intros H;unfold interp_carry in H;rewrite <- H. simpl;apply spec_w_WW. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + assert ([|l|] = wB - 1). + assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega. + rewrite H0;change ([|xh|] + -1) with ([|xh|] - 1). + generalize (spec_pred_c xh);destruct (w_pred_c xh) as [h|h]; + 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. + rewrite H2;unfold interp_carry;rewrite spec_ww_Bm1;rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. + 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|])) + 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. + 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; + try rewrite wwB_wBwB;ring. + Qed. + + 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. + 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) + 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. + 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. + 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. + + Lemma spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB. + Proof. + 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. + exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh 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. + rewrite (mod_wwB w_digits w_to_Z);trivial. + rewrite spec_pred;rewrite spec_w_Bm1;rewrite <- H0;trivial. + Qed. + + Lemma spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. + Proof. + 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|])) + 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. + rewrite spec_w_WW;rewrite (mod_wwB w_digits w_to_Z spec_to_Z). + rewrite spec_sub;trivial. + simpl ww_to_Z;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial. + Qed. + + Lemma spec_ww_sub_carry : + forall x y, [[ww_sub_carry x y]] = ([[x]] - [[y]] - 1) mod wwB. + Proof. + destruct y as [ |yh yl];simpl. + ring_simplify ([[x]] - 0);exact (spec_ww_pred x). + destruct x as [ |xh xl];simpl. + apply Zmod_unique with (-1). + 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) + 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. + rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub;trivial. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial. + Qed. + +(* End DoubleProof. *) + +End DoubleSub. + + + + + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v new file mode 100644 index 00000000..28d40094 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v @@ -0,0 +1,71 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(*i $Id: DoubleType.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Open Local Scope Z_scope. + +Definition base digits := Zpower 2 (Zpos digits). + +Section Carry. + + Variable A : Type. + + Inductive carry := + | C0 : A -> carry + | C1 : A -> carry. + + Definition interp_carry (sign:Z)(B:Z)(interp:A -> Z) c := + match c with + | C0 x => interp x + | C1 x => sign*B + interp x + end. + +End Carry. + +Section Zn2Z. + + Variable znz : Type. + + (** 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. + *) + + Inductive zn2z := + | W0 : zn2z + | WW : znz -> znz -> zn2z. + + Definition zn2z_to_Z (wB:Z) (w_to_Z:znz->Z) (x:zn2z) := + match x with + | W0 => 0 + | WW xh xl => w_to_Z xh * wB + w_to_Z xl + end. + +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). +*) + +Fixpoint word (w:Type) (n:nat) : Type := + match n with + | O => w + | S n => zn2z (word w n) + end. + |