summaryrefslogtreecommitdiff
path: root/theories/Numbers/Natural
diff options
context:
space:
mode:
Diffstat (limited to 'theories/Numbers/Natural')
-rw-r--r--theories/Numbers/Natural/Abstract/NAdd.v22
-rw-r--r--theories/Numbers/Natural/Abstract/NAddOrder.v10
-rw-r--r--theories/Numbers/Natural/Abstract/NAxioms.v58
-rw-r--r--theories/Numbers/Natural/Abstract/NBase.v59
-rw-r--r--theories/Numbers/Natural/Abstract/NBits.v1463
-rw-r--r--theories/Numbers/Natural/Abstract/NDefOps.v177
-rw-r--r--theories/Numbers/Natural/Abstract/NDiv.v50
-rw-r--r--theories/Numbers/Natural/Abstract/NGcd.v213
-rw-r--r--theories/Numbers/Natural/Abstract/NIso.v21
-rw-r--r--theories/Numbers/Natural/Abstract/NLcm.v290
-rw-r--r--theories/Numbers/Natural/Abstract/NLog.v23
-rw-r--r--theories/Numbers/Natural/Abstract/NMaxMin.v135
-rw-r--r--theories/Numbers/Natural/Abstract/NMulOrder.v22
-rw-r--r--theories/Numbers/Natural/Abstract/NOrder.v43
-rw-r--r--theories/Numbers/Natural/Abstract/NParity.v63
-rw-r--r--theories/Numbers/Natural/Abstract/NPow.v160
-rw-r--r--theories/Numbers/Natural/Abstract/NProperties.v21
-rw-r--r--theories/Numbers/Natural/Abstract/NSqrt.v75
-rw-r--r--theories/Numbers/Natural/Abstract/NStrongRec.v44
-rw-r--r--theories/Numbers/Natural/Abstract/NSub.v44
-rw-r--r--theories/Numbers/Natural/BigN/BigN.v107
-rw-r--r--theories/Numbers/Natural/BigN/NMake.v1448
-rw-r--r--theories/Numbers/Natural/BigN/NMake_gen.ml3511
-rw-r--r--theories/Numbers/Natural/BigN/Nbasic.v223
-rw-r--r--theories/Numbers/Natural/Binary/NBinary.v141
-rw-r--r--theories/Numbers/Natural/Peano/NPeano.v806
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSig.v76
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v375
28 files changed, 6146 insertions, 3534 deletions
diff --git a/theories/Numbers/Natural/Abstract/NAdd.v b/theories/Numbers/Natural/Abstract/NAdd.v
index 4185de95..72e09f15 100644
--- a/theories/Numbers/Natural/Abstract/NAdd.v
+++ b/theories/Numbers/Natural/Abstract/NAdd.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,12 +8,10 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NAdd.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Export NBase.
-Module NAddPropFunct (Import N : NAxiomsSig').
-Include NBasePropFunct N.
+Module NAddProp (Import N : NAxiomsMiniSig').
+Include NBaseProp N.
(** For theorems about [add] that are both valid for [N] and [Z], see [NZAdd] *)
(** Now comes theorems valid for natural numbers but not for Z *)
@@ -24,9 +22,9 @@ intros n m; induct n.
nzsimpl; intuition.
intros n IH. nzsimpl.
setoid_replace (S (n + m) == 0) with False by
- (apply -> neg_false; apply neq_succ_0).
+ (apply neg_false; apply neq_succ_0).
setoid_replace (S n == 0) with False by
- (apply -> neg_false; apply neq_succ_0). tauto.
+ (apply neg_false; apply neq_succ_0). tauto.
Qed.
Theorem eq_add_succ :
@@ -47,13 +45,13 @@ Qed.
Theorem eq_add_1 : forall n m,
n + m == 1 -> n == 1 /\ m == 0 \/ n == 0 /\ m == 1.
Proof.
-intros n m H.
+intros n m. rewrite one_succ. intro H.
assert (H1 : exists p, n + m == S p) by now exists 0.
-apply -> eq_add_succ in H1. destruct H1 as [[n' H1] | [m' H1]].
+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.
+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.
+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, m ~= S (n + m).
@@ -77,5 +75,5 @@ intros n m H; rewrite (add_comm n (P m));
rewrite (add_comm n m); now apply add_pred_l.
Qed.
-End NAddPropFunct.
+End NAddProp.
diff --git a/theories/Numbers/Natural/Abstract/NAddOrder.v b/theories/Numbers/Natural/Abstract/NAddOrder.v
index 0282a6b8..da41886f 100644
--- a/theories/Numbers/Natural/Abstract/NAddOrder.v
+++ b/theories/Numbers/Natural/Abstract/NAddOrder.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,12 +8,10 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NAddOrder.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Export NOrder.
-Module NAddOrderPropFunct (Import N : NAxiomsSig').
-Include NOrderPropFunct N.
+Module NAddOrderProp (Import N : NAxiomsMiniSig').
+Include NOrderProp N.
(** Theorems true for natural numbers, not for integers *)
@@ -45,4 +43,4 @@ Proof.
intros; apply add_nonneg_pos. apply le_0_l. assumption.
Qed.
-End NAddOrderPropFunct.
+End NAddOrderProp.
diff --git a/theories/Numbers/Natural/Abstract/NAxioms.v b/theories/Numbers/Natural/Abstract/NAxioms.v
index d1cc9972..ca6ccc1b 100644
--- a/theories/Numbers/Natural/Abstract/NAxioms.v
+++ b/theories/Numbers/Natural/Abstract/NAxioms.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,32 +8,60 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NAxioms.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
+Require Export Bool NZAxioms NZParity NZPow NZSqrt NZLog NZDiv NZGcd NZBits.
-Require Export NZAxioms.
+(** From [NZ], we obtain natural numbers just by stating that [pred 0] == 0 *)
-Set Implicit Arguments.
+Module Type NAxiom (Import NZ : NZDomainSig').
+ Axiom pred_0 : P 0 == 0.
+End NAxiom.
-Module Type NAxioms (Import NZ : NZDomainSig').
+Module Type NAxiomsMiniSig := NZOrdAxiomsSig <+ NAxiom.
+Module Type NAxiomsMiniSig' := NZOrdAxiomsSig' <+ NAxiom.
-Axiom pred_0 : P 0 == 0.
+(** Let's now add some more functions and their specification *)
-Parameter Inline recursion : forall A : Type, A -> (t -> A -> A) -> t -> A.
-Implicit Arguments recursion [A].
+(** Division Function : we reuse NZDiv.DivMod and NZDiv.NZDivCommon,
+ and add to that a N-specific constraint. *)
-Declare Instance recursion_wd (A : Type) (Aeq : relation A) :
- Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) (@recursion A).
+Module Type NDivSpecific (Import N : NAxiomsMiniSig')(Import DM : DivMod' N).
+ Axiom mod_upper_bound : forall a b, b ~= 0 -> a mod b < b.
+End NDivSpecific.
+
+(** For all other functions, the NZ axiomatizations are enough. *)
+
+(** We now group everything together. *)
+
+Module Type NAxiomsSig := NAxiomsMiniSig <+ OrderFunctions
+ <+ NZParity.NZParity <+ NZPow.NZPow <+ NZSqrt.NZSqrt <+ NZLog.NZLog2
+ <+ NZGcd.NZGcd <+ NZDiv.NZDiv <+ NZBits.NZBits <+ NZSquare.
+
+Module Type NAxiomsSig' := NAxiomsMiniSig' <+ OrderFunctions'
+ <+ NZParity.NZParity <+ NZPow.NZPow' <+ NZSqrt.NZSqrt' <+ NZLog.NZLog2
+ <+ NZGcd.NZGcd' <+ NZDiv.NZDiv' <+ NZBits.NZBits' <+ NZSquare.
+
+
+(** It could also be interesting to have a constructive recursor function. *)
+
+Module Type NAxiomsRec (Import NZ : NZDomainSig').
+
+Parameter Inline recursion : forall {A : Type}, A -> (t -> A -> A) -> t -> A.
+
+Declare Instance recursion_wd {A : Type} (Aeq : relation A) :
+ Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) recursion.
Axiom recursion_0 :
- forall (A : Type) (a : A) (f : t -> A -> A), recursion a f 0 = a.
+ forall {A} (a : A) (f : t -> A -> A), recursion a f 0 = a.
Axiom recursion_succ :
- forall (A : Type) (Aeq : relation A) (a : A) (f : t -> A -> A),
+ forall {A} (Aeq : relation A) (a : A) (f : t -> A -> A),
Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
forall n, Aeq (recursion a f (S n)) (f n (recursion a f n)).
-End NAxioms.
+End NAxiomsRec.
-Module Type NAxiomsSig := NZOrdAxiomsSig <+ NAxioms.
-Module Type NAxiomsSig' := NZOrdAxiomsSig' <+ NAxioms.
+Module Type NAxiomsRecSig := NAxiomsMiniSig <+ NAxiomsRec.
+Module Type NAxiomsRecSig' := NAxiomsMiniSig' <+ NAxiomsRec.
+Module Type NAxiomsFullSig := NAxiomsSig <+ NAxiomsRec.
+Module Type NAxiomsFullSig' := NAxiomsSig' <+ NAxiomsRec.
diff --git a/theories/Numbers/Natural/Abstract/NBase.v b/theories/Numbers/Natural/Abstract/NBase.v
index efaba960..ac8a0522 100644
--- a/theories/Numbers/Natural/Abstract/NBase.v
+++ b/theories/Numbers/Natural/Abstract/NBase.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,48 +8,23 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NBase.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Export Decidable.
Require Export NAxioms.
Require Import NZProperties.
-Module NBasePropFunct (Import N : NAxiomsSig').
+Module NBaseProp (Import N : NAxiomsMiniSig').
(** First, we import all known facts about both natural numbers and integers. *)
-Include NZPropFunct N.
-
-(** 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 : Type) (a b : A) (n : N.t) : A :=
- recursion a (fun _ _ => b) n.
-
-Implicit Arguments if_zero [A].
-
-Instance if_zero_wd (A : Type) :
- Proper (Logic.eq ==> Logic.eq ==> N.eq ==> Logic.eq) (@if_zero A).
-Proof.
-intros; unfold if_zero.
-repeat red; intros. apply recursion_wd; auto. repeat red; auto.
-Qed.
-
-Theorem if_zero_0 : forall (A : Type) (a b : A), if_zero a b 0 = a.
-Proof.
-unfold if_zero; intros; now rewrite recursion_0.
-Qed.
+Include NZProp N.
-Theorem if_zero_succ :
- forall (A : Type) (a b : A) (n : N.t), if_zero a b (S n) = b.
-Proof.
-intros; unfold if_zero.
-now rewrite recursion_succ.
-Qed.
+(** From [pred_0] and order facts, we can prove that 0 isn't a successor. *)
Theorem neq_succ_0 : forall n, S n ~= 0.
Proof.
-intros n H.
-generalize (Logic.eq_refl (if_zero false true 0)).
-rewrite <- H at 1. rewrite if_zero_0, if_zero_succ; discriminate.
+ intros n EQ.
+ assert (EQ' := pred_succ n).
+ rewrite EQ, pred_0 in EQ'.
+ rewrite <- EQ' in EQ.
+ now apply (neq_succ_diag_l 0).
Qed.
Theorem neq_0_succ : forall n, 0 ~= S n.
@@ -66,7 +41,7 @@ nzinduct n.
now apply eq_le_incl.
intro n; split.
apply le_le_succ_r.
-intro H; apply -> le_succ_r in H; destruct H as [H | H].
+intro H; apply le_succ_r in H; destruct H as [H | H].
assumption.
symmetry in H; false_hyp H neq_succ_0.
Qed.
@@ -119,12 +94,11 @@ Qed.
Theorem eq_pred_0 : forall n, P n == 0 <-> n == 0 \/ n == 1.
Proof.
cases n.
-rewrite pred_0. setoid_replace (0 == 1) with False using relation iff. tauto.
-split; intro H; [symmetry in H; false_hyp H neq_succ_0 | elim H].
+rewrite pred_0. now split; [left|].
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.
+split. intros H; right. now rewrite H, one_succ.
+intros [H|H]. elim (neq_succ_0 _ H).
+apply succ_inj_wd. now rewrite <- one_succ.
Qed.
Theorem succ_pred : forall n, n ~= 0 -> S (P n) == n.
@@ -155,6 +129,7 @@ Theorem pair_induction :
A 0 -> A 1 ->
(forall n, A n -> A (S n) -> A (S (S n))) -> forall n, A n.
Proof.
+rewrite one_succ.
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.
@@ -204,7 +179,7 @@ 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 | | | ].
+ [solve_proper | | | ].
-End NBasePropFunct.
+End NBaseProp.
diff --git a/theories/Numbers/Natural/Abstract/NBits.v b/theories/Numbers/Natural/Abstract/NBits.v
new file mode 100644
index 00000000..c66f003e
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NBits.v
@@ -0,0 +1,1463 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Bool NAxioms NSub NPow NDiv NParity NLog.
+
+(** Derived properties of bitwise operations *)
+
+Module Type NBitsProp
+ (Import A : NAxiomsSig')
+ (Import B : NSubProp A)
+ (Import C : NParityProp A B)
+ (Import D : NPowProp A B C)
+ (Import E : NDivProp A B)
+ (Import F : NLog2Prop A B C D).
+
+Include BoolEqualityFacts A.
+
+Ltac order_nz := try apply pow_nonzero; order'.
+Hint Rewrite div_0_l mod_0_l div_1_r mod_1_r : nz.
+
+(** Some properties of power and division *)
+
+Lemma pow_sub_r : forall a b c, a~=0 -> c<=b -> a^(b-c) == a^b / a^c.
+Proof.
+ intros a b c Ha H.
+ apply div_unique with 0.
+ generalize (pow_nonzero a c Ha) (le_0_l (a^c)); order'.
+ nzsimpl. now rewrite <- pow_add_r, add_comm, sub_add.
+Qed.
+
+Lemma pow_div_l : forall a b c, b~=0 -> a mod b == 0 ->
+ (a/b)^c == a^c / b^c.
+Proof.
+ intros a b c Hb H.
+ apply div_unique with 0.
+ generalize (pow_nonzero b c Hb) (le_0_l (b^c)); order'.
+ nzsimpl. rewrite <- pow_mul_l. f_equiv. now apply div_exact.
+Qed.
+
+(** An injection from bits [true] and [false] to numbers 1 and 0.
+ We declare it as a (local) coercion for shorter statements. *)
+
+Definition b2n (b:bool) := if b then 1 else 0.
+Local Coercion b2n : bool >-> t.
+
+Instance b2n_proper : Proper (Logic.eq ==> eq) b2n.
+Proof. solve_proper. Qed.
+
+Lemma exists_div2 a : exists a' (b:bool), a == 2*a' + b.
+Proof.
+ elim (Even_or_Odd a); [intros (a',H)| intros (a',H)].
+ exists a'. exists false. now nzsimpl.
+ exists a'. exists true. now simpl.
+Qed.
+
+(** We can compact [testbit_odd_0] [testbit_even_0]
+ [testbit_even_succ] [testbit_odd_succ] in only two lemmas. *)
+
+Lemma testbit_0_r a (b:bool) : testbit (2*a+b) 0 = b.
+Proof.
+ destruct b; simpl; rewrite ?add_0_r.
+ apply testbit_odd_0.
+ apply testbit_even_0.
+Qed.
+
+Lemma testbit_succ_r a (b:bool) n :
+ testbit (2*a+b) (succ n) = testbit a n.
+Proof.
+ destruct b; simpl; rewrite ?add_0_r.
+ apply testbit_odd_succ, le_0_l.
+ apply testbit_even_succ, le_0_l.
+Qed.
+
+(** Alternative caracterisations of [testbit] *)
+
+(** This concise equation could have been taken as specification
+ for testbit in the interface, but it would have been hard to
+ implement with little initial knowledge about div and mod *)
+
+Lemma testbit_spec' a n : a.[n] == (a / 2^n) mod 2.
+Proof.
+ revert a. induct n.
+ intros a. nzsimpl.
+ destruct (exists_div2 a) as (a' & b & H). rewrite H at 1.
+ rewrite testbit_0_r. apply mod_unique with a'; trivial.
+ destruct b; order'.
+ intros n IH a.
+ destruct (exists_div2 a) as (a' & b & H). rewrite H at 1.
+ rewrite testbit_succ_r, IH. f_equiv.
+ rewrite pow_succ_r', <- div_div by order_nz. f_equiv.
+ apply div_unique with b; trivial.
+ destruct b; order'.
+Qed.
+
+(** This caracterisation that uses only basic operations and
+ power was initially taken as specification for testbit.
+ We describe [a] as having a low part and a high part, with
+ the corresponding bit in the middle. This caracterisation
+ is moderatly complex to implement, but also moderately
+ usable... *)
+
+Lemma testbit_spec a n :
+ exists l h, 0<=l<2^n /\ a == l + (a.[n] + 2*h)*2^n.
+Proof.
+ exists (a mod 2^n). exists (a / 2^n / 2). split.
+ split; [apply le_0_l | apply mod_upper_bound; order_nz].
+ rewrite add_comm, mul_comm, (add_comm a.[n]).
+ rewrite (div_mod a (2^n)) at 1 by order_nz. do 2 f_equiv.
+ rewrite testbit_spec'. apply div_mod. order'.
+Qed.
+
+Lemma testbit_true : forall a n,
+ a.[n] = true <-> (a / 2^n) mod 2 == 1.
+Proof.
+ intros a n.
+ rewrite <- testbit_spec'; destruct a.[n]; split; simpl; now try order'.
+Qed.
+
+Lemma testbit_false : forall a n,
+ a.[n] = false <-> (a / 2^n) mod 2 == 0.
+Proof.
+ intros a n.
+ rewrite <- testbit_spec'; destruct a.[n]; split; simpl; now try order'.
+Qed.
+
+Lemma testbit_eqb : forall a n,
+ a.[n] = eqb ((a / 2^n) mod 2) 1.
+Proof.
+ intros a n.
+ apply eq_true_iff_eq. now rewrite testbit_true, eqb_eq.
+Qed.
+
+(** Results about the injection [b2n] *)
+
+Lemma b2n_inj : forall (a0 b0:bool), a0 == b0 -> a0 = b0.
+Proof.
+ intros [|] [|]; simpl; trivial; order'.
+Qed.
+
+Lemma add_b2n_double_div2 : forall (a0:bool) a, (a0+2*a)/2 == a.
+Proof.
+ intros a0 a. rewrite mul_comm, div_add by order'.
+ now rewrite div_small, add_0_l by (destruct a0; order').
+Qed.
+
+Lemma add_b2n_double_bit0 : forall (a0:bool) a, (a0+2*a).[0] = a0.
+Proof.
+ intros a0 a. apply b2n_inj.
+ rewrite testbit_spec'. nzsimpl. rewrite mul_comm, mod_add by order'.
+ now rewrite mod_small by (destruct a0; order').
+Qed.
+
+Lemma b2n_div2 : forall (a0:bool), a0/2 == 0.
+Proof.
+ intros a0. rewrite <- (add_b2n_double_div2 a0 0). now nzsimpl.
+Qed.
+
+Lemma b2n_bit0 : forall (a0:bool), a0.[0] = a0.
+Proof.
+ intros a0. rewrite <- (add_b2n_double_bit0 a0 0) at 2. now nzsimpl.
+Qed.
+
+(** The specification of testbit by low and high parts is complete *)
+
+Lemma testbit_unique : forall a n (a0:bool) l h,
+ l<2^n -> a == l + (a0 + 2*h)*2^n -> a.[n] = a0.
+Proof.
+ intros a n a0 l h Hl EQ.
+ apply b2n_inj. rewrite testbit_spec' by trivial.
+ symmetry. apply mod_unique with h. destruct a0; simpl; order'.
+ symmetry. apply div_unique with l; trivial.
+ now rewrite add_comm, (add_comm _ a0), mul_comm.
+Qed.
+
+(** All bits of number 0 are 0 *)
+
+Lemma bits_0 : forall n, 0.[n] = false.
+Proof.
+ intros n. apply testbit_false. nzsimpl; order_nz.
+Qed.
+
+(** Various ways to refer to the lowest bit of a number *)
+
+Lemma bit0_odd : forall a, a.[0] = odd a.
+Proof.
+ intros. symmetry.
+ destruct (exists_div2 a) as (a' & b & EQ).
+ rewrite EQ, testbit_0_r, add_comm, odd_add_mul_2.
+ destruct b; simpl; apply odd_1 || apply odd_0.
+Qed.
+
+Lemma bit0_eqb : forall a, a.[0] = eqb (a mod 2) 1.
+Proof.
+ intros a. rewrite testbit_eqb. now nzsimpl.
+Qed.
+
+Lemma bit0_mod : forall a, a.[0] == a mod 2.
+Proof.
+ intros a. rewrite testbit_spec'. now nzsimpl.
+Qed.
+
+(** Hence testing a bit is equivalent to shifting and testing parity *)
+
+Lemma testbit_odd : forall a n, a.[n] = odd (a>>n).
+Proof.
+ intros. now rewrite <- bit0_odd, shiftr_spec, add_0_l.
+Qed.
+
+(** [log2] gives the highest nonzero bit *)
+
+Lemma bit_log2 : forall a, a~=0 -> a.[log2 a] = true.
+Proof.
+ intros a Ha.
+ assert (Ha' : 0 < a) by (generalize (le_0_l a); order).
+ destruct (log2_spec_alt a Ha') as (r & EQ & (_,Hr)).
+ rewrite EQ at 1.
+ rewrite testbit_true, add_comm.
+ rewrite <- (mul_1_l (2^log2 a)) at 1.
+ rewrite div_add by order_nz.
+ rewrite div_small by trivial.
+ rewrite add_0_l. apply mod_small. order'.
+Qed.
+
+Lemma bits_above_log2 : forall a n, log2 a < n ->
+ a.[n] = false.
+Proof.
+ intros a n H.
+ rewrite testbit_false.
+ rewrite div_small. nzsimpl; order'.
+ apply log2_lt_cancel. rewrite log2_pow2; trivial using le_0_l.
+Qed.
+
+(** Hence the number of bits of [a] is [1+log2 a]
+ (see [Psize] and [Psize_pos]).
+*)
+
+(** Testing bits after division or multiplication by a power of two *)
+
+Lemma div2_bits : forall a n, (a/2).[n] = a.[S n].
+Proof.
+ intros. apply eq_true_iff_eq.
+ rewrite 2 testbit_true.
+ rewrite pow_succ_r by apply le_0_l.
+ now rewrite div_div by order_nz.
+Qed.
+
+Lemma div_pow2_bits : forall a n m, (a/2^n).[m] = a.[m+n].
+Proof.
+ intros a n. revert a. induct n.
+ intros a m. now nzsimpl.
+ intros n IH a m. nzsimpl; try apply le_0_l.
+ rewrite <- div_div by order_nz.
+ now rewrite IH, div2_bits.
+Qed.
+
+Lemma double_bits_succ : forall a n, (2*a).[S n] = a.[n].
+Proof.
+ intros. rewrite <- div2_bits. now rewrite mul_comm, div_mul by order'.
+Qed.
+
+Lemma mul_pow2_bits_add : forall a n m, (a*2^n).[m+n] = a.[m].
+Proof.
+ intros. rewrite <- div_pow2_bits. now rewrite div_mul by order_nz.
+Qed.
+
+Lemma mul_pow2_bits_high : forall a n m, n<=m -> (a*2^n).[m] = a.[m-n].
+Proof.
+ intros.
+ rewrite <- (sub_add n m) at 1 by order'.
+ now rewrite mul_pow2_bits_add.
+Qed.
+
+Lemma mul_pow2_bits_low : forall a n m, m<n -> (a*2^n).[m] = false.
+Proof.
+ intros. apply testbit_false.
+ rewrite <- (sub_add m n) by order'. rewrite pow_add_r, mul_assoc.
+ rewrite div_mul by order_nz.
+ rewrite <- (succ_pred (n-m)). rewrite pow_succ_r.
+ now rewrite (mul_comm 2), mul_assoc, mod_mul by order'.
+ apply lt_le_pred.
+ apply sub_gt in H. generalize (le_0_l (n-m)); order.
+ now apply sub_gt.
+Qed.
+
+(** Selecting the low part of a number can be done by a modulo *)
+
+Lemma mod_pow2_bits_high : forall a n m, n<=m ->
+ (a mod 2^n).[m] = false.
+Proof.
+ intros a n m H.
+ destruct (eq_0_gt_0_cases (a mod 2^n)) as [EQ|LT].
+ now rewrite EQ, bits_0.
+ apply bits_above_log2.
+ apply lt_le_trans with n; trivial.
+ apply log2_lt_pow2; trivial.
+ apply mod_upper_bound; order_nz.
+Qed.
+
+Lemma mod_pow2_bits_low : forall a n m, m<n ->
+ (a mod 2^n).[m] = a.[m].
+Proof.
+ intros a n m H.
+ rewrite testbit_eqb.
+ rewrite <- (mod_add _ (2^(P (n-m))*(a/2^n))) by order'.
+ rewrite <- div_add by order_nz.
+ rewrite (mul_comm _ 2), mul_assoc, <- pow_succ_r', succ_pred
+ by now apply sub_gt.
+ rewrite mul_comm, mul_assoc, <- pow_add_r, (add_comm m), sub_add
+ by order.
+ rewrite add_comm, <- div_mod by order_nz.
+ symmetry. apply testbit_eqb.
+Qed.
+
+(** We now prove that having the same bits implies equality.
+ For that we use a notion of equality over functional
+ streams of bits. *)
+
+Definition eqf (f g:t -> bool) := forall n:t, f n = g n.
+
+Instance eqf_equiv : Equivalence eqf.
+Proof.
+ split; congruence.
+Qed.
+
+Local Infix "===" := eqf (at level 70, no associativity).
+
+Instance testbit_eqf : Proper (eq==>eqf) testbit.
+Proof.
+ intros a a' Ha n. now rewrite Ha.
+Qed.
+
+(** Only zero corresponds to the always-false stream. *)
+
+Lemma bits_inj_0 :
+ forall a, (forall n, a.[n] = false) -> a == 0.
+Proof.
+ intros a H. destruct (eq_decidable a 0) as [EQ|NEQ]; trivial.
+ apply bit_log2 in NEQ. now rewrite H in NEQ.
+Qed.
+
+(** If two numbers produce the same stream of bits, they are equal. *)
+
+Lemma bits_inj : forall a b, testbit a === testbit b -> a == b.
+Proof.
+ intros a. pattern a.
+ apply strong_right_induction with 0;[solve_proper|clear a|apply le_0_l].
+ intros a _ IH b H.
+ destruct (eq_0_gt_0_cases a) as [EQ|LT].
+ rewrite EQ in H |- *. symmetry. apply bits_inj_0.
+ intros n. now rewrite <- H, bits_0.
+ rewrite (div_mod a 2), (div_mod b 2) by order'.
+ f_equiv; [ | now rewrite <- 2 bit0_mod, H].
+ f_equiv.
+ apply IH; trivial using le_0_l.
+ apply div_lt; order'.
+ intro n. rewrite 2 div2_bits. apply H.
+Qed.
+
+Lemma bits_inj_iff : forall a b, testbit a === testbit b <-> a == b.
+Proof.
+ split. apply bits_inj. intros EQ; now rewrite EQ.
+Qed.
+
+Hint Rewrite lxor_spec lor_spec land_spec ldiff_spec bits_0 : bitwise.
+
+Ltac bitwise := apply bits_inj; intros ?m; autorewrite with bitwise.
+
+(** The streams of bits that correspond to a natural numbers are
+ exactly the ones that are always 0 after some point *)
+
+Lemma are_bits : forall (f:t->bool), Proper (eq==>Logic.eq) f ->
+ ((exists n, f === testbit n) <->
+ (exists k, forall m, k<=m -> f m = false)).
+Proof.
+ intros f Hf. split.
+ intros (a,H).
+ exists (S (log2 a)). intros m Hm. apply le_succ_l in Hm.
+ rewrite H, bits_above_log2; trivial using lt_succ_diag_r.
+ intros (k,Hk).
+ revert f Hf Hk. induct k.
+ intros f Hf H0.
+ exists 0. intros m. rewrite bits_0, H0; trivial. apply le_0_l.
+ intros k IH f Hf Hk.
+ destruct (IH (fun m => f (S m))) as (n, Hn).
+ solve_proper.
+ intros m Hm. apply Hk. now rewrite <- succ_le_mono.
+ exists (f 0 + 2*n). intros m.
+ destruct (zero_or_succ m) as [Hm|(m', Hm)]; rewrite Hm.
+ symmetry. apply add_b2n_double_bit0.
+ rewrite Hn, <- div2_bits.
+ rewrite mul_comm, div_add, b2n_div2, add_0_l; trivial. order'.
+Qed.
+
+(** Properties of shifts *)
+
+Lemma shiftr_spec' : forall a n m, (a >> n).[m] = a.[m+n].
+Proof.
+ intros. apply shiftr_spec. apply le_0_l.
+Qed.
+
+Lemma shiftl_spec_high' : forall a n m, n<=m -> (a << n).[m] = a.[m-n].
+Proof.
+ intros. apply shiftl_spec_high; trivial. apply le_0_l.
+Qed.
+
+Lemma shiftr_div_pow2 : forall a n, a >> n == a / 2^n.
+Proof.
+ intros. bitwise. rewrite shiftr_spec'.
+ symmetry. apply div_pow2_bits.
+Qed.
+
+Lemma shiftl_mul_pow2 : forall a n, a << n == a * 2^n.
+Proof.
+ intros. bitwise.
+ destruct (le_gt_cases n m) as [H|H].
+ now rewrite shiftl_spec_high', mul_pow2_bits_high.
+ now rewrite shiftl_spec_low, mul_pow2_bits_low.
+Qed.
+
+Lemma shiftl_spec_alt : forall a n m, (a << n).[m+n] = a.[m].
+Proof.
+ intros. now rewrite shiftl_mul_pow2, mul_pow2_bits_add.
+Qed.
+
+Instance shiftr_wd : Proper (eq==>eq==>eq) shiftr.
+Proof.
+ intros a a' Ha b b' Hb. now rewrite 2 shiftr_div_pow2, Ha, Hb.
+Qed.
+
+Instance shiftl_wd : Proper (eq==>eq==>eq) shiftl.
+Proof.
+ intros a a' Ha b b' Hb. now rewrite 2 shiftl_mul_pow2, Ha, Hb.
+Qed.
+
+Lemma shiftl_shiftl : forall a n m,
+ (a << n) << m == a << (n+m).
+Proof.
+ intros. now rewrite !shiftl_mul_pow2, pow_add_r, mul_assoc.
+Qed.
+
+Lemma shiftr_shiftr : forall a n m,
+ (a >> n) >> m == a >> (n+m).
+Proof.
+ intros.
+ now rewrite !shiftr_div_pow2, pow_add_r, div_div by order_nz.
+Qed.
+
+Lemma shiftr_shiftl_l : forall a n m, m<=n ->
+ (a << n) >> m == a << (n-m).
+Proof.
+ intros.
+ rewrite shiftr_div_pow2, !shiftl_mul_pow2.
+ rewrite <- (sub_add m n) at 1 by trivial.
+ now rewrite pow_add_r, mul_assoc, div_mul by order_nz.
+Qed.
+
+Lemma shiftr_shiftl_r : forall a n m, n<=m ->
+ (a << n) >> m == a >> (m-n).
+Proof.
+ intros.
+ rewrite !shiftr_div_pow2, shiftl_mul_pow2.
+ rewrite <- (sub_add n m) at 1 by trivial.
+ rewrite pow_add_r, (mul_comm (2^(m-n))).
+ now rewrite <- div_div, div_mul by order_nz.
+Qed.
+
+(** shifts and constants *)
+
+Lemma shiftl_1_l : forall n, 1 << n == 2^n.
+Proof.
+ intros. now rewrite shiftl_mul_pow2, mul_1_l.
+Qed.
+
+Lemma shiftl_0_r : forall a, a << 0 == a.
+Proof.
+ intros. rewrite shiftl_mul_pow2. now nzsimpl.
+Qed.
+
+Lemma shiftr_0_r : forall a, a >> 0 == a.
+Proof.
+ intros. rewrite shiftr_div_pow2. now nzsimpl.
+Qed.
+
+Lemma shiftl_0_l : forall n, 0 << n == 0.
+Proof.
+ intros. rewrite shiftl_mul_pow2. now nzsimpl.
+Qed.
+
+Lemma shiftr_0_l : forall n, 0 >> n == 0.
+Proof.
+ intros. rewrite shiftr_div_pow2. nzsimpl; order_nz.
+Qed.
+
+Lemma shiftl_eq_0_iff : forall a n, a << n == 0 <-> a == 0.
+Proof.
+ intros a n. rewrite shiftl_mul_pow2. rewrite eq_mul_0. split.
+ intros [H | H]; trivial. contradict H; order_nz.
+ intros H. now left.
+Qed.
+
+Lemma shiftr_eq_0_iff : forall a n,
+ a >> n == 0 <-> a==0 \/ (0<a /\ log2 a < n).
+Proof.
+ intros a n.
+ rewrite shiftr_div_pow2, div_small_iff by order_nz.
+ destruct (eq_0_gt_0_cases a) as [EQ|LT].
+ rewrite EQ. split. now left. intros _.
+ assert (H : 2~=0) by order'.
+ generalize (pow_nonzero 2 n H) (le_0_l (2^n)); order.
+ rewrite log2_lt_pow2; trivial.
+ split. right; split; trivial. intros [H|[_ H]]; now order.
+Qed.
+
+Lemma shiftr_eq_0 : forall a n, log2 a < n -> a >> n == 0.
+Proof.
+ intros a n H. rewrite shiftr_eq_0_iff.
+ destruct (eq_0_gt_0_cases a) as [EQ|LT]. now left. right; now split.
+Qed.
+
+(** Properties of [div2]. *)
+
+Lemma div2_div : forall a, div2 a == a/2.
+Proof.
+ intros. rewrite div2_spec, shiftr_div_pow2. now nzsimpl.
+Qed.
+
+Instance div2_wd : Proper (eq==>eq) div2.
+Proof.
+ intros a a' Ha. now rewrite 2 div2_div, Ha.
+Qed.
+
+Lemma div2_odd : forall a, a == 2*(div2 a) + odd a.
+Proof.
+ intros a. rewrite div2_div, <- bit0_odd, bit0_mod.
+ apply div_mod. order'.
+Qed.
+
+(** Properties of [lxor] and others, directly deduced
+ from properties of [xorb] and others. *)
+
+Instance lxor_wd : Proper (eq ==> eq ==> eq) lxor.
+Proof.
+ intros a a' Ha b b' Hb. bitwise. now rewrite Ha, Hb.
+Qed.
+
+Instance land_wd : Proper (eq ==> eq ==> eq) land.
+Proof.
+ intros a a' Ha b b' Hb. bitwise. now rewrite Ha, Hb.
+Qed.
+
+Instance lor_wd : Proper (eq ==> eq ==> eq) lor.
+Proof.
+ intros a a' Ha b b' Hb. bitwise. now rewrite Ha, Hb.
+Qed.
+
+Instance ldiff_wd : Proper (eq ==> eq ==> eq) ldiff.
+Proof.
+ intros a a' Ha b b' Hb. bitwise. now rewrite Ha, Hb.
+Qed.
+
+Lemma lxor_eq : forall a a', lxor a a' == 0 -> a == a'.
+Proof.
+ intros a a' H. bitwise. apply xorb_eq.
+ now rewrite <- lxor_spec, H, bits_0.
+Qed.
+
+Lemma lxor_nilpotent : forall a, lxor a a == 0.
+Proof.
+ intros. bitwise. apply xorb_nilpotent.
+Qed.
+
+Lemma lxor_eq_0_iff : forall a a', lxor a a' == 0 <-> a == a'.
+Proof.
+ split. apply lxor_eq. intros EQ; rewrite EQ; apply lxor_nilpotent.
+Qed.
+
+Lemma lxor_0_l : forall a, lxor 0 a == a.
+Proof.
+ intros. bitwise. apply xorb_false_l.
+Qed.
+
+Lemma lxor_0_r : forall a, lxor a 0 == a.
+Proof.
+ intros. bitwise. apply xorb_false_r.
+Qed.
+
+Lemma lxor_comm : forall a b, lxor a b == lxor b a.
+Proof.
+ intros. bitwise. apply xorb_comm.
+Qed.
+
+Lemma lxor_assoc :
+ forall a b c, lxor (lxor a b) c == lxor a (lxor b c).
+Proof.
+ intros. bitwise. apply xorb_assoc.
+Qed.
+
+Lemma lor_0_l : forall a, lor 0 a == a.
+Proof.
+ intros. bitwise. trivial.
+Qed.
+
+Lemma lor_0_r : forall a, lor a 0 == a.
+Proof.
+ intros. bitwise. apply orb_false_r.
+Qed.
+
+Lemma lor_comm : forall a b, lor a b == lor b a.
+Proof.
+ intros. bitwise. apply orb_comm.
+Qed.
+
+Lemma lor_assoc :
+ forall a b c, lor a (lor b c) == lor (lor a b) c.
+Proof.
+ intros. bitwise. apply orb_assoc.
+Qed.
+
+Lemma lor_diag : forall a, lor a a == a.
+Proof.
+ intros. bitwise. apply orb_diag.
+Qed.
+
+Lemma lor_eq_0_l : forall a b, lor a b == 0 -> a == 0.
+Proof.
+ intros a b H. bitwise.
+ apply (orb_false_iff a.[m] b.[m]).
+ now rewrite <- lor_spec, H, bits_0.
+Qed.
+
+Lemma lor_eq_0_iff : forall a b, lor a b == 0 <-> a == 0 /\ b == 0.
+Proof.
+ intros a b. split.
+ split. now apply lor_eq_0_l in H.
+ rewrite lor_comm in H. now apply lor_eq_0_l in H.
+ intros (EQ,EQ'). now rewrite EQ, lor_0_l.
+Qed.
+
+Lemma land_0_l : forall a, land 0 a == 0.
+Proof.
+ intros. bitwise. trivial.
+Qed.
+
+Lemma land_0_r : forall a, land a 0 == 0.
+Proof.
+ intros. bitwise. apply andb_false_r.
+Qed.
+
+Lemma land_comm : forall a b, land a b == land b a.
+Proof.
+ intros. bitwise. apply andb_comm.
+Qed.
+
+Lemma land_assoc :
+ forall a b c, land a (land b c) == land (land a b) c.
+Proof.
+ intros. bitwise. apply andb_assoc.
+Qed.
+
+Lemma land_diag : forall a, land a a == a.
+Proof.
+ intros. bitwise. apply andb_diag.
+Qed.
+
+Lemma ldiff_0_l : forall a, ldiff 0 a == 0.
+Proof.
+ intros. bitwise. trivial.
+Qed.
+
+Lemma ldiff_0_r : forall a, ldiff a 0 == a.
+Proof.
+ intros. bitwise. now rewrite andb_true_r.
+Qed.
+
+Lemma ldiff_diag : forall a, ldiff a a == 0.
+Proof.
+ intros. bitwise. apply andb_negb_r.
+Qed.
+
+Lemma lor_land_distr_l : forall a b c,
+ lor (land a b) c == land (lor a c) (lor b c).
+Proof.
+ intros. bitwise. apply orb_andb_distrib_l.
+Qed.
+
+Lemma lor_land_distr_r : forall a b c,
+ lor a (land b c) == land (lor a b) (lor a c).
+Proof.
+ intros. bitwise. apply orb_andb_distrib_r.
+Qed.
+
+Lemma land_lor_distr_l : forall a b c,
+ land (lor a b) c == lor (land a c) (land b c).
+Proof.
+ intros. bitwise. apply andb_orb_distrib_l.
+Qed.
+
+Lemma land_lor_distr_r : forall a b c,
+ land a (lor b c) == lor (land a b) (land a c).
+Proof.
+ intros. bitwise. apply andb_orb_distrib_r.
+Qed.
+
+Lemma ldiff_ldiff_l : forall a b c,
+ ldiff (ldiff a b) c == ldiff a (lor b c).
+Proof.
+ intros. bitwise. now rewrite negb_orb, andb_assoc.
+Qed.
+
+Lemma lor_ldiff_and : forall a b,
+ lor (ldiff a b) (land a b) == a.
+Proof.
+ intros. bitwise.
+ now rewrite <- andb_orb_distrib_r, orb_comm, orb_negb_r, andb_true_r.
+Qed.
+
+Lemma land_ldiff : forall a b,
+ land (ldiff a b) b == 0.
+Proof.
+ intros. bitwise.
+ now rewrite <-andb_assoc, (andb_comm (negb _)), andb_negb_r, andb_false_r.
+Qed.
+
+(** Properties of [setbit] and [clearbit] *)
+
+Definition setbit a n := lor a (1<<n).
+Definition clearbit a n := ldiff a (1<<n).
+
+Lemma setbit_spec' : forall a n, setbit a n == lor a (2^n).
+Proof.
+ intros. unfold setbit. now rewrite shiftl_1_l.
+Qed.
+
+Lemma clearbit_spec' : forall a n, clearbit a n == ldiff a (2^n).
+Proof.
+ intros. unfold clearbit. now rewrite shiftl_1_l.
+Qed.
+
+Instance setbit_wd : Proper (eq==>eq==>eq) setbit.
+Proof. unfold setbit. solve_proper. Qed.
+
+Instance clearbit_wd : Proper (eq==>eq==>eq) clearbit.
+Proof. unfold clearbit. solve_proper. Qed.
+
+Lemma pow2_bits_true : forall n, (2^n).[n] = true.
+Proof.
+ intros. rewrite <- (mul_1_l (2^n)). rewrite <- (add_0_l n) at 2.
+ now rewrite mul_pow2_bits_add, bit0_odd, odd_1.
+Qed.
+
+Lemma pow2_bits_false : forall n m, n~=m -> (2^n).[m] = false.
+Proof.
+ intros.
+ rewrite <- (mul_1_l (2^n)).
+ destruct (le_gt_cases n m).
+ rewrite mul_pow2_bits_high; trivial.
+ rewrite <- (succ_pred (m-n)) by (apply sub_gt; order).
+ now rewrite <- div2_bits, div_small, bits_0 by order'.
+ rewrite mul_pow2_bits_low; trivial.
+Qed.
+
+Lemma pow2_bits_eqb : forall n m, (2^n).[m] = eqb n m.
+Proof.
+ intros. apply eq_true_iff_eq. rewrite eqb_eq. split.
+ destruct (eq_decidable n m) as [H|H]. trivial.
+ now rewrite (pow2_bits_false _ _ H).
+ intros EQ. rewrite EQ. apply pow2_bits_true.
+Qed.
+
+Lemma setbit_eqb : forall a n m,
+ (setbit a n).[m] = eqb n m || a.[m].
+Proof.
+ intros. now rewrite setbit_spec', lor_spec, pow2_bits_eqb, orb_comm.
+Qed.
+
+Lemma setbit_iff : forall a n m,
+ (setbit a n).[m] = true <-> n==m \/ a.[m] = true.
+Proof.
+ intros. now rewrite setbit_eqb, orb_true_iff, eqb_eq.
+Qed.
+
+Lemma setbit_eq : forall a n, (setbit a n).[n] = true.
+Proof.
+ intros. apply setbit_iff. now left.
+Qed.
+
+Lemma setbit_neq : forall a n m, n~=m ->
+ (setbit a n).[m] = a.[m].
+Proof.
+ intros a n m H. rewrite setbit_eqb.
+ rewrite <- eqb_eq in H. apply not_true_is_false in H. now rewrite H.
+Qed.
+
+Lemma clearbit_eqb : forall a n m,
+ (clearbit a n).[m] = a.[m] && negb (eqb n m).
+Proof.
+ intros. now rewrite clearbit_spec', ldiff_spec, pow2_bits_eqb.
+Qed.
+
+Lemma clearbit_iff : forall a n m,
+ (clearbit a n).[m] = true <-> a.[m] = true /\ n~=m.
+Proof.
+ intros. rewrite clearbit_eqb, andb_true_iff, <- eqb_eq.
+ now rewrite negb_true_iff, not_true_iff_false.
+Qed.
+
+Lemma clearbit_eq : forall a n, (clearbit a n).[n] = false.
+Proof.
+ intros. rewrite clearbit_eqb, (proj2 (eqb_eq _ _) (eq_refl n)).
+ apply andb_false_r.
+Qed.
+
+Lemma clearbit_neq : forall a n m, n~=m ->
+ (clearbit a n).[m] = a.[m].
+Proof.
+ intros a n m H. rewrite clearbit_eqb.
+ rewrite <- eqb_eq in H. apply not_true_is_false in H. rewrite H.
+ apply andb_true_r.
+Qed.
+
+(** Shifts of bitwise operations *)
+
+Lemma shiftl_lxor : forall a b n,
+ (lxor a b) << n == lxor (a << n) (b << n).
+Proof.
+ intros. bitwise.
+ destruct (le_gt_cases n m).
+ now rewrite !shiftl_spec_high', lxor_spec.
+ now rewrite !shiftl_spec_low.
+Qed.
+
+Lemma shiftr_lxor : forall a b n,
+ (lxor a b) >> n == lxor (a >> n) (b >> n).
+Proof.
+ intros. bitwise. now rewrite !shiftr_spec', lxor_spec.
+Qed.
+
+Lemma shiftl_land : forall a b n,
+ (land a b) << n == land (a << n) (b << n).
+Proof.
+ intros. bitwise.
+ destruct (le_gt_cases n m).
+ now rewrite !shiftl_spec_high', land_spec.
+ now rewrite !shiftl_spec_low.
+Qed.
+
+Lemma shiftr_land : forall a b n,
+ (land a b) >> n == land (a >> n) (b >> n).
+Proof.
+ intros. bitwise. now rewrite !shiftr_spec', land_spec.
+Qed.
+
+Lemma shiftl_lor : forall a b n,
+ (lor a b) << n == lor (a << n) (b << n).
+Proof.
+ intros. bitwise.
+ destruct (le_gt_cases n m).
+ now rewrite !shiftl_spec_high', lor_spec.
+ now rewrite !shiftl_spec_low.
+Qed.
+
+Lemma shiftr_lor : forall a b n,
+ (lor a b) >> n == lor (a >> n) (b >> n).
+Proof.
+ intros. bitwise. now rewrite !shiftr_spec', lor_spec.
+Qed.
+
+Lemma shiftl_ldiff : forall a b n,
+ (ldiff a b) << n == ldiff (a << n) (b << n).
+Proof.
+ intros. bitwise.
+ destruct (le_gt_cases n m).
+ now rewrite !shiftl_spec_high', ldiff_spec.
+ now rewrite !shiftl_spec_low.
+Qed.
+
+Lemma shiftr_ldiff : forall a b n,
+ (ldiff a b) >> n == ldiff (a >> n) (b >> n).
+Proof.
+ intros. bitwise. now rewrite !shiftr_spec', ldiff_spec.
+Qed.
+
+(** We cannot have a function complementing all bits of a number,
+ otherwise it would have an infinity of bit 1. Nonetheless,
+ we can design a bounded complement *)
+
+Definition ones n := P (1 << n).
+
+Definition lnot a n := lxor a (ones n).
+
+Instance ones_wd : Proper (eq==>eq) ones.
+Proof. unfold ones. solve_proper. Qed.
+
+Instance lnot_wd : Proper (eq==>eq==>eq) lnot.
+Proof. unfold lnot. solve_proper. Qed.
+
+Lemma ones_equiv : forall n, ones n == P (2^n).
+Proof.
+ intros; unfold ones; now rewrite shiftl_1_l.
+Qed.
+
+Lemma ones_add : forall n m, ones (m+n) == 2^m * ones n + ones m.
+Proof.
+ intros n m. rewrite !ones_equiv.
+ rewrite <- !sub_1_r, mul_sub_distr_l, mul_1_r, <- pow_add_r.
+ rewrite add_sub_assoc, sub_add. reflexivity.
+ apply pow_le_mono_r. order'.
+ rewrite <- (add_0_r m) at 1. apply add_le_mono_l, le_0_l.
+ rewrite <- (pow_0_r 2). apply pow_le_mono_r. order'. apply le_0_l.
+Qed.
+
+Lemma ones_div_pow2 : forall n m, m<=n -> ones n / 2^m == ones (n-m).
+Proof.
+ intros n m H. symmetry. apply div_unique with (ones m).
+ rewrite ones_equiv.
+ apply le_succ_l. rewrite succ_pred; order_nz.
+ rewrite <- (sub_add m n H) at 1. rewrite (add_comm _ m).
+ apply ones_add.
+Qed.
+
+Lemma ones_mod_pow2 : forall n m, m<=n -> (ones n) mod (2^m) == ones m.
+Proof.
+ intros n m H. symmetry. apply mod_unique with (ones (n-m)).
+ rewrite ones_equiv.
+ apply le_succ_l. rewrite succ_pred; order_nz.
+ rewrite <- (sub_add m n H) at 1. rewrite (add_comm _ m).
+ apply ones_add.
+Qed.
+
+Lemma ones_spec_low : forall n m, m<n -> (ones n).[m] = true.
+Proof.
+ intros. apply testbit_true. rewrite ones_div_pow2 by order.
+ rewrite <- (pow_1_r 2). rewrite ones_mod_pow2.
+ rewrite ones_equiv. now nzsimpl'.
+ apply le_add_le_sub_r. nzsimpl. now apply le_succ_l.
+Qed.
+
+Lemma ones_spec_high : forall n m, n<=m -> (ones n).[m] = false.
+Proof.
+ intros.
+ destruct (eq_0_gt_0_cases n) as [EQ|LT]; rewrite ones_equiv.
+ now rewrite EQ, pow_0_r, one_succ, pred_succ, bits_0.
+ apply bits_above_log2.
+ rewrite log2_pred_pow2; trivial. rewrite <-le_succ_l, succ_pred; order.
+Qed.
+
+Lemma ones_spec_iff : forall n m, (ones n).[m] = true <-> m<n.
+Proof.
+ intros. split. intros H.
+ apply lt_nge. intro H'. apply ones_spec_high in H'.
+ rewrite H in H'; discriminate.
+ apply ones_spec_low.
+Qed.
+
+Lemma lnot_spec_low : forall a n m, m<n ->
+ (lnot a n).[m] = negb a.[m].
+Proof.
+ intros. unfold lnot. now rewrite lxor_spec, ones_spec_low.
+Qed.
+
+Lemma lnot_spec_high : forall a n m, n<=m ->
+ (lnot a n).[m] = a.[m].
+Proof.
+ intros. unfold lnot. now rewrite lxor_spec, ones_spec_high, xorb_false_r.
+Qed.
+
+Lemma lnot_involutive : forall a n, lnot (lnot a n) n == a.
+Proof.
+ intros a n. bitwise.
+ destruct (le_gt_cases n m).
+ now rewrite 2 lnot_spec_high.
+ now rewrite 2 lnot_spec_low, negb_involutive.
+Qed.
+
+Lemma lnot_0_l : forall n, lnot 0 n == ones n.
+Proof.
+ intros. unfold lnot. apply lxor_0_l.
+Qed.
+
+Lemma lnot_ones : forall n, lnot (ones n) n == 0.
+Proof.
+ intros. unfold lnot. apply lxor_nilpotent.
+Qed.
+
+(** Bounded complement and other operations *)
+
+Lemma lor_ones_low : forall a n, log2 a < n ->
+ lor a (ones n) == ones n.
+Proof.
+ intros a n H. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ now rewrite ones_spec_low, orb_true_r.
+Qed.
+
+Lemma land_ones : forall a n, land a (ones n) == a mod 2^n.
+Proof.
+ intros a n. bitwise. destruct (le_gt_cases n m).
+ now rewrite ones_spec_high, mod_pow2_bits_high, andb_false_r.
+ now rewrite ones_spec_low, mod_pow2_bits_low, andb_true_r.
+Qed.
+
+Lemma land_ones_low : forall a n, log2 a < n ->
+ land a (ones n) == a.
+Proof.
+ intros; rewrite land_ones. apply mod_small.
+ apply log2_lt_cancel. rewrite log2_pow2; trivial using le_0_l.
+Qed.
+
+Lemma ldiff_ones_r : forall a n,
+ ldiff a (ones n) == (a >> n) << n.
+Proof.
+ intros a n. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, shiftl_spec_high', shiftr_spec'; trivial.
+ rewrite sub_add; trivial. apply andb_true_r.
+ now rewrite ones_spec_low, shiftl_spec_low, andb_false_r.
+Qed.
+
+Lemma ldiff_ones_r_low : forall a n, log2 a < n ->
+ ldiff a (ones n) == 0.
+Proof.
+ intros a n H. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ now rewrite ones_spec_low, andb_false_r.
+Qed.
+
+Lemma ldiff_ones_l_low : forall a n, log2 a < n ->
+ ldiff (ones n) a == lnot a n.
+Proof.
+ intros a n H. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, lnot_spec_high, bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ now rewrite ones_spec_low, lnot_spec_low.
+Qed.
+
+Lemma lor_lnot_diag : forall a n,
+ lor a (lnot a n) == lor a (ones n).
+Proof.
+ intros a n. bitwise.
+ destruct (le_gt_cases n m).
+ rewrite lnot_spec_high, ones_spec_high; trivial. now destruct a.[m].
+ rewrite lnot_spec_low, ones_spec_low; trivial. now destruct a.[m].
+Qed.
+
+Lemma lor_lnot_diag_low : forall a n, log2 a < n ->
+ lor a (lnot a n) == ones n.
+Proof.
+ intros a n H. now rewrite lor_lnot_diag, lor_ones_low.
+Qed.
+
+Lemma land_lnot_diag : forall a n,
+ land a (lnot a n) == ldiff a (ones n).
+Proof.
+ intros a n. bitwise.
+ destruct (le_gt_cases n m).
+ rewrite lnot_spec_high, ones_spec_high; trivial. now destruct a.[m].
+ rewrite lnot_spec_low, ones_spec_low; trivial. now destruct a.[m].
+Qed.
+
+Lemma land_lnot_diag_low : forall a n, log2 a < n ->
+ land a (lnot a n) == 0.
+Proof.
+ intros. now rewrite land_lnot_diag, ldiff_ones_r_low.
+Qed.
+
+Lemma lnot_lor_low : forall a b n, log2 a < n -> log2 b < n ->
+ lnot (lor a b) n == land (lnot a n) (lnot b n).
+Proof.
+ intros a b n Ha Hb. bitwise. destruct (le_gt_cases n m).
+ rewrite !lnot_spec_high, lor_spec, !bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ now apply lt_le_trans with n.
+ now rewrite !lnot_spec_low, lor_spec, negb_orb.
+Qed.
+
+Lemma lnot_land_low : forall a b n, log2 a < n -> log2 b < n ->
+ lnot (land a b) n == lor (lnot a n) (lnot b n).
+Proof.
+ intros a b n Ha Hb. bitwise. destruct (le_gt_cases n m).
+ rewrite !lnot_spec_high, land_spec, !bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ now apply lt_le_trans with n.
+ now rewrite !lnot_spec_low, land_spec, negb_andb.
+Qed.
+
+Lemma ldiff_land_low : forall a b n, log2 a < n ->
+ ldiff a b == land a (lnot b n).
+Proof.
+ intros a b n Ha. bitwise. destruct (le_gt_cases n m).
+ rewrite (bits_above_log2 a m). trivial.
+ now apply lt_le_trans with n.
+ rewrite !lnot_spec_low; trivial.
+Qed.
+
+Lemma lnot_ldiff_low : forall a b n, log2 a < n -> log2 b < n ->
+ lnot (ldiff a b) n == lor (lnot a n) b.
+Proof.
+ intros a b n Ha Hb. bitwise. destruct (le_gt_cases n m).
+ rewrite !lnot_spec_high, ldiff_spec, !bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ now apply lt_le_trans with n.
+ now rewrite !lnot_spec_low, ldiff_spec, negb_andb, negb_involutive.
+Qed.
+
+Lemma lxor_lnot_lnot : forall a b n,
+ lxor (lnot a n) (lnot b n) == lxor a b.
+Proof.
+ intros a b n. bitwise. destruct (le_gt_cases n m).
+ rewrite !lnot_spec_high; trivial.
+ rewrite !lnot_spec_low, xorb_negb_negb; trivial.
+Qed.
+
+Lemma lnot_lxor_l : forall a b n,
+ lnot (lxor a b) n == lxor (lnot a n) b.
+Proof.
+ intros a b n. bitwise. destruct (le_gt_cases n m).
+ rewrite !lnot_spec_high, lxor_spec; trivial.
+ rewrite !lnot_spec_low, lxor_spec, negb_xorb_l; trivial.
+Qed.
+
+Lemma lnot_lxor_r : forall a b n,
+ lnot (lxor a b) n == lxor a (lnot b n).
+Proof.
+ intros a b n. bitwise. destruct (le_gt_cases n m).
+ rewrite !lnot_spec_high, lxor_spec; trivial.
+ rewrite !lnot_spec_low, lxor_spec, negb_xorb_r; trivial.
+Qed.
+
+Lemma lxor_lor : forall a b, land a b == 0 ->
+ lxor a b == lor a b.
+Proof.
+ intros a b H. bitwise.
+ assert (a.[m] && b.[m] = false)
+ by now rewrite <- land_spec, H, bits_0.
+ now destruct a.[m], b.[m].
+Qed.
+
+(** Bitwise operations and log2 *)
+
+Lemma log2_bits_unique : forall a n,
+ a.[n] = true ->
+ (forall m, n<m -> a.[m] = false) ->
+ log2 a == n.
+Proof.
+ intros a n H H'.
+ destruct (eq_0_gt_0_cases a) as [Ha|Ha].
+ now rewrite Ha, bits_0 in H.
+ apply le_antisymm; apply le_ngt; intros LT.
+ specialize (H' _ LT). now rewrite bit_log2 in H' by order.
+ now rewrite bits_above_log2 in H by order.
+Qed.
+
+Lemma log2_shiftr : forall a n, log2 (a >> n) == log2 a - n.
+Proof.
+ intros a n.
+ destruct (eq_0_gt_0_cases a) as [Ha|Ha].
+ now rewrite Ha, shiftr_0_l, log2_nonpos, sub_0_l by order.
+ destruct (lt_ge_cases (log2 a) n).
+ rewrite shiftr_eq_0, log2_nonpos by order.
+ symmetry. rewrite sub_0_le; order.
+ apply log2_bits_unique.
+ now rewrite shiftr_spec', sub_add, bit_log2 by order.
+ intros m Hm.
+ rewrite shiftr_spec'; trivial. apply bits_above_log2; try order.
+ now apply lt_sub_lt_add_r.
+Qed.
+
+Lemma log2_shiftl : forall a n, a~=0 -> log2 (a << n) == log2 a + n.
+Proof.
+ intros a n Ha.
+ rewrite shiftl_mul_pow2, add_comm by trivial.
+ apply log2_mul_pow2. generalize (le_0_l a); order. apply le_0_l.
+Qed.
+
+Lemma log2_lor : forall a b,
+ log2 (lor a b) == max (log2 a) (log2 b).
+Proof.
+ assert (AUX : forall a b, a<=b -> log2 (lor a b) == log2 b).
+ intros a b H.
+ destruct (eq_0_gt_0_cases a) as [Ha|Ha]. now rewrite Ha, lor_0_l.
+ apply log2_bits_unique.
+ now rewrite lor_spec, bit_log2, orb_true_r by order.
+ intros m Hm. assert (H' := log2_le_mono _ _ H).
+ now rewrite lor_spec, 2 bits_above_log2 by order.
+ (* main *)
+ intros a b. destruct (le_ge_cases a b) as [H|H].
+ rewrite max_r by now apply log2_le_mono.
+ now apply AUX.
+ rewrite max_l by now apply log2_le_mono.
+ rewrite lor_comm. now apply AUX.
+Qed.
+
+Lemma log2_land : forall a b,
+ log2 (land a b) <= min (log2 a) (log2 b).
+Proof.
+ assert (AUX : forall a b, a<=b -> log2 (land a b) <= log2 a).
+ intros a b H.
+ apply le_ngt. intros H'.
+ destruct (eq_decidable (land a b) 0) as [EQ|NEQ].
+ rewrite EQ in H'. apply log2_lt_cancel in H'. generalize (le_0_l a); order.
+ generalize (bit_log2 (land a b) NEQ).
+ now rewrite land_spec, bits_above_log2.
+ (* main *)
+ intros a b.
+ destruct (le_ge_cases a b) as [H|H].
+ rewrite min_l by now apply log2_le_mono. now apply AUX.
+ rewrite min_r by now apply log2_le_mono. rewrite land_comm. now apply AUX.
+Qed.
+
+Lemma log2_lxor : forall a b,
+ log2 (lxor a b) <= max (log2 a) (log2 b).
+Proof.
+ assert (AUX : forall a b, a<=b -> log2 (lxor a b) <= log2 b).
+ intros a b H.
+ apply le_ngt. intros H'.
+ destruct (eq_decidable (lxor a b) 0) as [EQ|NEQ].
+ rewrite EQ in H'. apply log2_lt_cancel in H'. generalize (le_0_l a); order.
+ generalize (bit_log2 (lxor a b) NEQ).
+ rewrite lxor_spec, 2 bits_above_log2; try order. discriminate.
+ apply le_lt_trans with (log2 b); trivial. now apply log2_le_mono.
+ (* main *)
+ intros a b.
+ destruct (le_ge_cases a b) as [H|H].
+ rewrite max_r by now apply log2_le_mono. now apply AUX.
+ rewrite max_l by now apply log2_le_mono. rewrite lxor_comm. now apply AUX.
+Qed.
+
+(** Bitwise operations and arithmetical operations *)
+
+Local Notation xor3 a b c := (xorb (xorb a b) c).
+Local Notation lxor3 a b c := (lxor (lxor a b) c).
+
+Local Notation nextcarry a b c := ((a&&b) || (c && (a||b))).
+Local Notation lnextcarry a b c := (lor (land a b) (land c (lor a b))).
+
+Lemma add_bit0 : forall a b, (a+b).[0] = xorb a.[0] b.[0].
+Proof.
+ intros. now rewrite !bit0_odd, odd_add.
+Qed.
+
+Lemma add3_bit0 : forall a b c,
+ (a+b+c).[0] = xor3 a.[0] b.[0] c.[0].
+Proof.
+ intros. now rewrite !add_bit0.
+Qed.
+
+Lemma add3_bits_div2 : forall (a0 b0 c0 : bool),
+ (a0 + b0 + c0)/2 == nextcarry a0 b0 c0.
+Proof.
+ assert (H : 1+1 == 2) by now nzsimpl'.
+ intros [|] [|] [|]; simpl; rewrite ?add_0_l, ?add_0_r, ?H;
+ (apply div_same; order') || (apply div_small; order') || idtac.
+ symmetry. apply div_unique with 1. order'. now nzsimpl'.
+Qed.
+
+Lemma add_carry_div2 : forall a b (c0:bool),
+ (a + b + c0)/2 == a/2 + b/2 + nextcarry a.[0] b.[0] c0.
+Proof.
+ intros.
+ rewrite <- add3_bits_div2.
+ rewrite (add_comm ((a/2)+_)).
+ rewrite <- div_add by order'.
+ f_equiv.
+ rewrite <- !div2_div, mul_comm, mul_add_distr_l.
+ rewrite (div2_odd a), <- bit0_odd at 1. fold (b2n a.[0]).
+ rewrite (div2_odd b), <- bit0_odd at 1. fold (b2n b.[0]).
+ rewrite add_shuffle1.
+ rewrite <-(add_assoc _ _ c0). apply add_comm.
+Qed.
+
+(** The main result concerning addition: we express the bits of the sum
+ in term of bits of [a] and [b] and of some carry stream which is also
+ recursively determined by another equation.
+*)
+
+Lemma add_carry_bits : forall a b (c0:bool), exists c,
+ a+b+c0 == lxor3 a b c /\ c/2 == lnextcarry a b c /\ c.[0] = c0.
+Proof.
+ intros a b c0.
+ (* induction over some n such that [a<2^n] and [b<2^n] *)
+ set (n:=max a b).
+ assert (Ha : a<2^n).
+ apply lt_le_trans with (2^a). apply pow_gt_lin_r, lt_1_2.
+ apply pow_le_mono_r. order'. unfold n.
+ destruct (le_ge_cases a b); [rewrite max_r|rewrite max_l]; order'.
+ assert (Hb : b<2^n).
+ apply lt_le_trans with (2^b). apply pow_gt_lin_r, lt_1_2.
+ apply pow_le_mono_r. order'. unfold n.
+ destruct (le_ge_cases a b); [rewrite max_r|rewrite max_l]; order'.
+ clearbody n.
+ revert a b c0 Ha Hb. induct n.
+ (*base*)
+ intros a b c0. rewrite !pow_0_r, !one_succ, !lt_succ_r. intros Ha Hb.
+ exists c0.
+ setoid_replace a with 0 by (generalize (le_0_l a); order').
+ setoid_replace b with 0 by (generalize (le_0_l b); order').
+ rewrite !add_0_l, !lxor_0_l, !lor_0_r, !land_0_r, !lor_0_r.
+ rewrite b2n_div2, b2n_bit0; now repeat split.
+ (*step*)
+ intros n IH a b c0 Ha Hb.
+ set (c1:=nextcarry a.[0] b.[0] c0).
+ destruct (IH (a/2) (b/2) c1) as (c & IH1 & IH2 & Hc); clear IH.
+ apply div_lt_upper_bound; trivial. order'. now rewrite <- pow_succ_r'.
+ apply div_lt_upper_bound; trivial. order'. now rewrite <- pow_succ_r'.
+ exists (c0 + 2*c). repeat split.
+ (* - add *)
+ bitwise.
+ destruct (zero_or_succ m) as [EQ|[m' EQ]]; rewrite EQ; clear EQ.
+ now rewrite add_b2n_double_bit0, add3_bit0, b2n_bit0.
+ rewrite <- !div2_bits, <- 2 lxor_spec.
+ f_equiv.
+ rewrite add_b2n_double_div2, <- IH1. apply add_carry_div2.
+ (* - carry *)
+ rewrite add_b2n_double_div2.
+ bitwise.
+ destruct (zero_or_succ m) as [EQ|[m' EQ]]; rewrite EQ; clear EQ.
+ now rewrite add_b2n_double_bit0.
+ rewrite <- !div2_bits, IH2. autorewrite with bitwise.
+ now rewrite add_b2n_double_div2.
+ (* - carry0 *)
+ apply add_b2n_double_bit0.
+Qed.
+
+(** Particular case : the second bit of an addition *)
+
+Lemma add_bit1 : forall a b,
+ (a+b).[1] = xor3 a.[1] b.[1] (a.[0] && b.[0]).
+Proof.
+ intros a b.
+ destruct (add_carry_bits a b false) as (c & EQ1 & EQ2 & Hc).
+ simpl in EQ1; rewrite add_0_r in EQ1. rewrite EQ1.
+ autorewrite with bitwise. f_equal.
+ rewrite one_succ, <- div2_bits, EQ2.
+ autorewrite with bitwise.
+ rewrite Hc. simpl. apply orb_false_r.
+Qed.
+
+(** In an addition, there will be no carries iff there is
+ no common bits in the numbers to add *)
+
+Lemma nocarry_equiv : forall a b c,
+ c/2 == lnextcarry a b c -> c.[0] = false ->
+ (c == 0 <-> land a b == 0).
+Proof.
+ intros a b c H H'.
+ split. intros EQ; rewrite EQ in *.
+ rewrite div_0_l in H by order'.
+ symmetry in H. now apply lor_eq_0_l in H.
+ intros EQ. rewrite EQ, lor_0_l in H.
+ apply bits_inj_0.
+ induct n. trivial.
+ intros n IH.
+ rewrite <- div2_bits, H.
+ autorewrite with bitwise.
+ now rewrite IH.
+Qed.
+
+(** When there is no common bits, the addition is just a xor *)
+
+Lemma add_nocarry_lxor : forall a b, land a b == 0 ->
+ a+b == lxor a b.
+Proof.
+ intros a b H.
+ destruct (add_carry_bits a b false) as (c & EQ1 & EQ2 & Hc).
+ simpl in EQ1; rewrite add_0_r in EQ1. rewrite EQ1.
+ apply (nocarry_equiv a b c) in H; trivial.
+ rewrite H. now rewrite lxor_0_r.
+Qed.
+
+(** A null [ldiff] implies being smaller *)
+
+Lemma ldiff_le : forall a b, ldiff a b == 0 -> a <= b.
+Proof.
+ cut (forall n a b, a < 2^n -> ldiff a b == 0 -> a <= b).
+ intros H a b. apply (H a), pow_gt_lin_r; order'.
+ induct n.
+ intros a b Ha _. rewrite pow_0_r, one_succ, lt_succ_r in Ha.
+ assert (Ha' : a == 0) by (generalize (le_0_l a); order').
+ rewrite Ha'. apply le_0_l.
+ intros n IH a b Ha H.
+ assert (NEQ : 2 ~= 0) by order'.
+ rewrite (div_mod a 2 NEQ), (div_mod b 2 NEQ).
+ apply add_le_mono.
+ apply mul_le_mono_l.
+ apply IH.
+ apply div_lt_upper_bound; trivial. now rewrite <- pow_succ_r'.
+ rewrite <- (pow_1_r 2), <- 2 shiftr_div_pow2.
+ now rewrite <- shiftr_ldiff, H, shiftr_div_pow2, pow_1_r, div_0_l.
+ rewrite <- 2 bit0_mod.
+ apply bits_inj_iff in H. specialize (H 0).
+ rewrite ldiff_spec, bits_0 in H.
+ destruct a.[0], b.[0]; try discriminate; simpl; order'.
+Qed.
+
+(** Subtraction can be a ldiff when the opposite ldiff is null. *)
+
+Lemma sub_nocarry_ldiff : forall a b, ldiff b a == 0 ->
+ a-b == ldiff a b.
+Proof.
+ intros a b H.
+ apply add_cancel_r with b.
+ rewrite sub_add.
+ symmetry.
+ rewrite add_nocarry_lxor.
+ bitwise.
+ apply bits_inj_iff in H. specialize (H m).
+ rewrite ldiff_spec, bits_0 in H.
+ now destruct a.[m], b.[m].
+ apply land_ldiff.
+ now apply ldiff_le.
+Qed.
+
+(** We can express lnot in term of subtraction *)
+
+Lemma add_lnot_diag_low : forall a n, log2 a < n ->
+ a + lnot a n == ones n.
+Proof.
+ intros a n H.
+ assert (H' := land_lnot_diag_low a n H).
+ rewrite add_nocarry_lxor, lxor_lor by trivial.
+ now apply lor_lnot_diag_low.
+Qed.
+
+Lemma lnot_sub_low : forall a n, log2 a < n ->
+ lnot a n == ones n - a.
+Proof.
+ intros a n H.
+ now rewrite <- (add_lnot_diag_low a n H), add_comm, add_sub.
+Qed.
+
+(** Adding numbers with no common bits cannot lead to a much bigger number *)
+
+Lemma add_nocarry_lt_pow2 : forall a b n, land a b == 0 ->
+ a < 2^n -> b < 2^n -> a+b < 2^n.
+Proof.
+ intros a b n H Ha Hb.
+ rewrite add_nocarry_lxor by trivial.
+ apply div_small_iff. order_nz.
+ rewrite <- shiftr_div_pow2, shiftr_lxor, !shiftr_div_pow2.
+ rewrite 2 div_small by trivial.
+ apply lxor_0_l.
+Qed.
+
+Lemma add_nocarry_mod_lt_pow2 : forall a b n, land a b == 0 ->
+ a mod 2^n + b mod 2^n < 2^n.
+Proof.
+ intros a b n H.
+ apply add_nocarry_lt_pow2.
+ bitwise.
+ destruct (le_gt_cases n m).
+ now rewrite mod_pow2_bits_high.
+ now rewrite !mod_pow2_bits_low, <- land_spec, H, bits_0.
+ apply mod_upper_bound; order_nz.
+ apply mod_upper_bound; order_nz.
+Qed.
+
+End NBitsProp.
diff --git a/theories/Numbers/Natural/Abstract/NDefOps.v b/theories/Numbers/Natural/Abstract/NDefOps.v
index 7b38c148..ad7a9f3a 100644
--- a/theories/Numbers/Natural/Abstract/NDefOps.v
+++ b/theories/Numbers/Natural/Abstract/NDefOps.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,14 +8,41 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NDefOps.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import Bool. (* To get the orb and negb function *)
Require Import RelationPairs.
Require Export NStrongRec.
-Module NdefOpsPropFunct (Import N : NAxiomsSig').
-Include NStrongRecPropFunct N.
+(** In this module, we derive generic implementations of usual operators
+ just via the use of a [recursion] function. *)
+
+Module NdefOpsProp (Import N : NAxiomsRecSig').
+Include NStrongRecProp N.
+
+(** Nullity Test *)
+
+Definition if_zero (A : Type) (a b : A) (n : N.t) : A :=
+ recursion a (fun _ _ => b) n.
+
+Arguments if_zero [A] a b n.
+
+Instance if_zero_wd (A : Type) :
+ Proper (Logic.eq ==> Logic.eq ==> N.eq ==> Logic.eq) (@if_zero A).
+Proof.
+unfold if_zero. (* TODO : solve_proper : SLOW + BUG *)
+f_equiv'.
+Qed.
+
+Theorem if_zero_0 : forall (A : Type) (a b : A), if_zero a b 0 = a.
+Proof.
+unfold if_zero; intros; now rewrite recursion_0.
+Qed.
+
+Theorem if_zero_succ :
+ forall (A : Type) (a b : A) (n : N.t), if_zero a b (S n) = b.
+Proof.
+intros; unfold if_zero.
+now rewrite recursion_succ.
+Qed.
(*****************************************************)
(** Addition *)
@@ -24,17 +51,9 @@ Definition def_add (x y : N.t) := recursion y (fun _ => S) x.
Local Infix "+++" := def_add (at level 50, left associativity).
-Instance def_add_prewd : Proper (N.eq==>N.eq==>N.eq) (fun _ => S).
-Proof.
-intros _ _ _ p p' Epp'; now rewrite Epp'.
-Qed.
-
Instance def_add_wd : Proper (N.eq ==> N.eq ==> N.eq) def_add.
Proof.
-intros x x' Exx' y y' Eyy'. unfold def_add.
-(* TODO: why rewrite Exx' don't work here (or verrrry slowly) ? *)
-apply recursion_wd with (Aeq := N.eq); auto with *.
-apply def_add_prewd.
+unfold def_add. f_equiv'.
Qed.
Theorem def_add_0_l : forall y, 0 +++ y == y.
@@ -45,7 +64,7 @@ Qed.
Theorem def_add_succ_l : forall x y, S x +++ y == S (x +++ y).
Proof.
intros x y; unfold def_add.
-rewrite recursion_succ; auto with *.
+rewrite recursion_succ; f_equiv'.
Qed.
Theorem def_add_add : forall n m, n +++ m == n + m.
@@ -62,18 +81,10 @@ Definition def_mul (x y : N.t) := recursion 0 (fun _ p => p +++ x) y.
Local Infix "**" := def_mul (at level 40, left associativity).
-Instance def_mul_prewd :
- Proper (N.eq==>N.eq==>N.eq==>N.eq) (fun x _ p => p +++ x).
-Proof.
-repeat red; intros; now apply def_add_wd.
-Qed.
-
Instance def_mul_wd : Proper (N.eq ==> N.eq ==> N.eq) def_mul.
Proof.
-unfold def_mul.
-intros x x' Exx' y y' Eyy'.
-apply recursion_wd; auto with *.
-now apply def_mul_prewd.
+unfold def_mul. (* TODO : solve_proper SLOW + BUG *)
+f_equiv'.
Qed.
Theorem def_mul_0_r : forall x, x ** 0 == 0.
@@ -85,7 +96,7 @@ Theorem def_mul_succ_r : forall x y, x ** S y == x ** y +++ x.
Proof.
intros x y; unfold def_mul.
rewrite recursion_succ; auto with *.
-now apply def_mul_prewd.
+f_equiv'.
Qed.
Theorem def_mul_mul : forall n m, n ** m == n * m.
@@ -106,25 +117,9 @@ recursion
Local Infix "<<" := ltb (at level 70, no associativity).
-Instance ltb_prewd1 : Proper (N.eq==>Logic.eq) (if_zero false true).
-Proof.
-red; intros; apply if_zero_wd; auto.
-Qed.
-
-Instance ltb_prewd2 : Proper (N.eq==>(N.eq==>Logic.eq)==>N.eq==>Logic.eq)
- (fun _ f n => recursion false (fun n' _ => f n') n).
-Proof.
-repeat red; intros; simpl.
-apply recursion_wd; auto with *.
-repeat red; auto.
-Qed.
-
Instance ltb_wd : Proper (N.eq ==> N.eq ==> Logic.eq) ltb.
Proof.
-unfold ltb.
-intros n n' Hn m m' Hm.
-apply f_equiv; auto with *.
-apply recursion_wd; auto; [ apply ltb_prewd1 | apply ltb_prewd2 ].
+unfold ltb. f_equiv'.
Qed.
Theorem ltb_base : forall n, 0 << n = if_zero false true n.
@@ -136,11 +131,9 @@ Theorem ltb_step :
forall m n, S m << n = recursion false (fun n' _ => m << n') n.
Proof.
intros m n; unfold ltb at 1.
-apply f_equiv; auto with *.
-rewrite recursion_succ by (apply ltb_prewd1||apply ltb_prewd2).
-fold (ltb m).
-repeat red; intros. apply recursion_wd; auto.
-repeat red; intros; now apply ltb_wd.
+f_equiv.
+rewrite recursion_succ; f_equiv'.
+reflexivity.
Qed.
(* Above, we rewrite applications of function. Is it possible to rewrite
@@ -162,8 +155,7 @@ Qed.
Theorem succ_ltb_mono : forall n m, (S n << S m) = (n << m).
Proof.
intros n m.
-rewrite ltb_step. rewrite recursion_succ; try reflexivity.
-repeat red; intros; now apply ltb_wd.
+rewrite ltb_step. rewrite recursion_succ; f_equiv'.
Qed.
Theorem ltb_lt : forall n m, n << m = true <-> n < m.
@@ -188,9 +180,7 @@ Definition even (x : N.t) := recursion true (fun _ p => negb p) x.
Instance even_wd : Proper (N.eq==>Logic.eq) even.
Proof.
-intros n n' Hn. unfold even.
-apply recursion_wd; auto.
-congruence.
+unfold even. f_equiv'.
Qed.
Theorem even_0 : even 0 = true.
@@ -202,19 +192,12 @@ Qed.
Theorem even_succ : forall x, even (S x) = negb (even x).
Proof.
unfold even.
-intro x; rewrite recursion_succ; try reflexivity.
-congruence.
+intro x; rewrite recursion_succ; f_equiv'.
Qed.
(*****************************************************)
(** Division by 2 *)
-Local Notation "a <= b <= c" := (a<=b /\ b<=c).
-Local Notation "a <= b < c" := (a<=b /\ b<c).
-Local Notation "a < b <= c" := (a<b /\ b<=c).
-Local Notation "a < b < c" := (a<b /\ b<c).
-Local Notation "2" := (S 1).
-
Definition half_aux (x : N.t) : N.t * N.t :=
recursion (0, 0) (fun _ p => let (x1, x2) := p in (S x2, x1)) x.
@@ -223,14 +206,14 @@ Definition half (x : N.t) := snd (half_aux x).
Instance half_aux_wd : Proper (N.eq ==> N.eq*N.eq) half_aux.
Proof.
intros x x' Hx. unfold half_aux.
-apply recursion_wd; auto with *.
+f_equiv; trivial.
intros y y' Hy (u,v) (u',v') (Hu,Hv). compute in *.
rewrite Hu, Hv; auto with *.
Qed.
Instance half_wd : Proper (N.eq==>N.eq) half.
Proof.
-intros x x' Hx. unfold half. rewrite Hx; auto with *.
+unfold half. f_equiv'.
Qed.
Lemma half_aux_0 : half_aux 0 = (0,0).
@@ -245,8 +228,7 @@ intros.
remember (half_aux x) as h.
destruct h as (f,s); simpl in *.
unfold half_aux in *.
-rewrite recursion_succ, <- Heqh; simpl; auto.
-repeat red; intros; subst; auto.
+rewrite recursion_succ, <- Heqh; simpl; f_equiv'.
Qed.
Theorem half_aux_spec : forall n,
@@ -258,7 +240,7 @@ rewrite half_aux_0; simpl; rewrite add_0_l; auto with *.
intros.
rewrite half_aux_succ. simpl.
rewrite add_succ_l, add_comm; auto.
-apply succ_wd; auto.
+now f_equiv.
Qed.
Theorem half_aux_spec2 : forall n,
@@ -271,7 +253,7 @@ rewrite half_aux_0; simpl. auto with *.
intros.
rewrite half_aux_succ; simpl.
destruct H; auto with *.
-right; apply succ_wd; auto with *.
+right; now f_equiv.
Qed.
Theorem half_0 : half 0 == 0.
@@ -281,14 +263,14 @@ Qed.
Theorem half_1 : half 1 == 0.
Proof.
-unfold half. rewrite half_aux_succ, half_aux_0; simpl; auto with *.
+unfold half. rewrite one_succ, half_aux_succ, half_aux_0; simpl; auto with *.
Qed.
Theorem half_double : forall n,
n == 2 * half n \/ n == 1 + 2 * half n.
Proof.
intros. unfold half.
-nzsimpl.
+nzsimpl'.
destruct (half_aux_spec2 n) as [H|H]; [left|right].
rewrite <- H at 1. apply half_aux_spec.
rewrite <- add_succ_l. rewrite <- H at 1. apply half_aux_spec.
@@ -319,24 +301,23 @@ assert (LE : 0 <= half n) by apply le_0_l.
le_elim LE; auto.
destruct (half_double n) as [E|E];
rewrite <- LE, mul_0_r, ?add_0_r in E; rewrite E in LT.
-destruct (nlt_0_r _ LT).
-rewrite <- succ_lt_mono in LT.
-destruct (nlt_0_r _ LT).
+order'.
+order.
Qed.
Theorem half_decrease : forall n, 0 < n -> half n < n.
Proof.
intros n LT.
-destruct (half_double n) as [E|E]; rewrite E at 2;
- rewrite ?mul_succ_l, ?mul_0_l, ?add_0_l, ?add_assoc.
+destruct (half_double n) as [E|E]; rewrite E at 2; nzsimpl'.
rewrite <- add_0_l at 1.
rewrite <- add_lt_mono_r.
assert (LE : 0 <= half n) by apply le_0_l.
le_elim LE; auto.
rewrite <- LE, mul_0_r in E. rewrite E in LT. destruct (nlt_0_r _ LT).
+rewrite <- add_succ_l.
rewrite <- add_0_l at 1.
rewrite <- add_lt_mono_r.
-rewrite add_succ_l. apply lt_0_succ.
+apply lt_0_succ.
Qed.
@@ -347,17 +328,9 @@ Definition pow (n m : N.t) := recursion 1 (fun _ r => n*r) m.
Local Infix "^^" := pow (at level 30, right associativity).
-Instance pow_prewd :
- Proper (N.eq==>N.eq==>N.eq==>N.eq) (fun n _ r => n*r).
-Proof.
-intros n n' Hn x x' Hx y y' Hy. rewrite Hn, Hy; auto with *.
-Qed.
-
Instance pow_wd : Proper (N.eq==>N.eq==>N.eq) pow.
Proof.
-intros n n' Hn m m' Hm. unfold pow.
-apply recursion_wd; auto with *.
-now apply pow_prewd.
+unfold pow. f_equiv'.
Qed.
Lemma pow_0 : forall n, n^^0 == 1.
@@ -367,8 +340,7 @@ Qed.
Lemma pow_succ : forall n m, n^^(S m) == n*(n^^m).
Proof.
-intros. unfold pow. rewrite recursion_succ; auto with *.
-now apply pow_prewd.
+intros. unfold pow. rewrite recursion_succ; f_equiv'.
Qed.
@@ -389,15 +361,13 @@ Proof.
intros g g' Hg n n' Hn.
rewrite Hn.
destruct (n' << 2); auto with *.
-apply succ_wd.
-apply Hg. rewrite Hn; auto with *.
+f_equiv. apply Hg. now f_equiv.
Qed.
Instance log_wd : Proper (N.eq==>N.eq) log.
Proof.
intros x x' Exx'. unfold log.
-apply strong_rec_wd; auto with *.
-apply log_prewd.
+apply strong_rec_wd; f_equiv'.
Qed.
Lemma log_good_step : forall n h1 h2,
@@ -408,9 +378,9 @@ Proof.
intros n h1 h2 E.
destruct (n<<2) as [ ]_eqn:H.
auto with *.
-apply succ_wd, E, half_decrease.
-rewrite <- not_true_iff_false, ltb_lt, nlt_ge, le_succ_l in H.
-apply lt_succ_l; auto.
+f_equiv. apply E, half_decrease.
+rewrite two_succ, <- not_true_iff_false, ltb_lt, nlt_ge, le_succ_l in H.
+order'.
Qed.
Hint Resolve log_good_step.
@@ -441,14 +411,14 @@ intros n IH k Hk1 Hk2.
destruct (lt_ge_cases k 2) as [LT|LE].
(* base *)
rewrite log_init, pow_0 by auto.
-rewrite <- le_succ_l in Hk2.
+rewrite <- le_succ_l, <- one_succ in Hk2.
le_elim Hk2.
-rewrite <- nle_gt, le_succ_l in LT. destruct LT; auto.
+rewrite two_succ, <- nle_gt, le_succ_l in LT. destruct LT; auto.
rewrite <- Hk2.
rewrite half_1; auto using lt_0_1, le_refl.
(* step *)
rewrite log_step, pow_succ by auto.
-rewrite le_succ_l in LE.
+rewrite two_succ, le_succ_l in LE.
destruct (IH (half k)) as (IH1,IH2).
rewrite <- lt_succ_r. apply lt_le_trans with k; auto.
now apply half_decrease.
@@ -458,22 +428,13 @@ split.
rewrite <- le_succ_l in IH1.
apply mul_le_mono_l with (p:=2) in IH1.
eapply lt_le_trans; eauto.
-nzsimpl.
+nzsimpl'.
rewrite lt_succ_r.
eapply le_trans; [ eapply half_lower_bound | ].
-nzsimpl; apply le_refl.
+nzsimpl'; apply le_refl.
eapply le_trans; [ | eapply half_upper_bound ].
apply mul_le_mono_l; auto.
Qed.
-(** Later:
-
-Theorem log_mul : forall n m, 0 < n -> 0 < m ->
- log (n*m) == log n + log m.
-
-Theorem log_pow2 : forall n, log (2^^n) = n.
-
-*)
-
-End NdefOpsPropFunct.
+End NdefOpsProp.
diff --git a/theories/Numbers/Natural/Abstract/NDiv.v b/theories/Numbers/Natural/Abstract/NDiv.v
index 171530f0..6db8e448 100644
--- a/theories/Numbers/Natural/Abstract/NDiv.v
+++ b/theories/Numbers/Natural/Abstract/NDiv.v
@@ -1,40 +1,36 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(** Euclidean Division *)
+Require Import NAxioms NSub NZDiv.
-Require Import NAxioms NProperties NZDiv.
+(** Properties of Euclidean Division *)
-Module Type NDivSpecific (Import N : NAxiomsSig')(Import DM : DivMod' N).
- Axiom mod_upper_bound : forall a b, b ~= 0 -> a mod b < b.
-End NDivSpecific.
+Module Type NDivProp (Import N : NAxiomsSig')(Import NP : NSubProp N).
-Module Type NDivSig := NAxiomsSig <+ DivMod <+ NZDivCommon <+ NDivSpecific.
-Module Type NDivSig' := NAxiomsSig' <+ DivMod' <+ NZDivCommon <+ NDivSpecific.
+(** We benefit from what already exists for NZ *)
+Module Import Private_NZDiv := Nop <+ NZDivProp N N NP.
-Module NDivPropFunct (Import N : NDivSig')(Import NP : NPropSig N).
+Ltac auto' := try rewrite <- neq_0_lt_0; auto using le_0_l.
-(** We benefit from what already exists for NZ *)
+(** Let's now state again theorems, but without useless hypothesis. *)
- Module ND <: NZDiv N.
- Definition div := div.
- Definition modulo := modulo.
- Definition div_wd := div_wd.
- Definition mod_wd := mod_wd.
- Definition div_mod := div_mod.
- Lemma mod_bound : forall a b, 0<=a -> 0<b -> 0 <= a mod b < b.
- Proof. split. apply le_0_l. apply mod_upper_bound. order. Qed.
- End ND.
- Module Import NZDivP := NZDivPropFunct N NP ND.
+Lemma mod_upper_bound : forall a b, b ~= 0 -> a mod b < b.
+Proof. intros. apply mod_bound_pos; auto'. Qed.
- Ltac auto' := try rewrite <- neq_0_lt_0; auto using le_0_l.
+(** Another formulation of the main equation *)
-(** Let's now state again theorems, but without useless hypothesis. *)
+Lemma mod_eq :
+ forall a b, b~=0 -> a mod b == a - b*(a/b).
+Proof.
+intros.
+symmetry. apply add_sub_eq_l. symmetry.
+now apply div_mod.
+Qed.
(** Uniqueness theorems *)
@@ -51,6 +47,9 @@ Theorem mod_unique:
forall a b q r, r<b -> a == b*q + r -> r == a mod b.
Proof. intros. apply mod_unique with q; auto'. Qed.
+Theorem div_unique_exact: forall a b q, b~=0 -> a == b*q -> q == a/b.
+Proof. intros. apply div_unique_exact; auto'. Qed.
+
(** A division by itself returns 1 *)
Lemma div_same : forall a, a~=0 -> a/a == 1.
@@ -223,6 +222,10 @@ Lemma div_div : forall a b c, b~=0 -> c~=0 ->
(a/b)/c == a/(b*c).
Proof. intros. apply div_div; auto'. Qed.
+Lemma mod_mul_r : forall a b c, b~=0 -> c~=0 ->
+ a mod (b*c) == a mod b + b*((a/b) mod c).
+Proof. intros. apply mod_mul_r; auto'. Qed.
+
(** A last inequality: *)
Theorem div_mul_le:
@@ -235,5 +238,4 @@ Lemma mod_divides : forall a b, b~=0 ->
(a mod b == 0 <-> exists c, a == b*c).
Proof. intros. apply mod_divides; auto'. Qed.
-End NDivPropFunct.
-
+End NDivProp.
diff --git a/theories/Numbers/Natural/Abstract/NGcd.v b/theories/Numbers/Natural/Abstract/NGcd.v
new file mode 100644
index 00000000..ece369d8
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NGcd.v
@@ -0,0 +1,213 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Properties of the greatest common divisor *)
+
+Require Import NAxioms NSub NZGcd.
+
+Module Type NGcdProp
+ (Import A : NAxiomsSig')
+ (Import B : NSubProp A).
+
+ Include NZGcdProp A A B.
+
+(** Results concerning divisibility*)
+
+Definition divide_1_r n : (n | 1) -> n == 1
+ := divide_1_r_nonneg n (le_0_l n).
+
+Definition divide_antisym n m : (n | m) -> (m | n) -> n == m
+ := divide_antisym_nonneg n m (le_0_l n) (le_0_l m).
+
+Lemma divide_add_cancel_r : forall n m p, (n | m) -> (n | m + p) -> (n | p).
+Proof.
+ intros n m p (q,Hq) (r,Hr).
+ exists (r-q). rewrite mul_sub_distr_r, <- Hq, <- Hr.
+ now rewrite add_comm, add_sub.
+Qed.
+
+Lemma divide_sub_r : forall n m p, (n | m) -> (n | p) -> (n | m - p).
+Proof.
+ intros n m p H H'.
+ destruct (le_ge_cases m p) as [LE|LE].
+ apply sub_0_le in LE. rewrite LE. apply divide_0_r.
+ apply divide_add_cancel_r with p; trivial.
+ now rewrite add_comm, sub_add.
+Qed.
+
+(** Properties of gcd *)
+
+Definition gcd_0_l n : gcd 0 n == n := gcd_0_l_nonneg n (le_0_l n).
+Definition gcd_0_r n : gcd n 0 == n := gcd_0_r_nonneg n (le_0_l n).
+Definition gcd_diag n : gcd n n == n := gcd_diag_nonneg n (le_0_l n).
+Definition gcd_unique' n m p := gcd_unique n m p (le_0_l p).
+Definition gcd_unique_alt' n m p := gcd_unique_alt n m p (le_0_l p).
+Definition divide_gcd_iff' n m := divide_gcd_iff n m (le_0_l n).
+
+Lemma gcd_add_mult_diag_r : forall n m p, gcd n (m+p*n) == gcd n m.
+Proof.
+ intros. apply gcd_unique_alt'.
+ intros. rewrite gcd_divide_iff. split; intros (U,V); split; trivial.
+ apply divide_add_r; trivial. now apply divide_mul_r.
+ apply divide_add_cancel_r with (p*n); trivial.
+ now apply divide_mul_r. now rewrite add_comm.
+Qed.
+
+Lemma gcd_add_diag_r : forall n m, gcd n (m+n) == gcd n m.
+Proof.
+ intros n m. rewrite <- (mul_1_l n) at 2. apply gcd_add_mult_diag_r.
+Qed.
+
+Lemma gcd_sub_diag_r : forall n m, n<=m -> gcd n (m-n) == gcd n m.
+Proof.
+ intros n m H. symmetry.
+ rewrite <- (sub_add n m H) at 1. apply gcd_add_diag_r.
+Qed.
+
+(** On natural numbers, we should use a particular form
+ for the Bezout identity, since we don't have full subtraction. *)
+
+Definition Bezout n m p := exists a b, a*n == p + b*m.
+
+Instance Bezout_wd : Proper (eq==>eq==>eq==>iff) Bezout.
+Proof.
+ unfold Bezout. intros x x' Hx y y' Hy z z' Hz.
+ setoid_rewrite Hx. setoid_rewrite Hy. now setoid_rewrite Hz.
+Qed.
+
+Lemma bezout_1_gcd : forall n m, Bezout n m 1 -> gcd n m == 1.
+Proof.
+ intros n m (q & r & H).
+ apply gcd_unique; trivial using divide_1_l, le_0_1.
+ intros p Hn Hm.
+ apply divide_add_cancel_r with (r*m).
+ now apply divide_mul_r.
+ rewrite add_comm, <- H. now apply divide_mul_r.
+Qed.
+
+(** For strictly positive numbers, we have Bezout in the two directions. *)
+
+Lemma gcd_bezout_pos_pos : forall n, 0<n -> forall m, 0<m ->
+ Bezout n m (gcd n m) /\ Bezout m n (gcd n m).
+Proof.
+ intros n Hn. rewrite <- le_succ_l, <- one_succ in Hn.
+ pattern n. apply strong_right_induction with (z:=1); trivial.
+ unfold Bezout. solve_proper.
+ clear n Hn. intros n Hn IHn.
+ intros m Hm. rewrite <- le_succ_l, <- one_succ in Hm.
+ pattern m. apply strong_right_induction with (z:=1); trivial.
+ unfold Bezout. solve_proper.
+ clear m Hm. intros m Hm IHm.
+ destruct (lt_trichotomy n m) as [LT|[EQ|LT]].
+ (* n < m *)
+ destruct (IHm (m-n)) as ((a & b & EQ), (a' & b' & EQ')).
+ rewrite one_succ, le_succ_l.
+ apply lt_add_lt_sub_l; now nzsimpl.
+ apply sub_lt; order'.
+ split.
+ exists (a+b). exists b.
+ rewrite mul_add_distr_r, EQ, mul_sub_distr_l, <- add_assoc.
+ rewrite gcd_sub_diag_r by order.
+ rewrite sub_add. reflexivity. apply mul_le_mono_l; order.
+ exists a'. exists (a'+b').
+ rewrite gcd_sub_diag_r in EQ' by order.
+ rewrite (add_comm a'), mul_add_distr_r, add_assoc, <- EQ'.
+ rewrite mul_sub_distr_l, sub_add. reflexivity. apply mul_le_mono_l; order.
+ (* n = m *)
+ rewrite EQ. rewrite gcd_diag.
+ split.
+ exists 1. exists 0. now nzsimpl.
+ exists 1. exists 0. now nzsimpl.
+ (* m < n *)
+ rewrite gcd_comm, and_comm.
+ apply IHn; trivial.
+ now rewrite <- le_succ_l, <- one_succ.
+Qed.
+
+Lemma gcd_bezout_pos : forall n m, 0<n -> Bezout n m (gcd n m).
+Proof.
+ intros n m Hn.
+ destruct (eq_0_gt_0_cases m) as [EQ|LT].
+ rewrite EQ, gcd_0_r. exists 1. exists 0. now nzsimpl.
+ now apply gcd_bezout_pos_pos.
+Qed.
+
+(** For arbitrary natural numbers, we could only say that at least
+ one of the Bezout identities holds. *)
+
+Lemma gcd_bezout : forall n m,
+ Bezout n m (gcd n m) \/ Bezout m n (gcd n m).
+Proof.
+ intros n m.
+ destruct (eq_0_gt_0_cases n) as [EQ|LT].
+ right. rewrite EQ, gcd_0_l. exists 1. exists 0. now nzsimpl.
+ left. now apply gcd_bezout_pos.
+Qed.
+
+Lemma gcd_mul_mono_l :
+ forall n m p, gcd (p * n) (p * m) == p * gcd n m.
+Proof.
+ intros n m p.
+ apply gcd_unique'.
+ apply mul_divide_mono_l, gcd_divide_l.
+ apply mul_divide_mono_l, gcd_divide_r.
+ intros q H H'.
+ destruct (eq_0_gt_0_cases n) as [EQ|LT].
+ rewrite EQ in *. now rewrite gcd_0_l.
+ destruct (gcd_bezout_pos n m) as (a & b & EQ); trivial.
+ apply divide_add_cancel_r with (p*m*b).
+ now apply divide_mul_l.
+ rewrite <- mul_assoc, <- mul_add_distr_l, add_comm, (mul_comm m), <- EQ.
+ rewrite (mul_comm a), mul_assoc.
+ now apply divide_mul_l.
+Qed.
+
+Lemma gcd_mul_mono_r :
+ forall n m p, gcd (n*p) (m*p) == gcd n m * p.
+Proof.
+ intros. rewrite !(mul_comm _ p). apply gcd_mul_mono_l.
+Qed.
+
+Lemma gauss : forall n m p, (n | m * p) -> gcd n m == 1 -> (n | p).
+Proof.
+ intros n m p H G.
+ destruct (eq_0_gt_0_cases n) as [EQ|LT].
+ rewrite EQ in *. rewrite gcd_0_l in G. now rewrite <- (mul_1_l p), <- G.
+ destruct (gcd_bezout_pos n m) as (a & b & EQ); trivial.
+ rewrite G in EQ.
+ apply divide_add_cancel_r with (m*p*b).
+ now apply divide_mul_l.
+ rewrite (mul_comm _ b), mul_assoc. rewrite <- (mul_1_l p) at 2.
+ rewrite <- mul_add_distr_r, add_comm, <- EQ.
+ now apply divide_mul_l, divide_factor_r.
+Qed.
+
+Lemma divide_mul_split : forall n m p, n ~= 0 -> (n | m * p) ->
+ exists q r, n == q*r /\ (q | m) /\ (r | p).
+Proof.
+ intros n m p Hn H.
+ assert (G := gcd_nonneg n m). le_elim G.
+ destruct (gcd_divide_l n m) as (q,Hq).
+ exists (gcd n m). exists q.
+ split. now rewrite mul_comm.
+ split. apply gcd_divide_r.
+ destruct (gcd_divide_r n m) as (r,Hr).
+ rewrite Hr in H. rewrite Hq in H at 1.
+ rewrite mul_shuffle0 in H. apply mul_divide_cancel_r in H; [|order].
+ apply gauss with r; trivial.
+ apply mul_cancel_r with (gcd n m); [order|].
+ rewrite mul_1_l.
+ rewrite <- gcd_mul_mono_r, <- Hq, <- Hr; order.
+ symmetry in G. apply gcd_eq_0 in G. destruct G as (Hn',_); order.
+Qed.
+
+(** TODO : relation between gcd and division and modulo *)
+
+(** TODO : more about rel_prime (i.e. gcd == 1), about prime ... *)
+
+End NGcdProp.
diff --git a/theories/Numbers/Natural/Abstract/NIso.v b/theories/Numbers/Natural/Abstract/NIso.v
index d484a625..bcf746a7 100644
--- a/theories/Numbers/Natural/Abstract/NIso.v
+++ b/theories/Numbers/Natural/Abstract/NIso.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,11 +8,9 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NIso.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import NBase.
-Module Homomorphism (N1 N2 : NAxiomsSig).
+Module Homomorphism (N1 N2 : NAxiomsRecSig).
Local Notation "n == m" := (N2.eq n m) (at level 70, no associativity).
@@ -25,11 +23,8 @@ Definition natural_isomorphism : N1.t -> N2.t :=
Instance natural_isomorphism_wd : Proper (N1.eq ==> N2.eq) natural_isomorphism.
Proof.
unfold natural_isomorphism.
-intros n m Eqxy.
-apply N1.recursion_wd.
-reflexivity.
-intros _ _ _ y' y'' H. now apply N2.succ_wd.
-assumption.
+repeat red; intros. f_equiv; trivial.
+repeat red; intros. now f_equiv.
Qed.
Theorem natural_isomorphism_0 : natural_isomorphism N1.zero == N2.zero.
@@ -42,7 +37,7 @@ Theorem natural_isomorphism_succ :
Proof.
unfold natural_isomorphism.
intro n. rewrite N1.recursion_succ; auto with *.
-repeat red; intros. apply N2.succ_wd; auto.
+repeat red; intros. now f_equiv.
Qed.
Theorem hom_nat_iso : homomorphism natural_isomorphism.
@@ -53,9 +48,9 @@ Qed.
End Homomorphism.
-Module Inverse (N1 N2 : NAxiomsSig).
+Module Inverse (N1 N2 : NAxiomsRecSig).
-Module Import NBasePropMod1 := NBasePropFunct N1.
+Module Import NBasePropMod1 := NBaseProp N1.
(* This makes the tactic induct available. Since it is taken from
(NBasePropFunct NAxiomsMod1), it refers to induction on N1. *)
@@ -76,7 +71,7 @@ Qed.
End Inverse.
-Module Isomorphism (N1 N2 : NAxiomsSig).
+Module Isomorphism (N1 N2 : NAxiomsRecSig).
Module Hom12 := Homomorphism N1 N2.
Module Hom21 := Homomorphism N2 N1.
diff --git a/theories/Numbers/Natural/Abstract/NLcm.v b/theories/Numbers/Natural/Abstract/NLcm.v
new file mode 100644
index 00000000..1e8e678c
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NLcm.v
@@ -0,0 +1,290 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import NAxioms NSub NDiv NGcd.
+
+(** * Least Common Multiple *)
+
+(** Unlike other functions around, we will define lcm below instead of
+ axiomatizing it. Indeed, there is no "prior art" about lcm in the
+ standard library to be compliant with, and the generic definition
+ of lcm via gcd is quite reasonable.
+
+ By the way, we also state here some combined properties of div/mod
+ and gcd.
+*)
+
+Module Type NLcmProp
+ (Import A : NAxiomsSig')
+ (Import B : NSubProp A)
+ (Import C : NDivProp A B)
+ (Import D : NGcdProp A B).
+
+(** Divibility and modulo *)
+
+Lemma mod_divide : forall a b, b~=0 -> (a mod b == 0 <-> (b|a)).
+Proof.
+ intros a b Hb. split.
+ intros Hab. exists (a/b). rewrite mul_comm.
+ rewrite (div_mod a b Hb) at 1. rewrite Hab; now nzsimpl.
+ intros (c,Hc). rewrite Hc. now apply mod_mul.
+Qed.
+
+Lemma divide_div_mul_exact : forall a b c, b~=0 -> (b|a) ->
+ (c*a)/b == c*(a/b).
+Proof.
+ intros a b c Hb H.
+ apply mul_cancel_l with b; trivial.
+ rewrite mul_assoc, mul_shuffle0.
+ assert (H':=H). apply mod_divide, div_exact in H'; trivial.
+ rewrite <- H', (mul_comm a c).
+ symmetry. apply div_exact; trivial.
+ apply mod_divide; trivial.
+ now apply divide_mul_r.
+Qed.
+
+(** Gcd of divided elements, for exact divisions *)
+
+Lemma gcd_div_factor : forall a b c, c~=0 -> (c|a) -> (c|b) ->
+ gcd (a/c) (b/c) == (gcd a b)/c.
+Proof.
+ intros a b c Hc Ha Hb.
+ apply mul_cancel_l with c; try order.
+ assert (H:=gcd_greatest _ _ _ Ha Hb).
+ apply mod_divide, div_exact in H; try order.
+ rewrite <- H.
+ rewrite <- gcd_mul_mono_l; try order.
+ f_equiv; symmetry; apply div_exact; try order;
+ apply mod_divide; trivial; try order.
+Qed.
+
+Lemma gcd_div_gcd : forall a b g, g~=0 -> g == gcd a b ->
+ gcd (a/g) (b/g) == 1.
+Proof.
+ intros a b g NZ EQ. rewrite gcd_div_factor.
+ now rewrite <- EQ, div_same.
+ generalize (gcd_nonneg a b); order.
+ rewrite EQ; apply gcd_divide_l.
+ rewrite EQ; apply gcd_divide_r.
+Qed.
+
+(** The following equality is crucial for Euclid algorithm *)
+
+Lemma gcd_mod : forall a b, b~=0 -> gcd (a mod b) b == gcd b a.
+Proof.
+ intros a b Hb. rewrite (gcd_comm _ b).
+ rewrite <- (gcd_add_mult_diag_r b (a mod b) (a/b)).
+ now rewrite add_comm, mul_comm, <- div_mod.
+Qed.
+
+(** We now define lcm thanks to gcd:
+
+ lcm a b = a * (b / gcd a b)
+ = (a / gcd a b) * b
+ = (a*b) / gcd a b
+
+ Nota: [lcm 0 0] should be 0, which isn't garantee with the third
+ equation above.
+*)
+
+Definition lcm a b := a*(b/gcd a b).
+
+Instance lcm_wd : Proper (eq==>eq==>eq) lcm.
+Proof. unfold lcm. solve_proper. Qed.
+
+Lemma lcm_equiv1 : forall a b, gcd a b ~= 0 ->
+ a * (b / gcd a b) == (a*b)/gcd a b.
+Proof.
+ intros a b H. rewrite divide_div_mul_exact; try easy. apply gcd_divide_r.
+Qed.
+
+Lemma lcm_equiv2 : forall a b, gcd a b ~= 0 ->
+ (a / gcd a b) * b == (a*b)/gcd a b.
+Proof.
+ intros a b H. rewrite 2 (mul_comm _ b).
+ rewrite divide_div_mul_exact; try easy. apply gcd_divide_l.
+Qed.
+
+Lemma gcd_div_swap : forall a b,
+ (a / gcd a b) * b == a * (b / gcd a b).
+Proof.
+ intros a b. destruct (eq_decidable (gcd a b) 0) as [EQ|NEQ].
+ apply gcd_eq_0 in EQ. destruct EQ as (EQ,EQ'). rewrite EQ, EQ'. now nzsimpl.
+ now rewrite lcm_equiv1, <-lcm_equiv2.
+Qed.
+
+Lemma divide_lcm_l : forall a b, (a | lcm a b).
+Proof.
+ unfold lcm. intros a b. apply divide_factor_l.
+Qed.
+
+Lemma divide_lcm_r : forall a b, (b | lcm a b).
+Proof.
+ unfold lcm. intros a b. rewrite <- gcd_div_swap.
+ apply divide_factor_r.
+Qed.
+
+Lemma divide_div : forall a b c, a~=0 -> (a|b) -> (b|c) -> (b/a|c/a).
+Proof.
+ intros a b c Ha Hb (c',Hc). exists c'.
+ now rewrite <- divide_div_mul_exact, Hc.
+Qed.
+
+Lemma lcm_least : forall a b c,
+ (a | c) -> (b | c) -> (lcm a b | c).
+Proof.
+ intros a b c Ha Hb. unfold lcm.
+ destruct (eq_decidable (gcd a b) 0) as [EQ|NEQ].
+ apply gcd_eq_0 in EQ. destruct EQ as (EQ,EQ'). rewrite EQ in *. now nzsimpl.
+ assert (Ga := gcd_divide_l a b).
+ assert (Gb := gcd_divide_r a b).
+ set (g:=gcd a b) in *.
+ assert (Ha' := divide_div g a c NEQ Ga Ha).
+ assert (Hb' := divide_div g b c NEQ Gb Hb).
+ destruct Ha' as (a',Ha'). rewrite Ha', mul_comm in Hb'.
+ apply gauss in Hb'; [|apply gcd_div_gcd; unfold g; trivial using gcd_comm].
+ destruct Hb' as (b',Hb').
+ exists b'.
+ rewrite mul_shuffle3, <- Hb'.
+ rewrite (proj2 (div_exact c g NEQ)).
+ rewrite Ha', mul_shuffle3, (mul_comm a a'). f_equiv.
+ symmetry. apply div_exact; trivial.
+ apply mod_divide; trivial.
+ apply mod_divide; trivial. transitivity a; trivial.
+Qed.
+
+Lemma lcm_comm : forall a b, lcm a b == lcm b a.
+Proof.
+ intros a b. unfold lcm. rewrite (gcd_comm b), (mul_comm b).
+ now rewrite <- gcd_div_swap.
+Qed.
+
+Lemma lcm_divide_iff : forall n m p,
+ (lcm n m | p) <-> (n | p) /\ (m | p).
+Proof.
+ intros. split. split.
+ transitivity (lcm n m); trivial using divide_lcm_l.
+ transitivity (lcm n m); trivial using divide_lcm_r.
+ intros (H,H'). now apply lcm_least.
+Qed.
+
+Lemma lcm_unique : forall n m p,
+ 0<=p -> (n|p) -> (m|p) ->
+ (forall q, (n|q) -> (m|q) -> (p|q)) ->
+ lcm n m == p.
+Proof.
+ intros n m p Hp Hn Hm H.
+ apply divide_antisym; trivial.
+ now apply lcm_least.
+ apply H. apply divide_lcm_l. apply divide_lcm_r.
+Qed.
+
+Lemma lcm_unique_alt : forall n m p, 0<=p ->
+ (forall q, (p|q) <-> (n|q) /\ (m|q)) ->
+ lcm n m == p.
+Proof.
+ intros n m p Hp H.
+ apply lcm_unique; trivial.
+ apply H, divide_refl.
+ apply H, divide_refl.
+ intros. apply H. now split.
+Qed.
+
+Lemma lcm_assoc : forall n m p, lcm n (lcm m p) == lcm (lcm n m) p.
+Proof.
+ intros. apply lcm_unique_alt. apply le_0_l.
+ intros. now rewrite !lcm_divide_iff, and_assoc.
+Qed.
+
+Lemma lcm_0_l : forall n, lcm 0 n == 0.
+Proof.
+ intros. apply lcm_unique; trivial. order.
+ apply divide_refl.
+ apply divide_0_r.
+Qed.
+
+Lemma lcm_0_r : forall n, lcm n 0 == 0.
+Proof.
+ intros. now rewrite lcm_comm, lcm_0_l.
+Qed.
+
+Lemma lcm_1_l : forall n, lcm 1 n == n.
+Proof.
+ intros. apply lcm_unique; trivial using divide_1_l, le_0_l, divide_refl.
+Qed.
+
+Lemma lcm_1_r : forall n, lcm n 1 == n.
+Proof.
+ intros. now rewrite lcm_comm, lcm_1_l.
+Qed.
+
+Lemma lcm_diag : forall n, lcm n n == n.
+Proof.
+ intros. apply lcm_unique; trivial using divide_refl, le_0_l.
+Qed.
+
+Lemma lcm_eq_0 : forall n m, lcm n m == 0 <-> n == 0 \/ m == 0.
+Proof.
+ intros. split.
+ intros EQ.
+ apply eq_mul_0.
+ apply divide_0_l. rewrite <- EQ. apply lcm_least.
+ apply divide_factor_l. apply divide_factor_r.
+ destruct 1 as [EQ|EQ]; rewrite EQ. apply lcm_0_l. apply lcm_0_r.
+Qed.
+
+Lemma divide_lcm_eq_r : forall n m, (n|m) -> lcm n m == m.
+Proof.
+ intros n m H. apply lcm_unique_alt; trivial using le_0_l.
+ intros q. split. split; trivial. now transitivity m.
+ now destruct 1.
+Qed.
+
+Lemma divide_lcm_iff : forall n m, (n|m) <-> lcm n m == m.
+Proof.
+ intros n m. split. now apply divide_lcm_eq_r.
+ intros EQ. rewrite <- EQ. apply divide_lcm_l.
+Qed.
+
+Lemma lcm_mul_mono_l :
+ forall n m p, lcm (p * n) (p * m) == p * lcm n m.
+Proof.
+ intros n m p.
+ destruct (eq_decidable p 0) as [Hp|Hp].
+ rewrite Hp. nzsimpl. rewrite lcm_0_l. now nzsimpl.
+ destruct (eq_decidable (gcd n m) 0) as [Hg|Hg].
+ apply gcd_eq_0 in Hg. destruct Hg as (Hn,Hm); rewrite Hn, Hm.
+ nzsimpl. rewrite lcm_0_l. now nzsimpl.
+ unfold lcm.
+ rewrite gcd_mul_mono_l.
+ rewrite mul_assoc. f_equiv.
+ now rewrite div_mul_cancel_l.
+Qed.
+
+Lemma lcm_mul_mono_r :
+ forall n m p, lcm (n * p) (m * p) == lcm n m * p.
+Proof.
+ intros n m p. now rewrite !(mul_comm _ p), lcm_mul_mono_l, mul_comm.
+Qed.
+
+Lemma gcd_1_lcm_mul : forall n m, n~=0 -> m~=0 ->
+ (gcd n m == 1 <-> lcm n m == n*m).
+Proof.
+ intros n m Hn Hm. split; intros H.
+ unfold lcm. rewrite H. now rewrite div_1_r.
+ unfold lcm in *.
+ apply mul_cancel_l in H; trivial.
+ assert (Hg : gcd n m ~= 0) by (red; rewrite gcd_eq_0; destruct 1; order).
+ assert (H' := gcd_divide_r n m).
+ apply mod_divide in H'; trivial. apply div_exact in H'; trivial.
+ rewrite H in H'.
+ rewrite <- (mul_1_l m) in H' at 1.
+ now apply mul_cancel_r in H'.
+Qed.
+
+End NLcmProp.
diff --git a/theories/Numbers/Natural/Abstract/NLog.v b/theories/Numbers/Natural/Abstract/NLog.v
new file mode 100644
index 00000000..74827c6e
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NLog.v
@@ -0,0 +1,23 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Base-2 Logarithm Properties *)
+
+Require Import NAxioms NSub NPow NParity NZLog.
+
+Module Type NLog2Prop
+ (A : NAxiomsSig)
+ (B : NSubProp A)
+ (C : NParityProp A B)
+ (D : NPowProp A B C).
+
+ (** For the moment we simply reuse NZ properties *)
+
+ Include NZLog2Prop A A A B D.Private_NZPow.
+ Include NZLog2UpProp A A A B D.Private_NZPow.
+End NLog2Prop.
diff --git a/theories/Numbers/Natural/Abstract/NMaxMin.v b/theories/Numbers/Natural/Abstract/NMaxMin.v
new file mode 100644
index 00000000..cdff6dbc
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NMaxMin.v
@@ -0,0 +1,135 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import NAxioms NSub GenericMinMax.
+
+(** * Properties of minimum and maximum specific to natural numbers *)
+
+Module Type NMaxMinProp (Import N : NAxiomsMiniSig').
+Include NSubProp N.
+
+(** Zero *)
+
+Lemma max_0_l : forall n, max 0 n == n.
+Proof.
+ intros. apply max_r. apply le_0_l.
+Qed.
+
+Lemma max_0_r : forall n, max n 0 == n.
+Proof.
+ intros. apply max_l. apply le_0_l.
+Qed.
+
+Lemma min_0_l : forall n, min 0 n == 0.
+Proof.
+ intros. apply min_l. apply le_0_l.
+Qed.
+
+Lemma min_0_r : forall n, min n 0 == 0.
+Proof.
+ intros. apply min_r. apply le_0_l.
+Qed.
+
+(** The following results are concrete instances of [max_monotone]
+ and similar lemmas. *)
+
+(** Succ *)
+
+Lemma succ_max_distr : forall n m, S (max n m) == max (S n) (S m).
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 max_r | rewrite 2 max_l]; now rewrite <- ?succ_le_mono.
+Qed.
+
+Lemma succ_min_distr : forall n m, S (min n m) == min (S n) (S m).
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 min_l | rewrite 2 min_r]; now rewrite <- ?succ_le_mono.
+Qed.
+
+(** Add *)
+
+Lemma add_max_distr_l : forall n m p, max (p + n) (p + m) == p + max n m.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 max_r | rewrite 2 max_l]; now rewrite <- ?add_le_mono_l.
+Qed.
+
+Lemma add_max_distr_r : forall n m p, max (n + p) (m + p) == max n m + p.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 max_r | rewrite 2 max_l]; now rewrite <- ?add_le_mono_r.
+Qed.
+
+Lemma add_min_distr_l : forall n m p, min (p + n) (p + m) == p + min n m.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 min_l | rewrite 2 min_r]; now rewrite <- ?add_le_mono_l.
+Qed.
+
+Lemma add_min_distr_r : forall n m p, min (n + p) (m + p) == min n m + p.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 min_l | rewrite 2 min_r]; now rewrite <- ?add_le_mono_r.
+Qed.
+
+(** Mul *)
+
+Lemma mul_max_distr_l : forall n m p, max (p * n) (p * m) == p * max n m.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 max_r | rewrite 2 max_l]; try order; now apply mul_le_mono_l.
+Qed.
+
+Lemma mul_max_distr_r : forall n m p, max (n * p) (m * p) == max n m * p.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 max_r | rewrite 2 max_l]; try order; now apply mul_le_mono_r.
+Qed.
+
+Lemma mul_min_distr_l : forall n m p, min (p * n) (p * m) == p * min n m.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 min_l | rewrite 2 min_r]; try order; now apply mul_le_mono_l.
+Qed.
+
+Lemma mul_min_distr_r : forall n m p, min (n * p) (m * p) == min n m * p.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 min_l | rewrite 2 min_r]; try order; now apply mul_le_mono_r.
+Qed.
+
+(** Sub *)
+
+Lemma sub_max_distr_l : forall n m p, max (p - n) (p - m) == p - min n m.
+Proof.
+ intros. destruct (le_ge_cases n m).
+ rewrite min_l by trivial. apply max_l. now apply sub_le_mono_l.
+ rewrite min_r by trivial. apply max_r. now apply sub_le_mono_l.
+Qed.
+
+Lemma sub_max_distr_r : forall n m p, max (n - p) (m - p) == max n m - p.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 max_r | rewrite 2 max_l]; try order; now apply sub_le_mono_r.
+Qed.
+
+Lemma sub_min_distr_l : forall n m p, min (p - n) (p - m) == p - max n m.
+Proof.
+ intros. destruct (le_ge_cases n m).
+ rewrite max_r by trivial. apply min_r. now apply sub_le_mono_l.
+ rewrite max_l by trivial. apply min_l. now apply sub_le_mono_l.
+Qed.
+
+Lemma sub_min_distr_r : forall n m p, min (n - p) (m - p) == min n m - p.
+Proof.
+ intros. destruct (le_ge_cases n m);
+ [rewrite 2 min_l | rewrite 2 min_r]; try order; now apply sub_le_mono_r.
+Qed.
+
+End NMaxMinProp.
diff --git a/theories/Numbers/Natural/Abstract/NMulOrder.v b/theories/Numbers/Natural/Abstract/NMulOrder.v
index bdd4b674..1d6e8ba0 100644
--- a/theories/Numbers/Natural/Abstract/NMulOrder.v
+++ b/theories/Numbers/Natural/Abstract/NMulOrder.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,12 +8,10 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NMulOrder.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Export NAddOrder.
-Module NMulOrderPropFunct (Import N : NAxiomsSig').
-Include NAddOrderPropFunct N.
+Module NMulOrderProp (Import N : NAxiomsMiniSig').
+Include NAddOrderProp N.
(** Theorems that are either not valid on Z or have different proofs
on N and Z *)
@@ -55,7 +53,7 @@ Qed.
Theorem lt_0_mul' : forall n m, n * m > 0 <-> n > 0 /\ m > 0.
Proof.
intros n m; split; [intro H | intros [H1 H2]].
-apply -> lt_0_mul in H. destruct H as [[H1 H2] | [H1 H2]]. now split.
+apply lt_0_mul in H. destruct H as [[H1 H2] | [H1 H2]]. now split.
false_hyp H1 nlt_0_r.
now apply mul_pos_pos.
Qed.
@@ -67,14 +65,18 @@ Proof.
intros n m.
split; [| intros [H1 H2]; now rewrite H1, H2, mul_1_l].
intro H; destruct (lt_trichotomy n 1) as [H1 | [H1 | H1]].
-apply -> lt_1_r in H1. rewrite H1, mul_0_l in H. false_hyp H neq_0_succ.
+apply lt_1_r in H1. rewrite H1, mul_0_l in H. order'.
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.
+rewrite H2, mul_0_r in H. order'.
+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 m).
rewrite H in H3; false_hyp H3 lt_irrefl.
Qed.
-End NMulOrderPropFunct.
+(** Alternative name : *)
+
+Definition mul_eq_1 := eq_mul_1.
+
+End NMulOrderProp.
diff --git a/theories/Numbers/Natural/Abstract/NOrder.v b/theories/Numbers/Natural/Abstract/NOrder.v
index 17dd3466..8bba7d72 100644
--- a/theories/Numbers/Natural/Abstract/NOrder.v
+++ b/theories/Numbers/Natural/Abstract/NOrder.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,18 +8,16 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NOrder.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Export NAdd.
-Module NOrderPropFunct (Import N : NAxiomsSig').
-Include NAddPropFunct N.
+Module NOrderProp (Import N : NAxiomsMiniSig').
+Include NAddProp N.
(* Theorems that are true for natural numbers but not for integers *)
Theorem lt_wf_0 : well_founded lt.
Proof.
-setoid_replace lt with (fun n m => 0 <= n /\ n < m).
+setoid_replace lt with (fun n m => 0 <= n < m).
apply lt_wf.
intros x y; split.
intro H; split; [apply le_0_l | assumption]. now intros [_ H].
@@ -29,12 +27,12 @@ Defined.
Theorem nlt_0_r : forall n, ~ n < 0.
Proof.
-intro n; apply -> le_ngt. apply le_0_l.
+intro n; apply le_ngt. apply le_0_l.
Qed.
Theorem nle_succ_0 : forall n, ~ (S n <= 0).
Proof.
-intros n H; apply -> le_succ_l in H; false_hyp H nlt_0_r.
+intros n H; apply le_succ_l in H; false_hyp H nlt_0_r.
Qed.
Theorem le_0_r : forall n, n <= 0 <-> n == 0.
@@ -65,6 +63,7 @@ Qed.
Theorem zero_one : forall n, n == 0 \/ n == 1 \/ 1 < n.
Proof.
+setoid_rewrite one_succ.
induct n. now left.
cases n. intros; right; now left.
intros n IH. destruct IH as [H | [H | H]].
@@ -75,6 +74,7 @@ Qed.
Theorem lt_1_r : forall n, n < 1 <-> n == 0.
Proof.
+setoid_rewrite one_succ.
cases n.
split; intro; [reflexivity | apply lt_succ_diag_r].
intros n. rewrite <- succ_lt_mono.
@@ -83,6 +83,7 @@ Qed.
Theorem le_1_r : forall n, n <= 1 <-> n == 0 \/ n == 1.
Proof.
+setoid_rewrite one_succ.
cases n.
split; intro; [now left | apply le_succ_diag_r].
intro n. rewrite <- succ_le_mono, le_0_r, succ_inj_wd.
@@ -117,9 +118,9 @@ Proof.
intros Base Step; induct n.
intros; apply Base.
intros n IH m H. elim H using le_ind.
-solve_predicate_wd.
+solve_proper.
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.
+intros k H1 H2. apply le_succ_l in H1. apply lt_le_incl in H1. auto.
Qed.
Theorem lt_ind_rel :
@@ -131,7 +132,7 @@ 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.
+solve_proper.
apply Step; [| apply IH]; now apply lt_succ_diag_r.
intros k H1 H2. apply lt_succ_l in H1. auto.
Qed.
@@ -175,7 +176,7 @@ Theorem lt_le_pred : forall n m, n < m -> n <= P m.
Proof.
intro n; cases m.
intro H; false_hyp H nlt_0_r.
-intros m IH. rewrite pred_succ; now apply -> lt_succ_r.
+intros m IH. rewrite pred_succ; now apply lt_succ_r.
Qed.
Theorem lt_pred_le : forall n m, P n < m -> n <= m.
@@ -183,7 +184,7 @@ Theorem lt_pred_le : forall n m, P n < m -> n <= m.
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.
+intros n IH. rewrite pred_succ in IH. now apply le_succ_l.
Qed.
Theorem lt_pred_lt : forall n m, n < P m -> n < m.
@@ -200,7 +201,7 @@ Theorem pred_le_mono : forall n m, n <= m -> P n <= P m.
(* Converse is false for n == 1, m == 0 *)
Proof.
intros n m H; elim H using le_ind_rel.
-solve_relation_wd.
+solve_proper.
intro; rewrite pred_0; apply le_0_l.
intros p q H1 _; now do 2 rewrite pred_succ.
Qed.
@@ -208,12 +209,12 @@ Qed.
Theorem pred_lt_mono : forall n m, n ~= 0 -> (n < m <-> P n < P m).
Proof.
intros n m H1; split; intro H2.
-assert (m ~= 0). apply <- neq_0_lt_0. now apply lt_lt_0 with n.
+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 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.
+apply succ_lt_mono in H2. now do 2 rewrite succ_pred in H2.
Qed.
Theorem lt_succ_lt_pred : forall n m, S n < m <-> n < P m.
@@ -224,13 +225,13 @@ Qed.
Theorem le_succ_le_pred : forall n m, S n <= m -> n <= P m.
(* Converse is false for n == m == 0 *)
Proof.
-intros n m H. apply lt_le_pred. now apply -> le_succ_l.
+intros n m H. apply lt_le_pred. now apply le_succ_l.
Qed.
Theorem lt_pred_lt_succ : forall n m, P n < m -> n < S m.
(* Converse is false for n == m == 0 *)
Proof.
-intros n m H. apply <- lt_succ_r. now apply lt_pred_le.
+intros n m H. apply lt_succ_r. now apply lt_pred_le.
Qed.
Theorem le_pred_le_succ : forall n m, P n <= m <-> n <= S m.
@@ -240,5 +241,5 @@ rewrite pred_0. split; intro H; apply le_0_l.
intro n. rewrite pred_succ. apply succ_le_mono.
Qed.
-End NOrderPropFunct.
+End NOrderProp.
diff --git a/theories/Numbers/Natural/Abstract/NParity.v b/theories/Numbers/Natural/Abstract/NParity.v
new file mode 100644
index 00000000..6a1e20ce
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NParity.v
@@ -0,0 +1,63 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Bool NSub NZParity.
+
+(** Some additionnal properties of [even], [odd]. *)
+
+Module Type NParityProp (Import N : NAxiomsSig')(Import NP : NSubProp N).
+
+Include NZParityProp N N NP.
+
+Lemma odd_pred : forall n, n~=0 -> odd (P n) = even n.
+Proof.
+ intros. rewrite <- (succ_pred n) at 2 by trivial.
+ symmetry. apply even_succ.
+Qed.
+
+Lemma even_pred : forall n, n~=0 -> even (P n) = odd n.
+Proof.
+ intros. rewrite <- (succ_pred n) at 2 by trivial.
+ symmetry. apply odd_succ.
+Qed.
+
+Lemma even_sub : forall n m, m<=n -> even (n-m) = Bool.eqb (even n) (even m).
+Proof.
+ intros.
+ case_eq (even n); case_eq (even m);
+ rewrite <- ?negb_true_iff, ?negb_even, ?odd_spec, ?even_spec;
+ intros (m',Hm) (n',Hn).
+ exists (n'-m'). now rewrite mul_sub_distr_l, Hn, Hm.
+ exists (n'-m'-1).
+ rewrite !mul_sub_distr_l, Hn, Hm, sub_add_distr, mul_1_r.
+ rewrite two_succ at 5. rewrite <- (add_1_l 1). rewrite sub_add_distr.
+ symmetry. apply sub_add.
+ apply le_add_le_sub_l.
+ rewrite add_1_l, <- two_succ, <- (mul_1_r 2) at 1.
+ rewrite <- mul_sub_distr_l. rewrite <- mul_le_mono_pos_l by order'.
+ rewrite one_succ, le_succ_l. rewrite <- lt_add_lt_sub_l, add_0_r.
+ destruct (le_gt_cases n' m') as [LE|GT]; trivial.
+ generalize (double_below _ _ LE). order.
+ exists (n'-m'). rewrite mul_sub_distr_l, Hn, Hm.
+ apply add_sub_swap.
+ apply mul_le_mono_pos_l; try order'.
+ destruct (le_gt_cases m' n') as [LE|GT]; trivial.
+ generalize (double_above _ _ GT). order.
+ exists (n'-m'). rewrite Hm,Hn, mul_sub_distr_l.
+ rewrite sub_add_distr. rewrite add_sub_swap. apply add_sub.
+ apply succ_le_mono.
+ rewrite add_1_r in Hm,Hn. order.
+Qed.
+
+Lemma odd_sub : forall n m, m<=n -> odd (n-m) = xorb (odd n) (odd m).
+Proof.
+ intros. rewrite <- !negb_even. rewrite even_sub by trivial.
+ now destruct (even n), (even m).
+Qed.
+
+End NParityProp.
diff --git a/theories/Numbers/Natural/Abstract/NPow.v b/theories/Numbers/Natural/Abstract/NPow.v
new file mode 100644
index 00000000..07aee9c6
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NPow.v
@@ -0,0 +1,160 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Properties of the power function *)
+
+Require Import Bool NAxioms NSub NParity NZPow.
+
+(** Derived properties of power, specialized on natural numbers *)
+
+Module Type NPowProp
+ (Import A : NAxiomsSig')
+ (Import B : NSubProp A)
+ (Import C : NParityProp A B).
+
+ Module Import Private_NZPow := Nop <+ NZPowProp A A B.
+
+Ltac auto' := trivial; try rewrite <- neq_0_lt_0; auto using le_0_l.
+Ltac wrap l := intros; apply l; auto'.
+
+Lemma pow_succ_r' : forall a b, a^(S b) == a * a^b.
+Proof. wrap pow_succ_r. Qed.
+
+(** Power and basic constants *)
+
+Lemma pow_0_l : forall a, a~=0 -> 0^a == 0.
+Proof. wrap pow_0_l. Qed.
+
+Definition pow_1_r : forall a, a^1 == a
+ := pow_1_r.
+
+Lemma pow_1_l : forall a, 1^a == 1.
+Proof. wrap pow_1_l. Qed.
+
+Definition pow_2_r : forall a, a^2 == a*a
+ := pow_2_r.
+
+(** Power and addition, multiplication *)
+
+Lemma pow_add_r : forall a b c, a^(b+c) == a^b * a^c.
+Proof. wrap pow_add_r. Qed.
+
+Lemma pow_mul_l : forall a b c, (a*b)^c == a^c * b^c.
+Proof. wrap pow_mul_l. Qed.
+
+Lemma pow_mul_r : forall a b c, a^(b*c) == (a^b)^c.
+Proof. wrap pow_mul_r. Qed.
+
+(** Power and nullity *)
+
+Lemma pow_eq_0 : forall a b, b~=0 -> a^b == 0 -> a == 0.
+Proof. intros. apply (pow_eq_0 a b); trivial. auto'. Qed.
+
+Lemma pow_nonzero : forall a b, a~=0 -> a^b ~= 0.
+Proof. wrap pow_nonzero. Qed.
+
+Lemma pow_eq_0_iff : forall a b, a^b == 0 <-> b~=0 /\ a==0.
+Proof.
+ intros a b. split.
+ rewrite pow_eq_0_iff. intros [H |[H H']].
+ generalize (le_0_l b); order. split; order.
+ intros (Hb,Ha). rewrite Ha. now apply pow_0_l'.
+Qed.
+
+(** Monotonicity *)
+
+Lemma pow_lt_mono_l : forall a b c, c~=0 -> a<b -> a^c < b^c.
+Proof. wrap pow_lt_mono_l. Qed.
+
+Lemma pow_le_mono_l : forall a b c, a<=b -> a^c <= b^c.
+Proof. wrap pow_le_mono_l. Qed.
+
+Lemma pow_gt_1 : forall a b, 1<a -> b~=0 -> 1<a^b.
+Proof. wrap pow_gt_1. Qed.
+
+Lemma pow_lt_mono_r : forall a b c, 1<a -> b<c -> a^b < a^c.
+Proof. wrap pow_lt_mono_r. Qed.
+
+(** NB: since 0^0 > 0^1, the following result isn't valid with a=0 *)
+
+Lemma pow_le_mono_r : forall a b c, a~=0 -> b<=c -> a^b <= a^c.
+Proof. wrap pow_le_mono_r. Qed.
+
+Lemma pow_le_mono : forall a b c d, a~=0 -> a<=c -> b<=d ->
+ a^b <= c^d.
+Proof. wrap pow_le_mono. Qed.
+
+Definition pow_lt_mono : forall a b c d, 0<a<c -> 0<b<d ->
+ a^b < c^d
+ := pow_lt_mono.
+
+(** Injectivity *)
+
+Lemma pow_inj_l : forall a b c, c~=0 -> a^c == b^c -> a == b.
+Proof. intros; eapply pow_inj_l; eauto; auto'. Qed.
+
+Lemma pow_inj_r : forall a b c, 1<a -> a^b == a^c -> b == c.
+Proof. intros; eapply pow_inj_r; eauto; auto'. Qed.
+
+(** Monotonicity results, both ways *)
+
+Lemma pow_lt_mono_l_iff : forall a b c, c~=0 ->
+ (a<b <-> a^c < b^c).
+Proof. wrap pow_lt_mono_l_iff. Qed.
+
+Lemma pow_le_mono_l_iff : forall a b c, c~=0 ->
+ (a<=b <-> a^c <= b^c).
+Proof. wrap pow_le_mono_l_iff. Qed.
+
+Lemma pow_lt_mono_r_iff : forall a b c, 1<a ->
+ (b<c <-> a^b < a^c).
+Proof. wrap pow_lt_mono_r_iff. Qed.
+
+Lemma pow_le_mono_r_iff : forall a b c, 1<a ->
+ (b<=c <-> a^b <= a^c).
+Proof. wrap pow_le_mono_r_iff. Qed.
+
+(** For any a>1, the a^x function is above the identity function *)
+
+Lemma pow_gt_lin_r : forall a b, 1<a -> b < a^b.
+Proof. wrap pow_gt_lin_r. Qed.
+
+(** Someday, we should say something about the full Newton formula.
+ In the meantime, we can at least provide some inequalities about
+ (a+b)^c.
+*)
+
+Lemma pow_add_lower : forall a b c, c~=0 ->
+ a^c + b^c <= (a+b)^c.
+Proof. wrap pow_add_lower. Qed.
+
+(** This upper bound can also be seen as a convexity proof for x^c :
+ image of (a+b)/2 is below the middle of the images of a and b
+*)
+
+Lemma pow_add_upper : forall a b c, c~=0 ->
+ (a+b)^c <= 2^(pred c) * (a^c + b^c).
+Proof. wrap pow_add_upper. Qed.
+
+(** Power and parity *)
+
+Lemma even_pow : forall a b, b~=0 -> even (a^b) = even a.
+Proof.
+ intros a b Hb. rewrite neq_0_lt_0 in Hb.
+ apply lt_ind with (4:=Hb). solve_proper.
+ now nzsimpl.
+ clear b Hb. intros b Hb IH.
+ rewrite pow_succ_r', even_mul, IH. now destruct (even a).
+Qed.
+
+Lemma odd_pow : forall a b, b~=0 -> odd (a^b) = odd a.
+Proof.
+ intros. now rewrite <- !negb_even, even_pow.
+Qed.
+
+End NPowProp.
diff --git a/theories/Numbers/Natural/Abstract/NProperties.v b/theories/Numbers/Natural/Abstract/NProperties.v
index c9e05113..1edb6b51 100644
--- a/theories/Numbers/Natural/Abstract/NProperties.v
+++ b/theories/Numbers/Natural/Abstract/NProperties.v
@@ -1,22 +1,17 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: NProperties.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
+Require Export NAxioms.
+Require Import NMaxMin NParity NPow NSqrt NLog NDiv NGcd NLcm NBits.
-Require Export NAxioms NSub.
+(** This functor summarizes all known facts about N. *)
-(** This functor summarizes all known facts about N.
- For the moment it is only an alias to [NSubPropFunct], which
- subsumes all others.
-*)
-
-Module Type NPropSig := NSubPropFunct.
-
-Module NPropFunct (N:NAxiomsSig) <: NPropSig N.
- Include NPropSig N.
-End NPropFunct.
+Module Type NProp (N:NAxiomsSig) :=
+ NMaxMinProp N <+ NParityProp N <+ NPowProp N <+ NSqrtProp N
+ <+ NLog2Prop N <+ NDivProp N <+ NGcdProp N <+ NLcmProp N
+ <+ NBitsProp N.
diff --git a/theories/Numbers/Natural/Abstract/NSqrt.v b/theories/Numbers/Natural/Abstract/NSqrt.v
new file mode 100644
index 00000000..34b7d011
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NSqrt.v
@@ -0,0 +1,75 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** Properties of Square Root Function *)
+
+Require Import NAxioms NSub NZSqrt.
+
+Module NSqrtProp (Import A : NAxiomsSig')(Import B : NSubProp A).
+
+ Module Import Private_NZSqrt := Nop <+ NZSqrtProp A A B.
+
+ Ltac auto' := trivial; try rewrite <- neq_0_lt_0; auto using le_0_l.
+ Ltac wrap l := intros; apply l; auto'.
+
+ (** We redefine NZSqrt's results, without the non-negative hyps *)
+
+Lemma sqrt_spec' : forall a, √a*√a <= a < S (√a) * S (√a).
+Proof. wrap sqrt_spec. Qed.
+
+Definition sqrt_unique : forall a b, b*b<=a<(S b)*(S b) -> √a == b
+ := sqrt_unique.
+
+Lemma sqrt_square : forall a, √(a*a) == a.
+Proof. wrap sqrt_square. Qed.
+
+Definition sqrt_le_mono : forall a b, a<=b -> √a <= √b
+ := sqrt_le_mono.
+
+Definition sqrt_lt_cancel : forall a b, √a < √b -> a < b
+ := sqrt_lt_cancel.
+
+Lemma sqrt_le_square : forall a b, b*b<=a <-> b <= √a.
+Proof. wrap sqrt_le_square. Qed.
+
+Lemma sqrt_lt_square : forall a b, a<b*b <-> √a < b.
+Proof. wrap sqrt_lt_square. Qed.
+
+Definition sqrt_0 := sqrt_0.
+Definition sqrt_1 := sqrt_1.
+Definition sqrt_2 := sqrt_2.
+
+Definition sqrt_lt_lin : forall a, 1<a -> √a<a
+ := sqrt_lt_lin.
+
+Lemma sqrt_le_lin : forall a, √a<=a.
+Proof. wrap sqrt_le_lin. Qed.
+
+Definition sqrt_mul_below : forall a b, √a * √b <= √(a*b)
+ := sqrt_mul_below.
+
+Lemma sqrt_mul_above : forall a b, √(a*b) < S (√a) * S (√b).
+Proof. wrap sqrt_mul_above. Qed.
+
+Lemma sqrt_succ_le : forall a, √(S a) <= S (√a).
+Proof. wrap sqrt_succ_le. Qed.
+
+Lemma sqrt_succ_or : forall a, √(S a) == S (√a) \/ √(S a) == √a.
+Proof. wrap sqrt_succ_or. Qed.
+
+Definition sqrt_add_le : forall a b, √(a+b) <= √a + √b
+ := sqrt_add_le.
+
+Lemma add_sqrt_le : forall a b, √a + √b <= √(2*(a+b)).
+Proof. wrap add_sqrt_le. Qed.
+
+(** For the moment, we include stuff about [sqrt_up] with patching them. *)
+
+Include NZSqrtUpProp A A B Private_NZSqrt.
+
+End NSqrtProp.
diff --git a/theories/Numbers/Natural/Abstract/NStrongRec.v b/theories/Numbers/Natural/Abstract/NStrongRec.v
index d9a2427d..607746d5 100644
--- a/theories/Numbers/Natural/Abstract/NStrongRec.v
+++ b/theories/Numbers/Natural/Abstract/NStrongRec.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,15 +8,15 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NStrongRec.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
(** This file defined the strong (course-of-value, well-founded) recursion
and proves its properties *)
Require Export NSub.
-Module NStrongRecPropFunct (Import N : NAxiomsSig').
-Include NSubPropFunct N.
+Ltac f_equiv' := repeat progress (f_equiv; try intros ? ? ?; auto).
+
+Module NStrongRecProp (Import N : NAxiomsRecSig').
+Include NSubProp N.
Section StrongRecursion.
@@ -51,30 +51,18 @@ Proof.
reflexivity.
Qed.
-(** We need a result similar to [f_equal], but for setoid equalities. *)
-Lemma f_equiv : forall f g x y,
- (N.eq==>Aeq)%signature f g -> N.eq x y -> Aeq (f x) (g y).
-Proof.
-auto.
-Qed.
-
Instance strong_rec0_wd :
Proper (Aeq ==> ((N.eq ==> Aeq) ==> N.eq ==> Aeq) ==> N.eq ==> N.eq ==> Aeq)
strong_rec0.
Proof.
-unfold strong_rec0.
-repeat red; intros.
-apply f_equiv; auto.
-apply recursion_wd; try red; auto.
+unfold strong_rec0; f_equiv'.
Qed.
Instance strong_rec_wd :
Proper (Aeq ==> ((N.eq ==> Aeq) ==> N.eq ==> Aeq) ==> N.eq ==> Aeq) strong_rec.
Proof.
intros a a' Eaa' f f' Eff' n n' Enn'.
-rewrite !strong_rec_alt.
-apply strong_rec0_wd; auto.
-now rewrite Enn'.
+rewrite !strong_rec_alt; f_equiv'.
Qed.
Section FixPoint.
@@ -92,18 +80,16 @@ Lemma strong_rec0_succ : forall a n m,
Aeq (strong_rec0 a f (S n) m) (f (strong_rec0 a f n) m).
Proof.
intros. unfold strong_rec0.
-apply f_equiv; auto with *.
-rewrite recursion_succ; try (repeat red; auto with *; fail).
-apply f_wd.
-apply recursion_wd; try red; auto with *.
+f_equiv.
+rewrite recursion_succ; f_equiv'.
+reflexivity.
Qed.
Lemma strong_rec_0 : forall a,
Aeq (strong_rec a f 0) (f (fun _ => a) 0).
Proof.
-intros. rewrite strong_rec_alt, strong_rec0_succ.
-apply f_wd; auto with *.
-red; intros; rewrite strong_rec0_0; auto with *.
+intros. rewrite strong_rec_alt, strong_rec0_succ; f_equiv'.
+rewrite strong_rec0_0. reflexivity.
Qed.
(* We need an assumption saying that for every n, the step function (f h n)
@@ -158,7 +144,7 @@ intros.
transitivity (f (fun n => strong_rec0 a f (S n) n) n).
rewrite strong_rec_alt.
apply strong_rec0_fixpoint.
-apply f_wd; auto with *.
+f_equiv.
intros x x' Hx; rewrite strong_rec_alt, Hx; auto with *.
Qed.
@@ -204,7 +190,7 @@ Qed.
End FixPoint.
End StrongRecursion.
-Implicit Arguments strong_rec [A].
+Arguments strong_rec [A] a f n.
-End NStrongRecPropFunct.
+End NStrongRecProp.
diff --git a/theories/Numbers/Natural/Abstract/NSub.v b/theories/Numbers/Natural/Abstract/NSub.v
index c0be3114..d7143c67 100644
--- a/theories/Numbers/Natural/Abstract/NSub.v
+++ b/theories/Numbers/Natural/Abstract/NSub.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,12 +8,10 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NSub.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Export NMulOrder.
-Module Type NSubPropFunct (Import N : NAxiomsSig').
-Include NMulOrderPropFunct N.
+Module Type NSubProp (Import N : NAxiomsMiniSig').
+Include NMulOrderProp N.
Theorem sub_0_l : forall n, 0 - n == 0.
Proof.
@@ -37,7 +35,7 @@ Qed.
Theorem sub_gt : forall n m, n > m -> n - m ~= 0.
Proof.
intros n m H; elim H using lt_ind_rel; clear n m H.
-solve_relation_wd.
+solve_proper.
intro; rewrite sub_0_r; apply neq_succ_0.
intros; now rewrite sub_succ.
Qed.
@@ -47,8 +45,8 @@ 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).
+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.
@@ -205,6 +203,26 @@ Proof.
intros n m p. rewrite add_comm; apply lt_add_lt_sub_r.
Qed.
+Theorem sub_lt : forall n m, m <= n -> 0 < m -> n - m < n.
+Proof.
+intros n m LE LT.
+assert (LE' := le_sub_l n m). rewrite lt_eq_cases in LE'.
+destruct LE' as [LT'|EQ]. assumption.
+apply add_sub_eq_nz in EQ; [|order].
+rewrite (add_lt_mono_r _ _ n), add_0_l in LT. order.
+Qed.
+
+Lemma sub_le_mono_r : forall n m p, n <= m -> n-p <= m-p.
+Proof.
+ intros. rewrite le_sub_le_add_r. transitivity m. assumption. apply sub_add_le.
+Qed.
+
+Lemma sub_le_mono_l : forall n m p, n <= m -> p-m <= p-n.
+Proof.
+ intros. rewrite le_sub_le_add_r.
+ transitivity (p-n+n); [ apply sub_add_le | now apply add_le_mono_l].
+Qed.
+
(** Sub and mul *)
Theorem mul_pred_r : forall n m, n * (P m) == n * m - n.
@@ -224,10 +242,10 @@ 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).
+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.
@@ -298,5 +316,5 @@ Theorem add_dichotomy :
forall n m, (exists p, p + n == m) \/ (exists p, p + m == n).
Proof. exact le_alt_dichotomy. Qed.
-End NSubPropFunct.
+End NSubProp.
diff --git a/theories/Numbers/Natural/BigN/BigN.v b/theories/Numbers/Natural/BigN/BigN.v
index 7c480862..7f205b38 100644
--- a/theories/Numbers/Natural/BigN/BigN.v
+++ b/theories/Numbers/Natural/BigN/BigN.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,7 +12,7 @@
Require Export Int31.
Require Import CyclicAxioms Cyclic31 Ring31 NSig NSigNAxioms NMake
- NProperties NDiv GenericMinMax.
+ NProperties GenericMinMax.
(** The following [BigN] module regroups both the operations and
all the abstract properties:
@@ -21,73 +21,63 @@ Require Import CyclicAxioms Cyclic31 Ring31 NSig NSigNAxioms NMake
w.r.t. ZArith
- [NTypeIsNAxioms] shows (mainly) that these operations implement
the interface [NAxioms]
- - [NPropSig] adds all generic properties derived from [NAxioms]
- - [NDivPropFunct] provides generic properties of [div] and [mod].
+ - [NProp] adds all generic properties derived from [NAxioms]
- [MinMax*Properties] provides properties of [min] and [max].
*)
-Module BigN <: NType <: OrderedTypeFull <: TotalOrder :=
- NMake.Make Int31Cyclic <+ NTypeIsNAxioms
- <+ !NPropSig <+ !NDivPropFunct <+ HasEqBool2Dec
- <+ !MinMaxLogicalProperties <+ !MinMaxDecProperties.
+Delimit Scope bigN_scope with bigN.
+Module BigN <: NType <: OrderedTypeFull <: TotalOrder.
+ Include NMake.Make Int31Cyclic [scope abstract_scope to bigN_scope].
+ Bind Scope bigN_scope with t t'.
+ Include NTypeIsNAxioms
+ <+ NProp [no inline]
+ <+ HasEqBool2Dec [no inline]
+ <+ MinMaxLogicalProperties [no inline]
+ <+ MinMaxDecProperties [no inline].
+End BigN.
+
+(** Nota concerning scopes : for the first Include, we cannot bind
+ the scope bigN_scope to a type that doesn't exists yet.
+ We hence need to explicitely declare the scope substitution.
+ For the next Include, the abstract type t (in scope abstract_scope)
+ gets substituted to concrete BigN.t (in scope bigN_scope),
+ and the corresponding argument scope are fixed automatically.
+*)
(** 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_.
-(* Bind Scope has no retroactive effect, let's declare scopes by hand. *)
-Arguments Scope BigN.to_Z [bigN_scope].
-Arguments Scope BigN.succ [bigN_scope].
-Arguments Scope BigN.pred [bigN_scope].
-Arguments Scope BigN.square [bigN_scope].
-Arguments Scope BigN.add [bigN_scope bigN_scope].
-Arguments Scope BigN.sub [bigN_scope bigN_scope].
-Arguments Scope BigN.mul [bigN_scope bigN_scope].
-Arguments Scope BigN.div [bigN_scope bigN_scope].
-Arguments Scope BigN.eq [bigN_scope bigN_scope].
-Arguments Scope BigN.lt [bigN_scope bigN_scope].
-Arguments Scope BigN.le [bigN_scope bigN_scope].
-Arguments Scope BigN.eq [bigN_scope bigN_scope].
-Arguments Scope BigN.compare [bigN_scope bigN_scope].
-Arguments Scope BigN.min [bigN_scope bigN_scope].
-Arguments Scope BigN.max [bigN_scope bigN_scope].
-Arguments Scope BigN.eq_bool [bigN_scope bigN_scope].
-Arguments Scope BigN.power_pos [bigN_scope positive_scope].
-Arguments Scope BigN.power [bigN_scope N_scope].
-Arguments Scope BigN.sqrt [bigN_scope].
-Arguments Scope BigN.div_eucl [bigN_scope bigN_scope].
-Arguments Scope BigN.modulo [bigN_scope bigN_scope].
-Arguments Scope BigN.gcd [bigN_scope bigN_scope].
+Local Open Scope bigN_scope.
+Notation bigN := BigN.t.
+Bind Scope bigN_scope with bigN BigN.t BigN.t'.
+Arguments BigN.N0 _%int31.
Local Notation "0" := BigN.zero : bigN_scope. (* temporary notation *)
Local Notation "1" := BigN.one : bigN_scope. (* temporary notation *)
+Local Notation "2" := BigN.two : bigN_scope. (* temporary notation *)
Infix "+" := BigN.add : bigN_scope.
Infix "-" := BigN.sub : bigN_scope.
Infix "*" := BigN.mul : bigN_scope.
Infix "/" := BigN.div : bigN_scope.
-Infix "^" := BigN.power : bigN_scope.
+Infix "^" := BigN.pow : bigN_scope.
Infix "?=" := BigN.compare : bigN_scope.
+Infix "=?" := BigN.eqb (at level 70, no associativity) : bigN_scope.
+Infix "<=?" := BigN.leb (at level 70, no associativity) : bigN_scope.
+Infix "<?" := BigN.ltb (at level 70, no associativity) : bigN_scope.
Infix "==" := BigN.eq (at level 70, no associativity) : bigN_scope.
-Notation "x != y" := (~x==y)%bigN (at level 70, no associativity) : bigN_scope.
+Notation "x != y" := (~x==y) (at level 70, no associativity) : bigN_scope.
Infix "<" := BigN.lt : bigN_scope.
Infix "<=" := BigN.le : bigN_scope.
-Notation "x > y" := (BigN.lt y x)(only parsing) : bigN_scope.
-Notation "x >= y" := (BigN.le y x)(only parsing) : bigN_scope.
-Notation "x < y < z" := (x<y /\ y<z)%bigN : bigN_scope.
-Notation "x < y <= z" := (x<y /\ y<=z)%bigN : bigN_scope.
-Notation "x <= y < z" := (x<=y /\ y<z)%bigN : bigN_scope.
-Notation "x <= y <= z" := (x<=y /\ y<=z)%bigN : bigN_scope.
+Notation "x > y" := (y < x) (only parsing) : bigN_scope.
+Notation "x >= y" := (y <= x) (only parsing) : bigN_scope.
+Notation "x < y < z" := (x<y /\ y<z) : bigN_scope.
+Notation "x < y <= z" := (x<y /\ y<=z) : bigN_scope.
+Notation "x <= y < z" := (x<=y /\ y<z) : bigN_scope.
+Notation "x <= y <= z" := (x<=y /\ y<=z) : bigN_scope.
Notation "[ i ]" := (BigN.to_Z i) : bigN_scope.
Infix "mod" := BigN.modulo (at level 40, no associativity) : bigN_scope.
-Local Open Scope bigN_scope.
-
(** Example of reasoning about [BigN] *)
Theorem succ_pred: forall q : bigN,
@@ -107,24 +97,24 @@ exact BigN.mul_1_l. exact BigN.mul_0_l. exact BigN.mul_comm.
exact BigN.mul_assoc. exact BigN.mul_add_distr_r.
Qed.
-Lemma BigNeqb_correct : forall x y, BigN.eq_bool x y = true -> x==y.
+Lemma BigNeqb_correct : forall x y, (x =? y) = true -> x==y.
Proof. now apply BigN.eqb_eq. Qed.
-Lemma BigNpower : power_theory 1 BigN.mul BigN.eq (@id N) BigN.power.
+Lemma BigNpower : power_theory 1 BigN.mul BigN.eq BigN.of_N BigN.pow.
Proof.
constructor.
-intros. red. rewrite BigN.spec_power. unfold id.
-destruct Zpower_theory as [EQ]. rewrite EQ.
+intros. red. rewrite BigN.spec_pow, BigN.spec_of_N.
+rewrite Zpower_theory.(rpow_pow_N).
destruct n; simpl. reflexivity.
induction p; simpl; intros; BigN.zify; rewrite ?IHp; auto.
Qed.
Lemma BigNdiv : div_theory BigN.eq BigN.add BigN.mul (@id _)
- (fun a b => if BigN.eq_bool b 0 then (0,a) else BigN.div_eucl a b).
+ (fun a b => if b =? 0 then (0,a) else BigN.div_eucl a b).
Proof.
constructor. unfold id. intros a b.
BigN.zify.
-generalize (Zeq_bool_if [b] 0); destruct (Zeq_bool [b] 0).
+case Z.eqb_spec.
BigN.zify. auto with zarith.
intros NEQ.
generalize (BigN.spec_div_eucl a b).
@@ -163,6 +153,7 @@ Ltac isBigNcst t :=
end
| BigN.zero => constr:true
| BigN.one => constr:true
+ | BigN.two => constr:true
| _ => constr:false
end.
@@ -172,6 +163,12 @@ Ltac BigNcst t :=
| false => constr:NotConstant
end.
+Ltac BigN_to_N t :=
+ match isBigNcst t with
+ | true => eval vm_compute in (BigN.to_N t)
+ | false => constr:NotConstant
+ end.
+
Ltac Ncst t :=
match isNcst t with
| true => constr:t
@@ -183,11 +180,11 @@ Ltac Ncst t :=
Add Ring BigNr : BigNring
(decidable BigNeqb_correct,
constants [BigNcst],
- power_tac BigNpower [Ncst],
+ power_tac BigNpower [BigN_to_N],
div BigNdiv).
Section TestRing.
-Let test : forall x y, 1 + x*y + x^2 + 1 == 1*1 + 1 + y*x + 1*x*x.
+Let test : forall x y, 1 + x*y^1 + x^2 + 1 == 1*1 + 1 + y*x + 1*x*x.
intros. ring_simplify. reflexivity.
Qed.
End TestRing.
diff --git a/theories/Numbers/Natural/BigN/NMake.v b/theories/Numbers/Natural/BigN/NMake.v
index 2b70f1bb..952f6183 100644
--- a/theories/Numbers/Natural/BigN/NMake.v
+++ b/theories/Numbers/Natural/BigN/NMake.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -16,18 +16,176 @@
representation. The representation-dependent (and macro-generated) part
is now in [NMake_gen]. *)
-Require Import BigNumPrelude ZArith CyclicAxioms.
-Require Import Nbasic Wf_nat StreamMemo NSig NMake_gen.
+Require Import Bool BigNumPrelude ZArith Nnat Ndigits CyclicAxioms DoubleType
+ Nbasic Wf_nat StreamMemo NSig NMake_gen.
-Module Make (Import W0:CyclicType) <: NType.
+Module Make (W0:CyclicType) <: NType.
- (** Macro-generated part *)
+ (** Let's include the macro-generated part. Even if we can't functorize
+ things (due to Eval red_t below), the rest of the module only uses
+ elements mentionned in interface [NAbstract]. *)
Include NMake_gen.Make W0.
+ Open Scope Z_scope.
+
+ Local Notation "[ x ]" := (to_Z x).
+
+ Definition eq (x y : t) := [x] = [y].
+
+ Declare Reduction red_t :=
+ lazy beta iota delta
+ [iter_t reduce same_level mk_t mk_t_S succ_t dom_t dom_op].
+
+ Ltac red_t :=
+ match goal with |- ?u => let v := (eval red_t in u) in change v end.
+
+ (** * Generic results *)
+
+ Tactic Notation "destr_t" constr(x) "as" simple_intropattern(pat) :=
+ destruct (destr_t x) as pat; cbv zeta;
+ rewrite ?iter_mk_t, ?spec_mk_t, ?spec_reduce.
+
+ Lemma spec_same_level : forall A (P:Z->Z->A->Prop)
+ (f : forall n, dom_t n -> dom_t n -> A),
+ (forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y)) ->
+ forall x y, P [x] [y] (same_level f x y).
+ Proof.
+ intros. apply spec_same_level_dep with (P:=fun _ => P); auto.
+ Qed.
+
+ Theorem spec_pos: forall x, 0 <= [x].
+ Proof.
+ intros x. destr_t x as (n,x). now case (ZnZ.spec_to_Z x).
+ Qed.
+
+ Lemma digits_dom_op_incr : forall n m, (n<=m)%nat ->
+ (ZnZ.digits (dom_op n) <= ZnZ.digits (dom_op m))%positive.
+ Proof.
+ intros.
+ change (Zpos (ZnZ.digits (dom_op n)) <= Zpos (ZnZ.digits (dom_op m))).
+ rewrite !digits_dom_op, !Pshiftl_nat_Zpower.
+ apply Zmult_le_compat_l; auto with zarith.
+ apply Zpower_le_monotone2; auto with zarith.
+ Qed.
+
+ Definition to_N (x : t) := Z.to_N (to_Z x).
+
+ (** * Zero, One *)
+
+ Definition zero := mk_t O ZnZ.zero.
+ Definition one := mk_t O ZnZ.one.
+
+ Theorem spec_0: [zero] = 0.
+ Proof.
+ unfold zero. rewrite spec_mk_t. exact ZnZ.spec_0.
+ Qed.
+
+ Theorem spec_1: [one] = 1.
+ Proof.
+ unfold one. rewrite spec_mk_t. exact ZnZ.spec_1.
+ Qed.
+
+ (** * Successor *)
+
+ (** NB: it is crucial here and for the rest of this file to preserve
+ the let-in's. They allow to pre-compute once and for all the
+ field access to Z/nZ initial structures (when n=0..6). *)
+
+ Local Notation succn := (fun n =>
+ let op := dom_op n in
+ let succ_c := ZnZ.succ_c in
+ let one := ZnZ.one in
+ fun x => match succ_c x with
+ | C0 r => mk_t n r
+ | C1 r => mk_t_S n (WW one r)
+ end).
+
+ Definition succ : t -> t := Eval red_t in iter_t succn.
+
+ Lemma succ_fold : succ = iter_t succn.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_succ: forall n, [succ n] = [n] + 1.
+ Proof.
+ intros x. rewrite succ_fold. destr_t x as (n,x).
+ generalize (ZnZ.spec_succ_c x); case ZnZ.succ_c.
+ intros. rewrite spec_mk_t. assumption.
+ intros. unfold interp_carry in *.
+ rewrite spec_mk_t_S. simpl. rewrite ZnZ.spec_1. assumption.
+ Qed.
+
+ (** Two *)
+
+ (** Not really pretty, but since W0 might be Z/2Z, we're not sure
+ there's a proper 2 there. *)
+
+ Definition two := succ one.
+
+ Lemma spec_2 : [two] = 2.
+ Proof.
+ unfold two. now rewrite spec_succ, spec_1.
+ Qed.
+
+ (** * Addition *)
+
+ Local Notation addn := (fun n =>
+ let op := dom_op n in
+ let add_c := ZnZ.add_c in
+ let one := ZnZ.one in
+ fun x y =>match add_c x y with
+ | C0 r => mk_t n r
+ | C1 r => mk_t_S n (WW one r)
+ end).
+
+ Definition add : t -> t -> t := Eval red_t in same_level addn.
+
+ Lemma add_fold : add = same_level addn.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_add: forall x y, [add x y] = [x] + [y].
+ Proof.
+ intros x y. rewrite add_fold. apply spec_same_level; clear x y.
+ intros n x y. simpl.
+ generalize (ZnZ.spec_add_c x y); case ZnZ.add_c; intros z H.
+ rewrite spec_mk_t. assumption.
+ rewrite spec_mk_t_S. unfold interp_carry in H.
+ simpl. rewrite ZnZ.spec_1. assumption.
+ Qed.
(** * Predecessor *)
+ Local Notation predn := (fun n =>
+ let pred_c := ZnZ.pred_c in
+ fun x => match pred_c x with
+ | C0 r => reduce n r
+ | C1 _ => zero
+ end).
+
+ Definition pred : t -> t := Eval red_t in iter_t predn.
+
+ Lemma pred_fold : pred = iter_t predn.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_pred_pos : forall x, 0 < [x] -> [pred x] = [x] - 1.
+ Proof.
+ intros x. rewrite pred_fold. destr_t x as (n,x). intros H.
+ generalize (ZnZ.spec_pred_c x); case ZnZ.pred_c; intros y H'.
+ rewrite spec_reduce. assumption.
+ exfalso. unfold interp_carry in *.
+ generalize (ZnZ.spec_to_Z x) (ZnZ.spec_to_Z y); auto with zarith.
+ Qed.
+
+ Theorem spec_pred0 : forall x, [x] = 0 -> [pred x] = 0.
+ Proof.
+ intros x. rewrite pred_fold. destr_t x as (n,x). intros H.
+ generalize (ZnZ.spec_pred_c x); case ZnZ.pred_c; intros y H'.
+ rewrite spec_reduce.
+ unfold interp_carry in H'.
+ generalize (ZnZ.spec_to_Z y); auto with zarith.
+ exact spec_0.
+ Qed.
+
Lemma spec_pred : forall x, [pred x] = Zmax 0 ([x]-1).
Proof.
intros. destruct (Zle_lt_or_eq _ _ (spec_pos x)).
@@ -36,9 +194,42 @@ Module Make (Import W0:CyclicType) <: NType.
rewrite <- H; apply spec_pred0; auto.
Qed.
-
(** * Subtraction *)
+ Local Notation subn := (fun n =>
+ let sub_c := ZnZ.sub_c in
+ fun x y => match sub_c x y with
+ | C0 r => reduce n r
+ | C1 r => zero
+ end).
+
+ Definition sub : t -> t -> t := Eval red_t in same_level subn.
+
+ Lemma sub_fold : sub = same_level subn.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_sub_pos : forall x y, [y] <= [x] -> [sub x y] = [x] - [y].
+ Proof.
+ intros x y. rewrite sub_fold. apply spec_same_level. clear x y.
+ intros n x y. simpl.
+ generalize (ZnZ.spec_sub_c x y); case ZnZ.sub_c; intros z H LE.
+ rewrite spec_reduce. assumption.
+ unfold interp_carry in H.
+ exfalso.
+ generalize (ZnZ.spec_to_Z z); auto with zarith.
+ Qed.
+
+ Theorem spec_sub0 : forall x y, [x] < [y] -> [sub x y] = 0.
+ Proof.
+ intros x y. rewrite sub_fold. apply spec_same_level. clear x y.
+ intros n x y. simpl.
+ generalize (ZnZ.spec_sub_c x y); case ZnZ.sub_c; intros z H LE.
+ rewrite spec_reduce.
+ unfold interp_carry in H.
+ generalize (ZnZ.spec_to_Z z); auto with zarith.
+ exact spec_0.
+ Qed.
+
Lemma spec_sub : forall x y, [sub x y] = Zmax 0 ([x]-[y]).
Proof.
intros. destruct (Zle_or_lt [y] [x]).
@@ -48,35 +239,112 @@ Module Make (Import W0:CyclicType) <: NType.
(** * Comparison *)
- Theorem spec_compare : forall x y, compare x y = Zcompare [x] [y].
+ Definition comparen_m n :
+ forall m, word (dom_t n) (S m) -> dom_t n -> comparison :=
+ let op := dom_op n in
+ let zero := @ZnZ.zero _ op in
+ let compare := @ZnZ.compare _ op in
+ let compare0 := compare zero in
+ fun m => compare_mn_1 (dom_t n) (dom_t n) zero compare compare0 compare (S m).
+
+ Let spec_comparen_m:
+ forall n m (x : word (dom_t n) (S m)) (y : dom_t n),
+ comparen_m n m x y = Zcompare (eval n (S m) x) (ZnZ.to_Z y).
+ Proof.
+ intros n m x y.
+ unfold comparen_m, eval.
+ rewrite nmake_double.
+ apply spec_compare_mn_1.
+ exact ZnZ.spec_0.
+ intros. apply ZnZ.spec_compare.
+ exact ZnZ.spec_to_Z.
+ exact ZnZ.spec_compare.
+ exact ZnZ.spec_compare.
+ exact ZnZ.spec_to_Z.
+ Qed.
+
+ Definition comparenm n m wx wy :=
+ let mn := Max.max n m in
+ let d := diff n m in
+ let op := make_op mn in
+ ZnZ.compare
+ (castm (diff_r n m) (extend_tr wx (snd d)))
+ (castm (diff_l n m) (extend_tr wy (fst d))).
+
+ Local Notation compare_folded :=
+ (iter_sym _
+ (fun n => @ZnZ.compare _ (dom_op n))
+ comparen_m
+ comparenm
+ CompOpp).
+
+ Definition compare : t -> t -> comparison :=
+ Eval lazy beta iota delta [iter_sym dom_op dom_t comparen_m] in
+ compare_folded.
+
+ Lemma compare_fold : compare = compare_folded.
Proof.
- intros x y. generalize (spec_compare_aux x y); destruct compare;
- intros; symmetry; try rewrite Zcompare_Eq_iff_eq; assumption.
+ lazy beta iota delta [iter_sym dom_op dom_t comparen_m]. reflexivity.
Qed.
- Definition eq_bool x y :=
+(** TODO: no need for ZnZ.Spec_rect , Spec_ind, and so on... *)
+
+ Theorem spec_compare : forall x y,
+ compare x y = Zcompare [x] [y].
+ Proof.
+ intros x y. rewrite compare_fold. apply spec_iter_sym; clear x y.
+ intros. apply ZnZ.spec_compare.
+ intros. cbv beta zeta. apply spec_comparen_m.
+ intros n m x y; unfold comparenm.
+ rewrite (spec_cast_l n m x), (spec_cast_r n m y).
+ unfold to_Z; apply ZnZ.spec_compare.
+ intros. subst. apply Zcompare_antisym.
+ Qed.
+
+ Definition eqb (x y : t) : bool :=
match compare x y with
| Eq => true
| _ => false
end.
- Theorem spec_eq_bool : forall x y, eq_bool x y = Zeq_bool [x] [y].
+ Theorem spec_eqb x y : eqb x y = Z.eqb [x] [y].
Proof.
- intros. unfold eq_bool, Zeq_bool. rewrite spec_compare; reflexivity.
+ apply eq_iff_eq_true.
+ unfold eqb. rewrite Z.eqb_eq, <- Z.compare_eq_iff, spec_compare.
+ split; [now destruct Z.compare | now intros ->].
Qed.
- Theorem spec_eq_bool_aux: forall x y,
- if eq_bool x y then [x] = [y] else [x] <> [y].
+ Definition lt (n m : t) := [n] < [m].
+ Definition le (n m : t) := [n] <= [m].
+
+ Definition ltb (x y : t) : bool :=
+ match compare x y with
+ | Lt => true
+ | _ => false
+ end.
+
+ Theorem spec_ltb x y : ltb x y = Z.ltb [x] [y].
Proof.
- intros x y; unfold eq_bool.
- generalize (spec_compare_aux x y); case compare; auto with zarith.
+ apply eq_iff_eq_true.
+ rewrite Z.ltb_lt. unfold Z.lt, ltb. rewrite spec_compare.
+ split; [now destruct Z.compare | now intros ->].
Qed.
- Definition lt n m := [n] < [m].
- Definition le n m := [n] <= [m].
+ Definition leb (x y : t) : bool :=
+ match compare x y with
+ | Gt => false
+ | _ => true
+ end.
+
+ Theorem spec_leb x y : leb x y = Z.leb [x] [y].
+ Proof.
+ apply eq_iff_eq_true.
+ rewrite Z.leb_le. unfold Z.le, leb. rewrite spec_compare.
+ destruct Z.compare; split; try easy. now destruct 1.
+ Qed.
- 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.
+ Definition min (n m : t) : t := match compare n m with Gt => m | _ => n end.
+ Definition max (n m : t) : t := match compare n m with Lt => m | _ => n end.
Theorem spec_max : forall n m, [max n m] = Zmax [n] [m].
Proof.
@@ -88,46 +356,239 @@ Module Make (Import W0:CyclicType) <: NType.
intros. unfold min, Zmin. rewrite spec_compare; destruct Zcompare; reflexivity.
Qed.
+ (** * Multiplication *)
+
+ Definition wn_mul n : forall m, word (dom_t n) (S m) -> dom_t n -> t :=
+ let op := dom_op n in
+ let zero := @ZnZ.zero _ op in
+ let succ := @ZnZ.succ _ op in
+ let add_c := @ZnZ.add_c _ op in
+ let mul_c := @ZnZ.mul_c _ op in
+ let ww := @ZnZ.WW _ op in
+ let ow := @ZnZ.OW _ op in
+ let eq0 := @ZnZ.eq0 _ op in
+ let mul_add := @DoubleMul.w_mul_add _ zero succ add_c mul_c in
+ let mul_add_n1 := @DoubleMul.double_mul_add_n1 _ zero ww ow mul_add in
+ fun m x y =>
+ let (w,r) := mul_add_n1 (S m) x y zero in
+ if eq0 w then mk_t_w' n m r
+ else mk_t_w' n (S m) (WW (extend n m w) r).
+
+ Definition mulnm n m x y :=
+ let mn := Max.max n m in
+ let d := diff n m in
+ let op := make_op mn in
+ reduce_n (S mn) (ZnZ.mul_c
+ (castm (diff_r n m) (extend_tr x (snd d)))
+ (castm (diff_l n m) (extend_tr y (fst d)))).
+
+ Local Notation mul_folded :=
+ (iter_sym _
+ (fun n => let mul_c := ZnZ.mul_c in
+ fun x y => reduce (S n) (succ_t _ (mul_c x y)))
+ wn_mul
+ mulnm
+ (fun x => x)).
+
+ Definition mul : t -> t -> t :=
+ Eval lazy beta iota delta
+ [iter_sym dom_op dom_t reduce succ_t extend zeron
+ wn_mul DoubleMul.w_mul_add mk_t_w'] in
+ mul_folded.
+
+ Lemma mul_fold : mul = mul_folded.
+ Proof.
+ lazy beta iota delta
+ [iter_sym dom_op dom_t reduce succ_t extend zeron
+ wn_mul DoubleMul.w_mul_add mk_t_w']. reflexivity.
+ Qed.
- (** * Power *)
+ Lemma spec_muln:
+ forall n (x: word _ (S n)) y,
+ [Nn (S n) (ZnZ.mul_c (Ops:=make_op n) x y)] = [Nn n x] * [Nn n y].
+ Proof.
+ intros n x y; unfold to_Z.
+ rewrite <- ZnZ.spec_mul_c.
+ rewrite make_op_S.
+ case ZnZ.mul_c; auto.
+ Qed.
- Fixpoint power_pos (x:t) (p:positive) {struct p} : t :=
- match p with
- | xH => x
- | xO p => square (power_pos x p)
- | xI p => mul (square (power_pos x p)) x
- end.
+ Lemma spec_mul_add_n1: forall n m x y z,
+ let (q,r) := DoubleMul.double_mul_add_n1 ZnZ.zero ZnZ.WW ZnZ.OW
+ (DoubleMul.w_mul_add ZnZ.zero ZnZ.succ ZnZ.add_c ZnZ.mul_c)
+ (S m) x y z in
+ ZnZ.to_Z q * (base (ZnZ.digits (nmake_op _ (dom_op n) (S m))))
+ + eval n (S m) r =
+ eval n (S m) x * ZnZ.to_Z y + ZnZ.to_Z z.
+ Proof.
+ intros n m x y z.
+ rewrite digits_nmake.
+ unfold eval. rewrite nmake_double.
+ apply DoubleMul.spec_double_mul_add_n1.
+ apply ZnZ.spec_0.
+ exact ZnZ.spec_WW.
+ exact ZnZ.spec_OW.
+ apply DoubleCyclic.spec_mul_add.
+ Qed.
- Theorem spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
+ Lemma spec_wn_mul : forall n m x y,
+ [wn_mul n m x y] = (eval n (S m) x) * ZnZ.to_Z y.
Proof.
- intros x n; generalize x; elim n; clear n x; simpl power_pos.
- intros; rewrite spec_mul; rewrite spec_square; rewrite H.
- rewrite Zpos_xI; rewrite Zpower_exp; auto with zarith.
- rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.
- rewrite Zpower_2; rewrite Zpower_1_r; auto.
- intros; rewrite spec_square; rewrite H.
- rewrite Zpos_xO; auto with zarith.
- rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.
- rewrite Zpower_2; auto.
- intros; rewrite Zpower_1_r; auto.
+ intros; unfold wn_mul.
+ generalize (spec_mul_add_n1 n m x y ZnZ.zero).
+ case DoubleMul.double_mul_add_n1; intros q r Hqr.
+ rewrite ZnZ.spec_0, Zplus_0_r in Hqr. rewrite <- Hqr.
+ generalize (ZnZ.spec_eq0 q); case ZnZ.eq0; intros HH.
+ rewrite HH; auto. simpl. apply spec_mk_t_w'.
+ clear.
+ rewrite spec_mk_t_w'.
+ set (m' := S m) in *.
+ unfold eval.
+ rewrite nmake_WW. f_equal. f_equal.
+ rewrite <- spec_mk_t.
+ symmetry. apply spec_extend.
Qed.
- Definition power x (n:N) := match n with
- | BinNat.N0 => one
- | BinNat.Npos p => power_pos x p
- end.
+ Theorem spec_mul : forall x y, [mul x y] = [x] * [y].
+ Proof.
+ intros x y. rewrite mul_fold. apply spec_iter_sym; clear x y.
+ intros n x y. cbv zeta beta.
+ rewrite spec_reduce, spec_succ_t, <- ZnZ.spec_mul_c; auto.
+ apply spec_wn_mul.
+ intros n m x y; unfold mulnm. rewrite spec_reduce_n.
+ rewrite (spec_cast_l n m x), (spec_cast_r n m y).
+ apply spec_muln.
+ intros. rewrite Zmult_comm; auto.
+ Qed.
- Theorem spec_power: forall x n, [power x n] = [x] ^ Z_of_N n.
+ (** * Division by a smaller number *)
+
+ Definition wn_divn1 n :=
+ let op := dom_op n in
+ let zd := ZnZ.zdigits op in
+ let zero := @ZnZ.zero _ op in
+ let ww := @ZnZ.WW _ op in
+ let head0 := @ZnZ.head0 _ op in
+ let add_mul_div := @ZnZ.add_mul_div _ op in
+ let div21 := @ZnZ.div21 _ op in
+ let compare := @ZnZ.compare _ op in
+ let sub := @ZnZ.sub _ op in
+ let ddivn1 :=
+ DoubleDivn1.double_divn1 zd zero ww head0 add_mul_div div21 compare sub in
+ fun m x y => let (u,v) := ddivn1 (S m) x y in (mk_t_w' n m u, mk_t n v).
+
+ Let div_gtnm n m wx wy :=
+ let mn := Max.max n m in
+ let d := diff n m in
+ let op := make_op mn in
+ let (q, r):= ZnZ.div_gt
+ (castm (diff_r n m) (extend_tr wx (snd d)))
+ (castm (diff_l n m) (extend_tr wy (fst d))) in
+ (reduce_n mn q, reduce_n mn r).
+
+ Local Notation div_gt_folded :=
+ (iter _
+ (fun n => let div_gt := ZnZ.div_gt in
+ fun x y => let (u,v) := div_gt x y in (reduce n u, reduce n v))
+ (fun n =>
+ let div_gt := ZnZ.div_gt in
+ fun m x y =>
+ let y' := DoubleBase.get_low (zeron n) (S m) y in
+ let (u,v) := div_gt x y' in (reduce n u, reduce n v))
+ wn_divn1
+ div_gtnm).
+
+ Definition div_gt :=
+ Eval lazy beta iota delta
+ [iter dom_op dom_t reduce zeron wn_divn1 mk_t_w' mk_t] in
+ div_gt_folded.
+
+ Lemma div_gt_fold : div_gt = div_gt_folded.
Proof.
- destruct n; simpl. apply (spec_1 w0_spec).
- apply spec_power_pos.
+ lazy beta iota delta [iter dom_op dom_t reduce zeron wn_divn1 mk_t_w' mk_t].
+ reflexivity.
Qed.
+ Lemma spec_get_endn: forall n m x y,
+ eval n m x <= [mk_t n y] ->
+ [mk_t n (DoubleBase.get_low (zeron n) m x)] = eval n m x.
+ Proof.
+ intros n m x y H.
+ unfold eval. rewrite nmake_double.
+ rewrite spec_mk_t in *.
+ apply DoubleBase.spec_get_low.
+ apply spec_zeron.
+ exact ZnZ.spec_to_Z.
+ apply Zle_lt_trans with (ZnZ.to_Z y); auto.
+ rewrite <- nmake_double; auto.
+ case (ZnZ.spec_to_Z y); auto.
+ Qed.
- (** * Div *)
+ Let spec_divn1 n :=
+ DoubleDivn1.spec_double_divn1
+ (ZnZ.zdigits (dom_op n)) (ZnZ.zero:dom_t n)
+ ZnZ.WW ZnZ.head0
+ ZnZ.add_mul_div ZnZ.div21
+ ZnZ.compare ZnZ.sub ZnZ.to_Z
+ ZnZ.spec_to_Z
+ ZnZ.spec_zdigits
+ ZnZ.spec_0 ZnZ.spec_WW ZnZ.spec_head0
+ ZnZ.spec_add_mul_div ZnZ.spec_div21
+ ZnZ.spec_compare ZnZ.spec_sub.
+
+ Lemma spec_div_gt_aux : forall x y, [x] > [y] -> 0 < [y] ->
+ let (q,r) := div_gt x y in
+ [x] = [q] * [y] + [r] /\ 0 <= [r] < [y].
+ Proof.
+ intros x y. rewrite div_gt_fold. apply spec_iter; clear x y.
+ intros n x y H1 H2. simpl.
+ generalize (ZnZ.spec_div_gt x y H1 H2); case ZnZ.div_gt.
+ intros u v. rewrite 2 spec_reduce. auto.
+ intros n m x y H1 H2. cbv zeta beta.
+ generalize (ZnZ.spec_div_gt x
+ (DoubleBase.get_low (zeron n) (S m) y)).
+ case ZnZ.div_gt.
+ intros u v H3; repeat rewrite spec_reduce.
+ generalize (spec_get_endn n (S m) y x). rewrite !spec_mk_t. intros H4.
+ rewrite H4 in H3; auto with zarith.
+ intros n m x y H1 H2.
+ generalize (spec_divn1 n (S m) x y H2).
+ unfold wn_divn1; case DoubleDivn1.double_divn1.
+ intros u v H3.
+ rewrite spec_mk_t_w', spec_mk_t.
+ rewrite <- !nmake_double in H3; auto.
+ intros n m x y H1 H2; unfold div_gtnm.
+ generalize (ZnZ.spec_div_gt
+ (castm (diff_r n m)
+ (extend_tr x (snd (diff n m))))
+ (castm (diff_l n m)
+ (extend_tr y (fst (diff n m))))).
+ case ZnZ.div_gt.
+ intros xx yy HH.
+ repeat rewrite spec_reduce_n.
+ rewrite (spec_cast_l n m x), (spec_cast_r n m y).
+ unfold to_Z; apply HH.
+ rewrite (spec_cast_l n m x) in H1; auto.
+ rewrite (spec_cast_r n m y) in H1; auto.
+ rewrite (spec_cast_r n m y) in H2; auto.
+ Qed.
+
+ Theorem spec_div_gt: forall x y, [x] > [y] -> 0 < [y] ->
+ let (q,r) := div_gt x y in
+ [q] = [x] / [y] /\ [r] = [x] mod [y].
+ Proof.
+ intros x y H1 H2; generalize (spec_div_gt_aux x y H1 H2); case div_gt.
+ intros q r (H3, H4); split.
+ apply (Zdiv_unique [x] [y] [q] [r]); auto.
+ rewrite Zmult_comm; auto.
+ apply (Zmod_unique [x] [y] [q] [r]); auto.
+ rewrite Zmult_comm; auto.
+ Qed.
- Definition div_eucl x y :=
- if eq_bool y zero then (zero,zero) else
+ (** * General Division *)
+
+ Definition div_eucl (x y : t) : t * t :=
+ if eqb y zero then (zero,zero) else
match compare x y with
| Eq => (one, zero)
| Lt => (zero, x)
@@ -138,32 +599,27 @@ Module Make (Import W0:CyclicType) <: NType.
let (q,r) := div_eucl x y in
([q], [r]) = Zdiv_eucl [x] [y].
Proof.
- assert (F0: [zero] = 0).
- exact (spec_0 w0_spec).
- assert (F1: [one] = 1).
- exact (spec_1 w0_spec).
intros x y. unfold div_eucl.
- generalize (spec_eq_bool_aux y zero). destruct eq_bool; rewrite F0.
- intro H. rewrite H. destruct [x]; auto.
- intro H'.
- assert (0 < [y]) by (generalize (spec_pos y); auto with zarith).
+ rewrite spec_eqb, spec_compare, spec_0.
+ case Z.eqb_spec.
+ intros ->. rewrite spec_0. destruct [x]; auto.
+ intros H'.
+ assert (H : 0 < [y]) by (generalize (spec_pos y); auto with zarith).
clear H'.
- generalize (spec_compare_aux x y); case compare; try rewrite F0;
- try rewrite F1; intros; auto with zarith.
- rewrite H0; generalize (Z_div_same [y] (Zlt_gt _ _ H))
- (Z_mod_same [y] (Zlt_gt _ _ H));
+ case Zcompare_spec; intros Cmp;
+ rewrite ?spec_0, ?spec_1; intros; auto with zarith.
+ rewrite Cmp; generalize (Z_div_same [y] (Zlt_gt _ _ H))
+ (Z_mod_same [y] (Zlt_gt _ _ H));
unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.
- assert (F2: 0 <= [x] < [y]).
- generalize (spec_pos x); auto.
- generalize (Zdiv_small _ _ F2)
- (Zmod_small _ _ F2);
+ assert (LeLt: 0 <= [x] < [y]) by (generalize (spec_pos x); auto).
+ generalize (Zdiv_small _ _ LeLt) (Zmod_small _ _ LeLt);
unfold Zdiv, Zmod; case Zdiv_eucl; intros; subst; auto.
- generalize (spec_div_gt _ _ H0 H); auto.
+ generalize (spec_div_gt _ _ (Zlt_gt _ _ Cmp) H); auto.
unfold Zdiv, Zmod; case Zdiv_eucl; case div_gt.
intros a b c d (H1, H2); subst; auto.
Qed.
- Definition div x y := fst (div_eucl x y).
+ Definition div (x y : t) : t := fst (div_eucl x y).
Theorem spec_div:
forall x y, [div x y] = [x] / [y].
@@ -174,11 +630,90 @@ Module Make (Import W0:CyclicType) <: NType.
injection H; auto.
Qed.
+ (** * Modulo by a smaller number *)
+
+ Definition wn_modn1 n :=
+ let op := dom_op n in
+ let zd := ZnZ.zdigits op in
+ let zero := @ZnZ.zero _ op in
+ let head0 := @ZnZ.head0 _ op in
+ let add_mul_div := @ZnZ.add_mul_div _ op in
+ let div21 := @ZnZ.div21 _ op in
+ let compare := @ZnZ.compare _ op in
+ let sub := @ZnZ.sub _ op in
+ let dmodn1 :=
+ DoubleDivn1.double_modn1 zd zero head0 add_mul_div div21 compare sub in
+ fun m x y => reduce n (dmodn1 (S m) x y).
+
+ Let mod_gtnm n m wx wy :=
+ let mn := Max.max n m in
+ let d := diff n m in
+ let op := make_op mn in
+ reduce_n mn (ZnZ.modulo_gt
+ (castm (diff_r n m) (extend_tr wx (snd d)))
+ (castm (diff_l n m) (extend_tr wy (fst d)))).
+
+ Local Notation mod_gt_folded :=
+ (iter _
+ (fun n => let modulo_gt := ZnZ.modulo_gt in
+ fun x y => reduce n (modulo_gt x y))
+ (fun n => let modulo_gt := ZnZ.modulo_gt in
+ fun m x y =>
+ reduce n (modulo_gt x (DoubleBase.get_low (zeron n) (S m) y)))
+ wn_modn1
+ mod_gtnm).
+
+ Definition mod_gt :=
+ Eval lazy beta iota delta [iter dom_op dom_t reduce wn_modn1 zeron] in
+ mod_gt_folded.
+
+ Lemma mod_gt_fold : mod_gt = mod_gt_folded.
+ Proof.
+ lazy beta iota delta [iter dom_op dom_t reduce wn_modn1 zeron].
+ reflexivity.
+ Qed.
+
+ Let spec_modn1 n :=
+ DoubleDivn1.spec_double_modn1
+ (ZnZ.zdigits (dom_op n)) (ZnZ.zero:dom_t n)
+ ZnZ.WW ZnZ.head0
+ ZnZ.add_mul_div ZnZ.div21
+ ZnZ.compare ZnZ.sub ZnZ.to_Z
+ ZnZ.spec_to_Z
+ ZnZ.spec_zdigits
+ ZnZ.spec_0 ZnZ.spec_WW ZnZ.spec_head0
+ ZnZ.spec_add_mul_div ZnZ.spec_div21
+ ZnZ.spec_compare ZnZ.spec_sub.
+
+ Theorem spec_mod_gt:
+ forall x y, [x] > [y] -> 0 < [y] -> [mod_gt x y] = [x] mod [y].
+ Proof.
+ intros x y. rewrite mod_gt_fold. apply spec_iter; clear x y.
+ intros n x y H1 H2. simpl. rewrite spec_reduce.
+ exact (ZnZ.spec_modulo_gt x y H1 H2).
+ intros n m x y H1 H2. cbv zeta beta. rewrite spec_reduce.
+ rewrite <- spec_mk_t in H1.
+ rewrite <- (spec_get_endn n (S m) y x); auto with zarith.
+ rewrite spec_mk_t.
+ apply ZnZ.spec_modulo_gt; auto.
+ rewrite <- (spec_get_endn n (S m) y x), !spec_mk_t in H1; auto with zarith.
+ rewrite <- (spec_get_endn n (S m) y x), !spec_mk_t in H2; auto with zarith.
+ intros n m x y H1 H2. unfold wn_modn1. rewrite spec_reduce.
+ unfold eval; rewrite nmake_double.
+ apply (spec_modn1 n); auto.
+ intros n m x y H1 H2; unfold mod_gtnm.
+ repeat rewrite spec_reduce_n.
+ rewrite (spec_cast_l n m x), (spec_cast_r n m y).
+ unfold to_Z; apply ZnZ.spec_modulo_gt.
+ rewrite (spec_cast_l n m x) in H1; auto.
+ rewrite (spec_cast_r n m y) in H1; auto.
+ rewrite (spec_cast_r n m y) in H2; auto.
+ Qed.
- (** * Modulo *)
+ (** * General Modulo *)
- Definition modulo x y :=
- if eq_bool y zero then zero else
+ Definition modulo (x y : t) : t :=
+ if eqb y zero then zero else
match compare x y with
| Eq => zero
| Lt => x
@@ -188,24 +723,129 @@ Module Make (Import W0:CyclicType) <: NType.
Theorem spec_modulo:
forall x y, [modulo x y] = [x] mod [y].
Proof.
- assert (F0: [zero] = 0).
- exact (spec_0 w0_spec).
- assert (F1: [one] = 1).
- exact (spec_1 w0_spec).
intros x y. unfold modulo.
- generalize (spec_eq_bool_aux y zero). destruct eq_bool; rewrite F0.
- intro H; rewrite H. destruct [x]; auto.
+ rewrite spec_eqb, spec_compare, spec_0.
+ case Z.eqb_spec.
+ intros ->; rewrite spec_0. destruct [x]; auto.
intro H'.
assert (H : 0 < [y]) by (generalize (spec_pos y); auto with zarith).
clear H'.
- generalize (spec_compare_aux x y); case compare; try rewrite F0;
- try rewrite F1; intros; try split; auto with zarith.
+ case Zcompare_spec;
+ rewrite ?spec_0, ?spec_1; intros; try split; auto with zarith.
rewrite H0; apply sym_equal; apply Z_mod_same; auto with zarith.
apply sym_equal; apply Zmod_small; auto with zarith.
generalize (spec_pos x); auto with zarith.
- apply spec_mod_gt; auto.
+ apply spec_mod_gt; auto with zarith.
+ Qed.
+
+ (** * Square *)
+
+ Local Notation squaren := (fun n =>
+ let square_c := ZnZ.square_c in
+ fun x => reduce (S n) (succ_t _ (square_c x))).
+
+ Definition square : t -> t := Eval red_t in iter_t squaren.
+
+ Lemma square_fold : square = iter_t squaren.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_square: forall x, [square x] = [x] * [x].
+ Proof.
+ intros x. rewrite square_fold. destr_t x as (n,x).
+ rewrite spec_succ_t. exact (ZnZ.spec_square_c x).
+ Qed.
+
+ (** * Square Root *)
+
+ Local Notation sqrtn := (fun n =>
+ let sqrt := ZnZ.sqrt in
+ fun x => reduce n (sqrt x)).
+
+ Definition sqrt : t -> t := Eval red_t in iter_t sqrtn.
+
+ Lemma sqrt_fold : sqrt = iter_t sqrtn.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_sqrt_aux: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
+ Proof.
+ intros x. rewrite sqrt_fold. destr_t x as (n,x). exact (ZnZ.spec_sqrt x).
+ Qed.
+
+ Theorem spec_sqrt: forall x, [sqrt x] = Z.sqrt [x].
+ Proof.
+ intros x.
+ symmetry. apply Z.sqrt_unique.
+ rewrite <- ! Zpower_2. apply spec_sqrt_aux.
+ Qed.
+
+ (** * Power *)
+
+ Fixpoint pow_pos (x:t)(p:positive) : t :=
+ match p with
+ | xH => x
+ | xO p => square (pow_pos x p)
+ | xI p => mul (square (pow_pos x p)) x
+ end.
+
+ Theorem spec_pow_pos: forall x n, [pow_pos x n] = [x] ^ Zpos n.
+ Proof.
+ intros x n; generalize x; elim n; clear n x; simpl pow_pos.
+ intros; rewrite spec_mul; rewrite spec_square; rewrite H.
+ rewrite Zpos_xI; rewrite Zpower_exp; auto with zarith.
+ rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.
+ rewrite Zpower_2; rewrite Zpower_1_r; auto.
+ intros; rewrite spec_square; rewrite H.
+ rewrite Zpos_xO; auto with zarith.
+ rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.
+ rewrite Zpower_2; auto.
+ intros; rewrite Zpower_1_r; auto.
Qed.
+ Definition pow_N (x:t)(n:N) : t := match n with
+ | BinNat.N0 => one
+ | BinNat.Npos p => pow_pos x p
+ end.
+
+ Theorem spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z_of_N n.
+ Proof.
+ destruct n; simpl. apply spec_1.
+ apply spec_pow_pos.
+ Qed.
+
+ Definition pow (x y:t) : t := pow_N x (to_N y).
+
+ Theorem spec_pow : forall x y, [pow x y] = [x] ^ [y].
+ Proof.
+ intros. unfold pow, to_N.
+ now rewrite spec_pow_N, Z2N.id by apply spec_pos.
+ Qed.
+
+
+ (** * digits
+
+ Number of digits in the representation of a numbers
+ (including head zero's).
+ NB: This function isn't a morphism for setoid [eq].
+ *)
+
+ Local Notation digitsn := (fun n =>
+ let digits := ZnZ.digits (dom_op n) in
+ fun _ => digits).
+
+ Definition digits : t -> positive := Eval red_t in iter_t digitsn.
+
+ Lemma digits_fold : digits = iter_t digitsn.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_digits: forall x, 0 <= [x] < 2 ^ Zpos (digits x).
+ Proof.
+ intros x. rewrite digits_fold. destr_t x as (n,x). exact (ZnZ.spec_to_Z x).
+ Qed.
+
+ Lemma digits_level : forall x, digits x = ZnZ.digits (dom_op (level x)).
+ Proof.
+ intros x. rewrite digits_fold. unfold level. destr_t x as (n,x). reflexivity.
+ Qed.
(** * Gcd *)
@@ -226,15 +866,12 @@ Module Make (Import W0:CyclicType) <: NType.
Zis_gcd [a1] [b1] [cont a1 b1]) ->
Zis_gcd [a] [b] [gcd_gt_body a b cont].
Proof.
- assert (F1: [zero] = 0).
- unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.
intros a b cont p H2 H3 H4; unfold gcd_gt_body.
- generalize (spec_compare_aux b zero); case compare; try rewrite F1.
- intros HH; rewrite HH; apply Zis_gcd_0.
+ rewrite ! spec_compare, spec_0. case Zcompare_spec.
+ intros ->; apply Zis_gcd_0.
intros HH; absurd (0 <= [b]); auto with zarith.
case (spec_digits b); auto with zarith.
- intros H5; generalize (spec_compare_aux (mod_gt a b) zero);
- case compare; try rewrite F1.
+ intros H5; case Zcompare_spec.
intros H6; rewrite <- (Zmult_1_r [b]).
rewrite (Z_div_mod_eq [a] [b]); auto with zarith.
rewrite <- spec_mod_gt; auto with zarith.
@@ -273,7 +910,7 @@ Module Make (Import W0:CyclicType) <: NType.
intros HH; generalize H3; rewrite <- HH; simpl Zpower; auto with zarith.
Qed.
- Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) {struct p} : t :=
+ Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) : t :=
gcd_gt_body a b
(fun a b =>
match p with
@@ -310,12 +947,7 @@ Module Make (Import W0:CyclicType) <: NType.
(Zpos p + n - 1); auto with zarith.
intros a3 b3 H12 H13; apply H4; auto with zarith.
apply Zlt_le_trans with (1 := H12).
- case (Zle_or_lt 1 n); intros HH.
- apply Zpower_le_monotone; auto with zarith.
- apply Zle_trans with 0; auto with zarith.
- assert (HH1: n - 1 < 0); auto with zarith.
- generalize HH1; case (n - 1); auto with zarith.
- intros p1 HH2; discriminate.
+ apply Zpower_le_monotone2; auto with zarith.
intros n a b cont H H2 H3.
simpl gcd_gt_aux.
apply Zspec_gcd_gt_body with (n + 1); auto with zarith.
@@ -345,7 +977,7 @@ Module Make (Import W0:CyclicType) <: NType.
intros; apply False_ind; auto with zarith.
Qed.
- Definition gcd a b :=
+ Definition gcd (a b : t) : t :=
match compare a b with
| Eq => a
| Lt => gcd_gt b a
@@ -357,7 +989,7 @@ Module Make (Import W0:CyclicType) <: NType.
intros a b.
case (spec_digits a); intros H1 H2.
case (spec_digits b); intros H3 H4.
- unfold gcd; generalize (spec_compare_aux a b); case compare.
+ unfold gcd. rewrite spec_compare. case Zcompare_spec.
intros HH; rewrite HH; apply sym_equal; apply Zis_gcd_gcd; auto.
apply Zis_gcd_refl.
intros; apply trans_equal with (Zgcd [b] [a]).
@@ -365,13 +997,91 @@ Module Make (Import W0:CyclicType) <: NType.
apply Zis_gcd_gcd; auto with zarith.
apply Zgcd_is_pos.
apply Zis_gcd_sym; apply Zgcd_is_gcd.
- intros; apply spec_gcd_gt; auto.
+ intros; apply spec_gcd_gt; auto with zarith.
+ Qed.
+
+ (** * Parity test *)
+
+ Definition even : t -> bool := Eval red_t in
+ iter_t (fun n x => ZnZ.is_even x).
+
+ Definition odd x := negb (even x).
+
+ Lemma even_fold : even = iter_t (fun n x => ZnZ.is_even x).
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_even_aux: forall x,
+ if even x then [x] mod 2 = 0 else [x] mod 2 = 1.
+ Proof.
+ intros x. rewrite even_fold. destr_t x as (n,x).
+ exact (ZnZ.spec_is_even x).
+ Qed.
+
+ Theorem spec_even: forall x, even x = Zeven_bool [x].
+ Proof.
+ intros x. assert (H := spec_even_aux x). symmetry.
+ rewrite (Z_div_mod_eq_full [x] 2); auto with zarith.
+ destruct (even x); rewrite H, ?Zplus_0_r.
+ rewrite Zeven_bool_iff. apply Zeven_2p.
+ apply not_true_is_false. rewrite Zeven_bool_iff.
+ apply Zodd_not_Zeven. apply Zodd_2p_plus_1.
Qed.
+ Theorem spec_odd: forall x, odd x = Zodd_bool [x].
+ Proof.
+ intros x. unfold odd.
+ assert (H := spec_even_aux x). symmetry.
+ rewrite (Z_div_mod_eq_full [x] 2); auto with zarith.
+ destruct (even x); rewrite H, ?Zplus_0_r; simpl negb.
+ apply not_true_is_false. rewrite Zodd_bool_iff.
+ apply Zeven_not_Zodd. apply Zeven_2p.
+ apply Zodd_bool_iff. apply Zodd_2p_plus_1.
+ Qed.
(** * Conversion *)
- Definition of_N x :=
+ Definition pheight p :=
+ Peano.pred (nat_of_P (get_height (ZnZ.digits (dom_op 0)) (plength p))).
+
+ Theorem pheight_correct: forall p,
+ Zpos p < 2 ^ (Zpos (ZnZ.digits (dom_op 0)) * 2 ^ (Z_of_nat (pheight p))).
+ Proof.
+ intros p; unfold pheight.
+ assert (F1: forall x, Z_of_nat (Peano.pred (nat_of_P x)) = Zpos x - 1).
+ intros x.
+ assert (Zsucc (Z_of_nat (Peano.pred (nat_of_P x))) = Zpos x); auto with zarith.
+ rewrite <- inj_S.
+ rewrite <- (fun x => S_pred x 0); auto with zarith.
+ rewrite Zpos_eq_Z_of_nat_o_nat_of_P; auto.
+ apply lt_le_trans with 1%nat; auto with zarith.
+ exact (le_Pmult_nat x 1).
+ rewrite F1; clear F1.
+ assert (F2:= (get_height_correct (ZnZ.digits (dom_op 0)) (plength p))).
+ apply Zlt_le_trans with (Zpos (Psucc p)).
+ rewrite Zpos_succ_morphism; auto with zarith.
+ apply Zle_trans with (1 := plength_pred_correct (Psucc p)).
+ rewrite Ppred_succ.
+ apply Zpower_le_monotone2; auto with zarith.
+ Qed.
+
+ Definition of_pos (x:positive) : t :=
+ let n := pheight x in
+ reduce n (snd (ZnZ.of_pos x)).
+
+ Theorem spec_of_pos: forall x,
+ [of_pos x] = Zpos x.
+ Proof.
+ intros x; unfold of_pos.
+ rewrite spec_reduce.
+ simpl.
+ apply ZnZ.of_pos_correct.
+ unfold base.
+ apply Zlt_le_trans with (1 := pheight_correct x).
+ apply Zpower_le_monotone2; auto with zarith.
+ rewrite (digits_dom_op (_ _)), Pshiftl_nat_Zpower. auto with zarith.
+ Qed.
+
+ Definition of_N (x:N) : t :=
match x with
| BinNat.N0 => zero
| Npos p => of_pos p
@@ -381,51 +1091,437 @@ Module Make (Import W0:CyclicType) <: NType.
[of_N x] = Z_of_N x.
Proof.
intros x; case x.
- simpl of_N.
- unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.
+ simpl of_N. exact spec_0.
intros p; exact (spec_of_pos p).
Qed.
+ (** * [head0] and [tail0]
- (** * Shift *)
+ Number of zero at the beginning and at the end of
+ the representation of the number.
+ NB: these functions are not morphism for setoid [eq].
+ *)
- Definition shiftr n x :=
- match compare n (Ndigits x) with
- | Lt => unsafe_shiftr n x
- | _ => N0 w_0
- end.
+ Local Notation head0n := (fun n =>
+ let head0 := ZnZ.head0 in
+ fun x => reduce n (head0 x)).
+
+ Definition head0 : t -> t := Eval red_t in iter_t head0n.
+
+ Lemma head0_fold : head0 = iter_t head0n.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_head00: forall x, [x] = 0 -> [head0 x] = Zpos (digits x).
+ Proof.
+ intros x. rewrite head0_fold, digits_fold. destr_t x as (n,x).
+ exact (ZnZ.spec_head00 x).
+ Qed.
+
+ Lemma pow2_pos_minus_1 : forall z, 0<z -> 2^(z-1) = 2^z / 2.
+ Proof.
+ intros. apply Zdiv_unique with 0; auto with zarith.
+ change 2 with (2^1) at 2.
+ rewrite <- Zpower_exp; auto with zarith.
+ rewrite Zplus_0_r. f_equal. auto with zarith.
+ Qed.
- Theorem spec_shiftr: forall n x,
- [shiftr n x] = [x] / 2 ^ [n].
- Proof.
- intros n x; unfold shiftr;
- generalize (spec_compare_aux n (Ndigits x)); case compare; intros H.
- apply trans_equal with (1 := spec_0 w0_spec).
- apply sym_equal; apply Zdiv_small; rewrite H.
- rewrite spec_Ndigits; exact (spec_digits x).
- rewrite <- spec_unsafe_shiftr; auto with zarith.
- apply trans_equal with (1 := spec_0 w0_spec).
- apply sym_equal; apply Zdiv_small.
- rewrite spec_Ndigits in H; case (spec_digits x); intros H1 H2.
- split; auto.
- apply Zlt_le_trans with (1 := H2).
- apply Zpower_le_monotone; auto with zarith.
- Qed.
-
- Definition shiftl_aux_body cont n x :=
- match compare n (head0 x) with
- Gt => cont n (double_size x)
- | _ => unsafe_shiftl n x
+ Theorem spec_head0: forall x, 0 < [x] ->
+ 2 ^ (Zpos (digits x) - 1) <= 2 ^ [head0 x] * [x] < 2 ^ Zpos (digits x).
+ Proof.
+ intros x. rewrite pow2_pos_minus_1 by (red; auto).
+ rewrite head0_fold, digits_fold. destr_t x as (n,x). exact (ZnZ.spec_head0 x).
+ Qed.
+
+ Local Notation tail0n := (fun n =>
+ let tail0 := ZnZ.tail0 in
+ fun x => reduce n (tail0 x)).
+
+ Definition tail0 : t -> t := Eval red_t in iter_t tail0n.
+
+ Lemma tail0_fold : tail0 = iter_t tail0n.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_tail00: forall x, [x] = 0 -> [tail0 x] = Zpos (digits x).
+ Proof.
+ intros x. rewrite tail0_fold, digits_fold. destr_t x as (n,x).
+ exact (ZnZ.spec_tail00 x).
+ Qed.
+
+ Theorem spec_tail0: forall x,
+ 0 < [x] -> exists y, 0 <= y /\ [x] = (2 * y + 1) * 2 ^ [tail0 x].
+ Proof.
+ intros x. rewrite tail0_fold. destr_t x as (n,x). exact (ZnZ.spec_tail0 x).
+ Qed.
+
+ (** * [Ndigits]
+
+ Same as [digits] but encoded using large integers
+ NB: this function is not a morphism for setoid [eq].
+ *)
+
+ Local Notation Ndigitsn := (fun n =>
+ let d := reduce n (ZnZ.zdigits (dom_op n)) in
+ fun _ => d).
+
+ Definition Ndigits : t -> t := Eval red_t in iter_t Ndigitsn.
+
+ Lemma Ndigits_fold : Ndigits = iter_t Ndigitsn.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_Ndigits: forall x, [Ndigits x] = Zpos (digits x).
+ Proof.
+ intros x. rewrite Ndigits_fold, digits_fold. destr_t x as (n,x).
+ apply ZnZ.spec_zdigits.
+ Qed.
+
+ (** * Binary logarithm *)
+
+ Local Notation log2n := (fun n =>
+ let op := dom_op n in
+ let zdigits := ZnZ.zdigits op in
+ let head0 := ZnZ.head0 in
+ let sub_carry := ZnZ.sub_carry in
+ fun x => reduce n (sub_carry zdigits (head0 x))).
+
+ Definition log2 : t -> t := Eval red_t in
+ let log2 := iter_t log2n in
+ fun x => if eqb x zero then zero else log2 x.
+
+ Lemma log2_fold :
+ log2 = fun x => if eqb x zero then zero else iter_t log2n x.
+ Proof. red_t; reflexivity. Qed.
+
+ Lemma spec_log2_0 : forall x, [x] = 0 -> [log2 x] = 0.
+ Proof.
+ intros x H. rewrite log2_fold.
+ rewrite spec_eqb, H. rewrite spec_0. simpl. exact spec_0.
+ Qed.
+
+ Lemma head0_zdigits : forall n (x : dom_t n),
+ 0 < ZnZ.to_Z x ->
+ ZnZ.to_Z (ZnZ.head0 x) < ZnZ.to_Z (ZnZ.zdigits (dom_op n)).
+ Proof.
+ intros n x H.
+ destruct (ZnZ.spec_head0 x H) as (_,H0).
+ intros.
+ assert (H1 := ZnZ.spec_to_Z (ZnZ.head0 x)).
+ assert (H2 := ZnZ.spec_to_Z (ZnZ.zdigits (dom_op n))).
+ unfold base in *.
+ rewrite ZnZ.spec_zdigits in H2 |- *.
+ set (h := ZnZ.to_Z (ZnZ.head0 x)) in *; clearbody h.
+ set (d := ZnZ.digits (dom_op n)) in *; clearbody d.
+ destruct (Z_lt_le_dec h (Zpos d)); auto. exfalso.
+ assert (1 * 2^Zpos d <= ZnZ.to_Z x * 2^h).
+ apply Zmult_le_compat; auto with zarith.
+ apply Zpower_le_monotone2; auto with zarith.
+ rewrite Zmult_comm in H0. auto with zarith.
+ Qed.
+
+ Lemma spec_log2_pos : forall x, [x]<>0 ->
+ 2^[log2 x] <= [x] < 2^([log2 x]+1).
+ Proof.
+ intros x H. rewrite log2_fold.
+ rewrite spec_eqb. rewrite spec_0.
+ case Z.eqb_spec.
+ auto with zarith.
+ clear H.
+ destr_t x as (n,x). intros H.
+ rewrite ZnZ.spec_sub_carry.
+ assert (H0 := ZnZ.spec_to_Z x).
+ assert (H1 := ZnZ.spec_to_Z (ZnZ.head0 x)).
+ assert (H2 := ZnZ.spec_to_Z (ZnZ.zdigits (dom_op n))).
+ assert (H3 := head0_zdigits n x).
+ rewrite Zmod_small by auto with zarith.
+ rewrite (Z.mul_lt_mono_pos_l (2^(ZnZ.to_Z (ZnZ.head0 x))));
+ auto with zarith.
+ rewrite (Z.mul_le_mono_pos_l _ _ (2^(ZnZ.to_Z (ZnZ.head0 x))));
+ auto with zarith.
+ rewrite <- 2 Zpower_exp; auto with zarith.
+ rewrite Z.add_sub_assoc, Zplus_minus.
+ rewrite Z.sub_simpl_r, Zplus_minus.
+ rewrite ZnZ.spec_zdigits.
+ rewrite pow2_pos_minus_1 by (red; auto).
+ apply ZnZ.spec_head0; auto with zarith.
+ Qed.
+
+ Lemma spec_log2 : forall x, [log2 x] = Z.log2 [x].
+ Proof.
+ intros. destruct (Z_lt_ge_dec 0 [x]).
+ symmetry. apply Z.log2_unique. apply spec_pos.
+ apply spec_log2_pos. intro EQ; rewrite EQ in *; auto with zarith.
+ rewrite spec_log2_0. rewrite Z.log2_nonpos; auto with zarith.
+ generalize (spec_pos x); auto with zarith.
+ Qed.
+
+ Lemma log2_digits_head0 : forall x, 0 < [x] ->
+ [log2 x] = Zpos (digits x) - [head0 x] - 1.
+ Proof.
+ intros. rewrite log2_fold.
+ rewrite spec_eqb. rewrite spec_0.
+ case Z.eqb_spec.
+ auto with zarith.
+ intros _. revert H. rewrite digits_fold, head0_fold. destr_t x as (n,x).
+ rewrite ZnZ.spec_sub_carry.
+ intros.
+ generalize (head0_zdigits n x H).
+ generalize (ZnZ.spec_to_Z (ZnZ.head0 x)).
+ generalize (ZnZ.spec_to_Z (ZnZ.zdigits (dom_op n))).
+ rewrite ZnZ.spec_zdigits. intros. apply Zmod_small.
+ auto with zarith.
+ Qed.
+
+ (** * Right shift *)
+
+ Local Notation shiftrn := (fun n =>
+ let op := dom_op n in
+ let zdigits := ZnZ.zdigits op in
+ let sub_c := ZnZ.sub_c in
+ let add_mul_div := ZnZ.add_mul_div in
+ let zzero := ZnZ.zero in
+ fun x p => match sub_c zdigits p with
+ | C0 d => reduce n (add_mul_div d zzero x)
+ | C1 _ => zero
+ end).
+
+ Definition shiftr : t -> t -> t := Eval red_t in
+ same_level shiftrn.
+
+ Lemma shiftr_fold : shiftr = same_level shiftrn.
+ Proof. red_t; reflexivity. Qed.
+
+ Lemma div_pow2_bound :forall x y z,
+ 0 <= x -> 0 <= y -> x < z -> 0 <= x / 2 ^ y < z.
+ Proof.
+ intros x y z HH HH1 HH2.
+ split; auto with zarith.
+ apply Zle_lt_trans with (2 := HH2); auto with zarith.
+ apply Zdiv_le_upper_bound; auto with zarith.
+ pattern x at 1; replace x with (x * 2 ^ 0); auto with zarith.
+ apply Zmult_le_compat_l; auto.
+ apply Zpower_le_monotone2; auto with zarith.
+ rewrite Zpower_0_r; ring.
+ Qed.
+
+ Theorem spec_shiftr_pow2 : forall x n,
+ [shiftr x n] = [x] / 2 ^ [n].
+ Proof.
+ intros x y. rewrite shiftr_fold. apply spec_same_level. clear x y.
+ intros n x p. simpl.
+ assert (Hx := ZnZ.spec_to_Z x).
+ assert (Hy := ZnZ.spec_to_Z p).
+ generalize (ZnZ.spec_sub_c (ZnZ.zdigits (dom_op n)) p).
+ case ZnZ.sub_c; intros d H; unfold interp_carry in *; simpl.
+ (** Subtraction without underflow : [ p <= digits ] *)
+ rewrite spec_reduce.
+ rewrite ZnZ.spec_zdigits in H.
+ rewrite ZnZ.spec_add_mul_div by auto with zarith.
+ rewrite ZnZ.spec_0, Zmult_0_l, Zplus_0_l.
+ rewrite Zmod_small.
+ f_equal. f_equal. auto with zarith.
+ split. auto with zarith.
+ apply div_pow2_bound; auto with zarith.
+ (** Subtraction with underflow : [ digits < p ] *)
+ rewrite ZnZ.spec_0. symmetry.
+ apply Zdiv_small.
+ split; auto with zarith.
+ apply Zlt_le_trans with (base (ZnZ.digits (dom_op n))); auto with zarith.
+ unfold base. apply Zpower_le_monotone2; auto with zarith.
+ rewrite ZnZ.spec_zdigits in H.
+ generalize (ZnZ.spec_to_Z d); auto with zarith.
+ Qed.
+
+ Lemma spec_shiftr: forall x p, [shiftr x p] = Z.shiftr [x] [p].
+ Proof.
+ intros.
+ now rewrite spec_shiftr_pow2, Z.shiftr_div_pow2 by apply spec_pos.
+ Qed.
+
+ (** * Left shift *)
+
+ (** First an unsafe version, working correctly only if
+ the representation is large enough *)
+
+ Local Notation unsafe_shiftln := (fun n =>
+ let op := dom_op n in
+ let add_mul_div := ZnZ.add_mul_div in
+ let zero := ZnZ.zero in
+ fun x p => reduce n (add_mul_div p x zero)).
+
+ Definition unsafe_shiftl : t -> t -> t := Eval red_t in
+ same_level unsafe_shiftln.
+
+ Lemma unsafe_shiftl_fold : unsafe_shiftl = same_level unsafe_shiftln.
+ Proof. red_t; reflexivity. Qed.
+
+ Theorem spec_unsafe_shiftl_aux : forall x p K,
+ 0 <= K ->
+ [x] < 2^K ->
+ [p] + K <= Zpos (digits x) ->
+ [unsafe_shiftl x p] = [x] * 2 ^ [p].
+ Proof.
+ intros x p.
+ rewrite unsafe_shiftl_fold. rewrite digits_level.
+ apply spec_same_level_dep.
+ intros n m z z' r LE H K HK H1 H2. apply (H K); auto.
+ transitivity (Zpos (ZnZ.digits (dom_op n))); auto.
+ apply digits_dom_op_incr; auto.
+ clear x p.
+ intros n x p K HK Hx Hp. simpl. rewrite spec_reduce.
+ destruct (ZnZ.spec_to_Z x).
+ destruct (ZnZ.spec_to_Z p).
+ rewrite ZnZ.spec_add_mul_div by (omega with *).
+ rewrite ZnZ.spec_0, Zdiv_0_l, Zplus_0_r.
+ apply Zmod_small. unfold base.
+ split; auto with zarith.
+ rewrite Zmult_comm.
+ apply Zlt_le_trans with (2^(ZnZ.to_Z p + K)).
+ rewrite Zpower_exp; auto with zarith.
+ apply Zmult_lt_compat_l; auto with zarith.
+ apply Zpower_le_monotone2; auto with zarith.
+ Qed.
+
+ Theorem spec_unsafe_shiftl: forall x p,
+ [p] <= [head0 x] -> [unsafe_shiftl x p] = [x] * 2 ^ [p].
+ Proof.
+ intros.
+ destruct (Z_eq_dec [x] 0) as [EQ|NEQ].
+ (* [x] = 0 *)
+ apply spec_unsafe_shiftl_aux with 0; auto with zarith.
+ now rewrite EQ.
+ rewrite spec_head00 in *; auto with zarith.
+ (* [x] <> 0 *)
+ apply spec_unsafe_shiftl_aux with ([log2 x] + 1); auto with zarith.
+ generalize (spec_pos (log2 x)); auto with zarith.
+ destruct (spec_log2_pos x); auto with zarith.
+ rewrite log2_digits_head0; auto with zarith.
+ generalize (spec_pos x); auto with zarith.
+ Qed.
+
+ (** Then we define a function doubling the size of the representation
+ but without changing the value of the number. *)
+
+ Local Notation double_size_n := (fun n =>
+ let zero := ZnZ.zero in
+ fun x => mk_t_S n (WW zero x)).
+
+ Definition double_size : t -> t := Eval red_t in
+ iter_t double_size_n.
+
+ Lemma double_size_fold : double_size = iter_t double_size_n.
+ Proof. red_t; reflexivity. Qed.
+
+ Lemma double_size_level : forall x, level (double_size x) = S (level x).
+ Proof.
+ intros x. rewrite double_size_fold; unfold level at 2. destr_t x as (n,x).
+ apply mk_t_S_level.
+ Qed.
+
+ Theorem spec_double_size_digits:
+ forall x, Zpos (digits (double_size x)) = 2 * (Zpos (digits x)).
+ Proof.
+ intros x. rewrite ! digits_level, double_size_level.
+ rewrite 2 digits_dom_op, 2 Pshiftl_nat_Zpower,
+ inj_S, Zpower_Zsucc; auto with zarith.
+ ring.
+ Qed.
+
+ Theorem spec_double_size: forall x, [double_size x] = [x].
+ Proof.
+ intros x. rewrite double_size_fold. destr_t x as (n,x).
+ rewrite spec_mk_t_S. simpl. rewrite ZnZ.spec_0. auto with zarith.
+ Qed.
+
+ Theorem spec_double_size_head0:
+ forall x, 2 * [head0 x] <= [head0 (double_size x)].
+ Proof.
+ intros x.
+ assert (F1:= spec_pos (head0 x)).
+ assert (F2: 0 < Zpos (digits x)).
+ red; auto.
+ case (Zle_lt_or_eq _ _ (spec_pos x)); intros HH.
+ generalize HH; rewrite <- (spec_double_size x); intros HH1.
+ case (spec_head0 x HH); intros _ HH2.
+ case (spec_head0 _ HH1).
+ rewrite (spec_double_size x); rewrite (spec_double_size_digits x).
+ intros HH3 _.
+ case (Zle_or_lt ([head0 (double_size x)]) (2 * [head0 x])); auto; intros HH4.
+ absurd (2 ^ (2 * [head0 x] )* [x] < 2 ^ [head0 (double_size x)] * [x]); auto.
+ apply Zle_not_lt.
+ apply Zmult_le_compat_r; auto with zarith.
+ apply Zpower_le_monotone2; auto; auto with zarith.
+ assert (HH5: 2 ^[head0 x] <= 2 ^(Zpos (digits x) - 1)).
+ case (Zle_lt_or_eq 1 [x]); auto with zarith; intros HH5.
+ apply Zmult_le_reg_r with (2 ^ 1); auto with zarith.
+ rewrite <- (fun x y z => Zpower_exp x (y - z)); auto with zarith.
+ assert (tmp: forall x, x - 1 + 1 = x); [intros; ring | rewrite tmp; clear tmp].
+ apply Zle_trans with (2 := Zlt_le_weak _ _ HH2).
+ apply Zmult_le_compat_l; auto with zarith.
+ rewrite Zpower_1_r; auto with zarith.
+ apply Zpower_le_monotone2; auto with zarith.
+ case (Zle_or_lt (Zpos (digits x)) [head0 x]); auto with zarith; intros HH6.
+ absurd (2 ^ Zpos (digits x) <= 2 ^ [head0 x] * [x]); auto with zarith.
+ rewrite <- HH5; rewrite Zmult_1_r.
+ apply Zpower_le_monotone2; auto with zarith.
+ rewrite (Zmult_comm 2).
+ rewrite Zpower_mult; auto with zarith.
+ rewrite Zpower_2.
+ apply Zlt_le_trans with (2 := HH3).
+ rewrite <- Zmult_assoc.
+ replace (2 * Zpos (digits x) - 1) with
+ ((Zpos (digits x) - 1) + (Zpos (digits x))).
+ rewrite Zpower_exp; auto with zarith.
+ apply Zmult_lt_compat2; auto with zarith.
+ split; auto with zarith.
+ apply Zmult_lt_0_compat; auto with zarith.
+ rewrite Zpos_xO; ring.
+ apply Zlt_le_weak; auto.
+ repeat rewrite spec_head00; auto.
+ rewrite spec_double_size_digits.
+ rewrite Zpos_xO; auto with zarith.
+ rewrite spec_double_size; auto.
+ Qed.
+
+ Theorem spec_double_size_head0_pos:
+ forall x, 0 < [head0 (double_size x)].
+ Proof.
+ intros x.
+ assert (F: 0 < Zpos (digits x)).
+ red; auto.
+ case (Zle_lt_or_eq _ _ (spec_pos (head0 (double_size x)))); auto; intros F0.
+ case (Zle_lt_or_eq _ _ (spec_pos (head0 x))); intros F1.
+ apply Zlt_le_trans with (2 := (spec_double_size_head0 x)); auto with zarith.
+ case (Zle_lt_or_eq _ _ (spec_pos x)); intros F3.
+ generalize F3; rewrite <- (spec_double_size x); intros F4.
+ absurd (2 ^ (Zpos (xO (digits x)) - 1) < 2 ^ (Zpos (digits x))).
+ apply Zle_not_lt.
+ apply Zpower_le_monotone2; auto with zarith.
+ rewrite Zpos_xO; auto with zarith.
+ case (spec_head0 x F3).
+ rewrite <- F1; rewrite Zpower_0_r; rewrite Zmult_1_l; intros _ HH.
+ apply Zle_lt_trans with (2 := HH).
+ case (spec_head0 _ F4).
+ rewrite (spec_double_size x); rewrite (spec_double_size_digits x).
+ rewrite <- F0; rewrite Zpower_0_r; rewrite Zmult_1_l; auto.
+ generalize F1; rewrite (spec_head00 _ (sym_equal F3)); auto with zarith.
+ Qed.
+
+ (** Finally we iterate [double_size] enough before [unsafe_shiftl]
+ in order to get a fully correct [shiftl]. *)
+
+ Definition shiftl_aux_body cont x n :=
+ match compare n (head0 x) with
+ Gt => cont (double_size x) n
+ | _ => unsafe_shiftl x n
end.
- Theorem spec_shiftl_aux_body: forall n p x cont,
+ Theorem spec_shiftl_aux_body: forall n x p cont,
2^ Zpos p <= [head0 x] ->
(forall x, 2 ^ (Zpos p + 1) <= [head0 x]->
- [cont n x] = [x] * 2 ^ [n]) ->
- [shiftl_aux_body cont n x] = [x] * 2 ^ [n].
+ [cont x n] = [x] * 2 ^ [n]) ->
+ [shiftl_aux_body cont x n] = [x] * 2 ^ [n].
Proof.
- intros n p x cont H1 H2; unfold shiftl_aux_body.
- generalize (spec_compare_aux n (head0 x)); case compare; intros H.
+ intros n x p cont H1 H2; unfold shiftl_aux_body.
+ rewrite spec_compare; case Zcompare_spec; intros H.
apply spec_unsafe_shiftl; auto with zarith.
apply spec_unsafe_shiftl; auto with zarith.
rewrite H2.
@@ -435,22 +1531,22 @@ Module Make (Import W0:CyclicType) <: NType.
rewrite Zpower_1_r; apply Zmult_le_compat_l; auto with zarith.
Qed.
- Fixpoint shiftl_aux p cont n x {struct p} :=
+ Fixpoint shiftl_aux p cont x n :=
shiftl_aux_body
- (fun n x => match p with
- | xH => cont n x
- | xO p => shiftl_aux p (shiftl_aux p cont) n x
- | xI p => shiftl_aux p (shiftl_aux p cont) n x
- end) n x.
+ (fun x n => match p with
+ | xH => cont x n
+ | xO p => shiftl_aux p (shiftl_aux p cont) x n
+ | xI p => shiftl_aux p (shiftl_aux p cont) x n
+ end) x n.
- Theorem spec_shiftl_aux: forall p q n x cont,
+ Theorem spec_shiftl_aux: forall p q x n cont,
2 ^ (Zpos q) <= [head0 x] ->
(forall x, 2 ^ (Zpos p + Zpos q) <= [head0 x] ->
- [cont n x] = [x] * 2 ^ [n]) ->
- [shiftl_aux p cont n x] = [x] * 2 ^ [n].
+ [cont x n] = [x] * 2 ^ [n]) ->
+ [shiftl_aux p cont x n] = [x] * 2 ^ [n].
Proof.
intros p; elim p; unfold shiftl_aux; fold shiftl_aux; clear p.
- intros p Hrec q n x cont H1 H2.
+ intros p Hrec q x n cont H1 H2.
apply spec_shiftl_aux_body with (q); auto.
intros x1 H3; apply Hrec with (q + 1)%positive; auto.
intros x2 H4; apply Hrec with (p + q + 1)%positive; auto.
@@ -465,7 +1561,7 @@ Module Make (Import W0:CyclicType) <: NType.
apply spec_shiftl_aux_body with (q); auto.
intros x1 H3; apply Hrec with (q); auto.
apply Zle_trans with (2 := H3); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
+ apply Zpower_le_monotone2; auto with zarith.
intros x2 H4; apply Hrec with (p + q)%positive; auto.
intros x3 H5; apply H2.
rewrite (Zpos_xO p).
@@ -477,20 +1573,20 @@ Module Make (Import W0:CyclicType) <: NType.
rewrite Zplus_comm; auto.
Qed.
- Definition shiftl n x :=
+ Definition shiftl x n :=
shiftl_aux_body
(shiftl_aux_body
- (shiftl_aux (digits n) unsafe_shiftl)) n x.
+ (shiftl_aux (digits n) unsafe_shiftl)) x n.
- Theorem spec_shiftl: forall n x,
- [shiftl n x] = [x] * 2 ^ [n].
+ Theorem spec_shiftl_pow2 : forall x n,
+ [shiftl x n] = [x] * 2 ^ [n].
Proof.
- intros n x; unfold shiftl, shiftl_aux_body.
- generalize (spec_compare_aux n (head0 x)); case compare; intros H.
+ intros x n; unfold shiftl, shiftl_aux_body.
+ rewrite spec_compare; case Zcompare_spec; intros H.
apply spec_unsafe_shiftl; auto with zarith.
apply spec_unsafe_shiftl; auto with zarith.
rewrite <- (spec_double_size x).
- generalize (spec_compare_aux n (head0 (double_size x))); case compare; intros H1.
+ rewrite spec_compare; case Zcompare_spec; intros H1.
apply spec_unsafe_shiftl; auto with zarith.
apply spec_unsafe_shiftl; auto with zarith.
rewrite <- (spec_double_size (double_size x)).
@@ -504,21 +1600,67 @@ Module Make (Import W0:CyclicType) <: NType.
apply Zle_trans with (2 := H2).
apply Zle_trans with (2 ^ Zpos (digits n)); auto with zarith.
case (spec_digits n); auto with zarith.
- apply Zpower_le_monotone; auto with zarith.
+ apply Zpower_le_monotone2; auto with zarith.
Qed.
+ Lemma spec_shiftl: forall x p, [shiftl x p] = Z.shiftl [x] [p].
+ Proof.
+ intros.
+ now rewrite spec_shiftl_pow2, Z.shiftl_mul_pow2 by apply spec_pos.
+ Qed.
- (** * Zero and One *)
+ (** Other bitwise operations *)
- Theorem spec_0: [zero] = 0.
+ Definition testbit x n := odd (shiftr x n).
+
+ Lemma spec_testbit: forall x p, testbit x p = Z.testbit [x] [p].
Proof.
- exact (spec_0 w0_spec).
+ intros. unfold testbit. symmetry.
+ rewrite spec_odd, spec_shiftr. apply Z.testbit_odd.
Qed.
- Theorem spec_1: [one] = 1.
+ Definition div2 x := shiftr x one.
+
+ Lemma spec_div2: forall x, [div2 x] = Z.div2 [x].
Proof.
- exact (spec_1 w0_spec).
+ intros. unfold div2. symmetry.
+ rewrite spec_shiftr, spec_1. apply Z.div2_spec.
Qed.
+ (** TODO : provide efficient versions instead of just converting
+ from/to N (see with Laurent) *)
+
+ Definition lor x y := of_N (N.lor (to_N x) (to_N y)).
+ Definition land x y := of_N (N.land (to_N x) (to_N y)).
+ Definition ldiff x y := of_N (N.ldiff (to_N x) (to_N y)).
+ Definition lxor x y := of_N (N.lxor (to_N x) (to_N y)).
+
+ Lemma spec_land: forall x y, [land x y] = Z.land [x] [y].
+ Proof.
+ intros x y. unfold land. rewrite spec_of_N. unfold to_N.
+ generalize (spec_pos x), (spec_pos y).
+ destruct [x], [y]; trivial; (now destruct 1) || (now destruct 2).
+ Qed.
+
+ Lemma spec_lor: forall x y, [lor x y] = Z.lor [x] [y].
+ Proof.
+ intros x y. unfold lor. rewrite spec_of_N. unfold to_N.
+ generalize (spec_pos x), (spec_pos y).
+ destruct [x], [y]; trivial; (now destruct 1) || (now destruct 2).
+ Qed.
+
+ Lemma spec_ldiff: forall x y, [ldiff x y] = Z.ldiff [x] [y].
+ Proof.
+ intros x y. unfold ldiff. rewrite spec_of_N. unfold to_N.
+ generalize (spec_pos x), (spec_pos y).
+ destruct [x], [y]; trivial; (now destruct 1) || (now destruct 2).
+ Qed.
+
+ Lemma spec_lxor: forall x y, [lxor x y] = Z.lxor [x] [y].
+ Proof.
+ intros x y. unfold lxor. rewrite spec_of_N. unfold to_N.
+ generalize (spec_pos x), (spec_pos y).
+ destruct [x], [y]; trivial; (now destruct 1) || (now destruct 2).
+ Qed.
End Make.
diff --git a/theories/Numbers/Natural/BigN/NMake_gen.ml b/theories/Numbers/Natural/BigN/NMake_gen.ml
index 67a62c40..59d440c3 100644
--- a/theories/Numbers/Natural/BigN/NMake_gen.ml
+++ b/theories/Numbers/Natural/BigN/NMake_gen.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,100 +8,88 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NMake_gen.ml 14641 2011-11-06 11:59:10Z herbelin $ i*)
+(*S NMake_gen.ml : this file generates NMake_gen.v *)
-(*S NMake_gen.ml : this file generates NMake.v *)
-
-(*s The two parameters that control the generation: *)
+(*s The parameter 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) ^ ")"
+let rec iter_str n s = if n = 0 then "" else (iter_str (n-1) s) ^ 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 *)
+let rec iter_str_gen n f = if n < 0 then "" else (iter_str_gen (n-1) f) ^ (f n)
-(* 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
+let rec iter_name i j base sep =
+ if i >= j then base^(string_of_int i)
+ else (iter_name i (j-1) base sep)^sep^" "^base^(string_of_int j)
+let pr s = Printf.printf (s^^"\n")
(*s The actual printing *)
let _ =
- pr "(************************************************************************)";
- pr "(* v * The Coq Proof Assistant / The Coq Development Team *)";
- pr "(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)";
- 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 ZArith CyclicAxioms";
- pr " DoubleType DoubleMul DoubleDivn1 DoubleCyclic Nbasic";
- pr " Wf_nat StreamMemo.";
- pr "";
- pr "Module Make (Import W0:CyclicType).";
- pr "";
+pr
+"(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
+(* \\VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
+(************************************************************************)
- pr " Definition w0 := W0.w.";
- for i = 1 to size do
- pr " Definition w%i := zn2z w%i." i (i-1)
- done;
- pr "";
+(** * NMake_gen *)
- 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 "";
+(** From a cyclic Z/nZ representation to arbitrary precision natural numbers.*)
+
+(** Remark: File automatically generated by NMake_gen.ml, DO NOT EDIT ! *)
+
+Require Import BigNumPrelude ZArith Ndigits CyclicAxioms
+ DoubleType DoubleMul DoubleDivn1 DoubleCyclic Nbasic
+ Wf_nat StreamMemo.
+
+Module Make (W0:CyclicType) <: NAbstract.
+
+ (** * The word types *)
+";
+
+pr " Local Notation w0 := W0.t.";
+for i = 1 to size do
+ pr " Definition w%i := zn2z w%i." i (i-1)
+done;
+pr "";
+
+pr " (** * The operation type classes for the word types *)
+";
+
+pr " Local Notation w0_op := W0.ops.";
+for i = 1 to min 3 size do
+ pr " Instance w%i_op : ZnZ.Ops w%i := mk_zn2z_ops w%i_op." i i (i-1)
+done;
+for i = 4 to size do
+ pr " Instance w%i_op : ZnZ.Ops w%i := mk_zn2z_ops_karatsuba w%i_op." i i (i-1)
+done;
+for i = size+1 to size+3 do
+ pr " Instance w%i_op : ZnZ.Ops (word w%i %i) := mk_zn2z_ops_karatsuba w%i_op." i size (i-size) (i-1)
+done;
+pr "";
pr " Section Make_op.";
- pr " Variable mk : forall w', znz_op w' -> znz_op (zn2z w').";
+ pr " Variable mk : forall w', ZnZ.Ops w' -> ZnZ.Ops (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 " Fixpoint make_op_aux (n:nat) : ZnZ.Ops (word w%i (S n)):=" size;
+ pr " match n return ZnZ.Ops (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 " match n1 return ZnZ.Ops (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 " match n2 return ZnZ.Ops (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";
@@ -110,2565 +98,912 @@ let _ =
pr "";
pr " End Make_op.";
pr "";
- pr " Definition omake_op := make_op_aux mk_zn2z_op_karatsuba.";
+ pr " Definition omake_op := make_op_aux mk_zn2z_ops_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 " Instance make_op n : ZnZ.Ops (word w%i (S n))" size;
+ pr " := dmemo_get _ omake_op n make_op_list.";
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 "";
+pr " Ltac unfold_ops := unfold omake_op, make_op_aux, w%i_op, w%i_op." (size+3) (size+2);
- for i = 0 to size do
- pr " Definition one%i := w%i_op.(znz_1)." i i
- done;
- pr "";
+pr
+"
+ Lemma make_op_omake: forall n, make_op n = omake_op n.
+ Proof.
+ intros n; unfold make_op, make_op_list.
+ refine (dmemo_get_correct _ _ _).
+ Qed.
+ Theorem make_op_S: forall n,
+ make_op (S n) = mk_zn2z_ops_karatsuba (make_op n).
+ Proof.
+ intros n. do 2 rewrite make_op_omake.
+ revert n. fix IHn 1.
+ do 3 (destruct n; [unfold_ops; reflexivity|]).
+ simpl mk_zn2z_ops_karatsuba. simpl word in *.
+ rewrite <- (IHn n). auto.
+ Qed.
- pr " Definition zero := %s0 w_0." c;
- pr " Definition one := %s0 one0." c;
- pr "";
+ (** * The main type [t], isomorphic with [exists n, word w0 n] *)
+";
- pr " Definition to_Z x :=";
- pr " match x with";
+ pr " Inductive t' :=";
for i = 0 to size do
- pr " | %s%i wx => w%i_op.(znz_to_Z) wx" c i i
+ pr " | N%i : w%i -> t'" i i
done;
- pr " | %sn n wx => (make_op n).(znz_to_Z) wx" c;
- pr " end.";
+ pr " | Nn : forall n, word w%i (S n) -> t'." size;
pr "";
-
- pr " Open Scope Z_scope.";
- pr " Notation \"[ x ]\" := (to_Z x).";
- pr "";
-
- pr " Definition to_N x := Zabs_N (to_Z x).";
+ pr " Definition t := t'.";
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 " Bind Scope abstract_scope with t t'.";
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 " (** * A generic toolbox for building and deconstructing [t] *)";
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 " Local Notation SizePlus n := %sn%s."
+ (iter_str size "(S ") (iter_str size ")");
+ pr " Local Notation Size := (SizePlus O).";
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 "";
- 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 "";
- 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 "";
- 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 "";
- 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 "";
-
-
- 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 " Tactic Notation \"do_size\" tactic(t) := do %i t." (size+1);
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";
+ pr " Definition dom_t n := match n 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);
+ pr " | %i => w%i" i i;
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 " | %sn => word w%i n" (if size=0 then "" else "SizePlus ") size;
+ 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 "";
+pr
+" Instance dom_op n : ZnZ.Ops (dom_t n) | 10.
+ Proof.
+ do_size (destruct n; [simpl;auto with *|]).
+ unfold dom_t. auto with *.
+ Defined.
+";
- 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.";
+ pr " Definition iter_t {A:Type}(f : forall n, dom_t n -> A) : t -> A :=";
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;
+ pr " let f%i := f %i in" i i;
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 [";
+ pr " let fn n := f (SizePlus (S n)) in";
+ pr " fun x => match x with";
for i = 0 to size do
- pr0 "extend%i " i;
+ pr " | N%i wx => f%i wx" 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 " | Nn n wx => fn n wx";
pr " end.";
pr "";
- pp " Ltac zg_tac := try";
- pp " (red; simpl Zcompare; auto;";
- pp " let t := fresh \"H\" in (intros t; discriminate t)).";
- pp "";
- pp " Lemma spec_iter: forall x y, P [x] [y] (iter x y).";
- pp " Proof.";
- pp " intros x; case x; clear x; unfold iter.";
- 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";
+ pr " Definition mk_t (n:nat) : dom_t n -> t :=";
+ pr " match n as n' return dom_t n' -> t 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);
+ pr " | %i => N%i" i i;
done;
- pr " | %sn m wy => fnm n m wx wy" c;
- pr " end";
+ pr " | %s(S n) => Nn n" (if size=0 then "" else "SizePlus ");
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
+" Definition level := iter_t (fun n _ => n).
+ Inductive View_t : t -> Prop :=
+ Mk_t : forall n (x : dom_t n), View_t (mk_t n x).
+
+ Lemma destr_t : forall x, View_t x.
+ Proof.
+ intros x. generalize (Mk_t (level x)). destruct x; simpl; auto.
+ Defined.
+
+ Lemma iter_mk_t : forall A (f:forall n, dom_t n -> A),
+ forall n x, iter_t f (mk_t n x) = f n x.
+ Proof.
+ do_size (destruct n; try reflexivity).
+ Qed.
+
+ (** * Projection to ZArith *)
+
+ Definition to_Z : t -> Z :=
+ Eval lazy beta iota delta [iter_t dom_t dom_op] in
+ iter_t (fun _ x => ZnZ.to_Z x).
+
+ Notation \"[ x ]\" := (to_Z x).
+
+ Theorem spec_mk_t : forall n (x:dom_t n), [mk_t n x] = ZnZ.to_Z x.
+ Proof.
+ intros. change to_Z with (iter_t (fun _ x => ZnZ.to_Z x)).
+ rewrite iter_mk_t; auto.
+ Qed.
+
+ (** * Regular make op, without memoization or karatsuba
+
+ This will normally never be used for actual computations,
+ but only for specification purpose when using
+ [word (dom_t n) m] intermediate values. *)
+
+ Fixpoint nmake_op (ww:Type) (ww_op: ZnZ.Ops ww) (n: nat) :
+ ZnZ.Ops (word ww n) :=
+ match n return ZnZ.Ops (word ww n) with
+ O => ww_op
+ | S n1 => mk_zn2z_ops (nmake_op ww ww_op n1)
+ end.
+
+ Let eval n m := ZnZ.to_Z (Ops:=nmake_op _ (dom_op n) m).
+
+ Theorem nmake_op_S: forall ww (w_op: ZnZ.Ops ww) x,
+ nmake_op _ w_op (S x) = mk_zn2z_ops (nmake_op _ w_op x).
+ Proof.
+ auto.
+ Qed.
+
+ Theorem digits_nmake_S :forall n ww (w_op: ZnZ.Ops ww),
+ ZnZ.digits (nmake_op _ w_op (S n)) =
+ xO (ZnZ.digits (nmake_op _ w_op n)).
+ Proof.
+ auto.
+ Qed.
+
+ Theorem digits_nmake : forall n ww (w_op: ZnZ.Ops ww),
+ ZnZ.digits (nmake_op _ w_op n) = Pos.shiftl_nat (ZnZ.digits w_op) n.
+ Proof.
+ induction n. auto.
+ intros ww ww_op. rewrite Pshiftl_nat_S, <- IHn; auto.
+ Qed.
+
+ Theorem nmake_double: forall n ww (w_op: ZnZ.Ops ww),
+ ZnZ.to_Z (Ops:=nmake_op _ w_op n) =
+ @DoubleBase.double_to_Z _ (ZnZ.digits w_op) (ZnZ.to_Z (Ops:=w_op)) n.
+ Proof.
+ intros n; elim n; auto; clear n.
+ intros n Hrec ww ww_op; simpl DoubleBase.double_to_Z; unfold zn2z_to_Z.
+ rewrite <- Hrec; auto.
+ unfold DoubleBase.double_wB; rewrite <- digits_nmake; auto.
+ Qed.
+
+ Theorem nmake_WW: forall ww ww_op n xh xl,
+ (ZnZ.to_Z (Ops:=nmake_op ww ww_op (S n)) (WW xh xl) =
+ ZnZ.to_Z (Ops:=nmake_op ww ww_op n) xh *
+ base (ZnZ.digits (nmake_op ww ww_op n)) +
+ ZnZ.to_Z (Ops:=nmake_op ww ww_op n) xl)%%Z.
+ Proof.
+ auto.
+ Qed.
+
+ (** * The specification proofs for the word operators *)
+";
+
+ if size <> 0 then
+ pr " Typeclasses Opaque %s." (iter_name 1 size "w" "");
+ pr "";
+
+ pr " Instance w0_spec: ZnZ.Specs w0_op := W0.specs.";
+ for i = 1 to min 3 size do
+ pr " Instance w%i_spec: ZnZ.Specs w%i_op := mk_zn2z_specs w%i_spec." i i (i-1)
+ done;
+ for i = 4 to size do
+ pr " Instance w%i_spec: ZnZ.Specs w%i_op := mk_zn2z_specs_karatsuba w%i_spec." i i (i-1)
+ done;
+ pr " Instance w%i_spec: ZnZ.Specs w%i_op := mk_zn2z_specs_karatsuba w%i_spec." (size+1) (size+1) size;
+
+
+pr "
+ Instance wn_spec (n:nat) : ZnZ.Specs (make_op n).
+ Proof.
+ induction n.
+ rewrite make_op_omake; simpl; auto with *.
+ rewrite make_op_S. exact (mk_zn2z_specs_karatsuba IHn).
+ Qed.
+
+ Instance dom_spec n : ZnZ.Specs (dom_op n) | 10.
+ Proof.
+ do_size (destruct n; auto with *). apply wn_spec.
+ Qed.
+
+ Let make_op_WW : forall n x y,
+ (ZnZ.to_Z (Ops:=make_op (S n)) (WW x y) =
+ ZnZ.to_Z (Ops:=make_op n) x * base (ZnZ.digits (make_op n))
+ + ZnZ.to_Z (Ops:=make_op n) y)%%Z.
+ Proof.
+ intros n x y; rewrite make_op_S; auto.
+ Qed.
+
+ (** * Zero *)
+
+ Definition zero0 : w0 := ZnZ.zero.
+
+ Definition zeron n : dom_t n :=
+ match n with
+ | O => zero0
+ | SizePlus (S n) => W0
+ | _ => W0
+ end.
+
+ Lemma spec_zeron : forall n, ZnZ.to_Z (zeron n) = 0%%Z.
+ Proof.
+ do_size (destruct n; [exact ZnZ.spec_0|]).
+ destruct n; auto. simpl. rewrite make_op_S. exact ZnZ.spec_0.
+ Qed.
+
+ (** * Digits *)
+
+ Lemma digits_make_op_0 : forall n,
+ ZnZ.digits (make_op n) = Pos.shiftl_nat (ZnZ.digits (dom_op Size)) (S n).
+ Proof.
+ induction n.
+ auto.
+ replace (ZnZ.digits (make_op (S n))) with (xO (ZnZ.digits (make_op n))).
+ rewrite IHn; auto.
+ rewrite make_op_S; auto.
+ Qed.
+
+ Lemma digits_make_op : forall n,
+ ZnZ.digits (make_op n) = Pos.shiftl_nat (ZnZ.digits w0_op) (SizePlus (S n)).
+ Proof.
+ intros. rewrite digits_make_op_0.
+ replace (SizePlus (S n)) with (S n + Size) by (rewrite <- plus_comm; auto).
+ rewrite Pshiftl_nat_plus. auto.
+ Qed.
+
+ Lemma digits_dom_op : forall n,
+ ZnZ.digits (dom_op n) = Pos.shiftl_nat (ZnZ.digits w0_op) n.
+ Proof.
+ do_size (destruct n; try reflexivity).
+ exact (digits_make_op n).
+ Qed.
+
+ Lemma digits_dom_op_nmake : forall n m,
+ ZnZ.digits (dom_op (m+n)) = ZnZ.digits (nmake_op _ (dom_op n) m).
+ Proof.
+ intros. rewrite digits_nmake, 2 digits_dom_op. apply Pshiftl_nat_plus.
+ Qed.
+
+ (** * Conversion between [zn2z (dom_t n)] and [dom_t (S n)].
+
+ These two types are provably equal, but not convertible,
+ hence we need some work. We now avoid using generic casts
+ (i.e. rewrite via proof of equalities in types), since
+ proving things with them is a mess.
+ *)
+
+ Definition succ_t n : zn2z (dom_t n) -> dom_t (S n) :=
+ match n with
+ | SizePlus (S _) => fun x => x
+ | _ => fun x => x
+ end.
+
+ Lemma spec_succ_t : forall n x,
+ ZnZ.to_Z (succ_t n x) =
+ zn2z_to_Z (base (ZnZ.digits (dom_op n))) ZnZ.to_Z x.
+ Proof.
+ do_size (destruct n ; [reflexivity|]).
+ intros. simpl. rewrite make_op_S. simpl. auto.
+ Qed.
+
+ Definition pred_t n : dom_t (S n) -> zn2z (dom_t n) :=
+ match n with
+ | SizePlus (S _) => fun x => x
+ | _ => fun x => x
+ end.
+
+ Lemma succ_pred_t : forall n x, succ_t n (pred_t n x) = x.
+ Proof.
+ do_size (destruct n ; [reflexivity|]). reflexivity.
+ Qed.
+
+ (** We can hence project from [zn2z (dom_t n)] to [t] : *)
+
+ Definition mk_t_S n (x : zn2z (dom_t n)) : t :=
+ mk_t (S n) (succ_t n x).
+
+ Lemma spec_mk_t_S : forall n x,
+ [mk_t_S n x] = zn2z_to_Z (base (ZnZ.digits (dom_op n))) ZnZ.to_Z x.
+ Proof.
+ intros. unfold mk_t_S. rewrite spec_mk_t. apply spec_succ_t.
+ Qed.
+
+ Lemma mk_t_S_level : forall n x, level (mk_t_S n x) = S n.
+ Proof.
+ intros. unfold mk_t_S, level. rewrite iter_mk_t; auto.
+ Qed.
+
+ (** * Conversion from [word (dom_t n) m] to [dom_t (m+n)].
+
+ Things are more complex here. We start with a naive version
+ that breaks zn2z-trees and reconstruct them. Doing this is
+ quite unfortunate, but I don't know how to fully avoid that.
+ (cast someday ?). Then we build an optimized version where
+ all basic cases (n<=6 or m<=7) are nicely handled.
+ *)
+
+ Definition zn2z_map {A} {B} (f:A->B) (x:zn2z A) : zn2z B :=
+ match x with
+ | W0 => W0
+ | WW h l => WW (f h) (f l)
+ end.
+
+ Lemma zn2z_map_id : forall A f (x:zn2z A), (forall u, f u = u) ->
+ zn2z_map f x = x.
+ Proof.
+ destruct x; auto; intros.
+ simpl; f_equal; auto.
+ Qed.
+
+ (** The naive version *)
+
+ Fixpoint plus_t n m : word (dom_t n) m -> dom_t (m+n) :=
+ match m as m' return word (dom_t n) m' -> dom_t (m'+n) with
+ | O => fun x => x
+ | S m => fun x => succ_t _ (zn2z_map (plus_t n m) x)
+ end.
+
+ Theorem spec_plus_t : forall n m (x:word (dom_t n) m),
+ ZnZ.to_Z (plus_t n m x) = eval n m x.
+ Proof.
+ unfold eval.
+ induction m.
+ simpl; auto.
+ intros.
+ simpl plus_t; simpl plus. rewrite spec_succ_t.
+ destruct x.
+ simpl; auto.
+ fold word in w, w0.
+ simpl. rewrite 2 IHm. f_equal. f_equal. f_equal.
+ apply digits_dom_op_nmake.
+ Qed.
+
+ Definition mk_t_w n m (x:word (dom_t n) m) : t :=
+ mk_t (m+n) (plus_t n m x).
+
+ Theorem spec_mk_t_w : forall n m (x:word (dom_t n) m),
+ [mk_t_w n m x] = eval n m x.
+ Proof.
+ intros. unfold mk_t_w. rewrite spec_mk_t. apply spec_plus_t.
+ Qed.
+
+ (** The optimized version.
+
+ NB: the last particular case for m could depend on n,
+ but it's simplier to just expand everywhere up to m=7
+ (cf [mk_t_w'] later).
+ *)
+
+ Definition plus_t' n : forall m, word (dom_t n) m -> dom_t (m+n) :=
+ match n return (forall m, word (dom_t n) m -> dom_t (m+n)) with
+ | SizePlus (S n') as n => plus_t n
+ | _ as n =>
+ fun m => match m return (word (dom_t n) m -> dom_t (m+n)) with
+ | SizePlus (S (S m')) as m => plus_t n m
+ | _ => fun x => x
+ end
+ end.
+
+ Lemma plus_t_equiv : forall n m x,
+ plus_t' n m x = plus_t n m x.
+ Proof.
+ (do_size try destruct n); try reflexivity;
+ (do_size try destruct m); try destruct m; try reflexivity;
+ simpl; symmetry; repeat (intros; apply zn2z_map_id; trivial).
+ Qed.
+
+ Lemma spec_plus_t' : forall n m x,
+ ZnZ.to_Z (plus_t' n m x) = eval n m x.
+ Proof.
+ intros; rewrite plus_t_equiv. apply spec_plus_t.
+ Qed.
+
+ (** Particular cases [Nk x] = eval i j x with specific k,i,j
+ can be solved by the following tactic *)
+
+ Ltac solve_eval :=
+ intros; rewrite <- spec_plus_t'; unfold to_Z; simpl dom_op; reflexivity.
+
+ (** The last particular case that remains useful *)
+
+ Lemma spec_eval_size : forall n x, [Nn n x] = eval Size (S n) x.
+ Proof.
+ induction n.
+ solve_eval.
+ destruct x as [ | xh xl ].
+ simpl. unfold eval. rewrite make_op_S. rewrite nmake_op_S. auto.
+ simpl word in xh, xl |- *.
+ unfold to_Z in *. rewrite make_op_WW.
+ unfold eval in *. rewrite nmake_WW.
+ f_equal; auto.
+ f_equal; auto.
+ f_equal.
+ rewrite <- digits_dom_op_nmake. rewrite plus_comm; auto.
+ Qed.
+
+ (** An optimized [mk_t_w].
+
+ We could say mk_t_w' := mk_t _ (plus_t' n m x)
+ (TODO: WHY NOT, BTW ??).
+ Instead we directly define functions for all intersting [n],
+ reverting to naive [mk_t_w] at places that should normally
+ never be used (see [mul] and [div_gt]).
+ *)
+";
+
+for i = 0 to size-1 do
+let pattern = (iter_str (size+1-i) "(S ") ^ "_" ^ (iter_str (size+1-i) ")") in
+pr
+" Let mk_t_%iw m := Eval cbv beta zeta iota delta [ mk_t plus ] in
+ match m return word w%i (S m) -> t with
+ | %s as p => mk_t_w %i (S p)
+ | p => mk_t (%i+p)
+ end.
+" i i pattern i (i+1)
+done;
+
+pr
+" Let mk_t_w' n : forall m, word (dom_t n) (S m) -> t :=
+ match n return (forall m, word (dom_t n) (S m) -> t) with";
+for i = 0 to size-1 do pr " | %i => mk_t_%iw" i i done;
+pr
+" | Size => Nn
+ | _ as n' => fun m => mk_t_w n' (S m)
+ end.
+";
+
+pr
+" Ltac solve_spec_mk_t_w' :=
+ rewrite <- spec_plus_t';
+ match goal with _ : word (dom_t ?n) ?m |- _ => apply (spec_mk_t (n+m)) end.
+
+ Theorem spec_mk_t_w' :
+ forall n m x, [mk_t_w' n m x] = eval n (S m) x.
+ Proof.
+ intros.
+ repeat (apply spec_mk_t_w || (destruct n;
+ [repeat (apply spec_mk_t_w || (destruct m; [solve_spec_mk_t_w'|]))|])).
+ apply spec_eval_size.
+ Qed.
+
+ (** * Extend : injecting [dom_t n] into [word (dom_t n) (S m)] *)
+
+ Definition extend n m (x:dom_t n) : word (dom_t n) (S m) :=
+ DoubleBase.extend_aux m (WW (zeron n) x).
+
+ Lemma spec_extend : forall n m x,
+ [mk_t n x] = eval n (S m) (extend n m x).
+ Proof.
+ intros. unfold eval, extend.
+ rewrite spec_mk_t.
+ assert (H : forall (x:dom_t n),
+ (ZnZ.to_Z (zeron n) * base (ZnZ.digits (dom_op n)) + ZnZ.to_Z x =
+ ZnZ.to_Z x)%%Z).
+ clear; intros; rewrite spec_zeron; auto.
+ rewrite <- (@DoubleBase.spec_extend _
+ (WW (zeron n)) (ZnZ.digits (dom_op n)) ZnZ.to_Z H m x).
+ simpl. rewrite digits_nmake, <- nmake_double. auto.
+ Qed.
+
+ (** A particular case of extend, used in [same_level]:
+ [extend_size] is [extend Size] *)
+
+ Definition extend_size := DoubleBase.extend (WW (W0:dom_t Size)).
+
+ Lemma spec_extend_size : forall n x, [mk_t Size x] = [Nn n (extend_size n x)].
+ Proof.
+ intros. rewrite spec_eval_size. apply (spec_extend Size n).
+ Qed.
+
+ (** Misc results about extensions *)
+
+ Let spec_extend_WW : forall n x,
+ [Nn (S n) (WW W0 x)] = [Nn n x].
+ Proof.
+ intros n x.
+ set (N:=SizePlus (S n)).
+ change ([Nn (S n) (extend N 0 x)]=[mk_t N x]).
+ rewrite (spec_extend N 0).
+ solve_eval.
+ Qed.
+
+ Let spec_extend_tr: forall m n w,
+ [Nn (m + n) (extend_tr w m)] = [Nn n w].
+ Proof.
+ induction m; auto.
+ intros n x; simpl extend_tr.
+ simpl plus; rewrite spec_extend_WW; auto.
+ Qed.
+
+ Let spec_cast_l: forall n m x1,
+ [Nn n x1] =
+ [Nn (Max.max n m) (castm (diff_r n m) (extend_tr x1 (snd (diff n m))))].
+ Proof.
+ intros n m x1; case (diff_r n m); simpl castm.
+ rewrite spec_extend_tr; auto.
+ Qed.
+
+ Let spec_cast_r: forall n m x1,
+ [Nn m x1] =
+ [Nn (Max.max n m) (castm (diff_l n m) (extend_tr x1 (fst (diff n m))))].
+ Proof.
+ intros n m x1; case (diff_l n m); simpl castm.
+ rewrite spec_extend_tr; auto.
+ Qed.
+
+ Ltac unfold_lets :=
+ match goal with
+ | h : _ |- _ => unfold h; clear h; unfold_lets
+ | _ => idtac
+ end.
+
+ (** * [same_level]
+
+ Generic binary operator construction, by extending the smaller
+ argument to the level of the other.
+ *)
+
+ Section SameLevel.
+
+ Variable res: Type.
+ Variable P : Z -> Z -> res -> Prop.
+ Variable f : forall n, dom_t n -> dom_t n -> res.
+ Variable Pf : forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y).
+";
+
+for i = 0 to size do
+pr " Let f%i : w%i -> w%i -> res := f %i." i i i i
+done;
+pr
+" Let fn n := f (SizePlus (S n)).
+
+ Let Pf' :
+ forall n x y u v, u = [mk_t n x] -> v = [mk_t n y] -> P u v (f n x y).
+ Proof.
+ intros. subst. rewrite 2 spec_mk_t. apply Pf.
+ Qed.
+";
+
+let ext i j s =
+ if j <= i then s else Printf.sprintf "(extend %i %i %s)" i (j-i-1) s
+in
+
+pr " Notation same_level_folded := (fun x y => match x, y with";
+for i = 0 to size do
+ for j = 0 to size do
+ pr " | N%i wx, N%i wy => f%i %s %s" i j (max i j) (ext i j "wx") (ext j i "wy")
+ done;
+ pr " | N%i wx, Nn m wy => fn m (extend_size m %s) wy" i (ext i size "wx")
+done;
+for i = 0 to size do
+ pr " | Nn n wx, N%i wy => fn n wx (extend_size n %s)" i (ext i size "wy")
+done;
+pr
+" | Nn n wx, Nn m wy =>
+ let mn := Max.max n m in
+ let d := diff n m in
+ fn mn
+ (castm (diff_r n m) (extend_tr wx (snd d)))
+ (castm (diff_l n m) (extend_tr wy (fst d)))
+ end).
+";
+
+pr
+" Definition same_level := Eval lazy beta iota delta
+ [ DoubleBase.extend DoubleBase.extend_aux extend zeron ]
+ in same_level_folded.
+
+ Lemma spec_same_level_0: forall x y, P [x] [y] (same_level x y).
+ Proof.
+ change same_level with same_level_folded. unfold_lets.
+ destruct x, y; apply Pf'; simpl mk_t; rewrite <- ?spec_extend_size;
+ match goal with
+ | |- context [ extend ?n ?m _ ] => apply (spec_extend n m)
+ | |- context [ castm _ _ ] => apply spec_cast_l || apply spec_cast_r
+ | _ => reflexivity
+ end.
+ Qed.
+
+ End SameLevel.
+
+ Arguments same_level [res] f x y.
+
+ Theorem spec_same_level_dep :
+ forall res
+ (P : nat -> Z -> Z -> res -> Prop)
+ (Pantimon : forall n m z z' r, n <= m -> P m z z' r -> P n z z' r)
+ (f : forall n, dom_t n -> dom_t n -> res)
+ (Pf: forall n x y, P n (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y)),
+ forall x y, P (level x) [x] [y] (same_level f x y).
+ Proof.
+ intros res P Pantimon f Pf.
+ set (f' := fun n x y => (n, f n x y)).
+ set (P' := fun z z' r => P (fst r) z z' (snd r)).
+ assert (FST : forall x y, level x <= fst (same_level f' x y))
+ by (destruct x, y; simpl; omega with * ).
+ assert (SND : forall x y, same_level f x y = snd (same_level f' x y))
+ by (destruct x, y; reflexivity).
+ intros. eapply Pantimon; [eapply FST|].
+ rewrite SND. eapply (@spec_same_level_0 _ P' f'); eauto.
+ Qed.
+
+ (** * [iter]
+
+ Generic binary operator construction, by splitting the larger
+ argument in blocks and applying the smaller argument to them.
+ *)
+
+ Section Iter.
+
+ Variable res: Type.
+ Variable P: Z -> Z -> res -> Prop.
+
+ Variable f : forall n, dom_t n -> dom_t n -> res.
+ Variable Pf : forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y).
+
+ Variable fd : forall n m, dom_t n -> word (dom_t n) (S m) -> res.
+ Variable fg : forall n m, word (dom_t n) (S m) -> dom_t n -> res.
+ Variable Pfd : forall n m x y, P (ZnZ.to_Z x) (eval n (S m) y) (fd n m x y).
+ Variable Pfg : forall n m x y, P (eval n (S m) x) (ZnZ.to_Z y) (fg n m x y).
+
+ Variable fnm: forall n m, word (dom_t Size) (S n) -> word (dom_t Size) (S m) -> res.
+ Variable Pfnm: forall n m x y, P [Nn n x] [Nn m y] (fnm n m x y).
+
+ Let Pf' :
+ forall n x y u v, u = [mk_t n x] -> v = [mk_t n y] -> P u v (f n x y).
+ Proof.
+ intros. subst. rewrite 2 spec_mk_t. apply Pf.
+ Qed.
+
+ Let Pfd' : forall n m x y u v, u = [mk_t n x] -> v = eval n (S m) y ->
+ P u v (fd n m x y).
+ Proof.
+ intros. subst. rewrite spec_mk_t. apply Pfd.
+ Qed.
+
+ Let Pfg' : forall n m x y u v, u = eval n (S m) x -> v = [mk_t n y] ->
+ P u v (fg n m x y).
+ Proof.
+ intros. subst. rewrite spec_mk_t. apply Pfg.
+ Qed.
+";
+
+for i = 0 to size do
+pr " Let f%i := f %i." i i
+done;
+
+for i = 0 to size do
+pr " Let f%in := fd %i." i i;
+pr " Let fn%i := fg %i." i i;
+done;
+
+pr " Notation iter_folded := (fun x y => match x, y with";
+for i = 0 to size do
+ for j = 0 to size do
+ pr " | N%i wx, N%i wy => f%s wx wy" i j
+ (if i = j then string_of_int i
+ else if i < j then string_of_int i ^ "n " ^ string_of_int (j-i-1)
+ else "n" ^ string_of_int j ^ " " ^ string_of_int (i-j-1))
+ done;
+ pr " | N%i wx, Nn m wy => f%in m %s wy" i size (ext i size "wx")
+done;
+for i = 0 to size do
+ pr " | Nn n wx, N%i wy => fn%i n wx %s" i size (ext i size "wy")
+done;
+pr
+" | Nn n wx, Nn m wy => fnm n m wx wy
+ end).
+";
+
+pr
+" Definition iter := Eval lazy beta iota delta
+ [extend DoubleBase.extend DoubleBase.extend_aux zeron]
+ in iter_folded.
+
+ Lemma spec_iter: forall x y, P [x] [y] (iter x y).
+ Proof.
+ change iter with iter_folded; unfold_lets.
+ destruct x; destruct y; apply Pf' || apply Pfd' || apply Pfg' || apply Pfnm;
+ simpl mk_t;
+ match goal with
+ | |- ?x = ?x => reflexivity
+ | |- [Nn _ _] = _ => apply spec_eval_size
+ | |- context [extend ?n ?m _] => apply (spec_extend n m)
+ | _ => idtac
+ end;
+ unfold to_Z; rewrite <- spec_plus_t'; simpl dom_op; reflexivity.
+ Qed.
+
+ End Iter.
+";
+
+pr
+" Definition switch
+ (P:nat->Type)%s
+ (fn:forall n, P n) n :=
+ match n return P n with"
+ (iter_str_gen size (fun i -> Printf.sprintf "(f%i:P %i)" i i));
+for i = 0 to size do pr " | %i => f%i" i i done;
+pr
+" | n => fn n
+ end.
+";
+
+pr
+" Lemma spec_switch : forall P (f:forall n, P n) n,
+ switch P %sf n = f n.
+ Proof.
+ repeat (destruct n; try reflexivity).
+ Qed.
+" (iter_str_gen size (fun i -> Printf.sprintf "(f %i) " i));
+
+pr
+" (** * [iter_sym]
+
+ A variant of [iter] for symmetric functions, or pseudo-symmetric
+ functions (when f y x can be deduced from f x y).
+ *)
+
+ Section IterSym.
+
+ Variable res: Type.
+ Variable P: Z -> Z -> res -> Prop.
+
+ Variable f : forall n, dom_t n -> dom_t n -> res.
+ Variable Pf : forall n x y, P (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y).
+
+ Variable fg : forall n m, word (dom_t n) (S m) -> dom_t n -> res.
+ Variable Pfg : forall n m x y, P (eval n (S m) x) (ZnZ.to_Z y) (fg n m x y).
+
+ Variable fnm: forall n m, word (dom_t Size) (S n) -> word (dom_t Size) (S m) -> res.
+ Variable Pfnm: forall n m x y, P [Nn n x] [Nn m y] (fnm n m x y).
+
+ Variable opp: res -> res.
+ Variable Popp : forall u v r, P u v r -> P v u (opp r).
+";
+
+for i = 0 to size do
+pr " Let f%i := f %i." i i
+done;
+
+for i = 0 to size do
+pr " Let fn%i := fg %i." i i;
+done;
+
+pr " Let f' := switch _ %s f." (iter_name 0 size "f" "");
+pr " Let fg' := switch _ %s fg." (iter_name 0 size "fn" "");
+
+pr
+" Local Notation iter_sym_folded :=
+ (iter res f' (fun n m x y => opp (fg' n m y x)) fg' fnm).
+
+ Definition iter_sym :=
+ Eval lazy beta zeta iota delta [iter f' fg' switch] in iter_sym_folded.
+
+ Lemma spec_iter_sym: forall x y, P [x] [y] (iter_sym x y).
+ Proof.
+ intros. change iter_sym with iter_sym_folded. apply spec_iter; clear x y.
+ unfold_lets.
+ intros. rewrite spec_switch. auto.
+ intros. apply Popp. unfold_lets. rewrite spec_switch; auto.
+ intros. unfold_lets. rewrite spec_switch; auto.
+ auto.
+ Qed.
+
+ End IterSym.
+
+ (** * Reduction
+
+ [reduce] can be used instead of [mk_t], it will choose the
+ lowest possible level. NB: We only search and remove leftmost
+ W0's via ZnZ.eq0, any non-W0 block ends the process, even
+ if its value is 0.
+ *)
+
+ (** First, a direct version ... *)
+
+ Fixpoint red_t n : dom_t n -> t :=
+ match n return dom_t n -> t with
+ | O => N0
+ | S n => fun x =>
+ let x' := pred_t n x in
+ reduce_n1 _ _ (N0 zero0) ZnZ.eq0 (red_t n) (mk_t_S n) x'
+ end.
+
+ Lemma spec_red_t : forall n x, [red_t n x] = [mk_t n x].
+ Proof.
+ induction n.
+ reflexivity.
+ intros.
+ simpl red_t. unfold reduce_n1.
+ rewrite <- (succ_pred_t n x) at 2.
+ remember (pred_t n x) as x'.
+ rewrite spec_mk_t, spec_succ_t.
+ destruct x' as [ | xh xl]. simpl. apply ZnZ.spec_0.
+ generalize (ZnZ.spec_eq0 xh); case ZnZ.eq0; intros H.
+ rewrite IHn, spec_mk_t. simpl. rewrite H; auto.
+ apply spec_mk_t_S.
+ Qed.
+
+ (** ... then a specialized one *)
+";
+
+for i = 0 to size do
+pr " Definition eq0%i := @ZnZ.eq0 _ w%i_op." i i;
+done;
+
+pr "
+ Definition reduce_0 := N0.";
+for i = 1 to size do
+ pr " Definition reduce_%i :=" i;
+ pr " Eval lazy beta iota delta [reduce_n1] in";
+ pr " reduce_n1 _ _ (N0 zero0) eq0%i reduce_%i N%i." (i-1) (i-1) i
+done;
- 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 " Eval lazy beta iota delta [reduce_n1] in";
+ pr " reduce_n1 _ _ (N0 zero0) eq0%i reduce_%i (Nn 0)." size size;
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 "";
- 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.";
-
- 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_pos : forall x, 0 < [x] -> [pred x] = [x] - 1.";
- pa " Admitted.";
- pp " Proof.";
- pp " intros x; case x; unfold pred.";
- for i = 0 to size do
- pp " intros x1 H1; unfold w%i_pred_c;" i;
- pp " 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_pos : 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 => CompOpp (comparen_%i (S n) y x))" i;
- pr " (fun n => comparen_%i (S n))" i;
- done;
- pr " comparenm).";
- 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 CompOpp c with Eq => v = u | Lt => v < u | Gt => v > u end.";
- pp " Proof.";
- pp " intros c u v; case c; unfold CompOpp; auto with zarith.";
- pp " Qed.";
- pp "";
-
-
- pr " Theorem spec_compare_aux: forall x y,";
- pr " match compare x y with";
- pr " Eq => [x] = [y]";
- pr " | Lt => [x] < [y]";
- pr " | Gt => [x] > [y]";
- pr " end.";
- pa " Admitted.";
- pp " Proof.";
- pp " refine (spec_iter _ (fun x y res =>";
- pp " match res with";
- pp " 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 => CompOpp (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 " (***************************************************************)";
- 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 "";
-
- 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 " (** * 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 " (***************************************************************)";
- 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 " (** digits: a measure for gcd *)";
- 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 " (***************************************************************)";
- 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 " (***************************************************************)";
- 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 unsafe_shiftr%i n x := w%i_op.(znz_add_mul_div) (w%i_op.(znz_sub) w%i_op.(znz_zdigits) n) w%i_op.(znz_0) x." i i i i i;
- done;
- pr " Definition unsafe_shiftrn n p x := (make_op n).(znz_add_mul_div) ((make_op n).(znz_sub) (make_op n).(znz_zdigits) p) (make_op n).(znz_0) x.";
- pr "";
-
- pr " Definition unsafe_shiftr := Eval lazy beta delta [same_level] in";
- pr " same_level _ (fun n x => %s0 (unsafe_shiftr0 n x))" c;
- for i = 1 to size do
- pr " (fun n x => reduce_%i (unsafe_shiftr%i n x))" i i;
- done;
- pr " (fun n p x => reduce_n n (unsafe_shiftrn n p x)).";
- pr "";
-
-
- pr " Theorem spec_unsafe_shiftr: forall n x,";
- pr " [n] <= [Ndigits x] -> [unsafe_shiftr n x] = [x] / 2 ^ [n].";
- pa " Admitted.";
- pp " Proof.";
- pp " assert (F0: forall x y, x - (x - y) = y).";
- 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 unsafe_shiftr, same_level.";
- for i = 0 to size do
- pp " intros x y; case y; clear y.";
- for j = 0 to i - 1 do
- pp " intros y; unfold unsafe_shiftr%i, Ndigits." i;
- pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
- pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i j i;
- pp " rewrite (spec_zdigits w%i_spec)." i;
- 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 unsafe_shiftr%i, Ndigits." i;
- pp " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
- pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i i i;
- for j = i + 1 to size do
- pp " intros y; unfold unsafe_shiftr%i, Ndigits." j;
- pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
- pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." j j i;
- pp " try (apply sym_equal; exact (spec_extend%in%i x))." i j;
- done;
- if i == size then
- begin
- pp " intros m y; unfold unsafe_shiftrn, Ndigits.";
- pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
- pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." size;
- pp " try (apply sym_equal; exact (spec_extend%in m x))." size;
- end
- else
- begin
- pp " intros m y; unfold unsafe_shiftrn, Ndigits.";
- pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
- pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." i;
- pp " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x])." size i (size - i - 1) i;
- 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 unsafe_shiftrn, Ndigits; try rewrite spec_reduce_n.";
- for i = 0 to size do
- pp " try rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
- pp " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith." i;
- 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 "";
-
- (* Unsafe_Shiftl *)
- for i = 0 to size do
- pr " Definition unsafe_shiftl%i n x := w%i_op.(znz_add_mul_div) n x w%i_op.(znz_0)." i i i
- done;
- pr " Definition unsafe_shiftln n p x := (make_op n).(znz_add_mul_div) p x (make_op n).(znz_0).";
- pr " Definition unsafe_shiftl := Eval lazy beta delta [same_level] in";
- pr " same_level _ (fun n x => %s0 (unsafe_shiftl0 n x))" c;
- for i = 1 to size do
- pr " (fun n x => reduce_%i (unsafe_shiftl%i n x))" i i;
- done;
- pr " (fun n p x => reduce_n n (unsafe_shiftln n p x)).";
- pr "";
- pr "";
-
-
- pr " Theorem spec_unsafe_shiftl: forall n x,";
- pr " [n] <= [head0 x] -> [unsafe_shiftl n x] = [x] * 2 ^ [n].";
- pa " Admitted.";
- pp " Proof.";
- pp " assert (F0: forall x y, x - (x - y) = y).";
- 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 unsafe_shiftl, same_level.";
- for i = 0 to size do
- pp " intros x y; case y; clear y.";
- for j = 0 to i - 1 do
- pp " intros y; unfold unsafe_shiftl%i, head0." i;
- pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
- pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i j i;
- pp " rewrite (spec_zdigits w%i_spec)." i;
- 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 unsafe_shiftl%i, head0." i;
- pp " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
- pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." i i i;
- for j = i + 1 to size do
- pp " intros y; unfold unsafe_shiftl%i, head0." j;
- pp " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1." i j;
- pp " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith." j j i;
- pp " try (apply sym_equal; exact (spec_extend%in%i x))." i j;
- done;
- if i == size then
- begin
- pp " intros m y; unfold unsafe_shiftln, head0.";
- pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
- pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." size;
- pp " try (apply sym_equal; exact (spec_extend%in m x))." size;
- end
- else
- begin
- pp " intros m y; unfold unsafe_shiftln, head0.";
- pp " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.";
- pp " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith." i;
- pp " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x])." size i (size - i - 1) i;
- 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 unsafe_shiftln, head0; try rewrite spec_reduce_n.";
- for i = 0 to size do
- pp " try rewrite spec_reduce_%i; unfold to_Z; intros H1." i;
- pp " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith." i;
- 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 "";
-
- (* 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 "End Make.";
- pr "";
-
+ pr " Eval lazy beta iota delta [reduce_n] in";
+ pr " reduce_n _ _ (N0 zero0) reduce_%i Nn n." (size + 1);
+ pr "";
+
+pr " Definition reduce n : dom_t n -> t :=";
+pr " match n with";
+for i = 0 to size do
+pr " | %i => reduce_%i" i i;
+done;
+pr " | %s(S n) => reduce_n n" (if size=0 then "" else "SizePlus ");
+pr " end.";
+pr "";
+
+pr " Ltac unfold_red := unfold reduce, %s." (iter_name 1 size "reduce_" ",");
+
+pr "
+ Ltac solve_red :=
+ let H := fresh in let G := fresh in
+ match goal with
+ | |- ?P (S ?n) => assert (H:P n) by solve_red
+ | _ => idtac
+ end;
+ intros n G x; destruct (le_lt_eq_dec _ _ G) as [LT|EQ];
+ solve [
+ apply (H _ (lt_n_Sm_le _ _ LT)) |
+ inversion LT |
+ subst; change (reduce 0 x = red_t 0 x); reflexivity |
+ specialize (H (pred n)); subst; destruct x;
+ [|unfold_red; rewrite H; auto]; reflexivity
+ ].
+
+ Lemma reduce_equiv : forall n x, n <= Size -> reduce n x = red_t n x.
+ Proof.
+ set (P N := forall n, n <= N -> forall x, reduce n x = red_t n x).
+ intros n x H. revert n H x. change (P Size). solve_red.
+ Qed.
+
+ Lemma spec_reduce_n : forall n x, [reduce_n n x] = [Nn n x].
+ Proof.
+ assert (H : forall x, reduce_%i x = red_t (SizePlus 1) x).
+ destruct x; [|unfold reduce_%i; rewrite (reduce_equiv Size)]; auto.
+ induction n.
+ intros. rewrite H. apply spec_red_t.
+ destruct x as [|xh xl].
+ simpl. rewrite make_op_S. exact ZnZ.spec_0.
+ fold word in *.
+ destruct xh; auto.
+ simpl reduce_n.
+ rewrite IHn.
+ rewrite spec_extend_WW; auto.
+ Qed.
+" (size+1) (size+1);
+
+pr
+" Lemma spec_reduce : forall n x, [reduce n x] = ZnZ.to_Z x.
+ Proof.
+ do_size (destruct n;
+ [intros; rewrite reduce_equiv;[apply spec_red_t|auto with arith]|]).
+ apply spec_reduce_n.
+ Qed.
+
+End Make.
+";
diff --git a/theories/Numbers/Natural/BigN/Nbasic.v b/theories/Numbers/Natural/BigN/Nbasic.v
index cdd41647..4717d0b2 100644
--- a/theories/Numbers/Natural/BigN/Nbasic.v
+++ b/theories/Numbers/Natural/BigN/Nbasic.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,9 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: Nbasic.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-Require Import ZArith.
+Require Import ZArith Ndigits.
Require Import BigNumPrelude.
Require Import Max.
Require Import DoubleType.
@@ -18,6 +16,26 @@ Require Import DoubleBase.
Require Import CyclicAxioms.
Require Import DoubleCyclic.
+Arguments mk_zn2z_ops [t] ops.
+Arguments mk_zn2z_ops_karatsuba [t] ops.
+Arguments mk_zn2z_specs [t ops] specs.
+Arguments mk_zn2z_specs_karatsuba [t ops] specs.
+Arguments ZnZ.digits [t] Ops.
+Arguments ZnZ.zdigits [t] Ops.
+
+Lemma Pshiftl_nat_Zpower : forall n p,
+ Zpos (Pos.shiftl_nat p n) = Zpos p * 2 ^ Z.of_nat n.
+Proof.
+ intros.
+ rewrite Z.mul_comm.
+ induction n. simpl; auto.
+ transitivity (2 * (2 ^ Z.of_nat n * Zpos p)).
+ rewrite <- IHn. auto.
+ rewrite Z.mul_assoc.
+ rewrite inj_S.
+ rewrite <- Z.pow_succ_r; auto with zarith.
+Qed.
+
(* To compute the necessary height *)
Fixpoint plength (p: positive) : positive :=
@@ -212,8 +230,8 @@ Fixpoint extend_tr (n : nat) {struct n}: (word w (S (n + m))) :=
End ExtendMax.
-Implicit Arguments extend_tr[w m].
-Implicit Arguments castm[w m n].
+Arguments extend_tr [w m] v n.
+Arguments castm [w m n] H x.
@@ -287,11 +305,7 @@ Section CompareRec.
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.
+ compare0_m x = (w_to_Z w_0 ?= wm_to_Z x).
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.
@@ -308,29 +322,25 @@ Section CompareRec.
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.
+ compare0_mn n x = (0 ?= double_to_Z n x).
+ Proof.
intros n; elim n; clear n; auto.
- intros x; generalize (spec_compare0_m x); rewrite w_to_Z_0; auto.
+ intros x; rewrite spec_compare0_m; rewrite w_to_Z_0; auto.
intros n Hrec x; case x; unfold compare0_mn; fold compare0_mn; auto.
+ fold word in *.
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.
+ rewrite 2 Hrec.
+ simpl double_to_Z.
+ set (wB := DoubleBase.double_wB wm_base n).
+ case Zcompare_spec; intros Cmp.
+ rewrite <- Cmp. reflexivity.
+ symmetry. apply Zgt_lt, Zlt_gt. (* ;-) *)
+ assert (0 < wB).
+ unfold wB, 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.
+ case (double_to_Z_pos n xl); auto with zarith.
+ case (double_to_Z_pos n xh); intros; exfalso; omega.
Qed.
Fixpoint compare_mn_1 (n:nat) : word wm n -> w -> comparison :=
@@ -348,17 +358,9 @@ Section CompareRec.
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.
+ compare x y = Zcompare (w_to_Z x) (w_to_Z y).
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.
+ compare_m x y = Zcompare (wm_to_Z x) (w_to_Z y).
Variable wm_base_lt: forall x,
0 <= w_to_Z x < base (wm_base).
@@ -369,8 +371,8 @@ Section CompareRec.
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)).
+ rewrite Pshiftl_nat_S, base_xO.
+ set (u := base (Pos.shiftl_nat wm_base n)).
assert (0 < u).
unfold u, base; auto with zarith.
replace (u^2) with (u * u); simpl; auto with zarith.
@@ -380,26 +382,23 @@ Section CompareRec.
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.
+ compare_mn_1 n x y = Zcompare (double_to_Z n x) (w_to_Z y).
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.
+ intros y; rewrite spec_compare; rewrite w_to_Z_0. reflexivity.
+ intros xh xl y; simpl;
+ rewrite spec_compare0_mn, Hrec. case Zcompare_spec.
+ intros H1b.
rewrite <- H1b; rewrite Zmult_0_l; rewrite Zplus_0_l; auto.
- apply Hrec.
- apply Zlt_gt.
+ symmetry. 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.
+ case (double_to_Z_pos n xh); intros; exfalso; omega.
Qed.
End CompareRec.
@@ -433,22 +432,6 @@ Section AddS.
End AddS.
-
- Lemma spec_opp: forall u x y,
- match u with
- | Eq => y = x
- | Lt => y < x
- | Gt => y > x
- end ->
- match CompOpp 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.
@@ -474,34 +457,112 @@ End AddS.
Variable w: Type.
- Theorem digits_zop: forall w (x: znz_op w),
- znz_digits (mk_zn2z_op x) = xO (znz_digits x).
+ Theorem digits_zop: forall t (ops : ZnZ.Ops t),
+ ZnZ.digits (mk_zn2z_ops ops) = xO (ZnZ.digits ops).
+ Proof.
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).
+ Theorem digits_kzop: forall t (ops : ZnZ.Ops t),
+ ZnZ.digits (mk_zn2z_ops_karatsuba ops) = xO (ZnZ.digits ops).
+ Proof.
intros ww x; auto.
Qed.
- Theorem make_zop: forall w (x: znz_op w),
- znz_to_Z (mk_zn2z_op x) =
+ Theorem make_zop: forall t (ops : ZnZ.Ops t),
+ @ZnZ.to_Z _ (mk_zn2z_ops ops) =
fun z => match z with
- W0 => 0
- | WW xh xl => znz_to_Z x xh * base (znz_digits x)
- + znz_to_Z x xl
+ | W0 => 0
+ | WW xh xl => ZnZ.to_Z xh * base (ZnZ.digits ops)
+ + ZnZ.to_Z xl
end.
+ Proof.
intros ww x; auto.
Qed.
- Theorem make_kzop: forall w (x: znz_op w),
- znz_to_Z (mk_zn2z_op_karatsuba x) =
+ Theorem make_kzop: forall t (ops: ZnZ.Ops t),
+ @ZnZ.to_Z _ (mk_zn2z_ops_karatsuba ops) =
fun z => match z with
- W0 => 0
- | WW xh xl => znz_to_Z x xh * base (znz_digits x)
- + znz_to_Z x xl
+ | W0 => 0
+ | WW xh xl => ZnZ.to_Z xh * base (ZnZ.digits ops)
+ + ZnZ.to_Z xl
end.
+ Proof.
intros ww x; auto.
Qed.
End SimplOp.
+
+(** Abstract vision of a datatype of arbitrary-large numbers.
+ Concrete operations can be derived from these generic
+ fonctions, in particular from [iter_t] and [same_level].
+*)
+
+Module Type NAbstract.
+
+(** The domains: a sequence of [Z/nZ] structures *)
+
+Parameter dom_t : nat -> Type.
+Declare Instance dom_op n : ZnZ.Ops (dom_t n).
+Declare Instance dom_spec n : ZnZ.Specs (dom_op n).
+
+Axiom digits_dom_op : forall n,
+ ZnZ.digits (dom_op n) = Pos.shiftl_nat (ZnZ.digits (dom_op 0)) n.
+
+(** The type [t] of arbitrary-large numbers, with abstract constructor [mk_t]
+ and destructor [destr_t] and iterator [iter_t] *)
+
+Parameter t : Type.
+
+Parameter mk_t : forall (n:nat), dom_t n -> t.
+
+Inductive View_t : t -> Prop :=
+ Mk_t : forall n (x : dom_t n), View_t (mk_t n x).
+
+Axiom destr_t : forall x, View_t x. (* i.e. every x is a (mk_t n xw) *)
+
+Parameter iter_t : forall {A:Type}(f : forall n, dom_t n -> A), t -> A.
+
+Axiom iter_mk_t : forall A (f:forall n, dom_t n -> A),
+ forall n x, iter_t f (mk_t n x) = f n x.
+
+(** Conversion to [ZArith] *)
+
+Parameter to_Z : t -> Z.
+Local Notation "[ x ]" := (to_Z x).
+
+Axiom spec_mk_t : forall n x, [mk_t n x] = ZnZ.to_Z x.
+
+(** [reduce] is like [mk_t], but try to minimise the level of the number *)
+
+Parameter reduce : forall (n:nat), dom_t n -> t.
+Axiom spec_reduce : forall n x, [reduce n x] = ZnZ.to_Z x.
+
+(** Number of level in the tree representation of a number.
+ NB: This function isn't a morphism for setoid [eq]. *)
+
+Definition level := iter_t (fun n _ => n).
+
+(** [same_level] and its rich specification, indexed by [level] *)
+
+Parameter same_level : forall {A:Type}
+ (f : forall n, dom_t n -> dom_t n -> A), t -> t -> A.
+
+Axiom spec_same_level_dep :
+ forall res
+ (P : nat -> Z -> Z -> res -> Prop)
+ (Pantimon : forall n m z z' r, (n <= m)%nat -> P m z z' r -> P n z z' r)
+ (f : forall n, dom_t n -> dom_t n -> res)
+ (Pf: forall n x y, P n (ZnZ.to_Z x) (ZnZ.to_Z y) (f n x y)),
+ forall x y, P (level x) [x] [y] (same_level f x y).
+
+(** [mk_t_S] : building a number of the next level *)
+
+Parameter mk_t_S : forall (n:nat), zn2z (dom_t n) -> t.
+
+Axiom spec_mk_t_S : forall n (x:zn2z (dom_t n)),
+ [mk_t_S n x] = zn2z_to_Z (base (ZnZ.digits (dom_op n))) ZnZ.to_Z x.
+
+Axiom mk_t_S_level : forall n x, level (mk_t_S n x) = S n.
+
+End NAbstract.
diff --git a/theories/Numbers/Natural/Binary/NBinary.v b/theories/Numbers/Natural/Binary/NBinary.v
index 029fdfca..43ca67dd 100644
--- a/theories/Numbers/Natural/Binary/NBinary.v
+++ b/theories/Numbers/Natural/Binary/NBinary.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,150 +8,15 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NBinary.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
Require Import BinPos.
Require Export BinNat.
Require Import NAxioms NProperties.
Local Open Scope N_scope.
-(** * Implementation of [NAxiomsSig] module type via [BinNat.N] *)
-
-Module NBinaryAxiomsMod <: NAxiomsSig.
-
-(** Bi-directional induction. *)
-
-Theorem bi_induction :
- forall A : N -> Prop, Proper (eq==>iff) A ->
- A N0 -> (forall n, A n <-> A (Nsucc n)) -> forall n : N, A n.
-Proof.
-intros A A_wd A0 AS. apply Nrect. assumption. intros; now apply -> AS.
-Qed.
-
-(** Basic operations. *)
-
-Definition eq_equiv : Equivalence (@eq N) := eq_equivalence.
-Local Obligation Tactic := simpl_relation.
-Program Instance succ_wd : Proper (eq==>eq) Nsucc.
-Program Instance pred_wd : Proper (eq==>eq) Npred.
-Program Instance add_wd : Proper (eq==>eq==>eq) Nplus.
-Program Instance sub_wd : Proper (eq==>eq==>eq) Nminus.
-Program Instance mul_wd : Proper (eq==>eq==>eq) Nmult.
-
-Definition pred_succ := Npred_succ.
-Definition add_0_l := Nplus_0_l.
-Definition add_succ_l := Nplus_succ.
-Definition sub_0_r := Nminus_0_r.
-Definition sub_succ_r := Nminus_succ_r.
-Definition mul_0_l := Nmult_0_l.
-Definition mul_succ_l n m := eq_trans (Nmult_Sn_m n m) (Nplus_comm _ _).
-
-(** Order *)
-
-Program Instance lt_wd : Proper (eq==>eq==>iff) Nlt.
-
-Definition lt_eq_cases := Nle_lteq.
-Definition lt_irrefl := Nlt_irrefl.
-
-Theorem lt_succ_r : forall n m, n < (Nsucc m) <-> n <= m.
-Proof.
-intros n m; unfold Nlt, Nle; destruct n as [| p]; destruct m as [| q]; simpl;
-split; intro H; try reflexivity; try discriminate.
-destruct p; simpl; intros; discriminate. exfalso; now apply H.
-apply -> Pcompare_p_Sq in H. destruct H as [H | H].
-now rewrite H. now rewrite H, Pcompare_refl.
-apply <- Pcompare_p_Sq. case_eq ((p ?= q)%positive Eq); intro H1.
-right; now apply Pcompare_Eq_eq. now left. exfalso; now apply H.
-Qed.
-
-Theorem min_l : forall n m, n <= m -> Nmin n m = n.
-Proof.
-unfold Nmin, Nle; intros n m H.
-destruct (n ?= m); try reflexivity. now elim H.
-Qed.
-
-Theorem min_r : forall n m, m <= n -> Nmin n m = m.
-Proof.
-unfold Nmin, Nle; intros n m H.
-case_eq (n ?= m); intro H1; try reflexivity.
-now apply -> Ncompare_eq_correct.
-rewrite <- Ncompare_antisym, H1 in H; elim H; auto.
-Qed.
-
-Theorem max_l : forall n m, m <= n -> Nmax n m = n.
-Proof.
-unfold Nmax, Nle; intros n m H.
-case_eq (n ?= m); intro H1; try reflexivity.
-symmetry; now apply -> Ncompare_eq_correct.
-rewrite <- Ncompare_antisym, H1 in H; elim H; auto.
-Qed.
-
-Theorem max_r : forall n m : N, n <= m -> Nmax n m = m.
-Proof.
-unfold Nmax, Nle; intros n m H.
-destruct (n ?= m); try reflexivity. now elim H.
-Qed.
-
-(** Part specific to natural numbers, not integers. *)
-
-Theorem pred_0 : Npred 0 = 0.
-Proof.
-reflexivity.
-Qed.
-
-Definition recursion (A : Type) : A -> (N -> A -> A) -> N -> A :=
- Nrect (fun _ => A).
-Implicit Arguments recursion [A].
-
-Instance recursion_wd A (Aeq : relation A) :
- Proper (Aeq==>(eq==>Aeq==>Aeq)==>eq==>Aeq) (@recursion A).
-Proof.
-intros a a' Eaa' f f' Eff'.
-intro x; pattern x; apply Nrect.
-intros x' H; now rewrite <- H.
-clear x.
-intros x IH x' H; rewrite <- H.
-unfold recursion in *. do 2 rewrite Nrect_step.
-now apply Eff'; [| apply IH].
-Qed.
-
-Theorem recursion_0 :
- forall (A : Type) (a : A) (f : N -> A -> A), recursion a f N0 = a.
-Proof.
-intros A a f; unfold recursion; now rewrite Nrect_base.
-Qed.
-
-Theorem recursion_succ :
- forall (A : Type) (Aeq : relation A) (a : A) (f : N -> A -> A),
- Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
- forall n : N, Aeq (recursion a f (Nsucc n)) (f n (recursion a f n)).
-Proof.
-unfold recursion; intros A Aeq a f EAaa f_wd n; pattern n; apply Nrect.
-rewrite Nrect_step; rewrite Nrect_base; now apply f_wd.
-clear n; intro n; do 2 rewrite Nrect_step; intro IH. apply f_wd; [reflexivity|].
-now rewrite Nrect_step.
-Qed.
-
-(** The instantiation of operations.
- Placing them at the very end avoids having indirections in above lemmas. *)
-
-Definition t := N.
-Definition eq := @eq N.
-Definition zero := N0.
-Definition succ := Nsucc.
-Definition pred := Npred.
-Definition add := Nplus.
-Definition sub := Nminus.
-Definition mul := Nmult.
-Definition lt := Nlt.
-Definition le := Nle.
-Definition min := Nmin.
-Definition max := Nmax.
-
-End NBinaryAxiomsMod.
+(** * [BinNat.N] already implements [NAxiomSig] *)
-Module Export NBinaryPropMod := NPropFunct NBinaryAxiomsMod.
+Module N <: NAxiomsSig := N.
(*
Require Import NDefOps.
diff --git a/theories/Numbers/Natural/Peano/NPeano.v b/theories/Numbers/Natural/Peano/NPeano.v
index fbc63c04..d5df6329 100644
--- a/theories/Numbers/Natural/Peano/NPeano.v
+++ b/theories/Numbers/Natural/Peano/NPeano.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,13 +8,571 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NPeano.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
+Require Import
+ Bool Peano Peano_dec Compare_dec Plus Mult Minus Le Lt EqNat Div2 Wf_nat
+ NAxioms NProperties.
-Require Import Arith MinMax NAxioms NProperties.
+(** Functions not already defined *)
+
+Fixpoint leb n m :=
+ match n, m with
+ | O, _ => true
+ | _, O => false
+ | S n', S m' => leb n' m'
+ end.
+
+Definition ltb n m := leb (S n) m.
+
+Infix "<=?" := leb (at level 70) : nat_scope.
+Infix "<?" := ltb (at level 70) : nat_scope.
+
+Lemma leb_le n m : (n <=? m) = true <-> n <= m.
+Proof.
+ revert m.
+ induction n. split; auto with arith.
+ destruct m; simpl. now split.
+ rewrite IHn. split; auto with arith.
+Qed.
+
+Lemma ltb_lt n m : (n <? m) = true <-> n < m.
+Proof.
+ unfold ltb, lt. apply leb_le.
+Qed.
+
+Fixpoint pow n m :=
+ match m with
+ | O => 1
+ | S m => n * (pow n m)
+ end.
+
+Infix "^" := pow : nat_scope.
+
+Lemma pow_0_r : forall a, a^0 = 1.
+Proof. reflexivity. Qed.
+
+Lemma pow_succ_r : forall a b, 0<=b -> a^(S b) = a * a^b.
+Proof. reflexivity. Qed.
+
+Definition square n := n * n.
+
+Lemma square_spec n : square n = n * n.
+Proof. reflexivity. Qed.
+
+Definition Even n := exists m, n = 2*m.
+Definition Odd n := exists m, n = 2*m+1.
+
+Fixpoint even n :=
+ match n with
+ | O => true
+ | 1 => false
+ | S (S n') => even n'
+ end.
+
+Definition odd n := negb (even n).
+
+Lemma even_spec : forall n, even n = true <-> Even n.
+Proof.
+ fix 1.
+ destruct n as [|[|n]]; simpl; try rewrite even_spec; split.
+ now exists 0.
+ trivial.
+ discriminate.
+ intros (m,H). destruct m. discriminate.
+ simpl in H. rewrite <- plus_n_Sm in H. discriminate.
+ intros (m,H). exists (S m). rewrite H. simpl. now rewrite plus_n_Sm.
+ intros (m,H). destruct m. discriminate. exists m.
+ simpl in H. rewrite <- plus_n_Sm in H. inversion H. reflexivity.
+Qed.
+
+Lemma odd_spec : forall n, odd n = true <-> Odd n.
+Proof.
+ unfold odd.
+ fix 1.
+ destruct n as [|[|n]]; simpl; try rewrite odd_spec; split.
+ discriminate.
+ intros (m,H). rewrite <- plus_n_Sm in H; discriminate.
+ now exists 0.
+ trivial.
+ intros (m,H). exists (S m). rewrite H. simpl. now rewrite <- (plus_n_Sm m).
+ intros (m,H). destruct m. discriminate. exists m.
+ simpl in H. rewrite <- plus_n_Sm in H. inversion H. simpl.
+ now rewrite <- !plus_n_Sm, <- !plus_n_O.
+Qed.
+
+Lemma Even_equiv : forall n, Even n <-> Even.even n.
+Proof.
+ split. intros (p,->). apply Even.even_mult_l. do 3 constructor.
+ intros H. destruct (even_2n n H) as (p,->).
+ exists p. unfold double. simpl. now rewrite <- plus_n_O.
+Qed.
+
+Lemma Odd_equiv : forall n, Odd n <-> Even.odd n.
+Proof.
+ split. intros (p,->). rewrite <- plus_n_Sm, <- plus_n_O.
+ apply Even.odd_S. apply Even.even_mult_l. do 3 constructor.
+ intros H. destruct (odd_S2n n H) as (p,->).
+ exists p. unfold double. simpl. now rewrite <- plus_n_Sm, <- !plus_n_O.
+Qed.
+
+(* A linear, tail-recursive, division for nat.
+
+ In [divmod], [y] is the predecessor of the actual divisor,
+ and [u] is [y] minus the real remainder
+*)
+
+Fixpoint divmod x y q u :=
+ match x with
+ | 0 => (q,u)
+ | S x' => match u with
+ | 0 => divmod x' y (S q) y
+ | S u' => divmod x' y q u'
+ end
+ end.
+
+Definition div x y :=
+ match y with
+ | 0 => y
+ | S y' => fst (divmod x y' 0 y')
+ end.
+
+Definition modulo x y :=
+ match y with
+ | 0 => y
+ | S y' => y' - snd (divmod x y' 0 y')
+ end.
+
+Infix "/" := div : nat_scope.
+Infix "mod" := modulo (at level 40, no associativity) : nat_scope.
+
+Lemma divmod_spec : forall x y q u, u <= y ->
+ let (q',u') := divmod x y q u in
+ x + (S y)*q + (y-u) = (S y)*q' + (y-u') /\ u' <= y.
+Proof.
+ induction x. simpl. intuition.
+ intros y q u H. destruct u; simpl divmod.
+ generalize (IHx y (S q) y (le_n y)). destruct divmod as (q',u').
+ intros (EQ,LE); split; trivial.
+ rewrite <- EQ, <- minus_n_O, minus_diag, <- plus_n_O.
+ now rewrite !plus_Sn_m, plus_n_Sm, <- plus_assoc, mult_n_Sm.
+ generalize (IHx y q u (le_Sn_le _ _ H)). destruct divmod as (q',u').
+ intros (EQ,LE); split; trivial.
+ rewrite <- EQ.
+ rewrite !plus_Sn_m, plus_n_Sm. f_equal. now apply minus_Sn_m.
+Qed.
+
+Lemma div_mod : forall x y, y<>0 -> x = y*(x/y) + x mod y.
+Proof.
+ intros x y Hy.
+ destruct y; [ now elim Hy | clear Hy ].
+ unfold div, modulo.
+ generalize (divmod_spec x y 0 y (le_n y)).
+ destruct divmod as (q,u).
+ intros (U,V).
+ simpl in *.
+ now rewrite <- mult_n_O, minus_diag, <- !plus_n_O in U.
+Qed.
+
+Lemma mod_bound_pos : forall x y, 0<=x -> 0<y -> 0 <= x mod y < y.
+Proof.
+ intros x y Hx Hy. split. auto with arith.
+ destruct y; [ now elim Hy | clear Hy ].
+ unfold modulo.
+ apply le_n_S, le_minus.
+Qed.
+
+(** Square root *)
+
+(** The following square root function is linear (and tail-recursive).
+ With Peano representation, we can't do better. For faster algorithm,
+ see Psqrt/Zsqrt/Nsqrt...
+
+ We search the square root of n = k + p^2 + (q - r)
+ with q = 2p and 0<=r<=q. We start with p=q=r=0, hence
+ looking for the square root of n = k. Then we progressively
+ decrease k and r. When k = S k' and r=0, it means we can use (S p)
+ as new sqrt candidate, since (S k')+p^2+2p = k'+(S p)^2.
+ When k reaches 0, we have found the biggest p^2 square contained
+ in n, hence the square root of n is p.
+*)
+
+Fixpoint sqrt_iter k p q r :=
+ match k with
+ | O => p
+ | S k' => match r with
+ | O => sqrt_iter k' (S p) (S (S q)) (S (S q))
+ | S r' => sqrt_iter k' p q r'
+ end
+ end.
+
+Definition sqrt n := sqrt_iter n 0 0 0.
+
+Lemma sqrt_iter_spec : forall k p q r,
+ q = p+p -> r<=q ->
+ let s := sqrt_iter k p q r in
+ s*s <= k + p*p + (q - r) < (S s)*(S s).
+Proof.
+ induction k.
+ (* k = 0 *)
+ simpl; intros p q r Hq Hr.
+ split.
+ apply le_plus_l.
+ apply le_lt_n_Sm.
+ rewrite <- mult_n_Sm.
+ rewrite plus_assoc, (plus_comm p), <- plus_assoc.
+ apply plus_le_compat; trivial.
+ rewrite <- Hq. apply le_minus.
+ (* k = S k' *)
+ destruct r.
+ (* r = 0 *)
+ intros Hq _.
+ replace (S k + p*p + (q-0)) with (k + (S p)*(S p) + (S (S q) - S (S q))).
+ apply IHk.
+ simpl. rewrite <- plus_n_Sm. congruence.
+ auto with arith.
+ rewrite minus_diag, <- minus_n_O, <- plus_n_O. simpl.
+ rewrite <- plus_n_Sm; f_equal. rewrite <- plus_assoc; f_equal.
+ rewrite <- mult_n_Sm, (plus_comm p), <- plus_assoc. congruence.
+ (* r = S r' *)
+ intros Hq Hr.
+ replace (S k + p*p + (q-S r)) with (k + p*p + (q - r)).
+ apply IHk; auto with arith.
+ simpl. rewrite plus_n_Sm. f_equal. rewrite minus_Sn_m; auto.
+Qed.
+
+Lemma sqrt_spec : forall n,
+ (sqrt n)*(sqrt n) <= n < S (sqrt n) * S (sqrt n).
+Proof.
+ intros.
+ set (s:=sqrt n).
+ replace n with (n + 0*0 + (0-0)).
+ apply sqrt_iter_spec; auto.
+ simpl. now rewrite <- 2 plus_n_O.
+Qed.
+
+(** A linear tail-recursive base-2 logarithm
+
+ In [log2_iter], we maintain the logarithm [p] of the counter [q],
+ while [r] is the distance between [q] and the next power of 2,
+ more precisely [q + S r = 2^(S p)] and [r<2^p]. At each
+ recursive call, [q] goes up while [r] goes down. When [r]
+ is 0, we know that [q] has almost reached a power of 2,
+ and we increase [p] at the next call, while resetting [r]
+ to [q].
+
+ Graphically (numbers are [q], stars are [r]) :
+
+<<
+ 10
+ 9
+ 8
+ 7 *
+ 6 *
+ 5 ...
+ 4
+ 3 *
+ 2 *
+ 1 * *
+0 * * *
+>>
+
+ We stop when [k], the global downward counter reaches 0.
+ At that moment, [q] is the number we're considering (since
+ [k+q] is invariant), and [p] its logarithm.
+*)
+
+Fixpoint log2_iter k p q r :=
+ match k with
+ | O => p
+ | S k' => match r with
+ | O => log2_iter k' (S p) (S q) q
+ | S r' => log2_iter k' p (S q) r'
+ end
+ end.
+
+Definition log2 n := log2_iter (pred n) 0 1 0.
+
+Lemma log2_iter_spec : forall k p q r,
+ 2^(S p) = q + S r -> r < 2^p ->
+ let s := log2_iter k p q r in
+ 2^s <= k + q < 2^(S s).
+Proof.
+ induction k.
+ (* k = 0 *)
+ intros p q r EQ LT. simpl log2_iter. cbv zeta.
+ split.
+ rewrite plus_O_n.
+ apply plus_le_reg_l with (2^p).
+ simpl pow in EQ. rewrite <- plus_n_O in EQ. rewrite EQ.
+ rewrite plus_comm. apply plus_le_compat_r. now apply lt_le_S.
+ rewrite EQ, plus_comm. apply plus_lt_compat_l. apply lt_0_Sn.
+ (* k = S k' *)
+ intros p q r EQ LT. destruct r.
+ (* r = 0 *)
+ rewrite <- plus_n_Sm, <- plus_n_O in EQ.
+ rewrite plus_Sn_m, plus_n_Sm. apply IHk.
+ rewrite <- EQ. remember (S p) as p'; simpl. now rewrite <- plus_n_O.
+ unfold lt. now rewrite EQ.
+ (* r = S r' *)
+ rewrite plus_Sn_m, plus_n_Sm. apply IHk.
+ now rewrite plus_Sn_m, plus_n_Sm.
+ unfold lt.
+ now apply lt_le_weak.
+Qed.
+
+Lemma log2_spec : forall n, 0<n ->
+ 2^(log2 n) <= n < 2^(S (log2 n)).
+Proof.
+ intros.
+ set (s:=log2 n).
+ replace n with (pred n + 1).
+ apply log2_iter_spec; auto.
+ rewrite <- plus_n_Sm, <- plus_n_O.
+ symmetry. now apply S_pred with 0.
+Qed.
+
+Lemma log2_nonpos : forall n, n<=0 -> log2 n = 0.
+Proof.
+ inversion 1; now subst.
+Qed.
+
+(** * Gcd *)
+
+(** We use Euclid algorithm, which is normally not structural,
+ but Coq is now clever enough to accept this (behind modulo
+ there is a subtraction, which now preserves being a subterm)
+*)
+
+Fixpoint gcd a b :=
+ match a with
+ | O => b
+ | S a' => gcd (b mod (S a')) (S a')
+ end.
+
+Definition divide x y := exists z, y=z*x.
+Notation "( x | y )" := (divide x y) (at level 0) : nat_scope.
+
+Lemma gcd_divide : forall a b, (gcd a b | a) /\ (gcd a b | b).
+Proof.
+ fix 1.
+ intros [|a] b; simpl.
+ split.
+ now exists 0.
+ exists 1. simpl. now rewrite <- plus_n_O.
+ fold (b mod (S a)).
+ destruct (gcd_divide (b mod (S a)) (S a)) as (H,H').
+ set (a':=S a) in *.
+ split; auto.
+ rewrite (div_mod b a') at 2 by discriminate.
+ destruct H as (u,Hu), H' as (v,Hv).
+ rewrite mult_comm.
+ exists ((b/a')*v + u).
+ rewrite mult_plus_distr_r.
+ now rewrite <- mult_assoc, <- Hv, <- Hu.
+Qed.
+
+Lemma gcd_divide_l : forall a b, (gcd a b | a).
+Proof.
+ intros. apply gcd_divide.
+Qed.
+
+Lemma gcd_divide_r : forall a b, (gcd a b | b).
+Proof.
+ intros. apply gcd_divide.
+Qed.
+
+Lemma gcd_greatest : forall a b c, (c|a) -> (c|b) -> (c|gcd a b).
+Proof.
+ fix 1.
+ intros [|a] b; simpl; auto.
+ fold (b mod (S a)).
+ intros c H H'. apply gcd_greatest; auto.
+ set (a':=S a) in *.
+ rewrite (div_mod b a') in H' by discriminate.
+ destruct H as (u,Hu), H' as (v,Hv).
+ exists (v - (b/a')*u).
+ rewrite mult_comm in Hv.
+ now rewrite mult_minus_distr_r, <- Hv, <-mult_assoc, <-Hu, minus_plus.
+Qed.
+
+(** * Bitwise operations *)
+
+(** We provide here some bitwise operations for unary numbers.
+ Some might be really naive, they are just there for fullfiling
+ the same interface as other for natural representations. As
+ soon as binary representations such as NArith are available,
+ it is clearly better to convert to/from them and use their ops.
+*)
+
+Fixpoint testbit a n :=
+ match n with
+ | O => odd a
+ | S n => testbit (div2 a) n
+ end.
+
+Definition shiftl a n := iter_nat n _ double a.
+Definition shiftr a n := iter_nat n _ div2 a.
+
+Fixpoint bitwise (op:bool->bool->bool) n a b :=
+ match n with
+ | O => O
+ | S n' =>
+ (if op (odd a) (odd b) then 1 else 0) +
+ 2*(bitwise op n' (div2 a) (div2 b))
+ end.
+
+Definition land a b := bitwise andb a a b.
+Definition lor a b := bitwise orb (max a b) a b.
+Definition ldiff a b := bitwise (fun b b' => b && negb b') a a b.
+Definition lxor a b := bitwise xorb (max a b) a b.
+
+Lemma double_twice : forall n, double n = 2*n.
+Proof.
+ simpl; intros. now rewrite <- plus_n_O.
+Qed.
+
+Lemma testbit_0_l : forall n, testbit 0 n = false.
+Proof.
+ now induction n.
+Qed.
+
+Lemma testbit_odd_0 a : testbit (2*a+1) 0 = true.
+Proof.
+ unfold testbit. rewrite odd_spec. now exists a.
+Qed.
+
+Lemma testbit_even_0 a : testbit (2*a) 0 = false.
+Proof.
+ unfold testbit, odd. rewrite (proj2 (even_spec _)); trivial.
+ now exists a.
+Qed.
+
+Lemma testbit_odd_succ a n : testbit (2*a+1) (S n) = testbit a n.
+Proof.
+ unfold testbit; fold testbit.
+ rewrite <- plus_n_Sm, <- plus_n_O. f_equal.
+ apply div2_double_plus_one.
+Qed.
+
+Lemma testbit_even_succ a n : testbit (2*a) (S n) = testbit a n.
+Proof.
+ unfold testbit; fold testbit. f_equal. apply div2_double.
+Qed.
+
+Lemma shiftr_spec : forall a n m,
+ testbit (shiftr a n) m = testbit a (m+n).
+Proof.
+ induction n; intros m. trivial.
+ now rewrite <- plus_n_O.
+ now rewrite <- plus_n_Sm, <- plus_Sn_m, <- IHn.
+Qed.
+
+Lemma shiftl_spec_high : forall a n m, n<=m ->
+ testbit (shiftl a n) m = testbit a (m-n).
+Proof.
+ induction n; intros m H. trivial.
+ now rewrite <- minus_n_O.
+ destruct m. inversion H.
+ simpl. apply le_S_n in H.
+ change (shiftl a (S n)) with (double (shiftl a n)).
+ rewrite double_twice, div2_double. now apply IHn.
+Qed.
+
+Lemma shiftl_spec_low : forall a n m, m<n ->
+ testbit (shiftl a n) m = false.
+Proof.
+ induction n; intros m H. inversion H.
+ change (shiftl a (S n)) with (double (shiftl a n)).
+ destruct m; simpl.
+ unfold odd. apply negb_false_iff.
+ apply even_spec. exists (shiftl a n). apply double_twice.
+ rewrite double_twice, div2_double. apply IHn.
+ now apply lt_S_n.
+Qed.
+
+Lemma div2_bitwise : forall op n a b,
+ div2 (bitwise op (S n) a b) = bitwise op n (div2 a) (div2 b).
+Proof.
+ intros. unfold bitwise; fold bitwise.
+ destruct (op (odd a) (odd b)).
+ now rewrite div2_double_plus_one.
+ now rewrite plus_O_n, div2_double.
+Qed.
+
+Lemma odd_bitwise : forall op n a b,
+ odd (bitwise op (S n) a b) = op (odd a) (odd b).
+Proof.
+ intros. unfold bitwise; fold bitwise.
+ destruct (op (odd a) (odd b)).
+ apply odd_spec. rewrite plus_comm. eexists; eauto.
+ unfold odd. apply negb_false_iff. apply even_spec.
+ rewrite plus_O_n; eexists; eauto.
+Qed.
+
+Lemma div2_decr : forall a n, a <= S n -> div2 a <= n.
+Proof.
+ destruct a; intros. apply le_0_n.
+ apply le_trans with a.
+ apply lt_n_Sm_le, lt_div2, lt_0_Sn. now apply le_S_n.
+Qed.
+
+Lemma testbit_bitwise_1 : forall op, (forall b, op false b = false) ->
+ forall n m a b, a<=n ->
+ testbit (bitwise op n a b) m = op (testbit a m) (testbit b m).
+Proof.
+ intros op Hop.
+ induction n; intros m a b Ha.
+ simpl. inversion Ha; subst. now rewrite testbit_0_l.
+ destruct m.
+ apply odd_bitwise.
+ unfold testbit; fold testbit. rewrite div2_bitwise.
+ apply IHn; now apply div2_decr.
+Qed.
+
+Lemma testbit_bitwise_2 : forall op, op false false = false ->
+ forall n m a b, a<=n -> b<=n ->
+ testbit (bitwise op n a b) m = op (testbit a m) (testbit b m).
+Proof.
+ intros op Hop.
+ induction n; intros m a b Ha Hb.
+ simpl. inversion Ha; inversion Hb; subst. now rewrite testbit_0_l.
+ destruct m.
+ apply odd_bitwise.
+ unfold testbit; fold testbit. rewrite div2_bitwise.
+ apply IHn; now apply div2_decr.
+Qed.
+
+Lemma land_spec : forall a b n,
+ testbit (land a b) n = testbit a n && testbit b n.
+Proof.
+ intros. unfold land. apply testbit_bitwise_1; trivial.
+Qed.
+
+Lemma ldiff_spec : forall a b n,
+ testbit (ldiff a b) n = testbit a n && negb (testbit b n).
+Proof.
+ intros. unfold ldiff. apply testbit_bitwise_1; trivial.
+Qed.
+
+Lemma lor_spec : forall a b n,
+ testbit (lor a b) n = testbit a n || testbit b n.
+Proof.
+ intros. unfold lor. apply testbit_bitwise_2. trivial.
+ destruct (le_ge_dec a b). now rewrite max_r. now rewrite max_l.
+ destruct (le_ge_dec a b). now rewrite max_r. now rewrite max_l.
+Qed.
+
+Lemma lxor_spec : forall a b n,
+ testbit (lxor a b) n = xorb (testbit a n) (testbit b n).
+Proof.
+ intros. unfold lxor. apply testbit_bitwise_2. trivial.
+ destruct (le_ge_dec a b). now rewrite max_r. now rewrite max_l.
+ destruct (le_ge_dec a b). now rewrite max_r. now rewrite max_l.
+Qed.
(** * Implementation of [NAxiomsSig] by [nat] *)
-Module NPeanoAxiomsMod <: NAxiomsSig.
+Module Nat
+ <: NAxiomsSig <: UsualDecidableTypeFull <: OrderedTypeFull <: TotalOrder.
(** Bi-directional induction. *)
@@ -40,6 +598,16 @@ Proof.
reflexivity.
Qed.
+Theorem one_succ : 1 = S 0.
+Proof.
+reflexivity.
+Qed.
+
+Theorem two_succ : 2 = S 1.
+Proof.
+reflexivity.
+Qed.
+
Theorem add_0_l : forall n : nat, 0 + n = n.
Proof.
reflexivity.
@@ -57,7 +625,7 @@ Qed.
Theorem sub_succ_r : forall n m : nat, n - (S m) = pred (n - m).
Proof.
-intros n m; induction n m using nat_double_ind; simpl; auto. apply sub_0_r.
+induction n; destruct m; simpl; auto. apply sub_0_r.
Qed.
Theorem mul_0_l : forall n : nat, 0 * n = 0.
@@ -67,49 +635,32 @@ Qed.
Theorem mul_succ_l : forall n m : nat, S n * m = n * m + m.
Proof.
-intros n m; now rewrite plus_comm.
+assert (add_S_r : forall n m, n+S m = S(n+m)) by (induction n; auto).
+assert (add_comm : forall n m, n+m = m+n).
+ induction n; simpl; auto. intros; rewrite add_S_r; auto.
+intros n m; now rewrite add_comm.
Qed.
(** Order on natural numbers *)
Program Instance lt_wd : Proper (eq==>eq==>iff) lt.
-Theorem lt_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 lt_irrefl : forall n : nat, ~ (n < n).
-Proof.
-exact lt_irrefl.
-Qed.
-
Theorem lt_succ_r : forall n m : nat, n < S m <-> n <= m.
Proof.
-intros n m; split; [apply lt_n_Sm_le | apply le_lt_n_Sm].
+unfold lt; split. apply le_S_n. induction 1; auto.
Qed.
-Theorem min_l : forall n m : nat, n <= m -> min n m = n.
-Proof.
-exact min_l.
-Qed.
-
-Theorem min_r : forall n m : nat, m <= n -> min n m = m.
-Proof.
-exact min_r.
-Qed.
-Theorem max_l : forall n m : nat, m <= n -> max n m = n.
+Theorem lt_eq_cases : forall n m : nat, n <= m <-> n < m \/ n = m.
Proof.
-exact max_l.
+split.
+inversion 1; auto. rewrite lt_succ_r; auto.
+destruct 1; [|subst; auto]. rewrite <- lt_succ_r; auto.
Qed.
-Theorem max_r : forall n m : nat, n <= m -> max n m = m.
+Theorem lt_irrefl : forall n : nat, ~ (n < n).
Proof.
-exact max_r.
+induction n. intro H; inversion H. rewrite lt_succ_r; auto.
Qed.
(** Facts specific to natural numbers, not integers. *)
@@ -119,25 +670,26 @@ Proof.
reflexivity.
Qed.
-Definition recursion (A : Type) : A -> (nat -> A -> A) -> nat -> A :=
+(** Recursion fonction *)
+
+Definition recursion {A} : A -> (nat -> A -> A) -> nat -> A :=
nat_rect (fun _ => A).
-Implicit Arguments recursion [A].
-Instance recursion_wd (A : Type) (Aeq : relation A) :
- Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) (@recursion A).
+Instance recursion_wd {A} (Aeq : relation A) :
+ Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) recursion.
Proof.
intros a a' Ha f f' Hf n n' Hn. subst n'.
induction n; simpl; auto. apply Hf; auto.
Qed.
Theorem recursion_0 :
- forall (A : Type) (a : A) (f : nat -> A -> A), recursion a f 0 = a.
+ forall {A} (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),
+ forall {A} (Aeq : relation A) (a : A) (f : nat -> A -> A),
Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
forall n : nat, Aeq (recursion a f (S n)) (f n (recursion a f n)).
Proof.
@@ -149,7 +701,11 @@ Qed.
Definition t := nat.
Definition eq := @eq nat.
+Definition eqb := beq_nat.
+Definition compare := nat_compare.
Definition zero := 0.
+Definition one := 1.
+Definition two := 2.
Definition succ := S.
Definition pred := pred.
Definition add := plus.
@@ -157,81 +713,101 @@ Definition sub := minus.
Definition mul := mult.
Definition lt := lt.
Definition le := le.
+Definition ltb := ltb.
+Definition leb := leb.
+
Definition min := min.
Definition max := max.
-
-End NPeanoAxiomsMod.
-
-(** Now we apply the largest property functor *)
-
-Module Export NPeanoPropMod := NPropFunct NPeanoAxiomsMod.
-
-
-
-(** Euclidean Division *)
-
-Definition divF div x y := if leb y x then S (div (x-y) y) else 0.
-Definition modF mod x y := if leb y x then mod (x-y) y else x.
-Definition initF (_ _ : nat) := 0.
-
-Fixpoint loop {A} (F:A->A)(i:A) (n:nat) : A :=
- match n with
- | 0 => i
- | S n => F (loop F i n)
- end.
-
-Definition div x y := loop divF initF x x y.
-Definition modulo x y := loop modF initF x x y.
-Infix "/" := div : nat_scope.
-Infix "mod" := modulo (at level 40, no associativity) : nat_scope.
-
-Lemma div_mod : forall x y, y<>0 -> x = y*(x/y) + x mod y.
-Proof.
- cut (forall n x y, y<>0 -> x<=n ->
- x = y*(loop divF initF n x y) + (loop modF initF n x y)).
- intros H x y Hy. apply H; auto.
- induction n.
- simpl; unfold initF; simpl. intros. nzsimpl. auto with arith.
- simpl; unfold divF at 1, modF at 1.
- intros.
- destruct (leb y x) as [ ]_eqn:L;
- [apply leb_complete in L | apply leb_complete_conv in L].
- rewrite mul_succ_r, <- add_assoc, (add_comm y), add_assoc.
- rewrite <- IHn; auto.
- symmetry; apply sub_add; auto.
- rewrite <- NPeanoAxiomsMod.lt_succ_r.
- apply lt_le_trans with x; auto.
- apply lt_minus; auto. rewrite <- neq_0_lt_0; auto.
- nzsimpl; auto.
-Qed.
-
-Lemma mod_upper_bound : forall x y, y<>0 -> x mod y < y.
-Proof.
- cut (forall n x y, y<>0 -> x<=n -> loop modF initF n x y < y).
- intros H x y Hy. apply H; auto.
- induction n.
- simpl; unfold initF. intros. rewrite <- neq_0_lt_0; auto.
- simpl; unfold modF at 1.
- intros.
- destruct (leb y x) as [ ]_eqn:L;
- [apply leb_complete in L | apply leb_complete_conv in L]; auto.
- apply IHn; auto.
- rewrite <- NPeanoAxiomsMod.lt_succ_r.
- apply lt_le_trans with x; auto.
- apply lt_minus; auto. rewrite <- neq_0_lt_0; auto.
-Qed.
-
-Require Import NDiv.
-
-Module NDivMod <: NDivSig.
- Include NPeanoAxiomsMod.
- Definition div := div.
- Definition modulo := modulo.
- Definition div_mod := div_mod.
- Definition mod_upper_bound := mod_upper_bound.
- Local Obligation Tactic := simpl_relation.
- Program Instance div_wd : Proper (eq==>eq==>eq) div.
- Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
-End NDivMod.
-
-Module Export NDivPropMod := NDivPropFunct NDivMod NPeanoPropMod.
+Definition max_l := max_l.
+Definition max_r := max_r.
+Definition min_l := min_l.
+Definition min_r := min_r.
+
+Definition eqb_eq := beq_nat_true_iff.
+Definition compare_spec := nat_compare_spec.
+Definition eq_dec := eq_nat_dec.
+Definition leb_le := leb_le.
+Definition ltb_lt := ltb_lt.
+
+Definition Even := Even.
+Definition Odd := Odd.
+Definition even := even.
+Definition odd := odd.
+Definition even_spec := even_spec.
+Definition odd_spec := odd_spec.
+
+Program Instance pow_wd : Proper (eq==>eq==>eq) pow.
+Definition pow_0_r := pow_0_r.
+Definition pow_succ_r := pow_succ_r.
+Lemma pow_neg_r : forall a b, b<0 -> a^b = 0. inversion 1. Qed.
+Definition pow := pow.
+
+Definition square := square.
+Definition square_spec := square_spec.
+
+Definition log2_spec := log2_spec.
+Definition log2_nonpos := log2_nonpos.
+Definition log2 := log2.
+
+Definition sqrt_spec a (Ha:0<=a) := sqrt_spec a.
+Lemma sqrt_neg : forall a, a<0 -> sqrt a = 0. inversion 1. Qed.
+Definition sqrt := sqrt.
+
+Definition div := div.
+Definition modulo := modulo.
+Program Instance div_wd : Proper (eq==>eq==>eq) div.
+Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
+Definition div_mod := div_mod.
+Definition mod_bound_pos := mod_bound_pos.
+
+Definition divide := divide.
+Definition gcd := gcd.
+Definition gcd_divide_l := gcd_divide_l.
+Definition gcd_divide_r := gcd_divide_r.
+Definition gcd_greatest := gcd_greatest.
+Lemma gcd_nonneg : forall a b, 0<=gcd a b.
+Proof. intros. apply le_O_n. Qed.
+
+Definition testbit := testbit.
+Definition shiftl := shiftl.
+Definition shiftr := shiftr.
+Definition lxor := lxor.
+Definition land := land.
+Definition lor := lor.
+Definition ldiff := ldiff.
+Definition div2 := div2.
+
+Program Instance testbit_wd : Proper (eq==>eq==>Logic.eq) testbit.
+Definition testbit_odd_0 := testbit_odd_0.
+Definition testbit_even_0 := testbit_even_0.
+Definition testbit_odd_succ a n (_:0<=n) := testbit_odd_succ a n.
+Definition testbit_even_succ a n (_:0<=n) := testbit_even_succ a n.
+Lemma testbit_neg_r a n (H:n<0) : testbit a n = false.
+Proof. inversion H. Qed.
+Definition shiftl_spec_low := shiftl_spec_low.
+Definition shiftl_spec_high a n m (_:0<=m) := shiftl_spec_high a n m.
+Definition shiftr_spec a n m (_:0<=m) := shiftr_spec a n m.
+Definition lxor_spec := lxor_spec.
+Definition land_spec := land_spec.
+Definition lor_spec := lor_spec.
+Definition ldiff_spec := ldiff_spec.
+Definition div2_spec a : div2 a = shiftr a 1 := eq_refl _.
+
+(** Generic Properties *)
+
+Include NProp
+ <+ UsualMinMaxLogicalProperties <+ UsualMinMaxDecProperties.
+
+End Nat.
+
+(** [Nat] contains an [order] tactic for natural numbers *)
+
+(** Note that [Nat.order] is domain-agnostic: it will not prove
+ [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x=y]. *)
+
+Section TestOrder.
+ Let test : forall x y, x<=y -> y<=x -> x=y.
+ Proof.
+ Nat.order.
+ Qed.
+End TestOrder.
diff --git a/theories/Numbers/Natural/SpecViaZ/NSig.v b/theories/Numbers/Natural/SpecViaZ/NSig.v
index 7893a82d..aaf44ca6 100644
--- a/theories/Numbers/Natural/SpecViaZ/NSig.v
+++ b/theories/Numbers/Natural/SpecViaZ/NSig.v
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,9 +8,7 @@
(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *)
(************************************************************************)
-(*i $Id: NSig.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-Require Import ZArith Znumtheory.
+Require Import BinInt.
Open Scope Z_scope.
@@ -29,60 +27,83 @@ Module Type NType.
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).
+ Parameter spec_of_N: forall x, to_Z (of_N x) = Z.of_N x.
+ Definition to_N n := Z.to_N (to_Z n).
Definition eq n m := [n] = [m].
Definition lt n m := [n] < [m].
Definition le n m := [n] <= [m].
Parameter compare : t -> t -> comparison.
- Parameter eq_bool : t -> t -> bool.
+ Parameter eqb : t -> t -> bool.
+ Parameter ltb : t -> t -> bool.
+ Parameter leb : t -> t -> bool.
Parameter max : t -> t -> t.
Parameter min : t -> t -> t.
Parameter zero : t.
Parameter one : t.
+ Parameter two : t.
Parameter succ : t -> t.
Parameter pred : t -> t.
Parameter add : t -> t -> t.
Parameter sub : t -> t -> t.
Parameter mul : t -> t -> t.
Parameter square : t -> t.
- Parameter power_pos : t -> positive -> t.
- Parameter power : t -> N -> t.
+ Parameter pow_pos : t -> positive -> t.
+ Parameter pow_N : t -> N -> t.
+ Parameter pow : t -> t -> t.
Parameter sqrt : t -> t.
+ Parameter log2 : t -> t.
Parameter div_eucl : t -> t -> t * t.
Parameter div : t -> t -> t.
Parameter modulo : t -> t -> t.
Parameter gcd : t -> t -> t.
+ Parameter even : t -> bool.
+ Parameter odd : t -> bool.
+ Parameter testbit : t -> t -> bool.
Parameter shiftr : t -> t -> t.
Parameter shiftl : t -> t -> t.
- Parameter is_even : t -> bool.
+ Parameter land : t -> t -> t.
+ Parameter lor : t -> t -> t.
+ Parameter ldiff : t -> t -> t.
+ Parameter lxor : t -> t -> t.
+ Parameter div2 : t -> t.
- Parameter spec_compare: forall x y, compare x y = Zcompare [x] [y].
- Parameter spec_eq_bool: forall x y, eq_bool x y = Zeq_bool [x] [y].
- Parameter spec_max : forall x y, [max x y] = Zmax [x] [y].
- Parameter spec_min : forall x y, [min x y] = Zmin [x] [y].
+ Parameter spec_compare: forall x y, compare x y = ([x] ?= [y]).
+ Parameter spec_eqb : forall x y, eqb x y = ([x] =? [y]).
+ Parameter spec_ltb : forall x y, ltb x y = ([x] <? [y]).
+ Parameter spec_leb : forall x y, leb x y = ([x] <=? [y]).
+ Parameter spec_max : forall x y, [max x y] = Z.max [x] [y].
+ Parameter spec_min : forall x y, [min x y] = Z.min [x] [y].
Parameter spec_0: [zero] = 0.
Parameter spec_1: [one] = 1.
+ Parameter spec_2: [two] = 2.
Parameter spec_succ: forall n, [succ n] = [n] + 1.
Parameter spec_add: forall x y, [add x y] = [x] + [y].
- Parameter spec_pred: forall x, [pred x] = Zmax 0 ([x] - 1).
- Parameter spec_sub: forall x y, [sub x y] = Zmax 0 ([x] - [y]).
+ Parameter spec_pred: forall x, [pred x] = Z.max 0 ([x] - 1).
+ Parameter spec_sub: forall x y, [sub x y] = Z.max 0 ([x] - [y]).
Parameter spec_mul: forall x y, [mul x y] = [x] * [y].
- Parameter spec_square: forall x, [square x] = [x] * [x].
- Parameter spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
- Parameter spec_power: forall x n, [power x n] = [x] ^ Z_of_N n.
- Parameter spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
+ Parameter spec_square: forall x, [square x] = [x] * [x].
+ Parameter spec_pow_pos: forall x n, [pow_pos x n] = [x] ^ Zpos n.
+ Parameter spec_pow_N: forall x n, [pow_N x n] = [x] ^ Z.of_N n.
+ Parameter spec_pow: forall x n, [pow x n] = [x] ^ [n].
+ Parameter spec_sqrt: forall x, [sqrt x] = Z.sqrt [x].
+ Parameter spec_log2: forall x, [log2 x] = Z.log2 [x].
Parameter spec_div_eucl: forall x y,
- let (q,r) := div_eucl x y in ([q], [r]) = Zdiv_eucl [x] [y].
+ let (q,r) := div_eucl x y in ([q], [r]) = Z.div_eucl [x] [y].
Parameter spec_div: forall x y, [div x y] = [x] / [y].
Parameter spec_modulo: forall x y, [modulo x y] = [x] mod [y].
- Parameter spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b].
- Parameter spec_shiftr: forall p x, [shiftr p x] = [x] / 2^[p].
- Parameter spec_shiftl: forall p x, [shiftl p x] = [x] * 2^[p].
- Parameter spec_is_even: forall x,
- if is_even x then [x] mod 2 = 0 else [x] mod 2 = 1.
+ Parameter spec_gcd: forall a b, [gcd a b] = Z.gcd [a] [b].
+ Parameter spec_even: forall x, even x = Z.even [x].
+ Parameter spec_odd: forall x, odd x = Z.odd [x].
+ Parameter spec_testbit: forall x p, testbit x p = Z.testbit [x] [p].
+ Parameter spec_shiftr: forall x p, [shiftr x p] = Z.shiftr [x] [p].
+ Parameter spec_shiftl: forall x p, [shiftl x p] = Z.shiftl [x] [p].
+ Parameter spec_land: forall x y, [land x y] = Z.land [x] [y].
+ Parameter spec_lor: forall x y, [lor x y] = Z.lor [x] [y].
+ Parameter spec_ldiff: forall x y, [ldiff x y] = Z.ldiff [x] [y].
+ Parameter spec_lxor: forall x y, [lxor x y] = Z.lxor [x] [y].
+ Parameter spec_div2: forall x, [div2 x] = Z.div2 [x].
End NType.
@@ -90,9 +111,12 @@ Module Type NType_Notation (Import N:NType).
Notation "[ x ]" := (to_Z x).
Infix "==" := eq (at level 70).
Notation "0" := zero.
+ Notation "1" := one.
+ Notation "2" := two.
Infix "+" := add.
Infix "-" := sub.
Infix "*" := mul.
+ Infix "^" := pow.
Infix "<=" := le.
Infix "<" := lt.
End NType_Notation.
diff --git a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
index a0e096be..2c7884ac 100644
--- a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
+++ b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
@@ -1,27 +1,28 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2011 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: NSigNAxioms.v 14641 2011-11-06 11:59:10Z herbelin $ i*)
-
-Require Import ZArith Nnat NAxioms NDiv NSig.
+Require Import ZArith OrdersFacts Nnat NAxioms NSig.
(** * The interface [NSig.NType] implies the interface [NAxiomsSig] *)
-Module NTypeIsNAxioms (Import N : NType').
+Module NTypeIsNAxioms (Import NN : NType').
Hint Rewrite
- spec_0 spec_succ spec_add spec_mul spec_pred spec_sub
- spec_div spec_modulo spec_gcd spec_compare spec_eq_bool
- spec_max spec_min spec_power_pos spec_power
+ spec_0 spec_1 spec_2 spec_succ spec_add spec_mul spec_pred spec_sub
+ spec_div spec_modulo spec_gcd spec_compare spec_eqb spec_ltb spec_leb
+ spec_square spec_sqrt spec_log2 spec_max spec_min spec_pow_pos spec_pow_N
+ spec_pow spec_even spec_odd spec_testbit spec_shiftl spec_shiftr
+ spec_land spec_lor spec_ldiff spec_lxor spec_div2 spec_of_N
: nsimpl.
Ltac nsimpl := autorewrite with nsimpl.
-Ltac ncongruence := unfold eq; repeat red; intros; nsimpl; congruence.
-Ltac zify := unfold eq, lt, le in *; nsimpl.
+Ltac ncongruence := unfold eq, to_N; repeat red; intros; nsimpl; congruence.
+Ltac zify := unfold eq, lt, le, to_N in *; nsimpl.
+Ltac omega_pos n := generalize (spec_pos n); omega with *.
Local Obligation Tactic := ncongruence.
@@ -36,14 +37,29 @@ Program Instance mul_wd : Proper (eq==>eq==>eq) mul.
Theorem pred_succ : forall n, pred (succ n) == n.
Proof.
-intros. zify. generalize (spec_pos n); omega with *.
+intros. zify. omega_pos n.
Qed.
-Definition N_of_Z z := of_N (Zabs_N z).
+Theorem one_succ : 1 == succ 0.
+Proof.
+now zify.
+Qed.
+
+Theorem two_succ : 2 == succ 1.
+Proof.
+now zify.
+Qed.
+
+Definition N_of_Z z := of_N (Z.to_N z).
+
+Lemma spec_N_of_Z z : (0<=z)%Z -> [N_of_Z z] = z.
+Proof.
+ unfold N_of_Z. zify. apply Z2N.id.
+Qed.
Section Induction.
-Variable A : N.t -> Prop.
+Variable A : NN.t -> Prop.
Hypothesis A_wd : Proper (eq==>iff) A.
Hypothesis A0 : A 0.
Hypothesis AS : forall n, A n <-> A (succ n).
@@ -62,9 +78,7 @@ Proof.
intros z H1 H2.
unfold B in *. apply -> AS in H2.
setoid_replace (N_of_Z (z + 1)) with (succ (N_of_Z z)); auto.
-unfold eq. rewrite spec_succ.
-unfold N_of_Z.
-rewrite 2 spec_of_N, 2 Z_of_N_abs, 2 Zabs_eq; auto with zarith.
+unfold eq. rewrite spec_succ, 2 spec_N_of_Z; auto with zarith.
Qed.
Lemma B_holds : forall z : Z, (0 <= z)%Z -> B z.
@@ -76,9 +90,7 @@ Theorem bi_induction : forall n, A n.
Proof.
intro n. setoid_replace n with (N_of_Z (to_Z n)).
apply B_holds. apply spec_pos.
-red; unfold N_of_Z.
-rewrite spec_of_N, Z_of_N_abs, Zabs_eq; auto.
-apply spec_pos.
+red. now rewrite spec_N_of_Z by apply spec_pos.
Qed.
End Induction.
@@ -95,7 +107,7 @@ Qed.
Theorem sub_0_r : forall n, n - 0 == n.
Proof.
-intros. zify. generalize (spec_pos n); omega with *.
+intros. zify. omega_pos n.
Qed.
Theorem sub_succ_r : forall n m, n - (succ m) == pred (n - m).
@@ -115,39 +127,69 @@ Qed.
(** Order *)
-Lemma compare_spec : forall x y, CompSpec eq lt x y (compare x y).
+Lemma eqb_eq x y : eqb x y = true <-> x == y.
+Proof.
+ zify. apply Z.eqb_eq.
+Qed.
+
+Lemma leb_le x y : leb x y = true <-> x <= y.
+Proof.
+ zify. apply Z.leb_le.
+Qed.
+
+Lemma ltb_lt x y : ltb x y = true <-> x < y.
+Proof.
+ zify. apply Z.ltb_lt.
+Qed.
+
+Lemma compare_eq_iff n m : compare n m = Eq <-> n == m.
Proof.
- intros. zify. destruct (Zcompare_spec [x] [y]); auto.
+ intros. zify. apply Z.compare_eq_iff.
Qed.
-Definition eqb := eq_bool.
+Lemma compare_lt_iff n m : compare n m = Lt <-> n < m.
+Proof.
+ intros. zify. reflexivity.
+Qed.
-Lemma eqb_eq : forall x y, eq_bool x y = true <-> x == y.
+Lemma compare_le_iff n m : compare n m <> Gt <-> n <= m.
Proof.
- intros. zify. symmetry. apply Zeq_is_eq_bool.
+ intros. zify. reflexivity.
Qed.
+Lemma compare_antisym n m : compare m n = CompOpp (compare n m).
+Proof.
+ intros. zify. apply Z.compare_antisym.
+Qed.
+
+Include BoolOrderFacts NN NN NN [no inline].
+
Instance compare_wd : Proper (eq ==> eq ==> Logic.eq) compare.
Proof.
-intros x x' Hx y y' Hy. rewrite 2 spec_compare, Hx, Hy; intuition.
+intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
Qed.
-Instance lt_wd : Proper (eq ==> eq ==> iff) lt.
+Instance eqb_wd : Proper (eq ==> eq ==> Logic.eq) eqb.
Proof.
-intros x x' Hx y y' Hy; unfold lt; rewrite Hx, Hy; intuition.
+intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
Qed.
-Theorem lt_eq_cases : forall n m, n <= m <-> n < m \/ n == m.
+Instance ltb_wd : Proper (eq ==> eq ==> Logic.eq) ltb.
Proof.
-intros. zify. omega.
+intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
Qed.
-Theorem lt_irrefl : forall n, ~ n < n.
+Instance leb_wd : Proper (eq ==> eq ==> Logic.eq) leb.
Proof.
-intros. zify. omega.
+intros x x' Hx y y' Hy. zify. now rewrite Hx, Hy.
Qed.
-Theorem lt_succ_r : forall n m, n < (succ m) <-> n <= m.
+Instance lt_wd : Proper (eq ==> eq ==> iff) lt.
+Proof.
+intros x x' Hx y y' Hy; unfold lt; rewrite Hx, Hy; intuition.
+Qed.
+
+Theorem lt_succ_r : forall n m, n < succ m <-> n <= m.
Proof.
intros. zify. omega.
Qed.
@@ -179,6 +221,98 @@ Proof.
zify. auto.
Qed.
+(** Power *)
+
+Program Instance pow_wd : Proper (eq==>eq==>eq) pow.
+
+Lemma pow_0_r : forall a, a^0 == 1.
+Proof.
+ intros. now zify.
+Qed.
+
+Lemma pow_succ_r : forall a b, 0<=b -> a^(succ b) == a * a^b.
+Proof.
+ intros a b. zify. intros. now Z.nzsimpl.
+Qed.
+
+Lemma pow_neg_r : forall a b, b<0 -> a^b == 0.
+Proof.
+ intros a b. zify. intro Hb. exfalso. omega_pos b.
+Qed.
+
+Lemma pow_pow_N : forall a b, a^b == pow_N a (to_N b).
+Proof.
+ intros. zify. f_equal.
+ now rewrite Z2N.id by apply spec_pos.
+Qed.
+
+Lemma pow_N_pow : forall a b, pow_N a b == a^(of_N b).
+Proof.
+ intros. now zify.
+Qed.
+
+Lemma pow_pos_N : forall a p, pow_pos a p == pow_N a (Npos p).
+Proof.
+ intros. now zify.
+Qed.
+
+(** Square *)
+
+Lemma square_spec n : square n == n * n.
+Proof.
+ now zify.
+Qed.
+
+(** Sqrt *)
+
+Lemma sqrt_spec : forall n, 0<=n ->
+ (sqrt n)*(sqrt n) <= n /\ n < (succ (sqrt n))*(succ (sqrt n)).
+Proof.
+ intros n. zify. apply Z.sqrt_spec.
+Qed.
+
+Lemma sqrt_neg : forall n, n<0 -> sqrt n == 0.
+Proof.
+ intros n. zify. intro H. exfalso. omega_pos n.
+Qed.
+
+(** Log2 *)
+
+Lemma log2_spec : forall n, 0<n ->
+ 2^(log2 n) <= n /\ n < 2^(succ (log2 n)).
+Proof.
+ intros n. zify. change (Z.log2 [n]+1)%Z with (Z.succ (Z.log2 [n])).
+ apply Z.log2_spec.
+Qed.
+
+Lemma log2_nonpos : forall n, n<=0 -> log2 n == 0.
+Proof.
+ intros n. zify. apply Z.log2_nonpos.
+Qed.
+
+(** Even / Odd *)
+
+Definition Even n := exists m, n == 2*m.
+Definition Odd n := exists m, n == 2*m+1.
+
+Lemma even_spec n : even n = true <-> Even n.
+Proof.
+ unfold Even. zify. rewrite Z.even_spec.
+ split; intros (m,Hm).
+ - exists (N_of_Z m). zify. rewrite spec_N_of_Z; trivial. omega_pos n.
+ - exists [m]. revert Hm; now zify.
+Qed.
+
+Lemma odd_spec n : odd n = true <-> Odd n.
+Proof.
+ unfold Odd. zify. rewrite Z.odd_spec.
+ split; intros (m,Hm).
+ - exists (N_of_Z m). zify. rewrite spec_N_of_Z; trivial. omega_pos n.
+ - exists [m]. revert Hm; now zify.
+Qed.
+
+(** Div / Mod *)
+
Program Instance div_wd : Proper (eq==>eq==>eq) div.
Program Instance mod_wd : Proper (eq==>eq==>eq) modulo.
@@ -187,16 +321,131 @@ Proof.
intros a b. zify. intros. apply Z_div_mod_eq_full; auto.
Qed.
-Theorem mod_upper_bound : forall a b, ~b==0 -> modulo a b < b.
+Theorem mod_bound_pos : forall a b, 0<=a -> 0<b ->
+ 0 <= modulo a b /\ modulo a b < b.
+Proof.
+intros a b. zify. apply Z.mod_bound_pos.
+Qed.
+
+(** Gcd *)
+
+Definition divide n m := exists p, m == p*n.
+Local Notation "( x | y )" := (divide x y) (at level 0).
+
+Lemma spec_divide : forall n m, (n|m) <-> Z.divide [n] [m].
+Proof.
+ intros n m. split.
+ - intros (p,H). exists [p]. revert H; now zify.
+ - intros (z,H). exists (of_N (Z.abs_N z)). zify.
+ rewrite N2Z.inj_abs_N.
+ rewrite <- (Z.abs_eq [m]), <- (Z.abs_eq [n]) by apply spec_pos.
+ now rewrite H, Z.abs_mul.
+Qed.
+
+Lemma gcd_divide_l : forall n m, (gcd n m | n).
Proof.
-intros a b. zify. intros.
-destruct (Z_mod_lt [a] [b]); auto.
-generalize (spec_pos b); auto with zarith.
+ intros n m. apply spec_divide. zify. apply Z.gcd_divide_l.
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].
+Lemma gcd_divide_r : forall n m, (gcd n m | m).
+Proof.
+ intros n m. apply spec_divide. zify. apply Z.gcd_divide_r.
+Qed.
+
+Lemma gcd_greatest : forall n m p, (p|n) -> (p|m) -> (p|gcd n m).
+Proof.
+ intros n m p. rewrite !spec_divide. zify. apply Z.gcd_greatest.
+Qed.
+
+Lemma gcd_nonneg : forall n m, 0 <= gcd n m.
+Proof.
+ intros. zify. apply Z.gcd_nonneg.
+Qed.
+
+(** Bitwise operations *)
+
+Program Instance testbit_wd : Proper (eq==>eq==>Logic.eq) testbit.
+
+Lemma testbit_odd_0 : forall a, testbit (2*a+1) 0 = true.
+Proof.
+ intros. zify. apply Z.testbit_odd_0.
+Qed.
+
+Lemma testbit_even_0 : forall a, testbit (2*a) 0 = false.
+Proof.
+ intros. zify. apply Z.testbit_even_0.
+Qed.
+
+Lemma testbit_odd_succ : forall a n, 0<=n ->
+ testbit (2*a+1) (succ n) = testbit a n.
+Proof.
+ intros a n. zify. apply Z.testbit_odd_succ.
+Qed.
+
+Lemma testbit_even_succ : forall a n, 0<=n ->
+ testbit (2*a) (succ n) = testbit a n.
+Proof.
+ intros a n. zify. apply Z.testbit_even_succ.
+Qed.
+
+Lemma testbit_neg_r : forall a n, n<0 -> testbit a n = false.
+Proof.
+ intros a n. zify. apply Z.testbit_neg_r.
+Qed.
+
+Lemma shiftr_spec : forall a n m, 0<=m ->
+ testbit (shiftr a n) m = testbit a (m+n).
+Proof.
+ intros a n m. zify. apply Z.shiftr_spec.
+Qed.
+
+Lemma shiftl_spec_high : forall a n m, 0<=m -> n<=m ->
+ testbit (shiftl a n) m = testbit a (m-n).
+Proof.
+ intros a n m. zify. intros Hn H. rewrite Z.max_r by auto with zarith.
+ now apply Z.shiftl_spec_high.
+Qed.
+
+Lemma shiftl_spec_low : forall a n m, m<n ->
+ testbit (shiftl a n) m = false.
+Proof.
+ intros a n m. zify. intros H. now apply Z.shiftl_spec_low.
+Qed.
+
+Lemma land_spec : forall a b n,
+ testbit (land a b) n = testbit a n && testbit b n.
+Proof.
+ intros a n m. zify. now apply Z.land_spec.
+Qed.
+
+Lemma lor_spec : forall a b n,
+ testbit (lor a b) n = testbit a n || testbit b n.
+Proof.
+ intros a n m. zify. now apply Z.lor_spec.
+Qed.
+
+Lemma ldiff_spec : forall a b n,
+ testbit (ldiff a b) n = testbit a n && negb (testbit b n).
+Proof.
+ intros a n m. zify. now apply Z.ldiff_spec.
+Qed.
+
+Lemma lxor_spec : forall a b n,
+ testbit (lxor a b) n = xorb (testbit a n) (testbit b n).
+Proof.
+ intros a n m. zify. now apply Z.lxor_spec.
+Qed.
+
+Lemma div2_spec : forall a, div2 a == shiftr a 1.
+Proof.
+ intros a. zify. now apply Z.div2_spec.
+Qed.
+
+(** Recursion *)
+
+Definition recursion (A : Type) (a : A) (f : NN.t -> A -> A) (n : NN.t) :=
+ Nrect (fun _ => A) a (fun n a => f (NN.of_N n) a) (NN.to_N n).
+Arguments recursion [A] a f n.
Instance recursion_wd (A : Type) (Aeq : relation A) :
Proper (Aeq ==> (eq==>Aeq==>Aeq) ==> eq ==> Aeq) (@recursion A).
@@ -204,53 +453,35 @@ Proof.
unfold eq.
intros a a' Eaa' f f' Eff' x x' Exx'.
unfold recursion.
-unfold N.to_N.
+unfold NN.to_N.
rewrite <- Exx'; clear x' Exx'.
-replace (Zabs_N [x]) with (N_of_nat (Zabs_nat [x])).
-induction (Zabs_nat [x]).
+induction (Z.to_N [x]) using N.peano_ind.
simpl; auto.
-rewrite N_of_S, 2 Nrect_step; auto. apply Eff'; auto.
-destruct [x]; simpl; auto.
-change (nat_of_P p) with (nat_of_N (Npos p)); apply N_of_nat_of_N.
-change (nat_of_P p) with (nat_of_N (Npos p)); apply N_of_nat_of_N.
+rewrite 2 Nrect_step. now apply Eff'.
Qed.
Theorem recursion_0 :
- forall (A : Type) (a : A) (f : N.t -> A -> A), recursion a f 0 = a.
+ forall (A : Type) (a : A) (f : NN.t -> A -> A), recursion a f 0 = a.
Proof.
-intros A a f; unfold recursion, N.to_N; rewrite N.spec_0; simpl; auto.
+intros A a f; unfold recursion, NN.to_N; rewrite NN.spec_0; simpl; auto.
Qed.
Theorem recursion_succ :
- forall (A : Type) (Aeq : relation A) (a : A) (f : N.t -> A -> A),
+ forall (A : Type) (Aeq : relation A) (a : A) (f : NN.t -> A -> A),
Aeq a a -> Proper (eq==>Aeq==>Aeq) f ->
forall n, Aeq (recursion a f (succ n)) (f n (recursion a f n)).
Proof.
-unfold N.eq, recursion; intros A Aeq a f EAaa f_wd n.
-replace (N.to_N (succ n)) with (Nsucc (N.to_N n)).
+unfold eq, recursion; intros A Aeq a f EAaa f_wd n.
+replace (to_N (succ n)) with (N.succ (to_N n)) by
+ (zify; now rewrite <- Z2N.inj_succ by apply spec_pos).
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.
-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 (spec_pos n); auto with zarith.
-apply spec_pos; auto.
+zify. now rewrite Z2N.id by apply spec_pos.
+fold (recursion a f n). apply recursion_wd; auto. red; auto.
Qed.
End NTypeIsNAxioms.
-Module NType_NAxioms (N : NType)
- <: NAxiomsSig <: NDivSig <: HasCompare N <: HasEqBool N <: HasMinMax N
- := N <+ NTypeIsNAxioms.
+Module NType_NAxioms (NN : NType)
+ <: NAxiomsSig <: OrderFunctions NN <: HasMinMax NN
+ := NN <+ NTypeIsNAxioms.