diff options
author | Samuel Mimram <smimram@debian.org> | 2008-07-25 15:12:53 +0200 |
---|---|---|
committer | Samuel Mimram <smimram@debian.org> | 2008-07-25 15:12:53 +0200 |
commit | a0cfa4f118023d35b767a999d5a2ac4b082857b4 (patch) | |
tree | dabcac548e299fee1da464c93b3dba98484f45b1 /theories/Numbers | |
parent | 2281410e38ef99d025ea77194585a9bc019fdaa9 (diff) |
Imported Upstream version 8.2~beta3+dfsgupstream/8.2.beta3+dfsg
Diffstat (limited to 'theories/Numbers')
66 files changed, 29507 insertions, 0 deletions
diff --git a/theories/Numbers/BigNumPrelude.v b/theories/Numbers/BigNumPrelude.v new file mode 100644 index 00000000..9669eacd --- /dev/null +++ b/theories/Numbers/BigNumPrelude.v @@ -0,0 +1,372 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: BigNumPrelude.v 11013 2008-05-28 18:17:30Z letouzey $ i*) + +(** * BigNumPrelude *) + +(** Auxillary functions & theorems used for arbitrary precision efficient + numbers. *) + + +Require Import ArithRing. +Require Export ZArith. +Require Export Znumtheory. +Require Export Zpow_facts. + +(* *** Nota Bene *** + All results that were general enough has been moved in ZArith. + Only remain here specialized lemmas and compatibility elements. + (P.L. 5/11/2007). +*) + + +Open Local Scope Z_scope. + +(* For compatibility of scripts, weaker version of some lemmas of Zdiv *) + +Lemma Zlt0_not_eq : forall n, 0<n -> n<>0. +Proof. + auto with zarith. +Qed. + +Definition Zdiv_mult_cancel_r a b c H := Zdiv.Zdiv_mult_cancel_r a b c (Zlt0_not_eq _ H). +Definition Zdiv_mult_cancel_l a b c H := Zdiv.Zdiv_mult_cancel_r a b c (Zlt0_not_eq _ H). +Definition Z_div_plus_l a b c H := Zdiv.Z_div_plus_full_l a b c (Zlt0_not_eq _ H). + +(* Automation *) + +Hint Extern 2 (Zle _ _) => + (match goal with + |- Zpos _ <= Zpos _ => exact (refl_equal _) +| H: _ <= ?p |- _ <= ?p => apply Zle_trans with (2 := H) +| H: _ < ?p |- _ <= ?p => apply Zlt_le_weak; apply Zle_lt_trans with (2 := H) + end). + +Hint Extern 2 (Zlt _ _) => + (match goal with + |- Zpos _ < Zpos _ => exact (refl_equal _) +| H: _ <= ?p |- _ <= ?p => apply Zlt_le_trans with (2 := H) +| H: _ < ?p |- _ <= ?p => apply Zle_lt_trans with (2 := H) + end). + + +Hint Resolve Zlt_gt Zle_ge Z_div_pos: zarith. + +(************************************** + Properties of order and product + **************************************) + + Theorem beta_lex: forall a b c d beta, + a * beta + b <= c * beta + d -> + 0 <= b < beta -> 0 <= d < beta -> + a <= c. + Proof. + intros a b c d beta H1 (H3, H4) (H5, H6). + assert (a - c < 1); auto with zarith. + apply Zmult_lt_reg_r with beta; auto with zarith. + apply Zle_lt_trans with (d - b); auto with zarith. + rewrite Zmult_minus_distr_r; auto with zarith. + Qed. + + Theorem beta_lex_inv: forall a b c d beta, + a < c -> 0 <= b < beta -> + 0 <= d < beta -> + a * beta + b < c * beta + d. + Proof. + intros a b c d beta H1 (H3, H4) (H5, H6). + case (Zle_or_lt (c * beta + d) (a * beta + b)); auto with zarith. + intros H7; contradict H1;apply Zle_not_lt;apply beta_lex with (1 := H7);auto. + Qed. + + Lemma beta_mult : forall h l beta, + 0 <= h < beta -> 0 <= l < beta -> 0 <= h*beta+l < beta^2. + Proof. + intros h l beta H1 H2;split. auto with zarith. + rewrite <- (Zplus_0_r (beta^2)); rewrite Zpower_2; + apply beta_lex_inv;auto with zarith. + Qed. + + Lemma Zmult_lt_b : + forall b x y, 0 <= x < b -> 0 <= y < b -> 0 <= x * y <= b^2 - 2*b + 1. + Proof. + intros b x y (Hx1,Hx2) (Hy1,Hy2);split;auto with zarith. + apply Zle_trans with ((b-1)*(b-1)). + apply Zmult_le_compat;auto with zarith. + apply Zeq_le;ring. + Qed. + + Lemma sum_mul_carry : forall xh xl yh yl wc cc beta, + 1 < beta -> + 0 <= wc < beta -> + 0 <= xh < beta -> + 0 <= xl < beta -> + 0 <= yh < beta -> + 0 <= yl < beta -> + 0 <= cc < beta^2 -> + wc*beta^2 + cc = xh*yl + xl*yh -> + 0 <= wc <= 1. + Proof. + intros xh xl yh yl wc cc beta U H1 H2 H3 H4 H5 H6 H7. + assert (H8 := Zmult_lt_b beta xh yl H2 H5). + assert (H9 := Zmult_lt_b beta xl yh H3 H4). + split;auto with zarith. + apply beta_lex with (cc) (beta^2 - 2) (beta^2); auto with zarith. + Qed. + + Theorem mult_add_ineq: forall x y cross beta, + 0 <= x < beta -> + 0 <= y < beta -> + 0 <= cross < beta -> + 0 <= x * y + cross < beta^2. + Proof. + intros x y cross beta HH HH1 HH2. + split; auto with zarith. + apply Zle_lt_trans with ((beta-1)*(beta-1)+(beta-1)); auto with zarith. + apply Zplus_le_compat; auto with zarith. + apply Zmult_le_compat; auto with zarith. + repeat (rewrite Zmult_minus_distr_l || rewrite Zmult_minus_distr_r); + rewrite Zpower_2; auto with zarith. + Qed. + + Theorem mult_add_ineq2: forall x y c cross beta, + 0 <= x < beta -> + 0 <= y < beta -> + 0 <= c*beta + cross <= 2*beta - 2 -> + 0 <= x * y + (c*beta + cross) < beta^2. + Proof. + intros x y c cross beta HH HH1 HH2. + split; auto with zarith. + apply Zle_lt_trans with ((beta-1)*(beta-1)+(2*beta-2));auto with zarith. + apply Zplus_le_compat; auto with zarith. + apply Zmult_le_compat; auto with zarith. + repeat (rewrite Zmult_minus_distr_l || rewrite Zmult_minus_distr_r); + rewrite Zpower_2; auto with zarith. + Qed. + +Theorem mult_add_ineq3: forall x y c cross beta, + 0 <= x < beta -> + 0 <= y < beta -> + 0 <= cross <= beta - 2 -> + 0 <= c <= 1 -> + 0 <= x * y + (c*beta + cross) < beta^2. + Proof. + intros x y c cross beta HH HH1 HH2 HH3. + apply mult_add_ineq2;auto with zarith. + split;auto with zarith. + apply Zle_trans with (1*beta+cross);auto with zarith. + Qed. + +Hint Rewrite Zmult_1_r Zmult_0_r Zmult_1_l Zmult_0_l Zplus_0_l Zplus_0_r Zminus_0_r: rm10. + + +(************************************** + Properties of Zdiv and Zmod +**************************************) + +Theorem Zmod_le_first: forall a b, 0 <= a -> 0 < b -> 0 <= a mod b <= a. + Proof. + intros a b H H1;case (Z_mod_lt a b);auto with zarith;intros H2 H3;split;auto. + case (Zle_or_lt b a); intros H4; auto with zarith. + rewrite Zmod_small; auto with zarith. + Qed. + + + Theorem Zmod_distr: forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a -> + (2 ^a * r + t) mod (2 ^ b) = (2 ^a * r) mod (2 ^ b) + t. + Proof. + intros a b r t (H1, H2) H3 (H4, H5). + assert (t < 2 ^ b). + apply Zlt_le_trans with (1:= H5); auto with zarith. + apply Zpower_le_monotone; auto with zarith. + rewrite Zplus_mod; auto with zarith. + rewrite Zmod_small with (a := t); auto with zarith. + apply Zmod_small; auto with zarith. + split; auto with zarith. + assert (0 <= 2 ^a * r); auto with zarith. + apply Zplus_le_0_compat; auto with zarith. + match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end; + auto with zarith. + pattern (2 ^ b) at 2; replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); + try ring. + apply Zplus_le_lt_compat; auto with zarith. + replace b with ((b - a) + a); try ring. + rewrite Zpower_exp; auto with zarith. + pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a)); + try rewrite <- Zmult_minus_distr_r. + rewrite (Zmult_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l; + auto with zarith. + rewrite (Zmult_comm (2 ^a)); apply Zmult_le_compat_r; auto with zarith. + match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end; + auto with zarith. + Qed. + + Theorem Zmod_shift_r: + forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a -> + (r * 2 ^a + t) mod (2 ^ b) = (r * 2 ^a) mod (2 ^ b) + t. + Proof. + intros a b r t (H1, H2) H3 (H4, H5). + assert (t < 2 ^ b). + apply Zlt_le_trans with (1:= H5); auto with zarith. + apply Zpower_le_monotone; auto with zarith. + rewrite Zplus_mod; auto with zarith. + rewrite Zmod_small with (a := t); auto with zarith. + apply Zmod_small; auto with zarith. + split; auto with zarith. + assert (0 <= 2 ^a * r); auto with zarith. + apply Zplus_le_0_compat; auto with zarith. + match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end; + auto with zarith. + pattern (2 ^ b) at 2;replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); try ring. + apply Zplus_le_lt_compat; auto with zarith. + replace b with ((b - a) + a); try ring. + rewrite Zpower_exp; auto with zarith. + pattern (2 ^a) at 4; rewrite <- (Zmult_1_l (2 ^a)); + try rewrite <- Zmult_minus_distr_r. + repeat rewrite (fun x => Zmult_comm x (2 ^ a)); rewrite Zmult_mod_distr_l; + auto with zarith. + apply Zmult_le_compat_l; auto with zarith. + match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end; + auto with zarith. + Qed. + + Theorem Zdiv_shift_r: + forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a -> + (r * 2 ^a + t) / (2 ^ b) = (r * 2 ^a) / (2 ^ b). + Proof. + intros a b r t (H1, H2) H3 (H4, H5). + assert (Eq: t < 2 ^ b); auto with zarith. + apply Zlt_le_trans with (1 := H5); auto with zarith. + apply Zpower_le_monotone; auto with zarith. + pattern (r * 2 ^ a) at 1; rewrite Z_div_mod_eq with (b := 2 ^ b); + auto with zarith. + rewrite <- Zplus_assoc. + rewrite <- Zmod_shift_r; auto with zarith. + rewrite (Zmult_comm (2 ^ b)); rewrite Z_div_plus_full_l; auto with zarith. + rewrite (fun x y => @Zdiv_small (x mod y)); auto with zarith. + match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end; + auto with zarith. + Qed. + + + Lemma shift_unshift_mod : forall n p a, + 0 <= a < 2^n -> + 0 <= p <= n -> + a * 2^p = a / 2^(n - p) * 2^n + (a*2^p) mod 2^n. + Proof. + intros n p a H1 H2. + pattern (a*2^p) at 1;replace (a*2^p) with + (a*2^p/2^n * 2^n + a*2^p mod 2^n). + 2:symmetry;rewrite (Zmult_comm (a*2^p/2^n));apply Z_div_mod_eq. + replace (a * 2 ^ p / 2 ^ n) with (a / 2 ^ (n - p));trivial. + replace (2^n) with (2^(n-p)*2^p). + symmetry;apply Zdiv_mult_cancel_r. + destruct H1;trivial. + cut (0 < 2^p); auto with zarith. + rewrite <- Zpower_exp. + replace (n-p+p) with n;trivial. ring. + omega. omega. + apply Zlt_gt. apply Zpower_gt_0;auto with zarith. + Qed. + + + Lemma shift_unshift_mod_2 : forall n p a, (0<=p<=n)%Z -> + ((a * 2 ^ (n - p)) mod (2^n) / 2 ^ (n - p)) mod (2^n) = + a mod 2 ^ p. + Proof. + intros. + rewrite Zmod_small. + rewrite Zmod_eq by (auto with zarith). + unfold Zminus at 1. + rewrite Z_div_plus_l by (auto with zarith). + assert (2^n = 2^(n-p)*2^p). + rewrite <- Zpower_exp by (auto with zarith). + replace (n-p+p) with n; auto with zarith. + rewrite H0. + rewrite <- Zdiv_Zdiv, Z_div_mult by (auto with zarith). + rewrite (Zmult_comm (2^(n-p))), Zmult_assoc. + rewrite Zopp_mult_distr_l. + rewrite Z_div_mult by (auto with zarith). + symmetry; apply Zmod_eq; auto with zarith. + + remember (a * 2 ^ (n - p)) as b. + destruct (Z_mod_lt b (2^n)); auto with zarith. + split. + apply Z_div_pos; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + apply Zlt_le_trans with (2^n); auto with zarith. + rewrite <- (Zmult_1_r (2^n)) at 1. + apply Zmult_le_compat; auto with zarith. + cut (0 < 2 ^ (n-p)); auto with zarith. + Qed. + + Lemma div_le_0 : forall p x, 0 <= x -> 0 <= x / 2 ^ p. + Proof. + intros p x Hle;destruct (Z_le_gt_dec 0 p). + apply Zdiv_le_lower_bound;auto with zarith. + replace (2^p) with 0. + destruct x;compute;intro;discriminate. + destruct p;trivial;discriminate z. + Qed. + + Lemma div_lt : forall p x y, 0 <= x < y -> x / 2^p < y. + Proof. + intros p x y H;destruct (Z_le_gt_dec 0 p). + apply Zdiv_lt_upper_bound;auto with zarith. + apply Zlt_le_trans with y;auto with zarith. + rewrite <- (Zmult_1_r y);apply Zmult_le_compat;auto with zarith. + assert (0 < 2^p);auto with zarith. + replace (2^p) with 0. + destruct x;change (0<y);auto with zarith. + destruct p;trivial;discriminate z. + Qed. + + Theorem Zgcd_div_pos a b: + (0 < b)%Z -> (0 < Zgcd a b)%Z -> (0 < b / Zgcd a b)%Z. + Proof. + intros a b Ha Hg. + case (Zle_lt_or_eq 0 (b/Zgcd a b)); auto. + apply Z_div_pos; auto with zarith. + intros H; generalize Ha. + pattern b at 1; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite <- H; auto with zarith. + assert (F := (Zgcd_is_gcd a b)); inversion F; auto. + Qed. + +Theorem Zbounded_induction : + (forall Q : Z -> Prop, forall b : Z, + Q 0 -> + (forall n, 0 <= n -> n < b - 1 -> Q n -> Q (n + 1)) -> + forall n, 0 <= n -> n < b -> Q n)%Z. +Proof. +intros Q b Q0 QS. +set (Q' := fun n => (n < b /\ Q n) \/ (b <= n)). +assert (H : forall n, 0 <= n -> Q' n). +apply natlike_rec2; unfold Q'. +destruct (Zle_or_lt b 0) as [H | H]. now right. left; now split. +intros n H IH. destruct IH as [[IH1 IH2] | IH]. +destruct (Zle_or_lt (b - 1) n) as [H1 | H1]. +right; auto with zarith. +left. split; [auto with zarith | now apply (QS n)]. +right; auto with zarith. +unfold Q' in *; intros n H1 H2. destruct (H n H1) as [[H3 H4] | H3]. +assumption. apply Zle_not_lt in H3. false_hyp H2 H3. +Qed. + +Lemma Zsquare_le : forall x, x <= x*x. +Proof. +intros. +destruct (Z_lt_le_dec 0 x). +pattern x at 1; rewrite <- (Zmult_1_l x). +apply Zmult_le_compat; auto with zarith. +apply Zle_trans with 0; auto with zarith. +rewrite <- Zmult_opp_opp. +apply Zmult_le_0_compat; auto with zarith. +Qed. diff --git a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v new file mode 100644 index 00000000..528d78c3 --- /dev/null +++ b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v @@ -0,0 +1,375 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(* $Id: CyclicAxioms.v 11012 2008-05-28 16:34:43Z letouzey $ *) + +(** * 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. *) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import Znumtheory. +Require Import BigNumPrelude. +Require Import DoubleType. + +Open Local Scope Z_scope. + +(** First, a description via an operator record and a spec record. *) + +Section Z_nZ_Op. + + Variable znz : Type. + + Record znz_op := mk_znz_op { + + (* Conversion functions with Z *) + znz_digits : positive; + znz_zdigits: znz; + znz_to_Z : znz -> Z; + znz_of_pos : positive -> N * znz; (* Euclidean division by [2^digits] *) + znz_head0 : znz -> znz; (* number of digits 0 in front of the number *) + znz_tail0 : znz -> znz; (* number of digits 0 at the bottom of the number *) + + (* Basic numbers *) + znz_0 : znz; + znz_1 : znz; + znz_Bm1 : znz; (* [2^digits-1], which is equivalent to [-1] *) + + (* Comparison *) + znz_compare : znz -> znz -> comparison; + znz_eq0 : znz -> bool; + + (* Basic arithmetic operations *) + znz_opp_c : znz -> carry znz; + znz_opp : znz -> znz; + znz_opp_carry : znz -> znz; (* the carry is known to be -1 *) + + znz_succ_c : znz -> carry znz; + znz_add_c : znz -> znz -> carry znz; + znz_add_carry_c : znz -> znz -> carry znz; + znz_succ : znz -> znz; + znz_add : znz -> znz -> znz; + znz_add_carry : znz -> znz -> znz; + + znz_pred_c : znz -> carry znz; + znz_sub_c : znz -> znz -> carry znz; + znz_sub_carry_c : znz -> znz -> carry znz; + znz_pred : znz -> znz; + znz_sub : znz -> znz -> znz; + znz_sub_carry : znz -> znz -> znz; + + znz_mul_c : znz -> znz -> zn2z znz; + znz_mul : znz -> znz -> znz; + znz_square_c : znz -> zn2z znz; + + (* Special divisions operations *) + znz_div21 : znz -> znz -> znz -> znz*znz; + znz_div_gt : znz -> znz -> znz * znz; (* specialized version of [znz_div] *) + znz_div : znz -> znz -> znz * znz; + + znz_mod_gt : znz -> znz -> znz; (* specialized version of [znz_mod] *) + znz_mod : znz -> znz -> znz; + + znz_gcd_gt : znz -> znz -> znz; (* specialized version of [znz_gcd] *) + 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]: + [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] *) + znz_pos_mod : znz -> znz -> znz; + + znz_is_even : znz -> bool; + (* square root *) + znz_sqrt2 : znz -> znz -> znz * carry znz; + znz_sqrt : znz -> znz }. + +End Z_nZ_Op. + +Section Z_nZ_Spec. + Variable w : Type. + Variable w_op : znz_op w. + + Let w_digits := w_op.(znz_digits). + Let w_zdigits := w_op.(znz_zdigits). + Let w_to_Z := w_op.(znz_to_Z). + Let w_of_pos := w_op.(znz_of_pos). + Let w_head0 := w_op.(znz_head0). + Let w_tail0 := w_op.(znz_tail0). + + Let w0 := w_op.(znz_0). + Let w1 := w_op.(znz_1). + Let wBm1 := w_op.(znz_Bm1). + + Let w_compare := w_op.(znz_compare). + Let w_eq0 := w_op.(znz_eq0). + + Let w_opp_c := w_op.(znz_opp_c). + Let w_opp := w_op.(znz_opp). + Let w_opp_carry := w_op.(znz_opp_carry). + + Let w_succ_c := w_op.(znz_succ_c). + Let w_add_c := w_op.(znz_add_c). + Let w_add_carry_c := w_op.(znz_add_carry_c). + Let w_succ := w_op.(znz_succ). + Let w_add := w_op.(znz_add). + Let w_add_carry := w_op.(znz_add_carry). + + Let w_pred_c := w_op.(znz_pred_c). + Let w_sub_c := w_op.(znz_sub_c). + Let w_sub_carry_c := w_op.(znz_sub_carry_c). + Let w_pred := w_op.(znz_pred). + Let w_sub := w_op.(znz_sub). + Let w_sub_carry := w_op.(znz_sub_carry). + + Let w_mul_c := w_op.(znz_mul_c). + Let w_mul := w_op.(znz_mul). + Let w_square_c := w_op.(znz_square_c). + + Let w_div21 := w_op.(znz_div21). + Let w_div_gt := w_op.(znz_div_gt). + Let w_div := w_op.(znz_div). + + Let w_mod_gt := w_op.(znz_mod_gt). + Let w_mod := w_op.(znz_mod). + + Let w_gcd_gt := w_op.(znz_gcd_gt). + Let w_gcd := w_op.(znz_gcd). + + Let w_add_mul_div := w_op.(znz_add_mul_div). + + Let w_pos_mod := w_op.(znz_pos_mod). + + Let w_is_even := w_op.(znz_is_even). + Let w_sqrt2 := w_op.(znz_sqrt2). + Let w_sqrt := w_op.(znz_sqrt). + + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + + Let wB := base w_digits. + + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + + Notation "[-| c |]" := + (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). + + Notation "[|| x ||]" := + (zn2z_to_Z wB w_to_Z x) (at level 0, x at level 99). + + Record znz_spec := mk_znz_spec { + + (* Conversion functions with Z *) + spec_to_Z : forall x, 0 <= [| x |] < wB; + spec_of_pos : forall p, + Zpos p = (Z_of_N (fst (w_of_pos p)))*wB + [|(snd (w_of_pos p))|]; + spec_zdigits : [| w_zdigits |] = Zpos w_digits; + spec_more_than_1_digit: 1 < Zpos w_digits; + + (* Basic numbers *) + spec_0 : [|w0|] = 0; + spec_1 : [|w1|] = 1; + spec_Bm1 : [|wBm1|] = wB - 1; + + (* Comparison *) + spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end; + spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0; + (* Basic arithmetic operations *) + spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|]; + spec_opp : forall x, [|w_opp x|] = (-[|x|]) mod wB; + spec_opp_carry : forall x, [|w_opp_carry x|] = wB - [|x|] - 1; + + spec_succ_c : forall x, [+|w_succ_c x|] = [|x|] + 1; + spec_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]; + spec_add_carry_c : forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1; + spec_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB; + spec_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB; + spec_add_carry : + forall x y, [|w_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB; + + spec_pred_c : forall x, [-|w_pred_c x|] = [|x|] - 1; + spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|]; + spec_sub_carry_c : forall x y, [-|w_sub_carry_c x y|] = [|x|] - [|y|] - 1; + spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB; + spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB; + spec_sub_carry : + forall x y, [|w_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB; + + spec_mul_c : forall x y, [|| w_mul_c x y ||] = [|x|] * [|y|]; + spec_mul : forall x y, [|w_mul x y|] = ([|x|] * [|y|]) mod wB; + spec_square_c : forall x, [|| w_square_c x||] = [|x|] * [|x|]; + + (* Special divisions operations *) + spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]; + spec_div_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> + let (q,r) := w_div_gt a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]; + spec_div : forall a b, 0 < [|b|] -> + let (q,r) := w_div a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 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; + 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_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; + spec_pos_mod : forall w p, + [|w_pos_mod p w|] = [|w|] mod (2 ^ [|p|]); + (* sqrt *) + spec_is_even : forall x, + if w_is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1; + spec_sqrt2 : forall x y, + wB/ 4 <= [|x|] -> + let (s,r) := w_sqrt2 x y in + [||WW x y||] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]; + spec_sqrt : forall x, + [|w_sqrt x|] ^ 2 <= [|x|] < ([|w_sqrt x|] + 1) ^ 2 + }. + +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 := + if w_eq0 h then W0 else WW h w_0. + + Definition znz_0W l := + if w_eq0 l then W0 else WW w_0 l. + + Definition znz_WW h l := + if w_eq0 h then znz_0W l else WW h l. + + Lemma spec_W0 : forall h, + zn2z_to_Z wB w_to_Z (znz_W0 h) = (w_to_Z h)*wB. + Proof. + unfold zn2z_to_Z, znz_W0, w_to_Z; simpl; intros. + case_eq (w_eq0 h); intros. + rewrite (op_spec.(spec_eq0) _ H); auto. + unfold w_0; rewrite op_spec.(spec_0); auto with zarith. + Qed. + + 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. + case_eq (w_eq0 l); intros. + rewrite (op_spec.(spec_eq0) _ H); auto. + unfold w_0; rewrite op_spec.(spec_0); auto with zarith. + Qed. + + 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. + case_eq (w_eq0 h); intros. + rewrite (op_spec.(spec_eq0) _ H); auto. + rewrite spec_0W; auto. + simpl; auto. + Qed. + +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. + + Notation "[| x |]" := (znz_to_Z w_op x) (at level 0, x at level 99). + + Definition znz_of_Z (w:Type) (op:znz_op w) z := + match z with + | Zpos p => snd (op.(znz_of_pos) p) + | _ => op.(znz_0) + end. + + Theorem znz_of_pos_correct: + forall p, Zpos p < base (znz_digits w_op) -> [|(snd (znz_of_pos w_op p))|] = Zpos p. + intros p Hp. + generalize (spec_of_pos op_spec p). + case (znz_of_pos w_op p); intros n w1; simpl. + case n; simpl Npos; auto with zarith. + intros p1 Hp1; contradict Hp; apply Zle_not_lt. + rewrite Hp1; auto with zarith. + match goal with |- _ <= ?X + ?Y => + apply Zle_trans with X; auto with zarith + end. + match goal with |- ?X <= _ => + 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. + unfold base; auto with zarith. + case (spec_to_Z op_spec w1); auto with zarith. + Qed. + + Theorem znz_of_Z_correct: + forall p, 0 <= p < base (znz_digits w_op) -> [|znz_of_Z w_op p|] = p. + intros p; case p; simpl; try rewrite spec_0; auto. + intros; rewrite znz_of_pos_correct; auto with zarith. + intros p1 (H1, _); contradict H1; apply Zlt_not_le; red; simpl; auto. + Qed. +End znz_of_pos. + + +(** A modular specification grouping the earlier records. *) + +Module Type CyclicType. + Parameter w : Type. + Parameter w_op : znz_op w. + Parameter w_spec : znz_spec w_op. +End CyclicType. diff --git a/theories/Numbers/Cyclic/Abstract/NZCyclic.v b/theories/Numbers/Cyclic/Abstract/NZCyclic.v new file mode 100644 index 00000000..22f6d95b --- /dev/null +++ b/theories/Numbers/Cyclic/Abstract/NZCyclic.v @@ -0,0 +1,236 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZCyclic.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NZAxioms. +Require Import BigNumPrelude. +Require Import DoubleType. +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 power of 2. +*) + +Module NZCyclicAxiomsMod (Import Cyclic : CyclicType) <: NZAxiomsSig. + +Open Local Scope Z_scope. + +Definition NZ := 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)). + +Notation Local "[| 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). + +Theorem NZeq_equiv : equiv NZ NZeq. +Proof. +unfold equiv, reflexive, symmetric, transitive, NZeq; repeat split; intros; auto. +now transitivity [| y |]. +Qed. + +Add Relation NZ NZeq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) +as NZeq_rel. + +Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd. +Proof. +unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_succ). now rewrite H. +Qed. + +Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd. +Proof. +unfold NZeq; intros n m H. do 2 rewrite w_spec.(spec_pred). now rewrite H. +Qed. + +Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd. +Proof. +unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_add). +now rewrite H1, H2. +Qed. + +Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd. +Proof. +unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_sub). +now rewrite H1, H2. +Qed. + +Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd. +Proof. +unfold NZeq; intros n1 n2 H1 m1 m2 H2. do 2 rewrite w_spec.(spec_mul). +now rewrite H1, H2. +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'" := NZsucc : IntScope. +Notation "'P'" := NZpred : IntScope. +(*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. + +Theorem gt_wB_1 : 1 < wB. +Proof. +unfold base. +apply Zpower_gt_1; unfold Zlt; auto with zarith. +Qed. + +Theorem gt_wB_0 : 0 < wB. +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. +Proof. +intro n. +pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zplus_mod. +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. +Proof. +intro n. +pattern 1 at 2. replace 1 with (1 mod wB). rewrite <- Zminus_mod. +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 |]. +Proof. +intro n; rewrite Zmod_small. reflexivity. apply w_spec.(spec_to_Z). +Qed. + +Theorem NZpred_succ : forall n : NZ, P (S n) == n. +Proof. +intro n; unfold NZsucc, NZpred, NZeq. rewrite w_spec.(spec_pred), w_spec.(spec_succ). +rewrite <- NZpred_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. +Proof. +unfold NZeq, NZ_to_Z, Z_to_NZ. rewrite znz_of_Z_correct. +symmetry; apply w_spec.(spec_0). +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. +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. + +Let B (n : Z) := A (Z_to_NZ n). + +Lemma B0 : B 0. +Proof. +unfold B. now rewrite Z_to_NZ_0. +Qed. + +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). +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. +Qed. + +Lemma B_holds : forall n : Z, 0 <= n < wB -> B n. +Proof. +intros n [H1 H2]. +apply Zbounded_induction with wB. +apply B0. apply BS. assumption. assumption. +Qed. + +Theorem NZinduction : forall n : NZ, A n. +Proof. +intro n. setoid_replace n with (Z_to_NZ (NZ_to_Z n)) using relation NZeq. +apply B_holds. apply w_spec.(spec_to_Z). +unfold NZeq, NZ_to_Z, Z_to_NZ; rewrite znz_of_Z_correct. +reflexivity. +exact w_spec. +apply w_spec.(spec_to_Z). +Qed. + +End Induction. + +Theorem NZadd_0_l : forall n : NZ, 0 + n == n. +Proof. +intro n; unfold NZadd, NZ0, NZeq. rewrite w_spec.(spec_add). rewrite w_spec.(spec_0). +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). +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. +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. +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. +Qed. + +Theorem NZsub_succ_r : forall n m : NZ, 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. +Qed. + +Theorem NZmul_0_l : forall n : NZ, 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. +Qed. + +Theorem NZmul_succ_l : forall n m : NZ, (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. +now rewrite Zmult_plus_distr_l, Zmult_1_l. +Qed. + +End NZCyclicAxiomsMod. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v new file mode 100644 index 00000000..61d8d0fb --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleAdd.v @@ -0,0 +1,318 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleAdd.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleAdd. + Variable w : Type. + Variable w_0 : w. + Variable w_1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable ww_1 : zn2z w. + Variable w_succ_c : w -> carry w. + Variable w_add_c : w -> w -> carry w. + Variable w_add_carry_c : w -> w -> carry w. + Variable w_succ : w -> w. + Variable w_add : w -> w -> w. + Variable w_add_carry : w -> w -> w. + + Definition ww_succ_c x := + match x with + | W0 => C0 ww_1 + | WW xh xl => + match w_succ_c xl with + | C0 l => C0 (WW xh l) + | C1 l => + match w_succ_c xh with + | C0 h => C0 (WW h w_0) + | C1 h => C1 W0 + end + end + end. + + Definition ww_succ x := + match x with + | W0 => ww_1 + | WW xh xl => + match w_succ_c xl with + | C0 l => WW xh l + | C1 l => w_W0 (w_succ xh) + end + end. + + Definition ww_add_c x y := + match x, y with + | W0, _ => C0 y + | _, W0 => C0 x + | WW xh xl, WW yh yl => + match w_add_c xl yl with + | C0 l => + match w_add_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (w_WW h l) + end + | C1 l => + match w_add_carry_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (w_WW h l) + end + end + end. + + Variable R : Type. + Variable f0 f1 : zn2z w -> R. + + Definition ww_add_c_cont x y := + match x, y with + | W0, _ => f0 y + | _, W0 => f0 x + | WW xh xl, WW yh yl => + match w_add_c xl yl with + | C0 l => + match w_add_c xh yh with + | C0 h => f0 (WW h l) + | C1 h => f1 (w_WW h l) + end + | C1 l => + match w_add_carry_c xh yh with + | C0 h => f0 (WW h l) + | C1 h => f1 (w_WW h l) + end + end + end. + + (* ww_add et ww_add_carry conserve la forme normale s'il n'y a pas + de debordement *) + Definition ww_add x y := + match x, y with + | W0, _ => y + | _, W0 => x + | WW xh xl, WW yh yl => + match w_add_c xl yl with + | C0 l => WW (w_add xh yh) l + | C1 l => WW (w_add_carry xh yh) l + end + end. + + Definition ww_add_carry_c x y := + match x, y with + | W0, W0 => C0 ww_1 + | W0, WW yh yl => ww_succ_c (WW yh yl) + | WW xh xl, W0 => ww_succ_c (WW xh xl) + | WW xh xl, WW yh yl => + match w_add_carry_c xl yl with + | C0 l => + match w_add_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (WW h l) + end + | C1 l => + match w_add_carry_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (w_WW h l) + end + end + end. + + Definition ww_add_carry x y := + match x, y with + | W0, W0 => ww_1 + | W0, WW yh yl => ww_succ (WW yh yl) + | WW xh xl, W0 => ww_succ (WW xh xl) + | WW xh xl, WW yh yl => + match w_add_carry_c xl yl with + | C0 l => WW (w_add xh yh) l + | C1 l => WW (w_add_carry xh yh) l + end + end. + + (*Section DoubleProof.*) + Variable w_digits : positive. + Variable w_to_Z : w -> Z. + + + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. + Variable spec_w_succ_c : forall x, [+|w_succ_c x|] = [|x|] + 1. + Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. + Variable spec_w_add_carry_c : + forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1. + Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB. + Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB. + Variable spec_w_add_carry : + forall x y, [|w_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB. + + Lemma spec_ww_succ_c : forall x, [+[ww_succ_c x]] = [[x]] + 1. + Proof. + destruct x as [ |xh xl];simpl. apply spec_ww_1. + generalize (spec_w_succ_c xl);destruct (w_succ_c xl) as [l|l]; + intro H;unfold interp_carry in H. simpl;rewrite H;ring. + rewrite <- Zplus_assoc;rewrite <- H;rewrite Zmult_1_l. + assert ([|l|] = 0). generalize (spec_to_Z xl)(spec_to_Z l);omega. + rewrite H0;generalize (spec_w_succ_c xh);destruct (w_succ_c xh) as [h|h]; + intro H1;unfold interp_carry in H1. + simpl;rewrite H1;rewrite spec_w_0;ring. + unfold interp_carry;simpl ww_to_Z;rewrite wwB_wBwB. + assert ([|xh|] = wB - 1). generalize (spec_to_Z xh)(spec_to_Z h);omega. + rewrite H2;ring. + Qed. + + Lemma spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]]. + Proof. + destruct x as [ |xh xl];simpl;trivial. + destruct y as [ |yh yl];simpl. rewrite Zplus_0_r;trivial. + replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|])) + with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring. + generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l]; + intros H;unfold interp_carry in H;rewrite <- H. + generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; + intros H1;unfold interp_carry in *;rewrite <- H1. trivial. + repeat rewrite Zmult_1_l;rewrite spec_w_WW;rewrite wwB_wBwB; ring. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh) + as [h|h]; intros H1;unfold interp_carry in *;rewrite <- H1. + simpl;ring. + repeat rewrite Zmult_1_l;rewrite wwB_wBwB;rewrite spec_w_WW;ring. + Qed. + + Section Cont. + Variable P : zn2z w -> zn2z w -> R -> Prop. + Variable x y : zn2z w. + Variable spec_f0 : forall r, [[r]] = [[x]] + [[y]] -> P x y (f0 r). + Variable spec_f1 : forall r, wwB + [[r]] = [[x]] + [[y]] -> P x y (f1 r). + + Lemma spec_ww_add_c_cont : P x y (ww_add_c_cont x y). + Proof. + destruct x as [ |xh xl];simpl;trivial. + apply spec_f0;trivial. + destruct y as [ |yh yl];simpl. + apply spec_f0;simpl;rewrite Zplus_0_r;trivial. + generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l]; + intros H;unfold interp_carry in H. + generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; + intros H1;unfold interp_carry in *. + apply spec_f0. simpl;rewrite H;rewrite H1;ring. + apply spec_f1. simpl;rewrite spec_w_WW;rewrite H. + rewrite Zplus_assoc;rewrite wwB_wBwB. rewrite Zpower_2; rewrite <- Zmult_plus_distr_l. + rewrite Zmult_1_l in H1;rewrite H1;ring. + generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh) + as [h|h]; intros H1;unfold interp_carry in *. + apply spec_f0;simpl;rewrite H1. rewrite Zmult_plus_distr_l. + rewrite <- Zplus_assoc;rewrite H;ring. + apply spec_f1. simpl;rewrite spec_w_WW;rewrite wwB_wBwB. + rewrite Zplus_assoc; rewrite Zpower_2; rewrite <- Zmult_plus_distr_l. + rewrite Zmult_1_l in H1;rewrite H1. rewrite Zmult_plus_distr_l. + rewrite <- Zplus_assoc;rewrite H;ring. + Qed. + + End Cont. + + Lemma spec_ww_add_carry_c : + forall x y, [+[ww_add_carry_c x y]] = [[x]] + [[y]] + 1. + Proof. + destruct x as [ |xh xl];intro y;simpl. + exact (spec_ww_succ_c y). + destruct y as [ |yh yl];simpl. + rewrite Zplus_0_r;exact (spec_ww_succ_c (WW xh xl)). + replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1) + with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring. + generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) + as [l|l];intros H;unfold interp_carry in H;rewrite <- H. + generalize (spec_w_add_c xh yh);destruct (w_add_c xh yh) as [h|h]; + intros H1;unfold interp_carry in H1;rewrite <- H1. trivial. + unfold interp_carry;repeat rewrite Zmult_1_l;simpl;rewrite wwB_wBwB;ring. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + generalize (spec_w_add_carry_c xh yh);destruct (w_add_carry_c xh yh) + as [h|h];intros H1;unfold interp_carry in H1;rewrite <- H1. trivial. + unfold interp_carry;rewrite spec_w_WW; + repeat rewrite Zmult_1_l;simpl;rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_succ : forall x, [[ww_succ x]] = ([[x]] + 1) mod wwB. + Proof. + destruct x as [ |xh xl];simpl. + rewrite spec_ww_1;rewrite Zmod_small;trivial. + split;[intro;discriminate|apply wwB_pos]. + rewrite <- Zplus_assoc;generalize (spec_w_succ_c xl); + destruct (w_succ_c xl) as[l|l];intro H;unfold interp_carry in H;rewrite <-H. + rewrite Zmod_small;trivial. + rewrite wwB_wBwB;apply beta_mult;apply spec_to_Z. + assert ([|l|] = 0). clear spec_ww_1 spec_w_1 spec_w_0. + assert (H1:= spec_to_Z l); assert (H2:= spec_to_Z xl); omega. + rewrite H0;rewrite Zplus_0_r;rewrite <- Zmult_plus_distr_l;rewrite wwB_wBwB. + rewrite Zpower_2; rewrite Zmult_mod_distr_r;try apply lt_0_wB. + rewrite spec_w_W0;rewrite spec_w_succ;trivial. + Qed. + + Lemma spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB. + Proof. + destruct x as [ |xh xl];intros y;simpl. + rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial. + destruct y as [ |yh yl]. + change [[W0]] with 0;rewrite Zplus_0_r. + rewrite Zmod_small;trivial. + exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)). + simpl. replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|])) + with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|])). 2:ring. + generalize (spec_w_add_c xl yl);destruct (w_add_c xl yl) as [l|l]; + unfold interp_carry;intros H;simpl;rewrite <- H. + rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial. + Qed. + + Lemma spec_ww_add_carry : + forall x y, [[ww_add_carry x y]] = ([[x]] + [[y]] + 1) mod wwB. + Proof. + destruct x as [ |xh xl];intros y;simpl. + exact (spec_ww_succ y). + destruct y as [ |yh yl]. + change [[W0]] with 0;rewrite Zplus_0_r. exact (spec_ww_succ (WW xh xl)). + simpl;replace ([|xh|] * wB + [|xl|] + ([|yh|] * wB + [|yl|]) + 1) + with (([|xh|]+[|yh|])*wB + ([|xl|]+[|yl|]+1)). 2:ring. + generalize (spec_w_add_carry_c xl yl);destruct (w_add_carry_c xl yl) + as [l|l];unfold interp_carry;intros H;rewrite <- H;simpl ww_to_Z. + rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add;trivial. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite(mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_w_add_carry;trivial. + Qed. + +(* End DoubleProof. *) +End DoubleAdd. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v new file mode 100644 index 00000000..952516ac --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleBase.v @@ -0,0 +1,446 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleBase.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. + +Open Local Scope Z_scope. + +Section DoubleBase. + Variable w : Type. + Variable w_0 : w. + Variable w_1 : w. + Variable w_Bm1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_digits : positive. + Variable w_zdigits: w. + Variable w_add: w -> w -> zn2z w. + Variable w_to_Z : w -> Z. + Variable w_compare : w -> w -> comparison. + + Definition ww_digits := xO w_digits. + + Definition ww_zdigits := w_add w_zdigits w_zdigits. + + Definition ww_to_Z := zn2z_to_Z (base w_digits) w_to_Z. + + Definition ww_1 := WW w_0 w_1. + + Definition ww_Bm1 := WW w_Bm1 w_Bm1. + + Definition ww_WW xh xl : zn2z (zn2z w) := + match xh, xl with + | W0, W0 => W0 + | _, _ => WW xh xl + end. + + Definition ww_W0 h : zn2z (zn2z w) := + match h with + | W0 => W0 + | _ => WW h W0 + end. + + Definition ww_0W l : zn2z (zn2z w) := + match l with + | W0 => W0 + | _ => WW W0 l + end. + + Definition double_WW (n:nat) := + match n return word w n -> word w n -> word w (S n) with + | O => w_WW + | S n => + fun (h l : zn2z (word w n)) => + match h, l with + | W0, W0 => W0 + | _, _ => WW h l + end + end. + + Fixpoint double_digits (n:nat) : positive := + match n with + | O => w_digits + | S n => xO (double_digits n) + end. + + Definition double_wB n := base (double_digits n). + + Fixpoint double_to_Z (n:nat) : word w n -> Z := + match n return word w n -> Z with + | O => w_to_Z + | S n => zn2z_to_Z (double_wB n) (double_to_Z n) + end. + + Fixpoint extend_aux (n:nat) (x:zn2z w) {struct n}: word w (S n) := + match n return word w (S n) with + | O => x + | S n1 => WW W0 (extend_aux n1 x) + end. + + Definition extend (n:nat) (x:w) : word w (S n) := + let r := w_0W x in + match r with + | W0 => W0 + | _ => extend_aux n r + end. + + Definition double_0 n : word w n := + match n return word w n with + | O => w_0 + | S _ => W0 + end. + + Definition double_split (n:nat) (x:zn2z (word w n)) := + match x with + | W0 => + match n return word w n * word w n with + | O => (w_0,w_0) + | S _ => (W0, W0) + end + | WW h l => (h,l) + end. + + Definition ww_compare x y := + match x, y with + | W0, W0 => Eq + | W0, WW yh yl => + match w_compare w_0 yh with + | Eq => w_compare w_0 yl + | _ => Lt + end + | WW xh xl, W0 => + match w_compare xh w_0 with + | Eq => w_compare xl w_0 + | _ => Gt + end + | WW xh xl, WW yh yl => + match w_compare xh yh with + | Eq => w_compare xl yl + | Lt => Lt + | Gt => Gt + end + end. + + + (* Return the low part of the composed word*) + Fixpoint get_low (n : nat) {struct n}: + word w n -> w := + match n return (word w n -> w) with + | 0%nat => fun x => x + | S n1 => + fun x => + match x with + | W0 => w_0 + | WW _ x1 => get_low n1 x1 + end + end. + + + Section DoubleProof. + Notation wB := (base w_digits). + Notation wwB := (base ww_digits). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[[ x ]]" := (ww_to_Z x) (at level 0, x at level 99). + Notation "[+[ c ]]" := + (interp_carry 1 wwB ww_to_Z c) (at level 0, x at level 99). + Notation "[-[ c ]]" := + (interp_carry (-1) wwB ww_to_Z c) (at level 0, x at level 99). + Notation "[! n | x !]" := (double_to_Z n x) (at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_compare : forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + + Lemma wwB_wBwB : wwB = wB^2. + Proof. + unfold base, ww_digits;rewrite Zpower_2; rewrite (Zpos_xO w_digits). + replace (2 * Zpos w_digits) with (Zpos w_digits + Zpos w_digits). + apply Zpower_exp; unfold Zge;simpl;intros;discriminate. + ring. + Qed. + + Lemma spec_ww_1 : [[ww_1]] = 1. + Proof. simpl;rewrite spec_w_0;rewrite spec_w_1;ring. Qed. + + Lemma spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1. + Proof. simpl;rewrite spec_w_Bm1;rewrite wwB_wBwB;ring. Qed. + + Lemma lt_0_wB : 0 < wB. + Proof. + unfold base;apply Zpower_gt_0. unfold Zlt;reflexivity. + unfold Zle;intros H;discriminate H. + Qed. + + Lemma lt_0_wwB : 0 < wwB. + Proof. rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_lt_0_compat;apply lt_0_wB. Qed. + + Lemma wB_pos: 1 < wB. + Proof. + unfold base;apply Zlt_le_trans with (2^1). unfold Zlt;reflexivity. + apply Zpower_le_monotone. unfold Zlt;reflexivity. + split;unfold Zle;intros H. discriminate H. + clear spec_w_0W w_0W spec_w_Bm1 spec_to_Z spec_w_WW w_WW. + destruct w_digits; discriminate H. + Qed. + + Lemma wwB_pos: 1 < wwB. + Proof. + assert (H:= wB_pos);rewrite wwB_wBwB;rewrite <-(Zmult_1_r 1). + rewrite Zpower_2. + apply Zmult_lt_compat2;(split;[unfold Zlt;reflexivity|trivial]). + apply Zlt_le_weak;trivial. + Qed. + + Theorem wB_div_2: 2 * (wB / 2) = wB. + Proof. + clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W + spec_to_Z;unfold base. + assert (2 ^ Zpos w_digits = 2 * (2 ^ (Zpos w_digits - 1))). + pattern 2 at 2; rewrite <- Zpower_1_r. + rewrite <- Zpower_exp; auto with zarith. + f_equal; auto with zarith. + case w_digits; compute; intros; discriminate. + rewrite H; f_equal; auto with zarith. + rewrite Zmult_comm; apply Z_div_mult; auto with zarith. + Qed. + + Theorem wwB_div_2 : wwB / 2 = wB / 2 * wB. + Proof. + clear spec_w_0 w_0 spec_w_1 w_1 spec_w_Bm1 w_Bm1 spec_w_WW spec_w_0W + spec_to_Z. + rewrite wwB_wBwB; rewrite Zpower_2. + pattern wB at 1; rewrite <- wB_div_2; auto. + rewrite <- Zmult_assoc. + repeat (rewrite (Zmult_comm 2); rewrite Z_div_mult); auto with zarith. + Qed. + + Lemma mod_wwB : forall z x, + (z*wB + [|x|]) mod wwB = (z mod wB)*wB + [|x|]. + Proof. + intros z x. + rewrite Zplus_mod. + pattern wwB at 1;rewrite wwB_wBwB; rewrite Zpower_2. + rewrite Zmult_mod_distr_r;try apply lt_0_wB. + rewrite (Zmod_small [|x|]). + apply Zmod_small;rewrite wwB_wBwB;apply beta_mult;try apply spec_to_Z. + apply Z_mod_lt;apply Zlt_gt;apply lt_0_wB. + destruct (spec_to_Z x);split;trivial. + change [|x|] with (0*wB+[|x|]). rewrite wwB_wBwB. + rewrite Zpower_2;rewrite <- (Zplus_0_r (wB*wB));apply beta_lex_inv. + apply lt_0_wB. apply spec_to_Z. split;[apply Zle_refl | apply lt_0_wB]. + Qed. + + Lemma wB_div : forall x y, ([|x|] * wB + [|y|]) / wB = [|x|]. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + intros x y;unfold base;rewrite Zdiv_shift_r;auto with zarith. + rewrite Z_div_mult;auto with zarith. + destruct (spec_to_Z x);trivial. + Qed. + + Lemma wB_div_plus : forall x y p, + 0 <= p -> + ([|x|]*wB + [|y|]) / 2^(Zpos w_digits + p) = [|x|] / 2^p. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + intros x y p Hp;rewrite Zpower_exp;auto with zarith. + rewrite <- Zdiv_Zdiv;auto with zarith. + rewrite wB_div;trivial. + Qed. + + Lemma lt_wB_wwB : wB < wwB. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + unfold base;apply Zpower_lt_monotone;auto with zarith. + assert (0 < Zpos w_digits). compute;reflexivity. + unfold ww_digits;rewrite Zpos_xO;auto with zarith. + Qed. + + Lemma w_to_Z_wwB : forall x, x < wB -> x < wwB. + Proof. + intros x H;apply Zlt_trans with wB;trivial;apply lt_wB_wwB. + Qed. + + Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + destruct x as [ |h l];simpl. + split;[apply Zle_refl|apply lt_0_wwB]. + assert (H:=spec_to_Z h);assert (L:=spec_to_Z l);split. + apply Zplus_le_0_compat;auto with zarith. + rewrite <- (Zplus_0_r wwB);rewrite wwB_wBwB; rewrite Zpower_2; + apply beta_lex_inv;auto with zarith. + Qed. + + Lemma double_wB_wwB : forall n, double_wB n * double_wB n = double_wB (S n). + Proof. + intros n;unfold double_wB;simpl. + unfold base;rewrite (Zpos_xO (double_digits n)). + replace (2 * Zpos (double_digits n)) with + (Zpos (double_digits n) + Zpos (double_digits n)). + symmetry; apply Zpower_exp;intro;discriminate. + ring. + Qed. + + Lemma double_wB_pos: + forall n, 0 <= double_wB n. + Proof. + intros n; unfold double_wB, base; auto with zarith. + Qed. + + Lemma double_wB_more_digits: + forall n, wB <= double_wB n. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + intros n; elim n; clear n; auto. + unfold double_wB, double_digits; auto with zarith. + intros n H1; rewrite <- double_wB_wwB. + apply Zle_trans with (wB * 1). + rewrite Zmult_1_r; apply Zle_refl. + apply Zmult_le_compat; auto with zarith. + apply Zle_trans with wB; auto with zarith. + unfold base. + rewrite <- (Zpower_0_r 2). + apply Zpower_le_monotone2; auto with zarith. + unfold base; auto with zarith. + Qed. + + Lemma spec_double_to_Z : + forall n (x:word w n), 0 <= [!n | x!] < double_wB n. + Proof. + clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1. + induction n;intros. exact (spec_to_Z x). + unfold double_to_Z;fold double_to_Z. + destruct x;unfold zn2z_to_Z. + unfold double_wB,base;split;auto with zarith. + assert (U0:= IHn w0);assert (U1:= IHn w1). + split;auto with zarith. + apply Zlt_le_trans with ((double_wB n - 1) * double_wB n + double_wB n). + assert (double_to_Z n w0*double_wB n <= (double_wB n - 1)*double_wB n). + apply Zmult_le_compat_r;auto with zarith. + auto with zarith. + rewrite <- double_wB_wwB. + replace ((double_wB n - 1) * double_wB n + double_wB n) with (double_wB n * double_wB n); + [auto with zarith | ring]. + Qed. + + Lemma spec_get_low: + forall n x, + [!n | x!] < wB -> [|get_low n x|] = [!n | x!]. + Proof. + clear spec_w_1 spec_w_Bm1. + intros n; elim n; auto; clear n. + intros n Hrec x; case x; clear x; auto. + intros xx yy H1; simpl in H1. + assert (F1: [!n | xx!] = 0). + case (Zle_lt_or_eq 0 ([!n | xx!])); auto. + case (spec_double_to_Z n xx); auto. + intros F2. + assert (F3 := double_wB_more_digits n). + assert (F4: 0 <= [!n | yy!]). + case (spec_double_to_Z n yy); auto. + assert (F5: 1 * wB <= [!n | xx!] * double_wB n); + auto with zarith. + apply Zmult_le_compat; auto with zarith. + unfold base; auto with zarith. + simpl get_low; simpl double_to_Z. + generalize H1; clear H1. + rewrite F1; rewrite Zmult_0_l; rewrite Zplus_0_l. + intros H1; apply Hrec; auto. + Qed. + + Lemma spec_double_WW : forall n (h l : word w n), + [!S n|double_WW n h l!] = [!n|h!] * double_wB n + [!n|l!]. + Proof. + induction n;simpl;intros;trivial. + destruct h;auto. + destruct l;auto. + Qed. + + Lemma spec_extend_aux : forall n x, [!S n|extend_aux n x!] = [[x]]. + Proof. induction n;simpl;trivial. Qed. + + Lemma spec_extend : forall n x, [!S n|extend n x!] = [|x|]. + Proof. + intros n x;assert (H:= spec_w_0W x);unfold extend. + destruct (w_0W x);simpl;trivial. + rewrite <- H;exact (spec_extend_aux n (WW w0 w1)). + Qed. + + Lemma spec_double_0 : forall n, [!n|double_0 n!] = 0. + Proof. destruct n;trivial. Qed. + + Lemma spec_double_split : forall n x, + let (h,l) := double_split n x in + [!S n|x!] = [!n|h!] * double_wB n + [!n|l!]. + Proof. + destruct x;simpl;auto. + destruct n;simpl;trivial. + rewrite spec_w_0;trivial. + Qed. + + Lemma wB_lex_inv: forall a b c d, + a < c -> + a * wB + [|b|] < c * wB + [|d|]. + Proof. + intros a b c d H1; apply beta_lex_inv with (1 := H1); auto. + Qed. + + Lemma spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Proof. + destruct x as [ |xh xl];destruct y as [ |yh yl];simpl;trivial. + generalize (spec_w_compare w_0 yh);destruct (w_compare w_0 yh); + intros H;rewrite spec_w_0 in H. + rewrite <- H;simpl;rewrite <- spec_w_0;apply spec_w_compare. + change 0 with (0*wB+0);pattern 0 at 2;rewrite <- spec_w_0. + apply wB_lex_inv;trivial. + absurd (0 <= [|yh|]). apply Zgt_not_le;trivial. + destruct (spec_to_Z yh);trivial. + generalize (spec_w_compare xh w_0);destruct (w_compare xh w_0); + intros H;rewrite spec_w_0 in H. + rewrite H;simpl;rewrite <- spec_w_0;apply spec_w_compare. + absurd (0 <= [|xh|]). apply Zgt_not_le;apply Zlt_gt;trivial. + destruct (spec_to_Z xh);trivial. + apply Zlt_gt;change 0 with (0*wB+0);pattern 0 at 2;rewrite <- spec_w_0. + apply wB_lex_inv;apply Zgt_lt;trivial. + + generalize (spec_w_compare xh yh);destruct (w_compare xh yh);intros H. + rewrite H;generalize (spec_w_compare xl yl);destruct (w_compare xl yl); + intros H1;[rewrite H1|apply Zplus_lt_compat_l|apply Zplus_gt_compat_l]; + trivial. + apply wB_lex_inv;trivial. + apply Zlt_gt;apply wB_lex_inv;apply Zgt_lt;trivial. + Qed. + + + End DoubleProof. + +End DoubleBase. + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v new file mode 100644 index 00000000..cca32a59 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleCyclic.v @@ -0,0 +1,885 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleCyclic.v 11012 2008-05-28 16:34:43Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. +Require Import DoubleAdd. +Require Import DoubleSub. +Require Import DoubleMul. +Require Import DoubleSqrt. +Require Import DoubleLift. +Require Import DoubleDivn1. +Require Import DoubleDiv. +Require Import CyclicAxioms. + +Open Local Scope Z_scope. + + +Section Z_2nZ. + + Variable w : Type. + Variable w_op : znz_op w. + Let w_digits := w_op.(znz_digits). + Let w_zdigits := w_op.(znz_zdigits). + + Let w_to_Z := w_op.(znz_to_Z). + Let w_of_pos := w_op.(znz_of_pos). + Let w_head0 := w_op.(znz_head0). + Let w_tail0 := w_op.(znz_tail0). + + Let w_0 := w_op.(znz_0). + Let w_1 := w_op.(znz_1). + Let w_Bm1 := w_op.(znz_Bm1). + + Let w_compare := w_op.(znz_compare). + Let w_eq0 := w_op.(znz_eq0). + + Let w_opp_c := w_op.(znz_opp_c). + Let w_opp := w_op.(znz_opp). + Let w_opp_carry := w_op.(znz_opp_carry). + + Let w_succ_c := w_op.(znz_succ_c). + Let w_add_c := w_op.(znz_add_c). + Let w_add_carry_c := w_op.(znz_add_carry_c). + Let w_succ := w_op.(znz_succ). + Let w_add := w_op.(znz_add). + Let w_add_carry := w_op.(znz_add_carry). + + Let w_pred_c := w_op.(znz_pred_c). + Let w_sub_c := w_op.(znz_sub_c). + Let w_sub_carry_c := w_op.(znz_sub_carry_c). + Let w_pred := w_op.(znz_pred). + Let w_sub := w_op.(znz_sub). + Let w_sub_carry := w_op.(znz_sub_carry). + + + Let w_mul_c := w_op.(znz_mul_c). + Let w_mul := w_op.(znz_mul). + Let w_square_c := w_op.(znz_square_c). + + Let w_div21 := w_op.(znz_div21). + Let w_div_gt := w_op.(znz_div_gt). + Let w_div := w_op.(znz_div). + + Let w_mod_gt := w_op.(znz_mod_gt). + Let w_mod := w_op.(znz_mod). + + Let w_gcd_gt := w_op.(znz_gcd_gt). + Let w_gcd := w_op.(znz_gcd). + + Let w_add_mul_div := w_op.(znz_add_mul_div). + + Let w_pos_mod := w_op.(znz_pos_mod). + + Let w_is_even := w_op.(znz_is_even). + Let w_sqrt2 := w_op.(znz_sqrt2). + Let w_sqrt := w_op.(znz_sqrt). + + Let _zn2z := zn2z w. + + Let wB := base w_digits. + + Let w_Bm2 := w_pred w_Bm1. + + Let ww_1 := ww_1 w_0 w_1. + Let ww_Bm1 := ww_Bm1 w_Bm1. + + Let w_add2 a b := match w_add_c a b with C0 p => WW w_0 p | C1 p => WW w_1 p end. + + Let _ww_digits := xO w_digits. + + Let _ww_zdigits := w_add2 w_zdigits w_zdigits. + + Let to_Z := zn2z_to_Z wB w_to_Z. + + Let w_W0 := znz_W0 w_op. + Let w_0W := znz_0W w_op. + Let w_WW := znz_WW w_op. + + Let ww_of_pos p := + match w_of_pos p with + | (N0, l) => (N0, WW w_0 l) + | (Npos ph,l) => + let (n,h) := w_of_pos ph in (n, w_WW h l) + end. + + Let head0 := + Eval lazy beta delta [ww_head0] in + ww_head0 w_0 w_0W w_compare w_head0 w_add2 w_zdigits _ww_zdigits. + + Let tail0 := + Eval lazy beta delta [ww_tail0] in + ww_tail0 w_0 w_0W w_compare w_tail0 w_add2 w_zdigits _ww_zdigits. + + Let ww_WW := Eval lazy beta delta [ww_WW] in (@ww_WW w). + Let ww_0W := Eval lazy beta delta [ww_0W] in (@ww_0W w). + Let ww_W0 := Eval lazy beta delta [ww_W0] in (@ww_W0 w). + + (* ** Comparison ** *) + Let compare := + Eval lazy beta delta[ww_compare] in ww_compare w_0 w_compare. + + Let eq0 (x:zn2z w) := + match x with + | W0 => true + | _ => false + end. + + (* ** Opposites ** *) + Let opp_c := + Eval lazy beta delta [ww_opp_c] in ww_opp_c w_0 w_opp_c w_opp_carry. + + Let opp := + Eval lazy beta delta [ww_opp] in ww_opp w_0 w_opp_c w_opp_carry w_opp. + + Let opp_carry := + Eval lazy beta delta [ww_opp_carry] in ww_opp_carry w_WW ww_Bm1 w_opp_carry. + + (* ** Additions ** *) + + Let succ_c := + Eval lazy beta delta [ww_succ_c] in ww_succ_c w_0 ww_1 w_succ_c. + + Let add_c := + Eval lazy beta delta [ww_add_c] in ww_add_c w_WW w_add_c w_add_carry_c. + + Let add_carry_c := + Eval lazy beta iota delta [ww_add_carry_c ww_succ_c] in + ww_add_carry_c w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c. + + Let succ := + Eval lazy beta delta [ww_succ] in ww_succ w_W0 ww_1 w_succ_c w_succ. + + Let add := + Eval lazy beta delta [ww_add] in ww_add w_add_c w_add w_add_carry. + + Let add_carry := + Eval lazy beta iota delta [ww_add_carry ww_succ] in + ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ w_add w_add_carry. + + (* ** Subtractions ** *) + + Let pred_c := + Eval lazy beta delta [ww_pred_c] in ww_pred_c w_Bm1 w_WW ww_Bm1 w_pred_c. + + Let sub_c := + Eval lazy beta iota delta [ww_sub_c ww_opp_c] in + ww_sub_c w_0 w_WW w_opp_c w_opp_carry w_sub_c w_sub_carry_c. + + Let sub_carry_c := + Eval lazy beta iota delta [ww_sub_carry_c ww_pred_c ww_opp_carry] in + ww_sub_carry_c w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c w_sub_c w_sub_carry_c. + + Let pred := + Eval lazy beta delta [ww_pred] in ww_pred w_Bm1 w_WW ww_Bm1 w_pred_c w_pred. + + Let sub := + Eval lazy beta iota delta [ww_sub ww_opp] in + ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry. + + Let sub_carry := + Eval lazy beta iota delta [ww_sub_carry ww_pred ww_opp_carry] in + ww_sub_carry w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c w_sub_carry_c w_pred + w_sub w_sub_carry. + + + (* ** Multiplication ** *) + + Let mul_c := + Eval lazy beta iota delta [ww_mul_c double_mul_c] in + ww_mul_c w_0 w_1 w_WW w_W0 w_mul_c add_c add add_carry. + + Let karatsuba_c := + Eval lazy beta iota delta [ww_karatsuba_c double_mul_c kara_prod] in + ww_karatsuba_c w_0 w_1 w_WW w_W0 w_compare w_add w_sub w_mul_c + add_c add add_carry sub_c sub. + + Let mul := + Eval lazy beta delta [ww_mul] in + ww_mul w_W0 w_add w_mul_c w_mul add. + + Let square_c := + Eval lazy beta delta [ww_square_c] in + ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add add_carry. + + (* Division operation *) + + Let div32 := + Eval lazy beta iota delta [w_div32] in + w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c + w_add w_add_carry w_pred w_sub w_mul_c w_div21 sub_c. + + Let div21 := + Eval lazy beta iota delta [ww_div21] in + ww_div21 w_0 w_0W div32 ww_1 compare sub. + + Let low (p: zn2z w) := match p with WW _ p1 => p1 | _ => w_0 end. + + Let add_mul_div := + Eval lazy beta delta [ww_add_mul_div] in + ww_add_mul_div w_0 w_WW w_W0 w_0W compare w_add_mul_div sub w_zdigits low. + + Let div_gt := + Eval lazy beta delta [ww_div_gt] in + ww_div_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp + w_opp_carry w_sub_c w_sub w_sub_carry + w_div_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits. + + Let div := + Eval lazy beta delta [ww_div] in ww_div ww_1 compare div_gt. + + Let mod_gt := + Eval lazy beta delta [ww_mod_gt] in + ww_mod_gt w_0 w_WW w_0W w_compare w_eq0 w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry + w_mod_gt w_add_mul_div w_head0 w_div21 div32 _ww_zdigits add_mul_div w_zdigits. + + Let mod_ := + Eval lazy beta delta [ww_mod] in ww_mod compare mod_gt. + + Let pos_mod := + Eval lazy beta delta [ww_pos_mod] in + ww_pos_mod w_0 w_zdigits w_WW w_pos_mod compare w_0W low sub _ww_zdigits. + + Let is_even := + Eval lazy beta delta [ww_is_even] in ww_is_even w_is_even. + + Let sqrt2 := + Eval lazy beta delta [ww_sqrt2] in + ww_sqrt2 w_is_even w_compare w_0 w_1 w_Bm1 w_0W w_sub w_square_c + w_div21 w_add_mul_div w_zdigits w_add_c w_sqrt2 w_pred pred_c + pred add_c add sub_c add_mul_div. + + Let sqrt := + Eval lazy beta delta [ww_sqrt] in + ww_sqrt w_is_even w_0 w_sub w_add_mul_div w_zdigits + _ww_zdigits w_sqrt2 pred add_mul_div head0 compare low. + + Let gcd_gt_fix := + Eval cbv beta delta [ww_gcd_gt_aux ww_gcd_gt_body] in + ww_gcd_gt_aux w_0 w_WW w_0W w_compare w_opp_c w_opp w_opp_carry + w_sub_c w_sub w_sub_carry w_gcd_gt + w_add_mul_div w_head0 w_div21 div32 _ww_zdigits add_mul_div + w_zdigits. + + Let gcd_cont := + Eval lazy beta delta [gcd_cont] in gcd_cont ww_1 w_1 w_compare. + + Let gcd_gt := + Eval lazy beta delta [ww_gcd_gt] in + ww_gcd_gt w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont. + + Let gcd := + Eval lazy beta delta [ww_gcd] in + ww_gcd compare w_0 w_eq0 w_gcd_gt _ww_digits gcd_gt_fix gcd_cont. + + (* ** Record of operators on 2 words *) + + Definition mk_zn2z_op := + mk_znz_op _ww_digits _ww_zdigits + to_Z ww_of_pos head0 tail0 + W0 ww_1 ww_Bm1 + compare eq0 + opp_c opp opp_carry + succ_c add_c add_carry_c + succ add add_carry + pred_c sub_c sub_carry_c + pred sub sub_carry + mul_c mul square_c + div21 div_gt div + mod_gt mod_ + gcd_gt gcd + add_mul_div + pos_mod + is_even + sqrt2 + sqrt. + + Definition mk_zn2z_op_karatsuba := + mk_znz_op _ww_digits _ww_zdigits + to_Z ww_of_pos head0 tail0 + W0 ww_1 ww_Bm1 + compare eq0 + opp_c opp opp_carry + succ_c add_c add_carry_c + succ add add_carry + pred_c sub_c sub_carry_c + pred sub sub_carry + karatsuba_c mul square_c + div21 div_gt div + mod_gt mod_ + gcd_gt gcd + add_mul_div + pos_mod + is_even + sqrt2 + sqrt. + + (* Proof *) + Variable op_spec : znz_spec w_op. + + Hint Resolve + (spec_to_Z op_spec) + (spec_of_pos op_spec) + (spec_0 op_spec) + (spec_1 op_spec) + (spec_Bm1 op_spec) + (spec_compare op_spec) + (spec_eq0 op_spec) + (spec_opp_c op_spec) + (spec_opp op_spec) + (spec_opp_carry op_spec) + (spec_succ_c op_spec) + (spec_add_c op_spec) + (spec_add_carry_c op_spec) + (spec_succ op_spec) + (spec_add op_spec) + (spec_add_carry op_spec) + (spec_pred_c op_spec) + (spec_sub_c op_spec) + (spec_sub_carry_c op_spec) + (spec_pred op_spec) + (spec_sub op_spec) + (spec_sub_carry op_spec) + (spec_mul_c op_spec) + (spec_mul op_spec) + (spec_square_c op_spec) + (spec_div21 op_spec) + (spec_div_gt op_spec) + (spec_div op_spec) + (spec_mod_gt op_spec) + (spec_mod op_spec) + (spec_gcd_gt op_spec) + (spec_gcd op_spec) + (spec_head0 op_spec) + (spec_tail0 op_spec) + (spec_add_mul_div op_spec) + (spec_pos_mod) + (spec_is_even) + (spec_sqrt2) + (spec_sqrt) + (spec_W0 op_spec) + (spec_0W op_spec) + (spec_WW op_spec). + + Ltac wwauto := unfold ww_to_Z; auto. + + Let wwB := base _ww_digits. + + Notation "[| x |]" := (to_Z x) (at level 0, x at level 99). + + Notation "[+| c |]" := + (interp_carry 1 wwB to_Z c) (at level 0, x at level 99). + + Notation "[-| c |]" := + (interp_carry (-1) wwB to_Z c) (at level 0, x at level 99). + + Notation "[[ x ]]" := (zn2z_to_Z wwB to_Z x) (at level 0, x at level 99). + + Let spec_ww_to_Z : forall x, 0 <= [| x |] < wwB. + Proof. refine (spec_ww_to_Z w_digits w_to_Z _);auto. Qed. + + Let spec_ww_of_pos : forall p, + Zpos p = (Z_of_N (fst (ww_of_pos p)))*wwB + [|(snd (ww_of_pos p))|]. + Proof. + unfold ww_of_pos;intros. + assert (H:= spec_of_pos op_spec p);unfold w_of_pos; + destruct (znz_of_pos w_op p). simpl in H. + rewrite H;clear H;destruct n;simpl to_Z. + simpl;unfold w_to_Z,w_0;rewrite (spec_0 op_spec);trivial. + unfold Z_of_N; assert (H:= spec_of_pos op_spec p0); + destruct (znz_of_pos w_op p0). simpl in H. + rewrite H;unfold fst, snd,Z_of_N, to_Z. + rewrite (spec_WW op_spec). + replace wwB with (wB*wB). + unfold wB,w_to_Z,w_digits;clear H;destruct n;ring. + symmetry. rewrite <- Zpower_2; exact (wwB_wBwB w_digits). + Qed. + + Let spec_ww_0 : [|W0|] = 0. + Proof. reflexivity. Qed. + + Let spec_ww_1 : [|ww_1|] = 1. + Proof. refine (spec_ww_1 w_0 w_1 w_digits w_to_Z _ _);auto. Qed. + + Let spec_ww_Bm1 : [|ww_Bm1|] = wwB - 1. + Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed. + + Let spec_ww_compare : + forall x y, + match compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Proof. + refine (spec_ww_compare w_0 w_digits w_to_Z w_compare _ _ _);auto. + exact (spec_compare op_spec). + Qed. + + Let spec_ww_eq0 : forall x, eq0 x = true -> [|x|] = 0. + Proof. destruct x;simpl;intros;trivial;discriminate. Qed. + + Let spec_ww_opp_c : forall x, [-|opp_c x|] = -[|x|]. + Proof. + refine(spec_ww_opp_c w_0 w_0 W0 w_opp_c w_opp_carry w_digits w_to_Z _ _ _ _); + auto. + Qed. + + Let spec_ww_opp : forall x, [|opp x|] = (-[|x|]) mod wwB. + Proof. + refine(spec_ww_opp w_0 w_0 W0 w_opp_c w_opp_carry w_opp + w_digits w_to_Z _ _ _ _ _); + auto. + Qed. + + Let spec_ww_opp_carry : forall x, [|opp_carry x|] = wwB - [|x|] - 1. + Proof. + refine (spec_ww_opp_carry w_WW ww_Bm1 w_opp_carry w_digits w_to_Z _ _ _); + wwauto. + Qed. + + Let spec_ww_succ_c : forall x, [+|succ_c x|] = [|x|] + 1. + Proof. + refine (spec_ww_succ_c w_0 w_0 ww_1 w_succ_c w_digits w_to_Z _ _ _ _);auto. + Qed. + + Let spec_ww_add_c : forall x y, [+|add_c x y|] = [|x|] + [|y|]. + Proof. + refine (spec_ww_add_c w_WW w_add_c w_add_carry_c w_digits w_to_Z _ _ _);wwauto. + Qed. + + Let spec_ww_add_carry_c : forall x y, [+|add_carry_c x y|] = [|x|]+[|y|]+1. + Proof. + refine (spec_ww_add_carry_c w_0 w_0 w_WW ww_1 w_succ_c w_add_c w_add_carry_c + w_digits w_to_Z _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_succ : forall x, [|succ x|] = ([|x|] + 1) mod wwB. + Proof. + refine (spec_ww_succ w_W0 ww_1 w_succ_c w_succ w_digits w_to_Z _ _ _ _ _); + wwauto. + Qed. + + Let spec_ww_add : forall x y, [|add x y|] = ([|x|] + [|y|]) mod wwB. + Proof. + refine (spec_ww_add w_add_c w_add w_add_carry w_digits w_to_Z _ _ _ _);auto. + Qed. + + Let spec_ww_add_carry : forall x y, [|add_carry x y|]=([|x|]+[|y|]+1)mod wwB. + Proof. + refine (spec_ww_add_carry w_W0 ww_1 w_succ_c w_add_carry_c w_succ + w_add w_add_carry w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_pred_c : forall x, [-|pred_c x|] = [|x|] - 1. + Proof. + refine (spec_ww_pred_c w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_digits w_to_Z + _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_sub_c : forall x y, [-|sub_c x y|] = [|x|] - [|y|]. + Proof. + refine (spec_ww_sub_c w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c + w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_sub_carry_c : forall x y, [-|sub_carry_c x y|] = [|x|]-[|y|]-1. + Proof. + refine (spec_ww_sub_carry_c w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c + w_sub_c w_sub_carry_c w_digits w_to_Z _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_pred : forall x, [|pred x|] = ([|x|] - 1) mod wwB. + Proof. + refine (spec_ww_pred w_0 w_Bm1 w_WW ww_Bm1 w_pred_c w_pred w_digits w_to_Z + _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_sub : forall x y, [|sub x y|] = ([|x|] - [|y|]) mod wwB. + Proof. + refine (spec_ww_sub w_0 w_0 w_WW W0 w_opp_c w_opp_carry w_sub_c w_opp + w_sub w_sub_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_sub_carry : forall x y, [|sub_carry x y|]=([|x|]-[|y|]-1) mod wwB. + Proof. + refine (spec_ww_sub_carry w_0 w_Bm1 w_WW ww_Bm1 w_opp_carry w_pred_c + w_sub_carry_c w_pred w_sub w_sub_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _); + wwauto. + Qed. + + Let spec_ww_mul_c : forall x y, [[mul_c x y ]] = [|x|] * [|y|]. + Proof. + refine (spec_ww_mul_c w_0 w_1 w_WW w_W0 w_mul_c add_c add add_carry w_digits + w_to_Z _ _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_ww_karatsuba_c : forall x y, [[karatsuba_c x y ]] = [|x|] * [|y|]. + Proof. + refine (spec_ww_karatsuba_c _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + _ _ _ _ _ _ _ _ _ _ _ _); wwauto. + unfold w_digits; apply spec_more_than_1_digit; auto. + exact (spec_compare op_spec). + Qed. + + Let spec_ww_mul : forall x y, [|mul x y|] = ([|x|] * [|y|]) mod wwB. + Proof. + refine (spec_ww_mul w_W0 w_add w_mul_c w_mul add w_digits w_to_Z _ _ _ _ _); + wwauto. + Qed. + + Let spec_ww_square_c : forall x, [[square_c x]] = [|x|] * [|x|]. + Proof. + refine (spec_ww_square_c w_0 w_1 w_WW w_W0 w_mul_c w_square_c add_c add + add_carry w_digits w_to_Z _ _ _ _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_w_div32 : forall a1 a2 a3 b1 b2, + wB / 2 <= (w_to_Z b1) -> + [|WW a1 a2|] < [|WW b1 b2|] -> + let (q, r) := div32 a1 a2 a3 b1 b2 in + (w_to_Z a1) * wwB + (w_to_Z a2) * wB + (w_to_Z a3) = + (w_to_Z q) * ((w_to_Z b1)*wB + (w_to_Z b2)) + [|r|] /\ + 0 <= [|r|] < (w_to_Z b1)*wB + w_to_Z b2. + Proof. + refine (spec_w_div32 w_0 w_Bm1 w_Bm2 w_WW w_compare w_add_c w_add_carry_c + w_add w_add_carry w_pred w_sub w_mul_c w_div21 sub_c w_digits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + unfold w_Bm2, w_to_Z, w_pred, w_Bm1. + rewrite (spec_pred op_spec);rewrite (spec_Bm1 op_spec). + unfold w_digits;rewrite Zmod_small. ring. + assert (H:= wB_pos(znz_digits w_op)). omega. + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + Qed. + + Let spec_ww_div21 : forall a1 a2 b, + wwB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := div21 a1 a2 b in + [|a1|] *wwB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + refine (spec_ww_div21 w_0 w_0W div32 ww_1 compare sub w_digits w_to_Z + _ _ _ _ _ _ _);wwauto. + Qed. + + Let spec_add2: forall x y, + [|w_add2 x y|] = w_to_Z x + w_to_Z y. + unfold w_add2. + intros xh xl; generalize (spec_add_c op_spec xh xl). + unfold w_add_c; case znz_add_c; unfold interp_carry; simpl ww_to_Z. + intros w0 Hw0; simpl; unfold w_to_Z; rewrite Hw0. + unfold w_0; rewrite spec_0; simpl; auto with zarith. + intros w0; rewrite Zmult_1_l; simpl. + unfold w_to_Z, w_1; rewrite spec_1; auto with zarith. + rewrite Zmult_1_l; auto. + Qed. + + Let spec_low: forall x, + w_to_Z (low x) = [|x|] mod wB. + intros x; case x; simpl low. + unfold ww_to_Z, w_to_Z, w_0; rewrite (spec_0 op_spec); simpl. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + unfold wB, base; auto with zarith. + intros xh xl; simpl. + rewrite Zplus_comm; rewrite Z_mod_plus; auto with zarith. + rewrite Zmod_small; auto with zarith. + unfold wB, base; auto with zarith. + Qed. + + Let spec_ww_digits: + [|_ww_zdigits|] = Zpos (xO w_digits). + Proof. + unfold w_to_Z, _ww_zdigits. + rewrite spec_add2. + unfold w_to_Z, w_zdigits, w_digits. + rewrite spec_zdigits; auto. + rewrite Zpos_xO; auto with zarith. + Qed. + + + Let spec_ww_head00 : forall x, [|x|] = 0 -> [|head0 x|] = Zpos _ww_digits. + Proof. + refine (spec_ww_head00 w_0 w_0W + w_compare w_head0 w_add2 w_zdigits _ww_zdigits + w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); auto. + exact (spec_compare op_spec). + exact (spec_head00 op_spec). + exact (spec_zdigits op_spec). + Qed. + + Let spec_ww_head0 : forall x, 0 < [|x|] -> + wwB/ 2 <= 2 ^ [|head0 x|] * [|x|] < wwB. + Proof. + refine (spec_ww_head0 w_0 w_0W w_compare w_head0 + w_add2 w_zdigits _ww_zdigits + w_to_Z _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_zdigits op_spec). + Qed. + + Let spec_ww_tail00 : forall x, [|x|] = 0 -> [|tail0 x|] = Zpos _ww_digits. + Proof. + refine (spec_ww_tail00 w_0 w_0W + w_compare w_tail0 w_add2 w_zdigits _ww_zdigits + w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); wwauto. + exact (spec_compare op_spec). + exact (spec_tail00 op_spec). + exact (spec_zdigits op_spec). + Qed. + + + Let spec_ww_tail0 : forall x, 0 < [|x|] -> + exists y, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail0 x|]. + Proof. + refine (spec_ww_tail0 (w_digits := w_digits) w_0 w_0W w_compare w_tail0 + w_add2 w_zdigits _ww_zdigits w_to_Z _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_zdigits op_spec). + Qed. + + Lemma spec_ww_add_mul_div : forall x y p, + [|p|] <= Zpos _ww_digits -> + [| add_mul_div p x y |] = + ([|x|] * (2 ^ [|p|]) + + [|y|] / (2 ^ ((Zpos _ww_digits) - [|p|]))) mod wwB. + Proof. + refine (@spec_ww_add_mul_div w w_0 w_WW w_W0 w_0W compare w_add_mul_div + sub w_digits w_zdigits low w_to_Z + _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_zdigits op_spec). + Qed. + + Let spec_ww_div_gt : forall a b, + [|a|] > [|b|] -> 0 < [|b|] -> + let (q,r) := div_gt a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]. + Proof. +refine +(@spec_ww_div_gt w w_digits w_0 w_WW w_0W w_compare w_eq0 + w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt + w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ +). + exact (spec_0 op_spec). + exact (spec_to_Z op_spec). + wwauto. + wwauto. + exact (spec_compare op_spec). + exact (spec_eq0 op_spec). + exact (spec_opp_c op_spec). + exact (spec_opp op_spec). + exact (spec_opp_carry op_spec). + exact (spec_sub_c op_spec). + exact (spec_sub op_spec). + exact (spec_sub_carry op_spec). + exact (spec_div_gt op_spec). + exact (spec_add_mul_div op_spec). + exact (spec_head0 op_spec). + exact (spec_div21 op_spec). + exact spec_w_div32. + exact (spec_zdigits op_spec). + exact spec_ww_digits. + exact spec_ww_1. + exact spec_ww_add_mul_div. + Qed. + + Let spec_ww_div : forall a b, 0 < [|b|] -> + let (q,r) := div a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + refine (spec_ww_div w_digits ww_1 compare div_gt w_to_Z _ _ _ _);auto. + Qed. + + Let spec_ww_mod_gt : forall a b, + [|a|] > [|b|] -> 0 < [|b|] -> + [|mod_gt a b|] = [|a|] mod [|b|]. + Proof. + refine (@spec_ww_mod_gt w w_digits w_0 w_WW w_0W w_compare w_eq0 + w_opp_c w_opp w_opp_carry w_sub_c w_sub w_sub_carry w_div_gt w_mod_gt + w_add_mul_div w_head0 w_div21 div32 _ww_zdigits ww_1 add_mul_div + w_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_div_gt op_spec). + exact (spec_div21 op_spec). + exact (spec_zdigits op_spec). + exact spec_ww_add_mul_div. + Qed. + + Let spec_ww_mod : forall a b, 0 < [|b|] -> [|mod_ a b|] = [|a|] mod [|b|]. + Proof. + refine (spec_ww_mod w_digits W0 compare mod_gt w_to_Z _ _ _);auto. + Qed. + + Let spec_ww_gcd_gt : forall a b, [|a|] > [|b|] -> + Zis_gcd [|a|] [|b|] [|gcd_gt a b|]. + Proof. + refine (@spec_ww_gcd_gt w w_digits W0 w_to_Z _ + w_0 w_0 w_eq0 w_gcd_gt _ww_digits + _ gcd_gt_fix _ _ _ _ gcd_cont _);auto. + refine (@spec_ww_gcd_gt_aux w w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp + w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0 + w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + exact (spec_zdigits op_spec). + exact spec_ww_add_mul_div. + refine (@spec_gcd_cont w w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare + _ _);auto. + exact (spec_compare op_spec). + Qed. + + Let spec_ww_gcd : forall a b, Zis_gcd [|a|] [|b|] [|gcd a b|]. + Proof. + refine (@spec_ww_gcd w w_digits W0 compare w_to_Z _ _ w_0 w_0 w_eq0 w_gcd_gt + _ww_digits _ gcd_gt_fix _ _ _ _ gcd_cont _);auto. + refine (@spec_ww_gcd_gt_aux w w_digits w_0 w_WW w_0W w_compare w_opp_c w_opp + w_opp_carry w_sub_c w_sub w_sub_carry w_gcd_gt w_add_mul_div w_head0 + w_div21 div32 _ww_zdigits ww_1 add_mul_div w_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + exact (spec_zdigits op_spec). + exact spec_ww_add_mul_div. + refine (@spec_gcd_cont w w_digits ww_1 w_to_Z _ _ w_0 w_1 w_compare + _ _);auto. + exact (spec_compare op_spec). + Qed. + + Let spec_ww_is_even : forall x, + match is_even x with + true => [|x|] mod 2 = 0 + | false => [|x|] mod 2 = 1 + end. + Proof. + refine (@spec_ww_is_even w w_is_even w_0 w_1 w_Bm1 w_digits _ _ _ _ _); auto. + exact (spec_is_even op_spec). + Qed. + + Let spec_ww_sqrt2 : forall x y, + wwB/ 4 <= [|x|] -> + let (s,r) := sqrt2 x y in + [[WW x y]] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]. + Proof. + intros x y H. + refine (@spec_ww_sqrt2 w w_is_even w_compare w_0 w_1 w_Bm1 + w_0W w_sub w_square_c w_div21 w_add_mul_div w_digits w_zdigits + _ww_zdigits + w_add_c w_sqrt2 w_pred pred_c pred add_c add sub_c add_mul_div + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto. + exact (spec_zdigits op_spec). + exact (spec_more_than_1_digit op_spec). + exact (spec_is_even op_spec). + exact (spec_compare op_spec). + exact (spec_div21 op_spec). + exact (spec_ww_add_mul_div). + exact (spec_sqrt2 op_spec). + Qed. + + Let spec_ww_sqrt : forall x, + [|sqrt x|] ^ 2 <= [|x|] < ([|sqrt x|] + 1) ^ 2. + Proof. + refine (@spec_ww_sqrt w w_is_even w_0 w_1 w_Bm1 + w_sub w_add_mul_div w_digits w_zdigits _ww_zdigits + w_sqrt2 pred add_mul_div head0 compare + _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _); wwauto. + exact (spec_zdigits op_spec). + exact (spec_more_than_1_digit op_spec). + exact (spec_is_even op_spec). + exact (spec_ww_add_mul_div). + exact (spec_sqrt2 op_spec). + Qed. + + Lemma mk_znz2_spec : znz_spec mk_zn2z_op. + Proof. + apply mk_znz_spec;auto. + exact spec_ww_add_mul_div. + + refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW + w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_pos_mod op_spec). + exact (spec_zdigits op_spec). + unfold w_to_Z, w_zdigits. + rewrite (spec_zdigits op_spec). + rewrite <- Zpos_xO; exact spec_ww_digits. + Qed. + + Lemma mk_znz2_karatsuba_spec : znz_spec mk_zn2z_op_karatsuba. + Proof. + apply mk_znz_spec;auto. + exact spec_ww_add_mul_div. + refine (@spec_ww_pos_mod w w_0 w_digits w_zdigits w_WW + w_pos_mod compare w_0W low sub _ww_zdigits w_to_Z + _ _ _ _ _ _ _ _ _ _ _ _);wwauto. + exact (spec_pos_mod op_spec). + exact (spec_zdigits op_spec). + unfold w_to_Z, w_zdigits. + rewrite (spec_zdigits op_spec). + rewrite <- Zpos_xO; exact spec_ww_digits. + Qed. + +End Z_2nZ. + +Section MulAdd. + + Variable w: Type. + Variable op: znz_op w. + Variable sop: znz_spec op. + + Definition mul_add:= w_mul_add (znz_0 op) (znz_succ op) (znz_add_c op) (znz_mul_c op). + + Notation "[| x |]" := (znz_to_Z op x) (at level 0, x at level 99). + + Notation "[|| x ||]" := + (zn2z_to_Z (base (znz_digits op)) (znz_to_Z op) x) (at level 0, x at level 99). + + + Lemma spec_mul_add: forall x y z, + let (zh, zl) := mul_add x y z in + [||WW zh zl||] = [|x|] * [|y|] + [|z|]. + Proof. + intros x y z. + refine (spec_w_mul_add _ _ _ _ _ _ _ _ _ _ _ _ x y z); auto. + exact (spec_0 sop). + exact (spec_to_Z sop). + exact (spec_succ sop). + exact (spec_add_c sop). + exact (spec_mul_c sop). + Qed. + +End MulAdd. + + +(** Modular versions of DoubleCyclic *) + +Module DoubleCyclic (C:CyclicType) <: CyclicType. + Definition w := zn2z C.w. + Definition w_op := mk_zn2z_op C.w_op. + Definition w_spec := mk_znz2_spec C.w_spec. +End DoubleCyclic. + +Module DoubleCyclicKaratsuba (C:CyclicType) <: CyclicType. + Definition w := zn2z C.w. + Definition w_op := mk_zn2z_op_karatsuba C.w_op. + Definition w_spec := mk_znz2_karatsuba_spec C.w_spec. +End DoubleCyclicKaratsuba. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v new file mode 100644 index 00000000..075aef59 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDiv.v @@ -0,0 +1,1540 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleDiv.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. +Require Import DoubleDivn1. +Require Import DoubleAdd. +Require Import DoubleSub. + +Open Local Scope Z_scope. + +Ltac zarith := auto with zarith. + + +Section POS_MOD. + + Variable w:Type. + Variable w_0 : w. + Variable w_digits : positive. + Variable w_zdigits : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_pos_mod : w -> w -> w. + Variable w_compare : w -> w -> comparison. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable w_0W : w -> zn2z w. + Variable low: zn2z w -> w. + Variable ww_sub: zn2z w -> zn2z w -> zn2z w. + Variable ww_zdigits : zn2z w. + + + Definition ww_pos_mod p x := + let zdigits := w_0W w_zdigits in + match x with + | W0 => W0 + | WW xh xl => + match ww_compare p zdigits with + | Eq => w_WW w_0 xl + | Lt => w_WW w_0 (w_pos_mod (low p) xl) + | Gt => + match ww_compare p ww_zdigits with + | Lt => + let n := low (ww_sub p zdigits) in + w_WW (w_pos_mod n xh) xl + | _ => x + end + end + end. + + + Variable w_to_Z : w -> Z. + + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + + + Variable spec_w_0 : [|w_0|] = 0. + + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + + Variable spec_pos_mod : forall w p, + [|w_pos_mod p w|] = [|w|] mod (2 ^ [|p|]). + + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_sub: forall x y, + [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. + + Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits. + Variable spec_low: forall x, [| low x|] = [[x]] mod wB. + Variable spec_ww_zdigits : [[ww_zdigits]] = 2 * [|w_zdigits|]. + Variable spec_ww_digits : ww_digits w_digits = xO w_digits. + + + Hint Rewrite spec_w_0 spec_w_WW : w_rewrite. + + Lemma spec_ww_pos_mod : forall w p, + [[ww_pos_mod p w]] = [[w]] mod (2 ^ [[p]]). + assert (HHHHH:= lt_0_wB w_digits). + assert (F0: forall x y, x - y + y = x); auto with zarith. + intros w1 p; case (spec_to_w_Z p); intros HH1 HH2. + unfold ww_pos_mod; case w1. + simpl; rewrite Zmod_small; split; auto with zarith. + intros xh xl; generalize (spec_ww_compare p (w_0W w_zdigits)); + case ww_compare; + rewrite spec_w_0W; rewrite spec_zdigits; fold wB; + intros H1. + rewrite H1; simpl ww_to_Z. + autorewrite with w_rewrite rm10. + rewrite Zplus_mod; auto with zarith. + rewrite Z_mod_mult; auto with zarith. + autorewrite with rm10. + rewrite Zmod_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + autorewrite with w_rewrite rm10. + simpl ww_to_Z. + rewrite spec_pos_mod. + assert (HH0: [|low p|] = [[p]]). + rewrite spec_low. + apply Zmod_small; auto with zarith. + case (spec_to_w_Z p); intros HHH1 HHH2; split; auto with zarith. + apply Zlt_le_trans with (1 := H1). + unfold base; apply Zpower2_le_lin; auto with zarith. + rewrite HH0. + rewrite Zplus_mod; auto with zarith. + unfold base. + rewrite <- (F0 (Zpos w_digits) [[p]]). + rewrite Zpower_exp; auto with zarith. + rewrite Zmult_assoc. + rewrite Z_mod_mult; auto with zarith. + autorewrite with w_rewrite rm10. + rewrite Zmod_mod; auto with zarith. +generalize (spec_ww_compare p ww_zdigits); + case ww_compare; rewrite spec_ww_zdigits; + rewrite spec_zdigits; intros H2. + replace (2^[[p]]) with wwB. + rewrite Zmod_small; auto with zarith. + unfold base; rewrite H2. + rewrite spec_ww_digits; auto. + assert (HH0: [|low (ww_sub p (w_0W w_zdigits))|] = + [[p]] - Zpos w_digits). + rewrite spec_low. + rewrite spec_ww_sub. + rewrite spec_w_0W; rewrite spec_zdigits. + rewrite <- Zmod_div_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. + exists wB; unfold base; rewrite <- Zpower_exp; auto with zarith. + rewrite spec_ww_digits; + apply f_equal with (f := Zpower 2); rewrite Zpos_xO; auto with zarith. + simpl ww_to_Z; autorewrite with w_rewrite. + rewrite spec_pos_mod; rewrite HH0. + pattern [|xh|] at 2; + rewrite Z_div_mod_eq with (b := 2 ^ ([[p]] - Zpos w_digits)); + auto with zarith. + rewrite (fun x => (Zmult_comm (2 ^ x))); rewrite Zmult_plus_distr_l. + unfold base; rewrite <- Zmult_assoc; rewrite <- Zpower_exp; + auto with zarith. + rewrite F0; auto with zarith. + rewrite <- Zplus_assoc; rewrite Zplus_mod; auto with zarith. + rewrite Z_mod_mult; auto with zarith. + autorewrite with rm10. + rewrite Zmod_mod; auto with zarith. + apply sym_equal; apply Zmod_small; auto with zarith. + case (spec_to_Z xh); intros U1 U2. + case (spec_to_Z xl); intros U3 U4. + split; auto with zarith. + apply Zplus_le_0_compat; auto with zarith. + apply Zmult_le_0_compat; auto with zarith. + match goal with |- 0 <= ?X mod ?Y => + case (Z_mod_lt X Y); auto with zarith + end. + match goal with |- ?X mod ?Y * ?U + ?Z < ?T => + apply Zle_lt_trans with ((Y - 1) * U + Z ); + [case (Z_mod_lt X Y); auto with zarith | idtac] + end. + match goal with |- ?X * ?U + ?Y < ?Z => + apply Zle_lt_trans with (X * U + (U - 1)) + end. + apply Zplus_le_compat_l; auto with zarith. + case (spec_to_Z xl); unfold base; auto with zarith. + rewrite Zmult_minus_distr_r; rewrite <- Zpower_exp; auto with zarith. + rewrite F0; auto with zarith. + rewrite Zmod_small; auto with zarith. + case (spec_to_w_Z (WW xh xl)); intros U1 U2. + split; auto with zarith. + apply Zlt_le_trans with (1:= U2). + unfold base; rewrite spec_ww_digits. + apply Zpower_le_monotone; auto with zarith. + split; auto with zarith. + rewrite Zpos_xO; auto with zarith. + Qed. + +End POS_MOD. + +Section DoubleDiv32. + + Variable w : Type. + Variable w_0 : w. + Variable w_Bm1 : w. + Variable w_Bm2 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable w_add_c : w -> w -> carry w. + Variable w_add_carry_c : w -> w -> carry w. + Variable w_add : w -> w -> w. + Variable w_add_carry : w -> w -> w. + Variable w_pred : w -> w. + Variable w_sub : w -> w -> w. + Variable w_mul_c : w -> w -> zn2z w. + Variable w_div21 : w -> w -> w -> w*w. + Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w). + + Definition w_div32 a1 a2 a3 b1 b2 := + Eval lazy beta iota delta [ww_add_c_cont ww_add] in + match w_compare a1 b1 with + | Lt => + let (q,r) := w_div21 a1 a2 b1 in + match ww_sub_c (w_WW r a3) (w_mul_c q b2) with + | C0 r1 => (q,r1) + | C1 r1 => + let q := w_pred q in + ww_add_c_cont w_WW w_add_c w_add_carry_c + (fun r2=>(w_pred q, ww_add w_add_c w_add w_add_carry r2 (WW b1 b2))) + (fun r2 => (q,r2)) + r1 (WW b1 b2) + end + | Eq => + ww_add_c_cont w_WW w_add_c w_add_carry_c + (fun r => (w_Bm2, ww_add w_add_c w_add w_add_carry r (WW b1 b2))) + (fun r => (w_Bm1,r)) + (WW (w_sub a2 b2) a3) (WW b1 b2) + | Gt => (w_0, W0) (* cas absurde *) + end. + + (* Proof *) + + Variable w_digits : positive. + Variable w_to_Z : w -> Z. + + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_w_Bm2 : [|w_Bm2|] = wB - 2. + + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. + Variable spec_w_add_carry_c : + forall x y, [+|w_add_carry_c x y|] = [|x|] + [|y|] + 1. + + Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB. + Variable spec_w_add_carry : + forall x y, [|w_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB. + + Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB. + Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + + Variable spec_mul_c : forall x y, [[ w_mul_c x y ]] = [|x|] * [|y|]. + Variable spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + + Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Theorem wB_div2: forall x, wB/2 <= x -> wB <= 2 * x. + intros x H; rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith. + Qed. + + Lemma Zmult_lt_0_reg_r_2 : forall n m : Z, 0 <= n -> 0 < m * n -> 0 < m. + Proof. + intros n m H1 H2;apply Zmult_lt_0_reg_r with n;trivial. + destruct (Zle_lt_or_eq _ _ H1);trivial. + subst;rewrite Zmult_0_r in H2;discriminate H2. + Qed. + + Theorem spec_w_div32 : forall a1 a2 a3 b1 b2, + wB/2 <= [|b1|] -> + [[WW a1 a2]] < [[WW b1 b2]] -> + let (q,r) := w_div32 a1 a2 a3 b1 b2 in + [|a1|] * wwB + [|a2|] * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]. + Proof. + intros a1 a2 a3 b1 b2 Hle Hlt. + assert (U:= lt_0_wB w_digits); assert (U1:= lt_0_wwB w_digits). + Spec_w_to_Z a1;Spec_w_to_Z a2;Spec_w_to_Z a3;Spec_w_to_Z b1;Spec_w_to_Z b2. + rewrite wwB_wBwB; rewrite Zpower_2; rewrite Zmult_assoc;rewrite <- Zmult_plus_distr_l. + change (w_div32 a1 a2 a3 b1 b2) with + match w_compare a1 b1 with + | Lt => + let (q,r) := w_div21 a1 a2 b1 in + match ww_sub_c (w_WW r a3) (w_mul_c q b2) with + | C0 r1 => (q,r1) + | C1 r1 => + let q := w_pred q in + ww_add_c_cont w_WW w_add_c w_add_carry_c + (fun r2=>(w_pred q, ww_add w_add_c w_add w_add_carry r2 (WW b1 b2))) + (fun r2 => (q,r2)) + r1 (WW b1 b2) + end + | Eq => + ww_add_c_cont w_WW w_add_c w_add_carry_c + (fun r => (w_Bm2, ww_add w_add_c w_add w_add_carry r (WW b1 b2))) + (fun r => (w_Bm1,r)) + (WW (w_sub a2 b2) a3) (WW b1 b2) + | Gt => (w_0, W0) (* cas absurde *) + end. + assert (Hcmp:=spec_compare a1 b1);destruct (w_compare a1 b1). + simpl in Hlt. + rewrite Hcmp in Hlt;assert ([|a2|] < [|b2|]). omega. + assert ([[WW (w_sub a2 b2) a3]] = ([|a2|]-[|b2|])*wB + [|a3|] + wwB). + simpl;rewrite spec_sub. + assert ([|a2|] - [|b2|] = wB*(-1) + ([|a2|] - [|b2|] + wB)). ring. + assert (0 <= [|a2|] - [|b2|] + wB < wB). omega. + rewrite <-(Zmod_unique ([|a2|]-[|b2|]) wB (-1) ([|a2|]-[|b2|]+wB) H1 H0). + rewrite wwB_wBwB;ring. + assert (U2 := wB_pos w_digits). + eapply spec_ww_add_c_cont with (P := + fun (x y:zn2z w) (res:w*zn2z w) => + let (q, r) := res in + ([|a1|] * wB + [|a2|]) * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]);eauto. + rewrite H0;intros r. + repeat + (rewrite spec_ww_add;eauto || rewrite spec_w_Bm1 || rewrite spec_w_Bm2); + simpl ww_to_Z;try rewrite Zmult_1_l;intros H1. + assert (0<= ([[r]] + ([|b1|] * wB + [|b2|])) - wwB < [|b1|] * wB + [|b2|]). + Spec_ww_to_Z r;split;zarith. + rewrite H1. + assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB). + rewrite wwB_wBwB; rewrite Zpower_2; zarith. + assert (-wwB < ([|a2|] - [|b2|]) * wB + [|a3|] < 0). + split. apply Zlt_le_trans with (([|a2|] - [|b2|]) * wB);zarith. + rewrite wwB_wBwB;replace (-(wB^2)) with (-wB*wB);[zarith | ring]. + apply Zmult_lt_compat_r;zarith. + apply Zle_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith. + replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with + (([|a2|] - [|b2|] + 1) * wB + - 1);[zarith | ring]. + assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith. + replace 0 with (0*wB);zarith. + replace (([|a2|] - [|b2|]) * wB + [|a3|] + wwB + ([|b1|] * wB + [|b2|]) + + ([|b1|] * wB + [|b2|]) - wwB) with + (([|a2|] - [|b2|]) * wB + [|a3|] + 2*[|b1|] * wB + 2*[|b2|]); + [zarith | ring]. + rewrite <- (Zmod_unique ([[r]] + ([|b1|] * wB + [|b2|])) wwB + 1 ([[r]] + ([|b1|] * wB + [|b2|]) - wwB));zarith;try (ring;fail). + split. rewrite H1;rewrite Hcmp;ring. trivial. + Spec_ww_to_Z (WW b1 b2). simpl in HH4;zarith. + rewrite H0;intros r;repeat + (rewrite spec_w_Bm1 || rewrite spec_w_Bm2); + simpl ww_to_Z;try rewrite Zmult_1_l;intros H1. + assert ([[r]]=([|a2|]-[|b2|])*wB+[|a3|]+([|b1|]*wB+[|b2|])). zarith. + split. rewrite H2;rewrite Hcmp;ring. + split. Spec_ww_to_Z r;zarith. + rewrite H2. + assert (([|a2|] - [|b2|]) * wB + [|a3|] < 0);zarith. + apply Zle_lt_trans with (([|a2|] - [|b2|]) * wB + (wB -1));zarith. + replace ( ([|a2|] - [|b2|]) * wB + (wB - 1)) with + (([|a2|] - [|b2|] + 1) * wB + - 1);[zarith|ring]. + assert (([|a2|] - [|b2|] + 1) * wB <= 0);zarith. + replace 0 with (0*wB);zarith. + (* Cas Lt *) + assert (Hdiv21 := spec_div21 a2 Hle Hcmp); + destruct (w_div21 a1 a2 b1) as (q, r);destruct Hdiv21. + rewrite H. + assert (Hq := spec_to_Z q). + generalize + (spec_ww_sub_c (w_WW r a3) (w_mul_c q b2)); + destruct (ww_sub_c (w_WW r a3) (w_mul_c q b2)) + as [r1|r1];repeat (rewrite spec_w_WW || rewrite spec_mul_c); + unfold interp_carry;intros H1. + rewrite H1. + split. ring. split. + rewrite <- H1;destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z r1);trivial. + apply Zle_lt_trans with ([|r|] * wB + [|a3|]). + assert ( 0 <= [|q|] * [|b2|]);zarith. + apply beta_lex_inv;zarith. + assert ([[r1]] = [|r|] * wB + [|a3|] - [|q|] * [|b2|] + wwB). + rewrite <- H1;ring. + Spec_ww_to_Z r1; assert (0 <= [|r|]*wB). zarith. + assert (0 < [|q|] * [|b2|]). zarith. + assert (0 < [|q|]). + apply Zmult_lt_0_reg_r_2 with [|b2|];zarith. + eapply spec_ww_add_c_cont with (P := + fun (x y:zn2z w) (res:w*zn2z w) => + let (q0, r0) := res in + ([|q|] * [|b1|] + [|r|]) * wB + [|a3|] = + [|q0|] * ([|b1|] * wB + [|b2|]) + [[r0]] /\ + 0 <= [[r0]] < [|b1|] * wB + [|b2|]);eauto. + intros r2;repeat (rewrite spec_pred || rewrite spec_ww_add;eauto); + simpl ww_to_Z;intros H7. + assert (0 < [|q|] - 1). + assert (1 <= [|q|]). zarith. + destruct (Zle_lt_or_eq _ _ H6);zarith. + rewrite <- H8 in H2;rewrite H2 in H7. + assert (0 < [|b1|]*wB). apply Zmult_lt_0_compat;zarith. + Spec_ww_to_Z r2. zarith. + rewrite (Zmod_small ([|q|] -1));zarith. + rewrite (Zmod_small ([|q|] -1 -1));zarith. + assert ([[r2]] + ([|b1|] * wB + [|b2|]) = + wwB * 1 + + ([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2 * ([|b1|] * wB + [|b2|]))). + rewrite H7;rewrite H2;ring. + assert + ([|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|]) + < [|b1|]*wB + [|b2|]). + Spec_ww_to_Z r2;omega. + Spec_ww_to_Z (WW b1 b2). simpl in HH5. + assert + (0 <= [|r|]*wB + [|a3|] - [|q|]*[|b2|] + 2 * ([|b1|]*wB + [|b2|]) + < wwB). split;try omega. + replace (2*([|b1|]*wB+[|b2|])) with ((2*[|b1|])*wB+2*[|b2|]). 2:ring. + assert (H12:= wB_div2 Hle). assert (wwB <= 2 * [|b1|] * wB). + rewrite wwB_wBwB; rewrite Zpower_2; zarith. omega. + rewrite <- (Zmod_unique + ([[r2]] + ([|b1|] * wB + [|b2|])) + wwB + 1 + ([|r|] * wB + [|a3|] - [|q|] * [|b2|] + 2*([|b1|] * wB + [|b2|])) + H10 H8). + split. ring. zarith. + intros r2;repeat (rewrite spec_pred);simpl ww_to_Z;intros H7. + rewrite (Zmod_small ([|q|] -1));zarith. + split. + replace [[r2]] with ([[r1]] + ([|b1|] * wB + [|b2|]) -wwB). + rewrite H2; ring. rewrite <- H7; ring. + Spec_ww_to_Z r2;Spec_ww_to_Z r1. omega. + simpl in Hlt. + assert ([|a1|] * wB + [|a2|] <= [|b1|] * wB + [|b2|]). zarith. + assert (H1 := beta_lex _ _ _ _ _ H HH0 HH3). rewrite spec_w_0;simpl;zarith. + Qed. + + +End DoubleDiv32. + +Section DoubleDiv21. + Variable w : Type. + Variable w_0 : w. + + Variable w_0W : w -> zn2z w. + Variable w_div32 : w -> w -> w -> w -> w -> w * zn2z w. + + Variable ww_1 : zn2z w. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable ww_sub : zn2z w -> zn2z w -> zn2z w. + + + Definition ww_div21 a1 a2 b := + match a1 with + | W0 => + match ww_compare a2 b with + | Gt => (ww_1, ww_sub a2 b) + | Eq => (ww_1, W0) + | Lt => (W0, a2) + end + | WW a1h a1l => + match a2 with + | W0 => + match b with + | W0 => (W0,W0) (* cas absurde *) + | WW b1 b2 => + let (q1, r) := w_div32 a1h a1l w_0 b1 b2 in + match r with + | W0 => (WW q1 w_0, W0) + | WW r1 r2 => + let (q2, s) := w_div32 r1 r2 w_0 b1 b2 in + (WW q1 q2, s) + end + end + | WW a2h a2l => + match b with + | W0 => (W0,W0) (* cas absurde *) + | WW b1 b2 => + let (q1, r) := w_div32 a1h a1l a2h b1 b2 in + match r with + | W0 => (WW q1 w_0, w_0W a2l) + | WW r1 r2 => + let (q2, s) := w_div32 r1 r2 a2l b1 b2 in + (WW q1 q2, s) + end + end + end + end. + + (* Proof *) + + Variable w_digits : positive. + Variable w_to_Z : w -> Z. + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_w_div32 : forall a1 a2 a3 b1 b2, + wB/2 <= [|b1|] -> + [[WW a1 a2]] < [[WW b1 b2]] -> + let (q,r) := w_div32 a1 a2 a3 b1 b2 in + [|a1|] * wwB + [|a2|] * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]. + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. + + Theorem wwB_div: wwB = 2 * (wwB / 2). + Proof. + rewrite wwB_div_2; rewrite Zmult_assoc; rewrite wB_div_2; auto. + rewrite <- Zpower_2; apply wwB_wBwB. + Qed. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Theorem spec_ww_div21 : forall a1 a2 b, + wwB/2 <= [[b]] -> + [[a1]] < [[b]] -> + let (q,r) := ww_div21 a1 a2 b in + [[a1]] *wwB+[[a2]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]. + Proof. + assert (U:= lt_0_wB w_digits). + assert (U1:= lt_0_wwB w_digits). + intros a1 a2 b H Hlt; unfold ww_div21. + Spec_ww_to_Z b; assert (Eq: 0 < [[b]]). Spec_ww_to_Z a1;omega. + generalize Hlt H ;clear Hlt H;case a1. + intros H1 H2;simpl in H1;Spec_ww_to_Z a2; + match goal with |-context [ww_compare ?Y ?Z] => + generalize (spec_ww_compare Y Z); case (ww_compare Y Z) + end; simpl;try rewrite spec_ww_1;autorewrite with rm10; intros;zarith. + rewrite spec_ww_sub;simpl. rewrite Zmod_small;zarith. + split. ring. + assert (wwB <= 2*[[b]]);zarith. + rewrite wwB_div;zarith. + intros a1h a1l. Spec_w_to_Z a1h;Spec_w_to_Z a1l. Spec_ww_to_Z a2. + destruct a2 as [ |a3 a4]; + (destruct b as [ |b1 b2];[unfold Zle in Eq;discriminate Eq|idtac]); + try (Spec_w_to_Z a3; Spec_w_to_Z a4); Spec_w_to_Z b1; Spec_w_to_Z b2; + intros Hlt H; match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] => + generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U); + intros q1 r H0 + end; (assert (Eq1: wB / 2 <= [|b1|]);[ + apply (@beta_lex (wB / 2) 0 [|b1|] [|b2|] wB); auto with zarith; + autorewrite with rm10;repeat rewrite (Zmult_comm wB); + rewrite <- wwB_div_2; trivial + | generalize (H0 Eq1 Hlt);clear H0;destruct r as [ |r1 r2];simpl; + try rewrite spec_w_0; try rewrite spec_w_0W;repeat rewrite Zplus_0_r; + intros (H1,H2) ]). + split;[rewrite wwB_wBwB; rewrite Zpower_2 | trivial]. + rewrite Zmult_assoc;rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc; + rewrite <- Zpower_2; rewrite <- wwB_wBwB;rewrite H1;ring. + destruct H2 as (H2,H3);match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] => + generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U); + intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end. + split;[rewrite wwB_wBwB | trivial]. + rewrite Zpower_2. + rewrite Zmult_assoc;rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc; + rewrite <- Zpower_2. + rewrite <- wwB_wBwB;rewrite H1. + rewrite spec_w_0 in H4;rewrite Zplus_0_r in H4. + repeat rewrite Zmult_plus_distr_l. rewrite <- (Zmult_assoc [|r1|]). + rewrite <- Zpower_2; rewrite <- wwB_wBwB;rewrite H4;simpl;ring. + split;[rewrite wwB_wBwB | split;zarith]. + replace (([|a1h|] * wB + [|a1l|]) * wB^2 + ([|a3|] * wB + [|a4|])) + with (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB+ [|a4|]). + rewrite H1;ring. rewrite wwB_wBwB;ring. + change [|a4|] with (0*wB+[|a4|]);apply beta_lex_inv;zarith. + assert (1 <= wB/2);zarith. + assert (H_:= wB_pos w_digits);apply Zdiv_le_lower_bound;zarith. + destruct H2 as (H2,H3);match goal with |-context [w_div32 ?X ?Y ?Z ?T ?U] => + generalize (@spec_w_div32 X Y Z T U); case (w_div32 X Y Z T U); + intros q r H0;generalize (H0 Eq1 H3);clear H0;intros (H4,H5) end. + split;trivial. + replace (([|a1h|] * wB + [|a1l|]) * wwB + ([|a3|] * wB + [|a4|])) with + (([|a1h|] * wwB + [|a1l|] * wB + [|a3|])*wB + [|a4|]); + [rewrite H1 | rewrite wwB_wBwB;ring]. + replace (([|q1|]*([|b1|]*wB+[|b2|])+([|r1|]*wB+[|r2|]))*wB+[|a4|]) with + (([|q1|]*([|b1|]*wB+[|b2|]))*wB+([|r1|]*wwB+[|r2|]*wB+[|a4|])); + [rewrite H4;simpl|rewrite wwB_wBwB];ring. + Qed. + +End DoubleDiv21. + +Section DoubleDivGt. + Variable w : Type. + Variable w_digits : positive. + Variable w_0 : w. + + Variable w_WW : w -> w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable w_eq0 : w -> bool. + Variable w_opp_c : w -> carry w. + Variable w_opp w_opp_carry : w -> w. + Variable w_sub_c : w -> w -> carry w. + Variable w_sub w_sub_carry : w -> w -> w. + + Variable w_div_gt : w -> w -> w*w. + Variable w_mod_gt : w -> w -> w. + Variable w_gcd_gt : w -> w -> w. + Variable w_add_mul_div : w -> w -> w -> w. + Variable w_head0 : w -> w. + Variable w_div21 : w -> w -> w -> w * w. + Variable w_div32 : w -> w -> w -> w -> w -> w * zn2z w. + + + Variable _ww_zdigits : zn2z w. + Variable ww_1 : zn2z w. + Variable ww_add_mul_div : zn2z w -> zn2z w -> zn2z w -> zn2z w. + + Variable w_zdigits : w. + + Definition ww_div_gt_aux ah al bh bl := + Eval lazy beta iota delta [ww_sub ww_opp] in + let p := w_head0 bh in + match w_compare p w_0 with + | Gt => + let b1 := w_add_mul_div p bh bl in + let b2 := w_add_mul_div p bl w_0 in + let a1 := w_add_mul_div p w_0 ah in + let a2 := w_add_mul_div p ah al in + let a3 := w_add_mul_div p al w_0 in + let (q,r) := w_div32 a1 a2 a3 b1 b2 in + (WW w_0 q, ww_add_mul_div + (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c + w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r) + | _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c + w_opp w_sub w_sub_carry (WW ah al) (WW bh bl)) + end. + + Definition ww_div_gt a b := + Eval lazy beta iota delta [ww_div_gt_aux double_divn1 + double_divn1_p double_divn1_p_aux double_divn1_0 double_divn1_0_aux + double_split double_0 double_WW] in + match a, b with + | W0, _ => (W0,W0) + | _, W0 => (W0,W0) + | WW ah al, WW bh bl => + if w_eq0 ah then + let (q,r) := w_div_gt al bl in + (WW w_0 q, w_0W r) + else + match w_compare w_0 bh with + | Eq => + let(q,r):= + double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 a bl in + (q, w_0W r) + | Lt => ww_div_gt_aux ah al bh bl + | Gt => (W0,W0) (* cas absurde *) + end + end. + + Definition ww_mod_gt_aux ah al bh bl := + Eval lazy beta iota delta [ww_sub ww_opp] in + let p := w_head0 bh in + match w_compare p w_0 with + | Gt => + let b1 := w_add_mul_div p bh bl in + let b2 := w_add_mul_div p bl w_0 in + let a1 := w_add_mul_div p w_0 ah in + let a2 := w_add_mul_div p ah al in + let a3 := w_add_mul_div p al w_0 in + let (q,r) := w_div32 a1 a2 a3 b1 b2 in + ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c + w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r + | _ => + ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c + w_opp w_sub w_sub_carry (WW ah al) (WW bh bl) + end. + + Definition ww_mod_gt a b := + Eval lazy beta iota delta [ww_mod_gt_aux double_modn1 + double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux + double_split double_0 double_WW snd] in + match a, b with + | W0, _ => W0 + | _, W0 => W0 + | WW ah al, WW bh bl => + if w_eq0 ah then w_0W (w_mod_gt al bl) + else + match w_compare w_0 bh with + | Eq => + w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 a bl) + | Lt => ww_mod_gt_aux ah al bh bl + | Gt => W0 (* cas absurde *) + end + end. + + Definition ww_gcd_gt_body (cont: w->w->w->w->zn2z w) (ah al bh bl: w) := + Eval lazy beta iota delta [ww_mod_gt_aux double_modn1 + double_modn1_p double_modn1_p_aux double_modn1_0 double_modn1_0_aux + double_split double_0 double_WW snd] in + match w_compare w_0 bh with + | Eq => + match w_compare w_0 bl with + | Eq => WW ah al (* normalement n'arrive pas si forme normale *) + | Lt => + let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 (WW ah al) bl in + WW w_0 (w_gcd_gt bl m) + | Gt => W0 (* absurde *) + end + | Lt => + let m := ww_mod_gt_aux ah al bh bl in + match m with + | W0 => WW bh bl + | WW mh ml => + match w_compare w_0 mh with + | Eq => + match w_compare w_0 ml with + | Eq => WW bh bl + | _ => + let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 (WW bh bl) ml in + WW w_0 (w_gcd_gt ml r) + end + | Lt => + let r := ww_mod_gt_aux bh bl mh ml in + match r with + | W0 => m + | WW rh rl => cont mh ml rh rl + end + | Gt => W0 (* absurde *) + end + end + | Gt => W0 (* absurde *) + end. + + Fixpoint ww_gcd_gt_aux + (p:positive) (cont: w -> w -> w -> w -> zn2z w) (ah al bh bl : w) + {struct p} : zn2z w := + ww_gcd_gt_body + (fun mh ml rh rl => match p with + | xH => cont mh ml rh rl + | xO p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl + | xI p => ww_gcd_gt_aux p (ww_gcd_gt_aux p cont) mh ml rh rl + end) ah al bh bl. + + + (* Proof *) + + Variable w_to_Z : w -> Z. + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[-| c |]" := + (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0. + + Variable spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|]. + Variable spec_opp : forall x, [|w_opp x|] = (-[|x|]) mod wB. + Variable spec_opp_carry : forall x, [|w_opp_carry x|] = wB - [|x|] - 1. + + Variable spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|]. + Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + Variable spec_sub_carry : + forall x y, [|w_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB. + + Variable spec_div_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> + let (q,r) := w_div_gt a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Variable spec_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> + [|w_mod_gt a b|] = [|a|] mod [|b|]. + Variable spec_gcd_gt : forall a b, [|a|] > [|b|] -> + Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|]. + + Variable spec_add_mul_div : forall x y p, + [|p|] <= Zpos w_digits -> + [| w_add_mul_div p x y |] = + ([|x|] * (2 ^ ([|p|])) + + [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB. + Variable spec_head0 : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ [|w_head0 x|] * [|x|] < wB. + + Variable spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + + Variable spec_w_div32 : forall a1 a2 a3 b1 b2, + wB/2 <= [|b1|] -> + [[WW a1 a2]] < [[WW b1 b2]] -> + let (q,r) := w_div32 a1 a2 a3 b1 b2 in + [|a1|] * wwB + [|a2|] * wB + [|a3|] = + [|q|] * ([|b1|] * wB + [|b2|]) + [[r]] /\ + 0 <= [[r]] < [|b1|] * wB + [|b2|]. + + Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits. + + Variable spec_ww_digits_ : [[_ww_zdigits]] = Zpos (xO w_digits). + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_ww_add_mul_div : forall x y p, + [[p]] <= Zpos (xO w_digits) -> + [[ ww_add_mul_div p x y ]] = + ([[x]] * (2^[[p]]) + + [[y]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Lemma to_Z_div_minus_p : forall x p, + 0 < [|p|] < Zpos w_digits -> + 0 <= [|x|] / 2 ^ (Zpos w_digits - [|p|]) < 2 ^ [|p|]. + Proof. + intros x p H;Spec_w_to_Z x. + split. apply Zdiv_le_lower_bound;zarith. + apply Zdiv_lt_upper_bound;zarith. + rewrite <- Zpower_exp;zarith. + ring_simplify ([|p|] + (Zpos w_digits - [|p|])); unfold base in HH;zarith. + Qed. + Hint Resolve to_Z_div_minus_p : zarith. + + Lemma spec_ww_div_gt_aux : forall ah al bh bl, + [[WW ah al]] > [[WW bh bl]] -> + 0 < [|bh|] -> + let (q,r) := ww_div_gt_aux ah al bh bl in + [[WW ah al]] = [[q]] * [[WW bh bl]] + [[r]] /\ + 0 <= [[r]] < [[WW bh bl]]. + Proof. + intros ah al bh bl Hgt Hpos;unfold ww_div_gt_aux. + change + (let (q, r) := let p := w_head0 bh in + match w_compare p w_0 with + | Gt => + let b1 := w_add_mul_div p bh bl in + let b2 := w_add_mul_div p bl w_0 in + let a1 := w_add_mul_div p w_0 ah in + let a2 := w_add_mul_div p ah al in + let a3 := w_add_mul_div p al w_0 in + let (q,r) := w_div32 a1 a2 a3 b1 b2 in + (WW w_0 q, ww_add_mul_div + (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c + w_opp w_sub w_sub_carry _ww_zdigits (w_0W p)) W0 r) + | _ => (ww_1, ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c + w_opp w_sub w_sub_carry (WW ah al) (WW bh bl)) + end in [[WW ah al]]=[[q]]*[[WW bh bl]]+[[r]] /\ 0 <=[[r]]< [[WW bh bl]]). + assert (Hh := spec_head0 Hpos). + lazy zeta. + generalize (spec_compare (w_head0 bh) w_0); case w_compare; + rewrite spec_w_0; intros HH. + generalize Hh; rewrite HH; simpl Zpower; + rewrite Zmult_1_l; intros (HH1, HH2); clear HH. + assert (wwB <= 2*[[WW bh bl]]). + apply Zle_trans with (2*[|bh|]*wB). + rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat_r; zarith. + rewrite <- wB_div_2; apply Zmult_le_compat_l; zarith. + simpl ww_to_Z;rewrite Zmult_plus_distr_r;rewrite Zmult_assoc. + Spec_w_to_Z bl;zarith. + Spec_ww_to_Z (WW ah al). + rewrite spec_ww_sub;eauto. + simpl;rewrite spec_ww_1;rewrite Zmult_1_l;simpl. + simpl ww_to_Z in Hgt, H, HH;rewrite Zmod_small;split;zarith. + case (spec_to_Z (w_head0 bh)); auto with zarith. + assert ([|w_head0 bh|] < Zpos w_digits). + destruct (Z_lt_ge_dec [|w_head0 bh|] (Zpos w_digits));trivial. + elimtype False. + assert (2 ^ [|w_head0 bh|] * [|bh|] >= wB);auto with zarith. + apply Zle_ge; replace wB with (wB * 1);try ring. + Spec_w_to_Z bh;apply Zmult_le_compat;zarith. + unfold base;apply Zpower_le_monotone;zarith. + assert (HHHH : 0 < [|w_head0 bh|] < Zpos w_digits); auto with zarith. + assert (Hb:= Zlt_le_weak _ _ H). + generalize (spec_add_mul_div w_0 ah Hb) + (spec_add_mul_div ah al Hb) + (spec_add_mul_div al w_0 Hb) + (spec_add_mul_div bh bl Hb) + (spec_add_mul_div bl w_0 Hb); + rewrite spec_w_0; repeat rewrite Zmult_0_l;repeat rewrite Zplus_0_l; + rewrite Zdiv_0_l;repeat rewrite Zplus_0_r. + Spec_w_to_Z ah;Spec_w_to_Z bh. + unfold base;repeat rewrite Zmod_shift_r;zarith. + assert (H3:=to_Z_div_minus_p ah HHHH);assert(H4:=to_Z_div_minus_p al HHHH); + assert (H5:=to_Z_div_minus_p bl HHHH). + rewrite Zmult_comm in Hh. + assert (2^[|w_head0 bh|] < wB). unfold base;apply Zpower_lt_monotone;zarith. + unfold base in H0;rewrite Zmod_small;zarith. + fold wB; rewrite (Zmod_small ([|bh|] * 2 ^ [|w_head0 bh|]));zarith. + intros U1 U2 U3 V1 V2. + generalize (@spec_w_div32 (w_add_mul_div (w_head0 bh) w_0 ah) + (w_add_mul_div (w_head0 bh) ah al) + (w_add_mul_div (w_head0 bh) al w_0) + (w_add_mul_div (w_head0 bh) bh bl) + (w_add_mul_div (w_head0 bh) bl w_0)). + destruct (w_div32 (w_add_mul_div (w_head0 bh) w_0 ah) + (w_add_mul_div (w_head0 bh) ah al) + (w_add_mul_div (w_head0 bh) al w_0) + (w_add_mul_div (w_head0 bh) bh bl) + (w_add_mul_div (w_head0 bh) bl w_0)) as (q,r). + rewrite V1;rewrite V2. rewrite Zmult_plus_distr_l. + rewrite <- (Zplus_assoc ([|bh|] * 2 ^ [|w_head0 bh|] * wB)). + unfold base;rewrite <- shift_unshift_mod;zarith. fold wB. + replace ([|bh|] * 2 ^ [|w_head0 bh|] * wB + [|bl|] * 2 ^ [|w_head0 bh|]) with + ([[WW bh bl]] * 2^[|w_head0 bh|]). 2:simpl;ring. + fold wwB. rewrite wwB_wBwB. rewrite Zpower_2. rewrite U1;rewrite U2;rewrite U3. + rewrite Zmult_assoc. rewrite Zmult_plus_distr_l. + rewrite (Zplus_assoc ([|ah|] / 2^(Zpos(w_digits) - [|w_head0 bh|])*wB * wB)). + rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc. + unfold base;repeat rewrite <- shift_unshift_mod;zarith. fold wB. + replace ([|ah|] * 2 ^ [|w_head0 bh|] * wB + [|al|] * 2 ^ [|w_head0 bh|]) with + ([[WW ah al]] * 2^[|w_head0 bh|]). 2:simpl;ring. + intros Hd;destruct Hd;zarith. + simpl. apply beta_lex_inv;zarith. rewrite U1;rewrite V1. + assert ([|ah|] / 2 ^ (Zpos (w_digits) - [|w_head0 bh|]) < wB/2);zarith. + apply Zdiv_lt_upper_bound;zarith. + unfold base. + replace (2^Zpos (w_digits)) with (2^(Zpos (w_digits) - 1)*2). + rewrite Z_div_mult;zarith. rewrite <- Zpower_exp;zarith. + apply Zlt_le_trans with wB;zarith. + unfold base;apply Zpower_le_monotone;zarith. + pattern 2 at 2;replace 2 with (2^1);trivial. + rewrite <- Zpower_exp;zarith. ring_simplify (Zpos (w_digits) - 1 + 1);trivial. + change [[WW w_0 q]] with ([|w_0|]*wB+[|q|]);rewrite spec_w_0;rewrite + Zmult_0_l;rewrite Zplus_0_l. + replace [[ww_add_mul_div (ww_sub w_0 w_WW w_opp_c w_opp_carry w_sub_c w_opp w_sub w_sub_carry + _ww_zdigits (w_0W (w_head0 bh))) W0 r]] with ([[r]]/2^[|w_head0 bh|]). + assert (0 < 2^[|w_head0 bh|]). apply Zpower_gt_0;zarith. + split. + rewrite <- (Z_div_mult [[WW ah al]] (2^[|w_head0 bh|]));zarith. + rewrite H1;rewrite Zmult_assoc;apply Z_div_plus_l;trivial. + split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith. + rewrite spec_ww_add_mul_div. + rewrite spec_ww_sub; auto with zarith. + rewrite spec_ww_digits_. + change (Zpos (xO (w_digits))) with (2*Zpos (w_digits));zarith. + simpl ww_to_Z;rewrite Zmult_0_l;rewrite Zplus_0_l. + rewrite spec_w_0W. + rewrite (fun x y => Zmod_small (x-y)); auto with zarith. + ring_simplify (2 * Zpos w_digits - (2 * Zpos w_digits - [|w_head0 bh|])). + rewrite Zmod_small;zarith. + split;[apply Zdiv_le_lower_bound| apply Zdiv_lt_upper_bound];zarith. + Spec_ww_to_Z r. + apply Zlt_le_trans with wwB;zarith. + rewrite <- (Zmult_1_r wwB);apply Zmult_le_compat;zarith. + split; auto with zarith. + apply Zle_lt_trans with (2 * Zpos w_digits); auto with zarith. + unfold base, ww_digits; rewrite (Zpos_xO w_digits). + apply Zpower2_lt_lin; auto with zarith. + rewrite spec_ww_sub; auto with zarith. + rewrite spec_ww_digits_; rewrite spec_w_0W. + rewrite Zmod_small;zarith. + rewrite Zpos_xO; split; auto with zarith. + apply Zle_lt_trans with (2 * Zpos w_digits); auto with zarith. + unfold base, ww_digits; rewrite (Zpos_xO w_digits). + apply Zpower2_lt_lin; auto with zarith. + Qed. + + Lemma spec_ww_div_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + let (q,r) := ww_div_gt a b in + [[a]] = [[q]] * [[b]] + [[r]] /\ + 0 <= [[r]] < [[b]]. + Proof. + intros a b Hgt Hpos;unfold ww_div_gt. + change (let (q,r) := match a, b with + | W0, _ => (W0,W0) + | _, W0 => (W0,W0) + | WW ah al, WW bh bl => + if w_eq0 ah then + let (q,r) := w_div_gt al bl in + (WW w_0 q, w_0W r) + else + match w_compare w_0 bh with + | Eq => + let(q,r):= + double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 a bl in + (q, w_0W r) + | Lt => ww_div_gt_aux ah al bh bl + | Gt => (W0,W0) (* cas absurde *) + end + end in [[a]] = [[q]] * [[b]] + [[r]] /\ 0 <= [[r]] < [[b]]). + destruct a as [ |ah al]. simpl in Hgt;omega. + destruct b as [ |bh bl]. simpl in Hpos;omega. + Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl. + assert (H:=@spec_eq0 ah);destruct (w_eq0 ah). + simpl ww_to_Z;rewrite H;trivial. simpl in Hgt;rewrite H in Hgt;trivial. + assert ([|bh|] <= 0). + apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith. + assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;rewrite H1;simpl in Hgt. + simpl. simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos. + assert (H2:=spec_div_gt Hgt Hpos);destruct (w_div_gt al bl). + repeat rewrite spec_w_0W;simpl;rewrite spec_w_0;simpl;trivial. + clear H. + assert (Hcmp := spec_compare w_0 bh); destruct (w_compare w_0 bh). + rewrite spec_w_0 in Hcmp. change [[WW bh bl]] with ([|bh|]*wB+[|bl|]). + rewrite <- Hcmp;rewrite Zmult_0_l;rewrite Zplus_0_l. + simpl in Hpos;rewrite <- Hcmp in Hpos;simpl in Hpos. + assert (H2:= @spec_double_divn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div + w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 + spec_add_mul_div spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hpos). + unfold double_to_Z,double_wB,double_digits in H2. + destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 + (WW ah al) bl). + rewrite spec_w_0W;unfold ww_to_Z;trivial. + apply spec_ww_div_gt_aux;trivial. rewrite spec_w_0 in Hcmp;trivial. + rewrite spec_w_0 in Hcmp;elimtype False;omega. + Qed. + + Lemma spec_ww_mod_gt_aux_eq : forall ah al bh bl, + ww_mod_gt_aux ah al bh bl = snd (ww_div_gt_aux ah al bh bl). + Proof. + intros ah al bh bl. unfold ww_mod_gt_aux, ww_div_gt_aux. + case w_compare; auto. + case w_div32; auto. + Qed. + + Lemma spec_ww_mod_gt_aux : forall ah al bh bl, + [[WW ah al]] > [[WW bh bl]] -> + 0 < [|bh|] -> + [[ww_mod_gt_aux ah al bh bl]] = [[WW ah al]] mod [[WW bh bl]]. + Proof. + intros. rewrite spec_ww_mod_gt_aux_eq;trivial. + assert (H3 := spec_ww_div_gt_aux ah al bl H H0). + destruct (ww_div_gt_aux ah al bh bl) as (q,r);simpl. simpl in H,H3. + destruct H3;apply Zmod_unique with [[q]];zarith. + rewrite H1;ring. + Qed. + + Lemma spec_w_mod_gt_eq : forall a b, [|a|] > [|b|] -> 0 <[|b|] -> + [|w_mod_gt a b|] = [|snd (w_div_gt a b)|]. + Proof. + intros a b Hgt Hpos. + rewrite spec_mod_gt;trivial. + assert (H:=spec_div_gt Hgt Hpos). + destruct (w_div_gt a b) as (q,r);simpl. + rewrite Zmult_comm in H;destruct H. + symmetry;apply Zmod_unique with [|q|];trivial. + Qed. + + Lemma spec_ww_mod_gt_eq : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + [[ww_mod_gt a b]] = [[snd (ww_div_gt a b)]]. + Proof. + intros a b Hgt Hpos. + change (ww_mod_gt a b) with + (match a, b with + | W0, _ => W0 + | _, W0 => W0 + | WW ah al, WW bh bl => + if w_eq0 ah then w_0W (w_mod_gt al bl) + else + match w_compare w_0 bh with + | Eq => + w_0W (double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 a bl) + | Lt => ww_mod_gt_aux ah al bh bl + | Gt => W0 (* cas absurde *) + end end). + change (ww_div_gt a b) with + (match a, b with + | W0, _ => (W0,W0) + | _, W0 => (W0,W0) + | WW ah al, WW bh bl => + if w_eq0 ah then + let (q,r) := w_div_gt al bl in + (WW w_0 q, w_0W r) + else + match w_compare w_0 bh with + | Eq => + let(q,r):= + double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 a bl in + (q, w_0W r) + | Lt => ww_div_gt_aux ah al bh bl + | Gt => (W0,W0) (* cas absurde *) + end + end). + destruct a as [ |ah al];trivial. + destruct b as [ |bh bl];trivial. + Spec_w_to_Z ah; Spec_w_to_Z al; Spec_w_to_Z bh; Spec_w_to_Z bl. + assert (H:=@spec_eq0 ah);destruct (w_eq0 ah). + simpl in Hgt;rewrite H in Hgt;trivial. + assert ([|bh|] <= 0). + apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith. + assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt. + simpl in Hpos;rewrite H1 in Hpos;simpl in Hpos. + rewrite spec_w_0W;rewrite spec_w_mod_gt_eq;trivial. + destruct (w_div_gt al bl);simpl;rewrite spec_w_0W;trivial. + clear H. + assert (H2 := spec_compare w_0 bh);destruct (w_compare w_0 bh). + rewrite (@spec_double_modn1_aux w w_zdigits w_0 w_WW w_head0 w_add_mul_div + w_div21 w_compare w_sub w_to_Z spec_w_0 spec_compare 1 (WW ah al) bl). + destruct (double_divn1 w_zdigits w_0 w_WW w_head0 w_add_mul_div w_div21 w_compare w_sub 1 + (WW ah al) bl);simpl;trivial. + rewrite spec_ww_mod_gt_aux_eq;trivial;symmetry;trivial. + trivial. + Qed. + + Lemma spec_ww_mod_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + [[ww_mod_gt a b]] = [[a]] mod [[b]]. + Proof. + intros a b Hgt Hpos. + assert (H:= spec_ww_div_gt a b Hgt Hpos). + rewrite (spec_ww_mod_gt_eq a b Hgt Hpos). + destruct (ww_div_gt a b)as(q,r);destruct H. + apply Zmod_unique with[[q]];simpl;trivial. + rewrite Zmult_comm;trivial. + Qed. + + Lemma Zis_gcd_mod : forall a b d, + 0 < b -> Zis_gcd b (a mod b) d -> Zis_gcd a b d. + Proof. + intros a b d H H1; apply Zis_gcd_for_euclid with (a/b). + pattern a at 1;rewrite (Z_div_mod_eq a b). + ring_simplify (b * (a / b) + a mod b - a / b * b);trivial. zarith. + Qed. + + Lemma spec_ww_gcd_gt_aux_body : + forall ah al bh bl n cont, + [[WW bh bl]] <= 2^n -> + [[WW ah al]] > [[WW bh bl]] -> + (forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> [[WW yh yl]] <= 2^(n-1) -> + Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) -> + Zis_gcd [[WW ah al]] [[WW bh bl]] [[ww_gcd_gt_body cont ah al bh bl]]. + Proof. + intros ah al bh bl n cont Hlog Hgt Hcont. + change (ww_gcd_gt_body cont ah al bh bl) with (match w_compare w_0 bh with + | Eq => + match w_compare w_0 bl with + | Eq => WW ah al (* normalement n'arrive pas si forme normale *) + | Lt => + let m := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 (WW ah al) bl in + WW w_0 (w_gcd_gt bl m) + | Gt => W0 (* absurde *) + end + | Lt => + let m := ww_mod_gt_aux ah al bh bl in + match m with + | W0 => WW bh bl + | WW mh ml => + match w_compare w_0 mh with + | Eq => + match w_compare w_0 ml with + | Eq => WW bh bl + | _ => + let r := double_modn1 w_zdigits w_0 w_head0 w_add_mul_div w_div21 + w_compare w_sub 1 (WW bh bl) ml in + WW w_0 (w_gcd_gt ml r) + end + | Lt => + let r := ww_mod_gt_aux bh bl mh ml in + match r with + | W0 => m + | WW rh rl => cont mh ml rh rl + end + | Gt => W0 (* absurde *) + end + end + | Gt => W0 (* absurde *) + end). + assert (Hbh := spec_compare w_0 bh);destruct (w_compare w_0 bh). + simpl ww_to_Z in *. rewrite spec_w_0 in Hbh;rewrite <- Hbh; + rewrite Zmult_0_l;rewrite Zplus_0_l. + assert (Hbl := spec_compare w_0 bl); destruct (w_compare w_0 bl). + rewrite spec_w_0 in Hbl;rewrite <- Hbl;apply Zis_gcd_0. + simpl;rewrite spec_w_0;rewrite Zmult_0_l;rewrite Zplus_0_l. + rewrite spec_w_0 in Hbl. + apply Zis_gcd_mod;zarith. + change ([|ah|] * wB + [|al|]) with (double_to_Z w_digits w_to_Z 1 (WW ah al)). + rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div + w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 spec_add_mul_div + spec_div21 spec_compare spec_sub 1 (WW ah al) bl Hbl). + apply spec_gcd_gt. + rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial. + apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + destruct (Z_mod_lt x y);zarith end. + rewrite spec_w_0 in Hbl;Spec_w_to_Z bl;elimtype False;omega. + rewrite spec_w_0 in Hbh;assert (H:= spec_ww_mod_gt_aux _ _ _ Hgt Hbh). + assert (H2 : 0 < [[WW bh bl]]). + simpl;Spec_w_to_Z bl. apply Zlt_le_trans with ([|bh|]*wB);zarith. + apply Zmult_lt_0_compat;zarith. + apply Zis_gcd_mod;trivial. rewrite <- H. + simpl in *;destruct (ww_mod_gt_aux ah al bh bl) as [ |mh ml]. + simpl;apply Zis_gcd_0;zarith. + assert (Hmh := spec_compare w_0 mh);destruct (w_compare w_0 mh). + simpl;rewrite spec_w_0 in Hmh; rewrite <- Hmh;simpl. + assert (Hml := spec_compare w_0 ml);destruct (w_compare w_0 ml). + rewrite <- Hml;rewrite spec_w_0;simpl;apply Zis_gcd_0. + simpl;rewrite spec_w_0;simpl. + rewrite spec_w_0 in Hml. apply Zis_gcd_mod;zarith. + change ([|bh|] * wB + [|bl|]) with (double_to_Z w_digits w_to_Z 1 (WW bh bl)). + rewrite <- (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW w_head0 w_add_mul_div + w_div21 w_compare w_sub w_to_Z spec_to_Z spec_w_zdigits spec_w_0 spec_w_WW spec_head0 spec_add_mul_div + spec_div21 spec_compare spec_sub 1 (WW bh bl) ml Hml). + apply spec_gcd_gt. + rewrite (@spec_double_modn1 w w_digits w_zdigits w_0 w_WW); trivial. + apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + destruct (Z_mod_lt x y);zarith end. + rewrite spec_w_0 in Hml;Spec_w_to_Z ml;elimtype False;omega. + rewrite spec_w_0 in Hmh. assert ([[WW bh bl]] > [[WW mh ml]]). + rewrite H;simpl; apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + destruct (Z_mod_lt x y);zarith end. + assert (H1:= spec_ww_mod_gt_aux _ _ _ H0 Hmh). + assert (H3 : 0 < [[WW mh ml]]). + simpl;Spec_w_to_Z ml. apply Zlt_le_trans with ([|mh|]*wB);zarith. + apply Zmult_lt_0_compat;zarith. + apply Zis_gcd_mod;zarith. simpl in *;rewrite <- H1. + destruct (ww_mod_gt_aux bh bl mh ml) as [ |rh rl]. simpl; apply Zis_gcd_0. + simpl;apply Hcont. simpl in H1;rewrite H1. + apply Zlt_gt;match goal with | |- ?x mod ?y < ?y => + destruct (Z_mod_lt x y);zarith end. + apply Zle_trans with (2^n/2). + apply Zdiv_le_lower_bound;zarith. + apply Zle_trans with ([|bh|] * wB + [|bl|]);zarith. + assert (H3' := Z_div_mod_eq [[WW bh bl]] [[WW mh ml]] (Zlt_gt _ _ H3)). + assert (H4' : 0 <= [[WW bh bl]]/[[WW mh ml]]). + apply Zge_le;apply Z_div_ge0;zarith. simpl in *;rewrite H1. + pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3'. + destruct (Zle_lt_or_eq _ _ H4'). + assert (H6' : [[WW bh bl]] mod [[WW mh ml]] = + [[WW bh bl]] - [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])). + simpl;pattern ([|bh|] * wB + [|bl|]) at 2;rewrite H3';ring. simpl in H6'. + assert ([[WW mh ml]] <= [[WW mh ml]] * ([[WW bh bl]]/[[WW mh ml]])). + simpl;pattern ([|mh|]*wB+[|ml|]) at 1;rewrite <- Zmult_1_r;zarith. + simpl in *;assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in H8; + zarith. + assert (H8 := Z_mod_lt [[WW bh bl]] [[WW mh ml]]);simpl in *;zarith. + rewrite <- H4 in H3';rewrite Zmult_0_r in H3';simpl in H3';zarith. + pattern n at 1;replace n with (n-1+1);try ring. + rewrite Zpower_exp;zarith. change (2^1) with 2. + rewrite Z_div_mult;zarith. + assert (2^1 <= 2^n). change (2^1) with 2;zarith. + assert (H7 := @Zpower_le_monotone_inv 2 1 n);zarith. + rewrite spec_w_0 in Hmh;Spec_w_to_Z mh;elimtype False;zarith. + rewrite spec_w_0 in Hbh;Spec_w_to_Z bh;elimtype False;zarith. + Qed. + + Lemma spec_ww_gcd_gt_aux : + forall p cont n, + (forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> + [[WW yh yl]] <= 2^n -> + Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) -> + forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] -> + [[WW bh bl]] <= 2^(Zpos p + n) -> + Zis_gcd [[WW ah al]] [[WW bh bl]] + [[ww_gcd_gt_aux p cont ah al bh bl]]. + Proof. + induction p;intros cont n Hcont ah al bh bl Hgt Hs;simpl ww_gcd_gt_aux. + assert (0 < Zpos p). unfold Zlt;reflexivity. + apply spec_ww_gcd_gt_aux_body with (n := Zpos (xI p) + n); + trivial;rewrite Zpos_xI. + intros. apply IHp with (n := Zpos p + n);zarith. + intros. apply IHp with (n := n );zarith. + apply Zle_trans with (2 ^ (2* Zpos p + 1+ n -1));zarith. + apply Zpower_le_monotone2;zarith. + assert (0 < Zpos p). unfold Zlt;reflexivity. + apply spec_ww_gcd_gt_aux_body with (n := Zpos (xO p) + n );trivial. + rewrite (Zpos_xO p). + intros. apply IHp with (n := Zpos p + n - 1);zarith. + intros. apply IHp with (n := n -1 );zarith. + intros;apply Hcont;zarith. + apply Zle_trans with (2^(n-1));zarith. + apply Zpower_le_monotone2;zarith. + apply Zle_trans with (2 ^ (Zpos p + n -1));zarith. + apply Zpower_le_monotone2;zarith. + apply Zle_trans with (2 ^ (2*Zpos p + n -1));zarith. + apply Zpower_le_monotone2;zarith. + apply spec_ww_gcd_gt_aux_body with (n := n+1);trivial. + rewrite Zplus_comm;trivial. + ring_simplify (n + 1 - 1);trivial. + Qed. + +End DoubleDivGt. + +Section DoubleDiv. + + Variable w : Type. + Variable w_digits : positive. + Variable ww_1 : zn2z w. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + + Variable ww_div_gt : zn2z w -> zn2z w -> zn2z w * zn2z w. + Variable ww_mod_gt : zn2z w -> zn2z w -> zn2z w. + + Definition ww_div a b := + match ww_compare a b with + | Gt => ww_div_gt a b + | Eq => (ww_1, W0) + | Lt => (W0, a) + end. + + Definition ww_mod a b := + match ww_compare a b with + | Gt => ww_mod_gt a b + | Eq => W0 + | Lt => a + end. + + Variable w_to_Z : w -> Z. + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_ww_1 : [[ww_1]] = 1. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_div_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + let (q,r) := ww_div_gt a b in + [[a]] = [[q]] * [[b]] + [[r]] /\ + 0 <= [[r]] < [[b]]. + Variable spec_ww_mod_gt : forall a b, [[a]] > [[b]] -> 0 < [[b]] -> + [[ww_mod_gt a b]] = [[a]] mod [[b]]. + + Ltac Spec_w_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_to_Z x). + + Ltac Spec_ww_to_Z x := + let H:= fresh "HH" in + assert (H:= spec_ww_to_Z w_digits w_to_Z spec_to_Z x). + + Lemma spec_ww_div : forall a b, 0 < [[b]] -> + let (q,r) := ww_div a b in + [[a]] = [[q]] * [[b]] + [[r]] /\ + 0 <= [[r]] < [[b]]. + Proof. + intros a b Hpos;unfold ww_div. + assert (H:=spec_ww_compare a b);destruct (ww_compare a b). + simpl;rewrite spec_ww_1;split;zarith. + simpl;split;[ring|Spec_ww_to_Z a;zarith]. + apply spec_ww_div_gt;trivial. + Qed. + + Lemma spec_ww_mod : forall a b, 0 < [[b]] -> + [[ww_mod a b]] = [[a]] mod [[b]]. + Proof. + intros a b Hpos;unfold ww_mod. + assert (H := spec_ww_compare a b);destruct (ww_compare a b). + simpl;apply Zmod_unique with 1;try rewrite H;zarith. + Spec_ww_to_Z a;symmetry;apply Zmod_small;zarith. + apply spec_ww_mod_gt;trivial. + Qed. + + + Variable w_0 : w. + Variable w_1 : w. + Variable w_compare : w -> w -> comparison. + Variable w_eq0 : w -> bool. + Variable w_gcd_gt : w -> w -> w. + Variable _ww_digits : positive. + Variable spec_ww_digits_ : _ww_digits = xO w_digits. + Variable ww_gcd_gt_fix : + positive -> (w -> w -> w -> w -> zn2z w) -> + w -> w -> w -> w -> zn2z w. + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0. + Variable spec_gcd_gt : forall a b, [|a|] > [|b|] -> + Zis_gcd [|a|] [|b|] [|w_gcd_gt a b|]. + Variable spec_gcd_gt_fix : + forall p cont n, + (forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> + [[WW yh yl]] <= 2^n -> + Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]) -> + forall ah al bh bl , [[WW ah al]] > [[WW bh bl]] -> + [[WW bh bl]] <= 2^(Zpos p + n) -> + Zis_gcd [[WW ah al]] [[WW bh bl]] + [[ww_gcd_gt_fix p cont ah al bh bl]]. + + Definition gcd_cont (xh xl yh yl:w) := + match w_compare w_1 yl with + | Eq => ww_1 + | _ => WW xh xl + end. + + Lemma spec_gcd_cont : forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> + [[WW yh yl]] <= 1 -> + Zis_gcd [[WW xh xl]] [[WW yh yl]] [[gcd_cont xh xl yh yl]]. + Proof. + intros xh xl yh yl Hgt' Hle. simpl in Hle. + assert ([|yh|] = 0). + change 1 with (0*wB+1) in Hle. + assert (0 <= 1 < wB). split;zarith. apply wB_pos. + assert (H1:= beta_lex _ _ _ _ _ Hle (spec_to_Z yl) H). + Spec_w_to_Z yh;zarith. + unfold gcd_cont;assert (Hcmpy:=spec_compare w_1 yl); + rewrite spec_w_1 in Hcmpy. + simpl;rewrite H;simpl;destruct (w_compare w_1 yl). + rewrite spec_ww_1;rewrite <- Hcmpy;apply Zis_gcd_mod;zarith. + rewrite <- (Zmod_unique ([|xh|]*wB+[|xl|]) 1 ([|xh|]*wB+[|xl|]) 0);zarith. + rewrite H in Hle; elimtype False;zarith. + assert ([|yl|] = 0). Spec_w_to_Z yl;zarith. + rewrite H0;simpl;apply Zis_gcd_0;trivial. + Qed. + + + Variable cont : w -> w -> w -> w -> zn2z w. + Variable spec_cont : forall xh xl yh yl, + [[WW xh xl]] > [[WW yh yl]] -> + [[WW yh yl]] <= 1 -> + Zis_gcd [[WW xh xl]] [[WW yh yl]] [[cont xh xl yh yl]]. + + Definition ww_gcd_gt a b := + match a, b with + | W0, _ => b + | _, W0 => a + | WW ah al, WW bh bl => + if w_eq0 ah then (WW w_0 (w_gcd_gt al bl)) + else ww_gcd_gt_fix _ww_digits cont ah al bh bl + end. + + Definition ww_gcd a b := + Eval lazy beta delta [ww_gcd_gt] in + match ww_compare a b with + | Gt => ww_gcd_gt a b + | Eq => a + | Lt => ww_gcd_gt b a + end. + + Lemma spec_ww_gcd_gt : forall a b, [[a]] > [[b]] -> + Zis_gcd [[a]] [[b]] [[ww_gcd_gt a b]]. + Proof. + intros a b Hgt;unfold ww_gcd_gt. + destruct a as [ |ah al]. simpl;apply Zis_gcd_sym;apply Zis_gcd_0. + destruct b as [ |bh bl]. simpl;apply Zis_gcd_0. + simpl in Hgt. generalize (@spec_eq0 ah);destruct (w_eq0 ah);intros. + simpl;rewrite H in Hgt;trivial;rewrite H;trivial;rewrite spec_w_0;simpl. + assert ([|bh|] <= 0). + apply beta_lex with (d:=[|al|])(b:=[|bl|]) (beta := wB);zarith. + Spec_w_to_Z bh;assert ([|bh|] = 0);zarith. rewrite H1 in Hgt;simpl in Hgt. + rewrite H1;simpl;auto. clear H. + apply spec_gcd_gt_fix with (n:= 0);trivial. + rewrite Zplus_0_r;rewrite spec_ww_digits_. + change (2 ^ Zpos (xO w_digits)) with wwB. Spec_ww_to_Z (WW bh bl);zarith. + Qed. + + Lemma spec_ww_gcd : forall a b, Zis_gcd [[a]] [[b]] [[ww_gcd a b]]. + Proof. + intros a b. + change (ww_gcd a b) with + (match ww_compare a b with + | Gt => ww_gcd_gt a b + | Eq => a + | Lt => ww_gcd_gt b a + end). + assert (Hcmp := spec_ww_compare a b);destruct (ww_compare a b). + Spec_ww_to_Z b;rewrite Hcmp. + apply Zis_gcd_for_euclid with 1;zarith. + ring_simplify ([[b]] - 1 * [[b]]). apply Zis_gcd_0;zarith. + apply Zis_gcd_sym;apply spec_ww_gcd_gt;zarith. + apply spec_ww_gcd_gt;zarith. + Qed. + +End DoubleDiv. + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v new file mode 100644 index 00000000..d6f6a05f --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleDivn1.v @@ -0,0 +1,528 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleDivn1.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section GENDIVN1. + + Variable w : Type. + Variable w_digits : positive. + Variable w_zdigits : w. + Variable w_0 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_head0 : w -> w. + Variable w_add_mul_div : w -> w -> w -> w. + Variable w_div21 : w -> w -> w -> w * w. + Variable w_compare : w -> w -> comparison. + Variable w_sub : w -> w -> w. + + + + (* ** For proofs ** *) + Variable w_to_Z : w -> Z. + + Notation wB := (base w_digits). + + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x) + (at level 0, x at level 99). + Notation "[[ x ]]" := (zn2z_to_Z wB w_to_Z x) (at level 0, x at level 99). + + Variable spec_to_Z : forall x, 0 <= [| x |] < wB. + Variable spec_w_zdigits: [|w_zdigits|] = Zpos w_digits. + Variable spec_0 : [|w_0|] = 0. + Variable spec_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_head0 : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ [|w_head0 x|] * [|x|] < wB. + Variable spec_add_mul_div : forall x y p, + [|p|] <= Zpos w_digits -> + [| w_add_mul_div p x y |] = + ([|x|] * (2 ^ [|p|]) + + [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB. + Variable spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Variable spec_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_sub: forall x y, + [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + + + + Section DIVAUX. + Variable b2p : w. + Variable b2p_le : wB/2 <= [|b2p|]. + + Definition double_divn1_0_aux n (divn1: w -> word w n -> word w n * w) r h := + let (hh,hl) := double_split w_0 n h in + let (qh,rh) := divn1 r hh in + let (ql,rl) := divn1 rh hl in + (double_WW w_WW n qh ql, rl). + + Fixpoint double_divn1_0 (n:nat) : w -> word w n -> word w n * w := + match n return w -> word w n -> word w n * w with + | O => fun r x => w_div21 r x b2p + | S n => double_divn1_0_aux n (double_divn1_0 n) + end. + + Lemma spec_split : forall (n : nat) (x : zn2z (word w n)), + let (h, l) := double_split w_0 n x in + [!S n | x!] = [!n | h!] * double_wB w_digits n + [!n | l!]. + Proof (spec_double_split w_0 w_digits w_to_Z spec_0). + + Lemma spec_double_divn1_0 : forall n r a, + [|r|] < [|b2p|] -> + let (q,r') := double_divn1_0 n r a in + [|r|] * double_wB w_digits n + [!n|a!] = [!n|q!] * [|b2p|] + [|r'|] /\ + 0 <= [|r'|] < [|b2p|]. + Proof. + induction n;intros. + exact (spec_div21 a b2p_le H). + simpl (double_divn1_0 (S n) r a); unfold double_divn1_0_aux. + assert (H1 := spec_split n a);destruct (double_split w_0 n a) as (hh,hl). + rewrite H1. + assert (H2 := IHn r hh H);destruct (double_divn1_0 n r hh) as (qh,rh). + destruct H2. + assert ([|rh|] < [|b2p|]). omega. + assert (H4 := IHn rh hl H3);destruct (double_divn1_0 n rh hl) as (ql,rl). + destruct H4;split;trivial. + rewrite spec_double_WW;trivial. + rewrite <- double_wB_wwB. + rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite H0;rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc. + rewrite H4;ring. + Qed. + + Definition double_modn1_0_aux n (modn1:w -> word w n -> w) r h := + let (hh,hl) := double_split w_0 n h in modn1 (modn1 r hh) hl. + + Fixpoint double_modn1_0 (n:nat) : w -> word w n -> w := + match n return w -> word w n -> w with + | O => fun r x => snd (w_div21 r x b2p) + | S n => double_modn1_0_aux n (double_modn1_0 n) + end. + + Lemma spec_double_modn1_0 : forall n r x, + double_modn1_0 n r x = snd (double_divn1_0 n r x). + Proof. + induction n;simpl;intros;trivial. + unfold double_modn1_0_aux, double_divn1_0_aux. + destruct (double_split w_0 n x) as (hh,hl). + rewrite (IHn r hh). + destruct (double_divn1_0 n r hh) as (qh,rh);simpl. + rewrite IHn. destruct (double_divn1_0 n rh hl);trivial. + Qed. + + Variable p : w. + Variable p_bounded : [|p|] <= Zpos w_digits. + + Lemma spec_add_mul_divp : forall x y, + [| w_add_mul_div p x y |] = + ([|x|] * (2 ^ [|p|]) + + [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB. + Proof. + intros;apply spec_add_mul_div;auto. + Qed. + + Definition double_divn1_p_aux n + (divn1 : w -> word w n -> word w n -> word w n * w) r h l := + let (hh,hl) := double_split w_0 n h in + let (lh,ll) := double_split w_0 n l in + let (qh,rh) := divn1 r hh hl in + let (ql,rl) := divn1 rh hl lh in + (double_WW w_WW n qh ql, rl). + + Fixpoint double_divn1_p (n:nat) : w -> word w n -> word w n -> word w n * w := + match n return w -> word w n -> word w n -> word w n * w with + | O => fun r h l => w_div21 r (w_add_mul_div p h l) b2p + | S n => double_divn1_p_aux n (double_divn1_p n) + end. + + Lemma p_lt_double_digits : forall n, [|p|] <= Zpos (double_digits w_digits n). + Proof. +(* + induction n;simpl. destruct p_bounded;trivial. + case (spec_to_Z p); rewrite Zpos_xO;auto with zarith. +*) + induction n;simpl. trivial. + case (spec_to_Z p); rewrite Zpos_xO;auto with zarith. + Qed. + + Lemma spec_double_divn1_p : forall n r h l, + [|r|] < [|b2p|] -> + let (q,r') := double_divn1_p n r h l in + [|r|] * double_wB w_digits n + + ([!n|h!]*2^[|p|] + + [!n|l!] / (2^(Zpos(double_digits w_digits n) - [|p|]))) + mod double_wB w_digits n = [!n|q!] * [|b2p|] + [|r'|] /\ + 0 <= [|r'|] < [|b2p|]. + Proof. + case (spec_to_Z p); intros HH0 HH1. + induction n;intros. + simpl (double_divn1_p 0 r h l). + unfold double_to_Z, double_wB, double_digits. + rewrite <- spec_add_mul_divp. + exact (spec_div21 (w_add_mul_div p h l) b2p_le H). + simpl (double_divn1_p (S n) r h l). + unfold double_divn1_p_aux. + assert (H1 := spec_split n h);destruct (double_split w_0 n h) as (hh,hl). + rewrite H1. rewrite <- double_wB_wwB. + assert (H2 := spec_split n l);destruct (double_split w_0 n l) as (lh,ll). + rewrite H2. + replace ([|r|] * (double_wB w_digits n * double_wB w_digits n) + + (([!n|hh!] * double_wB w_digits n + [!n|hl!]) * 2 ^ [|p|] + + ([!n|lh!] * double_wB w_digits n + [!n|ll!]) / + 2^(Zpos (double_digits w_digits (S n)) - [|p|])) mod + (double_wB w_digits n * double_wB w_digits n)) with + (([|r|] * double_wB w_digits n + ([!n|hh!] * 2^[|p|] + + [!n|hl!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod + double_wB w_digits n) * double_wB w_digits n + + ([!n|hl!] * 2^[|p|] + + [!n|lh!] / 2^(Zpos (double_digits w_digits n) - [|p|])) mod + double_wB w_digits n). + generalize (IHn r hh hl H);destruct (double_divn1_p n r hh hl) as (qh,rh); + intros (H3,H4);rewrite H3. + assert ([|rh|] < [|b2p|]). omega. + replace (([!n|qh!] * [|b2p|] + [|rh|]) * double_wB w_digits n + + ([!n|hl!] * 2 ^ [|p|] + + [!n|lh!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])) mod + double_wB w_digits n) with + ([!n|qh!] * [|b2p|] *double_wB w_digits n + ([|rh|]*double_wB w_digits n + + ([!n|hl!] * 2 ^ [|p|] + + [!n|lh!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])) mod + double_wB w_digits n)). 2:ring. + generalize (IHn rh hl lh H0);destruct (double_divn1_p n rh hl lh) as (ql,rl); + intros (H5,H6);rewrite H5. + split;[rewrite spec_double_WW;trivial;ring|trivial]. + assert (Uhh := spec_double_to_Z w_digits w_to_Z spec_to_Z n hh); + unfold double_wB,base in Uhh. + assert (Uhl := spec_double_to_Z w_digits w_to_Z spec_to_Z n hl); + unfold double_wB,base in Uhl. + assert (Ulh := spec_double_to_Z w_digits w_to_Z spec_to_Z n lh); + unfold double_wB,base in Ulh. + assert (Ull := spec_double_to_Z w_digits w_to_Z spec_to_Z n ll); + unfold double_wB,base in Ull. + unfold double_wB,base. + assert (UU:=p_lt_double_digits n). + rewrite Zdiv_shift_r;auto with zarith. + 2:change (Zpos (double_digits w_digits (S n))) + with (2*Zpos (double_digits w_digits n));auto with zarith. + replace (2 ^ (Zpos (double_digits w_digits (S n)) - [|p|])) with + (2^(Zpos (double_digits w_digits n) - [|p|])*2^Zpos (double_digits w_digits n)). + rewrite Zdiv_mult_cancel_r;auto with zarith. + rewrite Zmult_plus_distr_l with (p:= 2^[|p|]). + pattern ([!n|hl!] * 2^[|p|]) at 2; + rewrite (shift_unshift_mod (Zpos(double_digits w_digits n))([|p|])([!n|hl!])); + auto with zarith. + rewrite Zplus_assoc. + replace + ([!n|hh!] * 2^Zpos (double_digits w_digits n)* 2^[|p|] + + ([!n|hl!] / 2^(Zpos (double_digits w_digits n)-[|p|])* + 2^Zpos(double_digits w_digits n))) + with + (([!n|hh!] *2^[|p|] + double_to_Z w_digits w_to_Z n hl / + 2^(Zpos (double_digits w_digits n)-[|p|])) + * 2^Zpos(double_digits w_digits n));try (ring;fail). + rewrite <- Zplus_assoc. + rewrite <- (Zmod_shift_r ([|p|]));auto with zarith. + replace + (2 ^ Zpos (double_digits w_digits n) * 2 ^ Zpos (double_digits w_digits n)) with + (2 ^ (Zpos (double_digits w_digits n) + Zpos (double_digits w_digits n))). + rewrite (Zmod_shift_r (Zpos (double_digits w_digits n)));auto with zarith. + replace (2 ^ (Zpos (double_digits w_digits n) + Zpos (double_digits w_digits n))) + with (2^Zpos(double_digits w_digits n) *2^Zpos(double_digits w_digits n)). + rewrite (Zmult_comm (([!n|hh!] * 2 ^ [|p|] + + [!n|hl!] / 2 ^ (Zpos (double_digits w_digits n) - [|p|])))). + rewrite Zmult_mod_distr_l;auto with zarith. + ring. + rewrite Zpower_exp;auto with zarith. + assert (0 < Zpos (double_digits w_digits n)). unfold Zlt;reflexivity. + auto with zarith. + apply Z_mod_lt;auto with zarith. + rewrite Zpower_exp;auto with zarith. + split;auto with zarith. + apply Zdiv_lt_upper_bound;auto with zarith. + rewrite <- Zpower_exp;auto with zarith. + replace ([|p|] + (Zpos (double_digits w_digits n) - [|p|])) with + (Zpos(double_digits w_digits n));auto with zarith. + rewrite <- Zpower_exp;auto with zarith. + replace (Zpos (double_digits w_digits (S n)) - [|p|]) with + (Zpos (double_digits w_digits n) - [|p|] + + Zpos (double_digits w_digits n));trivial. + change (Zpos (double_digits w_digits (S n))) with + (2*Zpos (double_digits w_digits n)). ring. + Qed. + + Definition double_modn1_p_aux n (modn1 : w -> word w n -> word w n -> w) r h l:= + let (hh,hl) := double_split w_0 n h in + let (lh,ll) := double_split w_0 n l in + modn1 (modn1 r hh hl) hl lh. + + Fixpoint double_modn1_p (n:nat) : w -> word w n -> word w n -> w := + match n return w -> word w n -> word w n -> w with + | O => fun r h l => snd (w_div21 r (w_add_mul_div p h l) b2p) + | S n => double_modn1_p_aux n (double_modn1_p n) + end. + + Lemma spec_double_modn1_p : forall n r h l , + double_modn1_p n r h l = snd (double_divn1_p n r h l). + Proof. + induction n;simpl;intros;trivial. + unfold double_modn1_p_aux, double_divn1_p_aux. + destruct(double_split w_0 n h)as(hh,hl);destruct(double_split w_0 n l) as (lh,ll). + rewrite (IHn r hh hl);destruct (double_divn1_p n r hh hl) as (qh,rh). + rewrite IHn;simpl;destruct (double_divn1_p n rh hl lh);trivial. + Qed. + + End DIVAUX. + + Fixpoint high (n:nat) : word w n -> w := + match n return word w n -> w with + | O => fun a => a + | S n => + fun (a:zn2z (word w n)) => + match a with + | W0 => w_0 + | WW h l => high n h + end + end. + + Lemma spec_double_digits:forall n, Zpos w_digits <= Zpos (double_digits w_digits n). + Proof. + induction n;simpl;auto with zarith. + change (Zpos (xO (double_digits w_digits n))) with + (2*Zpos (double_digits w_digits n)). + assert (0 < Zpos w_digits);auto with zarith. + exact (refl_equal Lt). + Qed. + + Lemma spec_high : forall n (x:word w n), + [|high n x|] = [!n|x!] / 2^(Zpos (double_digits w_digits n) - Zpos w_digits). + Proof. + induction n;intros. + unfold high,double_digits,double_to_Z. + replace (Zpos w_digits - Zpos w_digits) with 0;try ring. + simpl. rewrite <- (Zdiv_unique [|x|] 1 [|x|] 0);auto with zarith. + assert (U2 := spec_double_digits n). + assert (U3 : 0 < Zpos w_digits). exact (refl_equal Lt). + destruct x;unfold high;fold high. + unfold double_to_Z,zn2z_to_Z;rewrite spec_0. + rewrite Zdiv_0_l;trivial. + assert (U0 := spec_double_to_Z w_digits w_to_Z spec_to_Z n w0); + assert (U1 := spec_double_to_Z w_digits w_to_Z spec_to_Z n w1). + simpl [!S n|WW w0 w1!]. + unfold double_wB,base;rewrite Zdiv_shift_r;auto with zarith. + replace (2 ^ (Zpos (double_digits w_digits (S n)) - Zpos w_digits)) with + (2^(Zpos (double_digits w_digits n) - Zpos w_digits) * + 2^Zpos (double_digits w_digits n)). + rewrite Zdiv_mult_cancel_r;auto with zarith. + rewrite <- Zpower_exp;auto with zarith. + replace (Zpos (double_digits w_digits n) - Zpos w_digits + + Zpos (double_digits w_digits n)) with + (Zpos (double_digits w_digits (S n)) - Zpos w_digits);trivial. + change (Zpos (double_digits w_digits (S n))) with + (2*Zpos (double_digits w_digits n));ring. + change (Zpos (double_digits w_digits (S n))) with + (2*Zpos (double_digits w_digits n)); auto with zarith. + Qed. + + Definition double_divn1 (n:nat) (a:word w n) (b:w) := + let p := w_head0 b in + match w_compare p w_0 with + | Gt => + let b2p := w_add_mul_div p b w_0 in + let ha := high n a in + let k := w_sub w_zdigits p in + let lsr_n := w_add_mul_div k w_0 in + let r0 := w_add_mul_div p w_0 ha in + let (q,r) := double_divn1_p b2p p n r0 a (double_0 w_0 n) in + (q, lsr_n r) + | _ => double_divn1_0 b n w_0 a + end. + + Lemma spec_double_divn1 : forall n a b, + 0 < [|b|] -> + let (q,r) := double_divn1 n a b in + [!n|a!] = [!n|q!] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + intros n a b H. unfold double_divn1. + case (spec_head0 H); intros H0 H1. + case (spec_to_Z (w_head0 b)); intros HH1 HH2. + generalize (spec_compare (w_head0 b) w_0); case w_compare; + rewrite spec_0; intros H2; auto with zarith. + assert (Hv1: wB/2 <= [|b|]). + generalize H0; rewrite H2; rewrite Zpower_0_r; + rewrite Zmult_1_l; auto. + assert (Hv2: [|w_0|] < [|b|]). + rewrite spec_0; auto. + generalize (spec_double_divn1_0 Hv1 n a Hv2). + rewrite spec_0;rewrite Zmult_0_l; rewrite Zplus_0_l; auto. + contradict H2; auto with zarith. + assert (HHHH : 0 < [|w_head0 b|]); auto with zarith. + assert ([|w_head0 b|] < Zpos w_digits). + case (Zle_or_lt (Zpos w_digits) [|w_head0 b|]); auto; intros HH. + assert (2 ^ [|w_head0 b|] < wB). + apply Zle_lt_trans with (2 ^ [|w_head0 b|] * [|b|]);auto with zarith. + replace (2 ^ [|w_head0 b|]) with (2^[|w_head0 b|] * 1);try (ring;fail). + apply Zmult_le_compat;auto with zarith. + assert (wB <= 2^[|w_head0 b|]). + unfold base;apply Zpower_le_monotone;auto with zarith. omega. + assert ([|w_add_mul_div (w_head0 b) b w_0|] = + 2 ^ [|w_head0 b|] * [|b|]). + rewrite (spec_add_mul_div b w_0); auto with zarith. + rewrite spec_0;rewrite Zdiv_0_l; try omega. + rewrite Zplus_0_r; rewrite Zmult_comm. + rewrite Zmod_small; auto with zarith. + assert (H5 := spec_to_Z (high n a)). + assert + ([|w_add_mul_div (w_head0 b) w_0 (high n a)|] + <[|w_add_mul_div (w_head0 b) b w_0|]). + rewrite H4. + rewrite spec_add_mul_div;auto with zarith. + rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l. + assert (([|high n a|]/2^(Zpos w_digits - [|w_head0 b|])) < wB). + apply Zdiv_lt_upper_bound;auto with zarith. + apply Zlt_le_trans with wB;auto with zarith. + pattern wB at 1;replace wB with (wB*1);try ring. + apply Zmult_le_compat;auto with zarith. + assert (H6 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|])); + auto with zarith. + rewrite Zmod_small;auto with zarith. + apply Zdiv_lt_upper_bound;auto with zarith. + apply Zlt_le_trans with wB;auto with zarith. + apply Zle_trans with (2 ^ [|w_head0 b|] * [|b|] * 2). + rewrite <- wB_div_2; try omega. + apply Zmult_le_compat;auto with zarith. + pattern 2 at 1;rewrite <- Zpower_1_r. + apply Zpower_le_monotone;split;auto with zarith. + rewrite <- H4 in H0. + assert (Hb3: [|w_head0 b|] <= Zpos w_digits); auto with zarith. + assert (H7:= spec_double_divn1_p H0 Hb3 n a (double_0 w_0 n) H6). + destruct (double_divn1_p (w_add_mul_div (w_head0 b) b w_0) (w_head0 b) n + (w_add_mul_div (w_head0 b) w_0 (high n a)) a + (double_0 w_0 n)) as (q,r). + assert (U:= spec_double_digits n). + rewrite spec_double_0 in H7;trivial;rewrite Zdiv_0_l in H7. + rewrite Zplus_0_r in H7. + rewrite spec_add_mul_div in H7;auto with zarith. + rewrite spec_0 in H7;rewrite Zmult_0_l in H7;rewrite Zplus_0_l in H7. + assert (([|high n a|] / 2 ^ (Zpos w_digits - [|w_head0 b|])) mod wB + = [!n|a!] / 2^(Zpos (double_digits w_digits n) - [|w_head0 b|])). + rewrite Zmod_small;auto with zarith. + rewrite spec_high. rewrite Zdiv_Zdiv;auto with zarith. + rewrite <- Zpower_exp;auto with zarith. + replace (Zpos (double_digits w_digits n) - Zpos w_digits + + (Zpos w_digits - [|w_head0 b|])) + with (Zpos (double_digits w_digits n) - [|w_head0 b|]);trivial;ring. + assert (H8 := Zpower_gt_0 2 (Zpos w_digits - [|w_head0 b|]));auto with zarith. + split;auto with zarith. + apply Zle_lt_trans with ([|high n a|]);auto with zarith. + apply Zdiv_le_upper_bound;auto with zarith. + pattern ([|high n a|]) at 1;rewrite <- Zmult_1_r. + apply Zmult_le_compat;auto with zarith. + rewrite H8 in H7;unfold double_wB,base in H7. + rewrite <- shift_unshift_mod in H7;auto with zarith. + rewrite H4 in H7. + assert ([|w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r|] + = [|r|]/2^[|w_head0 b|]). + rewrite spec_add_mul_div. + rewrite spec_0;rewrite Zmult_0_l;rewrite Zplus_0_l. + replace (Zpos w_digits - [|w_sub w_zdigits (w_head0 b)|]) + with ([|w_head0 b|]). + rewrite Zmod_small;auto with zarith. + assert (H9 := spec_to_Z r). + split;auto with zarith. + apply Zle_lt_trans with ([|r|]);auto with zarith. + apply Zdiv_le_upper_bound;auto with zarith. + pattern ([|r|]) at 1;rewrite <- Zmult_1_r. + apply Zmult_le_compat;auto with zarith. + assert (H10 := Zpower_gt_0 2 ([|w_head0 b|]));auto with zarith. + rewrite spec_sub. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + case (spec_to_Z w_zdigits); auto with zarith. + rewrite spec_sub. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + case (spec_to_Z w_zdigits); auto with zarith. + case H7; intros H71 H72. + split. + rewrite <- (Z_div_mult [!n|a!] (2^[|w_head0 b|]));auto with zarith. + rewrite H71;rewrite H9. + replace ([!n|q!] * (2 ^ [|w_head0 b|] * [|b|])) + with ([!n|q!] *[|b|] * 2^[|w_head0 b|]); + try (ring;fail). + rewrite Z_div_plus_l;auto with zarith. + assert (H10 := spec_to_Z + (w_add_mul_div (w_sub w_zdigits (w_head0 b)) w_0 r));split; + auto with zarith. + rewrite H9. + apply Zdiv_lt_upper_bound;auto with zarith. + rewrite Zmult_comm;auto with zarith. + exact (spec_double_to_Z w_digits w_to_Z spec_to_Z n a). + Qed. + + + Definition double_modn1 (n:nat) (a:word w n) (b:w) := + let p := w_head0 b in + match w_compare p w_0 with + | Gt => + let b2p := w_add_mul_div p b w_0 in + let ha := high n a in + let k := w_sub w_zdigits p in + let lsr_n := w_add_mul_div k w_0 in + let r0 := w_add_mul_div p w_0 ha in + let r := double_modn1_p b2p p n r0 a (double_0 w_0 n) in + lsr_n r + | _ => double_modn1_0 b n w_0 a + end. + + Lemma spec_double_modn1_aux : forall n a b, + double_modn1 n a b = snd (double_divn1 n a b). + Proof. + intros n a b;unfold double_divn1,double_modn1. + generalize (spec_compare (w_head0 b) w_0); case w_compare; + rewrite spec_0; intros H2; auto with zarith. + apply spec_double_modn1_0. + apply spec_double_modn1_0. + rewrite spec_double_modn1_p. + destruct (double_divn1_p (w_add_mul_div (w_head0 b) b w_0) (w_head0 b) n + (w_add_mul_div (w_head0 b) w_0 (high n a)) a (double_0 w_0 n));simpl;trivial. + Qed. + + Lemma spec_double_modn1 : forall n a b, 0 < [|b|] -> + [|double_modn1 n a b|] = [!n|a!] mod [|b|]. + Proof. + intros n a b H;assert (H1 := spec_double_divn1 n a H). + assert (H2 := spec_double_modn1_aux n a b). + rewrite H2;destruct (double_divn1 n a b) as (q,r). + simpl;apply Zmod_unique with (double_to_Z w_digits w_to_Z n q);auto with zarith. + destruct H1 as (h1,h2);rewrite h1;ring. + Qed. + +End GENDIVN1. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v new file mode 100644 index 00000000..50c72487 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleLift.v @@ -0,0 +1,487 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleLift.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleLift. + Variable w : Type. + Variable w_0 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable w_head0 : w -> w. + Variable w_tail0 : w -> w. + Variable w_add: w -> w -> zn2z w. + Variable w_add_mul_div : w -> w -> w -> w. + Variable ww_sub: zn2z w -> zn2z w -> zn2z w. + Variable w_digits : positive. + Variable ww_Digits : positive. + Variable w_zdigits : w. + Variable ww_zdigits : zn2z w. + Variable low: zn2z w -> w. + + Definition ww_head0 x := + match x with + | W0 => ww_zdigits + | WW xh xl => + match w_compare w_0 xh with + | Eq => w_add w_zdigits (w_head0 xl) + | _ => w_0W (w_head0 xh) + end + end. + + + Definition ww_tail0 x := + match x with + | W0 => ww_zdigits + | WW xh xl => + match w_compare w_0 xl with + | Eq => w_add w_zdigits (w_tail0 xh) + | _ => w_0W (w_tail0 xl) + end + end. + + + (* 0 < p < ww_digits *) + Definition ww_add_mul_div p x y := + let zdigits := w_0W w_zdigits in + match x, y with + | W0, W0 => W0 + | W0, WW yh yl => + match ww_compare p zdigits with + | Eq => w_0W yh + | Lt => w_0W (w_add_mul_div (low p) w_0 yh) + | Gt => + let n := low (ww_sub p zdigits) in + w_WW (w_add_mul_div n w_0 yh) (w_add_mul_div n yh yl) + end + | WW xh xl, W0 => + match ww_compare p zdigits with + | Eq => w_W0 xl + | Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl w_0) + | Gt => + let n := low (ww_sub p zdigits) in + w_W0 (w_add_mul_div n xl w_0) + end + | WW xh xl, WW yh yl => + match ww_compare p zdigits with + | Eq => w_WW xl yh + | Lt => w_WW (w_add_mul_div (low p) xh xl) (w_add_mul_div (low p) xl yh) + | Gt => + let n := low (ww_sub p zdigits) in + w_WW (w_add_mul_div n xl yh) (w_add_mul_div n yh yl) + end + end. + + Section DoubleProof. + Variable w_to_Z : w -> Z. + + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_compare : forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_digits : ww_Digits = xO w_digits. + Variable spec_w_head00 : forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits. + Variable spec_w_head0 : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB. + Variable spec_w_tail00 : forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits. + Variable spec_w_tail0 : forall x, 0 < [|x|] -> + exists y, 0 <= y /\ [|x|] = (2* y + 1) * (2 ^ [|w_tail0 x|]). + Variable spec_w_add_mul_div : forall x y p, + [|p|] <= Zpos w_digits -> + [| w_add_mul_div p x y |] = + ([|x|] * (2 ^ [|p|]) + + [|y|] / (2 ^ ((Zpos w_digits) - [|p|]))) mod wB. + Variable spec_w_add: forall x y, + [[w_add x y]] = [|x|] + [|y|]. + Variable spec_ww_sub: forall x y, + [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. + + Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits. + Variable spec_low: forall x, [| low x|] = [[x]] mod wB. + + Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos ww_Digits. + + Hint Resolve div_le_0 div_lt w_to_Z_wwB: lift. + Ltac zarith := auto with zarith lift. + + Lemma spec_ww_head00 : forall x, [[x]] = 0 -> [[ww_head0 x]] = Zpos ww_Digits. + Proof. + intros x; case x; unfold ww_head0. + intros HH; rewrite spec_ww_zdigits; auto. + intros xh xl; simpl; intros Hx. + case (spec_to_Z xh); intros Hx1 Hx2. + case (spec_to_Z xl); intros Hy1 Hy2. + assert (F1: [|xh|] = 0). + case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + apply Zlt_le_trans with (1 := Hy3); auto with zarith. + pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]). + apply Zplus_le_compat_r; auto with zarith. + case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith. + apply Zmult_lt_0_compat; auto with zarith. + generalize (spec_compare w_0 xh); case w_compare. + intros H; simpl. + rewrite spec_w_add; rewrite spec_w_head00. + rewrite spec_zdigits; rewrite spec_ww_digits. + rewrite Zpos_xO; auto with zarith. + rewrite F1 in Hx; auto with zarith. + rewrite spec_w_0; auto with zarith. + rewrite spec_w_0; auto with zarith. + Qed. + + Lemma spec_ww_head0 : forall x, 0 < [[x]] -> + wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB. + Proof. + clear spec_ww_zdigits. + rewrite wwB_div_2;rewrite Zmult_comm;rewrite wwB_wBwB. + assert (U:= lt_0_wB w_digits); destruct x as [ |xh xl];simpl ww_to_Z;intros H. + unfold Zlt in H;discriminate H. + assert (H0 := spec_compare w_0 xh);rewrite spec_w_0 in H0. + destruct (w_compare w_0 xh). + rewrite <- H0. simpl Zplus. rewrite <- H0 in H;simpl in H. + case (spec_to_Z w_zdigits); + case (spec_to_Z (w_head0 xl)); intros HH1 HH2 HH3 HH4. + rewrite spec_w_add. + rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith. + case (spec_w_head0 H); intros H1 H2. + rewrite Zpower_2; fold wB; rewrite <- Zmult_assoc; split. + apply Zmult_le_compat_l; auto with zarith. + apply Zmult_lt_compat_l; auto with zarith. + assert (H1 := spec_w_head0 H0). + rewrite spec_w_0W. + split. + rewrite Zmult_plus_distr_r;rewrite Zmult_assoc. + apply Zle_trans with (2 ^ [|w_head0 xh|] * [|xh|] * wB). + rewrite Zmult_comm; zarith. + assert (0 <= 2 ^ [|w_head0 xh|] * [|xl|]);zarith. + assert (H2:=spec_to_Z xl);apply Zmult_le_0_compat;zarith. + case (spec_to_Z (w_head0 xh)); intros H2 _. + generalize ([|w_head0 xh|]) H1 H2;clear H1 H2; + intros p H1 H2. + assert (Eq1 : 2^p < wB). + rewrite <- (Zmult_1_r (2^p));apply Zle_lt_trans with (2^p*[|xh|]);zarith. + assert (Eq2: p < Zpos w_digits). + destruct (Zle_or_lt (Zpos w_digits) p);trivial;contradict Eq1. + apply Zle_not_lt;unfold base;apply Zpower_le_monotone;zarith. + assert (Zpos w_digits = p + (Zpos w_digits - p)). ring. + rewrite Zpower_2. + unfold base at 2;rewrite H3;rewrite Zpower_exp;zarith. + rewrite <- Zmult_assoc; apply Zmult_lt_compat_l; zarith. + rewrite <- (Zplus_0_r (2^(Zpos w_digits - p)*wB));apply beta_lex_inv;zarith. + apply Zmult_lt_reg_r with (2 ^ p); zarith. + rewrite <- Zpower_exp;zarith. + rewrite Zmult_comm;ring_simplify (Zpos w_digits - p + p);fold wB;zarith. + assert (H1 := spec_to_Z xh);zarith. + Qed. + + Lemma spec_ww_tail00 : forall x, [[x]] = 0 -> [[ww_tail0 x]] = Zpos ww_Digits. + Proof. + intros x; case x; unfold ww_tail0. + intros HH; rewrite spec_ww_zdigits; auto. + intros xh xl; simpl; intros Hx. + case (spec_to_Z xh); intros Hx1 Hx2. + case (spec_to_Z xl); intros Hy1 Hy2. + assert (F1: [|xh|] = 0). + case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + apply Zlt_le_trans with (1 := Hy3); auto with zarith. + pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]). + apply Zplus_le_compat_r; auto with zarith. + case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3. + absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith. + rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith. + apply Zmult_lt_0_compat; auto with zarith. + assert (F2: [|xl|] = 0). + rewrite F1 in Hx; auto with zarith. + generalize (spec_compare w_0 xl); case w_compare. + intros H; simpl. + rewrite spec_w_add; rewrite spec_w_tail00; auto. + rewrite spec_zdigits; rewrite spec_ww_digits. + rewrite Zpos_xO; auto with zarith. + rewrite spec_w_0; auto with zarith. + rewrite spec_w_0; auto with zarith. + Qed. + + Lemma spec_ww_tail0 : forall x, 0 < [[x]] -> + exists y, 0 <= y /\ [[x]] = (2 * y + 1) * 2 ^ [[ww_tail0 x]]. + Proof. + clear spec_ww_zdigits. + destruct x as [ |xh xl];simpl ww_to_Z;intros H. + unfold Zlt in H;discriminate H. + assert (H0 := spec_compare w_0 xl);rewrite spec_w_0 in H0. + destruct (w_compare w_0 xl). + rewrite <- H0; rewrite Zplus_0_r. + case (spec_to_Z (w_tail0 xh)); intros HH1 HH2. + generalize H; rewrite <- H0; rewrite Zplus_0_r; clear H; intros H. + case (@spec_w_tail0 xh). + apply Zmult_lt_reg_r with wB; auto with zarith. + unfold base; auto with zarith. + intros z (Hz1, Hz2); exists z; split; auto. + rewrite spec_w_add; rewrite (fun x => Zplus_comm [|x|]). + rewrite spec_zdigits; rewrite Zpower_exp; auto with zarith. + rewrite Zmult_assoc; rewrite <- Hz2; auto. + + case (spec_to_Z (w_tail0 xh)); intros HH1 HH2. + case (spec_w_tail0 H0); intros z (Hz1, Hz2). + assert (Hp: [|w_tail0 xl|] < Zpos w_digits). + case (Zle_or_lt (Zpos w_digits) [|w_tail0 xl|]); auto; intros H1. + absurd (2 ^ (Zpos w_digits) <= 2 ^ [|w_tail0 xl|]). + apply Zlt_not_le. + case (spec_to_Z xl); intros HH3 HH4. + apply Zle_lt_trans with (2 := HH4). + apply Zle_trans with (1 * 2 ^ [|w_tail0 xl|]); auto with zarith. + rewrite Hz2. + apply Zmult_le_compat_r; auto with zarith. + apply Zpower_le_monotone; auto with zarith. + exists ([|xh|] * (2 ^ ((Zpos w_digits - [|w_tail0 xl|]) - 1)) + z); split. + apply Zplus_le_0_compat; auto. + apply Zmult_le_0_compat; auto with zarith. + case (spec_to_Z xh); auto. + rewrite spec_w_0W. + rewrite (Zmult_plus_distr_r 2); rewrite <- Zplus_assoc. + rewrite Zmult_plus_distr_l; rewrite <- Hz2. + apply f_equal2 with (f := Zplus); auto. + rewrite (Zmult_comm 2). + repeat rewrite <- Zmult_assoc. + apply f_equal2 with (f := Zmult); auto. + case (spec_to_Z (w_tail0 xl)); intros HH3 HH4. + pattern 2 at 2; rewrite <- Zpower_1_r. + lazy beta; repeat rewrite <- Zpower_exp; auto with zarith. + unfold base; apply f_equal with (f := Zpower 2); auto with zarith. + + contradict H0; case (spec_to_Z xl); auto with zarith. + Qed. + + Hint Rewrite Zdiv_0_l Zmult_0_l Zplus_0_l Zmult_0_r Zplus_0_r + spec_w_W0 spec_w_0W spec_w_WW spec_w_0 + (wB_div w_digits w_to_Z spec_to_Z) + (wB_div_plus w_digits w_to_Z spec_to_Z) : w_rewrite. + Ltac w_rewrite := autorewrite with w_rewrite;trivial. + + Lemma spec_ww_add_mul_div_aux : forall xh xl yh yl p, + let zdigits := w_0W w_zdigits in + [[p]] <= Zpos (xO w_digits) -> + [[match ww_compare p zdigits with + | Eq => w_WW xl yh + | Lt => w_WW (w_add_mul_div (low p) xh xl) + (w_add_mul_div (low p) xl yh) + | Gt => + let n := low (ww_sub p zdigits) in + w_WW (w_add_mul_div n xl yh) (w_add_mul_div n yh yl) + end]] = + ([[WW xh xl]] * (2^[[p]]) + + [[WW yh yl]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB. + Proof. + clear spec_ww_zdigits. + intros xh xl yh yl p zdigits;assert (HwwB := wwB_pos w_digits). + case (spec_to_w_Z p); intros Hv1 Hv2. + replace (Zpos (xO w_digits)) with (Zpos w_digits + Zpos w_digits). + 2 : rewrite Zpos_xO;ring. + replace (Zpos w_digits + Zpos w_digits - [[p]]) with + (Zpos w_digits + (Zpos w_digits - [[p]])). 2:ring. + intros Hp; assert (Hxh := spec_to_Z xh);assert (Hxl:=spec_to_Z xl); + assert (Hx := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh xl)); + simpl in Hx;assert (Hyh := spec_to_Z yh);assert (Hyl:=spec_to_Z yl); + assert (Hy:=spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW yh yl));simpl in Hy. + generalize (spec_ww_compare p zdigits); case ww_compare; intros H1. + rewrite H1; unfold zdigits; rewrite spec_w_0W. + rewrite spec_zdigits; rewrite Zminus_diag; rewrite Zplus_0_r. + simpl ww_to_Z; w_rewrite;zarith. + fold wB. + rewrite Zmult_plus_distr_l;rewrite <- Zmult_assoc;rewrite <- Zplus_assoc. + rewrite <- Zpower_2. + rewrite <- wwB_wBwB;apply Zmod_unique with [|xh|]. + exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xl yh)). ring. + simpl ww_to_Z; w_rewrite;zarith. + assert (HH0: [|low p|] = [[p]]). + rewrite spec_low. + apply Zmod_small. + case (spec_to_w_Z p); intros HH1 HH2; split; auto. + generalize H1; unfold zdigits; rewrite spec_w_0W; + rewrite spec_zdigits; intros tmp. + apply Zlt_le_trans with (1 := tmp). + unfold base. + apply Zpower2_le_lin; auto with zarith. + 2: generalize H1; unfold zdigits; rewrite spec_w_0W; + rewrite spec_zdigits; auto with zarith. + generalize H1; unfold zdigits; rewrite spec_w_0W; + rewrite spec_zdigits; auto; clear H1; intros H1. + assert (HH: [|low p|] <= Zpos w_digits). + rewrite HH0; auto with zarith. + repeat rewrite spec_w_add_mul_div with (1 := HH). + rewrite HH0. + rewrite Zmult_plus_distr_l. + pattern ([|xl|] * 2 ^ [[p]]) at 2; + rewrite shift_unshift_mod with (n:= Zpos w_digits);fold wB;zarith. + replace ([|xh|] * wB * 2^[[p]]) with ([|xh|] * 2^[[p]] * wB). 2:ring. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. rewrite <- Zplus_assoc. + unfold base at 5;rewrite <- Zmod_shift_r;zarith. + unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits)); + fold wB;fold wwB;zarith. + rewrite wwB_wBwB;rewrite Zpower_2; rewrite Zmult_mod_distr_r;zarith. + unfold ww_digits;rewrite Zpos_xO;zarith. apply Z_mod_lt;zarith. + split;zarith. apply Zdiv_lt_upper_bound;zarith. + rewrite <- Zpower_exp;zarith. + ring_simplify ([[p]] + (Zpos w_digits - [[p]]));fold wB;zarith. + assert (Hv: [[p]] > Zpos w_digits). + generalize H1; clear H1. + unfold zdigits; rewrite spec_w_0W; rewrite spec_zdigits; auto. + clear H1. + assert (HH0: [|low (ww_sub p zdigits)|] = [[p]] - Zpos w_digits). + rewrite spec_low. + rewrite spec_ww_sub. + unfold zdigits; rewrite spec_w_0W; rewrite spec_zdigits. + rewrite <- Zmod_div_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + exists wB; unfold base. + unfold ww_digits; rewrite (Zpos_xO w_digits). + rewrite <- Zpower_exp; auto with zarith. + apply f_equal with (f := fun x => 2 ^ x); auto with zarith. + assert (HH: [|low (ww_sub p zdigits)|] <= Zpos w_digits). + rewrite HH0; auto with zarith. + replace (Zpos w_digits + (Zpos w_digits - [[p]])) with + (Zpos w_digits - ([[p]] - Zpos w_digits)); zarith. + lazy zeta; simpl ww_to_Z; w_rewrite;zarith. + repeat rewrite spec_w_add_mul_div;zarith. + rewrite HH0. + pattern wB at 5;replace wB with + (2^(([[p]] - Zpos w_digits) + + (Zpos w_digits - ([[p]] - Zpos w_digits)))). + rewrite Zpower_exp;zarith. rewrite Zmult_assoc. + rewrite Z_div_plus_l;zarith. + rewrite shift_unshift_mod with (a:= [|yh|]) (p:= [[p]] - Zpos w_digits) + (n := Zpos w_digits);zarith. fold wB. + set (u := [[p]] - Zpos w_digits). + replace [[p]] with (u + Zpos w_digits);zarith. + rewrite Zpower_exp;zarith. rewrite Zmult_assoc. fold wB. + repeat rewrite Zplus_assoc. rewrite <- Zmult_plus_distr_l. + repeat rewrite <- Zplus_assoc. + unfold base;rewrite Zmod_shift_r with (b:= Zpos (ww_digits w_digits)); + fold wB;fold wwB;zarith. + unfold base;rewrite Zmod_shift_r with (a:= Zpos w_digits) + (b:= Zpos w_digits);fold wB;fold wwB;zarith. + rewrite wwB_wBwB; rewrite Zpower_2; rewrite Zmult_mod_distr_r;zarith. + rewrite Zmult_plus_distr_l. + replace ([|xh|] * wB * 2 ^ u) with + ([|xh|]*2^u*wB). 2:ring. + repeat rewrite <- Zplus_assoc. + rewrite (Zplus_comm ([|xh|] * 2 ^ u * wB)). + rewrite Z_mod_plus;zarith. rewrite Z_mod_mult;zarith. + unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith. + unfold u; split;zarith. + split;zarith. unfold u; apply Zdiv_lt_upper_bound;zarith. + rewrite <- Zpower_exp;zarith. + fold u. + ring_simplify (u + (Zpos w_digits - u)); fold + wB;zarith. unfold ww_digits;rewrite Zpos_xO;zarith. + unfold base;rewrite <- Zmod_shift_r;zarith. fold base;apply Z_mod_lt;zarith. + unfold u; split;zarith. + unfold u; split;zarith. + apply Zdiv_lt_upper_bound;zarith. + rewrite <- Zpower_exp;zarith. + fold u. + ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith. + unfold u;zarith. + unfold u;zarith. + set (u := [[p]] - Zpos w_digits). + ring_simplify (u + (Zpos w_digits - u)); fold wB; auto with zarith. + Qed. + + Lemma spec_ww_add_mul_div : forall x y p, + [[p]] <= Zpos (xO w_digits) -> + [[ ww_add_mul_div p x y ]] = + ([[x]] * (2^[[p]]) + + [[y]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB. + Proof. + clear spec_ww_zdigits. + intros x y p H. + destruct x as [ |xh xl]; + [assert (H1 := @spec_ww_add_mul_div_aux w_0 w_0) + |assert (H1 := @spec_ww_add_mul_div_aux xh xl)]; + (destruct y as [ |yh yl]; + [generalize (H1 w_0 w_0 p H) | generalize (H1 yh yl p H)]; + clear H1;w_rewrite);simpl ww_add_mul_div. + replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial]. + intros Heq;rewrite <- Heq;clear Heq; auto. + generalize (spec_ww_compare p (w_0W w_zdigits)); + case ww_compare; intros H1; w_rewrite. + rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith. + generalize H1; w_rewrite; rewrite spec_zdigits; clear H1; intros H1. + assert (HH0: [|low p|] = [[p]]). + rewrite spec_low. + apply Zmod_small. + case (spec_to_w_Z p); intros HH1 HH2; split; auto. + apply Zlt_le_trans with (1 := H1). + unfold base; apply Zpower2_le_lin; auto with zarith. + rewrite HH0; auto with zarith. + replace [[WW w_0 w_0]] with 0;[w_rewrite|simpl;w_rewrite;trivial]. + intros Heq;rewrite <- Heq;clear Heq. + generalize (spec_ww_compare p (w_0W w_zdigits)); + case ww_compare; intros H1; w_rewrite. + rewrite (spec_w_add_mul_div w_0 w_0);w_rewrite;zarith. + rewrite Zpos_xO in H;zarith. + assert (HH: [|low (ww_sub p (w_0W w_zdigits)) |] = [[p]] - Zpos w_digits). + generalize H1; clear H1. + rewrite spec_low. + rewrite spec_ww_sub; w_rewrite; intros H1. + rewrite <- Zmod_div_mod; auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + unfold base; auto with zarith. + unfold base; auto with zarith. + exists wB; unfold base. + unfold ww_digits; rewrite (Zpos_xO w_digits). + rewrite <- Zpower_exp; auto with zarith. + apply f_equal with (f := fun x => 2 ^ x); auto with zarith. + case (spec_to_Z xh); auto with zarith. + Qed. + + End DoubleProof. + +End DoubleLift. + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v new file mode 100644 index 00000000..c7d83acc --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleMul.v @@ -0,0 +1,628 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleMul.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleMul. + Variable w : Type. + Variable w_0 : w. + Variable w_1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_compare : w -> w -> comparison. + Variable w_succ : w -> w. + Variable w_add_c : w -> w -> carry w. + Variable w_add : w -> w -> w. + Variable w_sub: w -> w -> w. + Variable w_mul_c : w -> w -> zn2z w. + Variable w_mul : w -> w -> w. + Variable w_square_c : w -> zn2z w. + Variable ww_add_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_add : zn2z w -> zn2z w -> zn2z w. + Variable ww_add_carry : zn2z w -> zn2z w -> zn2z w. + Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_sub : zn2z w -> zn2z w -> zn2z w. + + (* ** Multiplication ** *) + + (* (xh*B+xl) (yh*B + yl) + xh*yh = hh = |hhh|hhl|B2 + xh*yl +xl*yh = cc = |cch|ccl|B + xl*yl = ll = |llh|lll + *) + + Definition double_mul_c (cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w) x y := + match x, y with + | W0, _ => W0 + | _, W0 => W0 + | WW xh xl, WW yh yl => + let hh := w_mul_c xh yh in + let ll := w_mul_c xl yl in + let (wc,cc) := cross xh xl yh yl hh ll in + match cc with + | W0 => WW (ww_add hh (w_W0 wc)) ll + | WW cch ccl => + match ww_add_c (w_W0 ccl) ll with + | C0 l => WW (ww_add hh (w_WW wc cch)) l + | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l + end + end + end. + + Definition ww_mul_c := + double_mul_c + (fun xh xl yh yl hh ll=> + match ww_add_c (w_mul_c xh yl) (w_mul_c xl yh) with + | C0 cc => (w_0, cc) + | C1 cc => (w_1, cc) + end). + + Definition w_2 := w_add w_1 w_1. + + Definition kara_prod xh xl yh yl hh ll := + match ww_add_c hh ll with + C0 m => + match w_compare xl xh with + Eq => (w_0, m) + | Lt => + match w_compare yl yh with + Eq => (w_0, m) + | Lt => (w_0, ww_sub m (w_mul_c (w_sub xh xl) (w_sub yh yl))) + | Gt => match ww_add_c m (w_mul_c (w_sub xh xl) (w_sub yl yh)) with + C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1) + end + end + | Gt => + match w_compare yl yh with + Eq => (w_0, m) + | Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with + C1 m1 => (w_1, m1) | C0 m1 => (w_0, m1) + end + | Gt => (w_0, ww_sub m (w_mul_c (w_sub xl xh) (w_sub yl yh))) + end + end + | C1 m => + match w_compare xl xh with + Eq => (w_1, m) + | Lt => + match w_compare yl yh with + Eq => (w_1, m) + | Lt => match ww_sub_c m (w_mul_c (w_sub xh xl) (w_sub yh yl)) with + C0 m1 => (w_1, m1) | C1 m1 => (w_0, m1) + end + | Gt => match ww_add_c m (w_mul_c (w_sub xh xl) (w_sub yl yh)) with + C1 m1 => (w_2, m1) | C0 m1 => (w_1, m1) + end + end + | Gt => + match w_compare yl yh with + Eq => (w_1, m) + | Lt => match ww_add_c m (w_mul_c (w_sub xl xh) (w_sub yh yl)) with + C1 m1 => (w_2, m1) | C0 m1 => (w_1, m1) + end + | Gt => match ww_sub_c m (w_mul_c (w_sub xl xh) (w_sub yl yh)) with + C1 m1 => (w_0, m1) | C0 m1 => (w_1, m1) + end + end + end + end. + + Definition ww_karatsuba_c := double_mul_c kara_prod. + + Definition ww_mul x y := + match x, y with + | W0, _ => W0 + | _, W0 => W0 + | WW xh xl, WW yh yl => + let ccl := w_add (w_mul xh yl) (w_mul xl yh) in + ww_add (w_W0 ccl) (w_mul_c xl yl) + end. + + Definition ww_square_c x := + match x with + | W0 => W0 + | WW xh xl => + let hh := w_square_c xh in + let ll := w_square_c xl in + let xhxl := w_mul_c xh xl in + let (wc,cc) := + match ww_add_c xhxl xhxl with + | C0 cc => (w_0, cc) + | C1 cc => (w_1, cc) + end in + match cc with + | W0 => WW (ww_add hh (w_W0 wc)) ll + | WW cch ccl => + match ww_add_c (w_W0 ccl) ll with + | C0 l => WW (ww_add hh (w_WW wc cch)) l + | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l + end + end + end. + + Section DoubleMulAddn1. + Variable w_mul_add : w -> w -> w -> w * w. + + Fixpoint double_mul_add_n1 (n:nat) : word w n -> w -> w -> w * word w n := + match n return word w n -> w -> w -> w * word w n with + | O => w_mul_add + | S n1 => + let mul_add := double_mul_add_n1 n1 in + fun x y r => + match x with + | W0 => (w_0,extend w_0W n1 r) + | WW xh xl => + let (rl,l) := mul_add xl y r in + let (rh,h) := mul_add xh y rl in + (rh, double_WW w_WW n1 h l) + end + end. + + End DoubleMulAddn1. + + Section DoubleMulAddmn1. + Variable wn: Type. + Variable extend_n : w -> wn. + Variable wn_0W : wn -> zn2z wn. + Variable wn_WW : wn -> wn -> zn2z wn. + Variable w_mul_add_n1 : wn -> w -> w -> w*wn. + Fixpoint double_mul_add_mn1 (m:nat) : + word wn m -> w -> w -> w*word wn m := + match m return word wn m -> w -> w -> w*word wn m with + | O => w_mul_add_n1 + | S m1 => + let mul_add := double_mul_add_mn1 m1 in + fun x y r => + match x with + | W0 => (w_0,extend wn_0W m1 (extend_n r)) + | WW xh xl => + let (rl,l) := mul_add xl y r in + let (rh,h) := mul_add xh y rl in + (rh, double_WW wn_WW m1 h l) + end + end. + + End DoubleMulAddmn1. + + Definition w_mul_add x y r := + match w_mul_c x y with + | W0 => (w_0, r) + | WW h l => + match w_add_c l r with + | C0 lr => (h,lr) + | C1 lr => (w_succ h, lr) + end + end. + + + (*Section DoubleProof. *) + Variable w_digits : positive. + Variable w_to_Z : w -> Z. + + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + + Notation "[|| x ||]" := + (zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99). + + Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x) + (at level 0, x at level 99). + + Variable spec_more_than_1_digit: 1 < Zpos w_digits. + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_w_compare : + forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_w_succ : forall x, [|w_succ x|] = ([|x|] + 1) mod wB. + Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. + Variable spec_w_add : forall x y, [|w_add x y|] = ([|x|] + [|y|]) mod wB. + Variable spec_w_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + + Variable spec_w_mul_c : forall x y, [[ w_mul_c x y ]] = [|x|] * [|y|]. + Variable spec_w_mul : forall x y, [|w_mul x y|] = ([|x|] * [|y|]) mod wB. + Variable spec_w_square_c : forall x, [[ w_square_c x]] = [|x|] * [|x|]. + + Variable spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]]. + Variable spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB. + Variable spec_ww_add_carry : + forall x y, [[ww_add_carry x y]] = ([[x]] + [[y]] + 1) mod wwB. + Variable spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. + Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. + + + Lemma spec_ww_to_Z : forall x, 0 <= [[x]] < wwB. + Proof. intros x;apply spec_ww_to_Z;auto. Qed. + + Lemma spec_ww_to_Z_wBwB : forall x, 0 <= [[x]] < wB^2. + Proof. rewrite <- wwB_wBwB;apply spec_ww_to_Z. Qed. + + Hint Resolve spec_ww_to_Z spec_ww_to_Z_wBwB : mult. + Ltac zarith := auto with zarith mult. + + Lemma wBwB_lex: forall a b c d, + a * wB^2 + [[b]] <= c * wB^2 + [[d]] -> + a <= c. + Proof. + intros a b c d H; apply beta_lex with [[b]] [[d]] (wB^2);zarith. + Qed. + + Lemma wBwB_lex_inv: forall a b c d, + a < c -> + a * wB^2 + [[b]] < c * wB^2 + [[d]]. + Proof. + intros a b c d H; apply beta_lex_inv; zarith. + Qed. + + Lemma sum_mul_carry : forall xh xl yh yl wc cc, + [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] -> + 0 <= [|wc|] <= 1. + Proof. + intros. + apply (sum_mul_carry [|xh|] [|xl|] [|yh|] [|yl|] [|wc|][[cc]] wB);zarith. + apply wB_pos. + Qed. + + Theorem mult_add_ineq: forall xH yH crossH, + 0 <= [|xH|] * [|yH|] + [|crossH|] < wwB. + Proof. + intros;rewrite wwB_wBwB;apply mult_add_ineq;zarith. + Qed. + + Hint Resolve mult_add_ineq : mult. + + Lemma spec_mul_aux : forall xh xl yh yl wc (cc:zn2z w) hh ll, + [[hh]] = [|xh|] * [|yh|] -> + [[ll]] = [|xl|] * [|yl|] -> + [|wc|]*wB^2 + [[cc]] = [|xh|] * [|yl|] + [|xl|] * [|yh|] -> + [||match cc with + | W0 => WW (ww_add hh (w_W0 wc)) ll + | WW cch ccl => + match ww_add_c (w_W0 ccl) ll with + | C0 l => WW (ww_add hh (w_WW wc cch)) l + | C1 l => WW (ww_add_carry hh (w_WW wc cch)) l + end + end||] = ([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|]). + Proof. + intros;assert (U1 := wB_pos w_digits). + replace (([|xh|] * wB + [|xl|]) * ([|yh|] * wB + [|yl|])) with + ([|xh|]*[|yh|]*wB^2 + ([|xh|]*[|yl|] + [|xl|]*[|yh|])*wB + [|xl|]*[|yl|]). + 2:ring. rewrite <- H1;rewrite <- H;rewrite <- H0. + assert (H2 := sum_mul_carry _ _ _ _ _ _ H1). + destruct cc as [ | cch ccl]; simpl zn2z_to_Z; simpl ww_to_Z. + rewrite spec_ww_add;rewrite spec_w_W0;rewrite Zmod_small; + rewrite wwB_wBwB. ring. + rewrite <- (Zplus_0_r ([|wc|]*wB));rewrite H;apply mult_add_ineq3;zarith. + simpl ww_to_Z in H1. assert (U:=spec_to_Z cch). + assert ([|wc|]*wB + [|cch|] <= 2*wB - 3). + destruct (Z_le_gt_dec ([|wc|]*wB + [|cch|]) (2*wB - 3));trivial. + assert ([|xh|] * [|yl|] + [|xl|] * [|yh|] <= (2*wB - 4)*wB + 2). + ring_simplify ((2*wB - 4)*wB + 2). + assert (H4 := Zmult_lt_b _ _ _ (spec_to_Z xh) (spec_to_Z yl)). + assert (H5 := Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)). + omega. + generalize H3;clear H3;rewrite <- H1. + rewrite Zplus_assoc; rewrite Zpower_2; rewrite Zmult_assoc; + rewrite <- Zmult_plus_distr_l. + assert (((2 * wB - 4) + 2)*wB <= ([|wc|] * wB + [|cch|])*wB). + apply Zmult_le_compat;zarith. + rewrite Zmult_plus_distr_l in H3. + intros. assert (U2 := spec_to_Z ccl);omega. + generalize (spec_ww_add_c (w_W0 ccl) ll);destruct (ww_add_c (w_W0 ccl) ll) + as [l|l];unfold interp_carry;rewrite spec_w_W0;try rewrite Zmult_1_l; + simpl zn2z_to_Z; + try rewrite spec_ww_add;try rewrite spec_ww_add_carry;rewrite spec_w_WW; + rewrite Zmod_small;rewrite wwB_wBwB;intros. + rewrite H4;ring. rewrite H;apply mult_add_ineq2;zarith. + rewrite Zplus_assoc;rewrite Zmult_plus_distr_l. + rewrite Zmult_1_l;rewrite <- Zplus_assoc;rewrite H4;ring. + repeat rewrite <- Zplus_assoc;rewrite H;apply mult_add_ineq2;zarith. + Qed. + + Lemma spec_double_mul_c : forall cross:w->w->w->w->zn2z w -> zn2z w -> w*zn2z w, + (forall xh xl yh yl hh ll, + [[hh]] = [|xh|]*[|yh|] -> + [[ll]] = [|xl|]*[|yl|] -> + let (wc,cc) := cross xh xl yh yl hh ll in + [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]) -> + forall x y, [||double_mul_c cross x y||] = [[x]] * [[y]]. + Proof. + intros cross Hcross x y;destruct x as [ |xh xl];simpl;trivial. + destruct y as [ |yh yl];simpl. rewrite Zmult_0_r;trivial. + assert (H1:= spec_w_mul_c xh yh);assert (H2:= spec_w_mul_c xl yl). + generalize (Hcross _ _ _ _ _ _ H1 H2). + destruct (cross xh xl yh yl (w_mul_c xh yh) (w_mul_c xl yl)) as (wc,cc). + intros;apply spec_mul_aux;trivial. + rewrite <- wwB_wBwB;trivial. + Qed. + + Lemma spec_ww_mul_c : forall x y, [||ww_mul_c x y||] = [[x]] * [[y]]. + Proof. + intros x y;unfold ww_mul_c;apply spec_double_mul_c. + intros xh xl yh yl hh ll H1 H2. + generalize (spec_ww_add_c (w_mul_c xh yl) (w_mul_c xl yh)); + destruct (ww_add_c (w_mul_c xh yl) (w_mul_c xl yh)) as [c|c]; + unfold interp_carry;repeat rewrite spec_w_mul_c;intros H; + (rewrite spec_w_0 || rewrite spec_w_1);rewrite H;ring. + Qed. + + Lemma spec_w_2: [|w_2|] = 2. + unfold w_2; rewrite spec_w_add; rewrite spec_w_1; simpl. + apply Zmod_small; split; auto with zarith. + rewrite <- (Zpower_1_r 2); unfold base; apply Zpower_lt_monotone; auto with zarith. + Qed. + + Lemma kara_prod_aux : forall xh xl yh yl, + xh*yh + xl*yl - (xh-xl)*(yh-yl) = xh*yl + xl*yh. + Proof. intros;ring. Qed. + + Lemma spec_kara_prod : forall xh xl yh yl hh ll, + [[hh]] = [|xh|]*[|yh|] -> + [[ll]] = [|xl|]*[|yl|] -> + let (wc,cc) := kara_prod xh xl yh yl hh ll in + [|wc|]*wwB + [[cc]] = [|xh|]*[|yl|] + [|xl|]*[|yh|]. + Proof. + intros xh xl yh yl hh ll H H0; rewrite <- kara_prod_aux; + rewrite <- H; rewrite <- H0; unfold kara_prod. + assert (Hxh := (spec_to_Z xh)); assert (Hxl := (spec_to_Z xl)); + assert (Hyh := (spec_to_Z yh)); assert (Hyl := (spec_to_Z yl)). + generalize (spec_ww_add_c hh ll); case (ww_add_c hh ll); + intros z Hz; rewrite <- Hz; unfold interp_carry; assert (Hz1 := (spec_ww_to_Z z)). + generalize (spec_w_compare xl xh); case (w_compare xl xh); intros Hxlh; + try rewrite Hxlh; try rewrite spec_w_0; try (ring; fail). + generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh. + rewrite Hylh; rewrite spec_w_0; try (ring; fail). + rewrite spec_w_0; try (ring; fail). + repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + split; auto with zarith. + simpl in Hz; rewrite Hz; rewrite H; rewrite H0. + rewrite kara_prod_aux; apply Zplus_le_0_compat; apply Zmult_le_0_compat; auto with zarith. + apply Zle_lt_trans with ([[z]]-0); auto with zarith. + unfold Zminus; apply Zplus_le_compat_l; apply Zle_left_rev; simpl; rewrite Zopp_involutive. + apply Zmult_le_0_compat; auto with zarith. + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2; + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh. + rewrite Hylh; rewrite spec_w_0; try (ring; fail). + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_0; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_1; unfold interp_carry in Hz2; rewrite Hz2; + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_0; try (ring; fail). + repeat (rewrite spec_ww_sub || rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + split. + match goal with |- context[(?x - ?y) * (?z - ?t)] => + replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring] + end. + simpl in Hz; rewrite Hz; rewrite H; rewrite H0. + rewrite kara_prod_aux; apply Zplus_le_0_compat; apply Zmult_le_0_compat; auto with zarith. + apply Zle_lt_trans with ([[z]]-0); auto with zarith. + unfold Zminus; apply Zplus_le_compat_l; apply Zle_left_rev; simpl; rewrite Zopp_involutive. + apply Zmult_le_0_compat; auto with zarith. + (** there is a carry in hh + ll **) + rewrite Zmult_1_l. + generalize (spec_w_compare xl xh); case (w_compare xl xh); intros Hxlh; + try rewrite Hxlh; try rewrite spec_w_1; try (ring; fail). + generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh; + try rewrite Hylh; try rewrite spec_w_1; try (ring; fail). + match goal with |- context[ww_sub_c ?x ?y] => + generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l. + generalize Hz2; clear Hz2; unfold interp_carry. + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_2; unfold interp_carry in Hz2. + apply trans_equal with (wwB + (1 * wwB + [[z1]])). + ring. + rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + generalize (spec_w_compare yl yh); case (w_compare yl yh); intros Hylh; + try rewrite Hylh; try rewrite spec_w_1; try (ring; fail). + match goal with |- context[ww_add_c ?x ?y] => + generalize (spec_ww_add_c x y); case (ww_add_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_2; unfold interp_carry in Hz2. + apply trans_equal with (wwB + (1 * wwB + [[z1]])). + ring. + rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + match goal with |- context[ww_sub_c ?x ?y] => + generalize (spec_ww_sub_c x y); case (ww_sub_c x y); try rewrite spec_w_1; + intros z1 Hz2 + end. + simpl in Hz2; rewrite Hz2; repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + rewrite spec_w_0; rewrite Zmult_0_l; rewrite Zplus_0_l. + match goal with |- context[(?x - ?y) * (?z - ?t)] => + replace ((x - y) * (z - t)) with ((y - x) * (t - z)); [idtac | ring] + end. + generalize Hz2; clear Hz2; unfold interp_carry. + repeat (rewrite spec_w_sub || rewrite spec_w_mul_c). + repeat rewrite Zmod_small; auto with zarith; try (ring; fail). + Qed. + + Lemma sub_carry : forall xh xl yh yl z, + 0 <= z -> + [|xh|]*[|yl|] + [|xl|]*[|yh|] = wwB + z -> + z < wwB. + Proof. + intros xh xl yh yl z Hle Heq. + destruct (Z_le_gt_dec wwB z);auto with zarith. + generalize (Zmult_lt_b _ _ _ (spec_to_Z xh) (spec_to_Z yl)). + generalize (Zmult_lt_b _ _ _ (spec_to_Z xl) (spec_to_Z yh)). + rewrite <- wwB_wBwB;intros H1 H2. + assert (H3 := wB_pos w_digits). + assert (2*wB <= wwB). + rewrite wwB_wBwB; rewrite Zpower_2; apply Zmult_le_compat;zarith. + omega. + Qed. + + Ltac Spec_ww_to_Z x := + let H:= fresh "H" in + assert (H:= spec_ww_to_Z x). + + Ltac Zmult_lt_b x y := + let H := fresh "H" in + assert (H := Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)). + + Lemma spec_ww_karatsuba_c : forall x y, [||ww_karatsuba_c x y||]=[[x]]*[[y]]. + Proof. + intros x y; unfold ww_karatsuba_c;apply spec_double_mul_c. + intros; apply spec_kara_prod; auto. + Qed. + + Lemma spec_ww_mul : forall x y, [[ww_mul x y]] = [[x]]*[[y]] mod wwB. + Proof. + assert (U:= lt_0_wB w_digits). + assert (U1:= lt_0_wwB w_digits). + intros x y; case x; auto; intros xh xl. + case y; auto. + simpl; rewrite Zmult_0_r; rewrite Zmod_small; auto with zarith. + intros yh yl;simpl. + repeat (rewrite spec_ww_add || rewrite spec_w_W0 || rewrite spec_w_mul_c + || rewrite spec_w_add || rewrite spec_w_mul). + rewrite <- Zplus_mod; auto with zarith. + repeat (rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r). + rewrite <- Zmult_mod_distr_r; auto with zarith. + rewrite <- Zpower_2; rewrite <- wwB_wBwB; auto with zarith. + rewrite Zplus_mod; auto with zarith. + rewrite Zmod_mod; auto with zarith. + rewrite <- Zplus_mod; auto with zarith. + match goal with |- ?X mod _ = _ => + rewrite <- Z_mod_plus with (a := X) (b := [|xh|] * [|yh|]) + end; auto with zarith. + f_equal; auto; rewrite wwB_wBwB; ring. + Qed. + + Lemma spec_ww_square_c : forall x, [||ww_square_c x||] = [[x]]*[[x]]. + Proof. + destruct x as [ |xh xl];simpl;trivial. + case_eq match ww_add_c (w_mul_c xh xl) (w_mul_c xh xl) with + | C0 cc => (w_0, cc) + | C1 cc => (w_1, cc) + end;intros wc cc Heq. + apply (spec_mul_aux xh xl xh xl wc cc);trivial. + generalize Heq (spec_ww_add_c (w_mul_c xh xl) (w_mul_c xh xl));clear Heq. + rewrite spec_w_mul_c;destruct (ww_add_c (w_mul_c xh xl) (w_mul_c xh xl)); + unfold interp_carry;try rewrite Zmult_1_l;intros Heq Heq';inversion Heq; + rewrite (Zmult_comm [|xl|]);subst. + rewrite spec_w_0;rewrite Zmult_0_l;rewrite Zplus_0_l;trivial. + rewrite spec_w_1;rewrite Zmult_1_l;rewrite <- wwB_wBwB;trivial. + Qed. + + Section DoubleMulAddn1Proof. + + Variable w_mul_add : w -> w -> w -> w * w. + Variable spec_w_mul_add : forall x y r, + let (h,l):= w_mul_add x y r in + [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|]. + + Lemma spec_double_mul_add_n1 : forall n x y r, + let (h,l) := double_mul_add_n1 w_mul_add n x y r in + [|h|]*double_wB w_digits n + [!n|l!] = [!n|x!]*[|y|]+[|r|]. + Proof. + induction n;intros x y r;trivial. + exact (spec_w_mul_add x y r). + unfold double_mul_add_n1;destruct x as[ |xh xl]; + fold(double_mul_add_n1 w_mul_add). + rewrite spec_w_0;rewrite spec_extend;simpl;trivial. + assert(H:=IHn xl y r);destruct (double_mul_add_n1 w_mul_add n xl y r)as(rl,l). + assert(U:=IHn xh y rl);destruct(double_mul_add_n1 w_mul_add n xh y rl)as(rh,h). + rewrite <- double_wB_wwB. rewrite spec_double_WW;simpl;trivial. + rewrite Zmult_plus_distr_l;rewrite <- Zplus_assoc;rewrite <- H. + rewrite Zmult_assoc;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite U;ring. + Qed. + + End DoubleMulAddn1Proof. + + Lemma spec_w_mul_add : forall x y r, + let (h,l):= w_mul_add x y r in + [|h|]*wB+[|l|] = [|x|]*[|y|] + [|r|]. + Proof. + intros x y r;unfold w_mul_add;assert (H:=spec_w_mul_c x y); + destruct (w_mul_c x y) as [ |h l];simpl;rewrite <- H. + rewrite spec_w_0;trivial. + assert (U:=spec_w_add_c l r);destruct (w_add_c l r) as [lr|lr];unfold + interp_carry in U;try rewrite Zmult_1_l in H;simpl. + rewrite U;ring. rewrite spec_w_succ. rewrite Zmod_small. + rewrite <- Zplus_assoc;rewrite <- U;ring. + simpl in H;assert (H1:= Zmult_lt_b _ _ _ (spec_to_Z x) (spec_to_Z y)). + rewrite <- H in H1. + assert (H2:=spec_to_Z h);split;zarith. + case H1;clear H1;intro H1;clear H1. + replace (wB ^ 2 - 2 * wB) with ((wB - 2)*wB). 2:ring. + intros H0;assert (U1:= wB_pos w_digits). + assert (H1 := beta_lex _ _ _ _ _ H0 (spec_to_Z l));zarith. + Qed. + +(* End DoubleProof. *) + +End DoubleMul. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v new file mode 100644 index 00000000..043ff351 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSqrt.v @@ -0,0 +1,1389 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleSqrt.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleSqrt. + Variable w : Type. + Variable w_is_even : w -> bool. + Variable w_compare : w -> w -> comparison. + Variable w_0 : w. + Variable w_1 : w. + Variable w_Bm1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable w_W0 : w -> zn2z w. + Variable w_0W : w -> zn2z w. + Variable w_sub : w -> w -> w. + Variable w_sub_c : w -> w -> carry w. + Variable w_square_c : w -> zn2z w. + Variable w_div21 : w -> w -> w -> w * w. + Variable w_add_mul_div : w -> w -> w -> w. + Variable w_digits : positive. + Variable w_zdigits : w. + Variable ww_zdigits : zn2z w. + Variable w_add_c : w -> w -> carry w. + Variable w_sqrt2 : w -> w -> w * carry w. + Variable w_pred : w -> w. + Variable ww_pred_c : zn2z w -> carry (zn2z w). + Variable ww_pred : zn2z w -> zn2z w. + Variable ww_add_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_add : zn2z w -> zn2z w -> zn2z w. + Variable ww_sub_c : zn2z w -> zn2z w -> carry (zn2z w). + Variable ww_add_mul_div : zn2z w -> zn2z w -> zn2z w -> zn2z w. + Variable ww_head0 : zn2z w -> zn2z w. + Variable ww_compare : zn2z w -> zn2z w -> comparison. + Variable low : zn2z w -> w. + + Let wwBm1 := ww_Bm1 w_Bm1. + + Definition ww_is_even x := + match x with + | W0 => true + | WW xh xl => w_is_even xl + end. + + Let w_div21c x y z := + match w_compare x z with + | Eq => + match w_compare y z with + Eq => (C1 w_1, w_0) + | Gt => (C1 w_1, w_sub y z) + | Lt => (C1 w_0, y) + end + | Gt => + let x1 := w_sub x z in + let (q, r) := w_div21 x1 y z in + (C1 q, r) + | Lt => + let (q, r) := w_div21 x y z in + (C0 q, r) + end. + + Let w_div2s x y s := + match x with + C1 x1 => + let x2 := w_sub x1 s in + let (q, r) := w_div21c x2 y s in + match q with + C0 q1 => + if w_is_even q1 then + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), C0 r) + else + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), w_add_c r s) + | C1 q1 => + if w_is_even q1 then + (C1 (w_add_mul_div (w_pred w_zdigits) w_0 q1), C0 r) + else + (C1 (w_add_mul_div (w_pred w_zdigits) w_0 q1), w_add_c r s) + end + | C0 x1 => + let (q, r) := w_div21c x1 y s in + match q with + C0 q1 => + if w_is_even q1 then + (C0 (w_add_mul_div (w_pred w_zdigits) w_0 q1), C0 r) + else + (C0 (w_add_mul_div (w_pred w_zdigits) w_0 q1), w_add_c r s) + | C1 q1 => + if w_is_even q1 then + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), C0 r) + else + (C0 (w_add_mul_div (w_pred w_zdigits) w_1 q1), w_add_c r s) + end + end. + + Definition split x := + match x with + | W0 => (w_0,w_0) + | WW h l => (h,l) + end. + + Definition ww_sqrt2 x y := + let (x1, x2) := split x in + let (y1, y2) := split y in + let ( q, r) := w_sqrt2 x1 x2 in + let (q1, r1) := w_div2s r y1 q in + match q1 with + C0 q1 => + let q2 := w_square_c q1 in + let a := WW q q1 in + match r1 with + C1 r2 => + match ww_sub_c (WW r2 y2) q2 with + C0 r3 => (a, C1 r3) + | C1 r3 => (a, C0 r3) + end + | C0 r2 => + match ww_sub_c (WW r2 y2) q2 with + C0 r3 => (a, C0 r3) + | C1 r3 => + let a2 := ww_add_mul_div (w_0W w_1) a W0 in + match ww_pred_c a2 with + C0 a3 => + (ww_pred a, ww_add_c a3 r3) + | C1 a3 => + (ww_pred a, C0 (ww_add a3 r3)) + end + end + end + | C1 q1 => + let a1 := WW q w_Bm1 in + let a2 := ww_add_mul_div (w_0W w_1) a1 wwBm1 in + (a1, ww_add_c a2 y) + end. + + Definition ww_is_zero x := + match ww_compare W0 x with + Eq => true + | _ => false + end. + + Definition ww_head1 x := + let p := ww_head0 x in + if (ww_is_even p) then p else ww_pred p. + + Definition ww_sqrt x := + if (ww_is_zero x) then W0 + else + let p := ww_head1 x in + match ww_compare p W0 with + | Gt => + match ww_add_mul_div p x W0 with + W0 => W0 + | WW x1 x2 => + let (r, _) := w_sqrt2 x1 x2 in + WW w_0 (w_add_mul_div + (w_sub w_zdigits + (low (ww_add_mul_div (ww_pred ww_zdigits) + W0 p))) w_0 r) + end + | _ => + match x with + W0 => W0 + | WW x1 x2 => WW w_0 (fst (w_sqrt2 x1 x2)) + end + end. + + + Variable w_to_Z : w -> Z. + + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + + Notation "[|| x ||]" := + (zn2z_to_Z wwB (ww_to_Z w_digits w_to_Z) x) (at level 0, x at level 99). + + Notation "[! n | x !]" := (double_to_Z w_digits w_to_Z n x) + (at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_1 : [|w_1|] = 1. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_w_zdigits : [|w_zdigits|] = Zpos w_digits. + Variable spec_more_than_1_digit: 1 < Zpos w_digits. + + Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos (xO w_digits). + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_to_w_Z : forall x, 0 <= [[x]] < wwB. + + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + Variable spec_w_W0 : forall h, [[w_W0 h]] = [|h|] * wB. + Variable spec_w_0W : forall l, [[w_0W l]] = [|l|]. + Variable spec_w_is_even : forall x, + if w_is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1. + Variable spec_w_compare : forall x y, + match w_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Variable spec_w_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + Variable spec_w_square_c : forall x, [[ w_square_c x]] = [|x|] * [|x|]. + Variable spec_w_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := w_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Variable spec_w_add_mul_div : forall x y p, + [|p|] <= Zpos w_digits -> + [| w_add_mul_div p x y |] = + ([|x|] * (2 ^ [|p|]) + + [|y|] / (Zpower 2 ((Zpos w_digits) - [|p|]))) mod wB. + Variable spec_ww_add_mul_div : forall x y p, + [[p]] <= Zpos (xO w_digits) -> + [[ ww_add_mul_div p x y ]] = + ([[x]] * (2^ [[p]]) + + [[y]] / (2^ (Zpos (xO w_digits) - [[p]]))) mod wwB. + Variable spec_w_add_c : forall x y, [+|w_add_c x y|] = [|x|] + [|y|]. + Variable spec_ww_add : forall x y, [[ww_add x y]] = ([[x]] + [[y]]) mod wwB. + Variable spec_w_sqrt2 : forall x y, + wB/ 4 <= [|x|] -> + let (s,r) := w_sqrt2 x y in + [[WW x y]] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]. + Variable spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. + Variable spec_ww_pred_c : forall x, [-[ww_pred_c x]] = [[x]] - 1. + Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB. + Variable spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB. + Variable spec_ww_add_c : forall x y, [+[ww_add_c x y]] = [[x]] + [[y]]. + Variable spec_ww_compare : forall x y, + match ww_compare x y with + | Eq => [[x]] = [[y]] + | Lt => [[x]] < [[y]] + | Gt => [[x]] > [[y]] + end. + Variable spec_ww_head0 : forall x, 0 < [[x]] -> + wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB. + Variable spec_low: forall x, [|low x|] = [[x]] mod wB. + + Let spec_ww_Bm1 : [[wwBm1]] = wwB - 1. + Proof. refine (spec_ww_Bm1 w_Bm1 w_digits w_to_Z _);auto. Qed. + + + Hint Rewrite spec_w_0 spec_w_1 w_Bm1 spec_w_WW spec_w_sub + spec_w_div21 spec_w_add_mul_div spec_ww_Bm1 + spec_w_add_c spec_w_sqrt2: w_rewrite. + + Lemma spec_ww_is_even : forall x, + if ww_is_even x then [[x]] mod 2 = 0 else [[x]] mod 2 = 1. +clear spec_more_than_1_digit. +intros x; case x; simpl ww_is_even. + simpl. + rewrite Zmod_small; auto with zarith. + intros w1 w2; simpl. + unfold base. + rewrite Zplus_mod; auto with zarith. + rewrite (fun x y => (Zdivide_mod (x * y))); auto with zarith. + rewrite Zplus_0_l; rewrite Zmod_mod; auto with zarith. + apply spec_w_is_even; auto with zarith. + apply Zdivide_mult_r; apply Zpower_divide; auto with zarith. + red; simpl; auto. + Qed. + + + Theorem spec_w_div21c : forall a1 a2 b, + wB/2 <= [|b|] -> + let (q,r) := w_div21c a1 a2 b in + [|a1|] * wB + [|a2|] = [+|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]. + intros a1 a2 b Hb; unfold w_div21c. + assert (H: 0 < [|b|]); auto with zarith. + assert (U := wB_pos w_digits). + apply Zlt_le_trans with (2 := Hb); auto with zarith. + apply Zlt_le_trans with 1; auto with zarith. + apply Zdiv_le_lower_bound; auto with zarith. + repeat match goal with |- context[w_compare ?y ?z] => + generalize (spec_w_compare y z); + case (w_compare y z) + end. + intros H1 H2; split. + unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith. + rewrite H1; rewrite H2; ring. + autorewrite with w_rewrite; auto with zarith. + intros H1 H2; split. + unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith. + rewrite H2; ring. + destruct (spec_to_Z a2);auto with zarith. + intros H1 H2; split. + unfold interp_carry; autorewrite with w_rewrite rm10; auto with zarith. + rewrite H2; rewrite Zmod_small; auto with zarith. + ring. + destruct (spec_to_Z a2);auto with zarith. + rewrite spec_w_sub; auto with zarith. + destruct (spec_to_Z a2) as [H3 H4];auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + assert ([|a2|] < 2 * [|b|]); auto with zarith. + apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith. + rewrite wB_div_2; auto. + intros H1. + match goal with |- context[w_div21 ?y ?z ?t] => + generalize (@spec_w_div21 y z t Hb H1); + case (w_div21 y z t); simpl; autorewrite with w_rewrite; + auto + end. + intros H1. + assert (H2: [|w_sub a1 b|] < [|b|]). + rewrite spec_w_sub; auto with zarith. + rewrite Zmod_small; auto with zarith. + assert ([|a1|] < 2 * [|b|]); auto with zarith. + apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith. + rewrite wB_div_2; auto. + destruct (spec_to_Z a1);auto with zarith. + destruct (spec_to_Z a1);auto with zarith. + match goal with |- context[w_div21 ?y ?z ?t] => + generalize (@spec_w_div21 y z t Hb H2); + case (w_div21 y z t); autorewrite with w_rewrite; + auto + end. + intros w0 w1; replace [+|C1 w0|] with (wB + [|w0|]). + rewrite Zmod_small; auto with zarith. + intros (H3, H4); split; auto. + rewrite Zmult_plus_distr_l. + rewrite <- Zplus_assoc; rewrite <- H3; ring. + split; auto with zarith. + assert ([|a1|] < 2 * [|b|]); auto with zarith. + apply Zlt_le_trans with (2 * (wB / 2)); auto with zarith. + rewrite wB_div_2; auto. + destruct (spec_to_Z a1);auto with zarith. + destruct (spec_to_Z a1);auto with zarith. + simpl; case wB; auto. + Qed. + + Theorem C0_id: forall p, [+|C0 p|] = [|p|]. + intros p; simpl; auto. + Qed. + + Theorem add_mult_div_2: forall w, + [|w_add_mul_div (w_pred w_zdigits) w_0 w|] = [|w|] / 2. + intros w1. + assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1). + rewrite spec_pred; rewrite spec_w_zdigits. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. + rewrite spec_w_add_mul_div; auto with zarith. + autorewrite with w_rewrite rm10. + match goal with |- context[?X - ?Y] => + replace (X - Y) with 1 + end. + rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith. + destruct (spec_to_Z w1) as [H1 H2];auto with zarith. + split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + rewrite Hp; ring. + Qed. + + Theorem add_mult_div_2_plus_1: forall w, + [|w_add_mul_div (w_pred w_zdigits) w_1 w|] = + [|w|] / 2 + 2 ^ Zpos (w_digits - 1). + intros w1. + assert (Hp: [|w_pred w_zdigits|] = Zpos w_digits - 1). + rewrite spec_pred; rewrite spec_w_zdigits. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. + autorewrite with w_rewrite rm10; auto with zarith. + match goal with |- context[?X - ?Y] => + replace (X - Y) with 1 + end; rewrite Hp; try ring. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + rewrite Zpower_1_r; rewrite Zmod_small; auto with zarith. + destruct (spec_to_Z w1) as [H1 H2];auto with zarith. + split; auto with zarith. + unfold base. + match goal with |- _ < _ ^ ?X => + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite <- (tmp X); clear tmp + end. + rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith. + assert (tmp: forall p, 1 + (p -1) - 1 = p - 1); auto with zarith; + rewrite tmp; clear tmp; auto with zarith. + match goal with |- ?X + ?Y < _ => + assert (Y < X); auto with zarith + end. + apply Zdiv_lt_upper_bound; auto with zarith. + pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp; + auto with zarith. + assert (tmp: forall p, (p - 1) + 1 = p); auto with zarith; + rewrite tmp; clear tmp; auto with zarith. + Qed. + + Theorem add_mult_mult_2: forall w, + [|w_add_mul_div w_1 w w_0|] = 2 * [|w|] mod wB. + intros w1. + autorewrite with w_rewrite rm10; auto with zarith. + rewrite Zpower_1_r; auto with zarith. + rewrite Zmult_comm; auto. + Qed. + + Theorem ww_add_mult_mult_2: forall w, + [[ww_add_mul_div (w_0W w_1) w W0]] = 2 * [[w]] mod wwB. + intros w1. + rewrite spec_ww_add_mul_div; auto with zarith. + autorewrite with w_rewrite rm10. + rewrite spec_w_0W; rewrite spec_w_1. + rewrite Zpower_1_r; auto with zarith. + rewrite Zmult_comm; auto. + rewrite spec_w_0W; rewrite spec_w_1; auto with zarith. + red; simpl; intros; discriminate. + Qed. + + Theorem ww_add_mult_mult_2_plus_1: forall w, + [[ww_add_mul_div (w_0W w_1) w wwBm1]] = + (2 * [[w]] + 1) mod wwB. + intros w1. + rewrite spec_ww_add_mul_div; auto with zarith. + rewrite spec_w_0W; rewrite spec_w_1; auto with zarith. + rewrite Zpower_1_r; auto with zarith. + f_equal; auto. + rewrite Zmult_comm; f_equal; auto. + autorewrite with w_rewrite rm10. + unfold ww_digits, base. + apply sym_equal; apply Zdiv_unique with (r := 2 ^ (Zpos (ww_digits w_digits) - 1) -1); + auto with zarith. + unfold ww_digits; split; auto with zarith. + match goal with |- 0 <= ?X - 1 => + assert (0 < X); auto with zarith + end. + apply Zpower_gt_0; auto with zarith. + match goal with |- 0 <= ?X - 1 => + assert (0 < X); auto with zarith; red; reflexivity + end. + unfold ww_digits; autorewrite with rm10. + assert (tmp: forall p q r, p + (q - r) = p + q - r); auto with zarith; + rewrite tmp; clear tmp. + assert (tmp: forall p, p + p = 2 * p); auto with zarith; + rewrite tmp; clear tmp. + f_equal; auto. + pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp; + auto with zarith. + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite tmp; clear tmp; auto. + match goal with |- ?X - 1 >= 0 => + assert (0 < X); auto with zarith; red; reflexivity + end. + rewrite spec_w_0W; rewrite spec_w_1; auto with zarith. + red; simpl; intros; discriminate. + Qed. + + Theorem Zplus_mod_one: forall a1 b1, 0 < b1 -> (a1 + b1) mod b1 = a1 mod b1. + intros a1 b1 H; rewrite Zplus_mod; auto with zarith. + rewrite Z_mod_same; try rewrite Zplus_0_r; auto with zarith. + apply Zmod_mod; auto. + Qed. + + Lemma C1_plus_wB: forall x, [+|C1 x|] = wB + [|x|]. + unfold interp_carry; auto with zarith. + Qed. + + Theorem spec_w_div2s : forall a1 a2 b, + wB/2 <= [|b|] -> [+|a1|] <= 2 * [|b|] -> + let (q,r) := w_div2s a1 a2 b in + [+|a1|] * wB + [|a2|] = [+|q|] * (2 * [|b|]) + [+|r|] /\ 0 <= [+|r|] < 2 * [|b|]. + intros a1 a2 b H. + assert (HH: 0 < [|b|]); auto with zarith. + assert (U := wB_pos w_digits). + apply Zlt_le_trans with (2 := H); auto with zarith. + apply Zlt_le_trans with 1; auto with zarith. + apply Zdiv_le_lower_bound; auto with zarith. + unfold w_div2s; case a1; intros w0 H0. + match goal with |- context[w_div21c ?y ?z ?t] => + generalize (@spec_w_div21c y z t H); + case (w_div21c y z t); autorewrite with w_rewrite; + auto + end. + intros c w1; case c. + simpl interp_carry; intros w2 (Hw1, Hw2). + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat rewrite C0_id. + rewrite add_mult_div_2. + intros H1; split; auto with zarith. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + repeat rewrite C0_id. + rewrite add_mult_div_2. + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + intros w2; rewrite C1_plus_wB. + intros (Hw1, Hw2). + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat rewrite C0_id. + intros H1; split; auto with zarith. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1. + repeat rewrite C0_id. + rewrite add_mult_div_2_plus_1; unfold base. + match goal with |- context[_ ^ ?X] => + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + repeat rewrite C0_id. + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + rewrite add_mult_div_2_plus_1. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1. + unfold base. + match goal with |- context[_ ^ ?X] => + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + repeat rewrite C1_plus_wB in H0. + rewrite C1_plus_wB. + match goal with |- context[w_div21c ?y ?z ?t] => + generalize (@spec_w_div21c y z t H); + case (w_div21c y z t); autorewrite with w_rewrite; + auto + end. + intros c w1; case c. + intros w2 (Hw1, Hw2); rewrite C0_id in Hw1. + rewrite <- Zplus_mod_one in Hw1; auto with zarith. + rewrite Zmod_small in Hw1; auto with zarith. + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat rewrite C0_id. + intros H1; split; auto with zarith. + rewrite add_mult_div_2_plus_1. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; unfold base. + match goal with |- context[_ ^ ?X] => + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + repeat rewrite C0_id. + rewrite add_mult_div_2_plus_1. + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; unfold base. + match goal with |- context[_ ^ ?X] => + assert (tmp: forall p, 1 + (p - 1) = p); auto with zarith; + rewrite <- (tmp X); clear tmp; rewrite Zpower_exp; + try rewrite Zpower_1_r; auto with zarith + end. + rewrite Zpos_minus; auto with zarith. + rewrite Zmax_right; auto with zarith. + ring. + split; auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z w0);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + intros w2; rewrite C1_plus_wB. + rewrite <- Zplus_mod_one; auto with zarith. + rewrite Zmod_small; auto with zarith. + intros (Hw1, Hw2). + match goal with |- context[w_is_even ?y] => + generalize (spec_w_is_even y); + case (w_is_even y) + end. + repeat (rewrite C0_id || rewrite C1_plus_wB). + intros H1; split; auto with zarith. + rewrite add_mult_div_2. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + repeat (rewrite C0_id || rewrite C1_plus_wB). + rewrite spec_w_add_c; auto with zarith. + intros H1; split; auto with zarith. + rewrite add_mult_div_2. + replace (wB + [|w0|]) with ([|b|] + ([|w0|] - [|b|] + wB)); + auto with zarith. + rewrite Zmult_plus_distr_l; rewrite <- Zplus_assoc. + rewrite Hw1. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] 2); + auto with zarith. + rewrite H1; ring. + split; auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z w0);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + destruct (spec_to_Z b);auto with zarith. + Qed. + + Theorem wB_div_4: 4 * (wB / 4) = wB. + Proof. + unfold base. + assert (2 ^ Zpos w_digits = + 4 * (2 ^ (Zpos w_digits - 2))). + change 4 with (2 ^ 2). + rewrite <- Zpower_exp; auto with zarith. + f_equal; auto with zarith. + rewrite H. + rewrite (fun x => (Zmult_comm 4 (2 ^x))). + rewrite Z_div_mult; auto with zarith. + Qed. + + Theorem Zsquare_mult: forall p, p ^ 2 = p * p. + intros p; change 2 with (1 + 1); rewrite Zpower_exp; + try rewrite Zpower_1_r; auto with zarith. + Qed. + + Theorem Zsquare_pos: forall p, 0 <= p ^ 2. + intros p; case (Zle_or_lt 0 p); intros H1. + rewrite Zsquare_mult; apply Zmult_le_0_compat; auto with zarith. + rewrite Zsquare_mult; replace (p * p) with ((- p) * (- p)); try ring. + apply Zmult_le_0_compat; auto with zarith. + Qed. + + Lemma spec_split: forall x, + [|fst (split x)|] * wB + [|snd (split x)|] = [[x]]. + intros x; case x; simpl; autorewrite with w_rewrite; + auto with zarith. + Qed. + + Theorem mult_wwB: forall x y, [|x|] * [|y|] < wwB. + Proof. + intros x y; rewrite wwB_wBwB; rewrite Zpower_2. + generalize (spec_to_Z x); intros U. + generalize (spec_to_Z y); intros U1. + apply Zle_lt_trans with ((wB -1 ) * (wB - 1)); auto with zarith. + apply Zmult_le_compat; auto with zarith. + repeat (rewrite Zmult_minus_distr_r || rewrite Zmult_minus_distr_l); + auto with zarith. + Qed. + Hint Resolve mult_wwB. + + Lemma spec_ww_sqrt2 : forall x y, + wwB/ 4 <= [[x]] -> + let (s,r) := ww_sqrt2 x y in + [||WW x y||] = [[s]] ^ 2 + [+[r]] /\ + [+[r]] <= 2 * [[s]]. + intros x y H; unfold ww_sqrt2. + repeat match goal with |- context[split ?x] => + generalize (spec_split x); case (split x) + end; simpl fst; simpl snd. + intros w0 w1 Hw0 w2 w3 Hw1. + assert (U: wB/4 <= [|w2|]). + case (Zle_or_lt (wB / 4) [|w2|]); auto; intros H1. + contradict H; apply Zlt_not_le. + rewrite wwB_wBwB; rewrite Zpower_2. + pattern wB at 1; rewrite <- wB_div_4; rewrite <- Zmult_assoc; + rewrite Zmult_comm. + rewrite Z_div_mult; auto with zarith. + rewrite <- Hw1. + match goal with |- _ < ?X => + pattern X; rewrite <- Zplus_0_r; apply beta_lex_inv; + auto with zarith + end. + destruct (spec_to_Z w3);auto with zarith. + generalize (@spec_w_sqrt2 w2 w3 U); case (w_sqrt2 w2 w3). + intros w4 c (H1, H2). + assert (U1: wB/2 <= [|w4|]). + case (Zle_or_lt (wB/2) [|w4|]); auto with zarith. + intros U1. + assert (U2 : [|w4|] <= wB/2 -1); auto with zarith. + assert (U3 : [|w4|] ^ 2 <= wB/4 * wB - wB + 1); auto with zarith. + match goal with |- ?X ^ 2 <= ?Y => + rewrite Zsquare_mult; + replace Y with ((wB/2 - 1) * (wB/2 -1)) + end. + apply Zmult_le_compat; auto with zarith. + destruct (spec_to_Z w4);auto with zarith. + destruct (spec_to_Z w4);auto with zarith. + pattern wB at 4 5; rewrite <- wB_div_2. + rewrite Zmult_assoc. + replace ((wB / 4) * 2) with (wB / 2). + ring. + pattern wB at 1; rewrite <- wB_div_4. + change 4 with (2 * 2). + rewrite <- Zmult_assoc; rewrite (Zmult_comm 2). + rewrite Z_div_mult; try ring; auto with zarith. + assert (U4 : [+|c|] <= wB -2); auto with zarith. + apply Zle_trans with (1 := H2). + match goal with |- ?X <= ?Y => + replace Y with (2 * (wB/ 2 - 1)); auto with zarith + end. + pattern wB at 2; rewrite <- wB_div_2; auto with zarith. + match type of H1 with ?X = _ => + assert (U5: X < wB / 4 * wB) + end. + rewrite H1; auto with zarith. + contradict U; apply Zlt_not_le. + apply Zmult_lt_reg_r with wB; auto with zarith. + destruct (spec_to_Z w4);auto with zarith. + apply Zle_lt_trans with (2 := U5). + unfold ww_to_Z, zn2z_to_Z. + destruct (spec_to_Z w3);auto with zarith. + generalize (@spec_w_div2s c w0 w4 U1 H2). + case (w_div2s c w0 w4). + intros c0; case c0; intros w5; + repeat (rewrite C0_id || rewrite C1_plus_wB). + intros c1; case c1; intros w6; + repeat (rewrite C0_id || rewrite C1_plus_wB). + intros (H3, H4). + match goal with |- context [ww_sub_c ?y ?z] => + generalize (spec_ww_sub_c y z); case (ww_sub_c y z) + end. + intros z; change [-[C0 z]] with ([[z]]). + change [+[C0 z]] with ([[z]]). + intros H5; rewrite spec_w_square_c in H5; + auto. + split. + unfold zn2z_to_Z; rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite H5. + unfold ww_to_Z, zn2z_to_Z. + repeat rewrite Zsquare_mult; ring. + rewrite H5. + unfold ww_to_Z, zn2z_to_Z. + match goal with |- ?X - ?Y * ?Y <= _ => + assert (V := Zsquare_pos Y); + rewrite Zsquare_mult in V; + apply Zle_trans with X; auto with zarith; + clear V + end. + match goal with |- ?X * wB + ?Y <= 2 * (?Z * wB + ?T) => + apply Zle_trans with ((2 * Z - 1) * wB + wB); auto with zarith + end. + destruct (spec_to_Z w1);auto with zarith. + match goal with |- ?X <= _ => + replace X with (2 * [|w4|] * wB); auto with zarith + end. + rewrite Zmult_plus_distr_r; rewrite Zmult_assoc. + destruct (spec_to_Z w5); auto with zarith. + ring. + intros z; replace [-[C1 z]] with (- wwB + [[z]]). + 2: simpl; case wwB; auto with zarith. + intros H5; rewrite spec_w_square_c in H5; + auto. + match goal with |- context [ww_pred_c ?y] => + generalize (spec_ww_pred_c y); case (ww_pred_c y) + end. + intros z1; change [-[C0 z1]] with ([[z1]]). + rewrite ww_add_mult_mult_2. + rewrite spec_ww_add_c. + rewrite spec_ww_pred. + rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]); + auto with zarith. + intros Hz1; rewrite Zmod_small; auto with zarith. + match type of H5 with -?X + ?Y = ?Z => + assert (V: Y = Z + X); + try (rewrite <- H5; ring) + end. + split. + unfold zn2z_to_Z; rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite V. + rewrite Hz1. + unfold ww_to_Z; simpl zn2z_to_Z. + repeat rewrite Zsquare_mult; ring. + rewrite Hz1. + destruct (spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith. + assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)). + assert (0 < [[WW w4 w5]]); auto with zarith. + apply Zlt_le_trans with (wB/ 2 * wB + 0); auto with zarith. + autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith. + apply Zmult_lt_reg_r with 2; auto with zarith. + autorewrite with rm10. + rewrite Zmult_comm; rewrite wB_div_2; auto with zarith. + case (spec_to_Z w5);auto with zarith. + case (spec_to_Z w5);auto with zarith. + simpl. + assert (V2 := spec_to_Z w5);auto with zarith. + assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith. + split; auto with zarith. + assert (wwB <= 2 * [[WW w4 w5]]); auto with zarith. + apply Zle_trans with (2 * ([|w4|] * wB)). + rewrite wwB_wBwB; rewrite Zpower_2. + rewrite Zmult_assoc; apply Zmult_le_compat_r; auto with zarith. + rewrite <- wB_div_2; auto with zarith. + assert (V2 := spec_to_Z w5);auto with zarith. + simpl ww_to_Z; assert (V2 := spec_to_Z w5);auto with zarith. + assert (V1 := spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w4 w5)); auto with zarith. + intros z1; change [-[C1 z1]] with (-wwB + [[z1]]). + match goal with |- context[([+[C0 ?z]])] => + change [+[C0 z]] with ([[z]]) + end. + rewrite spec_ww_add; auto with zarith. + rewrite spec_ww_pred; auto with zarith. + rewrite ww_add_mult_mult_2. + rename V1 into VV1. + assert (VV2: 0 < [[WW w4 w5]]); auto with zarith. + apply Zlt_le_trans with (wB/ 2 * wB + 0); auto with zarith. + autorewrite with rm10; apply Zmult_lt_0_compat; auto with zarith. + apply Zmult_lt_reg_r with 2; auto with zarith. + autorewrite with rm10. + rewrite Zmult_comm; rewrite wB_div_2; auto with zarith. + assert (VV3 := spec_to_Z w5);auto with zarith. + assert (VV3 := spec_to_Z w5);auto with zarith. + simpl. + assert (VV3 := spec_to_Z w5);auto with zarith. + assert (VV3: wwB <= 2 * [[WW w4 w5]]); auto with zarith. + apply Zle_trans with (2 * ([|w4|] * wB)). + rewrite wwB_wBwB; rewrite Zpower_2. + rewrite Zmult_assoc; apply Zmult_le_compat_r; auto with zarith. + rewrite <- wB_div_2; auto with zarith. + case (spec_to_Z w5);auto with zarith. + simpl ww_to_Z; assert (V4 := spec_to_Z w5);auto with zarith. + rewrite <- Zmod_unique with (q := 1) (r := -wwB + 2 * [[WW w4 w5]]); + auto with zarith. + intros Hz1; rewrite Zmod_small; auto with zarith. + match type of H5 with -?X + ?Y = ?Z => + assert (V: Y = Z + X); + try (rewrite <- H5; ring) + end. + match type of Hz1 with -?X + ?Y = -?X + ?Z - 1 => + assert (V1: Y = Z - 1); + [replace (Z - 1) with (X + (-X + Z -1)); + [rewrite <- Hz1 | idtac]; ring + | idtac] + end. + rewrite <- Zmod_unique with (q := 1) (r := -wwB + [[z1]] + [[z]]); + auto with zarith. + unfold zn2z_to_Z; rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + split. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite V. + rewrite Hz1. + unfold ww_to_Z; simpl zn2z_to_Z. + repeat rewrite Zsquare_mult; ring. + assert (V2 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith. + assert (V2 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z);auto with zarith. + assert (V3 := spec_ww_to_Z w_digits w_to_Z spec_to_Z z1);auto with zarith. + split; auto with zarith. + rewrite (Zplus_comm (-wwB)); rewrite <- Zplus_assoc. + rewrite H5. + match goal with |- 0 <= ?X + (?Y - ?Z) => + apply Zle_trans with (X - Z); auto with zarith + end. + 2: generalize (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW w6 w1)); unfold ww_to_Z; auto with zarith. + rewrite V1. + match goal with |- 0 <= ?X - 1 - ?Y => + assert (Y < X); auto with zarith + end. + apply Zlt_le_trans with wwB; auto with zarith. + intros (H3, H4). + match goal with |- context [ww_sub_c ?y ?z] => + generalize (spec_ww_sub_c y z); case (ww_sub_c y z) + end. + intros z; change [-[C0 z]] with ([[z]]). + match goal with |- context[([+[C1 ?z]])] => + replace [+[C1 z]] with (wwB + [[z]]) + end. + 2: simpl; case wwB; auto. + intros H5; rewrite spec_w_square_c in H5; + auto. + split. + change ([||WW x y||]) with ([[x]] * wwB + [[y]]). + rewrite <- Hw1. + unfold ww_to_Z, zn2z_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite H5. + unfold ww_to_Z; simpl zn2z_to_Z. + rewrite wwB_wBwB. + repeat rewrite Zsquare_mult; ring. + simpl ww_to_Z. + rewrite H5. + simpl ww_to_Z. + rewrite wwB_wBwB; rewrite Zpower_2. + match goal with |- ?X * ?Y + (?Z * ?Y + ?T - ?U) <= _ => + apply Zle_trans with (X * Y + (Z * Y + T - 0)); + auto with zarith + end. + assert (V := Zsquare_pos [|w5|]); + rewrite Zsquare_mult in V; auto with zarith. + autorewrite with rm10. + match goal with |- _ <= 2 * (?U * ?V + ?W) => + apply Zle_trans with (2 * U * V + 0); + auto with zarith + end. + match goal with |- ?X * ?Y + (?Z * ?Y + ?T) <= _ => + replace (X * Y + (Z * Y + T)) with ((X + Z) * Y + T); + try ring + end. + apply Zlt_le_weak; apply beta_lex_inv; auto with zarith. + destruct (spec_to_Z w1);auto with zarith. + destruct (spec_to_Z w5);auto with zarith. + rewrite Zmult_plus_distr_r; auto with zarith. + rewrite Zmult_assoc; auto with zarith. + intros z; replace [-[C1 z]] with (- wwB + [[z]]). + 2: simpl; case wwB; auto with zarith. + intros H5; rewrite spec_w_square_c in H5; + auto. + match goal with |- context[([+[C0 ?z]])] => + change [+[C0 z]] with ([[z]]) + end. + match type of H5 with -?X + ?Y = ?Z => + assert (V: Y = Z + X); + try (rewrite <- H5; ring) + end. + change ([||WW x y||]) with ([[x]] * wwB + [[y]]). + simpl ww_to_Z. + rewrite <- Hw1. + simpl ww_to_Z in H1; rewrite H1. + rewrite <- Hw0. + split. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H3. + rewrite V. + simpl ww_to_Z. + rewrite wwB_wBwB. + repeat rewrite Zsquare_mult; ring. + rewrite V. + simpl ww_to_Z. + rewrite wwB_wBwB; rewrite Zpower_2. + match goal with |- (?Z * ?Y + ?T - ?U) + ?X * ?Y <= _ => + apply Zle_trans with ((Z * Y + T - 0) + X * Y); + auto with zarith + end. + assert (V1 := Zsquare_pos [|w5|]); + rewrite Zsquare_mult in V1; auto with zarith. + autorewrite with rm10. + match goal with |- _ <= 2 * (?U * ?V + ?W) => + apply Zle_trans with (2 * U * V + 0); + auto with zarith + end. + match goal with |- (?Z * ?Y + ?T) + ?X * ?Y <= _ => + replace ((Z * Y + T) + X * Y) with ((X + Z) * Y + T); + try ring + end. + apply Zlt_le_weak; apply beta_lex_inv; auto with zarith. + destruct (spec_to_Z w1);auto with zarith. + destruct (spec_to_Z w5);auto with zarith. + rewrite Zmult_plus_distr_r; auto with zarith. + rewrite Zmult_assoc; auto with zarith. + case Zle_lt_or_eq with (1 := H2); clear H2; intros H2. + intros c1 (H3, H4). + match type of H3 with ?X = ?Y => + absurd (X < Y) + end. + apply Zle_not_lt; rewrite <- H3; auto with zarith. + rewrite Zmult_plus_distr_l. + apply Zlt_le_trans with ((2 * [|w4|]) * wB + 0); + auto with zarith. + apply beta_lex_inv; auto with zarith. + destruct (spec_to_Z w0);auto with zarith. + assert (V1 := spec_to_Z w5);auto with zarith. + rewrite (Zmult_comm wB); auto with zarith. + assert (0 <= [|w5|] * (2 * [|w4|])); auto with zarith. + intros c1 (H3, H4); rewrite H2 in H3. + match type of H3 with ?X + ?Y = (?Z + ?T) * ?U + ?V => + assert (VV: (Y = (T * U) + V)); + [replace Y with ((X + Y) - X); + [rewrite H3; ring | ring] | idtac] + end. + assert (V1 := spec_to_Z w0);auto with zarith. + assert (V2 := spec_to_Z w5);auto with zarith. + case (Zle_lt_or_eq 0 [|w5|]); auto with zarith; intros V3. + match type of VV with ?X = ?Y => + absurd (X < Y) + end. + apply Zle_not_lt; rewrite <- VV; auto with zarith. + apply Zlt_le_trans with wB; auto with zarith. + match goal with |- _ <= ?X + _ => + apply Zle_trans with X; auto with zarith + end. + match goal with |- _ <= _ * ?X => + apply Zle_trans with (1 * X); auto with zarith + end. + autorewrite with rm10. + rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith. + rewrite <- V3 in VV; generalize VV; autorewrite with rm10; + clear VV; intros VV. + rewrite spec_ww_add_c; auto with zarith. + rewrite ww_add_mult_mult_2_plus_1. + match goal with |- context[?X mod wwB] => + rewrite <- Zmod_unique with (q := 1) (r := -wwB + X) + end; auto with zarith. + simpl ww_to_Z. + rewrite spec_w_Bm1; auto with zarith. + split. + change ([||WW x y||]) with ([[x]] * wwB + [[y]]). + rewrite <- Hw1. + simpl ww_to_Z in H1; rewrite H1. + rewrite <- Hw0. + match goal with |- (?X ^2 + ?Y) * wwB + (?Z * wB + ?T) = ?U => + apply trans_equal with ((X * wB) ^ 2 + (Y * wB + Z) * wB + T) + end. + repeat rewrite Zsquare_mult. + rewrite wwB_wBwB; ring. + rewrite H2. + rewrite wwB_wBwB. + repeat rewrite Zsquare_mult; ring. + assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z y);auto with zarith. + assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z y);auto with zarith. + simpl ww_to_Z; unfold ww_to_Z. + rewrite spec_w_Bm1; auto with zarith. + split. + rewrite wwB_wBwB; rewrite Zpower_2. + match goal with |- _ <= -?X + (2 * (?Z * ?T + ?U) + ?V) => + assert (X <= 2 * Z * T); auto with zarith + end. + apply Zmult_le_compat_r; auto with zarith. + rewrite <- wB_div_2; apply Zmult_le_compat_l; auto with zarith. + rewrite Zmult_plus_distr_r; auto with zarith. + rewrite Zmult_assoc; auto with zarith. + match goal with |- _ + ?X < _ => + replace X with ((2 * (([|w4|]) + 1) * wB) - 1); try ring + end. + assert (2 * ([|w4|] + 1) * wB <= 2 * wwB); auto with zarith. + rewrite <- Zmult_assoc; apply Zmult_le_compat_l; auto with zarith. + rewrite wwB_wBwB; rewrite Zpower_2. + apply Zmult_le_compat_r; auto with zarith. + case (spec_to_Z w4);auto with zarith. + Qed. + + Lemma spec_ww_is_zero: forall x, + if ww_is_zero x then [[x]] = 0 else 0 < [[x]]. + intro x; unfold ww_is_zero. + generalize (spec_ww_compare W0 x); case (ww_compare W0 x); + auto with zarith. + simpl ww_to_Z. + assert (V4 := spec_ww_to_Z w_digits w_to_Z spec_to_Z x);auto with zarith. + Qed. + + Lemma wwB_4_2: 2 * (wwB / 4) = wwB/ 2. + pattern wwB at 1; rewrite wwB_wBwB; rewrite Zpower_2. + rewrite <- wB_div_2. + match goal with |- context[(2 * ?X) * (2 * ?Z)] => + replace ((2 * X) * (2 * Z)) with ((X * Z) * 4); try ring + end. + rewrite Z_div_mult; auto with zarith. + rewrite Zmult_assoc; rewrite wB_div_2. + rewrite wwB_div_2; ring. + Qed. + + + Lemma spec_ww_head1 + : forall x : zn2z w, + (ww_is_even (ww_head1 x) = true) /\ + (0 < [[x]] -> wwB / 4 <= 2 ^ [[ww_head1 x]] * [[x]] < wwB). + assert (U := wB_pos w_digits). + intros x; unfold ww_head1. + generalize (spec_ww_is_even (ww_head0 x)); case_eq (ww_is_even (ww_head0 x)). + intros HH H1; rewrite HH; split; auto. + intros H2. + generalize (spec_ww_head0 x H2); case (ww_head0 x); autorewrite with rm10. + intros (H3, H4); split; auto with zarith. + apply Zle_trans with (2 := H3). + apply Zdiv_le_compat_l; auto with zarith. + intros xh xl (H3, H4); split; auto with zarith. + apply Zle_trans with (2 := H3). + apply Zdiv_le_compat_l; auto with zarith. + intros H1. + case (spec_to_w_Z (ww_head0 x)); intros Hv1 Hv2. + assert (Hp0: 0 < [[ww_head0 x]]). + generalize (spec_ww_is_even (ww_head0 x)); rewrite H1. + generalize Hv1; case [[ww_head0 x]]. + rewrite Zmod_small; auto with zarith. + intros; assert (0 < Zpos p); auto with zarith. + red; simpl; auto. + intros p H2; case H2; auto. + assert (Hp: [[ww_pred (ww_head0 x)]] = [[ww_head0 x]] - 1). + rewrite spec_ww_pred. + rewrite Zmod_small; auto with zarith. + intros H2; split. + generalize (spec_ww_is_even (ww_pred (ww_head0 x))); + case ww_is_even; auto. + rewrite Hp. + rewrite Zminus_mod; auto with zarith. + rewrite H2; repeat rewrite Zmod_small; auto with zarith. + intros H3; rewrite Hp. + case (spec_ww_head0 x); auto; intros Hv3 Hv4. + assert (Hu: forall u, 0 < u -> 2 * 2 ^ (u - 1) = 2 ^u). + intros u Hu. + pattern 2 at 1; rewrite <- Zpower_1_r. + rewrite <- Zpower_exp; auto with zarith. + ring_simplify (1 + (u - 1)); auto with zarith. + split; auto with zarith. + apply Zmult_le_reg_r with 2; auto with zarith. + repeat rewrite (fun x => Zmult_comm x 2). + rewrite wwB_4_2. + rewrite Zmult_assoc; rewrite Hu; auto with zarith. + apply Zle_lt_trans with (2 * 2 ^ ([[ww_head0 x]] - 1) * [[x]]); auto with zarith; + rewrite Hu; auto with zarith. + apply Zmult_le_compat_r; auto with zarith. + apply Zpower_le_monotone; auto with zarith. + Qed. + + Theorem wwB_4_wB_4: wwB / 4 = wB / 4 * wB. + apply sym_equal; apply Zdiv_unique with 0; + auto with zarith. + rewrite Zmult_assoc; rewrite wB_div_4; auto with zarith. + rewrite wwB_wBwB; ring. + Qed. + + Lemma spec_ww_sqrt : forall x, + [[ww_sqrt x]] ^ 2 <= [[x]] < ([[ww_sqrt x]] + 1) ^ 2. + assert (U := wB_pos w_digits). + intro x; unfold ww_sqrt. + generalize (spec_ww_is_zero x); case (ww_is_zero x). + simpl ww_to_Z; simpl Zpower; unfold Zpower_pos; simpl; + auto with zarith. + intros H1. + generalize (spec_ww_compare (ww_head1 x) W0); case ww_compare; + simpl ww_to_Z; autorewrite with rm10. + generalize H1; case x. + intros HH; contradict HH; simpl ww_to_Z; auto with zarith. + intros w0 w1; simpl ww_to_Z; autorewrite with w_rewrite rm10. + intros H2; case (spec_ww_head1 (WW w0 w1)); intros H3 H4 H5. + generalize (H4 H2); clear H4; rewrite H5; clear H5; autorewrite with rm10. + intros (H4, H5). + assert (V: wB/4 <= [|w0|]). + apply beta_lex with 0 [|w1|] wB; auto with zarith; autorewrite with rm10. + rewrite <- wwB_4_wB_4; auto. + generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith. + case (w_sqrt2 w0 w1); intros w2 c. + simpl ww_to_Z; simpl fst. + case c; unfold interp_carry; autorewrite with rm10. + intros w3 (H6, H7); rewrite H6. + assert (V1 := spec_to_Z w3);auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith. + match goal with |- ?X < ?Z => + replace Z with (X + 1); auto with zarith + end. + repeat rewrite Zsquare_mult; ring. + intros w3 (H6, H7); rewrite H6. + assert (V1 := spec_to_Z w3);auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with ([|w2|] ^2 + 2 * [|w2|]); auto with zarith. + match goal with |- ?X < ?Z => + replace Z with (X + 1); auto with zarith + end. + repeat rewrite Zsquare_mult; ring. + intros HH; case (spec_to_w_Z (ww_head1 x)); auto with zarith. + intros Hv1. + case (spec_ww_head1 x); intros Hp1 Hp2. + generalize (Hp2 H1); clear Hp2; intros Hp2. + assert (Hv2: [[ww_head1 x]] <= Zpos (xO w_digits)). + case (Zle_or_lt (Zpos (xO w_digits)) [[ww_head1 x]]); auto with zarith; intros HH1. + case Hp2; intros _ HH2; contradict HH2. + apply Zle_not_lt; unfold base. + apply Zle_trans with (2 ^ [[ww_head1 x]]). + apply Zpower_le_monotone; auto with zarith. + pattern (2 ^ [[ww_head1 x]]) at 1; + rewrite <- (Zmult_1_r (2 ^ [[ww_head1 x]])). + apply Zmult_le_compat_l; auto with zarith. + generalize (spec_ww_add_mul_div x W0 (ww_head1 x) Hv2); + case ww_add_mul_div. + simpl ww_to_Z; autorewrite with w_rewrite rm10. + rewrite Zmod_small; auto with zarith. + intros H2; case (Zmult_integral _ _ (sym_equal H2)); clear H2; intros H2. + rewrite H2; unfold Zpower, Zpower_pos; simpl; auto with zarith. + match type of H2 with ?X = ?Y => + absurd (Y < X); try (rewrite H2; auto with zarith; fail) + end. + apply Zpower_gt_0; auto with zarith. + split; auto with zarith. + case Hp2; intros _ tmp; apply Zle_lt_trans with (2 := tmp); + clear tmp. + rewrite Zmult_comm; apply Zmult_le_compat_r; auto with zarith. + assert (Hv0: [[ww_head1 x]] = 2 * ([[ww_head1 x]]/2)). + pattern [[ww_head1 x]] at 1; rewrite (Z_div_mod_eq [[ww_head1 x]] 2); + auto with zarith. + generalize (spec_ww_is_even (ww_head1 x)); rewrite Hp1; + intros tmp; rewrite tmp; rewrite Zplus_0_r; auto. + intros w0 w1; autorewrite with w_rewrite rm10. + rewrite Zmod_small; auto with zarith. + 2: rewrite Zmult_comm; auto with zarith. + intros H2. + assert (V: wB/4 <= [|w0|]). + apply beta_lex with 0 [|w1|] wB; auto with zarith; autorewrite with rm10. + simpl ww_to_Z in H2; rewrite H2. + rewrite <- wwB_4_wB_4; auto with zarith. + rewrite Zmult_comm; auto with zarith. + assert (V1 := spec_to_Z w1);auto with zarith. + generalize (@spec_w_sqrt2 w0 w1 V);auto with zarith. + case (w_sqrt2 w0 w1); intros w2 c. + case (spec_to_Z w2); intros HH1 HH2. + simpl ww_to_Z; simpl fst. + assert (Hv3: [[ww_pred ww_zdigits]] + = Zpos (xO w_digits) - 1). + rewrite spec_ww_pred; rewrite spec_ww_zdigits. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (Zpos (xO w_digits)); auto with zarith. + unfold base; apply Zpower2_le_lin; auto with zarith. + assert (Hv4: [[ww_head1 x]]/2 < wB). + apply Zle_lt_trans with (Zpos w_digits). + apply Zmult_le_reg_r with 2; auto with zarith. + repeat rewrite (fun x => Zmult_comm x 2). + rewrite <- Hv0; rewrite <- Zpos_xO; auto. + unfold base; apply Zpower2_lt_lin; auto with zarith. + assert (Hv5: [[(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))]] + = [[ww_head1 x]]/2). + rewrite spec_ww_add_mul_div. + simpl ww_to_Z; autorewrite with rm10. + rewrite Hv3. + ring_simplify (Zpos (xO w_digits) - (Zpos (xO w_digits) - 1)). + rewrite Zpower_1_r. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + apply Zlt_le_trans with (1 := Hv4); auto with zarith. + unfold base; apply Zpower_le_monotone; auto with zarith. + split; unfold ww_digits; try rewrite Zpos_xO; auto with zarith. + rewrite Hv3; auto with zarith. + assert (Hv6: [|low(ww_add_mul_div (ww_pred ww_zdigits) W0 (ww_head1 x))|] + = [[ww_head1 x]]/2). + rewrite spec_low. + rewrite Hv5; rewrite Zmod_small; auto with zarith. + rewrite spec_w_add_mul_div; auto with zarith. + rewrite spec_w_sub; auto with zarith. + rewrite spec_w_0. + simpl ww_to_Z; autorewrite with rm10. + rewrite Hv6; rewrite spec_w_zdigits. + rewrite (fun x y => Zmod_small (x - y)). + ring_simplify (Zpos w_digits - (Zpos w_digits - [[ww_head1 x]] / 2)). + rewrite Zmod_small. + simpl ww_to_Z in H2; rewrite H2; auto with zarith. + intros (H4, H5); split. + apply Zmult_le_reg_r with (2 ^ [[ww_head1 x]]); auto with zarith. + rewrite H4. + apply Zle_trans with ([|w2|] ^ 2); auto with zarith. + rewrite Zmult_comm. + pattern [[ww_head1 x]] at 1; + rewrite Hv0; auto with zarith. + rewrite (Zmult_comm 2); rewrite Zpower_mult; + auto with zarith. + assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2); + try (intros; repeat rewrite Zsquare_mult; ring); + rewrite tmp; clear tmp. + apply Zpower_le_monotone3; auto with zarith. + split; auto with zarith. + pattern [|w2|] at 2; + rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]] / 2))); + auto with zarith. + match goal with |- ?X <= ?X + ?Y => + assert (0 <= Y); auto with zarith + end. + case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]] / 2))); auto with zarith. + case c; unfold interp_carry; autorewrite with rm10; + intros w3; assert (V3 := spec_to_Z w3);auto with zarith. + apply Zmult_lt_reg_r with (2 ^ [[ww_head1 x]]); auto with zarith. + rewrite H4. + apply Zle_lt_trans with ([|w2|] ^ 2 + 2 * [|w2|]); auto with zarith. + apply Zlt_le_trans with (([|w2|] + 1) ^ 2); auto with zarith. + match goal with |- ?X < ?Y => + replace Y with (X + 1); auto with zarith + end. + repeat rewrite (Zsquare_mult); ring. + rewrite Zmult_comm. + pattern [[ww_head1 x]] at 1; rewrite Hv0. + rewrite (Zmult_comm 2); rewrite Zpower_mult; + auto with zarith. + assert (tmp: forall p q, p ^ 2 * q ^ 2 = (p * q) ^2); + try (intros; repeat rewrite Zsquare_mult; ring); + rewrite tmp; clear tmp. + apply Zpower_le_monotone3; auto with zarith. + split; auto with zarith. + pattern [|w2|] at 1; rewrite (Z_div_mod_eq [|w2|] (2 ^ ([[ww_head1 x]]/2))); + auto with zarith. + rewrite <- Zplus_assoc; rewrite Zmult_plus_distr_r. + autorewrite with rm10; apply Zplus_le_compat_l; auto with zarith. + case (Z_mod_lt [|w2|] (2 ^ ([[ww_head1 x]]/2))); auto with zarith. + split; auto with zarith. + apply Zle_lt_trans with ([|w2|]); auto with zarith. + apply Zdiv_le_upper_bound; auto with zarith. + pattern [|w2|] at 1; replace [|w2|] with ([|w2|] * 2 ^0); + auto with zarith. + apply Zmult_le_compat_l; auto with zarith. + apply Zpower_le_monotone; auto with zarith. + rewrite Zpower_0_r; autorewrite with rm10; auto. + split; auto with zarith. + rewrite Hv0 in Hv2; rewrite (Zpos_xO w_digits) in Hv2; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + rewrite spec_w_sub; auto with zarith. + rewrite Hv6; rewrite spec_w_zdigits; auto with zarith. + assert (Hv7: 0 < [[ww_head1 x]]/2); auto with zarith. + rewrite Zmod_small; auto with zarith. + split; auto with zarith. + assert ([[ww_head1 x]]/2 <= Zpos w_digits); auto with zarith. + apply Zmult_le_reg_r with 2; auto with zarith. + repeat rewrite (fun x => Zmult_comm x 2). + rewrite <- Hv0; rewrite <- Zpos_xO; auto with zarith. + apply Zle_lt_trans with (Zpos w_digits); auto with zarith. + unfold base; apply Zpower2_lt_lin; auto with zarith. + Qed. + +End DoubleSqrt. diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v new file mode 100644 index 00000000..269d62bb --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleSub.v @@ -0,0 +1,357 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleSub.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import DoubleBase. + +Open Local Scope Z_scope. + +Section DoubleSub. + Variable w : Type. + Variable w_0 : w. + Variable w_Bm1 : w. + Variable w_WW : w -> w -> zn2z w. + Variable ww_Bm1 : zn2z w. + Variable w_opp_c : w -> carry w. + Variable w_opp_carry : w -> w. + Variable w_pred_c : w -> carry w. + Variable w_sub_c : w -> w -> carry w. + Variable w_sub_carry_c : w -> w -> carry w. + Variable w_opp : w -> w. + Variable w_pred : w -> w. + Variable w_sub : w -> w -> w. + Variable w_sub_carry : w -> w -> w. + + (* ** Opposites ** *) + Definition ww_opp_c x := + match x with + | W0 => C0 W0 + | WW xh xl => + match w_opp_c xl with + | C0 _ => + match w_opp_c xh with + | C0 h => C0 W0 + | C1 h => C1 (WW h w_0) + end + | C1 l => C1 (WW (w_opp_carry xh) l) + end + end. + + Definition ww_opp x := + match x with + | W0 => W0 + | WW xh xl => + match w_opp_c xl with + | C0 _ => WW (w_opp xh) w_0 + | C1 l => WW (w_opp_carry xh) l + end + end. + + Definition ww_opp_carry x := + match x with + | W0 => ww_Bm1 + | WW xh xl => w_WW (w_opp_carry xh) (w_opp_carry xl) + end. + + Definition ww_pred_c x := + match x with + | W0 => C1 ww_Bm1 + | WW xh xl => + match w_pred_c xl with + | C0 l => C0 (w_WW xh l) + | C1 _ => + match w_pred_c xh with + | C0 h => C0 (WW h w_Bm1) + | C1 _ => C1 ww_Bm1 + end + end + end. + + Definition ww_pred x := + match x with + | W0 => ww_Bm1 + | WW xh xl => + match w_pred_c xl with + | C0 l => w_WW xh l + | C1 l => WW (w_pred xh) w_Bm1 + end + end. + + Definition ww_sub_c x y := + match y, x with + | W0, _ => C0 x + | WW yh yl, W0 => ww_opp_c (WW yh yl) + | WW yh yl, WW xh xl => + match w_sub_c xl yl with + | C0 l => + match w_sub_c xh yh with + | C0 h => C0 (w_WW h l) + | C1 h => C1 (WW h l) + end + | C1 l => + match w_sub_carry_c xh yh with + | C0 h => C0 (WW h l) + | C1 h => C1 (WW h l) + end + end + end. + + Definition ww_sub x y := + match y, x with + | W0, _ => x + | WW yh yl, W0 => ww_opp (WW yh yl) + | WW yh yl, WW xh xl => + match w_sub_c xl yl with + | C0 l => w_WW (w_sub xh yh) l + | C1 l => WW (w_sub_carry xh yh) l + end + end. + + Definition ww_sub_carry_c x y := + match y, x with + | W0, W0 => C1 ww_Bm1 + | W0, WW xh xl => ww_pred_c (WW xh xl) + | WW yh yl, W0 => C1 (ww_opp_carry (WW yh yl)) + | WW yh yl, WW xh xl => + match w_sub_carry_c xl yl with + | C0 l => + match w_sub_c xh yh with + | C0 h => C0 (w_WW h l) + | C1 h => C1 (WW h l) + end + | C1 l => + match w_sub_carry_c xh yh with + | C0 h => C0 (w_WW h l) + | C1 h => C1 (w_WW h l) + end + end + end. + + Definition ww_sub_carry x y := + match y, x with + | W0, W0 => ww_Bm1 + | W0, WW xh xl => ww_pred (WW xh xl) + | WW yh yl, W0 => ww_opp_carry (WW yh yl) + | WW yh yl, WW xh xl => + match w_sub_carry_c xl yl with + | C0 l => w_WW (w_sub xh yh) l + | C1 l => w_WW (w_sub_carry xh yh) l + end + end. + + (*Section DoubleProof.*) + Variable w_digits : positive. + Variable w_to_Z : w -> Z. + + + Notation wB := (base w_digits). + Notation wwB := (base (ww_digits w_digits)). + Notation "[| x |]" := (w_to_Z x) (at level 0, x at level 99). + Notation "[+| c |]" := + (interp_carry 1 wB w_to_Z c) (at level 0, x at level 99). + Notation "[-| c |]" := + (interp_carry (-1) wB w_to_Z c) (at level 0, x at level 99). + + Notation "[[ x ]]" := (ww_to_Z w_digits w_to_Z x)(at level 0, x at level 99). + Notation "[+[ c ]]" := + (interp_carry 1 wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + Notation "[-[ c ]]" := + (interp_carry (-1) wwB (ww_to_Z w_digits w_to_Z) c) + (at level 0, x at level 99). + + Variable spec_w_0 : [|w_0|] = 0. + Variable spec_w_Bm1 : [|w_Bm1|] = wB - 1. + Variable spec_ww_Bm1 : [[ww_Bm1]] = wwB - 1. + Variable spec_to_Z : forall x, 0 <= [|x|] < wB. + Variable spec_w_WW : forall h l, [[w_WW h l]] = [|h|] * wB + [|l|]. + + Variable spec_opp_c : forall x, [-|w_opp_c x|] = -[|x|]. + Variable spec_opp : forall x, [|w_opp x|] = (-[|x|]) mod wB. + Variable spec_opp_carry : forall x, [|w_opp_carry x|] = wB - [|x|] - 1. + + Variable spec_pred_c : forall x, [-|w_pred_c x|] = [|x|] - 1. + Variable spec_sub_c : forall x y, [-|w_sub_c x y|] = [|x|] - [|y|]. + Variable spec_sub_carry_c : + forall x y, [-|w_sub_carry_c x y|] = [|x|] - [|y|] - 1. + + Variable spec_pred : forall x, [|w_pred x|] = ([|x|] - 1) mod wB. + Variable spec_sub : forall x y, [|w_sub x y|] = ([|x|] - [|y|]) mod wB. + Variable spec_sub_carry : + forall x y, [|w_sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB. + + + Lemma spec_ww_opp_c : forall x, [-[ww_opp_c x]] = -[[x]]. + Proof. + destruct x as [ |xh xl];simpl. reflexivity. + rewrite Zopp_plus_distr;generalize (spec_opp_c xl);destruct (w_opp_c xl) + as [l|l];intros H;unfold interp_carry in H;rewrite <- H; + rewrite Zopp_mult_distr_l. + assert ([|l|] = 0). + assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega. + rewrite H0;generalize (spec_opp_c xh);destruct (w_opp_c xh) + as [h|h];intros H1;unfold interp_carry in *;rewrite <- H1. + assert ([|h|] = 0). + assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega. + rewrite H2;reflexivity. + simpl ww_to_Z;rewrite wwB_wBwB;rewrite spec_w_0;ring. + unfold interp_carry;simpl ww_to_Z;rewrite wwB_wBwB;rewrite spec_opp_carry; + ring. + Qed. + + Lemma spec_ww_opp : forall x, [[ww_opp x]] = (-[[x]]) mod wwB. + Proof. + destruct x as [ |xh xl];simpl. reflexivity. + rewrite Zopp_plus_distr;rewrite Zopp_mult_distr_l. + generalize (spec_opp_c xl);destruct (w_opp_c xl) + as [l|l];intros H;unfold interp_carry in H;rewrite <- H;simpl ww_to_Z. + rewrite spec_w_0;rewrite Zplus_0_r;rewrite wwB_wBwB. + assert ([|l|] = 0). + assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega. + rewrite H0;rewrite Zplus_0_r; rewrite Zpower_2; + rewrite Zmult_mod_distr_r;try apply lt_0_wB. + rewrite spec_opp;trivial. + apply Zmod_unique with (q:= -1). + exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW (w_opp_carry xh) l)). + rewrite spec_opp_carry;rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_opp_carry : forall x, [[ww_opp_carry x]] = wwB - [[x]] - 1. + Proof. + destruct x as [ |xh xl];simpl. rewrite spec_ww_Bm1;ring. + rewrite spec_w_WW;simpl;repeat rewrite spec_opp_carry;rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_pred_c : forall x, [-[ww_pred_c x]] = [[x]] - 1. + Proof. + destruct x as [ |xh xl];unfold ww_pred_c. + unfold interp_carry;rewrite spec_ww_Bm1;simpl ww_to_Z;ring. + simpl ww_to_Z;replace (([|xh|]*wB+[|xl|])-1) with ([|xh|]*wB+([|xl|]-1)). + 2:ring. generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l]; + intros H;unfold interp_carry in H;rewrite <- H. simpl;apply spec_w_WW. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + assert ([|l|] = wB - 1). + assert (H1:= spec_to_Z l);assert (H2 := spec_to_Z xl);omega. + rewrite H0;change ([|xh|] + -1) with ([|xh|] - 1). + generalize (spec_pred_c xh);destruct (w_pred_c xh) as [h|h]; + intros H1;unfold interp_carry in H1;rewrite <- H1. + simpl;rewrite spec_w_Bm1;ring. + assert ([|h|] = wB - 1). + assert (H3:= spec_to_Z h);assert (H2 := spec_to_Z xh);omega. + rewrite H2;unfold interp_carry;rewrite spec_ww_Bm1;rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_sub_c : forall x y, [-[ww_sub_c x y]] = [[x]] - [[y]]. + Proof. + destruct y as [ |yh yl];simpl. ring. + destruct x as [ |xh xl];simpl. exact (spec_ww_opp_c (WW yh yl)). + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|])) + with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|])). 2:ring. + generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl) as [l|l];intros H; + unfold interp_carry in H;rewrite <- H. + generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1; + unfold interp_carry in H1;rewrite <- H1;unfold interp_carry; + try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1). + generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h]; + intros H1;unfold interp_carry in *;rewrite <- H1;simpl ww_to_Z; + try rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_sub_carry_c : + forall x y, [-[ww_sub_carry_c x y]] = [[x]] - [[y]] - 1. + Proof. + destruct y as [ |yh yl];simpl. + unfold Zminus;simpl;rewrite Zplus_0_r;exact (spec_ww_pred_c x). + destruct x as [ |xh xl]. + unfold interp_carry;rewrite spec_w_WW;simpl ww_to_Z;rewrite wwB_wBwB; + repeat rewrite spec_opp_carry;ring. + simpl ww_to_Z. + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1) + with (([|xh|]-[|yh|])*wB + ([|xl|]-[|yl|]-1)). 2:ring. + generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl) + as [l|l];intros H;unfold interp_carry in H;rewrite <- H. + generalize (spec_sub_c xh yh);destruct (w_sub_c xh yh) as [h|h];intros H1; + unfold interp_carry in H1;rewrite <- H1;unfold interp_carry; + try rewrite spec_w_WW;simpl ww_to_Z;try rewrite wwB_wBwB;ring. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + change ([|xh|] - [|yh|] + -1) with ([|xh|] - [|yh|] - 1). + generalize (spec_sub_carry_c xh yh);destruct (w_sub_carry_c xh yh) as [h|h]; + intros H1;unfold interp_carry in *;rewrite <- H1;try rewrite spec_w_WW; + simpl ww_to_Z; try rewrite wwB_wBwB;ring. + Qed. + + Lemma spec_ww_pred : forall x, [[ww_pred x]] = ([[x]] - 1) mod wwB. + Proof. + destruct x as [ |xh xl];simpl. + apply Zmod_unique with (-1). apply spec_ww_to_Z;trivial. + rewrite spec_ww_Bm1;ring. + replace ([|xh|]*wB + [|xl|] - 1) with ([|xh|]*wB + ([|xl|] - 1)). 2:ring. + generalize (spec_pred_c xl);destruct (w_pred_c xl) as [l|l];intro H; + unfold interp_carry in H;rewrite <- H;simpl ww_to_Z. + rewrite Zmod_small. apply spec_w_WW. + exact (spec_ww_to_Z w_digits w_to_Z spec_to_Z (WW xh l)). + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + change ([|xh|] + -1) with ([|xh|] - 1). + assert ([|l|] = wB - 1). + assert (H1:= spec_to_Z l);assert (H2:= spec_to_Z xl);omega. + rewrite (mod_wwB w_digits w_to_Z);trivial. + rewrite spec_pred;rewrite spec_w_Bm1;rewrite <- H0;trivial. + Qed. + + Lemma spec_ww_sub : forall x y, [[ww_sub x y]] = ([[x]] - [[y]]) mod wwB. + Proof. + destruct y as [ |yh yl];simpl. + ring_simplify ([[x]] - 0);rewrite Zmod_small;trivial. apply spec_ww_to_Z;trivial. + destruct x as [ |xh xl];simpl. exact (spec_ww_opp (WW yh yl)). + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|])) + with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|])). 2:ring. + generalize (spec_sub_c xl yl);destruct (w_sub_c xl yl)as[l|l];intros H; + unfold interp_carry in H;rewrite <- H. + rewrite spec_w_WW;rewrite (mod_wwB w_digits w_to_Z spec_to_Z). + rewrite spec_sub;trivial. + simpl ww_to_Z;rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial. + Qed. + + Lemma spec_ww_sub_carry : + forall x y, [[ww_sub_carry x y]] = ([[x]] - [[y]] - 1) mod wwB. + Proof. + destruct y as [ |yh yl];simpl. + ring_simplify ([[x]] - 0);exact (spec_ww_pred x). + destruct x as [ |xh xl];simpl. + apply Zmod_unique with (-1). + apply spec_ww_to_Z;trivial. + fold (ww_opp_carry (WW yh yl)). + rewrite (spec_ww_opp_carry (WW yh yl));simpl ww_to_Z;ring. + replace ([|xh|] * wB + [|xl|] - ([|yh|] * wB + [|yl|]) - 1) + with (([|xh|] - [|yh|]) * wB + ([|xl|] - [|yl|] - 1)). 2:ring. + generalize (spec_sub_carry_c xl yl);destruct (w_sub_carry_c xl yl)as[l|l]; + intros H;unfold interp_carry in H;rewrite <- H;rewrite spec_w_WW. + rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub;trivial. + rewrite Zplus_assoc;rewrite <- Zmult_plus_distr_l. + rewrite (mod_wwB w_digits w_to_Z spec_to_Z);rewrite spec_sub_carry;trivial. + Qed. + +(* End DoubleProof. *) + +End DoubleSub. + + + + + diff --git a/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v new file mode 100644 index 00000000..28d40094 --- /dev/null +++ b/theories/Numbers/Cyclic/DoubleCyclic/DoubleType.v @@ -0,0 +1,71 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: DoubleType.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Set Implicit Arguments. + +Require Import ZArith. +Open Local Scope Z_scope. + +Definition base digits := Zpower 2 (Zpos digits). + +Section Carry. + + Variable A : Type. + + Inductive carry := + | C0 : A -> carry + | C1 : A -> carry. + + Definition interp_carry (sign:Z)(B:Z)(interp:A -> Z) c := + match c with + | C0 x => interp x + | C1 x => sign*B + interp x + end. + +End Carry. + +Section Zn2Z. + + Variable znz : Type. + + (** From a type [znz] representing a cyclic structure Z/nZ, + we produce a representation of Z/2nZ by pairs of elements of [znz] + (plus a special case for zero). High half of the new number comes + first. + *) + + Inductive zn2z := + | W0 : zn2z + | WW : znz -> znz -> zn2z. + + Definition zn2z_to_Z (wB:Z) (w_to_Z:znz->Z) (x:zn2z) := + match x with + | W0 => 0 + | WW xh xl => w_to_Z xh * wB + w_to_Z xl + end. + +End Zn2Z. + +Implicit Arguments W0 [znz]. + +(** From a cyclic representation [w], we iterate the [zn2z] construct + [n] times, gaining the type of binary trees of depth at most [n], + whose leafs are either W0 (if depth < n) or elements of w + (if depth = n). +*) + +Fixpoint word (w:Type) (n:nat) : Type := + match n with + | O => w + | S n => zn2z (word w n) + end. + diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v new file mode 100644 index 00000000..4d655eac --- /dev/null +++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v @@ -0,0 +1,2516 @@ +(************************************************************************) +(* 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: Cyclic31.v 11034 2008-06-02 08:15:34Z thery $ i*) + +(** * Int31 numbers defines indeed a cyclic structure : Z/(2^31)Z *) + +(** +Author: Arnaud Spiwack (+ Pierre Letouzey) +*) + +Require Import List. +Require Import Min. +Require Export Int31. +Require Import Znumtheory. +Require Import Zgcd_alt. +Require Import Zpow_facts. +Require Import BigNumPrelude. +Require Import CyclicAxioms. +Require Import ROmega. + +Open Scope nat_scope. +Open Scope int31_scope. + +Section Basics. + + (** * Basic results about [iszero], [shiftl], [shiftr] *) + + 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 + end. + reflexivity. + Qed. + + Lemma iszero_not_eq0 : forall x, iszero x = false -> x<>0. + Proof. + intros x H Eq; rewrite Eq in H; simpl in *; discriminate. + Qed. + + Lemma sneakl_shiftr : forall x, + x = sneakl (firstr x) (shiftr x). + Proof. + destruct x; simpl; auto. + Qed. + + Lemma sneakr_shiftl : forall x, + x = sneakr (firstl x) (shiftl x). + Proof. + destruct x; simpl; auto. + Qed. + + Lemma twice_zero : forall x, + twice x = 0 <-> twice_plus_one x = 1. + Proof. + destruct x; simpl in *; split; + intro H; injection H; intros; subst; auto. + Qed. + + 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. + destruct x; simpl in *; rewrite H; auto. + destruct x; simpl in *; rewrite H; auto. + Qed. + + + + (** * Iterated shift to the right *) + + Definition nshiftr n x := iter_nat n _ shiftr x. + + Lemma nshiftr_S : + forall n x, nshiftr (S n) x = shiftr (nshiftr n x). + Proof. + reflexivity. + Qed. + + Lemma nshiftr_S_tail : + forall n x, nshiftr (S n) x = nshiftr n (shiftr x). + Proof. + induction n; simpl; auto. + intros; rewrite nshiftr_S, IHn, nshiftr_S; auto. + Qed. + + Lemma nshiftr_n_0 : forall n, nshiftr n 0 = 0. + Proof. + induction n; simpl; auto. + rewrite nshiftr_S, IHn; auto. + Qed. + + Lemma nshiftr_size : forall x, nshiftr size x = 0. + Proof. + destruct x; simpl; auto. + Qed. + + Lemma nshiftr_above_size : forall k x, size<=k -> + nshiftr k x = 0. + Proof. + intros. + replace k with ((k-size)+size)%nat by omega. + induction (k-size)%nat; auto. + rewrite nshiftr_size; auto. + simpl; rewrite nshiftr_S, IHn; auto. + Qed. + + (** * Iterated shift to the left *) + + Definition nshiftl n x := iter_nat n _ shiftl x. + + Lemma nshiftl_S : + forall n x, nshiftl (S n) x = shiftl (nshiftl n x). + Proof. + reflexivity. + Qed. + + Lemma nshiftl_S_tail : + forall n x, nshiftl (S n) x = nshiftl n (shiftl x). + Proof. + induction n; simpl; auto. + intros; rewrite nshiftl_S, IHn, nshiftl_S; auto. + Qed. + + Lemma nshiftl_n_0 : forall n, nshiftl n 0 = 0. + Proof. + induction n; simpl; auto. + rewrite nshiftl_S, IHn; auto. + Qed. + + Lemma nshiftl_size : forall x, nshiftl size x = 0. + Proof. + destruct x; simpl; auto. + Qed. + + Lemma nshiftl_above_size : forall k x, size<=k -> + nshiftl k x = 0. + Proof. + intros. + replace k with ((k-size)+size)%nat by omega. + induction (k-size)%nat; auto. + rewrite nshiftl_size; auto. + simpl; rewrite nshiftl_S, IHn; auto. + Qed. + + Lemma firstr_firstl : + forall x, firstr x = firstl (nshiftl (pred size) x). + Proof. + destruct x; simpl; auto. + Qed. + + 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, + 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 -> + nshiftr n x = 0 -> nshiftr p x = 0. + Proof. + intros. + replace p with ((p-n)+n)%nat by omega. + induction (p-n)%nat. + simpl; auto. + simpl; rewrite nshiftr_S; rewrite IHn0; auto. + Qed. + + Lemma nshiftr_0_firstl : forall n x, n < size -> + nshiftr n x = 0 -> firstl x = D0. + Proof. + intros. + apply nshiftr_predsize_0_firstl. + apply nshiftr_0_propagates with n; auto; omega. + Qed. + + (** * Some induction principles over [int31] *) + + (** 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)) -> + forall x, P x. + Proof. + intros. + assert (forall n, n<=size -> P (nshiftr (size - n) x)). + induction n; intros. + rewrite nshiftr_size; auto. + rewrite sneakl_shiftr. + apply H0. + change (P (nshiftr (S (size - S n)) x)). + replace (S (size - S n))%nat with (size - n)%nat by omega. + apply IHn; omega. + 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)) -> + forall x, P x. + Proof. + induction x using int31_ind_sneakl; auto. + destruct d; auto. + Qed. + + + (** * 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 = + caserec (firstr x) (shiftr x) (recr_aux n A case0 caserec (shiftr x)). + Proof. + intros; simpl; rewrite H; auto. + Qed. + + Lemma recr_aux_converges : + forall n p x, n <= size -> n <= p -> + recr_aux n A case0 caserec (nshiftr (size - n) x) = + recr_aux p A case0 caserec (nshiftr (size - n) x). + Proof. + induction n. + simpl; intros. + rewrite nshiftr_size; destruct p; simpl; auto. + intros. + destruct p. + inversion H0. + unfold recr_aux; fold recr_aux. + destruct (iszero (nshiftr (size - S n) x)); auto. + f_equal. + change (shiftr (nshiftr (size - S n) x)) with (nshiftr (S (size - S n)) x). + replace (S (size - S n))%nat with (size - n)%nat by omega. + apply IHn; auto with arith. + Qed. + + 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. + unfold recr. + change x with (nshiftr (size - size) x). + rewrite (recr_aux_converges size (S size)); auto with arith. + rewrite recr_aux_eqn; auto. + Qed. + + (** [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) + (i:int31) : A := + match n with + | O => case0 + | S next => + 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. + + Lemma recrbis_aux_equiv : forall n x, + recrbis_aux n A case0 caserec x = recr_aux n A case0 caserec x. + Proof. + induction n; simpl; auto; intros. + case_eq (iszero x); intros; [ | f_equal; auto ]. + rewrite (iszero_eq0 _ H); simpl; auto. + replace (recrbis_aux n A case0 caserec 0) with case0; auto. + clear H IHn; induction n; simpl; congruence. + Qed. + + Lemma recrbis_equiv : forall x, + recrbis A case0 caserec x = recr A case0 caserec x. + Proof. + intros; apply recrbis_aux_equiv; auto. + Qed. + + End Recr. + + (** * Incrementation *) + + Section Incr. + + (** Variant of [incr] via [recrbis] *) + + Let Incr (b : digits) (si rec : int31) := + match b with + | D0 => sneakl D1 si + | D1 => sneakl D0 rec + end. + + Definition incrbis_aux n x := recrbis_aux n _ In Incr x. + + Lemma incrbis_aux_equiv : forall x, incrbis_aux size x = incr x. + Proof. + unfold incr, recr, incrbis_aux; fold Incr; intros. + apply recrbis_aux_equiv; auto. + Qed. + + (** Recursive equations satisfied by [incr] *) + + Lemma incr_eqn1 : + forall x, firstr x = D0 -> incr x = twice_plus_one (shiftr x). + Proof. + intros. + case_eq (iszero x); intros. + rewrite (iszero_eq0 _ H0); simpl; auto. + unfold incr; rewrite recr_eqn; fold incr; auto. + rewrite H; auto. + Qed. + + Lemma incr_eqn2 : + forall x, firstr x = D1 -> incr x = twice (incr (shiftr x)). + Proof. + intros. + case_eq (iszero x); intros. + rewrite (iszero_eq0 _ H0) in H; simpl in H; discriminate. + unfold incr; rewrite recr_eqn; fold incr; auto. + rewrite H; auto. + Qed. + + Lemma incr_twice : forall x, incr (twice x) = twice_plus_one x. + Proof. + intros. + rewrite incr_eqn1; destruct x; simpl; auto. + Qed. + + Lemma incr_twice_plus_one_firstl : + forall x, firstl x = D0 -> incr (twice_plus_one x) = twice (incr x). + Proof. + intros. + rewrite incr_eqn2; [ | destruct x; simpl; auto ]. + 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 + (see later). *) + + End Incr. + + (** * Conversion to [Z] : the [phi] function *) + + Section Phi. + + (** Variant of [phi] via [recrbis] *) + + 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. + Proof. + unfold phi, recr, phibis_aux; fold Phi; intros. + apply recrbis_aux_equiv; auto. + Qed. + + (** Recursive equations satisfied by [phi] *) + + Lemma phi_eqn1 : forall x, firstr x = D0 -> + phi x = Zdouble (phi (shiftr x)). + Proof. + intros. + case_eq (iszero x); intros. + rewrite (iszero_eq0 _ H0); simpl; auto. + intros; unfold phi; rewrite recr_eqn; fold phi; auto. + rewrite H; auto. + Qed. + + Lemma phi_eqn2 : forall x, firstr x = D1 -> + phi x = Zdouble_plus_one (phi (shiftr x)). + Proof. + intros. + case_eq (iszero x); intros. + rewrite (iszero_eq0 _ H0) in H; simpl in H; discriminate. + intros; unfold phi; rewrite recr_eqn; fold phi; auto. + rewrite H; auto. + Qed. + + Lemma phi_twice_firstl : forall x, firstl x = D0 -> + phi (twice x) = Zdouble (phi x). + Proof. + intros. + rewrite phi_eqn1; auto; [ | destruct x; auto ]. + f_equal; f_equal. + destruct x; simpl in *; rewrite H; auto. + Qed. + + Lemma phi_twice_plus_one_firstl : forall x, firstl x = D0 -> + phi (twice_plus_one x) = Zdouble_plus_one (phi x). + Proof. + intros. + rewrite phi_eqn2; auto; [ | destruct x; auto ]. + f_equal; f_equal. + destruct x; simpl in *; rewrite H; auto. + Qed. + + End Phi. + + (** [phi x] is positive and lower than [2^31] *) + + Lemma phibis_aux_pos : forall n x, (0 <= phibis_aux n x)%Z. + Proof. + induction n. + simpl; unfold phibis_aux; simpl; auto with zarith. + intros. + 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 -> + (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; + 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. + simpl; auto. + rewrite H0. + assert (H1 : n <= size) by omega. + specialize (IHn x H1). + set (y:=phibis_aux n (nshiftr (size - n) x)) in *. + rewrite inj_S, Zpower_Zsucc; auto with zarith. + case_eq (firstr (nshiftr (size - S n) x)); intros. + rewrite Zdouble_mult; auto with zarith. + rewrite Zdouble_plus_one_mult; auto with zarith. + Qed. + + Lemma phi_bounded : forall x, (0 <= phi x < 2 ^ (Z_of_nat size))%Z. + Proof. + intros. + rewrite <- phibis_aux_equiv. + split. + apply phibis_aux_pos. + change x with (nshiftr (size-size) x). + apply phibis_aux_bounded; auto. + Qed. + + 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. + intros. + unfold nshiftr in H; simpl in *. + unfold phibis_aux, recrbis_aux. + rewrite H, Zdouble_plus_one_mult; omega. + + intros. + remember (S n) as m. + unfold phibis_aux, recrbis_aux; fold recrbis_aux; + fold (phibis_aux m (shiftr x)). + subst m. + rewrite inj_S, Zpower_Zsucc; auto with zarith. + assert (2^(Z_of_nat n) <= phibis_aux (S n) (shiftr x))%Z. + apply IHn. + rewrite <- nshiftr_S_tail; auto. + destruct (firstr x). + 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 : + forall x, firstl x = D1 -> (2^(Z_of_nat (pred size)) <= phi x)%Z. + Proof. + intros. + generalize (phibis_aux_lowerbound (pred size) x). + rewrite <- firstl_firstr. + change (S (pred size)) with size; auto. + rewrite phibis_aux_equiv; auto. + Qed. + + (** * Equivalence modulo [2^n] *) + + Section EqShiftL. + + (** After killing [n] bits at the left, are the numbers equal ?*) + + Definition EqShiftL n x y := + nshiftl n x = nshiftl n y. + + Lemma EqShiftL_zero : forall x y, EqShiftL O x y <-> x = y. + Proof. + unfold EqShiftL; intros; unfold nshiftl; simpl; split; auto. + Qed. + + Lemma EqShiftL_size : forall k x y, size<=k -> EqShiftL k x y. + Proof. + red; intros; rewrite 2 nshiftl_above_size; auto. + Qed. + + Lemma EqShiftL_le : forall k k' x y, k <= k' -> + EqShiftL k x y -> EqShiftL k' x y. + Proof. + unfold EqShiftL; intros. + replace k' with ((k'-k)+k)%nat by omega. + remember (k'-k)%nat as n. + clear Heqn H k'. + induction n; simpl; auto. + rewrite 2 nshiftl_S; f_equal; auto. + Qed. + + 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. + unfold size. + auto with arith. + Qed. + + Lemma EqShiftL_twice : forall k x y, + EqShiftL k (twice x) (twice y) <-> EqShiftL (S k) x y. + Proof. + intros; unfold EqShiftL. + rewrite 2 nshiftl_S_tail; split; auto. + Qed. + + (** * From int31 to list of digits. *) + + (** Lower (=rightmost) bits comes first. *) + + Definition i2l := recrbis _ nil (fun d _ rec => d::rec). + + Lemma i2l_length : forall x, length (i2l x) = size. + Proof. + intros; reflexivity. + Qed. + + Fixpoint lshiftl l x := + match l with + | nil => x + | d::l => sneakl d (lshiftl l x) + end. + + Definition l2i l := lshiftl l On. + + Lemma l2i_i2l : forall x, l2i (i2l x) = x. + Proof. + destruct x; compute; auto. + Qed. + + 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, + i2l (sneakl d x) = d :: removelast (i2l x). + Proof. + destruct x; compute; auto. + Qed. + + Lemma i2l_l2i : forall l, length l = size -> + i2l (l2i l) = l. + Proof. + repeat (destruct l as [ |? l]; [intros; discriminate | ]). + destruct l; [ | intros; discriminate]. + intros _; compute; auto. + Qed. + + Fixpoint cstlist (A:Type)(a:A) n := + match n with + | O => nil + | S n => a::cstlist _ a n + end. + + Lemma i2l_nshiftl : forall n x, n<=size -> + i2l (nshiftl n x) = cstlist _ D0 n ++ firstn (size-n) (i2l x). + Proof. + induction n. + intros. + assert (firstn (size-0) (i2l x) = i2l x). + rewrite <- minus_n_O, <- (i2l_length x). + induction (i2l x); simpl; f_equal; auto. + rewrite H0; clear H0. + reflexivity. + + intros. + rewrite nshiftl_S. + unfold shiftl; rewrite i2l_sneakl. + simpl cstlist. + rewrite <- app_comm_cons; f_equal. + rewrite IHn; [ | omega]. + rewrite removelast_app. + f_equal. + replace (size-n)%nat with (S (size - S n))%nat by omega. + rewrite removelast_firstn; auto. + rewrite i2l_length; omega. + generalize (firstn_length (size-n) (i2l x)). + rewrite i2l_length. + intros H0 H1; rewrite H1 in H0. + rewrite min_l in H0 by omega. + simpl length in H0. + omega. + Qed. + + (** [i2l] can be used to define a relation equivalent to [EqShiftL] *) + + Lemma EqShiftL_i2l : forall k x y, + EqShiftL k x y <-> firstn (size-k) (i2l x) = firstn (size-k) (i2l y). + Proof. + intros. + destruct (le_lt_dec size k). + split; intros. + replace (size-k)%nat with O by omega. + unfold firstn; auto. + apply EqShiftL_size; auto. + + unfold EqShiftL. + assert (k <= size) by omega. + split; intros. + assert (i2l (nshiftl k x) = i2l (nshiftl k y)) by (f_equal; auto). + rewrite 2 i2l_nshiftl in H1; auto. + eapply app_inv_head; eauto. + assert (i2l (nshiftl k x) = i2l (nshiftl k y)). + rewrite 2 i2l_nshiftl; auto. + f_equal; auto. + rewrite <- (l2i_i2l (nshiftl k x)), <- (l2i_i2l (nshiftl k y)). + f_equal; auto. + Qed. + + (** This equivalence allows to prove easily the following delicate + result *) + + 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. + destruct (le_lt_dec size k). + split; intros; apply EqShiftL_size; auto. + + rewrite 2 EqShiftL_i2l. + unfold twice_plus_one. + rewrite 2 i2l_sneakl. + replace (size-k)%nat with (S (size - S k))%nat by omega. + remember (size - S k)%nat as n. + remember (i2l x) as lx. + remember (i2l y) as ly. + simpl. + rewrite 2 firstn_removelast. + split; intros. + injection H; auto. + f_equal; auto. + subst ly n; rewrite i2l_length; omega. + subst lx n; rewrite i2l_length; omega. + Qed. + + Lemma EqShiftL_shiftr : forall k x y, EqShiftL k x y -> + EqShiftL (S k) (shiftr x) (shiftr y). + Proof. + intros. + destruct (le_lt_dec size (S k)). + apply EqShiftL_size; auto. + case_eq (firstr x); intros. + rewrite <- EqShiftL_twice. + unfold twice; rewrite <- H0. + rewrite <- sneakl_shiftr. + rewrite (EqShiftL_firstr k x y); auto. + rewrite <- sneakl_shiftr; auto. + omega. + rewrite <- EqShiftL_twice_plus_one. + unfold twice_plus_one; rewrite <- H0. + rewrite <- sneakl_shiftr. + rewrite (EqShiftL_firstr k x y); auto. + rewrite <- sneakl_shiftr; auto. + omega. + Qed. + + Lemma EqShiftL_incrbis : forall n k x y, n<=size -> + (n+k=S size)%nat -> + 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). + subst k; apply EqShiftL_size; auto. + 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, + 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 : + forall x, incr (twice_plus_one x) = twice (incr x). + Proof. + intros. + rewrite incr_eqn2; [ | destruct x; simpl; auto]. + apply EqShiftL_incr. + red; destruct x; simpl; auto. + Qed. + + Lemma incr_firstr : forall x, firstr (incr x) <> firstr x. + Proof. + intros. + case_eq (firstr x); intros. + rewrite incr_eqn1; auto. + destruct (shiftr x); simpl; discriminate. + rewrite incr_eqn2; auto. + destruct (incr (shiftr x)); simpl; discriminate. + Qed. + + Lemma incr_inv : forall x y, + incr x = twice_plus_one y -> x = twice y. + Proof. + intros. + case_eq (iszero x); intros. + rewrite (iszero_eq0 _ H0) in *; simpl in *. + change (incr 0) with 1 in H. + symmetry; rewrite twice_zero; auto. + case_eq (firstr x); intros. + rewrite incr_eqn1 in H; auto. + clear H0; destruct x; destruct y; simpl in *. + injection H; intros; subst; auto. + elim (incr_firstr x). + rewrite H1, H; destruct y; simpl; auto. + Qed. + + (** * Conversion from [Z] : the [phi_inv] function *) + + (** First, recursive equations *) + + 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. + induction p; simpl. + rewrite 2 incr_twice; auto. + rewrite incr_twice, incr_twice_plus_one. + f_equal. + apply incr_inv; auto. + auto. + Qed. + + 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, + phi_inv (Zsucc z) = incr (phi_inv z). + Proof. + destruct z. + simpl; auto. + simpl; auto. + induction p; simpl; auto. + rewrite Pplus_one_succ_r, IHp, incr_twice_plus_one; auto. + rewrite incr_twice; auto. + simpl; auto. + destruct p; simpl; auto. + rewrite incr_twice; auto. + f_equal. + rewrite incr_twice_plus_one; auto. + induction p; simpl; auto. + rewrite incr_twice; auto. + f_equal. + rewrite incr_twice_plus_one; auto. + Qed. + + (** [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)) = + nshiftr (size-n) x. + Proof. + induction n. + intros; simpl. + rewrite nshiftr_size; auto. + intros. + 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. + rewrite H0. + case_eq (firstr (nshiftr (size - S n) x)); intros. + + rewrite phi_inv_double. + rewrite IHn by omega. + rewrite <- H0. + remember (nshiftr (size - S n) x) as y. + destruct y; simpl in H1; rewrite H1; auto. + + rewrite phi_inv_double_plus_one. + rewrite IHn by omega. + rewrite <- H0. + remember (nshiftr (size - S n) x) as y. + destruct y; simpl in H1; rewrite H1; auto. + Qed. + + Lemma phi_inv_phi : forall x, phi_inv (phi x) = x. + Proof. + intros. + rewrite <- phibis_aux_equiv. + replace x with (nshiftr (size - size) x) by auto. + apply phi_inv_phi_aux; auto. + Qed. + + (** The other composition [phi o phi_inv] is harder to prove correct. + In particular, an overflow can happen, so a modulo is needed. + For the moment, we proceed via several steps, the first one + being a detour to [positive_to_in31]. *) + + (** * [positive_to_int31] *) + + (** 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 := + match n with + | O => (Npos p, On) + | S n => match p with + | xO p => let (r,i) := p2ibis n p in (r, twice i) + | xI p => let (r,i) := p2ibis n p in (r, twice_plus_one i) + | xH => (N0, In) + end + end. + + Lemma p2ibis_bounded : forall n p, + nshiftr n (snd (p2ibis n p)) = 0. + Proof. + induction n. + simpl; intros; auto. + simpl; intros. + destruct p; simpl. + + specialize IHn with p. + destruct (p2ibis n p); simpl in *. + rewrite nshiftr_S_tail. + destruct (le_lt_dec size n). + rewrite nshiftr_above_size; auto. + assert (H:=nshiftr_0_firstl _ _ l IHn). + replace (shiftr (twice_plus_one i)) with i; auto. + destruct i; simpl in *; rewrite H; auto. + + specialize IHn with p. + destruct (p2ibis n p); simpl in *. + rewrite nshiftr_S_tail. + destruct (le_lt_dec size n). + rewrite nshiftr_above_size; auto. + assert (H:=nshiftr_0_firstl _ _ l IHn). + replace (shiftr (twice i)) with i; auto. + destruct i; simpl in *; rewrite H; auto. + + rewrite nshiftr_S_tail; auto. + 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) + + 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; + auto with zarith). + rewrite (Zmult_comm 2). + assert (n<=size) by omega. + destruct p; simpl; [ | | auto]; + specialize (IHn p H0); + generalize (p2ibis_bounded n p); + destruct (p2ibis n p) as (r,i); simpl in *; intros. + + change (Zpos p~1) with (2*Zpos p + 1)%Z. + rewrite phi_twice_plus_one_firstl, Zdouble_plus_one_mult. + rewrite IHn; ring. + apply (nshiftr_0_firstl n); auto; try omega. + + change (Zpos p~0) with (2*Zpos p)%Z. + rewrite phi_twice_firstl. + change (Zdouble (phi i)) with (2*(phi i))%Z. + rewrite IHn; ring. + apply (nshiftr_0_firstl n); auto; try omega. + Qed. + + (** We now prove that this [p2ibis] is related to [phi_inv_positive] *) + + 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; + 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, + phi (phi_inv_positive p) = (Zpos p) mod (2^(Z_of_nat size)). + Proof. + intros. + replace (phi_inv_positive p) with (snd (p2ibis size p)). + rewrite (p2ibis_spec size p) by auto. + rewrite Zplus_comm, Z_mod_plus. + symmetry; apply Zmod_small. + apply phi_bounded. + auto with zarith. + symmetry. + rewrite <- EqShiftL_zero. + apply (phi_inv_positive_p2ibis size p); auto. + Qed. + + (** Moreover, [p2ibis] is also related with [p2i] and hence with + [positive_to_int31]. *) + + Lemma double_twice_firstl : forall x, firstl x = D0 -> Twon*x = twice x. + Proof. + 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 -> + Twon*x+In = twice_plus_one x. + Proof. + intros. + rewrite double_twice_firstl; auto. + unfold add31. + 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 -> + 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; + f_equal; auto. + apply double_twice_plus_one_firstl. + apply (nshiftr_0_firstl n); auto; omega. + apply double_twice_firstl. + apply (nshiftr_0_firstl n); auto; omega. + Qed. + + Lemma positive_to_int31_phi_inv_positive : forall p, + snd (positive_to_int31 p) = phi_inv_positive p. + Proof. + intros; unfold positive_to_int31. + rewrite p2i_p2ibis; auto. + symmetry. + rewrite <- EqShiftL_zero. + 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) + + phi (snd (positive_to_int31 p)))%Z. + Proof. + unfold positive_to_int31. + intros; rewrite p2i_p2ibis; auto. + 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 + [phi o twice] and so one. *) + + Lemma phi_twice : forall x, + phi (twice x) = (Zdouble (phi x)) mod 2^(Z_of_nat size). + Proof. + intros. + pattern x at 1; rewrite <- (phi_inv_phi x). + rewrite <- phi_inv_double. + assert (0 <= Zdouble (phi x))%Z. + rewrite Zdouble_mult; generalize (phi_bounded x); omega. + destruct (Zdouble (phi x)). + simpl; auto. + apply phi_phi_inv_positive. + compute in H; elim H; auto. + Qed. + + Lemma phi_twice_plus_one : forall x, + phi (twice_plus_one x) = (Zdouble_plus_one (phi x)) mod 2^(Z_of_nat size). + Proof. + intros. + pattern x at 1; rewrite <- (phi_inv_phi x). + rewrite <- phi_inv_double_plus_one. + assert (0 <= Zdouble_plus_one (phi x))%Z. + rewrite Zdouble_plus_one_mult; generalize (phi_bounded x); omega. + destruct (Zdouble_plus_one (phi x)). + simpl; auto. + apply phi_phi_inv_positive. + compute in H; elim H; auto. + Qed. + + Lemma phi_incr : forall x, + phi (incr x) = (Zsucc (phi x)) mod 2^(Z_of_nat size). + 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; + generalize (phi_bounded x); omega. + destruct (Zsucc (phi x)). + simpl; auto. + apply phi_phi_inv_positive. + compute in H; elim H; auto. + Qed. + + (** With the previous results, we can deal with [phi o phi_inv] even + in the negative case *) + + Lemma phi_phi_inv_negative : + forall p, phi (incr (complement_negative p)) = (Zneg p) mod 2^(Z_of_nat size). + Proof. + induction p. + + simpl complement_negative. + rewrite phi_incr in IHp. + rewrite incr_twice, phi_twice_plus_one. + remember (phi (complement_negative p)) as q. + rewrite Zdouble_plus_one_mult. + replace (2*q+1)%Z with (2*(Zsucc q)-1)%Z by omega. + rewrite <- Zminus_mod_idemp_l, <- Zmult_mod_idemp_r, IHp. + rewrite Zmult_mod_idemp_r, Zminus_mod_idemp_l; auto with zarith. + + simpl complement_negative. + rewrite incr_twice_plus_one, phi_twice. + remember (phi (incr (complement_negative p))) as q. + rewrite Zdouble_mult, IHp, Zmult_mod_idemp_r; auto with zarith. + + simpl; auto. + Qed. + + Lemma phi_phi_inv : + forall z, phi (phi_inv z) = z mod 2 ^ (Z_of_nat size). + Proof. + destruct z. + simpl; auto. + apply phi_phi_inv_positive. + apply phi_phi_inv_negative. + Qed. + +End Basics. + + +Section Int31_Op. + +(** Nullity test *) +Let w_iszero i := match i ?= 0 with Eq => true | _ => false end. + +(** Modulo [2^p] *) +Let w_pos_mod p i := + match compare31 p 31 with + | Lt => addmuldiv31 p 0 (addmuldiv31 (31-p) i 0) + | _ => i + end. + +(** Parity test *) +Let w_iseven i := + let (_,r) := i/2 in + match r ?= 0 with Eq => true | _ => false end. + +Definition int31_op := (mk_znz_op + 31%positive (* number of digits *) + 31 (* number of digits *) + phi (* conversion to Z *) + positive_to_int31 (* positive -> N*int31 : p => N,i where p = N*2^31+phi i *) + head031 (* number of head 0 *) + tail031 (* number of tail 0 *) + (* Basic constructors *) + 0 + 1 + Tn (* 2^31 - 1 *) + (* Comparison *) + compare31 + w_iszero + (* Basic arithmetic operations *) + (fun i => 0 -c i) + (fun i => 0 - i) + (fun i => 0-i-1) + (fun i => i +c 1) + add31c + add31carryc + (fun i => i + 1) + add31 + (fun i j => i + j + 1) + (fun i => i -c 1) + sub31c + sub31carryc + (fun i => i - 1) + sub31 + (fun i j => i - j - 1) + mul31c + mul31 + (fun x => x *c x) + (* special (euclidian) division operations *) + div3121 + div31 (* this is supposed to be the special case of division a/b where a > b *) + div31 + (* euclidian division remainder *) + (* again special case for a > b *) + (fun i j => let (_,r) := i/j in r) + (fun i j => let (_,r) := i/j in r) + gcd31 (*gcd_gt*) + gcd31 (*gcd*) + (* shift operations *) + addmuldiv31 (*add_mul_div *) + (* modulo 2^p *) + w_pos_mod + (* is i even ? *) + w_iseven + (* square root operations *) + sqrt312 (* sqrt2 *) + sqrt31 (* sqrt *) +). + +End Int31_Op. + +Section Int31_Spec. + + Open Local 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. + Proof. + auto with zarith. + Qed. + + Notation "[+| c |]" := + (interp_carry 1 wB phi c) (at level 0, x at level 99). + + Notation "[-| c |]" := + (interp_carry (-1) wB phi c) (at level 0, x at level 99). + + Notation "[|| x ||]" := + (zn2z_to_Z wB phi x) (at level 0, x at level 99). + + Lemma spec_zdigits : [| 31 |] = 31. + Proof. + reflexivity. + Qed. + + Lemma spec_more_than_1_digit: 1 < 31. + Proof. + auto with zarith. + Qed. + + Lemma spec_0 : [| 0 |] = 0. + Proof. + reflexivity. + Qed. + + Lemma spec_1 : [| 1 |] = 1. + Proof. + reflexivity. + Qed. + + Lemma spec_Bm1 : [| Tn |] = wB - 1. + Proof. + reflexivity. + Qed. + + Lemma spec_compare : forall x y, + match (x ?= y)%int31 with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Proof. + clear; unfold compare31; simpl; intros. + case_eq ([|x|] ?= [|y|]); auto. + intros; apply Zcompare_Eq_eq; auto. + Qed. + + (** Addition *) + + Lemma spec_add_c : forall x y, [+|add31c x y|] = [|x|] + [|y|]. + Proof. + intros; unfold add31c, add31, interp_carry; rewrite phi_phi_inv. + generalize (phi_bounded x)(phi_bounded y); intros. + set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y. + + assert ((X+Y) mod wB ?= X+Y <> Eq -> [+|C1 (phi_inv (X+Y))|] = X+Y). + unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros. + destruct (Z_lt_le_dec (X+Y) wB). + contradict H1; auto using Zmod_small with zarith. + rewrite <- (Z_mod_plus_full (X+Y) (-1) wB). + rewrite Zmod_small; romega. + + generalize (Zcompare_Eq_eq ((X+Y) mod wB) (X+Y)); intros Heq. + destruct Zcompare; intros; + [ 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. + Qed. + + Lemma spec_add_carry_c : forall x y, [+|add31carryc x y|] = [|x|] + [|y|] + 1. + Proof. + intros. + unfold add31carryc, interp_carry; rewrite phi_phi_inv. + generalize (phi_bounded x)(phi_bounded y); intros. + set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y. + + assert ((X+Y+1) mod wB ?= X+Y+1 <> Eq -> [+|C1 (phi_inv (X+Y+1))|] = X+Y+1). + unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros. + destruct (Z_lt_le_dec (X+Y+1) wB). + contradict H1; auto using Zmod_small with zarith. + rewrite <- (Z_mod_plus_full (X+Y+1) (-1) wB). + rewrite Zmod_small; romega. + + generalize (Zcompare_Eq_eq ((X+Y+1) mod wB) (X+Y+1)); intros Heq. + destruct Zcompare; intros; + [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. + Qed. + + Lemma spec_add : forall x y, [|x+y|] = ([|x|] + [|y|]) mod wB. + Proof. + intros; apply phi_phi_inv. + Qed. + + Lemma spec_add_carry : + forall x y, [|x+y+1|] = ([|x|] + [|y|] + 1) mod wB. + Proof. + unfold add31; intros. + repeat rewrite phi_phi_inv. + apply Zplus_mod_idemp_l. + Qed. + + Lemma spec_succ : forall x, [|x+1|] = ([|x|] + 1) mod wB. + Proof. + intros; rewrite <- spec_1; apply spec_add. + Qed. + + (** Substraction *) + + Lemma spec_sub_c : forall x y, [-|sub31c x y|] = [|x|] - [|y|]. + Proof. + unfold sub31c, sub31, interp_carry; intros. + rewrite phi_phi_inv. + generalize (phi_bounded x)(phi_bounded y); intros. + set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y. + + assert ((X-Y) mod wB ?= X-Y <> Eq -> [-|C1 (phi_inv (X-Y))|] = X-Y). + unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros. + destruct (Z_lt_le_dec (X-Y) 0). + rewrite <- (Z_mod_plus_full (X-Y) 1 wB). + rewrite Zmod_small; romega. + contradict H1; apply Zmod_small; romega. + + generalize (Zcompare_Eq_eq ((X-Y) mod wB) (X-Y)); intros Heq. + destruct Zcompare; intros; + [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. + Qed. + + Lemma spec_sub_carry_c : forall x y, [-|sub31carryc x y|] = [|x|] - [|y|] - 1. + Proof. + unfold sub31carryc, sub31, interp_carry; intros. + rewrite phi_phi_inv. + generalize (phi_bounded x)(phi_bounded y); intros. + set (X:=[|x|]) in *; set (Y:=[|y|]) in *; clearbody X Y. + + assert ((X-Y-1) mod wB ?= X-Y-1 <> Eq -> [-|C1 (phi_inv (X-Y-1))|] = X-Y-1). + unfold interp_carry; rewrite phi_phi_inv, Zcompare_Eq_iff_eq; intros. + destruct (Z_lt_le_dec (X-Y-1) 0). + rewrite <- (Z_mod_plus_full (X-Y-1) 1 wB). + rewrite Zmod_small; romega. + contradict H1; apply Zmod_small; romega. + + generalize (Zcompare_Eq_eq ((X-Y-1) mod wB) (X-Y-1)); intros Heq. + destruct Zcompare; intros; + [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. + Qed. + + Lemma spec_sub : forall x y, [|x-y|] = ([|x|] - [|y|]) mod wB. + Proof. + intros; apply phi_phi_inv. + Qed. + + Lemma spec_sub_carry : + forall x y, [|x-y-1|] = ([|x|] - [|y|] - 1) mod wB. + Proof. + unfold sub31; intros. + repeat rewrite phi_phi_inv. + apply Zminus_mod_idemp_l. + Qed. + + Lemma spec_opp_c : forall x, [-|sub31c 0 x|] = -[|x|]. + Proof. + intros; apply spec_sub_c. + Qed. + + Lemma spec_opp : forall x, [|0 - x|] = (-[|x|]) mod wB. + Proof. + intros; apply phi_phi_inv. + Qed. + + Lemma spec_opp_carry : forall x, [|0 - x - 1|] = wB - [|x|] - 1. + Proof. + unfold sub31; intros. + repeat rewrite phi_phi_inv. + change [|1|] with 1; change [|0|] with 0. + rewrite <- (Z_mod_plus_full (0-[|x|]) 1 wB). + rewrite Zminus_mod_idemp_l. + rewrite Zmod_small; generalize (phi_bounded x); romega. + Qed. + + Lemma spec_pred_c : forall x, [-|sub31c x 1|] = [|x|] - 1. + Proof. + intros; apply spec_sub_c. + Qed. + + Lemma spec_pred : forall x, [|x-1|] = ([|x|] - 1) mod wB. + Proof. + intros; apply spec_sub. + Qed. + + (** Multiplication *) + + Lemma phi2_phi_inv2 : forall x, [||phi_inv2 x||] = x mod (wB^2). + Proof. + assert (forall z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2). + intros. + assert ((z/wB) mod wB = z/wB - (z/wB/wB)*wB). + rewrite (Z_div_mod_eq (z/wB) wB wB_pos) at 2; ring. + assert (z mod wB = z - (z/wB)*wB). + rewrite (Z_div_mod_eq z wB wB_pos) at 2; ring. + rewrite H. + rewrite H0 at 1. + ring_simplify. + rewrite Zdiv_Zdiv; auto with zarith. + rewrite (Z_div_mod_eq z (wB*wB)) at 2; auto with zarith. + change (wB*wB) with (wB^2); ring. + + unfold phi_inv2. + destruct x; unfold zn2z_to_Z; rewrite ?phi_phi_inv; + change base with wB; auto. + Qed. + + Lemma spec_mul_c : forall x y, [|| mul31c x y ||] = [|x|] * [|y|]. + Proof. + unfold mul31c; intros. + rewrite phi2_phi_inv2. + apply Zmod_small. + generalize (phi_bounded x)(phi_bounded y); intros. + change (wB^2) with (wB * wB). + auto using Zmult_lt_compat with zarith. + Qed. + + Lemma spec_mul : forall x y, [|x*y|] = ([|x|] * [|y|]) mod wB. + Proof. + intros; apply phi_phi_inv. + Qed. + + Lemma spec_square_c : forall x, [|| mul31c x x ||] = [|x|] * [|x|]. + Proof. + intros; apply spec_mul_c. + Qed. + + (** Division *) + + Lemma spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := div3121 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + unfold div3121; intros. + generalize (phi_bounded a1)(phi_bounded a2)(phi_bounded b); intros. + assert ([|b|]>0) by (auto with zarith). + generalize (Z_div_mod (phi2 a1 a2) [|b|] H4) (Z_div_pos (phi2 a1 a2) [|b|] H4). + unfold Zdiv; destruct (Zdiv_eucl (phi2 a1 a2) [|b|]); simpl. + rewrite ?phi_phi_inv. + destruct 1; intros. + unfold phi2 in *. + change base with wB; change base with wB in H5. + change (Zpower_pos 2 31) with wB; change (Zpower_pos 2 31) with wB in H. + rewrite H5, Zmult_comm. + replace (z0 mod wB) with z0 by (symmetry; apply Zmod_small; omega). + replace (z mod wB) with z; auto with zarith. + symmetry; apply Zmod_small. + split. + apply H7; change base with wB; auto with zarith. + apply Zmult_gt_0_lt_reg_r with [|b|]. + omega. + rewrite Zmult_comm. + apply Zle_lt_trans with ([|b|]*z+z0). + omega. + rewrite <- H5. + apply Zle_lt_trans with ([|a1|]*wB+(wB-1)). + omega. + replace ([|a1|]*wB+(wB-1)) with (wB*([|a1|]+1)-1) by ring. + assert (wB*([|a1|]+1) <= wB*[|b|]); try omega. + apply Zmult_le_compat; omega. + Qed. + + Lemma spec_div : forall a b, 0 < [|b|] -> + let (q,r) := div31 a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + unfold div31; intros. + assert ([|b|]>0) by (auto with zarith). + generalize (Z_div_mod [|a|] [|b|] H0) (Z_div_pos [|a|] [|b|] H0). + unfold Zdiv; destruct (Zdiv_eucl [|a|] [|b|]); simpl. + rewrite ?phi_phi_inv. + destruct 1; intros. + rewrite H1, Zmult_comm. + generalize (phi_bounded a)(phi_bounded b); intros. + replace (z0 mod wB) with z0 by (symmetry; apply Zmod_small; omega). + replace (z mod wB) with z; auto with zarith. + symmetry; apply Zmod_small. + split; auto with zarith. + apply Zle_lt_trans with [|a|]; auto with zarith. + rewrite H1. + apply Zle_trans with ([|b|]*z); try omega. + rewrite <- (Zmult_1_l z) at 1. + apply Zmult_le_compat; auto with zarith. + Qed. + + Lemma spec_mod : forall a b, 0 < [|b|] -> + [|let (_,r) := (a/b)%int31 in r|] = [|a|] mod [|b|]. + Proof. + unfold div31; intros. + assert ([|b|]>0) by (auto with zarith). + unfold Zmod. + generalize (Z_div_mod [|a|] [|b|] H0). + destruct (Zdiv_eucl [|a|] [|b|]); simpl. + rewrite ?phi_phi_inv. + destruct 1; intros. + generalize (phi_bounded b); intros. + apply Zmod_small; omega. + Qed. + + Lemma phi_gcd : forall i j, + [|gcd31 i j|] = Zgcdn (2*size) [|j|] [|i|]. + Proof. + unfold gcd31. + induction (2*size)%nat; intros. + reflexivity. + simpl. + unfold compare31. + change [|On|] with 0. + generalize (phi_bounded j)(phi_bounded i); intros. + case_eq [|j|]; intros. + simpl; intros. + generalize (Zabs_spec [|i|]); omega. + simpl. + rewrite IHn, H1; f_equal. + rewrite spec_mod, H1; auto. + rewrite H1; compute; auto. + rewrite H1 in H; destruct H as [H _]; compute in H; elim H; auto. + Qed. + + Lemma spec_gcd : forall a b, Zis_gcd [|a|] [|b|] [|gcd31 a b|]. + Proof. + intros. + rewrite phi_gcd. + apply Zis_gcd_sym. + apply Zgcdn_is_gcd. + unfold Zgcd_bound. + generalize (phi_bounded b). + destruct [|b|]. + unfold size; auto with zarith. + intros (_,H). + cut (Psize p <= size)%nat; [ omega | rewrite <- Zpower2_Psize; auto]. + intros (H,_); compute in H; elim H; auto. + Qed. + + Lemma iter_int31_iter_nat : forall A f i a, + iter_int31 i A f a = iter_nat (Zabs_nat [|i|]) A f a. + Proof. + intros. + unfold iter_int31. + rewrite <- recrbis_equiv; auto; unfold recrbis. + rewrite <- phibis_aux_equiv. + + revert i a; induction size. + simpl; auto. + simpl; intros. + 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; + 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 = + iter_nat (Zabs_nat (Zdouble_plus_one z)) A f a); f_equal. + rewrite Zdouble_plus_one_mult, Zmult_comm, <- Zplus_diag_eq_mult_2. + rewrite Zabs_nat_Zplus; auto with zarith. + rewrite Zabs_nat_Zplus; auto with zarith. + change (Zabs_nat 1) with 1%nat; omega. + Qed. + + 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, + addmuldiv31 p x y = addmuldiv31_alt (Zabs_nat [|p|]) x y. + Proof. + intros. + unfold addmuldiv31. + rewrite iter_int31_iter_nat. + set (n:=Zabs_nat [|p|]); clearbody n; clear p. + revert x y; induction n. + simpl; auto. + intros. + simpl addmuldiv31_alt. + replace (S n) with (n+1)%nat by (rewrite plus_comm; auto). + rewrite iter_nat_plus; simpl; auto. + Qed. + + Lemma spec_add_mul_div : forall x y p, [|p|] <= Zpos 31 -> + [| addmuldiv31 p x y |] = + ([|x|] * (2 ^ [|p|]) + [|y|] / (2 ^ ((Zpos 31) - [|p|]))) mod wB. + Proof. + intros. + rewrite addmuldiv31_equiv. + assert ([|p|] = Z_of_nat (Zabs_nat [|p|])). + rewrite inj_Zabs_nat; symmetry; apply Zabs_eq. + destruct (phi_bounded p); auto. + rewrite H0; rewrite H0 in H; clear H0; rewrite Zabs_nat_Z_of_nat. + set (n := Zabs_nat [|p|]) in *; clearbody n. + assert (n <= 31)%nat. + rewrite inj_le_iff; auto with zarith. + clear p H; revert x y. + + induction n. + simpl; intros. + change (Zpower_pos 2 31) with (2^31). + rewrite Zmult_1_r. + replace ([|y|] / 2^31) with 0. + rewrite Zplus_0_r. + symmetry; apply Zmod_small; apply phi_bounded. + symmetry; apply Zdiv_small; apply phi_bounded. + + simpl addmuldiv31_alt; intros. + rewrite IHn; [ | omega ]. + case_eq (firstl y); intros. + + rewrite phi_twice, Zdouble_mult. + rewrite phi_twice_firstl; auto. + change (Zdouble [|y|]) with (2*[|y|]). + rewrite inj_S, Zpower_Zsucc; auto with zarith. + rewrite Zplus_mod; rewrite Zmult_mod_idemp_l; rewrite <- Zplus_mod. + f_equal. + apply Zplus_eq_compat. + ring. + replace (31-Z_of_nat n) with (Zsucc(31-Zsucc(Z_of_nat n))) by ring. + rewrite Zpower_Zsucc, <- Zdiv_Zdiv; auto with zarith. + rewrite Zmult_comm, Z_div_mult; auto with zarith. + + rewrite phi_twice_plus_one, Zdouble_plus_one_mult. + rewrite phi_twice; auto. + change (Zdouble [|y|]) with (2*[|y|]). + rewrite inj_S, Zpower_Zsucc; auto with zarith. + rewrite Zplus_mod; rewrite Zmult_mod_idemp_l; rewrite <- Zplus_mod. + rewrite Zmult_plus_distr_l, Zmult_1_l, <- Zplus_assoc. + f_equal. + apply Zplus_eq_compat. + ring. + assert ((2*[|y|]) mod wB = 2*[|y|] - wB). + admit. + rewrite H1. + 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. + ring_simplify. + replace (31+-Z_of_nat n) with (Zsucc(31-Zsucc(Z_of_nat n))) by ring. + rewrite Zpower_Zsucc, <- Zdiv_Zdiv; auto with zarith. + rewrite Zmult_comm, Z_div_mult; auto with zarith. + Qed. + + Let w_pos_mod := int31_op.(znz_pos_mod). + + Lemma spec_pos_mod : forall w p, + [|w_pos_mod p w|] = [|w|] mod (2 ^ [|p|]). + Proof. + unfold w_pos_mod, znz_pos_mod, int31_op, compare31. + change [|31|] with 31%Z. + assert (forall w p, 31<=p -> [|w|] = [|w|] mod 2^p). + intros. + generalize (phi_bounded w). + symmetry; apply Zmod_small. + split; auto with zarith. + 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 | | + apply H; change ([|p|]>31)%Z in H0; auto with zarith ]. + change ([|p|]<31) in H0. + rewrite spec_add_mul_div by auto with zarith. + change [|0|] with 0%Z; rewrite Zmult_0_l, Zplus_0_l. + generalize (phi_bounded p)(phi_bounded w); intros. + assert (31-[|p|]<wB). + apply Zle_lt_trans with 31%Z; auto with zarith. + compute; auto. + assert ([|31-p|]=31-[|p|]). + unfold sub31; rewrite phi_phi_inv. + change [|31|] with 31%Z. + apply Zmod_small; auto with zarith. + rewrite spec_add_mul_div by (rewrite H4; auto with zarith). + change [|0|] with 0%Z; rewrite Zdiv_0_l, Zplus_0_r. + rewrite H4. + apply shift_unshift_mod_2; auto with zarith. + Qed. + + + (** Shift operations *) + + Lemma spec_head00: forall x, [|x|] = 0 -> [|head031 x|] = Zpos 31. + Proof. + intros. + generalize (phi_inv_phi x). + rewrite H; simpl. + intros H'; rewrite <- H'. + simpl; auto. + Qed. + + Fixpoint head031_alt n x := + match n with + | O => 0%nat + | S n => match firstl x with + | D0 => S (head031_alt n (shiftl x)) + | D1 => 0%nat + end + end. + + Lemma head031_equiv : + forall x, [|head031 x|] = Z_of_nat (head031_alt size x). + Proof. + intros. + case_eq (iszero x); intros. + rewrite (iszero_eq0 _ H). + simpl; auto. + + unfold head031, recl. + change On with (phi_inv (Z_of_nat (31-size))). + 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. + unfold head031_alt; fold head031_alt. + rewrite H. + assert ([|phi_inv (Z_of_nat (31-S n))|] = Z_of_nat (31 - S n)). + rewrite phi_phi_inv. + apply Zmod_small. + split. + change 0 with (Z_of_nat O); apply inj_le; omega. + apply Zle_lt_trans with (Z_of_nat 31). + apply inj_le; omega. + compute; auto. + case_eq (firstl x); intros; auto. + rewrite plus_Sn_m, plus_n_Sm. + replace (S (31 - S n)) with (31 - n)%nat by omega. + rewrite <- IHn; [ | omega | ]. + f_equal; f_equal. + unfold add31. + rewrite H1. + f_equal. + 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. + case_eq (iszero (shiftl x)); intros; auto. + rewrite (iszero_eq0 _ H0) in H; discriminate. + Qed. + + Lemma phi_nz : forall x, 0 < [|x|] <-> x <> 0%int31. + Proof. + split; intros. + red; intro; subst x; discriminate. + assert ([|x|]<>0%Z). + contradict H. + rewrite <- (phi_inv_phi x); rewrite H; auto. + generalize (phi_bounded x); auto with zarith. + Qed. + + Lemma spec_head0 : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ ([|head031 x|]) * [|x|] < wB. + Proof. + intros. + rewrite head031_equiv. + assert (nshiftl size x = 0%int31). + apply nshiftl_size. + revert x H H0. + unfold size at 2 5. + induction size. + simpl Z_of_nat. + intros. + compute in H0; rewrite H0 in H; discriminate. + + intros. + simpl head031_alt. + case_eq (firstl x); intros. + rewrite (inj_S (head031_alt n (shiftl x))), Zpower_Zsucc; auto with zarith. + rewrite <- Zmult_assoc, Zmult_comm, <- Zmult_assoc, <-(Zmult_comm 2). + rewrite <- Zdouble_mult, <- (phi_twice_firstl _ H1). + apply IHn. + + rewrite phi_nz; rewrite phi_nz in H; contradict H. + change twice with shiftl in H. + 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))). + apply phi_lowerbound; auto. + Qed. + + Lemma spec_tail00: forall x, [|x|] = 0 -> [|tail031 x|] = Zpos 31. + Proof. + intros. + generalize (phi_inv_phi x). + rewrite H; simpl. + intros H'; rewrite <- H'. + simpl; auto. + Qed. + + Fixpoint tail031_alt n x := + match n with + | O => 0%nat + | S n => match firstr x with + | D0 => S (tail031_alt n (shiftr x)) + | D1 => 0%nat + end + end. + + Lemma tail031_equiv : + forall x, [|tail031 x|] = Z_of_nat (tail031_alt size x). + Proof. + intros. + case_eq (iszero x); intros. + rewrite (iszero_eq0 _ H). + simpl; auto. + + unfold tail031, recr. + change On with (phi_inv (Z_of_nat (31-size))). + 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. + unfold tail031_alt; fold tail031_alt. + rewrite H. + assert ([|phi_inv (Z_of_nat (31-S n))|] = Z_of_nat (31 - S n)). + rewrite phi_phi_inv. + apply Zmod_small. + split. + change 0 with (Z_of_nat O); apply inj_le; omega. + apply Zle_lt_trans with (Z_of_nat 31). + apply inj_le; omega. + compute; auto. + case_eq (firstr x); intros; auto. + rewrite plus_Sn_m, plus_n_Sm. + replace (S (31 - S n)) with (31 - n)%nat by omega. + rewrite <- IHn; [ | omega | ]. + f_equal; f_equal. + unfold add31. + rewrite H1. + f_equal. + 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. + case_eq (iszero (shiftr x)); intros; auto. + rewrite (iszero_eq0 _ H0) in H; discriminate. + Qed. + + Lemma spec_tail0 : forall x, 0 < [|x|] -> + exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|tail031 x|]). + Proof. + intros. + rewrite tail031_equiv. + assert (nshiftr size x = 0%int31). + apply nshiftr_size. + revert x H H0. + induction size. + simpl Z_of_nat. + intros. + compute in H0; rewrite H0 in H; discriminate. + + intros. + simpl tail031_alt. + case_eq (firstr x); intros. + rewrite (inj_S (tail031_alt n (shiftr x))), Zpower_Zsucc; auto with zarith. + 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 + of a fortran program in boyer-moore *) + + Lemma quotient_by_2 a: a - 1 <= (a/2) + (a/2). + Proof. + intros a; 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 -> + (j * k) + j <= ((j + k)/2 + 1) ^ 2. + Proof. + intros j k 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)); + 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; + 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. + rewrite Zplus_comm, <- Z_div_plus_full_l; auto with zarith. + apply f_equal2 with (f := Zdiv); auto with zarith. + Qed. + + Lemma sqrt_main i j: 0 <= i -> 0 < j -> i < ((j + (i/j))/2 + 1) ^ 2. + Proof. + intros i j 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. + Qed. + + Lemma sqrt_init i: 1 < i -> i < (i/2 + 1) ^ 2. + Proof. + intros i 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. + rewrite Zpower_2, Z_div_plus_full_l; auto with zarith. + generalize (sqr_pos ((i - 2)/ 2)) (Z_div_pos (i - 2) 2). + rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r. + auto with zarith. + generalize (quotient_by_2 i). + rewrite Zpower_2 in H2 |- *; + repeat (rewrite Zmult_plus_distr_l || + rewrite Zmult_plus_distr_r || + rewrite Zmult_1_l || rewrite Zmult_1_r). + auto with zarith. + Qed. + + 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. + 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 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. + apply Z_mult_div_ge; auto with zarith. + Qed. + + (* George's trick *) + Inductive ZcompareSpec (i j: Z): comparison -> Prop := + ZcompareSpecEq: i = j -> ZcompareSpec i j Eq + | ZcompareSpecLt: i < j -> ZcompareSpec i j Lt + | ZcompareSpecGt: j < i -> ZcompareSpec i j Gt. + + Lemma Zcompare_spec i j: ZcompareSpec i j (i ?= j). + Proof. + intros i j; case_eq (Zcompare i j); intros H. + apply ZcompareSpecEq; apply Zcompare_Eq_eq; auto. + apply ZcompareSpecLt; auto. + apply ZcompareSpecGt; apply Zgt_lt; auto. + Qed. + + Lemma sqrt31_step_def rec i j: + sqrt31_step rec i j = + match (fst (i/j) ?= j)%int31 with + Lt => rec i (fst ((j + fst(i/j))/2))%int31 + | _ => j + end. + Proof. + intros rec i j; 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). + 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 -> + 2 * [|j|] < wB -> + (forall j1 : int31, + 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> + [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> + [|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2. + Proof. + assert (Hp2: 0 < [|2|]) by exact (refl_equal Lt). + intros rec i j 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 + (1 * 2 + (([|j|] - 2) + [|i|] / [|j|])); try ring. + rewrite Z_div_plus_full_l; auto with zarith. + assert (0 <= [|i|]/ [|j|]) by (apply Z_div_pos; auto with zarith). + assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / [|2|]) ; auto with zarith. + rewrite <- Hj1, Zdiv_1_r. + replace (1 + [|i|])%Z with (1 * 2 + ([|i|] - 1))%Z; try ring. + rewrite Z_div_plus_full_l; auto with zarith. + assert (0 <= ([|i|] - 1) /2)%Z by (apply Z_div_pos; auto with zarith). + change ([|2|]) with 2%Z; auto with zarith. + apply sqrt_test_false; auto with zarith. + rewrite spec_add, div31_phi; auto. + apply sym_equal; apply Zmod_small. + split; auto with zarith. + replace [|j + fst (i / j)%int31|] with ([|j|] + [|i|] / [|j|]). + apply sqrt_main; auto with zarith. + rewrite spec_add, div31_phi; auto. + apply sym_equal; apply Zmod_small. + split; auto with zarith. + Qed. + + Lemma iter31_sqrt_correct n rec i j: 0 < [|i|] -> 0 < [|j|] -> + [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < 2 ^ (Z_of_nat size) -> + (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] -> + [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < 2 ^ (Z_of_nat size) -> + [|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. + 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. + intros n Hrec rec i j Hi Hj Hij H31 HHrec. + apply sqrt31_step_correct; auto. + intros j1 Hj1 Hjp1; apply Hrec; auto with zarith. + intros j2 Hj2 H2j2 Hjp2 Hj31; apply Hrec; auto with zarith. + intros j3 Hj3 Hpj3. + apply HHrec; auto. + rewrite inj_S, Zpower_Zsucc. + apply Zle_trans with (2 ^Z_of_nat n + [|j2|]); auto with zarith. + apply Zle_0_nat. + Qed. + + Lemma spec_sqrt : forall x, + [|sqrt31 x|] ^ 2 <= [|x|] < ([|sqrt31 x|] + 1) ^ 2. + Proof. + intros i; unfold sqrt31. + generalize (spec_compare 1 i); case compare31; change [|1|] with 1; + intros Hi; auto with zarith. + repeat rewrite Zpower_2; auto with zarith. + apply iter31_sqrt_correct; auto with zarith. + rewrite div31_phi; change ([|2|]) with 2; auto with zarith. + replace ([|i|]) with (1 * 2 + ([|i|] - 2))%Z; try ring. + assert (0 <= ([|i|] - 2)/2)%Z by (apply Z_div_pos; auto with zarith). + rewrite Z_div_plus_full_l; auto with zarith. + rewrite div31_phi; change ([|2|]) with 2; auto with zarith. + apply sqrt_init; auto. + rewrite div31_phi; change ([|2|]) with 2; auto with zarith. + apply Zle_lt_trans with ([|i|]). + apply Z_mult_div_ge; auto with zarith. + case (phi_bounded i); auto. + intros j2 H1 H2; contradict H2; apply Zlt_not_le. + rewrite div31_phi; change ([|2|]) with 2; auto with zarith. + apply Zle_lt_trans with ([|i|]); auto with zarith. + assert (0 <= [|i|]/2)%Z by (apply Z_div_pos; auto with zarith). + apply Zle_trans with (2 * ([|i|]/2)); auto with zarith. + apply Z_mult_div_ge; auto with zarith. + case (phi_bounded i); unfold size; auto with zarith. + change [|0|] with 0; auto with zarith. + case (phi_bounded i); repeat rewrite Zpower_2; auto with zarith. + Qed. + + Lemma sqrt312_step_def rec ih il j: + sqrt312_step rec ih il j = + match (ih ?= j)%int31 with + Eq => j + | Gt => j + | _ => + match (fst (div3121 ih il j) ?= j)%int31 with + Lt => let m := match j +c fst (div3121 ih il j) with + C0 m1 => fst (m1/2)%int31 + | C1 m1 => (fst (m1/2) + v30)%int31 + end in rec ih il m + | _ => j + end + end. + Proof. + intros rec ih il j; unfold sqrt312_step; case div3121; intros. + simpl; case compare31; auto. + Qed. + + Lemma sqrt312_lower_bound ih il j: + phi2 ih il < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|]. + Proof. + intros ih il j H1. + case (phi_bounded j); intros Hbj _. + case (phi_bounded il); intros Hbil _. + case (phi_bounded ih); intros Hbih Hbih1. + assert (([|ih|] < [|j|] + 1)%Z); auto with zarith. + apply Zlt_square_simpl; auto with zarith. + repeat rewrite <-Zpower_2; apply Zle_lt_trans with (2 := H1). + apply Zle_trans with ([|ih|] * base)%Z; unfold phi2, base; + try rewrite Zpower_2; auto with zarith. + Qed. + + 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. + 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 -> + (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|] + 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. + assert (H1: ([|ih|] <= [|j|])%Z) by (apply sqrt312_lower_bound with il; auto). + case (phi_bounded ih); intros Hih1 _. + case (phi_bounded il); intros Hil1 _. + case (phi_bounded j); intros _ Hj1. + assert (Hp3: (0 < phi2 ih il)). + unfold phi2; apply Zlt_le_trans with ([|ih|] * base)%Z; auto with zarith. + apply Zmult_lt_0_compat; auto with zarith. + apply Zlt_le_trans with (2:= Hih); auto with zarith. + generalize (spec_compare ih j); case compare31; intros Hc1. + split; auto. + apply sqrt_test_true; auto. + unfold phi2, base; auto with zarith. + unfold phi2; rewrite Hc1. + assert (0 <= [|il|]/[|j|]) by (apply Z_div_pos; auto with zarith). + rewrite Zmult_comm, Z_div_plus_full_l; unfold base; auto with zarith. + unfold Zpower, Zpower_pos in Hj1; simpl in Hj1; auto with zarith. + case (Zle_or_lt (2 ^ 30) [|j|]); intros Hjj. + generalize (spec_compare (fst (div3121 ih il j)) j); case compare31; + rewrite div312_phi; auto; intros Hc; + try (split; auto; apply sqrt_test_true; auto with zarith; fail). + apply Hrec. + assert (Hf1: 0 <= phi2 ih il/ [|j|]) by (apply Z_div_pos; auto with zarith). + case (Zle_lt_or_eq 1 ([|j|])); auto with zarith; intros Hf2. + 2: contradict Hc; apply Zle_not_lt; rewrite <- Hf2, Zdiv_1_r; auto with zarith. + assert (Hf3: 0 < ([|j|] + phi2 ih il / [|j|]) / 2). + replace ([|j|] + phi2 ih il/ [|j|])%Z with + (1 * 2 + (([|j|] - 2) + phi2 ih il / [|j|])); try ring. + rewrite Z_div_plus_full_l; auto with zarith. + assert (0 <= ([|j|] - 2 + phi2 ih il / [|j|]) / 2) ; auto with zarith. + assert (Hf4: ([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]). + apply sqrt_test_false; auto with zarith. + generalize (spec_add_c j (fst (div3121 ih il j))). + unfold interp_carry; case add31c; intros r; + rewrite div312_phi; auto with zarith. + rewrite div31_phi; change [|2|] with 2%Z; auto with zarith. + intros HH; rewrite HH; clear HH; auto with zarith. + rewrite spec_add, div31_phi; change [|2|] with 2%Z; auto. + rewrite Zmult_1_l; intros HH. + rewrite Zplus_comm, <- Z_div_plus_full_l; auto with zarith. + change (phi v30 * 2) with (2 ^ Z_of_nat size). + rewrite HH, Zmod_small; auto with zarith. + replace (phi + match j +c fst (div3121 ih il j) with + | C0 m1 => fst (m1 / 2)%int31 + | C1 m1 => fst (m1 / 2)%int31 + v30 + end) with ((([|j|] + (phi2 ih il)/([|j|]))/2)). + apply sqrt_main; auto with zarith. + generalize (spec_add_c j (fst (div3121 ih il j))). + unfold interp_carry; case add31c; intros r; + rewrite div312_phi; auto with zarith. + rewrite div31_phi; auto with zarith. + intros HH; rewrite HH; auto with zarith. + intros HH; rewrite <- HH. + change (1 * 2 ^ Z_of_nat size) with (phi (v30) * 2). + rewrite Z_div_plus_full_l; auto with zarith. + rewrite Zplus_comm. + rewrite spec_add, Zmod_small. + rewrite div31_phi; auto. + split; auto with zarith. + case (phi_bounded (fst (r/2)%int31)); + case (phi_bounded v30); auto with zarith. + rewrite div31_phi; change (phi 2) with 2%Z; auto. + change (2 ^Z_of_nat size) with (base/2 + phi v30). + assert (phi r / 2 < base/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. + case (phi_bounded r); auto with zarith. + contradict Hij; apply Zle_not_lt. + assert ((1 + [|j|]) <= 2 ^ 30); auto with zarith. + apply Zle_trans with ((2 ^ 30) * (2 ^ 30)); auto with zarith. + assert (0 <= 1 + [|j|]); auto with zarith. + apply Zmult_le_compat; auto with zarith. + change ((2 ^ 30) * (2 ^ 30)) with ((2 ^ 29) * base). + apply Zle_trans with ([|ih|] * base); auto with zarith. + unfold phi2, base; auto with zarith. + split; auto. + apply sqrt_test_true; auto. + unfold phi2, base; auto with zarith. + apply Zle_ge; apply Zle_trans with (([|j|] * base)/[|j|]). + rewrite Zmult_comm, Z_div_mult; auto with zarith. + apply Zge_le; apply Z_div_ge; auto with zarith. + 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 -> + [|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|] + 1) ^ 2. + Proof. + intros n; 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. + intros n Hrec rec ih il j Hi Hj Hij HHrec. + apply sqrt312_step_correct; auto. + intros j1 Hj1 Hjp1; apply Hrec; auto with zarith. + intros j2 Hj2 H2j2 Hjp2; apply Hrec; auto with zarith. + intros j3 Hj3 Hpj3. + apply HHrec; auto. + rewrite inj_S, Zpower_Zsucc. + apply Zle_trans with (2 ^Z_of_nat n + [|j2|])%Z; auto with zarith. + apply Zle_0_nat. + Qed. + + Lemma spec_sqrt2 : forall x y, + wB/ 4 <= [|x|] -> + let (s,r) := sqrt312 x y in + [||WW x y||] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]. + 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 + (intros s; ring). + assert (Hb: 0 <= base) by (red; intros HH; discriminate). + assert (Hi2: phi2 ih il < (phi Tn + 1) ^ 2). + change ((phi Tn + 1) ^ 2) with (2^62). + apply Zle_lt_trans with ((2^31 -1) * base + (2^31 - 1)); auto with zarith. + 2: simpl; unfold Zpower_pos; simpl; auto with zarith. + case (phi_bounded ih); case (phi_bounded il); intros H1 H2 H3 H4. + unfold base, Zpower, Zpower_pos in H2,H4; simpl in H2,H4. + unfold phi2,Zpower, Zpower_pos; simpl iter_pos; auto with zarith. + case (iter312_sqrt_correct 31 (fun _ _ j => j) ih il Tn); auto with zarith. + change [|Tn|] with 2147483647; auto with zarith. + intros j1 _ HH; contradict HH. + apply Zlt_not_le. + change [|Tn|] with 2147483647; auto with zarith. + change (2 ^ Z_of_nat 31) with 2147483648; auto with zarith. + case (phi_bounded j1); auto with zarith. + set (s := iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn). + intros Hs1 Hs2. + generalize (spec_mul_c s s); case mul31c. + simpl zn2z_to_Z; intros HH. + assert ([|s|] = 0). + case (Zmult_integral _ _ (sym_equal HH)); auto. + contradict Hs2; apply Zle_not_lt; rewrite H. + change ((0 + 1) ^ 2) with 1. + apply Zle_trans with (2 ^ Z_of_nat size / 4 * base). + simpl; auto with zarith. + apply Zle_trans with ([|ih|] * base); auto with zarith. + unfold phi2; case (phi_bounded il); auto with zarith. + intros ih1 il1. + change [||WW ih1 il1||] with (phi2 ih1 il1). + intros Hihl1. + generalize (spec_sub_c il il1). + case sub31c; intros il2 Hil2. + simpl interp_carry in Hil2. + generalize (spec_compare ih ih1); case compare31. + unfold interp_carry. + intros H1; split. + rewrite Zpower_2, <- Hihl1. + unfold phi2; ring[Hil2 H1]. + replace [|il2|] with (phi2 ih il - phi2 ih1 il1). + rewrite Hihl1. + rewrite <-Hbin in Hs2; auto with zarith. + unfold phi2; rewrite H1, Hil2; ring. + unfold interp_carry. + intros H1; contradict Hs1. + apply Zlt_not_le; rewrite Zpower_2, <-Hihl1. + unfold phi2. + case (phi_bounded il); intros _ H2. + apply Zlt_le_trans with (([|ih|] + 1) * base + 0). + rewrite Zmult_plus_distr_l, Zplus_0_r; auto with zarith. + case (phi_bounded il1); intros H3 _. + apply Zplus_le_compat; auto with zarith. + unfold interp_carry; change (1 * 2 ^ Z_of_nat size) with base. + rewrite Zpower_2, <- Hihl1, Hil2. + intros H1. + case (Zle_lt_or_eq ([|ih1|] + 1) ([|ih|])); auto with zarith. + intros H2; contradict Hs2; apply Zle_not_lt. + replace (([|s|] + 1) ^ 2) with (phi2 ih1 il1 + 2 * [|s|] + 1). + unfold phi2. + case (phi_bounded il); intros Hpil _. + assert (Hl1l: [|il1|] <= [|il|]). + case (phi_bounded il2); rewrite Hil2; auto with zarith. + assert ([|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base); auto with zarith. + case (phi_bounded s); change (2 ^ Z_of_nat size) with base; intros _ Hps. + case (phi_bounded ih1); intros Hpih1 _; auto with zarith. + apply Zle_trans with (([|ih1|] + 2) * base); auto with zarith. + rewrite Zmult_plus_distr_l. + assert (2 * [|s|] + 1 <= 2 * base); auto with zarith. + rewrite Hihl1, Hbin; auto. + intros H2; split. + unfold phi2; rewrite <- H2; ring. + replace (base + ([|il|] - [|il1|])) with (phi2 ih il - ([|s|] * [|s|])). + rewrite <-Hbin in Hs2; auto with zarith. + rewrite <- Hihl1; unfold phi2; rewrite <- H2; ring. + unfold interp_carry in Hil2 |- *. + unfold interp_carry; change (1 * 2 ^ Z_of_nat size) with base. + assert (Hsih: [|ih - 1|] = [|ih|] - 1). + rewrite spec_sub, Zmod_small; auto; change [|1|] with 1. + case (phi_bounded ih); intros H1 H2. + generalize Hih; change (2 ^ Z_of_nat size / 4) with 536870912. + split; auto with zarith. + generalize (spec_compare (ih - 1) ih1); case compare31. + rewrite Hsih. + intros H1; split. + rewrite Zpower_2, <- Hihl1. + unfold phi2; rewrite <-H1. + apply trans_equal with ([|ih|] * base + [|il1|] + ([|il|] - [|il1|])). + ring. + rewrite <-Hil2. + change (2 ^ Z_of_nat size) with base; ring. + replace [|il2|] with (phi2 ih il - phi2 ih1 il1). + rewrite Hihl1. + rewrite <-Hbin in Hs2; auto with zarith. + unfold phi2. + rewrite <-H1. + ring_simplify. + apply trans_equal with (base + ([|il|] - [|il1|])). + ring. + rewrite <-Hil2. + change (2 ^ Z_of_nat size) with base; ring. + rewrite Hsih; intros H1. + assert (He: [|ih|] = [|ih1|]). + apply Zle_antisym; auto with zarith. + case (Zle_or_lt [|ih1|] [|ih|]); auto; intros H2. + contradict Hs1; apply Zlt_not_le; rewrite Zpower_2, <-Hihl1. + unfold phi2. + case (phi_bounded il); change (2 ^ Z_of_nat size) with base; + intros _ Hpil1. + apply Zlt_le_trans with (([|ih|] + 1) * base). + rewrite Zmult_plus_distr_l, Zmult_1_l; auto with zarith. + case (phi_bounded il1); intros Hpil2 _. + apply Zle_trans with (([|ih1|]) * base); auto with zarith. + rewrite Zpower_2, <-Hihl1; unfold phi2; rewrite <-He. + contradict Hs1; apply Zlt_not_le; rewrite Zpower_2, <-Hihl1. + unfold phi2; rewrite He. + assert (phi il - phi il1 < 0); auto with zarith. + rewrite <-Hil2. + case (phi_bounded il2); auto with zarith. + intros H1. + rewrite Zpower_2, <-Hihl1. + case (Zle_lt_or_eq ([|ih1|] + 2) [|ih|]); auto with zarith. + intros H2; contradict Hs2; apply Zle_not_lt. + replace (([|s|] + 1) ^ 2) with (phi2 ih1 il1 + 2 * [|s|] + 1). + unfold phi2. + assert ([|ih1|] * base + 2 * phi s + 1 <= [|ih|] * base + ([|il|] - [|il1|])); + auto with zarith. + rewrite <-Hil2. + change (-1 * 2 ^ Z_of_nat size) with (-base). + case (phi_bounded il2); intros Hpil2 _. + apply Zle_trans with ([|ih|] * base + - base); auto with zarith. + case (phi_bounded s); change (2 ^ Z_of_nat size) with base; intros _ Hps. + assert (2 * [|s|] + 1 <= 2 * base); auto with zarith. + apply Zle_trans with ([|ih1|] * base + 2 * base); auto with zarith. + assert (Hi: ([|ih1|] + 3) * base <= [|ih|] * base); auto with zarith. + rewrite Zmult_plus_distr_l in Hi; auto with zarith. + rewrite Hihl1, Hbin; auto. + intros H2; unfold phi2; rewrite <-H2. + split. + replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring. + rewrite <-Hil2. + change (-1 * 2 ^ Z_of_nat size) with (-base); ring. + replace (base + [|il2|]) with (phi2 ih il - phi2 ih1 il1). + rewrite Hihl1. + rewrite <-Hbin in Hs2; auto with zarith. + unfold phi2; rewrite <-H2. + replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring. + rewrite <-Hil2. + change (-1 * 2 ^ Z_of_nat size) with (-base); ring. + Qed. + + (** [iszero] *) + + Let w_eq0 := int31_op.(znz_eq0). + + Lemma spec_eq0 : forall x, w_eq0 x = true -> [|x|] = 0. + Proof. + clear; unfold w_eq0, znz_eq0; simpl. + unfold compare31; simpl; intros. + change [|0|] with 0 in H. + 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, + if w_is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1. + Proof. + unfold w_is_even; simpl; intros. + generalize (spec_div x 2). + destruct (x/2)%int31 as (q,r); intros. + unfold compare31. + change [|2|] with 2 in H. + change [|0|] with 0. + destruct H; auto with zarith. + replace ([|x|] mod 2) with [|r|]. + destruct H; auto with zarith. + case_eq ([|r|] ?= 0)%Z; intros. + apply Zcompare_Eq_eq; auto. + change ([|r|] < 0)%Z in H; auto with zarith. + change ([|r|] > 0)%Z in H; auto with zarith. + apply Zmod_unique with [|q|]; auto with zarith. + Qed. + + Definition int31_spec : znz_spec int31_op. + split. + exact phi_bounded. + exact positive_to_int31_spec. + exact spec_zdigits. + exact spec_more_than_1_digit. + + exact spec_0. + exact spec_1. + exact spec_Bm1. + + exact spec_compare. + exact spec_eq0. + + exact spec_opp_c. + exact spec_opp. + exact spec_opp_carry. + + exact spec_succ_c. + exact spec_add_c. + exact spec_add_carry_c. + exact spec_succ. + exact spec_add. + exact spec_add_carry. + + exact spec_pred_c. + exact spec_sub_c. + exact spec_sub_carry_c. + exact spec_pred. + exact spec_sub. + exact spec_sub_carry. + + exact spec_mul_c. + exact spec_mul. + exact spec_square_c. + + exact spec_div21. + intros; apply spec_div; auto. + exact spec_div. + + intros; unfold int31_op; simpl; apply spec_mod; auto. + exact spec_mod. + + intros; apply spec_gcd; auto. + exact spec_gcd. + + exact spec_head00. + exact spec_head0. + exact spec_tail00. + exact spec_tail0. + + exact spec_add_mul_div. + exact spec_pos_mod. + + exact spec_is_even. + exact spec_sqrt2. + exact spec_sqrt. + Qed. + +End Int31_Spec. + + +Module Int31Cyclic <: CyclicType. + Definition w := int31. + Definition w_op := int31_op. + Definition w_spec := int31_spec. +End Int31Cyclic. diff --git a/theories/Numbers/Cyclic/Int31/Int31.v b/theories/Numbers/Cyclic/Int31/Int31.v new file mode 100644 index 00000000..154b436b --- /dev/null +++ b/theories/Numbers/Cyclic/Int31/Int31.v @@ -0,0 +1,469 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: Int31.v 11072 2008-06-08 16:13:37Z herbelin $ i*) + +Require Import NaryFunctions. +Require Import Wf_nat. +Require Export ZArith. +Require Export DoubleType. + +Unset Boxed Definitions. + +(** * 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 + for this use of 31 is the underlying mecanism for hardware-efficient + computations by A. Spiwack. Apart from this, a switch to, say, + 63-bit integers is now just a matter of replacing every occurences + of 31 by 63. This is actually made possible by the use of + dependently-typed n-ary constructions for the inductive type + [int31], its constructor [I31] and any pattern matching on it. + If you modify this file, please preserve this genericity. *) + +Definition size := 31%nat. + +(** Digits *) + +Inductive digits : Type := D0 | D1. + +(** The type of 31-bit integers *) + +(** The type [int31] has a unique constructor [I31] that expects + 31 arguments of type [digits]. *) + +Inductive int31 : Type := I31 : nfun digits size int31. + +(* spiwack: Registration of the type of integers, so that the matchs in + the functions below perform dynamic decompilation (otherwise some segfault + occur when they are applied to one non-closed term and one closed term). *) +Register digits as int31 bits in "coq_int31" by True. +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. + +(** * Constants *) + +(** Zero is [I31 D0 ... D0] *) +Definition On : int31 := Eval compute in napply_cst _ _ D0 size I31. + +(** One is [I31 D0 ... D0 D1] *) +Definition In : int31 := Eval compute in (napply_cst _ _ D0 (size-1) I31) D1. + +(** The biggest integer is [I31 D1 ... D1], corresponding to [(2^size)-1] *) +Definition Tn : int31 := Eval compute in napply_cst _ _ D1 size I31. + +(** Two is [I31 D0 ... D0 D1 D0] *) +Definition Twon : int31 := Eval compute in (napply_cst _ _ D0 (size-2) I31) D1 D0. + +(** * Bits manipulation *) + + +(** [sneakr b x] shifts [x] to the right by one bit. + Rightmost digit is lost while leftmost digit becomes [b]. + 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. + Leftmost digit is lost while rightmost digit becomes [b]. + Pseudo-code is + [ match x with (I31 d0 ... dN) => I31 d1 ... dN b end ] +*) + +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 + consequences of [sneakl] and [sneakr]. *) + +Definition shiftl := sneakl D0. +Definition shiftr := sneakr D0. +Definition twice := sneakl D0. +Definition twice_plus_one := sneakl D1. + +(** [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 + int31_rect _ (fun d => napply_discard _ _ d (size-1)). + +(** [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 + int31_rect _ (napply_discard _ _ (fun d=>d) (size-1)). + +(** [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 + in int31_rect _ (nfold_bis _ _ f true size). + +(* 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 + (see below) *) + +Definition base := Eval compute in + iter_nat size Z Zdouble 1%Z. + +(** * Recursors *) + +Fixpoint recl_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A) + (i:int31) : A := + match n with + | O => case0 + | S next => + if iszero i then + case0 + else + let si := shiftl i in + 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) + (i:int31) : A := + match n with + | O => case0 + | S next => + if iszero i then + case0 + else + let si := shiftr i in + caserec (firstr i) si (recr_aux next A case0 caserec si) + end. + +Definition recl := recl_aux size. +Definition recr := recr_aux size. + +(** * Conversions *) + +(** From int31 to Z, we simply iterates [Zdouble] or [Zdouble_plus_one]. *) + +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) /\ + phi_inv 2n+1 = 2*(phi_inv n) + 1 ] *) + +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 *) + +Fixpoint complement_negative p := + match p with + | xI q => twice (complement_negative q) + | xO q => twice_plus_one (complement_negative q) + | xH => twice Tn + end. + +(** A simple incrementation function *) + +Definition incr : int31 -> int31 := + 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. *) + +Definition phi_inv : Z -> int31 := fun n => + match n with + | Z0 => On + | 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 + [zn2z int31] *) + +Definition phi_inv2 n := + match n with + | Z0 => W0 + | _ => WW (phi_inv (n/base)%Z) (phi_inv n) + end. + +(** [phi2] is similar to [phi] but takes a double word (two args) *) + +Definition phi2 nh nl := + ((phi nh)*base+(phi nl))%Z. + +(** * Addition *) + +(** Addition modulo [2^31] *) + +Definition add31 (n m : int31) := phi_inv ((phi n)+(phi m)). +Notation "n + m" := (add31 n m) : int31_scope. + +(** Addition with carry (the result is thus exact) *) + +(* spiwack : when executed in non-compiled*) +(* mode, (phi n)+(phi m) is computed twice*) +(* it may be considered to optimize it *) + +Definition add31c (n m : int31) := + let npm := n+m in + 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. + +(** Addition plus one with carry (the result is thus exact) *) + +Definition add31carryc (n m : int31) := + let npmpone_exact := ((phi n)+(phi m)+1)%Z in + let npmpone := phi_inv npmpone_exact in + match (phi npmpone ?= npmpone_exact)%Z with + | Eq => C0 npmpone + | _ => C1 npmpone + end. + +(** * Substraction *) + +(** Subtraction modulo [2^31] *) + +Definition sub31 (n m : int31) := phi_inv ((phi n)-(phi m)). +Notation "n - m" := (sub31 n m) : int31_scope. + +(** Subtraction with carry (thus exact) *) + +Definition sub31c (n m : int31) := + let nmm := n-m in + match (phi nmm ?= (phi n)-(phi m))%Z with + | Eq => C0 nmm + | _ => C1 nmm + end. +Notation "n '-c' m" := (sub31c n m) (at level 50, no associativity) : int31_scope. + +(** subtraction minus one with carry (thus exact) *) + +Definition sub31carryc (n m : int31) := + let nmmmone_exact := ((phi n)-(phi m)-1)%Z in + let nmmmone := phi_inv nmmmone_exact in + match (phi nmmmone ?= nmmmone_exact)%Z with + | Eq => C0 nmmmone + | _ => C1 nmmmone + end. + + +(** Multiplication *) + +(** multiplication modulo [2^31] *) + +Definition mul31 (n m : int31) := phi_inv ((phi n)*(phi m)). +Notation "n * m" := (mul31 n m) : int31_scope. + +(** multiplication with double word result (thus exact) *) + +Definition mul31c (n m : int31) := phi_inv2 ((phi n)*(phi m)). +Notation "n '*c' m" := (mul31c n m) (at level 40, no associativity) : int31_scope. + + +(** * Division *) + +(** Division of a double size word modulo [2^31] *) + +Definition div3121 (nh nl m : int31) := + let (q,r) := Zdiv_eucl (phi2 nh nl) (phi m) in + (phi_inv q, phi_inv r). + +(** Division modulo [2^31] *) + +Definition div31 (n m : int31) := + let (q,r) := Zdiv_eucl (phi n) (phi m) in + (phi_inv q, phi_inv r). +Notation "n / m" := (div31 n m) : int31_scope. + + +(** * Unsigned comparison *) + +Definition compare31 (n m : int31) := ((phi n)?=(phi m))%Z. +Notation "n ?= m" := (compare31 n m) (at level 70, no associativity) : int31_scope. + + +(** 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 + | D0 => fun x => rec (rec x) + | D1 => fun x => f (rec (rec x)) + end) + i. + +(** 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) + (fun ij => let (i,j) := ij in (sneakl (firstl j) i, shiftl j)) + (i,j) + in + res. + + +Register add31 as int31 plus in "coq_int31" by True. +Register add31c as int31 plusc in "coq_int31" by True. +Register add31carryc as int31 pluscarryc in "coq_int31" by True. +Register sub31 as int31 minus in "coq_int31" by True. +Register sub31c as int31 minusc in "coq_int31" by True. +Register sub31carryc as int31 minuscarryc in "coq_int31" by True. +Register mul31 as int31 times in "coq_int31" by True. +Register mul31c as int31 timesc in "coq_int31" by True. +Register div3121 as int31 div21 in "coq_int31" by True. +Register div31 as int31 div in "coq_int31" by True. +Register compare31 as int31 compare in "coq_int31" by True. +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 + | O => In + | S p => match j ?= On with + | Eq => i + | _ => euler p j (let (_, r ) := i/j in r) + end + end) + (2*size)%nat i j. + +(** Square root functions using newton iteration + we use a very naive upper-bound on the iteration + 2^31 instead of the usual 31. +**) + + + +Definition sqrt31_step (rec: int31 -> int31 -> int31) (i j: int31) := +Eval lazy delta [Twon] in + let (quo,_) := i/j in + match quo ?= j with + Lt => rec i (fst ((j + quo)/Twon)) + | _ => j + end. + +Fixpoint iter31_sqrt (n: nat) (rec: int31 -> int31 -> int31) + (i j: int31) {struct n} : int31 := + sqrt31_step + (match n with + O => rec + | S n => (iter31_sqrt n (iter31_sqrt n rec)) + end) i j. + +Definition sqrt31 i := +Eval lazy delta [On In Twon] in + match compare31 In i with + Gt => On + | Eq => In + | Lt => iter31_sqrt 31 (fun i j => j) i (fst (i/Twon)) + end. + +Definition v30 := Eval compute in (addmuldiv31 (phi_inv (Z_of_nat size - 1)) In On). + +Definition 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 | _ => + let (quo,_) := div3121 ih il j in + match quo ?= j with + Lt => let m := match j +c quo with + C0 m1 => fst (m1/Twon) + | C1 m1 => fst (m1/Twon) + v30 + end in rec ih il m + | _ => j + end end. + +Fixpoint iter312_sqrt (n: nat) + (rec: int31 -> int31 -> int31 -> int31) + (ih il j: int31) {struct n} : int31 := + sqrt312_step + (match n with + O => rec + | S n => (iter312_sqrt n (iter312_sqrt n rec)) + end) ih il j. + +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 => + match ih ?= ih1 with + Gt => (s, C1 il2) + | _ => (s, C0 il2) + end + | C1 il2 => + match (ih - In) ?= ih1 with (* we could parametrize ih - 1 *) + Gt => (s, C1 il2) + | _ => (s, C0 il2) + end + end + end. + + +Fixpoint p2i n p : (N*int31)%type := + match n with + | O => (Npos p, On) + | S n => match p with + | xO p => let (r,i) := p2i n p in (r, Twon*i) + | xI p => let (r,i) := p2i n p in (r, Twon*i+In) + | xH => (N0, In) + end + end. + +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 + 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 + | 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 + | 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. diff --git a/theories/Numbers/Cyclic/ZModulo/ZModulo.v b/theories/Numbers/Cyclic/ZModulo/ZModulo.v new file mode 100644 index 00000000..7c770e97 --- /dev/null +++ b/theories/Numbers/Cyclic/ZModulo/ZModulo.v @@ -0,0 +1,946 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + +(* $Id: ZModulo.v 11033 2008-06-01 22:56:50Z letouzey $ *) + +(** * 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 + implementation of CyclicAxioms, hence ensuring its coherence. *) + +Set Implicit Arguments. + +Require Import Bool. +Require Import ZArith. +Require Import Znumtheory. +Require Import BigNumPrelude. +Require Import DoubleType. +Require Import CyclicAxioms. + +Open Local Scope Z_scope. + +Section ZModulo. + + Variable digits : positive. + Hypothesis digits_ne_1 : digits <> 1%positive. + + Definition wB := base digits. + + Definition znz := Z. + Definition znz_digits := digits. + Definition znz_zdigits := Zpos digits. + Definition znz_to_Z x := x mod wB. + + Notation "[| x |]" := (znz_to_Z x) (at level 0, x at level 99). + + Notation "[+| c |]" := + (interp_carry 1 wB znz_to_Z c) (at level 0, x at level 99). + + Notation "[-| c |]" := + (interp_carry (-1) wB znz_to_Z c) (at level 0, x at level 99). + + Notation "[|| x ||]" := + (zn2z_to_Z wB znz_to_Z x) (at level 0, x at level 99). + + Lemma spec_more_than_1_digit: 1 < Zpos digits. + Proof. + unfold znz_digits. + generalize digits_ne_1; destruct digits; auto. + destruct 1; auto. + Qed. + Let digits_gt_1 := spec_more_than_1_digit. + + Lemma wB_pos : wB > 0. + Proof. + unfold wB, base; auto with zarith. + Qed. + Hint Resolve wB_pos. + + Lemma spec_to_Z_1 : forall x, 0 <= [|x|]. + Proof. + unfold znz_to_Z; intros; destruct (Z_mod_lt x wB wB_pos); auto. + Qed. + + Lemma spec_to_Z_2 : forall x, [|x|] < wB. + Proof. + unfold znz_to_Z; intros; destruct (Z_mod_lt x wB wB_pos); auto. + Qed. + Hint Resolve spec_to_Z_1 spec_to_Z_2. + + Lemma spec_to_Z : forall x, 0 <= [|x|] < wB. + Proof. + auto. + Qed. + + 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, + Zpos p = (Z_of_N (fst (znz_of_pos p)))*wB + [|(snd (znz_of_pos p))|]. + Proof. + intros; unfold znz_of_pos; simpl. + generalize (Z_div_mod_POS wB wB_pos p). + 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 + (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 + (destruct z; simpl; auto; elim H1; auto). + rewrite Zmult_comm; auto. + Qed. + + Lemma spec_zdigits : [|znz_zdigits|] = Zpos znz_digits. + Proof. + unfold znz_to_Z, znz_zdigits, znz_digits. + apply Zmod_small. + unfold wB, base. + split; auto with zarith. + apply Zpower2_lt_lin; auto with zarith. + Qed. + + 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. + apply Zmod_small; generalize wB_pos; auto with zarith. + Qed. + + Lemma spec_1 : [|znz_1|] = 1. + Proof. + unfold znz_to_Z, znz_1. + apply Zmod_small; split; auto with zarith. + unfold wB, base. + apply Zlt_trans with (Zpos digits); auto. + apply Zpower2_lt_lin; auto with zarith. + Qed. + + Lemma spec_Bm1 : [|znz_Bm1|] = wB - 1. + Proof. + unfold znz_to_Z, znz_Bm1. + apply Zmod_small; split; auto with zarith. + unfold wB, base. + cut (1 <= 2 ^ Zpos digits); auto with zarith. + apply Zle_trans with (Zpos digits); auto with zarith. + apply Zpower2_le_lin; auto with zarith. + Qed. + + Definition znz_compare x y := Zcompare [|x|] [|y|]. + + Lemma spec_compare : forall x y, + match znz_compare x y with + | Eq => [|x|] = [|y|] + | Lt => [|x|] < [|y|] + | Gt => [|x|] > [|y|] + end. + Proof. + intros; unfold znz_compare, Zlt, Zgt. + case_eq (Zcompare [|x|] [|y|]); auto. + intros; apply Zcompare_Eq_eq; auto. + Qed. + + 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 := + 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. + case_eq (znz_eq0 x); intros; unfold interp_carry. + fold [|x|]; rewrite (spec_eq0 x H); auto. + assert (x mod wB <> 0). + unfold znz_eq0, znz_to_Z in H. + intro H0; rewrite H0 in H; discriminate. + rewrite Z_mod_nz_opp_full; auto with zarith. + Qed. + + Lemma spec_opp : forall x, [|znz_opp x|] = (-[|x|]) mod wB. + Proof. + intros; unfold znz_opp, znz_to_Z; auto. + 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. + replace (- x - 1) with (- 1 - x) by omega. + rewrite <- Zminus_mod_idemp_r. + replace ( -1 - x mod wB) with (0 + ( -1 - x mod wB)) by omega. + rewrite <- (Z_mod_same_full wB). + rewrite Zplus_mod_idemp_l. + replace (wB + (-1 - x mod wB)) with (wB - x mod wB -1) by omega. + apply Zmod_small. + generalize (Z_mod_lt x wB wB_pos); omega. + Qed. + + 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 + if Z_lt_le_dec z wB then C0 z else C1 (z-wB). + + 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). + + Definition znz_succ := Zsucc. + Definition znz_add := Zplus. + Definition znz_add_carry x y := x + y + 1. + + Lemma Zmod_equal : + forall x y z, z>0 -> (x-y) mod z = 0 -> x mod z = y mod z. + Proof. + intros. + generalize (Z_div_mod_eq (x-y) z H); rewrite H0, Zplus_0_r. + remember ((x-y)/z) as k. + intros H1; symmetry in H1; rewrite <- Zeq_plus_swap in H1. + subst x. + rewrite Zplus_comm, Zmult_comm, Z_mod_plus; auto. + Qed. + + Lemma spec_succ_c : forall x, [+|znz_succ_c x|] = [|x|] + 1. + 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 + (apply Zmod_small; generalize wB_pos; omega). + rewrite <- Zminus_mod_idemp_l; rewrite Z_mod_same; simpl; auto. + apply Zmod_equal; auto. + + assert ((x+1) mod wB <> 0). + unfold znz_eq0, znz_to_Z in *; now destruct ((x+1) mod wB). + assert (x mod wB + 1 <> wB). + contradict H0. + rewrite Zeq_plus_swap in H0; simpl in H0. + rewrite <- Zplus_mod_idemp_l; rewrite H0. + replace (wB-1+1) with wB; auto with zarith; apply Z_mod_same; auto. + rewrite <- Zplus_mod_idemp_l. + apply Zmod_small. + generalize (Z_mod_lt x wB wB_pos); omega. + Qed. + + Lemma spec_add_c : forall x y, [+|znz_add_c x y|] = [|x|] + [|y|]. + Proof. + intros; unfold znz_add_c, znz_to_Z, interp_carry. + destruct Z_lt_le_dec. + apply Zmod_small; + generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + rewrite Zmult_1_l, Zplus_comm, Zeq_plus_swap. + apply Zmod_small; + generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + Qed. + + Lemma spec_add_carry_c : forall x y, [+|znz_add_carry_c x y|] = [|x|] + [|y|] + 1. + Proof. + intros; unfold znz_add_carry_c, znz_to_Z, interp_carry. + destruct Z_lt_le_dec. + apply Zmod_small; + generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + rewrite Zmult_1_l, Zplus_comm, Zeq_plus_swap. + apply Zmod_small; + generalize (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + Qed. + + Lemma spec_succ : forall x, [|znz_succ x|] = ([|x|] + 1) mod wB. + Proof. + intros; unfold znz_succ, znz_to_Z, Zsucc. + symmetry; apply Zplus_mod_idemp_l. + Qed. + + Lemma spec_add : forall x y, [|znz_add x y|] = ([|x|] + [|y|]) mod wB. + Proof. + intros; unfold znz_add, znz_to_Z; apply Zplus_mod. + Qed. + + Lemma spec_add_carry : + forall x y, [|znz_add_carry x y|] = ([|x|] + [|y|] + 1) mod wB. + Proof. + intros; unfold znz_add_carry, znz_to_Z. + rewrite <- Zplus_mod_idemp_l. + rewrite (Zplus_mod x y). + rewrite Zplus_mod_idemp_l; auto. + Qed. + + 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 + 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 + if Z_lt_le_dec z 0 then C1 (wB+z) else C0 z. + + Definition znz_pred := Zpred. + Definition znz_sub := Zminus. + Definition znz_sub_carry x y := x - y - 1. + + Lemma spec_pred_c : forall x, [-|znz_pred_c x|] = [|x|] - 1. + Proof. + intros; unfold znz_pred_c, znz_to_Z, interp_carry. + case_eq (znz_eq0 x); intros. + fold [|x|]; rewrite spec_eq0; auto. + replace ((wB-1) mod wB) with (wB-1); auto with zarith. + symmetry; apply Zmod_small; generalize wB_pos; omega. + + assert (x mod wB <> 0). + unfold znz_eq0, znz_to_Z in *; now destruct (x mod wB). + rewrite <- Zminus_mod_idemp_l. + apply Zmod_small. + generalize (Z_mod_lt x wB wB_pos); omega. + Qed. + + Lemma spec_sub_c : forall x y, [-|znz_sub_c x y|] = [|x|] - [|y|]. + 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 + (wB + (x mod wB - y mod wB)). + omega. + symmetry; apply Zmod_small. + generalize wB_pos (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + + apply Zmod_small. + generalize wB_pos (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + Qed. + + Lemma spec_sub_carry_c : forall x y, [-|znz_sub_carry_c x y|] = [|x|] - [|y|] - 1. + 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 + (wB + (x mod wB - y mod wB -1)). + omega. + symmetry; apply Zmod_small. + generalize wB_pos (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + + apply Zmod_small. + generalize wB_pos (Z_mod_lt x wB wB_pos) (Z_mod_lt y wB wB_pos); omega. + Qed. + + Lemma spec_pred : forall x, [|znz_pred x|] = ([|x|] - 1) mod wB. + Proof. + intros; unfold znz_pred, znz_to_Z, Zpred. + rewrite <- Zplus_mod_idemp_l; auto. + Qed. + + Lemma spec_sub : forall x y, [|znz_sub x y|] = ([|x|] - [|y|]) mod wB. + Proof. + intros; unfold znz_sub, znz_to_Z; apply Zminus_mod. + Qed. + + 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. + rewrite <- Zminus_mod_idemp_l. + rewrite (Zminus_mod x y). + rewrite Zminus_mod_idemp_l. + auto. + Qed. + + 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. + assert (Zdiv_eucl ([|x|]*[|y|]) wB = (([|x|]*[|y|])/wB,([|x|]*[|y|]) mod wB)). + unfold Zmod, Zdiv; destruct Zdiv_eucl; auto. + generalize (Z_div_mod ([|x|]*[|y|]) wB wB_pos); destruct Zdiv_eucl as (h,l). + destruct 1; injection H; clear H; intros. + rewrite H0. + assert ([|l|] = l). + apply Zmod_small; auto. + assert ([|h|] = h). + apply Zmod_small. + subst h. + split. + apply Z_div_pos; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + apply Zmult_lt_compat; auto with zarith. + clear H H0 H1 H2. + case_eq (znz_eq0 h); simpl; intros. + case_eq (znz_eq0 l); simpl; intros. + rewrite <- H3, <- H4, (spec_eq0 h), (spec_eq0 l); auto with zarith. + rewrite H3, H4; auto with zarith. + rewrite H3, H4; auto with zarith. + Qed. + + Lemma spec_mul : forall x y, [|znz_mul x y|] = ([|x|] * [|y|]) mod wB. + Proof. + intros; unfold znz_mul, znz_to_Z; apply Zmult_mod. + Qed. + + Lemma spec_square_c : forall x, [|| znz_square_c x||] = [|x|] * [|x|]. + Proof. + intros x; exact (spec_mul_c x x). + Qed. + + Definition znz_div x y := Zdiv_eucl [|x|] [|y|]. + + Lemma spec_div : forall a b, 0 < [|b|] -> + let (q,r) := znz_div a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + intros; unfold znz_div. + assert ([|b|]>0) by auto with zarith. + assert (Zdiv_eucl [|a|] [|b|] = ([|a|]/[|b|], [|a|] mod [|b|])). + unfold Zmod, Zdiv; destruct Zdiv_eucl; auto. + generalize (Z_div_mod [|a|] [|b|] H0). + 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|]; + auto with zarith. + assert ([|q|]=q). + apply Zmod_small. + subst q. + split. + apply Z_div_pos; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + apply Zlt_le_trans with (wB*1). + rewrite Zmult_1_r; auto with zarith. + apply Zmult_le_compat; generalize wB_pos; auto with zarith. + rewrite H5, H6; rewrite Zmult_comm; auto with zarith. + Qed. + + Definition znz_div_gt := znz_div. + + Lemma spec_div_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> + let (q,r) := znz_div_gt a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + intros. + apply spec_div; auto. + Qed. + + 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. + intros; unfold znz_mod. + apply Zmod_small. + assert ([|b|]>0) by auto with zarith. + generalize (Z_mod_lt [|a|] [|b|] H0) (Z_mod_lt b wB wB_pos). + fold [|b|]; omega. + Qed. + + Lemma spec_mod_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] -> + [|znz_mod_gt a b|] = [|a|] mod [|b|]. + Proof. + intros; apply spec_mod; auto. + Qed. + + Definition znz_gcd x y := Zgcd [|x|] [|y|]. + Definition znz_gcd_gt x y := Zgcd [|x|] [|y|]. + + Lemma Zgcd_bound : forall a b, 0<=a -> 0<=b -> Zgcd a b <= Zmax a b. + Proof. + intros. + generalize (Zgcd_is_gcd a b); inversion_clear 1. + destruct H2; destruct H3; clear H4. + assert (H3:=Zgcd_is_pos a b). + destruct (Z_eq_dec (Zgcd a b) 0). + rewrite e; generalize (Zmax_spec a b); omega. + assert (0 <= q). + apply Zmult_le_reg_r with (Zgcd a b); auto with zarith. + destruct (Z_eq_dec q 0). + + subst q; simpl in *; subst a; simpl; auto. + generalize (Zmax_spec 0 b) (Zabs_spec b); omega. + + apply Zle_trans with a. + rewrite H1 at 2. + rewrite <- (Zmult_1_l (Zgcd a b)) at 1. + apply Zmult_le_compat; auto with zarith. + generalize (Zmax_spec a b); omega. + Qed. + + Lemma spec_gcd : forall a b, Zis_gcd [|a|] [|b|] [|znz_gcd a b|]. + Proof. + intros; unfold znz_gcd. + generalize (Z_mod_lt a wB wB_pos)(Z_mod_lt b wB wB_pos); intros. + fold [|a|] in *; fold [|b|] in *. + replace ([|Zgcd [|a|] [|b|]|]) with (Zgcd [|a|] [|b|]). + apply Zgcd_is_gcd. + symmetry; apply Zmod_small. + split. + apply Zgcd_is_pos. + apply Zle_lt_trans with (Zmax [|a|] [|b|]). + apply Zgcd_bound; auto with zarith. + generalize (Zmax_spec [|a|] [|b|]); omega. + Qed. + + Lemma spec_gcd_gt : forall a b, [|a|] > [|b|] -> + Zis_gcd [|a|] [|b|] [|znz_gcd_gt a b|]. + Proof. + intros. apply spec_gcd; auto. + Qed. + + Definition znz_div21 a1 a2 b := + Zdiv_eucl ([|a1|]*wB+[|a2|]) [|b|]. + + Lemma spec_div21 : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := znz_div21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. + Proof. + intros; unfold znz_div21. + generalize (Z_mod_lt a1 wB wB_pos); fold [|a1|]; intros. + generalize (Z_mod_lt a2 wB wB_pos); fold [|a2|]; intros. + assert ([|b|]>0) by auto with zarith. + remember ([|a1|]*wB+[|a2|]) as a. + assert (Zdiv_eucl a [|b|] = (a/[|b|], a mod [|b|])). + unfold Zmod, Zdiv; destruct Zdiv_eucl; auto. + generalize (Z_div_mod a [|b|] H3). + 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|]; + auto with zarith. + assert ([|q|]=q). + apply Zmod_small. + subst q. + split. + 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. + rewrite H8, H9; rewrite Zmult_comm; auto with zarith. + Qed. + + Definition znz_add_mul_div p x y := + ([|x|] * (2 ^ [|p|]) + [|y|] / (2 ^ ((Zpos znz_digits) - [|p|]))). + Lemma spec_add_mul_div : forall x y p, + [|p|] <= Zpos znz_digits -> + [| znz_add_mul_div p x y |] = + ([|x|] * (2 ^ [|p|]) + + [|y|] / (2 ^ ((Zpos znz_digits) - [|p|]))) mod wB. + Proof. + intros; unfold znz_add_mul_div; auto. + Qed. + + Definition znz_pos_mod p w := [|w|] mod (2 ^ [|p|]). + Lemma spec_pos_mod : forall w p, + [|znz_pos_mod p w|] = [|w|] mod (2 ^ [|p|]). + Proof. + intros; unfold znz_pos_mod. + apply Zmod_small. + generalize (Z_mod_lt [|w|] (2 ^ [|p|])); intros. + split. + destruct H; auto with zarith. + apply Zle_lt_trans with [|w|]; auto with zarith. + apply Zmod_le; auto with zarith. + Qed. + + Definition znz_is_even x := + if Z_eq_dec ([|x|] mod 2) 0 then true else false. + + Lemma spec_is_even : forall x, + if znz_is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1. + Proof. + intros; unfold znz_is_even; destruct Z_eq_dec; auto. + generalize (Z_mod_lt [|x|] 2); omega. + Qed. + + Definition znz_sqrt x := Zsqrt_plain [|x|]. + Lemma spec_sqrt : forall x, + [|znz_sqrt x|] ^ 2 <= [|x|] < ([|znz_sqrt x|] + 1) ^ 2. + Proof. + intros. + unfold znz_sqrt. + repeat rewrite Zpower_2. + replace [|Zsqrt_plain [|x|]|] with (Zsqrt_plain [|x|]). + apply Zsqrt_interval; auto with zarith. + symmetry; apply Zmod_small. + split. + apply Zsqrt_plain_is_pos; auto with zarith. + + cut (Zsqrt_plain [|x|] <= (wB-1)); try omega. + rewrite <- (Zsqrt_square_id (wB-1)). + apply Zsqrt_le. + split; auto. + apply Zle_trans with (wB-1); auto with zarith. + generalize (spec_to_Z x); auto with zarith. + apply Zsquare_le. + generalize wB_pos; auto with zarith. + Qed. + + 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 + (s, if Z_lt_le_dec r wB then C0 r else C1 (r-wB)) + | Zneg _ => (0, C0 0) + end. + + Lemma spec_sqrt2 : forall x y, + wB/ 4 <= [|x|] -> + let (s,r) := znz_sqrt2 x y in + [||WW x y||] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]. + Proof. + intros; unfold znz_sqrt2. + simpl zn2z_to_Z. + remember ([|x|]*wB+[|y|]) as z. + destruct z. + auto with zarith. + destruct sqrtrempos; intros. + assert (s < wB). + destruct (Z_lt_le_dec s wB); auto. + assert (wB * wB <= Zpos p). + rewrite e. + apply Zle_trans with (s*s); try omega. + apply Zmult_le_compat; generalize wB_pos; auto with zarith. + assert (Zpos p < wB*wB). + rewrite Heqz. + replace (wB*wB) with ((wB-1)*wB+wB) by ring. + apply Zplus_le_lt_compat; auto with zarith. + apply Zmult_le_compat; auto with zarith. + generalize (spec_to_Z x); auto with zarith. + generalize wB_pos; auto with zarith. + omega. + replace [|s|] with s by (symmetry; apply Zmod_small; auto with zarith). + destruct Z_lt_le_dec; unfold interp_carry. + replace [|r|] with r by (symmetry; apply Zmod_small; auto with zarith). + rewrite Zpower_2; auto with zarith. + 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. + Qed. + + Lemma two_p_power2 : forall x, x>=0 -> two_p x = 2 ^ x. + Proof. + intros. + unfold two_p. + destruct x; simpl; auto. + apply two_power_pos_correct. + Qed. + + Definition znz_head0 x := match [|x|] with + | Z0 => znz_zdigits + | Zpos p => znz_zdigits - log_inf p - 1 + | _ => 0 + end. + + Lemma spec_head00: forall x, [|x|] = 0 -> [|znz_head0 x|] = Zpos znz_digits. + Proof. + unfold znz_head0; intros. + rewrite H; simpl. + apply spec_zdigits. + Qed. + + Lemma log_inf_bounded : forall x p, Zpos x < 2^p -> log_inf x < p. + Proof. + induction x; simpl; intros. + + assert (0 < p) by (destruct p; compute; auto with zarith; discriminate). + cut (log_inf x < p - 1); [omega| ]. + apply IHx. + change (Zpos x~1) with (2*(Zpos x)+1) in H. + replace p with (Zsucc (p-1)) in H; auto with zarith. + rewrite Zpower_Zsucc in H; auto with zarith. + + assert (0 < p) by (destruct p; compute; auto with zarith; discriminate). + cut (log_inf x < p - 1); [omega| ]. + apply IHx. + change (Zpos x~0) with (2*(Zpos x)) in H. + replace p with (Zsucc (p-1)) in H; auto with zarith. + rewrite Zpower_Zsucc in H; auto with zarith. + + simpl; intros; destruct p; compute; auto with zarith. + Qed. + + + Lemma spec_head0 : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ ([|znz_head0 x|]) * [|x|] < wB. + Proof. + intros; unfold znz_head0. + generalize (spec_to_Z x). + destruct [|x|]; try discriminate. + intros. + destruct (log_inf_correct p). + rewrite 2 two_p_power2 in H2; auto with zarith. + assert (0 <= znz_zdigits - log_inf p - 1 < wB). + split. + cut (log_inf p < znz_zdigits); try omega. + unfold znz_zdigits. + unfold wB, base in *. + apply log_inf_bounded; auto with zarith. + apply Zlt_trans with znz_zdigits. + omega. + unfold znz_zdigits, wB, base; apply Zpower2_lt_lin; auto with zarith. + + unfold znz_to_Z; rewrite (Zmod_small _ _ H3). + destruct H2. + split. + apply Zle_trans with (2^(znz_zdigits - log_inf p - 1)*(2^log_inf p)). + apply Zdiv_le_upper_bound; auto with zarith. + rewrite <- Zpower_exp; auto with zarith. + rewrite Zmult_comm; rewrite <- Zpower_Zsucc; auto with zarith. + replace (Zsucc (znz_zdigits - log_inf p -1 +log_inf p)) with znz_zdigits + by ring. + unfold wB, base, znz_zdigits; auto with zarith. + apply Zmult_le_compat; auto with zarith. + + 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. + replace (znz_zdigits - log_inf p -1 +Zsucc (log_inf p)) with znz_zdigits + by ring. + unfold wB, base, znz_zdigits; auto with zarith. + Qed. + + Fixpoint Ptail p := match p with + | xO p => (Ptail p)+1 + | _ => 0 + 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). + intros; simpl. + assert (d <> xH). + intro; subst. + compute in H; destruct p; discriminate. + assert (Zsucc (Zpos (Ppred d)) = Zpos d). + simpl; f_equal. + rewrite <- Pplus_one_succ_r. + destruct (Psucc_pred d); auto. + rewrite H1 in H0; elim H0; auto. + assert (Ptail p < Zpos (Ppred d)). + apply IHp. + apply Zmult_lt_reg_r with 2; auto with zarith. + rewrite (Zmult_comm (Zpos p)). + change (2 * Zpos p) with (Zpos p~0). + rewrite Zmult_comm. + rewrite <- Zpower_Zsucc; auto with zarith. + rewrite H1; auto. + rewrite <- H1; omega. + Qed. + + Definition znz_tail0 x := + match [|x|] with + | Z0 => znz_zdigits + | Zpos p => Ptail p + | Zneg _ => 0 + end. + + Lemma spec_tail00: forall x, [|x|] = 0 -> [|znz_tail0 x|] = Zpos znz_digits. + Proof. + unfold znz_tail0; intros. + rewrite H; simpl. + apply spec_zdigits. + Qed. + + Lemma spec_tail0 : forall x, 0 < [|x|] -> + exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|znz_tail0 x|]). + Proof. + intros; unfold znz_tail0. + generalize (spec_to_Z x). + destruct [|x|]; try discriminate; intros. + assert ([|Ptail p|] = Ptail p). + apply Zmod_small. + split; auto. + unfold wB, base in *. + apply Zlt_trans with (Zpos digits). + apply Ptail_bounded; auto with zarith. + apply Zpower2_lt_lin; auto with zarith. + rewrite H1. + + clear; induction p. + exists (Zpos p); simpl; rewrite Pmult_1_r; auto with zarith. + destruct IHp as (y & Yp & Ye). + exists y. + split; auto. + change (Zpos p~0) with (2*Zpos p). + rewrite Ye. + change (Ptail p~0) with (Zsucc (Ptail p)). + rewrite Zpower_Zsucc; auto; ring. + + exists 0; simpl; auto with zarith. + Qed. + + (** Let's now group everything in two records *) + + Definition zmod_op := mk_znz_op + (znz_digits : positive) + (znz_zdigits: znz) + (znz_to_Z : znz -> Z) + (znz_of_pos : positive -> N * znz) + (znz_head0 : znz -> znz) + (znz_tail0 : znz -> znz) + + (znz_0 : znz) + (znz_1 : znz) + (znz_Bm1 : znz) + + (znz_compare : znz -> znz -> comparison) + (znz_eq0 : znz -> bool) + + (znz_opp_c : znz -> carry znz) + (znz_opp : znz -> znz) + (znz_opp_carry : znz -> znz) + + (znz_succ_c : znz -> carry znz) + (znz_add_c : znz -> znz -> carry znz) + (znz_add_carry_c : znz -> znz -> carry znz) + (znz_succ : znz -> znz) + (znz_add : znz -> znz -> znz) + (znz_add_carry : znz -> znz -> znz) + + (znz_pred_c : znz -> carry znz) + (znz_sub_c : znz -> znz -> carry znz) + (znz_sub_carry_c : znz -> znz -> carry znz) + (znz_pred : znz -> znz) + (znz_sub : znz -> znz -> znz) + (znz_sub_carry : znz -> znz -> znz) + + (znz_mul_c : znz -> znz -> zn2z znz) + (znz_mul : znz -> znz -> znz) + (znz_square_c : znz -> zn2z znz) + + (znz_div21 : znz -> znz -> znz -> znz*znz) + (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_gcd_gt : znz -> znz -> znz) + (znz_gcd : znz -> znz -> znz) + (znz_add_mul_div : znz -> znz -> znz -> znz) + (znz_pos_mod : znz -> znz -> znz) + + (znz_is_even : znz -> bool) + (znz_sqrt2 : znz -> znz -> znz * carry znz) + (znz_sqrt : znz -> znz). + + Definition zmod_spec := mk_znz_spec zmod_op + spec_to_Z + spec_of_pos + spec_zdigits + 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_sqrt. + +End ZModulo. + +(** A modular version of the previous construction. *) + +Module Type PositiveNotOne. + Parameter p : positive. + Axiom not_one : p<> 1%positive. +End PositiveNotOne. + +Module ZModuloCyclicType (P:PositiveNotOne) <: CyclicType. + Definition w := Z. + Definition w_op := zmod_op P.p. + Definition w_spec := zmod_spec P.not_one. +End ZModuloCyclicType. + diff --git a/theories/Numbers/Integer/Abstract/ZAdd.v b/theories/Numbers/Integer/Abstract/ZAdd.v new file mode 100644 index 00000000..df941d90 --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZAdd.v @@ -0,0 +1,345 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZAdd.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export ZBase. + +Module ZAddPropFunct (Import ZAxiomsMod : ZAxiomsSig). +Module Export ZBasePropMod := ZBasePropFunct ZAxiomsMod. +Open Local Scope IntScope. + +Theorem Zadd_wd : + forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> n1 + m1 == n2 + m2. +Proof NZadd_wd. + +Theorem Zadd_0_l : forall n : Z, 0 + n == n. +Proof NZadd_0_l. + +Theorem Zadd_succ_l : forall n m : Z, (S n) + m == S (n + m). +Proof NZadd_succ_l. + +Theorem Zsub_0_r : forall n : Z, n - 0 == n. +Proof NZsub_0_r. + +Theorem Zsub_succ_r : forall n m : Z, n - (S m) == P (n - m). +Proof NZsub_succ_r. + +Theorem Zopp_0 : - 0 == 0. +Proof Zopp_0. + +Theorem Zopp_succ : forall n : Z, - (S n) == P (- n). +Proof Zopp_succ. + +(* Theorems that are valid for both natural numbers and integers *) + +Theorem Zadd_0_r : forall n : Z, n + 0 == n. +Proof NZadd_0_r. + +Theorem Zadd_succ_r : forall n m : Z, n + S m == S (n + m). +Proof NZadd_succ_r. + +Theorem Zadd_comm : forall n m : Z, n + m == m + n. +Proof NZadd_comm. + +Theorem Zadd_assoc : forall n m p : Z, n + (m + p) == (n + m) + p. +Proof NZadd_assoc. + +Theorem Zadd_shuffle1 : forall n m p q : Z, (n + m) + (p + q) == (n + p) + (m + q). +Proof NZadd_shuffle1. + +Theorem Zadd_shuffle2 : forall n m p q : Z, (n + m) + (p + q) == (n + q) + (m + p). +Proof NZadd_shuffle2. + +Theorem Zadd_1_l : forall n : Z, 1 + n == S n. +Proof NZadd_1_l. + +Theorem Zadd_1_r : forall n : Z, n + 1 == S n. +Proof NZadd_1_r. + +Theorem Zadd_cancel_l : forall n m p : Z, p + n == p + m <-> n == m. +Proof NZadd_cancel_l. + +Theorem Zadd_cancel_r : forall n m p : Z, n + p == m + p <-> n == m. +Proof NZadd_cancel_r. + +(* Theorems that are either not valid on N or have different proofs on N and Z *) + +Theorem Zadd_pred_l : forall n m : Z, P n + m == P (n + m). +Proof. +intros n m. +rewrite <- (Zsucc_pred n) at 2. +rewrite Zadd_succ_l. now rewrite Zpred_succ. +Qed. + +Theorem Zadd_pred_r : forall n m : Z, n + P m == P (n + m). +Proof. +intros n m; rewrite (Zadd_comm n (P m)), (Zadd_comm n m); +apply Zadd_pred_l. +Qed. + +Theorem Zadd_opp_r : forall n m : Z, n + (- m) == n - m. +Proof. +NZinduct m. +rewrite Zopp_0; rewrite Zsub_0_r; now rewrite Zadd_0_r. +intro m. rewrite Zopp_succ, Zsub_succ_r, Zadd_pred_r; now rewrite Zpred_inj_wd. +Qed. + +Theorem Zsub_0_l : forall n : Z, 0 - n == - n. +Proof. +intro n; rewrite <- Zadd_opp_r; now rewrite Zadd_0_l. +Qed. + +Theorem Zsub_succ_l : forall n m : Z, S n - m == S (n - m). +Proof. +intros n m; do 2 rewrite <- Zadd_opp_r; now rewrite Zadd_succ_l. +Qed. + +Theorem Zsub_pred_l : forall n m : Z, P n - m == P (n - m). +Proof. +intros n m. rewrite <- (Zsucc_pred n) at 2. +rewrite Zsub_succ_l; now rewrite Zpred_succ. +Qed. + +Theorem Zsub_pred_r : forall n m : Z, n - (P m) == S (n - m). +Proof. +intros n m. rewrite <- (Zsucc_pred m) at 2. +rewrite Zsub_succ_r; now rewrite Zsucc_pred. +Qed. + +Theorem Zopp_pred : forall n : Z, - (P n) == S (- n). +Proof. +intro n. rewrite <- (Zsucc_pred n) at 2. +rewrite Zopp_succ. now rewrite Zsucc_pred. +Qed. + +Theorem Zsub_diag : forall n : Z, n - n == 0. +Proof. +NZinduct n. +now rewrite Zsub_0_r. +intro n. rewrite Zsub_succ_r, Zsub_succ_l; now rewrite Zpred_succ. +Qed. + +Theorem Zadd_opp_diag_l : forall n : Z, - n + n == 0. +Proof. +intro n; now rewrite Zadd_comm, Zadd_opp_r, Zsub_diag. +Qed. + +Theorem Zadd_opp_diag_r : forall n : Z, n + (- n) == 0. +Proof. +intro n; rewrite Zadd_comm; apply Zadd_opp_diag_l. +Qed. + +Theorem Zadd_opp_l : forall n m : Z, - m + n == n - m. +Proof. +intros n m; rewrite <- Zadd_opp_r; now rewrite Zadd_comm. +Qed. + +Theorem Zadd_sub_assoc : forall n m p : Z, n + (m - p) == (n + m) - p. +Proof. +intros n m p; do 2 rewrite <- Zadd_opp_r; now rewrite Zadd_assoc. +Qed. + +Theorem Zopp_involutive : forall n : Z, - (- n) == n. +Proof. +NZinduct n. +now do 2 rewrite Zopp_0. +intro n. rewrite Zopp_succ, Zopp_pred; now rewrite Zsucc_inj_wd. +Qed. + +Theorem Zopp_add_distr : forall n m : Z, - (n + m) == - n + (- m). +Proof. +intros n m; NZinduct n. +rewrite Zopp_0; now do 2 rewrite Zadd_0_l. +intro n. rewrite Zadd_succ_l; do 2 rewrite Zopp_succ; rewrite Zadd_pred_l. +now rewrite Zpred_inj_wd. +Qed. + +Theorem Zopp_sub_distr : forall n m : Z, - (n - m) == - n + m. +Proof. +intros n m; rewrite <- Zadd_opp_r, Zopp_add_distr. +now rewrite Zopp_involutive. +Qed. + +Theorem Zopp_inj : forall n m : Z, - n == - m -> n == m. +Proof. +intros n m H. apply Zopp_wd in H. now do 2 rewrite Zopp_involutive in H. +Qed. + +Theorem Zopp_inj_wd : forall n m : Z, - n == - m <-> n == m. +Proof. +intros n m; split; [apply Zopp_inj | apply Zopp_wd]. +Qed. + +Theorem Zeq_opp_l : forall n m : Z, - n == m <-> n == - m. +Proof. +intros n m. now rewrite <- (Zopp_inj_wd (- n) m), Zopp_involutive. +Qed. + +Theorem Zeq_opp_r : forall n m : Z, n == - m <-> - n == m. +Proof. +symmetry; apply Zeq_opp_l. +Qed. + +Theorem Zsub_add_distr : forall n m p : Z, n - (m + p) == (n - m) - p. +Proof. +intros n m p; rewrite <- Zadd_opp_r, Zopp_add_distr, Zadd_assoc. +now do 2 rewrite Zadd_opp_r. +Qed. + +Theorem Zsub_sub_distr : forall n m p : Z, n - (m - p) == (n - m) + p. +Proof. +intros n m p; rewrite <- Zadd_opp_r, Zopp_sub_distr, Zadd_assoc. +now rewrite Zadd_opp_r. +Qed. + +Theorem sub_opp_l : forall n m : Z, - n - m == - m - n. +Proof. +intros n m. do 2 rewrite <- Zadd_opp_r. now rewrite Zadd_comm. +Qed. + +Theorem Zsub_opp_r : forall n m : Z, n - (- m) == n + m. +Proof. +intros n m; rewrite <- Zadd_opp_r; now rewrite Zopp_involutive. +Qed. + +Theorem Zadd_sub_swap : forall n m p : Z, n + m - p == n - p + m. +Proof. +intros n m p. rewrite <- Zadd_sub_assoc, <- (Zadd_opp_r n p), <- Zadd_assoc. +now rewrite Zadd_opp_l. +Qed. + +Theorem Zsub_cancel_l : forall n m p : Z, n - m == n - p <-> m == p. +Proof. +intros n m p. rewrite <- (Zadd_cancel_l (n - m) (n - p) (- n)). +do 2 rewrite Zadd_sub_assoc. rewrite Zadd_opp_diag_l; do 2 rewrite Zsub_0_l. +apply Zopp_inj_wd. +Qed. + +Theorem Zsub_cancel_r : forall n m p : Z, n - p == m - p <-> n == m. +Proof. +intros n m p. +stepl (n - p + p == m - p + p) by apply Zadd_cancel_r. +now do 2 rewrite <- Zsub_sub_distr, Zsub_diag, Zsub_0_r. +Qed. + +(* The next several theorems are devoted to moving terms from one side of +an equation to the other. The name contains the operation in the original +equation (add or sub) and the indication whether the left or right term +is moved. *) + +Theorem Zadd_move_l : forall n m p : Z, n + m == p <-> m == p - n. +Proof. +intros n m p. +stepl (n + m - n == p - n) by apply Zsub_cancel_r. +now rewrite Zadd_comm, <- Zadd_sub_assoc, Zsub_diag, Zadd_0_r. +Qed. + +Theorem Zadd_move_r : forall n m p : Z, n + m == p <-> n == p - m. +Proof. +intros n m p; rewrite Zadd_comm; now apply Zadd_move_l. +Qed. + +(* The two theorems above do not allow rewriting subformulas of the form +n - m == p to n == p + m since subtraction is in the right-hand side of +the equation. Hence the following two theorems. *) + +Theorem Zsub_move_l : forall n m p : Z, n - m == p <-> - m == p - n. +Proof. +intros n m p; rewrite <- (Zadd_opp_r n m); apply Zadd_move_l. +Qed. + +Theorem Zsub_move_r : forall n m p : Z, n - m == p <-> n == p + m. +Proof. +intros n m p; rewrite <- (Zadd_opp_r n m). now rewrite Zadd_move_r, Zsub_opp_r. +Qed. + +Theorem Zadd_move_0_l : forall n m : Z, n + m == 0 <-> m == - n. +Proof. +intros n m; now rewrite Zadd_move_l, Zsub_0_l. +Qed. + +Theorem Zadd_move_0_r : forall n m : Z, n + m == 0 <-> n == - m. +Proof. +intros n m; now rewrite Zadd_move_r, Zsub_0_l. +Qed. + +Theorem Zsub_move_0_l : forall n m : Z, n - m == 0 <-> - m == - n. +Proof. +intros n m. now rewrite Zsub_move_l, Zsub_0_l. +Qed. + +Theorem Zsub_move_0_r : forall n m : Z, n - m == 0 <-> n == m. +Proof. +intros n m. now rewrite Zsub_move_r, Zadd_0_l. +Qed. + +(* The following section is devoted to cancellation of like terms. The name +includes the first operator and the position of the term being canceled. *) + +Theorem Zadd_simpl_l : forall n m : Z, n + m - n == m. +Proof. +intros; now rewrite Zadd_sub_swap, Zsub_diag, Zadd_0_l. +Qed. + +Theorem Zadd_simpl_r : forall n m : Z, n + m - m == n. +Proof. +intros; now rewrite <- Zadd_sub_assoc, Zsub_diag, Zadd_0_r. +Qed. + +Theorem Zsub_simpl_l : forall n m : Z, - n - m + n == - m. +Proof. +intros; now rewrite <- Zadd_sub_swap, Zadd_opp_diag_l, Zsub_0_l. +Qed. + +Theorem Zsub_simpl_r : forall n m : Z, n - m + m == n. +Proof. +intros; now rewrite <- Zsub_sub_distr, Zsub_diag, Zsub_0_r. +Qed. + +(* Now we have two sums or differences; the name includes the two operators +and the position of the terms being canceled *) + +Theorem Zadd_add_simpl_l_l : forall n m p : Z, (n + m) - (n + p) == m - p. +Proof. +intros n m p. now rewrite (Zadd_comm n m), <- Zadd_sub_assoc, +Zsub_add_distr, Zsub_diag, Zsub_0_l, Zadd_opp_r. +Qed. + +Theorem Zadd_add_simpl_l_r : forall n m p : Z, (n + m) - (p + n) == m - p. +Proof. +intros n m p. rewrite (Zadd_comm p n); apply Zadd_add_simpl_l_l. +Qed. + +Theorem Zadd_add_simpl_r_l : forall n m p : Z, (n + m) - (m + p) == n - p. +Proof. +intros n m p. rewrite (Zadd_comm n m); apply Zadd_add_simpl_l_l. +Qed. + +Theorem Zadd_add_simpl_r_r : forall n m p : Z, (n + m) - (p + m) == n - p. +Proof. +intros n m p. rewrite (Zadd_comm p m); apply Zadd_add_simpl_r_l. +Qed. + +Theorem Zsub_add_simpl_r_l : forall n m p : Z, (n - m) + (m + p) == n + p. +Proof. +intros n m p. now rewrite <- Zsub_sub_distr, Zsub_add_distr, Zsub_diag, +Zsub_0_l, Zsub_opp_r. +Qed. + +Theorem Zsub_add_simpl_r_r : forall n m p : Z, (n - m) + (p + m) == n + p. +Proof. +intros n m p. rewrite (Zadd_comm p m); apply Zsub_add_simpl_r_l. +Qed. + +(* Of course, there are many other variants *) + +End ZAddPropFunct. + diff --git a/theories/Numbers/Integer/Abstract/ZAddOrder.v b/theories/Numbers/Integer/Abstract/ZAddOrder.v new file mode 100644 index 00000000..101ea634 --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZAddOrder.v @@ -0,0 +1,373 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZAddOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export ZLt. + +Module ZAddOrderPropFunct (Import ZAxiomsMod : ZAxiomsSig). +Module Export ZOrderPropMod := ZOrderPropFunct ZAxiomsMod. +Open Local Scope IntScope. + +(* Theorems that are true on both natural numbers and integers *) + +Theorem Zadd_lt_mono_l : forall n m p : Z, n < m <-> p + n < p + m. +Proof NZadd_lt_mono_l. + +Theorem Zadd_lt_mono_r : forall n m p : Z, n < m <-> n + p < m + p. +Proof NZadd_lt_mono_r. + +Theorem Zadd_lt_mono : forall n m p q : Z, n < m -> p < q -> n + p < m + q. +Proof NZadd_lt_mono. + +Theorem Zadd_le_mono_l : forall n m p : Z, n <= m <-> p + n <= p + m. +Proof NZadd_le_mono_l. + +Theorem Zadd_le_mono_r : forall n m p : Z, n <= m <-> n + p <= m + p. +Proof NZadd_le_mono_r. + +Theorem Zadd_le_mono : forall n m p q : Z, n <= m -> p <= q -> n + p <= m + q. +Proof NZadd_le_mono. + +Theorem Zadd_lt_le_mono : forall n m p q : Z, n < m -> p <= q -> n + p < m + q. +Proof NZadd_lt_le_mono. + +Theorem Zadd_le_lt_mono : forall n m p q : Z, n <= m -> p < q -> n + p < m + q. +Proof NZadd_le_lt_mono. + +Theorem Zadd_pos_pos : forall n m : Z, 0 < n -> 0 < m -> 0 < n + m. +Proof NZadd_pos_pos. + +Theorem Zadd_pos_nonneg : forall n m : Z, 0 < n -> 0 <= m -> 0 < n + m. +Proof NZadd_pos_nonneg. + +Theorem Zadd_nonneg_pos : forall n m : Z, 0 <= n -> 0 < m -> 0 < n + m. +Proof NZadd_nonneg_pos. + +Theorem Zadd_nonneg_nonneg : forall n m : Z, 0 <= n -> 0 <= m -> 0 <= n + m. +Proof NZadd_nonneg_nonneg. + +Theorem Zlt_add_pos_l : forall n m : Z, 0 < n -> m < n + m. +Proof NZlt_add_pos_l. + +Theorem Zlt_add_pos_r : forall n m : Z, 0 < n -> m < m + n. +Proof NZlt_add_pos_r. + +Theorem Zle_lt_add_lt : forall n m p q : Z, n <= m -> p + m < q + n -> p < q. +Proof NZle_lt_add_lt. + +Theorem Zlt_le_add_lt : forall n m p q : Z, n < m -> p + m <= q + n -> p < q. +Proof NZlt_le_add_lt. + +Theorem Zle_le_add_le : forall n m p q : Z, n <= m -> p + m <= q + n -> p <= q. +Proof NZle_le_add_le. + +Theorem Zadd_lt_cases : forall n m p q : Z, n + m < p + q -> n < p \/ m < q. +Proof NZadd_lt_cases. + +Theorem Zadd_le_cases : forall n m p q : Z, n + m <= p + q -> n <= p \/ m <= q. +Proof NZadd_le_cases. + +Theorem Zadd_neg_cases : forall n m : Z, n + m < 0 -> n < 0 \/ m < 0. +Proof NZadd_neg_cases. + +Theorem Zadd_pos_cases : forall n m : Z, 0 < n + m -> 0 < n \/ 0 < m. +Proof NZadd_pos_cases. + +Theorem Zadd_nonpos_cases : forall n m : Z, n + m <= 0 -> n <= 0 \/ m <= 0. +Proof NZadd_nonpos_cases. + +Theorem Zadd_nonneg_cases : forall n m : Z, 0 <= n + m -> 0 <= n \/ 0 <= m. +Proof NZadd_nonneg_cases. + +(* Theorems that are either not valid on N or have different proofs on N and Z *) + +Theorem Zadd_neg_neg : forall n m : Z, n < 0 -> m < 0 -> n + m < 0. +Proof. +intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_lt_mono. +Qed. + +Theorem Zadd_neg_nonpos : forall n m : Z, n < 0 -> m <= 0 -> n + m < 0. +Proof. +intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_lt_le_mono. +Qed. + +Theorem Zadd_nonpos_neg : forall n m : Z, n <= 0 -> m < 0 -> n + m < 0. +Proof. +intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_le_lt_mono. +Qed. + +Theorem Zadd_nonpos_nonpos : forall n m : Z, n <= 0 -> m <= 0 -> n + m <= 0. +Proof. +intros n m H1 H2. rewrite <- (Zadd_0_l 0). now apply Zadd_le_mono. +Qed. + +(** Sub and order *) + +Theorem Zlt_0_sub : forall n m : Z, 0 < m - n <-> n < m. +Proof. +intros n m. stepl (0 + n < m - n + n) by symmetry; apply Zadd_lt_mono_r. +rewrite Zadd_0_l; now rewrite Zsub_simpl_r. +Qed. + +Notation Zsub_pos := Zlt_0_sub (only parsing). + +Theorem Zle_0_sub : forall n m : Z, 0 <= m - n <-> n <= m. +Proof. +intros n m; stepl (0 + n <= m - n + n) by symmetry; apply Zadd_le_mono_r. +rewrite Zadd_0_l; now rewrite Zsub_simpl_r. +Qed. + +Notation Zsub_nonneg := Zle_0_sub (only parsing). + +Theorem Zlt_sub_0 : forall n m : Z, n - m < 0 <-> n < m. +Proof. +intros n m. stepl (n - m + m < 0 + m) by symmetry; apply Zadd_lt_mono_r. +rewrite Zadd_0_l; now rewrite Zsub_simpl_r. +Qed. + +Notation Zsub_neg := Zlt_sub_0 (only parsing). + +Theorem Zle_sub_0 : forall n m : Z, n - m <= 0 <-> n <= m. +Proof. +intros n m. stepl (n - m + m <= 0 + m) by symmetry; apply Zadd_le_mono_r. +rewrite Zadd_0_l; now rewrite Zsub_simpl_r. +Qed. + +Notation Zsub_nonpos := Zle_sub_0 (only parsing). + +Theorem Zopp_lt_mono : forall n m : Z, n < m <-> - m < - n. +Proof. +intros n m. stepr (m + - m < m + - n) by symmetry; apply Zadd_lt_mono_l. +do 2 rewrite Zadd_opp_r. rewrite Zsub_diag. symmetry; apply Zlt_0_sub. +Qed. + +Theorem Zopp_le_mono : forall n m : Z, n <= m <-> - m <= - n. +Proof. +intros n m. stepr (m + - m <= m + - n) by symmetry; apply Zadd_le_mono_l. +do 2 rewrite Zadd_opp_r. rewrite Zsub_diag. symmetry; apply Zle_0_sub. +Qed. + +Theorem Zopp_pos_neg : forall n : Z, 0 < - n <-> n < 0. +Proof. +intro n; rewrite (Zopp_lt_mono n 0); now rewrite Zopp_0. +Qed. + +Theorem Zopp_neg_pos : forall n : Z, - n < 0 <-> 0 < n. +Proof. +intro n. rewrite (Zopp_lt_mono 0 n). now rewrite Zopp_0. +Qed. + +Theorem Zopp_nonneg_nonpos : forall n : Z, 0 <= - n <-> n <= 0. +Proof. +intro n; rewrite (Zopp_le_mono n 0); now rewrite Zopp_0. +Qed. + +Theorem Zopp_nonpos_nonneg : forall n : Z, - n <= 0 <-> 0 <= n. +Proof. +intro n. rewrite (Zopp_le_mono 0 n). now rewrite Zopp_0. +Qed. + +Theorem Zsub_lt_mono_l : forall n m p : Z, n < m <-> p - m < p - n. +Proof. +intros n m p. do 2 rewrite <- Zadd_opp_r. rewrite <- Zadd_lt_mono_l. +apply Zopp_lt_mono. +Qed. + +Theorem Zsub_lt_mono_r : forall n m p : Z, n < m <-> n - p < m - p. +Proof. +intros n m p; do 2 rewrite <- Zadd_opp_r; apply Zadd_lt_mono_r. +Qed. + +Theorem Zsub_lt_mono : forall n m p q : Z, n < m -> q < p -> n - p < m - q. +Proof. +intros n m p q H1 H2. +apply NZlt_trans with (m - p); +[now apply -> Zsub_lt_mono_r | now apply -> Zsub_lt_mono_l]. +Qed. + +Theorem Zsub_le_mono_l : forall n m p : Z, n <= m <-> p - m <= p - n. +Proof. +intros n m p; do 2 rewrite <- Zadd_opp_r; rewrite <- Zadd_le_mono_l; +apply Zopp_le_mono. +Qed. + +Theorem Zsub_le_mono_r : forall n m p : Z, n <= m <-> n - p <= m - p. +Proof. +intros n m p; do 2 rewrite <- Zadd_opp_r; apply Zadd_le_mono_r. +Qed. + +Theorem Zsub_le_mono : forall n m p q : Z, n <= m -> q <= p -> n - p <= m - q. +Proof. +intros n m p q H1 H2. +apply NZle_trans with (m - p); +[now apply -> Zsub_le_mono_r | now apply -> Zsub_le_mono_l]. +Qed. + +Theorem Zsub_lt_le_mono : forall n m p q : Z, n < m -> q <= p -> n - p < m - q. +Proof. +intros n m p q H1 H2. +apply NZlt_le_trans with (m - p); +[now apply -> Zsub_lt_mono_r | now apply -> Zsub_le_mono_l]. +Qed. + +Theorem Zsub_le_lt_mono : forall n m p q : Z, n <= m -> q < p -> n - p < m - q. +Proof. +intros n m p q H1 H2. +apply NZle_lt_trans with (m - p); +[now apply -> Zsub_le_mono_r | now apply -> Zsub_lt_mono_l]. +Qed. + +Theorem Zle_lt_sub_lt : forall n m p q : Z, n <= m -> p - n < q - m -> p < q. +Proof. +intros n m p q H1 H2. apply (Zle_lt_add_lt (- m) (- n)); +[now apply -> Zopp_le_mono | now do 2 rewrite Zadd_opp_r]. +Qed. + +Theorem Zlt_le_sub_lt : forall n m p q : Z, n < m -> p - n <= q - m -> p < q. +Proof. +intros n m p q H1 H2. apply (Zlt_le_add_lt (- m) (- n)); +[now apply -> Zopp_lt_mono | now do 2 rewrite Zadd_opp_r]. +Qed. + +Theorem Zle_le_sub_lt : forall n m p q : Z, n <= m -> p - n <= q - m -> p <= q. +Proof. +intros n m p q H1 H2. apply (Zle_le_add_le (- m) (- n)); +[now apply -> Zopp_le_mono | now do 2 rewrite Zadd_opp_r]. +Qed. + +Theorem Zlt_add_lt_sub_r : forall n m p : Z, n + p < m <-> n < m - p. +Proof. +intros n m p. stepl (n + p - p < m - p) by symmetry; apply Zsub_lt_mono_r. +now rewrite Zadd_simpl_r. +Qed. + +Theorem Zle_add_le_sub_r : forall n m p : Z, n + p <= m <-> n <= m - p. +Proof. +intros n m p. stepl (n + p - p <= m - p) by symmetry; apply Zsub_le_mono_r. +now rewrite Zadd_simpl_r. +Qed. + +Theorem Zlt_add_lt_sub_l : forall n m p : Z, n + p < m <-> p < m - n. +Proof. +intros n m p. rewrite Zadd_comm; apply Zlt_add_lt_sub_r. +Qed. + +Theorem Zle_add_le_sub_l : forall n m p : Z, n + p <= m <-> p <= m - n. +Proof. +intros n m p. rewrite Zadd_comm; apply Zle_add_le_sub_r. +Qed. + +Theorem Zlt_sub_lt_add_r : forall n m p : Z, n - p < m <-> n < m + p. +Proof. +intros n m p. stepl (n - p + p < m + p) by symmetry; apply Zadd_lt_mono_r. +now rewrite Zsub_simpl_r. +Qed. + +Theorem Zle_sub_le_add_r : forall n m p : Z, n - p <= m <-> n <= m + p. +Proof. +intros n m p. stepl (n - p + p <= m + p) by symmetry; apply Zadd_le_mono_r. +now rewrite Zsub_simpl_r. +Qed. + +Theorem Zlt_sub_lt_add_l : forall n m p : Z, n - m < p <-> n < m + p. +Proof. +intros n m p. rewrite Zadd_comm; apply Zlt_sub_lt_add_r. +Qed. + +Theorem Zle_sub_le_add_l : forall n m p : Z, n - m <= p <-> n <= m + p. +Proof. +intros n m p. rewrite Zadd_comm; apply Zle_sub_le_add_r. +Qed. + +Theorem Zlt_sub_lt_add : forall n m p q : Z, n - m < p - q <-> n + q < m + p. +Proof. +intros n m p q. rewrite Zlt_sub_lt_add_l. rewrite Zadd_sub_assoc. +now rewrite <- Zlt_add_lt_sub_r. +Qed. + +Theorem Zle_sub_le_add : forall n m p q : Z, n - m <= p - q <-> n + q <= m + p. +Proof. +intros n m p q. rewrite Zle_sub_le_add_l. rewrite Zadd_sub_assoc. +now rewrite <- Zle_add_le_sub_r. +Qed. + +Theorem Zlt_sub_pos : forall n m : Z, 0 < m <-> n - m < n. +Proof. +intros n m. stepr (n - m < n - 0) by now rewrite Zsub_0_r. apply Zsub_lt_mono_l. +Qed. + +Theorem Zle_sub_nonneg : forall n m : Z, 0 <= m <-> n - m <= n. +Proof. +intros n m. stepr (n - m <= n - 0) by now rewrite Zsub_0_r. apply Zsub_le_mono_l. +Qed. + +Theorem Zsub_lt_cases : forall n m p q : Z, n - m < p - q -> n < m \/ q < p. +Proof. +intros n m p q H. rewrite Zlt_sub_lt_add in H. now apply Zadd_lt_cases. +Qed. + +Theorem Zsub_le_cases : forall n m p q : Z, n - m <= p - q -> n <= m \/ q <= p. +Proof. +intros n m p q H. rewrite Zle_sub_le_add in H. now apply Zadd_le_cases. +Qed. + +Theorem Zsub_neg_cases : forall n m : Z, n - m < 0 -> n < 0 \/ 0 < m. +Proof. +intros n m H; rewrite <- Zadd_opp_r in H. +setoid_replace (0 < m) with (- m < 0) using relation iff by (symmetry; apply Zopp_neg_pos). +now apply Zadd_neg_cases. +Qed. + +Theorem Zsub_pos_cases : forall n m : Z, 0 < n - m -> 0 < n \/ m < 0. +Proof. +intros n m H; rewrite <- Zadd_opp_r in H. +setoid_replace (m < 0) with (0 < - m) using relation iff by (symmetry; apply Zopp_pos_neg). +now apply Zadd_pos_cases. +Qed. + +Theorem Zsub_nonpos_cases : forall n m : Z, n - m <= 0 -> n <= 0 \/ 0 <= m. +Proof. +intros n m H; rewrite <- Zadd_opp_r in H. +setoid_replace (0 <= m) with (- m <= 0) using relation iff by (symmetry; apply Zopp_nonpos_nonneg). +now apply Zadd_nonpos_cases. +Qed. + +Theorem Zsub_nonneg_cases : forall n m : Z, 0 <= n - m -> 0 <= n \/ m <= 0. +Proof. +intros n m H; rewrite <- Zadd_opp_r in H. +setoid_replace (m <= 0) with (0 <= - m) using relation iff by (symmetry; apply Zopp_nonneg_nonpos). +now apply Zadd_nonneg_cases. +Qed. + +Section PosNeg. + +Variable P : Z -> Prop. +Hypothesis P_wd : predicate_wd Zeq P. + +Add Morphism P with signature Zeq ==> iff as P_morph. Proof. exact P_wd. Qed. + +Theorem Z0_pos_neg : + P 0 -> (forall n : Z, 0 < n -> P n /\ P (- n)) -> forall n : Z, P n. +Proof. +intros H1 H2 n. destruct (Zlt_trichotomy n 0) as [H3 | [H3 | H3]]. +apply <- Zopp_pos_neg in H3. apply H2 in H3. destruct H3 as [_ H3]. +now rewrite Zopp_involutive in H3. +now rewrite H3. +apply H2 in H3; now destruct H3. +Qed. + +End PosNeg. + +Ltac Z0_pos_neg n := induction_maker n ltac:(apply Z0_pos_neg). + +End ZAddOrderPropFunct. + + diff --git a/theories/Numbers/Integer/Abstract/ZAxioms.v b/theories/Numbers/Integer/Abstract/ZAxioms.v new file mode 100644 index 00000000..c4a4b6b8 --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZAxioms.v @@ -0,0 +1,65 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NZAxioms. + +Set Implicit Arguments. + +Module Type ZAxiomsSig. +Declare Module Export NZOrdAxiomsMod : NZOrdAxiomsSig. + +Delimit Scope IntScope with Int. +Notation Z := NZ. +Notation Zeq := NZeq. +Notation Z0 := NZ0. +Notation Z1 := (NZsucc NZ0). +Notation S := NZsucc. +Notation P := NZpred. +Notation Zadd := NZadd. +Notation Zmul := NZmul. +Notation Zsub := NZsub. +Notation Zlt := NZlt. +Notation Zle := NZle. +Notation Zmin := NZmin. +Notation Zmax := NZmax. +Notation "x == y" := (NZeq x y) (at level 70) : IntScope. +Notation "x ~= y" := (~ NZeq x y) (at level 70) : IntScope. +Notation "0" := NZ0 : IntScope. +Notation "1" := (NZsucc NZ0) : IntScope. +Notation "x + y" := (NZadd x y) : IntScope. +Notation "x - y" := (NZsub x y) : IntScope. +Notation "x * y" := (NZmul x y) : IntScope. +Notation "x < y" := (NZlt x y) : IntScope. +Notation "x <= y" := (NZle x y) : IntScope. +Notation "x > y" := (NZlt y x) (only parsing) : IntScope. +Notation "x >= y" := (NZle y x) (only parsing) : IntScope. + +Parameter Zopp : Z -> Z. + +(*Notation "- 1" := (Zopp 1) : IntScope. +Check (-1).*) + +Add Morphism Zopp with signature Zeq ==> Zeq as Zopp_wd. + +Notation "- x" := (Zopp x) (at level 35, right associativity) : IntScope. +Notation "- 1" := (Zopp (NZsucc NZ0)) : IntScope. + +Open Local Scope IntScope. + +(* Integers are obtained by postulating that every number has a predecessor *) +Axiom Zsucc_pred : forall n : Z, S (P n) == n. + +Axiom Zopp_0 : - 0 == 0. +Axiom Zopp_succ : forall n : Z, - (S n) == P (- n). + +End ZAxiomsSig. + diff --git a/theories/Numbers/Integer/Abstract/ZBase.v b/theories/Numbers/Integer/Abstract/ZBase.v new file mode 100644 index 00000000..29e18548 --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZBase.v @@ -0,0 +1,86 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZBase.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export Decidable. +Require Export ZAxioms. +Require Import NZMulOrder. + +Module ZBasePropFunct (Import ZAxiomsMod : ZAxiomsSig). + +(* Note: writing "Export" instead of "Import" on the previous line leads to +some warnings about hiding repeated declarations and results in the loss of +notations in Zadd and later *) + +Open Local Scope IntScope. + +Module Export NZMulOrderMod := NZMulOrderPropFunct NZOrdAxiomsMod. + +Theorem Zsucc_wd : forall n1 n2 : Z, n1 == n2 -> S n1 == S n2. +Proof NZsucc_wd. + +Theorem Zpred_wd : forall n1 n2 : Z, n1 == n2 -> P n1 == P n2. +Proof NZpred_wd. + +Theorem Zpred_succ : forall n : Z, P (S n) == n. +Proof NZpred_succ. + +Theorem Zeq_refl : forall n : Z, n == n. +Proof (proj1 NZeq_equiv). + +Theorem Zeq_symm : forall n m : Z, n == m -> m == n. +Proof (proj2 (proj2 NZeq_equiv)). + +Theorem Zeq_trans : forall n m p : Z, n == m -> m == p -> n == p. +Proof (proj1 (proj2 NZeq_equiv)). + +Theorem Zneq_symm : forall n m : Z, n ~= m -> m ~= n. +Proof NZneq_symm. + +Theorem Zsucc_inj : forall n1 n2 : Z, S n1 == S n2 -> n1 == n2. +Proof NZsucc_inj. + +Theorem Zsucc_inj_wd : forall n1 n2 : Z, S n1 == S n2 <-> n1 == n2. +Proof NZsucc_inj_wd. + +Theorem Zsucc_inj_wd_neg : forall n m : Z, S n ~= S m <-> n ~= m. +Proof NZsucc_inj_wd_neg. + +(* Decidability and stability of equality was proved only in NZOrder, but +since it does not mention order, we'll put it here *) + +Theorem Zeq_dec : forall n m : Z, decidable (n == m). +Proof NZeq_dec. + +Theorem Zeq_dne : forall n m : Z, ~ ~ n == m <-> n == m. +Proof NZeq_dne. + +Theorem Zcentral_induction : +forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, A z -> + (forall n : Z, A n <-> A (S n)) -> + forall n : Z, A n. +Proof NZcentral_induction. + +(* Theorems that are true for integers but not for natural numbers *) + +Theorem Zpred_inj : forall n m : Z, P n == P m -> n == m. +Proof. +intros n m H. apply NZsucc_wd in H. now do 2 rewrite Zsucc_pred in H. +Qed. + +Theorem Zpred_inj_wd : forall n1 n2 : Z, P n1 == P n2 <-> n1 == n2. +Proof. +intros n1 n2; split; [apply Zpred_inj | apply NZpred_wd]. +Qed. + +End ZBasePropFunct. + diff --git a/theories/Numbers/Integer/Abstract/ZDomain.v b/theories/Numbers/Integer/Abstract/ZDomain.v new file mode 100644 index 00000000..15beb2b9 --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZDomain.v @@ -0,0 +1,69 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZDomain.v 10934 2008-05-15 21:58:20Z letouzey $ i*) + +Require Export NumPrelude. + +Module Type ZDomainSignature. + +Parameter Inline Z : Set. +Parameter Inline Zeq : Z -> Z -> Prop. +Parameter Inline e : Z -> Z -> bool. + +Axiom eq_equiv_e : forall x y : Z, Zeq x y <-> e x y. +Axiom eq_equiv : equiv Z Zeq. + +Add Relation Z Zeq + reflexivity proved by (proj1 eq_equiv) + symmetry proved by (proj2 (proj2 eq_equiv)) + transitivity proved by (proj1 (proj2 eq_equiv)) +as eq_rel. + +Delimit Scope IntScope with Int. +Bind Scope IntScope with Z. +Notation "x == y" := (Zeq x y) (at level 70) : IntScope. +Notation "x # y" := (~ Zeq x y) (at level 70) : IntScope. + +End ZDomainSignature. + +Module ZDomainProperties (Import ZDomainModule : ZDomainSignature). +Open Local Scope IntScope. + +Add Morphism e with signature Zeq ==> Zeq ==> eq_bool as e_wd. +Proof. +intros x x' Exx' y y' Eyy'. +case_eq (e x y); case_eq (e x' y'); intros H1 H2; trivial. +assert (x == y); [apply <- eq_equiv_e; now rewrite H2 | +assert (x' == y'); [rewrite <- Exx'; now rewrite <- Eyy' | +rewrite <- H1; assert (H3 : e x' y'); [now apply -> eq_equiv_e | now inversion H3]]]. +assert (x' == y'); [apply <- eq_equiv_e; now rewrite H1 | +assert (x == y); [rewrite Exx'; now rewrite Eyy' | +rewrite <- H2; assert (H3 : e x y); [now apply -> eq_equiv_e | now inversion H3]]]. +Qed. + +Theorem neq_symm : forall n m, n # m -> m # n. +Proof. +intros n m H1 H2; symmetry in H2; false_hyp H2 H1. +Qed. + +Theorem ZE_stepl : forall x y z : Z, x == y -> x == z -> z == y. +Proof. +intros x y z H1 H2; now rewrite <- H1. +Qed. + +Declare Left Step ZE_stepl. + +(* The right step lemma is just transitivity of Zeq *) +Declare Right Step (proj1 (proj2 eq_equiv)). + +End ZDomainProperties. + + diff --git a/theories/Numbers/Integer/Abstract/ZLt.v b/theories/Numbers/Integer/Abstract/ZLt.v new file mode 100644 index 00000000..2a88a535 --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZLt.v @@ -0,0 +1,432 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZLt.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export ZMul. + +Module ZOrderPropFunct (Import ZAxiomsMod : ZAxiomsSig). +Module Export ZMulPropMod := ZMulPropFunct ZAxiomsMod. +Open Local Scope IntScope. + +(* Axioms *) + +Theorem Zlt_wd : + forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> (n1 < m1 <-> n2 < m2). +Proof NZlt_wd. + +Theorem Zle_wd : + forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> (n1 <= m1 <-> n2 <= m2). +Proof NZle_wd. + +Theorem Zmin_wd : + forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> Zmin n1 m1 == Zmin n2 m2. +Proof NZmin_wd. + +Theorem Zmax_wd : + forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> Zmax n1 m1 == Zmax n2 m2. +Proof NZmax_wd. + +Theorem Zlt_eq_cases : forall n m : Z, n <= m <-> n < m \/ n == m. +Proof NZlt_eq_cases. + +Theorem Zlt_irrefl : forall n : Z, ~ n < n. +Proof NZlt_irrefl. + +Theorem Zlt_succ_r : forall n m : Z, n < S m <-> n <= m. +Proof NZlt_succ_r. + +Theorem Zmin_l : forall n m : Z, n <= m -> Zmin n m == n. +Proof NZmin_l. + +Theorem Zmin_r : forall n m : Z, m <= n -> Zmin n m == m. +Proof NZmin_r. + +Theorem Zmax_l : forall n m : Z, m <= n -> Zmax n m == n. +Proof NZmax_l. + +Theorem Zmax_r : forall n m : Z, n <= m -> Zmax n m == m. +Proof NZmax_r. + +(* Renaming theorems from NZOrder.v *) + +Theorem Zlt_le_incl : forall n m : Z, n < m -> n <= m. +Proof NZlt_le_incl. + +Theorem Zlt_neq : forall n m : Z, n < m -> n ~= m. +Proof NZlt_neq. + +Theorem Zle_neq : forall n m : Z, n < m <-> n <= m /\ n ~= m. +Proof NZle_neq. + +Theorem Zle_refl : forall n : Z, n <= n. +Proof NZle_refl. + +Theorem Zlt_succ_diag_r : forall n : Z, n < S n. +Proof NZlt_succ_diag_r. + +Theorem Zle_succ_diag_r : forall n : Z, n <= S n. +Proof NZle_succ_diag_r. + +Theorem Zlt_0_1 : 0 < 1. +Proof NZlt_0_1. + +Theorem Zle_0_1 : 0 <= 1. +Proof NZle_0_1. + +Theorem Zlt_lt_succ_r : forall n m : Z, n < m -> n < S m. +Proof NZlt_lt_succ_r. + +Theorem Zle_le_succ_r : forall n m : Z, n <= m -> n <= S m. +Proof NZle_le_succ_r. + +Theorem Zle_succ_r : forall n m : Z, n <= S m <-> n <= m \/ n == S m. +Proof NZle_succ_r. + +Theorem Zneq_succ_diag_l : forall n : Z, S n ~= n. +Proof NZneq_succ_diag_l. + +Theorem Zneq_succ_diag_r : forall n : Z, n ~= S n. +Proof NZneq_succ_diag_r. + +Theorem Znlt_succ_diag_l : forall n : Z, ~ S n < n. +Proof NZnlt_succ_diag_l. + +Theorem Znle_succ_diag_l : forall n : Z, ~ S n <= n. +Proof NZnle_succ_diag_l. + +Theorem Zle_succ_l : forall n m : Z, S n <= m <-> n < m. +Proof NZle_succ_l. + +Theorem Zlt_succ_l : forall n m : Z, S n < m -> n < m. +Proof NZlt_succ_l. + +Theorem Zsucc_lt_mono : forall n m : Z, n < m <-> S n < S m. +Proof NZsucc_lt_mono. + +Theorem Zsucc_le_mono : forall n m : Z, n <= m <-> S n <= S m. +Proof NZsucc_le_mono. + +Theorem Zlt_asymm : forall n m, n < m -> ~ m < n. +Proof NZlt_asymm. + +Notation Zlt_ngt := Zlt_asymm (only parsing). + +Theorem Zlt_trans : forall n m p : Z, n < m -> m < p -> n < p. +Proof NZlt_trans. + +Theorem Zle_trans : forall n m p : Z, n <= m -> m <= p -> n <= p. +Proof NZle_trans. + +Theorem Zle_lt_trans : forall n m p : Z, n <= m -> m < p -> n < p. +Proof NZle_lt_trans. + +Theorem Zlt_le_trans : forall n m p : Z, n < m -> m <= p -> n < p. +Proof NZlt_le_trans. + +Theorem Zle_antisymm : forall n m : Z, n <= m -> m <= n -> n == m. +Proof NZle_antisymm. + +Theorem Zlt_1_l : forall n m : Z, 0 < n -> n < m -> 1 < m. +Proof NZlt_1_l. + +(** Trichotomy, decidability, and double negation elimination *) + +Theorem Zlt_trichotomy : forall n m : Z, n < m \/ n == m \/ m < n. +Proof NZlt_trichotomy. + +Notation Zlt_eq_gt_cases := Zlt_trichotomy (only parsing). + +Theorem Zlt_gt_cases : forall n m : Z, n ~= m <-> n < m \/ n > m. +Proof NZlt_gt_cases. + +Theorem Zle_gt_cases : forall n m : Z, n <= m \/ n > m. +Proof NZle_gt_cases. + +Theorem Zlt_ge_cases : forall n m : Z, n < m \/ n >= m. +Proof NZlt_ge_cases. + +Theorem Zle_ge_cases : forall n m : Z, n <= m \/ n >= m. +Proof NZle_ge_cases. + +(** Instances of the previous theorems for m == 0 *) + +Theorem Zneg_pos_cases : forall n : Z, n ~= 0 <-> n < 0 \/ n > 0. +Proof. +intro; apply Zlt_gt_cases. +Qed. + +Theorem Znonpos_pos_cases : forall n : Z, n <= 0 \/ n > 0. +Proof. +intro; apply Zle_gt_cases. +Qed. + +Theorem Zneg_nonneg_cases : forall n : Z, n < 0 \/ n >= 0. +Proof. +intro; apply Zlt_ge_cases. +Qed. + +Theorem Znonpos_nonneg_cases : forall n : Z, n <= 0 \/ n >= 0. +Proof. +intro; apply Zle_ge_cases. +Qed. + +Theorem Zle_ngt : forall n m : Z, n <= m <-> ~ n > m. +Proof NZle_ngt. + +Theorem Znlt_ge : forall n m : Z, ~ n < m <-> n >= m. +Proof NZnlt_ge. + +Theorem Zlt_dec : forall n m : Z, decidable (n < m). +Proof NZlt_dec. + +Theorem Zlt_dne : forall n m, ~ ~ n < m <-> n < m. +Proof NZlt_dne. + +Theorem Znle_gt : forall n m : Z, ~ n <= m <-> n > m. +Proof NZnle_gt. + +Theorem Zlt_nge : forall n m : Z, n < m <-> ~ n >= m. +Proof NZlt_nge. + +Theorem Zle_dec : forall n m : Z, decidable (n <= m). +Proof NZle_dec. + +Theorem Zle_dne : forall n m : Z, ~ ~ n <= m <-> n <= m. +Proof NZle_dne. + +Theorem Znlt_succ_r : forall n m : Z, ~ m < S n <-> n < m. +Proof NZnlt_succ_r. + +Theorem Zlt_exists_pred : + forall z n : Z, z < n -> exists k : Z, n == S k /\ z <= k. +Proof NZlt_exists_pred. + +Theorem Zlt_succ_iter_r : + forall (n : nat) (m : Z), m < NZsucc_iter (Datatypes.S n) m. +Proof NZlt_succ_iter_r. + +Theorem Zneq_succ_iter_l : + forall (n : nat) (m : Z), NZsucc_iter (Datatypes.S n) m ~= m. +Proof NZneq_succ_iter_l. + +(** Stronger variant of induction with assumptions n >= 0 (n < 0) +in the induction step *) + +Theorem Zright_induction : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, A z -> + (forall n : Z, z <= n -> A n -> A (S n)) -> + forall n : Z, z <= n -> A n. +Proof NZright_induction. + +Theorem Zleft_induction : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, A z -> + (forall n : Z, n < z -> A (S n) -> A n) -> + forall n : Z, n <= z -> A n. +Proof NZleft_induction. + +Theorem Zright_induction' : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, + (forall n : Z, n <= z -> A n) -> + (forall n : Z, z <= n -> A n -> A (S n)) -> + forall n : Z, A n. +Proof NZright_induction'. + +Theorem Zleft_induction' : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, + (forall n : Z, z <= n -> A n) -> + (forall n : Z, n < z -> A (S n) -> A n) -> + forall n : Z, A n. +Proof NZleft_induction'. + +Theorem Zstrong_right_induction : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, + (forall n : Z, z <= n -> (forall m : Z, z <= m -> m < n -> A m) -> A n) -> + forall n : Z, z <= n -> A n. +Proof NZstrong_right_induction. + +Theorem Zstrong_left_induction : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, + (forall n : Z, n <= z -> (forall m : Z, m <= z -> S n <= m -> A m) -> A n) -> + forall n : Z, n <= z -> A n. +Proof NZstrong_left_induction. + +Theorem Zstrong_right_induction' : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, + (forall n : Z, n <= z -> A n) -> + (forall n : Z, z <= n -> (forall m : Z, z <= m -> m < n -> A m) -> A n) -> + forall n : Z, A n. +Proof NZstrong_right_induction'. + +Theorem Zstrong_left_induction' : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, + (forall n : Z, z <= n -> A n) -> + (forall n : Z, n <= z -> (forall m : Z, m <= z -> S n <= m -> A m) -> A n) -> + forall n : Z, A n. +Proof NZstrong_left_induction'. + +Theorem Zorder_induction : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, A z -> + (forall n : Z, z <= n -> A n -> A (S n)) -> + (forall n : Z, n < z -> A (S n) -> A n) -> + forall n : Z, A n. +Proof NZorder_induction. + +Theorem Zorder_induction' : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall z : Z, A z -> + (forall n : Z, z <= n -> A n -> A (S n)) -> + (forall n : Z, n <= z -> A n -> A (P n)) -> + forall n : Z, A n. +Proof NZorder_induction'. + +Theorem Zorder_induction_0 : + forall A : Z -> Prop, predicate_wd Zeq A -> + A 0 -> + (forall n : Z, 0 <= n -> A n -> A (S n)) -> + (forall n : Z, n < 0 -> A (S n) -> A n) -> + forall n : Z, A n. +Proof NZorder_induction_0. + +Theorem Zorder_induction'_0 : + forall A : Z -> Prop, predicate_wd Zeq A -> + A 0 -> + (forall n : Z, 0 <= n -> A n -> A (S n)) -> + (forall n : Z, n <= 0 -> A n -> A (P n)) -> + forall n : Z, A n. +Proof NZorder_induction'_0. + +Ltac Zinduct n := induction_maker n ltac:(apply Zorder_induction_0). + +(** Elimintation principle for < *) + +Theorem Zlt_ind : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall n : Z, A (S n) -> + (forall m : Z, n < m -> A m -> A (S m)) -> forall m : Z, n < m -> A m. +Proof NZlt_ind. + +(** Elimintation principle for <= *) + +Theorem Zle_ind : + forall A : Z -> Prop, predicate_wd Zeq A -> + forall n : Z, A n -> + (forall m : Z, n <= m -> A m -> A (S m)) -> forall m : Z, n <= m -> A m. +Proof NZle_ind. + +(** Well-founded relations *) + +Theorem Zlt_wf : forall z : Z, well_founded (fun n m : Z => z <= n /\ n < m). +Proof NZlt_wf. + +Theorem Zgt_wf : forall z : Z, well_founded (fun n m : Z => m < n /\ n <= z). +Proof NZgt_wf. + +(* Theorems that are either not valid on N or have different proofs on N and Z *) + +Theorem Zlt_pred_l : forall n : Z, P n < n. +Proof. +intro n; rewrite <- (Zsucc_pred n) at 2; apply Zlt_succ_diag_r. +Qed. + +Theorem Zle_pred_l : forall n : Z, P n <= n. +Proof. +intro; apply Zlt_le_incl; apply Zlt_pred_l. +Qed. + +Theorem Zlt_le_pred : forall n m : Z, n < m <-> n <= P m. +Proof. +intros n m; rewrite <- (Zsucc_pred m); rewrite Zpred_succ. apply Zlt_succ_r. +Qed. + +Theorem Znle_pred_r : forall n : Z, ~ n <= P n. +Proof. +intro; rewrite <- Zlt_le_pred; apply Zlt_irrefl. +Qed. + +Theorem Zlt_pred_le : forall n m : Z, P n < m <-> n <= m. +Proof. +intros n m; rewrite <- (Zsucc_pred n) at 2. +symmetry; apply Zle_succ_l. +Qed. + +Theorem Zlt_lt_pred : forall n m : Z, n < m -> P n < m. +Proof. +intros; apply <- Zlt_pred_le; now apply Zlt_le_incl. +Qed. + +Theorem Zle_le_pred : forall n m : Z, n <= m -> P n <= m. +Proof. +intros; apply Zlt_le_incl; now apply <- Zlt_pred_le. +Qed. + +Theorem Zlt_pred_lt : forall n m : Z, n < P m -> n < m. +Proof. +intros n m H; apply Zlt_trans with (P m); [assumption | apply Zlt_pred_l]. +Qed. + +Theorem Zle_pred_lt : forall n m : Z, n <= P m -> n <= m. +Proof. +intros; apply Zlt_le_incl; now apply <- Zlt_le_pred. +Qed. + +Theorem Zpred_lt_mono : forall n m : Z, n < m <-> P n < P m. +Proof. +intros; rewrite Zlt_le_pred; symmetry; apply Zlt_pred_le. +Qed. + +Theorem Zpred_le_mono : forall n m : Z, n <= m <-> P n <= P m. +Proof. +intros; rewrite <- Zlt_pred_le; now rewrite Zlt_le_pred. +Qed. + +Theorem Zlt_succ_lt_pred : forall n m : Z, S n < m <-> n < P m. +Proof. +intros n m; now rewrite (Zpred_lt_mono (S n) m), Zpred_succ. +Qed. + +Theorem Zle_succ_le_pred : forall n m : Z, S n <= m <-> n <= P m. +Proof. +intros n m; now rewrite (Zpred_le_mono (S n) m), Zpred_succ. +Qed. + +Theorem Zlt_pred_lt_succ : forall n m : Z, P n < m <-> n < S m. +Proof. +intros; rewrite Zlt_pred_le; symmetry; apply Zlt_succ_r. +Qed. + +Theorem Zle_pred_lt_succ : forall n m : Z, P n <= m <-> n <= S m. +Proof. +intros n m; now rewrite (Zpred_le_mono n (S m)), Zpred_succ. +Qed. + +Theorem Zneq_pred_l : forall n : Z, P n ~= n. +Proof. +intro; apply Zlt_neq; apply Zlt_pred_l. +Qed. + +Theorem Zlt_n1_r : forall n m : Z, n < m -> m < 0 -> n < -1. +Proof. +intros n m H1 H2. apply -> Zlt_le_pred in H2. +setoid_replace (P 0) with (-1) in H2. now apply NZlt_le_trans with m. +apply <- Zeq_opp_r. now rewrite Zopp_pred, Zopp_0. +Qed. + +End ZOrderPropFunct. + diff --git a/theories/Numbers/Integer/Abstract/ZMul.v b/theories/Numbers/Integer/Abstract/ZMul.v new file mode 100644 index 00000000..c48d1b4c --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZMul.v @@ -0,0 +1,115 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZMul.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export ZAdd. + +Module ZMulPropFunct (Import ZAxiomsMod : ZAxiomsSig). +Module Export ZAddPropMod := ZAddPropFunct ZAxiomsMod. +Open Local Scope IntScope. + +Theorem Zmul_wd : + forall n1 n2 : Z, n1 == n2 -> forall m1 m2 : Z, m1 == m2 -> n1 * m1 == n2 * m2. +Proof NZmul_wd. + +Theorem Zmul_0_l : forall n : Z, 0 * n == 0. +Proof NZmul_0_l. + +Theorem Zmul_succ_l : forall n m : Z, (S n) * m == n * m + m. +Proof NZmul_succ_l. + +(* Theorems that are valid for both natural numbers and integers *) + +Theorem Zmul_0_r : forall n : Z, n * 0 == 0. +Proof NZmul_0_r. + +Theorem Zmul_succ_r : forall n m : Z, n * (S m) == n * m + n. +Proof NZmul_succ_r. + +Theorem Zmul_comm : forall n m : Z, n * m == m * n. +Proof NZmul_comm. + +Theorem Zmul_add_distr_r : forall n m p : Z, (n + m) * p == n * p + m * p. +Proof NZmul_add_distr_r. + +Theorem Zmul_add_distr_l : forall n m p : Z, n * (m + p) == n * m + n * p. +Proof NZmul_add_distr_l. + +(* A note on naming: right (correspondingly, left) distributivity happens +when the sum is multiplied by a number on the right (left), not when the +sum itself is the right (left) factor in the product (see planetmath.org +and mathworld.wolfram.com). In the old library BinInt, distributivity over +subtraction was named correctly, but distributivity over addition was named +incorrectly. The names in Isabelle/HOL library are also incorrect. *) + +Theorem Zmul_assoc : forall n m p : Z, n * (m * p) == (n * m) * p. +Proof NZmul_assoc. + +Theorem Zmul_1_l : forall n : Z, 1 * n == n. +Proof NZmul_1_l. + +Theorem Zmul_1_r : forall n : Z, n * 1 == n. +Proof NZmul_1_r. + +(* The following two theorems are true in an ordered ring, +but since they don't mention order, we'll put them here *) + +Theorem Zeq_mul_0 : forall n m : Z, n * m == 0 <-> n == 0 \/ m == 0. +Proof NZeq_mul_0. + +Theorem Zneq_mul_0 : forall n m : Z, n ~= 0 /\ m ~= 0 <-> n * m ~= 0. +Proof NZneq_mul_0. + +(* Theorems that are either not valid on N or have different proofs on N and Z *) + +Theorem Zmul_pred_r : forall n m : Z, n * (P m) == n * m - n. +Proof. +intros n m. +rewrite <- (Zsucc_pred m) at 2. +now rewrite Zmul_succ_r, <- Zadd_sub_assoc, Zsub_diag, Zadd_0_r. +Qed. + +Theorem Zmul_pred_l : forall n m : Z, (P n) * m == n * m - m. +Proof. +intros n m; rewrite (Zmul_comm (P n) m), (Zmul_comm n m). apply Zmul_pred_r. +Qed. + +Theorem Zmul_opp_l : forall n m : Z, (- n) * m == - (n * m). +Proof. +intros n m. apply -> Zadd_move_0_r. +now rewrite <- Zmul_add_distr_r, Zadd_opp_diag_l, Zmul_0_l. +Qed. + +Theorem Zmul_opp_r : forall n m : Z, n * (- m) == - (n * m). +Proof. +intros n m; rewrite (Zmul_comm n (- m)), (Zmul_comm n m); apply Zmul_opp_l. +Qed. + +Theorem Zmul_opp_opp : forall n m : Z, (- n) * (- m) == n * m. +Proof. +intros n m; now rewrite Zmul_opp_l, Zmul_opp_r, Zopp_involutive. +Qed. + +Theorem Zmul_sub_distr_l : forall n m p : Z, n * (m - p) == n * m - n * p. +Proof. +intros n m p. do 2 rewrite <- Zadd_opp_r. rewrite Zmul_add_distr_l. +now rewrite Zmul_opp_r. +Qed. + +Theorem Zmul_sub_distr_r : forall n m p : Z, (n - m) * p == n * p - m * p. +Proof. +intros n m p; rewrite (Zmul_comm (n - m) p), (Zmul_comm n p), (Zmul_comm m p); +now apply Zmul_sub_distr_l. +Qed. + +End ZMulPropFunct. + + diff --git a/theories/Numbers/Integer/Abstract/ZMulOrder.v b/theories/Numbers/Integer/Abstract/ZMulOrder.v new file mode 100644 index 00000000..e3f1d9aa --- /dev/null +++ b/theories/Numbers/Integer/Abstract/ZMulOrder.v @@ -0,0 +1,343 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZMulOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export ZAddOrder. + +Module ZMulOrderPropFunct (Import ZAxiomsMod : ZAxiomsSig). +Module Export ZAddOrderPropMod := ZAddOrderPropFunct ZAxiomsMod. +Open Local Scope IntScope. + +Theorem Zmul_lt_pred : + forall p q n m : Z, S p == q -> (p * n < p * m <-> q * n + m < q * m + n). +Proof NZmul_lt_pred. + +Theorem Zmul_lt_mono_pos_l : forall p n m : Z, 0 < p -> (n < m <-> p * n < p * m). +Proof NZmul_lt_mono_pos_l. + +Theorem Zmul_lt_mono_pos_r : forall p n m : Z, 0 < p -> (n < m <-> n * p < m * p). +Proof NZmul_lt_mono_pos_r. + +Theorem Zmul_lt_mono_neg_l : forall p n m : Z, p < 0 -> (n < m <-> p * m < p * n). +Proof NZmul_lt_mono_neg_l. + +Theorem Zmul_lt_mono_neg_r : forall p n m : Z, p < 0 -> (n < m <-> m * p < n * p). +Proof NZmul_lt_mono_neg_r. + +Theorem Zmul_le_mono_nonneg_l : forall n m p : Z, 0 <= p -> n <= m -> p * n <= p * m. +Proof NZmul_le_mono_nonneg_l. + +Theorem Zmul_le_mono_nonpos_l : forall n m p : Z, p <= 0 -> n <= m -> p * m <= p * n. +Proof NZmul_le_mono_nonpos_l. + +Theorem Zmul_le_mono_nonneg_r : forall n m p : Z, 0 <= p -> n <= m -> n * p <= m * p. +Proof NZmul_le_mono_nonneg_r. + +Theorem Zmul_le_mono_nonpos_r : forall n m p : Z, p <= 0 -> n <= m -> m * p <= n * p. +Proof NZmul_le_mono_nonpos_r. + +Theorem Zmul_cancel_l : forall n m p : Z, p ~= 0 -> (p * n == p * m <-> n == m). +Proof NZmul_cancel_l. + +Theorem Zmul_cancel_r : forall n m p : Z, p ~= 0 -> (n * p == m * p <-> n == m). +Proof NZmul_cancel_r. + +Theorem Zmul_id_l : forall n m : Z, m ~= 0 -> (n * m == m <-> n == 1). +Proof NZmul_id_l. + +Theorem Zmul_id_r : forall n m : Z, n ~= 0 -> (n * m == n <-> m == 1). +Proof NZmul_id_r. + +Theorem Zmul_le_mono_pos_l : forall n m p : Z, 0 < p -> (n <= m <-> p * n <= p * m). +Proof NZmul_le_mono_pos_l. + +Theorem Zmul_le_mono_pos_r : forall n m p : Z, 0 < p -> (n <= m <-> n * p <= m * p). +Proof NZmul_le_mono_pos_r. + +Theorem Zmul_le_mono_neg_l : forall n m p : Z, p < 0 -> (n <= m <-> p * m <= p * n). +Proof NZmul_le_mono_neg_l. + +Theorem Zmul_le_mono_neg_r : forall n m p : Z, p < 0 -> (n <= m <-> m * p <= n * p). +Proof NZmul_le_mono_neg_r. + +Theorem Zmul_lt_mono_nonneg : + forall n m p q : Z, 0 <= n -> n < m -> 0 <= p -> p < q -> n * p < m * q. +Proof NZmul_lt_mono_nonneg. + +Theorem Zmul_lt_mono_nonpos : + forall n m p q : Z, m <= 0 -> n < m -> q <= 0 -> p < q -> m * q < n * p. +Proof. +intros n m p q H1 H2 H3 H4. +apply Zle_lt_trans with (m * p). +apply Zmul_le_mono_nonpos_l; [assumption | now apply Zlt_le_incl]. +apply -> Zmul_lt_mono_neg_r; [assumption | now apply Zlt_le_trans with q]. +Qed. + +Theorem Zmul_le_mono_nonneg : + forall n m p q : Z, 0 <= n -> n <= m -> 0 <= p -> p <= q -> n * p <= m * q. +Proof NZmul_le_mono_nonneg. + +Theorem Zmul_le_mono_nonpos : + forall n m p q : Z, m <= 0 -> n <= m -> q <= 0 -> p <= q -> m * q <= n * p. +Proof. +intros n m p q H1 H2 H3 H4. +apply Zle_trans with (m * p). +now apply Zmul_le_mono_nonpos_l. +apply Zmul_le_mono_nonpos_r; [now apply Zle_trans with q | assumption]. +Qed. + +Theorem Zmul_pos_pos : forall n m : Z, 0 < n -> 0 < m -> 0 < n * m. +Proof NZmul_pos_pos. + +Theorem Zmul_neg_neg : forall n m : Z, n < 0 -> m < 0 -> 0 < n * m. +Proof NZmul_neg_neg. + +Theorem Zmul_pos_neg : forall n m : Z, 0 < n -> m < 0 -> n * m < 0. +Proof NZmul_pos_neg. + +Theorem Zmul_neg_pos : forall n m : Z, n < 0 -> 0 < m -> n * m < 0. +Proof NZmul_neg_pos. + +Theorem Zmul_nonneg_nonneg : forall n m : Z, 0 <= n -> 0 <= m -> 0 <= n * m. +Proof. +intros n m H1 H2. +rewrite <- (Zmul_0_l m). now apply Zmul_le_mono_nonneg_r. +Qed. + +Theorem Zmul_nonpos_nonpos : forall n m : Z, n <= 0 -> m <= 0 -> 0 <= n * m. +Proof. +intros n m H1 H2. +rewrite <- (Zmul_0_l m). now apply Zmul_le_mono_nonpos_r. +Qed. + +Theorem Zmul_nonneg_nonpos : forall n m : Z, 0 <= n -> m <= 0 -> n * m <= 0. +Proof. +intros n m H1 H2. +rewrite <- (Zmul_0_l m). now apply Zmul_le_mono_nonpos_r. +Qed. + +Theorem Zmul_nonpos_nonneg : forall n m : Z, n <= 0 -> 0 <= m -> n * m <= 0. +Proof. +intros; rewrite Zmul_comm; now apply Zmul_nonneg_nonpos. +Qed. + +Theorem Zlt_1_mul_pos : forall n m : Z, 1 < n -> 0 < m -> 1 < n * m. +Proof NZlt_1_mul_pos. + +Theorem Zeq_mul_0 : forall n m : Z, n * m == 0 <-> n == 0 \/ m == 0. +Proof NZeq_mul_0. + +Theorem Zneq_mul_0 : forall n m : Z, n ~= 0 /\ m ~= 0 <-> n * m ~= 0. +Proof NZneq_mul_0. + +Theorem Zeq_square_0 : forall n : Z, n * n == 0 <-> n == 0. +Proof NZeq_square_0. + +Theorem Zeq_mul_0_l : forall n m : Z, n * m == 0 -> m ~= 0 -> n == 0. +Proof NZeq_mul_0_l. + +Theorem Zeq_mul_0_r : forall n m : Z, n * m == 0 -> n ~= 0 -> m == 0. +Proof NZeq_mul_0_r. + +Theorem Zlt_0_mul : forall n m : Z, 0 < n * m <-> 0 < n /\ 0 < m \/ m < 0 /\ n < 0. +Proof NZlt_0_mul. + +Notation Zmul_pos := Zlt_0_mul (only parsing). + +Theorem Zlt_mul_0 : + forall n m : Z, n * m < 0 <-> n < 0 /\ m > 0 \/ n > 0 /\ m < 0. +Proof. +intros n m; split; [intro H | intros [[H1 H2] | [H1 H2]]]. +destruct (Zlt_trichotomy n 0) as [H1 | [H1 | H1]]; +[| rewrite H1 in H; rewrite Zmul_0_l in H; false_hyp H Zlt_irrefl |]; +(destruct (Zlt_trichotomy m 0) as [H2 | [H2 | H2]]; +[| rewrite H2 in H; rewrite Zmul_0_r in H; false_hyp H Zlt_irrefl |]); +try (left; now split); try (right; now split). +assert (H3 : n * m > 0) by now apply Zmul_neg_neg. +elimtype False; now apply (Zlt_asymm (n * m) 0). +assert (H3 : n * m > 0) by now apply Zmul_pos_pos. +elimtype False; now apply (Zlt_asymm (n * m) 0). +now apply Zmul_neg_pos. now apply Zmul_pos_neg. +Qed. + +Notation Zmul_neg := Zlt_mul_0 (only parsing). + +Theorem Zle_0_mul : + forall n m : Z, 0 <= n * m -> 0 <= n /\ 0 <= m \/ n <= 0 /\ m <= 0. +Proof. +assert (R : forall n : Z, 0 == n <-> n == 0) by (intros; split; apply Zeq_symm). +intros n m. repeat rewrite Zlt_eq_cases. repeat rewrite R. +rewrite Zlt_0_mul, Zeq_mul_0. +pose proof (Zlt_trichotomy n 0); pose proof (Zlt_trichotomy m 0). tauto. +Qed. + +Notation Zmul_nonneg := Zle_0_mul (only parsing). + +Theorem Zle_mul_0 : + forall n m : Z, n * m <= 0 -> 0 <= n /\ m <= 0 \/ n <= 0 /\ 0 <= m. +Proof. +assert (R : forall n : Z, 0 == n <-> n == 0) by (intros; split; apply Zeq_symm). +intros n m. repeat rewrite Zlt_eq_cases. repeat rewrite R. +rewrite Zlt_mul_0, Zeq_mul_0. +pose proof (Zlt_trichotomy n 0); pose proof (Zlt_trichotomy m 0). tauto. +Qed. + +Notation Zmul_nonpos := Zle_mul_0 (only parsing). + +Theorem Zle_0_square : forall n : Z, 0 <= n * n. +Proof. +intro n; destruct (Zneg_nonneg_cases n). +apply Zlt_le_incl; now apply Zmul_neg_neg. +now apply Zmul_nonneg_nonneg. +Qed. + +Notation Zsquare_nonneg := Zle_0_square (only parsing). + +Theorem Znlt_square_0 : forall n : Z, ~ n * n < 0. +Proof. +intros n H. apply -> Zlt_nge in H. apply H. apply Zsquare_nonneg. +Qed. + +Theorem Zsquare_lt_mono_nonneg : forall n m : Z, 0 <= n -> n < m -> n * n < m * m. +Proof NZsquare_lt_mono_nonneg. + +Theorem Zsquare_lt_mono_nonpos : forall n m : Z, n <= 0 -> m < n -> n * n < m * m. +Proof. +intros n m H1 H2. now apply Zmul_lt_mono_nonpos. +Qed. + +Theorem Zsquare_le_mono_nonneg : forall n m : Z, 0 <= n -> n <= m -> n * n <= m * m. +Proof NZsquare_le_mono_nonneg. + +Theorem Zsquare_le_mono_nonpos : forall n m : Z, n <= 0 -> m <= n -> n * n <= m * m. +Proof. +intros n m H1 H2. now apply Zmul_le_mono_nonpos. +Qed. + +Theorem Zsquare_lt_simpl_nonneg : forall n m : Z, 0 <= m -> n * n < m * m -> n < m. +Proof NZsquare_lt_simpl_nonneg. + +Theorem Zsquare_le_simpl_nonneg : forall n m : Z, 0 <= m -> n * n <= m * m -> n <= m. +Proof NZsquare_le_simpl_nonneg. + +Theorem Zsquare_lt_simpl_nonpos : forall n m : Z, m <= 0 -> n * n < m * m -> m < n. +Proof. +intros n m H1 H2. destruct (Zle_gt_cases n 0). +destruct (NZlt_ge_cases m n). +assumption. assert (F : m * m <= n * n) by now apply Zsquare_le_mono_nonpos. +apply -> NZle_ngt in F. false_hyp H2 F. +now apply Zle_lt_trans with 0. +Qed. + +Theorem Zsquare_le_simpl_nonpos : forall n m : NZ, m <= 0 -> n * n <= m * m -> m <= n. +Proof. +intros n m H1 H2. destruct (NZle_gt_cases n 0). +destruct (NZle_gt_cases m n). +assumption. assert (F : m * m < n * n) by now apply Zsquare_lt_mono_nonpos. +apply -> NZlt_nge in F. false_hyp H2 F. +apply Zlt_le_incl; now apply NZle_lt_trans with 0. +Qed. + +Theorem Zmul_2_mono_l : forall n m : Z, n < m -> 1 + (1 + 1) * n < (1 + 1) * m. +Proof NZmul_2_mono_l. + +Theorem Zlt_1_mul_neg : forall n m : Z, n < -1 -> m < 0 -> 1 < n * m. +Proof. +intros n m H1 H2. apply -> (NZmul_lt_mono_neg_r m) in H1. +apply <- Zopp_pos_neg in H2. rewrite Zmul_opp_l, Zmul_1_l in H1. +now apply Zlt_1_l with (- m). +assumption. +Qed. + +Theorem Zlt_mul_n1_neg : forall n m : Z, 1 < n -> m < 0 -> n * m < -1. +Proof. +intros n m H1 H2. apply -> (NZmul_lt_mono_neg_r m) in H1. +rewrite Zmul_1_l in H1. now apply Zlt_n1_r with m. +assumption. +Qed. + +Theorem Zlt_mul_n1_pos : forall n m : Z, n < -1 -> 0 < m -> n * m < -1. +Proof. +intros n m H1 H2. apply -> (NZmul_lt_mono_pos_r m) in H1. +rewrite Zmul_opp_l, Zmul_1_l in H1. +apply <- Zopp_neg_pos in H2. now apply Zlt_n1_r with (- m). +assumption. +Qed. + +Theorem Zlt_1_mul_l : forall n m : Z, 1 < n -> n * m < -1 \/ n * m == 0 \/ 1 < n * m. +Proof. +intros n m H; destruct (Zlt_trichotomy m 0) as [H1 | [H1 | H1]]. +left. now apply Zlt_mul_n1_neg. +right; left; now rewrite H1, Zmul_0_r. +right; right; now apply Zlt_1_mul_pos. +Qed. + +Theorem Zlt_n1_mul_r : forall n m : Z, n < -1 -> n * m < -1 \/ n * m == 0 \/ 1 < n * m. +Proof. +intros n m H; destruct (Zlt_trichotomy m 0) as [H1 | [H1 | H1]]. +right; right. now apply Zlt_1_mul_neg. +right; left; now rewrite H1, Zmul_0_r. +left. now apply Zlt_mul_n1_pos. +Qed. + +Theorem Zeq_mul_1 : forall n m : Z, n * m == 1 -> n == 1 \/ n == -1. +Proof. +assert (F : ~ 1 < -1). +intro H. +assert (H1 : -1 < 0). apply <- Zopp_neg_pos. apply Zlt_succ_diag_r. +assert (H2 : 1 < 0) by now apply Zlt_trans with (-1). false_hyp H2 Znlt_succ_diag_l. +Z0_pos_neg n. +intros m H; rewrite Zmul_0_l in H; false_hyp H Zneq_succ_diag_r. +intros n H; split; apply <- Zle_succ_l in H; le_elim H. +intros m H1; apply (Zlt_1_mul_l n m) in H. +rewrite H1 in H; destruct H as [H | [H | H]]. +false_hyp H F. false_hyp H Zneq_succ_diag_l. false_hyp H Zlt_irrefl. +intros; now left. +intros m H1; apply (Zlt_1_mul_l n m) in H. rewrite Zmul_opp_l in H1; +apply -> Zeq_opp_l in H1. rewrite H1 in H; destruct H as [H | [H | H]]. +false_hyp H Zlt_irrefl. apply -> Zeq_opp_l in H. rewrite Zopp_0 in H. +false_hyp H Zneq_succ_diag_l. false_hyp H F. +intros; right; symmetry; now apply Zopp_wd. +Qed. + +Theorem Zlt_mul_diag_l : forall n m : Z, n < 0 -> (1 < m <-> n * m < n). +Proof. +intros n m H. stepr (n * m < n * 1) by now rewrite Zmul_1_r. +now apply Zmul_lt_mono_neg_l. +Qed. + +Theorem Zlt_mul_diag_r : forall n m : Z, 0 < n -> (1 < m <-> n < n * m). +Proof. +intros n m H. stepr (n * 1 < n * m) by now rewrite Zmul_1_r. +now apply Zmul_lt_mono_pos_l. +Qed. + +Theorem Zle_mul_diag_l : forall n m : Z, n < 0 -> (1 <= m <-> n * m <= n). +Proof. +intros n m H. stepr (n * m <= n * 1) by now rewrite Zmul_1_r. +now apply Zmul_le_mono_neg_l. +Qed. + +Theorem Zle_mul_diag_r : forall n m : Z, 0 < n -> (1 <= m <-> n <= n * m). +Proof. +intros n m H. stepr (n * 1 <= n * m) by now rewrite Zmul_1_r. +now apply Zmul_le_mono_pos_l. +Qed. + +Theorem Zlt_mul_r : forall n m p : Z, 0 < n -> 1 < p -> n < m -> n < m * p. +Proof. +intros. stepl (n * 1) by now rewrite Zmul_1_r. +apply Zmul_lt_mono_nonneg. +now apply Zlt_le_incl. assumption. apply Zle_0_1. assumption. +Qed. + +End ZMulOrderPropFunct. + diff --git a/theories/Numbers/Integer/BigZ/BigZ.v b/theories/Numbers/Integer/BigZ/BigZ.v new file mode 100644 index 00000000..09abf424 --- /dev/null +++ b/theories/Numbers/Integer/BigZ/BigZ.v @@ -0,0 +1,109 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: BigZ.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export BigN. +Require Import ZMulOrder. +Require Import ZSig. +Require Import ZSigZAxioms. +Require Import ZMake. + +Module BigZ <: ZType := ZMake.Make BigN. + +(** Module [BigZ] implements [ZAxiomsSig] *) + +Module Export BigZAxiomsMod := ZSig_ZAxioms BigZ. +Module Export BigZMulOrderPropMod := ZMulOrderPropFunct BigZAxiomsMod. + +(** Notations about [BigZ] *) + +Notation bigZ := BigZ.t. + +Delimit Scope bigZ_scope with bigZ. +Bind Scope bigZ_scope with bigZ. +Bind Scope bigZ_scope with BigZ.t. +Bind Scope bigZ_scope with BigZ.t_. + +Notation Local "0" := BigZ.zero : bigZ_scope. +Infix "+" := BigZ.add : bigZ_scope. +Infix "-" := BigZ.sub : bigZ_scope. +Notation "- x" := (BigZ.opp x) : bigZ_scope. +Infix "*" := BigZ.mul : bigZ_scope. +Infix "/" := BigZ.div : bigZ_scope. +Infix "?=" := BigZ.compare : bigZ_scope. +Infix "==" := BigZ.eq (at level 70, no associativity) : bigZ_scope. +Infix "<" := BigZ.lt : bigZ_scope. +Infix "<=" := BigZ.le : bigZ_scope. +Notation "[ i ]" := (BigZ.to_Z i) : bigZ_scope. + +Open Scope bigZ_scope. + +(** Some additional results about [BigZ] *) + +Theorem spec_to_Z: forall n:bigZ, + BigN.to_Z (BigZ.to_N n) = ((Zsgn [n]) * [n])%Z. +Proof. +intros n; case n; simpl; intros p; + generalize (BigN.spec_pos p); case (BigN.to_Z p); auto. +intros p1 H1; case H1; auto. +intros p1 H1; case H1; auto. +Qed. + +Theorem spec_to_N n: + ([n] = Zsgn [n] * (BigN.to_Z (BigZ.to_N n)))%Z. +Proof. +intros n; case n; simpl; intros p; + generalize (BigN.spec_pos p); case (BigN.to_Z p); auto. +intros p1 H1; case H1; auto. +intros p1 H1; case H1; auto. +Qed. + +Theorem spec_to_Z_pos: forall n, (0 <= [n])%Z -> + BigN.to_Z (BigZ.to_N n) = [n]. +Proof. +intros n; case n; simpl; intros p; + generalize (BigN.spec_pos p); case (BigN.to_Z p); auto. +intros p1 _ H1; case H1; auto. +intros p1 H1; case H1; auto. +Qed. + +Lemma sub_opp : forall x y : bigZ, x - y == x + (- y). +Proof. +red; intros; zsimpl; auto. +Qed. + +Lemma add_opp : forall x : bigZ, x + (- x) == 0. +Proof. +red; intros; zsimpl; auto with zarith. +Qed. + +(** [BigZ] is a ring *) + +Lemma BigZring : + ring_theory BigZ.zero BigZ.one BigZ.add BigZ.mul BigZ.sub BigZ.opp BigZ.eq. +Proof. +constructor. +exact Zadd_0_l. +exact Zadd_comm. +exact Zadd_assoc. +exact Zmul_1_l. +exact Zmul_comm. +exact Zmul_assoc. +exact Zmul_add_distr_r. +exact sub_opp. +exact add_opp. +Qed. + +Add Ring BigZr : BigZring. + +(** Todo: tactic translating from [BigZ] to [Z] + omega *) + +(** Todo: micromega *) diff --git a/theories/Numbers/Integer/BigZ/ZMake.v b/theories/Numbers/Integer/BigZ/ZMake.v new file mode 100644 index 00000000..1f2b12bb --- /dev/null +++ b/theories/Numbers/Integer/BigZ/ZMake.v @@ -0,0 +1,491 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZMake.v 11027 2008-06-01 13:28:59Z letouzey $ i*) + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import NSig. +Require Import ZSig. + +Open Scope Z_scope. + +(** * ZMake + + A generic transformation from a structure of natural numbers + [NSig.NType] to a structure of integers [ZSig.ZType]. +*) + +Module Make (N:NType) <: ZType. + + Inductive t_ := + | Pos : N.t -> t_ + | Neg : N.t -> t_. + + Definition t := t_. + + Definition zero := Pos N.zero. + Definition one := Pos N.one. + Definition minus_one := Neg N.one. + + Definition of_Z x := + match x with + | Zpos x => Pos (N.of_N (Npos x)) + | Z0 => zero + | Zneg x => Neg (N.of_N (Npos x)) + end. + + Definition to_Z x := + match x with + | Pos nx => N.to_Z nx + | Neg nx => Zopp (N.to_Z nx) + end. + + Theorem spec_of_Z: forall x, to_Z (of_Z x) = x. + intros x; case x; unfold to_Z, of_Z, zero. + exact N.spec_0. + intros; rewrite N.spec_of_N; auto. + intros; rewrite N.spec_of_N; auto. + Qed. + + Definition eq x y := (to_Z x = to_Z y). + + Theorem spec_0: to_Z zero = 0. + exact N.spec_0. + Qed. + + Theorem spec_1: to_Z one = 1. + exact N.spec_1. + Qed. + + Theorem spec_m1: to_Z minus_one = -1. + simpl; rewrite N.spec_1; auto. + Qed. + + Definition compare x y := + match x, y with + | Pos nx, Pos ny => N.compare nx ny + | Pos nx, Neg ny => + match N.compare nx N.zero with + | Gt => Gt + | _ => N.compare ny N.zero + end + | Neg nx, Pos ny => + match N.compare N.zero nx with + | Lt => Lt + | _ => N.compare N.zero ny + end + | Neg nx, Neg ny => N.compare ny nx + end. + + Definition lt n m := compare n m = Lt. + Definition le n m := compare n m <> Gt. + Definition min n m := match compare n m with Gt => m | _ => n end. + Definition max n m := match compare n m with Lt => m | _ => n end. + + Theorem spec_compare: forall x y, + match compare x y with + Eq => to_Z x = to_Z y + | Lt => to_Z x < to_Z y + | Gt => to_Z x > to_Z y + end. + unfold compare, to_Z; intros x y; case x; case y; clear x y; + intros x y; auto; generalize (N.spec_pos x) (N.spec_pos y). + generalize (N.spec_compare y x); case N.compare; auto with zarith. + generalize (N.spec_compare y N.zero); case N.compare; + try rewrite N.spec_0; auto with zarith. + generalize (N.spec_compare x N.zero); case N.compare; + rewrite N.spec_0; auto with zarith. + generalize (N.spec_compare x N.zero); case N.compare; + rewrite N.spec_0; auto with zarith. + generalize (N.spec_compare N.zero y); case N.compare; + try rewrite N.spec_0; auto with zarith. + generalize (N.spec_compare N.zero x); case N.compare; + rewrite N.spec_0; auto with zarith. + generalize (N.spec_compare N.zero x); case N.compare; + rewrite N.spec_0; auto with zarith. + generalize (N.spec_compare x y); case N.compare; auto with zarith. + Qed. + + Definition eq_bool x y := + match compare x y with + | Eq => true + | _ => false + end. + + Theorem spec_eq_bool: forall x y, + if eq_bool x y then to_Z x = to_Z y else to_Z x <> to_Z y. + intros x y; unfold eq_bool; + generalize (spec_compare x y); case compare; auto with zarith. + Qed. + + Definition cmp_sign x y := + match x, y with + | Pos nx, Neg ny => + if N.eq_bool ny N.zero then Eq else Gt + | Neg nx, Pos ny => + if N.eq_bool nx N.zero then Eq else Lt + | _, _ => Eq + end. + + Theorem spec_cmp_sign: forall x y, + match cmp_sign x y with + | Gt => 0 <= to_Z x /\ to_Z y < 0 + | Lt => to_Z x < 0 /\ 0 <= to_Z y + | Eq => True + end. + Proof. + intros [x | x] [y | y]; unfold cmp_sign; auto. + generalize (N.spec_eq_bool y N.zero); case N.eq_bool; auto. + rewrite N.spec_0; unfold to_Z. + generalize (N.spec_pos x) (N.spec_pos y); auto with zarith. + generalize (N.spec_eq_bool x N.zero); case N.eq_bool; auto. + rewrite N.spec_0; unfold to_Z. + generalize (N.spec_pos x) (N.spec_pos y); auto with zarith. + Qed. + + Definition to_N x := + match x with + | Pos nx => nx + | Neg nx => nx + end. + + Definition abs x := Pos (to_N x). + + Theorem spec_abs: forall x, to_Z (abs x) = Zabs (to_Z x). + intros x; case x; clear x; intros x; assert (F:=N.spec_pos x). + simpl; rewrite Zabs_eq; auto. + simpl; rewrite Zabs_non_eq; simpl; auto with zarith. + Qed. + + Definition opp x := + match x with + | Pos nx => Neg nx + | Neg nx => Pos nx + end. + + Theorem spec_opp: forall x, to_Z (opp x) = - to_Z x. + intros x; case x; simpl; auto with zarith. + Qed. + + Definition succ x := + match x with + | Pos n => Pos (N.succ n) + | Neg n => + match N.compare N.zero n with + | Lt => Neg (N.pred n) + | _ => one + end + end. + + Theorem spec_succ: forall n, to_Z (succ n) = to_Z n + 1. + intros x; case x; clear x; intros x. + exact (N.spec_succ x). + simpl; generalize (N.spec_compare N.zero x); case N.compare; + rewrite N.spec_0; simpl. + intros HH; rewrite <- HH; rewrite N.spec_1; ring. + intros HH; rewrite N.spec_pred; auto with zarith. + generalize (N.spec_pos x); auto with zarith. + Qed. + + Definition add x y := + match x, y with + | Pos nx, Pos ny => Pos (N.add nx ny) + | Pos nx, Neg ny => + match N.compare nx ny with + | Gt => Pos (N.sub nx ny) + | Eq => zero + | Lt => Neg (N.sub ny nx) + end + | Neg nx, Pos ny => + match N.compare nx ny with + | Gt => Neg (N.sub nx ny) + | Eq => zero + | Lt => Pos (N.sub ny nx) + end + | Neg nx, Neg ny => Neg (N.add nx ny) + end. + + Theorem spec_add: forall x y, to_Z (add x y) = to_Z x + to_Z y. + unfold add, to_Z; intros [x | x] [y | y]. + exact (N.spec_add x y). + unfold zero; generalize (N.spec_compare x y); case N.compare. + rewrite N.spec_0; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + unfold zero; generalize (N.spec_compare x y); case N.compare. + rewrite N.spec_0; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + intros; rewrite N.spec_add; try ring; auto with zarith. + Qed. + + Definition pred x := + match x with + | Pos nx => + match N.compare N.zero nx with + | Lt => Pos (N.pred nx) + | _ => minus_one + end + | Neg nx => Neg (N.succ nx) + end. + + Theorem spec_pred: forall x, to_Z (pred x) = to_Z x - 1. + unfold pred, to_Z, minus_one; intros [x | x]. + generalize (N.spec_compare N.zero x); case N.compare; + rewrite N.spec_0; try rewrite N.spec_1; auto with zarith. + intros H; exact (N.spec_pred _ H). + generalize (N.spec_pos x); auto with zarith. + rewrite N.spec_succ; ring. + Qed. + + Definition sub x y := + match x, y with + | Pos nx, Pos ny => + match N.compare nx ny with + | Gt => Pos (N.sub nx ny) + | Eq => zero + | Lt => Neg (N.sub ny nx) + end + | Pos nx, Neg ny => Pos (N.add nx ny) + | Neg nx, Pos ny => Neg (N.add nx ny) + | Neg nx, Neg ny => + match N.compare nx ny with + | Gt => Neg (N.sub nx ny) + | Eq => zero + | Lt => Pos (N.sub ny nx) + end + end. + + Theorem spec_sub: forall x y, to_Z (sub x y) = to_Z x - to_Z y. + unfold sub, to_Z; intros [x | x] [y | y]. + unfold zero; generalize (N.spec_compare x y); case N.compare. + rewrite N.spec_0; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + rewrite N.spec_add; ring. + rewrite N.spec_add; ring. + unfold zero; generalize (N.spec_compare x y); case N.compare. + rewrite N.spec_0; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + intros; rewrite N.spec_sub; try ring; auto with zarith. + Qed. + + Definition mul x y := + match x, y with + | Pos nx, Pos ny => Pos (N.mul nx ny) + | Pos nx, Neg ny => Neg (N.mul nx ny) + | Neg nx, Pos ny => Neg (N.mul nx ny) + | Neg nx, Neg ny => Pos (N.mul nx ny) + end. + + + Theorem spec_mul: forall x y, to_Z (mul x y) = to_Z x * to_Z y. + unfold mul, to_Z; intros [x | x] [y | y]; rewrite N.spec_mul; ring. + Qed. + + Definition square x := + match x with + | Pos nx => Pos (N.square nx) + | Neg nx => Pos (N.square nx) + end. + + Theorem spec_square: forall x, to_Z (square x) = to_Z x * to_Z x. + unfold square, to_Z; intros [x | x]; rewrite N.spec_square; ring. + Qed. + + Definition power_pos x p := + match x with + | Pos nx => Pos (N.power_pos nx p) + | Neg nx => + match p with + | xH => x + | xO _ => Pos (N.power_pos nx p) + | xI _ => Neg (N.power_pos nx p) + end + end. + + Theorem spec_power_pos: forall x n, to_Z (power_pos x n) = to_Z x ^ Zpos n. + assert (F0: forall x, (-x)^2 = x^2). + intros x; rewrite Zpower_2; ring. + unfold power_pos, to_Z; intros [x | x] [p | p |]; + try rewrite N.spec_power_pos; try ring. + assert (F: 0 <= 2 * Zpos p). + assert (0 <= Zpos p); auto with zarith. + rewrite Zpos_xI; repeat rewrite Zpower_exp; auto with zarith. + repeat rewrite Zpower_mult; auto with zarith. + rewrite F0; ring. + assert (F: 0 <= 2 * Zpos p). + assert (0 <= Zpos p); auto with zarith. + rewrite Zpos_xO; repeat rewrite Zpower_exp; auto with zarith. + repeat rewrite Zpower_mult; auto with zarith. + rewrite F0; ring. + Qed. + + Definition sqrt x := + match x with + | Pos nx => Pos (N.sqrt nx) + | Neg nx => Neg N.zero + end. + + + Theorem spec_sqrt: forall x, 0 <= to_Z x -> + to_Z (sqrt x) ^ 2 <= to_Z x < (to_Z (sqrt x) + 1) ^ 2. + unfold to_Z, sqrt; intros [x | x] H. + exact (N.spec_sqrt x). + replace (N.to_Z x) with 0. + rewrite N.spec_0; simpl Zpower; unfold Zpower_pos, iter_pos; + auto with zarith. + generalize (N.spec_pos x); auto with zarith. + Qed. + + Definition div_eucl x y := + match x, y with + | Pos nx, Pos ny => + let (q, r) := N.div_eucl nx ny in + (Pos q, Pos r) + | Pos nx, Neg ny => + let (q, r) := N.div_eucl nx ny in + match N.compare N.zero r with + | Eq => (Neg q, zero) + | _ => (Neg (N.succ q), Neg (N.sub ny r)) + end + | Neg nx, Pos ny => + let (q, r) := N.div_eucl nx ny in + match N.compare N.zero r with + | Eq => (Neg q, zero) + | _ => (Neg (N.succ q), Pos (N.sub ny r)) + end + | Neg nx, Neg ny => + let (q, r) := N.div_eucl nx ny in + (Pos q, Neg r) + end. + + + Theorem spec_div_eucl: forall x y, + to_Z y <> 0 -> + let (q,r) := div_eucl x y in + (to_Z q, to_Z r) = Zdiv_eucl (to_Z x) (to_Z y). + unfold div_eucl, to_Z; intros [x | x] [y | y] H. + assert (H1: 0 < N.to_Z y). + generalize (N.spec_pos y); auto with zarith. + generalize (N.spec_div_eucl x y H1); case N.div_eucl; auto. + assert (HH: 0 < N.to_Z y). + generalize (N.spec_pos y); auto with zarith. + generalize (N.spec_div_eucl x y HH); case N.div_eucl; auto. + intros q r; generalize (N.spec_pos x) HH; unfold Zdiv_eucl; + case_eq (N.to_Z x); case_eq (N.to_Z y); + try (intros; apply False_ind; auto with zarith; fail). + intros p He1 He2 _ _ H1; injection H1; intros H2 H3. + generalize (N.spec_compare N.zero r); case N.compare; + unfold zero; rewrite N.spec_0; try rewrite H3; auto. + rewrite H2; intros; apply False_ind; auto with zarith. + rewrite H2; intros; apply False_ind; auto with zarith. + intros p _ _ _ H1; discriminate H1. + intros p He p1 He1 H1 _. + generalize (N.spec_compare N.zero r); case N.compare. + change (- Zpos p) with (Zneg p). + unfold zero; lazy zeta. + rewrite N.spec_0; intros H2; rewrite <- H2. + intros H3; rewrite <- H3; auto. + rewrite N.spec_0; intros H2. + change (- Zpos p) with (Zneg p); lazy iota beta. + intros H3; rewrite <- H3; auto. + rewrite N.spec_succ; rewrite N.spec_sub. + generalize H2; case (N.to_Z r). + intros; apply False_ind; auto with zarith. + intros p2 _; rewrite He; auto with zarith. + change (Zneg p) with (- (Zpos p)); apply f_equal2 with (f := @pair Z Z); ring. + intros p2 H4; discriminate H4. + assert (N.to_Z r = (Zpos p1 mod (Zpos p))). + unfold Zmod, Zdiv_eucl; rewrite <- H3; auto. + case (Z_mod_lt (Zpos p1) (Zpos p)); auto with zarith. + rewrite N.spec_0; intros H2; generalize (N.spec_pos r); + intros; apply False_ind; auto with zarith. + assert (HH: 0 < N.to_Z y). + generalize (N.spec_pos y); auto with zarith. + generalize (N.spec_div_eucl x y HH); case N.div_eucl; auto. + intros q r; generalize (N.spec_pos x) HH; unfold Zdiv_eucl; + case_eq (N.to_Z x); case_eq (N.to_Z y); + try (intros; apply False_ind; auto with zarith; fail). + intros p He1 He2 _ _ H1; injection H1; intros H2 H3. + generalize (N.spec_compare N.zero r); case N.compare; + unfold zero; rewrite N.spec_0; try rewrite H3; auto. + rewrite H2; intros; apply False_ind; auto with zarith. + rewrite H2; intros; apply False_ind; auto with zarith. + intros p _ _ _ H1; discriminate H1. + intros p He p1 He1 H1 _. + generalize (N.spec_compare N.zero r); case N.compare. + change (- Zpos p1) with (Zneg p1). + unfold zero; lazy zeta. + rewrite N.spec_0; intros H2; rewrite <- H2. + intros H3; rewrite <- H3; auto. + rewrite N.spec_0; intros H2. + change (- Zpos p1) with (Zneg p1); lazy iota beta. + intros H3; rewrite <- H3; auto. + rewrite N.spec_succ; rewrite N.spec_sub. + generalize H2; case (N.to_Z r). + intros; apply False_ind; auto with zarith. + intros p2 _; rewrite He; auto with zarith. + intros p2 H4; discriminate H4. + assert (N.to_Z r = (Zpos p1 mod (Zpos p))). + unfold Zmod, Zdiv_eucl; rewrite <- H3; auto. + case (Z_mod_lt (Zpos p1) (Zpos p)); auto with zarith. + rewrite N.spec_0; generalize (N.spec_pos r); intros; apply False_ind; auto with zarith. + assert (H1: 0 < N.to_Z y). + generalize (N.spec_pos y); auto with zarith. + generalize (N.spec_div_eucl x y H1); case N.div_eucl; auto. + intros q r; generalize (N.spec_pos x) H1; unfold Zdiv_eucl; + case_eq (N.to_Z x); case_eq (N.to_Z y); + try (intros; apply False_ind; auto with zarith; fail). + change (-0) with 0; lazy iota beta; auto. + intros p _ _ _ _ H2; injection H2. + intros H3 H4; rewrite H3; rewrite H4; auto. + intros p _ _ _ H2; discriminate H2. + intros p He p1 He1 _ _ H2. + change (- Zpos p1) with (Zneg p1); lazy iota beta. + change (- Zpos p) with (Zneg p); lazy iota beta. + rewrite <- H2; auto. + Qed. + + Definition div x y := fst (div_eucl x y). + + Definition spec_div: forall x y, + to_Z y <> 0 -> to_Z (div x y) = to_Z x / to_Z y. + intros x y H1; generalize (spec_div_eucl x y H1); unfold div, Zdiv. + case div_eucl; case Zdiv_eucl; simpl; auto. + intros q r q11 r1 H; injection H; auto. + Qed. + + Definition modulo x y := snd (div_eucl x y). + + Theorem spec_modulo: + forall x y, to_Z y <> 0 -> to_Z (modulo x y) = to_Z x mod to_Z y. + intros x y H1; generalize (spec_div_eucl x y H1); unfold modulo, Zmod. + case div_eucl; case Zdiv_eucl; simpl; auto. + intros q r q11 r1 H; injection H; auto. + Qed. + + Definition gcd x y := + match x, y with + | Pos nx, Pos ny => Pos (N.gcd nx ny) + | Pos nx, Neg ny => Pos (N.gcd nx ny) + | Neg nx, Pos ny => Pos (N.gcd nx ny) + | Neg nx, Neg ny => Pos (N.gcd nx ny) + end. + + Theorem spec_gcd: forall a b, to_Z (gcd a b) = Zgcd (to_Z a) (to_Z b). + unfold gcd, Zgcd, to_Z; intros [x | x] [y | y]; rewrite N.spec_gcd; unfold Zgcd; + auto; case N.to_Z; simpl; auto with zarith; + try rewrite Zabs_Zopp; auto; + case N.to_Z; simpl; auto with zarith. + Qed. + +End Make. diff --git a/theories/Numbers/Integer/Binary/ZBinary.v b/theories/Numbers/Integer/Binary/ZBinary.v new file mode 100644 index 00000000..66d2a96a --- /dev/null +++ b/theories/Numbers/Integer/Binary/ZBinary.v @@ -0,0 +1,249 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZBinary.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import ZMulOrder. +Require Import ZArith. + +Open Local Scope Z_scope. + +Module ZBinAxiomsMod <: ZAxiomsSig. +Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig. +Module Export NZAxiomsMod <: NZAxiomsSig. + +Definition NZ := Z. +Definition NZeq := (@eq Z). +Definition NZ0 := 0. +Definition NZsucc := Zsucc'. +Definition NZpred := Zpred'. +Definition NZadd := Zplus. +Definition NZsub := Zminus. +Definition NZmul := Zmult. + +Theorem NZeq_equiv : equiv Z NZeq. +Proof. +exact (@eq_equiv Z). +Qed. + +Add Relation Z NZeq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) +as NZeq_rel. + +Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd. +Proof. +congruence. +Qed. + +Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd. +Proof. +congruence. +Qed. + +Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd. +Proof. +congruence. +Qed. + +Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd. +Proof. +congruence. +Qed. + +Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd. +Proof. +congruence. +Qed. + +Theorem NZpred_succ : forall n : Z, NZpred (NZsucc n) = n. +Proof. +exact Zpred'_succ'. +Qed. + +Theorem NZinduction : + forall A : Z -> Prop, predicate_wd NZeq A -> + A 0 -> (forall n : Z, A n <-> A (NZsucc n)) -> forall n : Z, A n. +Proof. +intros A A_wd A0 AS n; apply Zind; clear n. +assumption. +intros; now apply -> AS. +intros n H. rewrite <- (Zsucc'_pred' n) in H. now apply <- AS. +Qed. + +Theorem NZadd_0_l : forall n : Z, 0 + n = n. +Proof. +exact Zplus_0_l. +Qed. + +Theorem NZadd_succ_l : forall n m : Z, (NZsucc n) + m = NZsucc (n + m). +Proof. +intros; do 2 rewrite <- Zsucc_succ'; apply Zplus_succ_l. +Qed. + +Theorem NZsub_0_r : forall n : Z, n - 0 = n. +Proof. +exact Zminus_0_r. +Qed. + +Theorem NZsub_succ_r : forall n m : Z, n - (NZsucc m) = NZpred (n - m). +Proof. +intros; rewrite <- Zsucc_succ'; rewrite <- Zpred_pred'; +apply Zminus_succ_r. +Qed. + +Theorem NZmul_0_l : forall n : Z, 0 * n = 0. +Proof. +reflexivity. +Qed. + +Theorem NZmul_succ_l : forall n m : Z, (NZsucc n) * m = n * m + m. +Proof. +intros; rewrite <- Zsucc_succ'; apply Zmult_succ_l. +Qed. + +End NZAxiomsMod. + +Definition NZlt := Zlt. +Definition NZle := Zle. +Definition NZmin := Zmin. +Definition NZmax := Zmax. + +Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd. +Proof. +unfold NZeq. intros n1 n2 H1 m1 m2 H2; rewrite H1; now rewrite H2. +Qed. + +Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd. +Proof. +unfold NZeq. intros n1 n2 H1 m1 m2 H2; rewrite H1; now rewrite H2. +Qed. + +Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd. +Proof. +congruence. +Qed. + +Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd. +Proof. +congruence. +Qed. + +Theorem NZlt_eq_cases : forall n m : Z, n <= m <-> n < m \/ n = m. +Proof. +intros n m; split. apply Zle_lt_or_eq. +intro H; destruct H as [H | H]. now apply Zlt_le_weak. rewrite H; apply Zle_refl. +Qed. + +Theorem NZlt_irrefl : forall n : Z, ~ n < n. +Proof. +exact Zlt_irrefl. +Qed. + +Theorem NZlt_succ_r : forall n m : Z, n < (NZsucc m) <-> n <= m. +Proof. +intros; unfold NZsucc; rewrite <- Zsucc_succ'; split; +[apply Zlt_succ_le | apply Zle_lt_succ]. +Qed. + +Theorem NZmin_l : forall n m : NZ, n <= m -> NZmin n m = n. +Proof. +unfold NZmin, Zmin, Zle; intros n m H. +destruct (n ?= m); try reflexivity. now elim H. +Qed. + +Theorem NZmin_r : forall n m : NZ, m <= n -> NZmin n m = m. +Proof. +unfold NZmin, Zmin, Zle; intros n m H. +case_eq (n ?= m); intro H1; try reflexivity. +now apply Zcompare_Eq_eq. +apply <- Zcompare_Gt_Lt_antisym in H1. now elim H. +Qed. + +Theorem NZmax_l : forall n m : NZ, m <= n -> NZmax n m = n. +Proof. +unfold NZmax, Zmax, Zle; intros n m H. +case_eq (n ?= m); intro H1; try reflexivity. +apply <- Zcompare_Gt_Lt_antisym in H1. now elim H. +Qed. + +Theorem NZmax_r : forall n m : NZ, n <= m -> NZmax n m = m. +Proof. +unfold NZmax, Zmax, Zle; intros n m H. +case_eq (n ?= m); intro H1. +now apply Zcompare_Eq_eq. reflexivity. now elim H. +Qed. + +End NZOrdAxiomsMod. + +Definition Zopp (x : Z) := +match x with +| Z0 => Z0 +| Zpos x => Zneg x +| Zneg x => Zpos x +end. + +Add Morphism Zopp with signature NZeq ==> NZeq as Zopp_wd. +Proof. +congruence. +Qed. + +Theorem Zsucc_pred : forall n : Z, NZsucc (NZpred n) = n. +Proof. +exact Zsucc'_pred'. +Qed. + +Theorem Zopp_0 : - 0 = 0. +Proof. +reflexivity. +Qed. + +Theorem Zopp_succ : forall n : Z, - (NZsucc n) = NZpred (- n). +Proof. +intro; rewrite <- Zsucc_succ'; rewrite <- Zpred_pred'. apply Zopp_succ. +Qed. + +End ZBinAxiomsMod. + +Module Export ZBinMulOrderPropMod := ZMulOrderPropFunct ZBinAxiomsMod. + +(** Z forms a ring *) + +(*Lemma Zring : ring_theory 0 1 NZadd NZmul NZsub Zopp NZeq. +Proof. +constructor. +exact Zadd_0_l. +exact Zadd_comm. +exact Zadd_assoc. +exact Zmul_1_l. +exact Zmul_comm. +exact Zmul_assoc. +exact Zmul_add_distr_r. +intros; now rewrite Zadd_opp_minus. +exact Zadd_opp_r. +Qed. + +Add Ring ZR : Zring.*) + + + +(* +Theorem eq_equiv_e : forall x y : Z, E x y <-> e x y. +Proof. +intros x y; unfold E, e, Zeq_bool; split; intro H. +rewrite H; now rewrite Zcompare_refl. +rewrite eq_true_unfold_pos in H. +assert (H1 : (x ?= y) = Eq). +case_eq (x ?= y); intro H1; rewrite H1 in H; simpl in H; +[reflexivity | discriminate H | discriminate H]. +now apply Zcompare_Eq_eq. +Qed. +*) diff --git a/theories/Numbers/Integer/NatPairs/ZNatPairs.v b/theories/Numbers/Integer/NatPairs/ZNatPairs.v new file mode 100644 index 00000000..8b3d815d --- /dev/null +++ b/theories/Numbers/Integer/NatPairs/ZNatPairs.v @@ -0,0 +1,422 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZNatPairs.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import NSub. (* The most complete file for natural numbers *) +Require Export ZMulOrder. (* The most complete file for integers *) +Require Export Ring. + +Module ZPairsAxiomsMod (Import NAxiomsMod : NAxiomsSig) <: ZAxiomsSig. +Module Import NPropMod := NSubPropFunct NAxiomsMod. (* Get all properties of natural numbers *) + +(* We do not declare ring in Natural/Abstract for two reasons. First, some +of the properties proved in NAdd and NMul are used in the new BinNat, +and it is in turn used in Ring. Using ring in Natural/Abstract would be +circular. It is possible, however, not to make BinNat dependent on +Numbers/Natural and prove the properties necessary for ring from scratch +(this is, of course, how it used to be). In addition, if we define semiring +structures in the implementation subdirectories of Natural, we are able to +specify binary natural numbers as the type of coefficients. For these +reasons we define an abstract semiring here. *) + +Open Local Scope NatScope. + +Lemma Nsemi_ring : semi_ring_theory 0 1 add mul Neq. +Proof. +constructor. +exact add_0_l. +exact add_comm. +exact add_assoc. +exact mul_1_l. +exact mul_0_l. +exact mul_comm. +exact mul_assoc. +exact mul_add_distr_r. +Qed. + +Add Ring NSR : Nsemi_ring. + +(* The definitios of functions (NZadd, NZmul, etc.) will be unfolded by +the properties functor. Since we don't want Zadd_comm to refer to unfolded +definitions of equality: fun p1 p2 : NZ => (fst p1 + snd p2) = (fst p2 + snd p1), +we will provide an extra layer of definitions. *) + +Definition Z := (N * N)%type. +Definition Z0 : Z := (0, 0). +Definition Zeq (p1 p2 : Z) := ((fst p1) + (snd p2) == (fst p2) + (snd p1)). +Definition Zsucc (n : Z) : Z := (S (fst n), snd n). +Definition Zpred (n : Z) : Z := (fst n, S (snd n)). + +(* We do not have Zpred (Zsucc n) = n but only Zpred (Zsucc n) == n. It +could be possible to consider as canonical only pairs where one of the +elements is 0, and make all operations convert canonical values into other +canonical values. In that case, we could get rid of setoids and arrive at +integers as signed natural numbers. *) + +Definition Zadd (n m : Z) : Z := ((fst n) + (fst m), (snd n) + (snd m)). +Definition Zsub (n m : Z) : Z := ((fst n) + (snd m), (snd n) + (fst m)). + +(* Unfortunately, the elements of the pair keep increasing, even during +subtraction *) + +Definition Zmul (n m : Z) : Z := + ((fst n) * (fst m) + (snd n) * (snd m), (fst n) * (snd m) + (snd n) * (fst m)). +Definition Zlt (n m : Z) := (fst n) + (snd m) < (fst m) + (snd n). +Definition Zle (n m : Z) := (fst n) + (snd m) <= (fst m) + (snd n). +Definition Zmin (n m : Z) := (min ((fst n) + (snd m)) ((fst m) + (snd n)), (snd n) + (snd m)). +Definition Zmax (n m : Z) := (max ((fst n) + (snd m)) ((fst m) + (snd n)), (snd n) + (snd m)). + +Delimit Scope IntScope with Int. +Bind Scope IntScope with Z. +Notation "x == y" := (Zeq x y) (at level 70) : IntScope. +Notation "x ~= y" := (~ Zeq x y) (at level 70) : IntScope. +Notation "0" := Z0 : IntScope. +Notation "1" := (Zsucc Z0) : IntScope. +Notation "x + y" := (Zadd x y) : IntScope. +Notation "x - y" := (Zsub x y) : IntScope. +Notation "x * y" := (Zmul x y) : IntScope. +Notation "x < y" := (Zlt x y) : IntScope. +Notation "x <= y" := (Zle x y) : IntScope. +Notation "x > y" := (Zlt y x) (only parsing) : IntScope. +Notation "x >= y" := (Zle y x) (only parsing) : IntScope. + +Notation Local N := NZ. +(* To remember N without having to use a long qualifying name. since NZ will be redefined *) +Notation Local NE := NZeq (only parsing). +Notation Local add_wd := NZadd_wd (only parsing). + +Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig. +Module Export NZAxiomsMod <: NZAxiomsSig. + +Definition NZ : Type := Z. +Definition NZeq := Zeq. +Definition NZ0 := Z0. +Definition NZsucc := Zsucc. +Definition NZpred := Zpred. +Definition NZadd := Zadd. +Definition NZsub := Zsub. +Definition NZmul := Zmul. + +Theorem ZE_refl : reflexive Z Zeq. +Proof. +unfold reflexive, Zeq. reflexivity. +Qed. + +Theorem ZE_symm : symmetric Z Zeq. +Proof. +unfold symmetric, Zeq; now symmetry. +Qed. + +Theorem ZE_trans : transitive Z Zeq. +Proof. +unfold transitive, Zeq. intros n m p H1 H2. +assert (H3 : (fst n + snd m) + (fst m + snd p) == (fst m + snd n) + (fst p + snd m)) +by now apply add_wd. +stepl ((fst n + snd p) + (fst m + snd m)) in H3 by ring. +stepr ((fst p + snd n) + (fst m + snd m)) in H3 by ring. +now apply -> add_cancel_r in H3. +Qed. + +Theorem NZeq_equiv : equiv Z Zeq. +Proof. +unfold equiv; repeat split; [apply ZE_refl | apply ZE_trans | apply ZE_symm]. +Qed. + +Add Relation Z Zeq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) +as NZeq_rel. + +Add Morphism (@pair N N) with signature NE ==> NE ==> Zeq as Zpair_wd. +Proof. +intros n1 n2 H1 m1 m2 H2; unfold Zeq; simpl; rewrite H1; now rewrite H2. +Qed. + +Add Morphism NZsucc with signature Zeq ==> Zeq as NZsucc_wd. +Proof. +unfold NZsucc, Zeq; intros n m H; simpl. +do 2 rewrite add_succ_l; now rewrite H. +Qed. + +Add Morphism NZpred with signature Zeq ==> Zeq as NZpred_wd. +Proof. +unfold NZpred, Zeq; intros n m H; simpl. +do 2 rewrite add_succ_r; now rewrite H. +Qed. + +Add Morphism NZadd with signature Zeq ==> Zeq ==> Zeq as NZadd_wd. +Proof. +unfold Zeq, NZadd; intros n1 m1 H1 n2 m2 H2; simpl. +assert (H3 : (fst n1 + snd m1) + (fst n2 + snd m2) == (fst m1 + snd n1) + (fst m2 + snd n2)) +by now apply add_wd. +stepl (fst n1 + snd m1 + (fst n2 + snd m2)) by ring. +now stepr (fst m1 + snd n1 + (fst m2 + snd n2)) by ring. +Qed. + +Add Morphism NZsub with signature Zeq ==> Zeq ==> Zeq as NZsub_wd. +Proof. +unfold Zeq, NZsub; intros n1 m1 H1 n2 m2 H2; simpl. +symmetry in H2. +assert (H3 : (fst n1 + snd m1) + (fst m2 + snd n2) == (fst m1 + snd n1) + (fst n2 + snd m2)) +by now apply add_wd. +stepl (fst n1 + snd m1 + (fst m2 + snd n2)) by ring. +now stepr (fst m1 + snd n1 + (fst n2 + snd m2)) by ring. +Qed. + +Add Morphism NZmul with signature Zeq ==> Zeq ==> Zeq as NZmul_wd. +Proof. +unfold NZmul, Zeq; intros n1 m1 H1 n2 m2 H2; simpl. +stepl (fst n1 * fst n2 + (snd n1 * snd n2 + fst m1 * snd m2 + snd m1 * fst m2)) by ring. +stepr (fst n1 * snd n2 + (fst m1 * fst m2 + snd m1 * snd m2 + snd n1 * fst n2)) by ring. +apply add_mul_repl_pair with (n := fst m2) (m := snd m2); [| now idtac]. +stepl (snd n1 * snd n2 + (fst n1 * fst m2 + fst m1 * snd m2 + snd m1 * fst m2)) by ring. +stepr (snd n1 * fst n2 + (fst n1 * snd m2 + fst m1 * fst m2 + snd m1 * snd m2)) by ring. +apply add_mul_repl_pair with (n := snd m2) (m := fst m2); +[| (stepl (fst n2 + snd m2) by ring); now stepr (fst m2 + snd n2) by ring]. +stepl (snd m2 * snd n1 + (fst n1 * fst m2 + fst m1 * snd m2 + snd m1 * fst m2)) by ring. +stepr (snd m2 * fst n1 + (snd n1 * fst m2 + fst m1 * fst m2 + snd m1 * snd m2)) by ring. +apply add_mul_repl_pair with (n := snd m1) (m := fst m1); +[ | (stepl (fst n1 + snd m1) by ring); now stepr (fst m1 + snd n1) by ring]. +stepl (fst m2 * fst n1 + (snd m2 * snd m1 + fst m1 * snd m2 + snd m1 * fst m2)) by ring. +stepr (fst m2 * snd n1 + (snd m2 * fst m1 + fst m1 * fst m2 + snd m1 * snd m2)) by ring. +apply add_mul_repl_pair with (n := fst m1) (m := snd m1); [| now idtac]. +ring. +Qed. + +Section Induction. +Open Scope NatScope. (* automatically closes at the end of the section *) +Variable A : Z -> Prop. +Hypothesis A_wd : predicate_wd Zeq A. + +Add Morphism A with signature Zeq ==> iff as A_morph. +Proof. +exact A_wd. +Qed. + +Theorem NZinduction : + A 0 -> (forall n : Z, A n <-> A (Zsucc n)) -> forall n : Z, A n. (* 0 is interpreted as in Z due to "Bind" directive *) +Proof. +intros A0 AS n; unfold NZ0, Zsucc, predicate_wd, fun_wd, Zeq in *. +destruct n as [n m]. +cut (forall p : N, A (p, 0)); [intro H1 |]. +cut (forall p : N, A (0, p)); [intro H2 |]. +destruct (add_dichotomy n m) as [[p H] | [p H]]. +rewrite (A_wd (n, m) (0, p)) by (rewrite add_0_l; now rewrite add_comm). +apply H2. +rewrite (A_wd (n, m) (p, 0)) by now rewrite add_0_r. apply H1. +induct p. assumption. intros p IH. +apply -> (A_wd (0, p) (1, S p)) in IH; [| now rewrite add_0_l, add_1_l]. +now apply <- AS. +induct p. assumption. intros p IH. +replace 0 with (snd (p, 0)); [| reflexivity]. +replace (S p) with (S (fst (p, 0))); [| reflexivity]. now apply -> AS. +Qed. + +End Induction. + +(* Time to prove theorems in the language of Z *) + +Open Local Scope IntScope. + +Theorem NZpred_succ : forall n : Z, Zpred (Zsucc n) == n. +Proof. +unfold NZpred, NZsucc, Zeq; intro n; simpl. +rewrite add_succ_l; now rewrite add_succ_r. +Qed. + +Theorem NZadd_0_l : forall n : Z, 0 + n == n. +Proof. +intro n; unfold NZadd, Zeq; simpl. now do 2 rewrite add_0_l. +Qed. + +Theorem NZadd_succ_l : forall n m : Z, (Zsucc n) + m == Zsucc (n + m). +Proof. +intros n m; unfold NZadd, Zeq; simpl. now do 2 rewrite add_succ_l. +Qed. + +Theorem NZsub_0_r : forall n : Z, n - 0 == n. +Proof. +intro n; unfold NZsub, Zeq; simpl. now do 2 rewrite add_0_r. +Qed. + +Theorem NZsub_succ_r : forall n m : Z, n - (Zsucc m) == Zpred (n - m). +Proof. +intros n m; unfold NZsub, Zeq; simpl. symmetry; now rewrite add_succ_r. +Qed. + +Theorem NZmul_0_l : forall n : Z, 0 * n == 0. +Proof. +intro n; unfold NZmul, Zeq; simpl. +repeat rewrite mul_0_l. now rewrite add_assoc. +Qed. + +Theorem NZmul_succ_l : forall n m : Z, (Zsucc n) * m == n * m + m. +Proof. +intros n m; unfold NZmul, NZsucc, Zeq; simpl. +do 2 rewrite mul_succ_l. ring. +Qed. + +End NZAxiomsMod. + +Definition NZlt := Zlt. +Definition NZle := Zle. +Definition NZmin := Zmin. +Definition NZmax := Zmax. + +Add Morphism NZlt with signature Zeq ==> Zeq ==> iff as NZlt_wd. +Proof. +unfold NZlt, Zlt, Zeq; intros n1 m1 H1 n2 m2 H2; simpl. split; intro H. +stepr (snd m1 + fst m2) by apply add_comm. +apply (add_lt_repl_pair (fst n1) (snd n1)); [| assumption]. +stepl (snd m2 + fst n1) by apply add_comm. +stepr (fst m2 + snd n1) by apply add_comm. +apply (add_lt_repl_pair (snd n2) (fst n2)). +now stepl (fst n1 + snd n2) by apply add_comm. +stepl (fst m2 + snd n2) by apply add_comm. now stepr (fst n2 + snd m2) by apply add_comm. +stepr (snd n1 + fst n2) by apply add_comm. +apply (add_lt_repl_pair (fst m1) (snd m1)); [| now symmetry]. +stepl (snd n2 + fst m1) by apply add_comm. +stepr (fst n2 + snd m1) by apply add_comm. +apply (add_lt_repl_pair (snd m2) (fst m2)). +now stepl (fst m1 + snd m2) by apply add_comm. +stepl (fst n2 + snd m2) by apply add_comm. now stepr (fst m2 + snd n2) by apply add_comm. +Qed. + +Add Morphism NZle with signature Zeq ==> Zeq ==> iff as NZle_wd. +Proof. +unfold NZle, Zle, Zeq; intros n1 m1 H1 n2 m2 H2; simpl. +do 2 rewrite lt_eq_cases. rewrite (NZlt_wd n1 m1 H1 n2 m2 H2). fold (m1 < m2)%Int. +fold (n1 == n2)%Int (m1 == m2)%Int; fold (n1 == m1)%Int in H1; fold (n2 == m2)%Int in H2. +now rewrite H1, H2. +Qed. + +Add Morphism NZmin with signature Zeq ==> Zeq ==> Zeq as NZmin_wd. +Proof. +intros n1 m1 H1 n2 m2 H2; unfold NZmin, Zeq; simpl. +destruct (le_ge_cases (fst n1 + snd n2) (fst n2 + snd n1)) as [H | H]. +rewrite (min_l (fst n1 + snd n2) (fst n2 + snd n1)) by assumption. +rewrite (min_l (fst m1 + snd m2) (fst m2 + snd m1)) by +now apply -> (NZle_wd n1 m1 H1 n2 m2 H2). +stepl ((fst n1 + snd m1) + (snd n2 + snd m2)) by ring. +unfold Zeq in H1. rewrite H1. ring. +rewrite (min_r (fst n1 + snd n2) (fst n2 + snd n1)) by assumption. +rewrite (min_r (fst m1 + snd m2) (fst m2 + snd m1)) by +now apply -> (NZle_wd n2 m2 H2 n1 m1 H1). +stepl ((fst n2 + snd m2) + (snd n1 + snd m1)) by ring. +unfold Zeq in H2. rewrite H2. ring. +Qed. + +Add Morphism NZmax with signature Zeq ==> Zeq ==> Zeq as NZmax_wd. +Proof. +intros n1 m1 H1 n2 m2 H2; unfold NZmax, Zeq; simpl. +destruct (le_ge_cases (fst n1 + snd n2) (fst n2 + snd n1)) as [H | H]. +rewrite (max_r (fst n1 + snd n2) (fst n2 + snd n1)) by assumption. +rewrite (max_r (fst m1 + snd m2) (fst m2 + snd m1)) by +now apply -> (NZle_wd n1 m1 H1 n2 m2 H2). +stepl ((fst n2 + snd m2) + (snd n1 + snd m1)) by ring. +unfold Zeq in H2. rewrite H2. ring. +rewrite (max_l (fst n1 + snd n2) (fst n2 + snd n1)) by assumption. +rewrite (max_l (fst m1 + snd m2) (fst m2 + snd m1)) by +now apply -> (NZle_wd n2 m2 H2 n1 m1 H1). +stepl ((fst n1 + snd m1) + (snd n2 + snd m2)) by ring. +unfold Zeq in H1. rewrite H1. ring. +Qed. + +Open Local Scope IntScope. + +Theorem NZlt_eq_cases : forall n m : Z, n <= m <-> n < m \/ n == m. +Proof. +intros n m; unfold Zlt, Zle, Zeq; simpl. apply lt_eq_cases. +Qed. + +Theorem NZlt_irrefl : forall n : Z, ~ (n < n). +Proof. +intros n; unfold Zlt, Zeq; simpl. apply lt_irrefl. +Qed. + +Theorem NZlt_succ_r : forall n m : Z, n < (Zsucc m) <-> n <= m. +Proof. +intros n m; unfold Zlt, Zle, Zeq; simpl. rewrite add_succ_l; apply lt_succ_r. +Qed. + +Theorem NZmin_l : forall n m : Z, n <= m -> Zmin n m == n. +Proof. +unfold Zmin, Zle, Zeq; simpl; intros n m H. +rewrite min_l by assumption. ring. +Qed. + +Theorem NZmin_r : forall n m : Z, m <= n -> Zmin n m == m. +Proof. +unfold Zmin, Zle, Zeq; simpl; intros n m H. +rewrite min_r by assumption. ring. +Qed. + +Theorem NZmax_l : forall n m : Z, m <= n -> Zmax n m == n. +Proof. +unfold Zmax, Zle, Zeq; simpl; intros n m H. +rewrite max_l by assumption. ring. +Qed. + +Theorem NZmax_r : forall n m : Z, n <= m -> Zmax n m == m. +Proof. +unfold Zmax, Zle, Zeq; simpl; intros n m H. +rewrite max_r by assumption. ring. +Qed. + +End NZOrdAxiomsMod. + +Definition Zopp (n : Z) : Z := (snd n, fst n). + +Notation "- x" := (Zopp x) : IntScope. + +Add Morphism Zopp with signature Zeq ==> Zeq as Zopp_wd. +Proof. +unfold Zeq; intros n m H; simpl. symmetry. +stepl (fst n + snd m) by apply add_comm. +now stepr (fst m + snd n) by apply add_comm. +Qed. + +Open Local Scope IntScope. + +Theorem Zsucc_pred : forall n : Z, Zsucc (Zpred n) == n. +Proof. +intro n; unfold Zsucc, Zpred, Zeq; simpl. +rewrite add_succ_l; now rewrite add_succ_r. +Qed. + +Theorem Zopp_0 : - 0 == 0. +Proof. +unfold Zopp, Zeq; simpl. now rewrite add_0_l. +Qed. + +Theorem Zopp_succ : forall n, - (Zsucc n) == Zpred (- n). +Proof. +reflexivity. +Qed. + +End ZPairsAxiomsMod. + +(* For example, let's build integers out of pairs of Peano natural numbers +and get their properties *) + +(* The following lines increase the compilation time at least twice *) +(* +Require Import NPeano. + +Module Export ZPairsPeanoAxiomsMod := ZPairsAxiomsMod NPeanoAxiomsMod. +Module Export ZPairsMulOrderPropMod := ZMulOrderPropFunct ZPairsPeanoAxiomsMod. + +Open Local Scope IntScope. + +Eval compute in (3, 5) * (4, 6). +*) + diff --git a/theories/Numbers/Integer/SpecViaZ/ZSig.v b/theories/Numbers/Integer/SpecViaZ/ZSig.v new file mode 100644 index 00000000..0af98c74 --- /dev/null +++ b/theories/Numbers/Integer/SpecViaZ/ZSig.v @@ -0,0 +1,117 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: ZSig.v 11027 2008-06-01 13:28:59Z letouzey $ i*) + +Require Import ZArith Znumtheory. + +Open Scope Z_scope. + +(** * ZSig *) + +(** Interface of a rich structure about integers. + Specifications are written via translation to Z. +*) + +Module Type ZType. + + Parameter t : Type. + + Parameter to_Z : t -> Z. + Notation "[ x ]" := (to_Z x). + + Definition eq x y := ([x] = [y]). + + Parameter of_Z : Z -> t. + Parameter spec_of_Z: forall x, to_Z (of_Z x) = x. + + Parameter zero : t. + Parameter one : t. + Parameter minus_one : t. + + Parameter spec_0: [zero] = 0. + Parameter spec_1: [one] = 1. + Parameter spec_m1: [minus_one] = -1. + + Parameter compare : t -> t -> comparison. + + Parameter spec_compare: forall x y, + match compare x y with + | Eq => [x] = [y] + | Lt => [x] < [y] + | Gt => [x] > [y] + end. + + Definition lt n m := compare n m = Lt. + Definition le n m := compare n m <> Gt. + Definition min n m := match compare n m with Gt => m | _ => n end. + Definition max n m := match compare n m with Lt => m | _ => n end. + + Parameter eq_bool : t -> t -> bool. + + Parameter spec_eq_bool: forall x y, + if eq_bool x y then [x] = [y] else [x] <> [y]. + + Parameter succ : t -> t. + + Parameter spec_succ: forall n, [succ n] = [n] + 1. + + Parameter add : t -> t -> t. + + Parameter spec_add: forall x y, [add x y] = [x] + [y]. + + Parameter pred : t -> t. + + Parameter spec_pred: forall x, [pred x] = [x] - 1. + + Parameter sub : t -> t -> t. + + Parameter spec_sub: forall x y, [sub x y] = [x] - [y]. + + Parameter opp : t -> t. + + Parameter spec_opp: forall x, [opp x] = - [x]. + + Parameter mul : t -> t -> t. + + Parameter spec_mul: forall x y, [mul x y] = [x] * [y]. + + Parameter square : t -> t. + + Parameter spec_square: forall x, [square x] = [x] * [x]. + + Parameter power_pos : t -> positive -> t. + + Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n. + + Parameter sqrt : t -> t. + + Parameter spec_sqrt: forall x, 0 <= [x] -> + [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2. + + Parameter div_eucl : t -> t -> t * t. + + Parameter spec_div_eucl: forall x y, [y] <> 0 -> + let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y]. + + Parameter div : t -> t -> t. + + Parameter spec_div: forall x y, [y] <> 0 -> [div x y] = [x] / [y]. + + Parameter modulo : t -> t -> t. + + Parameter spec_modulo: forall x y, [y] <> 0 -> + [modulo x y] = [x] mod [y]. + + Parameter gcd : t -> t -> t. + + Parameter spec_gcd: forall a b, [gcd a b] = Zgcd (to_Z a) (to_Z b). + +End ZType. diff --git a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v new file mode 100644 index 00000000..d7c56267 --- /dev/null +++ b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v @@ -0,0 +1,306 @@ +(************************************************************************) +(* 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: ZSigZAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import ZArith. +Require Import ZAxioms. +Require Import ZSig. + +(** * The interface [ZSig.ZType] implies the interface [ZAxiomsSig] *) + +Module ZSig_ZAxioms (Z:ZType) <: ZAxiomsSig. + +Delimit Scope IntScope with Int. +Bind Scope IntScope with Z.t. +Open Local Scope IntScope. +Notation "[ x ]" := (Z.to_Z x) : IntScope. +Infix "==" := Z.eq (at level 70) : IntScope. +Notation "0" := Z.zero : IntScope. +Infix "+" := Z.add : IntScope. +Infix "-" := Z.sub : IntScope. +Infix "*" := Z.mul : IntScope. +Notation "- x" := (Z.opp x) : IntScope. + +Hint Rewrite + Z.spec_0 Z.spec_1 Z.spec_add Z.spec_sub Z.spec_pred Z.spec_succ + Z.spec_mul Z.spec_opp Z.spec_of_Z : Zspec. + +Ltac zsimpl := unfold Z.eq in *; autorewrite with Zspec. + +Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig. +Module Export NZAxiomsMod <: NZAxiomsSig. + +Definition NZ := Z.t. +Definition NZeq := Z.eq. +Definition NZ0 := Z.zero. +Definition NZsucc := Z.succ. +Definition NZpred := Z.pred. +Definition NZadd := Z.add. +Definition NZsub := Z.sub. +Definition NZmul := Z.mul. + +Theorem NZeq_equiv : equiv Z.t Z.eq. +Proof. +repeat split; repeat red; intros; auto; congruence. +Qed. + +Add Relation Z.t Z.eq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) + as NZeq_rel. + +Add Morphism NZsucc with signature Z.eq ==> Z.eq as NZsucc_wd. +Proof. +intros; zsimpl; f_equal; assumption. +Qed. + +Add Morphism NZpred with signature Z.eq ==> Z.eq as NZpred_wd. +Proof. +intros; zsimpl; f_equal; assumption. +Qed. + +Add Morphism NZadd with signature Z.eq ==> Z.eq ==> Z.eq as NZadd_wd. +Proof. +intros; zsimpl; f_equal; assumption. +Qed. + +Add Morphism NZsub with signature Z.eq ==> Z.eq ==> Z.eq as NZsub_wd. +Proof. +intros; zsimpl; f_equal; assumption. +Qed. + +Add Morphism NZmul with signature Z.eq ==> Z.eq ==> Z.eq as NZmul_wd. +Proof. +intros; zsimpl; f_equal; assumption. +Qed. + +Theorem NZpred_succ : forall n, Z.pred (Z.succ n) == n. +Proof. +intros; zsimpl; auto with zarith. +Qed. + +Section Induction. + +Variable A : Z.t -> Prop. +Hypothesis A_wd : predicate_wd Z.eq A. +Hypothesis A0 : A 0. +Hypothesis AS : forall n, A n <-> A (Z.succ n). + +Add Morphism A with signature Z.eq ==> iff as A_morph. +Proof. apply A_wd. Qed. + +Let B (z : Z) := A (Z.of_Z z). + +Lemma B0 : B 0. +Proof. +unfold B; simpl. +rewrite <- (A_wd 0); auto. +zsimpl; auto. +Qed. + +Lemma BS : forall z : Z, B z -> B (z + 1). +Proof. +intros z H. +unfold B in *. apply -> AS in H. +setoid_replace (Z.of_Z (z + 1)) with (Z.succ (Z.of_Z z)); auto. +zsimpl; auto. +Qed. + +Lemma BP : forall z : Z, B z -> B (z - 1). +Proof. +intros z H. +unfold B in *. rewrite AS. +setoid_replace (Z.succ (Z.of_Z (z - 1))) with (Z.of_Z z); auto. +zsimpl; auto with zarith. +Qed. + +Lemma B_holds : forall z : Z, B z. +Proof. +intros; destruct (Z_lt_le_dec 0 z). +apply natlike_ind; auto with zarith. +apply B0. +intros; apply BS; auto. +replace z with (-(-z))%Z in * by (auto with zarith). +remember (-z)%Z as z'. +pattern z'; apply natlike_ind. +apply B0. +intros; rewrite Zopp_succ; unfold Zpred; apply BP; auto. +subst z'; auto with zarith. +Qed. + +Theorem NZinduction : forall n, A n. +Proof. +intro n. setoid_replace n with (Z.of_Z (Z.to_Z n)). +apply B_holds. +zsimpl; auto. +Qed. + +End Induction. + +Theorem NZadd_0_l : forall n, 0 + n == n. +Proof. +intros; zsimpl; auto with zarith. +Qed. + +Theorem NZadd_succ_l : forall n m, (Z.succ n) + m == Z.succ (n + m). +Proof. +intros; zsimpl; auto with zarith. +Qed. + +Theorem NZsub_0_r : forall n, n - 0 == n. +Proof. +intros; zsimpl; auto with zarith. +Qed. + +Theorem NZsub_succ_r : forall n m, n - (Z.succ m) == Z.pred (n - m). +Proof. +intros; zsimpl; auto with zarith. +Qed. + +Theorem NZmul_0_l : forall n, 0 * n == 0. +Proof. +intros; zsimpl; auto with zarith. +Qed. + +Theorem NZmul_succ_l : forall n m, (Z.succ n) * m == n * m + m. +Proof. +intros; zsimpl; ring. +Qed. + +End NZAxiomsMod. + +Definition NZlt := Z.lt. +Definition NZle := Z.le. +Definition NZmin := Z.min. +Definition NZmax := Z.max. + +Infix "<=" := Z.le : IntScope. +Infix "<" := Z.lt : IntScope. + +Lemma spec_compare_alt : forall x y, Z.compare x y = ([x] ?= [y])%Z. +Proof. + intros; generalize (Z.spec_compare x y). + destruct (Z.compare x y); auto. + intros H; rewrite H; symmetry; apply Zcompare_refl. +Qed. + +Lemma spec_lt : forall x y, (x<y) <-> ([x]<[y])%Z. +Proof. + intros; unfold Z.lt, Zlt; rewrite spec_compare_alt; intuition. +Qed. + +Lemma spec_le : forall x y, (x<=y) <-> ([x]<=[y])%Z. +Proof. + intros; unfold Z.le, Zle; rewrite spec_compare_alt; intuition. +Qed. + +Lemma spec_min : forall x y, [Z.min x y] = Zmin [x] [y]. +Proof. + intros; unfold Z.min, Zmin. + rewrite spec_compare_alt; destruct Zcompare; auto. +Qed. + +Lemma spec_max : forall x y, [Z.max x y] = Zmax [x] [y]. +Proof. + intros; unfold Z.max, Zmax. + rewrite spec_compare_alt; destruct Zcompare; auto. +Qed. + +Add Morphism Z.compare with signature Z.eq ==> Z.eq ==> (@eq comparison) as compare_wd. +Proof. +intros x x' Hx y y' Hy. +rewrite 2 spec_compare_alt; rewrite Hx, Hy; intuition. +Qed. + +Add Morphism Z.lt with signature Z.eq ==> Z.eq ==> iff as NZlt_wd. +Proof. +intros x x' Hx y y' Hy; unfold Z.lt; rewrite Hx, Hy; intuition. +Qed. + +Add Morphism Z.le with signature Z.eq ==> Z.eq ==> iff as NZle_wd. +Proof. +intros x x' Hx y y' Hy; unfold Z.le; rewrite Hx, Hy; intuition. +Qed. + +Add Morphism Z.min with signature Z.eq ==> Z.eq ==> Z.eq as NZmin_wd. +Proof. +intros; red; rewrite 2 spec_min; congruence. +Qed. + +Add Morphism Z.max with signature Z.eq ==> Z.eq ==> Z.eq as NZmax_wd. +Proof. +intros; red; rewrite 2 spec_max; congruence. +Qed. + +Theorem NZlt_eq_cases : forall n m, n <= m <-> n < m \/ n == m. +Proof. +intros. +unfold Z.eq; rewrite spec_lt, spec_le; omega. +Qed. + +Theorem NZlt_irrefl : forall n, ~ n < n. +Proof. +intros; rewrite spec_lt; auto with zarith. +Qed. + +Theorem NZlt_succ_r : forall n m, n < (Z.succ m) <-> n <= m. +Proof. +intros; rewrite spec_lt, spec_le, Z.spec_succ; omega. +Qed. + +Theorem NZmin_l : forall n m, n <= m -> Z.min n m == n. +Proof. +intros n m; unfold Z.eq; rewrite spec_le, spec_min. +generalize (Zmin_spec [n] [m]); omega. +Qed. + +Theorem NZmin_r : forall n m, m <= n -> Z.min n m == m. +Proof. +intros n m; unfold Z.eq; rewrite spec_le, spec_min. +generalize (Zmin_spec [n] [m]); omega. +Qed. + +Theorem NZmax_l : forall n m, m <= n -> Z.max n m == n. +Proof. +intros n m; unfold Z.eq; rewrite spec_le, spec_max. +generalize (Zmax_spec [n] [m]); omega. +Qed. + +Theorem NZmax_r : forall n m, n <= m -> Z.max n m == m. +Proof. +intros n m; unfold Z.eq; rewrite spec_le, spec_max. +generalize (Zmax_spec [n] [m]); omega. +Qed. + +End NZOrdAxiomsMod. + +Definition Zopp := Z.opp. + +Add Morphism Z.opp with signature Z.eq ==> Z.eq as Zopp_wd. +Proof. +intros; zsimpl; auto with zarith. +Qed. + +Theorem Zsucc_pred : forall n, Z.succ (Z.pred n) == n. +Proof. +red; intros; zsimpl; auto with zarith. +Qed. + +Theorem Zopp_0 : - 0 == 0. +Proof. +red; intros; zsimpl; auto with zarith. +Qed. + +Theorem Zopp_succ : forall n, - (Z.succ n) == Z.pred (- n). +Proof. +intros; zsimpl; auto with zarith. +Qed. + +End ZSig_ZAxioms. diff --git a/theories/Numbers/NaryFunctions.v b/theories/Numbers/NaryFunctions.v new file mode 100644 index 00000000..04a48d51 --- /dev/null +++ b/theories/Numbers/NaryFunctions.v @@ -0,0 +1,142 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) +(* Pierre Letouzey, Jerome Vouillon, PPS, Paris 7, 2008 *) +(************************************************************************) + +(*i $Id: NaryFunctions.v 10967 2008-05-22 12:59:38Z letouzey $ i*) + +Open Local Scope type_scope. + +Require Import List. + +(** * Generic dependently-typed operators about [n]-ary functions *) + +(** The type of [n]-ary function: [nfun A n B] is + [A -> ... -> A -> B] with [n] occurences of [A] in this type. *) + +Fixpoint nfun A n B := + match n with + | O => B + | S n => A -> (nfun A n B) + end. + +Notation " A ^^ n --> B " := (nfun A n B) + (at level 50, n at next level) : type_scope. + +(** [napply_cst _ _ a n f] iterates [n] times the application of a + particular constant [a] to the [n]-ary function [f]. *) + +Fixpoint napply_cst (A B:Type)(a:A) n : (A^^n-->B) -> B := + match n return (A^^n-->B) -> B with + | O => fun x => x + | S n => fun x => napply_cst _ _ a n (x a) + end. + + +(** A generic transformation from an n-ary function to another one.*) + +Fixpoint nfun_to_nfun (A B C:Type)(f:B -> C) n : + (A^^n-->B) -> (A^^n-->C) := + match n return (A^^n-->B) -> (A^^n-->C) with + | O => f + | S n => fun g a => nfun_to_nfun _ _ _ f n (g a) + end. + +(** [napply_except_last _ _ n f] expects [n] arguments of type [A], + applies [n-1] of them to [f] and discard the last one. *) + +Definition napply_except_last (A B:Type) := + nfun_to_nfun A B (A->B) (fun b a => b). + +(** [napply_then_last _ _ a n f] expects [n] arguments of type [A], + applies them to [f] and then apply [a] to the result. *) + +Definition napply_then_last (A B:Type)(a:A) := + nfun_to_nfun A (A->B) B (fun fab => fab a). + +(** [napply_discard _ b n] expects [n] arguments, discards then, + and returns [b]. *) + +Fixpoint napply_discard (A B:Type)(b:B) n : A^^n-->B := + match n return A^^n-->B with + | O => b + | S n => fun _ => napply_discard _ _ b n + end. + +(** A fold function *) + +Fixpoint nfold A B (f:A->B->B)(b:B) n : (A^^n-->B) := + match n return (A^^n-->B) with + | O => b + | S n => fun a => (nfold _ _ f (f a b) n) + end. + + +(** [n]-ary products : [nprod A n] is [A*...*A*unit], + with [n] occurrences of [A] in this type. *) + +Fixpoint nprod A n : Type := match n with + | O => unit + | S n => (A * nprod A n)%type +end. + +Notation "A ^ n" := (nprod A n) : type_scope. + +(** [n]-ary curryfication / uncurryfication *) + +Fixpoint ncurry (A B:Type) n : (A^n -> B) -> (A^^n-->B) := + match n return (A^n -> B) -> (A^^n-->B) with + | O => fun x => x tt + | S n => fun f a => ncurry _ _ n (fun p => f (a,p)) + end. + +Fixpoint nuncurry (A B:Type) n : (A^^n-->B) -> (A^n -> B) := + match n return (A^^n-->B) -> (A^n -> B) with + | O => fun x _ => x + | S n => fun f p => let (x,p) := p in nuncurry _ _ n (f x) p + end. + +(** Earlier functions can also be defined via [ncurry/nuncurry]. + For instance : *) + +Definition nfun_to_nfun_bis A B C (f:B->C) n : + (A^^n-->B) -> (A^^n-->C) := + fun anb => ncurry _ _ n (fun an => f ((nuncurry _ _ n anb) an)). + +(** We can also us it to obtain another [fold] function, + equivalent to the previous one, but with a nicer expansion + (see for instance Int31.iszero). *) + +Fixpoint nfold_bis A B (f:A->B->B)(b:B) n : (A^^n-->B) := + match n return (A^^n-->B) with + | O => b + | S n => fun a => + nfun_to_nfun_bis _ _ _ (f a) n (nfold_bis _ _ f b n) + end. + +(** From [nprod] to [list] *) + +Fixpoint nprod_to_list (A:Type) n : A^n -> list A := + match n with + | O => fun _ => nil + | S n => fun p => let (x,p) := p in x::(nprod_to_list _ n p) + end. + +(** From [list] to [nprod] *) + +Fixpoint nprod_of_list (A:Type)(l:list A) : A^(length l) := + match l return A^(length l) with + | nil => tt + | x::l => (x, nprod_of_list _ l) + end. + +(** This gives an additional way to write the fold *) + +Definition nfold_list (A B:Type)(f:A->B->B)(b:B) n : (A^^n-->B) := + ncurry _ _ n (fun p => fold_right f b (nprod_to_list _ _ p)). + diff --git a/theories/Numbers/NatInt/NZAdd.v b/theories/Numbers/NatInt/NZAdd.v new file mode 100644 index 00000000..c9bb5c95 --- /dev/null +++ b/theories/Numbers/NatInt/NZAdd.v @@ -0,0 +1,91 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZAdd.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import NZAxioms. +Require Import NZBase. + +Module NZAddPropFunct (Import NZAxiomsMod : NZAxiomsSig). +Module Export NZBasePropMod := NZBasePropFunct NZAxiomsMod. +Open Local Scope NatIntScope. + +Theorem NZadd_0_r : forall n : NZ, n + 0 == n. +Proof. +NZinduct n. now rewrite NZadd_0_l. +intro. rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd. +Qed. + +Theorem NZadd_succ_r : forall n m : NZ, n + S m == S (n + m). +Proof. +intros n m; NZinduct n. +now do 2 rewrite NZadd_0_l. +intro. repeat rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd. +Qed. + +Theorem NZadd_comm : forall n m : NZ, n + m == m + n. +Proof. +intros n m; NZinduct n. +rewrite NZadd_0_l; now rewrite NZadd_0_r. +intros n. rewrite NZadd_succ_l; rewrite NZadd_succ_r. now rewrite NZsucc_inj_wd. +Qed. + +Theorem NZadd_1_l : forall n : NZ, 1 + n == S n. +Proof. +intro n; rewrite NZadd_succ_l; now rewrite NZadd_0_l. +Qed. + +Theorem NZadd_1_r : forall n : NZ, n + 1 == S n. +Proof. +intro n; rewrite NZadd_comm; apply NZadd_1_l. +Qed. + +Theorem NZadd_assoc : forall n m p : NZ, n + (m + p) == (n + m) + p. +Proof. +intros n m p; NZinduct n. +now do 2 rewrite NZadd_0_l. +intro. do 3 rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd. +Qed. + +Theorem NZadd_shuffle1 : forall n m p q : NZ, (n + m) + (p + q) == (n + p) + (m + q). +Proof. +intros n m p q. +rewrite <- (NZadd_assoc n m (p + q)). rewrite (NZadd_comm m (p + q)). +rewrite <- (NZadd_assoc p q m). rewrite (NZadd_assoc n p (q + m)). +now rewrite (NZadd_comm q m). +Qed. + +Theorem NZadd_shuffle2 : forall n m p q : NZ, (n + m) + (p + q) == (n + q) + (m + p). +Proof. +intros n m p q. +rewrite <- (NZadd_assoc n m (p + q)). rewrite (NZadd_assoc m p q). +rewrite (NZadd_comm (m + p) q). now rewrite <- (NZadd_assoc n q (m + p)). +Qed. + +Theorem NZadd_cancel_l : forall n m p : NZ, p + n == p + m <-> n == m. +Proof. +intros n m p; NZinduct p. +now do 2 rewrite NZadd_0_l. +intro p. do 2 rewrite NZadd_succ_l. now rewrite NZsucc_inj_wd. +Qed. + +Theorem NZadd_cancel_r : forall n m p : NZ, n + p == m + p <-> n == m. +Proof. +intros n m p. rewrite (NZadd_comm n p); rewrite (NZadd_comm m p). +apply NZadd_cancel_l. +Qed. + +Theorem NZsub_1_r : forall n : NZ, n - 1 == P n. +Proof. +intro n; rewrite NZsub_succ_r; now rewrite NZsub_0_r. +Qed. + +End NZAddPropFunct. + diff --git a/theories/Numbers/NatInt/NZAddOrder.v b/theories/Numbers/NatInt/NZAddOrder.v new file mode 100644 index 00000000..50d1c42f --- /dev/null +++ b/theories/Numbers/NatInt/NZAddOrder.v @@ -0,0 +1,166 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZAddOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import NZAxioms. +Require Import NZOrder. + +Module NZAddOrderPropFunct (Import NZOrdAxiomsMod : NZOrdAxiomsSig). +Module Export NZOrderPropMod := NZOrderPropFunct NZOrdAxiomsMod. +Open Local Scope NatIntScope. + +Theorem NZadd_lt_mono_l : forall n m p : NZ, n < m <-> p + n < p + m. +Proof. +intros n m p; NZinduct p. +now do 2 rewrite NZadd_0_l. +intro p. do 2 rewrite NZadd_succ_l. now rewrite <- NZsucc_lt_mono. +Qed. + +Theorem NZadd_lt_mono_r : forall n m p : NZ, n < m <-> n + p < m + p. +Proof. +intros n m p. +rewrite (NZadd_comm n p); rewrite (NZadd_comm m p); apply NZadd_lt_mono_l. +Qed. + +Theorem NZadd_lt_mono : forall n m p q : NZ, n < m -> p < q -> n + p < m + q. +Proof. +intros n m p q H1 H2. +apply NZlt_trans with (m + p); +[now apply -> NZadd_lt_mono_r | now apply -> NZadd_lt_mono_l]. +Qed. + +Theorem NZadd_le_mono_l : forall n m p : NZ, n <= m <-> p + n <= p + m. +Proof. +intros n m p; NZinduct p. +now do 2 rewrite NZadd_0_l. +intro p. do 2 rewrite NZadd_succ_l. now rewrite <- NZsucc_le_mono. +Qed. + +Theorem NZadd_le_mono_r : forall n m p : NZ, n <= m <-> n + p <= m + p. +Proof. +intros n m p. +rewrite (NZadd_comm n p); rewrite (NZadd_comm m p); apply NZadd_le_mono_l. +Qed. + +Theorem NZadd_le_mono : forall n m p q : NZ, n <= m -> p <= q -> n + p <= m + q. +Proof. +intros n m p q H1 H2. +apply NZle_trans with (m + p); +[now apply -> NZadd_le_mono_r | now apply -> NZadd_le_mono_l]. +Qed. + +Theorem NZadd_lt_le_mono : forall n m p q : NZ, n < m -> p <= q -> n + p < m + q. +Proof. +intros n m p q H1 H2. +apply NZlt_le_trans with (m + p); +[now apply -> NZadd_lt_mono_r | now apply -> NZadd_le_mono_l]. +Qed. + +Theorem NZadd_le_lt_mono : forall n m p q : NZ, n <= m -> p < q -> n + p < m + q. +Proof. +intros n m p q H1 H2. +apply NZle_lt_trans with (m + p); +[now apply -> NZadd_le_mono_r | now apply -> NZadd_lt_mono_l]. +Qed. + +Theorem NZadd_pos_pos : forall n m : NZ, 0 < n -> 0 < m -> 0 < n + m. +Proof. +intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_lt_mono. +Qed. + +Theorem NZadd_pos_nonneg : forall n m : NZ, 0 < n -> 0 <= m -> 0 < n + m. +Proof. +intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_lt_le_mono. +Qed. + +Theorem NZadd_nonneg_pos : forall n m : NZ, 0 <= n -> 0 < m -> 0 < n + m. +Proof. +intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_le_lt_mono. +Qed. + +Theorem NZadd_nonneg_nonneg : forall n m : NZ, 0 <= n -> 0 <= m -> 0 <= n + m. +Proof. +intros n m H1 H2. rewrite <- (NZadd_0_l 0). now apply NZadd_le_mono. +Qed. + +Theorem NZlt_add_pos_l : forall n m : NZ, 0 < n -> m < n + m. +Proof. +intros n m H. apply -> (NZadd_lt_mono_r 0 n m) in H. +now rewrite NZadd_0_l in H. +Qed. + +Theorem NZlt_add_pos_r : forall n m : NZ, 0 < n -> m < m + n. +Proof. +intros; rewrite NZadd_comm; now apply NZlt_add_pos_l. +Qed. + +Theorem NZle_lt_add_lt : forall n m p q : NZ, n <= m -> p + m < q + n -> p < q. +Proof. +intros n m p q H1 H2. destruct (NZle_gt_cases q p); [| assumption]. +pose proof (NZadd_le_mono q p n m H H1) as H3. apply <- NZnle_gt in H2. +false_hyp H3 H2. +Qed. + +Theorem NZlt_le_add_lt : forall n m p q : NZ, n < m -> p + m <= q + n -> p < q. +Proof. +intros n m p q H1 H2. destruct (NZle_gt_cases q p); [| assumption]. +pose proof (NZadd_le_lt_mono q p n m H H1) as H3. apply <- NZnle_gt in H3. +false_hyp H2 H3. +Qed. + +Theorem NZle_le_add_le : forall n m p q : NZ, n <= m -> p + m <= q + n -> p <= q. +Proof. +intros n m p q H1 H2. destruct (NZle_gt_cases p q); [assumption |]. +pose proof (NZadd_lt_le_mono q p n m H H1) as H3. apply <- NZnle_gt in H3. +false_hyp H2 H3. +Qed. + +Theorem NZadd_lt_cases : forall n m p q : NZ, n + m < p + q -> n < p \/ m < q. +Proof. +intros n m p q H; +destruct (NZle_gt_cases p n) as [H1 | H1]. +destruct (NZle_gt_cases q m) as [H2 | H2]. +pose proof (NZadd_le_mono p n q m H1 H2) as H3. apply -> NZle_ngt in H3. +false_hyp H H3. +now right. now left. +Qed. + +Theorem NZadd_le_cases : forall n m p q : NZ, n + m <= p + q -> n <= p \/ m <= q. +Proof. +intros n m p q H. +destruct (NZle_gt_cases n p) as [H1 | H1]. now left. +destruct (NZle_gt_cases m q) as [H2 | H2]. now right. +assert (H3 : p + q < n + m) by now apply NZadd_lt_mono. +apply -> NZle_ngt in H. false_hyp H3 H. +Qed. + +Theorem NZadd_neg_cases : forall n m : NZ, n + m < 0 -> n < 0 \/ m < 0. +Proof. +intros n m H; apply NZadd_lt_cases; now rewrite NZadd_0_l. +Qed. + +Theorem NZadd_pos_cases : forall n m : NZ, 0 < n + m -> 0 < n \/ 0 < m. +Proof. +intros n m H; apply NZadd_lt_cases; now rewrite NZadd_0_l. +Qed. + +Theorem NZadd_nonpos_cases : forall n m : NZ, n + m <= 0 -> n <= 0 \/ m <= 0. +Proof. +intros n m H; apply NZadd_le_cases; now rewrite NZadd_0_l. +Qed. + +Theorem NZadd_nonneg_cases : forall n m : NZ, 0 <= n + m -> 0 <= n \/ 0 <= m. +Proof. +intros n m H; apply NZadd_le_cases; now rewrite NZadd_0_l. +Qed. + +End NZAddOrderPropFunct. + diff --git a/theories/Numbers/NatInt/NZAxioms.v b/theories/Numbers/NatInt/NZAxioms.v new file mode 100644 index 00000000..26933646 --- /dev/null +++ b/theories/Numbers/NatInt/NZAxioms.v @@ -0,0 +1,99 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NumPrelude. + +Module Type NZAxiomsSig. + +Parameter Inline NZ : Type. +Parameter Inline NZeq : NZ -> NZ -> Prop. +Parameter Inline NZ0 : NZ. +Parameter Inline NZsucc : NZ -> NZ. +Parameter Inline NZpred : NZ -> NZ. +Parameter Inline NZadd : NZ -> NZ -> NZ. +Parameter Inline NZsub : NZ -> NZ -> NZ. +Parameter Inline NZmul : NZ -> NZ -> NZ. + +(* Unary subtraction (opp) is not defined on natural numbers, so we have + it for integers only *) + +Axiom NZeq_equiv : equiv NZ NZeq. +Add Relation NZ NZeq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) +as NZeq_rel. + +Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd. +Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd. +Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd. +Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd. +Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd. + +Delimit Scope NatIntScope with NatInt. +Open Local Scope NatIntScope. +Notation "x == y" := (NZeq x y) (at level 70) : NatIntScope. +Notation "x ~= y" := (~ NZeq x y) (at level 70) : NatIntScope. +Notation "0" := NZ0 : NatIntScope. +Notation S := NZsucc. +Notation P := NZpred. +Notation "1" := (S 0) : NatIntScope. +Notation "x + y" := (NZadd x y) : NatIntScope. +Notation "x - y" := (NZsub x y) : NatIntScope. +Notation "x * y" := (NZmul x y) : NatIntScope. + +Axiom NZpred_succ : forall n : NZ, P (S n) == n. + +Axiom NZinduction : + forall A : NZ -> Prop, predicate_wd NZeq A -> + A 0 -> (forall n : NZ, A n <-> A (S n)) -> forall n : NZ, A n. + +Axiom NZadd_0_l : forall n : NZ, 0 + n == n. +Axiom NZadd_succ_l : forall n m : NZ, (S n) + m == S (n + m). + +Axiom NZsub_0_r : forall n : NZ, n - 0 == n. +Axiom NZsub_succ_r : forall n m : NZ, n - (S m) == P (n - m). + +Axiom NZmul_0_l : forall n : NZ, 0 * n == 0. +Axiom NZmul_succ_l : forall n m : NZ, S n * m == n * m + m. + +End NZAxiomsSig. + +Module Type NZOrdAxiomsSig. +Declare Module Export NZAxiomsMod : NZAxiomsSig. +Open Local Scope NatIntScope. + +Parameter Inline NZlt : NZ -> NZ -> Prop. +Parameter Inline NZle : NZ -> NZ -> Prop. +Parameter Inline NZmin : NZ -> NZ -> NZ. +Parameter Inline NZmax : NZ -> NZ -> NZ. + +Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd. +Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd. +Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd. +Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd. + +Notation "x < y" := (NZlt x y) : NatIntScope. +Notation "x <= y" := (NZle x y) : NatIntScope. +Notation "x > y" := (NZlt y x) (only parsing) : NatIntScope. +Notation "x >= y" := (NZle y x) (only parsing) : NatIntScope. + +Axiom NZlt_eq_cases : forall n m : NZ, n <= m <-> n < m \/ n == m. +Axiom NZlt_irrefl : forall n : NZ, ~ (n < n). +Axiom NZlt_succ_r : forall n m : NZ, n < S m <-> n <= m. + +Axiom NZmin_l : forall n m : NZ, n <= m -> NZmin n m == n. +Axiom NZmin_r : forall n m : NZ, m <= n -> NZmin n m == m. +Axiom NZmax_l : forall n m : NZ, m <= n -> NZmax n m == n. +Axiom NZmax_r : forall n m : NZ, n <= m -> NZmax n m == m. + +End NZOrdAxiomsSig. diff --git a/theories/Numbers/NatInt/NZBase.v b/theories/Numbers/NatInt/NZBase.v new file mode 100644 index 00000000..8b01e353 --- /dev/null +++ b/theories/Numbers/NatInt/NZBase.v @@ -0,0 +1,84 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZBase.v 10934 2008-05-15 21:58:20Z letouzey $ i*) + +Require Import NZAxioms. + +Module NZBasePropFunct (Import NZAxiomsMod : NZAxiomsSig). +Open Local Scope NatIntScope. + +Theorem NZneq_symm : forall n m : NZ, n ~= m -> m ~= n. +Proof. +intros n m H1 H2; symmetry in H2; false_hyp H2 H1. +Qed. + +Theorem NZE_stepl : forall x y z : NZ, x == y -> x == z -> z == y. +Proof. +intros x y z H1 H2; now rewrite <- H1. +Qed. + +Declare Left Step NZE_stepl. +(* The right step lemma is just the transitivity of NZeq *) +Declare Right Step (proj1 (proj2 NZeq_equiv)). + +Theorem NZsucc_inj : forall n1 n2 : NZ, S n1 == S n2 -> n1 == n2. +Proof. +intros n1 n2 H. +apply NZpred_wd in H. now do 2 rewrite NZpred_succ in H. +Qed. + +(* The following theorem is useful as an equivalence for proving +bidirectional induction steps *) +Theorem NZsucc_inj_wd : forall n1 n2 : NZ, S n1 == S n2 <-> n1 == n2. +Proof. +intros; split. +apply NZsucc_inj. +apply NZsucc_wd. +Qed. + +Theorem NZsucc_inj_wd_neg : forall n m : NZ, S n ~= S m <-> n ~= m. +Proof. +intros; now rewrite NZsucc_inj_wd. +Qed. + +(* We cannot prove that the predecessor is injective, nor that it is +left-inverse to the successor at this point *) + +Section CentralInduction. + +Variable A : predicate NZ. + +Hypothesis A_wd : predicate_wd NZeq A. + +Add Morphism A with signature NZeq ==> iff as A_morph. +Proof. apply A_wd. Qed. + +Theorem NZcentral_induction : + forall z : NZ, A z -> + (forall n : NZ, A n <-> A (S n)) -> + forall n : NZ, A n. +Proof. +intros z Base Step; revert Base; pattern z; apply NZinduction. +solve_predicate_wd. +intro; now apply NZinduction. +intro; pose proof (Step n); tauto. +Qed. + +End CentralInduction. + +Tactic Notation "NZinduct" ident(n) := + induction_maker n ltac:(apply NZinduction). + +Tactic Notation "NZinduct" ident(n) constr(u) := + induction_maker n ltac:(apply NZcentral_induction with (z := u)). + +End NZBasePropFunct. + diff --git a/theories/Numbers/NatInt/NZMul.v b/theories/Numbers/NatInt/NZMul.v new file mode 100644 index 00000000..fda8b7a3 --- /dev/null +++ b/theories/Numbers/NatInt/NZMul.v @@ -0,0 +1,80 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZMul.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import NZAxioms. +Require Import NZAdd. + +Module NZMulPropFunct (Import NZAxiomsMod : NZAxiomsSig). +Module Export NZAddPropMod := NZAddPropFunct NZAxiomsMod. +Open Local Scope NatIntScope. + +Theorem NZmul_0_r : forall n : NZ, n * 0 == 0. +Proof. +NZinduct n. +now rewrite NZmul_0_l. +intro. rewrite NZmul_succ_l. now rewrite NZadd_0_r. +Qed. + +Theorem NZmul_succ_r : forall n m : NZ, n * (S m) == n * m + n. +Proof. +intros n m; NZinduct n. +do 2 rewrite NZmul_0_l; now rewrite NZadd_0_l. +intro n. do 2 rewrite NZmul_succ_l. do 2 rewrite NZadd_succ_r. +rewrite NZsucc_inj_wd. rewrite <- (NZadd_assoc (n * m) m n). +rewrite (NZadd_comm m n). rewrite NZadd_assoc. +now rewrite NZadd_cancel_r. +Qed. + +Theorem NZmul_comm : forall n m : NZ, n * m == m * n. +Proof. +intros n m; NZinduct n. +rewrite NZmul_0_l; now rewrite NZmul_0_r. +intro. rewrite NZmul_succ_l; rewrite NZmul_succ_r. now rewrite NZadd_cancel_r. +Qed. + +Theorem NZmul_add_distr_r : forall n m p : NZ, (n + m) * p == n * p + m * p. +Proof. +intros n m p; NZinduct n. +rewrite NZmul_0_l. now do 2 rewrite NZadd_0_l. +intro n. rewrite NZadd_succ_l. do 2 rewrite NZmul_succ_l. +rewrite <- (NZadd_assoc (n * p) p (m * p)). +rewrite (NZadd_comm p (m * p)). rewrite (NZadd_assoc (n * p) (m * p) p). +now rewrite NZadd_cancel_r. +Qed. + +Theorem NZmul_add_distr_l : forall n m p : NZ, n * (m + p) == n * m + n * p. +Proof. +intros n m p. +rewrite (NZmul_comm n (m + p)). rewrite (NZmul_comm n m). +rewrite (NZmul_comm n p). apply NZmul_add_distr_r. +Qed. + +Theorem NZmul_assoc : forall n m p : NZ, n * (m * p) == (n * m) * p. +Proof. +intros n m p; NZinduct n. +now do 3 rewrite NZmul_0_l. +intro n. do 2 rewrite NZmul_succ_l. rewrite NZmul_add_distr_r. +now rewrite NZadd_cancel_r. +Qed. + +Theorem NZmul_1_l : forall n : NZ, 1 * n == n. +Proof. +intro n. rewrite NZmul_succ_l; rewrite NZmul_0_l. now rewrite NZadd_0_l. +Qed. + +Theorem NZmul_1_r : forall n : NZ, n * 1 == n. +Proof. +intro n; rewrite NZmul_comm; apply NZmul_1_l. +Qed. + +End NZMulPropFunct. + diff --git a/theories/Numbers/NatInt/NZMulOrder.v b/theories/Numbers/NatInt/NZMulOrder.v new file mode 100644 index 00000000..c707bf73 --- /dev/null +++ b/theories/Numbers/NatInt/NZMulOrder.v @@ -0,0 +1,310 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZMulOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import NZAxioms. +Require Import NZAddOrder. + +Module NZMulOrderPropFunct (Import NZOrdAxiomsMod : NZOrdAxiomsSig). +Module Export NZAddOrderPropMod := NZAddOrderPropFunct NZOrdAxiomsMod. +Open Local Scope NatIntScope. + +Theorem NZmul_lt_pred : + forall p q n m : NZ, S p == q -> (p * n < p * m <-> q * n + m < q * m + n). +Proof. +intros p q n m H. rewrite <- H. do 2 rewrite NZmul_succ_l. +rewrite <- (NZadd_assoc (p * n) n m). +rewrite <- (NZadd_assoc (p * m) m n). +rewrite (NZadd_comm n m). now rewrite <- NZadd_lt_mono_r. +Qed. + +Theorem NZmul_lt_mono_pos_l : forall p n m : NZ, 0 < p -> (n < m <-> p * n < p * m). +Proof. +NZord_induct p. +intros n m H; false_hyp H NZlt_irrefl. +intros p H IH n m H1. do 2 rewrite NZmul_succ_l. +le_elim H. assert (LR : forall n m : NZ, n < m -> p * n + n < p * m + m). +intros n1 m1 H2. apply NZadd_lt_mono; [now apply -> IH | assumption]. +split; [apply LR |]. intro H2. apply -> NZlt_dne; intro H3. +apply <- NZle_ngt in H3. le_elim H3. +apply NZlt_asymm in H2. apply H2. now apply LR. +rewrite H3 in H2; false_hyp H2 NZlt_irrefl. +rewrite <- H; do 2 rewrite NZmul_0_l; now do 2 rewrite NZadd_0_l. +intros p H1 _ n m H2. apply NZlt_asymm in H1. false_hyp H2 H1. +Qed. + +Theorem NZmul_lt_mono_pos_r : forall p n m : NZ, 0 < p -> (n < m <-> n * p < m * p). +Proof. +intros p n m. +rewrite (NZmul_comm n p); rewrite (NZmul_comm m p). now apply NZmul_lt_mono_pos_l. +Qed. + +Theorem NZmul_lt_mono_neg_l : forall p n m : NZ, p < 0 -> (n < m <-> p * m < p * n). +Proof. +NZord_induct p. +intros n m H; false_hyp H NZlt_irrefl. +intros p H1 _ n m H2. apply NZlt_succ_l in H2. apply <- NZnle_gt in H2. false_hyp H1 H2. +intros p H IH n m H1. apply <- NZle_succ_l in H. +le_elim H. assert (LR : forall n m : NZ, n < m -> p * m < p * n). +intros n1 m1 H2. apply (NZle_lt_add_lt n1 m1). +now apply NZlt_le_incl. do 2 rewrite <- NZmul_succ_l. now apply -> IH. +split; [apply LR |]. intro H2. apply -> NZlt_dne; intro H3. +apply <- NZle_ngt in H3. le_elim H3. +apply NZlt_asymm in H2. apply H2. now apply LR. +rewrite H3 in H2; false_hyp H2 NZlt_irrefl. +rewrite (NZmul_lt_pred p (S p)) by reflexivity. +rewrite H; do 2 rewrite NZmul_0_l; now do 2 rewrite NZadd_0_l. +Qed. + +Theorem NZmul_lt_mono_neg_r : forall p n m : NZ, p < 0 -> (n < m <-> m * p < n * p). +Proof. +intros p n m. +rewrite (NZmul_comm n p); rewrite (NZmul_comm m p). now apply NZmul_lt_mono_neg_l. +Qed. + +Theorem NZmul_le_mono_nonneg_l : forall n m p : NZ, 0 <= p -> n <= m -> p * n <= p * m. +Proof. +intros n m p H1 H2. le_elim H1. +le_elim H2. apply NZlt_le_incl. now apply -> NZmul_lt_mono_pos_l. +apply NZeq_le_incl; now rewrite H2. +apply NZeq_le_incl; rewrite <- H1; now do 2 rewrite NZmul_0_l. +Qed. + +Theorem NZmul_le_mono_nonpos_l : forall n m p : NZ, p <= 0 -> n <= m -> p * m <= p * n. +Proof. +intros n m p H1 H2. le_elim H1. +le_elim H2. apply NZlt_le_incl. now apply -> NZmul_lt_mono_neg_l. +apply NZeq_le_incl; now rewrite H2. +apply NZeq_le_incl; rewrite H1; now do 2 rewrite NZmul_0_l. +Qed. + +Theorem NZmul_le_mono_nonneg_r : forall n m p : NZ, 0 <= p -> n <= m -> n * p <= m * p. +Proof. +intros n m p H1 H2; rewrite (NZmul_comm n p); rewrite (NZmul_comm m p); +now apply NZmul_le_mono_nonneg_l. +Qed. + +Theorem NZmul_le_mono_nonpos_r : forall n m p : NZ, p <= 0 -> n <= m -> m * p <= n * p. +Proof. +intros n m p H1 H2; rewrite (NZmul_comm n p); rewrite (NZmul_comm m p); +now apply NZmul_le_mono_nonpos_l. +Qed. + +Theorem NZmul_cancel_l : forall n m p : NZ, p ~= 0 -> (p * n == p * m <-> n == m). +Proof. +intros n m p H; split; intro H1. +destruct (NZlt_trichotomy p 0) as [H2 | [H2 | H2]]. +apply -> NZeq_dne; intro H3. apply -> NZlt_gt_cases in H3. destruct H3 as [H3 | H3]. +assert (H4 : p * m < p * n); [now apply -> NZmul_lt_mono_neg_l |]. +rewrite H1 in H4; false_hyp H4 NZlt_irrefl. +assert (H4 : p * n < p * m); [now apply -> NZmul_lt_mono_neg_l |]. +rewrite H1 in H4; false_hyp H4 NZlt_irrefl. +false_hyp H2 H. +apply -> NZeq_dne; intro H3. apply -> NZlt_gt_cases in H3. destruct H3 as [H3 | H3]. +assert (H4 : p * n < p * m) by (now apply -> NZmul_lt_mono_pos_l). +rewrite H1 in H4; false_hyp H4 NZlt_irrefl. +assert (H4 : p * m < p * n) by (now apply -> NZmul_lt_mono_pos_l). +rewrite H1 in H4; false_hyp H4 NZlt_irrefl. +now rewrite H1. +Qed. + +Theorem NZmul_cancel_r : forall n m p : NZ, p ~= 0 -> (n * p == m * p <-> n == m). +Proof. +intros n m p. rewrite (NZmul_comm n p), (NZmul_comm m p); apply NZmul_cancel_l. +Qed. + +Theorem NZmul_id_l : forall n m : NZ, m ~= 0 -> (n * m == m <-> n == 1). +Proof. +intros n m H. +stepl (n * m == 1 * m) by now rewrite NZmul_1_l. now apply NZmul_cancel_r. +Qed. + +Theorem NZmul_id_r : forall n m : NZ, n ~= 0 -> (n * m == n <-> m == 1). +Proof. +intros n m; rewrite NZmul_comm; apply NZmul_id_l. +Qed. + +Theorem NZmul_le_mono_pos_l : forall n m p : NZ, 0 < p -> (n <= m <-> p * n <= p * m). +Proof. +intros n m p H; do 2 rewrite NZlt_eq_cases. +rewrite (NZmul_lt_mono_pos_l p n m) by assumption. +now rewrite -> (NZmul_cancel_l n m p) by +(intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl). +Qed. + +Theorem NZmul_le_mono_pos_r : forall n m p : NZ, 0 < p -> (n <= m <-> n * p <= m * p). +Proof. +intros n m p. rewrite (NZmul_comm n p); rewrite (NZmul_comm m p); +apply NZmul_le_mono_pos_l. +Qed. + +Theorem NZmul_le_mono_neg_l : forall n m p : NZ, p < 0 -> (n <= m <-> p * m <= p * n). +Proof. +intros n m p H; do 2 rewrite NZlt_eq_cases. +rewrite (NZmul_lt_mono_neg_l p n m); [| assumption]. +rewrite -> (NZmul_cancel_l m n p) by (intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl). +now setoid_replace (n == m) with (m == n) using relation iff by (split; now intro). +Qed. + +Theorem NZmul_le_mono_neg_r : forall n m p : NZ, p < 0 -> (n <= m <-> m * p <= n * p). +Proof. +intros n m p. rewrite (NZmul_comm n p); rewrite (NZmul_comm m p); +apply NZmul_le_mono_neg_l. +Qed. + +Theorem NZmul_lt_mono_nonneg : + forall n m p q : NZ, 0 <= n -> n < m -> 0 <= p -> p < q -> n * p < m * q. +Proof. +intros n m p q H1 H2 H3 H4. +apply NZle_lt_trans with (m * p). +apply NZmul_le_mono_nonneg_r; [assumption | now apply NZlt_le_incl]. +apply -> NZmul_lt_mono_pos_l; [assumption | now apply NZle_lt_trans with n]. +Qed. + +(* There are still many variants of the theorem above. One can assume 0 < n +or 0 < p or n <= m or p <= q. *) + +Theorem NZmul_le_mono_nonneg : + forall n m p q : NZ, 0 <= n -> n <= m -> 0 <= p -> p <= q -> n * p <= m * q. +Proof. +intros n m p q H1 H2 H3 H4. +le_elim H2; le_elim H4. +apply NZlt_le_incl; now apply NZmul_lt_mono_nonneg. +rewrite <- H4; apply NZmul_le_mono_nonneg_r; [assumption | now apply NZlt_le_incl]. +rewrite <- H2; apply NZmul_le_mono_nonneg_l; [assumption | now apply NZlt_le_incl]. +rewrite H2; rewrite H4; now apply NZeq_le_incl. +Qed. + +Theorem NZmul_pos_pos : forall n m : NZ, 0 < n -> 0 < m -> 0 < n * m. +Proof. +intros n m H1 H2. +rewrite <- (NZmul_0_l m). now apply -> NZmul_lt_mono_pos_r. +Qed. + +Theorem NZmul_neg_neg : forall n m : NZ, n < 0 -> m < 0 -> 0 < n * m. +Proof. +intros n m H1 H2. +rewrite <- (NZmul_0_l m). now apply -> NZmul_lt_mono_neg_r. +Qed. + +Theorem NZmul_pos_neg : forall n m : NZ, 0 < n -> m < 0 -> n * m < 0. +Proof. +intros n m H1 H2. +rewrite <- (NZmul_0_l m). now apply -> NZmul_lt_mono_neg_r. +Qed. + +Theorem NZmul_neg_pos : forall n m : NZ, n < 0 -> 0 < m -> n * m < 0. +Proof. +intros; rewrite NZmul_comm; now apply NZmul_pos_neg. +Qed. + +Theorem NZlt_1_mul_pos : forall n m : NZ, 1 < n -> 0 < m -> 1 < n * m. +Proof. +intros n m H1 H2. apply -> (NZmul_lt_mono_pos_r m) in H1. +rewrite NZmul_1_l in H1. now apply NZlt_1_l with m. +assumption. +Qed. + +Theorem NZeq_mul_0 : forall n m : NZ, n * m == 0 <-> n == 0 \/ m == 0. +Proof. +intros n m; split. +intro H; destruct (NZlt_trichotomy n 0) as [H1 | [H1 | H1]]; +destruct (NZlt_trichotomy m 0) as [H2 | [H2 | H2]]; +try (now right); try (now left). +elimtype False; now apply (NZlt_neq 0 (n * m)); [apply NZmul_neg_neg |]. +elimtype False; now apply (NZlt_neq (n * m) 0); [apply NZmul_neg_pos |]. +elimtype False; now apply (NZlt_neq (n * m) 0); [apply NZmul_pos_neg |]. +elimtype False; now apply (NZlt_neq 0 (n * m)); [apply NZmul_pos_pos |]. +intros [H | H]. now rewrite H, NZmul_0_l. now rewrite H, NZmul_0_r. +Qed. + +Theorem NZneq_mul_0 : forall n m : NZ, n ~= 0 /\ m ~= 0 <-> n * m ~= 0. +Proof. +intros n m; split; intro H. +intro H1; apply -> NZeq_mul_0 in H1. tauto. +split; intro H1; rewrite H1 in H; +(rewrite NZmul_0_l in H || rewrite NZmul_0_r in H); now apply H. +Qed. + +Theorem NZeq_square_0 : forall n : NZ, n * n == 0 <-> n == 0. +Proof. +intro n; rewrite NZeq_mul_0; tauto. +Qed. + +Theorem NZeq_mul_0_l : forall n m : NZ, n * m == 0 -> m ~= 0 -> n == 0. +Proof. +intros n m H1 H2. apply -> NZeq_mul_0 in H1. destruct H1 as [H1 | H1]. +assumption. false_hyp H1 H2. +Qed. + +Theorem NZeq_mul_0_r : forall n m : NZ, n * m == 0 -> n ~= 0 -> m == 0. +Proof. +intros n m H1 H2; apply -> NZeq_mul_0 in H1. destruct H1 as [H1 | H1]. +false_hyp H1 H2. assumption. +Qed. + +Theorem NZlt_0_mul : forall n m : NZ, 0 < n * m <-> (0 < n /\ 0 < m) \/ (m < 0 /\ n < 0). +Proof. +intros n m; split; [intro H | intros [[H1 H2] | [H1 H2]]]. +destruct (NZlt_trichotomy n 0) as [H1 | [H1 | H1]]; +[| rewrite H1 in H; rewrite NZmul_0_l in H; false_hyp H NZlt_irrefl |]; +(destruct (NZlt_trichotomy m 0) as [H2 | [H2 | H2]]; +[| rewrite H2 in H; rewrite NZmul_0_r in H; false_hyp H NZlt_irrefl |]); +try (left; now split); try (right; now split). +assert (H3 : n * m < 0) by now apply NZmul_neg_pos. +elimtype False; now apply (NZlt_asymm (n * m) 0). +assert (H3 : n * m < 0) by now apply NZmul_pos_neg. +elimtype False; now apply (NZlt_asymm (n * m) 0). +now apply NZmul_pos_pos. now apply NZmul_neg_neg. +Qed. + +Theorem NZsquare_lt_mono_nonneg : forall n m : NZ, 0 <= n -> n < m -> n * n < m * m. +Proof. +intros n m H1 H2. now apply NZmul_lt_mono_nonneg. +Qed. + +Theorem NZsquare_le_mono_nonneg : forall n m : NZ, 0 <= n -> n <= m -> n * n <= m * m. +Proof. +intros n m H1 H2. now apply NZmul_le_mono_nonneg. +Qed. + +(* The converse theorems require nonnegativity (or nonpositivity) of the +other variable *) + +Theorem NZsquare_lt_simpl_nonneg : forall n m : NZ, 0 <= m -> n * n < m * m -> n < m. +Proof. +intros n m H1 H2. destruct (NZlt_ge_cases n 0). +now apply NZlt_le_trans with 0. +destruct (NZlt_ge_cases n m). +assumption. assert (F : m * m <= n * n) by now apply NZsquare_le_mono_nonneg. +apply -> NZle_ngt in F. false_hyp H2 F. +Qed. + +Theorem NZsquare_le_simpl_nonneg : forall n m : NZ, 0 <= m -> n * n <= m * m -> n <= m. +Proof. +intros n m H1 H2. destruct (NZlt_ge_cases n 0). +apply NZlt_le_incl; now apply NZlt_le_trans with 0. +destruct (NZle_gt_cases n m). +assumption. assert (F : m * m < n * n) by now apply NZsquare_lt_mono_nonneg. +apply -> NZlt_nge in F. false_hyp H2 F. +Qed. + +Theorem NZmul_2_mono_l : forall n m : NZ, n < m -> 1 + (1 + 1) * n < (1 + 1) * m. +Proof. +intros n m H. apply <- NZle_succ_l in H. +apply -> (NZmul_le_mono_pos_l (S n) m (1 + 1)) in H. +repeat rewrite NZmul_add_distr_r in *; repeat rewrite NZmul_1_l in *. +repeat rewrite NZadd_succ_r in *. repeat rewrite NZadd_succ_l in *. rewrite NZadd_0_l. +now apply -> NZle_succ_l. +apply NZadd_pos_pos; now apply NZlt_succ_diag_r. +Qed. + +End NZMulOrderPropFunct. diff --git a/theories/Numbers/NatInt/NZOrder.v b/theories/Numbers/NatInt/NZOrder.v new file mode 100644 index 00000000..15004824 --- /dev/null +++ b/theories/Numbers/NatInt/NZOrder.v @@ -0,0 +1,666 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NZOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import NZAxioms. +Require Import NZMul. +Require Import Decidable. + +Module NZOrderPropFunct (Import NZOrdAxiomsMod : NZOrdAxiomsSig). +Module Export NZMulPropMod := NZMulPropFunct NZAxiomsMod. +Open Local Scope NatIntScope. + +Ltac le_elim H := rewrite NZlt_eq_cases in H; destruct H as [H | H]. + +Theorem NZlt_le_incl : forall n m : NZ, n < m -> n <= m. +Proof. +intros; apply <- NZlt_eq_cases; now left. +Qed. + +Theorem NZeq_le_incl : forall n m : NZ, n == m -> n <= m. +Proof. +intros; apply <- NZlt_eq_cases; now right. +Qed. + +Lemma NZlt_stepl : forall x y z : NZ, x < y -> x == z -> z < y. +Proof. +intros x y z H1 H2; now rewrite <- H2. +Qed. + +Lemma NZlt_stepr : forall x y z : NZ, x < y -> y == z -> x < z. +Proof. +intros x y z H1 H2; now rewrite <- H2. +Qed. + +Lemma NZle_stepl : forall x y z : NZ, x <= y -> x == z -> z <= y. +Proof. +intros x y z H1 H2; now rewrite <- H2. +Qed. + +Lemma NZle_stepr : forall x y z : NZ, x <= y -> y == z -> x <= z. +Proof. +intros x y z H1 H2; now rewrite <- H2. +Qed. + +Declare Left Step NZlt_stepl. +Declare Right Step NZlt_stepr. +Declare Left Step NZle_stepl. +Declare Right Step NZle_stepr. + +Theorem NZlt_neq : forall n m : NZ, n < m -> n ~= m. +Proof. +intros n m H1 H2; rewrite H2 in H1; false_hyp H1 NZlt_irrefl. +Qed. + +Theorem NZle_neq : forall n m : NZ, n < m <-> n <= m /\ n ~= m. +Proof. +intros n m; split; [intro H | intros [H1 H2]]. +split. now apply NZlt_le_incl. now apply NZlt_neq. +le_elim H1. assumption. false_hyp H1 H2. +Qed. + +Theorem NZle_refl : forall n : NZ, n <= n. +Proof. +intro; now apply NZeq_le_incl. +Qed. + +Theorem NZlt_succ_diag_r : forall n : NZ, n < S n. +Proof. +intro n. rewrite NZlt_succ_r. now apply NZeq_le_incl. +Qed. + +Theorem NZle_succ_diag_r : forall n : NZ, n <= S n. +Proof. +intro; apply NZlt_le_incl; apply NZlt_succ_diag_r. +Qed. + +Theorem NZlt_0_1 : 0 < 1. +Proof. +apply NZlt_succ_diag_r. +Qed. + +Theorem NZle_0_1 : 0 <= 1. +Proof. +apply NZle_succ_diag_r. +Qed. + +Theorem NZlt_lt_succ_r : forall n m : NZ, n < m -> n < S m. +Proof. +intros. rewrite NZlt_succ_r. now apply NZlt_le_incl. +Qed. + +Theorem NZle_le_succ_r : forall n m : NZ, n <= m -> n <= S m. +Proof. +intros n m H. rewrite <- NZlt_succ_r in H. now apply NZlt_le_incl. +Qed. + +Theorem NZle_succ_r : forall n m : NZ, n <= S m <-> n <= m \/ n == S m. +Proof. +intros n m; rewrite NZlt_eq_cases. now rewrite NZlt_succ_r. +Qed. + +(* The following theorem is a special case of neq_succ_iter_l below, +but we prove it separately *) + +Theorem NZneq_succ_diag_l : forall n : NZ, S n ~= n. +Proof. +intros n H. pose proof (NZlt_succ_diag_r n) as H1. rewrite H in H1. +false_hyp H1 NZlt_irrefl. +Qed. + +Theorem NZneq_succ_diag_r : forall n : NZ, n ~= S n. +Proof. +intro n; apply NZneq_symm; apply NZneq_succ_diag_l. +Qed. + +Theorem NZnlt_succ_diag_l : forall n : NZ, ~ S n < n. +Proof. +intros n H; apply NZlt_lt_succ_r in H. false_hyp H NZlt_irrefl. +Qed. + +Theorem NZnle_succ_diag_l : forall n : NZ, ~ S n <= n. +Proof. +intros n H; le_elim H. +false_hyp H NZnlt_succ_diag_l. false_hyp H NZneq_succ_diag_l. +Qed. + +Theorem NZle_succ_l : forall n m : NZ, S n <= m <-> n < m. +Proof. +intro n; NZinduct m n. +setoid_replace (n < n) with False using relation iff by + (apply -> neg_false; apply NZlt_irrefl). +now setoid_replace (S n <= n) with False using relation iff by + (apply -> neg_false; apply NZnle_succ_diag_l). +intro m. rewrite NZlt_succ_r. rewrite NZle_succ_r. +rewrite NZsucc_inj_wd. +rewrite (NZlt_eq_cases n m). +rewrite or_cancel_r. +reflexivity. +intros H1 H2; rewrite H2 in H1; false_hyp H1 NZnle_succ_diag_l. +apply NZlt_neq. +Qed. + +Theorem NZlt_succ_l : forall n m : NZ, S n < m -> n < m. +Proof. +intros n m H; apply -> NZle_succ_l; now apply NZlt_le_incl. +Qed. + +Theorem NZsucc_lt_mono : forall n m : NZ, n < m <-> S n < S m. +Proof. +intros n m. rewrite <- NZle_succ_l. symmetry. apply NZlt_succ_r. +Qed. + +Theorem NZsucc_le_mono : forall n m : NZ, n <= m <-> S n <= S m. +Proof. +intros n m. do 2 rewrite NZlt_eq_cases. +rewrite <- NZsucc_lt_mono; now rewrite NZsucc_inj_wd. +Qed. + +Theorem NZlt_asymm : forall n m, n < m -> ~ m < n. +Proof. +intros n m; NZinduct n m. +intros H _; false_hyp H NZlt_irrefl. +intro n; split; intros H H1 H2. +apply NZlt_succ_l in H1. apply -> NZlt_succ_r in H2. le_elim H2. +now apply H. rewrite H2 in H1; false_hyp H1 NZlt_irrefl. +apply NZlt_lt_succ_r in H2. apply <- NZle_succ_l in H1. le_elim H1. +now apply H. rewrite H1 in H2; false_hyp H2 NZlt_irrefl. +Qed. + +Theorem NZlt_trans : forall n m p : NZ, n < m -> m < p -> n < p. +Proof. +intros n m p; NZinduct p m. +intros _ H; false_hyp H NZlt_irrefl. +intro p. do 2 rewrite NZlt_succ_r. +split; intros H H1 H2. +apply NZlt_le_incl; le_elim H2; [now apply H | now rewrite H2 in H1]. +assert (n <= p) as H3. apply H. assumption. now apply NZlt_le_incl. +le_elim H3. assumption. rewrite <- H3 in H2. +elimtype False; now apply (NZlt_asymm n m). +Qed. + +Theorem NZle_trans : forall n m p : NZ, n <= m -> m <= p -> n <= p. +Proof. +intros n m p H1 H2; le_elim H1. +le_elim H2. apply NZlt_le_incl; now apply NZlt_trans with (m := m). +apply NZlt_le_incl; now rewrite <- H2. now rewrite H1. +Qed. + +Theorem NZle_lt_trans : forall n m p : NZ, n <= m -> m < p -> n < p. +Proof. +intros n m p H1 H2; le_elim H1. +now apply NZlt_trans with (m := m). now rewrite H1. +Qed. + +Theorem NZlt_le_trans : forall n m p : NZ, n < m -> m <= p -> n < p. +Proof. +intros n m p H1 H2; le_elim H2. +now apply NZlt_trans with (m := m). now rewrite <- H2. +Qed. + +Theorem NZle_antisymm : forall n m : NZ, n <= m -> m <= n -> n == m. +Proof. +intros n m H1 H2; now (le_elim H1; le_elim H2); +[elimtype False; apply (NZlt_asymm n m) | | |]. +Qed. + +Theorem NZlt_1_l : forall n m : NZ, 0 < n -> n < m -> 1 < m. +Proof. +intros n m H1 H2. apply <- NZle_succ_l in H1. now apply NZle_lt_trans with n. +Qed. + +(** Trichotomy, decidability, and double negation elimination *) + +Theorem NZlt_trichotomy : forall n m : NZ, n < m \/ n == m \/ m < n. +Proof. +intros n m; NZinduct n m. +right; now left. +intro n; rewrite NZlt_succ_r. stepr ((S n < m \/ S n == m) \/ m <= n) by tauto. +rewrite <- (NZlt_eq_cases (S n) m). +setoid_replace (n == m) with (m == n) using relation iff by now split. +stepl (n < m \/ m < n \/ m == n) by tauto. rewrite <- NZlt_eq_cases. +apply or_iff_compat_r. symmetry; apply NZle_succ_l. +Qed. + +(* Decidability of equality, even though true in each finite ring, does not +have a uniform proof. Otherwise, the proof for two fixed numbers would +reduce to a normal form that will say if the numbers are equal or not, +which cannot be true in all finite rings. Therefore, we prove decidability +in the presence of order. *) + +Theorem NZeq_dec : forall n m : NZ, decidable (n == m). +Proof. +intros n m; destruct (NZlt_trichotomy n m) as [H | [H | H]]. +right; intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl. +now left. +right; intro H1; rewrite H1 in H; false_hyp H NZlt_irrefl. +Qed. + +(* DNE stands for double-negation elimination *) + +Theorem NZeq_dne : forall n m, ~ ~ n == m <-> n == m. +Proof. +intros n m; split; intro H. +destruct (NZeq_dec n m) as [H1 | H1]. +assumption. false_hyp H1 H. +intro H1; now apply H1. +Qed. + +Theorem NZlt_gt_cases : forall n m : NZ, n ~= m <-> n < m \/ n > m. +Proof. +intros n m; split. +pose proof (NZlt_trichotomy n m); tauto. +intros H H1; destruct H as [H | H]; rewrite H1 in H; false_hyp H NZlt_irrefl. +Qed. + +Theorem NZle_gt_cases : forall n m : NZ, n <= m \/ n > m. +Proof. +intros n m; destruct (NZlt_trichotomy n m) as [H | [H | H]]. +left; now apply NZlt_le_incl. left; now apply NZeq_le_incl. now right. +Qed. + +(* The following theorem is cleary redundant, but helps not to +remember whether one has to say le_gt_cases or lt_ge_cases *) + +Theorem NZlt_ge_cases : forall n m : NZ, n < m \/ n >= m. +Proof. +intros n m; destruct (NZle_gt_cases m n); try (now left); try (now right). +Qed. + +Theorem NZle_ge_cases : forall n m : NZ, n <= m \/ n >= m. +Proof. +intros n m; destruct (NZle_gt_cases n m) as [H | H]. +now left. right; now apply NZlt_le_incl. +Qed. + +Theorem NZle_ngt : forall n m : NZ, n <= m <-> ~ n > m. +Proof. +intros n m. split; intro H; [intro H1 |]. +eapply NZle_lt_trans in H; [| eassumption ..]. false_hyp H NZlt_irrefl. +destruct (NZle_gt_cases n m) as [H1 | H1]. +assumption. false_hyp H1 H. +Qed. + +(* Redundant but useful *) + +Theorem NZnlt_ge : forall n m : NZ, ~ n < m <-> n >= m. +Proof. +intros n m; symmetry; apply NZle_ngt. +Qed. + +Theorem NZlt_dec : forall n m : NZ, decidable (n < m). +Proof. +intros n m; destruct (NZle_gt_cases m n); +[right; now apply -> NZle_ngt | now left]. +Qed. + +Theorem NZlt_dne : forall n m, ~ ~ n < m <-> n < m. +Proof. +intros n m; split; intro H; +[destruct (NZlt_dec n m) as [H1 | H1]; [assumption | false_hyp H1 H] | +intro H1; false_hyp H H1]. +Qed. + +Theorem NZnle_gt : forall n m : NZ, ~ n <= m <-> n > m. +Proof. +intros n m. rewrite NZle_ngt. apply NZlt_dne. +Qed. + +(* Redundant but useful *) + +Theorem NZlt_nge : forall n m : NZ, n < m <-> ~ n >= m. +Proof. +intros n m; symmetry; apply NZnle_gt. +Qed. + +Theorem NZle_dec : forall n m : NZ, decidable (n <= m). +Proof. +intros n m; destruct (NZle_gt_cases n m); +[now left | right; now apply <- NZnle_gt]. +Qed. + +Theorem NZle_dne : forall n m : NZ, ~ ~ n <= m <-> n <= m. +Proof. +intros n m; split; intro H; +[destruct (NZle_dec n m) as [H1 | H1]; [assumption | false_hyp H1 H] | +intro H1; false_hyp H H1]. +Qed. + +Theorem NZnlt_succ_r : forall n m : NZ, ~ m < S n <-> n < m. +Proof. +intros n m; rewrite NZlt_succ_r; apply NZnle_gt. +Qed. + +(* The difference between integers and natural numbers is that for +every integer there is a predecessor, which is not true for natural +numbers. However, for both classes, every number that is bigger than +some other number has a predecessor. The proof of this fact by regular +induction does not go through, so we need to use strong +(course-of-value) induction. *) + +Lemma NZlt_exists_pred_strong : + forall z n m : NZ, z < m -> m <= n -> exists k : NZ, m == S k /\ z <= k. +Proof. +intro z; NZinduct n z. +intros m H1 H2; apply <- NZnle_gt in H1; false_hyp H2 H1. +intro n; split; intros IH m H1 H2. +apply -> NZle_succ_r in H2; destruct H2 as [H2 | H2]. +now apply IH. exists n. now split; [| rewrite <- NZlt_succ_r; rewrite <- H2]. +apply IH. assumption. now apply NZle_le_succ_r. +Qed. + +Theorem NZlt_exists_pred : + forall z n : NZ, z < n -> exists k : NZ, n == S k /\ z <= k. +Proof. +intros z n H; apply NZlt_exists_pred_strong with (z := z) (n := n). +assumption. apply NZle_refl. +Qed. + +(** A corollary of having an order is that NZ is infinite *) + +(* This section about infinity of NZ relies on the type nat and can be +safely removed *) + +Definition NZsucc_iter (n : nat) (m : NZ) := + nat_rect (fun _ => NZ) m (fun _ l => S l) n. + +Theorem NZlt_succ_iter_r : + forall (n : nat) (m : NZ), m < NZsucc_iter (Datatypes.S n) m. +Proof. +intros n m; induction n as [| n IH]; simpl in *. +apply NZlt_succ_diag_r. now apply NZlt_lt_succ_r. +Qed. + +Theorem NZneq_succ_iter_l : + forall (n : nat) (m : NZ), NZsucc_iter (Datatypes.S n) m ~= m. +Proof. +intros n m H. pose proof (NZlt_succ_iter_r n m) as H1. rewrite H in H1. +false_hyp H1 NZlt_irrefl. +Qed. + +(* End of the section about the infinity of NZ *) + +(** Stronger variant of induction with assumptions n >= 0 (n < 0) +in the induction step *) + +Section Induction. + +Variable A : NZ -> Prop. +Hypothesis A_wd : predicate_wd NZeq A. + +Add Morphism A with signature NZeq ==> iff as A_morph. +Proof. apply A_wd. Qed. + +Section Center. + +Variable z : NZ. (* A z is the basis of induction *) + +Section RightInduction. + +Let A' (n : NZ) := forall m : NZ, z <= m -> m < n -> A m. +Let right_step := forall n : NZ, z <= n -> A n -> A (S n). +Let right_step' := forall n : NZ, z <= n -> A' n -> A n. +Let right_step'' := forall n : NZ, A' n <-> A' (S n). + +Lemma NZrs_rs' : A z -> right_step -> right_step'. +Proof. +intros Az RS n H1 H2. +le_elim H1. apply NZlt_exists_pred in H1. destruct H1 as [k [H3 H4]]. +rewrite H3. apply RS; [assumption | apply H2; [assumption | rewrite H3; apply NZlt_succ_diag_r]]. +rewrite <- H1; apply Az. +Qed. + +Lemma NZrs'_rs'' : right_step' -> right_step''. +Proof. +intros RS' n; split; intros H1 m H2 H3. +apply -> NZlt_succ_r in H3; le_elim H3; +[now apply H1 | rewrite H3 in *; now apply RS']. +apply H1; [assumption | now apply NZlt_lt_succ_r]. +Qed. + +Lemma NZrbase : A' z. +Proof. +intros m H1 H2. apply -> NZle_ngt in H1. false_hyp H2 H1. +Qed. + +Lemma NZA'A_right : (forall n : NZ, A' n) -> forall n : NZ, z <= n -> A n. +Proof. +intros H1 n H2. apply H1 with (n := S n); [assumption | apply NZlt_succ_diag_r]. +Qed. + +Theorem NZstrong_right_induction: right_step' -> forall n : NZ, z <= n -> A n. +Proof. +intro RS'; apply NZA'A_right; unfold A'; NZinduct n z; +[apply NZrbase | apply NZrs'_rs''; apply RS']. +Qed. + +Theorem NZright_induction : A z -> right_step -> forall n : NZ, z <= n -> A n. +Proof. +intros Az RS; apply NZstrong_right_induction; now apply NZrs_rs'. +Qed. + +Theorem NZright_induction' : + (forall n : NZ, n <= z -> A n) -> right_step -> forall n : NZ, A n. +Proof. +intros L R n. +destruct (NZlt_trichotomy n z) as [H | [H | H]]. +apply L; now apply NZlt_le_incl. +apply L; now apply NZeq_le_incl. +apply NZright_induction. apply L; now apply NZeq_le_incl. assumption. now apply NZlt_le_incl. +Qed. + +Theorem NZstrong_right_induction' : + (forall n : NZ, n <= z -> A n) -> right_step' -> forall n : NZ, A n. +Proof. +intros L R n. +destruct (NZlt_trichotomy n z) as [H | [H | H]]. +apply L; now apply NZlt_le_incl. +apply L; now apply NZeq_le_incl. +apply NZstrong_right_induction. assumption. now apply NZlt_le_incl. +Qed. + +End RightInduction. + +Section LeftInduction. + +Let A' (n : NZ) := forall m : NZ, m <= z -> n <= m -> A m. +Let left_step := forall n : NZ, n < z -> A (S n) -> A n. +Let left_step' := forall n : NZ, n <= z -> A' (S n) -> A n. +Let left_step'' := forall n : NZ, A' n <-> A' (S n). + +Lemma NZls_ls' : A z -> left_step -> left_step'. +Proof. +intros Az LS n H1 H2. le_elim H1. +apply LS; [assumption | apply H2; [now apply <- NZle_succ_l | now apply NZeq_le_incl]]. +rewrite H1; apply Az. +Qed. + +Lemma NZls'_ls'' : left_step' -> left_step''. +Proof. +intros LS' n; split; intros H1 m H2 H3. +apply -> NZle_succ_l in H3. apply NZlt_le_incl in H3. now apply H1. +le_elim H3. +apply <- NZle_succ_l in H3. now apply H1. +rewrite <- H3 in *; now apply LS'. +Qed. + +Lemma NZlbase : A' (S z). +Proof. +intros m H1 H2. apply -> NZle_succ_l in H2. +apply -> NZle_ngt in H1. false_hyp H2 H1. +Qed. + +Lemma NZA'A_left : (forall n : NZ, A' n) -> forall n : NZ, n <= z -> A n. +Proof. +intros H1 n H2. apply H1 with (n := n); [assumption | now apply NZeq_le_incl]. +Qed. + +Theorem NZstrong_left_induction: left_step' -> forall n : NZ, n <= z -> A n. +Proof. +intro LS'; apply NZA'A_left; unfold A'; NZinduct n (S z); +[apply NZlbase | apply NZls'_ls''; apply LS']. +Qed. + +Theorem NZleft_induction : A z -> left_step -> forall n : NZ, n <= z -> A n. +Proof. +intros Az LS; apply NZstrong_left_induction; now apply NZls_ls'. +Qed. + +Theorem NZleft_induction' : + (forall n : NZ, z <= n -> A n) -> left_step -> forall n : NZ, A n. +Proof. +intros R L n. +destruct (NZlt_trichotomy n z) as [H | [H | H]]. +apply NZleft_induction. apply R. now apply NZeq_le_incl. assumption. now apply NZlt_le_incl. +rewrite H; apply R; now apply NZeq_le_incl. +apply R; now apply NZlt_le_incl. +Qed. + +Theorem NZstrong_left_induction' : + (forall n : NZ, z <= n -> A n) -> left_step' -> forall n : NZ, A n. +Proof. +intros R L n. +destruct (NZlt_trichotomy n z) as [H | [H | H]]. +apply NZstrong_left_induction; auto. now apply NZlt_le_incl. +rewrite H; apply R; now apply NZeq_le_incl. +apply R; now apply NZlt_le_incl. +Qed. + +End LeftInduction. + +Theorem NZorder_induction : + A z -> + (forall n : NZ, z <= n -> A n -> A (S n)) -> + (forall n : NZ, n < z -> A (S n) -> A n) -> + forall n : NZ, A n. +Proof. +intros Az RS LS n. +destruct (NZlt_trichotomy n z) as [H | [H | H]]. +now apply NZleft_induction; [| | apply NZlt_le_incl]. +now rewrite H. +now apply NZright_induction; [| | apply NZlt_le_incl]. +Qed. + +Theorem NZorder_induction' : + A z -> + (forall n : NZ, z <= n -> A n -> A (S n)) -> + (forall n : NZ, n <= z -> A n -> A (P n)) -> + forall n : NZ, A n. +Proof. +intros Az AS AP n; apply NZorder_induction; try assumption. +intros m H1 H2. apply AP in H2; [| now apply <- NZle_succ_l]. +unfold predicate_wd, fun_wd in A_wd; apply -> (A_wd (P (S m)) m); +[assumption | apply NZpred_succ]. +Qed. + +End Center. + +Theorem NZorder_induction_0 : + A 0 -> + (forall n : NZ, 0 <= n -> A n -> A (S n)) -> + (forall n : NZ, n < 0 -> A (S n) -> A n) -> + forall n : NZ, A n. +Proof (NZorder_induction 0). + +Theorem NZorder_induction'_0 : + A 0 -> + (forall n : NZ, 0 <= n -> A n -> A (S n)) -> + (forall n : NZ, n <= 0 -> A n -> A (P n)) -> + forall n : NZ, A n. +Proof (NZorder_induction' 0). + +(** Elimintation principle for < *) + +Theorem NZlt_ind : forall (n : NZ), + A (S n) -> + (forall m : NZ, n < m -> A m -> A (S m)) -> + forall m : NZ, n < m -> A m. +Proof. +intros n H1 H2 m H3. +apply NZright_induction with (S n); [assumption | | now apply <- NZle_succ_l]. +intros; apply H2; try assumption. now apply -> NZle_succ_l. +Qed. + +(** Elimintation principle for <= *) + +Theorem NZle_ind : forall (n : NZ), + A n -> + (forall m : NZ, n <= m -> A m -> A (S m)) -> + forall m : NZ, n <= m -> A m. +Proof. +intros n H1 H2 m H3. +now apply NZright_induction with n. +Qed. + +End Induction. + +Tactic Notation "NZord_induct" ident(n) := + induction_maker n ltac:(apply NZorder_induction_0). + +Tactic Notation "NZord_induct" ident(n) constr(z) := + induction_maker n ltac:(apply NZorder_induction with z). + +Section WF. + +Variable z : NZ. + +Let Rlt (n m : NZ) := z <= n /\ n < m. +Let Rgt (n m : NZ) := m < n /\ n <= z. + +Add Morphism Rlt with signature NZeq ==> NZeq ==> iff as Rlt_wd. +Proof. +intros x1 x2 H1 x3 x4 H2; unfold Rlt; rewrite H1; now rewrite H2. +Qed. + +Add Morphism Rgt with signature NZeq ==> NZeq ==> iff as Rgt_wd. +Proof. +intros x1 x2 H1 x3 x4 H2; unfold Rgt; rewrite H1; now rewrite H2. +Qed. + +Lemma NZAcc_lt_wd : predicate_wd NZeq (Acc Rlt). +Proof. +unfold predicate_wd, fun_wd. +intros x1 x2 H; split; intro H1; destruct H1 as [H2]; +constructor; intros; apply H2; now (rewrite H || rewrite <- H). +Qed. + +Lemma NZAcc_gt_wd : predicate_wd NZeq (Acc Rgt). +Proof. +unfold predicate_wd, fun_wd. +intros x1 x2 H; split; intro H1; destruct H1 as [H2]; +constructor; intros; apply H2; now (rewrite H || rewrite <- H). +Qed. + +Theorem NZlt_wf : well_founded Rlt. +Proof. +unfold well_founded. +apply NZstrong_right_induction' with (z := z). +apply NZAcc_lt_wd. +intros n H; constructor; intros y [H1 H2]. +apply <- NZnle_gt in H2. elim H2. now apply NZle_trans with z. +intros n H1 H2; constructor; intros m [H3 H4]. now apply H2. +Qed. + +Theorem NZgt_wf : well_founded Rgt. +Proof. +unfold well_founded. +apply NZstrong_left_induction' with (z := z). +apply NZAcc_gt_wd. +intros n H; constructor; intros y [H1 H2]. +apply <- NZnle_gt in H2. elim H2. now apply NZle_lt_trans with n. +intros n H1 H2; constructor; intros m [H3 H4]. +apply H2. assumption. now apply <- NZle_succ_l. +Qed. + +End WF. + +End NZOrderPropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NAdd.v b/theories/Numbers/Natural/Abstract/NAdd.v new file mode 100644 index 00000000..f58b87d8 --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NAdd.v @@ -0,0 +1,156 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NAdd.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NBase. + +Module NAddPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NBasePropMod := NBasePropFunct NAxiomsMod. + +Open Local Scope NatScope. + +Theorem add_wd : + forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> n1 + m1 == n2 + m2. +Proof NZadd_wd. + +Theorem add_0_l : forall n : N, 0 + n == n. +Proof NZadd_0_l. + +Theorem add_succ_l : forall n m : N, (S n) + m == S (n + m). +Proof NZadd_succ_l. + +(** Theorems that are valid for both natural numbers and integers *) + +Theorem add_0_r : forall n : N, n + 0 == n. +Proof NZadd_0_r. + +Theorem add_succ_r : forall n m : N, n + S m == S (n + m). +Proof NZadd_succ_r. + +Theorem add_comm : forall n m : N, n + m == m + n. +Proof NZadd_comm. + +Theorem add_assoc : forall n m p : N, n + (m + p) == (n + m) + p. +Proof NZadd_assoc. + +Theorem add_shuffle1 : forall n m p q : N, (n + m) + (p + q) == (n + p) + (m + q). +Proof NZadd_shuffle1. + +Theorem add_shuffle2 : forall n m p q : N, (n + m) + (p + q) == (n + q) + (m + p). +Proof NZadd_shuffle2. + +Theorem add_1_l : forall n : N, 1 + n == S n. +Proof NZadd_1_l. + +Theorem add_1_r : forall n : N, n + 1 == S n. +Proof NZadd_1_r. + +Theorem add_cancel_l : forall n m p : N, p + n == p + m <-> n == m. +Proof NZadd_cancel_l. + +Theorem add_cancel_r : forall n m p : N, n + p == m + p <-> n == m. +Proof NZadd_cancel_r. + +(* Theorems that are valid for natural numbers but cannot be proved for Z *) + +Theorem eq_add_0 : forall n m : N, n + m == 0 <-> n == 0 /\ m == 0. +Proof. +intros n m; induct n. +(* The next command does not work with the axiom add_0_l from NAddSig *) +rewrite add_0_l. intuition reflexivity. +intros n IH. rewrite add_succ_l. +setoid_replace (S (n + m) == 0) with False using relation iff by + (apply -> neg_false; apply neq_succ_0). +setoid_replace (S n == 0) with False using relation iff by + (apply -> neg_false; apply neq_succ_0). tauto. +Qed. + +Theorem eq_add_succ : + forall n m : N, (exists p : N, n + m == S p) <-> + (exists n' : N, n == S n') \/ (exists m' : N, m == S m'). +Proof. +intros n m; cases n. +split; intro H. +destruct H as [p H]. rewrite add_0_l in H; right; now exists p. +destruct H as [[n' H] | [m' H]]. +symmetry in H; false_hyp H neq_succ_0. +exists m'; now rewrite add_0_l. +intro n; split; intro H. +left; now exists n. +exists (n + m); now rewrite add_succ_l. +Qed. + +Theorem eq_add_1 : forall n m : N, + n + m == 1 -> n == 1 /\ m == 0 \/ n == 0 /\ m == 1. +Proof. +intros n m H. +assert (H1 : exists p : N, n + m == S p) by now exists 0. +apply -> eq_add_succ in H1. destruct H1 as [[n' H1] | [m' H1]]. +left. rewrite H1 in H; rewrite add_succ_l in H; apply succ_inj in H. +apply -> eq_add_0 in H. destruct H as [H2 H3]; rewrite H2 in H1; now split. +right. rewrite H1 in H; rewrite add_succ_r in H; apply succ_inj in H. +apply -> eq_add_0 in H. destruct H as [H2 H3]; rewrite H3 in H1; now split. +Qed. + +Theorem succ_add_discr : forall n m : N, m ~= S (n + m). +Proof. +intro n; induct m. +apply neq_symm. apply neq_succ_0. +intros m IH H. apply succ_inj in H. rewrite add_succ_r in H. +unfold not in IH; now apply IH. +Qed. + +Theorem add_pred_l : forall n m : N, n ~= 0 -> P n + m == P (n + m). +Proof. +intros n m; cases n. +intro H; now elim H. +intros n IH; rewrite add_succ_l; now do 2 rewrite pred_succ. +Qed. + +Theorem add_pred_r : forall n m : N, m ~= 0 -> n + P m == P (n + m). +Proof. +intros n m H; rewrite (add_comm n (P m)); +rewrite (add_comm n m); now apply add_pred_l. +Qed. + +(* One could define n <= m as exists p : N, p + n == m. Then we have +dichotomy: + +forall n m : N, n <= m \/ m <= n, + +i.e., + +forall n m : N, (exists p : N, p + n == m) \/ (exists p : N, p + m == n) (1) + +We will need (1) in the proof of induction principle for integers +constructed as pairs of natural numbers. The formula (1) can be proved +using properties of order and truncated subtraction. Thus, p would be +m - n or n - m and (1) would hold by theorem sub_add from Sub.v +depending on whether n <= m or m <= n. However, in proving induction +for integers constructed from natural numbers we do not need to +require implementations of order and sub; it is enough to prove (1) +here. *) + +Theorem add_dichotomy : + forall n m : N, (exists p : N, p + n == m) \/ (exists p : N, p + m == n). +Proof. +intros n m; induct n. +left; exists m; apply add_0_r. +intros n IH. +destruct IH as [[p H] | [p H]]. +destruct (zero_or_succ p) as [H1 | [p' H1]]; rewrite H1 in H. +rewrite add_0_l in H. right; exists (S 0); rewrite H; rewrite add_succ_l; now rewrite add_0_l. +left; exists p'; rewrite add_succ_r; now rewrite add_succ_l in H. +right; exists (S p). rewrite add_succ_l; now rewrite H. +Qed. + +End NAddPropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NAddOrder.v b/theories/Numbers/Natural/Abstract/NAddOrder.v new file mode 100644 index 00000000..7024fd00 --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NAddOrder.v @@ -0,0 +1,114 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NAddOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NOrder. + +Module NAddOrderPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NOrderPropMod := NOrderPropFunct NAxiomsMod. +Open Local Scope NatScope. + +Theorem add_lt_mono_l : forall n m p : N, n < m <-> p + n < p + m. +Proof NZadd_lt_mono_l. + +Theorem add_lt_mono_r : forall n m p : N, n < m <-> n + p < m + p. +Proof NZadd_lt_mono_r. + +Theorem add_lt_mono : forall n m p q : N, n < m -> p < q -> n + p < m + q. +Proof NZadd_lt_mono. + +Theorem add_le_mono_l : forall n m p : N, n <= m <-> p + n <= p + m. +Proof NZadd_le_mono_l. + +Theorem add_le_mono_r : forall n m p : N, n <= m <-> n + p <= m + p. +Proof NZadd_le_mono_r. + +Theorem add_le_mono : forall n m p q : N, n <= m -> p <= q -> n + p <= m + q. +Proof NZadd_le_mono. + +Theorem add_lt_le_mono : forall n m p q : N, n < m -> p <= q -> n + p < m + q. +Proof NZadd_lt_le_mono. + +Theorem add_le_lt_mono : forall n m p q : N, n <= m -> p < q -> n + p < m + q. +Proof NZadd_le_lt_mono. + +Theorem add_pos_pos : forall n m : N, 0 < n -> 0 < m -> 0 < n + m. +Proof NZadd_pos_pos. + +Theorem lt_add_pos_l : forall n m : N, 0 < n -> m < n + m. +Proof NZlt_add_pos_l. + +Theorem lt_add_pos_r : forall n m : N, 0 < n -> m < m + n. +Proof NZlt_add_pos_r. + +Theorem le_lt_add_lt : forall n m p q : N, n <= m -> p + m < q + n -> p < q. +Proof NZle_lt_add_lt. + +Theorem lt_le_add_lt : forall n m p q : N, n < m -> p + m <= q + n -> p < q. +Proof NZlt_le_add_lt. + +Theorem le_le_add_le : forall n m p q : N, n <= m -> p + m <= q + n -> p <= q. +Proof NZle_le_add_le. + +Theorem add_lt_cases : forall n m p q : N, n + m < p + q -> n < p \/ m < q. +Proof NZadd_lt_cases. + +Theorem add_le_cases : forall n m p q : N, n + m <= p + q -> n <= p \/ m <= q. +Proof NZadd_le_cases. + +Theorem add_pos_cases : forall n m : N, 0 < n + m -> 0 < n \/ 0 < m. +Proof NZadd_pos_cases. + +(* Theorems true for natural numbers *) + +Theorem le_add_r : forall n m : N, n <= n + m. +Proof. +intro n; induct m. +rewrite add_0_r; now apply eq_le_incl. +intros m IH. rewrite add_succ_r; now apply le_le_succ_r. +Qed. + +Theorem lt_lt_add_r : forall n m p : N, n < m -> n < m + p. +Proof. +intros n m p H; rewrite <- (add_0_r n). +apply add_lt_le_mono; [assumption | apply le_0_l]. +Qed. + +Theorem lt_lt_add_l : forall n m p : N, n < m -> n < p + m. +Proof. +intros n m p; rewrite add_comm; apply lt_lt_add_r. +Qed. + +Theorem add_pos_l : forall n m : N, 0 < n -> 0 < n + m. +Proof. +intros; apply NZadd_pos_nonneg. assumption. apply le_0_l. +Qed. + +Theorem add_pos_r : forall n m : N, 0 < m -> 0 < n + m. +Proof. +intros; apply NZadd_nonneg_pos. apply le_0_l. assumption. +Qed. + +(* The following property is used to prove the correctness of the +definition of order on integers constructed from pairs of natural numbers *) + +Theorem add_lt_repl_pair : forall n m n' m' u v : N, + n + u < m + v -> n + m' == n' + m -> n' + u < m' + v. +Proof. +intros n m n' m' u v H1 H2. +symmetry in H2. assert (H3 : n' + m <= n + m') by now apply eq_le_incl. +pose proof (add_lt_le_mono _ _ _ _ H1 H3) as H4. +rewrite (add_shuffle2 n u), (add_shuffle1 m v), (add_comm m n) in H4. +do 2 rewrite <- add_assoc in H4. do 2 apply <- add_lt_mono_l in H4. +now rewrite (add_comm n' u), (add_comm m' v). +Qed. + +End NAddOrderPropFunct. diff --git a/theories/Numbers/Natural/Abstract/NAxioms.v b/theories/Numbers/Natural/Abstract/NAxioms.v new file mode 100644 index 00000000..750cc977 --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NAxioms.v @@ -0,0 +1,71 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NZAxioms. + +Set Implicit Arguments. + +Module Type NAxiomsSig. +Declare Module Export NZOrdAxiomsMod : NZOrdAxiomsSig. + +Delimit Scope NatScope with Nat. +Notation N := NZ. +Notation Neq := NZeq. +Notation N0 := NZ0. +Notation N1 := (NZsucc NZ0). +Notation S := NZsucc. +Notation P := NZpred. +Notation add := NZadd. +Notation mul := NZmul. +Notation sub := NZsub. +Notation lt := NZlt. +Notation le := NZle. +Notation min := NZmin. +Notation max := NZmax. +Notation "x == y" := (Neq x y) (at level 70) : NatScope. +Notation "x ~= y" := (~ Neq x y) (at level 70) : NatScope. +Notation "0" := NZ0 : NatScope. +Notation "1" := (NZsucc NZ0) : NatScope. +Notation "x + y" := (NZadd x y) : NatScope. +Notation "x - y" := (NZsub x y) : NatScope. +Notation "x * y" := (NZmul x y) : NatScope. +Notation "x < y" := (NZlt x y) : NatScope. +Notation "x <= y" := (NZle x y) : NatScope. +Notation "x > y" := (NZlt y x) (only parsing) : NatScope. +Notation "x >= y" := (NZle y x) (only parsing) : NatScope. + +Open Local Scope NatScope. + +Parameter Inline recursion : forall A : Type, A -> (N -> A -> A) -> N -> A. +Implicit Arguments recursion [A]. + +Axiom pred_0 : P 0 == 0. + +Axiom recursion_wd : forall (A : Type) (Aeq : relation A), + forall a a' : A, Aeq a a' -> + forall f f' : N -> A -> A, fun2_eq Neq Aeq Aeq f f' -> + forall x x' : N, x == x' -> + Aeq (recursion a f x) (recursion a' f' x'). + +Axiom recursion_0 : + forall (A : Type) (a : A) (f : N -> A -> A), recursion a f 0 = a. + +Axiom recursion_succ : + forall (A : Type) (Aeq : relation A) (a : A) (f : N -> A -> A), + Aeq a a -> fun2_wd Neq Aeq Aeq f -> + forall n : N, Aeq (recursion a f (S n)) (f n (recursion a f n)). + +(*Axiom dep_rec : + forall A : N -> Type, A 0 -> (forall n : N, A n -> A (S n)) -> forall n : N, A n.*) + +End NAxiomsSig. + diff --git a/theories/Numbers/Natural/Abstract/NBase.v b/theories/Numbers/Natural/Abstract/NBase.v new file mode 100644 index 00000000..3e4032b5 --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NBase.v @@ -0,0 +1,288 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NBase.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export Decidable. +Require Export NAxioms. +Require Import NZMulOrder. (* The last property functor on NZ, which subsumes all others *) + +Module NBasePropFunct (Import NAxiomsMod : NAxiomsSig). + +Open Local Scope NatScope. + +(* We call the last property functor on NZ, which includes all the previous +ones, to get all properties of NZ at once. This way we will include them +only one time. *) + +Module Export NZMulOrderMod := NZMulOrderPropFunct NZOrdAxiomsMod. + +(* Here we probably need to re-prove all axioms declared in NAxioms.v to +make sure that the definitions like N, S and add are unfolded in them, +since unfolding is done only inside a functor. In fact, we'll do it in the +files that prove the corresponding properties. In those files, we will also +rename properties proved in NZ files by removing NZ from their names. In +this way, one only has to consult, for example, NAdd.v to see all +available properties for add, i.e., one does not have to go to NAxioms.v +for axioms and NZAdd.v for theorems. *) + +Theorem succ_wd : forall n1 n2 : N, n1 == n2 -> S n1 == S n2. +Proof NZsucc_wd. + +Theorem pred_wd : forall n1 n2 : N, n1 == n2 -> P n1 == P n2. +Proof NZpred_wd. + +Theorem pred_succ : forall n : N, P (S n) == n. +Proof NZpred_succ. + +Theorem pred_0 : P 0 == 0. +Proof pred_0. + +Theorem Neq_refl : forall n : N, n == n. +Proof (proj1 NZeq_equiv). + +Theorem Neq_symm : forall n m : N, n == m -> m == n. +Proof (proj2 (proj2 NZeq_equiv)). + +Theorem Neq_trans : forall n m p : N, n == m -> m == p -> n == p. +Proof (proj1 (proj2 NZeq_equiv)). + +Theorem neq_symm : forall n m : N, n ~= m -> m ~= n. +Proof NZneq_symm. + +Theorem succ_inj : forall n1 n2 : N, S n1 == S n2 -> n1 == n2. +Proof NZsucc_inj. + +Theorem succ_inj_wd : forall n1 n2 : N, S n1 == S n2 <-> n1 == n2. +Proof NZsucc_inj_wd. + +Theorem succ_inj_wd_neg : forall n m : N, S n ~= S m <-> n ~= m. +Proof NZsucc_inj_wd_neg. + +(* Decidability and stability of equality was proved only in NZOrder, but +since it does not mention order, we'll put it here *) + +Theorem eq_dec : forall n m : N, decidable (n == m). +Proof NZeq_dec. + +Theorem eq_dne : forall n m : N, ~ ~ n == m <-> n == m. +Proof NZeq_dne. + +(* Now we prove that the successor of a number is not zero by defining a +function (by recursion) that maps 0 to false and the successor to true *) + +Definition if_zero (A : Set) (a b : A) (n : N) : A := + recursion a (fun _ _ => b) n. + +Add Parametric Morphism (A : Set) : (if_zero A) with signature (@eq _ ==> @eq _ ==> Neq ==> @eq _) as if_zero_wd. +Proof. +intros; unfold if_zero. apply recursion_wd with (Aeq := (@eq A)). +reflexivity. unfold fun2_eq; now intros. assumption. +Qed. + +Theorem if_zero_0 : forall (A : Set) (a b : A), if_zero A a b 0 = a. +Proof. +unfold if_zero; intros; now rewrite recursion_0. +Qed. + +Theorem if_zero_succ : forall (A : Set) (a b : A) (n : N), if_zero A a b (S n) = b. +Proof. +intros; unfold if_zero. +now rewrite (@recursion_succ A (@eq A)); [| | unfold fun2_wd; now intros]. +Qed. + +Implicit Arguments if_zero [A]. + +Theorem neq_succ_0 : forall n : N, S n ~= 0. +Proof. +intros n H. +assert (true = false); [| discriminate]. +replace true with (if_zero false true (S n)) by apply if_zero_succ. +pattern false at 2; replace false with (if_zero false true 0) by apply if_zero_0. +now rewrite H. +Qed. + +Theorem neq_0_succ : forall n : N, 0 ~= S n. +Proof. +intro n; apply neq_symm; apply neq_succ_0. +Qed. + +(* Next, we show that all numbers are nonnegative and recover regular induction +from the bidirectional induction on NZ *) + +Theorem le_0_l : forall n : N, 0 <= n. +Proof. +NZinduct n. +now apply NZeq_le_incl. +intro n; split. +apply NZle_le_succ_r. +intro H; apply -> NZle_succ_r in H; destruct H as [H | H]. +assumption. +symmetry in H; false_hyp H neq_succ_0. +Qed. + +Theorem induction : + forall A : N -> Prop, predicate_wd Neq A -> + A 0 -> (forall n : N, A n -> A (S n)) -> forall n : N, A n. +Proof. +intros A A_wd A0 AS n; apply NZright_induction with 0; try assumption. +intros; auto; apply le_0_l. apply le_0_l. +Qed. + +(* The theorems NZinduction, NZcentral_induction and the tactic NZinduct +refer to bidirectional induction, which is not useful on natural +numbers. Therefore, we define a new induction tactic for natural numbers. +We do not have to call "Declare Left Step" and "Declare Right Step" +commands again, since the data for stepl and stepr tactics is inherited +from NZ. *) + +Ltac induct n := induction_maker n ltac:(apply induction). + +Theorem case_analysis : + forall A : N -> Prop, predicate_wd Neq A -> + A 0 -> (forall n : N, A (S n)) -> forall n : N, A n. +Proof. +intros; apply induction; auto. +Qed. + +Ltac cases n := induction_maker n ltac:(apply case_analysis). + +Theorem neq_0 : ~ forall n, n == 0. +Proof. +intro H; apply (neq_succ_0 0). apply H. +Qed. + +Theorem neq_0_r : forall n, n ~= 0 <-> exists m, n == S m. +Proof. +cases n. split; intro H; +[now elim H | destruct H as [m H]; symmetry in H; false_hyp H neq_succ_0]. +intro n; split; intro H; [now exists n | apply neq_succ_0]. +Qed. + +Theorem zero_or_succ : forall n, n == 0 \/ exists m, n == S m. +Proof. +cases n. +now left. +intro n; right; now exists n. +Qed. + +Theorem eq_pred_0 : forall n : N, P n == 0 <-> n == 0 \/ n == 1. +Proof. +cases n. +rewrite pred_0. setoid_replace (0 == 1) with False using relation iff. tauto. +split; intro H; [symmetry in H; false_hyp H neq_succ_0 | elim H]. +intro n. rewrite pred_succ. +setoid_replace (S n == 0) with False using relation iff by + (apply -> neg_false; apply neq_succ_0). +rewrite succ_inj_wd. tauto. +Qed. + +Theorem succ_pred : forall n : N, n ~= 0 -> S (P n) == n. +Proof. +cases n. +intro H; elimtype False; now apply H. +intros; now rewrite pred_succ. +Qed. + +Theorem pred_inj : forall n m : N, n ~= 0 -> m ~= 0 -> P n == P m -> n == m. +Proof. +intros n m; cases n. +intros H; elimtype False; now apply H. +intros n _; cases m. +intros H; elimtype False; now apply H. +intros m H2 H3. do 2 rewrite pred_succ in H3. now rewrite H3. +Qed. + +(* The following induction principle is useful for reasoning about, e.g., +Fibonacci numbers *) + +Section PairInduction. + +Variable A : N -> Prop. +Hypothesis A_wd : predicate_wd Neq A. + +Add Morphism A with signature Neq ==> iff as A_morph. +Proof. +exact A_wd. +Qed. + +Theorem pair_induction : + A 0 -> A 1 -> + (forall n, A n -> A (S n) -> A (S (S n))) -> forall n, A n. +Proof. +intros until 3. +assert (D : forall n, A n /\ A (S n)); [ |intro n; exact (proj1 (D n))]. +induct n; [ | intros n [IH1 IH2]]; auto. +Qed. + +End PairInduction. + +(*Ltac pair_induct n := induction_maker n ltac:(apply pair_induction).*) + +(* The following is useful for reasoning about, e.g., Ackermann function *) +Section TwoDimensionalInduction. + +Variable R : N -> N -> Prop. +Hypothesis R_wd : relation_wd Neq Neq R. + +Add Morphism R with signature Neq ==> Neq ==> iff as R_morph. +Proof. +exact R_wd. +Qed. + +Theorem two_dim_induction : + R 0 0 -> + (forall n m, R n m -> R n (S m)) -> + (forall n, (forall m, R n m) -> R (S n) 0) -> forall n m, R n m. +Proof. +intros H1 H2 H3. induct n. +induct m. +exact H1. exact (H2 0). +intros n IH. induct m. +now apply H3. exact (H2 (S n)). +Qed. + +End TwoDimensionalInduction. + +(*Ltac two_dim_induct n m := + try intros until n; + try intros until m; + pattern n, m; apply two_dim_induction; clear n m; + [solve_relation_wd | | | ].*) + +Section DoubleInduction. + +Variable R : N -> N -> Prop. +Hypothesis R_wd : relation_wd Neq Neq R. + +Add Morphism R with signature Neq ==> Neq ==> iff as R_morph1. +Proof. +exact R_wd. +Qed. + +Theorem double_induction : + (forall m : N, R 0 m) -> + (forall n : N, R (S n) 0) -> + (forall n m : N, R n m -> R (S n) (S m)) -> forall n m : N, R n m. +Proof. +intros H1 H2 H3; induct n; auto. +intros n H; cases m; auto. +Qed. + +End DoubleInduction. + +Ltac double_induct n m := + try intros until n; + try intros until m; + pattern n, m; apply double_induction; clear n m; + [solve_relation_wd | | | ]. + +End NBasePropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NDefOps.v b/theories/Numbers/Natural/Abstract/NDefOps.v new file mode 100644 index 00000000..e15e4672 --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NDefOps.v @@ -0,0 +1,298 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NDefOps.v 11039 2008-06-02 23:26:13Z letouzey $ i*) + +Require Import Bool. (* To get the orb and negb function *) +Require Export NStrongRec. + +Module NdefOpsPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NStrongRecPropMod := NStrongRecPropFunct NAxiomsMod. +Open Local Scope NatScope. + +(*****************************************************) +(** Addition *) + +Definition def_add (x y : N) := recursion y (fun _ p => S p) x. + +Infix Local "++" := def_add (at level 50, left associativity). + +Add Morphism def_add with signature Neq ==> Neq ==> Neq as def_add_wd. +Proof. +unfold def_add. +intros x x' Exx' y y' Eyy'. +apply recursion_wd with (Aeq := Neq). +assumption. +unfold fun2_eq; intros _ _ _ p p' Epp'; now rewrite Epp'. +assumption. +Qed. + +Theorem def_add_0_l : forall y : N, 0 ++ y == y. +Proof. +intro y. unfold def_add. now rewrite recursion_0. +Qed. + +Theorem def_add_succ_l : forall x y : N, S x ++ y == S (x ++ y). +Proof. +intros x y; unfold def_add. +rewrite (@recursion_succ N Neq); try reflexivity. +unfold fun2_wd. intros _ _ _ m1 m2 H2. now rewrite H2. +Qed. + +Theorem def_add_add : forall n m : N, n ++ m == n + m. +Proof. +intros n m; induct n. +now rewrite def_add_0_l, add_0_l. +intros n H. now rewrite def_add_succ_l, add_succ_l, H. +Qed. + +(*****************************************************) +(** Multiplication *) + +Definition def_mul (x y : N) := recursion 0 (fun _ p => p ++ x) y. + +Infix Local "**" := def_mul (at level 40, left associativity). + +Lemma def_mul_step_wd : forall x : N, fun2_wd Neq Neq Neq (fun _ p => def_add p x). +Proof. +unfold fun2_wd. intros. now apply def_add_wd. +Qed. + +Lemma def_mul_step_equal : + forall x x' : N, x == x' -> + fun2_eq Neq Neq Neq (fun _ p => def_add p x) (fun x p => def_add p x'). +Proof. +unfold fun2_eq; intros; apply def_add_wd; assumption. +Qed. + +Add Morphism def_mul with signature Neq ==> Neq ==> Neq as def_mul_wd. +Proof. +unfold def_mul. +intros x x' Exx' y y' Eyy'. +apply recursion_wd with (Aeq := Neq). +reflexivity. apply def_mul_step_equal. assumption. assumption. +Qed. + +Theorem def_mul_0_r : forall x : N, x ** 0 == 0. +Proof. +intro. unfold def_mul. now rewrite recursion_0. +Qed. + +Theorem def_mul_succ_r : forall x y : N, x ** S y == x ** y ++ x. +Proof. +intros x y; unfold def_mul. +now rewrite (@recursion_succ N Neq); [| apply def_mul_step_wd |]. +Qed. + +Theorem def_mul_mul : forall n m : N, n ** m == n * m. +Proof. +intros n m; induct m. +now rewrite def_mul_0_r, mul_0_r. +intros m IH; now rewrite def_mul_succ_r, mul_succ_r, def_add_add, IH. +Qed. + +(*****************************************************) +(** Order *) + +Definition def_ltb (m : N) : N -> bool := +recursion + (if_zero false true) + (fun _ f => fun n => recursion false (fun n' _ => f n') n) + m. + +Infix Local "<<" := def_ltb (at level 70, no associativity). + +Lemma lt_base_wd : fun_wd Neq (@eq bool) (if_zero false true). +unfold fun_wd; intros; now apply if_zero_wd. +Qed. + +Lemma lt_step_wd : +fun2_wd Neq (fun_eq Neq (@eq bool)) (fun_eq Neq (@eq bool)) + (fun _ f => fun n => recursion false (fun n' _ => f n') n). +Proof. +unfold fun2_wd, fun_eq. +intros x x' Exx' f f' Eff' y y' Eyy'. +apply recursion_wd with (Aeq := @eq bool). +reflexivity. +unfold fun2_eq; intros; now apply Eff'. +assumption. +Qed. + +Lemma lt_curry_wd : + forall m m' : N, m == m' -> fun_eq Neq (@eq bool) (def_ltb m) (def_ltb m'). +Proof. +unfold def_ltb. +intros m m' Emm'. +apply recursion_wd with (Aeq := fun_eq Neq (@eq bool)). +apply lt_base_wd. +apply lt_step_wd. +assumption. +Qed. + +Add Morphism def_ltb with signature Neq ==> Neq ==> (@eq bool) as def_ltb_wd. +Proof. +intros; now apply lt_curry_wd. +Qed. + +Theorem def_ltb_base : forall n : N, 0 << n = if_zero false true n. +Proof. +intro n; unfold def_ltb; now rewrite recursion_0. +Qed. + +Theorem def_ltb_step : + forall m n : N, S m << n = recursion false (fun n' _ => m << n') n. +Proof. +intros m n; unfold def_ltb. +pose proof + (@recursion_succ + (N -> bool) + (fun_eq Neq (@eq bool)) + (if_zero false true) + (fun _ f => fun n => recursion false (fun n' _ => f n') n) + lt_base_wd + lt_step_wd + m n n) as H. +now rewrite H. +Qed. + +(* Above, we rewrite applications of function. Is it possible to rewrite +functions themselves, i.e., rewrite (recursion lt_base lt_step (S n)) to +lt_step n (recursion lt_base lt_step n)? *) + +Theorem def_ltb_0 : forall n : N, n << 0 = false. +Proof. +cases n. +rewrite def_ltb_base; now rewrite if_zero_0. +intro n; rewrite def_ltb_step. now rewrite recursion_0. +Qed. + +Theorem def_ltb_0_succ : forall n : N, 0 << S n = true. +Proof. +intro n; rewrite def_ltb_base; now rewrite if_zero_succ. +Qed. + +Theorem succ_def_ltb_mono : forall n m : N, (S n << S m) = (n << m). +Proof. +intros n m. +rewrite def_ltb_step. rewrite (@recursion_succ bool (@eq bool)); try reflexivity. +unfold fun2_wd; intros; now apply def_ltb_wd. +Qed. + +Theorem def_ltb_lt : forall n m : N, n << m = true <-> n < m. +Proof. +double_induct n m. +cases m. +rewrite def_ltb_0. split; intro H; [discriminate H | false_hyp H nlt_0_r]. +intro n. rewrite def_ltb_0_succ. split; intro; [apply lt_0_succ | reflexivity]. +intro n. rewrite def_ltb_0. split; intro H; [discriminate | false_hyp H nlt_0_r]. +intros n m. rewrite succ_def_ltb_mono. now rewrite <- succ_lt_mono. +Qed. + +(* +(*****************************************************) +(** Even *) + +Definition even (x : N) := recursion true (fun _ p => negb p) x. + +Lemma even_step_wd : fun2_wd Neq (@eq bool) (@eq bool) (fun x p => if p then false else true). +Proof. +unfold fun2_wd. +intros x x' Exx' b b' Ebb'. +unfold eq_bool; destruct b; destruct b'; now simpl. +Qed. + +Add Morphism even with signature Neq ==> (@eq bool) as even_wd. +Proof. +unfold even; intros. +apply recursion_wd with (A := bool) (Aeq := (@eq bool)). +now unfold eq_bool. +unfold fun2_eq. intros _ _ _ b b' Ebb'. unfold eq_bool; destruct b; destruct b'; now simpl. +assumption. +Qed. + +Theorem even_0 : even 0 = true. +Proof. +unfold even. +now rewrite recursion_0. +Qed. + +Theorem even_succ : forall x : N, even (S x) = negb (even x). +Proof. +unfold even. +intro x; rewrite (recursion_succ (@eq bool)); try reflexivity. +unfold fun2_wd. +intros _ _ _ b b' Ebb'. destruct b; destruct b'; now simpl. +Qed. + +(*****************************************************) +(** Division by 2 *) + +Definition half_aux (x : N) : N * N := + recursion (0, 0) (fun _ p => let (x1, x2) := p in ((S x2, x1))) x. + +Definition half (x : N) := snd (half_aux x). + +Definition E2 := prod_rel Neq Neq. + +Add Relation (prod N N) E2 +reflexivity proved by (prod_rel_refl N N Neq Neq E_equiv E_equiv) +symmetry proved by (prod_rel_symm N N Neq Neq E_equiv E_equiv) +transitivity proved by (prod_rel_trans N N Neq Neq E_equiv E_equiv) +as E2_rel. + +Lemma half_step_wd: fun2_wd Neq E2 E2 (fun _ p => let (x1, x2) := p in ((S x2, x1))). +Proof. +unfold fun2_wd, E2, prod_rel. +intros _ _ _ p1 p2 [H1 H2]. +destruct p1; destruct p2; simpl in *. +now split; [rewrite H2 |]. +Qed. + +Add Morphism half with signature Neq ==> Neq as half_wd. +Proof. +unfold half. +assert (H: forall x y, x == y -> E2 (half_aux x) (half_aux y)). +intros x y Exy; unfold half_aux; apply recursion_wd with (Aeq := E2); unfold E2. +unfold E2. +unfold prod_rel; simpl; now split. +unfold fun2_eq, prod_rel; simpl. +intros _ _ _ p1 p2; destruct p1; destruct p2; simpl. +intros [H1 H2]; split; [rewrite H2 | assumption]. reflexivity. assumption. +unfold E2, prod_rel in H. intros x y Exy; apply H in Exy. +exact (proj2 Exy). +Qed. + +(*****************************************************) +(** Logarithm for the base 2 *) + +Definition log (x : N) : N := +strong_rec 0 + (fun x g => + if (e x 0) then 0 + else if (e x 1) then 0 + else S (g (half x))) + x. + +Add Morphism log with signature Neq ==> Neq as log_wd. +Proof. +intros x x' Exx'. unfold log. +apply strong_rec_wd with (Aeq := Neq); try (reflexivity || assumption). +unfold fun2_eq. intros y y' Eyy' g g' Egg'. +assert (H : e y 0 = e y' 0); [now apply e_wd|]. +rewrite <- H; clear H. +assert (H : e y 1 = e y' 1); [now apply e_wd|]. +rewrite <- H; clear H. +assert (H : S (g (half y)) == S (g' (half y'))); +[apply succ_wd; apply Egg'; now apply half_wd|]. +now destruct (e y 0); destruct (e y 1). +Qed. +*) +End NdefOpsPropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NIso.v b/theories/Numbers/Natural/Abstract/NIso.v new file mode 100644 index 00000000..f6ccf3db --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NIso.v @@ -0,0 +1,122 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NIso.v 10934 2008-05-15 21:58:20Z letouzey $ i*) + +Require Import NBase. + +Module Homomorphism (NAxiomsMod1 NAxiomsMod2 : NAxiomsSig). + +Module NBasePropMod2 := NBasePropFunct NAxiomsMod2. + +Notation Local N1 := NAxiomsMod1.N. +Notation Local N2 := NAxiomsMod2.N. +Notation Local Eq1 := NAxiomsMod1.Neq. +Notation Local Eq2 := NAxiomsMod2.Neq. +Notation Local O1 := NAxiomsMod1.N0. +Notation Local O2 := NAxiomsMod2.N0. +Notation Local S1 := NAxiomsMod1.S. +Notation Local S2 := NAxiomsMod2.S. +Notation Local "n == m" := (Eq2 n m) (at level 70, no associativity). + +Definition homomorphism (f : N1 -> N2) : Prop := + f O1 == O2 /\ forall n : N1, f (S1 n) == S2 (f n). + +Definition natural_isomorphism : N1 -> N2 := + NAxiomsMod1.recursion O2 (fun (n : N1) (p : N2) => S2 p). + +Add Morphism natural_isomorphism with signature Eq1 ==> Eq2 as natural_isomorphism_wd. +Proof. +unfold natural_isomorphism. +intros n m Eqxy. +apply NAxiomsMod1.recursion_wd with (Aeq := Eq2). +reflexivity. +unfold fun2_eq. intros _ _ _ y' y'' H. now apply NBasePropMod2.succ_wd. +assumption. +Qed. + +Theorem natural_isomorphism_0 : natural_isomorphism O1 == O2. +Proof. +unfold natural_isomorphism; now rewrite NAxiomsMod1.recursion_0. +Qed. + +Theorem natural_isomorphism_succ : + forall n : N1, natural_isomorphism (S1 n) == S2 (natural_isomorphism n). +Proof. +unfold natural_isomorphism. +intro n. now rewrite (@NAxiomsMod1.recursion_succ N2 NAxiomsMod2.Neq) ; +[ | | unfold fun2_wd; intros; apply NBasePropMod2.succ_wd]. +Qed. + +Theorem hom_nat_iso : homomorphism natural_isomorphism. +Proof. +unfold homomorphism, natural_isomorphism; split; +[exact natural_isomorphism_0 | exact natural_isomorphism_succ]. +Qed. + +End Homomorphism. + +Module Inverse (NAxiomsMod1 NAxiomsMod2 : NAxiomsSig). + +Module Import NBasePropMod1 := NBasePropFunct NAxiomsMod1. +(* This makes the tactic induct available. Since it is taken from +(NBasePropFunct NAxiomsMod1), it refers to induction on N1. *) + +Module Hom12 := Homomorphism NAxiomsMod1 NAxiomsMod2. +Module Hom21 := Homomorphism NAxiomsMod2 NAxiomsMod1. + +Notation Local N1 := NAxiomsMod1.N. +Notation Local N2 := NAxiomsMod2.N. +Notation Local h12 := Hom12.natural_isomorphism. +Notation Local h21 := Hom21.natural_isomorphism. + +Notation Local "n == m" := (NAxiomsMod1.Neq n m) (at level 70, no associativity). + +Lemma inverse_nat_iso : forall n : N1, h21 (h12 n) == n. +Proof. +induct n. +now rewrite Hom12.natural_isomorphism_0, Hom21.natural_isomorphism_0. +intros n IH. +now rewrite Hom12.natural_isomorphism_succ, Hom21.natural_isomorphism_succ, IH. +Qed. + +End Inverse. + +Module Isomorphism (NAxiomsMod1 NAxiomsMod2 : NAxiomsSig). + +Module Hom12 := Homomorphism NAxiomsMod1 NAxiomsMod2. +Module Hom21 := Homomorphism NAxiomsMod2 NAxiomsMod1. + +Module Inverse12 := Inverse NAxiomsMod1 NAxiomsMod2. +Module Inverse21 := Inverse NAxiomsMod2 NAxiomsMod1. + +Notation Local N1 := NAxiomsMod1.N. +Notation Local N2 := NAxiomsMod2.N. +Notation Local Eq1 := NAxiomsMod1.Neq. +Notation Local Eq2 := NAxiomsMod2.Neq. +Notation Local h12 := Hom12.natural_isomorphism. +Notation Local h21 := Hom21.natural_isomorphism. + +Definition isomorphism (f1 : N1 -> N2) (f2 : N2 -> N1) : Prop := + Hom12.homomorphism f1 /\ Hom21.homomorphism f2 /\ + forall n : N1, Eq1 (f2 (f1 n)) n /\ + forall n : N2, Eq2 (f1 (f2 n)) n. + +Theorem iso_nat_iso : isomorphism h12 h21. +Proof. +unfold isomorphism. +split. apply Hom12.hom_nat_iso. +split. apply Hom21.hom_nat_iso. +split. apply Inverse12.inverse_nat_iso. +apply Inverse21.inverse_nat_iso. +Qed. + +End Isomorphism. + diff --git a/theories/Numbers/Natural/Abstract/NMul.v b/theories/Numbers/Natural/Abstract/NMul.v new file mode 100644 index 00000000..0b00f689 --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NMul.v @@ -0,0 +1,87 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NMul.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NAdd. + +Module NMulPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NAddPropMod := NAddPropFunct NAxiomsMod. +Open Local Scope NatScope. + +Theorem mul_wd : + forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> n1 * m1 == n2 * m2. +Proof NZmul_wd. + +Theorem mul_0_l : forall n : N, 0 * n == 0. +Proof NZmul_0_l. + +Theorem mul_succ_l : forall n m : N, (S n) * m == n * m + m. +Proof NZmul_succ_l. + +(** Theorems that are valid for both natural numbers and integers *) + +Theorem mul_0_r : forall n, n * 0 == 0. +Proof NZmul_0_r. + +Theorem mul_succ_r : forall n m, n * (S m) == n * m + n. +Proof NZmul_succ_r. + +Theorem mul_comm : forall n m : N, n * m == m * n. +Proof NZmul_comm. + +Theorem mul_add_distr_r : forall n m p : N, (n + m) * p == n * p + m * p. +Proof NZmul_add_distr_r. + +Theorem mul_add_distr_l : forall n m p : N, n * (m + p) == n * m + n * p. +Proof NZmul_add_distr_l. + +Theorem mul_assoc : forall n m p : N, n * (m * p) == (n * m) * p. +Proof NZmul_assoc. + +Theorem mul_1_l : forall n : N, 1 * n == n. +Proof NZmul_1_l. + +Theorem mul_1_r : forall n : N, n * 1 == n. +Proof NZmul_1_r. + +(* Theorems that cannot be proved in NZMul *) + +(* In proving the correctness of the definition of multiplication on +integers constructed from pairs of natural numbers, we'll need the +following fact about natural numbers: + +a * n + u == a * m + v -> n + m' == n' + m -> a * n' + u = a * m' + v + +Here n + m' == n' + m expresses equality of integers (n, m) and (n', m'), +since a pair (a, b) of natural numbers represents the integer a - b. On +integers, the formula above could be proved by moving a * m to the left, +factoring out a and replacing n - m by n' - m'. However, the formula is +required in the process of constructing integers, so it has to be proved +for natural numbers, where terms cannot be moved from one side of an +equation to the other. The proof uses the cancellation laws add_cancel_l +and add_cancel_r. *) + +Theorem add_mul_repl_pair : forall a n m n' m' u v : N, + a * n + u == a * m + v -> n + m' == n' + m -> a * n' + u == a * m' + v. +Proof. +intros a n m n' m' u v H1 H2. +apply (@NZmul_wd a a) in H2; [| reflexivity]. +do 2 rewrite mul_add_distr_l in H2. symmetry in H2. +pose proof (NZadd_wd _ _ H1 _ _ H2) as H3. +rewrite (add_shuffle1 (a * m)), (add_comm (a * m) (a * n)) in H3. +do 2 rewrite <- add_assoc in H3. apply -> add_cancel_l in H3. +rewrite (add_assoc u), (add_comm (a * m)) in H3. +apply -> add_cancel_r in H3. +now rewrite (add_comm (a * n') u), (add_comm (a * m') v). +Qed. + +End NMulPropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NMulOrder.v b/theories/Numbers/Natural/Abstract/NMulOrder.v new file mode 100644 index 00000000..aa21fb50 --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NMulOrder.v @@ -0,0 +1,131 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NMulOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NAddOrder. + +Module NMulOrderPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NAddOrderPropMod := NAddOrderPropFunct NAxiomsMod. +Open Local Scope NatScope. + +Theorem mul_lt_pred : + forall p q n m : N, S p == q -> (p * n < p * m <-> q * n + m < q * m + n). +Proof NZmul_lt_pred. + +Theorem mul_lt_mono_pos_l : forall p n m : N, 0 < p -> (n < m <-> p * n < p * m). +Proof NZmul_lt_mono_pos_l. + +Theorem mul_lt_mono_pos_r : forall p n m : N, 0 < p -> (n < m <-> n * p < m * p). +Proof NZmul_lt_mono_pos_r. + +Theorem mul_cancel_l : forall n m p : N, p ~= 0 -> (p * n == p * m <-> n == m). +Proof NZmul_cancel_l. + +Theorem mul_cancel_r : forall n m p : N, p ~= 0 -> (n * p == m * p <-> n == m). +Proof NZmul_cancel_r. + +Theorem mul_id_l : forall n m : N, m ~= 0 -> (n * m == m <-> n == 1). +Proof NZmul_id_l. + +Theorem mul_id_r : forall n m : N, n ~= 0 -> (n * m == n <-> m == 1). +Proof NZmul_id_r. + +Theorem mul_le_mono_pos_l : forall n m p : N, 0 < p -> (n <= m <-> p * n <= p * m). +Proof NZmul_le_mono_pos_l. + +Theorem mul_le_mono_pos_r : forall n m p : N, 0 < p -> (n <= m <-> n * p <= m * p). +Proof NZmul_le_mono_pos_r. + +Theorem mul_pos_pos : forall n m : N, 0 < n -> 0 < m -> 0 < n * m. +Proof NZmul_pos_pos. + +Theorem lt_1_mul_pos : forall n m : N, 1 < n -> 0 < m -> 1 < n * m. +Proof NZlt_1_mul_pos. + +Theorem eq_mul_0 : forall n m : N, n * m == 0 <-> n == 0 \/ m == 0. +Proof NZeq_mul_0. + +Theorem neq_mul_0 : forall n m : N, n ~= 0 /\ m ~= 0 <-> n * m ~= 0. +Proof NZneq_mul_0. + +Theorem eq_square_0 : forall n : N, n * n == 0 <-> n == 0. +Proof NZeq_square_0. + +Theorem eq_mul_0_l : forall n m : N, n * m == 0 -> m ~= 0 -> n == 0. +Proof NZeq_mul_0_l. + +Theorem eq_mul_0_r : forall n m : N, n * m == 0 -> n ~= 0 -> m == 0. +Proof NZeq_mul_0_r. + +Theorem square_lt_mono : forall n m : N, n < m <-> n * n < m * m. +Proof. +intros n m; split; intro; +[apply NZsquare_lt_mono_nonneg | apply NZsquare_lt_simpl_nonneg]; +try assumption; apply le_0_l. +Qed. + +Theorem square_le_mono : forall n m : N, n <= m <-> n * n <= m * m. +Proof. +intros n m; split; intro; +[apply NZsquare_le_mono_nonneg | apply NZsquare_le_simpl_nonneg]; +try assumption; apply le_0_l. +Qed. + +Theorem mul_2_mono_l : forall n m : N, n < m -> 1 + (1 + 1) * n < (1 + 1) * m. +Proof NZmul_2_mono_l. + +(* Theorems that are either not valid on Z or have different proofs on N and Z *) + +Theorem mul_le_mono_l : forall n m p : N, n <= m -> p * n <= p * m. +Proof. +intros; apply NZmul_le_mono_nonneg_l. apply le_0_l. assumption. +Qed. + +Theorem mul_le_mono_r : forall n m p : N, n <= m -> n * p <= m * p. +Proof. +intros; apply NZmul_le_mono_nonneg_r. apply le_0_l. assumption. +Qed. + +Theorem mul_lt_mono : forall n m p q : N, n < m -> p < q -> n * p < m * q. +Proof. +intros; apply NZmul_lt_mono_nonneg; try assumption; apply le_0_l. +Qed. + +Theorem mul_le_mono : forall n m p q : N, n <= m -> p <= q -> n * p <= m * q. +Proof. +intros; apply NZmul_le_mono_nonneg; try assumption; apply le_0_l. +Qed. + +Theorem lt_0_mul : forall n m : N, n * m > 0 <-> n > 0 /\ m > 0. +Proof. +intros n m; split; [intro H | intros [H1 H2]]. +apply -> NZlt_0_mul in H. destruct H as [[H1 H2] | [H1 H2]]. now split. false_hyp H1 nlt_0_r. +now apply NZmul_pos_pos. +Qed. + +Notation mul_pos := lt_0_mul (only parsing). + +Theorem eq_mul_1 : forall n m : N, n * m == 1 <-> n == 1 /\ m == 1. +Proof. +intros n m. +split; [| intros [H1 H2]; now rewrite H1, H2, mul_1_l]. +intro H; destruct (NZlt_trichotomy n 1) as [H1 | [H1 | H1]]. +apply -> lt_1_r in H1. rewrite H1, mul_0_l in H. false_hyp H neq_0_succ. +rewrite H1, mul_1_l in H; now split. +destruct (eq_0_gt_0_cases m) as [H2 | H2]. +rewrite H2, mul_0_r in H; false_hyp H neq_0_succ. +apply -> (mul_lt_mono_pos_r m) in H1; [| assumption]. rewrite mul_1_l in H1. +assert (H3 : 1 < n * m) by now apply (lt_1_l 0 m). +rewrite H in H3; false_hyp H3 lt_irrefl. +Qed. + +End NMulOrderPropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NOrder.v b/theories/Numbers/Natural/Abstract/NOrder.v new file mode 100644 index 00000000..826ffa2c --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NOrder.v @@ -0,0 +1,539 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NOrder.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NMul. + +Module NOrderPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NMulPropMod := NMulPropFunct NAxiomsMod. +Open Local Scope NatScope. + +(* The tactics le_less, le_equal and le_elim are inherited from NZOrder.v *) + +(* Axioms *) + +Theorem lt_wd : + forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> (n1 < m1 <-> n2 < m2). +Proof NZlt_wd. + +Theorem le_wd : + forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> (n1 <= m1 <-> n2 <= m2). +Proof NZle_wd. + +Theorem min_wd : + forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> min n1 m1 == min n2 m2. +Proof NZmin_wd. + +Theorem max_wd : + forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> max n1 m1 == max n2 m2. +Proof NZmax_wd. + +Theorem lt_eq_cases : forall n m : N, n <= m <-> n < m \/ n == m. +Proof NZlt_eq_cases. + +Theorem lt_irrefl : forall n : N, ~ n < n. +Proof NZlt_irrefl. + +Theorem lt_succ_r : forall n m : N, n < S m <-> n <= m. +Proof NZlt_succ_r. + +Theorem min_l : forall n m : N, n <= m -> min n m == n. +Proof NZmin_l. + +Theorem min_r : forall n m : N, m <= n -> min n m == m. +Proof NZmin_r. + +Theorem max_l : forall n m : N, m <= n -> max n m == n. +Proof NZmax_l. + +Theorem max_r : forall n m : N, n <= m -> max n m == m. +Proof NZmax_r. + +(* Renaming theorems from NZOrder.v *) + +Theorem lt_le_incl : forall n m : N, n < m -> n <= m. +Proof NZlt_le_incl. + +Theorem eq_le_incl : forall n m : N, n == m -> n <= m. +Proof NZeq_le_incl. + +Theorem lt_neq : forall n m : N, n < m -> n ~= m. +Proof NZlt_neq. + +Theorem le_neq : forall n m : N, n < m <-> n <= m /\ n ~= m. +Proof NZle_neq. + +Theorem le_refl : forall n : N, n <= n. +Proof NZle_refl. + +Theorem lt_succ_diag_r : forall n : N, n < S n. +Proof NZlt_succ_diag_r. + +Theorem le_succ_diag_r : forall n : N, n <= S n. +Proof NZle_succ_diag_r. + +Theorem lt_0_1 : 0 < 1. +Proof NZlt_0_1. + +Theorem le_0_1 : 0 <= 1. +Proof NZle_0_1. + +Theorem lt_lt_succ_r : forall n m : N, n < m -> n < S m. +Proof NZlt_lt_succ_r. + +Theorem le_le_succ_r : forall n m : N, n <= m -> n <= S m. +Proof NZle_le_succ_r. + +Theorem le_succ_r : forall n m : N, n <= S m <-> n <= m \/ n == S m. +Proof NZle_succ_r. + +Theorem neq_succ_diag_l : forall n : N, S n ~= n. +Proof NZneq_succ_diag_l. + +Theorem neq_succ_diag_r : forall n : N, n ~= S n. +Proof NZneq_succ_diag_r. + +Theorem nlt_succ_diag_l : forall n : N, ~ S n < n. +Proof NZnlt_succ_diag_l. + +Theorem nle_succ_diag_l : forall n : N, ~ S n <= n. +Proof NZnle_succ_diag_l. + +Theorem le_succ_l : forall n m : N, S n <= m <-> n < m. +Proof NZle_succ_l. + +Theorem lt_succ_l : forall n m : N, S n < m -> n < m. +Proof NZlt_succ_l. + +Theorem succ_lt_mono : forall n m : N, n < m <-> S n < S m. +Proof NZsucc_lt_mono. + +Theorem succ_le_mono : forall n m : N, n <= m <-> S n <= S m. +Proof NZsucc_le_mono. + +Theorem lt_asymm : forall n m : N, n < m -> ~ m < n. +Proof NZlt_asymm. + +Notation lt_ngt := lt_asymm (only parsing). + +Theorem lt_trans : forall n m p : N, n < m -> m < p -> n < p. +Proof NZlt_trans. + +Theorem le_trans : forall n m p : N, n <= m -> m <= p -> n <= p. +Proof NZle_trans. + +Theorem le_lt_trans : forall n m p : N, n <= m -> m < p -> n < p. +Proof NZle_lt_trans. + +Theorem lt_le_trans : forall n m p : N, n < m -> m <= p -> n < p. +Proof NZlt_le_trans. + +Theorem le_antisymm : forall n m : N, n <= m -> m <= n -> n == m. +Proof NZle_antisymm. + +(** Trichotomy, decidability, and double negation elimination *) + +Theorem lt_trichotomy : forall n m : N, n < m \/ n == m \/ m < n. +Proof NZlt_trichotomy. + +Notation lt_eq_gt_cases := lt_trichotomy (only parsing). + +Theorem lt_gt_cases : forall n m : N, n ~= m <-> n < m \/ n > m. +Proof NZlt_gt_cases. + +Theorem le_gt_cases : forall n m : N, n <= m \/ n > m. +Proof NZle_gt_cases. + +Theorem lt_ge_cases : forall n m : N, n < m \/ n >= m. +Proof NZlt_ge_cases. + +Theorem le_ge_cases : forall n m : N, n <= m \/ n >= m. +Proof NZle_ge_cases. + +Theorem le_ngt : forall n m : N, n <= m <-> ~ n > m. +Proof NZle_ngt. + +Theorem nlt_ge : forall n m : N, ~ n < m <-> n >= m. +Proof NZnlt_ge. + +Theorem lt_dec : forall n m : N, decidable (n < m). +Proof NZlt_dec. + +Theorem lt_dne : forall n m : N, ~ ~ n < m <-> n < m. +Proof NZlt_dne. + +Theorem nle_gt : forall n m : N, ~ n <= m <-> n > m. +Proof NZnle_gt. + +Theorem lt_nge : forall n m : N, n < m <-> ~ n >= m. +Proof NZlt_nge. + +Theorem le_dec : forall n m : N, decidable (n <= m). +Proof NZle_dec. + +Theorem le_dne : forall n m : N, ~ ~ n <= m <-> n <= m. +Proof NZle_dne. + +Theorem nlt_succ_r : forall n m : N, ~ m < S n <-> n < m. +Proof NZnlt_succ_r. + +Theorem lt_exists_pred : + forall z n : N, z < n -> exists k : N, n == S k /\ z <= k. +Proof NZlt_exists_pred. + +Theorem lt_succ_iter_r : + forall (n : nat) (m : N), m < NZsucc_iter (Datatypes.S n) m. +Proof NZlt_succ_iter_r. + +Theorem neq_succ_iter_l : + forall (n : nat) (m : N), NZsucc_iter (Datatypes.S n) m ~= m. +Proof NZneq_succ_iter_l. + +(** Stronger variant of induction with assumptions n >= 0 (n < 0) +in the induction step *) + +Theorem right_induction : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, A z -> + (forall n : N, z <= n -> A n -> A (S n)) -> + forall n : N, z <= n -> A n. +Proof NZright_induction. + +Theorem left_induction : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, A z -> + (forall n : N, n < z -> A (S n) -> A n) -> + forall n : N, n <= z -> A n. +Proof NZleft_induction. + +Theorem right_induction' : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, + (forall n : N, n <= z -> A n) -> + (forall n : N, z <= n -> A n -> A (S n)) -> + forall n : N, A n. +Proof NZright_induction'. + +Theorem left_induction' : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, + (forall n : N, z <= n -> A n) -> + (forall n : N, n < z -> A (S n) -> A n) -> + forall n : N, A n. +Proof NZleft_induction'. + +Theorem strong_right_induction : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, + (forall n : N, z <= n -> (forall m : N, z <= m -> m < n -> A m) -> A n) -> + forall n : N, z <= n -> A n. +Proof NZstrong_right_induction. + +Theorem strong_left_induction : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, + (forall n : N, n <= z -> (forall m : N, m <= z -> S n <= m -> A m) -> A n) -> + forall n : N, n <= z -> A n. +Proof NZstrong_left_induction. + +Theorem strong_right_induction' : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, + (forall n : N, n <= z -> A n) -> + (forall n : N, z <= n -> (forall m : N, z <= m -> m < n -> A m) -> A n) -> + forall n : N, A n. +Proof NZstrong_right_induction'. + +Theorem strong_left_induction' : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, + (forall n : N, z <= n -> A n) -> + (forall n : N, n <= z -> (forall m : N, m <= z -> S n <= m -> A m) -> A n) -> + forall n : N, A n. +Proof NZstrong_left_induction'. + +Theorem order_induction : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, A z -> + (forall n : N, z <= n -> A n -> A (S n)) -> + (forall n : N, n < z -> A (S n) -> A n) -> + forall n : N, A n. +Proof NZorder_induction. + +Theorem order_induction' : + forall A : N -> Prop, predicate_wd Neq A -> + forall z : N, A z -> + (forall n : N, z <= n -> A n -> A (S n)) -> + (forall n : N, n <= z -> A n -> A (P n)) -> + forall n : N, A n. +Proof NZorder_induction'. + +(* We don't need order_induction_0 and order_induction'_0 (see NZOrder and +ZOrder) since they boil down to regular induction *) + +(** Elimintation principle for < *) + +Theorem lt_ind : + forall A : N -> Prop, predicate_wd Neq A -> + forall n : N, + A (S n) -> + (forall m : N, n < m -> A m -> A (S m)) -> + forall m : N, n < m -> A m. +Proof NZlt_ind. + +(** Elimintation principle for <= *) + +Theorem le_ind : + forall A : N -> Prop, predicate_wd Neq A -> + forall n : N, + A n -> + (forall m : N, n <= m -> A m -> A (S m)) -> + forall m : N, n <= m -> A m. +Proof NZle_ind. + +(** Well-founded relations *) + +Theorem lt_wf : forall z : N, well_founded (fun n m : N => z <= n /\ n < m). +Proof NZlt_wf. + +Theorem gt_wf : forall z : N, well_founded (fun n m : N => m < n /\ n <= z). +Proof NZgt_wf. + +Theorem lt_wf_0 : well_founded lt. +Proof. +assert (H : relations_eq lt (fun n m : N => 0 <= n /\ n < m)). +intros x y; split. +intro H; split; [apply le_0_l | assumption]. now intros [_ H]. +rewrite H; apply lt_wf. +(* does not work: +setoid_replace lt with (fun n m : N => 0 <= n /\ n < m) using relation relations_eq.*) +Qed. + +(* Theorems that are true for natural numbers but not for integers *) + +(* "le_0_l : forall n : N, 0 <= n" was proved in NBase.v *) + +Theorem nlt_0_r : forall n : N, ~ n < 0. +Proof. +intro n; apply -> le_ngt. apply le_0_l. +Qed. + +Theorem nle_succ_0 : forall n : N, ~ (S n <= 0). +Proof. +intros n H; apply -> le_succ_l in H; false_hyp H nlt_0_r. +Qed. + +Theorem le_0_r : forall n : N, n <= 0 <-> n == 0. +Proof. +intros n; split; intro H. +le_elim H; [false_hyp H nlt_0_r | assumption]. +now apply eq_le_incl. +Qed. + +Theorem lt_0_succ : forall n : N, 0 < S n. +Proof. +induct n; [apply lt_succ_diag_r | intros n H; now apply lt_lt_succ_r]. +Qed. + +Theorem neq_0_lt_0 : forall n : N, n ~= 0 <-> 0 < n. +Proof. +cases n. +split; intro H; [now elim H | intro; now apply lt_irrefl with 0]. +intro n; split; intro H; [apply lt_0_succ | apply neq_succ_0]. +Qed. + +Theorem eq_0_gt_0_cases : forall n : N, n == 0 \/ 0 < n. +Proof. +cases n. +now left. +intro; right; apply lt_0_succ. +Qed. + +Theorem zero_one : forall n : N, n == 0 \/ n == 1 \/ 1 < n. +Proof. +induct n. now left. +cases n. intros; right; now left. +intros n IH. destruct IH as [H | [H | H]]. +false_hyp H neq_succ_0. +right; right. rewrite H. apply lt_succ_diag_r. +right; right. now apply lt_lt_succ_r. +Qed. + +Theorem lt_1_r : forall n : N, n < 1 <-> n == 0. +Proof. +cases n. +split; intro; [reflexivity | apply lt_succ_diag_r]. +intros n. rewrite <- succ_lt_mono. +split; intro H; [false_hyp H nlt_0_r | false_hyp H neq_succ_0]. +Qed. + +Theorem le_1_r : forall n : N, n <= 1 <-> n == 0 \/ n == 1. +Proof. +cases n. +split; intro; [now left | apply le_succ_diag_r]. +intro n. rewrite <- succ_le_mono, le_0_r, succ_inj_wd. +split; [intro; now right | intros [H | H]; [false_hyp H neq_succ_0 | assumption]]. +Qed. + +Theorem lt_lt_0 : forall n m : N, n < m -> 0 < m. +Proof. +intros n m; induct n. +trivial. +intros n IH H. apply IH; now apply lt_succ_l. +Qed. + +Theorem lt_1_l : forall n m p : N, n < m -> m < p -> 1 < p. +Proof. +intros n m p H1 H2. +apply le_lt_trans with m. apply <- le_succ_l. apply le_lt_trans with n. +apply le_0_l. assumption. assumption. +Qed. + +(** Elimination principlies for < and <= for relations *) + +Section RelElim. + +(* FIXME: Variable R : relation N. -- does not work *) + +Variable R : N -> N -> Prop. +Hypothesis R_wd : relation_wd Neq Neq R. + +Add Morphism R with signature Neq ==> Neq ==> iff as R_morph2. +Proof. apply R_wd. Qed. + +Theorem le_ind_rel : + (forall m : N, R 0 m) -> + (forall n m : N, n <= m -> R n m -> R (S n) (S m)) -> + forall n m : N, n <= m -> R n m. +Proof. +intros Base Step; induct n. +intros; apply Base. +intros n IH m H. elim H using le_ind. +solve_predicate_wd. +apply Step; [| apply IH]; now apply eq_le_incl. +intros k H1 H2. apply -> le_succ_l in H1. apply lt_le_incl in H1. auto. +Qed. + +Theorem lt_ind_rel : + (forall m : N, R 0 (S m)) -> + (forall n m : N, n < m -> R n m -> R (S n) (S m)) -> + forall n m : N, n < m -> R n m. +Proof. +intros Base Step; induct n. +intros m H. apply lt_exists_pred in H; destruct H as [m' [H _]]. +rewrite H; apply Base. +intros n IH m H. elim H using lt_ind. +solve_predicate_wd. +apply Step; [| apply IH]; now apply lt_succ_diag_r. +intros k H1 H2. apply lt_succ_l in H1. auto. +Qed. + +End RelElim. + +(** Predecessor and order *) + +Theorem succ_pred_pos : forall n : N, 0 < n -> S (P n) == n. +Proof. +intros n H; apply succ_pred; intro H1; rewrite H1 in H. +false_hyp H lt_irrefl. +Qed. + +Theorem le_pred_l : forall n : N, P n <= n. +Proof. +cases n. +rewrite pred_0; now apply eq_le_incl. +intros; rewrite pred_succ; apply le_succ_diag_r. +Qed. + +Theorem lt_pred_l : forall n : N, n ~= 0 -> P n < n. +Proof. +cases n. +intro H; elimtype False; now apply H. +intros; rewrite pred_succ; apply lt_succ_diag_r. +Qed. + +Theorem le_le_pred : forall n m : N, n <= m -> P n <= m. +Proof. +intros n m H; apply le_trans with n. apply le_pred_l. assumption. +Qed. + +Theorem lt_lt_pred : forall n m : N, n < m -> P n < m. +Proof. +intros n m H; apply le_lt_trans with n. apply le_pred_l. assumption. +Qed. + +Theorem lt_le_pred : forall n m : N, n < m -> n <= P m. (* Converse is false for n == m == 0 *) +Proof. +intro n; cases m. +intro H; false_hyp H nlt_0_r. +intros m IH. rewrite pred_succ; now apply -> lt_succ_r. +Qed. + +Theorem lt_pred_le : forall n m : N, P n < m -> n <= m. (* Converse is false for n == m == 0 *) +Proof. +intros n m; cases n. +rewrite pred_0; intro H; now apply lt_le_incl. +intros n IH. rewrite pred_succ in IH. now apply <- le_succ_l. +Qed. + +Theorem lt_pred_lt : forall n m : N, n < P m -> n < m. +Proof. +intros n m H; apply lt_le_trans with (P m); [assumption | apply le_pred_l]. +Qed. + +Theorem le_pred_le : forall n m : N, n <= P m -> n <= m. +Proof. +intros n m H; apply le_trans with (P m); [assumption | apply le_pred_l]. +Qed. + +Theorem pred_le_mono : forall n m : N, n <= m -> P n <= P m. (* Converse is false for n == 1, m == 0 *) +Proof. +intros n m H; elim H using le_ind_rel. +solve_relation_wd. +intro; rewrite pred_0; apply le_0_l. +intros p q H1 _; now do 2 rewrite pred_succ. +Qed. + +Theorem pred_lt_mono : forall n m : N, n ~= 0 -> (n < m <-> P n < P m). +Proof. +intros n m H1; split; intro H2. +assert (m ~= 0). apply <- neq_0_lt_0. now apply lt_lt_0 with n. +now rewrite <- (succ_pred n) in H2; rewrite <- (succ_pred m) in H2 ; +[apply <- succ_lt_mono | | |]. +assert (m ~= 0). apply <- neq_0_lt_0. apply lt_lt_0 with (P n). +apply lt_le_trans with (P m). assumption. apply le_pred_l. +apply -> succ_lt_mono in H2. now do 2 rewrite succ_pred in H2. +Qed. + +Theorem lt_succ_lt_pred : forall n m : N, S n < m <-> n < P m. +Proof. +intros n m. rewrite pred_lt_mono by apply neq_succ_0. now rewrite pred_succ. +Qed. + +Theorem le_succ_le_pred : forall n m : N, S n <= m -> n <= P m. (* Converse is false for n == m == 0 *) +Proof. +intros n m H. apply lt_le_pred. now apply -> le_succ_l. +Qed. + +Theorem lt_pred_lt_succ : forall n m : N, P n < m -> n < S m. (* Converse is false for n == m == 0 *) +Proof. +intros n m H. apply <- lt_succ_r. now apply lt_pred_le. +Qed. + +Theorem le_pred_le_succ : forall n m : N, P n <= m <-> n <= S m. +Proof. +intros n m; cases n. +rewrite pred_0. split; intro H; apply le_0_l. +intro n. rewrite pred_succ. apply succ_le_mono. +Qed. + +End NOrderPropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NStrongRec.v b/theories/Numbers/Natural/Abstract/NStrongRec.v new file mode 100644 index 00000000..031dbdea --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NStrongRec.v @@ -0,0 +1,133 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NStrongRec.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +(** This file defined the strong (course-of-value, well-founded) recursion +and proves its properties *) + +Require Export NSub. + +Module NStrongRecPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NSubPropMod := NSubPropFunct NAxiomsMod. +Open Local Scope NatScope. + +Section StrongRecursion. + +Variable A : Set. +Variable Aeq : relation A. + +Notation Local "x ==A y" := (Aeq x y) (at level 70, no associativity). + +Hypothesis Aeq_equiv : equiv A Aeq. + +Add Relation A Aeq + reflexivity proved by (proj1 Aeq_equiv) + symmetry proved by (proj2 (proj2 Aeq_equiv)) + transitivity proved by (proj1 (proj2 Aeq_equiv)) +as Aeq_rel. + +Definition strong_rec (a : A) (f : N -> (N -> A) -> A) (n : N) : A := +recursion + (fun _ : N => a) + (fun (m : N) (p : N -> A) (k : N) => f k p) + (S n) + n. + +Theorem strong_rec_wd : +forall a a' : A, a ==A a' -> + forall f f', fun2_eq Neq (fun_eq Neq Aeq) Aeq f f' -> + forall n n', n == n' -> + strong_rec a f n ==A strong_rec a' f' n'. +Proof. +intros a a' Eaa' f f' Eff' n n' Enn'. +(* First we prove that recursion (which is on type N -> A) returns +extensionally equal functions, and then we use the fact that n == n' *) +assert (H : fun_eq Neq Aeq + (recursion + (fun _ : N => a) + (fun (m : N) (p : N -> A) (k : N) => f k p) + (S n)) + (recursion + (fun _ : N => a') + (fun (m : N) (p : N -> A) (k : N) => f' k p) + (S n'))). +apply recursion_wd with (Aeq := fun_eq Neq Aeq). +unfold fun_eq; now intros. +unfold fun2_eq. intros y y' Eyy' p p' Epp'. unfold fun_eq. auto. +now rewrite Enn'. +unfold strong_rec. +now apply H. +Qed. + +(*Section FixPoint. + +Variable a : A. +Variable f : N -> (N -> A) -> A. + +Hypothesis f_wd : fun2_wd Neq (fun_eq Neq Aeq) Aeq f. + +Let g (n : N) : A := strong_rec a f n. + +Add Morphism g with signature Neq ==> Aeq as g_wd. +Proof. +intros n1 n2 H. unfold g. now apply strong_rec_wd. +Qed. + +Theorem NtoA_eq_symm : symmetric (N -> A) (fun_eq Neq Aeq). +Proof. +apply fun_eq_symm. +exact (proj2 (proj2 NZeq_equiv)). +exact (proj2 (proj2 Aeq_equiv)). +Qed. + +Theorem NtoA_eq_trans : transitive (N -> A) (fun_eq Neq Aeq). +Proof. +apply fun_eq_trans. +exact (proj1 NZeq_equiv). +exact (proj1 (proj2 NZeq_equiv)). +exact (proj1 (proj2 Aeq_equiv)). +Qed. + +Add Relation (N -> A) (fun_eq Neq Aeq) + symmetry proved by NtoA_eq_symm + transitivity proved by NtoA_eq_trans +as NtoA_eq_rel. + +Add Morphism f with signature Neq ==> (fun_eq Neq Aeq) ==> Aeq as f_morph. +Proof. +apply f_wd. +Qed. + +(* We need an assumption saying that for every n, the step function (f n h) +calls h only on the segment [0 ... n - 1]. This means that if h1 and h2 +coincide on values < n, then (f n h1) coincides with (f n h2) *) + +Hypothesis step_good : + forall (n : N) (h1 h2 : N -> A), + (forall m : N, m < n -> Aeq (h1 m) (h2 m)) -> Aeq (f n h1) (f n h2). + +(* Todo: +Theorem strong_rec_fixpoint : forall n : N, Aeq (g n) (f n g). +Proof. +apply induction. +unfold predicate_wd, fun_wd. +intros x y H. rewrite H. unfold fun_eq; apply g_wd. +reflexivity. +unfold g, strong_rec. +*) + +End FixPoint.*) +End StrongRecursion. + +Implicit Arguments strong_rec [A]. + +End NStrongRecPropFunct. + diff --git a/theories/Numbers/Natural/Abstract/NSub.v b/theories/Numbers/Natural/Abstract/NSub.v new file mode 100644 index 00000000..f67689dd --- /dev/null +++ b/theories/Numbers/Natural/Abstract/NSub.v @@ -0,0 +1,180 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NSub.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Export NMulOrder. + +Module NSubPropFunct (Import NAxiomsMod : NAxiomsSig). +Module Export NMulOrderPropMod := NMulOrderPropFunct NAxiomsMod. +Open Local Scope NatScope. + +Theorem sub_wd : + forall n1 n2 : N, n1 == n2 -> forall m1 m2 : N, m1 == m2 -> n1 - m1 == n2 - m2. +Proof NZsub_wd. + +Theorem sub_0_r : forall n : N, n - 0 == n. +Proof NZsub_0_r. + +Theorem sub_succ_r : forall n m : N, n - (S m) == P (n - m). +Proof NZsub_succ_r. + +Theorem sub_1_r : forall n : N, n - 1 == P n. +Proof. +intro n; rewrite sub_succ_r; now rewrite sub_0_r. +Qed. + +Theorem sub_0_l : forall n : N, 0 - n == 0. +Proof. +induct n. +apply sub_0_r. +intros n IH; rewrite sub_succ_r; rewrite IH. now apply pred_0. +Qed. + +Theorem sub_succ : forall n m : N, S n - S m == n - m. +Proof. +intro n; induct m. +rewrite sub_succ_r. do 2 rewrite sub_0_r. now rewrite pred_succ. +intros m IH. rewrite sub_succ_r. rewrite IH. now rewrite sub_succ_r. +Qed. + +Theorem sub_diag : forall n : N, n - n == 0. +Proof. +induct n. apply sub_0_r. intros n IH; rewrite sub_succ; now rewrite IH. +Qed. + +Theorem sub_gt : forall n m : N, n > m -> n - m ~= 0. +Proof. +intros n m H; elim H using lt_ind_rel; clear n m H. +solve_relation_wd. +intro; rewrite sub_0_r; apply neq_succ_0. +intros; now rewrite sub_succ. +Qed. + +Theorem add_sub_assoc : forall n m p : N, p <= m -> n + (m - p) == (n + m) - p. +Proof. +intros n m p; induct p. +intro; now do 2 rewrite sub_0_r. +intros p IH H. do 2 rewrite sub_succ_r. +rewrite <- IH by (apply lt_le_incl; now apply -> le_succ_l). +rewrite add_pred_r by (apply sub_gt; now apply -> le_succ_l). +reflexivity. +Qed. + +Theorem sub_succ_l : forall n m : N, n <= m -> S m - n == S (m - n). +Proof. +intros n m H. rewrite <- (add_1_l m). rewrite <- (add_1_l (m - n)). +symmetry; now apply add_sub_assoc. +Qed. + +Theorem add_sub : forall n m : N, (n + m) - m == n. +Proof. +intros n m. rewrite <- add_sub_assoc by (apply le_refl). +rewrite sub_diag; now rewrite add_0_r. +Qed. + +Theorem sub_add : forall n m : N, n <= m -> (m - n) + n == m. +Proof. +intros n m H. rewrite add_comm. rewrite add_sub_assoc by assumption. +rewrite add_comm. apply add_sub. +Qed. + +Theorem add_sub_eq_l : forall n m p : N, m + p == n -> n - m == p. +Proof. +intros n m p H. symmetry. +assert (H1 : m + p - m == n - m) by now rewrite H. +rewrite add_comm in H1. now rewrite add_sub in H1. +Qed. + +Theorem add_sub_eq_r : forall n m p : N, m + p == n -> n - p == m. +Proof. +intros n m p H; rewrite add_comm in H; now apply add_sub_eq_l. +Qed. + +(* This could be proved by adding m to both sides. Then the proof would +use add_sub_assoc and sub_0_le, which is proven below. *) + +Theorem add_sub_eq_nz : forall n m p : N, p ~= 0 -> n - m == p -> m + p == n. +Proof. +intros n m p H; double_induct n m. +intros m H1; rewrite sub_0_l in H1. symmetry in H1; false_hyp H1 H. +intro n; rewrite sub_0_r; now rewrite add_0_l. +intros n m IH H1. rewrite sub_succ in H1. apply IH in H1. +rewrite add_succ_l; now rewrite H1. +Qed. + +Theorem sub_add_distr : forall n m p : N, n - (m + p) == (n - m) - p. +Proof. +intros n m; induct p. +rewrite add_0_r; now rewrite sub_0_r. +intros p IH. rewrite add_succ_r; do 2 rewrite sub_succ_r. now rewrite IH. +Qed. + +Theorem add_sub_swap : forall n m p : N, p <= n -> n + m - p == n - p + m. +Proof. +intros n m p H. +rewrite (add_comm n m). +rewrite <- add_sub_assoc by assumption. +now rewrite (add_comm m (n - p)). +Qed. + +(** Sub and order *) + +Theorem le_sub_l : forall n m : N, n - m <= n. +Proof. +intro n; induct m. +rewrite sub_0_r; now apply eq_le_incl. +intros m IH. rewrite sub_succ_r. +apply le_trans with (n - m); [apply le_pred_l | assumption]. +Qed. + +Theorem sub_0_le : forall n m : N, n - m == 0 <-> n <= m. +Proof. +double_induct n m. +intro m; split; intro; [apply le_0_l | apply sub_0_l]. +intro m; rewrite sub_0_r; split; intro H; +[false_hyp H neq_succ_0 | false_hyp H nle_succ_0]. +intros n m H. rewrite <- succ_le_mono. now rewrite sub_succ. +Qed. + +(** Sub and mul *) + +Theorem mul_pred_r : forall n m : N, n * (P m) == n * m - n. +Proof. +intros n m; cases m. +now rewrite pred_0, mul_0_r, sub_0_l. +intro m; rewrite pred_succ, mul_succ_r, <- add_sub_assoc. +now rewrite sub_diag, add_0_r. +now apply eq_le_incl. +Qed. + +Theorem mul_sub_distr_r : forall n m p : N, (n - m) * p == n * p - m * p. +Proof. +intros n m p; induct n. +now rewrite sub_0_l, mul_0_l, sub_0_l. +intros n IH. destruct (le_gt_cases m n) as [H | H]. +rewrite sub_succ_l by assumption. do 2 rewrite mul_succ_l. +rewrite (add_comm ((n - m) * p) p), (add_comm (n * p) p). +rewrite <- (add_sub_assoc p (n * p) (m * p)) by now apply mul_le_mono_r. +now apply <- add_cancel_l. +assert (H1 : S n <= m); [now apply <- le_succ_l |]. +setoid_replace (S n - m) with 0 by now apply <- sub_0_le. +setoid_replace ((S n * p) - m * p) with 0 by (apply <- sub_0_le; now apply mul_le_mono_r). +apply mul_0_l. +Qed. + +Theorem mul_sub_distr_l : forall n m p : N, p * (n - m) == p * n - p * m. +Proof. +intros n m p; rewrite (mul_comm p (n - m)), (mul_comm p n), (mul_comm p m). +apply mul_sub_distr_r. +Qed. + +End NSubPropFunct. + diff --git a/theories/Numbers/Natural/BigN/BigN.v b/theories/Numbers/Natural/BigN/BigN.v new file mode 100644 index 00000000..0574c09f --- /dev/null +++ b/theories/Numbers/Natural/BigN/BigN.v @@ -0,0 +1,83 @@ +(************************************************************************) +(* 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: BigN.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +(** * Natural numbers in base 2^31 *) + +(** +Author: Arnaud Spiwack +*) + +Require Export Int31. +Require Import CyclicAxioms. +Require Import Cyclic31. +Require Import NSig. +Require Import NSigNAxioms. +Require Import NMake. +Require Import NSub. + +Module BigN <: NType := NMake.Make Int31Cyclic. + +(** Module [BigN] implements [NAxiomsSig] *) + +Module Export BigNAxiomsMod := NSig_NAxioms BigN. +Module Export BigNSubPropMod := NSubPropFunct BigNAxiomsMod. + +(** Notations about [BigN] *) + +Notation bigN := BigN.t. + +Delimit Scope bigN_scope with bigN. +Bind Scope bigN_scope with bigN. +Bind Scope bigN_scope with BigN.t. +Bind Scope bigN_scope with BigN.t_. + +Notation Local "0" := BigN.zero : bigN_scope. (* temporary notation *) +Infix "+" := BigN.add : bigN_scope. +Infix "-" := BigN.sub : bigN_scope. +Infix "*" := BigN.mul : bigN_scope. +Infix "/" := BigN.div : bigN_scope. +Infix "?=" := BigN.compare : bigN_scope. +Infix "==" := BigN.eq (at level 70, no associativity) : bigN_scope. +Infix "<" := BigN.lt : bigN_scope. +Infix "<=" := BigN.le : bigN_scope. +Notation "[ i ]" := (BigN.to_Z i) : bigN_scope. + +Open Scope bigN_scope. + +(** Example of reasoning about [BigN] *) + +Theorem succ_pred: forall q:bigN, + 0 < q -> BigN.succ (BigN.pred q) == q. +Proof. +intros; apply succ_pred. +intro H'; rewrite H' in H; discriminate. +Qed. + +(** [BigN] is a semi-ring *) + +Lemma BigNring : + semi_ring_theory BigN.zero BigN.one BigN.add BigN.mul BigN.eq. +Proof. +constructor. +exact add_0_l. +exact add_comm. +exact add_assoc. +exact mul_1_l. +exact mul_0_l. +exact mul_comm. +exact mul_assoc. +exact mul_add_distr_r. +Qed. + +Add Ring BigNr : BigNring. + +(** Todo: tactic translating from [BigN] to [Z] + omega *) + +(** Todo: micromega *) diff --git a/theories/Numbers/Natural/BigN/NMake_gen.ml b/theories/Numbers/Natural/BigN/NMake_gen.ml new file mode 100644 index 00000000..bd0fb5b1 --- /dev/null +++ b/theories/Numbers/Natural/BigN/NMake_gen.ml @@ -0,0 +1,3166 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NMake_gen.ml 11136 2008-06-18 10:41:34Z herbelin $ i*) + +(*S NMake_gen.ml : this file generates NMake.v *) + + +(*s The two parameters that control the generation: *) + +let size = 6 (* how many times should we repeat the Z/nZ --> Z/2nZ + process before relying on a generic construct *) +let gen_proof = true (* should we generate proofs ? *) + + +(*s Some utilities *) + +let t = "t" +let c = "N" +let pz n = if n == 0 then "w_0" else "W0" +let rec gen2 n = if n == 0 then "1" else if n == 1 then "2" + else "2 * " ^ (gen2 (n - 1)) +let rec genxO n s = + if n == 0 then s else " (xO" ^ (genxO (n - 1) s) ^ ")" + +(* NB: in ocaml >= 3.10, we could use Printf.ifprintf for printing to + /dev/null, but for being compatible with earlier ocaml and not + relying on system-dependent stuff like open_out "/dev/null", + let's use instead a magical hack *) + +(* Standard printer, with a final newline *) +let pr s = Printf.printf (s^^"\n") +(* Printing to /dev/null *) +let pn = (fun s -> Obj.magic (fun _ _ _ _ _ _ _ _ _ _ _ _ _ _ -> ()) + : ('a, out_channel, unit) format -> 'a) +(* Proof printer : prints iff gen_proof is true *) +let pp = if gen_proof then pr else pn +(* Printer for admitted parts : prints iff gen_proof is false *) +let pa = if not gen_proof then pr else pn +(* Same as before, but without the final newline *) +let pr0 = Printf.printf +let pp0 = if gen_proof then pr0 else pn + + +(*s The actual printing *) + +let _ = + + pr "(************************************************************************)"; + pr "(* v * The Coq Proof Assistant / The Coq Development Team *)"; + pr "(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)"; + pr "(* \\VV/ **************************************************************)"; + pr "(* // * This file is distributed under the terms of the *)"; + pr "(* * GNU Lesser General Public License Version 2.1 *)"; + pr "(************************************************************************)"; + pr "(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)"; + pr "(************************************************************************)"; + pr ""; + pr "(** * NMake *)"; + pr ""; + pr "(** From a cyclic Z/nZ representation to arbitrary precision natural numbers.*)"; + pr ""; + pr "(** Remark: File automatically generated by NMake_gen.ml, DO NOT EDIT ! *)"; + pr ""; + pr "Require Import BigNumPrelude."; + pr "Require Import ZArith."; + pr "Require Import CyclicAxioms."; + pr "Require Import DoubleType."; + pr "Require Import DoubleMul."; + pr "Require Import DoubleDivn1."; + pr "Require Import DoubleCyclic."; + pr "Require Import Nbasic."; + pr "Require Import Wf_nat."; + pr "Require Import StreamMemo."; + pr "Require Import NSig."; + pr ""; + pr "Module Make (Import W0:CyclicType) <: NType."; + pr ""; + + pr " Definition w0 := W0.w."; + for i = 1 to size do + pr " Definition w%i := zn2z w%i." i (i-1) + done; + pr ""; + + pr " Definition w0_op := W0.w_op."; + for i = 1 to 3 do + pr " Definition w%i_op := mk_zn2z_op w%i_op." i (i-1) + done; + for i = 4 to size + 3 do + pr " Definition w%i_op := mk_zn2z_op_karatsuba w%i_op." i (i-1) + done; + pr ""; + + pr " Section Make_op."; + pr " Variable mk : forall w', znz_op w' -> znz_op (zn2z w')."; + pr ""; + pr " Fixpoint make_op_aux (n:nat) : znz_op (word w%i (S n)):=" size; + pr " match n return znz_op (word w%i (S n)) with" size; + pr " | O => w%i_op" (size+1); + pr " | S n1 =>"; + pr " match n1 return znz_op (word w%i (S (S n1))) with" size; + pr " | O => w%i_op" (size+2); + pr " | S n2 =>"; + pr " match n2 return znz_op (word w%i (S (S (S n2)))) with" size; + pr " | O => w%i_op" (size+3); + pr " | S n3 => mk _ (mk _ (mk _ (make_op_aux n3)))"; + pr " end"; + pr " end"; + pr " end."; + pr ""; + pr " End Make_op."; + pr ""; + pr " Definition omake_op := make_op_aux mk_zn2z_op_karatsuba."; + pr ""; + pr ""; + pr " Definition make_op_list := dmemo_list _ omake_op."; + pr ""; + pr " Definition make_op n := dmemo_get _ omake_op n make_op_list."; + pr ""; + pr " Lemma make_op_omake: forall n, make_op n = omake_op n."; + pr " intros n; unfold make_op, make_op_list."; + pr " refine (dmemo_get_correct _ _ _)."; + pr " Qed."; + pr ""; + + pr " Inductive %s_ :=" t; + for i = 0 to size do + pr " | %s%i : w%i -> %s_" c i i t + done; + pr " | %sn : forall n, word w%i (S n) -> %s_." c size t; + pr ""; + pr " Definition %s := %s_." t t; + pr ""; + + pr " Definition w_0 := w0_op.(znz_0)."; + pr ""; + + for i = 0 to size do + pr " Definition one%i := w%i_op.(znz_1)." i i + done; + pr ""; + + + pr " Definition zero := %s0 w_0." c; + pr " Definition one := %s0 one0." c; + pr ""; + + pr " Definition to_Z x :="; + pr " match x with"; + for i = 0 to size do + pr " | %s%i wx => w%i_op.(znz_to_Z) wx" c i i + done; + pr " | %sn n wx => (make_op n).(znz_to_Z) wx" c; + pr " end."; + pr ""; + + pr " Open Scope Z_scope."; + pr " Notation \"[ x ]\" := (to_Z x)."; + pr ""; + + pr " Definition to_N x := Zabs_N (to_Z x)."; + pr ""; + + pr " Definition eq x y := (to_Z x = to_Z y)."; + pr ""; + + pp " (* Regular make op (no karatsuba) *)"; + pp " Fixpoint nmake_op (ww:Type) (ww_op: znz_op ww) (n: nat) : "; + pp " znz_op (word ww n) :="; + pp " match n return znz_op (word ww n) with "; + pp " O => ww_op"; + pp " | S n1 => mk_zn2z_op (nmake_op ww ww_op n1) "; + pp " end."; + pp ""; + pp " (* Simplification by rewriting for nmake_op *)"; + pp " Theorem nmake_op_S: forall ww (w_op: znz_op ww) x, "; + pp " nmake_op _ w_op (S x) = mk_zn2z_op (nmake_op _ w_op x)."; + pp " auto."; + pp " Qed."; + pp ""; + + + pr " (* Eval and extend functions for each level *)"; + for i = 0 to size do + pp " Let nmake_op%i := nmake_op _ w%i_op." i i; + pp " Let eval%in n := znz_to_Z (nmake_op%i n)." i i; + if i == 0 then + pr " Let extend%i := DoubleBase.extend (WW w_0)." i + else + pr " Let extend%i := DoubleBase.extend (WW (W0: w%i))." i i; + done; + pr ""; + + + pp " Theorem digits_doubled:forall n ww (w_op: znz_op ww), "; + pp " znz_digits (nmake_op _ w_op n) = "; + pp " DoubleBase.double_digits (znz_digits w_op) n."; + pp " Proof."; + pp " intros n; elim n; auto; clear n."; + pp " intros n Hrec ww ww_op; simpl DoubleBase.double_digits."; + pp " rewrite <- Hrec; auto."; + pp " Qed."; + pp ""; + pp " Theorem nmake_double: forall n ww (w_op: znz_op ww), "; + pp " znz_to_Z (nmake_op _ w_op n) ="; + pp " @DoubleBase.double_to_Z _ (znz_digits w_op) (znz_to_Z w_op) n."; + pp " Proof."; + pp " intros n; elim n; auto; clear n."; + pp " intros n Hrec ww ww_op; simpl DoubleBase.double_to_Z; unfold zn2z_to_Z."; + pp " rewrite <- Hrec; auto."; + pp " unfold DoubleBase.double_wB; rewrite <- digits_doubled; auto."; + pp " Qed."; + pp ""; + + + pp " Theorem digits_nmake:forall n ww (w_op: znz_op ww), "; + pp " znz_digits (nmake_op _ w_op (S n)) = "; + pp " xO (znz_digits (nmake_op _ w_op n))."; + pp " Proof."; + pp " auto."; + pp " Qed."; + pp ""; + + + pp " Theorem znz_nmake_op: forall ww ww_op n xh xl,"; + pp " znz_to_Z (nmake_op ww ww_op (S n)) (WW xh xl) ="; + pp " znz_to_Z (nmake_op ww ww_op n) xh *"; + pp " base (znz_digits (nmake_op ww ww_op n)) +"; + pp " znz_to_Z (nmake_op ww ww_op n) xl."; + pp " Proof."; + pp " auto."; + pp " Qed."; + pp ""; + + pp " Theorem make_op_S: forall n,"; + pp " make_op (S n) = mk_zn2z_op_karatsuba (make_op n)."; + pp " intro n."; + pp " do 2 rewrite make_op_omake."; + pp " pattern n; apply lt_wf_ind; clear n."; + pp " intros n; case n; clear n."; + pp " intros _; unfold omake_op, make_op_aux, w%i_op; apply refl_equal." (size + 2); + pp " intros n; case n; clear n."; + pp " intros _; unfold omake_op, make_op_aux, w%i_op; apply refl_equal." (size + 3); + pp " intros n; case n; clear n."; + pp " intros _; unfold omake_op, make_op_aux, w%i_op, w%i_op; apply refl_equal." (size + 3) (size + 2); + pp " intros n Hrec."; + pp " change (omake_op (S (S (S (S n))))) with"; + pp " (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (omake_op (S n)))))."; + pp " change (omake_op (S (S (S n)))) with"; + pp " (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (omake_op n))))."; + pp " rewrite Hrec; auto with arith."; + pp " Qed."; + pp " "; + + + for i = 1 to size + 2 do + pp " Let znz_to_Z_%i: forall x y," i; + pp " znz_to_Z w%i_op (WW x y) = " i; + pp " znz_to_Z w%i_op x * base (znz_digits w%i_op) + znz_to_Z w%i_op y." (i-1) (i-1) (i-1); + pp " Proof."; + pp " auto."; + pp " Qed. "; + pp ""; + done; + + pp " Let znz_to_Z_n: forall n x y,"; + pp " znz_to_Z (make_op (S n)) (WW x y) = "; + pp " znz_to_Z (make_op n) x * base (znz_digits (make_op n)) + znz_to_Z (make_op n) y."; + pp " Proof."; + pp " intros n x y; rewrite make_op_S; auto."; + pp " Qed. "; + pp ""; + + pp " Let w0_spec: znz_spec w0_op := W0.w_spec."; + for i = 1 to 3 do + pp " Let w%i_spec: znz_spec w%i_op := mk_znz2_spec w%i_spec." i i (i-1) + done; + for i = 4 to size + 3 do + pp " Let w%i_spec : znz_spec w%i_op := mk_znz2_karatsuba_spec w%i_spec." i i (i-1) + done; + pp ""; + + pp " Let wn_spec: forall n, znz_spec (make_op n)."; + pp " intros n; elim n; clear n."; + pp " exact w%i_spec." (size + 1); + pp " intros n Hrec; rewrite make_op_S."; + pp " exact (mk_znz2_karatsuba_spec Hrec)."; + pp " Qed."; + pp ""; + + for i = 0 to size do + pr " Definition w%i_eq0 := w%i_op.(znz_eq0)." i i; + pr " Let spec_w%i_eq0: forall x, if w%i_eq0 x then [%s%i x] = 0 else True." i i c i; + pa " Admitted."; + pp " Proof."; + pp " intros x; unfold w%i_eq0, to_Z; generalize (spec_eq0 w%i_spec x);" i i; + pp " case znz_eq0; auto."; + pp " Qed."; + pr ""; + done; + pr ""; + + + for i = 0 to size do + pp " Theorem digits_w%i: znz_digits w%i_op = znz_digits (nmake_op _ w0_op %i)." i i i; + if i == 0 then + pp " auto." + else + pp " rewrite digits_nmake; rewrite <- digits_w%i; auto." (i - 1); + pp " Qed."; + pp ""; + pp " Let spec_double_eval%in: forall n, eval%in n = DoubleBase.double_to_Z (znz_digits w%i_op) (znz_to_Z w%i_op) n." i i i i; + pp " Proof."; + pp " intros n; exact (nmake_double n w%i w%i_op)." i i; + pp " Qed."; + pp ""; + done; + + for i = 0 to size do + for j = 0 to (size - i) do + pp " Theorem digits_w%in%i: znz_digits w%i_op = znz_digits (nmake_op _ w%i_op %i)." i j (i + j) i j; + pp " Proof."; + if j == 0 then + if i == 0 then + pp " auto." + else + begin + pp " apply trans_equal with (xO (znz_digits w%i_op))." (i + j -1); + pp " auto."; + pp " unfold nmake_op; auto."; + end + else + begin + pp " apply trans_equal with (xO (znz_digits w%i_op))." (i + j -1); + pp " auto."; + pp " rewrite digits_nmake."; + pp " rewrite digits_w%in%i." i (j - 1); + pp " auto."; + end; + pp " Qed."; + pp ""; + pp " Let spec_eval%in%i: forall x, [%s%i x] = eval%in %i x." i j c (i + j) i j; + pp " Proof."; + if j == 0 then + pp " intros x; rewrite spec_double_eval%in; unfold DoubleBase.double_to_Z, to_Z; auto." i + else + begin + pp " intros x; case x."; + pp " auto."; + pp " intros xh xl; unfold to_Z; rewrite znz_to_Z_%i." (i + j); + pp " rewrite digits_w%in%i." i (j - 1); + pp " generalize (spec_eval%in%i); unfold to_Z; intros HH; repeat rewrite HH." i (j - 1); + pp " unfold eval%in, nmake_op%i." i i; + pp " rewrite (znz_nmake_op _ w%i_op %i); auto." i (j - 1); + end; + pp " Qed."; + if i + j <> size then + begin + pp " Let spec_extend%in%i: forall x, [%s%i x] = [%s%i (extend%i %i x)]." i (i + j + 1) c i c (i + j + 1) i j; + if j == 0 then + begin + pp " intros x; change (extend%i 0 x) with (WW (znz_0 w%i_op) x)." i (i + j); + pp " unfold to_Z; rewrite znz_to_Z_%i." (i + j + 1); + pp " rewrite (spec_0 w%i_spec); auto." (i + j); + end + else + begin + pp " intros x; change (extend%i %i x) with (WW (znz_0 w%i_op) (extend%i %i x))." i j (i + j) i (j - 1); + pp " unfold to_Z; rewrite znz_to_Z_%i." (i + j + 1); + pp " rewrite (spec_0 w%i_spec)." (i + j); + pp " generalize (spec_extend%in%i x); unfold to_Z." i (i + j); + pp " intros HH; rewrite <- HH; auto."; + end; + pp " Qed."; + pp ""; + end; + done; + + pp " Theorem digits_w%in%i: znz_digits w%i_op = znz_digits (nmake_op _ w%i_op %i)." i (size - i + 1) (size + 1) i (size - i + 1); + pp " Proof."; + pp " apply trans_equal with (xO (znz_digits w%i_op))." size; + pp " auto."; + pp " rewrite digits_nmake."; + pp " rewrite digits_w%in%i." i (size - i); + pp " auto."; + pp " Qed."; + pp ""; + + pp " Let spec_eval%in%i: forall x, [%sn 0 x] = eval%in %i x." i (size - i + 1) c i (size - i + 1); + pp " Proof."; + pp " intros x; case x."; + pp " auto."; + pp " intros xh xl; unfold to_Z; rewrite znz_to_Z_%i." (size + 1); + pp " rewrite digits_w%in%i." i (size - i); + pp " generalize (spec_eval%in%i); unfold to_Z; intros HH; repeat rewrite HH." i (size - i); + pp " unfold eval%in, nmake_op%i." i i; + pp " rewrite (znz_nmake_op _ w%i_op %i); auto." i (size - i); + pp " Qed."; + pp ""; + + pp " Let spec_eval%in%i: forall x, [%sn 1 x] = eval%in %i x." i (size - i + 2) c i (size - i + 2); + pp " intros x; case x."; + pp " auto."; + pp " intros xh xl; unfold to_Z; rewrite znz_to_Z_%i." (size + 2); + pp " rewrite digits_w%in%i." i (size + 1 - i); + pp " generalize (spec_eval%in%i); unfold to_Z; change (make_op 0) with (w%i_op); intros HH; repeat rewrite HH." i (size + 1 - i) (size + 1); + pp " unfold eval%in, nmake_op%i." i i; + pp " rewrite (znz_nmake_op _ w%i_op %i); auto." i (size + 1 - i); + pp " Qed."; + pp ""; + done; + + pp " Let digits_w%in: forall n," size; + pp " znz_digits (make_op n) = znz_digits (nmake_op _ w%i_op (S n))." size; + pp " intros n; elim n; clear n."; + pp " change (znz_digits (make_op 0)) with (xO (znz_digits w%i_op))." size; + pp " rewrite nmake_op_S; apply sym_equal; auto."; + pp " intros n Hrec."; + pp " replace (znz_digits (make_op (S n))) with (xO (znz_digits (make_op n)))."; + pp " rewrite Hrec."; + pp " rewrite nmake_op_S; apply sym_equal; auto."; + pp " rewrite make_op_S; apply sym_equal; auto."; + pp " Qed."; + pp ""; + + pp " Let spec_eval%in: forall n x, [%sn n x] = eval%in (S n) x." size c size; + pp " intros n; elim n; clear n."; + pp " exact spec_eval%in1." size; + pp " intros n Hrec x; case x; clear x."; + pp " unfold to_Z, eval%in, nmake_op%i." size size; + pp " rewrite make_op_S; rewrite nmake_op_S; auto."; + pp " intros xh xl."; + pp " unfold to_Z in Hrec |- *."; + pp " rewrite znz_to_Z_n."; + pp " rewrite digits_w%in." size; + pp " repeat rewrite Hrec."; + pp " unfold eval%in, nmake_op%i." size size; + pp " apply sym_equal; rewrite nmake_op_S; auto."; + pp " Qed."; + pp ""; + + pp " Let spec_extend%in: forall n x, [%s%i x] = [%sn n (extend%i n x)]." size c size c size ; + pp " intros n; elim n; clear n."; + pp " intros x; change (extend%i 0 x) with (WW (znz_0 w%i_op) x)." size size; + pp " unfold to_Z."; + pp " change (make_op 0) with w%i_op." (size + 1); + pp " rewrite znz_to_Z_%i; rewrite (spec_0 w%i_spec); auto." (size + 1) size; + pp " intros n Hrec x."; + pp " change (extend%i (S n) x) with (WW W0 (extend%i n x))." size size; + pp " unfold to_Z in Hrec |- *; rewrite znz_to_Z_n; auto."; + pp " rewrite <- Hrec."; + pp " replace (znz_to_Z (make_op n) W0) with 0; auto."; + pp " case n; auto; intros; rewrite make_op_S; auto."; + pp " Qed."; + pp ""; + + pr " Theorem spec_pos: forall x, 0 <= [x]."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; clear x."; + for i = 0 to size do + pp " intros x; case (spec_to_Z w%i_spec x); auto." i; + done; + pp " intros n x; case (spec_to_Z (wn_spec n) x); auto."; + pp " Qed."; + pr ""; + + pp " Let spec_extendn_0: forall n wx, [%sn n (extend n _ wx)] = [%sn 0 wx]." c c; + pp " intros n; elim n; auto."; + pp " intros n1 Hrec wx; simpl extend; rewrite <- Hrec; auto."; + pp " unfold to_Z."; + pp " case n1; auto; intros n2; repeat rewrite make_op_S; auto."; + pp " Qed."; + pp " Hint Rewrite spec_extendn_0: extr."; + pp ""; + pp " Let spec_extendn0_0: forall n wx, [%sn (S n) (WW W0 wx)] = [%sn n wx]." c c; + pp " Proof."; + pp " intros n x; unfold to_Z."; + pp " rewrite znz_to_Z_n."; + pp " rewrite <- (Zplus_0_l (znz_to_Z (make_op n) x))."; + pp " apply (f_equal2 Zplus); auto."; + pp " case n; auto."; + pp " intros n1; rewrite make_op_S; auto."; + pp " Qed."; + pp " Hint Rewrite spec_extendn_0: extr."; + pp ""; + pp " Let spec_extend_tr: forall m n (w: word _ (S n)),"; + pp " [%sn (m + n) (extend_tr w m)] = [%sn n w]." c c; + pp " Proof."; + pp " induction m; auto."; + pp " intros n x; simpl extend_tr."; + pp " simpl plus; rewrite spec_extendn0_0; auto."; + pp " Qed."; + pp " Hint Rewrite spec_extend_tr: extr."; + pp ""; + pp " Let spec_cast_l: forall n m x1,"; + pp " [%sn (Max.max n m)" c; + pp " (castm (diff_r n m) (extend_tr x1 (snd (diff n m))))] ="; + pp " [%sn n x1]." c; + pp " Proof."; + pp " intros n m x1; case (diff_r n m); simpl castm."; + pp " rewrite spec_extend_tr; auto."; + pp " Qed."; + pp " Hint Rewrite spec_cast_l: extr."; + pp ""; + pp " Let spec_cast_r: forall n m x1,"; + pp " [%sn (Max.max n m)" c; + pp " (castm (diff_l n m) (extend_tr x1 (fst (diff n m))))] ="; + pp " [%sn m x1]." c; + pp " Proof."; + pp " intros n m x1; case (diff_l n m); simpl castm."; + pp " rewrite spec_extend_tr; auto."; + pp " Qed."; + pp " Hint Rewrite spec_cast_r: extr."; + pp ""; + + + pr " Section LevelAndIter."; + pr ""; + pr " Variable res: Type."; + pr " Variable xxx: res."; + pr " Variable P: Z -> Z -> res -> Prop."; + pr " (* Abstraction function for each level *)"; + for i = 0 to size do + pr " Variable f%i: w%i -> w%i -> res." i i i; + pr " Variable f%in: forall n, w%i -> word w%i (S n) -> res." i i i; + pr " Variable fn%i: forall n, word w%i (S n) -> w%i -> res." i i i; + pp " Variable Pf%i: forall x y, P [%s%i x] [%s%i y] (f%i x y)." i c i c i i; + if i == size then + begin + pp " Variable Pf%in: forall n x y, P [%s%i x] (eval%in (S n) y) (f%in n x y)." i c i i i; + pp " Variable Pfn%i: forall n x y, P (eval%in (S n) x) [%s%i y] (fn%i n x y)." i i c i i; + end + else + begin + pp " Variable Pf%in: forall n x y, Z_of_nat n <= %i -> P [%s%i x] (eval%in (S n) y) (f%in n x y)." i (size - i) c i i i; + pp " Variable Pfn%i: forall n x y, Z_of_nat n <= %i -> P (eval%in (S n) x) [%s%i y] (fn%i n x y)." i (size - i) i c i i; + end; + pr ""; + done; + pr " Variable fnn: forall n, word w%i (S n) -> word w%i (S n) -> res." size size; + pp " Variable Pfnn: forall n x y, P [%sn n x] [%sn n y] (fnn n x y)." c c; + pr " Variable fnm: forall n m, word w%i (S n) -> word w%i (S m) -> res." size size; + pp " Variable Pfnm: forall n m x y, P [%sn n x] [%sn m y] (fnm n m x y)." c c; + pr ""; + pr " (* Special zero functions *)"; + pr " Variable f0t: t_ -> res."; + pp " Variable Pf0t: forall x, P 0 [x] (f0t x)."; + pr " Variable ft0: t_ -> res."; + pp " Variable Pft0: forall x, P [x] 0 (ft0 x)."; + pr ""; + + + pr " (* We level the two arguments before applying *)"; + pr " (* the functions at each leval *)"; + pr " Definition same_level (x y: t_): res :="; + pr0 " Eval lazy zeta beta iota delta ["; + for i = 0 to size do + pr0 "extend%i " i; + done; + pr ""; + pr " DoubleBase.extend DoubleBase.extend_aux"; + pr " ] in"; + pr " match x, y with"; + for i = 0 to size do + for j = 0 to i - 1 do + pr " | %s%i wx, %s%i wy => f%i wx (extend%i %i wy)" c i c j i j (i - j -1); + done; + pr " | %s%i wx, %s%i wy => f%i wx wy" c i c i i; + for j = i + 1 to size do + pr " | %s%i wx, %s%i wy => f%i (extend%i %i wx) wy" c i c j j i (j - i - 1); + done; + if i == size then + pr " | %s%i wx, %sn m wy => fnn m (extend%i m wx) wy" c size c size + else + pr " | %s%i wx, %sn m wy => fnn m (extend%i m (extend%i %i wx)) wy" c i c size i (size - i - 1); + done; + for i = 0 to size do + if i == size then + pr " | %sn n wx, %s%i wy => fnn n wx (extend%i n wy)" c c size size + else + pr " | %sn n wx, %s%i wy => fnn n wx (extend%i n (extend%i %i wy))" c c i size i (size - i - 1); + done; + pr " | %sn n wx, Nn m wy =>" c; + pr " let mn := Max.max n m in"; + pr " let d := diff n m in"; + pr " fnn mn"; + pr " (castm (diff_r n m) (extend_tr wx (snd d)))"; + pr " (castm (diff_l n m) (extend_tr wy (fst d)))"; + pr " end."; + pr ""; + + pp " Lemma spec_same_level: forall x y, P [x] [y] (same_level x y)."; + pp " Proof."; + pp " intros x; case x; clear x; unfold same_level."; + for i = 0 to size do + pp " intros x y; case y; clear y."; + for j = 0 to i - 1 do + pp " intros y; rewrite spec_extend%in%i; apply Pf%i." j i i; + done; + pp " intros y; apply Pf%i." i; + for j = i + 1 to size do + pp " intros y; rewrite spec_extend%in%i; apply Pf%i." i j j; + done; + if i == size then + pp " intros m y; rewrite (spec_extend%in m); apply Pfnn." size + else + pp " intros m y; rewrite spec_extend%in%i; rewrite (spec_extend%in m); apply Pfnn." i size size; + done; + pp " intros n x y; case y; clear y."; + for i = 0 to size do + if i == size then + pp " intros y; rewrite (spec_extend%in n); apply Pfnn." size + else + pp " intros y; rewrite spec_extend%in%i; rewrite (spec_extend%in n); apply Pfnn." i size size; + done; + pp " intros m y; rewrite <- (spec_cast_l n m x); "; + pp " rewrite <- (spec_cast_r n m y); apply Pfnn."; + pp " Qed."; + pp ""; + + pr " (* We level the two arguments before applying *)"; + pr " (* the functions at each level (special zero case) *)"; + pr " Definition same_level0 (x y: t_): res :="; + pr0 " Eval lazy zeta beta iota delta ["; + for i = 0 to size do + pr0 "extend%i " i; + done; + pr ""; + pr " DoubleBase.extend DoubleBase.extend_aux"; + pr " ] in"; + pr " match x with"; + for i = 0 to size do + pr " | %s%i wx =>" c i; + if i == 0 then + pr " if w0_eq0 wx then f0t y else"; + pr " match y with"; + for j = 0 to i - 1 do + pr " | %s%i wy =>" c j; + if j == 0 then + pr " if w0_eq0 wy then ft0 x else"; + pr " f%i wx (extend%i %i wy)" i j (i - j -1); + done; + pr " | %s%i wy => f%i wx wy" c i i; + for j = i + 1 to size do + pr " | %s%i wy => f%i (extend%i %i wx) wy" c j j i (j - i - 1); + done; + if i == size then + pr " | %sn m wy => fnn m (extend%i m wx) wy" c size + else + pr " | %sn m wy => fnn m (extend%i m (extend%i %i wx)) wy" c size i (size - i - 1); + pr" end"; + done; + pr " | %sn n wx =>" c; + pr " match y with"; + for i = 0 to size do + pr " | %s%i wy =>" c i; + if i == 0 then + pr " if w0_eq0 wy then ft0 x else"; + if i == size then + pr " fnn n wx (extend%i n wy)" size + else + pr " fnn n wx (extend%i n (extend%i %i wy))" size i (size - i - 1); + done; + pr " | %sn m wy =>" c; + pr " let mn := Max.max n m in"; + pr " let d := diff n m in"; + pr " fnn mn"; + pr " (castm (diff_r n m) (extend_tr wx (snd d)))"; + pr " (castm (diff_l n m) (extend_tr wy (fst d)))"; + pr " end"; + pr " end."; + pr ""; + + pp " Lemma spec_same_level0: forall x y, P [x] [y] (same_level0 x y)."; + pp " Proof."; + pp " intros x; case x; clear x; unfold same_level0."; + for i = 0 to size do + pp " intros x."; + if i == 0 then + begin + pp " generalize (spec_w0_eq0 x); case w0_eq0; intros H."; + pp " intros y; rewrite H; apply Pf0t."; + pp " clear H."; + end; + pp " intros y; case y; clear y."; + for j = 0 to i - 1 do + pp " intros y."; + if j == 0 then + begin + pp " generalize (spec_w0_eq0 y); case w0_eq0; intros H."; + pp " rewrite H; apply Pft0."; + pp " clear H."; + end; + pp " rewrite spec_extend%in%i; apply Pf%i." j i i; + done; + pp " intros y; apply Pf%i." i; + for j = i + 1 to size do + pp " intros y; rewrite spec_extend%in%i; apply Pf%i." i j j; + done; + if i == size then + pp " intros m y; rewrite (spec_extend%in m); apply Pfnn." size + else + pp " intros m y; rewrite spec_extend%in%i; rewrite (spec_extend%in m); apply Pfnn." i size size; + done; + pp " intros n x y; case y; clear y."; + for i = 0 to size do + pp " intros y."; + if i = 0 then + begin + pp " generalize (spec_w0_eq0 y); case w0_eq0; intros H."; + pp " rewrite H; apply Pft0."; + pp " clear H."; + end; + if i == size then + pp " rewrite (spec_extend%in n); apply Pfnn." size + else + pp " rewrite spec_extend%in%i; rewrite (spec_extend%in n); apply Pfnn." i size size; + done; + pp " intros m y; rewrite <- (spec_cast_l n m x); "; + pp " rewrite <- (spec_cast_r n m y); apply Pfnn."; + pp " Qed."; + pp ""; + + pr " (* We iter the smaller argument with the bigger *)"; + pr " Definition iter (x y: t_): res := "; + pr0 " Eval lazy zeta beta iota delta ["; + for i = 0 to size do + pr0 "extend%i " i; + done; + pr ""; + pr " DoubleBase.extend DoubleBase.extend_aux"; + pr " ] in"; + pr " match x, y with"; + for i = 0 to size do + for j = 0 to i - 1 do + pr " | %s%i wx, %s%i wy => fn%i %i wx wy" c i c j j (i - j - 1); + done; + pr " | %s%i wx, %s%i wy => f%i wx wy" c i c i i; + for j = i + 1 to size do + pr " | %s%i wx, %s%i wy => f%in %i wx wy" c i c j i (j - i - 1); + done; + if i == size then + pr " | %s%i wx, %sn m wy => f%in m wx wy" c size c size + else + pr " | %s%i wx, %sn m wy => f%in m (extend%i %i wx) wy" c i c size i (size - i - 1); + done; + for i = 0 to size do + if i == size then + pr " | %sn n wx, %s%i wy => fn%i n wx wy" c c size size + else + pr " | %sn n wx, %s%i wy => fn%i n wx (extend%i %i wy)" c c i size i (size - i - 1); + done; + pr " | %sn n wx, %sn m wy => fnm n m wx wy" c c; + pr " end."; + pr ""; + + pp " Ltac zg_tac := try"; + pp " (red; simpl Zcompare; auto;"; + pp " let t := fresh \"H\" in (intros t; discriminate t))."; + pp " Lemma spec_iter: forall x y, P [x] [y] (iter x y)."; + pp " Proof."; + pp " intros x; case x; clear x; unfold iter."; + for i = 0 to size do + pp " intros x y; case y; clear y."; + for j = 0 to i - 1 do + pp " intros y; rewrite spec_eval%in%i; apply (Pfn%i %i); zg_tac." j (i - j) j (i - j - 1); + done; + pp " intros y; apply Pf%i." i; + for j = i + 1 to size do + pp " intros y; rewrite spec_eval%in%i; apply (Pf%in %i); zg_tac." i (j - i) i (j - i - 1); + done; + if i == size then + pp " intros m y; rewrite spec_eval%in; apply Pf%in." size size + else + pp " intros m y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pf%in." i size size size; + done; + pp " intros n x y; case y; clear y."; + for i = 0 to size do + if i == size then + pp " intros y; rewrite spec_eval%in; apply Pfn%i." size size + else + pp " intros y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pfn%i." i size size size; + done; + pp " intros m y; apply Pfnm."; + pp " Qed."; + pp ""; + + + pr " (* We iter the smaller argument with the bigger (zero case) *)"; + pr " Definition iter0 (x y: t_): res :="; + pr0 " Eval lazy zeta beta iota delta ["; + for i = 0 to size do + pr0 "extend%i " i; + done; + pr ""; + pr " DoubleBase.extend DoubleBase.extend_aux"; + pr " ] in"; + pr " match x with"; + for i = 0 to size do + pr " | %s%i wx =>" c i; + if i == 0 then + pr " if w0_eq0 wx then f0t y else"; + pr " match y with"; + for j = 0 to i - 1 do + pr " | %s%i wy =>" c j; + if j == 0 then + pr " if w0_eq0 wy then ft0 x else"; + pr " fn%i %i wx wy" j (i - j - 1); + done; + pr " | %s%i wy => f%i wx wy" c i i; + for j = i + 1 to size do + pr " | %s%i wy => f%in %i wx wy" c j i (j - i - 1); + done; + if i == size then + pr " | %sn m wy => f%in m wx wy" c size + else + pr " | %sn m wy => f%in m (extend%i %i wx) wy" c size i (size - i - 1); + pr " end"; + done; + pr " | %sn n wx =>" c; + pr " match y with"; + for i = 0 to size do + pr " | %s%i wy =>" c i; + if i == 0 then + pr " if w0_eq0 wy then ft0 x else"; + if i == size then + pr " fn%i n wx wy" size + else + pr " fn%i n wx (extend%i %i wy)" size i (size - i - 1); + done; + pr " | %sn m wy => fnm n m wx wy" c; + pr " end"; + pr " end."; + pr ""; + + pp " Lemma spec_iter0: forall x y, P [x] [y] (iter0 x y)."; + pp " Proof."; + pp " intros x; case x; clear x; unfold iter0."; + for i = 0 to size do + pp " intros x."; + if i == 0 then + begin + pp " generalize (spec_w0_eq0 x); case w0_eq0; intros H."; + pp " intros y; rewrite H; apply Pf0t."; + pp " clear H."; + end; + pp " intros y; case y; clear y."; + for j = 0 to i - 1 do + pp " intros y."; + if j == 0 then + begin + pp " generalize (spec_w0_eq0 y); case w0_eq0; intros H."; + pp " rewrite H; apply Pft0."; + pp " clear H."; + end; + pp " rewrite spec_eval%in%i; apply (Pfn%i %i); zg_tac." j (i - j) j (i - j - 1); + done; + pp " intros y; apply Pf%i." i; + for j = i + 1 to size do + pp " intros y; rewrite spec_eval%in%i; apply (Pf%in %i); zg_tac." i (j - i) i (j - i - 1); + done; + if i == size then + pp " intros m y; rewrite spec_eval%in; apply Pf%in." size size + else + pp " intros m y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pf%in." i size size size; + done; + pp " intros n x y; case y; clear y."; + for i = 0 to size do + pp " intros y."; + if i = 0 then + begin + pp " generalize (spec_w0_eq0 y); case w0_eq0; intros H."; + pp " rewrite H; apply Pft0."; + pp " clear H."; + end; + if i == size then + pp " rewrite spec_eval%in; apply Pfn%i." size size + else + pp " rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pfn%i." i size size size; + done; + pp " intros m y; apply Pfnm."; + pp " Qed."; + pp ""; + + + pr " End LevelAndIter."; + pr ""; + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Reduction *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + pr " Definition reduce_0 (x:w) := %s0 x." c; + pr " Definition reduce_1 :="; + pr " Eval lazy beta iota delta[reduce_n1] in"; + pr " reduce_n1 _ _ zero w0_eq0 %s0 %s1." c c; + for i = 2 to size do + pr " Definition reduce_%i :=" i; + pr " Eval lazy beta iota delta[reduce_n1] in"; + pr " reduce_n1 _ _ zero w%i_eq0 reduce_%i %s%i." + (i-1) (i-1) c i + done; + pr " Definition reduce_%i :=" (size+1); + pr " Eval lazy beta iota delta[reduce_n1] in"; + pr " reduce_n1 _ _ zero w%i_eq0 reduce_%i (%sn 0)." + size size c; + + pr " Definition reduce_n n := "; + pr " Eval lazy beta iota delta[reduce_n] in"; + pr " reduce_n _ _ zero reduce_%i %sn n." (size + 1) c; + pr ""; + + pp " Let spec_reduce_0: forall x, [reduce_0 x] = [%s0 x]." c; + pp " Proof."; + pp " intros x; unfold to_Z, reduce_0."; + pp " auto."; + pp " Qed."; + pp " "; + + for i = 1 to size + 1 do + if i == size + 1 then + pp " Let spec_reduce_%i: forall x, [reduce_%i x] = [%sn 0 x]." i i c + else + pp " Let spec_reduce_%i: forall x, [reduce_%i x] = [%s%i x]." i i c i; + pp " Proof."; + pp " intros x; case x; unfold reduce_%i." i; + pp " exact (spec_0 w0_spec)."; + pp " intros x1 y1."; + pp " generalize (spec_w%i_eq0 x1); " (i - 1); + pp " case w%i_eq0; intros H1; auto." (i - 1); + if i <> 1 then + pp " rewrite spec_reduce_%i." (i - 1); + pp " unfold to_Z; rewrite znz_to_Z_%i." i; + pp " unfold to_Z in H1; rewrite H1; auto."; + pp " Qed."; + pp " "; + done; + + pp " Let spec_reduce_n: forall n x, [reduce_n n x] = [%sn n x]." c; + pp " Proof."; + pp " intros n; elim n; simpl reduce_n."; + pp " intros x; rewrite <- spec_reduce_%i; auto." (size + 1); + pp " intros n1 Hrec x; case x."; + pp " unfold to_Z; rewrite make_op_S; auto."; + pp " exact (spec_0 w0_spec)."; + pp " intros x1 y1; case x1; auto."; + pp " rewrite Hrec."; + pp " rewrite spec_extendn0_0; auto."; + pp " Qed."; + pp " "; + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Successor *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_succ_c := w%i_op.(znz_succ_c)." i i + done; + pr ""; + + for i = 0 to size do + pr " Definition w%i_succ := w%i_op.(znz_succ)." i i + done; + pr ""; + + pr " Definition succ x :="; + pr " match x with"; + for i = 0 to size-1 do + pr " | %s%i wx =>" c i; + pr " match w%i_succ_c wx with" i; + pr " | C0 r => %s%i r" c i; + pr " | C1 r => %s%i (WW one%i r)" c (i+1) i; + pr " end"; + done; + pr " | %s%i wx =>" c size; + pr " match w%i_succ_c wx with" size; + pr " | C0 r => %s%i r" c size; + pr " | C1 r => %sn 0 (WW one%i r)" c size ; + pr " end"; + pr " | %sn n wx =>" c; + pr " let op := make_op n in"; + pr " match op.(znz_succ_c) wx with"; + pr " | C0 r => %sn n r" c; + pr " | C1 r => %sn (S n) (WW op.(znz_1) r)" c; + pr " end"; + pr " end."; + pr ""; + + pr " Theorem spec_succ: forall n, [succ n] = [n] + 1."; + pa " Admitted."; + pp " Proof."; + pp " intros n; case n; unfold succ, to_Z."; + for i = 0 to size do + pp " intros n1; generalize (spec_succ_c w%i_spec n1);" i; + pp " unfold succ, to_Z, w%i_succ_c; case znz_succ_c; auto." i; + pp " intros ww H; rewrite <- H."; + pp " (rewrite znz_to_Z_%i; unfold interp_carry;" (i + 1); + pp " apply f_equal2 with (f := Zplus); auto;"; + pp " apply f_equal2 with (f := Zmult); auto;"; + pp " exact (spec_1 w%i_spec))." i; + done; + pp " intros k n1; generalize (spec_succ_c (wn_spec k) n1)."; + pp " unfold succ, to_Z; case znz_succ_c; auto."; + pp " intros ww H; rewrite <- H."; + pp " (rewrite (znz_to_Z_n k); unfold interp_carry;"; + pp " apply f_equal2 with (f := Zplus); auto;"; + pp " apply f_equal2 with (f := Zmult); auto;"; + pp " exact (spec_1 (wn_spec k)))."; + pp " Qed."; + pr ""; + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Adddition *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_add_c := znz_add_c w%i_op." i i; + pr " Definition w%i_add x y :=" i; + pr " match w%i_add_c x y with" i; + pr " | C0 r => %s%i r" c i; + if i == size then + pr " | C1 r => %sn 0 (WW one%i r)" c size + else + pr " | C1 r => %s%i (WW one%i r)" c (i + 1) i; + pr " end."; + pr ""; + done ; + pr " Definition addn n (x y : word w%i (S n)) :=" size; + pr " let op := make_op n in"; + pr " match op.(znz_add_c) x y with"; + pr " | C0 r => %sn n r" c; + pr " | C1 r => %sn (S n) (WW op.(znz_1) r) end." c; + pr ""; + + + for i = 0 to size do + pp " Let spec_w%i_add: forall x y, [w%i_add x y] = [%s%i x] + [%s%i y]." i i c i c i; + pp " Proof."; + pp " intros n m; unfold to_Z, w%i_add, w%i_add_c." i i; + pp " generalize (spec_add_c w%i_spec n m); case znz_add_c; auto." i; + pp " intros ww H; rewrite <- H."; + pp " rewrite znz_to_Z_%i; unfold interp_carry;" (i + 1); + pp " apply f_equal2 with (f := Zplus); auto;"; + pp " apply f_equal2 with (f := Zmult); auto;"; + pp " exact (spec_1 w%i_spec)." i; + pp " Qed."; + pp " Hint Rewrite spec_w%i_add: addr." i; + pp ""; + done; + pp " Let spec_wn_add: forall n x y, [addn n x y] = [%sn n x] + [%sn n y]." c c; + pp " Proof."; + pp " intros k n m; unfold to_Z, addn."; + pp " generalize (spec_add_c (wn_spec k) n m); case znz_add_c; auto."; + pp " intros ww H; rewrite <- H."; + pp " rewrite (znz_to_Z_n k); unfold interp_carry;"; + pp " apply f_equal2 with (f := Zplus); auto;"; + pp " apply f_equal2 with (f := Zmult); auto;"; + pp " exact (spec_1 (wn_spec k))."; + pp " Qed."; + pp " Hint Rewrite spec_wn_add: addr."; + + pr " Definition add := Eval lazy beta delta [same_level] in"; + pr0 " (same_level t_ "; + for i = 0 to size do + pr0 "w%i_add " i; + done; + pr "addn)."; + pr ""; + + pr " Theorem spec_add: forall x y, [add x y] = [x] + [y]."; + pa " Admitted."; + pp " Proof."; + pp " unfold add."; + pp " generalize (spec_same_level t_ (fun x y res => [res] = x + y))."; + pp " unfold same_level; intros HH; apply HH; clear HH."; + for i = 0 to size do + pp " exact spec_w%i_add." i; + done; + pp " exact spec_wn_add."; + pp " Qed."; + pr ""; + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Predecessor *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_pred_c := w%i_op.(znz_pred_c)." i i + done; + pr ""; + + pr " Definition pred x :="; + pr " match x with"; + for i = 0 to size do + pr " | %s%i wx =>" c i; + pr " match w%i_pred_c wx with" i; + pr " | C0 r => reduce_%i r" i; + pr " | C1 r => zero"; + pr " end"; + done; + pr " | %sn n wx =>" c; + pr " let op := make_op n in"; + pr " match op.(znz_pred_c) wx with"; + pr " | C0 r => reduce_n n r"; + pr " | C1 r => zero"; + pr " end"; + pr " end."; + pr ""; + + pr " Theorem spec_pred: forall x, 0 < [x] -> [pred x] = [x] - 1."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; unfold pred."; + for i = 0 to size do + pp " intros x1 H1; unfold w%i_pred_c; " i; + pp " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1." i; + pp " rewrite spec_reduce_%i; auto." i; + pp " unfold interp_carry; unfold to_Z."; + pp " case (spec_to_Z w%i_spec x1); intros HH1 HH2." i; + pp " case (spec_to_Z w%i_spec y1); intros HH3 HH4 HH5." i; + pp " assert (znz_to_Z w%i_op x1 - 1 < 0); auto with zarith." i; + pp " unfold to_Z in H1; auto with zarith."; + done; + pp " intros n x1 H1; "; + pp " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1."; + pp " rewrite spec_reduce_n; auto."; + pp " unfold interp_carry; unfold to_Z."; + pp " case (spec_to_Z (wn_spec n) x1); intros HH1 HH2."; + pp " case (spec_to_Z (wn_spec n) y1); intros HH3 HH4 HH5."; + pp " assert (znz_to_Z (make_op n) x1 - 1 < 0); auto with zarith."; + pp " unfold to_Z in H1; auto with zarith."; + pp " Qed."; + pp " "; + + pp " Let spec_pred0: forall x, [x] = 0 -> [pred x] = 0."; + pp " Proof."; + pp " intros x; case x; unfold pred."; + for i = 0 to size do + pp " intros x1 H1; unfold w%i_pred_c; " i; + pp " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1." i; + pp " unfold interp_carry; unfold to_Z."; + pp " unfold to_Z in H1; auto with zarith."; + pp " case (spec_to_Z w%i_spec y1); intros HH3 HH4; auto with zarith." i; + pp " intros; exact (spec_0 w0_spec)."; + done; + pp " intros n x1 H1; "; + pp " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1."; + pp " unfold interp_carry; unfold to_Z."; + pp " unfold to_Z in H1; auto with zarith."; + pp " case (spec_to_Z (wn_spec n) y1); intros HH3 HH4; auto with zarith."; + pp " intros; exact (spec_0 w0_spec)."; + pp " Qed."; + pr " "; + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Subtraction *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_sub_c := w%i_op.(znz_sub_c)." i i + done; + pr ""; + + for i = 0 to size do + pr " Definition w%i_sub x y :=" i; + pr " match w%i_sub_c x y with" i; + pr " | C0 r => reduce_%i r" i; + pr " | C1 r => zero"; + pr " end." + done; + pr ""; + + pr " Definition subn n (x y : word w%i (S n)) :=" size; + pr " let op := make_op n in"; + pr " match op.(znz_sub_c) x y with"; + pr " | C0 r => %sn n r" c; + pr " | C1 r => N0 w_0"; + pr " end."; + pr ""; + + for i = 0 to size do + pp " Let spec_w%i_sub: forall x y, [%s%i y] <= [%s%i x] -> [w%i_sub x y] = [%s%i x] - [%s%i y]." i c i c i i c i c i; + pp " Proof."; + pp " intros n m; unfold w%i_sub, w%i_sub_c." i i; + pp " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; " i; + if i == 0 then + pp " intros x; auto." + else + pp " intros x; try rewrite spec_reduce_%i; auto." i; + pp " unfold interp_carry; unfold zero, w_0, to_Z."; + pp " rewrite (spec_0 w0_spec)."; + pp " case (spec_to_Z w%i_spec x); intros; auto with zarith." i; + pp " Qed."; + pp ""; + done; + + pp " Let spec_wn_sub: forall n x y, [%sn n y] <= [%sn n x] -> [subn n x y] = [%sn n x] - [%sn n y]." c c c c; + pp " Proof."; + pp " intros k n m; unfold subn."; + pp " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; "; + pp " intros x; auto."; + pp " unfold interp_carry, to_Z."; + pp " case (spec_to_Z (wn_spec k) x); intros; auto with zarith."; + pp " Qed."; + pp ""; + + pr " Definition sub := Eval lazy beta delta [same_level] in"; + pr0 " (same_level t_ "; + for i = 0 to size do + pr0 "w%i_sub " i; + done; + pr "subn)."; + pr ""; + + pr " Theorem spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y]."; + pa " Admitted."; + pp " Proof."; + pp " unfold sub."; + pp " generalize (spec_same_level t_ (fun x y res => y <= x -> [res] = x - y))."; + pp " unfold same_level; intros HH; apply HH; clear HH."; + for i = 0 to size do + pp " exact spec_w%i_sub." i; + done; + pp " exact spec_wn_sub."; + pp " Qed."; + pr ""; + + for i = 0 to size do + pp " Let spec_w%i_sub0: forall x y, [%s%i x] < [%s%i y] -> [w%i_sub x y] = 0." i c i c i i; + pp " Proof."; + pp " intros n m; unfold w%i_sub, w%i_sub_c." i i; + pp " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; " i; + pp " intros x; unfold interp_carry."; + pp " unfold to_Z; case (spec_to_Z w%i_spec x); intros; auto with zarith." i; + pp " intros; unfold to_Z, zero, w_0; rewrite (spec_0 w0_spec); auto."; + pp " Qed."; + pp ""; + done; + + pp " Let spec_wn_sub0: forall n x y, [%sn n x] < [%sn n y] -> [subn n x y] = 0." c c; + pp " Proof."; + pp " intros k n m; unfold subn."; + pp " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; "; + pp " intros x; unfold interp_carry."; + pp " unfold to_Z; case (spec_to_Z (wn_spec k) x); intros; auto with zarith."; + pp " intros; unfold to_Z, w_0; rewrite (spec_0 (w0_spec)); auto."; + pp " Qed."; + pp ""; + + pr " Theorem spec_sub0: forall x y, [x] < [y] -> [sub x y] = 0."; + pa " Admitted."; + pp " Proof."; + pp " unfold sub."; + pp " generalize (spec_same_level t_ (fun x y res => x < y -> [res] = 0))."; + pp " unfold same_level; intros HH; apply HH; clear HH."; + for i = 0 to size do + pp " exact spec_w%i_sub0." i; + done; + pp " exact spec_wn_sub0."; + pp " Qed."; + pr ""; + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Comparison *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition compare_%i := w%i_op.(znz_compare)." i i; + pr " Definition comparen_%i :=" i; + pr " compare_mn_1 w%i w%i %s compare_%i (compare_%i %s) compare_%i." i i (pz i) i i (pz i) i + done; + pr ""; + + pr " Definition comparenm n m wx wy :="; + pr " let mn := Max.max n m in"; + pr " let d := diff n m in"; + pr " let op := make_op mn in"; + pr " op.(znz_compare)"; + pr " (castm (diff_r n m) (extend_tr wx (snd d)))"; + pr " (castm (diff_l n m) (extend_tr wy (fst d)))."; + pr ""; + + pr " Definition compare := Eval lazy beta delta [iter] in "; + pr " (iter _ "; + for i = 0 to size do + pr " compare_%i" i; + pr " (fun n x y => opp_compare (comparen_%i (S n) y x))" i; + pr " (fun n => comparen_%i (S n))" i; + done; + pr " comparenm)."; + pr ""; + + pr " Definition lt n m := compare n m = Lt."; + pr " Definition le n m := compare n m <> Gt."; + pr " Definition min n m := match compare n m with Gt => m | _ => n end."; + pr " Definition max n m := match compare n m with Lt => m | _ => n end."; + pr ""; + + for i = 0 to size do + pp " Let spec_compare_%i: forall x y," i; + pp " match compare_%i x y with " i; + pp " Eq => [%s%i x] = [%s%i y]" c i c i; + pp " | Lt => [%s%i x] < [%s%i y]" c i c i; + pp " | Gt => [%s%i x] > [%s%i y]" c i c i; + pp " end."; + pp " Proof."; + pp " unfold compare_%i, to_Z; exact (spec_compare w%i_spec)." i i; + pp " Qed."; + pp ""; + + pp " Let spec_comparen_%i:" i; + pp " forall (n : nat) (x : word w%i n) (y : w%i)," i i; + pp " match comparen_%i n x y with" i; + pp " | Eq => eval%in n x = [%s%i y]" i c i; + pp " | Lt => eval%in n x < [%s%i y]" i c i; + pp " | Gt => eval%in n x > [%s%i y]" i c i; + pp " end."; + pp " intros n x y."; + pp " unfold comparen_%i, to_Z; rewrite spec_double_eval%in." i i; + pp " apply spec_compare_mn_1."; + pp " exact (spec_0 w%i_spec)." i; + pp " intros x1; exact (spec_compare w%i_spec %s x1)." i (pz i); + pp " exact (spec_to_Z w%i_spec)." i; + pp " exact (spec_compare w%i_spec)." i; + pp " exact (spec_compare w%i_spec)." i; + pp " exact (spec_to_Z w%i_spec)." i; + pp " Qed."; + pp ""; + done; + + pp " Let spec_opp_compare: forall c (u v: Z),"; + pp " match c with Eq => u = v | Lt => u < v | Gt => u > v end ->"; + pp " match opp_compare c with Eq => v = u | Lt => v < u | Gt => v > u end."; + pp " Proof."; + pp " intros c u v; case c; unfold opp_compare; auto with zarith."; + pp " Qed."; + pp ""; + + + pr " Theorem spec_compare: forall x y,"; + pr " match compare x y with "; + pr " Eq => [x] = [y]"; + pr " | Lt => [x] < [y]"; + pr " | Gt => [x] > [y]"; + pr " end."; + pa " Admitted."; + pp " Proof."; + pp " refine (spec_iter _ (fun x y res => "; + pp " match res with "; + pp " Eq => x = y"; + pp " | Lt => x < y"; + pp " | Gt => x > y"; + pp " end)"; + for i = 0 to size do + pp " compare_%i" i; + pp " (fun n x y => opp_compare (comparen_%i (S n) y x))" i; + pp " (fun n => comparen_%i (S n)) _ _ _" i; + done; + pp " comparenm _)."; + + for i = 0 to size - 1 do + pp " exact spec_compare_%i." i; + pp " intros n x y H;apply spec_opp_compare; apply spec_comparen_%i." i; + pp " intros n x y H; exact (spec_comparen_%i (S n) x y)." i; + done; + pp " exact spec_compare_%i." size; + pp " intros n x y;apply spec_opp_compare; apply spec_comparen_%i." size; + pp " intros n; exact (spec_comparen_%i (S n))." size; + pp " intros n m x y; unfold comparenm."; + pp " rewrite <- (spec_cast_l n m x); rewrite <- (spec_cast_r n m y)."; + pp " unfold to_Z; apply (spec_compare (wn_spec (Max.max n m)))."; + pp " Qed."; + pr ""; + + pr " Definition eq_bool x y :="; + pr " match compare x y with"; + pr " | Eq => true"; + pr " | _ => false"; + pr " end."; + pr ""; + + + pr " Theorem spec_eq_bool: forall x y,"; + pr " if eq_bool x y then [x] = [y] else [x] <> [y]."; + pa " Admitted."; + pp " Proof."; + pp " intros x y; unfold eq_bool."; + pp " generalize (spec_compare x y); case compare; auto with zarith."; + pp " Qed."; + pr ""; + + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Multiplication *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_mul_c := w%i_op.(znz_mul_c)." i i + done; + pr ""; + + for i = 0 to size do + pr " Definition w%i_mul_add :=" i; + pr " Eval lazy beta delta [w_mul_add] in"; + pr " @w_mul_add w%i %s w%i_succ w%i_add_c w%i_mul_c." i (pz i) i i i + done; + pr ""; + + for i = 0 to size do + pr " Definition w%i_0W := znz_0W w%i_op." i i + done; + pr ""; + + for i = 0 to size do + pr " Definition w%i_WW := znz_WW w%i_op." i i + done; + pr ""; + + for i = 0 to size do + pr " Definition w%i_mul_add_n1 :=" i; + pr " @double_mul_add_n1 w%i %s w%i_WW w%i_0W w%i_mul_add." i (pz i) i i i + done; + pr ""; + + for i = 0 to size - 1 do + pr " Let to_Z%i n :=" i; + pr " match n return word w%i (S n) -> t_ with" i; + for j = 0 to size - i do + if (i + j) == size then + begin + pr " | %i%s => fun x => %sn 0 x" j "%nat" c; + pr " | %i%s => fun x => %sn 1 x" (j + 1) "%nat" c + end + else + pr " | %i%s => fun x => %s%i x" j "%nat" c (i + j + 1) + done; + pr " | _ => fun _ => N0 w_0"; + pr " end."; + pr ""; + done; + + + for i = 0 to size - 1 do + pp "Theorem to_Z%i_spec:" i; + pp " forall n x, Z_of_nat n <= %i -> [to_Z%i n x] = znz_to_Z (nmake_op _ w%i_op (S n)) x." (size + 1 - i) i i; + for j = 1 to size + 2 - i do + pp " intros n; case n; clear n."; + pp " unfold to_Z%i." i; + pp " intros x H; rewrite spec_eval%in%i; auto." i j; + done; + pp " intros n x."; + pp " repeat rewrite inj_S; unfold Zsucc; auto with zarith."; + pp " Qed."; + pp ""; + done; + + + for i = 0 to size do + pr " Definition w%i_mul n x y :=" i; + pr " let (w,r) := w%i_mul_add_n1 (S n) x y %s in" i (pz i); + if i == size then + begin + pr " if w%i_eq0 w then %sn n r" i c; + pr " else %sn (S n) (WW (extend%i n w) r)." c i; + end + else + begin + pr " if w%i_eq0 w then to_Z%i n r" i i; + pr " else to_Z%i (S n) (WW (extend%i n w) r)." i i; + end; + pr ""; + done; + + pr " Definition mulnm n m x y :="; + pr " let mn := Max.max n m in"; + pr " let d := diff n m in"; + pr " let op := make_op mn in"; + pr " reduce_n (S mn) (op.(znz_mul_c)"; + pr " (castm (diff_r n m) (extend_tr x (snd d)))"; + pr " (castm (diff_l n m) (extend_tr y (fst d))))."; + pr ""; + + pr " Definition mul := Eval lazy beta delta [iter0] in "; + pr " (iter0 t_ "; + for i = 0 to size do + pr " (fun x y => reduce_%i (w%i_mul_c x y)) " (i + 1) i; + pr " (fun n x y => w%i_mul n y x)" i; + pr " w%i_mul" i; + done; + pr " mulnm"; + pr " (fun _ => N0 w_0)"; + pr " (fun _ => N0 w_0)"; + pr " )."; + pr ""; + for i = 0 to size do + pp " Let spec_w%i_mul_add: forall x y z," i; + pp " let (q,r) := w%i_mul_add x y z in" i; + pp " znz_to_Z w%i_op q * (base (znz_digits w%i_op)) + znz_to_Z w%i_op r =" i i i; + pp " znz_to_Z w%i_op x * znz_to_Z w%i_op y + znz_to_Z w%i_op z :=" i i i ; + pp " (spec_mul_add w%i_spec)." i; + pp ""; + done; + + for i = 0 to size do + pp " Theorem spec_w%i_mul_add_n1: forall n x y z," i; + pp " let (q,r) := w%i_mul_add_n1 n x y z in" i; + pp " znz_to_Z w%i_op q * (base (znz_digits (nmake_op _ w%i_op n))) +" i i; + pp " znz_to_Z (nmake_op _ w%i_op n) r =" i; + pp " znz_to_Z (nmake_op _ w%i_op n) x * znz_to_Z w%i_op y +" i i; + pp " znz_to_Z w%i_op z." i; + pp " Proof."; + pp " intros n x y z; unfold w%i_mul_add_n1." i; + pp " rewrite nmake_double."; + pp " rewrite digits_doubled."; + pp " change (base (DoubleBase.double_digits (znz_digits w%i_op) n)) with" i; + pp " (DoubleBase.double_wB (znz_digits w%i_op) n)." i; + pp " apply spec_double_mul_add_n1; auto."; + if i == 0 then pp " exact (spec_0 w%i_spec)." i; + pp " exact (spec_WW w%i_spec)." i; + pp " exact (spec_0W w%i_spec)." i; + pp " exact (spec_mul_add w%i_spec)." i; + pp " Qed."; + pp ""; + done; + + pp " Lemma nmake_op_WW: forall ww ww1 n x y,"; + pp " znz_to_Z (nmake_op ww ww1 (S n)) (WW x y) ="; + pp " znz_to_Z (nmake_op ww ww1 n) x * base (znz_digits (nmake_op ww ww1 n)) +"; + pp " znz_to_Z (nmake_op ww ww1 n) y."; + pp " auto."; + pp " Qed."; + pp ""; + + for i = 0 to size do + pp " Lemma extend%in_spec: forall n x1," i; + pp " znz_to_Z (nmake_op _ w%i_op (S n)) (extend%i n x1) = " i i; + pp " znz_to_Z w%i_op x1." i; + pp " Proof."; + pp " intros n1 x2; rewrite nmake_double."; + pp " unfold extend%i." i; + pp " rewrite DoubleBase.spec_extend; auto."; + if i == 0 then + pp " intros l; simpl; unfold w_0; rewrite (spec_0 w0_spec); ring."; + pp " Qed."; + pp ""; + done; + + pp " Lemma spec_muln:"; + pp " forall n (x: word _ (S n)) y,"; + pp " [%sn (S n) (znz_mul_c (make_op n) x y)] = [%sn n x] * [%sn n y]." c c c; + pp " Proof."; + pp " intros n x y; unfold to_Z."; + pp " rewrite <- (spec_mul_c (wn_spec n))."; + pp " rewrite make_op_S."; + pp " case znz_mul_c; auto."; + pp " Qed."; + + pr " Theorem spec_mul: forall x y, [mul x y] = [x] * [y]."; + pa " Admitted."; + pp " Proof."; + for i = 0 to size do + pp " assert(F%i: " i; + pp " forall n x y,"; + if i <> size then + pp0 " Z_of_nat n <= %i -> " (size - i); + pp " [w%i_mul n x y] = eval%in (S n) x * [%s%i y])." i i c i; + if i == size then + pp " intros n x y; unfold w%i_mul." i + else + pp " intros n x y H; unfold w%i_mul." i; + pp " generalize (spec_w%i_mul_add_n1 (S n) x y %s)." i (pz i); + pp " case w%i_mul_add_n1; intros x1 y1." i; + pp " change (znz_to_Z (nmake_op _ w%i_op (S n)) x) with (eval%in (S n) x)." i i; + pp " change (znz_to_Z w%i_op y) with ([%s%i y])." i c i; + if i == 0 then + pp " unfold w_0; rewrite (spec_0 w0_spec); rewrite Zplus_0_r." + else + pp " change (znz_to_Z w%i_op W0) with 0; rewrite Zplus_0_r." i; + pp " intros H1; rewrite <- H1; clear H1."; + pp " generalize (spec_w%i_eq0 x1); case w%i_eq0; intros HH." i i; + pp " unfold to_Z in HH; rewrite HH."; + if i == size then + begin + pp " rewrite spec_eval%in; unfold eval%in, nmake_op%i; auto." i i i; + pp " rewrite spec_eval%in; unfold eval%in, nmake_op%i." i i i + end + else + begin + pp " rewrite to_Z%i_spec; auto with zarith." i; + pp " rewrite to_Z%i_spec; try (rewrite inj_S; auto with zarith)." i + end; + pp " rewrite nmake_op_WW; rewrite extend%in_spec; auto." i; + done; + pp " refine (spec_iter0 t_ (fun x y res => [res] = x * y)"; + for i = 0 to size do + pp " (fun x y => reduce_%i (w%i_mul_c x y)) " (i + 1) i; + pp " (fun n x y => w%i_mul n y x)" i; + pp " w%i_mul _ _ _" i; + done; + pp " mulnm _"; + pp " (fun _ => N0 w_0) _"; + pp " (fun _ => N0 w_0) _"; + pp " )."; + for i = 0 to size do + pp " intros x y; rewrite spec_reduce_%i." (i + 1); + pp " unfold w%i_mul_c, to_Z." i; + pp " generalize (spec_mul_c w%i_spec x y)." i; + pp " intros HH; rewrite <- HH; clear HH; auto."; + if i == size then + begin + pp " intros n x y; rewrite F%i; auto with zarith." i; + pp " intros n x y; rewrite F%i; auto with zarith. " i; + end + else + begin + pp " intros n x y H; rewrite F%i; auto with zarith." i; + pp " intros n x y H; rewrite F%i; auto with zarith. " i; + end; + done; + pp " intros n m x y; unfold mulnm."; + pp " rewrite spec_reduce_n."; + pp " rewrite <- (spec_cast_l n m x)."; + pp " rewrite <- (spec_cast_r n m y)."; + pp " rewrite spec_muln; rewrite spec_cast_l; rewrite spec_cast_r; auto."; + pp " intros x; unfold to_Z, w_0; rewrite (spec_0 w0_spec); ring."; + pp " intros x; unfold to_Z, w_0; rewrite (spec_0 w0_spec); ring."; + pp " Qed."; + pr ""; + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Square *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_square_c := w%i_op.(znz_square_c)." i i + done; + pr ""; + + pr " Definition square x :="; + pr " match x with"; + pr " | %s0 wx => reduce_1 (w0_square_c wx)" c; + for i = 1 to size - 1 do + pr " | %s%i wx => %s%i (w%i_square_c wx)" c i c (i+1) i + done; + pr " | %s%i wx => %sn 0 (w%i_square_c wx)" c size c size; + pr " | %sn n wx =>" c; + pr " let op := make_op n in"; + pr " %sn (S n) (op.(znz_square_c) wx)" c; + pr " end."; + pr ""; + + pr " Theorem spec_square: forall x, [square x] = [x] * [x]."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; unfold square; clear x."; + pp " intros x; rewrite spec_reduce_1; unfold to_Z."; + pp " exact (spec_square_c w%i_spec x)." 0; + for i = 1 to size do + pp " intros x; unfold to_Z."; + pp " exact (spec_square_c w%i_spec x)." i; + done; + pp " intros n x; unfold to_Z."; + pp " rewrite make_op_S."; + pp " exact (spec_square_c (wn_spec n) x)."; + pp "Qed."; + pr ""; + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Power *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + pr " Fixpoint power_pos (x:%s) (p:positive) {struct p} : %s :=" t t; + pr " match p with"; + pr " | xH => x"; + pr " | xO p => square (power_pos x p)"; + pr " | xI p => mul (square (power_pos x p)) x"; + pr " end."; + pr ""; + + pr " Theorem spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n."; + pa " Admitted."; + pp " Proof."; + pp " intros x n; generalize x; elim n; clear n x; simpl power_pos."; + pp " intros; rewrite spec_mul; rewrite spec_square; rewrite H."; + pp " rewrite Zpos_xI; rewrite Zpower_exp; auto with zarith."; + pp " rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith."; + pp " rewrite Zpower_2; rewrite Zpower_1_r; auto."; + pp " intros; rewrite spec_square; rewrite H."; + pp " rewrite Zpos_xO; auto with zarith."; + pp " rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith."; + pp " rewrite Zpower_2; auto."; + pp " intros; rewrite Zpower_1_r; auto."; + pp " Qed."; + pp ""; + pr ""; + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Square root *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_sqrt := w%i_op.(znz_sqrt)." i i + done; + pr ""; + + pr " Definition sqrt x :="; + pr " match x with"; + for i = 0 to size do + pr " | %s%i wx => reduce_%i (w%i_sqrt wx)" c i i i; + done; + pr " | %sn n wx =>" c; + pr " let op := make_op n in"; + pr " reduce_n n (op.(znz_sqrt) wx)"; + pr " end."; + pr ""; + + pr " Theorem spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2."; + pa " Admitted."; + pp " Proof."; + pp " intros x; unfold sqrt; case x; clear x."; + for i = 0 to size do + pp " intros x; rewrite spec_reduce_%i; exact (spec_sqrt w%i_spec x)." i i; + done; + pp " intros n x; rewrite spec_reduce_n; exact (spec_sqrt (wn_spec n) x)."; + pp " Qed."; + pr ""; + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Division *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_div_gt := w%i_op.(znz_div_gt)." i i + done; + pr ""; + + pp " Let spec_divn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) := "; + pp " (spec_double_divn1 "; + pp " ww_op.(znz_zdigits) ww_op.(znz_0)"; + pp " (znz_WW ww_op) ww_op.(znz_head0)"; + pp " ww_op.(znz_add_mul_div) ww_op.(znz_div21)"; + pp " ww_op.(znz_compare) ww_op.(znz_sub) (znz_to_Z ww_op)"; + pp " (spec_to_Z ww_spec) "; + pp " (spec_zdigits ww_spec)"; + pp " (spec_0 ww_spec) (spec_WW ww_spec) (spec_head0 ww_spec)"; + pp " (spec_add_mul_div ww_spec) (spec_div21 ww_spec) "; + pp " (CyclicAxioms.spec_compare ww_spec) (CyclicAxioms.spec_sub ww_spec))."; + pp ""; + + for i = 0 to size do + pr " Definition w%i_divn1 n x y :=" i; + pr " let (u, v) :="; + pr " double_divn1 w%i_op.(znz_zdigits) w%i_op.(znz_0)" i i; + pr " (znz_WW w%i_op) w%i_op.(znz_head0)" i i; + pr " w%i_op.(znz_add_mul_div) w%i_op.(znz_div21)" i i; + pr " w%i_op.(znz_compare) w%i_op.(znz_sub) (S n) x y in" i i; + if i == size then + pr " (%sn _ u, %s%i v)." c c i + else + pr " (to_Z%i _ u, %s%i v)." i c i; + done; + pr ""; + + for i = 0 to size do + pp " Lemma spec_get_end%i: forall n x y," i; + pp " eval%in n x <= [%s%i y] -> " i c i; + pp " [%s%i (DoubleBase.get_low %s n x)] = eval%in n x." c i (pz i) i; + pp " Proof."; + pp " intros n x y H."; + pp " rewrite spec_double_eval%in; unfold to_Z." i; + pp " apply DoubleBase.spec_get_low."; + pp " exact (spec_0 w%i_spec)." i; + pp " exact (spec_to_Z w%i_spec)." i; + pp " apply Zle_lt_trans with [%s%i y]; auto." c i; + pp " rewrite <- spec_double_eval%in; auto." i; + pp " unfold to_Z; case (spec_to_Z w%i_spec y); auto." i; + pp " Qed."; + pp ""; + done; + + for i = 0 to size do + pr " Let div_gt%i x y := let (u,v) := (w%i_div_gt x y) in (reduce_%i u, reduce_%i v)." i i i i; + done; + pr ""; + + + pr " Let div_gtnm n m wx wy :="; + pr " let mn := Max.max n m in"; + pr " let d := diff n m in"; + pr " let op := make_op mn in"; + pr " let (q, r):= op.(znz_div_gt)"; + pr " (castm (diff_r n m) (extend_tr wx (snd d)))"; + pr " (castm (diff_l n m) (extend_tr wy (fst d))) in"; + pr " (reduce_n mn q, reduce_n mn r)."; + pr ""; + + pr " Definition div_gt := Eval lazy beta delta [iter] in"; + pr " (iter _ "; + for i = 0 to size do + pr " div_gt%i" i; + pr " (fun n x y => div_gt%i x (DoubleBase.get_low %s (S n) y))" i (pz i); + pr " w%i_divn1" i; + done; + pr " div_gtnm)."; + pr ""; + + pr " Theorem spec_div_gt: forall x y,"; + pr " [x] > [y] -> 0 < [y] ->"; + pr " let (q,r) := div_gt x y in"; + pr " [q] = [x] / [y] /\\ [r] = [x] mod [y]."; + pa " Admitted."; + pp " Proof."; + pp " assert (FO:"; + pp " forall x y, [x] > [y] -> 0 < [y] ->"; + pp " let (q,r) := div_gt x y in"; + pp " [x] = [q] * [y] + [r] /\\ 0 <= [r] < [y])."; + pp " refine (spec_iter (t_*t_) (fun x y res => x > y -> 0 < y ->"; + pp " let (q,r) := res in"; + pp " x = [q] * y + [r] /\\ 0 <= [r] < y)"; + for i = 0 to size do + pp " div_gt%i" i; + pp " (fun n x y => div_gt%i x (DoubleBase.get_low %s (S n) y))" i (pz i); + pp " w%i_divn1 _ _ _" i; + done; + pp " div_gtnm _)."; + for i = 0 to size do + pp " intros x y H1 H2; unfold div_gt%i, w%i_div_gt." i i; + pp " generalize (spec_div_gt w%i_spec x y H1 H2); case znz_div_gt." i; + pp " intros xx yy; repeat rewrite spec_reduce_%i; auto." i; + if i == size then + pp " intros n x y H2 H3; unfold div_gt%i, w%i_div_gt." i i + else + pp " intros n x y H1 H2 H3; unfold div_gt%i, w%i_div_gt." i i; + pp " generalize (spec_div_gt w%i_spec x " i; + pp " (DoubleBase.get_low %s (S n) y))." (pz i); + pp0 " "; + for j = 0 to i do + pp0 "unfold w%i; " (i-j); + done; + pp "case znz_div_gt."; + pp " intros xx yy H4; repeat rewrite spec_reduce_%i." i; + pp " generalize (spec_get_end%i (S n) y x); unfold to_Z; intros H5." i; + pp " unfold to_Z in H2; rewrite H5 in H4; auto with zarith."; + if i == size then + pp " intros n x y H2 H3." + else + pp " intros n x y H1 H2 H3."; + pp " generalize"; + pp " (spec_divn1 w%i w%i_op w%i_spec (S n) x y H3)." i i i; + pp0 " unfold w%i_divn1; " i; + for j = 0 to i do + pp0 "unfold w%i; " (i-j); + done; + pp "case double_divn1."; + pp " intros xx yy H4."; + if i == size then + begin + pp " repeat rewrite <- spec_double_eval%in in H4; auto." i; + pp " rewrite spec_eval%in; auto." i; + end + else + begin + pp " rewrite to_Z%i_spec; auto with zarith." i; + pp " repeat rewrite <- spec_double_eval%in in H4; auto." i; + end; + done; + pp " intros n m x y H1 H2; unfold div_gtnm."; + pp " generalize (spec_div_gt (wn_spec (Max.max n m))"; + pp " (castm (diff_r n m)"; + pp " (extend_tr x (snd (diff n m))))"; + pp " (castm (diff_l n m)"; + pp " (extend_tr y (fst (diff n m)))))."; + pp " case znz_div_gt."; + pp " intros xx yy HH."; + pp " repeat rewrite spec_reduce_n."; + pp " rewrite <- (spec_cast_l n m x)."; + pp " rewrite <- (spec_cast_r n m y)."; + pp " unfold to_Z; apply HH."; + pp " rewrite <- (spec_cast_l n m x) in H1; auto."; + pp " rewrite <- (spec_cast_r n m y) in H1; auto."; + pp " rewrite <- (spec_cast_r n m y) in H2; auto."; + pp " intros x y H1 H2; generalize (FO x y H1 H2); case div_gt."; + pp " intros q r (H3, H4); split."; + pp " apply (Zdiv_unique [x] [y] [q] [r]); auto."; + pp " rewrite Zmult_comm; auto."; + pp " apply (Zmod_unique [x] [y] [q] [r]); auto."; + pp " rewrite Zmult_comm; auto."; + pp " Qed."; + pr ""; + + pr " Definition div_eucl x y :="; + pr " match compare x y with"; + pr " | Eq => (one, zero)"; + pr " | Lt => (zero, x)"; + pr " | Gt => div_gt x y"; + pr " end."; + pr ""; + + pr " Theorem spec_div_eucl: forall x y,"; + pr " 0 < [y] ->"; + pr " let (q,r) := div_eucl x y in"; + pr " ([q], [r]) = Zdiv_eucl [x] [y]."; + pa " Admitted."; + pp " Proof."; + pp " assert (F0: [zero] = 0)."; + pp " exact (spec_0 w0_spec)."; + pp " assert (F1: [one] = 1)."; + pp " exact (spec_1 w0_spec)."; + pp " intros x y H; generalize (spec_compare x y);"; + pp " unfold div_eucl; case compare; try rewrite F0;"; + pp " try rewrite F1; intros; auto with zarith."; + pp " rewrite H0; generalize (Z_div_same [y] (Zlt_gt _ _ H))"; + pp " (Z_mod_same [y] (Zlt_gt _ _ H));"; + pp " unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto."; + pp " assert (F2: 0 <= [x] < [y])."; + pp " generalize (spec_pos x); auto."; + pp " generalize (Zdiv_small _ _ F2)"; + pp " (Zmod_small _ _ F2);"; + pp " unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto."; + pp " generalize (spec_div_gt _ _ H0 H); auto."; + pp " unfold Zdiv, Zmod; case Zdiv_eucl; case div_gt."; + pp " intros a b c d (H1, H2); subst; auto."; + pp " Qed."; + pr ""; + + pr " Definition div x y := fst (div_eucl x y)."; + pr ""; + + pr " Theorem spec_div:"; + pr " forall x y, 0 < [y] -> [div x y] = [x] / [y]."; + pa " Admitted."; + pp " Proof."; + pp " intros x y H1; unfold div; generalize (spec_div_eucl x y H1);"; + pp " case div_eucl; simpl fst."; + pp " intros xx yy; unfold Zdiv; case Zdiv_eucl; intros qq rr H; "; + pp " injection H; auto."; + pp " Qed."; + pr ""; + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Modulo *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + for i = 0 to size do + pr " Definition w%i_mod_gt := w%i_op.(znz_mod_gt)." i i + done; + pr ""; + + for i = 0 to size do + pr " Definition w%i_modn1 :=" i; + pr " double_modn1 w%i_op.(znz_zdigits) w%i_op.(znz_0)" i i; + pr " w%i_op.(znz_head0) w%i_op.(znz_add_mul_div) w%i_op.(znz_div21)" i i i; + pr " w%i_op.(znz_compare) w%i_op.(znz_sub)." i i; + done; + pr ""; + + pr " Let mod_gtnm n m wx wy :="; + pr " let mn := Max.max n m in"; + pr " let d := diff n m in"; + pr " let op := make_op mn in"; + pr " reduce_n mn (op.(znz_mod_gt)"; + pr " (castm (diff_r n m) (extend_tr wx (snd d)))"; + pr " (castm (diff_l n m) (extend_tr wy (fst d))))."; + pr ""; + + pr " Definition mod_gt := Eval lazy beta delta[iter] in"; + pr " (iter _ "; + for i = 0 to size do + pr " (fun x y => reduce_%i (w%i_mod_gt x y))" i i; + pr " (fun n x y => reduce_%i (w%i_mod_gt x (DoubleBase.get_low %s (S n) y)))" i i (pz i); + pr " (fun n x y => reduce_%i (w%i_modn1 (S n) x y))" i i; + done; + pr " mod_gtnm)."; + pr ""; + + pp " Let spec_modn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) := "; + pp " (spec_double_modn1 "; + pp " ww_op.(znz_zdigits) ww_op.(znz_0)"; + pp " (znz_WW ww_op) ww_op.(znz_head0)"; + pp " ww_op.(znz_add_mul_div) ww_op.(znz_div21)"; + pp " ww_op.(znz_compare) ww_op.(znz_sub) (znz_to_Z ww_op)"; + pp " (spec_to_Z ww_spec) "; + pp " (spec_zdigits ww_spec)"; + pp " (spec_0 ww_spec) (spec_WW ww_spec) (spec_head0 ww_spec)"; + pp " (spec_add_mul_div ww_spec) (spec_div21 ww_spec) "; + pp " (CyclicAxioms.spec_compare ww_spec) (CyclicAxioms.spec_sub ww_spec))."; + pp ""; + + pr " Theorem spec_mod_gt:"; + pr " forall x y, [x] > [y] -> 0 < [y] -> [mod_gt x y] = [x] mod [y]."; + pa " Admitted."; + pp " Proof."; + pp " refine (spec_iter _ (fun x y res => x > y -> 0 < y ->"; + pp " [res] = x mod y)"; + for i = 0 to size do + pp " (fun x y => reduce_%i (w%i_mod_gt x y))" i i; + pp " (fun n x y => reduce_%i (w%i_mod_gt x (DoubleBase.get_low %s (S n) y)))" i i (pz i); + pp " (fun n x y => reduce_%i (w%i_modn1 (S n) x y)) _ _ _" i i; + done; + pp " mod_gtnm _)."; + for i = 0 to size do + pp " intros x y H1 H2; rewrite spec_reduce_%i." i; + pp " exact (spec_mod_gt w%i_spec x y H1 H2)." i; + if i == size then + pp " intros n x y H2 H3; rewrite spec_reduce_%i." i + else + pp " intros n x y H1 H2 H3; rewrite spec_reduce_%i." i; + pp " unfold w%i_mod_gt." i; + pp " rewrite <- (spec_get_end%i (S n) y x); auto with zarith." i; + pp " unfold to_Z; apply (spec_mod_gt w%i_spec); auto." i; + pp " rewrite <- (spec_get_end%i (S n) y x) in H2; auto with zarith." i; + pp " rewrite <- (spec_get_end%i (S n) y x) in H3; auto with zarith." i; + if i == size then + pp " intros n x y H2 H3; rewrite spec_reduce_%i." i + else + pp " intros n x y H1 H2 H3; rewrite spec_reduce_%i." i; + pp " unfold w%i_modn1, to_Z; rewrite spec_double_eval%in." i i; + pp " apply (spec_modn1 _ _ w%i_spec); auto." i; + done; + pp " intros n m x y H1 H2; unfold mod_gtnm."; + pp " repeat rewrite spec_reduce_n."; + pp " rewrite <- (spec_cast_l n m x)."; + pp " rewrite <- (spec_cast_r n m y)."; + pp " unfold to_Z; apply (spec_mod_gt (wn_spec (Max.max n m)))."; + pp " rewrite <- (spec_cast_l n m x) in H1; auto."; + pp " rewrite <- (spec_cast_r n m y) in H1; auto."; + pp " rewrite <- (spec_cast_r n m y) in H2; auto."; + pp " Qed."; + pr ""; + + pr " Definition modulo x y := "; + pr " match compare x y with"; + pr " | Eq => zero"; + pr " | Lt => x"; + pr " | Gt => mod_gt x y"; + pr " end."; + pr ""; + + pr " Theorem spec_modulo:"; + pr " forall x y, 0 < [y] -> [modulo x y] = [x] mod [y]."; + pa " Admitted."; + pp " Proof."; + pp " assert (F0: [zero] = 0)."; + pp " exact (spec_0 w0_spec)."; + pp " assert (F1: [one] = 1)."; + pp " exact (spec_1 w0_spec)."; + pp " intros x y H; generalize (spec_compare x y);"; + pp " unfold modulo; case compare; try rewrite F0;"; + pp " try rewrite F1; intros; try split; auto with zarith."; + pp " rewrite H0; apply sym_equal; apply Z_mod_same; auto with zarith."; + pp " apply sym_equal; apply Zmod_small; auto with zarith."; + pp " generalize (spec_pos x); auto with zarith."; + pp " apply spec_mod_gt; auto."; + pp " Qed."; + pr ""; + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Gcd *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + pr " Definition digits x :="; + pr " match x with"; + for i = 0 to size do + pr " | %s%i _ => w%i_op.(znz_digits)" c i i; + done; + pr " | %sn n _ => (make_op n).(znz_digits)" c; + pr " end."; + pr ""; + + pr " Theorem spec_digits: forall x, 0 <= [x] < 2 ^ Zpos (digits x)."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; clear x."; + for i = 0 to size do + pp " intros x; unfold to_Z, digits;"; + pp " generalize (spec_to_Z w%i_spec x); unfold base; intros H; exact H." i; + done; + pp " intros n x; unfold to_Z, digits;"; + pp " generalize (spec_to_Z (wn_spec n) x); unfold base; intros H; exact H."; + pp " Qed."; + pr ""; + + pr " Definition gcd_gt_body a b cont :="; + pr " match compare b zero with"; + pr " | Gt =>"; + pr " let r := mod_gt a b in"; + pr " match compare r zero with"; + pr " | Gt => cont r (mod_gt b r)"; + pr " | _ => b"; + pr " end"; + pr " | _ => a"; + pr " end."; + pr ""; + + pp " Theorem Zspec_gcd_gt_body: forall a b cont p,"; + pp " [a] > [b] -> [a] < 2 ^ p ->"; + pp " (forall a1 b1, [a1] < 2 ^ (p - 1) -> [a1] > [b1] ->"; + pp " Zis_gcd [a1] [b1] [cont a1 b1]) -> "; + pp " Zis_gcd [a] [b] [gcd_gt_body a b cont]."; + pp " Proof."; + pp " assert (F1: [zero] = 0)."; + pp " unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto."; + pp " intros a b cont p H2 H3 H4; unfold gcd_gt_body."; + pp " generalize (spec_compare b zero); case compare; try rewrite F1."; + pp " intros HH; rewrite HH; apply Zis_gcd_0."; + pp " intros HH; absurd (0 <= [b]); auto with zarith."; + pp " case (spec_digits b); auto with zarith."; + pp " intros H5; generalize (spec_compare (mod_gt a b) zero); "; + pp " case compare; try rewrite F1."; + pp " intros H6; rewrite <- (Zmult_1_r [b])."; + pp " rewrite (Z_div_mod_eq [a] [b]); auto with zarith."; + pp " rewrite <- spec_mod_gt; auto with zarith."; + pp " rewrite H6; rewrite Zplus_0_r."; + pp " apply Zis_gcd_mult; apply Zis_gcd_1."; + pp " intros; apply False_ind."; + pp " case (spec_digits (mod_gt a b)); auto with zarith."; + pp " intros H6; apply DoubleDiv.Zis_gcd_mod; auto with zarith."; + pp " apply DoubleDiv.Zis_gcd_mod; auto with zarith."; + pp " rewrite <- spec_mod_gt; auto with zarith."; + pp " assert (F2: [b] > [mod_gt a b])."; + pp " case (Z_mod_lt [a] [b]); auto with zarith."; + pp " repeat rewrite <- spec_mod_gt; auto with zarith."; + pp " assert (F3: [mod_gt a b] > [mod_gt b (mod_gt a b)])."; + pp " case (Z_mod_lt [b] [mod_gt a b]); auto with zarith."; + pp " rewrite <- spec_mod_gt; auto with zarith."; + pp " repeat rewrite <- spec_mod_gt; auto with zarith."; + pp " apply H4; auto with zarith."; + pp " apply Zmult_lt_reg_r with 2; auto with zarith."; + pp " apply Zle_lt_trans with ([b] + [mod_gt a b]); auto with zarith."; + pp " apply Zle_lt_trans with (([a]/[b]) * [b] + [mod_gt a b]); auto with zarith."; + pp " apply Zplus_le_compat_r."; + pp " pattern [b] at 1; rewrite <- (Zmult_1_l [b])."; + pp " apply Zmult_le_compat_r; auto with zarith."; + pp " case (Zle_lt_or_eq 0 ([a]/[b])); auto with zarith."; + pp " intros HH; rewrite (Z_div_mod_eq [a] [b]) in H2;"; + pp " try rewrite <- HH in H2; auto with zarith."; + pp " case (Z_mod_lt [a] [b]); auto with zarith."; + pp " rewrite Zmult_comm; rewrite spec_mod_gt; auto with zarith."; + pp " rewrite <- Z_div_mod_eq; auto with zarith."; + pp " pattern 2 at 2; rewrite <- (Zpower_1_r 2)."; + pp " rewrite <- Zpower_exp; auto with zarith."; + pp " ring_simplify (p - 1 + 1); auto."; + pp " case (Zle_lt_or_eq 0 p); auto with zarith."; + pp " generalize H3; case p; simpl Zpower; auto with zarith."; + pp " intros HH; generalize H3; rewrite <- HH; simpl Zpower; auto with zarith."; + pp " Qed."; + pp ""; + + pr " Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) {struct p} : t :="; + pr " gcd_gt_body a b"; + pr " (fun a b =>"; + pr " match p with"; + pr " | xH => cont a b"; + pr " | xO p => gcd_gt_aux p (gcd_gt_aux p cont) a b"; + pr " | xI p => gcd_gt_aux p (gcd_gt_aux p cont) a b"; + pr " end)."; + pr ""; + + pp " Theorem Zspec_gcd_gt_aux: forall p n a b cont,"; + pp " [a] > [b] -> [a] < 2 ^ (Zpos p + n) ->"; + pp " (forall a1 b1, [a1] < 2 ^ n -> [a1] > [b1] ->"; + pp " Zis_gcd [a1] [b1] [cont a1 b1]) ->"; + pp " Zis_gcd [a] [b] [gcd_gt_aux p cont a b]."; + pp " intros p; elim p; clear p."; + pp " intros p Hrec n a b cont H2 H3 H4."; + pp " unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xI p) + n); auto."; + pp " intros a1 b1 H6 H7."; + pp " apply Hrec with (Zpos p + n); auto."; + pp " replace (Zpos p + (Zpos p + n)) with"; + pp " (Zpos (xI p) + n - 1); auto."; + pp " rewrite Zpos_xI; ring."; + pp " intros a2 b2 H9 H10."; + pp " apply Hrec with n; auto."; + pp " intros p Hrec n a b cont H2 H3 H4."; + pp " unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xO p) + n); auto."; + pp " intros a1 b1 H6 H7."; + pp " apply Hrec with (Zpos p + n - 1); auto."; + pp " replace (Zpos p + (Zpos p + n - 1)) with"; + pp " (Zpos (xO p) + n - 1); auto."; + pp " rewrite Zpos_xO; ring."; + pp " intros a2 b2 H9 H10."; + pp " apply Hrec with (n - 1); auto."; + pp " replace (Zpos p + (n - 1)) with"; + pp " (Zpos p + n - 1); auto with zarith."; + pp " intros a3 b3 H12 H13; apply H4; auto with zarith."; + pp " apply Zlt_le_trans with (1 := H12)."; + pp " case (Zle_or_lt 1 n); intros HH."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " apply Zle_trans with 0; auto with zarith."; + pp " assert (HH1: n - 1 < 0); auto with zarith."; + pp " generalize HH1; case (n - 1); auto with zarith."; + pp " intros p1 HH2; discriminate."; + pp " intros n a b cont H H2 H3."; + pp " simpl gcd_gt_aux."; + pp " apply Zspec_gcd_gt_body with (n + 1); auto with zarith."; + pp " rewrite Zplus_comm; auto."; + pp " intros a1 b1 H5 H6; apply H3; auto."; + pp " replace n with (n + 1 - 1); auto; try ring."; + pp " Qed."; + pp ""; + + pr " Definition gcd_cont a b :="; + pr " match compare one b with"; + pr " | Eq => one"; + pr " | _ => a"; + pr " end."; + pr ""; + + pr " Definition gcd_gt a b := gcd_gt_aux (digits a) gcd_cont a b."; + pr ""; + + pr " Theorem spec_gcd_gt: forall a b,"; + pr " [a] > [b] -> [gcd_gt a b] = Zgcd [a] [b]."; + pa " Admitted."; + pp " Proof."; + pp " intros a b H2."; + pp " case (spec_digits (gcd_gt a b)); intros H3 H4."; + pp " case (spec_digits a); intros H5 H6."; + pp " apply sym_equal; apply Zis_gcd_gcd; auto with zarith."; + pp " unfold gcd_gt; apply Zspec_gcd_gt_aux with 0; auto with zarith."; + pp " intros a1 a2; rewrite Zpower_0_r."; + pp " case (spec_digits a2); intros H7 H8;"; + pp " intros; apply False_ind; auto with zarith."; + pp " Qed."; + pr ""; + + pr " Definition gcd a b :="; + pr " match compare a b with"; + pr " | Eq => a"; + pr " | Lt => gcd_gt b a"; + pr " | Gt => gcd_gt a b"; + pr " end."; + pr ""; + + pr " Theorem spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b]."; + pa " Admitted."; + pp " Proof."; + pp " intros a b."; + pp " case (spec_digits a); intros H1 H2."; + pp " case (spec_digits b); intros H3 H4."; + pp " unfold gcd; generalize (spec_compare a b); case compare."; + pp " intros HH; rewrite HH; apply sym_equal; apply Zis_gcd_gcd; auto."; + pp " apply Zis_gcd_refl."; + pp " intros; apply trans_equal with (Zgcd [b] [a])."; + pp " apply spec_gcd_gt; auto with zarith."; + pp " apply Zis_gcd_gcd; auto with zarith."; + pp " apply Zgcd_is_pos."; + pp " apply Zis_gcd_sym; apply Zgcd_is_gcd."; + pp " intros; apply spec_gcd_gt; auto."; + pp " Qed."; + pr ""; + + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Conversion *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + pr " Definition pheight p := "; + pr " Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (plength p)))."; + pr ""; + + pr " Theorem pheight_correct: forall p, "; + pr " Zpos p < 2 ^ (Zpos (znz_digits w0_op) * 2 ^ (Z_of_nat (pheight p)))."; + pr " Proof."; + pr " intros p; unfold pheight."; + pr " assert (F1: forall x, Z_of_nat (Peano.pred (nat_of_P x)) = Zpos x - 1)."; + pr " intros x."; + pr " assert (Zsucc (Z_of_nat (Peano.pred (nat_of_P x))) = Zpos x); auto with zarith."; + pr " rewrite <- inj_S."; + pr " rewrite <- (fun x => S_pred x 0); auto with zarith."; + pr " rewrite Zpos_eq_Z_of_nat_o_nat_of_P; auto."; + pr " apply lt_le_trans with 1%snat; auto with zarith." "%"; + pr " exact (le_Pmult_nat x 1)."; + pr " rewrite F1; clear F1."; + pr " assert (F2:= (get_height_correct (znz_digits w0_op) (plength p)))."; + pr " apply Zlt_le_trans with (Zpos (Psucc p))."; + pr " rewrite Zpos_succ_morphism; auto with zarith."; + pr " apply Zle_trans with (1 := plength_pred_correct (Psucc p))."; + pr " rewrite Ppred_succ."; + pr " apply Zpower_le_monotone; auto with zarith."; + pr " Qed."; + pr ""; + + pr " Definition of_pos x :="; + pr " let h := pheight x in"; + pr " match h with"; + for i = 0 to size do + pr " | %i%snat => reduce_%i (snd (w%i_op.(znz_of_pos) x))" i "%" i i; + done; + pr " | _ =>"; + pr " let n := minus h %i in" (size + 1); + pr " reduce_n n (snd ((make_op n).(znz_of_pos) x))"; + pr " end."; + pr ""; + + pr " Theorem spec_of_pos: forall x,"; + pr " [of_pos x] = Zpos x."; + pa " Admitted."; + pp " Proof."; + pp " assert (F := spec_more_than_1_digit w0_spec)."; + pp " intros x; unfold of_pos; case_eq (pheight x)."; + for i = 0 to size do + if i <> 0 then + pp " intros n; case n; clear n."; + pp " intros H1; rewrite spec_reduce_%i; unfold to_Z." i; + pp " apply (znz_of_pos_correct w%i_spec)." i; + pp " apply Zlt_le_trans with (1 := pheight_correct x)."; + pp " rewrite H1; simpl Z_of_nat; change (2^%i) with (%s)." i (gen2 i); + pp " unfold base."; + pp " apply Zpower_le_monotone; split; auto with zarith."; + if i <> 0 then + begin + pp " rewrite Zmult_comm; repeat rewrite <- Zmult_assoc."; + pp " repeat rewrite <- Zpos_xO."; + pp " refine (Zle_refl _)."; + end; + done; + pp " intros n."; + pp " intros H1; rewrite spec_reduce_n; unfold to_Z."; + pp " simpl minus; rewrite <- minus_n_O."; + pp " apply (znz_of_pos_correct (wn_spec n))."; + pp " apply Zlt_le_trans with (1 := pheight_correct x)."; + pp " unfold base."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " split; auto with zarith."; + pp " rewrite H1."; + pp " elim n; clear n H1."; + pp " simpl Z_of_nat; change (2^%i) with (%s)." (size + 1) (gen2 (size + 1)); + pp " rewrite Zmult_comm; repeat rewrite <- Zmult_assoc."; + pp " repeat rewrite <- Zpos_xO."; + pp " refine (Zle_refl _)."; + pp " intros n Hrec."; + pp " rewrite make_op_S."; + pp " change (@znz_digits (word _ (S (S n))) (mk_zn2z_op_karatsuba (make_op n))) with"; + pp " (xO (znz_digits (make_op n)))."; + pp " rewrite (fun x y => (Zpos_xO (@znz_digits x y)))."; + pp " rewrite inj_S; unfold Zsucc."; + pp " rewrite Zplus_comm; rewrite Zpower_exp; auto with zarith."; + pp " rewrite Zpower_1_r."; + pp " assert (tmp: forall x y z, x * (y * z) = y * (x * z));"; + pp " [intros; ring | rewrite tmp; clear tmp]."; + pp " apply Zmult_le_compat_l; auto with zarith."; + pp " Qed."; + pr ""; + + pr " Definition of_N x :="; + pr " match x with"; + pr " | BinNat.N0 => zero"; + pr " | Npos p => of_pos p"; + pr " end."; + pr ""; + + pr " Theorem spec_of_N: forall x,"; + pr " [of_N x] = Z_of_N x."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x."; + pp " simpl of_N."; + pp " unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto."; + pp " intros p; exact (spec_of_pos p)."; + pp " Qed."; + pr ""; + + pr " (***************************************************************)"; + pr " (* *)"; + pr " (* Shift *)"; + pr " (* *)"; + pr " (***************************************************************)"; + pr ""; + + (* Head0 *) + pr " Definition head0 w := match w with"; + for i = 0 to size do + pr " | %s%i w=> reduce_%i (w%i_op.(znz_head0) w)" c i i i; + done; + pr " | %sn n w=> reduce_n n ((make_op n).(znz_head0) w)" c; + pr " end."; + pr ""; + + pr " Theorem spec_head00: forall x, [x] = 0 ->[head0 x] = Zpos (digits x)."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; unfold head0; clear x."; + for i = 0 to size do + pp " intros x; rewrite spec_reduce_%i; exact (spec_head00 w%i_spec x)." i i; + done; + pp " intros n x; rewrite spec_reduce_n; exact (spec_head00 (wn_spec n) x)."; + pp " Qed."; + pr " "; + + pr " Theorem spec_head0: forall x, 0 < [x] ->"; + pr " 2 ^ (Zpos (digits x) - 1) <= 2 ^ [head0 x] * [x] < 2 ^ Zpos (digits x)."; + pa " Admitted."; + pp " Proof."; + pp " assert (F0: forall x, (x - 1) + 1 = x)."; + pp " intros; ring. "; + pp " intros x; case x; unfold digits, head0; clear x."; + for i = 0 to size do + pp " intros x Hx; rewrite spec_reduce_%i." i; + pp " assert (F1:= spec_more_than_1_digit w%i_spec)." i; + pp " generalize (spec_head0 w%i_spec x Hx)." i; + pp " unfold base."; + pp " pattern (Zpos (znz_digits w%i_op)) at 1; " i; + pp " rewrite <- (fun x => (F0 (Zpos x)))."; + pp " rewrite Zpower_exp; auto with zarith."; + pp " rewrite Zpower_1_r; rewrite Z_div_mult; auto with zarith."; + done; + pp " intros n x Hx; rewrite spec_reduce_n."; + pp " assert (F1:= spec_more_than_1_digit (wn_spec n))."; + pp " generalize (spec_head0 (wn_spec n) x Hx)."; + pp " unfold base."; + pp " pattern (Zpos (znz_digits (make_op n))) at 1; "; + pp " rewrite <- (fun x => (F0 (Zpos x)))."; + pp " rewrite Zpower_exp; auto with zarith."; + pp " rewrite Zpower_1_r; rewrite Z_div_mult; auto with zarith."; + pp " Qed."; + pr ""; + + + (* Tail0 *) + pr " Definition tail0 w := match w with"; + for i = 0 to size do + pr " | %s%i w=> reduce_%i (w%i_op.(znz_tail0) w)" c i i i; + done; + pr " | %sn n w=> reduce_n n ((make_op n).(znz_tail0) w)" c; + pr " end."; + pr ""; + + + pr " Theorem spec_tail00: forall x, [x] = 0 ->[tail0 x] = Zpos (digits x)."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; unfold tail0; clear x."; + for i = 0 to size do + pp " intros x; rewrite spec_reduce_%i; exact (spec_tail00 w%i_spec x)." i i; + done; + pp " intros n x; rewrite spec_reduce_n; exact (spec_tail00 (wn_spec n) x)."; + pp " Qed."; + pr " "; + + + pr " Theorem spec_tail0: forall x,"; + pr " 0 < [x] -> exists y, 0 <= y /\\ [x] = (2 * y + 1) * 2 ^ [tail0 x]."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; clear x; unfold tail0."; + for i = 0 to size do + pp " intros x Hx; rewrite spec_reduce_%i; exact (spec_tail0 w%i_spec x Hx)." i i; + done; + pp " intros n x Hx; rewrite spec_reduce_n; exact (spec_tail0 (wn_spec n) x Hx)."; + pp " Qed."; + pr ""; + + + (* Number of digits *) + pr " Definition %sdigits x :=" c; + pr " match x with"; + pr " | %s0 _ => %s0 w0_op.(znz_zdigits)" c c; + for i = 1 to size do + pr " | %s%i _ => reduce_%i w%i_op.(znz_zdigits)" c i i i; + done; + pr " | %sn n _ => reduce_n n (make_op n).(znz_zdigits)" c; + pr " end."; + pr ""; + + pr " Theorem spec_Ndigits: forall x, [Ndigits x] = Zpos (digits x)."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; clear x; unfold Ndigits, digits."; + for i = 0 to size do + pp " intros _; try rewrite spec_reduce_%i; exact (spec_zdigits w%i_spec)." i i; + done; + pp " intros n _; try rewrite spec_reduce_n; exact (spec_zdigits (wn_spec n))."; + pp " Qed."; + pr ""; + + + (* Shiftr *) + for i = 0 to size do + pr " Definition shiftr%i n x := w%i_op.(znz_add_mul_div) (w%i_op.(znz_sub) w%i_op.(znz_zdigits) n) w%i_op.(znz_0) x." i i i i i; + done; + pr " Definition shiftrn n p x := (make_op n).(znz_add_mul_div) ((make_op n).(znz_sub) (make_op n).(znz_zdigits) p) (make_op n).(znz_0) x."; + pr ""; + + pr " Definition shiftr := Eval lazy beta delta [same_level] in "; + pr " same_level _ (fun n x => %s0 (shiftr0 n x))" c; + for i = 1 to size do + pr " (fun n x => reduce_%i (shiftr%i n x))" i i; + done; + pr " (fun n p x => reduce_n n (shiftrn n p x))."; + pr ""; + + + pr " Theorem spec_shiftr: forall n x,"; + pr " [n] <= [Ndigits x] -> [shiftr n x] = [x] / 2 ^ [n]."; + pa " Admitted."; + pp " Proof."; + pp " assert (F0: forall x y, x - (x - y) = y)."; + pp " intros; ring."; + pp " assert (F2: forall x y z, 0 <= x -> 0 <= y -> x < z -> 0 <= x / 2 ^ y < z)."; + pp " intros x y z HH HH1 HH2."; + pp " split; auto with zarith."; + pp " apply Zle_lt_trans with (2 := HH2); auto with zarith."; + pp " apply Zdiv_le_upper_bound; auto with zarith."; + pp " pattern x at 1; replace x with (x * 2 ^ 0); auto with zarith."; + pp " apply Zmult_le_compat_l; auto."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " rewrite Zpower_0_r; ring."; + pp " assert (F3: forall x y, 0 <= y -> y <= x -> 0 <= x - y < 2 ^ x)."; + pp " intros xx y HH HH1."; + pp " split; auto with zarith."; + pp " apply Zle_lt_trans with xx; auto with zarith."; + pp " apply Zpower2_lt_lin; auto with zarith."; + pp " assert (F4: forall ww ww1 ww2 "; + pp " (ww_op: znz_op ww) (ww1_op: znz_op ww1) (ww2_op: znz_op ww2)"; + pp " xx yy xx1 yy1,"; + pp " znz_to_Z ww2_op yy <= znz_to_Z ww1_op (znz_zdigits ww1_op) ->"; + pp " znz_to_Z ww1_op (znz_zdigits ww1_op) <= znz_to_Z ww_op (znz_zdigits ww_op) ->"; + pp " znz_spec ww_op -> znz_spec ww1_op -> znz_spec ww2_op ->"; + pp " znz_to_Z ww_op xx1 = znz_to_Z ww1_op xx ->"; + pp " znz_to_Z ww_op yy1 = znz_to_Z ww2_op yy ->"; + pp " znz_to_Z ww_op"; + pp " (znz_add_mul_div ww_op (znz_sub ww_op (znz_zdigits ww_op) yy1)"; + pp " (znz_0 ww_op) xx1) = znz_to_Z ww1_op xx / 2 ^ znz_to_Z ww2_op yy)."; + pp " intros ww ww1 ww2 ww_op ww1_op ww2_op xx yy xx1 yy1 Hl Hl1 Hw Hw1 Hw2 Hx Hy."; + pp " case (spec_to_Z Hw xx1); auto with zarith; intros HH1 HH2."; + pp " case (spec_to_Z Hw yy1); auto with zarith; intros HH3 HH4."; + pp " rewrite <- Hx."; + pp " rewrite <- Hy."; + pp " generalize (spec_add_mul_div Hw"; + pp " (znz_0 ww_op) xx1"; + pp " (znz_sub ww_op (znz_zdigits ww_op) "; + pp " yy1)"; + pp " )."; + pp " rewrite (spec_0 Hw)."; + pp " rewrite Zmult_0_l; rewrite Zplus_0_l."; + pp " rewrite (CyclicAxioms.spec_sub Hw)."; + pp " rewrite Zmod_small; auto with zarith."; + pp " rewrite (spec_zdigits Hw)."; + pp " rewrite F0."; + pp " rewrite Zmod_small; auto with zarith."; + pp " unfold base; rewrite (spec_zdigits Hw) in Hl1 |- *;"; + pp " auto with zarith."; + pp " assert (F5: forall n m, (n <= m)%snat ->" "%"; + pp " Zpos (znz_digits (make_op n)) <= Zpos (znz_digits (make_op m)))."; + pp " intros n m HH; elim HH; clear m HH; auto with zarith."; + pp " intros m HH Hrec; apply Zle_trans with (1 := Hrec)."; + pp " rewrite make_op_S."; + pp " match goal with |- Zpos ?Y <= ?X => change X with (Zpos (xO Y)) end."; + pp " rewrite Zpos_xO."; + pp " assert (0 <= Zpos (znz_digits (make_op n))); auto with zarith."; + pp " assert (F6: forall n, Zpos (znz_digits w%i_op) <= Zpos (znz_digits (make_op n)))." size; + pp " intros n ; apply Zle_trans with (Zpos (znz_digits (make_op 0)))."; + pp " change (znz_digits (make_op 0)) with (xO (znz_digits w%i_op))." size; + pp " rewrite Zpos_xO."; + pp " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith." size; + pp " apply F5; auto with arith."; + pp " intros x; case x; clear x; unfold shiftr, same_level."; + for i = 0 to size do + pp " intros x y; case y; clear y."; + for j = 0 to i - 1 do + pp " intros y; unfold shiftr%i, Ndigits." i; + pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j; + pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i j i; + pp " rewrite (spec_zdigits w%i_spec)." i; + pp " rewrite (spec_zdigits w%i_spec)." j; + pp " change (znz_digits w%i_op) with %s." i (genxO (i - j) (" (znz_digits w"^(string_of_int j)^"_op)")); + pp " repeat rewrite (fun x => Zpos_xO (xO x))."; + pp " repeat rewrite (fun x y => Zpos_xO (@znz_digits x y))."; + pp " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith." j; + pp " try (apply sym_equal; exact (spec_extend%in%i y))." j i; + + done; + pp " intros y; unfold shiftr%i, Ndigits." i; + pp " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i; + pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i i i; + for j = i + 1 to size do + pp " intros y; unfold shiftr%i, Ndigits." j; + pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j; + pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." j j i; + pp " try (apply sym_equal; exact (spec_extend%in%i x))." i j; + done; + if i == size then + begin + pp " intros m y; unfold shiftrn, Ndigits."; + pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1."; + pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." size; + pp " try (apply sym_equal; exact (spec_extend%in m x))." size; + end + else + begin + pp " intros m y; unfold shiftrn, Ndigits."; + pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1."; + pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." i; + pp " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x])." size i (size - i - 1) i; + pp " rewrite <- (spec_extend%in m); rewrite <- spec_extend%in%i; auto." size i size; + end + done; + pp " intros n x y; case y; clear y;"; + pp " intros y; unfold shiftrn, Ndigits; try rewrite spec_reduce_n."; + for i = 0 to size do + pp " try rewrite spec_reduce_%i; unfold to_Z; intros H1." i; + pp " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith." i; + pp " rewrite (spec_zdigits w%i_spec)." i; + pp " rewrite (spec_zdigits (wn_spec n))."; + pp " apply Zle_trans with (2 := F6 n)."; + pp " change (znz_digits w%i_op) with %s." size (genxO (size - i) ("(znz_digits w" ^ (string_of_int i) ^ "_op)")); + pp " repeat rewrite (fun x => Zpos_xO (xO x))."; + pp " repeat rewrite (fun x y => Zpos_xO (@znz_digits x y))."; + pp " assert (H: 0 <= Zpos (znz_digits w%i_op)); auto with zarith." i; + if i == size then + pp " change ([Nn n (extend%i n y)] = [N%i y])." size i + else + pp " change ([Nn n (extend%i n (extend%i %i y))] = [N%i y])." size i (size - i - 1) i; + pp " rewrite <- (spec_extend%in n); auto." size; + if i <> size then + pp " try (rewrite <- spec_extend%in%i; auto)." i size; + done; + pp " generalize y; clear y; intros m y."; + pp " rewrite spec_reduce_n; unfold to_Z; intros H1."; + pp " apply F4 with (3:=(wn_spec (Max.max n m)))(4:=wn_spec m)(5:=wn_spec n); auto with zarith."; + pp " rewrite (spec_zdigits (wn_spec m))."; + pp " rewrite (spec_zdigits (wn_spec (Max.max n m)))."; + pp " apply F5; auto with arith."; + pp " exact (spec_cast_r n m y)."; + pp " exact (spec_cast_l n m x)."; + pp " Qed."; + pr ""; + + pr " Definition safe_shiftr n x := "; + pr " match compare n (Ndigits x) with"; + pr " | Lt => shiftr n x "; + pr " | _ => %s0 w_0" c; + pr " end."; + pr ""; + + + pr " Theorem spec_safe_shiftr: forall n x,"; + pr " [safe_shiftr n x] = [x] / 2 ^ [n]."; + pa " Admitted."; + pp " Proof."; + pp " intros n x; unfold safe_shiftr;"; + pp " generalize (spec_compare n (Ndigits x)); case compare; intros H."; + pp " apply trans_equal with (1 := spec_0 w0_spec)."; + pp " apply sym_equal; apply Zdiv_small; rewrite H."; + pp " rewrite spec_Ndigits; exact (spec_digits x)."; + pp " rewrite <- spec_shiftr; auto with zarith."; + pp " apply trans_equal with (1 := spec_0 w0_spec)."; + pp " apply sym_equal; apply Zdiv_small."; + pp " rewrite spec_Ndigits in H; case (spec_digits x); intros H1 H2."; + pp " split; auto."; + pp " apply Zlt_le_trans with (1 := H2)."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " Qed."; + pr ""; + + pr ""; + + (* Shiftl *) + for i = 0 to size do + pr " Definition shiftl%i n x := w%i_op.(znz_add_mul_div) n x w%i_op.(znz_0)." i i i + done; + pr " Definition shiftln n p x := (make_op n).(znz_add_mul_div) p x (make_op n).(znz_0)."; + pr " Definition shiftl := Eval lazy beta delta [same_level] in"; + pr " same_level _ (fun n x => %s0 (shiftl0 n x))" c; + for i = 1 to size do + pr " (fun n x => reduce_%i (shiftl%i n x))" i i; + done; + pr " (fun n p x => reduce_n n (shiftln n p x))."; + pr ""; + pr ""; + + + pr " Theorem spec_shiftl: forall n x,"; + pr " [n] <= [head0 x] -> [shiftl n x] = [x] * 2 ^ [n]."; + pa " Admitted."; + pp " Proof."; + pp " assert (F0: forall x y, x - (x - y) = y)."; + pp " intros; ring."; + pp " assert (F2: forall x y z, 0 <= x -> 0 <= y -> x < z -> 0 <= x / 2 ^ y < z)."; + pp " intros x y z HH HH1 HH2."; + pp " split; auto with zarith."; + pp " apply Zle_lt_trans with (2 := HH2); auto with zarith."; + pp " apply Zdiv_le_upper_bound; auto with zarith."; + pp " pattern x at 1; replace x with (x * 2 ^ 0); auto with zarith."; + pp " apply Zmult_le_compat_l; auto."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " rewrite Zpower_0_r; ring."; + pp " assert (F3: forall x y, 0 <= y -> y <= x -> 0 <= x - y < 2 ^ x)."; + pp " intros xx y HH HH1."; + pp " split; auto with zarith."; + pp " apply Zle_lt_trans with xx; auto with zarith."; + pp " apply Zpower2_lt_lin; auto with zarith."; + pp " assert (F4: forall ww ww1 ww2 "; + pp " (ww_op: znz_op ww) (ww1_op: znz_op ww1) (ww2_op: znz_op ww2)"; + pp " xx yy xx1 yy1,"; + pp " znz_to_Z ww2_op yy <= znz_to_Z ww1_op (znz_head0 ww1_op xx) ->"; + pp " znz_to_Z ww1_op (znz_zdigits ww1_op) <= znz_to_Z ww_op (znz_zdigits ww_op) ->"; + pp " znz_spec ww_op -> znz_spec ww1_op -> znz_spec ww2_op ->"; + pp " znz_to_Z ww_op xx1 = znz_to_Z ww1_op xx ->"; + pp " znz_to_Z ww_op yy1 = znz_to_Z ww2_op yy ->"; + pp " znz_to_Z ww_op"; + pp " (znz_add_mul_div ww_op yy1"; + pp " xx1 (znz_0 ww_op)) = znz_to_Z ww1_op xx * 2 ^ znz_to_Z ww2_op yy)."; + pp " intros ww ww1 ww2 ww_op ww1_op ww2_op xx yy xx1 yy1 Hl Hl1 Hw Hw1 Hw2 Hx Hy."; + pp " case (spec_to_Z Hw xx1); auto with zarith; intros HH1 HH2."; + pp " case (spec_to_Z Hw yy1); auto with zarith; intros HH3 HH4."; + pp " rewrite <- Hx."; + pp " rewrite <- Hy."; + pp " generalize (spec_add_mul_div Hw xx1 (znz_0 ww_op) yy1)."; + pp " rewrite (spec_0 Hw)."; + pp " assert (F1: znz_to_Z ww1_op (znz_head0 ww1_op xx) <= Zpos (znz_digits ww1_op))."; + pp " case (Zle_lt_or_eq _ _ HH1); intros HH5."; + pp " apply Zlt_le_weak."; + pp " case (CyclicAxioms.spec_head0 Hw1 xx)."; + pp " rewrite <- Hx; auto."; + pp " intros _ Hu; unfold base in Hu."; + pp " case (Zle_or_lt (Zpos (znz_digits ww1_op))"; + pp " (znz_to_Z ww1_op (znz_head0 ww1_op xx))); auto; intros H1."; + pp " absurd (2 ^ (Zpos (znz_digits ww1_op)) <= 2 ^ (znz_to_Z ww1_op (znz_head0 ww1_op xx)))."; + pp " apply Zlt_not_le."; + pp " case (spec_to_Z Hw1 xx); intros HHx3 HHx4."; + pp " rewrite <- (Zmult_1_r (2 ^ znz_to_Z ww1_op (znz_head0 ww1_op xx)))."; + pp " apply Zle_lt_trans with (2 := Hu)."; + pp " apply Zmult_le_compat_l; auto with zarith."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " rewrite (CyclicAxioms.spec_head00 Hw1 xx); auto with zarith."; + pp " rewrite Zdiv_0_l; auto with zarith."; + pp " rewrite Zplus_0_r."; + pp " case (Zle_lt_or_eq _ _ HH1); intros HH5."; + pp " rewrite Zmod_small; auto with zarith."; + pp " intros HH; apply HH."; + pp " rewrite Hy; apply Zle_trans with (1:= Hl)."; + pp " rewrite <- (spec_zdigits Hw). "; + pp " apply Zle_trans with (2 := Hl1); auto."; + pp " rewrite (spec_zdigits Hw1); auto with zarith."; + pp " split; auto with zarith ."; + pp " apply Zlt_le_trans with (base (znz_digits ww1_op))."; + pp " rewrite Hx."; + pp " case (CyclicAxioms.spec_head0 Hw1 xx); auto."; + pp " rewrite <- Hx; auto."; + pp " intros _ Hu; rewrite Zmult_comm in Hu."; + pp " apply Zle_lt_trans with (2 := Hu)."; + pp " apply Zmult_le_compat_l; auto with zarith."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " unfold base; apply Zpower_le_monotone; auto with zarith."; + pp " split; auto with zarith."; + pp " rewrite <- (spec_zdigits Hw); auto with zarith."; + pp " rewrite <- (spec_zdigits Hw1); auto with zarith."; + pp " rewrite <- HH5."; + pp " rewrite Zmult_0_l."; + pp " rewrite Zmod_small; auto with zarith."; + pp " intros HH; apply HH."; + pp " rewrite Hy; apply Zle_trans with (1 := Hl)."; + pp " rewrite (CyclicAxioms.spec_head00 Hw1 xx); auto with zarith."; + pp " rewrite <- (spec_zdigits Hw); auto with zarith."; + pp " rewrite <- (spec_zdigits Hw1); auto with zarith."; + pp " assert (F5: forall n m, (n <= m)%snat ->" "%"; + pp " Zpos (znz_digits (make_op n)) <= Zpos (znz_digits (make_op m)))."; + pp " intros n m HH; elim HH; clear m HH; auto with zarith."; + pp " intros m HH Hrec; apply Zle_trans with (1 := Hrec)."; + pp " rewrite make_op_S."; + pp " match goal with |- Zpos ?Y <= ?X => change X with (Zpos (xO Y)) end."; + pp " rewrite Zpos_xO."; + pp " assert (0 <= Zpos (znz_digits (make_op n))); auto with zarith."; + pp " assert (F6: forall n, Zpos (znz_digits w%i_op) <= Zpos (znz_digits (make_op n)))." size; + pp " intros n ; apply Zle_trans with (Zpos (znz_digits (make_op 0)))."; + pp " change (znz_digits (make_op 0)) with (xO (znz_digits w%i_op))." size; + pp " rewrite Zpos_xO."; + pp " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith." size; + pp " apply F5; auto with arith."; + pp " intros x; case x; clear x; unfold shiftl, same_level."; + for i = 0 to size do + pp " intros x y; case y; clear y."; + for j = 0 to i - 1 do + pp " intros y; unfold shiftl%i, head0." i; + pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j; + pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i j i; + pp " rewrite (spec_zdigits w%i_spec)." i; + pp " rewrite (spec_zdigits w%i_spec)." j; + pp " change (znz_digits w%i_op) with %s." i (genxO (i - j) (" (znz_digits w"^(string_of_int j)^"_op)")); + pp " repeat rewrite (fun x => Zpos_xO (xO x))."; + pp " repeat rewrite (fun x y => Zpos_xO (@znz_digits x y))."; + pp " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith." j; + pp " try (apply sym_equal; exact (spec_extend%in%i y))." j i; + done; + pp " intros y; unfold shiftl%i, head0." i; + pp " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i; + pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i i i; + for j = i + 1 to size do + pp " intros y; unfold shiftl%i, head0." j; + pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j; + pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." j j i; + pp " try (apply sym_equal; exact (spec_extend%in%i x))." i j; + done; + if i == size then + begin + pp " intros m y; unfold shiftln, head0."; + pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1."; + pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." size; + pp " try (apply sym_equal; exact (spec_extend%in m x))." size; + end + else + begin + pp " intros m y; unfold shiftln, head0."; + pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1."; + pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." i; + pp " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x])." size i (size - i - 1) i; + pp " rewrite <- (spec_extend%in m); rewrite <- spec_extend%in%i; auto." size i size; + end + done; + pp " intros n x y; case y; clear y;"; + pp " intros y; unfold shiftln, head0; try rewrite spec_reduce_n."; + for i = 0 to size do + pp " try rewrite spec_reduce_%i; unfold to_Z; intros H1." i; + pp " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith." i; + pp " rewrite (spec_zdigits w%i_spec)." i; + pp " rewrite (spec_zdigits (wn_spec n))."; + pp " apply Zle_trans with (2 := F6 n)."; + pp " change (znz_digits w%i_op) with %s." size (genxO (size - i) ("(znz_digits w" ^ (string_of_int i) ^ "_op)")); + pp " repeat rewrite (fun x => Zpos_xO (xO x))."; + pp " repeat rewrite (fun x y => Zpos_xO (@znz_digits x y))."; + pp " assert (H: 0 <= Zpos (znz_digits w%i_op)); auto with zarith." i; + if i == size then + pp " change ([Nn n (extend%i n y)] = [N%i y])." size i + else + pp " change ([Nn n (extend%i n (extend%i %i y))] = [N%i y])." size i (size - i - 1) i; + pp " rewrite <- (spec_extend%in n); auto." size; + if i <> size then + pp " try (rewrite <- spec_extend%in%i; auto)." i size; + done; + pp " generalize y; clear y; intros m y."; + pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1."; + pp " apply F4 with (3:=(wn_spec (Max.max n m)))(4:=wn_spec m)(5:=wn_spec n); auto with zarith."; + pp " rewrite (spec_zdigits (wn_spec m))."; + pp " rewrite (spec_zdigits (wn_spec (Max.max n m)))."; + pp " apply F5; auto with arith."; + pp " exact (spec_cast_r n m y)."; + pp " exact (spec_cast_l n m x)."; + pp " Qed."; + pr ""; + + (* Double size *) + pr " Definition double_size w := match w with"; + for i = 0 to size-1 do + pr " | %s%i x => %s%i (WW (znz_0 w%i_op) x)" c i c (i + 1) i; + done; + pr " | %s%i x => %sn 0 (WW (znz_0 w%i_op) x)" c size c size; + pr " | %sn n x => %sn (S n) (WW (znz_0 (make_op n)) x)" c c; + pr " end."; + pr ""; + + pr " Theorem spec_double_size_digits: "; + pr " forall x, digits (double_size x) = xO (digits x)."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; unfold double_size, digits; clear x; auto."; + pp " intros n x; rewrite make_op_S; auto."; + pp " Qed."; + pr ""; + + + pr " Theorem spec_double_size: forall x, [double_size x] = [x]."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; unfold double_size; clear x."; + for i = 0 to size do + pp " intros x; unfold to_Z, make_op; "; + pp " rewrite znz_to_Z_%i; rewrite (spec_0 w%i_spec); auto with zarith." (i + 1) i; + done; + pp " intros n x; unfold to_Z;"; + pp " generalize (znz_to_Z_n n); simpl word."; + pp " intros HH; rewrite HH; clear HH."; + pp " generalize (spec_0 (wn_spec n)); simpl word."; + pp " intros HH; rewrite HH; clear HH; auto with zarith."; + pp " Qed."; + pr ""; + + + pr " Theorem spec_double_size_head0: "; + pr " forall x, 2 * [head0 x] <= [head0 (double_size x)]."; + pa " Admitted."; + pp " Proof."; + pp " intros x."; + pp " assert (F1:= spec_pos (head0 x))."; + pp " assert (F2: 0 < Zpos (digits x))."; + pp " red; auto."; + pp " case (Zle_lt_or_eq _ _ (spec_pos x)); intros HH."; + pp " generalize HH; rewrite <- (spec_double_size x); intros HH1."; + pp " case (spec_head0 x HH); intros _ HH2."; + pp " case (spec_head0 _ HH1)."; + pp " rewrite (spec_double_size x); rewrite (spec_double_size_digits x)."; + pp " intros HH3 _."; + pp " case (Zle_or_lt ([head0 (double_size x)]) (2 * [head0 x])); auto; intros HH4."; + pp " absurd (2 ^ (2 * [head0 x] )* [x] < 2 ^ [head0 (double_size x)] * [x]); auto."; + pp " apply Zle_not_lt."; + pp " apply Zmult_le_compat_r; auto with zarith."; + pp " apply Zpower_le_monotone; auto; auto with zarith."; + pp " generalize (spec_pos (head0 (double_size x))); auto with zarith."; + pp " assert (HH5: 2 ^[head0 x] <= 2 ^(Zpos (digits x) - 1))."; + pp " case (Zle_lt_or_eq 1 [x]); auto with zarith; intros HH5."; + pp " apply Zmult_le_reg_r with (2 ^ 1); auto with zarith."; + pp " rewrite <- (fun x y z => Zpower_exp x (y - z)); auto with zarith."; + pp " assert (tmp: forall x, x - 1 + 1 = x); [intros; ring | rewrite tmp; clear tmp]."; + pp " apply Zle_trans with (2 := Zlt_le_weak _ _ HH2)."; + pp " apply Zmult_le_compat_l; auto with zarith."; + pp " rewrite Zpower_1_r; auto with zarith."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " split; auto with zarith. "; + pp " case (Zle_or_lt (Zpos (digits x)) [head0 x]); auto with zarith; intros HH6."; + pp " absurd (2 ^ Zpos (digits x) <= 2 ^ [head0 x] * [x]); auto with zarith."; + pp " rewrite <- HH5; rewrite Zmult_1_r."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " rewrite (Zmult_comm 2)."; + pp " rewrite Zpower_mult; auto with zarith."; + pp " rewrite Zpower_2."; + pp " apply Zlt_le_trans with (2 := HH3)."; + pp " rewrite <- Zmult_assoc."; + pp " replace (Zpos (xO (digits x)) - 1) with"; + pp " ((Zpos (digits x) - 1) + (Zpos (digits x)))."; + pp " rewrite Zpower_exp; auto with zarith."; + pp " apply Zmult_lt_compat2; auto with zarith."; + pp " split; auto with zarith."; + pp " apply Zmult_lt_0_compat; auto with zarith."; + pp " rewrite Zpos_xO; ring."; + pp " apply Zlt_le_weak; auto."; + pp " repeat rewrite spec_head00; auto."; + pp " rewrite spec_double_size_digits."; + pp " rewrite Zpos_xO; auto with zarith."; + pp " rewrite spec_double_size; auto."; + pp " Qed."; + pr ""; + + pr " Theorem spec_double_size_head0_pos: "; + pr " forall x, 0 < [head0 (double_size x)]."; + pa " Admitted."; + pp " Proof."; + pp " intros x."; + pp " assert (F: 0 < Zpos (digits x))."; + pp " red; auto."; + pp " case (Zle_lt_or_eq _ _ (spec_pos (head0 (double_size x)))); auto; intros F0."; + pp " case (Zle_lt_or_eq _ _ (spec_pos (head0 x))); intros F1."; + pp " apply Zlt_le_trans with (2 := (spec_double_size_head0 x)); auto with zarith."; + pp " case (Zle_lt_or_eq _ _ (spec_pos x)); intros F3."; + pp " generalize F3; rewrite <- (spec_double_size x); intros F4."; + pp " absurd (2 ^ (Zpos (xO (digits x)) - 1) < 2 ^ (Zpos (digits x)))."; + pp " apply Zle_not_lt."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " split; auto with zarith."; + pp " rewrite Zpos_xO; auto with zarith."; + pp " case (spec_head0 x F3)."; + pp " rewrite <- F1; rewrite Zpower_0_r; rewrite Zmult_1_l; intros _ HH."; + pp " apply Zle_lt_trans with (2 := HH)."; + pp " case (spec_head0 _ F4)."; + pp " rewrite (spec_double_size x); rewrite (spec_double_size_digits x)."; + pp " rewrite <- F0; rewrite Zpower_0_r; rewrite Zmult_1_l; auto."; + pp " generalize F1; rewrite (spec_head00 _ (sym_equal F3)); auto with zarith."; + pp " Qed."; + pr ""; + + + (* Safe shiftl *) + + pr " Definition safe_shiftl_aux_body cont n x :="; + pr " match compare n (head0 x) with"; + pr " Gt => cont n (double_size x)"; + pr " | _ => shiftl n x"; + pr " end."; + pr ""; + + pr " Theorem spec_safe_shift_aux_body: forall n p x cont,"; + pr " 2^ Zpos p <= [head0 x] ->"; + pr " (forall x, 2 ^ (Zpos p + 1) <= [head0 x]->"; + pr " [cont n x] = [x] * 2 ^ [n]) ->"; + pr " [safe_shiftl_aux_body cont n x] = [x] * 2 ^ [n]."; + pa " Admitted."; + pp " Proof."; + pp " intros n p x cont H1 H2; unfold safe_shiftl_aux_body."; + pp " generalize (spec_compare n (head0 x)); case compare; intros H."; + pp " apply spec_shiftl; auto with zarith."; + pp " apply spec_shiftl; auto with zarith."; + pp " rewrite H2."; + pp " rewrite spec_double_size; auto."; + pp " rewrite Zplus_comm; rewrite Zpower_exp; auto with zarith."; + pp " apply Zle_trans with (2 := spec_double_size_head0 x)."; + pp " rewrite Zpower_1_r; apply Zmult_le_compat_l; auto with zarith."; + pp " Qed."; + pr ""; + + pr " Fixpoint safe_shiftl_aux p cont n x {struct p} :="; + pr " safe_shiftl_aux_body "; + pr " (fun n x => match p with"; + pr " | xH => cont n x"; + pr " | xO p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x"; + pr " | xI p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x"; + pr " end) n x."; + pr ""; + + pr " Theorem spec_safe_shift_aux: forall p q n x cont,"; + pr " 2 ^ (Zpos q) <= [head0 x] ->"; + pr " (forall x, 2 ^ (Zpos p + Zpos q) <= [head0 x] ->"; + pr " [cont n x] = [x] * 2 ^ [n]) -> "; + pr " [safe_shiftl_aux p cont n x] = [x] * 2 ^ [n]."; + pa " Admitted."; + pp " Proof."; + pp " intros p; elim p; unfold safe_shiftl_aux; fold safe_shiftl_aux; clear p."; + pp " intros p Hrec q n x cont H1 H2."; + pp " apply spec_safe_shift_aux_body with (q); auto."; + pp " intros x1 H3; apply Hrec with (q + 1)%spositive; auto." "%"; + pp " intros x2 H4; apply Hrec with (p + q + 1)%spositive; auto." "%"; + pp " rewrite <- Pplus_assoc."; + pp " rewrite Zpos_plus_distr; auto."; + pp " intros x3 H5; apply H2."; + pp " rewrite Zpos_xI."; + pp " replace (2 * Zpos p + 1 + Zpos q) with (Zpos p + Zpos (p + q + 1));"; + pp " auto."; + pp " repeat rewrite Zpos_plus_distr; ring."; + pp " intros p Hrec q n x cont H1 H2."; + pp " apply spec_safe_shift_aux_body with (q); auto."; + pp " intros x1 H3; apply Hrec with (q); auto."; + pp " apply Zle_trans with (2 := H3); auto with zarith."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " intros x2 H4; apply Hrec with (p + q)%spositive; auto." "%"; + pp " intros x3 H5; apply H2."; + pp " rewrite (Zpos_xO p)."; + pp " replace (2 * Zpos p + Zpos q) with (Zpos p + Zpos (p + q));"; + pp " auto."; + pp " repeat rewrite Zpos_plus_distr; ring."; + pp " intros q n x cont H1 H2."; + pp " apply spec_safe_shift_aux_body with (q); auto."; + pp " rewrite Zplus_comm; auto."; + pp " Qed."; + pr ""; + + + pr " Definition safe_shiftl n x :="; + pr " safe_shiftl_aux_body"; + pr " (safe_shiftl_aux_body"; + pr " (safe_shiftl_aux (digits n) shiftl)) n x."; + pr ""; + + pr " Theorem spec_safe_shift: forall n x,"; + pr " [safe_shiftl n x] = [x] * 2 ^ [n]."; + pa " Admitted."; + pp " Proof."; + pp " intros n x; unfold safe_shiftl, safe_shiftl_aux_body."; + pp " generalize (spec_compare n (head0 x)); case compare; intros H."; + pp " apply spec_shiftl; auto with zarith."; + pp " apply spec_shiftl; auto with zarith."; + pp " rewrite <- (spec_double_size x)."; + pp " generalize (spec_compare n (head0 (double_size x))); case compare; intros H1."; + pp " apply spec_shiftl; auto with zarith."; + pp " apply spec_shiftl; auto with zarith."; + pp " rewrite <- (spec_double_size (double_size x))."; + pp " apply spec_safe_shift_aux with 1%spositive." "%"; + pp " apply Zle_trans with (2 := spec_double_size_head0 (double_size x))."; + pp " replace (2 ^ 1) with (2 * 1)."; + pp " apply Zmult_le_compat_l; auto with zarith."; + pp " generalize (spec_double_size_head0_pos x); auto with zarith."; + pp " rewrite Zpower_1_r; ring."; + pp " intros x1 H2; apply spec_shiftl."; + pp " apply Zle_trans with (2 := H2)."; + pp " apply Zle_trans with (2 ^ Zpos (digits n)); auto with zarith."; + pp " case (spec_digits n); auto with zarith."; + pp " apply Zpower_le_monotone; auto with zarith."; + pp " Qed."; + pr ""; + + (* even *) + pr " Definition is_even x :="; + pr " match x with"; + for i = 0 to size do + pr " | %s%i wx => w%i_op.(znz_is_even) wx" c i i + done; + pr " | %sn n wx => (make_op n).(znz_is_even) wx" c; + pr " end."; + pr ""; + + + pr " Theorem spec_is_even: forall x,"; + pr " if is_even x then [x] mod 2 = 0 else [x] mod 2 = 1."; + pa " Admitted."; + pp " Proof."; + pp " intros x; case x; unfold is_even, to_Z; clear x."; + for i = 0 to size do + pp " intros x; exact (spec_is_even w%i_spec x)." i; + done; + pp " intros n x; exact (spec_is_even (wn_spec n) x)."; + pp " Qed."; + pr ""; + + pr " Theorem spec_0: [zero] = 0."; + pa " Admitted."; + pp " Proof."; + pp " exact (spec_0 w0_spec)."; + pp " Qed."; + pr ""; + + pr " Theorem spec_1: [one] = 1."; + pa " Admitted."; + pp " Proof."; + pp " exact (spec_1 w0_spec)."; + pp " Qed."; + pr ""; + + pr "End Make."; + pr ""; + diff --git a/theories/Numbers/Natural/BigN/Nbasic.v b/theories/Numbers/Natural/BigN/Nbasic.v new file mode 100644 index 00000000..ae2cfd30 --- /dev/null +++ b/theories/Numbers/Natural/BigN/Nbasic.v @@ -0,0 +1,514 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: Nbasic.v 10964 2008-05-22 11:08:13Z letouzey $ i*) + +Require Import ZArith. +Require Import BigNumPrelude. +Require Import Max. +Require Import DoubleType. +Require Import DoubleBase. +Require Import CyclicAxioms. +Require Import DoubleCyclic. + +(* To compute the necessary height *) + +Fixpoint plength (p: positive) : positive := + match p with + xH => xH + | xO p1 => Psucc (plength p1) + | xI p1 => Psucc (plength p1) + end. + +Theorem plength_correct: forall p, (Zpos p < 2 ^ Zpos (plength p))%Z. +assert (F: (forall p, 2 ^ (Zpos (Psucc p)) = 2 * 2 ^ Zpos p)%Z). +intros p; replace (Zpos (Psucc p)) with (1 + Zpos p)%Z. +rewrite Zpower_exp; auto with zarith. +rewrite Zpos_succ_morphism; unfold Zsucc; auto with zarith. +intros p; elim p; simpl plength; auto. +intros p1 Hp1; rewrite F; repeat rewrite Zpos_xI. +assert (tmp: (forall p, 2 * p = p + p)%Z); + try repeat rewrite tmp; auto with zarith. +intros p1 Hp1; rewrite F; rewrite (Zpos_xO p1). +assert (tmp: (forall p, 2 * p = p + p)%Z); + try repeat rewrite tmp; auto with zarith. +rewrite Zpower_1_r; auto with zarith. +Qed. + +Theorem plength_pred_correct: forall p, (Zpos p <= 2 ^ Zpos (plength (Ppred p)))%Z. +intros p; case (Psucc_pred p); intros H1. +subst; simpl plength. +rewrite Zpower_1_r; auto with zarith. +pattern p at 1; rewrite <- H1. +rewrite Zpos_succ_morphism; unfold Zsucc; auto with zarith. +generalize (plength_correct (Ppred p)); auto with zarith. +Qed. + +Definition Pdiv p q := + match Zdiv (Zpos p) (Zpos q) with + Zpos q1 => match (Zpos p) - (Zpos q) * (Zpos q1) with + Z0 => q1 + | _ => (Psucc q1) + end + | _ => xH + end. + +Theorem Pdiv_le: forall p q, + Zpos p <= Zpos q * Zpos (Pdiv p q). +intros p q. +unfold Pdiv. +assert (H1: Zpos q > 0); auto with zarith. +assert (H1b: Zpos p >= 0); auto with zarith. +generalize (Z_div_ge0 (Zpos p) (Zpos q) H1 H1b). +generalize (Z_div_mod_eq (Zpos p) (Zpos q) H1); case Zdiv. + intros HH _; rewrite HH; rewrite Zmult_0_r; rewrite Zmult_1_r; simpl. +case (Z_mod_lt (Zpos p) (Zpos q) H1); auto with zarith. +intros q1 H2. +replace (Zpos p - Zpos q * Zpos q1) with (Zpos p mod Zpos q). + 2: pattern (Zpos p) at 2; rewrite H2; auto with zarith. +generalize H2 (Z_mod_lt (Zpos p) (Zpos q) H1); clear H2; + case Zmod. + intros HH _; rewrite HH; auto with zarith. + intros r1 HH (_,HH1); rewrite HH; rewrite Zpos_succ_morphism. + unfold Zsucc; rewrite Zmult_plus_distr_r; auto with zarith. + intros r1 _ (HH,_); case HH; auto. +intros q1 HH; rewrite HH. +unfold Zge; simpl Zcompare; intros HH1; case HH1; auto. +Qed. + +Definition is_one p := match p with xH => true | _ => false end. + +Theorem is_one_one: forall p, is_one p = true -> p = xH. +intros p; case p; auto; intros p1 H1; discriminate H1. +Qed. + +Definition get_height digits p := + let r := Pdiv p digits in + if is_one r then xH else Psucc (plength (Ppred r)). + +Theorem get_height_correct: + forall digits N, + Zpos N <= Zpos digits * (2 ^ (Zpos (get_height digits N) -1)). +intros digits N. +unfold get_height. +assert (H1 := Pdiv_le N digits). +case_eq (is_one (Pdiv N digits)); intros H2. +rewrite (is_one_one _ H2) in H1. +rewrite Zmult_1_r in H1. +change (2^(1-1))%Z with 1; rewrite Zmult_1_r; auto. +clear H2. +apply Zle_trans with (1 := H1). +apply Zmult_le_compat_l; auto with zarith. +rewrite Zpos_succ_morphism; unfold Zsucc. +rewrite Zplus_comm; rewrite Zminus_plus. +apply plength_pred_correct. +Qed. + +Definition zn2z_word_comm : forall w n, zn2z (word w n) = word (zn2z w) n. + fix zn2z_word_comm 2. + intros w n; case n. + reflexivity. + intros n0;simpl. + case (zn2z_word_comm w n0). + reflexivity. +Defined. + +Fixpoint extend (n:nat) {struct n} : forall w:Type, zn2z w -> word w (S n) := + match n return forall w:Type, zn2z w -> word w (S n) with + | O => fun w x => x + | S m => + let aux := extend m in + fun w x => WW W0 (aux w x) + end. + +Section ExtendMax. + +Open Scope nat_scope. + +Fixpoint plusnS (n m: nat) {struct n} : (n + S m = S (n + m))%nat := + match n return (n + S m = S (n + m))%nat with + | 0 => refl_equal (S m) + | S n1 => + let v := S (S n1 + m) in + eq_ind_r (fun n => S n = v) (refl_equal v) (plusnS n1 m) + end. + +Fixpoint plusn0 n : n + 0 = n := + match n return (n + 0 = n) with + | 0 => refl_equal 0 + | S n1 => + let v := S n1 in + eq_ind_r (fun n : nat => S n = v) (refl_equal v) (plusn0 n1) + end. + + Fixpoint diff (m n: nat) {struct m}: nat * nat := + match m, n with + O, n => (O, n) + | m, O => (m, O) + | S m1, S n1 => diff m1 n1 + end. + +Fixpoint diff_l (m n : nat) {struct m} : fst (diff m n) + n = max m n := + match m return fst (diff m n) + n = max m n with + | 0 => + match n return (n = max 0 n) with + | 0 => refl_equal _ + | S n0 => refl_equal _ + end + | S m1 => + match n return (fst (diff (S m1) n) + n = max (S m1) n) + with + | 0 => plusn0 _ + | S n1 => + let v := fst (diff m1 n1) + n1 in + let v1 := fst (diff m1 n1) + S n1 in + eq_ind v (fun n => v1 = S n) + (eq_ind v1 (fun n => v1 = n) (refl_equal v1) (S v) (plusnS _ _)) + _ (diff_l _ _) + end + end. + +Fixpoint diff_r (m n: nat) {struct m}: snd (diff m n) + m = max m n := + match m return (snd (diff m n) + m = max m n) with + | 0 => + match n return (snd (diff 0 n) + 0 = max 0 n) with + | 0 => refl_equal _ + | S _ => plusn0 _ + end + | S m => + match n return (snd (diff (S m) n) + S m = max (S m) n) with + | 0 => refl_equal (snd (diff (S m) 0) + S m) + | S n1 => + let v := S (max m n1) in + eq_ind_r (fun n => n = v) + (eq_ind_r (fun n => S n = v) + (refl_equal v) (diff_r _ _)) (plusnS _ _) + end + end. + + Variable w: Type. + + Definition castm (m n: nat) (H: m = n) (x: word w (S m)): + (word w (S n)) := + match H in (_ = y) return (word w (S y)) with + | refl_equal => x + end. + +Variable m: nat. +Variable v: (word w (S m)). + +Fixpoint extend_tr (n : nat) {struct n}: (word w (S (n + m))) := + match n return (word w (S (n + m))) with + | O => v + | S n1 => WW W0 (extend_tr n1) + end. + +End ExtendMax. + +Implicit Arguments extend_tr[w m]. +Implicit Arguments castm[w m n]. + + + +Section Reduce. + + Variable w : Type. + Variable nT : Type. + Variable N0 : nT. + Variable eq0 : w -> bool. + Variable reduce_n : w -> nT. + Variable zn2z_to_Nt : zn2z w -> nT. + + Definition reduce_n1 (x:zn2z w) := + match x with + | W0 => N0 + | WW xh xl => + if eq0 xh then reduce_n xl + else zn2z_to_Nt x + end. + +End Reduce. + +Section ReduceRec. + + Variable w : Type. + Variable nT : Type. + Variable N0 : nT. + Variable reduce_1n : zn2z w -> nT. + Variable c : forall n, word w (S n) -> nT. + + Fixpoint reduce_n (n:nat) : word w (S n) -> nT := + match n return word w (S n) -> nT with + | O => reduce_1n + | S m => fun x => + match x with + | W0 => N0 + | WW xh xl => + match xh with + | W0 => @reduce_n m xl + | _ => @c (S m) x + end + end + end. + +End ReduceRec. + +Definition opp_compare cmp := + match cmp with + | Lt => Gt + | Eq => Eq + | Gt => Lt + end. + +Section CompareRec. + + Variable wm w : Type. + Variable w_0 : w. + Variable compare : w -> w -> comparison. + Variable compare0_m : wm -> comparison. + Variable compare_m : wm -> w -> comparison. + + Fixpoint compare0_mn (n:nat) : word wm n -> comparison := + match n return word wm n -> comparison with + | O => compare0_m + | S m => fun x => + match x with + | W0 => Eq + | WW xh xl => + match compare0_mn m xh with + | Eq => compare0_mn m xl + | r => Lt + end + end + end. + + Variable wm_base: positive. + Variable wm_to_Z: wm -> Z. + Variable w_to_Z: w -> Z. + Variable w_to_Z_0: w_to_Z w_0 = 0. + Variable spec_compare0_m: forall x, + match compare0_m x with + Eq => w_to_Z w_0 = wm_to_Z x + | Lt => w_to_Z w_0 < wm_to_Z x + | Gt => w_to_Z w_0 > wm_to_Z x + end. + Variable wm_to_Z_pos: forall x, 0 <= wm_to_Z x < base wm_base. + + Let double_to_Z := double_to_Z wm_base wm_to_Z. + Let double_wB := double_wB wm_base. + + Lemma base_xO: forall n, base (xO n) = (base n)^2. + Proof. + intros n1; unfold base. + rewrite (Zpos_xO n1); rewrite Zmult_comm; rewrite Zpower_mult; auto with zarith. + Qed. + + Let double_to_Z_pos: forall n x, 0 <= double_to_Z n x < double_wB n := + (spec_double_to_Z wm_base wm_to_Z wm_to_Z_pos). + + + Lemma spec_compare0_mn: forall n x, + match compare0_mn n x with + Eq => 0 = double_to_Z n x + | Lt => 0 < double_to_Z n x + | Gt => 0 > double_to_Z n x + end. + Proof. + intros n; elim n; clear n; auto. + intros x; generalize (spec_compare0_m x); rewrite w_to_Z_0; auto. + intros n Hrec x; case x; unfold compare0_mn; fold compare0_mn; auto. + intros xh xl. + generalize (Hrec xh); case compare0_mn; auto. + generalize (Hrec xl); case compare0_mn; auto. + simpl double_to_Z; intros H1 H2; rewrite H1; rewrite <- H2; auto. + simpl double_to_Z; intros H1 H2; rewrite <- H2; auto. + case (double_to_Z_pos n xl); auto with zarith. + intros H1; simpl double_to_Z. + set (u := DoubleBase.double_wB wm_base n). + case (double_to_Z_pos n xl); intros H2 H3. + assert (0 < u); auto with zarith. + unfold u, DoubleBase.double_wB, base; auto with zarith. + change 0 with (0 + 0); apply Zplus_lt_le_compat; auto with zarith. + apply Zmult_lt_0_compat; auto with zarith. + case (double_to_Z_pos n xh); auto with zarith. + Qed. + + Fixpoint compare_mn_1 (n:nat) : word wm n -> w -> comparison := + match n return word wm n -> w -> comparison with + | O => compare_m + | S m => fun x y => + match x with + | W0 => compare w_0 y + | WW xh xl => + match compare0_mn m xh with + | Eq => compare_mn_1 m xl y + | r => Gt + end + end + end. + + Variable spec_compare: forall x y, + match compare x y with + Eq => w_to_Z x = w_to_Z y + | Lt => w_to_Z x < w_to_Z y + | Gt => w_to_Z x > w_to_Z y + end. + Variable spec_compare_m: forall x y, + match compare_m x y with + Eq => wm_to_Z x = w_to_Z y + | Lt => wm_to_Z x < w_to_Z y + | Gt => wm_to_Z x > w_to_Z y + end. + Variable wm_base_lt: forall x, + 0 <= w_to_Z x < base (wm_base). + + Let double_wB_lt: forall n x, + 0 <= w_to_Z x < (double_wB n). + Proof. + intros n x; elim n; simpl; auto; clear n. + intros n (H0, H); split; auto. + apply Zlt_le_trans with (1:= H). + unfold double_wB, DoubleBase.double_wB; simpl. + rewrite base_xO. + set (u := base (double_digits wm_base n)). + assert (0 < u). + unfold u, base; auto with zarith. + replace (u^2) with (u * u); simpl; auto with zarith. + apply Zle_trans with (1 * u); auto with zarith. + unfold Zpower_pos; simpl; ring. + Qed. + + + Lemma spec_compare_mn_1: forall n x y, + match compare_mn_1 n x y with + Eq => double_to_Z n x = w_to_Z y + | Lt => double_to_Z n x < w_to_Z y + | Gt => double_to_Z n x > w_to_Z y + end. + Proof. + intros n; elim n; simpl; auto; clear n. + intros n Hrec x; case x; clear x; auto. + intros y; generalize (spec_compare w_0 y); rewrite w_to_Z_0; case compare; auto. + intros xh xl y; simpl; generalize (spec_compare0_mn n xh); case compare0_mn; intros H1b. + rewrite <- H1b; rewrite Zmult_0_l; rewrite Zplus_0_l; auto. + apply Hrec. + apply Zlt_gt. + case (double_wB_lt n y); intros _ H0. + apply Zlt_le_trans with (1:= H0). + fold double_wB. + case (double_to_Z_pos n xl); intros H1 H2. + apply Zle_trans with (double_to_Z n xh * double_wB n); auto with zarith. + apply Zle_trans with (1 * double_wB n); auto with zarith. + case (double_to_Z_pos n xh); auto with zarith. + Qed. + +End CompareRec. + + +Section AddS. + + Variable w wm : Type. + Variable incr : wm -> carry wm. + Variable addr : w -> wm -> carry wm. + Variable injr : w -> zn2z wm. + + Variable w_0 u: w. + Fixpoint injs (n:nat): word w (S n) := + match n return (word w (S n)) with + O => WW w_0 u + | S n1 => (WW W0 (injs n1)) + end. + + Definition adds x y := + match y with + W0 => C0 (injr x) + | WW hy ly => match addr x ly with + C0 z => C0 (WW hy z) + | C1 z => match incr hy with + C0 z1 => C0 (WW z1 z) + | C1 z1 => C1 (WW z1 z) + end + end + end. + +End AddS. + + + Lemma spec_opp: forall u x y, + match u with + | Eq => y = x + | Lt => y < x + | Gt => y > x + end -> + match opp_compare u with + | Eq => x = y + | Lt => x < y + | Gt => x > y + end. + Proof. + intros u x y; case u; simpl; auto with zarith. + Qed. + + Fixpoint length_pos x := + match x with xH => O | xO x1 => S (length_pos x1) | xI x1 => S (length_pos x1) end. + + Theorem length_pos_lt: forall x y, + (length_pos x < length_pos y)%nat -> Zpos x < Zpos y. + Proof. + intros x; elim x; clear x; [intros x1 Hrec | intros x1 Hrec | idtac]; + intros y; case y; clear y; intros y1 H || intros H; simpl length_pos; + try (rewrite (Zpos_xI x1) || rewrite (Zpos_xO x1)); + try (rewrite (Zpos_xI y1) || rewrite (Zpos_xO y1)); + try (inversion H; fail); + try (assert (Zpos x1 < Zpos y1); [apply Hrec; apply lt_S_n | idtac]; auto with zarith); + assert (0 < Zpos y1); auto with zarith; red; auto. + Qed. + + Theorem cancel_app: forall A B (f g: A -> B) x, f = g -> f x = g x. + Proof. + intros A B f g x H; rewrite H; auto. + Qed. + + + Section SimplOp. + + Variable w: Type. + + Theorem digits_zop: forall w (x: znz_op w), + znz_digits (mk_zn2z_op x) = xO (znz_digits x). + intros ww x; auto. + Qed. + + Theorem digits_kzop: forall w (x: znz_op w), + znz_digits (mk_zn2z_op_karatsuba x) = xO (znz_digits x). + intros ww x; auto. + Qed. + + Theorem make_zop: forall w (x: znz_op w), + znz_to_Z (mk_zn2z_op x) = + fun z => match z with + W0 => 0 + | WW xh xl => znz_to_Z x xh * base (znz_digits x) + + znz_to_Z x xl + end. + intros ww x; auto. + Qed. + + Theorem make_kzop: forall w (x: znz_op w), + znz_to_Z (mk_zn2z_op_karatsuba x) = + fun z => match z with + W0 => 0 + | WW xh xl => znz_to_Z x xh * base (znz_digits x) + + znz_to_Z x xl + end. + intros ww x; auto. + Qed. + + End SimplOp. diff --git a/theories/Numbers/Natural/Binary/NBinDefs.v b/theories/Numbers/Natural/Binary/NBinDefs.v new file mode 100644 index 00000000..fc2bd2df --- /dev/null +++ b/theories/Numbers/Natural/Binary/NBinDefs.v @@ -0,0 +1,267 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NBinDefs.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import BinPos. +Require Export BinNat. +Require Import NSub. + +Open Local Scope N_scope. + +(** Implementation of [NAxiomsSig] module type via [BinNat.N] *) + +Module NBinaryAxiomsMod <: NAxiomsSig. +Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig. +Module Export NZAxiomsMod <: NZAxiomsSig. + +Definition NZ := N. +Definition NZeq := @eq N. +Definition NZ0 := N0. +Definition NZsucc := Nsucc. +Definition NZpred := Npred. +Definition NZadd := Nplus. +Definition NZsub := Nminus. +Definition NZmul := Nmult. + +Theorem NZeq_equiv : equiv N NZeq. +Proof (eq_equiv N). + +Add Relation N NZeq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) +as NZeq_rel. + +Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd. +Proof. +congruence. +Qed. + +Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd. +Proof. +congruence. +Qed. + +Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd. +Proof. +congruence. +Qed. + +Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd. +Proof. +congruence. +Qed. + +Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd. +Proof. +congruence. +Qed. + +Theorem NZinduction : + forall A : NZ -> Prop, predicate_wd NZeq A -> + A N0 -> (forall n, A n <-> A (NZsucc n)) -> forall n : NZ, A n. +Proof. +intros A A_wd A0 AS. apply Nrect. assumption. intros; now apply -> AS. +Qed. + +Theorem NZpred_succ : forall n : NZ, NZpred (NZsucc n) = n. +Proof. +destruct n as [| p]; simpl. reflexivity. +case_eq (Psucc p); try (intros q H; rewrite <- H; now rewrite Ppred_succ). +intro H; false_hyp H Psucc_not_one. +Qed. + +Theorem NZadd_0_l : forall n : NZ, N0 + n = n. +Proof. +reflexivity. +Qed. + +Theorem NZadd_succ_l : forall n m : NZ, (NZsucc n) + m = NZsucc (n + m). +Proof. +destruct n; destruct m. +simpl in |- *; reflexivity. +unfold NZsucc, NZadd, Nsucc, Nplus. rewrite <- Pplus_one_succ_l; reflexivity. +simpl in |- *; reflexivity. +simpl in |- *; rewrite Pplus_succ_permute_l; reflexivity. +Qed. + +Theorem NZsub_0_r : forall n : NZ, n - N0 = n. +Proof. +now destruct n. +Qed. + +Theorem NZsub_succ_r : forall n m : NZ, n - (NZsucc m) = NZpred (n - m). +Proof. +destruct n as [| p]; destruct m as [| q]; try reflexivity. +now destruct p. +simpl. rewrite Pminus_mask_succ_r, Pminus_mask_carry_spec. +now destruct (Pminus_mask p q) as [| r |]; [| destruct r |]. +Qed. + +Theorem NZmul_0_l : forall n : NZ, N0 * n = N0. +Proof. +destruct n; reflexivity. +Qed. + +Theorem NZmul_succ_l : forall n m : NZ, (NZsucc n) * m = n * m + m. +Proof. +destruct n as [| n]; destruct m as [| m]; simpl; try reflexivity. +now rewrite Pmult_Sn_m, Pplus_comm. +Qed. + +End NZAxiomsMod. + +Definition NZlt := Nlt. +Definition NZle := Nle. +Definition NZmin := Nmin. +Definition NZmax := Nmax. + +Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd. +Proof. +unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2. +Qed. + +Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd. +Proof. +unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2. +Qed. + +Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd. +Proof. +congruence. +Qed. + +Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd. +Proof. +congruence. +Qed. + +Theorem NZlt_eq_cases : forall n m : N, n <= m <-> n < m \/ n = m. +Proof. +intros n m. unfold Nle, Nlt. rewrite <- Ncompare_eq_correct. +destruct (n ?= m); split; intro H1; (try discriminate); try (now left); try now right. +now elim H1. destruct H1; discriminate. +Qed. + +Theorem NZlt_irrefl : forall n : NZ, ~ n < n. +Proof. +intro n; unfold Nlt; now rewrite Ncompare_refl. +Qed. + +Theorem NZlt_succ_r : forall n m : NZ, n < (NZsucc m) <-> n <= m. +Proof. +intros n m; unfold Nlt, Nle; destruct n as [| p]; destruct m as [| q]; simpl; +split; intro H; try reflexivity; try discriminate. +destruct p; simpl; intros; discriminate. elimtype False; now apply H. +apply -> Pcompare_p_Sq in H. destruct H as [H | H]. +now rewrite H. now rewrite H, Pcompare_refl. +apply <- Pcompare_p_Sq. case_eq ((p ?= q)%positive Eq); intro H1. +right; now apply Pcompare_Eq_eq. now left. elimtype False; now apply H. +Qed. + +Theorem NZmin_l : forall n m : N, n <= m -> NZmin n m = n. +Proof. +unfold NZmin, Nmin, Nle; intros n m H. +destruct (n ?= m); try reflexivity. now elim H. +Qed. + +Theorem NZmin_r : forall n m : N, m <= n -> NZmin n m = m. +Proof. +unfold NZmin, Nmin, Nle; intros n m H. +case_eq (n ?= m); intro H1; try reflexivity. +now apply -> Ncompare_eq_correct. +rewrite <- Ncompare_antisym, H1 in H; elim H; auto. +Qed. + +Theorem NZmax_l : forall n m : N, m <= n -> NZmax n m = n. +Proof. +unfold NZmax, Nmax, Nle; intros n m H. +case_eq (n ?= m); intro H1; try reflexivity. +symmetry; now apply -> Ncompare_eq_correct. +rewrite <- Ncompare_antisym, H1 in H; elim H; auto. +Qed. + +Theorem NZmax_r : forall n m : N, n <= m -> NZmax n m = m. +Proof. +unfold NZmax, Nmax, Nle; intros n m H. +destruct (n ?= m); try reflexivity. now elim H. +Qed. + +End NZOrdAxiomsMod. + +Definition recursion (A : Type) (a : A) (f : N -> A -> A) (n : N) := + Nrect (fun _ => A) a f n. +Implicit Arguments recursion [A]. + +Theorem pred_0 : Npred N0 = N0. +Proof. +reflexivity. +Qed. + +Theorem recursion_wd : +forall (A : Type) (Aeq : relation A), + forall a a' : A, Aeq a a' -> + forall f f' : N -> A -> A, fun2_eq NZeq Aeq Aeq f f' -> + forall x x' : N, x = x' -> + Aeq (recursion a f x) (recursion a' f' x'). +Proof. +unfold fun2_wd, NZeq, fun2_eq. +intros A Aeq a a' Eaa' f f' Eff'. +intro x; pattern x; apply Nrect. +intros x' H; now rewrite <- H. +clear x. +intros x IH x' H; rewrite <- H. +unfold recursion in *. do 2 rewrite Nrect_step. +now apply Eff'; [| apply IH]. +Qed. + +Theorem recursion_0 : + forall (A : Type) (a : A) (f : N -> A -> A), recursion a f N0 = a. +Proof. +intros A a f; unfold recursion; now rewrite Nrect_base. +Qed. + +Theorem recursion_succ : + forall (A : Type) (Aeq : relation A) (a : A) (f : N -> A -> A), + Aeq a a -> fun2_wd NZeq Aeq Aeq f -> + forall n : N, Aeq (recursion a f (Nsucc n)) (f n (recursion a f n)). +Proof. +unfold NZeq, recursion, fun2_wd; intros A Aeq a f EAaa f_wd n; pattern n; apply Nrect. +rewrite Nrect_step; rewrite Nrect_base; now apply f_wd. +clear n; intro n; do 2 rewrite Nrect_step; intro IH. apply f_wd; [reflexivity|]. +now rewrite Nrect_step. +Qed. + +End NBinaryAxiomsMod. + +Module Export NBinarySubPropMod := NSubPropFunct NBinaryAxiomsMod. + +(* Some fun comparing the efficiency of the generic log defined +by strong (course-of-value) recursion and the log defined by recursion +on notation *) +(* Time Eval compute in (log 100000). *) (* 98 sec *) + +(* +Fixpoint binposlog (p : positive) : N := +match p with +| xH => 0 +| xO p' => Nsucc (binposlog p') +| xI p' => Nsucc (binposlog p') +end. + +Definition binlog (n : N) : N := +match n with +| 0 => 0 +| Npos p => binposlog p +end. +*) +(* Eval compute in (binlog 1000000000000000000). *) (* Works very fast *) + diff --git a/theories/Numbers/Natural/Binary/NBinary.v b/theories/Numbers/Natural/Binary/NBinary.v new file mode 100644 index 00000000..2c99128d --- /dev/null +++ b/theories/Numbers/Natural/Binary/NBinary.v @@ -0,0 +1,15 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NBinary.v 10934 2008-05-15 21:58:20Z letouzey $ i*) + +Require Export NBinDefs. +Require Export NArithRing. + diff --git a/theories/Numbers/Natural/Peano/NPeano.v b/theories/Numbers/Natural/Peano/NPeano.v new file mode 100644 index 00000000..1c83da45 --- /dev/null +++ b/theories/Numbers/Natural/Peano/NPeano.v @@ -0,0 +1,220 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NPeano.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import Arith. +Require Import Min. +Require Import Max. +Require Import NSub. + +Module NPeanoAxiomsMod <: NAxiomsSig. +Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig. +Module Export NZAxiomsMod <: NZAxiomsSig. + +Definition NZ := nat. +Definition NZeq := (@eq nat). +Definition NZ0 := 0. +Definition NZsucc := S. +Definition NZpred := pred. +Definition NZadd := plus. +Definition NZsub := minus. +Definition NZmul := mult. + +Theorem NZeq_equiv : equiv nat NZeq. +Proof (eq_equiv nat). + +Add Relation nat NZeq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) +as NZeq_rel. + +(* If we say "Add Relation nat (@eq nat)" instead of "Add Relation nat NZeq" +then the theorem generated for succ_wd below is forall x, succ x = succ x, +which does not match the axioms in NAxiomsSig *) + +Add Morphism NZsucc with signature NZeq ==> NZeq as NZsucc_wd. +Proof. +congruence. +Qed. + +Add Morphism NZpred with signature NZeq ==> NZeq as NZpred_wd. +Proof. +congruence. +Qed. + +Add Morphism NZadd with signature NZeq ==> NZeq ==> NZeq as NZadd_wd. +Proof. +congruence. +Qed. + +Add Morphism NZsub with signature NZeq ==> NZeq ==> NZeq as NZsub_wd. +Proof. +congruence. +Qed. + +Add Morphism NZmul with signature NZeq ==> NZeq ==> NZeq as NZmul_wd. +Proof. +congruence. +Qed. + +Theorem NZinduction : + forall A : nat -> Prop, predicate_wd (@eq nat) A -> + A 0 -> (forall n : nat, A n <-> A (S n)) -> forall n : nat, A n. +Proof. +intros A A_wd A0 AS. apply nat_ind. assumption. intros; now apply -> AS. +Qed. + +Theorem NZpred_succ : forall n : nat, pred (S n) = n. +Proof. +reflexivity. +Qed. + +Theorem NZadd_0_l : forall n : nat, 0 + n = n. +Proof. +reflexivity. +Qed. + +Theorem NZadd_succ_l : forall n m : nat, (S n) + m = S (n + m). +Proof. +reflexivity. +Qed. + +Theorem NZsub_0_r : forall n : nat, n - 0 = n. +Proof. +intro n; now destruct n. +Qed. + +Theorem NZsub_succ_r : forall n m : nat, n - (S m) = pred (n - m). +Proof. +intros n m; induction n m using nat_double_ind; simpl; auto. apply NZsub_0_r. +Qed. + +Theorem NZmul_0_l : forall n : nat, 0 * n = 0. +Proof. +reflexivity. +Qed. + +Theorem NZmul_succ_l : forall n m : nat, S n * m = n * m + m. +Proof. +intros n m; now rewrite plus_comm. +Qed. + +End NZAxiomsMod. + +Definition NZlt := lt. +Definition NZle := le. +Definition NZmin := min. +Definition NZmax := max. + +Add Morphism NZlt with signature NZeq ==> NZeq ==> iff as NZlt_wd. +Proof. +unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2. +Qed. + +Add Morphism NZle with signature NZeq ==> NZeq ==> iff as NZle_wd. +Proof. +unfold NZeq; intros x1 x2 H1 y1 y2 H2; rewrite H1; now rewrite H2. +Qed. + +Add Morphism NZmin with signature NZeq ==> NZeq ==> NZeq as NZmin_wd. +Proof. +congruence. +Qed. + +Add Morphism NZmax with signature NZeq ==> NZeq ==> NZeq as NZmax_wd. +Proof. +congruence. +Qed. + +Theorem NZlt_eq_cases : forall n m : nat, n <= m <-> n < m \/ n = m. +Proof. +intros n m; split. +apply le_lt_or_eq. +intro H; destruct H as [H | H]. +now apply lt_le_weak. rewrite H; apply le_refl. +Qed. + +Theorem NZlt_irrefl : forall n : nat, ~ (n < n). +Proof. +exact lt_irrefl. +Qed. + +Theorem NZlt_succ_r : forall n m : nat, n < S m <-> n <= m. +Proof. +intros n m; split; [apply lt_n_Sm_le | apply le_lt_n_Sm]. +Qed. + +Theorem NZmin_l : forall n m : nat, n <= m -> NZmin n m = n. +Proof. +exact min_l. +Qed. + +Theorem NZmin_r : forall n m : nat, m <= n -> NZmin n m = m. +Proof. +exact min_r. +Qed. + +Theorem NZmax_l : forall n m : nat, m <= n -> NZmax n m = n. +Proof. +exact max_l. +Qed. + +Theorem NZmax_r : forall n m : nat, n <= m -> NZmax n m = m. +Proof. +exact max_r. +Qed. + +End NZOrdAxiomsMod. + +Definition recursion : forall A : Type, A -> (nat -> A -> A) -> nat -> A := + fun A : Type => nat_rect (fun _ => A). +Implicit Arguments recursion [A]. + +Theorem succ_neq_0 : forall n : nat, S n <> 0. +Proof. +intros; discriminate. +Qed. + +Theorem pred_0 : pred 0 = 0. +Proof. +reflexivity. +Qed. + +Theorem recursion_wd : forall (A : Type) (Aeq : relation A), + forall a a' : A, Aeq a a' -> + forall f f' : nat -> A -> A, fun2_eq (@eq nat) Aeq Aeq f f' -> + forall n n' : nat, n = n' -> + Aeq (recursion a f n) (recursion a' f' n'). +Proof. +unfold fun2_eq; induction n; intros n' Enn'; rewrite <- Enn' in *; simpl; auto. +Qed. + +Theorem recursion_0 : + forall (A : Type) (a : A) (f : nat -> A -> A), recursion a f 0 = a. +Proof. +reflexivity. +Qed. + +Theorem recursion_succ : + forall (A : Type) (Aeq : relation A) (a : A) (f : nat -> A -> A), + Aeq a a -> fun2_wd (@eq nat) Aeq Aeq f -> + forall n : nat, Aeq (recursion a f (S n)) (f n (recursion a f n)). +Proof. +induction n; simpl; auto. +Qed. + +End NPeanoAxiomsMod. + +(* Now we apply the largest property functor *) + +Module Export NPeanoSubPropMod := NSubPropFunct NPeanoAxiomsMod. + diff --git a/theories/Numbers/Natural/SpecViaZ/NSig.v b/theories/Numbers/Natural/SpecViaZ/NSig.v new file mode 100644 index 00000000..0275d1e1 --- /dev/null +++ b/theories/Numbers/Natural/SpecViaZ/NSig.v @@ -0,0 +1,115 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NSig.v 11027 2008-06-01 13:28:59Z letouzey $ i*) + +Require Import ZArith Znumtheory. + +Open Scope Z_scope. + +(** * NSig *) + +(** Interface of a rich structure about natural numbers. + Specifications are written via translation to Z. +*) + +Module Type NType. + + Parameter t : Type. + + Parameter to_Z : t -> Z. + Notation "[ x ]" := (to_Z x). + Parameter spec_pos: forall x, 0 <= [x]. + + Parameter of_N : N -> t. + Parameter spec_of_N: forall x, to_Z (of_N x) = Z_of_N x. + Definition to_N n := Zabs_N (to_Z n). + + Definition eq n m := ([n] = [m]). + + Parameter zero : t. + Parameter one : t. + + Parameter spec_0: [zero] = 0. + Parameter spec_1: [one] = 1. + + Parameter compare : t -> t -> comparison. + + Parameter spec_compare: forall x y, + match compare x y with + | Eq => [x] = [y] + | Lt => [x] < [y] + | Gt => [x] > [y] + end. + + Definition lt n m := compare n m = Lt. + Definition le n m := compare n m <> Gt. + Definition min n m := match compare n m with Gt => m | _ => n end. + Definition max n m := match compare n m with Lt => m | _ => n end. + + Parameter eq_bool : t -> t -> bool. + + Parameter spec_eq_bool: forall x y, + if eq_bool x y then [x] = [y] else [x] <> [y]. + + Parameter succ : t -> t. + + Parameter spec_succ: forall n, [succ n] = [n] + 1. + + Parameter add : t -> t -> t. + + Parameter spec_add: forall x y, [add x y] = [x] + [y]. + + Parameter pred : t -> t. + + Parameter spec_pred: forall x, 0 < [x] -> [pred x] = [x] - 1. + Parameter spec_pred0: forall x, [x] = 0 -> [pred x] = 0. + + Parameter sub : t -> t -> t. + + Parameter spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y]. + Parameter spec_sub0: forall x y, [x] < [y]-> [sub x y] = 0. + + Parameter mul : t -> t -> t. + + Parameter spec_mul: forall x y, [mul x y] = [x] * [y]. + + Parameter square : t -> t. + + Parameter spec_square: forall x, [square x] = [x] * [x]. + + Parameter power_pos : t -> positive -> t. + + Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n. + + Parameter sqrt : t -> t. + + Parameter spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2. + + Parameter div_eucl : t -> t -> t * t. + + Parameter spec_div_eucl: forall x y, + 0 < [y] -> + let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y]. + + Parameter div : t -> t -> t. + + Parameter spec_div: forall x y, 0 < [y] -> [div x y] = [x] / [y]. + + Parameter modulo : t -> t -> t. + + Parameter spec_modulo: + forall x y, 0 < [y] -> [modulo x y] = [x] mod [y]. + + Parameter gcd : t -> t -> t. + + Parameter spec_gcd: forall a b, [gcd a b] = Zgcd (to_Z a) (to_Z b). + +End NType. diff --git a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v new file mode 100644 index 00000000..fe068437 --- /dev/null +++ b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v @@ -0,0 +1,356 @@ +(************************************************************************) +(* 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: NSigNAxioms.v 11040 2008-06-03 00:04:16Z letouzey $ i*) + +Require Import ZArith. +Require Import Nnat. +Require Import NAxioms. +Require Import NSig. + +(** * The interface [NSig.NType] implies the interface [NAxiomsSig] *) + +Module NSig_NAxioms (N:NType) <: NAxiomsSig. + +Delimit Scope IntScope with Int. +Bind Scope IntScope with N.t. +Open Local Scope IntScope. +Notation "[ x ]" := (N.to_Z x) : IntScope. +Infix "==" := N.eq (at level 70) : IntScope. +Notation "0" := N.zero : IntScope. +Infix "+" := N.add : IntScope. +Infix "-" := N.sub : IntScope. +Infix "*" := N.mul : IntScope. + +Module Export NZOrdAxiomsMod <: NZOrdAxiomsSig. +Module Export NZAxiomsMod <: NZAxiomsSig. + +Definition NZ := N.t. +Definition NZeq := N.eq. +Definition NZ0 := N.zero. +Definition NZsucc := N.succ. +Definition NZpred := N.pred. +Definition NZadd := N.add. +Definition NZsub := N.sub. +Definition NZmul := N.mul. + +Theorem NZeq_equiv : equiv N.t N.eq. +Proof. +repeat split; repeat red; intros; auto; congruence. +Qed. + +Add Relation N.t N.eq + reflexivity proved by (proj1 NZeq_equiv) + symmetry proved by (proj2 (proj2 NZeq_equiv)) + transitivity proved by (proj1 (proj2 NZeq_equiv)) + as NZeq_rel. + +Add Morphism NZsucc with signature N.eq ==> N.eq as NZsucc_wd. +Proof. +unfold N.eq; intros; rewrite 2 N.spec_succ; f_equal; auto. +Qed. + +Add Morphism NZpred with signature N.eq ==> N.eq as NZpred_wd. +Proof. +unfold N.eq; intros. +generalize (N.spec_pos y) (N.spec_pos x) (N.spec_eq_bool x 0). +destruct N.eq_bool; rewrite N.spec_0; intros. +rewrite 2 N.spec_pred0; congruence. +rewrite 2 N.spec_pred; f_equal; auto; try omega. +Qed. + +Add Morphism NZadd with signature N.eq ==> N.eq ==> N.eq as NZadd_wd. +Proof. +unfold N.eq; intros; rewrite 2 N.spec_add; f_equal; auto. +Qed. + +Add Morphism NZsub with signature N.eq ==> N.eq ==> N.eq as NZsub_wd. +Proof. +unfold N.eq; intros x x' Hx y y' Hy. +destruct (Z_lt_le_dec [x] [y]). +rewrite 2 N.spec_sub0; f_equal; congruence. +rewrite 2 N.spec_sub; f_equal; congruence. +Qed. + +Add Morphism NZmul with signature N.eq ==> N.eq ==> N.eq as NZmul_wd. +Proof. +unfold N.eq; intros; rewrite 2 N.spec_mul; f_equal; auto. +Qed. + +Theorem NZpred_succ : forall n, N.pred (N.succ n) == n. +Proof. +unfold N.eq; intros. +rewrite N.spec_pred; rewrite N.spec_succ. +omega. +generalize (N.spec_pos n); omega. +Qed. + +Definition N_of_Z z := N.of_N (Zabs_N z). + +Section Induction. + +Variable A : N.t -> Prop. +Hypothesis A_wd : predicate_wd N.eq A. +Hypothesis A0 : A 0. +Hypothesis AS : forall n, A n <-> A (N.succ n). + +Add Morphism A with signature N.eq ==> iff as A_morph. +Proof. apply A_wd. Qed. + +Let B (z : Z) := A (N_of_Z z). + +Lemma B0 : B 0. +Proof. +unfold B, N_of_Z; simpl. +rewrite <- (A_wd 0); auto. +red; rewrite N.spec_0, N.spec_of_N; auto. +Qed. + +Lemma BS : forall z : Z, (0 <= z)%Z -> B z -> B (z + 1). +Proof. +intros z H1 H2. +unfold B in *. apply -> AS in H2. +setoid_replace (N_of_Z (z + 1)) with (N.succ (N_of_Z z)); auto. +unfold N.eq. rewrite N.spec_succ. +unfold N_of_Z. +rewrite 2 N.spec_of_N, 2 Z_of_N_abs, 2 Zabs_eq; auto with zarith. +Qed. + +Lemma B_holds : forall z : Z, (0 <= z)%Z -> B z. +Proof. +exact (natlike_ind B B0 BS). +Qed. + +Theorem NZinduction : forall n, A n. +Proof. +intro n. setoid_replace n with (N_of_Z (N.to_Z n)). +apply B_holds. apply N.spec_pos. +red; unfold N_of_Z. +rewrite N.spec_of_N, Z_of_N_abs, Zabs_eq; auto. +apply N.spec_pos. +Qed. + +End Induction. + +Theorem NZadd_0_l : forall n, 0 + n == n. +Proof. +intros; red; rewrite N.spec_add, N.spec_0; auto with zarith. +Qed. + +Theorem NZadd_succ_l : forall n m, (N.succ n) + m == N.succ (n + m). +Proof. +intros; red; rewrite N.spec_add, 2 N.spec_succ, N.spec_add; auto with zarith. +Qed. + +Theorem NZsub_0_r : forall n, n - 0 == n. +Proof. +intros; red; rewrite N.spec_sub; rewrite N.spec_0; auto with zarith. +apply N.spec_pos. +Qed. + +Theorem NZsub_succ_r : forall n m, n - (N.succ m) == N.pred (n - m). +Proof. +intros; red. +destruct (Z_lt_le_dec [n] [N.succ m]) as [H|H]. +rewrite N.spec_sub0; auto. +rewrite N.spec_succ in H. +rewrite N.spec_pred0; auto. +destruct (Z_eq_dec [n] [m]). +rewrite N.spec_sub; auto with zarith. +rewrite N.spec_sub0; auto with zarith. + +rewrite N.spec_sub, N.spec_succ in *; auto. +rewrite N.spec_pred, N.spec_sub; auto with zarith. +rewrite N.spec_sub; auto with zarith. +Qed. + +Theorem NZmul_0_l : forall n, 0 * n == 0. +Proof. +intros; red. +rewrite N.spec_mul, N.spec_0; auto with zarith. +Qed. + +Theorem NZmul_succ_l : forall n m, (N.succ n) * m == n * m + m. +Proof. +intros; red. +rewrite N.spec_add, 2 N.spec_mul, N.spec_succ; ring. +Qed. + +End NZAxiomsMod. + +Definition NZlt := N.lt. +Definition NZle := N.le. +Definition NZmin := N.min. +Definition NZmax := N.max. + +Infix "<=" := N.le : IntScope. +Infix "<" := N.lt : IntScope. + +Lemma spec_compare_alt : forall x y, N.compare x y = ([x] ?= [y])%Z. +Proof. + intros; generalize (N.spec_compare x y). + destruct (N.compare x y); auto. + intros H; rewrite H; symmetry; apply Zcompare_refl. +Qed. + +Lemma spec_lt : forall x y, (x<y) <-> ([x]<[y])%Z. +Proof. + intros; unfold N.lt, Zlt; rewrite spec_compare_alt; intuition. +Qed. + +Lemma spec_le : forall x y, (x<=y) <-> ([x]<=[y])%Z. +Proof. + intros; unfold N.le, Zle; rewrite spec_compare_alt; intuition. +Qed. + +Lemma spec_min : forall x y, [N.min x y] = Zmin [x] [y]. +Proof. + intros; unfold N.min, Zmin. + rewrite spec_compare_alt; destruct Zcompare; auto. +Qed. + +Lemma spec_max : forall x y, [N.max x y] = Zmax [x] [y]. +Proof. + intros; unfold N.max, Zmax. + rewrite spec_compare_alt; destruct Zcompare; auto. +Qed. + +Add Morphism N.compare with signature N.eq ==> N.eq ==> (@eq comparison) as compare_wd. +Proof. +intros x x' Hx y y' Hy. +rewrite 2 spec_compare_alt; rewrite Hx, Hy; intuition. +Qed. + +Add Morphism N.lt with signature N.eq ==> N.eq ==> iff as NZlt_wd. +Proof. +intros x x' Hx y y' Hy; unfold N.lt; rewrite Hx, Hy; intuition. +Qed. + +Add Morphism N.le with signature N.eq ==> N.eq ==> iff as NZle_wd. +Proof. +intros x x' Hx y y' Hy; unfold N.le; rewrite Hx, Hy; intuition. +Qed. + +Add Morphism N.min with signature N.eq ==> N.eq ==> N.eq as NZmin_wd. +Proof. +intros; red; rewrite 2 spec_min; congruence. +Qed. + +Add Morphism N.max with signature N.eq ==> N.eq ==> N.eq as NZmax_wd. +Proof. +intros; red; rewrite 2 spec_max; congruence. +Qed. + +Theorem NZlt_eq_cases : forall n m, n <= m <-> n < m \/ n == m. +Proof. +intros. +unfold N.eq; rewrite spec_lt, spec_le; omega. +Qed. + +Theorem NZlt_irrefl : forall n, ~ n < n. +Proof. +intros; rewrite spec_lt; auto with zarith. +Qed. + +Theorem NZlt_succ_r : forall n m, n < (N.succ m) <-> n <= m. +Proof. +intros; rewrite spec_lt, spec_le, N.spec_succ; omega. +Qed. + +Theorem NZmin_l : forall n m, n <= m -> N.min n m == n. +Proof. +intros n m; unfold N.eq; rewrite spec_le, spec_min. +generalize (Zmin_spec [n] [m]); omega. +Qed. + +Theorem NZmin_r : forall n m, m <= n -> N.min n m == m. +Proof. +intros n m; unfold N.eq; rewrite spec_le, spec_min. +generalize (Zmin_spec [n] [m]); omega. +Qed. + +Theorem NZmax_l : forall n m, m <= n -> N.max n m == n. +Proof. +intros n m; unfold N.eq; rewrite spec_le, spec_max. +generalize (Zmax_spec [n] [m]); omega. +Qed. + +Theorem NZmax_r : forall n m, n <= m -> N.max n m == m. +Proof. +intros n m; unfold N.eq; rewrite spec_le, spec_max. +generalize (Zmax_spec [n] [m]); omega. +Qed. + +End NZOrdAxiomsMod. + +Theorem pred_0 : N.pred 0 == 0. +Proof. +red; rewrite N.spec_pred0; rewrite N.spec_0; auto. +Qed. + +Definition recursion (A : Type) (a : A) (f : N.t -> A -> A) (n : N.t) := + Nrect (fun _ => A) a (fun n a => f (N.of_N n) a) (N.to_N n). +Implicit Arguments recursion [A]. + +Theorem recursion_wd : +forall (A : Type) (Aeq : relation A), + forall a a' : A, Aeq a a' -> + forall f f' : N.t -> A -> A, fun2_eq N.eq Aeq Aeq f f' -> + forall x x' : N.t, x == x' -> + Aeq (recursion a f x) (recursion a' f' x'). +Proof. +unfold fun2_wd, N.eq, fun2_eq. +intros A Aeq a a' Eaa' f f' Eff' x x' Exx'. +unfold recursion. +unfold N.to_N. +rewrite <- Exx'; clear x' Exx'. +replace (Zabs_N [x]) with (N_of_nat (Zabs_nat [x])). +induction (Zabs_nat [x]). +simpl; auto. +rewrite N_of_S, 2 Nrect_step; auto. +destruct [x]; simpl; auto. +change (nat_of_P p) with (nat_of_N (Npos p)); apply N_of_nat_of_N. +change (nat_of_P p) with (nat_of_N (Npos p)); apply N_of_nat_of_N. +Qed. + +Theorem recursion_0 : + forall (A : Type) (a : A) (f : N.t -> A -> A), recursion a f 0 = a. +Proof. +intros A a f; unfold recursion, N.to_N; rewrite N.spec_0; simpl; auto. +Qed. + +Theorem recursion_succ : + forall (A : Type) (Aeq : relation A) (a : A) (f : N.t -> A -> A), + Aeq a a -> fun2_wd N.eq Aeq Aeq f -> + forall n, Aeq (recursion a f (N.succ n)) (f n (recursion a f n)). +Proof. +unfold N.eq, recursion, fun2_wd; intros A Aeq a f EAaa f_wd n. +replace (N.to_N (N.succ n)) with (Nsucc (N.to_N n)). +rewrite Nrect_step. +apply f_wd; auto. +unfold N.to_N. +rewrite N.spec_of_N, Z_of_N_abs, Zabs_eq; auto. + apply N.spec_pos. + +fold (recursion a f n). +apply recursion_wd; auto. +red; auto. +red; auto. +unfold N.to_N. + +rewrite N.spec_succ. +change ([n]+1)%Z with (Zsucc [n]). +apply Z_of_N_eq_rev. +rewrite Z_of_N_succ. +rewrite 2 Z_of_N_abs. +rewrite 2 Zabs_eq; auto. +generalize (N.spec_pos n); auto with zarith. +apply N.spec_pos; auto. +Qed. + +End NSig_NAxioms. diff --git a/theories/Numbers/NumPrelude.v b/theories/Numbers/NumPrelude.v new file mode 100644 index 00000000..fdccf214 --- /dev/null +++ b/theories/Numbers/NumPrelude.v @@ -0,0 +1,267 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Evgeny Makarov, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: NumPrelude.v 10943 2008-05-19 08:45:13Z letouzey $ i*) + +Require Export Setoid. + +Set Implicit Arguments. +(* +Contents: +- Coercion from bool to Prop +- Extension of the tactics stepl and stepr +- Extentional properties of predicates, relations and functions + (well-definedness and equality) +- Relations on cartesian product +- Miscellaneous +*) + +(** Coercion from bool to Prop *) + +(*Definition eq_bool := (@eq bool).*) + +(*Inductive eq_true : bool -> Prop := is_eq_true : eq_true true.*) +(* This has been added to theories/Datatypes.v *) +(*Coercion eq_true : bool >-> Sortclass.*) + +(*Theorem eq_true_unfold_pos : forall b : bool, b <-> b = true. +Proof. +intro b; split; intro H. now inversion H. now rewrite H. +Qed. + +Theorem eq_true_unfold_neg : forall b : bool, ~ b <-> b = false. +Proof. +intros b; destruct b; simpl; rewrite eq_true_unfold_pos. +split; intro H; [elim (H (refl_equal true)) | discriminate H]. +split; intro H; [reflexivity | discriminate]. +Qed. + +Theorem eq_true_or : forall b1 b2 : bool, b1 || b2 <-> b1 \/ b2. +Proof. +destruct b1; destruct b2; simpl; tauto. +Qed. + +Theorem eq_true_and : forall b1 b2 : bool, b1 && b2 <-> b1 /\ b2. +Proof. +destruct b1; destruct b2; simpl; tauto. +Qed. + +Theorem eq_true_neg : forall b : bool, negb b <-> ~ b. +Proof. +destruct b; simpl; rewrite eq_true_unfold_pos; rewrite eq_true_unfold_neg; +split; now intro. +Qed. + +Theorem eq_true_iff : forall b1 b2 : bool, b1 = b2 <-> (b1 <-> b2). +Proof. +intros b1 b2; split; intro H. +now rewrite H. +destruct b1; destruct b2; simpl; try reflexivity. +apply -> eq_true_unfold_neg. rewrite H. now intro. +symmetry; apply -> eq_true_unfold_neg. rewrite <- H; now intro. +Qed.*) + +(** Extension of the tactics stepl and stepr to make them +applicable to hypotheses *) + +Tactic Notation "stepl" constr(t1') "in" hyp(H) := +match (type of H) with +| ?R ?t1 ?t2 => + let H1 := fresh in + cut (R t1' t2); [clear H; intro H | stepl t1; [assumption |]] +| _ => fail 1 ": the hypothesis" H "does not have the form (R t1 t2)" +end. + +Tactic Notation "stepl" constr(t1') "in" hyp(H) "by" tactic(r) := stepl t1' in H; [| r]. + +Tactic Notation "stepr" constr(t2') "in" hyp(H) := +match (type of H) with +| ?R ?t1 ?t2 => + let H1 := fresh in + cut (R t1 t2'); [clear H; intro H | stepr t2; [assumption |]] +| _ => fail 1 ": the hypothesis" H "does not have the form (R t1 t2)" +end. + +Tactic Notation "stepr" constr(t2') "in" hyp(H) "by" tactic(r) := stepr t2' in H; [| r]. + +(** Extentional properties of predicates, relations and functions *) + +Definition predicate (A : Type) := A -> Prop. + +Section ExtensionalProperties. + +Variables A B C : Type. +Variable Aeq : relation A. +Variable Beq : relation B. +Variable Ceq : relation C. + +(* "wd" stands for "well-defined" *) + +Definition fun_wd (f : A -> B) := forall x y : A, Aeq x y -> Beq (f x) (f y). + +Definition fun2_wd (f : A -> B -> C) := + forall x x' : A, Aeq x x' -> forall y y' : B, Beq y y' -> Ceq (f x y) (f x' y'). + +Definition fun_eq : relation (A -> B) := + fun f f' => forall x x' : A, Aeq x x' -> Beq (f x) (f' x'). + +(* Note that reflexivity of fun_eq means that every function +is well-defined w.r.t. Aeq and Beq, i.e., +forall x x' : A, Aeq x x' -> Beq (f x) (f x') *) + +Definition fun2_eq (f f' : A -> B -> C) := + forall x x' : A, Aeq x x' -> forall y y' : B, Beq y y' -> Ceq (f x y) (f' x' y'). + +End ExtensionalProperties. + +(* The following definitions instantiate Beq or Ceq to iff; therefore, they +have to be outside the ExtensionalProperties section *) + +Definition predicate_wd (A : Type) (Aeq : relation A) := fun_wd Aeq iff. + +Definition relation_wd (A B : Type) (Aeq : relation A) (Beq : relation B) := + fun2_wd Aeq Beq iff. + +Definition relations_eq (A B : Type) (R1 R2 : A -> B -> Prop) := + forall (x : A) (y : B), R1 x y <-> R2 x y. + +Theorem relations_eq_equiv : + forall (A B : Type), equiv (A -> B -> Prop) (@relations_eq A B). +Proof. +intros A B; unfold equiv. split; [| split]; +unfold reflexive, symmetric, transitive, relations_eq. +reflexivity. +intros R1 R2 R3 H1 H2 x y; rewrite H1; apply H2. +now symmetry. +Qed. + +Add Parametric Relation (A B : Type) : (A -> B -> Prop) (@relations_eq A B) + reflexivity proved by (proj1 (relations_eq_equiv A B)) + symmetry proved by (proj2 (proj2 (relations_eq_equiv A B))) + transitivity proved by (proj1 (proj2 (relations_eq_equiv A B))) +as relations_eq_rel. + +Add Parametric Morphism (A : Type) : (@well_founded A) with signature (@relations_eq A A) ==> iff as well_founded_wd. +Proof. +unfold relations_eq, well_founded; intros R1 R2 H; +split; intros H1 a; induction (H1 a) as [x H2 H3]; constructor; +intros y H4; apply H3; [now apply <- H | now apply -> H]. +Qed. + +(* solve_predicate_wd solves the goal [predicate_wd P] for P consisting of +morhisms and quatifiers *) + +Ltac solve_predicate_wd := +unfold predicate_wd; +let x := fresh "x" in +let y := fresh "y" in +let H := fresh "H" in + intros x y H; setoid_rewrite H; reflexivity. + +(* solve_relation_wd solves the goal [relation_wd R] for R consisting of +morhisms and quatifiers *) + +Ltac solve_relation_wd := +unfold relation_wd, fun2_wd; +let x1 := fresh "x" in +let y1 := fresh "y" in +let H1 := fresh "H" in +let x2 := fresh "x" in +let y2 := fresh "y" in +let H2 := fresh "H" in + intros x1 y1 H1 x2 y2 H2; + rewrite H1; setoid_rewrite H2; reflexivity. + +(* The following tactic uses solve_predicate_wd to solve the goals +relating to well-defidedness that are produced by applying induction. +We declare it to take the tactic that applies the induction theorem +and not the induction theorem itself because the tactic may, for +example, supply additional arguments, as does NZinduct_center in +NZBase.v *) + +Ltac induction_maker n t := + try intros until n; + pattern n; t; clear n; + [solve_predicate_wd | ..]. + +(** Relations on cartesian product. Used in MiscFunct for defining +functions whose domain is a product of sets by primitive recursion *) + +Section RelationOnProduct. + +Variables A B : Set. +Variable Aeq : relation A. +Variable Beq : relation B. + +Hypothesis EA_equiv : equiv A Aeq. +Hypothesis EB_equiv : equiv B Beq. + +Definition prod_rel : relation (A * B) := + fun p1 p2 => Aeq (fst p1) (fst p2) /\ Beq (snd p1) (snd p2). + +Lemma prod_rel_refl : reflexive (A * B) prod_rel. +Proof. +unfold reflexive, prod_rel. +destruct x; split; [apply (proj1 EA_equiv) | apply (proj1 EB_equiv)]; simpl. +Qed. + +Lemma prod_rel_symm : symmetric (A * B) prod_rel. +Proof. +unfold symmetric, prod_rel. +destruct x; destruct y; +split; [apply (proj2 (proj2 EA_equiv)) | apply (proj2 (proj2 EB_equiv))]; simpl in *; tauto. +Qed. + +Lemma prod_rel_trans : transitive (A * B) prod_rel. +Proof. +unfold transitive, prod_rel. +destruct x; destruct y; destruct z; simpl. +intros; split; [apply (proj1 (proj2 EA_equiv)) with (y := a0) | +apply (proj1 (proj2 EB_equiv)) with (y := b0)]; tauto. +Qed. + +Theorem prod_rel_equiv : equiv (A * B) prod_rel. +Proof. +unfold equiv; split; [exact prod_rel_refl | split; [exact prod_rel_trans | exact prod_rel_symm]]. +Qed. + +End RelationOnProduct. + +Implicit Arguments prod_rel [A B]. +Implicit Arguments prod_rel_equiv [A B]. + +(** Miscellaneous *) + +(*Definition comp_bool (x y : comparison) : bool := +match x, y with +| Lt, Lt => true +| Eq, Eq => true +| Gt, Gt => true +| _, _ => false +end. + +Theorem comp_bool_correct : forall x y : comparison, + comp_bool x y <-> x = y. +Proof. +destruct x; destruct y; simpl; split; now intro. +Qed.*) + +Lemma eq_equiv : forall A : Set, equiv A (@eq A). +Proof. +intro A; unfold equiv, reflexive, symmetric, transitive. +repeat split; [exact (@trans_eq A) | exact (@sym_eq A)]. +(* It is interesting how the tactic split proves reflexivity *) +Qed. + +(*Add Relation (fun A : Set => A) LE_Set + reflexivity proved by (fun A : Set => (proj1 (eq_equiv A))) + symmetry proved by (fun A : Set => (proj2 (proj2 (eq_equiv A)))) + transitivity proved by (fun A : Set => (proj1 (proj2 (eq_equiv A)))) +as EA_rel.*) diff --git a/theories/Numbers/Rational/BigQ/BigQ.v b/theories/Numbers/Rational/BigQ/BigQ.v new file mode 100644 index 00000000..39e120f7 --- /dev/null +++ b/theories/Numbers/Rational/BigQ/BigQ.v @@ -0,0 +1,35 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: BigQ.v 11028 2008-06-01 17:34:19Z letouzey $ i*) + +Require Export QMake_base. +Require Import QpMake. +Require Import QvMake. +Require Import Q0Make. +Require Import QifMake. +Require Import QbiMake. + +(* We choose for Q the implemention with + multiple representation of 0: 0, 1/0, 2/0 etc *) + +Module BigQ <: QSig.QType := Q0. + +Notation bigQ := BigQ.t. + +Delimit Scope bigQ_scope with bigQ. +Bind Scope bigQ_scope with bigQ. +Bind Scope bigQ_scope with BigQ.t. + +Notation " i + j " := (BigQ.add i j) : bigQ_scope. +Notation " i - j " := (BigQ.sub i j) : bigQ_scope. +Notation " i * j " := (BigQ.mul i j) : bigQ_scope. +Notation " i / j " := (BigQ.div i j) : bigQ_scope. +Notation " i ?= j " := (BigQ.compare i j) : bigQ_scope. diff --git a/theories/Numbers/Rational/BigQ/Q0Make.v b/theories/Numbers/Rational/BigQ/Q0Make.v new file mode 100644 index 00000000..93f52c03 --- /dev/null +++ b/theories/Numbers/Rational/BigQ/Q0Make.v @@ -0,0 +1,1412 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: Q0Make.v 11028 2008-06-01 17:34:19Z letouzey $ i*) + +Require Import Bool. +Require Import ZArith. +Require Import Znumtheory. +Require Import BigNumPrelude. +Require Import Arith. +Require Export BigN. +Require Export BigZ. +Require Import QArith. +Require Import Qcanon. +Require Import Qpower. +Require Import QSig. +Require Import QMake_base. + +Module Q0 <: QType. + + Import BinInt Zorder. + + (** The notation of a rational number is either an integer x, + interpreted as itself or a pair (x,y) of an integer x and a natural + number y interpreted as x/y. The pairs (x,0) and (0,y) are all + interpreted as 0. *) + + Definition t := q_type. + + (** Specification with respect to [QArith] *) + + Open Local Scope Q_scope. + + Definition of_Z x: t := Qz (BigZ.of_Z x). + + Definition of_Q q: t := + match q with x # y => + Qq (BigZ.of_Z x) (BigN.of_N (Npos y)) + end. + + Definition to_Q (q: t) := + match q with + Qz x => BigZ.to_Z x # 1 + |Qq x y => if BigN.eq_bool y BigN.zero then 0 + else BigZ.to_Z x # Z2P (BigN.to_Z y) + end. + + Notation "[ x ]" := (to_Q x). + + Theorem strong_spec_of_Q: forall q: Q, [of_Q q] = q. + Proof. + intros (x,y); simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigN.spec_of_pos; intros HH; discriminate HH. + rewrite BigZ.spec_of_Z; simpl. + rewrite (BigN.spec_of_pos); auto. + Qed. + + Theorem spec_of_Q: forall q: Q, [of_Q q] == q. + Proof. + intros; rewrite strong_spec_of_Q; red; auto. + Qed. + + Definition eq x y := [x] == [y]. + + Definition zero: t := Qz BigZ.zero. + Definition one: t := Qz BigZ.one. + Definition minus_one: t := Qz BigZ.minus_one. + + Lemma spec_0: [zero] == 0. + Proof. + reflexivity. + Qed. + + Lemma spec_1: [one] == 1. + Proof. + reflexivity. + Qed. + + Lemma spec_m1: [minus_one] == -(1). + Proof. + reflexivity. + Qed. + + Definition opp (x: t): t := + match x with + | Qz zx => Qz (BigZ.opp zx) + | Qq nx dx => Qq (BigZ.opp nx) dx + end. + + Theorem strong_spec_opp: forall q, [opp q] = -[q]. + Proof. + intros [z | x y]; simpl. + rewrite BigZ.spec_opp; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigZ.spec_opp; auto. + Qed. + + Theorem spec_opp : forall q, [opp q] == -[q]. + Proof. + intros; rewrite strong_spec_opp; red; auto. + Qed. + + Definition compare (x y: t) := + match x, y with + | Qz zx, Qz zy => BigZ.compare zx zy + | Qz zx, Qq ny dy => + if BigN.eq_bool dy BigN.zero then BigZ.compare zx BigZ.zero + else BigZ.compare (BigZ.mul zx (BigZ.Pos dy)) ny + | Qq nx dx, Qz zy => + if BigN.eq_bool dx BigN.zero then BigZ.compare BigZ.zero zy + else BigZ.compare nx (BigZ.mul zy (BigZ.Pos dx)) + | Qq nx dx, Qq ny dy => + match BigN.eq_bool dx BigN.zero, BigN.eq_bool dy BigN.zero with + | true, true => Eq + | true, false => BigZ.compare BigZ.zero ny + | false, true => BigZ.compare nx BigZ.zero + | false, false => BigZ.compare (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) + end + end. + + Theorem spec_compare: forall q1 q2, (compare q1 q2) = ([q1] ?= [q2]). + Proof. + intros [z1 | x1 y1] [z2 | x2 y2]; + unfold Qcompare, compare, to_Q, Qnum, Qden. + repeat rewrite Zmult_1_r. + generalize (BigZ.spec_compare z1 z2); case BigZ.compare; intros H; auto. + rewrite H; rewrite Zcompare_refl; auto. + rewrite Zmult_1_r. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + rewrite Zmult_1_r; generalize (BigZ.spec_compare z1 BigZ.zero); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH1; rewrite HH1; rewrite Zcompare_refl; auto. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (z1 * BigZ.Pos y2) x2)%bigZ; case BigZ.compare; + rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare BigZ.zero z2); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH1; rewrite <- HH1; rewrite Zcompare_refl; auto. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y1); auto with zarith. + rewrite Zmult_1_r. + generalize (BigZ.spec_compare x1 (z2 * BigZ.Pos y1))%bigZ; case BigZ.compare; + rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + rewrite Zcompare_refl; auto. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare BigZ.zero x2); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH2; rewrite <- HH2; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare x1 BigZ.zero)%bigZ; case BigZ.compare; + auto; rewrite BigZ.spec_0. + intros HH2; rewrite <- HH2; rewrite Zcompare_refl; auto. + repeat rewrite Z2P_correct. + 2: generalize (BigN.spec_pos y1); auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (x1 * BigZ.Pos y2) + (x2 * BigZ.Pos y1))%bigZ; case BigZ.compare; + repeat rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + Qed. + + Definition lt n m := compare n m = Lt. + Definition le n m := compare n m <> Gt. + Definition min n m := match compare n m with Gt => m | _ => n end. + Definition max n m := match compare n m with Lt => m | _ => n end. + +(* Je pense que cette fonction normalise bien ... *) + Definition norm n d: t := + let gcd := BigN.gcd (BigZ.to_N n) d in + match BigN.compare BigN.one gcd with + | Lt => + let n := BigZ.div n (BigZ.Pos gcd) in + let d := BigN.div d gcd in + match BigN.compare d BigN.one with + | Gt => Qq n d + | Eq => Qz n + | Lt => zero + end + | Eq => Qq n d + | Gt => zero (* gcd = 0 => both numbers are 0 *) + end. + + Theorem spec_norm: forall n q, [norm n q] == [Qq n q]. + Proof. + intros p q; unfold norm. + assert (Hp := BigN.spec_pos (BigZ.to_N p)). + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; auto; rewrite BigN.spec_1; rewrite BigN.spec_gcd; intros H1. + apply Qeq_refl. + generalize (BigN.spec_pos (q / BigN.gcd (BigZ.to_N p) q)%bigN). + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; auto; rewrite BigN.spec_1; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; intros H2 HH. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H3; simpl; + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + auto with zarith. + generalize H2; rewrite H3; + rewrite Zdiv_0_l; auto with zarith. + generalize H1 H2 H3 (BigN.spec_pos q); clear H1 H2 H3. + rewrite spec_to_N. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 H4; rewrite Z2P_correct; auto with zarith. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H3; simpl. + case H3. + generalize H1 H2 H3 HH; clear H1 H2 H3 HH. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 HH. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto with zarith. + case (Zle_lt_or_eq _ _ HH); auto with zarith. + intros HH1; rewrite <- HH1; ring. + generalize (Zgcd_is_gcd a b); intros HH1; inversion HH1; auto. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; intros H3. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H4. + case H3; rewrite H4; rewrite Zdiv_0_l; auto with zarith. + simpl. + assert (FF := BigN.spec_pos q). + rewrite Z2P_correct; auto with zarith. + rewrite <- BigN.spec_gcd; rewrite <- BigN.spec_div; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; auto with zarith. + simpl; rewrite BigZ.spec_div; simpl. + rewrite BigN.spec_gcd; auto with zarith. + generalize H1 H2 H3 H4 HH FF; clear H1 H2 H3 H4 HH FF. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 H4 HH FF. + rewrite spec_to_N; fold a. + rewrite Zgcd_div_swap; auto with zarith. + rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_gcd; auto with zarith. + case (Zle_lt_or_eq _ _ + (BigN.spec_pos (BigN.gcd (BigZ.to_N p) q))); + rewrite BigN.spec_gcd; auto with zarith. + intros; apply False_ind; auto with zarith. + intros HH2; assert (FF1 := Zgcd_inv_0_l _ _ (sym_equal HH2)). + assert (FF2 := Zgcd_inv_0_l _ _ (sym_equal HH2)). + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H2; simpl. + rewrite spec_to_N. + rewrite FF2; ring. + Qed. + + + Definition add (x y: t): t := + match x with + | Qz zx => + match y with + | Qz zy => Qz (BigZ.add zx zy) + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else Qq (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + end + | Qq nx dx => + if BigN.eq_bool dx BigN.zero then y + else match y with + | Qz zy => Qq (BigZ.add nx (BigZ.mul zy (BigZ.Pos dx))) dx + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + Qq n d + end + end. + + Theorem spec_add : forall x y, [add x y] == [x] + [y]. + Proof. + intros [x | nx dx] [y | ny dy]; unfold Qplus; simpl. + rewrite BigZ.spec_add; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + assert (F1:= BigN.spec_pos dy). + rewrite Zmult_1_r; red; simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH; simpl; try ring. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH1; simpl; try ring. + case HH; auto. + rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul; simpl; auto. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH; simpl; try ring. + rewrite Zmult_1_r; apply Qeq_refl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH1; simpl; try ring. + case HH; auto. + rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul; simpl; auto. + rewrite Zmult_1_r; rewrite Pmult_1_r. + apply Qeq_refl. + assert (F1:= BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + apply Qeq_refl. + case HH2; auto. + simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + case HH2; auto. + case HH1; auto. + rewrite Zmult_1_r; apply Qeq_refl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + simpl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + case HH; auto. + rewrite Zmult_1_r; rewrite Zplus_0_r; rewrite Pmult_1_r. + apply Qeq_refl. + simpl. + generalize (BigN.spec_eq_bool (dx * dy)%bigN BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_mul; + rewrite BigN.spec_0; intros HH2. + (case (Zmult_integral _ _ HH2); intros HH3); + [case HH| case HH1]; auto. + rewrite BigZ.spec_add; repeat rewrite BigZ.spec_mul; simpl. + assert (Fx: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (Fy: (0 < BigN.to_Z dy)%Z). + generalize (BigN.spec_pos dy); auto with zarith. + red; simpl; rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto with zarith. + apply Zmult_lt_0_compat; auto. + Qed. + + Definition add_norm (x y: t): t := + match x with + | Qz zx => + match y with + | Qz zy => Qz (BigZ.add zx zy) + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else norm (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + end + | Qq nx dx => + if BigN.eq_bool dx BigN.zero then y + else match y with + | Qz zy => norm (BigZ.add nx (BigZ.mul zy (BigZ.Pos dx))) dx + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + norm n d + end + end. + + Theorem spec_add_norm : forall x y, [add_norm x y] == [x] + [y]. + Proof. + intros x y; rewrite <- spec_add; auto. + case x; case y; clear x y; unfold add_norm, add. + intros; apply Qeq_refl. + intros p1 n p2. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + simpl. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + apply Qeq_refl. + apply Qeq_refl. + intros p1 p2 n. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + apply Qeq_refl. + intros p1 q1 p2 q2. + generalize (BigN.spec_eq_bool q2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + apply Qeq_refl. + generalize (BigN.spec_eq_bool q1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + apply Qeq_refl. + Qed. + + Definition sub x y := add x (opp y). + + Theorem spec_sub : forall x y, [sub x y] == [x] - [y]. + Proof. + intros x y; unfold sub; rewrite spec_add; auto. + rewrite spec_opp; ring. + Qed. + + Definition sub_norm x y := add_norm x (opp y). + + Theorem spec_sub_norm : forall x y, [sub_norm x y] == [x] - [y]. + Proof. + intros x y; unfold sub_norm; rewrite spec_add_norm; auto. + rewrite spec_opp; ring. + Qed. + + Definition mul (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => Qq (BigZ.mul zx ny) dy + | Qq nx dx, Qz zy => Qq (BigZ.mul nx zy) dx + | Qq nx dx, Qq ny dy => Qq (BigZ.mul nx ny) (BigN.mul dx dy) + end. + + Theorem spec_mul : forall x y, [mul x y] == [x] * [y]. + Proof. + intros [x | nx dx] [y | ny dy]; unfold Qmult; simpl. + rewrite BigZ.spec_mul; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH1. + red; simpl; ring. + rewrite BigZ.spec_mul; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH1. + red; simpl; ring. + rewrite BigZ.spec_mul; rewrite Pmult_1_r. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_mul; + intros HH1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH2. + red; simpl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH3. + red; simpl; ring. + case (Zmult_integral _ _ HH1); intros HH. + case HH2; auto. + case HH3; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH2. + case HH1; rewrite HH2; ring. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH3. + case HH1; rewrite HH3; ring. + rewrite BigZ.spec_mul. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto. + apply Qeq_refl. + generalize (BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_pos dy); auto with zarith. + Qed. + +Definition mul_norm (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => + if BigZ.eq_bool zx BigZ.zero then zero + else + let gcd := BigN.gcd (BigZ.to_N zx) dy in + match BigN.compare gcd BigN.one with + Gt => + let zx := BigZ.div zx (BigZ.Pos gcd) in + let d := BigN.div dy gcd in + if BigN.eq_bool d BigN.one then Qz (BigZ.mul zx ny) + else Qq (BigZ.mul zx ny) d + | _ => Qq (BigZ.mul zx ny) dy + end + | Qq nx dx, Qz zy => + if BigZ.eq_bool zy BigZ.zero then zero + else + let gcd := BigN.gcd (BigZ.to_N zy) dx in + match BigN.compare gcd BigN.one with + Gt => + let zy := BigZ.div zy (BigZ.Pos gcd) in + let d := BigN.div dx gcd in + if BigN.eq_bool d BigN.one then Qz (BigZ.mul zy nx) + else Qq (BigZ.mul zy nx) d + | _ => Qq (BigZ.mul zy nx) dx + end + | Qq nx dx, Qq ny dy => + let (nx, dy) := + let gcd := BigN.gcd (BigZ.to_N nx) dy in + match BigN.compare gcd BigN.one with + Gt => (BigZ.div nx (BigZ.Pos gcd), BigN.div dy gcd) + | _ => (nx, dy) + end in + let (ny, dx) := + let gcd := BigN.gcd (BigZ.to_N ny) dx in + match BigN.compare gcd BigN.one with + Gt => (BigZ.div ny (BigZ.Pos gcd), BigN.div dx gcd) + | _ => (ny, dx) + end in + let d := (BigN.mul dx dy) in + if BigN.eq_bool d BigN.one then Qz (BigZ.mul ny nx) + else Qq (BigZ.mul ny nx) d + end. + + Theorem spec_mul_norm : forall x y, [mul_norm x y] == [x] * [y]. + Proof. + intros x y; rewrite <- spec_mul; auto. + unfold mul_norm, mul; case x; case y; clear x y. + intros; apply Qeq_refl. + intros p1 n p2. + set (a := BigN.to_Z (BigZ.to_N p2)). + set (b := BigN.to_Z n). + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; unfold zero, to_Q; repeat rewrite BigZ.spec_0; intros H. + case BigN.eq_bool; try apply Qeq_refl. + rewrite BigZ.spec_mul; rewrite H. + red; simpl; ring. + assert (F: (0 < a)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos (BigZ.to_N p2))); auto. + intros H1; case H; rewrite spec_to_N; rewrite <- H1; ring. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; rewrite BigN.spec_gcd; + fold a b; intros H1. + apply Qeq_refl. + apply Qeq_refl. + assert (F0 : (0 < (Zgcd a b))%Z). + apply Zlt_trans with 1%Z. + red; auto. + apply Zgt_lt; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; + fold a b; intros H2. + assert (F1: b = Zgcd a b). + pattern b at 1; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); + auto with zarith. + rewrite H2; ring. + assert (FF := Zgcd_is_gcd a b); inversion FF; auto. + assert (F2: (0 < b)%Z). + rewrite F1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; fold b; intros H3. + rewrite H3 in F2; discriminate F2. + rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + rewrite BigZ.spec_mul. + red; simpl; rewrite Z2P_correct; auto. + rewrite Zmult_1_r; rewrite spec_to_N; fold a b. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p1)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; fold a b; auto; intros H3. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H4. + apply Qeq_refl. + case H4; fold b. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite H3; ring. + assert (FF := Zgcd_is_gcd a b); inversion FF; auto. + simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; fold b; intros H4. + case H3; rewrite H4; rewrite Zdiv_0_l; auto. + rewrite BigZ.spec_mul; rewrite BigZ.spec_div; simpl; + rewrite BigN.spec_gcd; fold a b; auto with zarith. + assert (F1: (0 < b)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos n)); fold b; auto with zarith. + red; simpl. + rewrite BigZ.spec_mul. + repeat rewrite Z2P_correct; auto. + rewrite spec_to_N; fold a. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p1)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + ring. + apply Zgcd_div_pos; auto. + intros p1 p2 n. + set (a := BigN.to_Z (BigZ.to_N p1)). + set (b := BigN.to_Z n). + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; unfold zero, to_Q; repeat rewrite BigZ.spec_0; intros H. + case BigN.eq_bool; try apply Qeq_refl. + rewrite BigZ.spec_mul; rewrite H. + red; simpl; ring. + assert (F: (0 < a)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos (BigZ.to_N p1))); auto. + intros H1; case H; rewrite spec_to_N; rewrite <- H1; ring. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; rewrite BigN.spec_gcd; + fold a b; intros H1. + repeat rewrite BigZ.spec_mul; rewrite Zmult_comm. + apply Qeq_refl. + repeat rewrite BigZ.spec_mul; rewrite Zmult_comm. + apply Qeq_refl. + assert (F0 : (0 < (Zgcd a b))%Z). + apply Zlt_trans with 1%Z. + red; auto. + apply Zgt_lt; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; + fold a b; intros H2. + assert (F1: b = Zgcd a b). + pattern b at 1; rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); + auto with zarith. + rewrite H2; ring. + assert (FF := Zgcd_is_gcd a b); inversion FF; auto. + assert (F2: (0 < b)%Z). + rewrite F1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; fold b; intros H3. + rewrite H3 in F2; discriminate F2. + rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + rewrite BigZ.spec_mul. + red; simpl; rewrite Z2P_correct; auto. + rewrite Zmult_1_r; rewrite spec_to_N; fold a b. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p2)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; fold a b; auto; intros H3. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H4. + apply Qeq_refl. + case H4; fold b. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite H3; ring. + assert (FF := Zgcd_is_gcd a b); inversion FF; auto. + simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; fold b; intros H4. + case H3; rewrite H4; rewrite Zdiv_0_l; auto. + rewrite BigZ.spec_mul; rewrite BigZ.spec_div; simpl; + rewrite BigN.spec_gcd; fold a b; auto with zarith. + assert (F1: (0 < b)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos n)); fold b; auto with zarith. + red; simpl. + rewrite BigZ.spec_mul. + repeat rewrite Z2P_correct; auto. + rewrite spec_to_N; fold a. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p2)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + ring. + apply Zgcd_div_pos; auto. + set (f := fun p t => + match (BigN.gcd (BigZ.to_N p) t ?= BigN.one)%bigN with + | Eq => (p, t) + | Lt => (p, t) + | Gt => + ((p / BigZ.Pos (BigN.gcd (BigZ.to_N p) t))%bigZ, + (t / BigN.gcd (BigZ.to_N p) t)%bigN) + end). + assert (F: forall p t, + let (n, d) := f p t in [Qq p t] == [Qq n d]). + intros p t1; unfold f. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; rewrite BigN.spec_gcd; intros H1. + apply Qeq_refl. + apply Qeq_refl. + set (a := BigN.to_Z (BigZ.to_N p)). + set (b := BigN.to_Z t1). + fold a b in H1. + assert (F0 : (0 < (Zgcd a b))%Z). + apply Zlt_trans with 1%Z. + red; auto. + apply Zgt_lt; auto. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; fold b; intros HH1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; fold b; intros HH2. + simpl; ring. + case HH2. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; fold a b; auto. + rewrite HH1; rewrite Zdiv_0_l; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; + rewrite BigN.spec_div; rewrite BigN.spec_gcd; fold a b; auto; + intros HH2. + case HH1. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite HH2; ring. + assert (FF := Zgcd_is_gcd a b); inversion FF; auto. + simpl. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; fold a b; auto with zarith. + assert (F1: (0 < b)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos t1)); fold b; auto with zarith. + intros HH; case HH1; auto. + repeat rewrite Z2P_correct; auto. + rewrite spec_to_N; fold a. + rewrite Zgcd_div_swap; auto. + apply Zgcd_div_pos; auto. + intros HH; rewrite HH in F0; discriminate F0. + intros p1 n1 p2 n2. + change ([let (nx , dy) := f p2 n1 in + let (ny, dx) := f p1 n2 in + if BigN.eq_bool (dx * dy)%bigN BigN.one + then Qz (ny * nx) + else Qq (ny * nx) (dx * dy)] == [Qq (p2 * p1) (n2 * n1)]). + generalize (F p2 n1) (F p1 n2). + case f; case f. + intros u1 u2 v1 v2 Hu1 Hv1. + apply Qeq_trans with [mul (Qq p2 n1) (Qq p1 n2)]. + rewrite spec_mul; rewrite Hu1; rewrite Hv1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; rewrite BigN.spec_mul; intros HH1. + assert (F1: BigN.to_Z u2 = 1%Z). + case (Zmult_1_inversion_l _ _ HH1); auto. + generalize (BigN.spec_pos u2); auto with zarith. + assert (F2: BigN.to_Z v2 = 1%Z). + rewrite Zmult_comm in HH1. + case (Zmult_1_inversion_l _ _ HH1); auto. + generalize (BigN.spec_pos v2); auto with zarith. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1. + rewrite H1 in F2; discriminate F2. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2. + rewrite H2 in F1; discriminate F1. + simpl; rewrite BigZ.spec_mul. + rewrite F1; rewrite F2; simpl; ring. + rewrite Qmult_comm; rewrite <- spec_mul. + apply Qeq_refl. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_mul; + rewrite Zmult_comm; intros H1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_mul; intros H2; auto. + case H2; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_mul; intros H2; auto. + case H1; auto. + Qed. + + +Definition inv (x: t): t := + match x with + | Qz (BigZ.Pos n) => Qq BigZ.one n + | Qz (BigZ.Neg n) => Qq BigZ.minus_one n + | Qq (BigZ.Pos n) d => Qq (BigZ.Pos d) n + | Qq (BigZ.Neg n) d => Qq (BigZ.Neg d) n + end. + + Theorem spec_inv : forall x, [inv x] == /[x]. + Proof. + intros [ [x | x] | [nx | nx] dx]; unfold inv, Qinv; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + rewrite H1; apply Qeq_refl. + generalize H1 (BigN.spec_pos x); case (BigN.to_Z x); auto. + intros HH; case HH; auto. + intros; red; simpl; auto. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + rewrite H1; apply Qeq_refl. + generalize H1 (BigN.spec_pos x); case (BigN.to_Z x); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl; auto. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + apply Qeq_refl. + rewrite H1; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + rewrite H2; red; simpl; auto. + generalize H1 (BigN.spec_pos nx); case (BigN.to_Z nx); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + apply Qeq_refl. + rewrite H1; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + rewrite H2; red; simpl; auto. + generalize H1 (BigN.spec_pos nx); case (BigN.to_Z nx); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl. + assert (tmp: forall x, Zneg x = Zopp (Zpos x)); auto. + rewrite tmp. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + ring. + generalize (BigN.spec_pos dx); auto with zarith. + intros p _ HH; case HH; auto. + Qed. + +Definition inv_norm (x: t): t := + match x with + | Qz (BigZ.Pos n) => + match BigN.compare n BigN.one with + Gt => Qq BigZ.one n + | _ => x + end + | Qz (BigZ.Neg n) => + match BigN.compare n BigN.one with + Gt => Qq BigZ.minus_one n + | _ => x + end + | Qq (BigZ.Pos n) d => + match BigN.compare n BigN.one with + Gt => Qq (BigZ.Pos d) n + | Eq => Qz (BigZ.Pos d) + | Lt => Qz (BigZ.zero) + end + | Qq (BigZ.Neg n) d => + match BigN.compare n BigN.one with + Gt => Qq (BigZ.Neg d) n + | Eq => Qz (BigZ.Neg d) + | Lt => Qz (BigZ.zero) + end + end. + + Theorem spec_inv_norm : forall x, [inv_norm x] == /[x]. + Proof. + intros [ [x | x] | [nx | nx] dx]; unfold inv_norm, Qinv. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H. + simpl; rewrite H; apply Qeq_refl. + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); simpl. + generalize H; case BigN.to_Z. + intros _ HH; discriminate HH. + intros p; case p; auto. + intros p1 HH; discriminate HH. + intros p1 HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; discriminate HH. + intros HH; rewrite <- HH. + apply Qeq_refl. + generalize H; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1. + rewrite H1; intros HH; discriminate. + generalize H; case BigN.to_Z. + intros HH; discriminate HH. + intros; red; simpl; auto. + intros p HH; discriminate HH. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H. + simpl; rewrite H; apply Qeq_refl. + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); simpl. + generalize H; case BigN.to_Z. + intros _ HH; discriminate HH. + intros p; case p; auto. + intros p1 HH; discriminate HH. + intros p1 HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; discriminate HH. + intros HH; rewrite <- HH. + apply Qeq_refl. + generalize H; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1. + rewrite H1; intros HH; discriminate. + generalize H; case BigN.to_Z. + intros HH; discriminate HH. + intros; red; simpl; auto. + intros p HH; discriminate HH. + simpl Qnum. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; simpl. + case BigN.compare; red; simpl; auto. + rewrite H1; auto. + case BigN.eq_bool; auto. + simpl; rewrite H1; auto. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H2. + rewrite H2. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + red; simpl. + rewrite Zmult_1_r; rewrite Pmult_1_r; rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + generalize H2 (BigN.spec_pos nx); case (BigN.to_Z nx). + intros; apply Qeq_refl. + intros p; case p; clear p. + intros p HH; discriminate HH. + intros p HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + simpl; generalize H2; case (BigN.to_Z nx). + intros HH; discriminate HH. + intros p Hp. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H4. + rewrite H4 in H2; discriminate H2. + red; simpl. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + intros p HH; discriminate HH. + simpl Qnum. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; simpl. + case BigN.compare; red; simpl; auto. + rewrite H1; auto. + case BigN.eq_bool; auto. + simpl; rewrite H1; auto. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H2. + rewrite H2. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + red; simpl. + assert (tmp: forall x, Zneg x = Zopp (Zpos x)); auto. + rewrite tmp. + rewrite Zmult_1_r; rewrite Pmult_1_r; rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + generalize H2 (BigN.spec_pos nx); case (BigN.to_Z nx). + intros; apply Qeq_refl. + intros p; case p; clear p. + intros p HH; discriminate HH. + intros p HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + simpl; generalize H2; case (BigN.to_Z nx). + intros HH; discriminate HH. + intros p Hp. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H4. + rewrite H4 in H2; discriminate H2. + red; simpl. + assert (tmp: forall x, Zneg x = Zopp (Zpos x)); auto. + rewrite tmp. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + ring. + generalize (BigN.spec_pos dx); auto with zarith. + intros p HH; discriminate HH. + Qed. + + Definition div x y := mul x (inv y). + + Theorem spec_div x y: [div x y] == [x] / [y]. + Proof. + intros x y; unfold div; rewrite spec_mul; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Definition div_norm x y := mul_norm x (inv y). + + Theorem spec_div_norm x y: [div_norm x y] == [x] / [y]. + Proof. + intros x y; unfold div_norm; rewrite spec_mul_norm; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Definition square (x: t): t := + match x with + | Qz zx => Qz (BigZ.square zx) + | Qq nx dx => Qq (BigZ.square nx) (BigN.square dx) + end. + + Theorem spec_square : forall x, [square x] == [x] ^ 2. + Proof. + intros [ x | nx dx]; unfold square. + red; simpl; rewrite BigZ.spec_square; auto with zarith. + simpl Qpower. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H. + red; simpl. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_square; + intros H1. + case H1; rewrite H; auto. + red; simpl. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_square; + intros H1. + case H; case (Zmult_integral _ _ H1); auto. + simpl. + rewrite BigZ.spec_square. + rewrite Zpos_mult_morphism. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto. + generalize (BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_pos dx); auto with zarith. + Qed. + + Definition power_pos (x: t) p: t := + match x with + | Qz zx => Qz (BigZ.power_pos zx p) + | Qq nx dx => Qq (BigZ.power_pos nx p) (BigN.power_pos dx p) + end. + + Theorem spec_power_pos : forall x p, [power_pos x p] == [x] ^ Zpos p. + Proof. + intros [x | nx dx] p; unfold power_pos. + unfold power_pos; red; simpl. + generalize (Qpower_decomp p (BigZ.to_Z x) 1). + unfold Qeq; simpl. + rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Zmult_1_r. + intros H; rewrite H. + rewrite BigZ.spec_power_pos; simpl; ring. + simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_power_pos; intros H1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H2. + elim p; simpl. + intros; red; simpl; auto. + intros p1 Hp1; rewrite <- Hp1; red; simpl; auto. + apply Qeq_refl. + case H2; generalize H1. + elim p; simpl. + intros p1 Hrec. + change (xI p1) with (1 + (xO p1))%positive. + rewrite Zpower_pos_is_exp; rewrite Zpower_pos_1_r. + intros HH; case (Zmult_integral _ _ HH); auto. + rewrite <- Pplus_diag. + rewrite Zpower_pos_is_exp. + intros HH1; case (Zmult_integral _ _ HH1); auto. + intros p1 Hrec. + rewrite <- Pplus_diag. + rewrite Zpower_pos_is_exp. + intros HH1; case (Zmult_integral _ _ HH1); auto. + rewrite Zpower_pos_1_r; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H2. + case H1; rewrite H2; auto. + simpl; rewrite Zpower_pos_0_l; auto. + assert (F1: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (F2: (0 < BigN.to_Z dx ^ ' p)%Z). + unfold Zpower; apply Zpower_pos_pos; auto. + unfold power_pos; red; simpl. + generalize (Qpower_decomp p (BigZ.to_Z nx) + (Z2P (BigN.to_Z dx))). + unfold Qeq; simpl. + repeat rewrite Z2P_correct; auto. + unfold Qeq; simpl; intros HH. + rewrite HH. + rewrite BigZ.spec_power_pos; simpl; ring. + Qed. + + (** Interaction with [Qcanon.Qc] *) + + Open Scope Qc_scope. + + Definition of_Qc q := of_Q (this q). + + Definition to_Qc q := !!(to_Q q). + + Notation "[[ x ]]" := (to_Qc x). + + Theorem spec_of_Qc: forall q, [[of_Qc q]] = q. + Proof. + intros (x, Hx); unfold of_Qc, to_Qc; simpl. + apply Qc_decomp; simpl. + intros. + rewrite <- H0 at 2; apply Qred_complete. + apply spec_of_Q. + Qed. + + Theorem spec_oppc: forall q, [[opp q]] = -[[q]]. + Proof. + intros q; unfold Qcopp, to_Qc, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + rewrite spec_opp. + rewrite <- Qred_opp. + rewrite Qred_correct; red; auto. + Qed. + + Theorem spec_comparec: forall q1 q2, + compare q1 q2 = ([[q1]] ?= [[q2]]). + Proof. + unfold Qccompare, to_Qc. + intros q1 q2; rewrite spec_compare; simpl; auto. + apply Qcompare_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_addc x y: + [[add x y]] = [[x]] + [[y]]. + Proof. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_add_normc x y: + [[add_norm x y]] = [[x]] + [[y]]. + Proof. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add_norm; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_subc x y: [[sub x y]] = [[x]] - [[y]]. + Proof. + intros x y; unfold sub; rewrite spec_addc; auto. + rewrite spec_oppc; ring. + Qed. + + Theorem spec_sub_normc x y: + [[sub_norm x y]] = [[x]] - [[y]]. + intros x y; unfold sub_norm; rewrite spec_add_normc; auto. + rewrite spec_oppc; ring. + Qed. + + Theorem spec_mulc x y: + [[mul x y]] = [[x]] * [[y]]. + Proof. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_mul_normc x y: + [[mul_norm x y]] = [[x]] * [[y]]. + Proof. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul_norm; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_invc x: + [[inv x]] = /[[x]]. + Proof. + intros x; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv; auto. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_inv_normc x: + [[inv_norm x]] = /[[x]]. + Proof. + intros x; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv_norm; auto. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_divc x y: [[div x y]] = [[x]] / [[y]]. + Proof. + intros x y; unfold div; rewrite spec_mulc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + Theorem spec_div_normc x y: [[div_norm x y]] = [[x]] / [[y]]. + Proof. + intros x y; unfold div_norm; rewrite spec_mul_normc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + Theorem spec_squarec x: [[square x]] = [[x]]^2. + Proof. + intros x; unfold to_Qc. + apply trans_equal with (!! ([x]^2)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_square; auto. + simpl Qcpower. + replace (!! [x] * 1) with (!![x]); try ring. + simpl. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Theorem spec_power_posc x p: + [[power_pos x p]] = [[x]] ^ nat_of_P p. + Proof. + intros x p; unfold to_Qc. + apply trans_equal with (!! ([x]^Zpos p)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_power_pos; auto. + pattern p; apply Pind; clear p. + simpl; ring. + intros p Hrec. + rewrite nat_of_P_succ_morphism; simpl Qcpower. + rewrite <- Hrec. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; + unfold this. + apply Qred_complete. + assert (F: [x] ^ ' Psucc p == [x] * [x] ^ ' p). + simpl; case x; simpl; clear x Hrec. + intros x; simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + intros nx dx. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + unfold Qpower_positive. + assert (tmp: forall p, pow_pos Qmult 0%Q p = 0%Q). + intros p1; elim p1; simpl; auto; clear p1. + intros p1 Hp1; rewrite Hp1; auto. + intros p1 Hp1; rewrite Hp1; auto. + repeat rewrite tmp; intros; red; simpl; auto. + intros H1. + assert (F1: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + repeat rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto. + 2: apply Zpower_pos_pos; auto. + 2: apply Zpower_pos_pos; auto. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + rewrite F. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + +End Q0. diff --git a/theories/Numbers/Rational/BigQ/QMake_base.v b/theories/Numbers/Rational/BigQ/QMake_base.v new file mode 100644 index 00000000..547e74b7 --- /dev/null +++ b/theories/Numbers/Rational/BigQ/QMake_base.v @@ -0,0 +1,34 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(* $Id: QMake_base.v 10964 2008-05-22 11:08:13Z letouzey $ *) + +(** * An implementation of rational numbers based on big integers *) + +Require Export BigN. +Require Export BigZ. + +(* Basic type for Q: a Z or a pair of a Z and an N *) + +Inductive q_type := + | Qz : BigZ.t -> q_type + | Qq : BigZ.t -> BigN.t -> q_type. + +Definition print_type x := + match x with + | Qz _ => Z + | _ => (Z*Z)%type + end. + +Definition print x := + match x return print_type x with + | Qz zx => BigZ.to_Z zx + | Qq nx dx => (BigZ.to_Z nx, BigN.to_Z dx) + end. diff --git a/theories/Numbers/Rational/BigQ/QbiMake.v b/theories/Numbers/Rational/BigQ/QbiMake.v new file mode 100644 index 00000000..699f383e --- /dev/null +++ b/theories/Numbers/Rational/BigQ/QbiMake.v @@ -0,0 +1,1066 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: QbiMake.v 11027 2008-06-01 13:28:59Z letouzey $ i*) + +Require Import Bool. +Require Import ZArith. +Require Import Znumtheory. +Require Import BigNumPrelude. +Require Import Arith. +Require Export BigN. +Require Export BigZ. +Require Import QArith. +Require Import Qcanon. +Require Import Qpower. +Require Import QMake_base. + +Module Qbi. + + Import BinInt Zorder. + Open Local Scope Q_scope. + Open Local Scope Qc_scope. + + (** The notation of a rational number is either an integer x, + interpreted as itself or a pair (x,y) of an integer x and a naturel + number y interpreted as x/y. The pairs (x,0) and (0,y) are all + interpreted as 0. *) + + Definition t := q_type. + + Definition zero: t := Qz BigZ.zero. + Definition one: t := Qz BigZ.one. + Definition minus_one: t := Qz BigZ.minus_one. + + Definition of_Z x: t := Qz (BigZ.of_Z x). + + + Definition of_Q q: t := + match q with x # y => + Qq (BigZ.of_Z x) (BigN.of_N (Npos y)) + end. + + Definition of_Qc q := of_Q (this q). + + Definition to_Q (q: t) := + match q with + Qz x => BigZ.to_Z x # 1 + |Qq x y => if BigN.eq_bool y BigN.zero then 0%Q + else BigZ.to_Z x # Z2P (BigN.to_Z y) + end. + + Definition to_Qc q := !!(to_Q q). + + Notation "[[ x ]]" := (to_Qc x). + + Notation "[ x ]" := (to_Q x). + + Theorem spec_to_Q: forall q: Q, [of_Q q] = q. + intros (x,y); simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigN.spec_of_pos; intros HH; discriminate HH. + rewrite BigZ.spec_of_Z; simpl. + rewrite (BigN.spec_of_pos); auto. + Qed. + + Theorem spec_to_Qc: forall q, [[of_Qc q]] = q. + intros (x, Hx); unfold of_Qc, to_Qc; simpl. + apply Qc_decomp; simpl. + intros; rewrite spec_to_Q; auto. + Qed. + + Definition opp (x: t): t := + match x with + | Qz zx => Qz (BigZ.opp zx) + | Qq nx dx => Qq (BigZ.opp nx) dx + end. + + Theorem spec_opp: forall q, ([opp q] = -[q])%Q. + intros [z | x y]; simpl. + rewrite BigZ.spec_opp; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigZ.spec_opp; auto. + Qed. + + Theorem spec_oppc: forall q, [[opp q]] = -[[q]]. + intros q; unfold Qcopp, to_Qc, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + rewrite spec_opp. + rewrite <- Qred_opp. + rewrite Qred_involutive; auto. + Qed. + + + Definition compare (x y: t) := + match x, y with + | Qz zx, Qz zy => BigZ.compare zx zy + | Qz zx, Qq ny dy => + if BigN.eq_bool dy BigN.zero then BigZ.compare zx BigZ.zero + else + match BigZ.cmp_sign zx ny with + | Lt => Lt + | Gt => Gt + | Eq => BigZ.compare (BigZ.mul zx (BigZ.Pos dy)) ny + end + | Qq nx dx, Qz zy => + if BigN.eq_bool dx BigN.zero then BigZ.compare BigZ.zero zy + else + match BigZ.cmp_sign nx zy with + | Lt => Lt + | Gt => Gt + | Eq => BigZ.compare nx (BigZ.mul zy (BigZ.Pos dx)) + end + | Qq nx dx, Qq ny dy => + match BigN.eq_bool dx BigN.zero, BigN.eq_bool dy BigN.zero with + | true, true => Eq + | true, false => BigZ.compare BigZ.zero ny + | false, true => BigZ.compare nx BigZ.zero + | false, false => + match BigZ.cmp_sign nx ny with + | Lt => Lt + | Gt => Gt + | Eq => BigZ.compare (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) + end + end + end. + + Theorem spec_compare: forall q1 q2, + compare q1 q2 = ([q1] ?= [q2])%Q. + intros [z1 | x1 y1] [z2 | x2 y2]; + unfold Qcompare, compare, to_Q, Qnum, Qden. + repeat rewrite Zmult_1_r. + generalize (BigZ.spec_compare z1 z2); case BigZ.compare; intros H; auto. + rewrite H; rewrite Zcompare_refl; auto. + rewrite Zmult_1_r. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + rewrite Zmult_1_r; generalize (BigZ.spec_compare z1 BigZ.zero); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH1; rewrite HH1; rewrite Zcompare_refl; auto. + set (a := BigZ.to_Z z1); set (b := BigZ.to_Z x2); + set (c := BigN.to_Z y2); fold c in HH. + assert (F: (0 < c)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos y2)); fold c; auto. + intros H1; case HH; rewrite <- H1; auto. + rewrite Z2P_correct; auto with zarith. + generalize (BigZ.spec_cmp_sign z1 x2); case BigZ.cmp_sign; fold a b c. + intros _; generalize (BigZ.spec_compare (z1 * BigZ.Pos y2)%bigZ x2); + case BigZ.compare; rewrite BigZ.spec_mul; simpl; fold a b c; auto. + intros H1; rewrite H1; rewrite Zcompare_refl; auto. + intros (H1, H2); apply sym_equal; change (a * c < b)%Z. + apply Zlt_le_trans with (2 := H2). + change 0%Z with (0 * c)%Z. + apply Zmult_lt_compat_r; auto with zarith. + intros (H1, H2); apply sym_equal; change (a * c > b)%Z. + apply Zlt_gt. + apply Zlt_le_trans with (1 := H2). + change 0%Z with (0 * c)%Z. + apply Zmult_le_compat_r; auto with zarith. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare BigZ.zero z2); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH1; rewrite <- HH1; rewrite Zcompare_refl; auto. + set (a := BigZ.to_Z z2); set (b := BigZ.to_Z x1); + set (c := BigN.to_Z y1); fold c in HH. + assert (F: (0 < c)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos y1)); fold c; auto. + intros H1; case HH; rewrite <- H1; auto. + rewrite Zmult_1_r; rewrite Z2P_correct; auto with zarith. + generalize (BigZ.spec_cmp_sign x1 z2); case BigZ.cmp_sign; fold a b c. + intros _; generalize (BigZ.spec_compare x1 (z2 * BigZ.Pos y1)%bigZ); + case BigZ.compare; rewrite BigZ.spec_mul; simpl; fold a b c; auto. + intros H1; rewrite H1; rewrite Zcompare_refl; auto. + intros (H1, H2); apply sym_equal; change (b < a * c)%Z. + apply Zlt_le_trans with (1 := H1). + change 0%Z with (0 * c)%Z. + apply Zmult_le_compat_r; auto with zarith. + intros (H1, H2); apply sym_equal; change (b > a * c)%Z. + apply Zlt_gt. + apply Zlt_le_trans with (2 := H1). + change 0%Z with (0 * c)%Z. + apply Zmult_lt_compat_r; auto with zarith. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + rewrite Zcompare_refl; auto. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare BigZ.zero x2); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH2; rewrite <- HH2; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare x1 BigZ.zero)%bigZ; case BigZ.compare; + auto; rewrite BigZ.spec_0. + intros HH2; rewrite <- HH2; rewrite Zcompare_refl; auto. + set (a := BigZ.to_Z x1); set (b := BigZ.to_Z x2); + set (c1 := BigN.to_Z y1); set (c2 := BigN.to_Z y2). + fold c1 in HH; fold c2 in HH1. + assert (F1: (0 < c1)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos y1)); fold c1; auto. + intros H1; case HH; rewrite <- H1; auto. + assert (F2: (0 < c2)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos y2)); fold c2; auto. + intros H1; case HH1; rewrite <- H1; auto. + repeat rewrite Z2P_correct; auto. + generalize (BigZ.spec_cmp_sign x1 x2); case BigZ.cmp_sign. + intros _; generalize (BigZ.spec_compare (x1 * BigZ.Pos y2)%bigZ + (x2 * BigZ.Pos y1)%bigZ); + case BigZ.compare; rewrite BigZ.spec_mul; simpl; fold a b c1 c2; auto. + rewrite BigZ.spec_mul; simpl; fold a b c1; intros HH2; rewrite HH2; + rewrite Zcompare_refl; auto. + rewrite BigZ.spec_mul; simpl; auto. + rewrite BigZ.spec_mul; simpl; auto. + fold a b; intros (H1, H2); apply sym_equal; change (a * c2 < b * c1)%Z. + apply Zlt_le_trans with 0%Z. + change 0%Z with (0 * c2)%Z. + apply Zmult_lt_compat_r; auto with zarith. + apply Zmult_le_0_compat; auto with zarith. + fold a b; intros (H1, H2); apply sym_equal; change (a * c2 > b * c1)%Z. + apply Zlt_gt; apply Zlt_le_trans with 0%Z. + change 0%Z with (0 * c1)%Z. + apply Zmult_lt_compat_r; auto with zarith. + apply Zmult_le_0_compat; auto with zarith. + Qed. + + + Definition do_norm_n n := + match n with + | BigN.N0 _ => false + | BigN.N1 _ => false + | BigN.N2 _ => false + | BigN.N3 _ => false + | BigN.N4 _ => false + | BigN.N5 _ => false + | BigN.N6 _ => false + | _ => true + end. + + Definition do_norm_z z := + match z with + | BigZ.Pos n => do_norm_n n + | BigZ.Neg n => do_norm_n n + end. + +(* Je pense que cette fonction normalise bien ... *) + Definition norm n d: t := + if andb (do_norm_z n) (do_norm_n d) then + let gcd := BigN.gcd (BigZ.to_N n) d in + match BigN.compare BigN.one gcd with + | Lt => + let n := BigZ.div n (BigZ.Pos gcd) in + let d := BigN.div d gcd in + match BigN.compare d BigN.one with + | Gt => Qq n d + | Eq => Qz n + | Lt => zero + end + | Eq => Qq n d + | Gt => zero (* gcd = 0 => both numbers are 0 *) + end + else Qq n d. + + Theorem spec_norm: forall n q, + ([norm n q] == [Qq n q])%Q. + intros p q; unfold norm. + case do_norm_z; simpl andb. + 2: apply Qeq_refl. + case do_norm_n. + 2: apply Qeq_refl. + assert (Hp := BigN.spec_pos (BigZ.to_N p)). + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; auto; rewrite BigN.spec_1; rewrite BigN.spec_gcd; intros H1. + apply Qeq_refl. + generalize (BigN.spec_pos (q / BigN.gcd (BigZ.to_N p) q)%bigN). + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; auto; rewrite BigN.spec_1; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; intros H2 HH. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H3; simpl; + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + auto with zarith. + generalize H2; rewrite H3; + rewrite Zdiv_0_l; auto with zarith. + generalize H1 H2 H3 (BigN.spec_pos q); clear H1 H2 H3. + rewrite spec_to_N. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 H4; rewrite Z2P_correct; auto with zarith. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H3; simpl. + case H3. + generalize H1 H2 H3 HH; clear H1 H2 H3 HH. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 HH. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto with zarith. + case (Zle_lt_or_eq _ _ HH); auto with zarith. + intros HH1; rewrite <- HH1; ring. + generalize (Zgcd_is_gcd a b); intros HH1; inversion HH1; auto. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; intros H3. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H4. + case H3; rewrite H4; rewrite Zdiv_0_l; auto with zarith. + simpl. + assert (FF := BigN.spec_pos q). + rewrite Z2P_correct; auto with zarith. + rewrite <- BigN.spec_gcd; rewrite <- BigN.spec_div; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; auto with zarith. + simpl; rewrite BigZ.spec_div; simpl. + rewrite BigN.spec_gcd; auto with zarith. + generalize H1 H2 H3 H4 HH FF; clear H1 H2 H3 H4 HH FF. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 H4 HH FF. + rewrite spec_to_N; fold a. + rewrite Zgcd_div_swap; auto with zarith. + rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_gcd; auto with zarith. + case (Zle_lt_or_eq _ _ + (BigN.spec_pos (BigN.gcd (BigZ.to_N p) q))); + rewrite BigN.spec_gcd; auto with zarith. + intros; apply False_ind; auto with zarith. + intros HH2; assert (FF1 := Zgcd_inv_0_l _ _ (sym_equal HH2)). + assert (FF2 := Zgcd_inv_0_l _ _ (sym_equal HH2)). + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H2; simpl. + rewrite spec_to_N. + rewrite FF2; ring. + Qed. + + Definition add (x y: t): t := + match x with + | Qz zx => + match y with + | Qz zy => Qz (BigZ.add zx zy) + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else Qq (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + end + | Qq nx dx => + if BigN.eq_bool dx BigN.zero then y + else match y with + | Qz zy => Qq (BigZ.add nx (BigZ.mul zy (BigZ.Pos dx))) dx + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else + if BigN.eq_bool dx dy then + let n := BigZ.add nx ny in + Qq n dx + else + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + Qq n d + end + end. + + + + Theorem spec_add x y: + ([add x y] == [x] + [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qplus; simpl. + rewrite BigZ.spec_add; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + assert (F1:= BigN.spec_pos dy). + rewrite Zmult_1_r; red; simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH; simpl; try ring. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH1; simpl; try ring. + case HH; auto. + rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul; simpl; auto. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH; simpl; try ring. + rewrite Zmult_1_r; apply Qeq_refl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH1; simpl; try ring. + case HH; auto. + rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul; simpl; auto. + rewrite Zmult_1_r; rewrite Pmult_1_r. + apply Qeq_refl. + assert (F1:= BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + apply Qeq_refl. + case HH2; auto. + simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + case HH2; auto. + case HH1; auto. + rewrite Zmult_1_r; apply Qeq_refl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + simpl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + case HH; auto. + rewrite Zmult_1_r; rewrite Zplus_0_r; rewrite Pmult_1_r. + apply Qeq_refl. + simpl. + generalize (BigN.spec_eq_bool (dx * dy)%bigN BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_mul; + rewrite BigN.spec_0; intros HH2. + (case (Zmult_integral _ _ HH2); intros HH3); + [case HH| case HH1]; auto. + generalize (BigN.spec_eq_bool dx dy); + case BigN.eq_bool; intros HH3. + rewrite <- HH3. + assert (Fx: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + red; simpl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH4. + case HH; auto. + simpl; rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; repeat rewrite BigZ.spec_mul; simpl. + ring. + assert (Fx: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (Fy: (0 < BigN.to_Z dy)%Z). + generalize (BigN.spec_pos dy); auto with zarith. + red; simpl; rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_mul; + rewrite BigN.spec_0; intros H3; simpl. + absurd (0 < 0)%Z; auto with zarith. + rewrite BigZ.spec_add; repeat rewrite BigZ.spec_mul; simpl. + repeat rewrite Z2P_correct; auto with zarith. + apply Zmult_lt_0_compat; auto. + Qed. + + Theorem spec_addc x y: + [[add x y]] = [[x]] + [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition add_norm (x y: t): t := + match x with + | Qz zx => + match y with + | Qz zy => Qz (BigZ.add zx zy) + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else + norm (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + end + | Qq nx dx => + if BigN.eq_bool dx BigN.zero then y + else match y with + | Qz zy => norm (BigZ.add nx (BigZ.mul zy (BigZ.Pos dx))) dx + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else + if BigN.eq_bool dx dy then + let n := BigZ.add nx ny in + norm n dx + else + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + norm n d + end + end. + + Theorem spec_add_norm x y: + ([add_norm x y] == [x] + [y])%Q. + intros x y; rewrite <- spec_add; auto. + case x; case y; clear x y; unfold add_norm, add. + intros; apply Qeq_refl. + intros p1 n p2. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + simpl. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + apply Qeq_refl. + apply Qeq_refl. + intros p1 p2 n. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + apply Qeq_refl. + intros p1 q1 p2 q2. + generalize (BigN.spec_eq_bool q2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + apply Qeq_refl. + generalize (BigN.spec_eq_bool q1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; intros HH3; + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end; apply Qeq_refl. + Qed. + + Theorem spec_add_normc x y: + [[add_norm x y]] = [[x]] + [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add_norm; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition sub x y := add x (opp y). + + Theorem spec_sub x y: + ([sub x y] == [x] - [y])%Q. + intros x y; unfold sub; rewrite spec_add; auto. + rewrite spec_opp; ring. + Qed. + + Theorem spec_subc x y: [[sub x y]] = [[x]] - [[y]]. + intros x y; unfold sub; rewrite spec_addc; auto. + rewrite spec_oppc; ring. + Qed. + + Definition sub_norm x y := add_norm x (opp y). + + Theorem spec_sub_norm x y: + ([sub_norm x y] == [x] - [y])%Q. + intros x y; unfold sub_norm; rewrite spec_add_norm; auto. + rewrite spec_opp; ring. + Qed. + + Theorem spec_sub_normc x y: + [[sub_norm x y]] = [[x]] - [[y]]. + intros x y; unfold sub_norm; rewrite spec_add_normc; auto. + rewrite spec_oppc; ring. + Qed. + + Definition mul (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => Qq (BigZ.mul zx ny) dy + | Qq nx dx, Qz zy => Qq (BigZ.mul nx zy) dx + | Qq nx dx, Qq ny dy => Qq (BigZ.mul nx ny) (BigN.mul dx dy) + end. + + Theorem spec_mul x y: ([mul x y] == [x] * [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qmult; simpl. + rewrite BigZ.spec_mul; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH1. + red; simpl; ring. + rewrite BigZ.spec_mul; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH1. + red; simpl; ring. + rewrite BigZ.spec_mul; rewrite Pmult_1_r. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_mul; + intros HH1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH2. + red; simpl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH3. + red; simpl; ring. + case (Zmult_integral _ _ HH1); intros HH. + case HH2; auto. + case HH3; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH2. + case HH1; rewrite HH2; ring. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH3. + case HH1; rewrite HH3; ring. + rewrite BigZ.spec_mul. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto. + apply Qeq_refl. + generalize (BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_pos dy); auto with zarith. + Qed. + + Theorem spec_mulc x y: + [[mul x y]] = [[x]] * [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition mul_norm (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => mul (Qz ny) (norm zx dy) + | Qq nx dx, Qz zy => mul (Qz nx) (norm zy dx) + | Qq nx dx, Qq ny dy => mul (norm nx dy) (norm ny dx) + end. + + Theorem spec_mul_norm x y: + ([mul_norm x y] == [x] * [y])%Q. + intros x y; rewrite <- spec_mul; auto. + unfold mul_norm; case x; case y; clear x y. + intros; apply Qeq_refl. + intros p1 n p2. + repeat rewrite spec_mul. + match goal with |- ?Z == _ => + match Z with context id [norm ?X ?Y] => + let y := context id [Qq X Y] in + apply Qeq_trans with y; [repeat apply Qmult_comp; + repeat apply Qplus_comp; repeat apply Qeq_refl; + apply spec_norm | idtac] + end + end. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH; simpl; ring. + intros p1 p2 n. + repeat rewrite spec_mul. + match goal with |- ?Z == _ => + match Z with context id [norm ?X ?Y] => + let y := context id [Qq X Y] in + apply Qeq_trans with y; [repeat apply Qmult_comp; + repeat apply Qplus_comp; repeat apply Qeq_refl; + apply spec_norm | idtac] + end + end. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH; simpl; try ring. + rewrite Pmult_1_r; auto. + intros p1 n1 p2 n2. + repeat rewrite spec_mul. + repeat match goal with |- ?Z == _ => + match Z with context id [norm ?X ?Y] => + let y := context id [Qq X Y] in + apply Qeq_trans with y; [repeat apply Qmult_comp; + repeat apply Qplus_comp; repeat apply Qeq_refl; + apply spec_norm | idtac] + end + end. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; try ring. + repeat rewrite Zpos_mult_morphism; ring. + Qed. + + Theorem spec_mul_normc x y: + [[mul_norm x y]] = [[x]] * [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul_norm; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition inv (x: t): t := + match x with + | Qz (BigZ.Pos n) => Qq BigZ.one n + | Qz (BigZ.Neg n) => Qq BigZ.minus_one n + | Qq (BigZ.Pos n) d => Qq (BigZ.Pos d) n + | Qq (BigZ.Neg n) d => Qq (BigZ.Neg d) n + end. + + + Theorem spec_inv x: + ([inv x] == /[x])%Q. + intros [ [x | x] | [nx | nx] dx]; unfold inv, Qinv; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + rewrite H1; apply Qeq_refl. + generalize H1 (BigN.spec_pos x); case (BigN.to_Z x); auto. + intros HH; case HH; auto. + intros; red; simpl; auto. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + rewrite H1; apply Qeq_refl. + generalize H1 (BigN.spec_pos x); case (BigN.to_Z x); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl; auto. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + apply Qeq_refl. + rewrite H1; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + rewrite H2; red; simpl; auto. + generalize H1 (BigN.spec_pos nx); case (BigN.to_Z nx); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + apply Qeq_refl. + rewrite H1; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + rewrite H2; red; simpl; auto. + generalize H1 (BigN.spec_pos nx); case (BigN.to_Z nx); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl. + assert (tmp: forall x, Zneg x = Zopp (Zpos x)); auto. + rewrite tmp. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + ring. + generalize (BigN.spec_pos dx); auto with zarith. + intros p _ HH; case HH; auto. + Qed. + + Theorem spec_invc x: + [[inv x]] = /[[x]]. + intros x; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv; auto. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition inv_norm (x: t): t := + match x with + | Qz (BigZ.Pos n) => + if BigN.eq_bool n BigN.zero then zero else Qq BigZ.one n + | Qz (BigZ.Neg n) => + if BigN.eq_bool n BigN.zero then zero else Qq BigZ.minus_one n + | Qq (BigZ.Pos n) d => + if BigN.eq_bool n BigN.zero then zero else Qq (BigZ.Pos d) n + | Qq (BigZ.Neg n) d => + if BigN.eq_bool n BigN.zero then zero else Qq (BigZ.Neg d) n + end. + + Theorem spec_inv_norm x: ([inv_norm x] == /[x])%Q. + intros x; rewrite <- spec_inv; generalize x; clear x. + intros [ [x | x] | [nx | nx] dx]; unfold inv_norm, inv; + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; try apply Qeq_refl; + red; simpl; + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; auto; + case H2; auto. + Qed. + + Theorem spec_inv_normc x: + [[inv_norm x]] = /[[x]]. + intros x; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv_norm; auto. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + + Definition div x y := mul x (inv y). + + Theorem spec_div x y: ([div x y] == [x] / [y])%Q. + intros x y; unfold div; rewrite spec_mul; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Theorem spec_divc x y: [[div x y]] = [[x]] / [[y]]. + intros x y; unfold div; rewrite spec_mulc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + Definition div_norm x y := mul_norm x (inv y). + + Theorem spec_div_norm x y: ([div_norm x y] == [x] / [y])%Q. + intros x y; unfold div_norm; rewrite spec_mul_norm; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Theorem spec_div_normc x y: [[div_norm x y]] = [[x]] / [[y]]. + intros x y; unfold div_norm; rewrite spec_mul_normc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + + Definition square (x: t): t := + match x with + | Qz zx => Qz (BigZ.square zx) + | Qq nx dx => Qq (BigZ.square nx) (BigN.square dx) + end. + + + Theorem spec_square x: ([square x] == [x] ^ 2)%Q. + intros [ x | nx dx]; unfold square. + red; simpl; rewrite BigZ.spec_square; auto with zarith. + simpl Qpower. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H. + red; simpl. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_square; + intros H1. + case H1; rewrite H; auto. + red; simpl. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_square; + intros H1. + case H; case (Zmult_integral _ _ H1); auto. + simpl. + rewrite BigZ.spec_square. + rewrite Zpos_mult_morphism. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto. + generalize (BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_pos dx); auto with zarith. + Qed. + + Theorem spec_squarec x: [[square x]] = [[x]]^2. + intros x; unfold to_Qc. + apply trans_equal with (!! ([x]^2)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_square; auto. + simpl Qcpower. + replace (!! [x] * 1) with (!![x]); try ring. + simpl. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition power_pos (x: t) p: t := + match x with + | Qz zx => Qz (BigZ.power_pos zx p) + | Qq nx dx => Qq (BigZ.power_pos nx p) (BigN.power_pos dx p) + end. + + Theorem spec_power_pos x p: ([power_pos x p] == [x] ^ Zpos p)%Q. + Proof. + intros [x | nx dx] p; unfold power_pos. + unfold power_pos; red; simpl. + generalize (Qpower_decomp p (BigZ.to_Z x) 1). + unfold Qeq; simpl. + rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Zmult_1_r. + intros H; rewrite H. + rewrite BigZ.spec_power_pos; simpl; ring. + simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_power_pos; intros H1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H2. + elim p; simpl. + intros; red; simpl; auto. + intros p1 Hp1; rewrite <- Hp1; red; simpl; auto. + apply Qeq_refl. + case H2; generalize H1. + elim p; simpl. + intros p1 Hrec. + change (xI p1) with (1 + (xO p1))%positive. + rewrite Zpower_pos_is_exp; rewrite Zpower_pos_1_r. + intros HH; case (Zmult_integral _ _ HH); auto. + rewrite <- Pplus_diag. + rewrite Zpower_pos_is_exp. + intros HH1; case (Zmult_integral _ _ HH1); auto. + intros p1 Hrec. + rewrite <- Pplus_diag. + rewrite Zpower_pos_is_exp. + intros HH1; case (Zmult_integral _ _ HH1); auto. + rewrite Zpower_pos_1_r; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H2. + case H1; rewrite H2; auto. + simpl; rewrite Zpower_pos_0_l; auto. + assert (F1: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (F2: (0 < BigN.to_Z dx ^ ' p)%Z). + unfold Zpower; apply Zpower_pos_pos; auto. + unfold power_pos; red; simpl. + generalize (Qpower_decomp p (BigZ.to_Z nx) + (Z2P (BigN.to_Z dx))). + unfold Qeq; simpl. + repeat rewrite Z2P_correct; auto. + unfold Qeq; simpl; intros HH. + rewrite HH. + rewrite BigZ.spec_power_pos; simpl; ring. + Qed. + + Theorem spec_power_posc x p: + [[power_pos x p]] = [[x]] ^ nat_of_P p. + intros x p; unfold to_Qc. + apply trans_equal with (!! ([x]^Zpos p)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_power_pos; auto. + pattern p; apply Pind; clear p. + simpl; ring. + intros p Hrec. + rewrite nat_of_P_succ_morphism; simpl Qcpower. + rewrite <- Hrec. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; + unfold this. + apply Qred_complete. + assert (F: [x] ^ ' Psucc p == [x] * [x] ^ ' p). + simpl; case x; simpl; clear x Hrec. + intros x; simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + intros nx dx. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + unfold Qpower_positive. + assert (tmp: forall p, pow_pos Qmult 0%Q p = 0%Q). + intros p1; elim p1; simpl; auto; clear p1. + intros p1 Hp1; rewrite Hp1; auto. + intros p1 Hp1; rewrite Hp1; auto. + repeat rewrite tmp; intros; red; simpl; auto. + intros H1. + assert (F1: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + repeat rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto. + 2: apply Zpower_pos_pos; auto. + 2: apply Zpower_pos_pos; auto. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + rewrite F. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + +End Qbi. diff --git a/theories/Numbers/Rational/BigQ/QifMake.v b/theories/Numbers/Rational/BigQ/QifMake.v new file mode 100644 index 00000000..1d8ecc94 --- /dev/null +++ b/theories/Numbers/Rational/BigQ/QifMake.v @@ -0,0 +1,979 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: QifMake.v 11027 2008-06-01 13:28:59Z letouzey $ i*) + +Require Import Bool. +Require Import ZArith. +Require Import Znumtheory. +Require Import BigNumPrelude. +Require Import Arith. +Require Export BigN. +Require Export BigZ. +Require Import QArith. +Require Import Qcanon. +Require Import Qpower. +Require Import QMake_base. + +Module Qif. + + Import BinInt. + Open Local Scope Q_scope. + Open Local Scope Qc_scope. + + (** The notation of a rational number is either an integer x, + interpreted as itself or a pair (x,y) of an integer x and a naturel + number y interpreted as x/y. The pairs (x,0) and (0,y) are all + interpreted as 0. *) + + Definition t := q_type. + + Definition zero: t := Qz BigZ.zero. + Definition one: t := Qz BigZ.one. + Definition minus_one: t := Qz BigZ.minus_one. + + Definition of_Z x: t := Qz (BigZ.of_Z x). + + Definition of_Q q: t := + match q with x # y => + Qq (BigZ.of_Z x) (BigN.of_N (Npos y)) + end. + + Definition of_Qc q := of_Q (this q). + + Definition to_Q (q: t) := + match q with + Qz x => BigZ.to_Z x # 1 + |Qq x y => if BigN.eq_bool y BigN.zero then 0%Q + else BigZ.to_Z x # Z2P (BigN.to_Z y) + end. + + Definition to_Qc q := !!(to_Q q). + + Notation "[[ x ]]" := (to_Qc x). + + Notation "[ x ]" := (to_Q x). + + Theorem spec_to_Q: forall q: Q, [of_Q q] = q. + intros (x,y); simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigN.spec_of_pos; intros HH; discriminate HH. + rewrite BigZ.spec_of_Z; simpl. + rewrite (BigN.spec_of_pos); auto. + Qed. + + Theorem spec_to_Qc: forall q, [[of_Qc q]] = q. + intros (x, Hx); unfold of_Qc, to_Qc; simpl. + apply Qc_decomp; simpl. + intros; rewrite spec_to_Q; auto. + Qed. + + Definition opp (x: t): t := + match x with + | Qz zx => Qz (BigZ.opp zx) + | Qq nx dx => Qq (BigZ.opp nx) dx + end. + + Theorem spec_opp: forall q, ([opp q] = -[q])%Q. + intros [z | x y]; simpl. + rewrite BigZ.spec_opp; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigZ.spec_opp; auto. + Qed. + + Theorem spec_oppc: forall q, [[opp q]] = -[[q]]. + intros q; unfold Qcopp, to_Qc, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + rewrite spec_opp. + rewrite <- Qred_opp. + rewrite Qred_involutive; auto. + Qed. + + Definition compare (x y: t) := + match x, y with + | Qz zx, Qz zy => BigZ.compare zx zy + | Qz zx, Qq ny dy => + if BigN.eq_bool dy BigN.zero then BigZ.compare zx BigZ.zero + else BigZ.compare (BigZ.mul zx (BigZ.Pos dy)) ny + | Qq nx dx, Qz zy => + if BigN.eq_bool dx BigN.zero then BigZ.compare BigZ.zero zy + else BigZ.compare nx (BigZ.mul zy (BigZ.Pos dx)) + | Qq nx dx, Qq ny dy => + match BigN.eq_bool dx BigN.zero, BigN.eq_bool dy BigN.zero with + | true, true => Eq + | true, false => BigZ.compare BigZ.zero ny + | false, true => BigZ.compare nx BigZ.zero + | false, false => BigZ.compare (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) + end + end. + + Theorem spec_compare: forall q1 q2, + compare q1 q2 = ([q1] ?= [q2])%Q. + intros [z1 | x1 y1] [z2 | x2 y2]; + unfold Qcompare, compare, to_Q, Qnum, Qden. + repeat rewrite Zmult_1_r. + generalize (BigZ.spec_compare z1 z2); case BigZ.compare; intros H; auto. + rewrite H; rewrite Zcompare_refl; auto. + rewrite Zmult_1_r. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + rewrite Zmult_1_r; generalize (BigZ.spec_compare z1 BigZ.zero); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH1; rewrite HH1; rewrite Zcompare_refl; auto. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (z1 * BigZ.Pos y2) x2)%bigZ; case BigZ.compare; + rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare BigZ.zero z2); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH1; rewrite <- HH1; rewrite Zcompare_refl; auto. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y1); auto with zarith. + rewrite Zmult_1_r. + generalize (BigZ.spec_compare x1 (z2 * BigZ.Pos y1))%bigZ; case BigZ.compare; + rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + rewrite Zcompare_refl; auto. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare BigZ.zero x2); + case BigZ.compare; auto. + rewrite BigZ.spec_0; intros HH2; rewrite <- HH2; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + rewrite Zmult_0_l; rewrite Zmult_1_r. + generalize (BigZ.spec_compare x1 BigZ.zero)%bigZ; case BigZ.compare; + auto; rewrite BigZ.spec_0. + intros HH2; rewrite <- HH2; rewrite Zcompare_refl; auto. + repeat rewrite Z2P_correct. + 2: generalize (BigN.spec_pos y1); auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (x1 * BigZ.Pos y2) + (x2 * BigZ.Pos y1))%bigZ; case BigZ.compare; + repeat rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + Qed. + + Definition do_norm_n n := + match n with + | BigN.N0 _ => false + | BigN.N1 _ => false + | BigN.N2 _ => false + | BigN.N3 _ => false + | BigN.N4 _ => false + | BigN.N5 _ => false + | BigN.N6 _ => false + | _ => true + end. + + Definition do_norm_z z := + match z with + | BigZ.Pos n => do_norm_n n + | BigZ.Neg n => do_norm_n n + end. + +(* Je pense que cette fonction normalise bien ... *) + Definition norm n d: t := + if andb (do_norm_z n) (do_norm_n d) then + let gcd := BigN.gcd (BigZ.to_N n) d in + match BigN.compare BigN.one gcd with + | Lt => + let n := BigZ.div n (BigZ.Pos gcd) in + let d := BigN.div d gcd in + match BigN.compare d BigN.one with + | Gt => Qq n d + | Eq => Qz n + | Lt => zero + end + | Eq => Qq n d + | Gt => zero (* gcd = 0 => both numbers are 0 *) + end + else Qq n d. + + Theorem spec_norm: forall n q, + ([norm n q] == [Qq n q])%Q. + intros p q; unfold norm. + case do_norm_z; simpl andb. + 2: apply Qeq_refl. + case do_norm_n. + 2: apply Qeq_refl. + assert (Hp := BigN.spec_pos (BigZ.to_N p)). + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; auto; rewrite BigN.spec_1; rewrite BigN.spec_gcd; intros H1. + apply Qeq_refl. + generalize (BigN.spec_pos (q / BigN.gcd (BigZ.to_N p) q)%bigN). + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; auto; rewrite BigN.spec_1; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; intros H2 HH. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H3; simpl; + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + auto with zarith. + generalize H2; rewrite H3; + rewrite Zdiv_0_l; auto with zarith. + generalize H1 H2 H3 (BigN.spec_pos q); clear H1 H2 H3. + rewrite spec_to_N. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 H4; rewrite Z2P_correct; auto with zarith. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H3; simpl. + case H3. + generalize H1 H2 H3 HH; clear H1 H2 H3 HH. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 HH. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto with zarith. + case (Zle_lt_or_eq _ _ HH); auto with zarith. + intros HH1; rewrite <- HH1; ring. + generalize (Zgcd_is_gcd a b); intros HH1; inversion HH1; auto. + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith; intros H3. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H4. + case H3; rewrite H4; rewrite Zdiv_0_l; auto with zarith. + simpl. + assert (FF := BigN.spec_pos q). + rewrite Z2P_correct; auto with zarith. + rewrite <- BigN.spec_gcd; rewrite <- BigN.spec_div; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; auto with zarith. + simpl; rewrite BigZ.spec_div; simpl. + rewrite BigN.spec_gcd; auto with zarith. + generalize H1 H2 H3 H4 HH FF; clear H1 H2 H3 H4 HH FF. + set (a := (BigN.to_Z (BigZ.to_N p))). + set (b := (BigN.to_Z q)). + intros H1 H2 H3 H4 HH FF. + rewrite spec_to_N; fold a. + rewrite Zgcd_div_swap; auto with zarith. + rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_div; + rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_gcd; auto with zarith. + case (Zle_lt_or_eq _ _ + (BigN.spec_pos (BigN.gcd (BigZ.to_N p) q))); + rewrite BigN.spec_gcd; auto with zarith. + intros; apply False_ind; auto with zarith. + intros HH2; assert (FF1 := Zgcd_inv_0_l _ _ (sym_equal HH2)). + assert (FF2 := Zgcd_inv_0_l _ _ (sym_equal HH2)). + red; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H2; simpl. + rewrite spec_to_N. + rewrite FF2; ring. + Qed. + + + Definition add (x y: t): t := + match x with + | Qz zx => + match y with + | Qz zy => Qz (BigZ.add zx zy) + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else Qq (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + end + | Qq nx dx => + if BigN.eq_bool dx BigN.zero then y + else match y with + | Qz zy => Qq (BigZ.add nx (BigZ.mul zy (BigZ.Pos dx))) dx + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + Qq n d + end + end. + + + Theorem spec_add x y: + ([add x y] == [x] + [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qplus; simpl. + rewrite BigZ.spec_add; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + assert (F1:= BigN.spec_pos dy). + rewrite Zmult_1_r; red; simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH; simpl; try ring. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH1; simpl; try ring. + case HH; auto. + rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul; simpl; auto. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH; simpl; try ring. + rewrite Zmult_1_r; apply Qeq_refl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; + rewrite BigN.spec_0; intros HH1; simpl; try ring. + case HH; auto. + rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul; simpl; auto. + rewrite Zmult_1_r; rewrite Pmult_1_r. + apply Qeq_refl. + assert (F1:= BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + apply Qeq_refl. + case HH2; auto. + simpl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + case HH2; auto. + case HH1; auto. + rewrite Zmult_1_r; apply Qeq_refl. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + simpl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + case HH; auto. + rewrite Zmult_1_r; rewrite Zplus_0_r; rewrite Pmult_1_r. + apply Qeq_refl. + simpl. + generalize (BigN.spec_eq_bool (dx * dy)%bigN BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_mul; + rewrite BigN.spec_0; intros HH2. + (case (Zmult_integral _ _ HH2); intros HH3); + [case HH| case HH1]; auto. + rewrite BigZ.spec_add; repeat rewrite BigZ.spec_mul; simpl. + assert (Fx: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (Fy: (0 < BigN.to_Z dy)%Z). + generalize (BigN.spec_pos dy); auto with zarith. + red; simpl; rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto with zarith. + apply Zmult_lt_0_compat; auto. + Qed. + + Theorem spec_addc x y: + [[add x y]] = [[x]] + [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition add_norm (x y: t): t := + match x with + | Qz zx => + match y with + | Qz zy => Qz (BigZ.add zx zy) + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else norm (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + end + | Qq nx dx => + if BigN.eq_bool dx BigN.zero then y + else match y with + | Qz zy => norm (BigZ.add nx (BigZ.mul zy (BigZ.Pos dx))) dx + | Qq ny dy => + if BigN.eq_bool dy BigN.zero then x + else + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + norm n d + end + end. + + Theorem spec_add_norm x y: + ([add_norm x y] == [x] + [y])%Q. + intros x y; rewrite <- spec_add; auto. + case x; case y; clear x y; unfold add_norm, add. + intros; apply Qeq_refl. + intros p1 n p2. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + simpl. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + apply Qeq_refl. + apply Qeq_refl. + intros p1 p2 n. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + apply Qeq_refl. + intros p1 q1 p2 q2. + generalize (BigN.spec_eq_bool q2 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH1. + apply Qeq_refl. + generalize (BigN.spec_eq_bool q1 BigN.zero); + case BigN.eq_bool; rewrite BigN.spec_0; intros HH2. + apply Qeq_refl. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + apply Qeq_refl. + Qed. + + Theorem spec_add_normc x y: + [[add_norm x y]] = [[x]] + [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add_norm; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition sub x y := add x (opp y). + + + Theorem spec_sub x y: + ([sub x y] == [x] - [y])%Q. + intros x y; unfold sub; rewrite spec_add; auto. + rewrite spec_opp; ring. + Qed. + + Theorem spec_subc x y: [[sub x y]] = [[x]] - [[y]]. + intros x y; unfold sub; rewrite spec_addc; auto. + rewrite spec_oppc; ring. + Qed. + + Definition sub_norm x y := add_norm x (opp y). + + Theorem spec_sub_norm x y: + ([sub_norm x y] == [x] - [y])%Q. + intros x y; unfold sub_norm; rewrite spec_add_norm; auto. + rewrite spec_opp; ring. + Qed. + + Theorem spec_sub_normc x y: + [[sub_norm x y]] = [[x]] - [[y]]. + intros x y; unfold sub_norm; rewrite spec_add_normc; auto. + rewrite spec_oppc; ring. + Qed. + + Definition mul (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => Qq (BigZ.mul zx ny) dy + | Qq nx dx, Qz zy => Qq (BigZ.mul nx zy) dx + | Qq nx dx, Qq ny dy => Qq (BigZ.mul nx ny) (BigN.mul dx dy) + end. + + + Theorem spec_mul x y: ([mul x y] == [x] * [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qmult; simpl. + rewrite BigZ.spec_mul; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH1. + red; simpl; ring. + rewrite BigZ.spec_mul; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH1. + red; simpl; ring. + rewrite BigZ.spec_mul; rewrite Pmult_1_r. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; rewrite BigN.spec_mul; + intros HH1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH2. + red; simpl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH3. + red; simpl; ring. + case (Zmult_integral _ _ HH1); intros HH. + case HH2; auto. + case HH3; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH2. + case HH1; rewrite HH2; ring. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros HH3. + case HH1; rewrite HH3; ring. + rewrite BigZ.spec_mul. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto. + apply Qeq_refl. + generalize (BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_pos dy); auto with zarith. + Qed. + + Theorem spec_mulc x y: + [[mul x y]] = [[x]] * [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + + Definition mul_norm (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => norm (BigZ.mul zx ny) dy + | Qq nx dx, Qz zy => norm (BigZ.mul nx zy) dx + | Qq nx dx, Qq ny dy => norm (BigZ.mul nx ny) (BigN.mul dx dy) + end. + + Theorem spec_mul_norm x y: + ([mul_norm x y] == [x] * [y])%Q. + intros x y; rewrite <- spec_mul; auto. + unfold mul_norm, mul; case x; case y; clear x y. + intros; apply Qeq_refl. + intros p1 n p2. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end; apply Qeq_refl. + intros p1 p2 n. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end; apply Qeq_refl. + intros p1 n1 p2 n2. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end; apply Qeq_refl. + Qed. + + Theorem spec_mul_normc x y: + [[mul_norm x y]] = [[x]] * [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul_norm; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + + + Definition inv (x: t): t := + match x with + | Qz (BigZ.Pos n) => Qq BigZ.one n + | Qz (BigZ.Neg n) => Qq BigZ.minus_one n + | Qq (BigZ.Pos n) d => Qq (BigZ.Pos d) n + | Qq (BigZ.Neg n) d => Qq (BigZ.Neg d) n + end. + + Theorem spec_inv x: + ([inv x] == /[x])%Q. + intros [ [x | x] | [nx | nx] dx]; unfold inv, Qinv; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + rewrite H1; apply Qeq_refl. + generalize H1 (BigN.spec_pos x); case (BigN.to_Z x); auto. + intros HH; case HH; auto. + intros; red; simpl; auto. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + rewrite H1; apply Qeq_refl. + generalize H1 (BigN.spec_pos x); case (BigN.to_Z x); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl; auto. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + apply Qeq_refl. + rewrite H1; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + rewrite H2; red; simpl; auto. + generalize H1 (BigN.spec_pos nx); case (BigN.to_Z nx); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + apply Qeq_refl. + rewrite H1; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H2; simpl; auto. + rewrite H2; red; simpl; auto. + generalize H1 (BigN.spec_pos nx); case (BigN.to_Z nx); simpl; + auto. + intros HH; case HH; auto. + intros; red; simpl. + assert (tmp: forall x, Zneg x = Zopp (Zpos x)); auto. + rewrite tmp. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + ring. + generalize (BigN.spec_pos dx); auto with zarith. + intros p _ HH; case HH; auto. + Qed. + + Theorem spec_invc x: + [[inv x]] = /[[x]]. + intros x; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv; auto. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + +Definition inv_norm (x: t): t := + match x with + | Qz (BigZ.Pos n) => + match BigN.compare n BigN.one with + Gt => Qq BigZ.one n + | _ => x + end + | Qz (BigZ.Neg n) => + match BigN.compare n BigN.one with + Gt => Qq BigZ.minus_one n + | _ => x + end + | Qq (BigZ.Pos n) d => + match BigN.compare n BigN.one with + Gt => Qq (BigZ.Pos d) n + | Eq => Qz (BigZ.Pos d) + | Lt => Qz (BigZ.zero) + end + | Qq (BigZ.Neg n) d => + match BigN.compare n BigN.one with + Gt => Qq (BigZ.Neg d) n + | Eq => Qz (BigZ.Neg d) + | Lt => Qz (BigZ.zero) + end + end. + + Theorem spec_inv_norm x: ([inv_norm x] == /[x])%Q. + intros [ [x | x] | [nx | nx] dx]; unfold inv_norm, Qinv. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H. + simpl; rewrite H; apply Qeq_refl. + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); simpl. + generalize H; case BigN.to_Z. + intros _ HH; discriminate HH. + intros p; case p; auto. + intros p1 HH; discriminate HH. + intros p1 HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; discriminate HH. + intros HH; rewrite <- HH. + apply Qeq_refl. + generalize H; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1. + rewrite H1; intros HH; discriminate. + generalize H; case BigN.to_Z. + intros HH; discriminate HH. + intros; red; simpl; auto. + intros p HH; discriminate HH. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H. + simpl; rewrite H; apply Qeq_refl. + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); simpl. + generalize H; case BigN.to_Z. + intros _ HH; discriminate HH. + intros p; case p; auto. + intros p1 HH; discriminate HH. + intros p1 HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; discriminate HH. + intros HH; rewrite <- HH. + apply Qeq_refl. + generalize H; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1. + rewrite H1; intros HH; discriminate. + generalize H; case BigN.to_Z. + intros HH; discriminate HH. + intros; red; simpl; auto. + intros p HH; discriminate HH. + simpl Qnum. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; simpl. + case BigN.compare; red; simpl; auto. + rewrite H1; auto. + case BigN.eq_bool; auto. + simpl; rewrite H1; auto. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H2. + rewrite H2. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + red; simpl. + rewrite Zmult_1_r; rewrite Pmult_1_r; rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + generalize H2 (BigN.spec_pos nx); case (BigN.to_Z nx). + intros; apply Qeq_refl. + intros p; case p; clear p. + intros p HH; discriminate HH. + intros p HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + simpl; generalize H2; case (BigN.to_Z nx). + intros HH; discriminate HH. + intros p Hp. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H4. + rewrite H4 in H2; discriminate H2. + red; simpl. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + intros p HH; discriminate HH. + simpl Qnum. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1; simpl. + case BigN.compare; red; simpl; auto. + rewrite H1; auto. + case BigN.eq_bool; auto. + simpl; rewrite H1; auto. + match goal with |- context[BigN.compare ?X ?Y] => + generalize (BigN.spec_compare X Y); case BigN.compare + end; rewrite BigN.spec_1; intros H2. + rewrite H2. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + red; simpl. + assert (tmp: forall x, Zneg x = Zopp (Zpos x)); auto. + rewrite tmp. + rewrite Zmult_1_r; rewrite Pmult_1_r; rewrite Z2P_correct; auto. + generalize (BigN.spec_pos dx); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + generalize H2 (BigN.spec_pos nx); case (BigN.to_Z nx). + intros; apply Qeq_refl. + intros p; case p; clear p. + intros p HH; discriminate HH. + intros p HH; discriminate HH. + intros HH; discriminate HH. + intros p _ HH; case HH; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H3. + case H1; auto. + simpl; generalize H2; case (BigN.to_Z nx). + intros HH; discriminate HH. + intros p Hp. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H4. + rewrite H4 in H2; discriminate H2. + red; simpl. + assert (tmp: forall x, Zneg x = Zopp (Zpos x)); auto. + rewrite tmp. + rewrite Zpos_mult_morphism. + rewrite Z2P_correct; auto. + ring. + generalize (BigN.spec_pos dx); auto with zarith. + intros p HH; discriminate HH. + Qed. + + Theorem spec_inv_normc x: + [[inv_norm x]] = /[[x]]. + intros x; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv_norm; auto. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + + Definition div x y := mul x (inv y). + + Theorem spec_div x y: ([div x y] == [x] / [y])%Q. + intros x y; unfold div; rewrite spec_mul; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Theorem spec_divc x y: [[div x y]] = [[x]] / [[y]]. + intros x y; unfold div; rewrite spec_mulc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + Definition div_norm x y := mul_norm x (inv y). + + Theorem spec_div_norm x y: ([div_norm x y] == [x] / [y])%Q. + intros x y; unfold div_norm; rewrite spec_mul_norm; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Theorem spec_div_normc x y: [[div_norm x y]] = [[x]] / [[y]]. + intros x y; unfold div_norm; rewrite spec_mul_normc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + + Definition square (x: t): t := + match x with + | Qz zx => Qz (BigZ.square zx) + | Qq nx dx => Qq (BigZ.square nx) (BigN.square dx) + end. + + Theorem spec_square x: ([square x] == [x] ^ 2)%Q. + intros [ x | nx dx]; unfold square. + red; simpl; rewrite BigZ.spec_square; auto with zarith. + simpl Qpower. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; intros H. + red; simpl. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_square; + intros H1. + case H1; rewrite H; auto. + red; simpl. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_square; + intros H1. + case H; case (Zmult_integral _ _ H1); auto. + simpl. + rewrite BigZ.spec_square. + rewrite Zpos_mult_morphism. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto. + generalize (BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_pos dx); auto with zarith. + Qed. + + Theorem spec_squarec x: [[square x]] = [[x]]^2. + intros x; unfold to_Qc. + apply trans_equal with (!! ([x]^2)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_square; auto. + simpl Qcpower. + replace (!! [x] * 1) with (!![x]); try ring. + simpl. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + +End Qif. diff --git a/theories/Numbers/Rational/BigQ/QpMake.v b/theories/Numbers/Rational/BigQ/QpMake.v new file mode 100644 index 00000000..ac3ca47a --- /dev/null +++ b/theories/Numbers/Rational/BigQ/QpMake.v @@ -0,0 +1,901 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: QpMake.v 11027 2008-06-01 13:28:59Z letouzey $ i*) + +Require Import Bool. +Require Import ZArith. +Require Import Znumtheory. +Require Import BigNumPrelude. +Require Import Arith. +Require Export BigN. +Require Export BigZ. +Require Import QArith. +Require Import Qcanon. +Require Import Qpower. +Require Import QMake_base. + +Notation Nspec_lt := BigNAxiomsMod.NZOrdAxiomsMod.spec_lt. +Notation Nspec_le := BigNAxiomsMod.NZOrdAxiomsMod.spec_le. + +Module Qp. + + (** The notation of a rational number is either an integer x, + interpreted as itself or a pair (x,y) of an integer x and a naturel + number y interpreted as x/(y+1). *) + + Definition t := q_type. + + Definition zero: t := Qz BigZ.zero. + Definition one: t := Qz BigZ.one. + Definition minus_one: t := Qz BigZ.minus_one. + + Definition of_Z x: t := Qz (BigZ.of_Z x). + + Definition d_to_Z d := BigZ.Pos (BigN.succ d). + + Definition of_Q q: t := + match q with x # y => + Qq (BigZ.of_Z x) (BigN.pred (BigN.of_N (Npos y))) + end. + + Definition of_Qc q := of_Q (this q). + + Definition to_Q (q: t) := + match q with + Qz x => BigZ.to_Z x # 1 + |Qq x y => BigZ.to_Z x # Z2P (BigN.to_Z (BigN.succ y)) + end. + + Definition to_Qc q := !!(to_Q q). + + Notation "[[ x ]]" := (to_Qc x). + + Notation "[ x ]" := (to_Q x). + + Theorem spec_to_Q: forall q: Q, [of_Q q] = q. + intros (x,y); simpl. + rewrite BigZ.spec_of_Z; auto. + rewrite BigN.spec_succ; simpl. simpl. + rewrite BigN.spec_pred; rewrite (BigN.spec_of_pos). + replace (Zpos y - 1 + 1)%Z with (Zpos y); auto; ring. + red; auto. + Qed. + + Theorem spec_to_Qc: forall q, [[of_Qc q]] = q. + intros (x, Hx); unfold of_Qc, to_Qc; simpl. + apply Qc_decomp; simpl. + intros; rewrite spec_to_Q; auto. + Qed. + + Definition opp (x: t): t := + match x with + | Qz zx => Qz (BigZ.opp zx) + | Qq nx dx => Qq (BigZ.opp nx) dx + end. + + + Theorem spec_opp: forall q, ([opp q] = -[q])%Q. + intros [z | x y]; simpl. + rewrite BigZ.spec_opp; auto. + rewrite BigZ.spec_opp; auto. + Qed. + + + Theorem spec_oppc: forall q, [[opp q]] = -[[q]]. + intros q; unfold Qcopp, to_Qc, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + rewrite spec_opp. + rewrite <- Qred_opp. + rewrite Qred_involutive; auto. + Qed. + + Definition compare (x y: t) := + match x, y with + | Qz zx, Qz zy => BigZ.compare zx zy + | Qz zx, Qq ny dy => BigZ.compare (BigZ.mul zx (d_to_Z dy)) ny + | Qq nx dy, Qz zy => BigZ.compare nx (BigZ.mul zy (d_to_Z dy)) + | Qq nx dx, Qq ny dy => + BigZ.compare (BigZ.mul nx (d_to_Z dy)) (BigZ.mul ny (d_to_Z dx)) + end. + + Theorem spec_compare: forall q1 q2, + compare q1 q2 = ([q1] ?= [q2])%Q. + intros [z1 | x1 y1] [z2 | x2 y2]; unfold Qcompare; simpl. + repeat rewrite Zmult_1_r. + generalize (BigZ.spec_compare z1 z2); case BigZ.compare; intros H; auto. + rewrite H; rewrite Zcompare_refl; auto. + rewrite Zmult_1_r. + rewrite BigN.spec_succ. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (z1 * d_to_Z y2) x2)%bigZ; case BigZ.compare; + intros H; rewrite <- H. + rewrite BigZ.spec_mul; unfold d_to_Z; simpl. + rewrite BigN.spec_succ. + rewrite Zcompare_refl; auto. + rewrite BigZ.spec_mul; unfold d_to_Z; simpl. + rewrite BigN.spec_succ; auto. + rewrite BigZ.spec_mul; unfold d_to_Z; simpl. + rewrite BigN.spec_succ; auto. + rewrite Zmult_1_r. + rewrite BigN.spec_succ. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y1); auto with zarith. + generalize (BigZ.spec_compare x1 (z2 * d_to_Z y1))%bigZ; case BigZ.compare; + rewrite BigZ.spec_mul; unfold d_to_Z; simpl; + rewrite BigN.spec_succ; intros H; auto. + rewrite H; rewrite Zcompare_refl; auto. + repeat rewrite BigN.spec_succ; auto. + repeat rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y1); auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (x1 * d_to_Z y2) + (x2 * d_to_Z y1))%bigZ; case BigZ.compare; + repeat rewrite BigZ.spec_mul; unfold d_to_Z; simpl; + repeat rewrite BigN.spec_succ; intros H; auto. + rewrite H; auto. + rewrite Zcompare_refl; auto. + Qed. + + + Theorem spec_comparec: forall q1 q2, + compare q1 q2 = ([[q1]] ?= [[q2]]). + unfold Qccompare, to_Qc. + intros q1 q2; rewrite spec_compare; simpl. + apply Qcompare_comp; apply Qeq_sym; apply Qred_correct. + Qed. + +(* Inv d > 0, Pour la forme normal unique on veut d > 1 *) + Definition norm n d: t := + if BigZ.eq_bool n BigZ.zero then zero + else + let gcd := BigN.gcd (BigZ.to_N n) d in + if BigN.eq_bool gcd BigN.one then Qq n (BigN.pred d) + else + let n := BigZ.div n (BigZ.Pos gcd) in + let d := BigN.div d gcd in + if BigN.eq_bool d BigN.one then Qz n + else Qq n (BigN.pred d). + + Theorem spec_norm: forall n q, + ((0 < BigN.to_Z q)%Z -> [norm n q] == [Qq n (BigN.pred q)])%Q. + intros p q; unfold norm; intros Hq. + assert (Hp := BigN.spec_pos (BigZ.to_N p)). + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; auto; rewrite BigZ.spec_0; intros H1. + red; simpl; rewrite H1; ring. + case (Zle_lt_or_eq _ _ Hp); clear Hp; intros Hp. + case (Zle_lt_or_eq _ _ + (Zgcd_is_pos (BigN.to_Z (BigZ.to_N p)) (BigN.to_Z q))); intros H4. + 2: generalize Hq; rewrite (Zgcd_inv_0_r _ _ (sym_equal H4)); auto with zarith. + 2: red; simpl; auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_1; intros H2. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_1. + red; simpl. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite Zmult_1_r. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + rewrite Z2P_correct; auto with zarith. + rewrite spec_to_N; intros; rewrite Zgcd_div_swap; auto. + rewrite H; ring. + intros H3. + red; simpl. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + assert (F: (0 < BigN.to_Z (q / BigN.gcd (BigZ.to_N p) q)%bigN)%Z). + rewrite BigN.spec_div; auto with zarith. + rewrite BigN.spec_gcd. + apply Zgcd_div_pos; auto. + rewrite BigN.spec_gcd; auto. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + rewrite Z2P_correct; auto. + rewrite Z2P_correct; auto. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite spec_to_N; apply Zgcd_div_swap; auto. + case H1; rewrite spec_to_N; rewrite <- Hp; ring. + Qed. + + Theorem spec_normc: forall n q, + (0 < BigN.to_Z q)%Z -> [[norm n q]] = [[Qq n (BigN.pred q)]]. + intros n q H; unfold to_Qc, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_norm; auto. + Qed. + + Definition add (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.add zx zy) + | Qz zx, Qq ny dy => Qq (BigZ.add (BigZ.mul zx (d_to_Z dy)) ny) dy + | Qq nx dx, Qz zy => Qq (BigZ.add nx (BigZ.mul zy (d_to_Z dx))) dx + | Qq nx dx, Qq ny dy => + let dx' := BigN.succ dx in + let dy' := BigN.succ dy in + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy')) (BigZ.mul ny (BigZ.Pos dx')) in + let d := BigN.pred (BigN.mul dx' dy') in + Qq n d + end. + + Theorem spec_d_to_Z: forall dy, + (BigZ.to_Z (d_to_Z dy) = BigN.to_Z dy + 1)%Z. + intros dy; unfold d_to_Z; simpl. + rewrite BigN.spec_succ; auto. + Qed. + + Theorem spec_succ_pos: forall p, + (0 < BigN.to_Z (BigN.succ p))%Z. + intros p; rewrite BigN.spec_succ; + generalize (BigN.spec_pos p); auto with zarith. + Qed. + + Theorem spec_add x y: ([add x y] == [x] + [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qplus; simpl. + rewrite BigZ.spec_add; repeat rewrite Zmult_1_r; auto. + apply Qeq_refl; auto. + assert (F1:= BigN.spec_pos dy). + rewrite Zmult_1_r. + simpl; rewrite Z2P_correct; rewrite BigN.spec_succ; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul. + rewrite spec_d_to_Z; apply Qeq_refl. + assert (F1:= BigN.spec_pos dx). + rewrite Zmult_1_r; rewrite Pmult_1_r. + simpl; rewrite Z2P_correct; rewrite BigN.spec_succ; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul. + rewrite spec_d_to_Z; apply Qeq_refl. + repeat rewrite BigN.spec_succ. + assert (Fx: (0 < BigN.to_Z dx + 1)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (Fy: (0 < BigN.to_Z dy + 1)%Z). + generalize (BigN.spec_pos dy); auto with zarith. + repeat rewrite BigN.spec_pred. + rewrite BigZ.spec_add; repeat rewrite BigN.spec_mul; + repeat rewrite BigN.spec_succ. + assert (tmp: forall x, (x-1+1 = x)%Z); [intros; ring | rewrite tmp; clear tmp]. + repeat rewrite Z2P_correct; auto. + repeat rewrite BigZ.spec_mul; simpl. + repeat rewrite BigN.spec_succ. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto; apply Qeq_refl. + rewrite BigN.spec_mul; repeat rewrite BigN.spec_succ; auto with zarith. + apply Zmult_lt_0_compat; auto. + Qed. + + Theorem spec_addc x y: [[add x y]] = [[x]] + [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition add_norm (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.add zx zy) + | Qz zx, Qq ny dy => + let d := BigN.succ dy in + norm (BigZ.add (BigZ.mul zx (BigZ.Pos d)) ny) d + | Qq nx dx, Qz zy => + let d := BigN.succ dx in + norm (BigZ.add (BigZ.mul zy (BigZ.Pos d)) nx) d + | Qq nx dx, Qq ny dy => + let dx' := BigN.succ dx in + let dy' := BigN.succ dy in + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy')) (BigZ.mul ny (BigZ.Pos dx')) in + let d := BigN.mul dx' dy' in + norm n d + end. + + Theorem spec_add_norm x y: ([add_norm x y] == [x] + [y])%Q. + intros x y; rewrite <- spec_add. + unfold add_norm, add; case x; case y. + intros; apply Qeq_refl. + intros p1 n p2. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X (BigN.pred Y)]); + [apply spec_norm | idtac] + end. + rewrite BigN.spec_succ; generalize (BigN.spec_pos n); auto with zarith. + simpl. + repeat rewrite BigZ.spec_add. + repeat rewrite BigZ.spec_mul; simpl. + rewrite BigN.succ_pred; try apply Qeq_refl; apply lt_0_succ. + intros p1 n p2. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X (BigN.pred Y)]); + [apply spec_norm | idtac] + end. + rewrite BigN.spec_succ; generalize (BigN.spec_pos p2); auto with zarith. + simpl. + repeat rewrite BigZ.spec_add. + repeat rewrite BigZ.spec_mul; simpl. + rewrite BinInt.Zplus_comm. + rewrite BigN.succ_pred; try apply Qeq_refl; apply lt_0_succ. + intros p1 q1 p2 q2. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X (BigN.pred Y)]); + [apply spec_norm | idtac] + end; try apply Qeq_refl. + rewrite BigN.spec_mul. + apply Zmult_lt_0_compat; apply spec_succ_pos. + Qed. + + Theorem spec_add_normc x y: [[add_norm x y]] = [[x]] + [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add_norm. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition sub (x y: t): t := add x (opp y). + + Theorem spec_sub x y: ([sub x y] == [x] - [y])%Q. + intros x y; unfold sub; rewrite spec_add. + rewrite spec_opp; ring. + Qed. + + Theorem spec_subc x y: [[sub x y]] = [[x]] - [[y]]. + intros x y; unfold sub; rewrite spec_addc. + rewrite spec_oppc; ring. + Qed. + + Definition sub_norm x y := add_norm x (opp y). + + Theorem spec_sub_norm x y: ([sub_norm x y] == [x] - [y])%Q. + intros x y; unfold sub_norm; rewrite spec_add_norm. + rewrite spec_opp; ring. + Qed. + + Theorem spec_sub_normc x y: [[sub_norm x y]] = [[x]] - [[y]]. + intros x y; unfold sub_norm; rewrite spec_add_normc. + rewrite spec_oppc; ring. + Qed. + + + Definition mul (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => Qq (BigZ.mul zx ny) dy + | Qq nx dx, Qz zy => Qq (BigZ.mul nx zy) dx + | Qq nx dx, Qq ny dy => + Qq (BigZ.mul nx ny) (BigN.pred (BigN.mul (BigN.succ dx) (BigN.succ dy))) + end. + + Theorem spec_mul x y: ([mul x y] == [x] * [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qmult; simpl. + rewrite BigZ.spec_mul; repeat rewrite Zmult_1_r; auto. + apply Qeq_refl; auto. + rewrite BigZ.spec_mul; apply Qeq_refl. + rewrite BigZ.spec_mul; rewrite Pmult_1_r; auto. + apply Qeq_refl; auto. + assert (F1:= spec_succ_pos dx). + assert (F2:= spec_succ_pos dy). + rewrite BigN.succ_pred. + rewrite BigN.spec_mul; rewrite BigZ.spec_mul. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto; apply Qeq_refl. + rewrite Nspec_lt, BigN.spec_0, BigN.spec_mul; auto. + apply Zmult_lt_0_compat; apply spec_succ_pos. + Qed. + + Theorem spec_mulc x y: [[mul x y]] = [[x]] * [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition mul_norm (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => + if BigZ.eq_bool zx BigZ.zero then zero + else + let d := BigN.succ dy in + let gcd := BigN.gcd (BigZ.to_N zx) d in + if BigN.eq_bool gcd BigN.one then Qq (BigZ.mul zx ny) dy + else + let zx := BigZ.div zx (BigZ.Pos gcd) in + let d := BigN.div d gcd in + if BigN.eq_bool d BigN.one then Qz (BigZ.mul zx ny) + else Qq (BigZ.mul zx ny) (BigN.pred d) + | Qq nx dx, Qz zy => + if BigZ.eq_bool zy BigZ.zero then zero + else + let d := BigN.succ dx in + let gcd := BigN.gcd (BigZ.to_N zy) d in + if BigN.eq_bool gcd BigN.one then Qq (BigZ.mul zy nx) dx + else + let zy := BigZ.div zy (BigZ.Pos gcd) in + let d := BigN.div d gcd in + if BigN.eq_bool d BigN.one then Qz (BigZ.mul zy nx) + else Qq (BigZ.mul zy nx) (BigN.pred d) + | Qq nx dx, Qq ny dy => + norm (BigZ.mul nx ny) (BigN.mul (BigN.succ dx) (BigN.succ dy)) + end. + + Theorem spec_mul_norm x y: ([mul_norm x y] == [x] * [y])%Q. + intros x y; rewrite <- spec_mul. + unfold mul_norm, mul; case x; case y. + intros; apply Qeq_refl. + intros p1 n p2. + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; unfold zero, to_Q; repeat rewrite BigZ.spec_0; intros H. + rewrite BigZ.spec_mul; rewrite H; red; auto. + assert (F: (0 < BigN.to_Z (BigZ.to_N p2))%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos (BigZ.to_N p2))); auto. + intros H1; case H; rewrite spec_to_N; rewrite <- H1; ring. + assert (F1: (0 < BigN.to_Z (BigN.succ n))%Z). + rewrite BigN.spec_succ; generalize (BigN.spec_pos n); auto with zarith. + assert (F2: (0 < Zgcd (BigN.to_Z (BigZ.to_N p2)) (BigN.to_Z (BigN.succ n)))%Z). + case (Zle_lt_or_eq _ _ (Zgcd_is_pos (BigN.to_Z (BigZ.to_N p2)) + (BigN.to_Z (BigN.succ n)))); intros H3; auto. + generalize F; rewrite (Zgcd_inv_0_l _ _ (sym_equal H3)); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; intros H1. + intros; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; + auto with zarith. + intros H2. + red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite spec_to_N. + rewrite Zmult_1_r; repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p1)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + intros H2. + red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite (spec_to_N p2). + case (Zle_lt_or_eq _ _ + (BigN.spec_pos (BigN.succ n / + BigN.gcd (BigZ.to_N p2) + (BigN.succ n)))%bigN); intros F3. + rewrite BigN.succ_pred; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p1)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto; try ring. + rewrite Nspec_lt, BigN.spec_0; auto. + apply False_ind; generalize F1. + rewrite (Zdivide_Zdiv_eq + (Zgcd (BigN.to_Z (BigZ.to_N p2)) (BigN.to_Z (BigN.succ n))) + (BigN.to_Z (BigN.succ n))); auto. + generalize F3; rewrite BigN.spec_div; rewrite BigN.spec_gcd; + auto with zarith. + intros HH; rewrite <- HH; auto with zarith. + assert (FF:= Zgcd_is_gcd (BigN.to_Z (BigZ.to_N p2)) + (BigN.to_Z (BigN.succ n))); inversion FF; auto. + intros p1 p2 n. + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; unfold zero, to_Q; repeat rewrite BigZ.spec_0; intros H. + rewrite BigZ.spec_mul; rewrite H; red; simpl; ring. + assert (F: (0 < BigN.to_Z (BigZ.to_N p1))%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos (BigZ.to_N p1))); auto. + intros H1; case H; rewrite spec_to_N; rewrite <- H1; ring. + assert (F1: (0 < BigN.to_Z (BigN.succ n))%Z). + rewrite BigN.spec_succ; generalize (BigN.spec_pos n); auto with zarith. + assert (F2: (0 < Zgcd (BigN.to_Z (BigZ.to_N p1)) (BigN.to_Z (BigN.succ n)))%Z). + case (Zle_lt_or_eq _ _ (Zgcd_is_pos (BigN.to_Z (BigZ.to_N p1)) + (BigN.to_Z (BigN.succ n)))); intros H3; auto. + generalize F; rewrite (Zgcd_inv_0_l _ _ (sym_equal H3)); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; intros H1. + intros; repeat rewrite BigZ.spec_mul; rewrite Zmult_comm; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; + auto with zarith. + intros H2. + red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite spec_to_N. + rewrite Zmult_1_r; repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p2)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + intros H2. + red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite (spec_to_N p1). + case (Zle_lt_or_eq _ _ + (BigN.spec_pos (BigN.succ n / + BigN.gcd (BigZ.to_N p1) + (BigN.succ n)))%bigN); intros F3. + rewrite BigN.succ_pred; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p2)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto; try ring. + rewrite Nspec_lt, BigN.spec_0; auto. + apply False_ind; generalize F1. + rewrite (Zdivide_Zdiv_eq + (Zgcd (BigN.to_Z (BigZ.to_N p1)) (BigN.to_Z (BigN.succ n))) + (BigN.to_Z (BigN.succ n))); auto. + generalize F3; rewrite BigN.spec_div; rewrite BigN.spec_gcd; + auto with zarith. + intros HH; rewrite <- HH; auto with zarith. + assert (FF:= Zgcd_is_gcd (BigN.to_Z (BigZ.to_N p1)) + (BigN.to_Z (BigN.succ n))); inversion FF; auto. + intros p1 n1 p2 n2. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X (BigN.pred Y)]); + [apply spec_norm | idtac] + end; try apply Qeq_refl. + rewrite BigN.spec_mul. + apply Zmult_lt_0_compat; rewrite BigN.spec_succ; + generalize (BigN.spec_pos n1) (BigN.spec_pos n2); auto with zarith. + Qed. + + Theorem spec_mul_normc x y: [[mul_norm x y]] = [[x]] * [[y]]. + intros x y; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul_norm. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition inv (x: t): t := + match x with + | Qz (BigZ.Pos n) => + if BigN.eq_bool n BigN.zero then zero else Qq BigZ.one (BigN.pred n) + | Qz (BigZ.Neg n) => + if BigN.eq_bool n BigN.zero then zero else Qq BigZ.minus_one (BigN.pred n) + | Qq (BigZ.Pos n) d => + if BigN.eq_bool n BigN.zero then zero else Qq (BigZ.Pos (BigN.succ d)) (BigN.pred n) + | Qq (BigZ.Neg n) d => + if BigN.eq_bool n BigN.zero then zero else Qq (BigZ.Neg (BigN.succ d)) (BigN.pred n) + end. + + Theorem spec_inv x: ([inv x] == /[x])%Q. + intros [ [x | x] | [nx | nx] dx]; unfold inv. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z x)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); auto with zarith. + unfold to_Q; rewrite BigZ.spec_1. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + red; unfold Qinv; simpl. + generalize F; case BigN.to_Z; auto with zarith. + intros p Hp; discriminate Hp. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z x)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); auto with zarith. + red; unfold Qinv; simpl. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + generalize F; case BigN.to_Z; simpl; auto with zarith. + intros p Hp; discriminate Hp. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z nx)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos nx)); auto with zarith. + red; unfold Qinv; simpl. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + rewrite BigN.spec_succ; rewrite Z2P_correct; auto with zarith. + generalize F; case BigN.to_Z; auto with zarith. + intros p Hp; discriminate Hp. + generalize (BigN.spec_pos dx); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z nx)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos nx)); auto with zarith. + red; unfold Qinv; simpl. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + rewrite BigN.spec_succ; rewrite Z2P_correct; auto with zarith. + generalize F; case BigN.to_Z; auto with zarith. + simpl; intros. + match goal with |- (?X = Zneg ?Y)%Z => + replace (Zneg Y) with (-(Zpos Y))%Z; + try rewrite Z2P_correct; auto with zarith + end. + rewrite Zpos_mult_morphism; + rewrite Z2P_correct; auto with zarith; try ring. + generalize (BigN.spec_pos dx); auto with zarith. + intros p Hp; discriminate Hp. + generalize (BigN.spec_pos dx); auto with zarith. + Qed. + + Theorem spec_invc x: [[inv x]] = /[[x]]. + intros x; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + +Definition inv_norm x := + match x with + | Qz (BigZ.Pos n) => + if BigN.eq_bool n BigN.zero then zero else + if BigN.eq_bool n BigN.one then x else Qq BigZ.one (BigN.pred n) + | Qz (BigZ.Neg n) => + if BigN.eq_bool n BigN.zero then zero else + if BigN.eq_bool n BigN.one then x else Qq BigZ.minus_one (BigN.pred n) + | Qq (BigZ.Pos n) d => let d := BigN.succ d in + if BigN.eq_bool n BigN.zero then zero else + if BigN.eq_bool n BigN.one then Qz (BigZ.Pos d) + else Qq (BigZ.Pos d) (BigN.pred n) + | Qq (BigZ.Neg n) d => let d := BigN.succ d in + if BigN.eq_bool n BigN.zero then zero else + if BigN.eq_bool n BigN.one then Qz (BigZ.Neg d) + else Qq (BigZ.Neg d) (BigN.pred n) + end. + + Theorem spec_inv_norm x: ([inv_norm x] == /[x])%Q. + intros x; rewrite <- spec_inv. + (case x; clear x); [intros [x | x] | intros nx dx]; + unfold inv_norm, inv. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + apply Qeq_refl. + assert (F: (0 < BigN.to_Z x)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; intros H1. + red; simpl. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + rewrite Z2P_correct; try rewrite H1; auto with zarith. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + apply Qeq_refl. + assert (F: (0 < BigN.to_Z x)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; intros H1. + red; simpl. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + rewrite Z2P_correct; try rewrite H1; auto with zarith. + apply Qeq_refl. + case nx; clear nx; intros nx. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; intros H1. + red; simpl. + rewrite BigN.succ_pred; try rewrite H1; auto with zarith. + rewrite Nspec_lt, BigN.spec_0, H1; auto with zarith. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; intros H1. + red; simpl. + rewrite BigN.succ_pred; try rewrite H1; auto with zarith. + rewrite Nspec_lt, BigN.spec_0, H1; auto with zarith. + apply Qeq_refl. + Qed. + + + Definition div x y := mul x (inv y). + + Theorem spec_div x y: ([div x y] == [x] / [y])%Q. + intros x y; unfold div; rewrite spec_mul; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Theorem spec_divc x y: [[div x y]] = [[x]] / [[y]]. + intros x y; unfold div; rewrite spec_mulc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + Definition div_norm x y := mul_norm x (inv y). + + Theorem spec_div_norm x y: ([div_norm x y] == [x] / [y])%Q. + intros x y; unfold div_norm; rewrite spec_mul_norm; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + Qed. + + Theorem spec_div_normc x y: [[div_norm x y]] = [[x]] / [[y]]. + intros x y; unfold div_norm; rewrite spec_mul_normc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + Qed. + + + Definition square (x: t): t := + match x with + | Qz zx => Qz (BigZ.square zx) + | Qq nx dx => Qq (BigZ.square nx) (BigN.pred (BigN.square (BigN.succ dx))) + end. + + Theorem spec_square x: ([square x] == [x] ^ 2)%Q. + intros [ x | nx dx]; unfold square. + red; simpl; rewrite BigZ.spec_square; auto with zarith. + red; simpl; rewrite BigZ.spec_square; auto with zarith. + assert (F: (0 < BigN.to_Z (BigN.succ dx))%Z). + rewrite BigN.spec_succ; + case (Zle_lt_or_eq _ _ (BigN.spec_pos dx)); auto with zarith. + assert (F1 : (0 < BigN.to_Z (BigN.square (BigN.succ dx)))%Z). + rewrite BigN.spec_square; apply Zmult_lt_0_compat; + auto with zarith. + rewrite BigN.succ_pred by (rewrite Nspec_lt, BigN.spec_0; auto). + rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto with zarith. + repeat rewrite BigN.spec_succ; auto with zarith. + rewrite BigN.spec_square; auto with zarith. + repeat rewrite BigN.spec_succ; auto with zarith. + Qed. + + Theorem spec_squarec x: [[square x]] = [[x]]^2. + intros x; unfold to_Qc. + apply trans_equal with (!! ([x]^2)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_square. + simpl Qcpower. + replace (!! [x] * 1) with (!![x]); try ring. + simpl. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition power_pos (x: t) p: t := + match x with + | Qz zx => Qz (BigZ.power_pos zx p) + | Qq nx dx => Qq (BigZ.power_pos nx p) (BigN.pred (BigN.power_pos (BigN.succ dx) p)) + end. + + + Theorem spec_power_pos x p: ([power_pos x p] == [x] ^ Zpos p)%Q. + Proof. + intros [x | nx dx] p; unfold power_pos. + unfold power_pos; red; simpl. + generalize (Qpower_decomp p (BigZ.to_Z x) 1). + unfold Qeq; simpl. + rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Zmult_1_r. + intros H; rewrite H. + rewrite BigZ.spec_power_pos; simpl; ring. + assert (F1: (0 < BigN.to_Z (BigN.succ dx))%Z). + rewrite BigN.spec_succ; + generalize (BigN.spec_pos dx); auto with zarith. + assert (F2: (0 < BigN.to_Z (BigN.succ dx) ^ ' p)%Z). + unfold Zpower; apply Zpower_pos_pos; auto. + unfold power_pos; red; simpl. + rewrite BigN.succ_pred, BigN.spec_power_pos. + rewrite Z2P_correct; auto. + generalize (Qpower_decomp p (BigZ.to_Z nx) + (Z2P (BigN.to_Z (BigN.succ dx)))). + unfold Qeq; simpl. + repeat rewrite Z2P_correct; auto. + unfold Qeq; simpl; intros HH. + rewrite HH. + rewrite BigZ.spec_power_pos; simpl; ring. + rewrite Nspec_lt, BigN.spec_0, BigN.spec_power_pos; auto. + Qed. + + Theorem spec_power_posc x p: [[power_pos x p]] = [[x]] ^ nat_of_P p. + intros x p; unfold to_Qc. + apply trans_equal with (!! ([x]^Zpos p)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_power_pos. + pattern p; apply Pind; clear p. + simpl; ring. + intros p Hrec. + rewrite nat_of_P_succ_morphism; simpl Qcpower. + rewrite <- Hrec. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; + unfold this. + apply Qred_complete. + assert (F: [x] ^ ' Psucc p == [x] * [x] ^ ' p). + simpl; case x; simpl; clear x Hrec. + intros x; simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + intros nx dx; simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + assert (F1: (0 < BigN.to_Z (BigN.succ dx))%Z). + rewrite BigN.spec_succ; generalize (BigN.spec_pos dx); + auto with zarith. + repeat rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto. + 2: apply Zpower_pos_pos; auto. + 2: apply Zpower_pos_pos; auto. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + rewrite F. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + +End Qp. diff --git a/theories/Numbers/Rational/BigQ/QvMake.v b/theories/Numbers/Rational/BigQ/QvMake.v new file mode 100644 index 00000000..4523e241 --- /dev/null +++ b/theories/Numbers/Rational/BigQ/QvMake.v @@ -0,0 +1,1151 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +(*i $Id: QvMake.v 11027 2008-06-01 13:28:59Z letouzey $ i*) + +Require Import Bool. +Require Import ZArith. +Require Import Znumtheory. +Require Import BigNumPrelude. +Require Import Arith. +Require Export BigN. +Require Export BigZ. +Require Import QArith. +Require Import Qcanon. +Require Import Qpower. +Require Import QMake_base. + +Module Qv. + + Import BinInt Zorder. + Open Local Scope Q_scope. + Open Local Scope Qc_scope. + + (** The notation of a rational number is either an integer x, + interpreted as itself or a pair (x,y) of an integer x and a naturel + number y interpreted as x/y. All functions maintain the invariant + that y is never zero. *) + + Definition t := q_type. + + Definition zero: t := Qz BigZ.zero. + Definition one: t := Qz BigZ.one. + Definition minus_one: t := Qz BigZ.minus_one. + + Definition of_Z x: t := Qz (BigZ.of_Z x). + + Definition wf x := + match x with + | Qz _ => True + | Qq n d => if BigN.eq_bool d BigN.zero then False else True + end. + + Definition of_Q q: t := + match q with x # y => + Qq (BigZ.of_Z x) (BigN.of_N (Npos y)) + end. + + Definition of_Qc q := of_Q (this q). + + Definition to_Q (q: t) := + match q with + Qz x => BigZ.to_Z x # 1 + |Qq x y => BigZ.to_Z x # Z2P (BigN.to_Z y) + end. + + Definition to_Qc q := !!(to_Q q). + + Notation "[[ x ]]" := (to_Qc x). + + Notation "[ x ]" := (to_Q x). + + Theorem spec_to_Q: forall q: Q, [of_Q q] = q. + intros (x,y); simpl. + rewrite BigZ.spec_of_Z; simpl. + rewrite (BigN.spec_of_pos); auto. + Qed. + + Theorem spec_to_Qc: forall q, [[of_Qc q]] = q. + intros (x, Hx); unfold of_Qc, to_Qc; simpl. + apply Qc_decomp; simpl. + intros; rewrite spec_to_Q; auto. + Qed. + + Definition opp (x: t): t := + match x with + | Qz zx => Qz (BigZ.opp zx) + | Qq nx dx => Qq (BigZ.opp nx) dx + end. + + Theorem wf_opp: forall x, wf x -> wf (opp x). + intros [zx | nx dx]; unfold opp, wf; auto. + Qed. + + Theorem spec_opp: forall q, ([opp q] = -[q])%Q. + intros [z | x y]; simpl. + rewrite BigZ.spec_opp; auto. + rewrite BigZ.spec_opp; auto. + Qed. + + Theorem spec_oppc: forall q, [[opp q]] = -[[q]]. + intros q; unfold Qcopp, to_Qc, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + rewrite spec_opp. + rewrite <- Qred_opp. + rewrite Qred_involutive; auto. + Qed. + + (* Les fonctions doivent assurer que si leur arguments sont valides alors + le resultat est correct et valide (si c'est un Q) + *) + + Definition compare (x y: t) := + match x, y with + | Qz zx, Qz zy => BigZ.compare zx zy + | Qz zx, Qq ny dy => BigZ.compare (BigZ.mul zx (BigZ.Pos dy)) ny + | Qq nx dx, Qz zy => BigZ.compare nx (BigZ.mul zy (BigZ.Pos dx)) + | Qq nx dx, Qq ny dy => BigZ.compare (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) + end. + + Theorem spec_compare: forall q1 q2, wf q1 -> wf q2 -> + compare q1 q2 = ([q1] ?= [q2])%Q. + intros [z1 | x1 y1] [z2 | x2 y2]; + unfold Qcompare, compare, to_Q, Qnum, Qden, wf. + repeat rewrite Zmult_1_r. + generalize (BigZ.spec_compare z1 z2); case BigZ.compare; intros H; auto. + rewrite H; rewrite Zcompare_refl; auto. + rewrite Zmult_1_r. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool. + intros _ _ HH; case HH. + rewrite BigN.spec_0; intros HH _ _. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (z1 * BigZ.Pos y2) x2)%bigZ; case BigZ.compare; + rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH _ _. + rewrite Z2P_correct; auto with zarith. + 2: generalize (BigN.spec_pos y1); auto with zarith. + rewrite Zmult_1_r. + generalize (BigZ.spec_compare x1 (z2 * BigZ.Pos y1))%bigZ; case BigZ.compare; + rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + generalize (BigN.spec_eq_bool y1 BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH1. + generalize (BigN.spec_eq_bool y2 BigN.zero); + case BigN.eq_bool. + intros _ _ HH; case HH. + rewrite BigN.spec_0; intros HH2 _ _. + repeat rewrite Z2P_correct. + 2: generalize (BigN.spec_pos y1); auto with zarith. + 2: generalize (BigN.spec_pos y2); auto with zarith. + generalize (BigZ.spec_compare (x1 * BigZ.Pos y2) + (x2 * BigZ.Pos y1))%bigZ; case BigZ.compare; + repeat rewrite BigZ.spec_mul; simpl; intros H; apply sym_equal; auto. + rewrite H; rewrite Zcompare_refl; auto. + Qed. + + Theorem spec_comparec: forall q1 q2, wf q1 -> wf q2 -> + compare q1 q2 = ([[q1]] ?= [[q2]]). + unfold Qccompare, to_Qc. + intros q1 q2 Hq1 Hq2; rewrite spec_compare; simpl; auto. + apply Qcompare_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition norm n d: t := + if BigZ.eq_bool n BigZ.zero then zero + else + let gcd := BigN.gcd (BigZ.to_N n) d in + if BigN.eq_bool gcd BigN.one then Qq n d + else + let n := BigZ.div n (BigZ.Pos gcd) in + let d := BigN.div d gcd in + if BigN.eq_bool d BigN.one then Qz n + else Qq n d. + + Theorem wf_norm: forall n q, + (BigN.to_Z q <> 0)%Z -> wf (norm n q). + intros p q; unfold norm, wf; intros Hq. + assert (Hp := BigN.spec_pos (BigZ.to_N p)). + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; auto; rewrite BigZ.spec_0; intros H1. + simpl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_1. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + set (a := BigN.to_Z (BigZ.to_N p)). + set (b := (BigN.to_Z q)). + assert (F: (0 < Zgcd a b)%Z). + case (Zle_lt_or_eq _ _ (Zgcd_is_pos a b)); auto. + intros HH1; case Hq; apply (Zgcd_inv_0_r _ _ (sym_equal HH1)). + rewrite BigN.spec_div; rewrite BigN.spec_gcd; auto; fold a; fold b. + intros H; case Hq; fold b. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite H; auto with zarith. + assert (F1:= Zgcd_is_gcd a b); inversion F1; auto. + Qed. + + Theorem spec_norm: forall n q, + ((0 < BigN.to_Z q)%Z -> [norm n q] == [Qq n q])%Q. + intros p q; unfold norm; intros Hq. + assert (Hp := BigN.spec_pos (BigZ.to_N p)). + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; auto; rewrite BigZ.spec_0; intros H1. + red; simpl; rewrite H1; ring. + case (Zle_lt_or_eq _ _ Hp); clear Hp; intros Hp. + case (Zle_lt_or_eq _ _ + (Zgcd_is_pos (BigN.to_Z (BigZ.to_N p)) (BigN.to_Z q))); intros H4. + 2: generalize Hq; rewrite (Zgcd_inv_0_r _ _ (sym_equal H4)); auto with zarith. + 2: red; simpl; auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_1; intros H2. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_1. + red; simpl. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite Zmult_1_r. + rewrite Z2P_correct; auto with zarith. + rewrite spec_to_N; intros; rewrite Zgcd_div_swap; auto. + rewrite H; ring. + intros H3. + red; simpl. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + assert (F: (0 < BigN.to_Z (q / BigN.gcd (BigZ.to_N p) q)%bigN)%Z). + rewrite BigN.spec_div; auto with zarith. + rewrite BigN.spec_gcd. + apply Zgcd_div_pos; auto. + rewrite BigN.spec_gcd; auto. + rewrite Z2P_correct; auto. + rewrite Z2P_correct; auto. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; auto with zarith. + rewrite spec_to_N; apply Zgcd_div_swap; auto. + case H1; rewrite spec_to_N; rewrite <- Hp; ring. + Qed. + + Theorem spec_normc: forall n q, + (0 < BigN.to_Z q)%Z -> [[norm n q]] = [[Qq n q]]. + intros n q H; unfold to_Qc, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_norm; auto. + Qed. + + Definition add (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.add zx zy) + | Qz zx, Qq ny dy => Qq (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + | Qq nx dx, Qz zy => Qq (BigZ.add nx (BigZ.mul zy (BigZ.Pos dx))) dx + | Qq nx dx, Qq ny dy => + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + Qq n d + end. + + Theorem wf_add: forall x y, wf x -> wf y -> wf (add x y). + intros [zx | nx dx] [zy | ny dy]; unfold add, wf; auto. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_mul. + intros H1 H2 H3. + case (Zmult_integral _ _ H1); auto with zarith. + Qed. + + Theorem spec_add x y: wf x -> wf y -> + ([add x y] == [x] + [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qplus; simpl. + rewrite BigZ.spec_add; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + assert (F1:= BigN.spec_pos dy). + rewrite Zmult_1_r. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool. + intros _ _ HH; case HH. + rewrite BigN.spec_0; intros HH _ _. + rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul. + simpl; apply Qeq_refl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH _ _. + assert (F1:= BigN.spec_pos dx). + rewrite Zmult_1_r; rewrite Pmult_1_r. + simpl; rewrite Z2P_correct; auto with zarith. + rewrite BigZ.spec_add; rewrite BigZ.spec_mul; simpl. + apply Qeq_refl. + generalize (BigN.spec_eq_bool dx BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH1. + generalize (BigN.spec_eq_bool dy BigN.zero); + case BigN.eq_bool. + intros _ _ HH; case HH. + rewrite BigN.spec_0; intros HH2 _ _. + assert (Fx: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (Fy: (0 < BigN.to_Z dy)%Z). + generalize (BigN.spec_pos dy); auto with zarith. + rewrite BigZ.spec_add; repeat rewrite BigN.spec_mul. + red; simpl. + rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto. + repeat rewrite BigZ.spec_mul; simpl; auto. + apply Zmult_lt_0_compat; auto. + Qed. + + Theorem spec_addc x y: wf x -> wf y -> + [[add x y]] = [[x]] + [[y]]. + intros x y H1 H2; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition add_norm (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.add zx zy) + | Qz zx, Qq ny dy => + norm (BigZ.add (BigZ.mul zx (BigZ.Pos dy)) ny) dy + | Qq nx dx, Qz zy => + norm (BigZ.add (BigZ.mul zy (BigZ.Pos dx)) nx) dx + | Qq nx dx, Qq ny dy => + let n := BigZ.add (BigZ.mul nx (BigZ.Pos dy)) (BigZ.mul ny (BigZ.Pos dx)) in + let d := BigN.mul dx dy in + norm n d + end. + + Theorem wf_add_norm: forall x y, wf x -> wf y -> wf (add_norm x y). + intros [zx | nx dx] [zy | ny dy]; unfold add_norm; auto. + intros HH1 HH2; apply wf_norm. + generalize HH2; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + intros HH1 HH2; apply wf_norm. + generalize HH1; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + intros HH1 HH2; apply wf_norm. + rewrite BigN.spec_mul; intros HH3. + case (Zmult_integral _ _ HH3). + generalize HH1; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + generalize HH2; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + Qed. + + Theorem spec_add_norm x y: wf x -> wf y -> + ([add_norm x y] == [x] + [y])%Q. + intros x y H1 H2; rewrite <- spec_add; auto. + generalize H1 H2; unfold add_norm, add, wf; case x; case y; clear H1 H2. + intros; apply Qeq_refl. + intros p1 n p2 _. + generalize (BigN.spec_eq_bool n BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH _. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + generalize (BigN.spec_pos n); auto with zarith. + simpl. + repeat rewrite BigZ.spec_add. + repeat rewrite BigZ.spec_mul; simpl. + apply Qeq_refl. + intros p1 n p2. + generalize (BigN.spec_eq_bool p2 BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH _ _. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end. + generalize (BigN.spec_pos p2); auto with zarith. + simpl. + repeat rewrite BigZ.spec_add. + repeat rewrite BigZ.spec_mul; simpl. + rewrite Zplus_comm. + apply Qeq_refl. + intros p1 q1 p2 q2. + generalize (BigN.spec_eq_bool q2 BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH1 _. + generalize (BigN.spec_eq_bool q1 BigN.zero); + case BigN.eq_bool. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH2 _. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end; try apply Qeq_refl. + rewrite BigN.spec_mul. + apply Zmult_lt_0_compat. + generalize (BigN.spec_pos q2); auto with zarith. + generalize (BigN.spec_pos q1); auto with zarith. + Qed. + + Theorem spec_add_normc x y: wf x -> wf y -> + [[add_norm x y]] = [[x]] + [[y]]. + intros x y Hx Hy; unfold to_Qc. + apply trans_equal with (!! ([x] + [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_add_norm; auto. + unfold Qcplus, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qplus_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition sub x y := add x (opp y). + + Theorem wf_sub x y: wf x -> wf y -> wf (sub x y). + intros x y Hx Hy; unfold sub; apply wf_add; auto. + apply wf_opp; auto. + Qed. + + Theorem spec_sub x y: wf x -> wf y -> + ([sub x y] == [x] - [y])%Q. + intros x y Hx Hy; unfold sub; rewrite spec_add; auto. + rewrite spec_opp; ring. + apply wf_opp; auto. + Qed. + + Theorem spec_subc x y: wf x -> wf y -> + [[sub x y]] = [[x]] - [[y]]. + intros x y Hx Hy; unfold sub; rewrite spec_addc; auto. + rewrite spec_oppc; ring. + apply wf_opp; auto. + Qed. + + Definition sub_norm x y := add_norm x (opp y). + + Theorem wf_sub_norm x y: wf x -> wf y -> wf (sub_norm x y). + intros x y Hx Hy; unfold sub_norm; apply wf_add_norm; auto. + apply wf_opp; auto. + Qed. + + Theorem spec_sub_norm x y: wf x -> wf y -> + ([sub_norm x y] == [x] - [y])%Q. + intros x y Hx Hy; unfold sub_norm; rewrite spec_add_norm; auto. + rewrite spec_opp; ring. + apply wf_opp; auto. + Qed. + + Theorem spec_sub_normc x y: wf x -> wf y -> + [[sub_norm x y]] = [[x]] - [[y]]. + intros x y Hx Hy; unfold sub_norm; rewrite spec_add_normc; auto. + rewrite spec_oppc; ring. + apply wf_opp; auto. + Qed. + + Definition mul (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => Qq (BigZ.mul zx ny) dy + | Qq nx dx, Qz zy => Qq (BigZ.mul nx zy) dx + | Qq nx dx, Qq ny dy => + Qq (BigZ.mul nx ny) (BigN.mul dx dy) + end. + + Theorem wf_mul: forall x y, wf x -> wf y -> wf (mul x y). + intros [zx | nx dx] [zy | ny dy]; unfold mul, wf; auto. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_mul. + intros H1 H2 H3. + case (Zmult_integral _ _ H1); auto with zarith. + Qed. + + Theorem spec_mul x y: wf x -> wf y -> ([mul x y] == [x] * [y])%Q. + intros [x | nx dx] [y | ny dy]; unfold Qmult; simpl. + rewrite BigZ.spec_mul; repeat rewrite Zmult_1_r; auto. + intros; apply Qeq_refl; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + intros _ _ HH; case HH. + rewrite BigN.spec_0; intros HH1 _ _. + rewrite BigZ.spec_mul; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + intros _ HH; case HH. + rewrite BigN.spec_0; intros HH1 _ _. + rewrite BigZ.spec_mul; rewrite Pmult_1_r. + apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + intros _ HH; case HH. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + intros _ _ _ HH; case HH. + rewrite BigN.spec_0; intros H1 H2 _ _. + rewrite BigZ.spec_mul; rewrite BigN.spec_mul. + assert (tmp: + (forall a b, 0 < a -> 0 < b -> Z2P (a * b) = (Z2P a * Z2P b)%positive)%Z). + intros [|a|a] [|b|b]; simpl; auto; intros; apply False_ind; auto with zarith. + rewrite tmp; auto. + apply Qeq_refl. + generalize (BigN.spec_pos dx); auto with zarith. + generalize (BigN.spec_pos dy); auto with zarith. + Qed. + + Theorem spec_mulc x y: wf x -> wf y -> + [[mul x y]] = [[x]] * [[y]]. + intros x y Hx Hy; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition mul_norm (x y: t): t := + match x, y with + | Qz zx, Qz zy => Qz (BigZ.mul zx zy) + | Qz zx, Qq ny dy => + if BigZ.eq_bool zx BigZ.zero then zero + else + let gcd := BigN.gcd (BigZ.to_N zx) dy in + if BigN.eq_bool gcd BigN.one then Qq (BigZ.mul zx ny) dy + else + let zx := BigZ.div zx (BigZ.Pos gcd) in + let d := BigN.div dy gcd in + if BigN.eq_bool d BigN.one then Qz (BigZ.mul zx ny) + else Qq (BigZ.mul zx ny) d + | Qq nx dx, Qz zy => + if BigZ.eq_bool zy BigZ.zero then zero + else + let gcd := BigN.gcd (BigZ.to_N zy) dx in + if BigN.eq_bool gcd BigN.one then Qq (BigZ.mul zy nx) dx + else + let zy := BigZ.div zy (BigZ.Pos gcd) in + let d := BigN.div dx gcd in + if BigN.eq_bool d BigN.one then Qz (BigZ.mul zy nx) + else Qq (BigZ.mul zy nx) d + | Qq nx dx, Qq ny dy => norm (BigZ.mul nx ny) (BigN.mul dx dy) + end. + + Theorem wf_mul_norm: forall x y, wf x -> wf y -> wf (mul_norm x y). + intros [zx | nx dx] [zy | ny dy]; unfold mul_norm; auto. + intros HH1 HH2. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + rewrite BigN.spec_1; rewrite BigZ.spec_0. + intros H1 H2; unfold wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + rewrite BigN.spec_0. + set (a := BigN.to_Z (BigZ.to_N zx)). + set (b := (BigN.to_Z dy)). + assert (F: (0 < Zgcd a b)%Z). + case (Zle_lt_or_eq _ _ (Zgcd_is_pos a b)); auto. + intros HH3; case H2; rewrite spec_to_N; fold a. + rewrite (Zgcd_inv_0_l _ _ (sym_equal HH3)); try ring. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; fold a; fold b; auto. + intros H. + generalize HH2; simpl wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + rewrite BigN.spec_0; intros HH; case HH; fold b. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite H; auto with zarith. + assert (F1:= Zgcd_is_gcd a b); inversion F1; auto. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + rewrite BigN.spec_1; rewrite BigN.spec_gcd. + intros HH1 H1 H2. + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; auto. + rewrite BigN.spec_1; rewrite BigN.spec_gcd. + intros HH1 H1 H2. + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; auto. + rewrite BigZ.spec_0. + intros HH2. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + set (a := BigN.to_Z (BigZ.to_N zy)). + set (b := (BigN.to_Z dx)). + assert (F: (0 < Zgcd a b)%Z). + case (Zle_lt_or_eq _ _ (Zgcd_is_pos a b)); auto. + intros HH3; case HH2; rewrite spec_to_N; fold a. + rewrite (Zgcd_inv_0_l _ _ (sym_equal HH3)); try ring. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; fold a; fold b; auto. + intros H; unfold wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + rewrite BigN.spec_0. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; fold a; fold b; auto. + intros HH; generalize H1; simpl wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + rewrite BigN.spec_0. + intros HH3; case HH3; fold b. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite HH; auto with zarith. + assert (F1:= Zgcd_is_gcd a b); inversion F1; auto. + intros HH1 HH2; apply wf_norm. + rewrite BigN.spec_mul; intros HH3. + case (Zmult_integral _ _ HH3). + generalize HH1; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + generalize HH2; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + Qed. + + Theorem spec_mul_norm x y: wf x -> wf y -> + ([mul_norm x y] == [x] * [y])%Q. + intros x y Hx Hy; rewrite <- spec_mul; auto. + unfold mul_norm, mul; generalize Hx Hy; case x; case y; clear Hx Hy. + intros; apply Qeq_refl. + intros p1 n p2 Hx Hy. + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; unfold zero, to_Q; repeat rewrite BigZ.spec_0; intros H. + rewrite BigZ.spec_mul; rewrite H; red; auto. + assert (F: (0 < BigN.to_Z (BigZ.to_N p2))%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos (BigZ.to_N p2))); auto. + intros H1; case H; rewrite spec_to_N; rewrite <- H1; ring. + assert (F1: (0 < BigN.to_Z n)%Z). + generalize Hy; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto. + intros _ HH; case HH. + rewrite BigN.spec_0; generalize (BigN.spec_pos n); auto with zarith. + set (a := BigN.to_Z (BigZ.to_N p2)). + set (b := BigN.to_Z n). + assert (F2: (0 < Zgcd a b )%Z). + case (Zle_lt_or_eq _ _ (Zgcd_is_pos a b)); intros H3; auto. + generalize F; fold a; rewrite (Zgcd_inv_0_l _ _ (sym_equal H3)); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; try rewrite BigN.spec_gcd; + fold a b; intros H1. + intros; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; + auto with zarith; fold a b; intros H2. + red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite spec_to_N; fold a; fold b. + rewrite Zmult_1_r; repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p1)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + intros H2; red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite Z2P_correct; auto with zarith. + rewrite (spec_to_N p2); fold a b. + rewrite Z2P_correct; auto with zarith. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p1)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto; try ring. + case (Zle_lt_or_eq _ _ + (BigN.spec_pos (n / + BigN.gcd (BigZ.to_N p2) + n))%bigN); + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + intros H3. + apply False_ind; generalize F1. + generalize Hy; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; auto with zarith. + intros HH; case HH; fold b. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + rewrite <- H3; ring. + assert (FF:= Zgcd_is_gcd a b); inversion FF; auto. + intros p1 p2 n Hx Hy. + match goal with |- context[BigZ.eq_bool ?X ?Y] => + generalize (BigZ.spec_eq_bool X Y); case BigZ.eq_bool + end; unfold zero, to_Q; repeat rewrite BigZ.spec_0; intros H. + rewrite BigZ.spec_mul; rewrite H; red; simpl; ring. + set (a := BigN.to_Z (BigZ.to_N p1)). + set (b := BigN.to_Z n). + assert (F: (0 < a)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos (BigZ.to_N p1))); auto. + intros H1; case H; rewrite spec_to_N; rewrite <- H1; ring. + assert (F1: (0 < b)%Z). + generalize Hx; unfold wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; auto with zarith. + generalize (BigN.spec_pos n); fold b; auto with zarith. + assert (F2: (0 < Zgcd a b)%Z). + case (Zle_lt_or_eq _ _ (Zgcd_is_pos a b)); intros H3; auto. + generalize F; rewrite (Zgcd_inv_0_l _ _ (sym_equal H3)); auto with zarith. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1; rewrite BigN.spec_gcd; fold a b; intros H1. + intros; repeat rewrite BigZ.spec_mul; rewrite Zmult_comm; apply Qeq_refl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_1. + rewrite BigN.spec_div; rewrite BigN.spec_gcd; + auto with zarith. + fold a b; intros H2. + red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite spec_to_N; fold a b. + rewrite Zmult_1_r; repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p2)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto with zarith. + rewrite H2; ring. + intros H2. + red; simpl. + repeat rewrite BigZ.spec_mul. + rewrite BigZ.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + rewrite Z2P_correct; auto with zarith. + rewrite (spec_to_N p1); fold a b. + case (Zle_lt_or_eq _ _ + (BigN.spec_pos (n / BigN.gcd (BigZ.to_N p1) n))%bigN); intros F3. + rewrite Z2P_correct; auto with zarith. + rewrite BigN.spec_div; simpl; rewrite BigN.spec_gcd; + fold a b; auto with zarith. + repeat rewrite <- Zmult_assoc. + rewrite (Zmult_comm (BigZ.to_Z p2)). + repeat rewrite Zmult_assoc. + rewrite Zgcd_div_swap; auto; try ring. + apply False_ind; generalize F1. + rewrite (Zdivide_Zdiv_eq (Zgcd a b) b); auto. + generalize F3; rewrite BigN.spec_div; rewrite BigN.spec_gcd; fold a b; + auto with zarith. + intros HH; rewrite <- HH; auto with zarith. + assert (FF:= Zgcd_is_gcd a b); inversion FF; auto. + intros p1 n1 p2 n2 Hn1 Hn2. + match goal with |- [norm ?X ?Y] == _ => + apply Qeq_trans with ([Qq X Y]); + [apply spec_norm | idtac] + end; try apply Qeq_refl. + rewrite BigN.spec_mul. + apply Zmult_lt_0_compat. + generalize Hn1; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; auto with zarith. + generalize (BigN.spec_pos n1) (BigN.spec_pos n2); auto with zarith. + generalize Hn2; simpl. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; auto with zarith. + generalize (BigN.spec_pos n1) (BigN.spec_pos n2); auto with zarith. + Qed. + + Theorem spec_mul_normc x y: wf x -> wf y -> + [[mul_norm x y]] = [[x]] * [[y]]. + intros x y Hx Hy; unfold to_Qc. + apply trans_equal with (!! ([x] * [y])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_mul_norm; auto. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + Definition inv (x: t): t := + match x with + | Qz (BigZ.Pos n) => + if BigN.eq_bool n BigN.zero then zero else Qq BigZ.one n + | Qz (BigZ.Neg n) => + if BigN.eq_bool n BigN.zero then zero else Qq BigZ.minus_one n + | Qq (BigZ.Pos n) d => + if BigN.eq_bool n BigN.zero then zero else Qq (BigZ.Pos d) n + | Qq (BigZ.Neg n) d => + if BigN.eq_bool n BigN.zero then zero else Qq (BigZ.Neg d) n + end. + + + Theorem wf_inv: forall x, wf x -> wf (inv x). + intros [ zx | nx dx]; unfold inv, wf; auto. + case zx; clear zx. + intros nx. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_mul. + intros nx. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0; rewrite BigN.spec_mul. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + intros _ HH; case HH. + intros H1 _. + case nx; clear nx. + intros nx. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; simpl; auto. + intros nx. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; simpl; auto. + Qed. + + Theorem spec_inv x: wf x -> + ([inv x] == /[x])%Q. + intros [ [x | x] _ | [nx | nx] dx]; unfold inv. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z x)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); auto with zarith. + unfold to_Q; rewrite BigZ.spec_1. + red; unfold Qinv; simpl. + generalize F; case BigN.to_Z; auto with zarith. + intros p Hp; discriminate Hp. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z x)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos x)); auto with zarith. + red; unfold Qinv; simpl. + generalize F; case BigN.to_Z; simpl; auto with zarith. + intros p Hp; discriminate Hp. + simpl wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1. + intros HH; case HH. + intros _. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z nx)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos nx)); auto with zarith. + red; unfold Qinv; simpl. + rewrite Z2P_correct; auto with zarith. + generalize F; case BigN.to_Z; auto with zarith. + intros p Hp; discriminate Hp. + generalize (BigN.spec_pos dx); auto with zarith. + simpl wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H1. + intros HH; case HH. + intros _. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; rewrite BigN.spec_0; intros H. + unfold zero, to_Q; rewrite BigZ.spec_0. + unfold BigZ.to_Z; rewrite H; apply Qeq_refl. + assert (F: (0 < BigN.to_Z nx)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos nx)); auto with zarith. + red; unfold Qinv; simpl. + rewrite Z2P_correct; auto with zarith. + generalize F; case BigN.to_Z; auto with zarith. + simpl; intros. + match goal with |- (?X = Zneg ?Y)%Z => + replace (Zneg Y) with (Zopp (Zpos Y)); + try rewrite Z2P_correct; auto with zarith + end. + rewrite Zpos_mult_morphism; + rewrite Z2P_correct; auto with zarith; try ring. + generalize (BigN.spec_pos dx); auto with zarith. + intros p Hp; discriminate Hp. + generalize (BigN.spec_pos dx); auto with zarith. + Qed. + + Theorem spec_invc x: wf x -> + [[inv x]] = /[[x]]. + intros x Hx; unfold to_Qc. + apply trans_equal with (!! (/[x])). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_inv; auto. + unfold Qcinv, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qinv_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + + Definition div x y := mul x (inv y). + + Theorem wf_div x y: wf x -> wf y -> wf (div x y). + intros x y Hx Hy; unfold div; apply wf_mul; auto. + apply wf_inv; auto. + Qed. + + Theorem spec_div x y: wf x -> wf y -> + ([div x y] == [x] / [y])%Q. + intros x y Hx Hy; unfold div; rewrite spec_mul; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + apply wf_inv; auto. + Qed. + + Theorem spec_divc x y: wf x -> wf y -> + [[div x y]] = [[x]] / [[y]]. + intros x y Hx Hy; unfold div; rewrite spec_mulc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + apply wf_inv; auto. + Qed. + + Definition div_norm x y := mul_norm x (inv y). + + Theorem wf_div_norm x y: wf x -> wf y -> wf (div_norm x y). + intros x y Hx Hy; unfold div_norm; apply wf_mul_norm; auto. + apply wf_inv; auto. + Qed. + + Theorem spec_div_norm x y: wf x -> wf y -> + ([div_norm x y] == [x] / [y])%Q. + intros x y Hx Hy; unfold div_norm; rewrite spec_mul_norm; auto. + unfold Qdiv; apply Qmult_comp. + apply Qeq_refl. + apply spec_inv; auto. + apply wf_inv; auto. + Qed. + + Theorem spec_div_normc x y: wf x -> wf y -> + [[div_norm x y]] = [[x]] / [[y]]. + intros x y Hx Hy; unfold div_norm; rewrite spec_mul_normc; auto. + unfold Qcdiv; apply f_equal2 with (f := Qcmult); auto. + apply spec_invc; auto. + apply wf_inv; auto. + Qed. + + Definition square (x: t): t := + match x with + | Qz zx => Qz (BigZ.square zx) + | Qq nx dx => Qq (BigZ.square nx) (BigN.square dx) + end. + + Theorem wf_square: forall x, wf x -> wf (square x). + intros [ zx | nx dx]; unfold square, wf; auto. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigN.spec_square; intros H1 H2; case H2. + case (Zmult_integral _ _ H1); auto. + Qed. + + Theorem spec_square x: wf x -> ([square x] == [x] ^ 2)%Q. + intros [ x | nx dx]; unfold square. + intros _. + red; simpl; rewrite BigZ.spec_square; auto with zarith. + unfold wf. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + intros _ HH; case HH. + intros H1 _. + red; simpl; rewrite BigZ.spec_square; auto with zarith. + assert (F: (0 < BigN.to_Z dx)%Z). + case (Zle_lt_or_eq _ _ (BigN.spec_pos dx)); auto with zarith. + assert (F1 : (0 < BigN.to_Z (BigN.square dx))%Z). + rewrite BigN.spec_square; apply Zmult_lt_0_compat; + auto with zarith. + rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto with zarith. + rewrite BigN.spec_square; auto with zarith. + Qed. + + Theorem spec_squarec x: wf x -> [[square x]] = [[x]]^2. + intros x Hx; unfold to_Qc. + apply trans_equal with (!! ([x]^2)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_square; auto. + simpl Qcpower. + replace (!! [x] * 1) with (!![x]); try ring. + simpl. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + + + Definition power_pos (x: t) p: t := + match x with + | Qz zx => Qz (BigZ.power_pos zx p) + | Qq nx dx => Qq (BigZ.power_pos nx p) (BigN.power_pos dx p) + end. + + Theorem wf_power_pos: forall x p, wf x -> wf (power_pos x p). + intros [ zx | nx dx] p; unfold power_pos, wf; auto. + repeat match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + rewrite BigN.spec_power_pos; simpl. + intros H1 H2 _. + case (Zle_lt_or_eq _ _ (BigN.spec_pos dx)); auto with zarith. + intros H3; generalize (Zpower_pos_pos _ p H3); auto with zarith. + Qed. + + Theorem spec_power_pos x p: wf x -> ([power_pos x p] == [x] ^ Zpos p)%Q. + Proof. + intros [x | nx dx] p; unfold power_pos. + intros _; unfold power_pos; red; simpl. + generalize (Qpower_decomp p (BigZ.to_Z x) 1). + unfold Qeq; simpl. + rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Zmult_1_r. + intros H; rewrite H. + rewrite BigZ.spec_power_pos; simpl; ring. + unfold wf. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + intros _ HH; case HH. + intros H1 _. + assert (F1: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + assert (F2: (0 < BigN.to_Z dx ^ ' p)%Z). + unfold Zpower; apply Zpower_pos_pos; auto. + unfold power_pos; red; simpl. + rewrite Z2P_correct; rewrite BigN.spec_power_pos; auto. + generalize (Qpower_decomp p (BigZ.to_Z nx) + (Z2P (BigN.to_Z dx))). + unfold Qeq; simpl. + repeat rewrite Z2P_correct; auto. + unfold Qeq; simpl; intros HH. + rewrite HH. + rewrite BigZ.spec_power_pos; simpl; ring. + Qed. + + Theorem spec_power_posc x p: wf x -> + [[power_pos x p]] = [[x]] ^ nat_of_P p. + intros x p Hx; unfold to_Qc. + apply trans_equal with (!! ([x]^Zpos p)). + unfold Q2Qc. + apply Qc_decomp; intros _ _; unfold this. + apply Qred_complete; apply spec_power_pos; auto. + pattern p; apply Pind; clear p. + simpl; ring. + intros p Hrec. + rewrite nat_of_P_succ_morphism; simpl Qcpower. + rewrite <- Hrec. + unfold Qcmult, Q2Qc. + apply Qc_decomp; intros _ _; + unfold this. + apply Qred_complete. + assert (F: [x] ^ ' Psucc p == [x] * [x] ^ ' p). + simpl; generalize Hx; case x; simpl; clear x Hx Hrec. + intros x _; simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + intros nx dx. + match goal with |- context[BigN.eq_bool ?X ?Y] => + generalize (BigN.spec_eq_bool X Y); case BigN.eq_bool + end; auto; rewrite BigN.spec_0. + intros _ HH; case HH. + intros H1 _. + assert (F1: (0 < BigN.to_Z dx)%Z). + generalize (BigN.spec_pos dx); auto with zarith. + simpl; repeat rewrite Qpower_decomp; simpl. + red; simpl; repeat rewrite Zpower_pos_1_l; simpl Z2P. + rewrite Pplus_one_succ_l. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + repeat rewrite Zpos_mult_morphism. + repeat rewrite Z2P_correct; auto. + 2: apply Zpower_pos_pos; auto. + 2: apply Zpower_pos_pos; auto. + rewrite Zpower_pos_is_exp. + rewrite Zpower_pos_1_r; auto. + rewrite F. + apply Qmult_comp; apply Qeq_sym; apply Qred_correct. + Qed. + +End Qv. + diff --git a/theories/Numbers/Rational/SpecViaQ/QSig.v b/theories/Numbers/Rational/SpecViaQ/QSig.v new file mode 100644 index 00000000..a488c7c6 --- /dev/null +++ b/theories/Numbers/Rational/SpecViaQ/QSig.v @@ -0,0 +1,84 @@ +(************************************************************************) +(* 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: QSig.v 11028 2008-06-01 17:34:19Z letouzey $ i*) + +Require Import QArith Qpower. + +Open Scope Q_scope. + +(** * QSig *) + +(** Interface of a rich structure about rational numbers. + Specifications are written via translation to Q. +*) + +Module Type QType. + + Parameter t : Type. + + Parameter to_Q : t -> Q. + Notation "[ x ]" := (to_Q x). + + Definition eq x y := [x] == [y]. + + Parameter of_Q : Q -> t. + Parameter spec_of_Q: forall x, to_Q (of_Q x) == x. + + Parameter zero : t. + Parameter one : t. + Parameter minus_one : t. + + Parameter spec_0: [zero] == 0. + Parameter spec_1: [one] == 1. + Parameter spec_m1: [minus_one] == -(1). + + Parameter compare : t -> t -> comparison. + + Parameter spec_compare: forall x y, compare x y = ([x] ?= [y]). + + Definition lt n m := compare n m = Lt. + Definition le n m := compare n m <> Gt. + Definition min n m := match compare n m with Gt => m | _ => n end. + Definition max n m := match compare n m with Lt => m | _ => n end. + + Parameter add : t -> t -> t. + + Parameter spec_add: forall x y, [add x y] == [x] + [y]. + + Parameter sub : t -> t -> t. + + Parameter spec_sub: forall x y, [sub x y] == [x] - [y]. + + Parameter opp : t -> t. + + Parameter spec_opp: forall x, [opp x] == - [x]. + + Parameter mul : t -> t -> t. + + Parameter spec_mul: forall x y, [mul x y] == [x] * [y]. + + Parameter square : t -> t. + + Parameter spec_square: forall x, [square x] == [x] ^ 2. + + Parameter inv : t -> t. + + Parameter spec_inv : forall x, [inv x] == / [x]. + + Parameter div : t -> t -> t. + + Parameter spec_div: forall x y, [div x y] == [x] / [y]. + + Parameter power_pos : t -> positive -> t. + + Parameter spec_power_pos: forall x n, [power_pos x n] == [x] ^ Zpos n. + +End QType. + +(* TODO: add norm function and variants, add eq_bool, what about Qc ? *)
\ No newline at end of file |