diff options
Diffstat (limited to 'theories/Numbers/Cyclic')
-rw-r--r-- | theories/Numbers/Cyclic/Abstract/CyclicAxioms.v | 159 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/Abstract/NZCyclic.v | 173 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v | 74 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v | 94 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v | 168 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v | 324 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v | 144 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v | 66 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v | 114 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v | 94 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v | 76 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v | 18 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/Int31/Cyclic31.v | 464 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/Int31/Int31.v | 141 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/Int31/Ring31.v | 103 | ||||
-rw-r--r-- | theories/Numbers/Cyclic/ZModulo/ZModulo.v | 227 |
16 files changed, 1319 insertions, 1120 deletions
diff --git a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v index 528d78c3..51df2fa3 100644 --- a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v +++ b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v @@ -8,12 +8,12 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(* $Id: CyclicAxioms.v 11012 2008-05-28 16:34:43Z letouzey $ *) +(* $Id$ *) (** * Signature and specification of a bounded integer structure *) -(** This file specifies how to represent [Z/nZ] when [n=2^d], - [d] being the number of digits of these bounded integers. *) +(** This file specifies how to represent [Z/nZ] when [n=2^d], + [d] being the number of digits of these bounded integers. *) Set Implicit Arguments. @@ -22,7 +22,7 @@ Require Import Znumtheory. Require Import BigNumPrelude. Require Import DoubleType. -Open Local Scope Z_scope. +Local Open Scope Z_scope. (** First, a description via an operator record and a spec record. *) @@ -33,7 +33,7 @@ Section Z_nZ_Op. Record znz_op := mk_znz_op { (* Conversion functions with Z *) - znz_digits : positive; + znz_digits : positive; znz_zdigits: znz; znz_to_Z : znz -> Z; znz_of_pos : positive -> N * znz; (* Euclidean division by [2^digits] *) @@ -78,12 +78,12 @@ Section Z_nZ_Op. znz_div : znz -> znz -> znz * znz; znz_mod_gt : znz -> znz -> znz; (* specialized version of [znz_mod] *) - znz_mod : znz -> znz -> znz; + znz_mod : znz -> znz -> znz; znz_gcd_gt : znz -> znz -> znz; (* specialized version of [znz_gcd] *) - znz_gcd : znz -> znz -> znz; + znz_gcd : znz -> znz -> znz; (* [znz_add_mul_div p i j] is a combination of the [(digits-p)] - low bits of [i] above the [p] high bits of [j]: + low bits of [i] above the [p] high bits of [j]: [znz_add_mul_div p i j = i*2^p+j/2^(digits-p)] *) znz_add_mul_div : znz -> znz -> znz -> znz; (* [znz_pos_mod p i] is [i mod 2^p] *) @@ -135,7 +135,7 @@ Section Z_nZ_Spec. Let w_mul_c := w_op.(znz_mul_c). Let w_mul := w_op.(znz_mul). Let w_square_c := w_op.(znz_square_c). - + Let w_div21 := w_op.(znz_div21). Let w_div_gt := w_op.(znz_div_gt). Let w_div := w_op.(znz_div). @@ -229,25 +229,25 @@ Section Z_nZ_Spec. spec_div : forall a b, 0 < [|b|] -> let (q,r) := w_div a b in [|a|] = [|q|] * [|b|] + [|r|] /\ - 0 <= [|r|] < [|b|]; - + 0 <= [|r|] < [|b|]; + spec_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> [|w_mod_gt a b|] = [|a|] mod [|b|]; spec_mod : forall a b, 0 < [|b|] -> [|w_mod a b|] = [|a|] mod [|b|]; - + spec_gcd_gt : forall a b, [|a|] > [|b|] -> Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|]; spec_gcd : forall a b, Zis_gcd [|a|] [|b|] [|w_gcd a b|]; - + (* shift operations *) spec_head00: forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits; spec_head0 : forall x, 0 < [|x|] -> - wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB; + wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB; spec_tail00: forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits; - spec_tail0 : forall x, 0 < [|x|] -> - exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|w_tail0 x|]) ; + spec_tail0 : forall x, 0 < [|x|] -> + exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|w_tail0 x|]) ; spec_add_mul_div : forall x y p, [|p|] <= Zpos w_digits -> [| w_add_mul_div p x y |] = @@ -272,23 +272,23 @@ End Z_nZ_Spec. (** Generic construction of double words *) Section WW. - + Variable w : Type. Variable w_op : znz_op w. Variable op_spec : znz_spec w_op. - + Let wB := base w_op.(znz_digits). Let w_to_Z := w_op.(znz_to_Z). Let w_eq0 := w_op.(znz_eq0). Let w_0 := w_op.(znz_0). - Definition znz_W0 h := + Definition znz_W0 h := if w_eq0 h then W0 else WW h w_0. - Definition znz_0W l := + Definition znz_0W l := if w_eq0 l then W0 else WW w_0 l. - Definition znz_WW h l := + Definition znz_WW h l := if w_eq0 h then znz_0W l else WW h l. Lemma spec_W0 : forall h, @@ -300,7 +300,7 @@ Section WW. unfold w_0; rewrite op_spec.(spec_0); auto with zarith. Qed. - Lemma spec_0W : forall l, + Lemma spec_0W : forall l, zn2z_to_Z wB w_to_Z (znz_0W l) = w_to_Z l. Proof. unfold zn2z_to_Z, znz_0W, w_to_Z; simpl; intros. @@ -309,7 +309,7 @@ Section WW. unfold w_0; rewrite op_spec.(spec_0); auto with zarith. Qed. - Lemma spec_WW : forall h l, + Lemma spec_WW : forall h l, zn2z_to_Z wB w_to_Z (znz_WW h l) = (w_to_Z h)*wB + w_to_Z l. Proof. unfold znz_WW, w_to_Z; simpl; intros. @@ -324,7 +324,7 @@ End WW. (** Injecting [Z] numbers into a cyclic structure *) Section znz_of_pos. - + Variable w : Type. Variable w_op : znz_op w. Variable op_spec : znz_spec w_op. @@ -349,7 +349,7 @@ Section znz_of_pos. apply Zle_trans with X; auto with zarith end. match goal with |- ?X <= _ => - pattern X at 1; rewrite <- (Zmult_1_l); + pattern X at 1; rewrite <- (Zmult_1_l); apply Zmult_le_compat_r; auto with zarith end. case p1; simpl; intros; red; simpl; intros; discriminate. @@ -373,3 +373,112 @@ Module Type CyclicType. Parameter w_op : znz_op w. Parameter w_spec : znz_spec w_op. End CyclicType. + + +(** A Cyclic structure can be seen as a ring *) + +Module CyclicRing (Import Cyclic : CyclicType). + +Definition t := w. + +Local Notation "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99). + +Definition eq (n m : t) := [| n |] = [| m |]. +Definition zero : t := w_op.(znz_0). +Definition one := w_op.(znz_1). +Definition add := w_op.(znz_add). +Definition sub := w_op.(znz_sub). +Definition mul := w_op.(znz_mul). +Definition opp := w_op.(znz_opp). + +Local Infix "==" := eq (at level 70). +Local Notation "0" := zero. +Local Notation "1" := one. +Local Infix "+" := add. +Local Infix "-" := sub. +Local Infix "*" := mul. +Local Notation "!!" := (base (znz_digits w_op)). + +Hint Rewrite + w_spec.(spec_0) w_spec.(spec_1) + w_spec.(spec_add) w_spec.(spec_mul) w_spec.(spec_opp) w_spec.(spec_sub) + : cyclic. + +Ltac zify := + unfold eq, zero, one, add, sub, mul, opp in *; autorewrite with cyclic. + +Lemma add_0_l : forall x, 0 + x == x. +Proof. +intros. zify. rewrite Zplus_0_l. +apply Zmod_small. apply w_spec.(spec_to_Z). +Qed. + +Lemma add_comm : forall x y, x + y == y + x. +Proof. +intros. zify. now rewrite Zplus_comm. +Qed. + +Lemma add_assoc : forall x y z, x + (y + z) == x + y + z. +Proof. +intros. zify. now rewrite Zplus_mod_idemp_r, Zplus_mod_idemp_l, Zplus_assoc. +Qed. + +Lemma mul_1_l : forall x, 1 * x == x. +Proof. +intros. zify. rewrite Zmult_1_l. +apply Zmod_small. apply w_spec.(spec_to_Z). +Qed. + +Lemma mul_comm : forall x y, x * y == y * x. +Proof. +intros. zify. now rewrite Zmult_comm. +Qed. + +Lemma mul_assoc : forall x y z, x * (y * z) == x * y * z. +Proof. +intros. zify. now rewrite Zmult_mod_idemp_r, Zmult_mod_idemp_l, Zmult_assoc. +Qed. + +Lemma mul_add_distr_r : forall x y z, (x+y)*z == x*z + y*z. +Proof. +intros. zify. now rewrite <- Zplus_mod, Zmult_mod_idemp_l, Zmult_plus_distr_l. +Qed. + +Lemma add_opp_r : forall x y, x + opp y == x-y. +Proof. +intros. zify. rewrite <- Zminus_mod_idemp_r. unfold Zminus. +destruct (Z_eq_dec ([|y|] mod !!) 0) as [EQ|NEQ]. +rewrite Z_mod_zero_opp_full, EQ, 2 Zplus_0_r; auto. +rewrite Z_mod_nz_opp_full by auto. +rewrite <- Zplus_mod_idemp_r, <- Zminus_mod_idemp_l. +rewrite Z_mod_same_full. simpl. now rewrite Zplus_mod_idemp_r. +Qed. + +Lemma add_opp_diag_r : forall x, x + opp x == 0. +Proof. +intros. red. rewrite add_opp_r. zify. now rewrite Zminus_diag, Zmod_0_l. +Qed. + +Lemma CyclicRing : ring_theory 0 1 add mul sub opp eq. +Proof. +constructor. +exact add_0_l. exact add_comm. exact add_assoc. +exact mul_1_l. exact mul_comm. exact mul_assoc. +exact mul_add_distr_r. +symmetry. apply add_opp_r. +exact add_opp_diag_r. +Qed. + +Definition eqb x y := + match w_op.(znz_compare) x y with Eq => true | _ => false end. + +Lemma eqb_eq : forall x y, eqb x y = true <-> x == y. +Proof. + intros. unfold eqb, eq. generalize (w_spec.(spec_compare) x y). + destruct (w_op.(znz_compare) x y); intuition; try discriminate. +Qed. + +Lemma eqb_correct : forall x y, eqb x y = true -> x==y. +Proof. now apply eqb_eq. Qed. + +End CyclicRing.
\ No newline at end of file diff --git a/theories/Numbers/Cyclic/Abstract/NZCyclic.v b/theories/Numbers/Cyclic/Abstract/NZCyclic.v index fb3f0cef..517e48ad 100644 --- a/theories/Numbers/Cyclic/Abstract/NZCyclic.v +++ b/theories/Numbers/Cyclic/Abstract/NZCyclic.v @@ -8,7 +8,7 @@ (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) -(*i $Id: NZCyclic.v 11238 2008-07-19 09:34:03Z herbelin $ i*) +(*i $Id$ i*) Require Export NZAxioms. Require Import BigNumPrelude. @@ -17,89 +17,79 @@ Require Import CyclicAxioms. (** * From [CyclicType] to [NZAxiomsSig] *) -(** A [Z/nZ] representation given by a module type [CyclicType] - implements [NZAxiomsSig], e.g. the common properties between - N and Z with no ordering. Notice that the [n] in [Z/nZ] is +(** A [Z/nZ] representation given by a module type [CyclicType] + implements [NZAxiomsSig], e.g. the common properties between + N and Z with no ordering. Notice that the [n] in [Z/nZ] is a power of 2. *) Module NZCyclicAxiomsMod (Import Cyclic : CyclicType) <: NZAxiomsSig. -Open Local Scope Z_scope. +Local Open Scope Z_scope. -Definition NZ := w. +Definition t := w. -Definition NZ_to_Z : NZ -> Z := znz_to_Z w_op. -Definition Z_to_NZ : Z -> NZ := znz_of_Z w_op. -Notation Local wB := (base w_op.(znz_digits)). +Definition NZ_to_Z : t -> Z := znz_to_Z w_op. +Definition Z_to_NZ : Z -> t := znz_of_Z w_op. +Local Notation wB := (base w_op.(znz_digits)). -Notation Local "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99). +Local Notation "[| x |]" := (w_op.(znz_to_Z) x) (at level 0, x at level 99). -Definition NZeq (n m : NZ) := [| n |] = [| m |]. -Definition NZ0 := w_op.(znz_0). -Definition NZsucc := w_op.(znz_succ). -Definition NZpred := w_op.(znz_pred). -Definition NZadd := w_op.(znz_add). -Definition NZsub := w_op.(znz_sub). -Definition NZmul := w_op.(znz_mul). +Definition eq (n m : t) := [| n |] = [| m |]. +Definition zero := w_op.(znz_0). +Definition succ := w_op.(znz_succ). +Definition pred := w_op.(znz_pred). +Definition add := w_op.(znz_add). +Definition sub := w_op.(znz_sub). +Definition mul := w_op.(znz_mul). -Theorem NZeq_equiv : equiv NZ NZeq. -Proof. -unfold equiv, reflexive, symmetric, transitive, NZeq; repeat split; intros; auto. -now transitivity [| y |]. -Qed. +Local Infix "==" := eq (at level 70). +Local Notation "0" := zero. +Local Notation S := succ. +Local Notation P := pred. +Local Infix "+" := add. +Local Infix "-" := sub. +Local Infix "*" := mul. -Add Relation NZ NZeq - reflexivity proved by (proj1 NZeq_equiv) - symmetry proved by (proj2 (proj2 NZeq_equiv)) - transitivity proved by (proj1 (proj2 NZeq_equiv)) -as NZeq_rel. +Hint Rewrite w_spec.(spec_0) w_spec.(spec_succ) w_spec.(spec_pred) + w_spec.(spec_add) w_spec.(spec_mul) w_spec.(spec_sub) : w. +Ltac wsimpl := + unfold eq, zero, succ, pred, add, sub, mul; autorewrite with w. +Ltac wcongruence := repeat red; intros; wsimpl; congruence. -Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd. +Instance eq_equiv : Equivalence eq. Proof. -unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_succ). now rewrite H. +unfold eq. firstorder. Qed. -Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd. +Instance succ_wd : Proper (eq ==> eq) succ. Proof. -unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_pred). now rewrite H. +wcongruence. Qed. -Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd. +Instance pred_wd : Proper (eq ==> eq) pred. Proof. -unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_add). -now rewrite H1, H2. +wcongruence. Qed. -Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd. +Instance add_wd : Proper (eq ==> eq ==> eq) add. Proof. -unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_sub). -now rewrite H1, H2. +wcongruence. Qed. -Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd. +Instance sub_wd : Proper (eq ==> eq ==> eq) sub. Proof. -unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_mul). -now rewrite H1, H2. +wcongruence. Qed. -Delimit Scope IntScope with Int. -Bind Scope IntScope with NZ. -Open Local Scope IntScope. -Notation "x == y" := (NZeq x y) (at level 70) : IntScope. -Notation "x ~= y" := (~ NZeq x y) (at level 70) : IntScope. -Notation "0" := NZ0 : IntScope. -Notation S x := (NZsucc x). -Notation P x := (NZpred x). -(*Notation "1" := (S 0) : IntScope.*) -Notation "x + y" := (NZadd x y) : IntScope. -Notation "x - y" := (NZsub x y) : IntScope. -Notation "x * y" := (NZmul x y) : IntScope. +Instance mul_wd : Proper (eq ==> eq ==> eq) mul. +Proof. +wcongruence. +Qed. Theorem gt_wB_1 : 1 < wB. Proof. -unfold base. -apply Zpower_gt_1; unfold Zlt; auto with zarith. +unfold base. apply Zpower_gt_1; unfold Zlt; auto with zarith. Qed. Theorem gt_wB_0 : 0 < wB. @@ -107,7 +97,7 @@ Proof. pose proof gt_wB_1; auto with zarith. Qed. -Lemma NZsucc_mod_wB : forall n : Z, (n + 1) mod wB = ((n mod wB) + 1) mod wB. +Lemma succ_mod_wB : forall n : Z, (n + 1) mod wB = ((n mod wB) + 1) mod wB. Proof. intro n. pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zplus_mod. @@ -115,7 +105,7 @@ reflexivity. now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]]. Qed. -Lemma NZpred_mod_wB : forall n : Z, (n - 1) mod wB = ((n mod wB) - 1) mod wB. +Lemma pred_mod_wB : forall n : Z, (n - 1) mod wB = ((n mod wB) - 1) mod wB. Proof. intro n. pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zminus_mod. @@ -123,34 +113,32 @@ reflexivity. now rewrite Zmod_small; [ | split; [auto with zarith | apply gt_wB_1]]. Qed. -Lemma NZ_to_Z_mod : forall n : NZ, [| n |] mod wB = [| n |]. +Lemma NZ_to_Z_mod : forall n, [| n |] mod wB = [| n |]. Proof. intro n; rewrite Zmod_small. reflexivity. apply w_spec.(spec_to_Z). Qed. -Theorem NZpred_succ : forall n : NZ, P (S n) == n. +Theorem pred_succ : forall n, P (S n) == n. Proof. -intro n; unfold NZsucc, NZpred, NZeq. rewrite w_spec.(spec_pred), w_spec.(spec_succ). -rewrite <- NZpred_mod_wB. +intro n. wsimpl. +rewrite <- pred_mod_wB. replace ([| n |] + 1 - 1)%Z with [| n |] by auto with zarith. apply NZ_to_Z_mod. Qed. -Lemma Z_to_NZ_0 : Z_to_NZ 0%Z == 0%Int. +Lemma Z_to_NZ_0 : Z_to_NZ 0%Z == 0. Proof. -unfold NZeq, NZ_to_Z, Z_to_NZ. rewrite znz_of_Z_correct. -symmetry; apply w_spec.(spec_0). +unfold NZ_to_Z, Z_to_NZ. wsimpl. +rewrite znz_of_Z_correct; auto. exact w_spec. split; [auto with zarith |apply gt_wB_0]. Qed. Section Induction. -Variable A : NZ -> Prop. -Hypothesis A_wd : predicate_wd NZeq A. +Variable A : t -> Prop. +Hypothesis A_wd : Proper (eq ==> iff) A. Hypothesis A0 : A 0. -Hypothesis AS : forall n : NZ, A n <-> A (S n). (* Below, we use only -> direction *) - -Add Morphism A with signature NZeq ==> iff as A_morph. -Proof. apply A_wd. Qed. +Hypothesis AS : forall n, A n <-> A (S n). + (* Below, we use only -> direction *) Let B (n : Z) := A (Z_to_NZ n). @@ -163,8 +151,8 @@ Lemma BS : forall n : Z, 0 <= n -> n < wB - 1 -> B n -> B (n + 1). Proof. intros n H1 H2 H3. unfold B in *. apply -> AS in H3. -setoid_replace (Z_to_NZ (n + 1)) with (S (Z_to_NZ n)) using relation NZeq. assumption. -unfold NZeq. rewrite w_spec.(spec_succ). +setoid_replace (Z_to_NZ (n + 1)) with (S (Z_to_NZ n)). assumption. +wsimpl. unfold NZ_to_Z, Z_to_NZ. do 2 (rewrite znz_of_Z_correct; [ | exact w_spec | auto with zarith]). symmetry; apply Zmod_small; auto with zarith. @@ -177,11 +165,11 @@ apply Zbounded_induction with wB. apply B0. apply BS. assumption. assumption. Qed. -Theorem NZinduction : forall n : NZ, A n. +Theorem bi_induction : forall n, A n. Proof. -intro n. setoid_replace n with (Z_to_NZ (NZ_to_Z n)) using relation NZeq. +intro n. setoid_replace n with (Z_to_NZ (NZ_to_Z n)). apply B_holds. apply w_spec.(spec_to_Z). -unfold NZeq, NZ_to_Z, Z_to_NZ; rewrite znz_of_Z_correct. +unfold eq, NZ_to_Z, Z_to_NZ; rewrite znz_of_Z_correct. reflexivity. exact w_spec. apply w_spec.(spec_to_Z). @@ -189,47 +177,40 @@ Qed. End Induction. -Theorem NZadd_0_l : forall n : NZ, 0 + n == n. +Theorem add_0_l : forall n, 0 + n == n. Proof. -intro n; unfold NZadd, NZ0, NZeq. rewrite w_spec.(spec_add). rewrite w_spec.(spec_0). +intro n. wsimpl. rewrite Zplus_0_l. rewrite Zmod_small; [reflexivity | apply w_spec.(spec_to_Z)]. Qed. -Theorem NZadd_succ_l : forall n m : NZ, (S n) + m == S (n + m). +Theorem add_succ_l : forall n m, (S n) + m == S (n + m). Proof. -intros n m; unfold NZadd, NZsucc, NZeq. rewrite w_spec.(spec_add). -do 2 rewrite w_spec.(spec_succ). rewrite w_spec.(spec_add). -rewrite NZsucc_mod_wB. repeat rewrite Zplus_mod_idemp_l; try apply gt_wB_0. +intros n m. wsimpl. +rewrite succ_mod_wB. repeat rewrite Zplus_mod_idemp_l; try apply gt_wB_0. rewrite <- (Zplus_assoc ([| n |] mod wB) 1 [| m |]). rewrite Zplus_mod_idemp_l. rewrite (Zplus_comm 1 [| m |]); now rewrite Zplus_assoc. Qed. -Theorem NZsub_0_r : forall n : NZ, n - 0 == n. +Theorem sub_0_r : forall n, n - 0 == n. Proof. -intro n; unfold NZsub, NZ0, NZeq. rewrite w_spec.(spec_sub). -rewrite w_spec.(spec_0). rewrite Zminus_0_r. apply NZ_to_Z_mod. +intro n. wsimpl. rewrite Zminus_0_r. apply NZ_to_Z_mod. Qed. -Theorem NZsub_succ_r : forall n m : NZ, n - (S m) == P (n - m). +Theorem sub_succ_r : forall n m, n - (S m) == P (n - m). Proof. -intros n m; unfold NZsub, NZsucc, NZpred, NZeq. -rewrite w_spec.(spec_pred). do 2 rewrite w_spec.(spec_sub). -rewrite w_spec.(spec_succ). rewrite Zminus_mod_idemp_r. -rewrite Zminus_mod_idemp_l. -now replace ([|n|] - ([|m|] + 1))%Z with ([|n|] - [|m|] - 1)%Z by auto with zarith. +intros n m. wsimpl. rewrite Zminus_mod_idemp_r, Zminus_mod_idemp_l. +now replace ([|n|] - ([|m|] + 1))%Z with ([|n|] - [|m|] - 1)%Z + by auto with zarith. Qed. -Theorem NZmul_0_l : forall n : NZ, 0 * n == 0. +Theorem mul_0_l : forall n, 0 * n == 0. Proof. -intro n; unfold NZmul, NZ0, NZ, NZeq. rewrite w_spec.(spec_mul). -rewrite w_spec.(spec_0). now rewrite Zmult_0_l. +intro n. wsimpl. now rewrite Zmult_0_l. Qed. -Theorem NZmul_succ_l : forall n m : NZ, (S n) * m == n * m + m. +Theorem mul_succ_l : forall n m, (S n) * m == n * m + m. Proof. -intros n m; unfold NZmul, NZsucc, NZadd, NZeq. rewrite w_spec.(spec_mul). -rewrite w_spec.(spec_add), w_spec.(spec_mul), w_spec.(spec_succ). -rewrite Zplus_mod_idemp_l, Zmult_mod_idemp_l. +intros n m. wsimpl. rewrite Zplus_mod_idemp_l, Zmult_mod_idemp_l. now rewrite Zmult_plus_distr_l, Zmult_1_l. Qed. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v index 61d8d0fb..aa798e1c 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleAdd.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -17,7 +17,7 @@ Require Import BigNumPrelude. Require Import DoubleType. Require Import DoubleBase. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section DoubleAdd. Variable w : Type. @@ -36,10 +36,10 @@ Section DoubleAdd. Definition ww_succ_c x := match x with | W0 => C0 ww_1 - | WW xh xl => + | WW xh xl => match w_succ_c xl with | C0 l => C0 (WW xh l) - | C1 l => + | C1 l => match w_succ_c xh with | C0 h => C0 (WW h w_0) | C1 h => C1 W0 @@ -47,13 +47,13 @@ Section DoubleAdd. end end. - Definition ww_succ x := + Definition ww_succ x := match x with | W0 => ww_1 | WW xh xl => match w_succ_c xl with | C0 l => WW xh l - | C1 l => w_W0 (w_succ xh) + | C1 l => w_W0 (w_succ xh) end end. @@ -63,12 +63,12 @@ Section DoubleAdd. | _, W0 => C0 x | WW xh xl, WW yh yl => match w_add_c xl yl with - | C0 l => + | C0 l => match w_add_c xh yh with | C0 h => C0 (WW h l) | C1 h => C1 (w_WW h l) - end - | C1 l => + end + | C1 l => match w_add_carry_c xh yh with | C0 h => C0 (WW h l) | C1 h => C1 (w_WW h l) @@ -85,12 +85,12 @@ Section DoubleAdd. | _, W0 => f0 x | WW xh xl, WW yh yl => match w_add_c xl yl with - | C0 l => + | C0 l => match w_add_c xh yh with | C0 h => f0 (WW h l) | C1 h => f1 (w_WW h l) - end - | C1 l => + end + | C1 l => match w_add_carry_c xh yh with | C0 h => f0 (WW h l) | C1 h => f1 (w_WW h l) @@ -118,12 +118,12 @@ Section DoubleAdd. | WW xh xl, W0 => ww_succ_c (WW xh xl) | WW xh xl, WW yh yl => match w_add_carry_c xl yl with - | C0 l => + | C0 l => match w_add_c xh yh with | C0 h => C0 (WW h l) | C1 h => C1 (WW h l) end - | C1 l => + | C1 l => match w_add_carry_c xh yh with | C0 h => C0 (WW h l) | C1 h => C1 (w_WW h l) @@ -131,7 +131,7 @@ Section DoubleAdd. end end. - Definition ww_add_carry x y := + Definition ww_add_carry x y := match x, y with | W0, W0 => ww_1 | W0, WW yh yl => ww_succ (WW yh yl) @@ -146,7 +146,7 @@ Section DoubleAdd. (*Section DoubleProof.*) Variable w_digits : positive. Variable w_to_Z : w -> Z. - + Notation wB := (base w_digits). Notation wwB := (base (ww_digits w_digits)). @@ -157,11 +157,11 @@ Section DoubleAdd. (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). - Notation "[+[ c ]]" := - (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). - Notation "[-[ c ]]" := - (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). Variable spec_w_0 : [|w_0|] = 0. @@ -172,7 +172,7 @@ Section DoubleAdd. Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. Variable spec_w_succ_c : forall x, [+|w_succ_c x|] = [|x|] + 1. Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. - Variable spec_w_add_carry_c : + Variable spec_w_add_carry_c : forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1. Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB. Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB. @@ -187,11 +187,11 @@ Section DoubleAdd. rewrite <- Zplus_assoc;rewrite <- H;rewrite Zmult_1_l. assert ([|l|] = 0). generalize (spec_to_Z xl)(spec_to_Z l);omega. rewrite H0;generalize (spec_w_succ_c xh);destruct (w_succ_c xh) as [h|h]; - intro H1;unfold interp_carry in H1. + intro H1;unfold interp_carry in H1. simpl;rewrite H1;rewrite spec_w_0;ring. unfold interp_carry;simpl ww_to_Z;rewrite wwB_wBwB. assert ([|xh|] = wB - 1). generalize (spec_to_Z xh)(spec_to_Z h);omega. - rewrite H2;ring. + rewrite H2;ring. Qed. Lemma spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]]. @@ -222,12 +222,12 @@ Section DoubleAdd. Proof. destruct x as [ |xh xl];simpl;trivial. apply spec_f0;trivial. - destruct y as [ |yh yl];simpl. + destruct y as [ |yh yl];simpl. apply spec_f0;simpl;rewrite Zplus_0_r;trivial. generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l]; intros H;unfold interp_carry in H. generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; - intros H1;unfold interp_carry in *. + intros H1;unfold interp_carry in *. apply spec_f0. simpl;rewrite H;rewrite H1;ring. apply spec_f1. simpl;rewrite spec_w_WW;rewrite H. rewrite Zplus_assoc;rewrite wwB_wBwB. rewrite Zpower_2; rewrite <- Zmult_plus_distr_l. @@ -236,12 +236,12 @@ Section DoubleAdd. as [h|h]; intros H1;unfold interp_carry in *. apply spec_f0;simpl;rewrite H1. rewrite Zmult_plus_distr_l. rewrite <- Zplus_assoc;rewrite H;ring. - apply spec_f1. simpl;rewrite spec_w_WW;rewrite wwB_wBwB. - rewrite Zplus_assoc; rewrite Zpower_2; rewrite <- Zmult_plus_distr_l. + apply spec_f1. simpl;rewrite spec_w_WW;rewrite wwB_wBwB. + rewrite Zplus_assoc; rewrite Zpower_2; rewrite <- Zmult_plus_distr_l. rewrite Zmult_1_l in H1;rewrite H1. rewrite Zmult_plus_distr_l. rewrite <- Zplus_assoc;rewrite H;ring. Qed. - + End Cont. Lemma spec_ww_add_carry_c : @@ -251,16 +251,16 @@ Section DoubleAdd. exact (spec_ww_succ_c y). destruct y as [ |yh yl];simpl. rewrite Zplus_0_r;exact (spec_ww_succ_c (WW xh xl)). - replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1) + replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1) with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring. - generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) + generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) as [l|l];intros H;unfold interp_carry in H;rewrite <- H. generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; intros H1;unfold interp_carry in H1;rewrite <- H1. trivial. unfold interp_carry;repeat rewrite Zmult_1_l;simpl;rewrite wwB_wBwB;ring. rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. - generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh) - as [h|h];intros H1;unfold interp_carry in H1;rewrite <- H1. trivial. + generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh) + as [h|h];intros H1;unfold interp_carry in H1;rewrite <- H1. trivial. unfold interp_carry;rewrite spec_w_WW; repeat rewrite Zmult_1_l;simpl;rewrite wwB_wBwB;ring. Qed. @@ -287,9 +287,9 @@ Section DoubleAdd. rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial. destruct y as [ |yh yl]. change [[W0]] with 0;rewrite Zplus_0_r. - rewrite Zmod_small;trivial. + rewrite Zmod_small;trivial. exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)). - simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|])) + simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|])) with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring. generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l]; unfold interp_carry;intros H;simpl;rewrite <- H. @@ -305,14 +305,14 @@ Section DoubleAdd. exact (spec_ww_succ y). destruct y as [ |yh yl]. change [[W0]] with 0;rewrite Zplus_0_r. exact (spec_ww_succ (WW xh xl)). - simpl;replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1) + simpl;replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1) with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring. - generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) + generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) as [l|l];unfold interp_carry;intros H;rewrite <- H;simpl ww_to_Z. rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial. rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial. - Qed. + Qed. (* End DoubleProof. *) End DoubleAdd. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v index 952516ac..88c34915 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleBase.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -16,7 +16,7 @@ Require Import ZArith. Require Import BigNumPrelude. Require Import DoubleType. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section DoubleBase. Variable w : Type. @@ -29,8 +29,8 @@ Section DoubleBase. Variable w_zdigits: w. Variable w_add: w -> w -> zn2z w. Variable w_to_Z : w -> Z. - Variable w_compare : w -> w -> comparison. - + Variable w_compare : w -> w -> comparison. + Definition ww_digits := xO w_digits. Definition ww_zdigits := w_add w_zdigits w_zdigits. @@ -46,7 +46,7 @@ Section DoubleBase. | W0, W0 => W0 | _, _ => WW xh xl end. - + Definition ww_W0 h : zn2z (zn2z w) := match h with | W0 => W0 @@ -58,10 +58,10 @@ Section DoubleBase. | W0 => W0 | _ => WW W0 l end. - - Definition double_WW (n:nat) := - match n return word w n -> word w n -> word w (S n) with - | O => w_WW + + Definition double_WW (n:nat) := + match n return word w n -> word w n -> word w (S n) with + | O => w_WW | S n => fun (h l : zn2z (word w n)) => match h, l with @@ -70,8 +70,8 @@ Section DoubleBase. end end. - Fixpoint double_digits (n:nat) : positive := - match n with + Fixpoint double_digits (n:nat) : positive := + match n with | O => w_digits | S n => xO (double_digits n) end. @@ -80,7 +80,7 @@ Section DoubleBase. Fixpoint double_to_Z (n:nat) : word w n -> Z := match n return word w n -> Z with - | O => w_to_Z + | O => w_to_Z | S n => zn2z_to_Z (double_wB n) (double_to_Z n) end. @@ -98,21 +98,21 @@ Section DoubleBase. end. Definition double_0 n : word w n := - match n return word w n with + match n return word w n with | O => w_0 | S _ => W0 end. - + Definition double_split (n:nat) (x:zn2z (word w n)) := - match x with - | W0 => - match n return word w n * word w n with + match x with + | W0 => + match n return word w n * word w n with | O => (w_0,w_0) | S _ => (W0, W0) end | WW h l => (h,l) end. - + Definition ww_compare x y := match x, y with | W0, W0 => Eq @@ -148,15 +148,15 @@ Section DoubleBase. end end. - + Section DoubleProof. Notation wB := (base w_digits). Notation wwB := (base ww_digits). Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). Notation "[[ x ]]" := (ww_to_Z x) (at level 0, x at level 99). - Notation "[+[ c ]]" := + Notation "[+[ c ]]" := (interp_carry 1 wwB ww_to_Z c) (at level 0, x at level 99). - Notation "[-[ c ]]" := + Notation "[-[ c ]]" := (interp_carry (-1) wwB ww_to_Z c) (at level 0, x at level 99). Notation "[! n | x !]" := (double_to_Z n x) (at level 0, x at level 99). @@ -188,7 +188,7 @@ Section DoubleBase. Proof. simpl;rewrite spec_w_Bm1;rewrite wwB_wBwB;ring. Qed. Lemma lt_0_wB : 0 < wB. - Proof. + Proof. unfold base;apply Zpower_gt_0. unfold Zlt;reflexivity. unfold Zle;intros H;discriminate H. Qed. @@ -197,25 +197,25 @@ Section DoubleBase. Proof. rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_lt_0_compat;apply lt_0_wB. Qed. Lemma wB_pos: 1 < wB. - Proof. + Proof. unfold base;apply Zlt_le_trans with (2^1). unfold Zlt;reflexivity. apply Zpower_le_monotone. unfold Zlt;reflexivity. split;unfold Zle;intros H. discriminate H. clear spec_w_0W w_0W spec_w_Bm1 spec_to_Z spec_w_WW w_WW. destruct w_digits; discriminate H. Qed. - - Lemma wwB_pos: 1 < wwB. + + Lemma wwB_pos: 1 < wwB. Proof. assert (H:= wB_pos);rewrite wwB_wBwB;rewrite <-(Zmult_1_r 1). rewrite Zpower_2. apply Zmult_lt_compat2;(split;[unfold Zlt;reflexivity|trivial]). - apply Zlt_le_weak;trivial. + apply Zlt_le_weak;trivial. Qed. Theorem wB_div_2: 2 * (wB / 2) = wB. Proof. - clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W + clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W spec_to_Z;unfold base. assert (2 ^ Zpos w_digits = 2 * (2 ^ (Zpos w_digits - 1))). pattern 2 at 2; rewrite <- Zpower_1_r. @@ -228,7 +228,7 @@ Section DoubleBase. Theorem wwB_div_2 : wwB / 2 = wB / 2 * wB. Proof. - clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W + clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W spec_to_Z. rewrite wwB_wBwB; rewrite Zpower_2. pattern wB at 1; rewrite <- wB_div_2; auto. @@ -236,11 +236,11 @@ Section DoubleBase. repeat (rewrite (Zmult_comm 2); rewrite Z_div_mult); auto with zarith. Qed. - Lemma mod_wwB : forall z x, + Lemma mod_wwB : forall z x, (z*wB + [|x|]) mod wwB = (z mod wB)*wB + [|x|]. Proof. intros z x. - rewrite Zplus_mod. + rewrite Zplus_mod. pattern wwB at 1;rewrite wwB_wBwB; rewrite Zpower_2. rewrite Zmult_mod_distr_r;try apply lt_0_wB. rewrite (Zmod_small [|x|]). @@ -260,8 +260,8 @@ Section DoubleBase. destruct (spec_to_Z x);trivial. Qed. - Lemma wB_div_plus : forall x y p, - 0 <= p -> + Lemma wB_div_plus : forall x y p, + 0 <= p -> ([|x|]*wB + [|y|]) / 2^(Zpos w_digits + p) = [|x|] / 2^p. Proof. clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. @@ -277,7 +277,7 @@ Section DoubleBase. assert (0 < Zpos w_digits). compute;reflexivity. unfold ww_digits;rewrite Zpos_xO;auto with zarith. Qed. - + Lemma w_to_Z_wwB : forall x, x < wB -> x < wwB. Proof. intros x H;apply Zlt_trans with wB;trivial;apply lt_wB_wwB. @@ -298,7 +298,7 @@ Section DoubleBase. Proof. intros n;unfold double_wB;simpl. unfold base;rewrite (Zpos_xO (double_digits n)). - replace (2 * Zpos (double_digits n)) with + replace (2 * Zpos (double_digits n)) with (Zpos (double_digits n) + Zpos (double_digits n)). symmetry; apply Zpower_exp;intro;discriminate. ring. @@ -327,7 +327,7 @@ Section DoubleBase. unfold base; auto with zarith. Qed. - Lemma spec_double_to_Z : + Lemma spec_double_to_Z : forall n (x:word w n), 0 <= [!n | x!] < double_wB n. Proof. clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. @@ -347,7 +347,7 @@ Section DoubleBase. Qed. Lemma spec_get_low: - forall n x, + forall n x, [!n | x!] < wB -> [|get_low n x|] = [!n | x!]. Proof. clear spec_w_1 spec_w_Bm1. @@ -380,19 +380,19 @@ Section DoubleBase. Qed. Lemma spec_extend_aux : forall n x, [!S n|extend_aux n x!] = [[x]]. - Proof. induction n;simpl;trivial. Qed. + Proof. induction n;simpl;trivial. Qed. Lemma spec_extend : forall n x, [!S n|extend n x!] = [|x|]. - Proof. + Proof. intros n x;assert (H:= spec_w_0W x);unfold extend. - destruct (w_0W x);simpl;trivial. + destruct (w_0W x);simpl;trivial. rewrite <- H;exact (spec_extend_aux n (WW w0 w1)). Qed. Lemma spec_double_0 : forall n, [!n|double_0 n!] = 0. Proof. destruct n;trivial. Qed. - Lemma spec_double_split : forall n x, + Lemma spec_double_split : forall n x, let (h,l) := double_split n x in [!S n|x!] = [!n|h!] * double_wB n + [!n|l!]. Proof. @@ -401,9 +401,9 @@ Section DoubleBase. rewrite spec_w_0;trivial. Qed. - Lemma wB_lex_inv: forall a b c d, - a < c -> - a * wB + [|b|] < c * wB + [|d|]. + Lemma wB_lex_inv: forall a b c d, + a < c -> + a * wB + [|b|] < c * wB + [|d|]. Proof. intros a b c d H1; apply beta_lex_inv with (1 := H1); auto. Qed. @@ -420,7 +420,7 @@ Section DoubleBase. intros H;rewrite spec_w_0 in H. rewrite <- H;simpl;rewrite <- spec_w_0;apply spec_w_compare. change 0 with (0*wB+0);pattern 0 at 2;rewrite <- spec_w_0. - apply wB_lex_inv;trivial. + apply wB_lex_inv;trivial. absurd (0 <= [|yh|]). apply Zgt_not_le;trivial. destruct (spec_to_Z yh);trivial. generalize (spec_w_compare xh w_0);destruct (w_compare xh w_0); @@ -429,8 +429,8 @@ Section DoubleBase. absurd (0 <= [|xh|]). apply Zgt_not_le;apply Zlt_gt;trivial. destruct (spec_to_Z xh);trivial. apply Zlt_gt;change 0 with (0*wB+0);pattern 0 at 2;rewrite <- spec_w_0. - apply wB_lex_inv;apply Zgt_lt;trivial. - + apply wB_lex_inv;apply Zgt_lt;trivial. + generalize (spec_w_compare xh yh);destruct (w_compare xh yh);intros H. rewrite H;generalize (spec_w_compare xl yl);destruct (w_compare xl yl); intros H1;[rewrite H1|apply Zplus_lt_compat_l|apply Zplus_gt_compat_l]; @@ -439,7 +439,7 @@ Section DoubleBase. apply Zlt_gt;apply wB_lex_inv;apply Zgt_lt;trivial. Qed. - + End DoubleProof. End DoubleBase. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v index cca32a59..eea29e7c 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleCyclic.v 11012 2008-05-28 16:34:43Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -22,10 +22,10 @@ Require Import DoubleMul. Require Import DoubleSqrt. Require Import DoubleLift. Require Import DoubleDivn1. -Require Import DoubleDiv. +Require Import DoubleDiv. Require Import CyclicAxioms. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section Z_2nZ. @@ -80,7 +80,7 @@ Section Z_2nZ. Let w_gcd_gt := w_op.(znz_gcd_gt). Let w_gcd := w_op.(znz_gcd). - Let w_add_mul_div := w_op.(znz_add_mul_div). + Let w_add_mul_div := w_op.(znz_add_mul_div). Let w_pos_mod := w_op.(znz_pos_mod). @@ -93,7 +93,7 @@ Section Z_2nZ. Let wB := base w_digits. Let w_Bm2 := w_pred w_Bm1. - + Let ww_1 := ww_1 w_0 w_1. Let ww_Bm1 := ww_Bm1 w_Bm1. @@ -112,16 +112,16 @@ Section Z_2nZ. Let ww_of_pos p := match w_of_pos p with | (N0, l) => (N0, WW w_0 l) - | (Npos ph,l) => + | (Npos ph,l) => let (n,h) := w_of_pos ph in (n, w_WW h l) end. Let head0 := - Eval lazy beta delta [ww_head0] in + Eval lazy beta delta [ww_head0] in ww_head0 w_0 w_0W w_compare w_head0 w_add2 w_zdigits _ww_zdigits. Let tail0 := - Eval lazy beta delta [ww_tail0] in + Eval lazy beta delta [ww_tail0] in ww_tail0 w_0 w_0W w_compare w_tail0 w_add2 w_zdigits _ww_zdigits. Let ww_WW := Eval lazy beta delta [ww_WW] in (@ww_WW w). @@ -132,7 +132,7 @@ Section Z_2nZ. Let compare := Eval lazy beta delta[ww_compare] in ww_compare w_0 w_compare. - Let eq0 (x:zn2z w) := + Let eq0 (x:zn2z w) := match x with | W0 => true | _ => false @@ -147,7 +147,7 @@ Section Z_2nZ. Let opp_carry := Eval lazy beta delta [ww_opp_carry] in ww_opp_carry w_WW ww_Bm1 w_opp_carry. - + (* ** Additions ** *) Let succ_c := @@ -157,16 +157,16 @@ Section Z_2nZ. Eval lazy beta delta [ww_add_c] in ww_add_c w_WW w_add_c w_add_carry_c. Let add_carry_c := - Eval lazy beta iota delta [ww_add_carry_c ww_succ_c] in + Eval lazy beta iota delta [ww_add_carry_c ww_succ_c] in ww_add_carry_c w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c. - Let succ := + Let succ := Eval lazy beta delta [ww_succ] in ww_succ w_W0 ww_1 w_succ_c w_succ. Let add := Eval lazy beta delta [ww_add] in ww_add w_add_c w_add w_add_carry. - Let add_carry := + Let add_carry := Eval lazy beta iota delta [ww_add_carry ww_succ] in ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ w_add w_add_carry. @@ -174,9 +174,9 @@ Section Z_2nZ. Let pred_c := Eval lazy beta delta [ww_pred_c] in ww_pred_c w_Bm1 w_WW ww_Bm1 w_pred_c. - + Let sub_c := - Eval lazy beta iota delta [ww_sub_c ww_opp_c] in + Eval lazy beta iota delta [ww_sub_c ww_opp_c] in ww_sub_c w_0 w_WW w_opp_c w_opp_carry w_sub_c w_sub_carry_c. Let sub_carry_c := @@ -186,8 +186,8 @@ Section Z_2nZ. Let pred := Eval lazy beta delta [ww_pred] in ww_pred w_Bm1 w_WW ww_Bm1 w_pred_c w_pred. - Let sub := - Eval lazy beta iota delta [ww_sub ww_opp] in + Let sub := + Eval lazy beta iota delta [ww_sub ww_opp] in ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry. Let sub_carry := @@ -204,7 +204,7 @@ Section Z_2nZ. Let karatsuba_c := Eval lazy beta iota delta [ww_karatsuba_c double_mul_c kara_prod] in - ww_karatsuba_c w_0 w_1 w_WW w_W0 w_compare w_add w_sub w_mul_c + ww_karatsuba_c w_0 w_1 w_WW w_W0 w_compare w_add w_sub w_mul_c add_c add add_carry sub_c sub. Let mul := @@ -219,7 +219,7 @@ Section Z_2nZ. Let div32 := Eval lazy beta iota delta [w_div32] in - w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c + w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c w_add w_add_carry w_pred w_sub w_mul_c w_div21 sub_c. Let div21 := @@ -234,40 +234,40 @@ Section Z_2nZ. Let div_gt := Eval lazy beta delta [ww_div_gt] in - ww_div_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp + ww_div_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits. Let div := Eval lazy beta delta [ww_div] in ww_div ww_1 compare div_gt. - + Let mod_gt := Eval lazy beta delta [ww_mod_gt] in ww_mod_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_mod_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits add_mul_div w_zdigits. - Let mod_ := + Let mod_ := Eval lazy beta delta [ww_mod] in ww_mod compare mod_gt. - Let pos_mod := - Eval lazy beta delta [ww_pos_mod] in + Let pos_mod := + Eval lazy beta delta [ww_pos_mod] in ww_pos_mod w_0 w_zdigits w_WW w_pos_mod compare w_0W low sub _ww_zdigits. - Let is_even := + Let is_even := Eval lazy beta delta [ww_is_even] in ww_is_even w_is_even. - Let sqrt2 := + Let sqrt2 := Eval lazy beta delta [ww_sqrt2] in ww_sqrt2 w_is_even w_compare w_0 w_1 w_Bm1 w_0W w_sub w_square_c w_div21 w_add_mul_div w_zdigits w_add_c w_sqrt2 w_pred pred_c pred add_c add sub_c add_mul_div. - Let sqrt := + Let sqrt := Eval lazy beta delta [ww_sqrt] in ww_sqrt w_is_even w_0 w_sub w_add_mul_div w_zdigits _ww_zdigits w_sqrt2 pred add_mul_div head0 compare low. - Let gcd_gt_fix := + Let gcd_gt_fix := Eval cbv beta delta [ww_gcd_gt_aux ww_gcd_gt_body] in ww_gcd_gt_aux w_0 w_WW w_0W w_compare w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt @@ -278,7 +278,7 @@ Section Z_2nZ. Eval lazy beta delta [gcd_cont] in gcd_cont ww_1 w_1 w_compare. Let gcd_gt := - Eval lazy beta delta [ww_gcd_gt] in + Eval lazy beta delta [ww_gcd_gt] in ww_gcd_gt w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont. Let gcd := @@ -286,18 +286,18 @@ Section Z_2nZ. ww_gcd compare w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont. (* ** Record of operators on 2 words *) - - Definition mk_zn2z_op := + + Definition mk_zn2z_op := mk_znz_op _ww_digits _ww_zdigits to_Z ww_of_pos head0 tail0 W0 ww_1 ww_Bm1 compare eq0 opp_c opp opp_carry - succ_c add_c add_carry_c - succ add add_carry - pred_c sub_c sub_carry_c + succ_c add_c add_carry_c + succ add add_carry + pred_c sub_c sub_carry_c pred sub sub_carry - mul_c mul square_c + mul_c mul square_c div21 div_gt div mod_gt mod_ gcd_gt gcd @@ -307,17 +307,17 @@ Section Z_2nZ. sqrt2 sqrt. - Definition mk_zn2z_op_karatsuba := + Definition mk_zn2z_op_karatsuba := mk_znz_op _ww_digits _ww_zdigits to_Z ww_of_pos head0 tail0 W0 ww_1 ww_Bm1 compare eq0 opp_c opp opp_carry - succ_c add_c add_carry_c - succ add add_carry - pred_c sub_c sub_carry_c + succ_c add_c add_carry_c + succ add add_carry + pred_c sub_c sub_carry_c pred sub sub_carry - karatsuba_c mul square_c + karatsuba_c mul square_c div21 div_gt div mod_gt mod_ gcd_gt gcd @@ -330,7 +330,7 @@ Section Z_2nZ. (* Proof *) Variable op_spec : znz_spec w_op. - Hint Resolve + Hint Resolve (spec_to_Z op_spec) (spec_of_pos op_spec) (spec_0 op_spec) @@ -358,13 +358,13 @@ Section Z_2nZ. (spec_square_c op_spec) (spec_div21 op_spec) (spec_div_gt op_spec) - (spec_div op_spec) + (spec_div op_spec) (spec_mod_gt op_spec) - (spec_mod op_spec) + (spec_mod op_spec) (spec_gcd_gt op_spec) - (spec_gcd op_spec) - (spec_head0 op_spec) - (spec_tail0 op_spec) + (spec_gcd op_spec) + (spec_head0 op_spec) + (spec_tail0 op_spec) (spec_add_mul_div op_spec) (spec_pos_mod) (spec_is_even) @@ -417,20 +417,20 @@ Section Z_2nZ. Let spec_ww_Bm1 : [|ww_Bm1|] = wwB - 1. Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed. - Let spec_ww_compare : + Let spec_ww_compare : forall x y, match compare x y with | Eq => [|x|] = [|y|] | Lt => [|x|] < [|y|] | Gt => [|x|] > [|y|] end. - Proof. - refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);auto. - exact (spec_compare op_spec). + Proof. + refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);auto. + exact (spec_compare op_spec). Qed. Let spec_ww_eq0 : forall x, eq0 x = true -> [|x|] = 0. - Proof. destruct x;simpl;intros;trivial;discriminate. Qed. + Proof. destruct x;simpl;intros;trivial;discriminate. Qed. Let spec_ww_opp_c : forall x, [-|opp_c x|] = -[|x|]. Proof. @@ -440,7 +440,7 @@ Section Z_2nZ. Let spec_ww_opp : forall x, [|opp x|] = (-[|x|]) mod wwB. Proof. - refine(spec_ww_opp w_0 w_0 W0 w_opp_c w_opp_carry w_opp + refine(spec_ww_opp w_0 w_0 W0 w_opp_c w_opp_carry w_opp w_digits w_to_Z _ _ _ _ _); auto. Qed. @@ -480,25 +480,25 @@ Section Z_2nZ. Let spec_ww_add_carry : forall x y, [|add_carry x y|]=([|x|]+[|y|]+1)mod wwB. Proof. - refine (spec_ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ + refine (spec_ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ w_add w_add_carry w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto. Qed. Let spec_ww_pred_c : forall x, [-|pred_c x|] = [|x|] - 1. Proof. - refine (spec_ww_pred_c w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_digits w_to_Z + refine (spec_ww_pred_c w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_digits w_to_Z _ _ _ _ _);wwauto. Qed. Let spec_ww_sub_c : forall x y, [-|sub_c x y|] = [|x|] - [|y|]. Proof. - refine (spec_ww_sub_c w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c + refine (spec_ww_sub_c w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _);wwauto. Qed. Let spec_ww_sub_carry_c : forall x y, [-|sub_carry_c x y|] = [|x|]-[|y|]-1. Proof. - refine (spec_ww_sub_carry_c w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c + refine (spec_ww_sub_carry_c w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c w_sub_c w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto. Qed. @@ -533,17 +533,17 @@ Section Z_2nZ. _ _ _ _ _ _ _ _ _ _ _ _); wwauto. unfold w_digits; apply spec_more_than_1_digit; auto. exact (spec_compare op_spec). - Qed. + Qed. Let spec_ww_mul : forall x y, [|mul x y|] = ([|x|] * [|y|]) mod wwB. Proof. refine (spec_ww_mul w_W0 w_add w_mul_c w_mul add w_digits w_to_Z _ _ _ _ _); - wwauto. + wwauto. Qed. Let spec_ww_square_c : forall x, [[square_c x]] = [|x|] * [|x|]. Proof. - refine (spec_ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add + refine (spec_ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add add_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _);wwauto. Qed. @@ -574,7 +574,7 @@ Section Z_2nZ. 0 <= [|r|] < [|b|]. Proof. refine (spec_ww_div21 w_0 w_0W div32 ww_1 compare sub w_digits w_to_Z - _ _ _ _ _ _ _);wwauto. + _ _ _ _ _ _ _);wwauto. Qed. Let spec_add2: forall x y, @@ -602,7 +602,7 @@ Section Z_2nZ. unfold wB, base; auto with zarith. Qed. - Let spec_ww_digits: + Let spec_ww_digits: [|_ww_zdigits|] = Zpos (xO w_digits). Proof. unfold w_to_Z, _ww_zdigits. @@ -615,7 +615,7 @@ Section Z_2nZ. Let spec_ww_head00 : forall x, [|x|] = 0 -> [|head0 x|] = Zpos _ww_digits. Proof. - refine (spec_ww_head00 w_0 w_0W + refine (spec_ww_head00 w_0 w_0W w_compare w_head0 w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); auto. exact (spec_compare op_spec). @@ -626,8 +626,8 @@ Section Z_2nZ. Let spec_ww_head0 : forall x, 0 < [|x|] -> wwB/ 2 <= 2 ^ [|head0 x|] * [|x|] < wwB. Proof. - refine (spec_ww_head0 w_0 w_0W w_compare w_head0 - w_add2 w_zdigits _ww_zdigits + refine (spec_ww_head0 w_0 w_0W w_compare w_head0 + w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ _ _ _ _);wwauto. exact (spec_compare op_spec). exact (spec_zdigits op_spec). @@ -635,7 +635,7 @@ Section Z_2nZ. Let spec_ww_tail00 : forall x, [|x|] = 0 -> [|tail0 x|] = Zpos _ww_digits. Proof. - refine (spec_ww_tail00 w_0 w_0W + refine (spec_ww_tail00 w_0 w_0W w_compare w_tail0 w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); wwauto. exact (spec_compare op_spec). @@ -647,7 +647,7 @@ Section Z_2nZ. Let spec_ww_tail0 : forall x, 0 < [|x|] -> exists y, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail0 x|]. Proof. - refine (spec_ww_tail0 (w_digits := w_digits) w_0 w_0W w_compare w_tail0 + refine (spec_ww_tail0 (w_digits := w_digits) w_0 w_0W w_compare w_tail0 w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ _ _ _ _);wwauto. exact (spec_compare op_spec). exact (spec_zdigits op_spec). @@ -659,19 +659,19 @@ Section Z_2nZ. ([|x|] * (2 ^ [|p|]) + [|y|] / (2 ^ ((Zpos _ww_digits) - [|p|]))) mod wwB. Proof. - refine (@spec_ww_add_mul_div w w_0 w_WW w_W0 w_0W compare w_add_mul_div + refine (@spec_ww_add_mul_div w w_0 w_WW w_W0 w_0W compare w_add_mul_div sub w_digits w_zdigits low w_to_Z _ _ _ _ _ _ _ _ _ _ _);wwauto. exact (spec_zdigits op_spec). Qed. - Let spec_ww_div_gt : forall a b, + Let spec_ww_div_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> let (q,r) := div_gt a b in [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]. Proof. -refine -(@spec_ww_div_gt w w_digits w_0 w_WW w_0W w_compare w_eq0 +refine +(@spec_ww_div_gt w w_digits w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ @@ -707,14 +707,14 @@ refine refine (spec_ww_div w_digits ww_1 compare div_gt w_to_Z _ _ _ _);auto. Qed. - Let spec_ww_mod_gt : forall a b, + Let spec_ww_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> [|mod_gt a b|] = [|a|] mod [|b|]. Proof. - refine (@spec_ww_mod_gt w w_digits w_0 w_WW w_0W w_compare w_eq0 + refine (@spec_ww_mod_gt w w_digits w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt w_mod_gt - w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div - w_zdigits w_to_Z + w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div + w_zdigits w_to_Z _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. exact (spec_compare op_spec). exact (spec_div_gt op_spec). @@ -731,12 +731,12 @@ refine Let spec_ww_gcd_gt : forall a b, [|a|] > [|b|] -> Zis_gcd [|a|] [|b|] [|gcd_gt a b|]. Proof. - refine (@spec_ww_gcd_gt w w_digits W0 w_to_Z _ + refine (@spec_ww_gcd_gt w w_digits W0 w_to_Z _ w_0 w_0 w_eq0 w_gcd_gt _ww_digits _ gcd_gt_fix _ _ _ _ gcd_cont _);auto. refine (@spec_ww_gcd_gt_aux w w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0 - w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z + w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. exact (spec_compare op_spec). exact (spec_div21 op_spec). @@ -753,7 +753,7 @@ refine _ww_digits _ gcd_gt_fix _ _ _ _ gcd_cont _);auto. refine (@spec_ww_gcd_gt_aux w w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0 - w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z + w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. exact (spec_compare op_spec). exact (spec_div21 op_spec). @@ -798,7 +798,7 @@ refine Let spec_ww_sqrt : forall x, [|sqrt x|] ^ 2 <= [|x|] < ([|sqrt x|] + 1) ^ 2. Proof. - refine (@spec_ww_sqrt w w_is_even w_0 w_1 w_Bm1 + refine (@spec_ww_sqrt w w_is_even w_0 w_1 w_Bm1 w_sub w_add_mul_div w_digits w_zdigits _ww_zdigits w_sqrt2 pred add_mul_div head0 compare _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto. @@ -814,7 +814,7 @@ refine apply mk_znz_spec;auto. exact spec_ww_add_mul_div. - refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW + refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z _ _ _ _ _ _ _ _ _ _ _ _);wwauto. exact (spec_pos_mod op_spec). @@ -828,7 +828,7 @@ refine Proof. apply mk_znz_spec;auto. exact spec_ww_add_mul_div. - refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW + refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z _ _ _ _ _ _ _ _ _ _ _ _);wwauto. exact (spec_pos_mod op_spec). @@ -838,10 +838,10 @@ refine rewrite <- Zpos_xO; exact spec_ww_digits. Qed. -End Z_2nZ. - +End Z_2nZ. + Section MulAdd. - + Variable w: Type. Variable op: znz_op w. Variable sop: znz_spec op. @@ -870,7 +870,7 @@ Section MulAdd. End MulAdd. -(** Modular versions of DoubleCyclic *) +(** Modular versions of DoubleCyclic *) Module DoubleCyclic (C:CyclicType) <: CyclicType. Definition w := zn2z C.w. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v index 075aef59..9204b4e0 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleDiv.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -20,7 +20,7 @@ Require Import DoubleDivn1. Require Import DoubleAdd. Require Import DoubleSub. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Ltac zarith := auto with zarith. @@ -41,13 +41,13 @@ Section POS_MOD. Variable ww_zdigits : zn2z w. - Definition ww_pos_mod p x := + Definition ww_pos_mod p x := let zdigits := w_0W w_zdigits in match x with | W0 => W0 | WW xh xl => match ww_compare p zdigits with - | Eq => w_WW w_0 xl + | Eq => w_WW w_0 xl | Lt => w_WW w_0 (w_pos_mod (low p) xl) | Gt => match ww_compare p ww_zdigits with @@ -87,7 +87,7 @@ Section POS_MOD. | Lt => [[x]] < [[y]] | Gt => [[x]] > [[y]] end. - Variable spec_ww_sub: forall x y, + Variable spec_ww_sub: forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits. @@ -106,7 +106,7 @@ Section POS_MOD. unfold ww_pos_mod; case w1. simpl; rewrite Zmod_small; split; auto with zarith. intros xh xl; generalize (spec_ww_compare p (w_0W w_zdigits)); - case ww_compare; + case ww_compare; rewrite spec_w_0W; rewrite spec_zdigits; fold wB; intros H1. rewrite H1; simpl ww_to_Z. @@ -135,13 +135,13 @@ Section POS_MOD. autorewrite with w_rewrite rm10. rewrite Zmod_mod; auto with zarith. generalize (spec_ww_compare p ww_zdigits); - case ww_compare; rewrite spec_ww_zdigits; + case ww_compare; rewrite spec_ww_zdigits; rewrite spec_zdigits; intros H2. replace (2^[[p]]) with wwB. rewrite Zmod_small; auto with zarith. unfold base; rewrite H2. rewrite spec_ww_digits; auto. - assert (HH0: [|low (ww_sub p (w_0W w_zdigits))|] = + assert (HH0: [|low (ww_sub p (w_0W w_zdigits))|] = [[p]] - Zpos w_digits). rewrite spec_low. rewrite spec_ww_sub. @@ -152,11 +152,11 @@ generalize (spec_ww_compare p ww_zdigits); apply Zlt_le_trans with (Zpos w_digits); auto with zarith. unfold base; apply Zpower2_le_lin; auto with zarith. exists wB; unfold base; rewrite <- Zpower_exp; auto with zarith. - rewrite spec_ww_digits; + rewrite spec_ww_digits; apply f_equal with (f := Zpower 2); rewrite Zpos_xO; auto with zarith. simpl ww_to_Z; autorewrite with w_rewrite. rewrite spec_pos_mod; rewrite HH0. - pattern [|xh|] at 2; + pattern [|xh|] at 2; rewrite Z_div_mod_eq with (b := 2 ^ ([[p]] - Zpos w_digits)); auto with zarith. rewrite (fun x => (Zmult_comm (2 ^ x))); rewrite Zmult_plus_distr_l. @@ -196,7 +196,7 @@ generalize (spec_ww_compare p ww_zdigits); split; auto with zarith. rewrite Zpos_xO; auto with zarith. Qed. - + End POS_MOD. Section DoubleDiv32. @@ -222,24 +222,24 @@ Section DoubleDiv32. match w_compare a1 b1 with | Lt => let (q,r) := w_div21 a1 a2 b1 in - match ww_sub_c (w_WW r a3) (w_mul_c q b2) with + match ww_sub_c (w_WW r a3) (w_mul_c q b2) with | C0 r1 => (q,r1) | C1 r1 => let q := w_pred q in - ww_add_c_cont w_WW w_add_c w_add_carry_c + ww_add_c_cont w_WW w_add_c w_add_carry_c (fun r2=>(w_pred q, ww_add w_add_c w_add w_add_carry r2 (WW b1 b2))) (fun r2 => (q,r2)) r1 (WW b1 b2) end | Eq => - ww_add_c_cont w_WW w_add_c w_add_carry_c + ww_add_c_cont w_WW w_add_c w_add_carry_c (fun r => (w_Bm2, ww_add w_add_c w_add w_add_carry r (WW b1 b2))) (fun r => (w_Bm1,r)) (WW (w_sub a2 b2) a3) (WW b1 b2) | Gt => (w_0, W0) (* cas absurde *) end. - (* Proof *) + (* Proof *) Variable w_digits : positive. Variable w_to_Z : w -> Z. @@ -253,8 +253,8 @@ Section DoubleDiv32. (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). - Notation "[-[ c ]]" := - (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). @@ -273,7 +273,7 @@ Section DoubleDiv32. | Gt => [|x|] > [|y|] end. Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. - Variable spec_w_add_carry_c : + Variable spec_w_add_carry_c : forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1. Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB. @@ -315,8 +315,8 @@ Section DoubleDiv32. wB/2 <= [|b1|] -> [[WW a1 a2]] < [[WW b1 b2]] -> let (q,r) := w_div32 a1 a2 a3 b1 b2 in - [|a1|] * wwB + [|a2|] * wB + [|a3|] = - [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + [|a1|] * wwB + [|a2|] * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ 0 <= [[r]] < [|b1|] * wB + [|b2|]. Proof. intros a1 a2 a3 b1 b2 Hle Hlt. @@ -327,17 +327,17 @@ Section DoubleDiv32. match w_compare a1 b1 with | Lt => let (q,r) := w_div21 a1 a2 b1 in - match ww_sub_c (w_WW r a3) (w_mul_c q b2) with + match ww_sub_c (w_WW r a3) (w_mul_c q b2) with | C0 r1 => (q,r1) | C1 r1 => let q := w_pred q in - ww_add_c_cont w_WW w_add_c w_add_carry_c + ww_add_c_cont w_WW w_add_c w_add_carry_c (fun r2=>(w_pred q, ww_add w_add_c w_add w_add_carry r2 (WW b1 b2))) (fun r2 => (q,r2)) r1 (WW b1 b2) end | Eq => - ww_add_c_cont w_WW w_add_c w_add_carry_c + ww_add_c_cont w_WW w_add_c w_add_carry_c (fun r => (w_Bm2, ww_add w_add_c w_add w_add_carry r (WW b1 b2))) (fun r => (w_Bm1,r)) (WW (w_sub a2 b2) a3) (WW b1 b2) @@ -360,7 +360,7 @@ Section DoubleDiv32. [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ 0 <= [[r]] < [|b1|] * wB + [|b2|]);eauto. rewrite H0;intros r. - repeat + repeat (rewrite spec_ww_add;eauto || rewrite spec_w_Bm1 || rewrite spec_w_Bm2); simpl ww_to_Z;try rewrite Zmult_1_l;intros H1. assert (0<= ([[r]] + ([|b1|] * wB + [|b2|])) - wwB < [|b1|] * wB + [|b2|]). @@ -385,7 +385,7 @@ Section DoubleDiv32. 1 ([[r]] + ([|b1|] * wB + [|b2|]) - wwB));zarith;try (ring;fail). split. rewrite H1;rewrite Hcmp;ring. trivial. Spec_ww_to_Z (WW b1 b2). simpl in HH4;zarith. - rewrite H0;intros r;repeat + rewrite H0;intros r;repeat (rewrite spec_w_Bm1 || rewrite spec_w_Bm2); simpl ww_to_Z;try rewrite Zmult_1_l;intros H1. assert ([[r]]=([|a2|]-[|b2|])*wB+[|a3|]+([|b1|]*wB+[|b2|])). zarith. @@ -409,7 +409,7 @@ Section DoubleDiv32. as [r1|r1];repeat (rewrite spec_w_WW || rewrite spec_mul_c); unfold interp_carry;intros H1. rewrite H1. - split. ring. split. + split. ring. split. rewrite <- H1;destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z r1);trivial. apply Zle_lt_trans with ([|r|] * wB + [|a3|]). assert ( 0 <= [|q|] * [|b2|]);zarith. @@ -418,7 +418,7 @@ Section DoubleDiv32. rewrite <- H1;ring. Spec_ww_to_Z r1; assert (0 <= [|r|]*wB). zarith. assert (0 < [|q|] * [|b2|]). zarith. - assert (0 < [|q|]). + assert (0 < [|q|]). apply Zmult_lt_0_reg_r_2 with [|b2|];zarith. eapply spec_ww_add_c_cont with (P := fun (x y:zn2z w) (res:w*zn2z w) => @@ -440,18 +440,18 @@ Section DoubleDiv32. wwB * 1 + ([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2 * ([|b1|] * wB + [|b2|]))). rewrite H7;rewrite H2;ring. - assert - ([|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|]) + assert + ([|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|]) < [|b1|]*wB + [|b2|]). Spec_ww_to_Z r2;omega. Spec_ww_to_Z (WW b1 b2). simpl in HH5. - assert - (0 <= [|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|]) + assert + (0 <= [|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|]) < wwB). split;try omega. replace (2*([|b1|]*wB+[|b2|])) with ((2*[|b1|])*wB+2*[|b2|]). 2:ring. assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB). rewrite wwB_wBwB; rewrite Zpower_2; zarith. omega. - rewrite <- (Zmod_unique + rewrite <- (Zmod_unique ([[r2]] + ([|b1|] * wB + [|b2|])) wwB 1 @@ -486,7 +486,7 @@ Section DoubleDiv21. Definition ww_div21 a1 a2 b := match a1 with - | W0 => + | W0 => match ww_compare a2 b with | Gt => (ww_1, ww_sub a2 b) | Eq => (ww_1, W0) @@ -529,8 +529,8 @@ Section DoubleDiv21. Notation wwB := (base (ww_digits w_digits)). Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). - Notation "[-[ c ]]" := - (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). Variable spec_w_0 : [|w_0|] = 0. @@ -540,8 +540,8 @@ Section DoubleDiv21. wB/2 <= [|b1|] -> [[WW a1 a2]] < [[WW b1 b2]] -> let (q,r) := w_div32 a1 a2 a3 b1 b2 in - [|a1|] * wwB + [|a2|] * wB + [|a3|] = - [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + [|a1|] * wwB + [|a2|] * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ 0 <= [[r]] < [|b1|] * wB + [|b2|]. Variable spec_ww_1 : [[ww_1]] = 1. Variable spec_ww_compare : forall x y, @@ -591,10 +591,10 @@ Section DoubleDiv21. intros Hlt H; match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] => generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U); intros q1 r H0 - end; (assert (Eq1: wB / 2 <= [|b1|]);[ + end; (assert (Eq1: wB / 2 <= [|b1|]);[ apply (@beta_lex (wB / 2) 0 [|b1|] [|b2|] wB); auto with zarith; autorewrite with rm10;repeat rewrite (Zmult_comm wB); - rewrite <- wwB_div_2; trivial + rewrite <- wwB_div_2; trivial | generalize (H0 Eq1 Hlt);clear H0;destruct r as [ |r1 r2];simpl; try rewrite spec_w_0; try rewrite spec_w_0W;repeat rewrite Zplus_0_r; intros (H1,H2) ]). @@ -611,10 +611,10 @@ Section DoubleDiv21. rewrite <- wwB_wBwB;rewrite H1. rewrite spec_w_0 in H4;rewrite Zplus_0_r in H4. repeat rewrite Zmult_plus_distr_l. rewrite <- (Zmult_assoc [|r1|]). - rewrite <- Zpower_2; rewrite <- wwB_wBwB;rewrite H4;simpl;ring. + rewrite <- Zpower_2; rewrite <- wwB_wBwB;rewrite H4;simpl;ring. split;[rewrite wwB_wBwB | split;zarith]. - replace (([|a1h|] * wB + [|a1l|]) * wB^2 + ([|a3|] * wB + [|a4|])) - with (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB+ [|a4|]). + replace (([|a1h|] * wB + [|a1l|]) * wB^2 + ([|a3|] * wB + [|a4|])) + with (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB+ [|a4|]). rewrite H1;ring. rewrite wwB_wBwB;ring. change [|a4|] with (0*wB+[|a4|]);apply beta_lex_inv;zarith. assert (1 <= wB/2);zarith. @@ -624,7 +624,7 @@ Section DoubleDiv21. intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end. split;trivial. replace (([|a1h|] * wB + [|a1l|]) * wwB + ([|a3|] * wB + [|a4|])) with - (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB + [|a4|]); + (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB + [|a4|]); [rewrite H1 | rewrite wwB_wBwB;ring]. replace (([|q1|]*([|b1|]*wB+[|b2|])+([|r1|]*wB+[|r2|]))*wB+[|a4|]) with (([|q1|]*([|b1|]*wB+[|b2|]))*wB+([|r1|]*wwB+[|r2|]*wB+[|a4|])); @@ -666,22 +666,22 @@ Section DoubleDivGt. Eval lazy beta iota delta [ww_sub ww_opp] in let p := w_head0 bh in match w_compare p w_0 with - | Gt => + | Gt => let b1 := w_add_mul_div p bh bl in let b2 := w_add_mul_div p bl w_0 in let a1 := w_add_mul_div p w_0 ah in let a2 := w_add_mul_div p ah al in let a3 := w_add_mul_div p al w_0 in let (q,r) := w_div32 a1 a2 a3 b1 b2 in - (WW w_0 q, ww_add_mul_div + (WW w_0 q, ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r) | _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry (WW ah al) (WW bh bl)) end. - Definition ww_div_gt a b := - Eval lazy beta iota delta [ww_div_gt_aux double_divn1 + Definition ww_div_gt a b := + Eval lazy beta iota delta [ww_div_gt_aux double_divn1 double_divn1_p double_divn1_p_aux double_divn1_0 double_divn1_0_aux double_split double_0 double_WW] in match a, b with @@ -691,11 +691,11 @@ Section DoubleDivGt. if w_eq0 ah then let (q,r) := w_div_gt al bl in (WW w_0 q, w_0W r) - else + else match w_compare w_0 bh with - | Eq => + | Eq => let(q,r):= - double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1 a bl in (q, w_0W r) | Lt => ww_div_gt_aux ah al bh bl @@ -707,7 +707,7 @@ Section DoubleDivGt. Eval lazy beta iota delta [ww_sub ww_opp] in let p := w_head0 bh in match w_compare p w_0 with - | Gt => + | Gt => let b1 := w_add_mul_div p bh bl in let b2 := w_add_mul_div p bl w_0 in let a1 := w_add_mul_div p w_0 ah in @@ -716,13 +716,13 @@ Section DoubleDivGt. let (q,r) := w_div32 a1 a2 a3 b1 b2 in ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r - | _ => + | _ => ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry (WW ah al) (WW bh bl) end. - Definition ww_mod_gt a b := - Eval lazy beta iota delta [ww_mod_gt_aux double_modn1 + Definition ww_mod_gt a b := + Eval lazy beta iota delta [ww_mod_gt_aux double_modn1 double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux double_split double_0 double_WW snd] in match a, b with @@ -730,10 +730,10 @@ Section DoubleDivGt. | _, W0 => W0 | WW ah al, WW bh bl => if w_eq0 ah then w_0W (w_mod_gt al bl) - else + else match w_compare w_0 bh with - | Eq => - w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + | Eq => + w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 w_compare w_sub 1 a bl) | Lt => ww_mod_gt_aux ah al bh bl | Gt => W0 (* cas absurde *) @@ -741,14 +741,14 @@ Section DoubleDivGt. end. Definition ww_gcd_gt_body (cont: w->w->w->w->zn2z w) (ah al bh bl: w) := - Eval lazy beta iota delta [ww_mod_gt_aux double_modn1 + Eval lazy beta iota delta [ww_mod_gt_aux double_modn1 double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux double_split double_0 double_WW snd] in match w_compare w_0 bh with | Eq => match w_compare w_0 bl with | Eq => WW ah al (* normalement n'arrive pas si forme normale *) - | Lt => + | Lt => let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 w_compare w_sub 1 (WW ah al) bl in WW w_0 (w_gcd_gt bl m) @@ -757,14 +757,14 @@ Section DoubleDivGt. | Lt => let m := ww_mod_gt_aux ah al bh bl in match m with - | W0 => WW bh bl + | W0 => WW bh bl | WW mh ml => match w_compare w_0 mh with | Eq => match w_compare w_0 ml with | Eq => WW bh bl - | _ => - let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + | _ => + let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 w_compare w_sub 1 (WW bh bl) ml in WW w_0 (w_gcd_gt ml r) end @@ -779,18 +779,18 @@ Section DoubleDivGt. end | Gt => W0 (* absurde *) end. - - Fixpoint ww_gcd_gt_aux - (p:positive) (cont: w -> w -> w -> w -> zn2z w) (ah al bh bl : w) + + Fixpoint ww_gcd_gt_aux + (p:positive) (cont: w -> w -> w -> w -> zn2z w) (ah al bh bl : w) {struct p} : zn2z w := - ww_gcd_gt_body + ww_gcd_gt_body (fun mh ml rh rl => match p with | xH => cont mh ml rh rl | xO p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl | xI p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl end) ah al bh bl. - + (* Proof *) Variable w_to_Z : w -> Z. @@ -816,7 +816,7 @@ Section DoubleDivGt. | Gt => [|x|] > [|y|] end. Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0. - + Variable spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|]. Variable spec_opp : forall x, [|w_opp x|] = (-[|x|]) mod wB. Variable spec_opp_carry : forall x, [|w_opp_carry x|] = wB - [|x|] - 1. @@ -854,8 +854,8 @@ Section DoubleDivGt. wB/2 <= [|b1|] -> [[WW a1 a2]] < [[WW b1 b2]] -> let (q,r) := w_div32 a1 a2 a3 b1 b2 in - [|a1|] * wwB + [|a2|] * wB + [|a3|] = - [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + [|a1|] * wwB + [|a2|] * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ 0 <= [[r]] < [|b1|] * wB + [|b2|]. Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits. @@ -899,14 +899,14 @@ Section DoubleDivGt. change (let (q, r) := let p := w_head0 bh in match w_compare p w_0 with - | Gt => + | Gt => let b1 := w_add_mul_div p bh bl in let b2 := w_add_mul_div p bl w_0 in let a1 := w_add_mul_div p w_0 ah in let a2 := w_add_mul_div p ah al in let a3 := w_add_mul_div p al w_0 in let (q,r) := w_div32 a1 a2 a3 b1 b2 in - (WW w_0 q, ww_add_mul_div + (WW w_0 q, ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r) | _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c @@ -931,7 +931,7 @@ Section DoubleDivGt. case (spec_to_Z (w_head0 bh)); auto with zarith. assert ([|w_head0 bh|] < Zpos w_digits). destruct (Z_lt_ge_dec [|w_head0 bh|] (Zpos w_digits));trivial. - elimtype False. + exfalso. assert (2 ^ [|w_head0 bh|] * [|bh|] >= wB);auto with zarith. apply Zle_ge; replace wB with (wB * 1);try ring. Spec_w_to_Z bh;apply Zmult_le_compat;zarith. @@ -945,11 +945,11 @@ Section DoubleDivGt. (spec_add_mul_div bl w_0 Hb); rewrite spec_w_0; repeat rewrite Zmult_0_l;repeat rewrite Zplus_0_l; rewrite Zdiv_0_l;repeat rewrite Zplus_0_r. - Spec_w_to_Z ah;Spec_w_to_Z bh. + Spec_w_to_Z ah;Spec_w_to_Z bh. unfold base;repeat rewrite Zmod_shift_r;zarith. assert (H3:=to_Z_div_minus_p ah HHHH);assert(H4:=to_Z_div_minus_p al HHHH); assert (H5:=to_Z_div_minus_p bl HHHH). - rewrite Zmult_comm in Hh. + rewrite Zmult_comm in Hh. assert (2^[|w_head0 bh|] < wB). unfold base;apply Zpower_lt_monotone;zarith. unfold base in H0;rewrite Zmod_small;zarith. fold wB; rewrite (Zmod_small ([|bh|] * 2 ^ [|w_head0 bh|]));zarith. @@ -964,15 +964,15 @@ Section DoubleDivGt. (w_add_mul_div (w_head0 bh) al w_0) (w_add_mul_div (w_head0 bh) bh bl) (w_add_mul_div (w_head0 bh) bl w_0)) as (q,r). - rewrite V1;rewrite V2. rewrite Zmult_plus_distr_l. - rewrite <- (Zplus_assoc ([|bh|] * 2 ^ [|w_head0 bh|] * wB)). + rewrite V1;rewrite V2. rewrite Zmult_plus_distr_l. + rewrite <- (Zplus_assoc ([|bh|] * 2 ^ [|w_head0 bh|] * wB)). unfold base;rewrite <- shift_unshift_mod;zarith. fold wB. replace ([|bh|] * 2 ^ [|w_head0 bh|] * wB + [|bl|] * 2 ^ [|w_head0 bh|]) with ([[WW bh bl]] * 2^[|w_head0 bh|]). 2:simpl;ring. fold wwB. rewrite wwB_wBwB. rewrite Zpower_2. rewrite U1;rewrite U2;rewrite U3. - rewrite Zmult_assoc. rewrite Zmult_plus_distr_l. + rewrite Zmult_assoc. rewrite Zmult_plus_distr_l. rewrite (Zplus_assoc ([|ah|] / 2^(Zpos(w_digits) - [|w_head0 bh|])*wB * wB)). - rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc. + rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc. unfold base;repeat rewrite <- shift_unshift_mod;zarith. fold wB. replace ([|ah|] * 2 ^ [|w_head0 bh|] * wB + [|al|] * 2 ^ [|w_head0 bh|]) with ([[WW ah al]] * 2^[|w_head0 bh|]). 2:simpl;ring. @@ -1027,7 +1027,7 @@ Section DoubleDivGt. [[a]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]. Proof. - intros a b Hgt Hpos;unfold ww_div_gt. + intros a b Hgt Hpos;unfold ww_div_gt. change (let (q,r) := match a, b with | W0, _ => (W0,W0) | _, W0 => (W0,W0) @@ -1035,23 +1035,23 @@ Section DoubleDivGt. if w_eq0 ah then let (q,r) := w_div_gt al bl in (WW w_0 q, w_0W r) - else + else match w_compare w_0 bh with - | Eq => + | Eq => let(q,r):= - double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1 a bl in (q, w_0W r) | Lt => ww_div_gt_aux ah al bh bl | Gt => (W0,W0) (* cas absurde *) end - end in [[a]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]). + end in [[a]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]). destruct a as [ |ah al]. simpl in Hgt;omega. destruct b as [ |bh bl]. simpl in Hpos;omega. Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl. assert (H:=@spec_eq0 ah);destruct (w_eq0 ah). simpl ww_to_Z;rewrite H;trivial. simpl in Hgt;rewrite H in Hgt;trivial. - assert ([|bh|] <= 0). + assert ([|bh|] <= 0). apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith. assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;rewrite H1;simpl in Hgt. simpl. simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos. @@ -1066,12 +1066,12 @@ Section DoubleDivGt. w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 spec_add_mul_div spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hpos). unfold double_to_Z,double_wB,double_digits in H2. - destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1 (WW ah al) bl). rewrite spec_w_0W;unfold ww_to_Z;trivial. apply spec_ww_div_gt_aux;trivial. rewrite spec_w_0 in Hcmp;trivial. - rewrite spec_w_0 in Hcmp;elimtype False;omega. + rewrite spec_w_0 in Hcmp;exfalso;omega. Qed. Lemma spec_ww_mod_gt_aux_eq : forall ah al bh bl, @@ -1104,26 +1104,26 @@ Section DoubleDivGt. rewrite Zmult_comm in H;destruct H. symmetry;apply Zmod_unique with [|q|];trivial. Qed. - + Lemma spec_ww_mod_gt_eq : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> [[ww_mod_gt a b]] = [[snd (ww_div_gt a b)]]. Proof. intros a b Hgt Hpos. - change (ww_mod_gt a b) with + change (ww_mod_gt a b) with (match a, b with | W0, _ => W0 | _, W0 => W0 | WW ah al, WW bh bl => if w_eq0 ah then w_0W (w_mod_gt al bl) - else + else match w_compare w_0 bh with - | Eq => - w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + | Eq => + w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 w_compare w_sub 1 a bl) | Lt => ww_mod_gt_aux ah al bh bl | Gt => W0 (* cas absurde *) end end). - change (ww_div_gt a b) with + change (ww_div_gt a b) with (match a, b with | W0, _ => (W0,W0) | _, W0 => (W0,W0) @@ -1131,11 +1131,11 @@ Section DoubleDivGt. if w_eq0 ah then let (q,r) := w_div_gt al bl in (WW w_0 q, w_0W r) - else + else match w_compare w_0 bh with - | Eq => + | Eq => let(q,r):= - double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1 a bl in (q, w_0W r) | Lt => ww_div_gt_aux ah al bh bl @@ -1147,7 +1147,7 @@ Section DoubleDivGt. Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl. assert (H:=@spec_eq0 ah);destruct (w_eq0 ah). simpl in Hgt;rewrite H in Hgt;trivial. - assert ([|bh|] <= 0). + assert ([|bh|] <= 0). apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith. assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt. simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos. @@ -1155,7 +1155,7 @@ Section DoubleDivGt. destruct (w_div_gt al bl);simpl;rewrite spec_w_0W;trivial. clear H. assert (H2 := spec_compare w_0 bh);destruct (w_compare w_0 bh). - rewrite (@spec_double_modn1_aux w w_zdigits w_0 w_WW w_head0 w_add_mul_div + rewrite (@spec_double_modn1_aux w w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub w_to_Z spec_w_0 spec_compare 1 (WW ah al) bl). destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1 (WW ah al) bl);simpl;trivial. @@ -1174,7 +1174,7 @@ Section DoubleDivGt. rewrite Zmult_comm;trivial. Qed. - Lemma Zis_gcd_mod : forall a b d, + Lemma Zis_gcd_mod : forall a b d, 0 < b -> Zis_gcd b (a mod b) d -> Zis_gcd a b d. Proof. intros a b d H H1; apply Zis_gcd_for_euclid with (a/b). @@ -1182,12 +1182,12 @@ Section DoubleDivGt. ring_simplify (b * (a / b) + a mod b - a / b * b);trivial. zarith. Qed. - Lemma spec_ww_gcd_gt_aux_body : + Lemma spec_ww_gcd_gt_aux_body : forall ah al bh bl n cont, - [[WW bh bl]] <= 2^n -> + [[WW bh bl]] <= 2^n -> [[WW ah al]] > [[WW bh bl]] -> - (forall xh xl yh yl, - [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^(n-1) -> + (forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^(n-1) -> Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) -> Zis_gcd [[WW ah al]] [[WW bh bl]] [[ww_gcd_gt_body cont ah al bh bl]]. Proof. @@ -1196,7 +1196,7 @@ Section DoubleDivGt. | Eq => match w_compare w_0 bl with | Eq => WW ah al (* normalement n'arrive pas si forme normale *) - | Lt => + | Lt => let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 w_compare w_sub 1 (WW ah al) bl in WW w_0 (w_gcd_gt bl m) @@ -1205,14 +1205,14 @@ Section DoubleDivGt. | Lt => let m := ww_mod_gt_aux ah al bh bl in match m with - | W0 => WW bh bl + | W0 => WW bh bl | WW mh ml => match w_compare w_0 mh with | Eq => match w_compare w_0 ml with | Eq => WW bh bl - | _ => - let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + | _ => + let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 w_compare w_sub 1 (WW bh bl) ml in WW w_0 (w_gcd_gt ml r) end @@ -1227,10 +1227,10 @@ Section DoubleDivGt. end | Gt => W0 (* absurde *) end). - assert (Hbh := spec_compare w_0 bh);destruct (w_compare w_0 bh). + assert (Hbh := spec_compare w_0 bh);destruct (w_compare w_0 bh). simpl ww_to_Z in *. rewrite spec_w_0 in Hbh;rewrite <- Hbh; rewrite Zmult_0_l;rewrite Zplus_0_l. - assert (Hbl := spec_compare w_0 bl); destruct (w_compare w_0 bl). + assert (Hbl := spec_compare w_0 bl); destruct (w_compare w_0 bl). rewrite spec_w_0 in Hbl;rewrite <- Hbl;apply Zis_gcd_0. simpl;rewrite spec_w_0;rewrite Zmult_0_l;rewrite Zplus_0_l. rewrite spec_w_0 in Hbl. @@ -1239,54 +1239,54 @@ Section DoubleDivGt. rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 spec_add_mul_div spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hbl). - apply spec_gcd_gt. - rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial. - apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + apply spec_gcd_gt. + rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial. + apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => destruct (Z_mod_lt x y);zarith end. - rewrite spec_w_0 in Hbl;Spec_w_to_Z bl;elimtype False;omega. + rewrite spec_w_0 in Hbl;Spec_w_to_Z bl;exfalso;omega. rewrite spec_w_0 in Hbh;assert (H:= spec_ww_mod_gt_aux _ _ _ Hgt Hbh). - assert (H2 : 0 < [[WW bh bl]]). + assert (H2 : 0 < [[WW bh bl]]). simpl;Spec_w_to_Z bl. apply Zlt_le_trans with ([|bh|]*wB);zarith. apply Zmult_lt_0_compat;zarith. apply Zis_gcd_mod;trivial. rewrite <- H. simpl in *;destruct (ww_mod_gt_aux ah al bh bl) as [ |mh ml]. - simpl;apply Zis_gcd_0;zarith. - assert (Hmh := spec_compare w_0 mh);destruct (w_compare w_0 mh). + simpl;apply Zis_gcd_0;zarith. + assert (Hmh := spec_compare w_0 mh);destruct (w_compare w_0 mh). simpl;rewrite spec_w_0 in Hmh; rewrite <- Hmh;simpl. - assert (Hml := spec_compare w_0 ml);destruct (w_compare w_0 ml). + assert (Hml := spec_compare w_0 ml);destruct (w_compare w_0 ml). rewrite <- Hml;rewrite spec_w_0;simpl;apply Zis_gcd_0. - simpl;rewrite spec_w_0;simpl. + simpl;rewrite spec_w_0;simpl. rewrite spec_w_0 in Hml. apply Zis_gcd_mod;zarith. change ([|bh|] * wB + [|bl|]) with (double_to_Z w_digits w_to_Z 1 (WW bh bl)). rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 spec_add_mul_div spec_div21 spec_compare spec_sub 1 (WW bh bl) ml Hml). - apply spec_gcd_gt. - rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial. - apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + apply spec_gcd_gt. + rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial. + apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => destruct (Z_mod_lt x y);zarith end. - rewrite spec_w_0 in Hml;Spec_w_to_Z ml;elimtype False;omega. + rewrite spec_w_0 in Hml;Spec_w_to_Z ml;exfalso;omega. rewrite spec_w_0 in Hmh. assert ([[WW bh bl]] > [[WW mh ml]]). - rewrite H;simpl; apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + rewrite H;simpl; apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => destruct (Z_mod_lt x y);zarith end. assert (H1:= spec_ww_mod_gt_aux _ _ _ H0 Hmh). - assert (H3 : 0 < [[WW mh ml]]). + assert (H3 : 0 < [[WW mh ml]]). simpl;Spec_w_to_Z ml. apply Zlt_le_trans with ([|mh|]*wB);zarith. apply Zmult_lt_0_compat;zarith. apply Zis_gcd_mod;zarith. simpl in *;rewrite <- H1. destruct (ww_mod_gt_aux bh bl mh ml) as [ |rh rl]. simpl; apply Zis_gcd_0. simpl;apply Hcont. simpl in H1;rewrite H1. - apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => destruct (Z_mod_lt x y);zarith end. - apply Zle_trans with (2^n/2). - apply Zdiv_le_lower_bound;zarith. + apply Zle_trans with (2^n/2). + apply Zdiv_le_lower_bound;zarith. apply Zle_trans with ([|bh|] * wB + [|bl|]);zarith. assert (H3' := Z_div_mod_eq [[WW bh bl]] [[WW mh ml]] (Zlt_gt _ _ H3)). assert (H4' : 0 <= [[WW bh bl]]/[[WW mh ml]]). apply Zge_le;apply Z_div_ge0;zarith. simpl in *;rewrite H1. pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3'. destruct (Zle_lt_or_eq _ _ H4'). - assert (H6' : [[WW bh bl]] mod [[WW mh ml]] = + assert (H6' : [[WW bh bl]] mod [[WW mh ml]] = [[WW bh bl]] - [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])). simpl;pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3';ring. simpl in H6'. assert ([[WW mh ml]] <= [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])). @@ -1300,14 +1300,14 @@ Section DoubleDivGt. rewrite Z_div_mult;zarith. assert (2^1 <= 2^n). change (2^1) with 2;zarith. assert (H7 := @Zpower_le_monotone_inv 2 1 n);zarith. - rewrite spec_w_0 in Hmh;Spec_w_to_Z mh;elimtype False;zarith. - rewrite spec_w_0 in Hbh;Spec_w_to_Z bh;elimtype False;zarith. + rewrite spec_w_0 in Hmh;Spec_w_to_Z mh;exfalso;zarith. + rewrite spec_w_0 in Hbh;Spec_w_to_Z bh;exfalso;zarith. Qed. - Lemma spec_ww_gcd_gt_aux : + Lemma spec_ww_gcd_gt_aux : forall p cont n, - (forall xh xl yh yl, - [[WW xh xl]] > [[WW yh yl]] -> + (forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^n -> Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) -> forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] -> @@ -1334,7 +1334,7 @@ Section DoubleDivGt. apply Zle_trans with (2 ^ (Zpos p + n -1));zarith. apply Zpower_le_monotone2;zarith. apply Zle_trans with (2 ^ (2*Zpos p + n -1));zarith. - apply Zpower_le_monotone2;zarith. + apply Zpower_le_monotone2;zarith. apply spec_ww_gcd_gt_aux_body with (n := n+1);trivial. rewrite Zplus_comm;trivial. ring_simplify (n + 1 - 1);trivial. @@ -1352,16 +1352,16 @@ Section DoubleDiv. Variable ww_div_gt : zn2z w -> zn2z w -> zn2z w * zn2z w. Variable ww_mod_gt : zn2z w -> zn2z w -> zn2z w. - Definition ww_div a b := - match ww_compare a b with - | Gt => ww_div_gt a b + Definition ww_div a b := + match ww_compare a b with + | Gt => ww_div_gt a b | Eq => (ww_1, W0) | Lt => (W0, a) end. - Definition ww_mod a b := - match ww_compare a b with - | Gt => ww_mod_gt a b + Definition ww_mod a b := + match ww_compare a b with + | Gt => ww_mod_gt a b | Eq => W0 | Lt => a end. @@ -1401,7 +1401,7 @@ Section DoubleDiv. Proof. intros a b Hpos;unfold ww_div. assert (H:=spec_ww_compare a b);destruct (ww_compare a b). - simpl;rewrite spec_ww_1;split;zarith. + simpl;rewrite spec_ww_1;split;zarith. simpl;split;[ring|Spec_ww_to_Z a;zarith]. apply spec_ww_div_gt;trivial. Qed. @@ -1409,7 +1409,7 @@ Section DoubleDiv. Lemma spec_ww_mod : forall a b, 0 < [[b]] -> [[ww_mod a b]] = [[a]] mod [[b]]. Proof. - intros a b Hpos;unfold ww_mod. + intros a b Hpos;unfold ww_mod. assert (H := spec_ww_compare a b);destruct (ww_compare a b). simpl;apply Zmod_unique with 1;try rewrite H;zarith. Spec_ww_to_Z a;symmetry;apply Zmod_small;zarith. @@ -1424,8 +1424,8 @@ Section DoubleDiv. Variable w_gcd_gt : w -> w -> w. Variable _ww_digits : positive. Variable spec_ww_digits_ : _ww_digits = xO w_digits. - Variable ww_gcd_gt_fix : - positive -> (w -> w -> w -> w -> zn2z w) -> + Variable ww_gcd_gt_fix : + positive -> (w -> w -> w -> w -> zn2z w) -> w -> w -> w -> w -> zn2z w. Variable spec_w_0 : [|w_0|] = 0. @@ -1440,10 +1440,10 @@ Section DoubleDiv. Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0. Variable spec_gcd_gt : forall a b, [|a|] > [|b|] -> Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|]. - Variable spec_gcd_gt_fix : + Variable spec_gcd_gt_fix : forall p cont n, - (forall xh xl yh yl, - [[WW xh xl]] > [[WW yh yl]] -> + (forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^n -> Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) -> forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] -> @@ -1451,20 +1451,20 @@ Section DoubleDiv. Zis_gcd [[WW ah al]] [[WW bh bl]] [[ww_gcd_gt_fix p cont ah al bh bl]]. - Definition gcd_cont (xh xl yh yl:w) := + Definition gcd_cont (xh xl yh yl:w) := match w_compare w_1 yl with - | Eq => ww_1 + | Eq => ww_1 | _ => WW xh xl end. - Lemma spec_gcd_cont : forall xh xl yh yl, - [[WW xh xl]] > [[WW yh yl]] -> + Lemma spec_gcd_cont : forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 1 -> Zis_gcd [[WW xh xl]] [[WW yh yl]] [[gcd_cont xh xl yh yl]]. Proof. intros xh xl yh yl Hgt' Hle. simpl in Hle. assert ([|yh|] = 0). - change 1 with (0*wB+1) in Hle. + change 1 with (0*wB+1) in Hle. assert (0 <= 1 < wB). split;zarith. apply wB_pos. assert (H1:= beta_lex _ _ _ _ _ Hle (spec_to_Z yl) H). Spec_w_to_Z yh;zarith. @@ -1473,20 +1473,20 @@ Section DoubleDiv. simpl;rewrite H;simpl;destruct (w_compare w_1 yl). rewrite spec_ww_1;rewrite <- Hcmpy;apply Zis_gcd_mod;zarith. rewrite <- (Zmod_unique ([|xh|]*wB+[|xl|]) 1 ([|xh|]*wB+[|xl|]) 0);zarith. - rewrite H in Hle; elimtype False;zarith. + rewrite H in Hle; exfalso;zarith. assert ([|yl|] = 0). Spec_w_to_Z yl;zarith. rewrite H0;simpl;apply Zis_gcd_0;trivial. Qed. - + Variable cont : w -> w -> w -> w -> zn2z w. - Variable spec_cont : forall xh xl yh yl, - [[WW xh xl]] > [[WW yh yl]] -> + Variable spec_cont : forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 1 -> Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]. - - Definition ww_gcd_gt a b := - match a, b with + + Definition ww_gcd_gt a b := + match a, b with | W0, _ => b | _, W0 => a | WW ah al, WW bh bl => @@ -1509,8 +1509,8 @@ Section DoubleDiv. destruct a as [ |ah al]. simpl;apply Zis_gcd_sym;apply Zis_gcd_0. destruct b as [ |bh bl]. simpl;apply Zis_gcd_0. simpl in Hgt. generalize (@spec_eq0 ah);destruct (w_eq0 ah);intros. - simpl;rewrite H in Hgt;trivial;rewrite H;trivial;rewrite spec_w_0;simpl. - assert ([|bh|] <= 0). + simpl;rewrite H in Hgt;trivial;rewrite H;trivial;rewrite spec_w_0;simpl. + assert ([|bh|] <= 0). apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith. Spec_w_to_Z bh;assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt. rewrite H1;simpl;auto. clear H. @@ -1522,7 +1522,7 @@ Section DoubleDiv. Lemma spec_ww_gcd : forall a b, Zis_gcd [[a]] [[b]] [[ww_gcd a b]]. Proof. intros a b. - change (ww_gcd a b) with + change (ww_gcd a b) with (match ww_compare a b with | Gt => ww_gcd_gt a b | Eq => a diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v index d6f6a05f..386bbb9e 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleDivn1.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -17,7 +17,7 @@ Require Import BigNumPrelude. Require Import DoubleType. Require Import DoubleBase. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section GENDIVN1. @@ -31,19 +31,19 @@ Section GENDIVN1. Variable w_div21 : w -> w -> w -> w * w. Variable w_compare : w -> w -> comparison. Variable w_sub : w -> w -> w. - - + + (* ** For proofs ** *) Variable w_to_Z : w -> Z. - - Notation wB := (base w_digits). + + Notation wB := (base w_digits). Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). - Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x) + Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x) (at level 0, x at level 99). Notation "[[ x ]]" := (zn2z_to_Z wB w_to_Z x) (at level 0, x at level 99). - + Variable spec_to_Z : forall x, 0 <= [| x |] < wB. Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits. Variable spec_0 : [|w_0|] = 0. @@ -68,10 +68,10 @@ Section GENDIVN1. | Lt => [|x|] < [|y|] | Gt => [|x|] > [|y|] end. - Variable spec_sub: forall x y, + Variable spec_sub: forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. - + Section DIVAUX. Variable b2p : w. @@ -85,10 +85,10 @@ Section GENDIVN1. Fixpoint double_divn1_0 (n:nat) : w -> word w n -> word w n * w := match n return w -> word w n -> word w n * w with - | O => fun r x => w_div21 r x b2p - | S n => double_divn1_0_aux n (double_divn1_0 n) + | O => fun r x => w_div21 r x b2p + | S n => double_divn1_0_aux n (double_divn1_0 n) end. - + Lemma spec_split : forall (n : nat) (x : zn2z (word w n)), let (h, l) := double_split w_0 n x in [!S n | x!] = [!n | h!] * double_wB w_digits n + [!n | l!]. @@ -132,11 +132,11 @@ Section GENDIVN1. induction n;simpl;intros;trivial. unfold double_modn1_0_aux, double_divn1_0_aux. destruct (double_split w_0 n x) as (hh,hl). - rewrite (IHn r hh). + rewrite (IHn r hh). destruct (double_divn1_0 n r hh) as (qh,rh);simpl. rewrite IHn. destruct (double_divn1_0 n rh hl);trivial. Qed. - + Variable p : w. Variable p_bounded : [|p|] <= Zpos w_digits. @@ -148,18 +148,18 @@ Section GENDIVN1. intros;apply spec_add_mul_div;auto. Qed. - Definition double_divn1_p_aux n - (divn1 : w -> word w n -> word w n -> word w n * w) r h l := + Definition double_divn1_p_aux n + (divn1 : w -> word w n -> word w n -> word w n * w) r h l := let (hh,hl) := double_split w_0 n h in - let (lh,ll) := double_split w_0 n l in + let (lh,ll) := double_split w_0 n l in let (qh,rh) := divn1 r hh hl in let (ql,rl) := divn1 rh hl lh in (double_WW w_WW n qh ql, rl). Fixpoint double_divn1_p (n:nat) : w -> word w n -> word w n -> word w n * w := match n return w -> word w n -> word w n -> word w n * w with - | O => fun r h l => w_div21 r (w_add_mul_div p h l) b2p - | S n => double_divn1_p_aux n (double_divn1_p n) + | O => fun r h l => w_div21 r (w_add_mul_div p h l) b2p + | S n => double_divn1_p_aux n (double_divn1_p n) end. Lemma p_lt_double_digits : forall n, [|p|] <= Zpos (double_digits w_digits n). @@ -175,8 +175,8 @@ Section GENDIVN1. Lemma spec_double_divn1_p : forall n r h l, [|r|] < [|b2p|] -> let (q,r') := double_divn1_p n r h l in - [|r|] * double_wB w_digits n + - ([!n|h!]*2^[|p|] + + [|r|] * double_wB w_digits n + + ([!n|h!]*2^[|p|] + [!n|l!] / (2^(Zpos(double_digits w_digits n) - [|p|]))) mod double_wB w_digits n = [!n|q!] * [|b2p|] + [|r'|] /\ 0 <= [|r'|] < [|b2p|]. @@ -198,26 +198,26 @@ Section GENDIVN1. ([!n|lh!] * double_wB w_digits n + [!n|ll!]) / 2^(Zpos (double_digits w_digits (S n)) - [|p|])) mod (double_wB w_digits n * double_wB w_digits n)) with - (([|r|] * double_wB w_digits n + ([!n|hh!] * 2^[|p|] + + (([|r|] * double_wB w_digits n + ([!n|hh!] * 2^[|p|] + [!n|hl!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod double_wB w_digits n) * double_wB w_digits n + - ([!n|hl!] * 2^[|p|] + - [!n|lh!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod + ([!n|hl!] * 2^[|p|] + + [!n|lh!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod double_wB w_digits n). generalize (IHn r hh hl H);destruct (double_divn1_p n r hh hl) as (qh,rh); intros (H3,H4);rewrite H3. - assert ([|rh|] < [|b2p|]). omega. + assert ([|rh|] < [|b2p|]). omega. replace (([!n|qh!] * [|b2p|] + [|rh|]) * double_wB w_digits n + ([!n|hl!] * 2 ^ [|p|] + [!n|lh!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])) mod - double_wB w_digits n) with + double_wB w_digits n) with ([!n|qh!] * [|b2p|] *double_wB w_digits n + ([|rh|]*double_wB w_digits n + ([!n|hl!] * 2 ^ [|p|] + [!n|lh!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])) mod double_wB w_digits n)). 2:ring. generalize (IHn rh hl lh H0);destruct (double_divn1_p n rh hl lh) as (ql,rl); intros (H5,H6);rewrite H5. - split;[rewrite spec_double_WW;trivial;ring|trivial]. + split;[rewrite spec_double_WW;trivial;ring|trivial]. assert (Uhh := spec_double_to_Z w_digits w_to_Z spec_to_Z n hh); unfold double_wB,base in Uhh. assert (Uhl := spec_double_to_Z w_digits w_to_Z spec_to_Z n hl); @@ -228,37 +228,37 @@ Section GENDIVN1. unfold double_wB,base in Ull. unfold double_wB,base. assert (UU:=p_lt_double_digits n). - rewrite Zdiv_shift_r;auto with zarith. - 2:change (Zpos (double_digits w_digits (S n))) + rewrite Zdiv_shift_r;auto with zarith. + 2:change (Zpos (double_digits w_digits (S n))) with (2*Zpos (double_digits w_digits n));auto with zarith. replace (2 ^ (Zpos (double_digits w_digits (S n)) - [|p|])) with (2^(Zpos (double_digits w_digits n) - [|p|])*2^Zpos (double_digits w_digits n)). rewrite Zdiv_mult_cancel_r;auto with zarith. - rewrite Zmult_plus_distr_l with (p:= 2^[|p|]). + rewrite Zmult_plus_distr_l with (p:= 2^[|p|]). pattern ([!n|hl!] * 2^[|p|]) at 2; rewrite (shift_unshift_mod (Zpos(double_digits w_digits n))([|p|])([!n|hl!])); auto with zarith. - rewrite Zplus_assoc. - replace + rewrite Zplus_assoc. + replace ([!n|hh!] * 2^Zpos (double_digits w_digits n)* 2^[|p|] + ([!n|hl!] / 2^(Zpos (double_digits w_digits n)-[|p|])* 2^Zpos(double_digits w_digits n))) - with - (([!n|hh!] *2^[|p|] + double_to_Z w_digits w_to_Z n hl / + with + (([!n|hh!] *2^[|p|] + double_to_Z w_digits w_to_Z n hl / 2^(Zpos (double_digits w_digits n)-[|p|])) * 2^Zpos(double_digits w_digits n));try (ring;fail). rewrite <- Zplus_assoc. rewrite <- (Zmod_shift_r ([|p|]));auto with zarith. - replace + replace (2 ^ Zpos (double_digits w_digits n) * 2 ^ Zpos (double_digits w_digits n)) with (2 ^ (Zpos (double_digits w_digits n) + Zpos (double_digits w_digits n))). rewrite (Zmod_shift_r (Zpos (double_digits w_digits n)));auto with zarith. replace (2 ^ (Zpos (double_digits w_digits n) + Zpos (double_digits w_digits n))) - with (2^Zpos(double_digits w_digits n) *2^Zpos(double_digits w_digits n)). + with (2^Zpos(double_digits w_digits n) *2^Zpos(double_digits w_digits n)). rewrite (Zmult_comm (([!n|hh!] * 2 ^ [|p|] + [!n|hl!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])))). rewrite Zmult_mod_distr_l;auto with zarith. - ring. + ring. rewrite Zpower_exp;auto with zarith. assert (0 < Zpos (double_digits w_digits n)). unfold Zlt;reflexivity. auto with zarith. @@ -267,24 +267,24 @@ Section GENDIVN1. split;auto with zarith. apply Zdiv_lt_upper_bound;auto with zarith. rewrite <- Zpower_exp;auto with zarith. - replace ([|p|] + (Zpos (double_digits w_digits n) - [|p|])) with + replace ([|p|] + (Zpos (double_digits w_digits n) - [|p|])) with (Zpos(double_digits w_digits n));auto with zarith. rewrite <- Zpower_exp;auto with zarith. - replace (Zpos (double_digits w_digits (S n)) - [|p|]) with - (Zpos (double_digits w_digits n) - [|p|] + + replace (Zpos (double_digits w_digits (S n)) - [|p|]) with + (Zpos (double_digits w_digits n) - [|p|] + Zpos (double_digits w_digits n));trivial. - change (Zpos (double_digits w_digits (S n))) with + change (Zpos (double_digits w_digits (S n))) with (2*Zpos (double_digits w_digits n)). ring. Qed. Definition double_modn1_p_aux n (modn1 : w -> word w n -> word w n -> w) r h l:= let (hh,hl) := double_split w_0 n h in - let (lh,ll) := double_split w_0 n l in + let (lh,ll) := double_split w_0 n l in modn1 (modn1 r hh hl) hl lh. Fixpoint double_modn1_p (n:nat) : w -> word w n -> word w n -> w := match n return w -> word w n -> word w n -> w with - | O => fun r h l => snd (w_div21 r (w_add_mul_div p h l) b2p) + | O => fun r h l => snd (w_div21 r (w_add_mul_div p h l) b2p) | S n => double_modn1_p_aux n (double_modn1_p n) end. @@ -302,8 +302,8 @@ Section GENDIVN1. Fixpoint high (n:nat) : word w n -> w := match n return word w n -> w with - | O => fun a => a - | S n => + | O => fun a => a + | S n => fun (a:zn2z (word w n)) => match a with | W0 => w_0 @@ -314,20 +314,20 @@ Section GENDIVN1. Lemma spec_double_digits:forall n, Zpos w_digits <= Zpos (double_digits w_digits n). Proof. induction n;simpl;auto with zarith. - change (Zpos (xO (double_digits w_digits n))) with + change (Zpos (xO (double_digits w_digits n))) with (2*Zpos (double_digits w_digits n)). assert (0 < Zpos w_digits);auto with zarith. exact (refl_equal Lt). Qed. - Lemma spec_high : forall n (x:word w n), + Lemma spec_high : forall n (x:word w n), [|high n x|] = [!n|x!] / 2^(Zpos (double_digits w_digits n) - Zpos w_digits). Proof. induction n;intros. unfold high,double_digits,double_to_Z. replace (Zpos w_digits - Zpos w_digits) with 0;try ring. simpl. rewrite <- (Zdiv_unique [|x|] 1 [|x|] 0);auto with zarith. - assert (U2 := spec_double_digits n). + assert (U2 := spec_double_digits n). assert (U3 : 0 < Zpos w_digits). exact (refl_equal Lt). destruct x;unfold high;fold high. unfold double_to_Z,zn2z_to_Z;rewrite spec_0. @@ -337,31 +337,31 @@ Section GENDIVN1. simpl [!S n|WW w0 w1!]. unfold double_wB,base;rewrite Zdiv_shift_r;auto with zarith. replace (2 ^ (Zpos (double_digits w_digits (S n)) - Zpos w_digits)) with - (2^(Zpos (double_digits w_digits n) - Zpos w_digits) * + (2^(Zpos (double_digits w_digits n) - Zpos w_digits) * 2^Zpos (double_digits w_digits n)). rewrite Zdiv_mult_cancel_r;auto with zarith. rewrite <- Zpower_exp;auto with zarith. - replace (Zpos (double_digits w_digits n) - Zpos w_digits + + replace (Zpos (double_digits w_digits n) - Zpos w_digits + Zpos (double_digits w_digits n)) with (Zpos (double_digits w_digits (S n)) - Zpos w_digits);trivial. - change (Zpos (double_digits w_digits (S n))) with + change (Zpos (double_digits w_digits (S n))) with (2*Zpos (double_digits w_digits n));ring. - change (Zpos (double_digits w_digits (S n))) with + change (Zpos (double_digits w_digits (S n))) with (2*Zpos (double_digits w_digits n)); auto with zarith. Qed. - - Definition double_divn1 (n:nat) (a:word w n) (b:w) := + + Definition double_divn1 (n:nat) (a:word w n) (b:w) := let p := w_head0 b in match w_compare p w_0 with | Gt => let b2p := w_add_mul_div p b w_0 in let ha := high n a in let k := w_sub w_zdigits p in - let lsr_n := w_add_mul_div k w_0 in + let lsr_n := w_add_mul_div k w_0 in let r0 := w_add_mul_div p w_0 ha in let (q,r) := double_divn1_p b2p p n r0 a (double_0 w_0 n) in (q, lsr_n r) - | _ => double_divn1_0 b n w_0 a + | _ => double_divn1_0 b n w_0 a end. Lemma spec_double_divn1 : forall n a b, @@ -392,21 +392,21 @@ Section GENDIVN1. apply Zmult_le_compat;auto with zarith. assert (wB <= 2^[|w_head0 b|]). unfold base;apply Zpower_le_monotone;auto with zarith. omega. - assert ([|w_add_mul_div (w_head0 b) b w_0|] = + assert ([|w_add_mul_div (w_head0 b) b w_0|] = 2 ^ [|w_head0 b|] * [|b|]). rewrite (spec_add_mul_div b w_0); auto with zarith. rewrite spec_0;rewrite Zdiv_0_l; try omega. rewrite Zplus_0_r; rewrite Zmult_comm. rewrite Zmod_small; auto with zarith. assert (H5 := spec_to_Z (high n a)). - assert + assert ([|w_add_mul_div (w_head0 b) w_0 (high n a)|] <[|w_add_mul_div (w_head0 b) b w_0|]). rewrite H4. rewrite spec_add_mul_div;auto with zarith. rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l. assert (([|high n a|]/2^(Zpos w_digits - [|w_head0 b|])) < wB). - apply Zdiv_lt_upper_bound;auto with zarith. + apply Zdiv_lt_upper_bound;auto with zarith. apply Zlt_le_trans with wB;auto with zarith. pattern wB at 1;replace wB with (wB*1);try ring. apply Zmult_le_compat;auto with zarith. @@ -420,8 +420,8 @@ Section GENDIVN1. apply Zmult_le_compat;auto with zarith. pattern 2 at 1;rewrite <- Zpower_1_r. apply Zpower_le_monotone;split;auto with zarith. - rewrite <- H4 in H0. - assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith. + rewrite <- H4 in H0. + assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith. assert (H7:= spec_double_divn1_p H0 Hb3 n a (double_0 w_0 n) H6). destruct (double_divn1_p (w_add_mul_div (w_head0 b) b w_0) (w_head0 b) n (w_add_mul_div (w_head0 b) w_0 (high n a)) a @@ -436,7 +436,7 @@ Section GENDIVN1. rewrite Zmod_small;auto with zarith. rewrite spec_high. rewrite Zdiv_Zdiv;auto with zarith. rewrite <- Zpower_exp;auto with zarith. - replace (Zpos (double_digits w_digits n) - Zpos w_digits + + replace (Zpos (double_digits w_digits n) - Zpos w_digits + (Zpos w_digits - [|w_head0 b|])) with (Zpos (double_digits w_digits n) - [|w_head0 b|]);trivial;ring. assert (H8 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith. @@ -448,11 +448,11 @@ Section GENDIVN1. rewrite H8 in H7;unfold double_wB,base in H7. rewrite <- shift_unshift_mod in H7;auto with zarith. rewrite H4 in H7. - assert ([|w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r|] + assert ([|w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r|] = [|r|]/2^[|w_head0 b|]). rewrite spec_add_mul_div. rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l. - replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|]) + replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|]) with ([|w_head0 b|]). rewrite Zmod_small;auto with zarith. assert (H9 := spec_to_Z r). @@ -474,11 +474,11 @@ Section GENDIVN1. split. rewrite <- (Z_div_mult [!n|a!] (2^[|w_head0 b|]));auto with zarith. rewrite H71;rewrite H9. - replace ([!n|q!] * (2 ^ [|w_head0 b|] * [|b|])) + replace ([!n|q!] * (2 ^ [|w_head0 b|] * [|b|])) with ([!n|q!] *[|b|] * 2^[|w_head0 b|]); try (ring;fail). rewrite Z_div_plus_l;auto with zarith. - assert (H10 := spec_to_Z + assert (H10 := spec_to_Z (w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r));split; auto with zarith. rewrite H9. @@ -487,19 +487,19 @@ Section GENDIVN1. exact (spec_double_to_Z w_digits w_to_Z spec_to_Z n a). Qed. - - Definition double_modn1 (n:nat) (a:word w n) (b:w) := + + Definition double_modn1 (n:nat) (a:word w n) (b:w) := let p := w_head0 b in match w_compare p w_0 with | Gt => let b2p := w_add_mul_div p b w_0 in let ha := high n a in let k := w_sub w_zdigits p in - let lsr_n := w_add_mul_div k w_0 in + let lsr_n := w_add_mul_div k w_0 in let r0 := w_add_mul_div p w_0 ha in let r := double_modn1_p b2p p n r0 a (double_0 w_0 n) in lsr_n r - | _ => double_modn1_0 b n w_0 a + | _ => double_modn1_0 b n w_0 a end. Lemma spec_double_modn1_aux : forall n a b, @@ -525,4 +525,4 @@ Section GENDIVN1. destruct H1 as (h1,h2);rewrite h1;ring. Qed. -End GENDIVN1. +End GENDIVN1. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v index 50c72487..21e694e5 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleLift.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -17,7 +17,7 @@ Require Import BigNumPrelude. Require Import DoubleType. Require Import DoubleBase. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section DoubleLift. Variable w : Type. @@ -61,13 +61,13 @@ Section DoubleLift. (* 0 < p < ww_digits *) - Definition ww_add_mul_div p x y := + Definition ww_add_mul_div p x y := let zdigits := w_0W w_zdigits in match x, y with | W0, W0 => W0 | W0, WW yh yl => match ww_compare p zdigits with - | Eq => w_0W yh + | Eq => w_0W yh | Lt => w_0W (w_add_mul_div (low p) w_0 yh) | Gt => let n := low (ww_sub p zdigits) in @@ -75,15 +75,15 @@ Section DoubleLift. end | WW xh xl, W0 => match ww_compare p zdigits with - | Eq => w_W0 xl + | Eq => w_W0 xl | Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl w_0) | Gt => let n := low (ww_sub p zdigits) in - w_W0 (w_add_mul_div n xl w_0) + w_W0 (w_add_mul_div n xl w_0) end | WW xh xl, WW yh yl => match ww_compare p zdigits with - | Eq => w_WW xl yh + | Eq => w_WW xl yh | Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl yh) | Gt => let n := low (ww_sub p zdigits) in @@ -93,7 +93,7 @@ Section DoubleLift. Section DoubleProof. Variable w_to_Z : w -> Z. - + Notation wB := (base w_digits). Notation wwB := (base (ww_digits w_digits)). Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). @@ -122,21 +122,21 @@ Section DoubleLift. Variable spec_w_head0 : forall x, 0 < [|x|] -> wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB. Variable spec_w_tail00 : forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits. - Variable spec_w_tail0 : forall x, 0 < [|x|] -> + Variable spec_w_tail0 : forall x, 0 < [|x|] -> exists y, 0 <= y /\ [|x|] = (2* y + 1) * (2 ^ [|w_tail0 x|]). Variable spec_w_add_mul_div : forall x y p, [|p|] <= Zpos w_digits -> [| w_add_mul_div p x y |] = ([|x|] * (2 ^ [|p|]) + [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB. - Variable spec_w_add: forall x y, + Variable spec_w_add: forall x y, [[w_add x y]] = [|x|] + [|y|]. - Variable spec_ww_sub: forall x y, + Variable spec_ww_sub: forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits. Variable spec_low: forall x, [| low x|] = [[x]] mod wB. - + Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos ww_Digits. Hint Resolve div_le_0 div_lt w_to_Z_wwB: lift. @@ -168,7 +168,7 @@ Section DoubleLift. rewrite spec_w_0; auto with zarith. rewrite spec_w_0; auto with zarith. Qed. - + Lemma spec_ww_head0 : forall x, 0 < [[x]] -> wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB. Proof. @@ -179,7 +179,7 @@ Section DoubleLift. assert (H0 := spec_compare w_0 xh);rewrite spec_w_0 in H0. destruct (w_compare w_0 xh). rewrite <- H0. simpl Zplus. rewrite <- H0 in H;simpl in H. - case (spec_to_Z w_zdigits); + case (spec_to_Z w_zdigits); case (spec_to_Z (w_head0 xl)); intros HH1 HH2 HH3 HH4. rewrite spec_w_add. rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith. @@ -209,7 +209,7 @@ Section DoubleLift. rewrite <- Zmult_assoc; apply Zmult_lt_compat_l; zarith. rewrite <- (Zplus_0_r (2^(Zpos w_digits - p)*wB));apply beta_lex_inv;zarith. apply Zmult_lt_reg_r with (2 ^ p); zarith. - rewrite <- Zpower_exp;zarith. + rewrite <- Zpower_exp;zarith. rewrite Zmult_comm;ring_simplify (Zpos w_digits - p + p);fold wB;zarith. assert (H1 := spec_to_Z xh);zarith. Qed. @@ -293,8 +293,8 @@ Section DoubleLift. Qed. Hint Rewrite Zdiv_0_l Zmult_0_l Zplus_0_l Zmult_0_r Zplus_0_r - spec_w_W0 spec_w_0W spec_w_WW spec_w_0 - (wB_div w_digits w_to_Z spec_to_Z) + spec_w_W0 spec_w_0W spec_w_WW spec_w_0 + (wB_div w_digits w_to_Z spec_to_Z) (wB_div_plus w_digits w_to_Z spec_to_Z) : w_rewrite. Ltac w_rewrite := autorewrite with w_rewrite;trivial. @@ -303,12 +303,12 @@ Section DoubleLift. [[p]] <= Zpos (xO w_digits) -> [[match ww_compare p zdigits with | Eq => w_WW xl yh - | Lt => w_WW (w_add_mul_div (low p) xh xl) + | Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl yh) | Gt => let n := low (ww_sub p zdigits) in w_WW (w_add_mul_div n xl yh) (w_add_mul_div n yh yl) - end]] = + end]] = ([[WW xh xl]] * (2^[[p]]) + [[WW yh yl]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB. Proof. @@ -317,7 +317,7 @@ Section DoubleLift. case (spec_to_w_Z p); intros Hv1 Hv2. replace (Zpos (xO w_digits)) with (Zpos w_digits + Zpos w_digits). 2 : rewrite Zpos_xO;ring. - replace (Zpos w_digits + Zpos w_digits - [[p]]) with + replace (Zpos w_digits + Zpos w_digits - [[p]]) with (Zpos w_digits + (Zpos w_digits - [[p]])). 2:ring. intros Hp; assert (Hxh := spec_to_Z xh);assert (Hxl:=spec_to_Z xl); assert (Hx := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)); @@ -330,7 +330,7 @@ Section DoubleLift. fold wB. rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;rewrite <- Zplus_assoc. rewrite <- Zpower_2. - rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|]. + rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|]. exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xl yh)). ring. simpl ww_to_Z; w_rewrite;zarith. assert (HH0: [|low p|] = [[p]]). @@ -353,7 +353,7 @@ Section DoubleLift. rewrite Zmult_plus_distr_l. pattern ([|xl|] * 2 ^ [[p]]) at 2; rewrite shift_unshift_mod with (n:= Zpos w_digits);fold wB;zarith. - replace ([|xh|] * wB * 2^[[p]]) with ([|xh|] * 2^[[p]] * wB). 2:ring. + replace ([|xh|] * wB * 2^[[p]]) with ([|xh|] * 2^[[p]] * wB). 2:ring. rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc. unfold base at 5;rewrite <- Zmod_shift_r;zarith. unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits)); @@ -387,8 +387,8 @@ Section DoubleLift. lazy zeta; simpl ww_to_Z; w_rewrite;zarith. repeat rewrite spec_w_add_mul_div;zarith. rewrite HH0. - pattern wB at 5;replace wB with - (2^(([[p]] - Zpos w_digits) + pattern wB at 5;replace wB with + (2^(([[p]] - Zpos w_digits) + (Zpos w_digits - ([[p]] - Zpos w_digits)))). rewrite Zpower_exp;zarith. rewrite Zmult_assoc. rewrite Z_div_plus_l;zarith. @@ -401,28 +401,28 @@ Section DoubleLift. repeat rewrite <- Zplus_assoc. unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits)); fold wB;fold wwB;zarith. - unfold base;rewrite Zmod_shift_r with (a:= Zpos w_digits) + unfold base;rewrite Zmod_shift_r with (a:= Zpos w_digits) (b:= Zpos w_digits);fold wB;fold wwB;zarith. rewrite wwB_wBwB; rewrite Zpower_2; rewrite Zmult_mod_distr_r;zarith. rewrite Zmult_plus_distr_l. - replace ([|xh|] * wB * 2 ^ u) with + replace ([|xh|] * wB * 2 ^ u) with ([|xh|]*2^u*wB). 2:ring. - repeat rewrite <- Zplus_assoc. + repeat rewrite <- Zplus_assoc. rewrite (Zplus_comm ([|xh|] * 2 ^ u * wB)). rewrite Z_mod_plus;zarith. rewrite Z_mod_mult;zarith. unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith. - unfold u; split;zarith. + unfold u; split;zarith. split;zarith. unfold u; apply Zdiv_lt_upper_bound;zarith. rewrite <- Zpower_exp;zarith. - fold u. - ring_simplify (u + (Zpos w_digits - u)); fold + fold u. + ring_simplify (u + (Zpos w_digits - u)); fold wB;zarith. unfold ww_digits;rewrite Zpos_xO;zarith. unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith. unfold u; split;zarith. unfold u; split;zarith. apply Zdiv_lt_upper_bound;zarith. rewrite <- Zpower_exp;zarith. - fold u. + fold u. ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith. unfold u;zarith. unfold u;zarith. @@ -446,7 +446,7 @@ Section DoubleLift. clear H1;w_rewrite);simpl ww_add_mul_div. replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial]. intros Heq;rewrite <- Heq;clear Heq; auto. - generalize (spec_ww_compare p (w_0W w_zdigits)); + generalize (spec_ww_compare p (w_0W w_zdigits)); case ww_compare; intros H1; w_rewrite. rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith. generalize H1; w_rewrite; rewrite spec_zdigits; clear H1; intros H1. @@ -459,7 +459,7 @@ Section DoubleLift. rewrite HH0; auto with zarith. replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial]. intros Heq;rewrite <- Heq;clear Heq. - generalize (spec_ww_compare p (w_0W w_zdigits)); + generalize (spec_ww_compare p (w_0W w_zdigits)); case ww_compare; intros H1; w_rewrite. rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith. rewrite Zpos_xO in H;zarith. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v index c7d83acc..7090c76a 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleMul.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -17,7 +17,7 @@ Require Import BigNumPrelude. Require Import DoubleType. Require Import DoubleBase. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section DoubleMul. Variable w : Type. @@ -45,7 +45,7 @@ Section DoubleMul. (* (xh*B+xl) (yh*B + yl) xh*yh = hh = |hhh|hhl|B2 xh*yl +xl*yh = cc = |cch|ccl|B - xl*yl = ll = |llh|lll + xl*yl = ll = |llh|lll *) Definition double_mul_c (cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w) x y := @@ -56,7 +56,7 @@ Section DoubleMul. let hh := w_mul_c xh yh in let ll := w_mul_c xl yl in let (wc,cc) := cross xh xl yh yl hh ll in - match cc with + match cc with | W0 => WW (ww_add hh (w_W0 wc)) ll | WW cch ccl => match ww_add_c (w_W0 ccl) ll with @@ -67,8 +67,8 @@ Section DoubleMul. end. Definition ww_mul_c := - double_mul_c - (fun xh xl yh yl hh ll=> + double_mul_c + (fun xh xl yh yl hh ll=> match ww_add_c (w_mul_c xh yl) (w_mul_c xl yh) with | C0 cc => (w_0, cc) | C1 cc => (w_1, cc) @@ -77,11 +77,11 @@ Section DoubleMul. Definition w_2 := w_add w_1 w_1. Definition kara_prod xh xl yh yl hh ll := - match ww_add_c hh ll with + match ww_add_c hh ll with C0 m => match w_compare xl xh with Eq => (w_0, m) - | Lt => + | Lt => match w_compare yl yh with Eq => (w_0, m) | Lt => (w_0, ww_sub m (w_mul_c (w_sub xh xl) (w_sub yh yl))) @@ -89,7 +89,7 @@ Section DoubleMul. C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1) end end - | Gt => + | Gt => match w_compare yl yh with Eq => (w_0, m) | Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with @@ -101,17 +101,17 @@ Section DoubleMul. | C1 m => match w_compare xl xh with Eq => (w_1, m) - | Lt => + | Lt => match w_compare yl yh with Eq => (w_1, m) | Lt => match ww_sub_c m (w_mul_c (w_sub xh xl) (w_sub yh yl)) with C0 m1 => (w_1, m1) | C1 m1 => (w_0, m1) - end + end | Gt => match ww_add_c m (w_mul_c (w_sub xh xl) (w_sub yl yh)) with C1 m1 => (w_2, m1) | C0 m1 => (w_1, m1) end end - | Gt => + | Gt => match w_compare yl yh with Eq => (w_1, m) | Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with @@ -129,8 +129,8 @@ Section DoubleMul. Definition ww_mul x y := match x, y with | W0, _ => W0 - | _, W0 => W0 - | WW xh xl, WW yh yl => + | _, W0 => W0 + | WW xh xl, WW yh yl => let ccl := w_add (w_mul xh yl) (w_mul xl yh) in ww_add (w_W0 ccl) (w_mul_c xl yl) end. @@ -161,9 +161,9 @@ Section DoubleMul. Variable w_mul_add : w -> w -> w -> w * w. Fixpoint double_mul_add_n1 (n:nat) : word w n -> w -> w -> w * word w n := - match n return word w n -> w -> w -> w * word w n with - | O => w_mul_add - | S n1 => + match n return word w n -> w -> w -> w * word w n with + | O => w_mul_add + | S n1 => let mul_add := double_mul_add_n1 n1 in fun x y r => match x with @@ -183,11 +183,11 @@ Section DoubleMul. Variable wn_0W : wn -> zn2z wn. Variable wn_WW : wn -> wn -> zn2z wn. Variable w_mul_add_n1 : wn -> w -> w -> w*wn. - Fixpoint double_mul_add_mn1 (m:nat) : + Fixpoint double_mul_add_mn1 (m:nat) : word wn m -> w -> w -> w*word wn m := - match m return word wn m -> w -> w -> w*word wn m with - | O => w_mul_add_n1 - | S m1 => + match m return word wn m -> w -> w -> w*word wn m with + | O => w_mul_add_n1 + | S m1 => let mul_add := double_mul_add_mn1 m1 in fun x y r => match x with @@ -207,11 +207,11 @@ Section DoubleMul. | WW h l => match w_add_c l r with | C0 lr => (h,lr) - | C1 lr => (w_succ h, lr) + | C1 lr => (w_succ h, lr) end end. - + (*Section DoubleProof. *) Variable w_digits : positive. Variable w_to_Z : w -> Z. @@ -225,11 +225,11 @@ Section DoubleMul. (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). - Notation "[+[ c ]]" := - (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). - Notation "[-[ c ]]" := - (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). Notation "[|| x ||]" := @@ -269,8 +269,8 @@ Section DoubleMul. forall x y, [[ww_add_carry x y]] = ([[x]] + [[y]] + 1) mod wwB. Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. - - + + Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB. Proof. intros x;apply spec_ww_to_Z;auto. Qed. @@ -281,21 +281,21 @@ Section DoubleMul. Ltac zarith := auto with zarith mult. Lemma wBwB_lex: forall a b c d, - a * wB^2 + [[b]] <= c * wB^2 + [[d]] -> + a * wB^2 + [[b]] <= c * wB^2 + [[d]] -> a <= c. - Proof. + Proof. intros a b c d H; apply beta_lex with [[b]] [[d]] (wB^2);zarith. Qed. - Lemma wBwB_lex_inv: forall a b c d, - a < c -> - a * wB^2 + [[b]] < c * wB^2 + [[d]]. + Lemma wBwB_lex_inv: forall a b c d, + a < c -> + a * wB^2 + [[b]] < c * wB^2 + [[d]]. Proof. intros a b c d H; apply beta_lex_inv; zarith. Qed. Lemma sum_mul_carry : forall xh xl yh yl wc cc, - [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] -> + [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] -> 0 <= [|wc|] <= 1. Proof. intros. @@ -303,14 +303,14 @@ Section DoubleMul. apply wB_pos. Qed. - Theorem mult_add_ineq: forall xH yH crossH, + Theorem mult_add_ineq: forall xH yH crossH, 0 <= [|xH|] * [|yH|] + [|crossH|] < wwB. Proof. intros;rewrite wwB_wBwB;apply mult_add_ineq;zarith. Qed. - + Hint Resolve mult_add_ineq : mult. - + Lemma spec_mul_aux : forall xh xl yh yl wc (cc:zn2z w) hh ll, [[hh]] = [|xh|] * [|yh|] -> [[ll]] = [|xl|] * [|yl|] -> @@ -325,9 +325,9 @@ Section DoubleMul. end||] = ([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|]). Proof. intros;assert (U1 := wB_pos w_digits). - replace (([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|])) with + replace (([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|])) with ([|xh|]*[|yh|]*wB^2 + ([|xh|]*[|yl|] + [|xl|]*[|yh|])*wB + [|xl|]*[|yl|]). - 2:ring. rewrite <- H1;rewrite <- H;rewrite <- H0. + 2:ring. rewrite <- H1;rewrite <- H;rewrite <- H0. assert (H2 := sum_mul_carry _ _ _ _ _ _ H1). destruct cc as [ | cch ccl]; simpl zn2z_to_Z; simpl ww_to_Z. rewrite spec_ww_add;rewrite spec_w_W0;rewrite Zmod_small; @@ -346,7 +346,7 @@ Section DoubleMul. rewrite <- Zmult_plus_distr_l. assert (((2 * wB - 4) + 2)*wB <= ([|wc|] * wB + [|cch|])*wB). apply Zmult_le_compat;zarith. - rewrite Zmult_plus_distr_l in H3. + rewrite Zmult_plus_distr_l in H3. intros. assert (U2 := spec_to_Z ccl);omega. generalize (spec_ww_add_c (w_W0 ccl) ll);destruct (ww_add_c (w_W0 ccl) ll) as [l|l];unfold interp_carry;rewrite spec_w_W0;try rewrite Zmult_1_l; @@ -363,8 +363,8 @@ Section DoubleMul. (forall xh xl yh yl hh ll, [[hh]] = [|xh|]*[|yh|] -> [[ll]] = [|xl|]*[|yl|] -> - let (wc,cc) := cross xh xl yh yl hh ll in - [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]) -> + let (wc,cc) := cross xh xl yh yl hh ll in + [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]) -> forall x y, [||double_mul_c cross x y||] = [[x]] * [[y]]. Proof. intros cross Hcross x y;destruct x as [ |xh xl];simpl;trivial. @@ -376,7 +376,7 @@ Section DoubleMul. rewrite <- wwB_wBwB;trivial. Qed. - Lemma spec_ww_mul_c : forall x y, [||ww_mul_c x y||] = [[x]] * [[y]]. + Lemma spec_ww_mul_c : forall x y, [||ww_mul_c x y||] = [[x]] * [[y]]. Proof. intros x y;unfold ww_mul_c;apply spec_double_mul_c. intros xh xl yh yl hh ll H1 H2. @@ -402,9 +402,9 @@ Section DoubleMul. let (wc,cc) := kara_prod xh xl yh yl hh ll in [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]. Proof. - intros xh xl yh yl hh ll H H0; rewrite <- kara_prod_aux; + intros xh xl yh yl hh ll H H0; rewrite <- kara_prod_aux; rewrite <- H; rewrite <- H0; unfold kara_prod. - assert (Hxh := (spec_to_Z xh)); assert (Hxl := (spec_to_Z xl)); + assert (Hxh := (spec_to_Z xh)); assert (Hxl := (spec_to_Z xl)); assert (Hyh := (spec_to_Z yh)); assert (Hyl := (spec_to_Z yl)). generalize (spec_ww_add_c hh ll); case (ww_add_c hh ll); intros z Hz; rewrite <- Hz; unfold interp_carry; assert (Hz1 := (spec_ww_to_Z z)). @@ -412,7 +412,7 @@ Section DoubleMul. try rewrite Hxlh; try rewrite spec_w_0; try (ring; fail). generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh. rewrite Hylh; rewrite spec_w_0; try (ring; fail). - rewrite spec_w_0; try (ring; fail). + rewrite spec_w_0; try (ring; fail). repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c). repeat rewrite Zmod_small; auto with zarith; try (ring; fail). split; auto with zarith. @@ -508,8 +508,8 @@ Section DoubleMul. repeat rewrite Zmod_small; auto with zarith; try (ring; fail). Qed. - Lemma sub_carry : forall xh xl yh yl z, - 0 <= z -> + Lemma sub_carry : forall xh xl yh yl z, + 0 <= z -> [|xh|]*[|yl|] + [|xl|]*[|yh|] = wwB + z -> z < wwB. Proof. @@ -519,7 +519,7 @@ Section DoubleMul. generalize (Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)). rewrite <- wwB_wBwB;intros H1 H2. assert (H3 := wB_pos w_digits). - assert (2*wB <= wwB). + assert (2*wB <= wwB). rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat;zarith. omega. Qed. @@ -528,7 +528,7 @@ Section DoubleMul. let H:= fresh "H" in assert (H:= spec_ww_to_Z x). - Ltac Zmult_lt_b x y := + Ltac Zmult_lt_b x y := let H := fresh "H" in assert (H := Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)). @@ -582,7 +582,7 @@ Section DoubleMul. Variable w_mul_add : w -> w -> w -> w * w. Variable spec_w_mul_add : forall x y r, let (h,l):= w_mul_add x y r in - [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|]. + [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|]. Lemma spec_double_mul_add_n1 : forall n x y r, let (h,l) := double_mul_add_n1 w_mul_add n x y r in @@ -590,7 +590,7 @@ Section DoubleMul. Proof. induction n;intros x y r;trivial. exact (spec_w_mul_add x y r). - unfold double_mul_add_n1;destruct x as[ |xh xl]; + unfold double_mul_add_n1;destruct x as[ |xh xl]; fold(double_mul_add_n1 w_mul_add). rewrite spec_w_0;rewrite spec_extend;simpl;trivial. assert(H:=IHn xl y r);destruct (double_mul_add_n1 w_mul_add n xl y r)as(rl,l). @@ -599,13 +599,13 @@ Section DoubleMul. rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc;rewrite <- H. rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. rewrite U;ring. - Qed. - + Qed. + End DoubleMulAddn1Proof. - Lemma spec_w_mul_add : forall x y r, + Lemma spec_w_mul_add : forall x y r, let (h,l):= w_mul_add x y r in - [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|]. + [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|]. Proof. intros x y r;unfold w_mul_add;assert (H:=spec_w_mul_c x y); destruct (w_mul_c x y) as [ |h l];simpl;rewrite <- H. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v index 043ff351..83a2e717 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleSqrt.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -17,7 +17,7 @@ Require Import BigNumPrelude. Require Import DoubleType. Require Import DoubleBase. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section DoubleSqrt. Variable w : Type. @@ -52,7 +52,7 @@ Section DoubleSqrt. Let wwBm1 := ww_Bm1 w_Bm1. - Definition ww_is_even x := + Definition ww_is_even x := match x with | W0 => true | WW xh xl => w_is_even xl @@ -62,7 +62,7 @@ Section DoubleSqrt. match w_compare x z with | Eq => match w_compare y z with - Eq => (C1 w_1, w_0) + Eq => (C1 w_1, w_0) | Gt => (C1 w_1, w_sub y z) | Lt => (C1 w_0, y) end @@ -120,7 +120,7 @@ Section DoubleSqrt. let ( q, r) := w_sqrt2 x1 x2 in let (q1, r1) := w_div2s r y1 q in match q1 with - C0 q1 => + C0 q1 => let q2 := w_square_c q1 in let a := WW q q1 in match r1 with @@ -132,9 +132,9 @@ Section DoubleSqrt. | C0 r2 => match ww_sub_c (WW r2 y2) q2 with C0 r3 => (a, C0 r3) - | C1 r3 => + | C1 r3 => let a2 := ww_add_mul_div (w_0W w_1) a W0 in - match ww_pred_c a2 with + match ww_pred_c a2 with C0 a3 => (ww_pred a, ww_add_c a3 r3) | C1 a3 => @@ -166,20 +166,20 @@ Section DoubleSqrt. | Gt => match ww_add_mul_div p x W0 with W0 => W0 - | WW x1 x2 => + | WW x1 x2 => let (r, _) := w_sqrt2 x1 x2 in - WW w_0 (w_add_mul_div - (w_sub w_zdigits + WW w_0 (w_add_mul_div + (w_sub w_zdigits (low (ww_add_mul_div (ww_pred ww_zdigits) W0 p))) w_0 r) end - | _ => + | _ => match x with W0 => W0 | WW x1 x2 => WW w_0 (fst (w_sqrt2 x1 x2)) end end. - + Variable w_to_Z : w -> Z. @@ -192,11 +192,11 @@ Section DoubleSqrt. (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). - Notation "[+[ c ]]" := - (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). - Notation "[-[ c ]]" := - (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). Notation "[|| x ||]" := @@ -269,14 +269,12 @@ Section DoubleSqrt. Let spec_ww_Bm1 : [[wwBm1]] = wwB - 1. Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed. - - Hint Rewrite spec_w_0 spec_w_1 w_Bm1 spec_w_WW spec_w_sub - spec_w_div21 spec_w_add_mul_div spec_ww_Bm1 - spec_w_add_c spec_w_sqrt2: w_rewrite. + Hint Rewrite spec_w_0 spec_w_1 spec_w_WW spec_w_sub + spec_w_add_mul_div spec_ww_Bm1 spec_w_add_c : w_rewrite. Lemma spec_ww_is_even : forall x, if ww_is_even x then [[x]] mod 2 = 0 else [[x]] mod 2 = 1. -clear spec_more_than_1_digit. +clear spec_more_than_1_digit. intros x; case x; simpl ww_is_even. simpl. rewrite Zmod_small; auto with zarith. @@ -379,8 +377,8 @@ intros x; case x; simpl ww_is_even. end. rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith. destruct (spec_to_Z w1) as [H1 H2];auto with zarith. - split; auto with zarith. - apply Zdiv_lt_upper_bound; auto with zarith. + split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. rewrite Hp; ring. Qed. @@ -402,7 +400,7 @@ intros x; case x; simpl ww_is_even. rewrite Zmax_right; auto with zarith. rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith. destruct (spec_to_Z w1) as [H1 H2];auto with zarith. - split; auto with zarith. + split; auto with zarith. unfold base. match goal with |- _ < _ ^ ?X => assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; @@ -434,7 +432,7 @@ intros x; case x; simpl ww_is_even. intros w1. rewrite spec_ww_add_mul_div; auto with zarith. autorewrite with w_rewrite rm10. - rewrite spec_w_0W; rewrite spec_w_1. + rewrite spec_w_0W; rewrite spec_w_1. rewrite Zpower_1_r; auto with zarith. rewrite Zmult_comm; auto. rewrite spec_w_0W; rewrite spec_w_1; auto with zarith. @@ -458,7 +456,7 @@ intros x; case x; simpl ww_is_even. match goal with |- 0 <= ?X - 1 => assert (0 < X); auto with zarith end. - apply Zpower_gt_0; auto with zarith. + apply Zpower_gt_0; auto with zarith. match goal with |- 0 <= ?X - 1 => assert (0 < X); auto with zarith; red; reflexivity end. @@ -542,7 +540,7 @@ intros x; case x; simpl ww_is_even. rewrite add_mult_div_2_plus_1; unfold base. match goal with |- context[_ ^ ?X] => assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; - rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith end. rewrite Zpos_minus; auto with zarith. @@ -559,7 +557,7 @@ intros x; case x; simpl ww_is_even. unfold base. match goal with |- context[_ ^ ?X] => assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; - rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith end. rewrite Zpos_minus; auto with zarith. @@ -592,7 +590,7 @@ intros x; case x; simpl ww_is_even. rewrite H1; unfold base. match goal with |- context[_ ^ ?X] => assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; - rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith end. rewrite Zpos_minus; auto with zarith. @@ -611,7 +609,7 @@ intros x; case x; simpl ww_is_even. rewrite H1; unfold base. match goal with |- context[_ ^ ?X] => assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; - rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith end. rewrite Zpos_minus; auto with zarith. @@ -682,7 +680,7 @@ intros x; case x; simpl ww_is_even. rewrite Zsquare_mult; replace (p * p) with ((- p) * (- p)); try ring. apply Zmult_le_0_compat; auto with zarith. Qed. - + Lemma spec_split: forall x, [|fst (split x)|] * wB + [|snd (split x)|] = [[x]]. intros x; case x; simpl; autorewrite with w_rewrite; @@ -751,7 +749,7 @@ intros x; case x; simpl ww_is_even. match goal with |- ?X <= ?Y => replace Y with (2 * (wB/ 2 - 1)); auto with zarith end. - pattern wB at 2; rewrite <- wB_div_2; auto with zarith. + pattern wB at 2; rewrite <- wB_div_2; auto with zarith. match type of H1 with ?X = _ => assert (U5: X < wB / 4 * wB) end. @@ -764,9 +762,9 @@ intros x; case x; simpl ww_is_even. destruct (spec_to_Z w3);auto with zarith. generalize (@spec_w_div2s c w0 w4 U1 H2). case (w_div2s c w0 w4). - intros c0; case c0; intros w5; + intros c0; case c0; intros w5; repeat (rewrite C0_id || rewrite C1_plus_wB). - intros c1; case c1; intros w6; + intros c1; case c1; intros w6; repeat (rewrite C0_id || rewrite C1_plus_wB). intros (H3, H4). match goal with |- context [ww_sub_c ?y ?z] => @@ -1038,7 +1036,7 @@ intros x; case x; simpl ww_is_even. end. apply Zle_not_lt; rewrite <- H3; auto with zarith. rewrite Zmult_plus_distr_l. - apply Zlt_le_trans with ((2 * [|w4|]) * wB + 0); + apply Zlt_le_trans with ((2 * [|w4|]) * wB + 0); auto with zarith. apply beta_lex_inv; auto with zarith. destruct (spec_to_Z w0);auto with zarith. @@ -1119,9 +1117,9 @@ intros x; case x; simpl ww_is_even. auto with zarith. simpl ww_to_Z. assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z x);auto with zarith. - Qed. - - Lemma wwB_4_2: 2 * (wwB / 4) = wwB/ 2. + Qed. + + Lemma wwB_4_2: 2 * (wwB / 4) = wwB/ 2. pattern wwB at 1; rewrite wwB_wBwB; rewrite Zpower_2. rewrite <- wB_div_2. match goal with |- context[(2 * ?X) * (2 * ?Z)] => @@ -1134,7 +1132,7 @@ intros x; case x; simpl ww_is_even. Lemma spec_ww_head1 - : forall x : zn2z w, + : forall x : zn2z w, (ww_is_even (ww_head1 x) = true) /\ (0 < [[x]] -> wwB / 4 <= 2 ^ [[ww_head1 x]] * [[x]] < wwB). assert (U := wB_pos w_digits). @@ -1167,7 +1165,7 @@ intros x; case x; simpl ww_is_even. rewrite Hp. rewrite Zminus_mod; auto with zarith. rewrite H2; repeat rewrite Zmod_small; auto with zarith. - intros H3; rewrite Hp. + intros H3; rewrite Hp. case (spec_ww_head0 x); auto; intros Hv3 Hv4. assert (Hu: forall u, 0 < u -> 2 * 2 ^ (u - 1) = 2 ^u). intros u Hu. @@ -1189,7 +1187,7 @@ intros x; case x; simpl ww_is_even. apply sym_equal; apply Zdiv_unique with 0; auto with zarith. rewrite Zmult_assoc; rewrite wB_div_4; auto with zarith. - rewrite wwB_wBwB; ring. + rewrite wwB_wBwB; ring. Qed. Lemma spec_ww_sqrt : forall x, @@ -1198,14 +1196,14 @@ intros x; case x; simpl ww_is_even. intro x; unfold ww_sqrt. generalize (spec_ww_is_zero x); case (ww_is_zero x). simpl ww_to_Z; simpl Zpower; unfold Zpower_pos; simpl; - auto with zarith. + auto with zarith. intros H1. generalize (spec_ww_compare (ww_head1 x) W0); case ww_compare; simpl ww_to_Z; autorewrite with rm10. generalize H1; case x. intros HH; contradict HH; simpl ww_to_Z; auto with zarith. intros w0 w1; simpl ww_to_Z; autorewrite with w_rewrite rm10. - intros H2; case (spec_ww_head1 (WW w0 w1)); intros H3 H4 H5. + intros H2; case (spec_ww_head1 (WW w0 w1)); intros H3 H4 H5. generalize (H4 H2); clear H4; rewrite H5; clear H5; autorewrite with rm10. intros (H4, H5). assert (V: wB/4 <= [|w0|]). @@ -1241,7 +1239,7 @@ intros x; case x; simpl ww_is_even. apply Zle_not_lt; unfold base. apply Zle_trans with (2 ^ [[ww_head1 x]]). apply Zpower_le_monotone; auto with zarith. - pattern (2 ^ [[ww_head1 x]]) at 1; + pattern (2 ^ [[ww_head1 x]]) at 1; rewrite <- (Zmult_1_r (2 ^ [[ww_head1 x]])). apply Zmult_le_compat_l; auto with zarith. generalize (spec_ww_add_mul_div x W0 (ww_head1 x) Hv2); @@ -1283,13 +1281,13 @@ intros x; case x; simpl ww_is_even. rewrite Zmod_small; auto with zarith. split; auto with zarith. apply Zlt_le_trans with (Zpos (xO w_digits)); auto with zarith. - unfold base; apply Zpower2_le_lin; auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. assert (Hv4: [[ww_head1 x]]/2 < wB). apply Zle_lt_trans with (Zpos w_digits). apply Zmult_le_reg_r with 2; auto with zarith. repeat rewrite (fun x => Zmult_comm x 2). rewrite <- Hv0; rewrite <- Zpos_xO; auto. - unfold base; apply Zpower2_lt_lin; auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. assert (Hv5: [[(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))]] = [[ww_head1 x]]/2). rewrite spec_ww_add_mul_div. @@ -1330,14 +1328,14 @@ intros x; case x; simpl ww_is_even. rewrite tmp; clear tmp. apply Zpower_le_monotone3; auto with zarith. split; auto with zarith. - pattern [|w2|] at 2; + pattern [|w2|] at 2; rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]] / 2))); auto with zarith. match goal with |- ?X <= ?X + ?Y => assert (0 <= Y); auto with zarith end. case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]] / 2))); auto with zarith. - case c; unfold interp_carry; autorewrite with rm10; + case c; unfold interp_carry; autorewrite with rm10; intros w3; assert (V3 := spec_to_Z w3);auto with zarith. apply Zmult_lt_reg_r with (2 ^ [[ww_head1 x]]); auto with zarith. rewrite H4. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v index 269d62bb..a7e55671 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleSub.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. @@ -17,7 +17,7 @@ Require Import BigNumPrelude. Require Import DoubleType. Require Import DoubleBase. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section DoubleSub. Variable w : Type. @@ -39,7 +39,7 @@ Section DoubleSub. Definition ww_opp_c x := match x with | W0 => C0 W0 - | WW xh xl => + | WW xh xl => match w_opp_c xl with | C0 _ => match w_opp_c xh with @@ -53,7 +53,7 @@ Section DoubleSub. Definition ww_opp x := match x with | W0 => W0 - | WW xh xl => + | WW xh xl => match w_opp_c xl with | C0 _ => WW (w_opp xh) w_0 | C1 l => WW (w_opp_carry xh) l @@ -72,14 +72,14 @@ Section DoubleSub. | WW xh xl => match w_pred_c xl with | C0 l => C0 (w_WW xh l) - | C1 _ => - match w_pred_c xh with + | C1 _ => + match w_pred_c xh with | C0 h => C0 (WW h w_Bm1) | C1 _ => C1 ww_Bm1 end end end. - + Definition ww_pred x := match x with | W0 => ww_Bm1 @@ -89,19 +89,19 @@ Section DoubleSub. | C1 l => WW (w_pred xh) w_Bm1 end end. - + Definition ww_sub_c x y := match y, x with | W0, _ => C0 x | WW yh yl, W0 => ww_opp_c (WW yh yl) | WW yh yl, WW xh xl => match w_sub_c xl yl with - | C0 l => + | C0 l => match w_sub_c xh yh with | C0 h => C0 (w_WW h l) | C1 h => C1 (WW h l) end - | C1 l => + | C1 l => match w_sub_carry_c xh yh with | C0 h => C0 (WW h l) | C1 h => C1 (WW h l) @@ -109,7 +109,7 @@ Section DoubleSub. end end. - Definition ww_sub x y := + Definition ww_sub x y := match y, x with | W0, _ => x | WW yh yl, W0 => ww_opp (WW yh yl) @@ -127,7 +127,7 @@ Section DoubleSub. | WW yh yl, W0 => C1 (ww_opp_carry (WW yh yl)) | WW yh yl, WW xh xl => match w_sub_carry_c xl yl with - | C0 l => + | C0 l => match w_sub_c xh yh with | C0 h => C0 (w_WW h l) | C1 h => C1 (WW h l) @@ -155,7 +155,7 @@ Section DoubleSub. (*Section DoubleProof.*) Variable w_digits : positive. Variable w_to_Z : w -> Z. - + Notation wB := (base w_digits). Notation wwB := (base (ww_digits w_digits)). @@ -166,13 +166,13 @@ Section DoubleSub. (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). - Notation "[+[ c ]]" := - (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). - Notation "[-[ c ]]" := - (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) (at level 0, x at level 99). - + Variable spec_w_0 : [|w_0|] = 0. Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. Variable spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1. @@ -187,7 +187,7 @@ Section DoubleSub. Variable spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|]. Variable spec_sub_carry_c : forall x y, [-|w_sub_carry_c x y|] = [|x|] - [|y|] - 1. - + Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB. Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. Variable spec_sub_carry : @@ -197,12 +197,12 @@ Section DoubleSub. Lemma spec_ww_opp_c : forall x, [-[ww_opp_c x]] = -[[x]]. Proof. destruct x as [ |xh xl];simpl. reflexivity. - rewrite Zopp_plus_distr;generalize (spec_opp_c xl);destruct (w_opp_c xl) + rewrite Zopp_plus_distr;generalize (spec_opp_c xl);destruct (w_opp_c xl) as [l|l];intros H;unfold interp_carry in H;rewrite <- H; - rewrite Zopp_mult_distr_l. + rewrite Zopp_mult_distr_l. assert ([|l|] = 0). assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega. - rewrite H0;generalize (spec_opp_c xh);destruct (w_opp_c xh) + rewrite H0;generalize (spec_opp_c xh);destruct (w_opp_c xh) as [h|h];intros H1;unfold interp_carry in *;rewrite <- H1. assert ([|h|] = 0). assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega. @@ -216,7 +216,7 @@ Section DoubleSub. Proof. destruct x as [ |xh xl];simpl. reflexivity. rewrite Zopp_plus_distr;rewrite Zopp_mult_distr_l. - generalize (spec_opp_c xl);destruct (w_opp_c xl) + generalize (spec_opp_c xl);destruct (w_opp_c xl) as [l|l];intros H;unfold interp_carry in H;rewrite <- H;simpl ww_to_Z. rewrite spec_w_0;rewrite Zplus_0_r;rewrite wwB_wBwB. assert ([|l|] = 0). @@ -247,7 +247,7 @@ Section DoubleSub. assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega. rewrite H0;change ([|xh|] + -1) with ([|xh|] - 1). generalize (spec_pred_c xh);destruct (w_pred_c xh) as [h|h]; - intros H1;unfold interp_carry in H1;rewrite <- H1. + intros H1;unfold interp_carry in H1;rewrite <- H1. simpl;rewrite spec_w_Bm1;ring. assert ([|h|] = wB - 1). assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega. @@ -258,14 +258,14 @@ Section DoubleSub. Proof. destruct y as [ |yh yl];simpl. ring. destruct x as [ |xh xl];simpl. exact (spec_ww_opp_c (WW yh yl)). - replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|])) + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|])) with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|])). 2:ring. generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl) as [l|l];intros H; unfold interp_carry in H;rewrite <- H. generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1; unfold interp_carry in H1;rewrite <- H1;unfold interp_carry; try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring. - rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1). generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h]; intros H1;unfold interp_carry in *;rewrite <- H1;simpl ww_to_Z; @@ -275,37 +275,37 @@ Section DoubleSub. Lemma spec_ww_sub_carry_c : forall x y, [-[ww_sub_carry_c x y]] = [[x]] - [[y]] - 1. Proof. - destruct y as [ |yh yl];simpl. + destruct y as [ |yh yl];simpl. unfold Zminus;simpl;rewrite Zplus_0_r;exact (spec_ww_pred_c x). destruct x as [ |xh xl]. unfold interp_carry;rewrite spec_w_WW;simpl ww_to_Z;rewrite wwB_wBwB; repeat rewrite spec_opp_carry;ring. simpl ww_to_Z. - replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1) + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1) with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|]-1)). 2:ring. - generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl) + generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl) as [l|l];intros H;unfold interp_carry in H;rewrite <- H. generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1; unfold interp_carry in H1;rewrite <- H1;unfold interp_carry; try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring. - rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1). generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h]; intros H1;unfold interp_carry in *;rewrite <- H1;try rewrite spec_w_WW; simpl ww_to_Z; try rewrite wwB_wBwB;ring. - Qed. - + Qed. + Lemma spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB. Proof. - destruct x as [ |xh xl];simpl. + destruct x as [ |xh xl];simpl. apply Zmod_unique with (-1). apply spec_ww_to_Z;trivial. rewrite spec_ww_Bm1;ring. replace ([|xh|]*wB + [|xl|] - 1) with ([|xh|]*wB + ([|xl|] - 1)). 2:ring. generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l];intro H; unfold interp_carry in H;rewrite <- H;simpl ww_to_Z. - rewrite Zmod_small. apply spec_w_WW. + rewrite Zmod_small. apply spec_w_WW. exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh l)). - rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. change ([|xh|] + -1) with ([|xh|] - 1). assert ([|l|] = wB - 1). assert (H1:= spec_to_Z l);assert (H2:= spec_to_Z xl);omega. @@ -318,7 +318,7 @@ Section DoubleSub. destruct y as [ |yh yl];simpl. ring_simplify ([[x]] - 0);rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial. destruct x as [ |xh xl];simpl. exact (spec_ww_opp (WW yh yl)). - replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|])) + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|])) with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|])). 2:ring. generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl)as[l|l];intros H; unfold interp_carry in H;rewrite <- H. @@ -338,7 +338,7 @@ Section DoubleSub. apply spec_ww_to_Z;trivial. fold (ww_opp_carry (WW yh yl)). rewrite (spec_ww_opp_carry (WW yh yl));simpl ww_to_Z;ring. - replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1) + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1) with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|] - 1)). 2:ring. generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl)as[l|l]; intros H;unfold interp_carry in H;rewrite <- H;rewrite spec_w_WW. @@ -354,4 +354,4 @@ End DoubleSub. - + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v index 28d40094..88cbb484 100644 --- a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v @@ -8,12 +8,12 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: DoubleType.v 10964 2008-05-22 11:08:13Z letouzey $ i*) +(*i $Id$ i*) Set Implicit Arguments. Require Import ZArith. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Definition base digits := Zpower 2 (Zpos digits). @@ -37,10 +37,10 @@ Section Zn2Z. Variable znz : Type. - (** From a type [znz] representing a cyclic structure Z/nZ, + (** From a type [znz] representing a cyclic structure Z/nZ, we produce a representation of Z/2nZ by pairs of elements of [znz] - (plus a special case for zero). High half of the new number comes - first. + (plus a special case for zero). High half of the new number comes + first. *) Inductive zn2z := @@ -57,10 +57,10 @@ End Zn2Z. Implicit Arguments W0 [znz]. -(** From a cyclic representation [w], we iterate the [zn2z] construct - [n] times, gaining the type of binary trees of depth at most [n], - whose leafs are either W0 (if depth < n) or elements of w - (if depth = n). +(** From a cyclic representation [w], we iterate the [zn2z] construct + [n] times, gaining the type of binary trees of depth at most [n], + whose leafs are either W0 (if depth < n) or elements of w + (if depth = n). *) Fixpoint word (w:Type) (n:nat) : Type := diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v index 6da1c6ec..8addf5b9 100644 --- a/theories/Numbers/Cyclic/Int31/Cyclic31.v +++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(*i $Id: Cyclic31.v 11907 2009-02-10 23:54:28Z letouzey $ i*) +(*i $Id$ i*) (** * Int31 numbers defines indeed a cyclic structure : Z/(2^31)Z *) @@ -24,8 +24,8 @@ Require Import BigNumPrelude. Require Import CyclicAxioms. Require Import ROmega. -Open Scope nat_scope. -Open Scope int31_scope. +Local Open Scope nat_scope. +Local Open Scope int31_scope. Section Basics. @@ -34,9 +34,9 @@ Section Basics. Lemma iszero_eq0 : forall x, iszero x = true -> x=0. Proof. destruct x; simpl; intros. - repeat - match goal with H:(if ?d then _ else _) = true |- _ => - destruct d; try discriminate + repeat + match goal with H:(if ?d then _ else _) = true |- _ => + destruct d; try discriminate end. reflexivity. Qed. @@ -46,26 +46,26 @@ Section Basics. intros x H Eq; rewrite Eq in H; simpl in *; discriminate. Qed. - Lemma sneakl_shiftr : forall x, + Lemma sneakl_shiftr : forall x, x = sneakl (firstr x) (shiftr x). Proof. destruct x; simpl; auto. Qed. - Lemma sneakr_shiftl : forall x, + Lemma sneakr_shiftl : forall x, x = sneakr (firstl x) (shiftl x). Proof. destruct x; simpl; auto. Qed. - Lemma twice_zero : forall x, + Lemma twice_zero : forall x, twice x = 0 <-> twice_plus_one x = 1. Proof. - destruct x; simpl in *; split; + destruct x; simpl in *; split; intro H; injection H; intros; subst; auto. Qed. - Lemma twice_or_twice_plus_one : forall x, + Lemma twice_or_twice_plus_one : forall x, x = twice (shiftr x) \/ x = twice_plus_one (shiftr x). Proof. intros; case_eq (firstr x); intros. @@ -79,13 +79,13 @@ Section Basics. Definition nshiftr n x := iter_nat n _ shiftr x. - Lemma nshiftr_S : + Lemma nshiftr_S : forall n x, nshiftr (S n) x = shiftr (nshiftr n x). Proof. reflexivity. Qed. - Lemma nshiftr_S_tail : + Lemma nshiftr_S_tail : forall n x, nshiftr (S n) x = nshiftr n (shiftr x). Proof. induction n; simpl; auto. @@ -103,7 +103,7 @@ Section Basics. destruct x; simpl; auto. Qed. - Lemma nshiftr_above_size : forall k x, size<=k -> + Lemma nshiftr_above_size : forall k x, size<=k -> nshiftr k x = 0. Proof. intros. @@ -117,13 +117,13 @@ Section Basics. Definition nshiftl n x := iter_nat n _ shiftl x. - Lemma nshiftl_S : + Lemma nshiftl_S : forall n x, nshiftl (S n) x = shiftl (nshiftl n x). Proof. reflexivity. Qed. - Lemma nshiftl_S_tail : + Lemma nshiftl_S_tail : forall n x, nshiftl (S n) x = nshiftl n (shiftl x). Proof. induction n; simpl; auto. @@ -141,7 +141,7 @@ Section Basics. destruct x; simpl; auto. Qed. - Lemma nshiftl_above_size : forall k x, size<=k -> + Lemma nshiftl_above_size : forall k x, size<=k -> nshiftl k x = 0. Proof. intros. @@ -151,27 +151,27 @@ Section Basics. simpl; rewrite nshiftl_S, IHn; auto. Qed. - Lemma firstr_firstl : + Lemma firstr_firstl : forall x, firstr x = firstl (nshiftl (pred size) x). Proof. destruct x; simpl; auto. Qed. - Lemma firstl_firstr : + Lemma firstl_firstr : forall x, firstl x = firstr (nshiftr (pred size) x). Proof. destruct x; simpl; auto. Qed. - + (** More advanced results about [nshiftr] *) - Lemma nshiftr_predsize_0_firstl : forall x, + Lemma nshiftr_predsize_0_firstl : forall x, nshiftr (pred size) x = 0 -> firstl x = D0. Proof. destruct x; compute; intros H; injection H; intros; subst; auto. Qed. - Lemma nshiftr_0_propagates : forall n p x, n <= p -> + Lemma nshiftr_0_propagates : forall n p x, n <= p -> nshiftr n x = 0 -> nshiftr p x = 0. Proof. intros. @@ -181,7 +181,7 @@ Section Basics. simpl; rewrite nshiftr_S; rewrite IHn0; auto. Qed. - Lemma nshiftr_0_firstl : forall n x, n < size -> + Lemma nshiftr_0_firstl : forall n x, n < size -> nshiftr n x = 0 -> firstl x = D0. Proof. intros. @@ -194,8 +194,8 @@ Section Basics. (** Not used for the moment. Are they really useful ? *) Lemma int31_ind_sneakl : forall P : int31->Prop, - P 0 -> - (forall x d, P x -> P (sneakl d x)) -> + P 0 -> + (forall x d, P x -> P (sneakl d x)) -> forall x, P x. Proof. intros. @@ -210,10 +210,10 @@ Section Basics. change x with (nshiftr (size-size) x); auto. Qed. - Lemma int31_ind_twice : forall P : int31->Prop, - P 0 -> - (forall x, P x -> P (twice x)) -> - (forall x, P x -> P (twice_plus_one x)) -> + Lemma int31_ind_twice : forall P : int31->Prop, + P 0 -> + (forall x, P x -> P (twice x)) -> + (forall x, P x -> P (twice_plus_one x)) -> forall x, P x. Proof. induction x using int31_ind_sneakl; auto. @@ -224,21 +224,21 @@ Section Basics. (** * Some generic results about [recr] *) Section Recr. - + (** [recr] satisfies the fixpoint equation used for its definition. *) Variable (A:Type)(case0:A)(caserec:digits->int31->A->A). - - Lemma recr_aux_eqn : forall n x, iszero x = false -> - recr_aux (S n) A case0 caserec x = + + Lemma recr_aux_eqn : forall n x, iszero x = false -> + recr_aux (S n) A case0 caserec x = caserec (firstr x) (shiftr x) (recr_aux n A case0 caserec (shiftr x)). Proof. intros; simpl; rewrite H; auto. Qed. - Lemma recr_aux_converges : + Lemma recr_aux_converges : forall n p x, n <= size -> n <= p -> - recr_aux n A case0 caserec (nshiftr (size - n) x) = + recr_aux n A case0 caserec (nshiftr (size - n) x) = recr_aux p A case0 caserec (nshiftr (size - n) x). Proof. induction n. @@ -255,8 +255,8 @@ Section Basics. apply IHn; auto with arith. Qed. - Lemma recr_eqn : forall x, iszero x = false -> - recr A case0 caserec x = + Lemma recr_eqn : forall x, iszero x = false -> + recr A case0 caserec x = caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x)). Proof. intros. @@ -265,11 +265,11 @@ Section Basics. rewrite (recr_aux_converges size (S size)); auto with arith. rewrite recr_aux_eqn; auto. Qed. - - (** [recr] is usually equivalent to a variant [recrbis] + + (** [recr] is usually equivalent to a variant [recrbis] written without [iszero] check. *) - Fixpoint recrbis_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A) + Fixpoint recrbis_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A) (i:int31) : A := match n with | O => case0 @@ -277,7 +277,7 @@ Section Basics. let si := shiftr i in caserec (firstr i) si (recrbis_aux next A case0 caserec si) end. - + Definition recrbis := recrbis_aux size. Hypothesis case0_caserec : caserec D0 0 case0 = case0. @@ -291,8 +291,8 @@ Section Basics. replace (recrbis_aux n A case0 caserec 0) with case0; auto. clear H IHn; induction n; simpl; congruence. Qed. - - Lemma recrbis_equiv : forall x, + + Lemma recrbis_equiv : forall x, recrbis A case0 caserec x = recr A case0 caserec x. Proof. intros; apply recrbis_aux_equiv; auto. @@ -348,7 +348,7 @@ Section Basics. rewrite incr_eqn1; destruct x; simpl; auto. Qed. - Lemma incr_twice_plus_one_firstl : + Lemma incr_twice_plus_one_firstl : forall x, firstl x = D0 -> incr (twice_plus_one x) = twice (incr x). Proof. intros. @@ -356,9 +356,9 @@ Section Basics. f_equal; f_equal. destruct x; simpl in *; rewrite H; auto. Qed. - - (** The previous result is actually true even without the - constraint on [firstl], but this is harder to prove + + (** The previous result is actually true even without the + constraint on [firstl], but this is harder to prove (see later). *) End Incr. @@ -369,9 +369,9 @@ Section Basics. (** Variant of [phi] via [recrbis] *) - Let Phi := fun b (_:int31) => + Let Phi := fun b (_:int31) => match b with D0 => Zdouble | D1 => Zdouble_plus_one end. - + Definition phibis_aux n x := recrbis_aux n _ Z0 Phi x. Lemma phibis_aux_equiv : forall x, phibis_aux size x = phi x. @@ -382,7 +382,7 @@ Section Basics. (** Recursive equations satisfied by [phi] *) - Lemma phi_eqn1 : forall x, firstr x = D0 -> + Lemma phi_eqn1 : forall x, firstr x = D0 -> phi x = Zdouble (phi (shiftr x)). Proof. intros. @@ -392,7 +392,7 @@ Section Basics. rewrite H; auto. Qed. - Lemma phi_eqn2 : forall x, firstr x = D1 -> + Lemma phi_eqn2 : forall x, firstr x = D1 -> phi x = Zdouble_plus_one (phi (shiftr x)). Proof. intros. @@ -402,7 +402,7 @@ Section Basics. rewrite H; auto. Qed. - Lemma phi_twice_firstl : forall x, firstl x = D0 -> + Lemma phi_twice_firstl : forall x, firstl x = D0 -> phi (twice x) = Zdouble (phi x). Proof. intros. @@ -411,7 +411,7 @@ Section Basics. destruct x; simpl in *; rewrite H; auto. Qed. - Lemma phi_twice_plus_one_firstl : forall x, firstl x = D0 -> + Lemma phi_twice_plus_one_firstl : forall x, firstl x = D0 -> phi (twice_plus_one x) = Zdouble_plus_one (phi x). Proof. intros. @@ -427,23 +427,23 @@ Section Basics. Lemma phibis_aux_pos : forall n x, (0 <= phibis_aux n x)%Z. Proof. induction n. - simpl; unfold phibis_aux; simpl; auto with zarith. + simpl; unfold phibis_aux; simpl; auto with zarith. intros. - unfold phibis_aux, recrbis_aux; fold recrbis_aux; + unfold phibis_aux, recrbis_aux; fold recrbis_aux; fold (phibis_aux n (shiftr x)). destruct (firstr x). specialize IHn with (shiftr x); rewrite Zdouble_mult; omega. specialize IHn with (shiftr x); rewrite Zdouble_plus_one_mult; omega. Qed. - Lemma phibis_aux_bounded : - forall n x, n <= size -> + Lemma phibis_aux_bounded : + forall n x, n <= size -> (phibis_aux n (nshiftr (size-n) x) < 2 ^ (Z_of_nat n))%Z. Proof. induction n. simpl; unfold phibis_aux; simpl; auto with zarith. intros. - unfold phibis_aux, recrbis_aux; fold recrbis_aux; + unfold phibis_aux, recrbis_aux; fold recrbis_aux; fold (phibis_aux n (shiftr (nshiftr (size - S n) x))). assert (shiftr (nshiftr (size - S n) x) = nshiftr (size-n) x). replace (size - n)%nat with (S (size - (S n))) by omega. @@ -468,8 +468,8 @@ Section Basics. apply phibis_aux_bounded; auto. Qed. - Lemma phibis_aux_lowerbound : - forall n x, firstr (nshiftr n x) = D1 -> + Lemma phibis_aux_lowerbound : + forall n x, firstr (nshiftr n x) = D1 -> (2 ^ Z_of_nat n <= phibis_aux (S n) x)%Z. Proof. induction n. @@ -480,7 +480,7 @@ Section Basics. intros. remember (S n) as m. - unfold phibis_aux, recrbis_aux; fold recrbis_aux; + unfold phibis_aux, recrbis_aux; fold recrbis_aux; fold (phibis_aux m (shiftr x)). subst m. rewrite inj_S, Zpower_Zsucc; auto with zarith. @@ -488,13 +488,13 @@ Section Basics. apply IHn. rewrite <- nshiftr_S_tail; auto. destruct (firstr x). - change (Zdouble (phibis_aux (S n) (shiftr x))) with + change (Zdouble (phibis_aux (S n) (shiftr x))) with (2*(phibis_aux (S n) (shiftr x)))%Z. omega. rewrite Zdouble_plus_one_mult; omega. Qed. - Lemma phi_lowerbound : + Lemma phi_lowerbound : forall x, firstl x = D1 -> (2^(Z_of_nat (pred size)) <= phi x)%Z. Proof. intros. @@ -508,9 +508,9 @@ Section Basics. Section EqShiftL. - (** After killing [n] bits at the left, are the numbers equal ?*) + (** After killing [n] bits at the left, are the numbers equal ?*) - Definition EqShiftL n x y := + Definition EqShiftL n x y := nshiftl n x = nshiftl n y. Lemma EqShiftL_zero : forall x y, EqShiftL O x y <-> x = y. @@ -523,7 +523,7 @@ Section Basics. red; intros; rewrite 2 nshiftl_above_size; auto. Qed. - Lemma EqShiftL_le : forall k k' x y, k <= k' -> + Lemma EqShiftL_le : forall k k' x y, k <= k' -> EqShiftL k x y -> EqShiftL k' x y. Proof. unfold EqShiftL; intros. @@ -534,18 +534,18 @@ Section Basics. rewrite 2 nshiftl_S; f_equal; auto. Qed. - Lemma EqShiftL_firstr : forall k x y, k < size -> + Lemma EqShiftL_firstr : forall k x y, k < size -> EqShiftL k x y -> firstr x = firstr y. Proof. intros. rewrite 2 firstr_firstl. f_equal. - apply EqShiftL_le with k; auto. + apply EqShiftL_le with k; auto. unfold size. auto with arith. Qed. - Lemma EqShiftL_twice : forall k x y, + Lemma EqShiftL_twice : forall k x y, EqShiftL k (twice x) (twice y) <-> EqShiftL (S k) x y. Proof. intros; unfold EqShiftL. @@ -553,7 +553,7 @@ Section Basics. Qed. (** * From int31 to list of digits. *) - + (** Lower (=rightmost) bits comes first. *) Definition i2l := recrbis _ nil (fun d _ rec => d::rec). @@ -561,10 +561,10 @@ Section Basics. Lemma i2l_length : forall x, length (i2l x) = size. Proof. intros; reflexivity. - Qed. + Qed. - Fixpoint lshiftl l x := - match l with + Fixpoint lshiftl l x := + match l with | nil => x | d::l => sneakl d (lshiftl l x) end. @@ -576,19 +576,19 @@ Section Basics. destruct x; compute; auto. Qed. - Lemma i2l_sneakr : forall x d, + Lemma i2l_sneakr : forall x d, i2l (sneakr d x) = tail (i2l x) ++ d::nil. Proof. destruct x; compute; auto. Qed. - Lemma i2l_sneakl : forall x d, + Lemma i2l_sneakl : forall x d, i2l (sneakl d x) = d :: removelast (i2l x). Proof. destruct x; compute; auto. Qed. - Lemma i2l_l2i : forall l, length l = size -> + Lemma i2l_l2i : forall l, length l = size -> i2l (l2i l) = l. Proof. repeat (destruct l as [ |? l]; [intros; discriminate | ]). @@ -596,9 +596,9 @@ Section Basics. intros _; compute; auto. Qed. - Fixpoint cstlist (A:Type)(a:A) n := - match n with - | O => nil + Fixpoint cstlist (A:Type)(a:A) n := + match n with + | O => nil | S n => a::cstlist _ a n end. @@ -612,7 +612,7 @@ Section Basics. induction (i2l x); simpl; f_equal; auto. rewrite H0; clear H0. reflexivity. - + intros. rewrite nshiftl_S. unfold shiftl; rewrite i2l_sneakl. @@ -657,10 +657,10 @@ Section Basics. f_equal; auto. Qed. - (** This equivalence allows to prove easily the following delicate + (** This equivalence allows to prove easily the following delicate result *) - Lemma EqShiftL_twice_plus_one : forall k x y, + Lemma EqShiftL_twice_plus_one : forall k x y, EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y. Proof. intros. @@ -683,7 +683,7 @@ Section Basics. subst lx n; rewrite i2l_length; omega. Qed. - Lemma EqShiftL_shiftr : forall k x y, EqShiftL k x y -> + Lemma EqShiftL_shiftr : forall k x y, EqShiftL k x y -> EqShiftL (S k) (shiftr x) (shiftr y). Proof. intros. @@ -704,41 +704,41 @@ Section Basics. omega. Qed. - Lemma EqShiftL_incrbis : forall n k x y, n<=size -> + Lemma EqShiftL_incrbis : forall n k x y, n<=size -> (n+k=S size)%nat -> - EqShiftL k x y -> + EqShiftL k x y -> EqShiftL k (incrbis_aux n x) (incrbis_aux n y). Proof. induction n; simpl; intros. red; auto. - destruct (eq_nat_dec k size). + destruct (eq_nat_dec k size). subst k; apply EqShiftL_size; auto. - unfold incrbis_aux; simpl; + unfold incrbis_aux; simpl; fold (incrbis_aux n (shiftr x)); fold (incrbis_aux n (shiftr y)). rewrite (EqShiftL_firstr k x y); auto; try omega. case_eq (firstr y); intros. rewrite EqShiftL_twice_plus_one. apply EqShiftL_shiftr; auto. - + rewrite EqShiftL_twice. apply IHn; try omega. apply EqShiftL_shiftr; auto. Qed. - Lemma EqShiftL_incr : forall x y, + Lemma EqShiftL_incr : forall x y, EqShiftL 1 x y -> EqShiftL 1 (incr x) (incr y). Proof. intros. rewrite <- 2 incrbis_aux_equiv. apply EqShiftL_incrbis; auto. Qed. - + End EqShiftL. (** * More equations about [incr] *) - Lemma incr_twice_plus_one : + Lemma incr_twice_plus_one : forall x, incr (twice_plus_one x) = twice (incr x). Proof. intros. @@ -757,7 +757,7 @@ Section Basics. destruct (incr (shiftr x)); simpl; discriminate. Qed. - Lemma incr_inv : forall x y, + Lemma incr_inv : forall x y, incr x = twice_plus_one y -> x = twice y. Proof. intros. @@ -777,7 +777,7 @@ Section Basics. (** First, recursive equations *) - Lemma phi_inv_double_plus_one : forall z, + Lemma phi_inv_double_plus_one : forall z, phi_inv (Zdouble_plus_one z) = twice_plus_one (phi_inv z). Proof. destruct z; simpl; auto. @@ -789,14 +789,14 @@ Section Basics. auto. Qed. - Lemma phi_inv_double : forall z, + Lemma phi_inv_double : forall z, phi_inv (Zdouble z) = twice (phi_inv z). Proof. destruct z; simpl; auto. rewrite incr_twice_plus_one; auto. Qed. - Lemma phi_inv_incr : forall z, + Lemma phi_inv_incr : forall z, phi_inv (Zsucc z) = incr (phi_inv z). Proof. destruct z. @@ -816,19 +816,19 @@ Section Basics. rewrite incr_twice_plus_one; auto. Qed. - (** [phi_inv o inv], the always-exact and easy-to-prove trip : + (** [phi_inv o inv], the always-exact and easy-to-prove trip : from int31 to Z and then back to int31. *) - Lemma phi_inv_phi_aux : - forall n x, n <= size -> - phi_inv (phibis_aux n (nshiftr (size-n) x)) = + Lemma phi_inv_phi_aux : + forall n x, n <= size -> + phi_inv (phibis_aux n (nshiftr (size-n) x)) = nshiftr (size-n) x. Proof. induction n. intros; simpl. rewrite nshiftr_size; auto. intros. - unfold phibis_aux, recrbis_aux; fold recrbis_aux; + unfold phibis_aux, recrbis_aux; fold recrbis_aux; fold (phibis_aux n (shiftr (nshiftr (size-S n) x))). assert (shiftr (nshiftr (size - S n) x) = nshiftr (size-n) x). replace (size - n)%nat with (S (size - (S n))); auto; omega. @@ -863,10 +863,10 @@ Section Basics. (** * [positive_to_int31] *) - (** A variant of [p2i] with [twice] and [twice_plus_one] instead of + (** A variant of [p2i] with [twice] and [twice_plus_one] instead of [2*i] and [2*i+1] *) - Fixpoint p2ibis n p : (N*int31)%type := + Fixpoint p2ibis n p : (N*int31)%type := match n with | O => (Npos p, On) | S n => match p with @@ -876,7 +876,7 @@ Section Basics. end end. - Lemma p2ibis_bounded : forall n p, + Lemma p2ibis_bounded : forall n p, nshiftr n (snd (p2ibis n p)) = 0. Proof. induction n. @@ -906,20 +906,20 @@ Section Basics. replace (shiftr In) with 0; auto. apply nshiftr_n_0. Qed. - + Lemma p2ibis_spec : forall n p, n<=size -> - Zpos p = ((Z_of_N (fst (p2ibis n p)))*2^(Z_of_nat n) + + Zpos p = ((Z_of_N (fst (p2ibis n p)))*2^(Z_of_nat n) + phi (snd (p2ibis n p)))%Z. Proof. induction n; intros. simpl; rewrite Pmult_1_r; auto. - replace (2^(Z_of_nat (S n)))%Z with (2*2^(Z_of_nat n))%Z by - (rewrite <- Zpower_Zsucc, <- Zpos_P_of_succ_nat; + replace (2^(Z_of_nat (S n)))%Z with (2*2^(Z_of_nat n))%Z by + (rewrite <- Zpower_Zsucc, <- Zpos_P_of_succ_nat; auto with zarith). rewrite (Zmult_comm 2). assert (n<=size) by omega. - destruct p; simpl; [ | | auto]; - specialize (IHn p H0); + destruct p; simpl; [ | | auto]; + specialize (IHn p H0); generalize (p2ibis_bounded n p); destruct (p2ibis n p) as (r,i); simpl in *; intros. @@ -937,25 +937,25 @@ Section Basics. (** We now prove that this [p2ibis] is related to [phi_inv_positive] *) - Lemma phi_inv_positive_p2ibis : forall n p, (n<=size)%nat -> + Lemma phi_inv_positive_p2ibis : forall n p, (n<=size)%nat -> EqShiftL (size-n) (phi_inv_positive p) (snd (p2ibis n p)). Proof. induction n. intros. apply EqShiftL_size; auto. intros. - simpl p2ibis; destruct p; [ | | red; auto]; - specialize IHn with p; - destruct (p2ibis n p); simpl snd in *; simpl phi_inv_positive; - rewrite ?EqShiftL_twice_plus_one, ?EqShiftL_twice; - replace (S (size - S n))%nat with (size - n)%nat by omega; + simpl p2ibis; destruct p; [ | | red; auto]; + specialize IHn with p; + destruct (p2ibis n p); simpl snd in *; simpl phi_inv_positive; + rewrite ?EqShiftL_twice_plus_one, ?EqShiftL_twice; + replace (S (size - S n))%nat with (size - n)%nat by omega; apply IHn; omega. Qed. (** This gives the expected result about [phi o phi_inv], at least for the positive case. *) - Lemma phi_phi_inv_positive : forall p, + Lemma phi_phi_inv_positive : forall p, phi (phi_inv_positive p) = (Zpos p) mod (2^(Z_of_nat size)). Proof. intros. @@ -975,12 +975,12 @@ Section Basics. Lemma double_twice_firstl : forall x, firstl x = D0 -> Twon*x = twice x. Proof. - intros. + intros. unfold mul31. rewrite <- Zdouble_mult, <- phi_twice_firstl, phi_inv_phi; auto. Qed. - Lemma double_twice_plus_one_firstl : forall x, firstl x = D0 -> + Lemma double_twice_plus_one_firstl : forall x, firstl x = D0 -> Twon*x+In = twice_plus_one x. Proof. intros. @@ -989,14 +989,14 @@ Section Basics. rewrite phi_twice_firstl, <- Zdouble_plus_one_mult, <- phi_twice_plus_one_firstl, phi_inv_phi; auto. Qed. - - Lemma p2i_p2ibis : forall n p, (n<=size)%nat -> + + Lemma p2i_p2ibis : forall n p, (n<=size)%nat -> p2i n p = p2ibis n p. Proof. induction n; simpl; auto; intros. - destruct p; auto; specialize IHn with p; - generalize (p2ibis_bounded n p); - rewrite IHn; try omega; destruct (p2ibis n p); simpl; intros; + destruct p; auto; specialize IHn with p; + generalize (p2ibis_bounded n p); + rewrite IHn; try omega; destruct (p2ibis n p); simpl; intros; f_equal; auto. apply double_twice_plus_one_firstl. apply (nshiftr_0_firstl n); auto; omega. @@ -1004,7 +1004,7 @@ Section Basics. apply (nshiftr_0_firstl n); auto; omega. Qed. - Lemma positive_to_int31_phi_inv_positive : forall p, + Lemma positive_to_int31_phi_inv_positive : forall p, snd (positive_to_int31 p) = phi_inv_positive p. Proof. intros; unfold positive_to_int31. @@ -1014,8 +1014,8 @@ Section Basics. apply (phi_inv_positive_p2ibis size); auto. Qed. - Lemma positive_to_int31_spec : forall p, - Zpos p = ((Z_of_N (fst (positive_to_int31 p)))*2^(Z_of_nat size) + + Lemma positive_to_int31_spec : forall p, + Zpos p = ((Z_of_N (fst (positive_to_int31 p)))*2^(Z_of_nat size) + phi (snd (positive_to_int31 p)))%Z. Proof. unfold positive_to_int31. @@ -1023,11 +1023,11 @@ Section Basics. apply p2ibis_spec; auto. Qed. - (** Thanks to the result about [phi o phi_inv_positive], we can - now establish easily the most general results about + (** Thanks to the result about [phi o phi_inv_positive], we can + now establish easily the most general results about [phi o twice] and so one. *) - - Lemma phi_twice : forall x, + + Lemma phi_twice : forall x, phi (twice x) = (Zdouble (phi x)) mod 2^(Z_of_nat size). Proof. intros. @@ -1041,7 +1041,7 @@ Section Basics. compute in H; elim H; auto. Qed. - Lemma phi_twice_plus_one : forall x, + Lemma phi_twice_plus_one : forall x, phi (twice_plus_one x) = (Zdouble_plus_one (phi x)) mod 2^(Z_of_nat size). Proof. intros. @@ -1055,14 +1055,14 @@ Section Basics. compute in H; elim H; auto. Qed. - Lemma phi_incr : forall x, + Lemma phi_incr : forall x, phi (incr x) = (Zsucc (phi x)) mod 2^(Z_of_nat size). Proof. intros. pattern x at 1; rewrite <- (phi_inv_phi x). rewrite <- phi_inv_incr. assert (0 <= Zsucc (phi x))%Z. - change (Zsucc (phi x)) with ((phi x)+1)%Z; + change (Zsucc (phi x)) with ((phi x)+1)%Z; generalize (phi_bounded x); omega. destruct (Zsucc (phi x)). simpl; auto. @@ -1070,10 +1070,10 @@ Section Basics. compute in H; elim H; auto. Qed. - (** With the previous results, we can deal with [phi o phi_inv] even + (** With the previous results, we can deal with [phi o phi_inv] even in the negative case *) - Lemma phi_phi_inv_negative : + Lemma phi_phi_inv_negative : forall p, phi (incr (complement_negative p)) = (Zneg p) mod 2^(Z_of_nat size). Proof. induction p. @@ -1091,11 +1091,11 @@ Section Basics. rewrite incr_twice_plus_one, phi_twice. remember (phi (incr (complement_negative p))) as q. rewrite Zdouble_mult, IHp, Zmult_mod_idemp_r; auto with zarith. - + simpl; auto. Qed. - Lemma phi_phi_inv : + Lemma phi_phi_inv : forall z, phi (phi_inv z) = z mod 2 ^ (Z_of_nat size). Proof. destruct z. @@ -1120,7 +1120,7 @@ Let w_pos_mod p i := end. (** Parity test *) -Let w_iseven i := +Let w_iseven i := let (_,r) := i/2 in match r ?= 0 with Eq => true | _ => false end. @@ -1140,7 +1140,7 @@ Definition int31_op := (mk_znz_op w_iszero (* Basic arithmetic operations *) (fun i => 0 -c i) - (fun i => 0 - i) + opp31 (fun i => 0-i-1) (fun i => i +c 1) add31c @@ -1181,14 +1181,14 @@ Definition int31_op := (mk_znz_op End Int31_Op. Section Int31_Spec. - - Open Local Scope Z_scope. + + Local Open Scope Z_scope. Notation "[| x |]" := (phi x) (at level 0, x at level 99). - Notation Local wB := (2 ^ (Z_of_nat size)). - - Lemma wB_pos : wB > 0. + Local Notation wB := (2 ^ (Z_of_nat size)). + + Lemma wB_pos : wB > 0. Proof. auto with zarith. Qed. @@ -1216,12 +1216,12 @@ Section Int31_Spec. Proof. reflexivity. Qed. - + Lemma spec_1 : [| 1 |] = 1. Proof. reflexivity. Qed. - + Lemma spec_Bm1 : [| Tn |] = wB - 1. Proof. reflexivity. @@ -1252,16 +1252,16 @@ Section Int31_Spec. destruct (Z_lt_le_dec (X+Y) wB). contradict H1; auto using Zmod_small with zarith. rewrite <- (Z_mod_plus_full (X+Y) (-1) wB). - rewrite Zmod_small; romega. + rewrite Zmod_small; romega. generalize (Zcompare_Eq_eq ((X+Y) mod wB) (X+Y)); intros Heq. - destruct Zcompare; intros; + destruct Zcompare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed. Lemma spec_succ_c : forall x, [+|add31c x 1|] = [|x|] + 1. Proof. - intros; apply spec_add_c. + intros; apply spec_add_c. Qed. Lemma spec_add_carry_c : forall x y, [+|add31carryc x y|] = [|x|] + [|y|] + 1. @@ -1279,7 +1279,7 @@ Section Int31_Spec. rewrite Zmod_small; romega. generalize (Zcompare_Eq_eq ((X+Y+1) mod wB) (X+Y+1)); intros Heq. - destruct Zcompare; intros; + destruct Zcompare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed. @@ -1304,7 +1304,7 @@ Section Int31_Spec. (** Substraction *) Lemma spec_sub_c : forall x y, [-|sub31c x y|] = [|x|] - [|y|]. - Proof. + Proof. unfold sub31c, sub31, interp_carry; intros. rewrite phi_phi_inv. generalize (phi_bounded x)(phi_bounded y); intros. @@ -1337,7 +1337,7 @@ Section Int31_Spec. contradict H1; apply Zmod_small; romega. generalize (Zcompare_Eq_eq ((X-Y-1) mod wB) (X-Y-1)); intros Heq. - destruct Zcompare; intros; + destruct Zcompare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed. @@ -1355,7 +1355,7 @@ Section Int31_Spec. Qed. Lemma spec_opp_c : forall x, [-|sub31c 0 x|] = -[|x|]. - Proof. + Proof. intros; apply spec_sub_c. Qed. @@ -1402,7 +1402,7 @@ Section Int31_Spec. change (wB*wB) with (wB^2); ring. unfold phi_inv2. - destruct x; unfold zn2z_to_Z; rewrite ?phi_phi_inv; + destruct x; unfold zn2z_to_Z; rewrite ?phi_phi_inv; change base with wB; auto. Qed. @@ -1426,7 +1426,7 @@ Section Int31_Spec. intros; apply spec_mul_c. Qed. - (** Division *) + (** Division *) Lemma spec_div21 : forall a1 a2 b, wB/2 <= [|b|] -> @@ -1537,7 +1537,7 @@ Section Int31_Spec. intros (H,_); compute in H; elim H; auto. Qed. - Lemma iter_int31_iter_nat : forall A f i a, + Lemma iter_int31_iter_nat : forall A f i a, iter_int31 i A f a = iter_nat (Zabs_nat [|i|]) A f a. Proof. intros. @@ -1548,17 +1548,17 @@ Section Int31_Spec. revert i a; induction size. simpl; auto. simpl; intros. - case_eq (firstr i); intros H; rewrite 2 IHn; + case_eq (firstr i); intros H; rewrite 2 IHn; unfold phibis_aux; simpl; rewrite H; fold (phibis_aux n (shiftr i)); - generalize (phibis_aux_pos n (shiftr i)); intros; - set (z := phibis_aux n (shiftr i)) in *; clearbody z; + generalize (phibis_aux_pos n (shiftr i)); intros; + set (z := phibis_aux n (shiftr i)) in *; clearbody z; rewrite <- iter_nat_plus. f_equal. rewrite Zdouble_mult, Zmult_comm, <- Zplus_diag_eq_mult_2. symmetry; apply Zabs_nat_Zplus; auto with zarith. - change (iter_nat (S (Zabs_nat z + Zabs_nat z)) A f a = + change (iter_nat (S (Zabs_nat z + Zabs_nat z)) A f a = iter_nat (Zabs_nat (Zdouble_plus_one z)) A f a); f_equal. rewrite Zdouble_plus_one_mult, Zmult_comm, <- Zplus_diag_eq_mult_2. rewrite Zabs_nat_Zplus; auto with zarith. @@ -1566,13 +1566,13 @@ Section Int31_Spec. change (Zabs_nat 1) with 1%nat; omega. Qed. - Fixpoint addmuldiv31_alt n i j := - match n with - | O => i + Fixpoint addmuldiv31_alt n i j := + match n with + | O => i | S n => addmuldiv31_alt n (sneakl (firstl j) i) (shiftl j) end. - Lemma addmuldiv31_equiv : forall p x y, + Lemma addmuldiv31_equiv : forall p x y, addmuldiv31 p x y = addmuldiv31_alt (Zabs_nat [|p|]) x y. Proof. intros. @@ -1588,7 +1588,7 @@ Section Int31_Spec. Qed. Lemma spec_add_mul_div : forall x y p, [|p|] <= Zpos 31 -> - [| addmuldiv31 p x y |] = + [| addmuldiv31 p x y |] = ([|x|] * (2 ^ [|p|]) + [|y|] / (2 ^ ((Zpos 31) - [|p|]))) mod wB. Proof. intros. @@ -1626,7 +1626,7 @@ Section Int31_Spec. replace (31-Z_of_nat n) with (Zsucc(31-Zsucc(Z_of_nat n))) by ring. rewrite Zpower_Zsucc, <- Zdiv_Zdiv; auto with zarith. rewrite Zmult_comm, Z_div_mult; auto with zarith. - + rewrite phi_twice_plus_one, Zdouble_plus_one_mult. rewrite phi_twice; auto. change (Zdouble [|y|]) with (2*[|y|]). @@ -1644,7 +1644,7 @@ Section Int31_Spec. unfold wB'. rewrite <- Zpower_Zsucc, <- inj_S by (auto with zarith). f_equal. rewrite H1. - replace wB with (2^(Z_of_nat n)*2^(31-Z_of_nat n)) by + replace wB with (2^(Z_of_nat n)*2^(31-Z_of_nat n)) by (rewrite <- Zpower_exp; auto with zarith; f_equal; unfold size; ring). unfold Zminus; rewrite Zopp_mult_distr_l. rewrite Z_div_plus; auto with zarith. @@ -1669,8 +1669,8 @@ Section Int31_Spec. apply Zlt_le_trans with wB; auto with zarith. apply Zpower_le_monotone; auto with zarith. intros. - case_eq ([|p|] ?= 31); intros; - [ apply H; rewrite (Zcompare_Eq_eq _ _ H0); auto with zarith | | + case_eq ([|p|] ?= 31); intros; + [ apply H; rewrite (Zcompare_Eq_eq _ _ H0); auto with zarith | | apply H; change ([|p|]>31)%Z in H0; auto with zarith ]. change ([|p|]<31) in H0. rewrite spec_add_mul_div by auto with zarith. @@ -1701,16 +1701,16 @@ Section Int31_Spec. simpl; auto. Qed. - Fixpoint head031_alt n x := - match n with + Fixpoint head031_alt n x := + match n with | O => 0%nat - | S n => match firstl x with + | S n => match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0%nat end end. - Lemma head031_equiv : + Lemma head031_equiv : forall x, [|head031 x|] = Z_of_nat (head031_alt size x). Proof. intros. @@ -1720,10 +1720,10 @@ Section Int31_Spec. unfold head031, recl. change On with (phi_inv (Z_of_nat (31-size))). - replace (head031_alt size x) with + replace (head031_alt size x) with (head031_alt size x + (31 - size))%nat by (apply plus_0_r; auto). assert (size <= 31)%nat by auto with arith. - + revert x H; induction size; intros. simpl; auto. unfold recl_aux; fold recl_aux. @@ -1748,7 +1748,7 @@ Section Int31_Spec. change [|In|] with 1. replace (31-n)%nat with (S (31 - S n))%nat by omega. rewrite inj_S; ring. - + clear - H H2. rewrite (sneakr_shiftl x) in H. rewrite H2 in H. @@ -1793,7 +1793,7 @@ Section Int31_Spec. rewrite (sneakr_shiftl x), H1, H; auto. rewrite <- nshiftl_S_tail; auto. - + change (2^(Z_of_nat 0)) with 1; rewrite Zmult_1_l. generalize (phi_bounded x); unfold size; split; auto with zarith. change (2^(Z_of_nat 31)/2) with (2^(Z_of_nat (pred size))). @@ -1809,16 +1809,16 @@ Section Int31_Spec. simpl; auto. Qed. - Fixpoint tail031_alt n x := - match n with + Fixpoint tail031_alt n x := + match n with | O => 0%nat - | S n => match firstr x with + | S n => match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0%nat end end. - Lemma tail031_equiv : + Lemma tail031_equiv : forall x, [|tail031 x|] = Z_of_nat (tail031_alt size x). Proof. intros. @@ -1828,10 +1828,10 @@ Section Int31_Spec. unfold tail031, recr. change On with (phi_inv (Z_of_nat (31-size))). - replace (tail031_alt size x) with + replace (tail031_alt size x) with (tail031_alt size x + (31 - size))%nat by (apply plus_0_r; auto). assert (size <= 31)%nat by auto with arith. - + revert x H; induction size; intros. simpl; auto. unfold recr_aux; fold recr_aux. @@ -1856,7 +1856,7 @@ Section Int31_Spec. change [|In|] with 1. replace (31-n)%nat with (S (31 - S n))%nat by omega. rewrite inj_S; ring. - + clear - H H2. rewrite (sneakl_shiftr x) in H. rewrite H2 in H. @@ -1864,7 +1864,7 @@ Section Int31_Spec. rewrite (iszero_eq0 _ H0) in H; discriminate. Qed. - Lemma spec_tail0 : forall x, 0 < [|x|] -> + Lemma spec_tail0 : forall x, 0 < [|x|] -> exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|tail031 x|]). Proof. intros. @@ -1882,23 +1882,23 @@ Section Int31_Spec. case_eq (firstr x); intros. rewrite (inj_S (tail031_alt n (shiftr x))), Zpower_Zsucc; auto with zarith. destruct (IHn (shiftr x)) as (y & Hy1 & Hy2). - + rewrite phi_nz; rewrite phi_nz in H; contradict H. rewrite (sneakl_shiftr x), H1, H; auto. rewrite <- nshiftr_S_tail; auto. - + exists y; split; auto. rewrite phi_eqn1; auto. rewrite Zdouble_mult, Hy2; ring. - + exists [|shiftr x|]. split. generalize (phi_bounded (shiftr x)); auto with zarith. rewrite phi_eqn2; auto. rewrite Zdouble_plus_one_mult; simpl; ring. Qed. - + (* Sqrt *) (* Direct transcription of an old proof @@ -1906,27 +1906,27 @@ Section Int31_Spec. Lemma quotient_by_2 a: a - 1 <= (a/2) + (a/2). Proof. - intros a; case (Z_mod_lt a 2); auto with zarith. + case (Z_mod_lt a 2); auto with zarith. intros H1; rewrite Zmod_eq_full; auto with zarith. Qed. - Lemma sqrt_main_trick j k: 0 <= j -> 0 <= k -> + Lemma sqrt_main_trick j k: 0 <= j -> 0 <= k -> (j * k) + j <= ((j + k)/2 + 1) ^ 2. Proof. - intros j k Hj; generalize Hj k; pattern j; apply natlike_ind; + intros Hj; generalize Hj k; pattern j; apply natlike_ind; auto; clear k j Hj. intros _ k Hk; repeat rewrite Zplus_0_l. apply Zmult_le_0_compat; generalize (Z_div_pos k 2); auto with zarith. intros j Hj Hrec _ k Hk; pattern k; apply natlike_ind; auto; clear k Hk. rewrite Zmult_0_r, Zplus_0_r, Zplus_0_l. - generalize (sqr_pos (Zsucc j / 2)) (quotient_by_2 (Zsucc j)); + generalize (sqr_pos (Zsucc j / 2)) (quotient_by_2 (Zsucc j)); unfold Zsucc. rewrite Zpower_2, Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r. auto with zarith. intros k Hk _. replace ((Zsucc j + Zsucc k) / 2) with ((j + k)/2 + 1). generalize (Hrec Hj k Hk) (quotient_by_2 (j + k)). - unfold Zsucc; repeat rewrite Zpower_2; + unfold Zsucc; repeat rewrite Zpower_2; repeat rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r. repeat rewrite Zmult_1_l; repeat rewrite Zmult_1_r. auto with zarith. @@ -1936,7 +1936,7 @@ Section Int31_Spec. Lemma sqrt_main i j: 0 <= i -> 0 < j -> i < ((j + (i/j))/2 + 1) ^ 2. Proof. - intros i j Hi Hj. + intros Hi Hj. assert (Hij: 0 <= i/j) by (apply Z_div_pos; auto with zarith). apply Zlt_le_trans with (2 := sqrt_main_trick _ _ (Zlt_le_weak _ _ Hj) Hij). pattern i at 1; rewrite (Z_div_mod_eq i j); case (Z_mod_lt i j); auto with zarith. @@ -1944,7 +1944,7 @@ Section Int31_Spec. Lemma sqrt_init i: 1 < i -> i < (i/2 + 1) ^ 2. Proof. - intros i Hi. + intros Hi. assert (H1: 0 <= i - 2) by auto with zarith. assert (H2: 1 <= (i / 2) ^ 2); auto with zarith. replace i with (1* 2 + (i - 2)); auto with zarith. @@ -1962,14 +1962,14 @@ Section Int31_Spec. Lemma sqrt_test_true i j: 0 <= i -> 0 < j -> i/j >= j -> j ^ 2 <= i. Proof. - intros i j Hi Hj Hd; rewrite Zpower_2. + intros Hi Hj Hd; rewrite Zpower_2. apply Zle_trans with (j * (i/j)); auto with zarith. apply Z_mult_div_ge; auto with zarith. Qed. Lemma sqrt_test_false i j: 0 <= i -> 0 < j -> i/j < j -> (j + (i/j))/2 < j. Proof. - intros i j Hi Hj H; case (Zle_or_lt j ((j + (i/j))/2)); auto. + intros Hi Hj H; case (Zle_or_lt j ((j + (i/j))/2)); auto. intros H1; contradict H; apply Zle_not_lt. assert (2 * j <= j + (i/j)); auto with zarith. apply Zle_trans with (2 * ((j + (i/j))/2)); auto with zarith. @@ -1984,32 +1984,32 @@ Section Int31_Spec. Lemma Zcompare_spec i j: ZcompareSpec i j (i ?= j). Proof. - intros i j; case_eq (Zcompare i j); intros H. + case_eq (Zcompare i j); intros H. apply ZcompareSpecEq; apply Zcompare_Eq_eq; auto. apply ZcompareSpecLt; auto. apply ZcompareSpecGt; apply Zgt_lt; auto. Qed. Lemma sqrt31_step_def rec i j: - sqrt31_step rec i j = + sqrt31_step rec i j = match (fst (i/j) ?= j)%int31 with Lt => rec i (fst ((j + fst(i/j))/2))%int31 | _ => j end. Proof. - intros rec i j; unfold sqrt31_step; case div31; intros. + unfold sqrt31_step; case div31; intros. simpl; case compare31; auto. Qed. Lemma div31_phi i j: 0 < [|j|] -> [|fst (i/j)%int31|] = [|i|]/[|j|]. - intros i j Hj; generalize (spec_div i j Hj). + intros Hj; generalize (spec_div i j Hj). case div31; intros q r; simpl fst. intros (H1,H2); apply Zdiv_unique with [|r|]; auto with zarith. rewrite H1; ring. Qed. - Lemma sqrt31_step_correct rec i j: - 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> + Lemma sqrt31_step_correct rec i j: + 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> @@ -2017,15 +2017,15 @@ Section Int31_Spec. [|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2. Proof. assert (Hp2: 0 < [|2|]) by exact (refl_equal Lt). - intros rec i j Hi Hj Hij H31 Hrec; rewrite sqrt31_step_def. - generalize (spec_compare (fst (i/j)%int31) j); case compare31; + intros Hi Hj Hij H31 Hrec; rewrite sqrt31_step_def. + generalize (spec_compare (fst (i/j)%int31) j); case compare31; rewrite div31_phi; auto; intros Hc; try (split; auto; apply sqrt_test_true; auto with zarith; fail). apply Hrec; repeat rewrite div31_phi; auto with zarith. replace [|(j + fst (i / j)%int31)|] with ([|j|] + [|i|] / [|j|]). split. case (Zle_lt_or_eq 1 [|j|]); auto with zarith; intros Hj1. - replace ([|j|] + [|i|]/[|j|]) with + replace ([|j|] + [|i|]/[|j|]) with (1 * 2 + (([|j|] - 2) + [|i|] / [|j|])); try ring. rewrite Z_div_plus_full_l; auto with zarith. assert (0 <= [|i|]/ [|j|]) by (apply Z_div_pos; auto with zarith). @@ -2048,12 +2048,12 @@ Section Int31_Spec. Lemma iter31_sqrt_correct n rec i j: 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < 2 ^ (Z_of_nat size) -> - (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] -> + (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < 2 ^ (Z_of_nat size) -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2. Proof. - intros n; elim n; unfold iter31_sqrt; fold iter31_sqrt; clear n. + revert rec i j; elim n; unfold iter31_sqrt; fold iter31_sqrt; clear n. intros rec i j Hi Hj Hij H31 Hrec; apply sqrt31_step_correct; auto with zarith. intros; apply Hrec; auto with zarith. rewrite Zpower_0_r; auto with zarith. @@ -2098,7 +2098,7 @@ Section Int31_Spec. Qed. Lemma sqrt312_step_def rec ih il j: - sqrt312_step rec ih il j = + sqrt312_step rec ih il j = match (ih ?= j)%int31 with Eq => j | Gt => j @@ -2112,14 +2112,14 @@ Section Int31_Spec. end end. Proof. - intros rec ih il j; unfold sqrt312_step; case div3121; intros. + unfold sqrt312_step; case div3121; intros. simpl; case compare31; auto. Qed. - Lemma sqrt312_lower_bound ih il j: + Lemma sqrt312_lower_bound ih il j: phi2 ih il < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|]. Proof. - intros ih il j H1. + intros H1. case (phi_bounded j); intros Hbj _. case (phi_bounded il); intros Hbil _. case (phi_bounded ih); intros Hbih Hbih1. @@ -2133,22 +2133,22 @@ Section Int31_Spec. Lemma div312_phi ih il j: (2^30 <= [|j|] -> [|ih|] < [|j|] -> [|fst (div3121 ih il j)|] = phi2 ih il/[|j|])%Z. Proof. - intros ih il j Hj Hj1. + intros Hj Hj1. generalize (spec_div21 ih il j Hj Hj1). case div3121; intros q r (Hq, Hr). apply Zdiv_unique with (phi r); auto with zarith. simpl fst; apply trans_equal with (1 := Hq); ring. Qed. - Lemma sqrt312_step_correct rec ih il j: - 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> + Lemma sqrt312_step_correct rec ih il j: + 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> - [|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il + [|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step rec ih il j|] + 1) ^ 2. Proof. assert (Hp2: (0 < [|2|])%Z) by exact (refl_equal Lt). - intros rec ih il j Hih Hj Hij Hrec; rewrite sqrt312_step_def. + intros Hih Hj Hij Hrec; rewrite sqrt312_step_def. assert (H1: ([|ih|] <= [|j|])%Z) by (apply sqrt312_lower_bound with il; auto). case (phi_bounded ih); intros Hih1 _. case (phi_bounded il); intros Hil1 _. @@ -2174,7 +2174,7 @@ Section Int31_Spec. case (Zle_lt_or_eq 1 ([|j|])); auto with zarith; intros Hf2. 2: contradict Hc; apply Zle_not_lt; rewrite <- Hf2, Zdiv_1_r; auto with zarith. assert (Hf3: 0 < ([|j|] + phi2 ih il / [|j|]) / 2). - replace ([|j|] + phi2 ih il/ [|j|])%Z with + replace ([|j|] + phi2 ih il/ [|j|])%Z with (1 * 2 + (([|j|] - 2) + phi2 ih il / [|j|])); try ring. rewrite Z_div_plus_full_l; auto with zarith. assert (0 <= ([|j|] - 2 + phi2 ih il / [|j|]) / 2) ; auto with zarith. @@ -2213,7 +2213,7 @@ Section Int31_Spec. rewrite div31_phi; change (phi 2) with 2%Z; auto. change (2 ^Z_of_nat size) with (base/2 + phi v30). assert (phi r / 2 < base/2); auto with zarith. - apply Zmult_gt_0_lt_reg_r with 2; auto with zarith. + apply Zmult_gt_0_lt_reg_r with 2; auto with zarith. change (base/2 * 2) with base. apply Zle_lt_trans with (phi r). rewrite Zmult_comm; apply Z_mult_div_ge; auto with zarith. @@ -2234,15 +2234,15 @@ Section Int31_Spec. apply Zge_le; apply Z_div_ge; auto with zarith. Qed. - Lemma iter312_sqrt_correct n rec ih il j: - 2^29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> - (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] -> - phi2 ih il < ([|j1|] + 1) ^ 2 -> + Lemma iter312_sqrt_correct n rec ih il j: + 2^29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> + (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] -> + phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> - [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il + [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j|] + 1) ^ 2. Proof. - intros n; elim n; unfold iter312_sqrt; fold iter312_sqrt; clear n. + revert rec ih il j; elim n; unfold iter312_sqrt; fold iter312_sqrt; clear n. intros rec ih il j Hi Hj Hij Hrec; apply sqrt312_step_correct; auto with zarith. intros; apply Hrec; auto with zarith. rewrite Zpower_0_r; auto with zarith. @@ -2265,7 +2265,7 @@ Section Int31_Spec. Proof. intros ih il Hih; unfold sqrt312. change [||WW ih il||] with (phi2 ih il). - assert (Hbin: forall s, s * s + 2* s + 1 = (s + 1) ^ 2) by + assert (Hbin: forall s, s * s + 2* s + 1 = (s + 1) ^ 2) by (intros s; ring). assert (Hb: 0 <= base) by (red; intros HH; discriminate). assert (Hi2: phi2 ih il < (phi Tn + 1) ^ 2). @@ -2428,9 +2428,9 @@ Section Int31_Spec. apply Zcompare_Eq_eq. now destruct ([|x|] ?= 0). Qed. - + (* Even *) - + Let w_is_even := int31_op.(znz_is_even). Lemma spec_is_even : forall x, @@ -2460,13 +2460,13 @@ Section Int31_Spec. exact spec_more_than_1_digit. exact spec_0. - exact spec_1. + exact spec_1. exact spec_Bm1. exact spec_compare. exact spec_eq0. - exact spec_opp_c. + exact spec_opp_c. exact spec_opp. exact spec_opp_carry. @@ -2500,7 +2500,7 @@ Section Int31_Spec. exact spec_head00. exact spec_head0. - exact spec_tail00. + exact spec_tail00. exact spec_tail0. exact spec_add_mul_div. diff --git a/theories/Numbers/Cyclic/Int31/Int31.v b/theories/Numbers/Cyclic/Int31/Int31.v index 154b436b..cc224254 100644 --- a/theories/Numbers/Cyclic/Int31/Int31.v +++ b/theories/Numbers/Cyclic/Int31/Int31.v @@ -8,7 +8,7 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) -(*i $Id: Int31.v 11072 2008-06-08 16:13:37Z herbelin $ i*) +(*i $Id$ i*) Require Import NaryFunctions. Require Import Wf_nat. @@ -17,7 +17,7 @@ Require Export DoubleType. Unset Boxed Definitions. -(** * 31-bit integers *) +(** * 31-bit integers *) (** This file contains basic definitions of a 31-bit integer arithmetic. In fact it is more general than that. The only reason @@ -36,11 +36,13 @@ Definition size := 31%nat. Inductive digits : Type := D0 | D1. (** The type of 31-bit integers *) - -(** The type [int31] has a unique constructor [I31] that expects + +(** The type [int31] has a unique constructor [I31] that expects 31 arguments of type [digits]. *) -Inductive int31 : Type := I31 : nfun digits size int31. +Definition digits31 t := Eval compute in nfun digits size t. + +Inductive int31 : Type := I31 : digits31 int31. (* spiwack: Registration of the type of integers, so that the matchs in the functions below perform dynamic decompilation (otherwise some segfault @@ -50,7 +52,7 @@ Register int31 as int31 type in "coq_int31" by True. Delimit Scope int31_scope with int31. Bind Scope int31_scope with int31. -Open Scope int31_scope. +Local Open Scope int31_scope. (** * Constants *) @@ -69,26 +71,26 @@ Definition Twon : int31 := Eval compute in (napply_cst _ _ D0 (size-2) I31) D1 D (** * Bits manipulation *) -(** [sneakr b x] shifts [x] to the right by one bit. +(** [sneakr b x] shifts [x] to the right by one bit. Rightmost digit is lost while leftmost digit becomes [b]. - Pseudo-code is + Pseudo-code is [ match x with (I31 d0 ... dN) => I31 b d0 ... d(N-1) end ] *) Definition sneakr : digits -> int31 -> int31 := Eval compute in fun b => int31_rect _ (napply_except_last _ _ (size-1) (I31 b)). -(** [sneakl b x] shifts [x] to the left by one bit. +(** [sneakl b x] shifts [x] to the left by one bit. Leftmost digit is lost while rightmost digit becomes [b]. - Pseudo-code is + Pseudo-code is [ match x with (I31 d0 ... dN) => I31 d1 ... dN b end ] *) -Definition sneakl : digits -> int31 -> int31 := Eval compute in +Definition sneakl : digits -> int31 -> int31 := Eval compute in fun b => int31_rect _ (fun _ => napply_then_last _ _ b (size-1) I31). -(** [shiftl], [shiftr], [twice] and [twice_plus_one] are direct +(** [shiftl], [shiftr], [twice] and [twice_plus_one] are direct consequences of [sneakl] and [sneakr]. *) Definition shiftl := sneakl D0. @@ -96,31 +98,31 @@ Definition shiftr := sneakr D0. Definition twice := sneakl D0. Definition twice_plus_one := sneakl D1. -(** [firstl x] returns the leftmost digit of number [x]. +(** [firstl x] returns the leftmost digit of number [x]. Pseudo-code is [ match x with (I31 d0 ... dN) => d0 end ] *) -Definition firstl : int31 -> digits := Eval compute in +Definition firstl : int31 -> digits := Eval compute in int31_rect _ (fun d => napply_discard _ _ d (size-1)). -(** [firstr x] returns the rightmost digit of number [x]. +(** [firstr x] returns the rightmost digit of number [x]. Pseudo-code is [ match x with (I31 d0 ... dN) => dN end ] *) -Definition firstr : int31 -> digits := Eval compute in +Definition firstr : int31 -> digits := Eval compute in int31_rect _ (napply_discard _ _ (fun d=>d) (size-1)). -(** [iszero x] is true iff [x = I31 D0 ... D0]. Pseudo-code is +(** [iszero x] is true iff [x = I31 D0 ... D0]. Pseudo-code is [ match x with (I31 D0 ... D0) => true | _ => false end ] *) -Definition iszero : int31 -> bool := Eval compute in - let f d b := match d with D0 => b | D1 => false end +Definition iszero : int31 -> bool := Eval compute in + let f d b := match d with D0 => b | D1 => false end in int31_rect _ (nfold_bis _ _ f true size). -(* NB: DO NOT transform the above match in a nicer (if then else). +(* NB: DO NOT transform the above match in a nicer (if then else). It seems to work, but later "unfold iszero" takes forever. *) -(** [base] is [2^31], obtained via iterations of [Zdouble]. - It can also be seen as the smallest b > 0 s.t. phi_inv b = 0 +(** [base] is [2^31], obtained via iterations of [Zdouble]. + It can also be seen as the smallest b > 0 s.t. phi_inv b = 0 (see below) *) Definition base := Eval compute in @@ -140,7 +142,7 @@ Fixpoint recl_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A) caserec (firstl i) si (recl_aux next A case0 caserec si) end. -Fixpoint recr_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A) +Fixpoint recr_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A) (i:int31) : A := match n with | O => case0 @@ -159,22 +161,22 @@ Definition recr := recr_aux size. (** From int31 to Z, we simply iterates [Zdouble] or [Zdouble_plus_one]. *) -Definition phi : int31 -> Z := +Definition phi : int31 -> Z := recr Z (0%Z) (fun b _ => match b with D0 => Zdouble | D1 => Zdouble_plus_one end). -(** From positive to int31. An abstract definition could be : - [ phi_inv (2n) = 2*(phi_inv n) /\ +(** From positive to int31. An abstract definition could be : + [ phi_inv (2n) = 2*(phi_inv n) /\ phi_inv 2n+1 = 2*(phi_inv n) + 1 ] *) -Fixpoint phi_inv_positive p := +Fixpoint phi_inv_positive p := match p with | xI q => twice_plus_one (phi_inv_positive q) | xO q => twice (phi_inv_positive q) | xH => In end. -(** The negative part : 2-complement *) +(** The negative part : 2-complement *) Fixpoint complement_negative p := match p with @@ -186,9 +188,9 @@ Fixpoint complement_negative p := (** A simple incrementation function *) Definition incr : int31 -> int31 := - recr int31 In - (fun b si rec => match b with - | D0 => sneakl D1 si + recr int31 In + (fun b si rec => match b with + | D0 => sneakl D1 si | D1 => sneakl D0 rec end). (** We can now define the conversion from Z to int31. *) @@ -196,11 +198,11 @@ Definition incr : int31 -> int31 := Definition phi_inv : Z -> int31 := fun n => match n with | Z0 => On - | Zpos p => phi_inv_positive p + | Zpos p => phi_inv_positive p | Zneg p => incr (complement_negative p) end. -(** [phi_inv2] is similar to [phi_inv] but returns a double word +(** [phi_inv2] is similar to [phi_inv] but returns a double word [zn2z int31] *) Definition phi_inv2 n := @@ -211,7 +213,7 @@ Definition phi_inv2 n := (** [phi2] is similar to [phi] but takes a double word (two args) *) -Definition phi2 nh nl := +Definition phi2 nh nl := ((phi nh)*base+(phi nl))%Z. (** * Addition *) @@ -227,11 +229,11 @@ Notation "n + m" := (add31 n m) : int31_scope. (* mode, (phi n)+(phi m) is computed twice*) (* it may be considered to optimize it *) -Definition add31c (n m : int31) := +Definition add31c (n m : int31) := let npm := n+m in - match (phi npm ?= (phi n)+(phi m))%Z with - | Eq => C0 npm - | _ => C1 npm + match (phi npm ?= (phi n)+(phi m))%Z with + | Eq => C0 npm + | _ => C1 npm end. Notation "n '+c' m" := (add31c n m) (at level 50, no associativity) : int31_scope. @@ -254,7 +256,7 @@ Notation "n - m" := (sub31 n m) : int31_scope. (** Subtraction with carry (thus exact) *) -Definition sub31c (n m : int31) := +Definition sub31c (n m : int31) := let nmm := n-m in match (phi nmm ?= (phi n)-(phi m))%Z with | Eq => C0 nmm @@ -272,6 +274,10 @@ Definition sub31carryc (n m : int31) := | _ => C1 nmmmone end. +(** Opposite *) + +Definition opp31 x := On - x. +Notation "- x" := (opp31 x) : int31_scope. (** Multiplication *) @@ -290,13 +296,13 @@ Notation "n '*c' m" := (mul31c n m) (at level 40, no associativity) : int31_scop (** Division of a double size word modulo [2^31] *) -Definition div3121 (nh nl m : int31) := +Definition div3121 (nh nl m : int31) := let (q,r) := Zdiv_eucl (phi2 nh nl) (phi m) in (phi_inv q, phi_inv r). (** Division modulo [2^31] *) -Definition div31 (n m : int31) := +Definition div31 (n m : int31) := let (q,r) := Zdiv_eucl (phi n) (phi m) in (phi_inv q, phi_inv r). Notation "n / m" := (div31 n m) : int31_scope. @@ -307,13 +313,16 @@ Notation "n / m" := (div31 n m) : int31_scope. Definition compare31 (n m : int31) := ((phi n)?=(phi m))%Z. Notation "n ?= m" := (compare31 n m) (at level 70, no associativity) : int31_scope. +Definition eqb31 (n m : int31) := + match n ?= m with Eq => true | _ => false end. + -(** Computing the [i]-th iterate of a function: +(** Computing the [i]-th iterate of a function: [iter_int31 i A f = f^i] *) Definition iter_int31 i A f := - recr (A->A) (fun x => x) - (fun b si rec => match b with + recr (A->A) (fun x => x) + (fun b si rec => match b with | D0 => fun x => rec (rec x) | D1 => fun x => f (rec (rec x)) end) @@ -322,9 +331,9 @@ Definition iter_int31 i A f := (** Combining the [(31-p)] low bits of [i] above the [p] high bits of [j]: [addmuldiv31 p i j = i*2^p+j/2^(31-p)] (modulo [2^31]) *) -Definition addmuldiv31 p i j := - let (res, _ ) := - iter_int31 p (int31*int31) +Definition addmuldiv31 p i j := + let (res, _ ) := + iter_int31 p (int31*int31) (fun ij => let (i,j) := ij in (sneakl (firstl j) i, shiftl j)) (i,j) in @@ -346,7 +355,7 @@ Register addmuldiv31 as int31 addmuldiv in "coq_int31" by True. Definition gcd31 (i j:int31) := (fix euler (guard:nat) (i j:int31) {struct guard} := - match guard with + match guard with | O => In | S p => match j ?= On with | Eq => i @@ -370,17 +379,17 @@ Eval lazy delta [Twon] in | _ => j end. -Fixpoint iter31_sqrt (n: nat) (rec: int31 -> int31 -> int31) +Fixpoint iter31_sqrt (n: nat) (rec: int31 -> int31 -> int31) (i j: int31) {struct n} : int31 := - sqrt31_step + sqrt31_step (match n with O => rec | S n => (iter31_sqrt n (iter31_sqrt n rec)) end) i j. -Definition sqrt31 i := +Definition sqrt31 i := Eval lazy delta [On In Twon] in - match compare31 In i with + match compare31 In i with Gt => On | Eq => In | Lt => iter31_sqrt 31 (fun i j => j) i (fst (i/Twon)) @@ -388,7 +397,7 @@ Eval lazy delta [On In Twon] in Definition v30 := Eval compute in (addmuldiv31 (phi_inv (Z_of_nat size - 1)) In On). -Definition sqrt312_step (rec: int31 -> int31 -> int31 -> int31) +Definition sqrt312_step (rec: int31 -> int31 -> int31 -> int31) (ih il j: int31) := Eval lazy delta [Twon v30] in match ih ?= j with Eq => j | Gt => j | _ => @@ -401,28 +410,28 @@ Eval lazy delta [Twon v30] in | _ => j end end. -Fixpoint iter312_sqrt (n: nat) - (rec: int31 -> int31 -> int31 -> int31) +Fixpoint iter312_sqrt (n: nat) + (rec: int31 -> int31 -> int31 -> int31) (ih il j: int31) {struct n} : int31 := - sqrt312_step + sqrt312_step (match n with O => rec | S n => (iter312_sqrt n (iter312_sqrt n rec)) end) ih il j. -Definition sqrt312 ih il := +Definition sqrt312 ih il := Eval lazy delta [On In] in let s := iter312_sqrt 31 (fun ih il j => j) ih il Tn in match s *c s with W0 => (On, C0 On) (* impossible *) | WW ih1 il1 => match il -c il1 with - C0 il2 => + C0 il2 => match ih ?= ih1 with Gt => (s, C1 il2) | _ => (s, C0 il2) end - | C1 il2 => + | C1 il2 => match (ih - In) ?= ih1 with (* we could parametrize ih - 1 *) Gt => (s, C1 il2) | _ => (s, C0 il2) @@ -431,7 +440,7 @@ Eval lazy delta [On In] in end. -Fixpoint p2i n p : (N*int31)%type := +Fixpoint p2i n p : (N*int31)%type := match n with | O => (Npos p, On) | S n => match p with @@ -444,26 +453,26 @@ Fixpoint p2i n p : (N*int31)%type := Definition positive_to_int31 (p:positive) := p2i size p. (** Constant 31 converted into type int31. - It is used as default answer for numbers of zeros + It is used as default answer for numbers of zeros in [head0] and [tail0] *) Definition T31 : int31 := Eval compute in phi_inv (Z_of_nat size). Definition head031 (i:int31) := - recl _ (fun _ => T31) - (fun b si rec n => match b with + recl _ (fun _ => T31) + (fun b si rec n => match b with | D0 => rec (add31 n In) | D1 => n end) i On. Definition tail031 (i:int31) := - recr _ (fun _ => T31) - (fun b si rec n => match b with + recr _ (fun _ => T31) + (fun b si rec n => match b with | D0 => rec (add31 n In) | D1 => n end) i On. Register head031 as int31 head0 in "coq_int31" by True. -Register tail031 as int31 tail0 in "coq_int31" by True. +Register tail031 as int31 tail0 in "coq_int31" by True. diff --git a/theories/Numbers/Cyclic/Int31/Ring31.v b/theories/Numbers/Cyclic/Int31/Ring31.v new file mode 100644 index 00000000..2ec406b0 --- /dev/null +++ b/theories/Numbers/Cyclic/Int31/Ring31.v @@ -0,0 +1,103 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(*i $Id$ i*) + +(** * Int31 numbers defines Z/(2^31)Z, and can hence be equipped + with a ring structure and a ring tactic *) + +Require Import Int31 Cyclic31 CyclicAxioms. + +Local Open Scope int31_scope. + +(** Detection of constants *) + +Local Open Scope list_scope. + +Ltac isInt31cst_lst l := + match l with + | nil => constr:true + | ?t::?l => match t with + | D1 => isInt31cst_lst l + | D0 => isInt31cst_lst l + | _ => constr:false + end + | _ => constr:false + end. + +Ltac isInt31cst t := + match t with + | I31 ?i0 ?i1 ?i2 ?i3 ?i4 ?i5 ?i6 ?i7 ?i8 ?i9 ?i10 + ?i11 ?i12 ?i13 ?i14 ?i15 ?i16 ?i17 ?i18 ?i19 ?i20 + ?i21 ?i22 ?i23 ?i24 ?i25 ?i26 ?i27 ?i28 ?i29 ?i30 => + let l := + constr:(i0::i1::i2::i3::i4::i5::i6::i7::i8::i9::i10 + ::i11::i12::i13::i14::i15::i16::i17::i18::i19::i20 + ::i21::i22::i23::i24::i25::i26::i27::i28::i29::i30::nil) + in isInt31cst_lst l + | Int31.On => constr:true + | Int31.In => constr:true + | Int31.Tn => constr:true + | Int31.Twon => constr:true + | _ => constr:false + end. + +Ltac Int31cst t := + match isInt31cst t with + | true => constr:t + | false => constr:NotConstant + end. + +(** The generic ring structure inferred from the Cyclic structure *) + +Module Int31ring := CyclicRing Int31Cyclic. + +(** Unlike in the generic [CyclicRing], we can use Leibniz here. *) + +Lemma Int31_canonic : forall x y, phi x = phi y -> x = y. +Proof. + intros x y EQ. + now rewrite <- (phi_inv_phi x), <- (phi_inv_phi y), EQ. +Qed. + +Lemma ring_theory_switch_eq : + forall A (R R':A->A->Prop) zero one add mul sub opp, + (forall x y : A, R x y -> R' x y) -> + ring_theory zero one add mul sub opp R -> + ring_theory zero one add mul sub opp R'. +Proof. +intros A R R' zero one add mul sub opp Impl Ring. +constructor; intros; apply Impl; apply Ring. +Qed. + +Lemma Int31Ring : ring_theory 0 1 add31 mul31 sub31 opp31 Logic.eq. +Proof. +exact (ring_theory_switch_eq _ _ _ _ _ _ _ _ _ Int31_canonic Int31ring.CyclicRing). +Qed. + +Lemma eqb31_eq : forall x y, eqb31 x y = true <-> x=y. +Proof. +unfold eqb31. intros x y. +generalize (Cyclic31.spec_compare x y). +destruct (x ?= y); intuition; subst; auto with zarith; try discriminate. +apply Int31_canonic; auto. +Qed. + +Lemma eqb31_correct : forall x y, eqb31 x y = true -> x=y. +Proof. now apply eqb31_eq. Qed. + +Add Ring Int31Ring : Int31Ring + (decidable eqb31_correct, + constants [Int31cst]). + +Section TestRing. +Let test : forall x y, 1 + x*y + x*x + 1 = 1*1 + 1 + y*x + 1*x*x. +intros. ring. +Qed. +End TestRing. + diff --git a/theories/Numbers/Cyclic/ZModulo/ZModulo.v b/theories/Numbers/Cyclic/ZModulo/ZModulo.v index 7c770e97..4f0f6c7c 100644 --- a/theories/Numbers/Cyclic/ZModulo/ZModulo.v +++ b/theories/Numbers/Cyclic/ZModulo/ZModulo.v @@ -6,13 +6,13 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* $Id: ZModulo.v 11033 2008-06-01 22:56:50Z letouzey $ *) +(* $Id$ *) -(** * Type [Z] viewed modulo a particular constant corresponds to [Z/nZ] +(** * Type [Z] viewed modulo a particular constant corresponds to [Z/nZ] as defined abstractly in CyclicAxioms. *) -(** Even if the construction provided here is not reused for building - the efficient arbitrary precision numbers, it provides a simple +(** Even if the construction provided here is not reused for building + the efficient arbitrary precision numbers, it provides a simple implementation of CyclicAxioms, hence ensuring its coherence. *) Set Implicit Arguments. @@ -24,7 +24,7 @@ Require Import BigNumPrelude. Require Import DoubleType. Require Import CyclicAxioms. -Open Local Scope Z_scope. +Local Open Scope Z_scope. Section ZModulo. @@ -56,9 +56,9 @@ Section ZModulo. destruct 1; auto. Qed. Let digits_gt_1 := spec_more_than_1_digit. - + Lemma wB_pos : wB > 0. - Proof. + Proof. unfold wB, base; auto with zarith. Qed. Hint Resolve wB_pos. @@ -79,7 +79,7 @@ Section ZModulo. auto. Qed. - Definition znz_of_pos x := + Definition znz_of_pos x := let (q,r) := Zdiv_eucl_POS x wB in (N_of_Z q, r). Lemma spec_of_pos : forall p, @@ -90,10 +90,10 @@ Section ZModulo. destruct (Zdiv_eucl_POS p wB); simpl; destruct 1. unfold znz_to_Z; rewrite Zmod_small; auto. assert (0 <= z). - replace z with (Zpos p / wB) by + replace z with (Zpos p / wB) by (symmetry; apply Zdiv_unique with z0; auto). apply Z_div_pos; auto with zarith. - replace (Z_of_N (N_of_Z z)) with z by + replace (Z_of_N (N_of_Z z)) with z by (destruct z; simpl; auto; elim H1; auto). rewrite Zmult_comm; auto. Qed. @@ -110,7 +110,7 @@ Section ZModulo. Definition znz_0 := 0. Definition znz_1 := 1. Definition znz_Bm1 := wB - 1. - + Lemma spec_0 : [|znz_0|] = 0. Proof. unfold znz_to_Z, znz_0. @@ -121,7 +121,7 @@ Section ZModulo. Proof. unfold znz_to_Z, znz_1. apply Zmod_small; split; auto with zarith. - unfold wB, base. + unfold wB, base. apply Zlt_trans with (Zpos digits); auto. apply Zpower2_lt_lin; auto with zarith. Qed. @@ -138,7 +138,7 @@ Section ZModulo. Definition znz_compare x y := Zcompare [|x|] [|y|]. - Lemma spec_compare : forall x y, + Lemma spec_compare : forall x y, match znz_compare x y with | Eq => [|x|] = [|y|] | Lt => [|x|] < [|y|] @@ -150,19 +150,19 @@ Section ZModulo. intros; apply Zcompare_Eq_eq; auto. Qed. - Definition znz_eq0 x := + Definition znz_eq0 x := match [|x|] with Z0 => true | _ => false end. - + Lemma spec_eq0 : forall x, znz_eq0 x = true -> [|x|] = 0. Proof. unfold znz_eq0; intros; now destruct [|x|]. Qed. - Definition znz_opp_c x := + Definition znz_opp_c x := if znz_eq0 x then C0 0 else C1 (- x). Definition znz_opp x := - x. Definition znz_opp_carry x := - x - 1. - + Lemma spec_opp_c : forall x, [-|znz_opp_c x|] = -[|x|]. Proof. intros; unfold znz_opp_c, znz_to_Z; auto. @@ -180,7 +180,7 @@ Section ZModulo. change ((- x) mod wB = (0 - (x mod wB)) mod wB). rewrite Zminus_mod_idemp_r; simpl; auto. Qed. - + Lemma spec_opp_carry : forall x, [|znz_opp_carry x|] = wB - [|x|] - 1. Proof. intros; unfold znz_opp_carry, znz_to_Z; auto. @@ -194,15 +194,15 @@ Section ZModulo. generalize (Z_mod_lt x wB wB_pos); omega. Qed. - Definition znz_succ_c x := - let y := Zsucc x in + Definition znz_succ_c x := + let y := Zsucc x in if znz_eq0 y then C1 0 else C0 y. - Definition znz_add_c x y := - let z := [|x|] + [|y|] in + Definition znz_add_c x y := + let z := [|x|] + [|y|] in if Z_lt_le_dec z wB then C0 z else C1 (z-wB). - Definition znz_add_carry_c x y := + Definition znz_add_carry_c x y := let z := [|x|]+[|y|]+1 in if Z_lt_le_dec z wB then C0 z else C1 (z-wB). @@ -210,7 +210,7 @@ Section ZModulo. Definition znz_add := Zplus. Definition znz_add_carry x y := x + y + 1. - Lemma Zmod_equal : + Lemma Zmod_equal : forall x y z, z>0 -> (x-y) mod z = 0 -> x mod z = y mod z. Proof. intros. @@ -225,12 +225,12 @@ Section ZModulo. Proof. intros; unfold znz_succ_c, znz_to_Z, Zsucc. case_eq (znz_eq0 (x+1)); intros; unfold interp_carry. - + rewrite Zmult_1_l. replace (wB + 0 mod wB) with wB by auto with zarith. symmetry; rewrite Zeq_plus_swap. assert ((x+1) mod wB = 0) by (apply spec_eq0; auto). - replace (wB-1) with ((wB-1) mod wB) by + replace (wB-1) with ((wB-1) mod wB) by (apply Zmod_small; generalize wB_pos; omega). rewrite <- Zminus_mod_idemp_l; rewrite Z_mod_same; simpl; auto. apply Zmod_equal; auto. @@ -289,15 +289,15 @@ Section ZModulo. rewrite Zplus_mod_idemp_l; auto. Qed. - Definition znz_pred_c x := + Definition znz_pred_c x := if znz_eq0 x then C1 (wB-1) else C0 (x-1). - Definition znz_sub_c x y := - let z := [|x|]-[|y|] in + Definition znz_sub_c x y := + let z := [|x|]-[|y|] in if Z_lt_le_dec z 0 then C1 (wB+z) else C0 z. - Definition znz_sub_carry_c x y := - let z := [|x|]-[|y|]-1 in + Definition znz_sub_carry_c x y := + let z := [|x|]-[|y|]-1 in if Z_lt_le_dec z 0 then C1 (wB+z) else C0 z. Definition znz_pred := Zpred. @@ -323,7 +323,7 @@ Section ZModulo. Proof. intros; unfold znz_sub_c, znz_to_Z, interp_carry. destruct Z_lt_le_dec. - replace ((wB + (x mod wB - y mod wB)) mod wB) with + replace ((wB + (x mod wB - y mod wB)) mod wB) with (wB + (x mod wB - y mod wB)). omega. symmetry; apply Zmod_small. @@ -337,7 +337,7 @@ Section ZModulo. Proof. intros; unfold znz_sub_carry_c, znz_to_Z, interp_carry. destruct Z_lt_le_dec. - replace ((wB + (x mod wB - y mod wB - 1)) mod wB) with + replace ((wB + (x mod wB - y mod wB - 1)) mod wB) with (wB + (x mod wB - y mod wB -1)). omega. symmetry; apply Zmod_small. @@ -358,7 +358,7 @@ Section ZModulo. intros; unfold znz_sub, znz_to_Z; apply Zminus_mod. Qed. - Lemma spec_sub_carry : + Lemma spec_sub_carry : forall x y, [|znz_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB. Proof. intros; unfold znz_sub_carry, znz_to_Z. @@ -367,15 +367,15 @@ Section ZModulo. rewrite Zminus_mod_idemp_l. auto. Qed. - - Definition znz_mul_c x y := + + Definition znz_mul_c x y := let (h,l) := Zdiv_eucl ([|x|]*[|y|]) wB in if znz_eq0 h then if znz_eq0 l then W0 else WW h l else WW h l. Definition znz_mul := Zmult. Definition znz_square_c x := znz_mul_c x x. - + Lemma spec_mul_c : forall x y, [|| znz_mul_c x y ||] = [|x|] * [|y|]. Proof. intros; unfold znz_mul_c, zn2z_to_Z. @@ -426,7 +426,7 @@ Section ZModulo. destruct Zdiv_eucl as (q,r); destruct 1; intros. injection H1; clear H1; intros. assert ([|r|]=r). - apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|]; + apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|]; auto with zarith. assert ([|q|]=q). apply Zmod_small. @@ -453,7 +453,7 @@ Section ZModulo. Definition znz_mod x y := [|x|] mod [|y|]. Definition znz_mod_gt x y := [|x|] mod [|y|]. - + Lemma spec_mod : forall a b, 0 < [|b|] -> [|znz_mod a b|] = [|a|] mod [|b|]. Proof. @@ -469,7 +469,7 @@ Section ZModulo. Proof. intros; apply spec_mod; auto. Qed. - + Definition znz_gcd x y := Zgcd [|x|] [|y|]. Definition znz_gcd_gt x y := Zgcd [|x|] [|y|]. @@ -516,7 +516,7 @@ Section ZModulo. intros. apply spec_gcd; auto. Qed. - Definition znz_div21 a1 a2 b := + Definition znz_div21 a1 a2 b := Zdiv_eucl ([|a1|]*wB+[|a2|]) [|b|]. Lemma spec_div21 : forall a1 a2 b, @@ -537,7 +537,7 @@ Section ZModulo. destruct Zdiv_eucl as (q,r); destruct 1; intros. injection H4; clear H4; intros. assert ([|r|]=r). - apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|]; + apply Zmod_small; generalize (Z_mod_lt b wB wB_pos); fold [|b|]; auto with zarith. assert ([|q|]=q). apply Zmod_small. @@ -546,7 +546,6 @@ Section ZModulo. apply Z_div_pos; auto with zarith. subst a; auto with zarith. apply Zdiv_lt_upper_bound; auto with zarith. - subst a; auto with zarith. subst a. replace (wB*[|b|]) with (([|b|]-1)*wB + wB) by ring. apply Zlt_le_trans with ([|a1|]*wB+wB); auto with zarith. @@ -577,7 +576,7 @@ Section ZModulo. apply Zmod_le; auto with zarith. Qed. - Definition znz_is_even x := + Definition znz_is_even x := if Z_eq_dec ([|x|] mod 2) 0 then true else false. Lemma spec_is_even : forall x, @@ -587,7 +586,7 @@ Section ZModulo. generalize (Z_mod_lt [|x|] 2); omega. Qed. - Definition znz_sqrt x := Zsqrt_plain [|x|]. + Definition znz_sqrt x := Zsqrt_plain [|x|]. Lemma spec_sqrt : forall x, [|znz_sqrt x|] ^ 2 <= [|x|] < ([|znz_sqrt x|] + 1) ^ 2. Proof. @@ -610,12 +609,12 @@ Section ZModulo. generalize wB_pos; auto with zarith. Qed. - Definition znz_sqrt2 x y := - let z := [|x|]*wB+[|y|] in - match z with + Definition znz_sqrt2 x y := + let z := [|x|]*wB+[|y|] in + match z with | Z0 => (0, C0 0) - | Zpos p => - let (s,r,_,_) := sqrtrempos p in + | Zpos p => + let (s,r,_,_) := sqrtrempos p in (s, if Z_lt_le_dec r wB then C0 r else C1 (r-wB)) | Zneg _ => (0, C0 0) end. @@ -652,7 +651,7 @@ Section ZModulo. rewrite Zpower_2; auto with zarith. replace [|r-wB|] with (r-wB) by (symmetry; apply Zmod_small; auto with zarith). rewrite Zpower_2; omega. - + assert (0<=Zneg p). rewrite Heqz; generalize wB_pos; auto with zarith. compute in H0; elim H0; auto. @@ -666,8 +665,8 @@ Section ZModulo. apply two_power_pos_correct. Qed. - Definition znz_head0 x := match [|x|] with - | Z0 => znz_zdigits + Definition znz_head0 x := match [|x|] with + | Z0 => znz_zdigits | Zpos p => znz_zdigits - log_inf p - 1 | _ => 0 end. @@ -696,7 +695,7 @@ Section ZModulo. change (Zpos x~0) with (2*(Zpos x)) in H. replace p with (Zsucc (p-1)) in H; auto with zarith. rewrite Zpower_Zsucc in H; auto with zarith. - + simpl; intros; destruct p; compute; auto with zarith. Qed. @@ -731,8 +730,8 @@ Section ZModulo. by ring. unfold wB, base, znz_zdigits; auto with zarith. apply Zmult_le_compat; auto with zarith. - - apply Zlt_le_trans + + apply Zlt_le_trans with (2^(znz_zdigits - log_inf p - 1)*(2^(Zsucc (log_inf p)))). apply Zmult_lt_compat_l; auto with zarith. rewrite <- Zpower_exp; auto with zarith. @@ -741,17 +740,17 @@ Section ZModulo. unfold wB, base, znz_zdigits; auto with zarith. Qed. - Fixpoint Ptail p := match p with + Fixpoint Ptail p := match p with | xO p => (Ptail p)+1 | _ => 0 - end. + end. Lemma Ptail_pos : forall p, 0 <= Ptail p. Proof. induction p; simpl; auto with zarith. Qed. Hint Resolve Ptail_pos. - + Lemma Ptail_bounded : forall p d, Zpos p < 2^(Zpos d) -> Ptail p < Zpos d. Proof. induction p; try (compute; auto; fail). @@ -776,7 +775,7 @@ Section ZModulo. Qed. Definition znz_tail0 x := - match [|x|] with + match [|x|] with | Z0 => znz_zdigits | Zpos p => Ptail p | Zneg _ => 0 @@ -789,7 +788,7 @@ Section ZModulo. apply spec_zdigits. Qed. - Lemma spec_tail0 : forall x, 0 < [|x|] -> + Lemma spec_tail0 : forall x, 0 < [|x|] -> exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|znz_tail0 x|]). Proof. intros; unfold znz_tail0. @@ -819,7 +818,7 @@ Section ZModulo. (** Let's now group everything in two records *) - Definition zmod_op := mk_znz_op + Definition zmod_op := mk_znz_op (znz_digits : positive) (znz_zdigits: znz) (znz_to_Z : znz -> Z) @@ -860,11 +859,11 @@ Section ZModulo. (znz_div_gt : znz -> znz -> znz * znz) (znz_div : znz -> znz -> znz * znz) - (znz_mod_gt : znz -> znz -> znz) - (znz_mod : znz -> znz -> znz) + (znz_mod_gt : znz -> znz -> znz) + (znz_mod : znz -> znz -> znz) (znz_gcd_gt : znz -> znz -> znz) - (znz_gcd : znz -> znz -> znz) + (znz_gcd : znz -> znz -> znz) (znz_add_mul_div : znz -> znz -> znz -> znz) (znz_pos_mod : znz -> znz -> znz) @@ -879,54 +878,54 @@ Section ZModulo. spec_more_than_1_digit spec_0 - spec_1 - spec_Bm1 - - spec_compare - spec_eq0 - - spec_opp_c - spec_opp - spec_opp_carry - - spec_succ_c - spec_add_c - spec_add_carry_c - spec_succ - spec_add - spec_add_carry - - spec_pred_c - spec_sub_c - spec_sub_carry_c - spec_pred - spec_sub - spec_sub_carry - - spec_mul_c - spec_mul - spec_square_c - - spec_div21 - spec_div_gt - spec_div - - spec_mod_gt - spec_mod - - spec_gcd_gt - spec_gcd - - spec_head00 - spec_head0 - spec_tail00 - spec_tail0 - - spec_add_mul_div - spec_pos_mod - - spec_is_even - spec_sqrt2 + spec_1 + spec_Bm1 + + spec_compare + spec_eq0 + + spec_opp_c + spec_opp + spec_opp_carry + + spec_succ_c + spec_add_c + spec_add_carry_c + spec_succ + spec_add + spec_add_carry + + spec_pred_c + spec_sub_c + spec_sub_carry_c + spec_pred + spec_sub + spec_sub_carry + + spec_mul_c + spec_mul + spec_square_c + + spec_div21 + spec_div_gt + spec_div + + spec_mod_gt + spec_mod + + spec_gcd_gt + spec_gcd + + spec_head00 + spec_head0 + spec_tail00 + spec_tail0 + + spec_add_mul_div + spec_pos_mod + + spec_is_even + spec_sqrt2 spec_sqrt. End ZModulo. @@ -935,7 +934,7 @@ End ZModulo. Module Type PositiveNotOne. Parameter p : positive. - Axiom not_one : p<> 1%positive. + Axiom not_one : p<> 1%positive. End PositiveNotOne. Module ZModuloCyclicType (P:PositiveNotOne) <: CyclicType. |