aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-12-06 15:47:32 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-12-06 15:47:32 +0000
commit9764ebbb67edf73a147c536a3c4f4ed0f1a7ce9e (patch)
tree881218364deec8873c06ca90c00134ae4cac724c
parentcb74dea69e7de85f427719019bc23ed3c974c8f3 (diff)
Numbers and bitwise functions.
See NatInt/NZBits.v for the common axiomatization of bitwise functions over naturals / integers. Some specs aren't pretty, but easier to prove, see alternate statements in property functors {N,Z}Bits. Negative numbers are considered via the two's complement convention. We provide implementations for N (in Ndigits.v), for nat (quite dummy, just for completeness), for Z (new file Zdigits_def), for BigN (for the moment partly by converting to N, to be improved soon) and for BigZ. NOTA: For BigN.shiftl and BigN.shiftr, the two arguments are now in the reversed order (for consistency with the rest of the world): for instance BigN.shiftl 1 10 is 2^10. NOTA2: Zeven.Zdiv2 is _not_ doing (Zdiv _ 2), but rather (Zquot _ 2) on negative numbers. For the moment I've kept it intact, and have just added a Zdiv2' which is truly equivalent to (Zdiv _ 2). To reorganize someday ? git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@13689 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--CHANGES7
-rw-r--r--doc/stdlib/index-list.html.template5
-rw-r--r--theories/Bool/Bool.v15
-rw-r--r--theories/NArith/BinNat.v178
-rw-r--r--theories/NArith/Ndigits.v304
-rw-r--r--theories/NArith/Ndist.v2
-rw-r--r--theories/Numbers/Integer/Abstract/ZAdd.v32
-rw-r--r--theories/Numbers/Integer/Abstract/ZAddOrder.v110
-rw-r--r--theories/Numbers/Integer/Abstract/ZAxioms.v35
-rw-r--r--theories/Numbers/Integer/Abstract/ZBase.v7
-rw-r--r--theories/Numbers/Integer/Abstract/ZBits.v1908
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivEucl.v17
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivFloor.v17
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivTrunc.v15
-rw-r--r--theories/Numbers/Integer/Abstract/ZGcd.v2
-rw-r--r--theories/Numbers/Integer/Abstract/ZLcm.v37
-rw-r--r--theories/Numbers/Integer/Abstract/ZLt.v4
-rw-r--r--theories/Numbers/Integer/Abstract/ZMulOrder.v16
-rw-r--r--theories/Numbers/Integer/Abstract/ZParity.v160
-rw-r--r--theories/Numbers/Integer/Abstract/ZPow.v26
-rw-r--r--theories/Numbers/Integer/Abstract/ZProperties.v5
-rw-r--r--theories/Numbers/Integer/Abstract/ZSgnAbs.v38
-rw-r--r--theories/Numbers/Integer/BigZ/BigZ.v13
-rw-r--r--theories/Numbers/Integer/BigZ/ZMake.v178
-rw-r--r--theories/Numbers/Integer/Binary/ZBinary.v24
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSig.v16
-rw-r--r--theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v69
-rw-r--r--theories/Numbers/NatInt/NZBits.v76
-rw-r--r--theories/Numbers/NatInt/NZDiv.v12
-rw-r--r--theories/Numbers/NatInt/NZParity.v263
-rw-r--r--theories/Numbers/NatInt/NZPow.v38
-rw-r--r--theories/Numbers/Natural/Abstract/NAxioms.v27
-rw-r--r--theories/Numbers/Natural/Abstract/NBits.v1422
-rw-r--r--theories/Numbers/Natural/Abstract/NDiv.v4
-rw-r--r--theories/Numbers/Natural/Abstract/NParity.v159
-rw-r--r--theories/Numbers/Natural/Abstract/NPow.v19
-rw-r--r--theories/Numbers/Natural/Abstract/NProperties.v5
-rw-r--r--theories/Numbers/Natural/BigN/BigN.v19
-rw-r--r--theories/Numbers/Natural/BigN/NMake.v192
-rw-r--r--theories/Numbers/Natural/BigN/NMake_gen.ml4
-rw-r--r--theories/Numbers/Natural/BigN/Nbasic.v2
-rw-r--r--theories/Numbers/Natural/Binary/NBinary.v24
-rw-r--r--theories/Numbers/Natural/Peano/NPeano.v223
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSig.v20
-rw-r--r--theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v89
-rw-r--r--theories/Numbers/vo.itarget4
-rw-r--r--theories/PArith/BinPos.v20
-rw-r--r--theories/Structures/Equalities.v13
-rw-r--r--theories/Structures/EqualitiesFacts.v15
-rw-r--r--theories/ZArith/ZArith.v3
-rw-r--r--theories/ZArith/Zdigits_def.v420
-rw-r--r--theories/ZArith/Zdiv.v29
-rw-r--r--theories/ZArith/Zeven.v137
-rw-r--r--theories/ZArith/Znat.v50
-rw-r--r--theories/ZArith/Zquot.v50
-rw-r--r--theories/ZArith/vo.itarget1
56 files changed, 5818 insertions, 762 deletions
diff --git a/CHANGES b/CHANGES
index 07b8dfea4..c7e5b30d7 100644
--- a/CHANGES
+++ b/CHANGES
@@ -26,8 +26,11 @@ Libraries
or lemmas may have moved. The alternative division (Trunc convention
instead of Floor) is now named Zquot (noted ÷) and Zrem by analogy
with Haskell. TODO: say more later.
-- When creating BigN, the macro-generated part NMake_gen is much smaller,
- improvements of the generic part NMake.
+- When creating BigN, the macro-generated part NMake_gen is much smaller.
+ The generic part NMake has been reworked and improved. Some changes
+ may introduce incompatibilities. In particular, the order of the arguments
+ for BigN.shiftl and BigN.shiftr is now reversed: the number to shift now
+ comes first. By default, the power function now takes two BigN.
Internal infrastructure
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index 02c1e928d..1a3663cf8 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -197,6 +197,7 @@ through the <tt>Require Import</tt> command.</p>
theories/ZArith/Znumtheory.v
theories/ZArith/Int.v
theories/ZArith/Zpow_facts.v
+ theories/ZArith/Zdigits_def.v
theories/ZArith/Zdigits.v
</dd>
@@ -241,10 +242,12 @@ through the <tt>Require Import</tt> command.</p>
theories/Numbers/NatInt/NZOrder.v
theories/Numbers/NatInt/NZDomain.v
theories/Numbers/NatInt/NZProperties.v
+ theories/Numbers/NatInt/NZParity.v
theories/Numbers/NatInt/NZPow.v
theories/Numbers/NatInt/NZSqrt.v
theories/Numbers/NatInt/NZLog.v
theories/Numbers/NatInt/NZGcd.v
+ theories/Numbers/NatInt/NZBits.v
</dd>
</dt>
@@ -291,6 +294,7 @@ through the <tt>Require Import</tt> command.</p>
theories/Numbers/Natural/Abstract/NLog.v
theories/Numbers/Natural/Abstract/NGcd.v
theories/Numbers/Natural/Abstract/NLcm.v
+ theories/Numbers/Natural/Abstract/NBits.v
theories/Numbers/Natural/Abstract/NProperties.v
theories/Numbers/Natural/Binary/NBinary.v
theories/Numbers/Natural/Peano/NPeano.v
@@ -319,6 +323,7 @@ through the <tt>Require Import</tt> command.</p>
theories/Numbers/Integer/Abstract/ZPow.v
theories/Numbers/Integer/Abstract/ZGcd.v
theories/Numbers/Integer/Abstract/ZLcm.v
+ theories/Numbers/Integer/Abstract/ZBits.v
theories/Numbers/Integer/Abstract/ZProperties.v
theories/Numbers/Integer/Abstract/ZDivEucl.v
theories/Numbers/Integer/Abstract/ZDivFloor.v
diff --git a/theories/Bool/Bool.v b/theories/Bool/Bool.v
index f4649be04..437ce5726 100644
--- a/theories/Bool/Bool.v
+++ b/theories/Bool/Bool.v
@@ -555,6 +555,21 @@ Proof.
destr_bool.
Qed.
+Lemma negb_xorb_l : forall b b', negb (xorb b b') = xorb (negb b) b'.
+Proof.
+ destruct b,b'; trivial.
+Qed.
+
+Lemma negb_xorb_r : forall b b', negb (xorb b b') = xorb b (negb b').
+Proof.
+ destruct b,b'; trivial.
+Qed.
+
+Lemma xorb_negb_negb : forall b b', xorb (negb b) (negb b') = xorb b b'.
+Proof.
+ destruct b,b'; trivial.
+Qed.
+
(** Lemmas about the [b = true] embedding of [bool] to [Prop] *)
Lemma eq_iff_eq_true : forall b1 b2, b1 = b2 <-> (b1 = true <-> b2 = true).
diff --git a/theories/NArith/BinNat.v b/theories/NArith/BinNat.v
index 51c5b462b..81e2e06e4 100644
--- a/theories/NArith/BinNat.v
+++ b/theories/NArith/BinNat.v
@@ -27,6 +27,13 @@ Arguments Scope Npos [positive_scope].
Local Open Scope N_scope.
+(** Some local ad-hoc notation, since no interpretation of numerical
+ constants is available yet. *)
+
+Local Notation "0" := N0 : N_scope.
+Local Notation "1" := (Npos 1) : N_scope.
+Local Notation "2" := (Npos 2) : N_scope.
+
Definition Ndiscr : forall n:N, { p:positive | n = Npos p } + { n = N0 }.
Proof.
destruct n; auto.
@@ -37,42 +44,59 @@ Defined.
Definition Ndouble_plus_one x :=
match x with
- | N0 => Npos 1
- | Npos p => Npos (xI p)
+ | 0 => Npos 1
+ | Npos p => Npos p~1
end.
(** Operation x -> 2*x *)
Definition Ndouble n :=
match n with
- | N0 => N0
- | Npos p => Npos (xO p)
+ | 0 => 0
+ | Npos p => Npos p~0
end.
(** Successor *)
Definition Nsucc n :=
match n with
- | N0 => Npos 1
+ | 0 => Npos 1
| Npos p => Npos (Psucc p)
end.
(** Predecessor *)
Definition Npred (n : N) := match n with
-| N0 => N0
+| 0 => 0
| Npos p => match p with
- | xH => N0
+ | xH => 0
| _ => Npos (Ppred p)
end
end.
+(** The successor of a N can be seen as a positive *)
+
+Definition Nsucc_pos (n : N) : positive :=
+ match n with
+ | N0 => 1%positive
+ | Npos p => Psucc p
+ end.
+
+(** Similarly, the predecessor of a positive, seen as a N *)
+
+Definition Ppred_N (p:positive) : N :=
+ match p with
+ | 1 => N0
+ | p~1 => Npos (p~0)
+ | p~0 => Npos (Pdouble_minus_one p)
+ end%positive.
+
(** Addition *)
Definition Nplus n m :=
match n, m with
- | N0, _ => m
- | _, N0 => n
+ | 0, _ => m
+ | _, 0 => n
| Npos p, Npos q => Npos (p + q)
end.
@@ -82,12 +106,12 @@ Infix "+" := Nplus : N_scope.
Definition Nminus (n m : N) :=
match n, m with
-| N0, _ => N0
-| n, N0 => n
+| 0, _ => 0
+| n, 0 => n
| Npos n', Npos m' =>
match Pminus_mask n' m' with
| IsPos p => Npos p
- | _ => N0
+ | _ => 0
end
end.
@@ -97,8 +121,8 @@ Infix "-" := Nminus : N_scope.
Definition Nmult n m :=
match n, m with
- | N0, _ => N0
- | _, N0 => N0
+ | 0, _ => 0
+ | _, 0 => 0
| Npos p, Npos q => Npos (p * q)
end.
@@ -108,7 +132,7 @@ Infix "*" := Nmult : N_scope.
Definition Neqb n m :=
match n, m with
- | N0, N0 => true
+ | 0, 0 => true
| Npos n, Npos m => Peqb n m
| _,_ => false
end.
@@ -117,9 +141,9 @@ Definition Neqb n m :=
Definition Ncompare n m :=
match n, m with
- | N0, N0 => Eq
- | N0, Npos m' => Lt
- | Npos n', N0 => Gt
+ | 0, 0 => Eq
+ | 0, Npos m' => Lt
+ | Npos n', 0 => Gt
| Npos n', Npos m' => (n' ?= m')%positive Eq
end.
@@ -162,7 +186,7 @@ Definition nat_of_N (a:N) :=
Definition N_of_nat (n:nat) :=
match n with
- | O => N0
+ | O => 0
| S n' => Npos (P_of_succ_nat n')
end.
@@ -178,43 +202,43 @@ Defined.
Lemma N_ind_double :
forall (a:N) (P:N -> Prop),
- P N0 ->
+ P 0 ->
(forall a, P a -> P (Ndouble a)) ->
(forall a, P a -> P (Ndouble_plus_one a)) -> P a.
Proof.
- intros; elim a. trivial.
- simple induction p. intros.
- apply (H1 (Npos p0)); trivial.
- intros; apply (H0 (Npos p0)); trivial.
- intros; apply (H1 N0); assumption.
+ intros a P P0 P2 PS2. destruct a as [|p]. trivial.
+ induction p as [p IHp|p IHp| ].
+ now apply (PS2 (Npos p)).
+ now apply (P2 (Npos p)).
+ now apply (PS2 0).
Qed.
Lemma N_rec_double :
forall (a:N) (P:N -> Set),
- P N0 ->
+ P 0 ->
(forall a, P a -> P (Ndouble a)) ->
(forall a, P a -> P (Ndouble_plus_one a)) -> P a.
Proof.
- intros; elim a. trivial.
- simple induction p. intros.
- apply (H1 (Npos p0)); trivial.
- intros; apply (H0 (Npos p0)); trivial.
- intros; apply (H1 N0); assumption.
+ intros a P P0 P2 PS2. destruct a as [|p]. trivial.
+ induction p as [p IHp|p IHp| ].
+ now apply (PS2 (Npos p)).
+ now apply (P2 (Npos p)).
+ now apply (PS2 0).
Qed.
(** Peano induction on binary natural numbers *)
Definition Nrect
- (P : N -> Type) (a : P N0)
+ (P : N -> Type) (a : P 0)
(f : forall n : N, P n -> P (Nsucc n)) (n : N) : P n :=
let f' (p : positive) (x : P (Npos p)) := f (Npos p) x in
let P' (p : positive) := P (Npos p) in
match n return (P n) with
-| N0 => a
-| Npos p => Prect P' (f N0 a) f' p
+| 0 => a
+| Npos p => Prect P' (f 0 a) f' p
end.
-Theorem Nrect_base : forall P a f, Nrect P a f N0 = a.
+Theorem Nrect_base : forall P a f, Nrect P a f 0 = a.
Proof.
intros P a f; simpl; reflexivity.
Qed.
@@ -229,7 +253,7 @@ Definition Nind (P : N -> Prop) := Nrect P.
Definition Nrec (P : N -> Set) := Nrect P.
-Theorem Nrec_base : forall P a f, Nrec P a f N0 = a.
+Theorem Nrec_base : forall P a f, Nrec P a f 0 = a.
Proof.
intros P a f; unfold Nrec; apply Nrect_base.
Qed.
@@ -248,14 +272,43 @@ case_eq (Psucc p); try (intros q H; rewrite <- H; now rewrite Ppred_succ).
intro H; false_hyp H Psucc_not_one.
Qed.
+Theorem Npred_minus : forall n, Npred n = Nminus n (Npos 1).
+Proof.
+ intros [|[p|p|]]; trivial.
+Qed.
+
+Lemma Nsucc_pred : forall n, n<>0 -> Nsucc (Npred n) = n.
+Proof.
+ intros [|n] H; (now destruct H) || clear H.
+ rewrite Npred_minus. simpl. destruct n; simpl; trivial.
+ f_equal; apply Psucc_o_double_minus_one_eq_xO.
+Qed.
+
+(** Properties of mixed successor and predecessor. *)
+
+Lemma Ppred_N_spec : forall p, Ppred_N p = Npred (Npos p).
+Proof.
+ now destruct p.
+Qed.
+
+Lemma Nsucc_pos_spec : forall n, Npos (Nsucc_pos n) = Nsucc n.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma Ppred_Nsucc : forall n, Ppred_N (Nsucc_pos n) = n.
+Proof.
+ intros. now rewrite Ppred_N_spec, Nsucc_pos_spec, Npred_succ.
+Qed.
+
(** Properties of addition *)
-Theorem Nplus_0_l : forall n:N, N0 + n = n.
+Theorem Nplus_0_l : forall n:N, 0 + n = n.
Proof.
reflexivity.
Qed.
-Theorem Nplus_0_r : forall n:N, n + N0 = n.
+Theorem Nplus_0_r : forall n:N, n + 0 = n.
Proof.
destruct n; reflexivity.
Qed.
@@ -285,7 +338,7 @@ destruct n, m.
simpl; rewrite Pplus_succ_permute_l; reflexivity.
Qed.
-Theorem Nsucc_0 : forall n : N, Nsucc n <> N0.
+Theorem Nsucc_0 : forall n : N, Nsucc n <> 0.
Proof.
now destruct n.
Qed.
@@ -308,7 +361,7 @@ Qed.
(** Properties of subtraction. *)
-Lemma Nminus_N0_Nle : forall n n' : N, n - n' = N0 <-> n <= n'.
+Lemma Nminus_N0_Nle : forall n n' : N, n - n' = 0 <-> n <= n'.
Proof.
intros [| p] [| q]; unfold Nle; simpl;
split; intro H; try easy.
@@ -319,7 +372,7 @@ subst. now rewrite Pminus_mask_diag.
now rewrite Pminus_mask_Lt.
Qed.
-Theorem Nminus_0_r : forall n : N, n - N0 = n.
+Theorem Nminus_0_r : forall n : N, n - 0 = n.
Proof.
now destruct n.
Qed.
@@ -332,14 +385,9 @@ simpl. rewrite Pminus_mask_succ_r, Pminus_mask_carry_spec.
now destruct (Pminus_mask p q) as [|[r|r|]|].
Qed.
-Theorem Npred_minus : forall n, Npred n = Nminus n (Npos 1).
-Proof.
- intros [|[p|p|]]; trivial.
-Qed.
-
(** Properties of multiplication *)
-Theorem Nmult_0_l : forall n:N, N0 * n = N0.
+Theorem Nmult_0_l : forall n:N, 0 * n = 0.
Proof.
reflexivity.
Qed.
@@ -386,7 +434,7 @@ Proof.
intros. rewrite ! (Nmult_comm p); apply Nmult_plus_distr_r.
Qed.
-Theorem Nmult_reg_r : forall n m p:N, p <> N0 -> n * p = m * p -> n = m.
+Theorem Nmult_reg_r : forall n m p:N, p <> 0 -> n * p = m * p -> n = m.
Proof.
destruct p; intros Hp H.
contradiction Hp; reflexivity.
@@ -405,6 +453,11 @@ Qed.
(** Properties of comparison *)
+Lemma Nle_0 : forall n, 0<=n.
+Proof.
+ now destruct n.
+Qed.
+
Lemma Ncompare_refl : forall n, (n ?= n) = Eq.
Proof.
destruct n; simpl; auto.
@@ -524,7 +577,7 @@ Qed.
(** 0 is the least natural number *)
-Theorem Ncompare_0 : forall n : N, Ncompare n N0 <> Lt.
+Theorem Ncompare_0 : forall n : N, Ncompare n 0 <> Lt.
Proof.
destruct n; discriminate.
Qed.
@@ -533,8 +586,8 @@ Qed.
Definition Ndiv2 (n:N) :=
match n with
- | N0 => N0
- | Npos 1 => N0
+ | 0 => 0
+ | Npos 1 => 0
| Npos (p~0) => Npos p
| Npos (p~1) => Npos p
end.
@@ -565,14 +618,14 @@ Qed.
Definition Npow n p :=
match p, n with
- | N0, _ => Npos 1
- | _, N0 => N0
+ | 0, _ => Npos 1
+ | _, 0 => 0
| Npos p, Npos q => Npos (Ppow q p)
end.
Infix "^" := Npow : N_scope.
-Lemma Npow_0_r : forall n, n ^ N0 = Npos 1.
+Lemma Npow_0_r : forall n, n ^ 0 = Npos 1.
Proof. reflexivity. Qed.
Lemma Npow_succ_r : forall n p, n^(Nsucc p) = n * n^p.
@@ -585,13 +638,13 @@ Qed.
Definition Nlog2 n :=
match n with
- | N0 => N0
- | Npos 1 => N0
+ | 0 => 0
+ | Npos 1 => 0
| Npos (p~0) => Npos (Psize_pos p)
| Npos (p~1) => Npos (Psize_pos p)
end.
-Lemma Nlog2_spec : forall n, N0 < n ->
+Lemma Nlog2_spec : forall n, 0 < n ->
(Npos 2)^(Nlog2 n) <= n < (Npos 2)^(Nsucc (Nlog2 n)).
Proof.
intros [|[p|p|]] H; discriminate H || clear H; simpl; split.
@@ -605,7 +658,7 @@ Proof.
reflexivity.
Qed.
-Lemma Nlog2_nonpos : forall n, n<=N0 -> Nlog2 n = N0.
+Lemma Nlog2_nonpos : forall n, n<=0 -> Nlog2 n = 0.
Proof.
intros [|n] Hn. reflexivity. now destruct Hn.
Qed.
@@ -614,19 +667,16 @@ Qed.
Definition Neven n :=
match n with
- | N0 => true
+ | 0 => true
| Npos (xO _) => true
| _ => false
end.
Definition Nodd n := negb (Neven n).
-Local Notation "1" := (Npos 1) : N_scope.
-Local Notation "2" := (Npos 2) : N_scope.
-
Lemma Neven_spec : forall n, Neven n = true <-> exists m, n=2*m.
Proof.
destruct n.
- split. now exists N0.
+ split. now exists 0.
trivial.
destruct p; simpl; split; trivial; try discriminate.
intros (m,H). now destruct m.
@@ -642,5 +692,5 @@ Proof.
destruct p; simpl; split; trivial; try discriminate.
exists (Npos p). reflexivity.
intros (m,H). now destruct m.
- exists N0. reflexivity.
+ exists 0. reflexivity.
Qed.
diff --git a/theories/NArith/Ndigits.v b/theories/NArith/Ndigits.v
index 98e88c6a2..0dd2fceaa 100644
--- a/theories/NArith/Ndigits.v
+++ b/theories/NArith/Ndigits.v
@@ -7,7 +7,7 @@
(************************************************************************)
Require Import Bool Morphisms Setoid Bvector BinPos BinNat Wf_nat
- Pnat Nnat Ndiv_def.
+ Pnat Nnat Ndiv_def Compare_dec Lt Minus.
Local Open Scope positive_scope.
@@ -192,12 +192,79 @@ Proof.
destruct a. trivial. apply Pbit_Ptestbit.
Qed.
+(** Correctness proof for [Ntestbit].
+
+ Ideally, we would say that (Ntestbit a n) is (a/2^n) mod 2
+ but that requires results about division we don't have yet.
+ Instead, we use a longer but simplier specification, and
+ obtain the nice equation later as a derived property.
+*)
+
+Lemma Nsuccdouble_bounds : forall n p, n<p -> 1+2*n<2*p.
+Proof.
+ intros [|n] [|p] H; try easy.
+ change (n<p)%positive in H. apply Ple_succ_l in H.
+ change (n~1 < p~0)%positive. apply Ple_succ_l. exact H.
+Qed.
+
+Lemma Ntestbit_spec : forall a n,
+ exists l, exists h, 0<=l<2^n /\
+ a = l + ((if Ntestbit a n then 1 else 0) + 2*h)*2^n.
+Proof.
+ intros [|a] n.
+ exists 0. exists 0. simpl; repeat split; now destruct n.
+ revert n. induction a as [a IH|a IH| ]; destruct n.
+ (* a~1, n=0 *)
+ exists 0. exists (Npos a). simpl. repeat split; now rewrite ?Pmult_1_r.
+ (* a~1 n>0 *)
+ destruct (IH (Npred (Npos p))) as (l & h & (_,H) & EQ). clear IH.
+ exists (1+2*l). exists h.
+ set (b := if Ntestbit (Npos a) (Npred (Npos p)) then 1 else 0) in EQ.
+ change (if Ntestbit _ _ then 1 else 0) with b.
+ rewrite <- (Nsucc_pred (Npos p)), Npow_succ_r by discriminate.
+ set (p' := Npred (Npos p)) in *.
+ split. split. apply Nle_0. now apply Nsuccdouble_bounds.
+ change (Npos a~1) with (1+2*(Npos a)). rewrite EQ.
+ rewrite <-Nplus_assoc. f_equal.
+ rewrite Nmult_plus_distr_l. f_equal.
+ now rewrite !Nmult_assoc, (Nmult_comm 2).
+ (* a~0 n=0 *)
+ exists 0. exists (Npos a). simpl. repeat split; now rewrite ?Pmult_1_r.
+ (* a~0 n>0 *)
+ destruct (IH (Npred (Npos p))) as (l & h & (_,H) & EQ). clear IH.
+ exists (2*l). exists h.
+ set (b := if Ntestbit (Npos a) (Npred (Npos p)) then 1 else 0) in EQ.
+ change (if Ntestbit _ _ then 1 else 0) with b.
+ rewrite <- (Nsucc_pred (Npos p)), !Npow_succ_r by discriminate.
+ set (p' := Npred (Npos p)) in *.
+ split. split. apply Nle_0. now destruct l, (2^p').
+ change (Npos a~0) with (2*(Npos a)). rewrite EQ.
+ rewrite Nmult_plus_distr_l. f_equal.
+ now rewrite !Nmult_assoc, (Nmult_comm 2).
+ (* 1 n=0 *)
+ exists 0. exists 0. simpl. now repeat split.
+ (* 1 n>0 *)
+ exists 1. exists 0. simpl. repeat split. easy. now apply Ppow_gt_1.
+Qed.
+
(** Equivalence of shifts, N and nat versions *)
+Lemma Nshiftr_nat_S : forall a n,
+ Nshiftr_nat a (S n) = Ndiv2 (Nshiftr_nat a n).
+Proof.
+ reflexivity.
+Qed.
+
+Lemma Nshiftl_nat_S : forall a n,
+ Nshiftl_nat a (S n) = Ndouble (Nshiftl_nat a n).
+Proof.
+ reflexivity.
+Qed.
+
Lemma Nshiftr_nat_equiv :
forall a n, Nshiftr_nat a (nat_of_N n) = Nshiftr a n.
Proof.
- intros a [|n]; simpl; unfold Nshiftr_nat.
+ intros a [|n]; simpl. unfold Nshiftr_nat.
trivial.
symmetry. apply iter_nat_of_P.
Qed.
@@ -224,166 +291,199 @@ Qed.
(** Correctness proofs for shifts *)
-Lemma Nshiftl_mult_pow : forall a n, Nshiftl a n = a * 2^n.
+Lemma Nshiftr_nat_spec : forall a n m,
+ Nbit (Nshiftr_nat a n) m = Nbit a (m+n).
Proof.
- intros [|a] [|n]; simpl; trivial.
- now rewrite Pmult_1_r.
- f_equal.
- set (f x := Pmult x a).
- rewrite Pmult_comm. fold (f (2^n))%positive.
- change a with (f xH).
- unfold Ppow. symmetry. now apply iter_pos_swap_gen.
+ induction n; intros m.
+ now rewrite <- plus_n_O.
+ simpl. rewrite <- plus_n_Sm, <- plus_Sn_m, <- IHn, Nshiftr_nat_S.
+ destruct (Nshiftr_nat a n) as [|[p|p|]]; simpl; trivial.
Qed.
-(*
-Lemma Nshiftr_div_pow : forall a n, Nshiftr a n = a / 2^n.
-*)
-
-(** Equality over functional streams of bits *)
+Lemma Nshiftr_spec : forall a n m,
+ Ntestbit (Nshiftr a n) m = Ntestbit a (m+n).
+Proof.
+ intros.
+ rewrite <- Nshiftr_nat_equiv, <- !Nbit_Ntestbit, nat_of_Nplus.
+ apply Nshiftr_nat_spec.
+Qed.
-Definition eqf (f g:nat -> bool) := forall n:nat, f n = g n.
+Lemma Nshiftl_nat_spec_high : forall a n m, (n<=m)%nat ->
+ Nbit (Nshiftl_nat a n) m = Nbit a (m-n).
+Proof.
+ induction n; intros m H.
+ now rewrite <- minus_n_O.
+ destruct m. inversion H. apply le_S_n in H.
+ simpl. rewrite <- IHn, Nshiftl_nat_S; trivial.
+ destruct (Nshiftl_nat a n) as [|[p|p|]]; simpl; trivial.
+Qed.
-Instance eqf_equiv : Equivalence eqf.
+Lemma Nshiftl_spec_high : forall a n m, n<=m ->
+ Ntestbit (Nshiftl a n) m = Ntestbit a (m-n).
Proof.
- split; congruence.
+ intros.
+ rewrite <- Nshiftl_nat_equiv, <- !Nbit_Ntestbit, nat_of_Nminus.
+ apply Nshiftl_nat_spec_high.
+ apply nat_compare_le. now rewrite <- nat_of_Ncompare.
Qed.
-Local Infix "==" := eqf (at level 70, no associativity).
+Lemma Nshiftl_nat_spec_low : forall a n m, (m<n)%nat ->
+ Nbit (Nshiftl_nat a n) m = false.
+Proof.
+ induction n; intros m H. inversion H.
+ rewrite Nshiftl_nat_S.
+ destruct m.
+ destruct (Nshiftl_nat a n); trivial.
+ specialize (IHn m (lt_S_n _ _ H)).
+ destruct (Nshiftl_nat a n); trivial.
+Qed.
-(** If two numbers produce the same stream of bits, they are equal. *)
+Lemma Nshiftl_spec_low : forall a n m, m<n ->
+ Ntestbit (Nshiftl a n) m = false.
+Proof.
+ intros.
+ rewrite <- Nshiftl_nat_equiv, <- Nbit_Ntestbit.
+ apply Nshiftl_nat_spec_low.
+ apply nat_compare_lt. now rewrite <- nat_of_Ncompare.
+Qed.
-Local Notation Step H := (fun n => H (S n)).
+(** Semantics of bitwise operations *)
-Lemma Pbit_faithful_0 : forall p, ~(Pbit p == (fun _ => false)).
+Lemma Pxor_semantics : forall p p' n,
+ Nbit (Pxor p p') n = xorb (Pbit p n) (Pbit p' n).
Proof.
- induction p as [p IHp|p IHp| ]; intros H; try discriminate (H O).
- apply (IHp (Step H)).
+ induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial; simpl;
+ (specialize (IHp p'); destruct Pxor; trivial; apply (IHp n)) ||
+ now destruct Pbit.
Qed.
-Lemma Pbit_faithful : forall p p', Pbit p == Pbit p' -> p = p'.
+Lemma Nxor_semantics : forall a a' n,
+ Nbit (Nxor a a') n = xorb (Nbit a n) (Nbit a' n).
Proof.
- induction p as [p IHp|p IHp| ]; intros [p'|p'|] H; trivial;
- try discriminate (H O).
- f_equal. apply (IHp _ (Step H)).
- destruct (Pbit_faithful_0 _ (Step H)).
- f_equal. apply (IHp _ (Step H)).
- symmetry in H. destruct (Pbit_faithful_0 _ (Step H)).
+ intros [|p] [|p'] n; simpl; trivial.
+ now destruct Pbit.
+ now destruct Pbit.
+ apply Pxor_semantics.
Qed.
-Lemma Nbit_faithful : forall n n', Nbit n == Nbit n' -> n = n'.
+Lemma Nxor_spec : forall a a' n,
+ Ntestbit (Nxor a a') n = xorb (Ntestbit a n) (Ntestbit a' n).
Proof.
- intros [|p] [|p'] H; trivial.
- symmetry in H. destruct (Pbit_faithful_0 _ H).
- destruct (Pbit_faithful_0 _ H).
- f_equal. apply Pbit_faithful, H.
+ intros. rewrite <- !Nbit_Ntestbit. apply Nxor_semantics.
Qed.
-Lemma Nbit_faithful_iff : forall n n', Nbit n == Nbit n' <-> n = n'.
+Lemma Por_semantics : forall p p' n,
+ Pbit (Por p p') n = (Pbit p n) || (Pbit p' n).
Proof.
- split. apply Nbit_faithful. intros; now subst.
+ induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial;
+ apply (IHp p' n) || now rewrite orb_false_r.
Qed.
-
-(** Bit operations for functional streams of bits *)
-
-Definition orf (f g:nat -> bool) (n:nat) := (f n) || (g n).
-Definition andf (f g:nat -> bool) (n:nat) := (f n) && (g n).
-Definition difff (f g:nat -> bool) (n:nat) := (f n) && negb (g n).
-Definition xorf (f g:nat -> bool) (n:nat) := xorb (f n) (g n).
-
-Instance eqf_orf : Proper (eqf==>eqf==>eqf) orf.
+Lemma Nor_semantics : forall a a' n,
+ Nbit (Nor a a') n = (Nbit a n) || (Nbit a' n).
Proof.
- unfold orf. congruence.
+ intros [|p] [|p'] n; simpl; trivial.
+ now rewrite orb_false_r.
+ apply Por_semantics.
Qed.
-Instance eqf_andf : Proper (eqf==>eqf==>eqf) andf.
+Lemma Nor_spec : forall a a' n,
+ Ntestbit (Nor a a') n = (Ntestbit a n) || (Ntestbit a' n).
Proof.
- unfold andf. congruence.
+ intros. rewrite <- !Nbit_Ntestbit. apply Nor_semantics.
Qed.
-Instance eqf_difff : Proper (eqf==>eqf==>eqf) difff.
+Lemma Pand_semantics : forall p p' n,
+ Nbit (Pand p p') n = (Pbit p n) && (Pbit p' n).
Proof.
- unfold difff. congruence.
+ induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial; simpl;
+ (specialize (IHp p'); destruct Pand; trivial; apply (IHp n)) ||
+ now rewrite andb_false_r.
Qed.
-Instance eqf_xorf : Proper (eqf==>eqf==>eqf) xorf.
+Lemma Nand_semantics : forall a a' n,
+ Nbit (Nand a a') n = (Nbit a n) && (Nbit a' n).
Proof.
- unfold xorf. congruence.
+ intros [|p] [|p'] n; simpl; trivial.
+ now rewrite andb_false_r.
+ apply Pand_semantics.
Qed.
-(** We now describe the semantics of [Nxor] and other bitwise
- operations in terms of bit streams. *)
-
-Lemma Pxor_semantics : forall p p',
- Nbit (Pxor p p') == xorf (Pbit p) (Pbit p').
+Lemma Nand_spec : forall a a' n,
+ Ntestbit (Nand a a') n = (Ntestbit a n) && (Ntestbit a' n).
Proof.
- induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial; simpl;
- (specialize (IHp p'); destruct Pxor; trivial; apply (IHp n)) ||
- (unfold xorf; now rewrite ?xorb_false_l, ?xorb_false_r).
+ intros. rewrite <- !Nbit_Ntestbit. apply Nand_semantics.
Qed.
-Lemma Nxor_semantics : forall a a',
- Nbit (Nxor a a') == xorf (Nbit a) (Nbit a').
+Lemma Pdiff_semantics : forall p p' n,
+ Nbit (Pdiff p p') n = (Pbit p n) && negb (Pbit p' n).
Proof.
- intros [|p] [|p'] n; simpl; unfold xorf; trivial.
- now rewrite xorb_false_l.
- now rewrite xorb_false_r.
- apply Pxor_semantics.
+ induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial; simpl;
+ (specialize (IHp p'); destruct Pdiff; trivial; apply (IHp n)) ||
+ now rewrite andb_true_r.
Qed.
-Lemma Por_semantics : forall p p',
- Pbit (Por p p') == orf (Pbit p) (Pbit p').
+Lemma Ndiff_semantics : forall a a' n,
+ Nbit (Ndiff a a') n = (Nbit a n) && negb (Nbit a' n).
Proof.
- induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial;
- unfold orf; apply (IHp p' n) || now rewrite orb_false_r.
+ intros [|p] [|p'] n; simpl; trivial.
+ simpl. now rewrite andb_true_r.
+ apply Pdiff_semantics.
Qed.
-Lemma Nor_semantics : forall a a',
- Nbit (Nor a a') == orf (Nbit a) (Nbit a').
+Lemma Ndiff_spec : forall a a' n,
+ Ntestbit (Ndiff a a') n = (Ntestbit a n) && negb (Ntestbit a' n).
Proof.
- intros [|p] [|p'] n; simpl; unfold orf; trivial.
- now rewrite orb_false_r.
- apply Por_semantics.
+ intros. rewrite <- !Nbit_Ntestbit. apply Ndiff_semantics.
Qed.
-Lemma Pand_semantics : forall p p',
- Nbit (Pand p p') == andf (Pbit p) (Pbit p').
+
+(** Equality over functional streams of bits *)
+
+Definition eqf (f g:nat -> bool) := forall n:nat, f n = g n.
+
+Program Instance eqf_equiv : Equivalence eqf.
+
+Local Infix "==" := eqf (at level 70, no associativity).
+
+(** If two numbers produce the same stream of bits, they are equal. *)
+
+Local Notation Step H := (fun n => H (S n)).
+
+Lemma Pbit_faithful_0 : forall p, ~(Pbit p == (fun _ => false)).
Proof.
- induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial; simpl;
- (specialize (IHp p'); destruct Pand; trivial; apply (IHp n)) ||
- (unfold andf; now rewrite andb_false_r).
+ induction p as [p IHp|p IHp| ]; intros H; try discriminate (H O).
+ apply (IHp (Step H)).
Qed.
-Lemma Nand_semantics : forall a a',
- Nbit (Nand a a') == andf (Nbit a) (Nbit a').
+Lemma Pbit_faithful : forall p p', Pbit p == Pbit p' -> p = p'.
Proof.
- intros [|p] [|p'] n; simpl; unfold andf; trivial.
- now rewrite andb_false_r.
- apply Pand_semantics.
+ induction p as [p IHp|p IHp| ]; intros [p'|p'|] H; trivial;
+ try discriminate (H O).
+ f_equal. apply (IHp _ (Step H)).
+ destruct (Pbit_faithful_0 _ (Step H)).
+ f_equal. apply (IHp _ (Step H)).
+ symmetry in H. destruct (Pbit_faithful_0 _ (Step H)).
Qed.
-Lemma Pdiff_semantics : forall p p',
- Nbit (Pdiff p p') == difff (Pbit p) (Pbit p').
+Lemma Nbit_faithful : forall n n', Nbit n == Nbit n' -> n = n'.
Proof.
- induction p as [p IHp|p IHp|]; intros [p'|p'|] [|n]; trivial; simpl;
- (specialize (IHp p'); destruct Pdiff; trivial; apply (IHp n)) ||
- (unfold difff; simpl; now rewrite andb_true_r).
+ intros [|p] [|p'] H; trivial.
+ symmetry in H. destruct (Pbit_faithful_0 _ H).
+ destruct (Pbit_faithful_0 _ H).
+ f_equal. apply Pbit_faithful, H.
Qed.
-Lemma Ndiff_semantics : forall a a',
- Nbit (Ndiff a a') == difff (Nbit a) (Nbit a').
+Lemma Nbit_faithful_iff : forall n n', Nbit n == Nbit n' <-> n = n'.
Proof.
- intros [|p] [|p'] n; simpl; unfold difff; trivial.
- simpl. now rewrite andb_true_r.
- apply Pdiff_semantics.
+ split. apply Nbit_faithful. intros; now subst.
Qed.
Hint Rewrite Nxor_semantics Nor_semantics
Nand_semantics Ndiff_semantics : bitwise_semantics.
Ltac bitwise_semantics :=
- apply Nbit_faithful; autorewrite with bitwise_semantics;
- intro n; unfold xorf, orf, andf, difff.
+ apply Nbit_faithful; intro n; autorewrite with bitwise_semantics.
(** Now, we can easily deduce properties of [Nxor] and others
from properties of [xorb] and others. *)
@@ -391,7 +491,7 @@ Ltac bitwise_semantics :=
Lemma Nxor_eq : forall a a', Nxor a a' = 0 -> a = a'.
Proof.
intros a a' H. bitwise_semantics. apply xorb_eq.
- rewrite <- Nbit_faithful_iff, Nxor_semantics in H. apply H.
+ now rewrite <- Nxor_semantics, H.
Qed.
Lemma Nxor_nilpotent : forall a, Nxor a a = 0.
@@ -593,7 +693,7 @@ Lemma Nxor_bit0 :
forall a a':N, Nbit0 (Nxor a a') = xorb (Nbit0 a) (Nbit0 a').
Proof.
intros. rewrite <- Nbit0_correct, (Nxor_semantics a a' O).
- unfold xorf. rewrite Nbit0_correct, Nbit0_correct. reflexivity.
+ rewrite Nbit0_correct, Nbit0_correct. reflexivity.
Qed.
Lemma Nxor_div2 :
@@ -601,7 +701,7 @@ Lemma Nxor_div2 :
Proof.
intros. apply Nbit_faithful. unfold eqf. intro.
rewrite (Nxor_semantics (Ndiv2 a) (Ndiv2 a') n), Ndiv2_correct, (Nxor_semantics a a' (S n)).
- unfold xorf. rewrite 2! Ndiv2_correct.
+ rewrite 2! Ndiv2_correct.
reflexivity.
Qed.
diff --git a/theories/NArith/Ndist.v b/theories/NArith/Ndist.v
index 0b61718f8..9d399f5cd 100644
--- a/theories/NArith/Ndist.v
+++ b/theories/NArith/Ndist.v
@@ -301,7 +301,7 @@ Proof.
cut (forall a'':N, Nxor (Npos p) a' = a'' -> Nbit a'' k = false).
intros. apply H1. reflexivity.
intro a''. case a''. intro. reflexivity.
- intros. rewrite <- H1. rewrite (Nxor_semantics (Npos p) a' k). unfold xorf in |- *.
+ intros. rewrite <- H1. rewrite (Nxor_semantics (Npos p) a' k).
rewrite
(Nplength_zeros (Npos p) (Pplength p)
(refl_equal (Nplength (Npos p))) k H0).
diff --git a/theories/Numbers/Integer/Abstract/ZAdd.v b/theories/Numbers/Integer/Abstract/ZAdd.v
index 2c386a980..f95dcc76f 100644
--- a/theories/Numbers/Integer/Abstract/ZAdd.v
+++ b/theories/Numbers/Integer/Abstract/ZAdd.v
@@ -16,23 +16,24 @@ Include ZBaseProp Z.
(** Theorems that are either not valid on N or have different proofs
on N and Z *)
+Hint Rewrite opp_0 : nz.
+
Theorem add_pred_l : forall n m, P n + m == P (n + m).
Proof.
intros n m.
rewrite <- (succ_pred n) at 2.
-rewrite add_succ_l. now rewrite pred_succ.
+now rewrite add_succ_l, pred_succ.
Qed.
Theorem add_pred_r : forall n m, n + P m == P (n + m).
Proof.
-intros n m; rewrite (add_comm n (P m)), (add_comm n m);
-apply add_pred_l.
+intros n m; rewrite 2 (add_comm n); apply add_pred_l.
Qed.
Theorem add_opp_r : forall n m, n + (- m) == n - m.
Proof.
nzinduct m.
-rewrite opp_0; rewrite sub_0_r; now rewrite add_0_r.
+now nzsimpl.
intro m. rewrite opp_succ, sub_succ_r, add_pred_r; now rewrite pred_inj_wd.
Qed.
@@ -43,7 +44,7 @@ Qed.
Theorem sub_succ_l : forall n m, S n - m == S (n - m).
Proof.
-intros n m; do 2 rewrite <- add_opp_r; now rewrite add_succ_l.
+intros n m; rewrite <- 2 add_opp_r; now rewrite add_succ_l.
Qed.
Theorem sub_pred_l : forall n m, P n - m == P (n - m).
@@ -67,7 +68,7 @@ Qed.
Theorem sub_diag : forall n, n - n == 0.
Proof.
nzinduct n.
-now rewrite sub_0_r.
+now nzsimpl.
intro n. rewrite sub_succ_r, sub_succ_l; now rewrite pred_succ.
Qed.
@@ -88,20 +89,20 @@ Qed.
Theorem add_sub_assoc : forall n m p, n + (m - p) == (n + m) - p.
Proof.
-intros n m p; do 2 rewrite <- add_opp_r; now rewrite add_assoc.
+intros n m p; rewrite <- 2 add_opp_r; now rewrite add_assoc.
Qed.
Theorem opp_involutive : forall n, - (- n) == n.
Proof.
nzinduct n.
-now do 2 rewrite opp_0.
+now nzsimpl.
intro n. rewrite opp_succ, opp_pred; now rewrite succ_inj_wd.
Qed.
Theorem opp_add_distr : forall n m, - (n + m) == - n + (- m).
Proof.
intros n m; nzinduct n.
-rewrite opp_0; now do 2 rewrite add_0_l.
+now nzsimpl.
intro n. rewrite add_succ_l; do 2 rewrite opp_succ; rewrite add_pred_l.
now rewrite pred_inj_wd.
Qed.
@@ -114,7 +115,7 @@ Qed.
Theorem opp_inj : forall n m, - n == - m -> n == m.
Proof.
-intros n m H. apply opp_wd in H. now do 2 rewrite opp_involutive in H.
+intros n m H. apply opp_wd in H. now rewrite 2 opp_involutive in H.
Qed.
Theorem opp_inj_wd : forall n m, - n == - m <-> n == m.
@@ -135,7 +136,7 @@ Qed.
Theorem sub_add_distr : forall n m p, n - (m + p) == (n - m) - p.
Proof.
intros n m p; rewrite <- add_opp_r, opp_add_distr, add_assoc.
-now do 2 rewrite add_opp_r.
+now rewrite 2 add_opp_r.
Qed.
Theorem sub_sub_distr : forall n m p, n - (m - p) == (n - m) + p.
@@ -146,7 +147,7 @@ Qed.
Theorem sub_opp_l : forall n m, - n - m == - m - n.
Proof.
-intros n m. do 2 rewrite <- add_opp_r. now rewrite add_comm.
+intros n m. rewrite <- 2 add_opp_r. now rewrite add_comm.
Qed.
Theorem sub_opp_r : forall n m, n - (- m) == n + m.
@@ -163,7 +164,7 @@ Qed.
Theorem sub_cancel_l : forall n m p, n - m == n - p <-> m == p.
Proof.
intros n m p. rewrite <- (add_cancel_l (n - m) (n - p) (- n)).
-do 2 rewrite add_sub_assoc. rewrite add_opp_diag_l; do 2 rewrite sub_0_l.
+rewrite 2 add_sub_assoc. rewrite add_opp_diag_l; rewrite 2 sub_0_l.
apply opp_inj_wd.
Qed.
@@ -250,6 +251,11 @@ Proof.
intros; now rewrite <- sub_sub_distr, sub_diag, sub_0_r.
Qed.
+Theorem sub_add : forall n m, m - n + n == m.
+Proof.
+ intros. now rewrite <- add_sub_swap, add_simpl_r.
+Qed.
+
(** Now we have two sums or differences; the name includes the two
operators and the position of the terms being canceled *)
diff --git a/theories/Numbers/Integer/Abstract/ZAddOrder.v b/theories/Numbers/Integer/Abstract/ZAddOrder.v
index b5ca908b4..423cdf585 100644
--- a/theories/Numbers/Integer/Abstract/ZAddOrder.v
+++ b/theories/Numbers/Integer/Abstract/ZAddOrder.v
@@ -18,173 +18,163 @@ Include ZOrderProp Z.
Theorem add_neg_neg : forall n m, n < 0 -> m < 0 -> n + m < 0.
Proof.
-intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_lt_mono.
+intros. rewrite <- (add_0_l 0). now apply add_lt_mono.
Qed.
Theorem add_neg_nonpos : forall n m, n < 0 -> m <= 0 -> n + m < 0.
Proof.
-intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_lt_le_mono.
+intros. rewrite <- (add_0_l 0). now apply add_lt_le_mono.
Qed.
Theorem add_nonpos_neg : forall n m, n <= 0 -> m < 0 -> n + m < 0.
Proof.
-intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_le_lt_mono.
+intros. rewrite <- (add_0_l 0). now apply add_le_lt_mono.
Qed.
Theorem add_nonpos_nonpos : forall n m, n <= 0 -> m <= 0 -> n + m <= 0.
Proof.
-intros n m H1 H2. rewrite <- (add_0_l 0). now apply add_le_mono.
+intros. rewrite <- (add_0_l 0). now apply add_le_mono.
Qed.
(** Sub and order *)
Theorem lt_0_sub : forall n m, 0 < m - n <-> n < m.
Proof.
-intros n m. stepl (0 + n < m - n + n) by symmetry; apply add_lt_mono_r.
-rewrite add_0_l; now rewrite sub_simpl_r.
+intros n m. now rewrite (add_lt_mono_r _ _ n), add_0_l, sub_simpl_r.
Qed.
Notation sub_pos := lt_0_sub (only parsing).
Theorem le_0_sub : forall n m, 0 <= m - n <-> n <= m.
Proof.
-intros n m; stepl (0 + n <= m - n + n) by symmetry; apply add_le_mono_r.
-rewrite add_0_l; now rewrite sub_simpl_r.
+intros n m. now rewrite (add_le_mono_r _ _ n), add_0_l, sub_simpl_r.
Qed.
Notation sub_nonneg := le_0_sub (only parsing).
Theorem lt_sub_0 : forall n m, n - m < 0 <-> n < m.
Proof.
-intros n m. stepl (n - m + m < 0 + m) by symmetry; apply add_lt_mono_r.
-rewrite add_0_l; now rewrite sub_simpl_r.
+intros n m. now rewrite (add_lt_mono_r _ _ m), add_0_l, sub_simpl_r.
Qed.
Notation sub_neg := lt_sub_0 (only parsing).
Theorem le_sub_0 : forall n m, n - m <= 0 <-> n <= m.
Proof.
-intros n m. stepl (n - m + m <= 0 + m) by symmetry; apply add_le_mono_r.
-rewrite add_0_l; now rewrite sub_simpl_r.
+intros n m. now rewrite (add_le_mono_r _ _ m), add_0_l, sub_simpl_r.
Qed.
Notation sub_nonpos := le_sub_0 (only parsing).
Theorem opp_lt_mono : forall n m, n < m <-> - m < - n.
Proof.
-intros n m. stepr (m + - m < m + - n) by symmetry; apply add_lt_mono_l.
-do 2 rewrite add_opp_r. rewrite sub_diag. symmetry; apply lt_0_sub.
+intros n m. now rewrite <- lt_0_sub, <- add_opp_l, <- sub_opp_r, lt_0_sub.
Qed.
Theorem opp_le_mono : forall n m, n <= m <-> - m <= - n.
Proof.
-intros n m. stepr (m + - m <= m + - n) by symmetry; apply add_le_mono_l.
-do 2 rewrite add_opp_r. rewrite sub_diag. symmetry; apply le_0_sub.
+intros n m. now rewrite <- le_0_sub, <- add_opp_l, <- sub_opp_r, le_0_sub.
Qed.
Theorem opp_pos_neg : forall n, 0 < - n <-> n < 0.
Proof.
-intro n; rewrite (opp_lt_mono n 0); now rewrite opp_0.
+intro n; now rewrite (opp_lt_mono n 0), opp_0.
Qed.
Theorem opp_neg_pos : forall n, - n < 0 <-> 0 < n.
Proof.
-intro n. rewrite (opp_lt_mono 0 n). now rewrite opp_0.
+intro n. now rewrite (opp_lt_mono 0 n), opp_0.
Qed.
Theorem opp_nonneg_nonpos : forall n, 0 <= - n <-> n <= 0.
Proof.
-intro n; rewrite (opp_le_mono n 0); now rewrite opp_0.
+intro n; now rewrite (opp_le_mono n 0), opp_0.
Qed.
Theorem opp_nonpos_nonneg : forall n, - n <= 0 <-> 0 <= n.
Proof.
-intro n. rewrite (opp_le_mono 0 n). now rewrite opp_0.
+intro n. now rewrite (opp_le_mono 0 n), opp_0.
Qed.
-Theorem lt_m1_0 : -(1) < 0.
+Theorem lt_m1_0 : -1 < 0.
Proof.
apply opp_neg_pos, lt_0_1.
Qed.
Theorem sub_lt_mono_l : forall n m p, n < m <-> p - m < p - n.
Proof.
-intros n m p. do 2 rewrite <- add_opp_r. rewrite <- add_lt_mono_l.
-apply opp_lt_mono.
+intros. now rewrite <- 2 add_opp_r, <- add_lt_mono_l, opp_lt_mono.
Qed.
Theorem sub_lt_mono_r : forall n m p, n < m <-> n - p < m - p.
Proof.
-intros n m p; do 2 rewrite <- add_opp_r; apply add_lt_mono_r.
+intros. now rewrite <- 2 add_opp_r, add_lt_mono_r.
Qed.
Theorem sub_lt_mono : forall n m p q, n < m -> q < p -> n - p < m - q.
Proof.
intros n m p q H1 H2.
apply lt_trans with (m - p);
-[now apply -> sub_lt_mono_r | now apply -> sub_lt_mono_l].
+[now apply sub_lt_mono_r | now apply sub_lt_mono_l].
Qed.
Theorem sub_le_mono_l : forall n m p, n <= m <-> p - m <= p - n.
Proof.
-intros n m p; do 2 rewrite <- add_opp_r; rewrite <- add_le_mono_l;
-apply opp_le_mono.
+intros. now rewrite <- 2 add_opp_r, <- add_le_mono_l, opp_le_mono.
Qed.
Theorem sub_le_mono_r : forall n m p, n <= m <-> n - p <= m - p.
Proof.
-intros n m p; do 2 rewrite <- add_opp_r; apply add_le_mono_r.
+intros. now rewrite <- 2 add_opp_r, add_le_mono_r.
Qed.
Theorem sub_le_mono : forall n m p q, n <= m -> q <= p -> n - p <= m - q.
Proof.
intros n m p q H1 H2.
apply le_trans with (m - p);
-[now apply -> sub_le_mono_r | now apply -> sub_le_mono_l].
+[now apply sub_le_mono_r | now apply sub_le_mono_l].
Qed.
Theorem sub_lt_le_mono : forall n m p q, n < m -> q <= p -> n - p < m - q.
Proof.
intros n m p q H1 H2.
apply lt_le_trans with (m - p);
-[now apply -> sub_lt_mono_r | now apply -> sub_le_mono_l].
+[now apply sub_lt_mono_r | now apply sub_le_mono_l].
Qed.
Theorem sub_le_lt_mono : forall n m p q, n <= m -> q < p -> n - p < m - q.
Proof.
intros n m p q H1 H2.
apply le_lt_trans with (m - p);
-[now apply -> sub_le_mono_r | now apply -> sub_lt_mono_l].
+[now apply sub_le_mono_r | now apply sub_lt_mono_l].
Qed.
Theorem le_lt_sub_lt : forall n m p q, n <= m -> p - n < q - m -> p < q.
Proof.
intros n m p q H1 H2. apply (le_lt_add_lt (- m) (- n));
-[now apply -> opp_le_mono | now do 2 rewrite add_opp_r].
+[now apply -> opp_le_mono | now rewrite 2 add_opp_r].
Qed.
Theorem lt_le_sub_lt : forall n m p q, n < m -> p - n <= q - m -> p < q.
Proof.
intros n m p q H1 H2. apply (lt_le_add_lt (- m) (- n));
-[now apply -> opp_lt_mono | now do 2 rewrite add_opp_r].
+[now apply -> opp_lt_mono | now rewrite 2 add_opp_r].
Qed.
Theorem le_le_sub_lt : forall n m p q, n <= m -> p - n <= q - m -> p <= q.
Proof.
intros n m p q H1 H2. apply (le_le_add_le (- m) (- n));
-[now apply -> opp_le_mono | now do 2 rewrite add_opp_r].
+[now apply -> opp_le_mono | now rewrite 2 add_opp_r].
Qed.
Theorem lt_add_lt_sub_r : forall n m p, n + p < m <-> n < m - p.
Proof.
-intros n m p. stepl (n + p - p < m - p) by symmetry; apply sub_lt_mono_r.
-now rewrite add_simpl_r.
+intros n m p. now rewrite (sub_lt_mono_r _ _ p), add_simpl_r.
Qed.
Theorem le_add_le_sub_r : forall n m p, n + p <= m <-> n <= m - p.
Proof.
-intros n m p. stepl (n + p - p <= m - p) by symmetry; apply sub_le_mono_r.
-now rewrite add_simpl_r.
+intros n m p. now rewrite (sub_le_mono_r _ _ p), add_simpl_r.
Qed.
Theorem lt_add_lt_sub_l : forall n m p, n + p < m <-> p < m - n.
@@ -199,14 +189,12 @@ Qed.
Theorem lt_sub_lt_add_r : forall n m p, n - p < m <-> n < m + p.
Proof.
-intros n m p. stepl (n - p + p < m + p) by symmetry; apply add_lt_mono_r.
-now rewrite sub_simpl_r.
+intros n m p. now rewrite (add_lt_mono_r _ _ p), sub_simpl_r.
Qed.
Theorem le_sub_le_add_r : forall n m p, n - p <= m <-> n <= m + p.
Proof.
-intros n m p. stepl (n - p + p <= m + p) by symmetry; apply add_le_mono_r.
-now rewrite sub_simpl_r.
+intros n m p. now rewrite (add_le_mono_r _ _ p), sub_simpl_r.
Qed.
Theorem lt_sub_lt_add_l : forall n m p, n - m < p <-> n < m + p.
@@ -221,74 +209,68 @@ Qed.
Theorem lt_sub_lt_add : forall n m p q, n - m < p - q <-> n + q < m + p.
Proof.
-intros n m p q. rewrite lt_sub_lt_add_l. rewrite add_sub_assoc.
-now rewrite <- lt_add_lt_sub_r.
+intros n m p q. now rewrite lt_sub_lt_add_l, add_sub_assoc, <- lt_add_lt_sub_r.
Qed.
Theorem le_sub_le_add : forall n m p q, n - m <= p - q <-> n + q <= m + p.
Proof.
-intros n m p q. rewrite le_sub_le_add_l. rewrite add_sub_assoc.
-now rewrite <- le_add_le_sub_r.
+intros n m p q. now rewrite le_sub_le_add_l, add_sub_assoc, <- le_add_le_sub_r.
Qed.
Theorem lt_sub_pos : forall n m, 0 < m <-> n - m < n.
Proof.
-intros n m. stepr (n - m < n - 0) by now rewrite sub_0_r. apply sub_lt_mono_l.
+intros n m. now rewrite (sub_lt_mono_l _ _ n), sub_0_r.
Qed.
Theorem le_sub_nonneg : forall n m, 0 <= m <-> n - m <= n.
Proof.
-intros n m. stepr (n - m <= n - 0) by now rewrite sub_0_r. apply sub_le_mono_l.
+intros n m. now rewrite (sub_le_mono_l _ _ n), sub_0_r.
Qed.
Theorem sub_lt_cases : forall n m p q, n - m < p - q -> n < m \/ q < p.
Proof.
-intros n m p q H. rewrite lt_sub_lt_add in H. now apply add_lt_cases.
+intros. now apply add_lt_cases, lt_sub_lt_add.
Qed.
Theorem sub_le_cases : forall n m p q, n - m <= p - q -> n <= m \/ q <= p.
Proof.
-intros n m p q H. rewrite le_sub_le_add in H. now apply add_le_cases.
+intros. now apply add_le_cases, le_sub_le_add.
Qed.
Theorem sub_neg_cases : forall n m, n - m < 0 -> n < 0 \/ 0 < m.
Proof.
-intros n m H; rewrite <- add_opp_r in H.
-setoid_replace (0 < m) with (- m < 0) using relation iff by (symmetry; apply opp_neg_pos).
-now apply add_neg_cases.
+intros.
+rewrite <- (opp_neg_pos m). apply add_neg_cases. now rewrite add_opp_r.
Qed.
Theorem sub_pos_cases : forall n m, 0 < n - m -> 0 < n \/ m < 0.
Proof.
-intros n m H; rewrite <- add_opp_r in H.
-setoid_replace (m < 0) with (0 < - m) using relation iff by (symmetry; apply opp_pos_neg).
-now apply add_pos_cases.
+intros.
+rewrite <- (opp_pos_neg m). apply add_pos_cases. now rewrite add_opp_r.
Qed.
Theorem sub_nonpos_cases : forall n m, n - m <= 0 -> n <= 0 \/ 0 <= m.
Proof.
-intros n m H; rewrite <- add_opp_r in H.
-setoid_replace (0 <= m) with (- m <= 0) using relation iff by (symmetry; apply opp_nonpos_nonneg).
-now apply add_nonpos_cases.
+intros.
+rewrite <- (opp_nonpos_nonneg m). apply add_nonpos_cases. now rewrite add_opp_r.
Qed.
Theorem sub_nonneg_cases : forall n m, 0 <= n - m -> 0 <= n \/ m <= 0.
Proof.
-intros n m H; rewrite <- add_opp_r in H.
-setoid_replace (m <= 0) with (0 <= - m) using relation iff by (symmetry; apply opp_nonneg_nonpos).
-now apply add_nonneg_cases.
+intros.
+rewrite <- (opp_nonneg_nonpos m). apply add_nonneg_cases. now rewrite add_opp_r.
Qed.
Section PosNeg.
Variable P : Z.t -> Prop.
-Hypothesis P_wd : Proper (Z.eq ==> iff) P.
+Hypothesis P_wd : Proper (eq ==> iff) P.
Theorem zero_pos_neg :
P 0 -> (forall n, 0 < n -> P n /\ P (- n)) -> forall n, P n.
Proof.
intros H1 H2 n. destruct (lt_trichotomy n 0) as [H3 | [H3 | H3]].
-apply <- opp_pos_neg in H3. apply H2 in H3. destruct H3 as [_ H3].
+apply opp_pos_neg, H2 in H3. destruct H3 as [_ H3].
now rewrite opp_involutive in H3.
now rewrite H3.
apply H2 in H3; now destruct H3.
diff --git a/theories/Numbers/Integer/Abstract/ZAxioms.v b/theories/Numbers/Integer/Abstract/ZAxioms.v
index f177755fa..237ac93ef 100644
--- a/theories/Numbers/Integer/Abstract/ZAxioms.v
+++ b/theories/Numbers/Integer/Abstract/ZAxioms.v
@@ -9,7 +9,7 @@
(************************************************************************)
Require Export NZAxioms.
-Require Import NZPow NZSqrt NZLog NZGcd NZDiv.
+Require Import Bool NZParity NZPow NZSqrt NZLog NZGcd NZDiv NZBits.
(** We obtain integers by postulating that successor of predecessor
is identity. *)
@@ -38,8 +38,15 @@ Module Type IsOpp (Import Z : NZAxiomsSig')(Import O : Opp' Z).
Axiom opp_succ : forall n, - (S n) == P (- n).
End IsOpp.
+Module Type OppCstNotation (Import A : NZAxiomsSig)(Import B : Opp A).
+ Notation "- 1" := (opp one).
+ Notation "- 2" := (opp two).
+End OppCstNotation.
+
Module Type ZAxiomsMiniSig := NZOrdAxiomsSig <+ ZAxiom <+ Opp <+ IsOpp.
-Module Type ZAxiomsMiniSig' := NZOrdAxiomsSig' <+ ZAxiom <+ Opp' <+ IsOpp.
+Module Type ZAxiomsMiniSig' := NZOrdAxiomsSig' <+ ZAxiom <+ Opp' <+ IsOpp
+ <+ OppCstNotation.
+
(** Other functions and their specifications *)
@@ -57,19 +64,9 @@ Module Type HasSgn (Import Z : ZAxiomsMiniSig').
Parameter Inline sgn : t -> t.
Axiom sgn_null : forall n, n==0 -> sgn n == 0.
Axiom sgn_pos : forall n, 0<n -> sgn n == 1.
- Axiom sgn_neg : forall n, n<0 -> sgn n == -(1).
+ Axiom sgn_neg : forall n, n<0 -> sgn n == -1.
End HasSgn.
-(** Parity functions *)
-
-Module Type Parity (Import Z : ZAxiomsMiniSig').
- Parameter Inline even odd : t -> bool.
- Definition Even n := exists m, n == 2*m.
- Definition Odd n := exists m, n == 2*m+1.
- Axiom even_spec : forall n, even n = true <-> Even n.
- Axiom odd_spec : forall n, odd n = true <-> Odd n.
-End Parity.
-
(** Divisions *)
(** First, the usual Coq convention of Truncated-Toward-Bottom
@@ -110,16 +107,18 @@ End QuotRemSpec.
Module Type ZQuot (Z:ZAxiomsMiniSig) := QuotRem Z <+ QuotRemSpec Z.
Module Type ZQuot' (Z:ZAxiomsMiniSig) := QuotRem' Z <+ QuotRemSpec Z.
-(** For pow sqrt log2 gcd, the NZ axiomatizations are enough. *)
+(** For all other functions, the NZ axiomatizations are enough. *)
(** Let's group everything *)
Module Type ZAxiomsSig :=
- ZAxiomsMiniSig <+ HasCompare <+ HasAbs <+ HasSgn <+ Parity
+ ZAxiomsMiniSig <+ HasCompare <+ HasEqBool <+ HasAbs <+ HasSgn
+ <+ NZParity.NZParity
<+ NZPow.NZPow <+ NZSqrt.NZSqrt <+ NZLog.NZLog2 <+ NZGcd.NZGcd
- <+ ZDiv <+ ZQuot.
+ <+ ZDiv <+ ZQuot <+ NZBits.NZBits.
Module Type ZAxiomsSig' :=
- ZAxiomsMiniSig' <+ HasCompare <+ HasAbs <+ HasSgn <+ Parity
+ ZAxiomsMiniSig' <+ HasCompare <+ HasEqBool <+ HasAbs <+ HasSgn
+ <+ NZParity.NZParity
<+ NZPow.NZPow' <+ NZSqrt.NZSqrt' <+ NZLog.NZLog2 <+ NZGcd.NZGcd'
- <+ ZDiv' <+ ZQuot'.
+ <+ ZDiv' <+ ZQuot' <+ NZBits.NZBits'.
diff --git a/theories/Numbers/Integer/Abstract/ZBase.v b/theories/Numbers/Integer/Abstract/ZBase.v
index f9bd8dba3..4afc10201 100644
--- a/theories/Numbers/Integer/Abstract/ZBase.v
+++ b/theories/Numbers/Integer/Abstract/ZBase.v
@@ -19,7 +19,7 @@ Include NZProp Z.
Theorem pred_inj : forall n m, P n == P m -> n == m.
Proof.
-intros n m H. apply succ_wd in H. now do 2 rewrite succ_pred in H.
+intros n m H. apply succ_wd in H. now rewrite 2 succ_pred in H.
Qed.
Theorem pred_inj_wd : forall n1 n2, P n1 == P n2 <-> n1 == n2.
@@ -27,5 +27,10 @@ Proof.
intros n1 n2; split; [apply pred_inj | apply pred_wd].
Qed.
+Lemma succ_m1 : S (-1) == 0.
+Proof.
+ now rewrite one_succ, opp_succ, opp_0, succ_pred.
+Qed.
+
End ZBaseProp.
diff --git a/theories/Numbers/Integer/Abstract/ZBits.v b/theories/Numbers/Integer/Abstract/ZBits.v
new file mode 100644
index 000000000..44cd08e67
--- /dev/null
+++ b/theories/Numbers/Integer/Abstract/ZBits.v
@@ -0,0 +1,1908 @@
+(************************************************************************)
+(* 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 ZAxioms ZMulOrder ZPow ZDivFloor ZSgnAbs ZParity NZLog.
+
+(** Derived properties of bitwise operations *)
+
+Module Type ZBitsProp
+ (Import A : ZAxiomsSig')
+ (Import B : ZMulOrderProp A)
+ (Import C : ZParityProp A B)
+ (Import D : ZSgnAbsProp A B)
+ (Import E : ZPowProp A B C D)
+ (Import F : ZDivProp A B D)
+ (Import G : NZLog2Prop A A A B E).
+
+Include BoolEqualityFacts A.
+
+Ltac order_nz := try apply pow_nonzero; order'.
+Ltac order_pos' := try apply abs_nonneg; order_pos.
+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 -> 0<=c<=b -> a^(b-c) == a^b / a^c.
+Proof.
+ intros a b c Ha (H,H'). rewrite <- (sub_simpl_r b c) at 2.
+ rewrite pow_add_r; trivial.
+ rewrite div_mul. reflexivity.
+ now apply pow_nonzero.
+ now apply le_0_sub.
+Qed.
+
+Lemma pow_div_l : forall a b c, b~=0 -> 0<=c -> a mod b == 0 ->
+ (a/b)^c == a^c / b^c.
+Proof.
+ intros a b c Hb Hc H. rewrite (div_mod a b Hb) at 2.
+ rewrite H, add_0_r, pow_mul_l, mul_comm, div_mul. reflexivity.
+ now apply pow_nonzero.
+Qed.
+
+(** An injection from bits [true] and [false] to numbers 1 and 0.
+ We declare it as a (local) coercion for shorter statements. *)
+
+Definition b2z (b:bool) := if b then 1 else 0.
+Local Coercion b2z : bool >-> t.
+
+(** Alternative caracterisations of [testbit] *)
+
+Lemma testbit_spec' : forall a n, 0<=n -> a.[n] == (a / 2^n) mod 2.
+Proof.
+ intros a n Hn.
+ destruct (testbit_spec a n Hn) as (l & h & H & EQ). fold (b2z a.[n]) in EQ.
+ apply mod_unique with h. left. destruct a.[n]; split; simpl; order'.
+ symmetry. apply div_unique with l. now left.
+ now rewrite add_comm, mul_comm, (add_comm (2*h)).
+Qed.
+
+Lemma testbit_true : forall a n, 0<=n ->
+ (a.[n] = true <-> (a / 2^n) mod 2 == 1).
+Proof.
+ intros a n Hn.
+ rewrite <- testbit_spec' by trivial.
+ destruct a.[n]; split; simpl; now try order'.
+Qed.
+
+Lemma testbit_false : forall a n, 0<=n ->
+ (a.[n] = false <-> (a / 2^n) mod 2 == 0).
+Proof.
+ intros a n Hn.
+ rewrite <- testbit_spec' by trivial.
+ destruct a.[n]; split; simpl; now try order'.
+Qed.
+
+Lemma testbit_eqb : forall a n, 0<=n ->
+ a.[n] = eqb ((a / 2^n) mod 2) 1.
+Proof.
+ intros a n Hn.
+ apply eq_true_iff_eq. now rewrite testbit_true, eqb_eq.
+Qed.
+
+(** testbit is hence a morphism *)
+
+Instance testbit_wd : Proper (eq==>eq==>Logic.eq) testbit.
+Proof.
+ intros a a' Ha n n' Hn.
+ destruct (le_gt_cases 0 n), (le_gt_cases 0 n'); try order.
+ now rewrite 2 testbit_eqb, Ha, Hn by trivial.
+ now rewrite 2 testbit_neg_r by trivial.
+Qed.
+
+(** Results about the injection [b2z] *)
+
+Lemma b2z_inj : forall (a0 b0:bool), a0 == b0 -> a0 = b0.
+Proof.
+ intros [|] [|]; simpl; trivial; order'.
+Qed.
+
+Lemma add_b2z_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; split; simpl; order').
+Qed.
+
+Lemma add_b2z_double_bit0 : forall (a0:bool) a, (a0+2*a).[0] = a0.
+Proof.
+ intros a0 a. apply b2z_inj.
+ rewrite testbit_spec' by order.
+ nzsimpl. rewrite mul_comm, mod_add by order'.
+ now rewrite mod_small by (destruct a0; split; simpl; order').
+Qed.
+
+Lemma b2z_div2 : forall (a0:bool), a0/2 == 0.
+Proof.
+ intros a0. rewrite <- (add_b2z_double_div2 a0 0). now nzsimpl.
+Qed.
+
+Lemma b2z_bit0 : forall (a0:bool), a0.[0] = a0.
+Proof.
+ intros a0. rewrite <- (add_b2z_double_bit0 a0 0) at 2. now nzsimpl.
+Qed.
+
+(** The initial specification of testbit is complete *)
+
+Lemma testbit_unique : forall a n (a0:bool) l h,
+ 0<=l<2^n -> a == l + (a0 + 2*h)*2^n -> a.[n] = a0.
+Proof.
+ intros a n a0 l h Hl EQ.
+ assert (0<=n).
+ destruct (le_gt_cases 0 n) as [Hn|Hn]; trivial.
+ rewrite pow_neg_r in Hl by trivial. destruct Hl; order.
+ apply b2z_inj. rewrite testbit_spec' by trivial.
+ symmetry. apply mod_unique with h.
+ left; destruct a0; simpl; split; order'.
+ symmetry. apply div_unique with l.
+ now left.
+ 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.
+ destruct (le_gt_cases 0 n).
+ apply testbit_false; trivial. nzsimpl; order_nz.
+ now apply testbit_neg_r.
+Qed.
+
+(** For negative numbers, we are actually doing two's complement *)
+
+Lemma bits_opp : forall a n, 0<=n -> (-a).[n] = negb (P a).[n].
+Proof.
+ intros a n Hn.
+ destruct (testbit_spec (-a) n Hn) as (l & h & Hl & EQ).
+ fold (b2z (-a).[n]) in EQ.
+ apply negb_sym.
+ apply testbit_unique with (2^n-l-1) (-h-1).
+ split.
+ apply lt_succ_r. rewrite sub_1_r, succ_pred. now apply lt_0_sub.
+ apply le_succ_l. rewrite sub_1_r, succ_pred. apply le_sub_le_add_r.
+ rewrite <- (add_0_r (2^n)) at 1. now apply add_le_mono_l.
+ rewrite <- add_sub_swap, sub_1_r. apply pred_wd.
+ apply opp_inj. rewrite opp_add_distr, opp_sub_distr.
+ rewrite (add_comm _ l), <- add_assoc.
+ rewrite EQ at 1. apply add_cancel_l.
+ rewrite <- opp_add_distr.
+ rewrite <- (mul_1_l (2^n)) at 2. rewrite <- mul_add_distr_r.
+ rewrite <- mul_opp_l.
+ apply mul_wd; try reflexivity.
+ rewrite !opp_add_distr.
+ rewrite <- mul_opp_r.
+ rewrite opp_sub_distr, opp_involutive.
+ rewrite (add_comm h).
+ rewrite mul_add_distr_l.
+ rewrite !add_assoc.
+ apply add_cancel_r.
+ rewrite mul_1_r.
+ rewrite add_comm, add_assoc, !add_opp_r, sub_1_r, two_succ, pred_succ.
+ destruct (-a).[n]; simpl. now rewrite sub_0_r. now nzsimpl'.
+Qed.
+
+(** All bits of number (-1) are 1 *)
+
+Lemma bits_m1 : forall n, 0<=n -> (-1).[n] = true.
+Proof.
+ intros. now rewrite bits_opp, one_succ, pred_succ, bits_0.
+Qed.
+
+(** Various ways to refer to the lowest bit of a number *)
+
+Lemma bit0_mod : forall a, a.[0] == a mod 2.
+Proof.
+ intros a. rewrite testbit_spec' by order. now nzsimpl.
+Qed.
+
+Lemma bit0_eqb : forall a, a.[0] = eqb (a mod 2) 1.
+Proof.
+ intros a. rewrite testbit_eqb by order. now nzsimpl.
+Qed.
+
+Lemma bit0_odd : forall a, a.[0] = odd a.
+Proof.
+ intros. rewrite bit0_eqb by order.
+ apply eq_true_iff_eq. rewrite eqb_eq, odd_spec. split.
+ intros H. exists (a/2). rewrite <- H. apply div_mod. order'.
+ intros (b,H).
+ rewrite H, add_comm, mul_comm, mod_add, mod_small; try split; order'.
+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 of positive numbers *)
+
+Lemma bit_log2 : forall a, 0<a -> a.[log2 a] = true.
+Proof.
+ intros a Ha.
+ assert (Ha' := log2_nonneg a).
+ destruct (log2_spec_alt a Ha) as (r & EQ & Hr).
+ rewrite EQ at 1.
+ rewrite testbit_true, add_comm by trivial.
+ rewrite <- (mul_1_l (2^log2 a)) at 1.
+ rewrite div_add by order_nz.
+ rewrite div_small; trivial.
+ rewrite add_0_l. apply mod_small. split; order'.
+Qed.
+
+Lemma bits_above_log2 : forall a n, 0<=a -> log2 a < n ->
+ a.[n] = false.
+Proof.
+ intros a n Ha H.
+ assert (Hn : 0<=n).
+ transitivity (log2 a). apply log2_nonneg. order'.
+ rewrite testbit_false by trivial.
+ rewrite div_small. nzsimpl; order'.
+ split. order. apply log2_lt_cancel. now rewrite log2_pow2.
+Qed.
+
+(** Hence the number of bits of [a] is [1+log2 a]
+ (see [Psize] and [Psize_pos]).
+*)
+
+(** For negative numbers, things are the other ways around:
+ log2 gives the highest zero bit (for numbers below -1).
+*)
+
+Lemma bit_log2_neg : forall a, a < -1 -> a.[log2 (P (-a))] = false.
+Proof.
+ intros a Ha.
+ rewrite <- (opp_involutive a) at 1.
+ rewrite bits_opp.
+ apply negb_false_iff.
+ apply bit_log2.
+ apply opp_lt_mono in Ha. rewrite opp_involutive in Ha.
+ apply lt_succ_lt_pred. now rewrite <- one_succ.
+ apply log2_nonneg.
+Qed.
+
+Lemma bits_above_log2_neg : forall a n, a < 0 -> log2 (P (-a)) < n ->
+ a.[n] = true.
+Proof.
+ intros a n Ha H.
+ assert (Hn : 0<=n).
+ transitivity (log2 (P (-a))). apply log2_nonneg. order'.
+ rewrite <- (opp_involutive a), bits_opp, negb_true_iff by trivial.
+ apply bits_above_log2; trivial.
+ now rewrite <- opp_succ, opp_nonneg_nonpos, le_succ_l.
+Qed.
+
+(** Accesing a high enough bit of a number gives its sign *)
+
+Lemma bits_iff_nonneg : forall a n, log2 (abs a) < n ->
+ (0<=a <-> a.[n] = false).
+Proof.
+ intros a n Hn. split; intros H.
+ rewrite abs_eq in Hn; trivial. now apply bits_above_log2.
+ destruct (le_gt_cases 0 a); trivial.
+ rewrite abs_neq in Hn by order.
+ rewrite bits_above_log2_neg in H; try easy.
+ apply le_lt_trans with (log2 (-a)); trivial.
+ apply log2_le_mono. apply le_pred_l.
+Qed.
+
+Lemma bits_iff_nonneg' : forall a,
+ 0<=a <-> a.[S (log2 (abs a))] = false.
+Proof.
+ intros. apply bits_iff_nonneg. apply lt_succ_diag_r.
+Qed.
+
+Lemma bits_iff_nonneg_ex : forall a,
+ 0<=a <-> (exists k, forall m, k<m -> a.[m] = false).
+Proof.
+ intros a. split.
+ intros Ha. exists (log2 a). intros m Hm. now apply bits_above_log2.
+ intros (k,Hk). destruct (le_gt_cases k (log2 (abs a))).
+ now apply bits_iff_nonneg', Hk, lt_succ_r.
+ apply (bits_iff_nonneg a (S k)).
+ now apply lt_succ_r, lt_le_incl.
+ apply Hk. apply lt_succ_diag_r.
+Qed.
+
+Lemma bits_iff_neg : forall a n, log2 (abs a) < n ->
+ (a<0 <-> a.[n] = true).
+Proof.
+ intros a n Hn.
+ now rewrite lt_nge, <- not_false_iff_true, (bits_iff_nonneg a n).
+Qed.
+
+Lemma bits_iff_neg' : forall a, a<0 <-> a.[S (log2 (abs a))] = true.
+Proof.
+ intros. apply bits_iff_neg. apply lt_succ_diag_r.
+Qed.
+
+Lemma bits_iff_neg_ex : forall a,
+ a<0 <-> (exists k, forall m, k<m -> a.[m] = true).
+Proof.
+ intros a. split.
+ intros Ha. exists (log2 (P (-a))). intros m Hm. now apply bits_above_log2_neg.
+ intros (k,Hk). destruct (le_gt_cases k (log2 (abs a))).
+ now apply bits_iff_neg', Hk, lt_succ_r.
+ apply (bits_iff_neg a (S k)).
+ now apply lt_succ_r, lt_le_incl.
+ apply Hk. apply lt_succ_diag_r.
+Qed.
+
+(** Testing bits after division or multiplication by a power of two *)
+
+Lemma div2_bits : forall a n, 0<=n -> (a/2).[n] = a.[S n].
+Proof.
+ intros a n Hn.
+ apply eq_true_iff_eq. rewrite 2 testbit_true by order_pos.
+ rewrite pow_succ_r by trivial.
+ now rewrite div_div by order_pos.
+Qed.
+
+Lemma div_pow2_bits : forall a n m, 0<=n -> 0<=m -> (a/2^n).[m] = a.[m+n].
+Proof.
+ intros a n m Hn. revert a m. apply le_ind with (4:=Hn).
+ solve_predicate_wd.
+ intros a m Hm. now nzsimpl.
+ clear n Hn. intros n Hn IH a m Hm. nzsimpl; trivial.
+ rewrite <- div_div by order_pos.
+ now rewrite IH, div2_bits by order_pos.
+Qed.
+
+Lemma double_bits_succ : forall a n, (2*a).[S n] = a.[n].
+Proof.
+ intros a n.
+ destruct (le_gt_cases 0 n) as [Hn|Hn].
+ now rewrite <- div2_bits, mul_comm, div_mul by order'.
+ rewrite (testbit_neg_r a n Hn).
+ apply le_succ_l, le_lteq in Hn. destruct Hn as [Hn|Hn].
+ now rewrite testbit_neg_r.
+ now rewrite Hn, bit0_odd, odd_mul, odd_2.
+Qed.
+
+Lemma double_bits : forall a n, (2*a).[n] = a.[P n].
+Proof.
+ intros a n. rewrite <- (succ_pred n) at 1. apply double_bits_succ.
+Qed.
+
+Lemma mul_pow2_bits_add : forall a n m, 0<=n -> (a*2^n).[n+m] = a.[m].
+Proof.
+ intros a n m Hn. revert a m. apply le_ind with (4:=Hn).
+ solve_predicate_wd.
+ intros a m. now nzsimpl.
+ clear n Hn. intros n Hn IH a m. nzsimpl; trivial.
+ rewrite mul_assoc, (mul_comm _ 2), <- mul_assoc.
+ now rewrite double_bits_succ.
+Qed.
+
+Lemma mul_pow2_bits : forall a n m, 0<=n -> (a*2^n).[m] = a.[m-n].
+Proof.
+ intros.
+ rewrite <- (add_simpl_r m n) at 1. rewrite add_sub_swap, add_comm.
+ now apply mul_pow2_bits_add.
+Qed.
+
+Lemma mul_pow2_bits_low : forall a n m, m<n -> (a*2^n).[m] = false.
+Proof.
+ intros.
+ destruct (le_gt_cases 0 n).
+ rewrite mul_pow2_bits by trivial.
+ apply testbit_neg_r. now apply lt_sub_0.
+ now rewrite pow_neg_r, mul_0_r, bits_0.
+Qed.
+
+(** Selecting the low part of a number can be done by a modulo *)
+
+Lemma mod_pow2_bits_high : forall a n m, 0<=n<=m ->
+ (a mod 2^n).[m] = false.
+Proof.
+ intros a n m (Hn,H).
+ destruct (mod_pos_bound a (2^n)) as [LE LT]. order_pos.
+ apply le_lteq in LE; destruct LE as [LT'|EQ].
+ apply bits_above_log2; try order.
+ apply lt_le_trans with n; trivial.
+ apply log2_lt_pow2; trivial.
+ now rewrite <-EQ, bits_0.
+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.
+ destruct (le_gt_cases 0 m) as [Hm|Hm]; [|now rewrite !testbit_neg_r].
+ rewrite testbit_eqb; trivial.
+ 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.
+ rewrite mul_comm, mul_assoc, <- pow_add_r, (add_comm m), sub_add; trivial.
+ rewrite add_comm, <- div_mod by order_nz.
+ symmetry. apply testbit_eqb; trivial.
+ apply le_0_sub; order.
+ now apply lt_le_pred, lt_0_sub.
+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 (lt_trichotomy a 0) as [Ha|[Ha|Ha]]; trivial.
+ apply (bits_above_log2_neg a (S (log2 (P (-a))))) in Ha.
+ now rewrite H in Ha.
+ apply lt_succ_diag_r.
+ apply bit_log2 in Ha. now rewrite H in Ha.
+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.
+ assert (AUX : forall n, 0<=n -> forall a b,
+ 0<=a<2^n -> testbit a === testbit b -> a == b).
+ intros n Hn. apply le_ind with (4:=Hn).
+ solve_predicate_wd.
+ intros a b Ha H. rewrite pow_0_r, one_succ, lt_succ_r in Ha.
+ assert (Ha' : a == 0) by (destruct Ha; order).
+ rewrite Ha' in *.
+ symmetry. apply bits_inj_0.
+ intros m. now rewrite <- H, bits_0.
+ clear n Hn. intros n Hn IH a b (Ha,Ha') H.
+ rewrite (div_mod a 2), (div_mod b 2) by order'.
+ apply add_wd; [ | now rewrite <- 2 bit0_mod, H].
+ apply mul_wd. reflexivity.
+ apply IH.
+ split. apply div_pos; order'.
+ apply div_lt_upper_bound. order'. now rewrite <- pow_succ_r.
+ intros m.
+ destruct (le_gt_cases 0 m).
+ rewrite 2 div2_bits by trivial. apply H.
+ now rewrite 2 testbit_neg_r.
+ intros a b H.
+ destruct (le_gt_cases 0 a) as [Ha|Ha].
+ apply (AUX a); trivial. split; trivial.
+ apply pow_gt_lin_r; order'.
+ apply succ_inj, opp_inj.
+ assert (0 <= - S a).
+ apply opp_le_mono. now rewrite opp_involutive, opp_0, le_succ_l.
+ apply (AUX (-(S a))); trivial. split; trivial.
+ apply pow_gt_lin_r; order'.
+ intros m. destruct (le_gt_cases 0 m).
+ now rewrite 2 bits_opp, 2 pred_succ, H.
+ now rewrite 2 testbit_neg_r.
+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.
+
+(** In fact, checking the bits at positive indexes is enough. *)
+
+Lemma bits_inj' : forall a b,
+ (forall n, 0<=n -> a.[n] = b.[n]) -> a == b.
+Proof.
+ intros a b H. apply bits_inj.
+ intros n. destruct (le_gt_cases 0 n).
+ now apply H.
+ now rewrite 2 testbit_neg_r.
+Qed.
+
+Lemma bits_inj_iff' : forall a b, (forall n, 0<=n -> a.[n] = b.[n]) <-> a == b.
+Proof.
+ split. apply bits_inj'. intros EQ n Hn; now rewrite EQ.
+Qed.
+
+Ltac bitwise := apply bits_inj'; intros ?m ?Hm; autorewrite with bitwise.
+
+Hint Rewrite lxor_spec lor_spec land_spec ldiff_spec bits_0 : bitwise.
+
+(** The streams of bits that correspond to a numbers are
+ exactly the ones which are stationary after some point. *)
+
+Lemma are_bits : forall (f:t->bool), Proper (eq==>Logic.eq) f ->
+ ((exists n, forall m, 0<=m -> f m = n.[m]) <->
+ (exists k, forall m, k<=m -> f m = f k)).
+Proof.
+ intros f Hf. split.
+ intros (a,H).
+ destruct (le_gt_cases 0 a).
+ exists (S (log2 a)). intros m Hm. apply le_succ_l in Hm.
+ rewrite 2 H, 2 bits_above_log2; trivial using lt_succ_diag_r.
+ order_pos. apply le_trans with (log2 a); order_pos.
+ exists (S (log2 (P (-a)))). intros m Hm. apply le_succ_l in Hm.
+ rewrite 2 H, 2 bits_above_log2_neg; trivial using lt_succ_diag_r.
+ order_pos. apply le_trans with (log2 (P (-a))); order_pos.
+ intros (k,Hk).
+ destruct (lt_ge_cases k 0) as [LT|LE].
+ case_eq (f 0); intros H0.
+ exists (-1). intros m Hm. rewrite bits_m1, Hk by order.
+ symmetry; rewrite <- H0. apply Hk; order.
+ exists 0. intros m Hm. rewrite bits_0, Hk by order.
+ symmetry; rewrite <- H0. apply Hk; order.
+ revert f Hf Hk. apply le_ind with (4:=LE).
+ (* compat : solve_predicat_wd fails here *)
+ apply proper_sym_impl_iff. exact eq_sym.
+ clear k LE. intros k k' Hk IH f Hf H. apply IH; trivial.
+ now setoid_rewrite Hk.
+ (* /compat *)
+ intros f Hf H0. destruct (f 0).
+ exists (-1). intros m Hm. now rewrite bits_m1, H0.
+ exists 0. intros m Hm. now rewrite bits_0, H0.
+ clear k LE. intros k LE IH f Hf Hk.
+ destruct (IH (fun m => f (S m))) as (n, Hn).
+ solve_predicate_wd.
+ intros m Hm. apply Hk. now rewrite <- succ_le_mono.
+ exists (f 0 + 2*n). intros m Hm.
+ apply le_lteq in Hm. destruct Hm as [Hm|Hm].
+ rewrite <- (succ_pred m), Hn, <- div2_bits.
+ rewrite mul_comm, div_add, b2z_div2, add_0_l; trivial. order'.
+ now rewrite <- lt_succ_r, succ_pred.
+ now rewrite <- lt_succ_r, succ_pred.
+ rewrite <- Hm.
+ symmetry. apply add_b2z_double_bit0.
+Qed.
+
+(** * Properties of shifts *)
+
+(** First, a unified specification for [shiftl] : the [shiftl_spec]
+ below (combined with [testbit_neg_r]) is equivalent to
+ [shiftl_spec_low] and [shiftl_spec_high]. *)
+
+Lemma shiftl_spec : forall a n m, 0<=m -> (a << n).[m] = a.[m-n].
+Proof.
+ intros.
+ destruct (le_gt_cases n m).
+ now apply shiftl_spec_high.
+ rewrite shiftl_spec_low, testbit_neg_r; trivial. now apply lt_sub_0.
+Qed.
+
+(** A shiftl by a negative number is a shiftr, and vice-versa *)
+
+Lemma shiftr_opp_r : forall a n, a >> (-n) == a << n.
+Proof.
+ intros. bitwise. now rewrite shiftr_spec, shiftl_spec, add_opp_r.
+Qed.
+
+Lemma shiftl_opp_r : forall a n, a << (-n) == a >> n.
+Proof.
+ intros. bitwise. now rewrite shiftr_spec, shiftl_spec, sub_opp_r.
+Qed.
+
+(** Shifts correspond to multiplication or division by a power of two *)
+
+Lemma shiftr_div_pow2 : forall a n, 0<=n -> a >> n == a / 2^n.
+Proof.
+ intros. bitwise. now rewrite shiftr_spec, div_pow2_bits.
+Qed.
+
+Lemma shiftr_mul_pow2 : forall a n, n<=0 -> a >> n == a * 2^(-n).
+Proof.
+ intros. bitwise. rewrite shiftr_spec, mul_pow2_bits; trivial.
+ now rewrite sub_opp_r.
+ now apply opp_nonneg_nonpos.
+Qed.
+
+Lemma shiftl_mul_pow2 : forall a n, 0<=n -> a << n == a * 2^n.
+Proof.
+ intros. bitwise. now rewrite shiftl_spec, mul_pow2_bits.
+Qed.
+
+Lemma shiftl_div_pow2 : forall a n, n<=0 -> a << n == a / 2^(-n).
+Proof.
+ intros. bitwise. rewrite shiftl_spec, div_pow2_bits; trivial.
+ now rewrite add_opp_r.
+ now apply opp_nonneg_nonpos.
+Qed.
+
+(** Shifts are morphisms *)
+
+Instance shiftr_wd : Proper (eq==>eq==>eq) shiftr.
+Proof.
+ intros a a' Ha n n' Hn.
+ destruct (le_ge_cases n 0) as [H|H]; assert (H':=H); rewrite Hn in H'.
+ now rewrite 2 shiftr_mul_pow2, Ha, Hn.
+ now rewrite 2 shiftr_div_pow2, Ha, Hn.
+Qed.
+
+Instance shiftl_wd : Proper (eq==>eq==>eq) shiftl.
+Proof.
+ intros a a' Ha n n' Hn. now rewrite <- 2 shiftr_opp_r, Ha, Hn.
+Qed.
+
+(** We could also have specified shiftl with an addition on the left. *)
+
+Lemma shiftl_spec_alt : forall a n m, 0<=n -> (a << n).[m+n] = a.[m].
+Proof.
+ intros. now rewrite shiftl_mul_pow2, mul_pow2_bits, add_simpl_r.
+Qed.
+
+(** Chaining several shifts. The only case for which
+ there isn't any simple expression is a true shiftr
+ followed by a true shiftl.
+*)
+
+Lemma shiftl_shiftl : forall a n m, 0<=n ->
+ (a << n) << m == a << (n+m).
+Proof.
+ intros a n p Hn. bitwise.
+ rewrite 2 (shiftl_spec _ _ m) by trivial.
+ rewrite add_comm, sub_add_distr.
+ destruct (le_gt_cases 0 (m-p)) as [H|H].
+ now rewrite shiftl_spec.
+ rewrite 2 testbit_neg_r; trivial.
+ apply lt_sub_0. now apply lt_le_trans with 0.
+Qed.
+
+Lemma shiftr_shiftl_l : forall a n m, 0<=n ->
+ (a << n) >> m == a << (n-m).
+Proof.
+ intros. now rewrite <- shiftl_opp_r, shiftl_shiftl, add_opp_r.
+Qed.
+
+Lemma shiftr_shiftl_r : forall a n m, 0<=n ->
+ (a << n) >> m == a >> (m-n).
+Proof.
+ intros. now rewrite <- 2 shiftl_opp_r, shiftl_shiftl, opp_sub_distr, add_comm.
+Qed.
+
+Lemma shiftr_shiftr : forall a n m, 0<=n -> 0<=m ->
+ (a >> n) >> m == a >> (n+m).
+Proof.
+ intros a n m Hn Hm.
+ now rewrite !shiftr_div_pow2, pow_add_r, div_div by order_pos.
+Qed.
+
+(** shifts and constants *)
+
+Lemma shiftl_1_l : forall n, 1 << n == 2^n.
+Proof.
+ intros n. destruct (le_gt_cases 0 n).
+ now rewrite shiftl_mul_pow2, mul_1_l.
+ rewrite shiftl_div_pow2, div_1_l, pow_neg_r; try order.
+ apply pow_gt_1. order'. now apply opp_pos_neg.
+Qed.
+
+Lemma shiftl_0_r : forall a, a << 0 == a.
+Proof.
+ intros. rewrite shiftl_mul_pow2 by order. now nzsimpl.
+Qed.
+
+Lemma shiftr_0_r : forall a, a >> 0 == a.
+Proof.
+ intros. now rewrite <- shiftl_opp_r, opp_0, shiftl_0_r.
+Qed.
+
+Lemma shiftl_0_l : forall n, 0 << n == 0.
+Proof.
+ intros.
+ destruct (le_ge_cases 0 n).
+ rewrite shiftl_mul_pow2 by trivial. now nzsimpl.
+ rewrite shiftl_div_pow2 by trivial.
+ rewrite <- opp_nonneg_nonpos in H. nzsimpl; order_nz.
+Qed.
+
+Lemma shiftr_0_l : forall n, 0 >> n == 0.
+Proof.
+ intros. now rewrite <- shiftl_opp_r, shiftl_0_l.
+Qed.
+
+Lemma shiftl_eq_0_iff : forall a n, 0<=n -> (a << n == 0 <-> a == 0).
+Proof.
+ intros a n Hn.
+ rewrite shiftl_mul_pow2 by trivial. 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.
+ destruct (le_gt_cases 0 n) as [Hn|Hn].
+ rewrite shiftr_div_pow2, div_small_iff by order_nz.
+ destruct (lt_trichotomy a 0) as [LT|[EQ|LT]].
+ split.
+ intros [(H,_)|(H,H')]. order. generalize (pow_nonneg 2 n le_0_2); order.
+ intros [H|(H,H')]; order.
+ rewrite EQ. split. now left. intros _; left. split; order_pos.
+ split. intros [(H,H')|(H,H')]; right. split; trivial.
+ apply log2_lt_pow2; trivial.
+ generalize (pow_nonneg 2 n le_0_2); order.
+ intros [H|(H,H')]. order. left.
+ split. order. now apply log2_lt_pow2.
+ rewrite shiftr_mul_pow2 by order. rewrite eq_mul_0.
+ split; intros [H|H].
+ now left.
+ elim (pow_nonzero 2 (-n)); try apply opp_nonneg_nonpos; order'.
+ now left.
+ destruct H. generalize (log2_nonneg a); order.
+Qed.
+
+Lemma shiftr_eq_0 : forall a n, 0<=a -> log2 a < n -> a >> n == 0.
+Proof.
+ intros a n Ha H.
+ apply shiftr_eq_0_iff.
+ apply le_lteq in Ha. destruct Ha as [Ha|Ha].
+ right. now split. now left.
+Qed.
+
+(** Properties of [div2]. *)
+
+Lemma div2_div : forall a, div2 a == a/2.
+Proof.
+ intros. rewrite div2_spec, shiftr_div_pow2. now nzsimpl. order'.
+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.
+ intros a a' Ha n n' Hn. unfold setbit. now rewrite Ha, Hn.
+Qed.
+
+Instance clearbit_wd : Proper (eq==>eq==>eq) clearbit.
+Proof.
+ intros a a' Ha n n' Hn. unfold clearbit. now rewrite Ha, Hn.
+Qed.
+
+Lemma pow2_bits_true : forall n, 0<=n -> (2^n).[n] = true.
+Proof.
+ intros. rewrite <- (mul_1_l (2^n)).
+ now rewrite mul_pow2_bits, sub_diag, bit0_odd, odd_1.
+Qed.
+
+Lemma pow2_bits_false : forall n m, n~=m -> (2^n).[m] = false.
+Proof.
+ intros.
+ destruct (le_gt_cases 0 n); [|now rewrite pow_neg_r, bits_0].
+ destruct (le_gt_cases n m).
+ rewrite <- (mul_1_l (2^n)), mul_pow2_bits; trivial.
+ rewrite <- (succ_pred (m-n)), <- div2_bits.
+ now rewrite div_small, bits_0 by (split; order').
+ rewrite <- lt_succ_r, succ_pred, lt_0_sub. order.
+ rewrite <- (mul_1_l (2^n)), mul_pow2_bits_low; trivial.
+Qed.
+
+Lemma pow2_bits_eqb : forall n m, 0<=n -> (2^n).[m] = eqb n m.
+Proof.
+ intros n m Hn. 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; order.
+Qed.
+
+Lemma setbit_eqb : forall a n m, 0<=n ->
+ (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, 0<=n ->
+ ((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, 0<=n -> (setbit a n).[n] = true.
+Proof.
+ intros. apply setbit_iff; trivial. now left.
+Qed.
+
+Lemma setbit_neq : forall a n m, 0<=n -> n~=m ->
+ (setbit a n).[m] = a.[m].
+Proof.
+ intros a n m Hn H. rewrite setbit_eqb; trivial.
+ 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.
+ destruct (le_gt_cases 0 m); [| now rewrite 2 testbit_neg_r].
+ rewrite clearbit_spec', ldiff_spec. f_equal. f_equal.
+ destruct (le_gt_cases 0 n) as [Hn|Hn].
+ now apply pow2_bits_eqb.
+ symmetry. rewrite pow_neg_r, bits_0, <- not_true_iff_false, eqb_eq; order.
+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. now rewrite !shiftl_spec, lxor_spec.
+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. now rewrite !shiftl_spec, land_spec.
+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. now rewrite !shiftl_spec, lor_spec.
+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. now rewrite !shiftl_spec, ldiff_spec.
+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.
+
+(** For integers, we do have a binary complement function *)
+
+Definition lnot a := P (-a).
+
+Instance lnot_wd : Proper (eq==>eq) lnot.
+Proof. intros a a' Ha; unfold lnot; now rewrite Ha. Qed.
+
+Lemma lnot_spec : forall a n, 0<=n -> (lnot a).[n] = negb a.[n].
+Proof.
+ intros. unfold lnot. rewrite <- (opp_involutive a) at 2.
+ rewrite bits_opp, negb_involutive; trivial.
+Qed.
+
+Lemma lnot_involutive : forall a, lnot (lnot a) == a.
+Proof.
+ intros a. bitwise. now rewrite 2 lnot_spec, negb_involutive.
+Qed.
+
+Lemma lnot_0 : lnot 0 == -1.
+Proof.
+ unfold lnot. now rewrite opp_0, <- sub_1_r, sub_0_l.
+Qed.
+
+Lemma lnot_m1 : lnot (-1) == 0.
+Proof.
+ unfold lnot. now rewrite opp_involutive, one_succ, pred_succ.
+Qed.
+
+(** Complement and other operations *)
+
+Lemma lor_m1_r : forall a, lor a (-1) == -1.
+Proof.
+ intros. bitwise. now rewrite bits_m1, orb_true_r.
+Qed.
+
+Lemma lor_m1_l : forall a, lor (-1) a == -1.
+Proof.
+ intros. now rewrite lor_comm, lor_m1_r.
+Qed.
+
+Lemma land_m1_r : forall a, land a (-1) == a.
+Proof.
+ intros. bitwise. now rewrite bits_m1, andb_true_r.
+Qed.
+
+Lemma land_m1_l : forall a, land (-1) a == a.
+Proof.
+ intros. now rewrite land_comm, land_m1_r.
+Qed.
+
+Lemma ldiff_m1_r : forall a, ldiff a (-1) == 0.
+Proof.
+ intros. bitwise. now rewrite bits_m1, andb_false_r.
+Qed.
+
+Lemma ldiff_m1_l : forall a, ldiff (-1) a == lnot a.
+Proof.
+ intros. bitwise. now rewrite lnot_spec, bits_m1.
+Qed.
+
+Lemma lor_lnot_diag : forall a, lor a (lnot a) == -1.
+Proof.
+ intros a. bitwise. rewrite lnot_spec, bits_m1; trivial.
+ now destruct a.[m].
+Qed.
+
+Lemma add_lnot_diag : forall a, a + lnot a == -1.
+Proof.
+ intros a. unfold lnot.
+ now rewrite add_pred_r, add_opp_r, sub_diag, one_succ, opp_succ, opp_0.
+Qed.
+
+Lemma ldiff_land : forall a b, ldiff a b == land a (lnot b).
+Proof.
+ intros. bitwise. now rewrite lnot_spec.
+Qed.
+
+Lemma land_lnot_diag : forall a, land a (lnot a) == 0.
+Proof.
+ intros. now rewrite <- ldiff_land, ldiff_diag.
+Qed.
+
+Lemma lnot_lor : forall a b, lnot (lor a b) == land (lnot a) (lnot b).
+Proof.
+ intros a b. bitwise. now rewrite !lnot_spec, lor_spec, negb_orb.
+Qed.
+
+Lemma lnot_land : forall a b, lnot (land a b) == lor (lnot a) (lnot b).
+Proof.
+ intros a b. bitwise. now rewrite !lnot_spec, land_spec, negb_andb.
+Qed.
+
+Lemma lnot_ldiff : forall a b, lnot (ldiff a b) == lor (lnot a) b.
+Proof.
+ intros a b. bitwise.
+ now rewrite !lnot_spec, ldiff_spec, negb_andb, negb_involutive.
+Qed.
+
+Lemma lxor_lnot_lnot : forall a b, lxor (lnot a) (lnot b) == lxor a b.
+Proof.
+ intros a b. bitwise. now rewrite !lnot_spec, xorb_negb_negb.
+Qed.
+
+Lemma lnot_lxor_l : forall a b, lnot (lxor a b) == lxor (lnot a) b.
+Proof.
+ intros a b. bitwise. now rewrite !lnot_spec, !lxor_spec, negb_xorb_l.
+Qed.
+
+Lemma lnot_lxor_r : forall a b, lnot (lxor a b) == lxor a (lnot b).
+Proof.
+ intros a b. bitwise. now rewrite !lnot_spec, !lxor_spec, negb_xorb_r.
+Qed.
+
+Lemma lxor_m1_r : forall a, lxor a (-1) == lnot a.
+Proof.
+ intros. now rewrite <- (lxor_0_r (lnot a)), <- lnot_m1, lxor_lnot_lnot.
+Qed.
+
+Lemma lxor_m1_l : forall a, lxor (-1) a == lnot a.
+Proof.
+ intros. now rewrite lxor_comm, lxor_m1_r.
+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.
+
+Lemma lnot_shiftr : forall a n, 0<=n -> lnot (a >> n) == (lnot a) >> n.
+Proof.
+ intros a n Hn. bitwise.
+ now rewrite lnot_spec, 2 shiftr_spec, lnot_spec by order_pos.
+Qed.
+
+(** [(ones n)] is [2^n-1], the number with [n] digits 1 *)
+
+Definition ones n := P (1<<n).
+
+Instance ones_wd : Proper (eq==>eq) ones.
+Proof. intros a a' Ha; unfold ones; now rewrite Ha. Qed.
+
+Lemma ones_equiv : forall n, ones n == P (2^n).
+Proof.
+ intros. unfold ones.
+ destruct (le_gt_cases 0 n).
+ now rewrite shiftl_mul_pow2, mul_1_l.
+ apply pred_wd. rewrite pow_neg_r; trivial.
+ rewrite <- shiftr_opp_r. apply shiftr_eq_0_iff. right; split.
+ order'. rewrite log2_1. now apply opp_pos_neg.
+Qed.
+
+Lemma ones_add : forall n m, 0<=n -> 0<=m ->
+ ones (m+n) == 2^m * ones n + ones m.
+Proof.
+ intros n m Hn Hm. rewrite !ones_equiv.
+ rewrite <- !sub_1_r, mul_sub_distr_l, mul_1_r, <- pow_add_r by trivial.
+ rewrite add_sub_assoc, sub_add. reflexivity.
+Qed.
+
+Lemma ones_div_pow2 : forall n m, 0<=m<=n -> ones n / 2^m == ones (n-m).
+Proof.
+ intros n m (Hm,H). symmetry. apply div_unique with (ones m).
+ left. rewrite ones_equiv. split.
+ rewrite <- lt_succ_r, succ_pred. order_pos.
+ now rewrite <- le_succ_l, succ_pred.
+ rewrite <- (sub_add m n) at 1. rewrite (add_comm _ m).
+ apply ones_add; trivial. now apply le_0_sub.
+Qed.
+
+Lemma ones_mod_pow2 : forall n m, 0<=m<=n -> (ones n) mod (2^m) == ones m.
+Proof.
+ intros n m (Hm,H). symmetry. apply mod_unique with (ones (n-m)).
+ left. rewrite ones_equiv. split.
+ rewrite <- lt_succ_r, succ_pred. order_pos.
+ now rewrite <- le_succ_l, succ_pred.
+ rewrite <- (sub_add m n) at 1. rewrite (add_comm _ m).
+ apply ones_add; trivial. now apply le_0_sub.
+Qed.
+
+Lemma ones_spec_low : forall n m, 0<=m<n -> (ones n).[m] = true.
+Proof.
+ intros n m (Hm,H). apply testbit_true; trivial.
+ rewrite ones_div_pow2 by (split; order).
+ rewrite <- (pow_1_r 2). rewrite ones_mod_pow2.
+ rewrite ones_equiv. now nzsimpl'.
+ split. order'. apply le_add_le_sub_r. nzsimpl. now apply le_succ_l.
+Qed.
+
+Lemma ones_spec_high : forall n m, 0<=n<=m -> (ones n).[m] = false.
+Proof.
+ intros n m (Hn,H). apply le_lteq in Hn. destruct Hn as [Hn|Hn].
+ apply bits_above_log2; rewrite ones_equiv.
+ rewrite <-lt_succ_r, succ_pred; order_pos.
+ rewrite log2_pred_pow2; trivial. now rewrite <-le_succ_l, succ_pred.
+ rewrite ones_equiv. now rewrite <- Hn, pow_0_r, one_succ, pred_succ, bits_0.
+Qed.
+
+Lemma ones_spec_iff : forall n m, 0<=n ->
+ ((ones n).[m] = true <-> 0<=m<n).
+Proof.
+ intros n m Hn. split. intros H.
+ destruct (lt_ge_cases m 0) as [Hm|Hm].
+ now rewrite testbit_neg_r in H.
+ split; trivial. apply lt_nge. intro H'. rewrite ones_spec_high in H.
+ discriminate. now split.
+ apply ones_spec_low.
+Qed.
+
+Lemma lor_ones_low : forall a n, 0<=a -> log2 a < n ->
+ lor a (ones n) == ones n.
+Proof.
+ intros a n Ha H. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, bits_above_log2; try split; trivial.
+ now apply lt_le_trans with n.
+ apply le_trans with (log2 a); order_pos.
+ rewrite ones_spec_low, orb_true_r; try split; trivial.
+Qed.
+
+Lemma land_ones : forall a n, 0<=n -> land a (ones n) == a mod 2^n.
+Proof.
+ intros a n Hn. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, mod_pow2_bits_high, andb_false_r;
+ try split; trivial.
+ rewrite ones_spec_low, mod_pow2_bits_low, andb_true_r;
+ try split; trivial.
+Qed.
+
+Lemma land_ones_low : forall a n, 0<=a -> log2 a < n ->
+ land a (ones n) == a.
+Proof.
+ intros a n Ha H.
+ assert (Hn : 0<=n) by (generalize (log2_nonneg a); order).
+ rewrite land_ones; trivial. apply mod_small.
+ split; trivial.
+ apply log2_lt_cancel. now rewrite log2_pow2.
+Qed.
+
+Lemma ldiff_ones_r : forall a n, 0<=n ->
+ ldiff a (ones n) == (a >> n) << n.
+Proof.
+ intros a n Hn. 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 apply le_0_sub.
+ now split.
+ rewrite ones_spec_low, shiftl_spec_low, andb_false_r;
+ try split; trivial.
+Qed.
+
+Lemma ldiff_ones_r_low : forall a n, 0<=a -> log2 a < n ->
+ ldiff a (ones n) == 0.
+Proof.
+ intros a n Ha H. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ split; trivial. now apply le_trans with (log2 a); order_pos.
+ rewrite ones_spec_low, andb_false_r; try split; trivial.
+Qed.
+
+Lemma ldiff_ones_l_low : forall a n, 0<=a -> log2 a < n ->
+ ldiff (ones n) a == lxor a (ones n).
+Proof.
+ intros a n Ha H. bitwise. destruct (le_gt_cases n m).
+ rewrite ones_spec_high, bits_above_log2; trivial.
+ now apply lt_le_trans with n.
+ split; trivial. now apply le_trans with (log2 a); order_pos.
+ rewrite ones_spec_low, xorb_true_r; try split; trivial.
+Qed.
+
+(** Bitwise operations and sign *)
+
+Lemma shiftl_nonneg : forall a n, 0 <= (a << n) <-> 0 <= a.
+Proof.
+ intros a n.
+ destruct (le_ge_cases 0 n) as [Hn|Hn].
+ (* 0<=n *)
+ rewrite 2 bits_iff_nonneg_ex. split; intros (k,Hk).
+ exists (k-n). intros m Hm.
+ destruct (le_gt_cases 0 m); [|now rewrite testbit_neg_r].
+ rewrite <- (add_simpl_r m n), <- (shiftl_spec a n) by order_pos.
+ apply Hk. now apply lt_sub_lt_add_r.
+ exists (k+n). intros m Hm.
+ destruct (le_gt_cases 0 m); [|now rewrite testbit_neg_r].
+ rewrite shiftl_spec by trivial. apply Hk. now apply lt_add_lt_sub_r.
+ (* n<=0*)
+ rewrite <- shiftr_opp_r, 2 bits_iff_nonneg_ex. split; intros (k,Hk).
+ destruct (le_gt_cases 0 k).
+ exists (k-n). intros m Hm. apply lt_sub_lt_add_r in Hm.
+ rewrite <- (add_simpl_r m n), <- add_opp_r, <- (shiftr_spec a (-n)).
+ now apply Hk. order.
+ assert (EQ : a >> (-n) == 0).
+ apply bits_inj'. intros m Hm. rewrite bits_0. apply Hk; order.
+ apply shiftr_eq_0_iff in EQ.
+ rewrite <- bits_iff_nonneg_ex. destruct EQ as [EQ|[LT _]]; order.
+ exists (k+n). intros m Hm.
+ destruct (le_gt_cases 0 m); [|now rewrite testbit_neg_r].
+ rewrite shiftr_spec by trivial. apply Hk.
+ rewrite add_opp_r. now apply lt_add_lt_sub_r.
+Qed.
+
+Lemma shiftl_neg : forall a n, (a << n) < 0 <-> a < 0.
+Proof.
+ intros a n. now rewrite 2 lt_nge, shiftl_nonneg.
+Qed.
+
+Lemma shiftr_nonneg : forall a n, 0 <= (a >> n) <-> 0 <= a.
+Proof.
+ intros. rewrite <- shiftl_opp_r. apply shiftl_nonneg.
+Qed.
+
+Lemma shiftr_neg : forall a n, (a >> n) < 0 <-> a < 0.
+Proof.
+ intros a n. now rewrite 2 lt_nge, shiftr_nonneg.
+Qed.
+
+Lemma div2_nonneg : forall a, 0 <= div2 a <-> 0 <= a.
+Proof.
+ intros. rewrite div2_spec. apply shiftr_nonneg.
+Qed.
+
+Lemma div2_neg : forall a, div2 a < 0 <-> a < 0.
+Proof.
+ intros a. now rewrite 2 lt_nge, div2_nonneg.
+Qed.
+
+Lemma lor_nonneg : forall a b, 0 <= lor a b <-> 0<=a /\ 0<=b.
+Proof.
+ intros a b.
+ rewrite 3 bits_iff_nonneg_ex. split. intros (k,Hk).
+ split; exists k; intros m Hm; apply (orb_false_elim a.[m] b.[m]);
+ rewrite <- lor_spec; now apply Hk.
+ intros ((k,Hk),(k',Hk')).
+ destruct (le_ge_cases k k'); [ exists k' | exists k ];
+ intros m Hm; rewrite lor_spec, Hk, Hk'; trivial; order.
+Qed.
+
+Lemma lor_neg : forall a b, lor a b < 0 <-> a < 0 \/ b < 0.
+Proof.
+ intros a b. rewrite 3 lt_nge, lor_nonneg. split.
+ apply not_and. apply le_decidable.
+ now intros [H|H] (H',H'').
+Qed.
+
+Lemma lnot_nonneg : forall a, 0 <= lnot a <-> a < 0.
+Proof.
+ intros a; unfold lnot.
+ now rewrite <- opp_succ, opp_nonneg_nonpos, le_succ_l.
+Qed.
+
+Lemma lnot_neg : forall a, lnot a < 0 <-> 0 <= a.
+Proof.
+ intros a. now rewrite le_ngt, lt_nge, lnot_nonneg.
+Qed.
+
+Lemma land_nonneg : forall a b, 0 <= land a b <-> 0<=a \/ 0<=b.
+Proof.
+ intros a b.
+ now rewrite <- (lnot_involutive (land a b)), lnot_land, lnot_nonneg,
+ lor_neg, !lnot_neg.
+Qed.
+
+Lemma land_neg : forall a b, land a b < 0 <-> a < 0 /\ b < 0.
+Proof.
+ intros a b.
+ now rewrite <- (lnot_involutive (land a b)), lnot_land, lnot_neg,
+ lor_nonneg, !lnot_nonneg.
+Qed.
+
+Lemma ldiff_nonneg : forall a b, 0 <= ldiff a b <-> 0<=a \/ b<0.
+Proof.
+ intros. now rewrite ldiff_land, land_nonneg, lnot_nonneg.
+Qed.
+
+Lemma ldiff_neg : forall a b, ldiff a b < 0 <-> a<0 /\ 0<=b.
+Proof.
+ intros. now rewrite ldiff_land, land_neg, lnot_neg.
+Qed.
+
+Lemma lxor_nonneg : forall a b, 0 <= lxor a b <-> (0<=a <-> 0<=b).
+Proof.
+ assert (H : forall a b, 0<=a -> 0<=b -> 0<=lxor a b).
+ intros a b. rewrite 3 bits_iff_nonneg_ex. intros (k,Hk) (k', Hk').
+ destruct (le_ge_cases k k'); [ exists k' | exists k];
+ intros m Hm; rewrite lxor_spec, Hk, Hk'; trivial; order.
+ assert (H' : forall a b, 0<=a -> b<0 -> lxor a b<0).
+ intros a b. rewrite bits_iff_nonneg_ex, 2 bits_iff_neg_ex.
+ intros (k,Hk) (k', Hk').
+ destruct (le_ge_cases k k'); [ exists k' | exists k];
+ intros m Hm; rewrite lxor_spec, Hk, Hk'; trivial; order.
+ intros a b.
+ split.
+ intros Hab. split.
+ intros Ha. destruct (le_gt_cases 0 b) as [Hb|Hb]; trivial.
+ generalize (H' _ _ Ha Hb). order.
+ intros Hb. destruct (le_gt_cases 0 a) as [Ha|Ha]; trivial.
+ generalize (H' _ _ Hb Ha). rewrite lxor_comm. order.
+ intros E.
+ destruct (le_gt_cases 0 a) as [Ha|Ha]. apply H; trivial. apply E; trivial.
+ destruct (le_gt_cases 0 b) as [Hb|Hb]. apply H; trivial. apply E; trivial.
+ rewrite <- lxor_lnot_lnot. apply H; now apply lnot_nonneg.
+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 (lt_trichotomy a 0) as [Ha|[Ha|Ha]].
+ (* a < 0 *)
+ destruct (proj1 (bits_iff_neg_ex a) Ha) as (k,Hk).
+ destruct (le_gt_cases n k).
+ specialize (Hk (S k) (lt_succ_diag_r _)).
+ rewrite H' in Hk. discriminate. apply lt_succ_r; order.
+ specialize (H' (S n) (lt_succ_diag_r _)).
+ rewrite Hk in H'. discriminate. apply lt_succ_r; order.
+ (* a = 0 *)
+ now rewrite Ha, bits_0 in H.
+ (* 0 < a *)
+ apply le_antisymm; apply le_ngt; intros LT.
+ specialize (H' _ LT). now rewrite bit_log2 in H'.
+ now rewrite bits_above_log2 in H by order.
+Qed.
+
+Lemma log2_shiftr : forall a n, 0<a -> log2 (a >> n) == max 0 (log2 a - n).
+Proof.
+ intros a n Ha.
+ destruct (le_gt_cases 0 (log2 a - n));
+ [rewrite max_r | rewrite max_l]; try order.
+ apply log2_bits_unique.
+ now rewrite shiftr_spec, sub_add, bit_log2.
+ intros m Hm.
+ destruct (le_gt_cases 0 m); [|now rewrite testbit_neg_r].
+ rewrite shiftr_spec; trivial. apply bits_above_log2; try order.
+ now apply lt_sub_lt_add_r.
+ rewrite lt_sub_lt_add_r, add_0_l in H.
+ apply log2_nonpos. apply le_lteq; right.
+ apply shiftr_eq_0_iff. right. now split.
+Qed.
+
+Lemma log2_shiftl : forall a n, 0<a -> 0<=n -> log2 (a << n) == log2 a + n.
+Proof.
+ intros a n Ha Hn.
+ rewrite shiftl_mul_pow2, add_comm by trivial.
+ now apply log2_mul_pow2.
+Qed.
+
+Lemma log2_shiftl' : forall a n, 0<a -> log2 (a << n) == max 0 (log2 a + n).
+Proof.
+ intros a n Ha.
+ rewrite <- shiftr_opp_r, log2_shiftr by trivial.
+ destruct (le_gt_cases 0 (log2 a + n));
+ [rewrite 2 max_r | rewrite 2 max_l]; rewrite ?sub_opp_r; try order.
+Qed.
+
+Lemma log2_lor : forall a b, 0<=a -> 0<=b ->
+ log2 (lor a b) == max (log2 a) (log2 b).
+Proof.
+ assert (AUX : forall a b, 0<=a -> a<=b -> log2 (lor a b) == log2 b).
+ intros a b Ha H.
+ apply le_lteq in Ha; destruct Ha 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 Ha Hb. 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, 0<=a -> 0<=b ->
+ log2 (land a b) <= min (log2 a) (log2 b).
+Proof.
+ assert (AUX : forall a b, 0<=a -> a<=b -> log2 (land a b) <= log2 a).
+ intros a b Ha Hb.
+ apply le_ngt. intros H'.
+ assert (LE : 0 <= land a b) by (apply land_nonneg; now left).
+ apply le_lteq in LE. destruct LE as [LT|EQ].
+ generalize (bit_log2 (land a b) LT).
+ now rewrite land_spec, bits_above_log2.
+ rewrite <- EQ in H'. apply log2_lt_cancel in H'; order.
+ (* main *)
+ intros a b Ha Hb.
+ 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, 0<=a -> 0<=b ->
+ log2 (lxor a b) <= max (log2 a) (log2 b).
+Proof.
+ assert (AUX : forall a b, 0<=a -> a<=b -> log2 (lxor a b) <= log2 b).
+ intros a b Ha Hb.
+ apply le_ngt. intros H'.
+ assert (LE : 0 <= lxor a b) by (apply lxor_nonneg; split; order).
+ apply le_lteq in LE. destruct LE as [LT|EQ].
+ generalize (bit_log2 (lxor a b) LT).
+ rewrite lxor_spec, 2 bits_above_log2; try order. discriminate.
+ apply le_lt_trans with (log2 b); trivial. now apply log2_le_mono.
+ rewrite <- EQ in H'. apply log2_lt_cancel in H'; order.
+ (* main *)
+ intros a b Ha Hb.
+ 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; split; order') || idtac.
+ symmetry. apply div_unique with 1. left; split; 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'.
+ apply div_wd; try easy.
+ rewrite <- !div2_div, mul_comm, mul_add_distr_l.
+ rewrite (div2_odd a), <- bit0_odd at 1.
+ rewrite (div2_odd b), <- bit0_odd at 1.
+ 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_aux : forall n, 0<=n ->
+ forall a b (c0:bool), -(2^n) <= a < 2^n -> -(2^n) <= b < 2^n ->
+ exists c,
+ a+b+c0 == lxor3 a b c /\ c/2 == lnextcarry a b c /\ c.[0] = c0.
+Proof.
+ intros n Hn. apply le_ind with (4:=Hn).
+ solve_predicate_wd.
+ (* base *)
+ intros a b c0. rewrite !pow_0_r, !one_succ, !lt_succ_r, <- !one_succ.
+ intros (Ha1,Ha2) (Hb1,Hb2).
+ apply le_lteq in Ha1; apply le_lteq in Hb1.
+ rewrite <- le_succ_l, succ_m1 in Ha1, Hb1.
+ destruct Ha1 as [Ha1|Ha1]; destruct Hb1 as [Hb1|Hb1].
+ (* base, a = 0, b = 0 *)
+ exists c0.
+ rewrite (le_antisymm _ _ Ha2 Ha1), (le_antisymm _ _ Hb2 Hb1).
+ rewrite !add_0_l, !lxor_0_l, !lor_0_r, !land_0_r, !lor_0_r.
+ rewrite b2z_div2, b2z_bit0; now repeat split.
+ (* base, a = 0, b = -1 *)
+ exists (-c0). rewrite <- Hb1, (le_antisymm _ _ Ha2 Ha1). repeat split.
+ rewrite add_0_l, lxor_0_l, lxor_m1_l.
+ unfold lnot. now rewrite opp_involutive, add_comm, add_opp_r, sub_1_r.
+ rewrite land_0_l, !lor_0_l, land_m1_r.
+ symmetry. apply div_unique with c0. left; destruct c0; simpl; split; order'.
+ now rewrite two_succ, mul_succ_l, mul_1_l, add_opp_r, sub_add.
+ rewrite bit0_odd, odd_opp; destruct c0; simpl; apply odd_1 || apply odd_0.
+ (* base, a = -1, b = 0 *)
+ exists (-c0). rewrite <- Ha1, (le_antisymm _ _ Hb2 Hb1). repeat split.
+ rewrite add_0_r, lxor_0_r, lxor_m1_l.
+ unfold lnot. now rewrite opp_involutive, add_comm, add_opp_r, sub_1_r.
+ rewrite land_0_r, lor_0_r, lor_0_l, land_m1_r.
+ symmetry. apply div_unique with c0. left; destruct c0; simpl; split; order'.
+ now rewrite two_succ, mul_succ_l, mul_1_l, add_opp_r, sub_add.
+ rewrite bit0_odd, odd_opp; destruct c0; simpl; apply odd_1 || apply odd_0.
+ (* base, a = -1, b = -1 *)
+ exists (c0 + 2*(-1)). rewrite <- Ha1, <- Hb1. repeat split.
+ rewrite lxor_m1_l, lnot_m1, lxor_0_l.
+ now rewrite two_succ, mul_succ_l, mul_1_l, add_comm, add_assoc.
+ rewrite land_m1_l, lor_m1_l.
+ apply add_b2z_double_div2.
+ apply add_b2z_double_bit0.
+ (* step *)
+ clear n Hn. intros n Hn 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.
+ split.
+ apply div_le_lower_bound. order'. now rewrite mul_opp_r, <- pow_succ_r.
+ apply div_lt_upper_bound. order'. now rewrite <- pow_succ_r.
+ split.
+ apply div_le_lower_bound. order'. now rewrite mul_opp_r, <- pow_succ_r.
+ apply div_lt_upper_bound. order'. now rewrite <- pow_succ_r.
+ exists (c0 + 2*c). repeat split.
+ (* step, add *)
+ bitwise.
+ apply le_lteq in Hm. destruct Hm as [Hm|Hm].
+ rewrite <- (succ_pred m), lt_succ_r in Hm.
+ rewrite <- (succ_pred m), <- !div2_bits, <- 2 lxor_spec by trivial.
+ apply testbit_wd; try easy.
+ rewrite add_b2z_double_div2, <- IH1. apply add_carry_div2.
+ rewrite <- Hm.
+ now rewrite add_b2z_double_bit0, add3_bit0, b2z_bit0.
+ (* step, carry *)
+ rewrite add_b2z_double_div2.
+ bitwise.
+ apply le_lteq in Hm. destruct Hm as [Hm|Hm].
+ rewrite <- (succ_pred m), lt_succ_r in Hm.
+ rewrite <- (succ_pred m), <- !div2_bits, IH2 by trivial.
+ autorewrite with bitwise. now rewrite add_b2z_double_div2.
+ rewrite <- Hm.
+ now rewrite add_b2z_double_bit0.
+ (* step, carry0 *)
+ apply add_b2z_double_bit0.
+Qed.
+
+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.
+ set (n := max (abs a) (abs b)).
+ apply (add_carry_bits_aux n).
+ (* positivity *)
+ unfold n.
+ destruct (le_ge_cases (abs a) (abs b));
+ [rewrite max_r|rewrite max_l]; order_pos'.
+ (* bound for a *)
+ assert (Ha : abs a < 2^n).
+ apply lt_le_trans with (2^(abs a)). apply pow_gt_lin_r; order_pos'.
+ apply pow_le_mono_r. order'. unfold n.
+ destruct (le_ge_cases (abs a) (abs b));
+ [rewrite max_r|rewrite max_l]; try order.
+ apply abs_lt in Ha. destruct Ha; split; order.
+ (* bound for b *)
+ assert (Hb : abs b < 2^n).
+ apply lt_le_trans with (2^(abs b)). apply pow_gt_lin_r; order_pos'.
+ apply pow_le_mono_r. order'. unfold n.
+ destruct (le_ge_cases (abs a) (abs b));
+ [rewrite max_r|rewrite max_l]; try order.
+ apply abs_lt in Hb. destruct Hb; split; order.
+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 by order.
+ 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'. intros n Hn. rewrite bits_0.
+ apply le_ind with (4:=Hn).
+ solve_predicate_wd.
+ trivial.
+ clear n Hn. intros n Hn IH.
+ rewrite <- div2_bits, H; trivial.
+ 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, 0<=b -> ldiff a b == 0 -> 0 <= a <= b.
+Proof.
+ assert (AUX : forall n, 0<=n ->
+ forall a b, 0 <= a < 2^n -> 0<=b -> ldiff a b == 0 -> a <= b).
+ intros n Hn. apply le_ind with (4:=Hn); clear n Hn.
+ solve_predicate_wd.
+ intros a b Ha Hb _. rewrite pow_0_r, one_succ, lt_succ_r in Ha.
+ setoid_replace a with 0 by (destruct Ha; order'); trivial.
+ intros n Hn IH a b (Ha,Ha') Hb 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_pos_l; try order'.
+ apply IH.
+ split. apply div_pos; order'.
+ apply div_lt_upper_bound; try order'. now rewrite <- pow_succ_r.
+ apply div_pos; order'.
+ rewrite <- (pow_1_r 2), <- 2 shiftr_div_pow2 by order'.
+ rewrite <- shiftr_ldiff, H, shiftr_div_pow2, pow_1_r, div_0_l; order'.
+ 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'.
+ (* main *)
+ intros a b Hb Hd.
+ assert (Ha : 0<=a).
+ apply le_ngt; intros Ha'. apply (lt_irrefl 0). rewrite <- Hd at 1.
+ apply ldiff_neg. now split.
+ split; trivial. apply (AUX a); try split; trivial. apply pow_gt_lin_r; 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; trivial.
+ 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.
+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.
+ destruct (le_gt_cases a 0) as [Ha'|Ha'].
+ apply le_lt_trans with (0+b). now apply add_le_mono_r. now nzsimpl.
+ destruct (le_gt_cases b 0) as [Hb'|Hb'].
+ apply le_lt_trans with (a+0). now apply add_le_mono_l. now nzsimpl.
+ rewrite add_nocarry_lxor by order.
+ destruct (lt_ge_cases 0 (lxor a b)); [|apply le_lt_trans with 0; order_pos].
+ apply log2_lt_pow2; trivial.
+ apply log2_lt_pow2 in Ha; trivial.
+ apply log2_lt_pow2 in Hb; trivial.
+ apply le_lt_trans with (max (log2 a) (log2 b)).
+ apply log2_lxor; order.
+ destruct (le_ge_cases (log2 a) (log2 b));
+ [rewrite max_r|rewrite max_l]; order.
+Qed.
+
+Lemma add_nocarry_mod_lt_pow2 : forall a b n, 0<=n -> land a b == 0 ->
+ a mod 2^n + b mod 2^n < 2^n.
+Proof.
+ intros a b n Hn H.
+ apply add_nocarry_lt_pow2.
+ bitwise.
+ destruct (le_gt_cases n m).
+ rewrite mod_pow2_bits_high; now split.
+ now rewrite !mod_pow2_bits_low, <- land_spec, H, bits_0.
+ apply mod_pos_bound; order_pos.
+ apply mod_pos_bound; order_pos.
+Qed.
+
+End ZBitsProp.
diff --git a/theories/Numbers/Integer/Abstract/ZDivEucl.v b/theories/Numbers/Integer/Abstract/ZDivEucl.v
index 070003972..c8fd29a54 100644
--- a/theories/Numbers/Integer/Abstract/ZDivEucl.v
+++ b/theories/Numbers/Integer/Abstract/ZDivEucl.v
@@ -575,6 +575,23 @@ Proof.
apply div_mod; order.
Qed.
+(** Similarly, the following result doesn't always hold for negative
+ [b] and [c]. For instance [3 mod (-2*-2)) = 3] while
+ [3 mod (-2) + (-2)*((3/-2) mod -2) = -1].
+*)
+
+Lemma mod_mul_r : forall a b c, 0<b -> 0<c ->
+ a mod (b*c) == a mod b + b*((a/b) mod c).
+Proof.
+ intros a b c Hb Hc.
+ apply add_cancel_l with (b*c*(a/(b*c))).
+ rewrite <- div_mod by (apply neq_mul_0; split; order).
+ rewrite <- div_div by trivial.
+ rewrite add_assoc, add_shuffle0, <- mul_assoc, <- mul_add_distr_l.
+ rewrite <- div_mod by order.
+ apply div_mod; order.
+Qed.
+
(** A last inequality: *)
Theorem div_mul_le:
diff --git a/theories/Numbers/Integer/Abstract/ZDivFloor.v b/theories/Numbers/Integer/Abstract/ZDivFloor.v
index 5f52f046f..017f995cc 100644
--- a/theories/Numbers/Integer/Abstract/ZDivFloor.v
+++ b/theories/Numbers/Integer/Abstract/ZDivFloor.v
@@ -618,6 +618,23 @@ Proof.
apply div_mod; order.
Qed.
+(** Similarly, the following result doesn't always hold for negative
+ [b] and [c]. For instance [3 mod (-2*-2)) = 3] while
+ [3 mod (-2) + (-2)*((3/-2) mod -2) = -1].
+*)
+
+Lemma rem_mul_r : forall a b c, 0<b -> 0<c ->
+ a mod (b*c) == a mod b + b*((a/b) mod c).
+Proof.
+ intros a b c Hb Hc.
+ apply add_cancel_l with (b*c*(a/(b*c))).
+ rewrite <- div_mod by (apply neq_mul_0; split; order).
+ rewrite <- div_div by trivial.
+ rewrite add_assoc, add_shuffle0, <- mul_assoc, <- mul_add_distr_l.
+ rewrite <- div_mod by order.
+ apply div_mod; order.
+Qed.
+
(** A last inequality: *)
Theorem div_mul_le:
diff --git a/theories/Numbers/Integer/Abstract/ZDivTrunc.v b/theories/Numbers/Integer/Abstract/ZDivTrunc.v
index e33265762..09f9e023e 100644
--- a/theories/Numbers/Integer/Abstract/ZDivTrunc.v
+++ b/theories/Numbers/Integer/Abstract/ZDivTrunc.v
@@ -584,8 +584,7 @@ destruct (le_0_mul _ _ Hab) as [(Ha,Hb)|(Ha,Hb)];
setoid_replace b with 0 by order. rewrite rem_0_l by order. nzsimpl; order.
Qed.
-
-(** Conversely, the following result needs less restrictions here. *)
+(** Conversely, the following results need less restrictions here. *)
Lemma quot_quot : forall a b c, b~=0 -> c~=0 ->
(a÷b)÷c == a÷(b*c).
@@ -605,6 +604,18 @@ apply opp_inj. rewrite <- 3 quot_opp_l; try order. apply Aux2; order.
rewrite <- neq_mul_0. tauto.
Qed.
+Lemma mod_mul_r : forall a b c, b~=0 -> c~=0 ->
+ a rem (b*c) == a rem b + b*((a÷b) rem c).
+Proof.
+ intros a b c Hb Hc.
+ apply add_cancel_l with (b*c*(a÷(b*c))).
+ rewrite <- quot_rem by (apply neq_mul_0; split; order).
+ rewrite <- quot_quot by trivial.
+ rewrite add_assoc, add_shuffle0, <- mul_assoc, <- mul_add_distr_l.
+ rewrite <- quot_rem by order.
+ apply quot_rem; order.
+Qed.
+
(** A last inequality: *)
Theorem quot_mul_le:
diff --git a/theories/Numbers/Integer/Abstract/ZGcd.v b/theories/Numbers/Integer/Abstract/ZGcd.v
index 6b1d4675e..8e128215d 100644
--- a/theories/Numbers/Integer/Abstract/ZGcd.v
+++ b/theories/Numbers/Integer/Abstract/ZGcd.v
@@ -51,7 +51,7 @@ Proof.
now apply divide_abs_l.
Qed.
-Lemma divide_1_r : forall n, (n | 1) -> n==1 \/ n==-(1).
+Lemma divide_1_r : forall n, (n | 1) -> n==1 \/ n==-1.
Proof.
intros n (m,Hm). now apply eq_mul_1 with m.
Qed.
diff --git a/theories/Numbers/Integer/Abstract/ZLcm.v b/theories/Numbers/Integer/Abstract/ZLcm.v
index 27bd5962c..bbf7d893a 100644
--- a/theories/Numbers/Integer/Abstract/ZLcm.v
+++ b/theories/Numbers/Integer/Abstract/ZLcm.v
@@ -43,6 +43,43 @@ Proof.
apply quot_rem. order.
Qed.
+(** We can use the sign rule to have an relation between divisions. *)
+
+Lemma quot_div : forall a b, b~=0 ->
+ a÷b == (sgn a)*(sgn b)*(abs a / abs b).
+Proof.
+ assert (AUX : forall a b, 0<b -> a÷b == (sgn a)*(sgn b)*(abs a / abs b)).
+ intros a b Hb. rewrite (sgn_pos b), (abs_eq b), mul_1_r by order.
+ destruct (lt_trichotomy 0 a) as [Ha|[Ha|Ha]].
+ rewrite sgn_pos, abs_eq, mul_1_l, quot_div_nonneg; order.
+ rewrite <- Ha, abs_0, sgn_0, quot_0_l, div_0_l, mul_0_l; order.
+ rewrite sgn_neg, abs_neq, mul_opp_l, mul_1_l, eq_opp_r, <-quot_opp_l
+ by order.
+ apply quot_div_nonneg; trivial. apply opp_nonneg_nonpos; order.
+ (* main *)
+ intros a b Hb.
+ apply neg_pos_cases in Hb. destruct Hb as [Hb|Hb]; [|now apply AUX].
+ rewrite <- (opp_involutive b) at 1. rewrite quot_opp_r.
+ rewrite AUX, abs_opp, sgn_opp, mul_opp_r, mul_opp_l, opp_involutive.
+ reflexivity.
+ now apply opp_pos_neg.
+ rewrite eq_opp_l, opp_0; order.
+Qed.
+
+Lemma rem_mod : forall a b, b~=0 ->
+ a rem b == (sgn a) * ((abs a) mod (abs b)).
+Proof.
+ intros a b Hb.
+ rewrite <- rem_abs_r by trivial.
+ assert (Hb' := proj2 (abs_pos b) Hb).
+ destruct (lt_trichotomy 0 a) as [Ha|[Ha|Ha]].
+ rewrite (abs_eq a), sgn_pos, mul_1_l, rem_mod_nonneg; order.
+ rewrite <- Ha, abs_0, sgn_0, mod_0_l, rem_0_l, mul_0_l; order.
+ rewrite sgn_neg, (abs_neq a), mul_opp_l, mul_1_l, eq_opp_r, <-rem_opp_l
+ by order.
+ apply rem_mod_nonneg; trivial. apply opp_nonneg_nonpos; order.
+Qed.
+
(** Modulo and remainder are null at the same place,
and this correspond to the divisibility relation. *)
diff --git a/theories/Numbers/Integer/Abstract/ZLt.v b/theories/Numbers/Integer/Abstract/ZLt.v
index 23eac0e69..5431b4a10 100644
--- a/theories/Numbers/Integer/Abstract/ZLt.v
+++ b/theories/Numbers/Integer/Abstract/ZLt.v
@@ -121,10 +121,10 @@ Proof.
intro; apply lt_neq; apply lt_pred_l.
Qed.
-Theorem lt_n1_r : forall n m, n < m -> m < 0 -> n < -(1).
+Theorem lt_m1_r : forall n m, n < m -> m < 0 -> n < -1.
Proof.
intros n m H1 H2. apply -> lt_le_pred in H2.
-setoid_replace (P 0) with (-(1)) in H2. now apply lt_le_trans with m.
+setoid_replace (P 0) with (-1) in H2. now apply lt_le_trans with m.
apply <- eq_opp_r. now rewrite one_succ, opp_pred, opp_0.
Qed.
diff --git a/theories/Numbers/Integer/Abstract/ZMulOrder.v b/theories/Numbers/Integer/Abstract/ZMulOrder.v
index 1010a0f2f..4c0a9a2ca 100644
--- a/theories/Numbers/Integer/Abstract/ZMulOrder.v
+++ b/theories/Numbers/Integer/Abstract/ZMulOrder.v
@@ -13,8 +13,6 @@ Require Export ZAddOrder.
Module Type ZMulOrderProp (Import Z : ZAxiomsMiniSig').
Include ZAddOrderProp Z.
-Local Notation "- 1" := (-(1)).
-
Theorem mul_lt_mono_nonpos :
forall n m p q, m <= 0 -> n < m -> q <= 0 -> p < q -> m * q < n * p.
Proof.
@@ -135,18 +133,18 @@ now apply lt_1_l with (- m).
assumption.
Qed.
-Theorem lt_mul_n1_neg : forall n m, 1 < n -> m < 0 -> n * m < -1.
+Theorem lt_mul_m1_neg : forall n m, 1 < n -> m < 0 -> n * m < -1.
Proof.
intros n m H1 H2. apply -> (mul_lt_mono_neg_r m) in H1.
-rewrite mul_1_l in H1. now apply lt_n1_r with m.
+rewrite mul_1_l in H1. now apply lt_m1_r with m.
assumption.
Qed.
-Theorem lt_mul_n1_pos : forall n m, n < -1 -> 0 < m -> n * m < -1.
+Theorem lt_mul_m1_pos : forall n m, n < -1 -> 0 < m -> n * m < -1.
Proof.
intros n m H1 H2. apply -> (mul_lt_mono_pos_r m) in H1.
rewrite mul_opp_l, mul_1_l in H1.
-apply <- opp_neg_pos in H2. now apply lt_n1_r with (- m).
+apply <- opp_neg_pos in H2. now apply lt_m1_r with (- m).
assumption.
Qed.
@@ -154,18 +152,18 @@ Theorem lt_1_mul_l : forall n m, 1 < n ->
n * m < -1 \/ n * m == 0 \/ 1 < n * m.
Proof.
intros n m H; destruct (lt_trichotomy m 0) as [H1 | [H1 | H1]].
-left. now apply lt_mul_n1_neg.
+left. now apply lt_mul_m1_neg.
right; left; now rewrite H1, mul_0_r.
right; right; now apply lt_1_mul_pos.
Qed.
-Theorem lt_n1_mul_r : forall n m, n < -1 ->
+Theorem lt_m1_mul_r : forall n m, n < -1 ->
n * m < -1 \/ n * m == 0 \/ 1 < n * m.
Proof.
intros n m H; destruct (lt_trichotomy m 0) as [H1 | [H1 | H1]].
right; right. now apply lt_1_mul_neg.
right; left; now rewrite H1, mul_0_r.
-left. now apply lt_mul_n1_pos.
+left. now apply lt_mul_m1_pos.
Qed.
Theorem eq_mul_1 : forall n m, n * m == 1 -> n == 1 \/ n == -1.
diff --git a/theories/Numbers/Integer/Abstract/ZParity.v b/theories/Numbers/Integer/Abstract/ZParity.v
index 4c752043c..5c7e9eb14 100644
--- a/theories/Numbers/Integer/Abstract/ZParity.v
+++ b/theories/Numbers/Integer/Abstract/ZParity.v
@@ -6,167 +6,23 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-Require Import Bool ZMulOrder.
+Require Import Bool ZMulOrder NZParity.
-(** Properties of [even] and [odd]. *)
-
-(** NB: most parts of [NParity] and [ZParity] are common,
- but it is difficult to share them in NZ, since
- initial proofs [Even_or_Odd] and [Even_Odd_False] must
- be proved differently *)
+(** Some more properties of [even] and [odd]. *)
Module Type ZParityProp (Import Z : ZAxiomsSig')
(Import ZP : ZMulOrderProp Z).
-Instance Even_wd : Proper (eq==>iff) Even.
-Proof. unfold Even. solve_predicate_wd. Qed.
-
-Instance Odd_wd : Proper (eq==>iff) Odd.
-Proof. unfold Odd. solve_predicate_wd. Qed.
-
-Instance even_wd : Proper (eq==>Logic.eq) even.
-Proof.
- intros x x' EQ. rewrite eq_iff_eq_true, 2 even_spec. now apply Even_wd.
-Qed.
-
-Instance odd_wd : Proper (eq==>Logic.eq) odd.
-Proof.
- intros x x' EQ. rewrite eq_iff_eq_true, 2 odd_spec. now apply Odd_wd.
-Qed.
-
-Lemma Even_or_Odd : forall x, Even x \/ Odd x.
-Proof.
- nzinduct x.
- left. exists 0. now nzsimpl.
- intros x.
- split; intros [(y,H)|(y,H)].
- right. exists y. rewrite H. now nzsimpl.
- left. exists (S y). rewrite H. now nzsimpl'.
- right. exists (P y). rewrite <- succ_inj_wd. rewrite H.
- nzsimpl'. now rewrite <- add_succ_l, <- add_succ_r, succ_pred.
- left. exists y. rewrite <- succ_inj_wd. rewrite H. now nzsimpl.
-Qed.
-
-Lemma double_below : forall n m, n<=m -> 2*n < 2*m+1.
-Proof.
- intros. nzsimpl'. apply lt_succ_r. now apply add_le_mono.
-Qed.
-
-Lemma double_above : forall n m, n<m -> 2*n+1 < 2*m.
-Proof.
- intros. nzsimpl'.
- rewrite <- le_succ_l, <- add_succ_l, <- add_succ_r.
- apply add_le_mono; now apply le_succ_l.
-Qed.
-
-Lemma Even_Odd_False : forall x, Even x -> Odd x -> False.
-Proof.
-intros x (y,E) (z,O). rewrite O in E; clear O.
-destruct (le_gt_cases y z) as [LE|GT].
-generalize (double_below _ _ LE); order.
-generalize (double_above _ _ GT); order.
-Qed.
-
-Lemma orb_even_odd : forall n, orb (even n) (odd n) = true.
-Proof.
- intros.
- destruct (Even_or_Odd n) as [H|H].
- rewrite <- even_spec in H. now rewrite H.
- rewrite <- odd_spec in H. now rewrite H, orb_true_r.
-Qed.
-
-Lemma negb_odd_even : forall n, negb (odd n) = even n.
-Proof.
- intros.
- generalize (Even_or_Odd n) (Even_Odd_False n).
- rewrite <- even_spec, <- odd_spec.
- destruct (odd n), (even n); simpl; intuition.
-Qed.
-
-Lemma negb_even_odd : forall n, negb (even n) = odd n.
-Proof.
- intros. rewrite <- negb_odd_even. apply negb_involutive.
-Qed.
-
-Lemma even_0 : even 0 = true.
-Proof.
- rewrite even_spec. exists 0. now nzsimpl.
-Qed.
-
-Lemma odd_1 : odd 1 = true.
-Proof.
- rewrite odd_spec. exists 0. now nzsimpl'.
-Qed.
-
-Lemma Odd_succ_Even : forall n, Odd (S n) <-> Even n.
-Proof.
- split; intros (m,H).
- exists m. apply succ_inj. now rewrite add_1_r in H.
- exists m. rewrite add_1_r. now apply succ_wd.
-Qed.
-
-Lemma odd_succ_even : forall n, odd (S n) = even n.
-Proof.
- intros. apply eq_iff_eq_true. rewrite even_spec, odd_spec.
- apply Odd_succ_Even.
-Qed.
-
-Lemma even_succ_odd : forall n, even (S n) = odd n.
-Proof.
- intros. now rewrite <- negb_odd_even, odd_succ_even, negb_even_odd.
-Qed.
-
-Lemma Even_succ_Odd : forall n, Even (S n) <-> Odd n.
-Proof.
- intros. now rewrite <- even_spec, even_succ_odd, odd_spec.
-Qed.
-
-Lemma odd_pred_even : forall n, odd (P n) = even n.
-Proof.
- intros. rewrite <- (succ_pred n) at 2. symmetry. apply even_succ_odd.
-Qed.
-
-Lemma even_pred_odd : forall n, even (P n) = odd n.
-Proof.
- intros. rewrite <- (succ_pred n) at 2. symmetry. apply odd_succ_even.
-Qed.
-
-Lemma even_add : forall n m, 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, ?odd_spec, ?even_spec;
- intros (m',Hm) (n',Hn).
- exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm.
- exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm, add_assoc.
- exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm, add_shuffle0.
- exists (n'+m'+1). rewrite Hm,Hn. nzsimpl'. now rewrite add_shuffle1.
-Qed.
-
-Lemma odd_add : forall n m, odd (n+m) = xorb (odd n) (odd m).
-Proof.
- intros. rewrite <- !negb_even_odd. rewrite even_add.
- now destruct (even n), (even m).
-Qed.
+Include NZParityProp Z Z ZP.
-Lemma even_mul : forall n m, even (mul n m) = even n || even m.
+Lemma odd_pred : forall n, odd (P n) = even n.
Proof.
- intros.
- case_eq (even n); simpl; rewrite ?even_spec.
- intros (n',Hn). exists (n'*m). now rewrite Hn, mul_assoc.
- case_eq (even m); simpl; rewrite ?even_spec.
- intros (m',Hm). exists (n*m'). now rewrite Hm, !mul_assoc, (mul_comm 2).
- (* odd / odd *)
- rewrite <- !negb_true_iff, !negb_even_odd, !odd_spec.
- intros (m',Hm) (n',Hn). exists (n'*2*m' +n'+m').
- rewrite Hn,Hm, !mul_add_distr_l, !mul_add_distr_r, !mul_1_l, !mul_1_r.
- now rewrite add_shuffle1, add_assoc, !mul_assoc.
+ intros. rewrite <- (succ_pred n) at 2. symmetry. apply even_succ.
Qed.
-Lemma odd_mul : forall n m, odd (mul n m) = odd n && odd m.
+Lemma even_pred : forall n, even (P n) = odd n.
Proof.
- intros. rewrite <- !negb_even_odd. rewrite even_mul.
- now destruct (even n), (even m).
+ intros. rewrite <- (succ_pred n) at 2. symmetry. apply odd_succ.
Qed.
Lemma even_opp : forall n, even (-n) = even n.
@@ -180,7 +36,7 @@ Qed.
Lemma odd_opp : forall n, odd (-n) = odd n.
Proof.
- intros. rewrite <- !negb_even_odd. now rewrite even_opp.
+ intros. rewrite <- !negb_even. now rewrite even_opp.
Qed.
Lemma even_sub : forall n m, even (n-m) = Bool.eqb (even n) (even m).
diff --git a/theories/Numbers/Integer/Abstract/ZPow.v b/theories/Numbers/Integer/Abstract/ZPow.v
index 8ea012250..682c680cc 100644
--- a/theories/Numbers/Integer/Abstract/ZPow.v
+++ b/theories/Numbers/Integer/Abstract/ZPow.v
@@ -30,7 +30,7 @@ Qed.
Lemma odd_pow : forall a b, 0<b -> odd (a^b) = odd a.
Proof.
- intros. now rewrite <- !negb_even_odd, even_pow.
+ intros. now rewrite <- !negb_even, even_pow.
Qed.
(** Properties of power of negative numbers *)
@@ -80,8 +80,7 @@ Proof.
rewrite <- EQ'. nzsimpl.
destruct (le_gt_cases 0 b).
apply pow_0_l.
- assert (b~=0) by
- (contradict H; now rewrite H, <-odd_spec, <-negb_even_odd, even_0).
+ assert (b~=0) by (contradict H; now rewrite H, <-odd_spec, odd_0).
order.
now rewrite pow_neg_r.
rewrite abs_neq by order.
@@ -95,8 +94,7 @@ Proof.
destruct (sgn_spec a) as [(LT,EQ)|[(EQ',EQ)|(LT,EQ)]]; rewrite EQ.
apply sgn_pos. apply pow_pos_nonneg; trivial.
rewrite <- EQ'. rewrite pow_0_l. apply sgn_0.
- assert (b~=0) by
- (contradict H; now rewrite H, <-odd_spec, <-negb_even_odd, even_0).
+ assert (b~=0) by (contradict H; now rewrite H, <-odd_spec, odd_0).
order.
apply sgn_neg.
rewrite <- (opp_involutive a). rewrite pow_opp_odd by trivial.
@@ -105,4 +103,22 @@ Proof.
now apply opp_pos_neg.
Qed.
+Lemma abs_pow : forall a b, abs (a^b) == (abs a)^b.
+Proof.
+ intros a b.
+ destruct (Even_or_Odd b).
+ rewrite pow_even_abs by trivial.
+ apply abs_eq, pow_nonneg, abs_nonneg.
+ rewrite pow_odd_abs_sgn by trivial.
+ rewrite abs_mul.
+ destruct (lt_trichotomy 0 a) as [Ha|[Ha|Ha]].
+ rewrite (sgn_pos a), (abs_eq 1), mul_1_l by order'.
+ apply abs_eq, pow_nonneg, abs_nonneg.
+ rewrite <- Ha, sgn_0, abs_0, mul_0_l.
+ symmetry. apply pow_0_l'. intro Hb. rewrite Hb in H.
+ apply (Even_Odd_False 0); trivial. exists 0; now nzsimpl.
+ rewrite (sgn_neg a), abs_opp, (abs_eq 1), mul_1_l by order'.
+ apply abs_eq, pow_nonneg, abs_nonneg.
+Qed.
+
End ZPowProp.
diff --git a/theories/Numbers/Integer/Abstract/ZProperties.v b/theories/Numbers/Integer/Abstract/ZProperties.v
index d04443d9c..c04551960 100644
--- a/theories/Numbers/Integer/Abstract/ZProperties.v
+++ b/theories/Numbers/Integer/Abstract/ZProperties.v
@@ -7,7 +7,7 @@
(************************************************************************)
Require Export ZAxioms ZMaxMin ZSgnAbs ZParity ZPow ZDivTrunc ZDivFloor
- ZGcd ZLcm NZLog NZSqrt.
+ ZGcd ZLcm NZLog NZSqrt ZBits.
(** This functor summarizes all known facts about Z. *)
@@ -15,4 +15,5 @@ Module Type ZProp (Z:ZAxiomsSig) :=
ZMaxMinProp Z <+ ZSgnAbsProp Z <+ ZParityProp Z <+ ZPowProp Z
<+ NZSqrtProp Z Z <+ NZSqrtUpProp Z Z
<+ NZLog2Prop Z Z Z <+ NZLog2UpProp Z Z Z
- <+ ZDivProp Z <+ ZQuotProp Z <+ ZGcdProp Z <+ ZLcmProp Z.
+ <+ ZDivProp Z <+ ZQuotProp Z <+ ZGcdProp Z <+ ZLcmProp Z
+ <+ ZBitsProp Z.
diff --git a/theories/Numbers/Integer/Abstract/ZSgnAbs.v b/theories/Numbers/Integer/Abstract/ZSgnAbs.v
index 6d90bc0fd..b2f6cc84d 100644
--- a/theories/Numbers/Integer/Abstract/ZSgnAbs.v
+++ b/theories/Numbers/Integer/Abstract/ZSgnAbs.v
@@ -37,12 +37,12 @@ Module Type ZDecAxiomsSig' := ZAxiomsMiniSig' <+ HasCompare.
Module Type GenericSgn (Import Z : ZDecAxiomsSig')
(Import ZP : ZMulOrderProp Z) <: HasSgn Z.
Definition sgn n :=
- match compare 0 n with Eq => 0 | Lt => 1 | Gt => -(1) end.
+ match compare 0 n with Eq => 0 | Lt => 1 | Gt => -1 end.
Lemma sgn_null : forall n, n==0 -> sgn n == 0.
Proof. unfold sgn; intros. destruct (compare_spec 0 n); order. Qed.
Lemma sgn_pos : forall n, 0<n -> sgn n == 1.
Proof. unfold sgn; intros. destruct (compare_spec 0 n); order. Qed.
- Lemma sgn_neg : forall n, n<0 -> sgn n == -(1).
+ Lemma sgn_neg : forall n, n<0 -> sgn n == -1.
Proof. unfold sgn; intros. destruct (compare_spec 0 n); order. Qed.
End GenericSgn.
@@ -168,6 +168,28 @@ Proof.
rewrite EQn, EQ, opp_inj_wd, eq_opp_l, or_comm. apply abs_eq_or_opp.
Qed.
+Lemma abs_lt : forall a b, abs a < b <-> -b < a < b.
+Proof.
+ intros a b.
+ destruct (abs_spec a) as [[LE EQ]|[LT EQ]]; rewrite EQ; clear EQ.
+ split; try split; try destruct 1; try order.
+ apply lt_le_trans with 0; trivial. apply opp_neg_pos; order.
+ rewrite opp_lt_mono, opp_involutive.
+ split; try split; try destruct 1; try order.
+ apply lt_le_trans with 0; trivial. apply opp_nonpos_nonneg; order.
+Qed.
+
+Lemma abs_le : forall a b, abs a <= b <-> -b <= a <= b.
+Proof.
+ intros a b.
+ destruct (abs_spec a) as [[LE EQ]|[LT EQ]]; rewrite EQ; clear EQ.
+ split; try split; try destruct 1; try order.
+ apply le_trans with 0; trivial. apply opp_nonpos_nonneg; order.
+ rewrite opp_le_mono, opp_involutive.
+ split; try split; try destruct 1; try order.
+ apply le_trans with 0. order. apply opp_nonpos_nonneg; order.
+Qed.
+
(** Triangular inequality *)
Lemma abs_triangle : forall n m, abs (n + m) <= abs n + abs m.
@@ -234,7 +256,7 @@ Qed.
Lemma sgn_spec : forall n,
0 < n /\ sgn n == 1 \/
0 == n /\ sgn n == 0 \/
- 0 > n /\ sgn n == -(1).
+ 0 > n /\ sgn n == -1.
Proof.
intros n.
destruct_sgn n; [left|right;left|right;right]; auto with relations.
@@ -249,7 +271,7 @@ Lemma sgn_pos_iff : forall n, sgn n == 1 <-> 0<n.
Proof.
split; try apply sgn_pos. destruct_sgn n; auto.
intros. elim (lt_neq 0 1); auto. apply lt_0_1.
- intros. elim (lt_neq (-(1)) 1); auto.
+ intros. elim (lt_neq (-1) 1); auto.
apply lt_trans with 0. rewrite opp_neg_pos. apply lt_0_1. apply lt_0_1.
Qed.
@@ -257,16 +279,16 @@ Lemma sgn_null_iff : forall n, sgn n == 0 <-> n==0.
Proof.
split; try apply sgn_null. destruct_sgn n; auto with relations.
intros. elim (lt_neq 0 1); auto with relations. apply lt_0_1.
- intros. elim (lt_neq (-(1)) 0); auto.
+ intros. elim (lt_neq (-1) 0); auto.
rewrite opp_neg_pos. apply lt_0_1.
Qed.
-Lemma sgn_neg_iff : forall n, sgn n == -(1) <-> n<0.
+Lemma sgn_neg_iff : forall n, sgn n == -1 <-> n<0.
Proof.
split; try apply sgn_neg. destruct_sgn n; auto with relations.
- intros. elim (lt_neq (-(1)) 1); auto with relations.
+ intros. elim (lt_neq (-1) 1); auto with relations.
apply lt_trans with 0. rewrite opp_neg_pos. apply lt_0_1. apply lt_0_1.
- intros. elim (lt_neq (-(1)) 0); auto with relations.
+ intros. elim (lt_neq (-1) 0); auto with relations.
rewrite opp_neg_pos. apply lt_0_1.
Qed.
diff --git a/theories/Numbers/Integer/BigZ/BigZ.v b/theories/Numbers/Integer/BigZ/BigZ.v
index 6153ccc75..583491386 100644
--- a/theories/Numbers/Integer/BigZ/BigZ.v
+++ b/theories/Numbers/Integer/BigZ/BigZ.v
@@ -67,8 +67,21 @@ Arguments Scope BigZ.modulo [bigZ_scope bigZ_scope].
Arguments Scope BigZ.quot [bigZ_scope bigZ_scope].
Arguments Scope BigZ.rem [bigZ_scope bigZ_scope].
Arguments Scope BigZ.gcd [bigZ_scope bigZ_scope].
+Arguments Scope BigZ.lcm [bigZ_scope bigZ_scope].
Arguments Scope BigZ.even [bigZ_scope].
Arguments Scope BigZ.odd [bigZ_scope].
+Arguments Scope BigN.testbit [bigZ_scope bigZ_scope].
+Arguments Scope BigN.shiftl [bigZ_scope bigZ_scope].
+Arguments Scope BigN.shiftr [bigZ_scope bigZ_scope].
+Arguments Scope BigN.lor [bigZ_scope bigZ_scope].
+Arguments Scope BigN.land [bigZ_scope bigZ_scope].
+Arguments Scope BigN.ldiff [bigZ_scope bigZ_scope].
+Arguments Scope BigN.lxor [bigZ_scope bigZ_scope].
+Arguments Scope BigN.setbit [bigZ_scope bigZ_scope].
+Arguments Scope BigN.clearbit [bigZ_scope bigZ_scope].
+Arguments Scope BigN.lnot [bigZ_scope].
+Arguments Scope BigN.div2 [bigZ_scope].
+Arguments Scope BigN.ones [bigZ_scope].
Local Notation "0" := BigZ.zero : bigZ_scope.
Local Notation "1" := BigZ.one : bigZ_scope.
diff --git a/theories/Numbers/Integer/BigZ/ZMake.v b/theories/Numbers/Integer/BigZ/ZMake.v
index 4c4eb6c10..1327c1923 100644
--- a/theories/Numbers/Integer/BigZ/ZMake.v
+++ b/theories/Numbers/Integer/BigZ/ZMake.v
@@ -546,4 +546,182 @@ Module Make (N:NType) <: ZType.
destruct (N.to_Z n) as [|p|p]; now try destruct p.
Qed.
+ Definition norm_pos z :=
+ match z with
+ | Pos _ => z
+ | Neg n => if N.eq_bool n N.zero then Pos n else z
+ end.
+
+ Definition testbit a n :=
+ match norm_pos n, norm_pos a with
+ | Pos p, Pos a => N.testbit a p
+ | Pos p, Neg a => negb (N.testbit (N.pred a) p)
+ | Neg p, _ => false
+ end.
+
+ Definition shiftl a n :=
+ match norm_pos a, n with
+ | Pos a, Pos n => Pos (N.shiftl a n)
+ | Pos a, Neg n => Pos (N.shiftr a n)
+ | Neg a, Pos n => Neg (N.shiftl a n)
+ | Neg a, Neg n => Neg (N.succ (N.shiftr (N.pred a) n))
+ end.
+
+ Definition shiftr a n := shiftl a (opp n).
+
+ Definition lor a b :=
+ match norm_pos a, norm_pos b with
+ | Pos a, Pos b => Pos (N.lor a b)
+ | Neg a, Pos b => Neg (N.succ (N.ldiff (N.pred a) b))
+ | Pos a, Neg b => Neg (N.succ (N.ldiff (N.pred b) a))
+ | Neg a, Neg b => Neg (N.succ (N.land (N.pred a) (N.pred b)))
+ end.
+
+ Definition land a b :=
+ match norm_pos a, norm_pos b with
+ | Pos a, Pos b => Pos (N.land a b)
+ | Neg a, Pos b => Pos (N.ldiff b (N.pred a))
+ | Pos a, Neg b => Pos (N.ldiff a (N.pred b))
+ | Neg a, Neg b => Neg (N.succ (N.lor (N.pred a) (N.pred b)))
+ end.
+
+ Definition ldiff a b :=
+ match norm_pos a, norm_pos b with
+ | Pos a, Pos b => Pos (N.ldiff a b)
+ | Neg a, Pos b => Neg (N.succ (N.lor (N.pred a) b))
+ | Pos a, Neg b => Pos (N.land a (N.pred b))
+ | Neg a, Neg b => Pos (N.ldiff (N.pred b) (N.pred a))
+ end.
+
+ Definition lxor a b :=
+ match norm_pos a, norm_pos b with
+ | Pos a, Pos b => Pos (N.lxor a b)
+ | Neg a, Pos b => Neg (N.succ (N.lxor (N.pred a) b))
+ | Pos a, Neg b => Neg (N.succ (N.lxor a (N.pred b)))
+ | Neg a, Neg b => Pos (N.lxor (N.pred a) (N.pred b))
+ end.
+
+ Definition div2 x := shiftr x one.
+
+ Lemma Zlnot_alt1 : forall x, -(x+1) = Z.lnot x.
+ Proof.
+ unfold Z.lnot, Zpred; auto with zarith.
+ Qed.
+
+ Lemma Zlnot_alt2 : forall x, Z.lnot (x-1) = -x.
+ Proof.
+ unfold Z.lnot, Zpred; auto with zarith.
+ Qed.
+
+ Lemma Zlnot_alt3 : forall x, Z.lnot (-x) = x-1.
+ Proof.
+ unfold Z.lnot, Zpred; auto with zarith.
+ Qed.
+
+ Lemma spec_norm_pos : forall x, to_Z (norm_pos x) = to_Z x.
+ Proof.
+ intros [x|x]; simpl; trivial.
+ rewrite N.spec_eq_bool, N.spec_0.
+ assert (H := Zeq_bool_if (N.to_Z x) 0).
+ destruct Zeq_bool; simpl; auto with zarith.
+ Qed.
+
+ Lemma spec_norm_pos_pos : forall x y, norm_pos x = Neg y ->
+ 0 < N.to_Z y.
+ Proof.
+ intros [x|x] y; simpl; try easy.
+ rewrite N.spec_eq_bool, N.spec_0.
+ assert (H := Zeq_bool_if (N.to_Z x) 0).
+ destruct Zeq_bool; simpl; try easy.
+ inversion 1; subst. generalize (N.spec_pos y); auto with zarith.
+ Qed.
+
+ Ltac destr_norm_pos x :=
+ rewrite <- (spec_norm_pos x);
+ let H := fresh in
+ let x' := fresh x in
+ assert (H := spec_norm_pos_pos x);
+ destruct (norm_pos x) as [x'|x'];
+ specialize (H x' (eq_refl _)) || clear H.
+
+ Lemma spec_testbit: forall x p, testbit x p = Ztestbit (to_Z x) (to_Z p).
+ Proof.
+ intros x p. unfold testbit.
+ destr_norm_pos p; simpl. destr_norm_pos x; simpl.
+ apply N.spec_testbit.
+ rewrite N.spec_testbit, N.spec_pred, Zmax_r by auto with zarith.
+ symmetry. apply Z.bits_opp. apply N.spec_pos.
+ symmetry. apply Ztestbit_neg_r; auto with zarith.
+ Qed.
+
+ Lemma spec_shiftl: forall x p, to_Z (shiftl x p) = Zshiftl (to_Z x) (to_Z p).
+ Proof.
+ intros x p. unfold shiftl.
+ destr_norm_pos x; destruct p as [p|p]; simpl;
+ assert (Hp := N.spec_pos p).
+ apply N.spec_shiftl.
+ rewrite Z.shiftl_opp_r. apply N.spec_shiftr.
+ rewrite !N.spec_shiftl.
+ rewrite !Z.shiftl_mul_pow2 by apply N.spec_pos.
+ apply Zopp_mult_distr_l.
+ rewrite Z.shiftl_opp_r, N.spec_succ, N.spec_shiftr, N.spec_pred, Zmax_r
+ by auto with zarith.
+ now rewrite Zlnot_alt1, Z.lnot_shiftr, Zlnot_alt2.
+ Qed.
+
+ Lemma spec_shiftr: forall x p, to_Z (shiftr x p) = Zshiftr (to_Z x) (to_Z p).
+ Proof.
+ intros. unfold shiftr. rewrite spec_shiftl, spec_opp.
+ apply Z.shiftl_opp_r.
+ Qed.
+
+ Lemma spec_land: forall x y, to_Z (land x y) = Zand (to_Z x) (to_Z y).
+ Proof.
+ intros x y. unfold land.
+ destr_norm_pos x; destr_norm_pos y; simpl;
+ rewrite ?N.spec_succ, ?N.spec_land, ?N.spec_ldiff, ?N.spec_lor,
+ ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1; auto with zarith.
+ now rewrite Z.ldiff_land, Zlnot_alt2.
+ now rewrite Z.ldiff_land, Z.land_comm, Zlnot_alt2.
+ now rewrite Z.lnot_lor, !Zlnot_alt2.
+ Qed.
+
+ Lemma spec_lor: forall x y, to_Z (lor x y) = Zor (to_Z x) (to_Z y).
+ Proof.
+ intros x y. unfold lor.
+ destr_norm_pos x; destr_norm_pos y; simpl;
+ rewrite ?N.spec_succ, ?N.spec_land, ?N.spec_ldiff, ?N.spec_lor,
+ ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1; auto with zarith.
+ now rewrite Z.lnot_ldiff, Z.lor_comm, Zlnot_alt2.
+ now rewrite Z.lnot_ldiff, Zlnot_alt2.
+ now rewrite Z.lnot_land, !Zlnot_alt2.
+ Qed.
+
+ Lemma spec_ldiff: forall x y, to_Z (ldiff x y) = Zdiff (to_Z x) (to_Z y).
+ Proof.
+ intros x y. unfold ldiff.
+ destr_norm_pos x; destr_norm_pos y; simpl;
+ rewrite ?N.spec_succ, ?N.spec_land, ?N.spec_ldiff, ?N.spec_lor,
+ ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1; auto with zarith.
+ now rewrite Z.ldiff_land, Zlnot_alt3.
+ now rewrite Z.lnot_lor, Z.ldiff_land, <- Zlnot_alt2.
+ now rewrite 2 Z.ldiff_land, Zlnot_alt2, Z.land_comm, Zlnot_alt3.
+ Qed.
+
+ Lemma spec_lxor: forall x y, to_Z (lxor x y) = Zxor (to_Z x) (to_Z y).
+ Proof.
+ intros x y. unfold lxor.
+ destr_norm_pos x; destr_norm_pos y; simpl;
+ rewrite ?N.spec_succ, ?N.spec_lxor, ?N.spec_pred, ?Zmax_r, ?Zlnot_alt1;
+ auto with zarith.
+ now rewrite !Z.lnot_lxor_r, Zlnot_alt2.
+ now rewrite !Z.lnot_lxor_l, Zlnot_alt2.
+ now rewrite <- Z.lxor_lnot_lnot, !Zlnot_alt2.
+ Qed.
+
+ Lemma spec_div2: forall x, to_Z (div2 x) = Zdiv2' (to_Z x).
+ Proof.
+ intros x. unfold div2. now rewrite spec_shiftr, Zdiv2'_spec, spec_1.
+ Qed.
+
End Make.
diff --git a/theories/Numbers/Integer/Binary/ZBinary.v b/theories/Numbers/Integer/Binary/ZBinary.v
index 5f38d57b8..eab33051d 100644
--- a/theories/Numbers/Integer/Binary/ZBinary.v
+++ b/theories/Numbers/Integer/Binary/ZBinary.v
@@ -10,7 +10,7 @@
Require Import ZAxioms ZProperties BinInt Zcompare Zorder ZArith_dec
- Zbool Zeven Zsqrt_def Zpow_def Zlog_def Zgcd_def Zdiv_def.
+ Zbool Zeven Zsqrt_def Zpow_def Zlog_def Zgcd_def Zdiv_def Zdigits_def.
Local Open Scope Z_scope.
@@ -191,6 +191,28 @@ Definition rem_opp_r := fun a b (_:b<>0) => Zrem_opp_r a b.
Definition quot := Zquot.
Definition rem := Zrem.
+(** Bitwise operations *)
+
+Definition testbit_spec := Ztestbit_spec.
+Definition testbit_neg_r := Ztestbit_neg_r.
+Definition shiftr_spec := Zshiftr_spec.
+Definition shiftl_spec_low := Zshiftl_spec_low.
+Definition shiftl_spec_high := Zshiftl_spec_high.
+Definition land_spec := Zand_spec.
+Definition lor_spec := Zor_spec.
+Definition ldiff_spec := Zdiff_spec.
+Definition lxor_spec := Zxor_spec.
+Definition div2_spec := Zdiv2'_spec.
+
+Definition testbit := Ztestbit.
+Definition shiftl := Zshiftl.
+Definition shiftr := Zshiftr.
+Definition land := Zand.
+Definition lor := Zor.
+Definition ldiff := Zdiff.
+Definition lxor := Zxor.
+Definition div2 := Zdiv2'.
+
(** We define [eq] only here to avoid refering to this [eq] above. *)
Definition eq := (@eq Z).
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSig.v b/theories/Numbers/Integer/SpecViaZ/ZSig.v
index 1c06b0b8e..c33915449 100644
--- a/theories/Numbers/Integer/SpecViaZ/ZSig.v
+++ b/theories/Numbers/Integer/SpecViaZ/ZSig.v
@@ -62,6 +62,14 @@ Module Type ZType.
Parameter abs : 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 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].
@@ -94,6 +102,14 @@ Module Type ZType.
Parameter spec_abs : forall x, [abs x] = Zabs [x].
Parameter spec_even : forall x, even x = Zeven_bool [x].
Parameter spec_odd : forall x, odd x = Zodd_bool [x].
+ Parameter spec_testbit: forall x p, testbit x p = Ztestbit [x] [p].
+ Parameter spec_shiftr: forall x p, [shiftr x p] = Zshiftr [x] [p].
+ Parameter spec_shiftl: forall x p, [shiftl x p] = Zshiftl [x] [p].
+ Parameter spec_land: forall x y, [land x y] = Zand [x] [y].
+ Parameter spec_lor: forall x y, [lor x y] = Zor [x] [y].
+ Parameter spec_ldiff: forall x y, [ldiff x y] = Zdiff [x] [y].
+ Parameter spec_lxor: forall x y, [lxor x y] = Zxor [x] [y].
+ Parameter spec_div2: forall x, [div2 x] = Zdiv2' [x].
End ZType.
diff --git a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
index 62b79fc3a..f8fa84283 100644
--- a/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
+++ b/theories/Numbers/Integer/SpecViaZ/ZSigZAxioms.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-Require Import ZArith Nnat ZAxioms ZSig.
+Require Import Bool ZArith Nnat ZAxioms ZSig.
(** * The interface [ZSig.ZType] implies the interface [ZAxiomsSig] *)
@@ -17,6 +17,8 @@ Hint Rewrite
spec_mul spec_opp spec_of_Z spec_div spec_modulo spec_sqrt
spec_compare spec_eq_bool spec_max spec_min spec_abs spec_sgn
spec_pow spec_log2 spec_even spec_odd spec_gcd spec_quot spec_rem
+ spec_testbit spec_shiftl spec_shiftr
+ spec_land spec_lor spec_ldiff spec_lxor spec_div2
: zsimpl.
Ltac zsimpl := autorewrite with zsimpl.
@@ -407,6 +409,71 @@ Proof.
intros. zify. apply Zgcd_nonneg.
Qed.
+(** Bitwise operations *)
+
+Lemma testbit_spec : forall a n, 0<=n ->
+ exists l, exists h, (0<=l /\ l<2^n) /\
+ a == l + ((if testbit a n then 1 else 0) + 2*h)*2^n.
+Proof.
+ intros a n. zify. intros H.
+ destruct (Ztestbit_spec [a] [n] H) as (l & h & Hl & EQ).
+ exists (of_Z l), (of_Z h).
+ destruct Ztestbit; zify; now split.
+Qed.
+
+Lemma testbit_neg_r : forall a n, n<0 -> testbit a n = false.
+Proof.
+ intros a n. zify. apply Ztestbit_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 Zshiftr_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.
+ now apply Zshiftl_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 Zshiftl_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 Zand_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 Zor_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 Zdiff_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 Zxor_spec.
+Qed.
+
+Lemma div2_spec : forall a, div2 a == shiftr a 1.
+Proof.
+ intros a. zify. now apply Zdiv2'_spec.
+Qed.
+
End ZTypeIsZAxioms.
Module ZType_ZAxioms (Z : ZType)
diff --git a/theories/Numbers/NatInt/NZBits.v b/theories/Numbers/NatInt/NZBits.v
new file mode 100644
index 000000000..072daa273
--- /dev/null
+++ b/theories/Numbers/NatInt/NZBits.v
@@ -0,0 +1,76 @@
+(************************************************************************)
+(* 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 NZAxioms NZMulOrder NZParity NZPow NZDiv NZLog.
+
+(** Axiomatization of some bitwise operations *)
+
+Module Type Bits (Import A : Typ).
+ Parameter Inline testbit : t -> t -> bool.
+ Parameters Inline shiftl shiftr land lor ldiff lxor : t -> t -> t.
+ Parameter Inline div2 : t -> t.
+End Bits.
+
+Module Type BitsNotation (Import A : Typ)(Import B : Bits A).
+ Notation "a .[ n ]" := (testbit a n) (at level 5, format "a .[ n ]").
+ Infix ">>" := shiftr (at level 30, no associativity).
+ Infix "<<" := shiftl (at level 30, no associativity).
+End BitsNotation.
+
+Module Type Bits' (A:Typ) := Bits A <+ BitsNotation A.
+
+(** For specifying [testbit], we do not rely on division and modulo,
+ since such a specification won't be easy to prove for particular
+ implementations: advanced properties of / and mod won't be
+ available at that moment. Instead, we decompose the number in
+ low and high part, with the considered bit in the middle.
+
+ Interestingly, this specification also holds for negative numbers,
+ (with a positive low part and a negative high part), and this will
+ correspond to a two's complement representation.
+
+ Moreover, we arbitrary choose false as result of [testbit] at
+ negative bit indexes (if they exist).
+*)
+
+Module Type NZBitsSpec
+ (Import A : NZOrdAxiomsSig')(Import B : Pow' A)(Import C : Bits' A).
+
+ Axiom testbit_spec : forall a n, 0<=n ->
+ exists l, exists h, 0<=l<2^n /\
+ a == l + ((if a.[n] then 1 else 0) + 2*h)*2^n.
+ Axiom testbit_neg_r : forall a n, n<0 -> a.[n] = false.
+
+ Axiom shiftr_spec : forall a n m, 0<=m -> (a >> n).[m] = a.[m+n].
+ Axiom shiftl_spec_high : forall a n m, 0<=m -> n<=m -> (a << n).[m] = a.[m-n].
+ Axiom shiftl_spec_low : forall a n m, m<n -> (a << n).[m] = false.
+
+ Axiom land_spec : forall a b n, (land a b).[n] = a.[n] && b.[n].
+ Axiom lor_spec : forall a b n, (lor a b).[n] = a.[n] || b.[n].
+ Axiom ldiff_spec : forall a b n, (ldiff a b).[n] = a.[n] && negb b.[n].
+ Axiom lxor_spec : forall a b n, (lxor a b).[n] = xorb a.[n] b.[n].
+ Axiom div2_spec : forall a, div2 a == a >> 1.
+
+End NZBitsSpec.
+
+Module Type NZBits (A:NZOrdAxiomsSig)(B:Pow A) := Bits A <+ NZBitsSpec A B.
+Module Type NZBits' (A:NZOrdAxiomsSig)(B:Pow A) := Bits' A <+ NZBitsSpec A B.
+
+(** In the functor of properties will also be defined:
+ - [setbit : t -> t -> t ] defined as [lor a (1<<n)].
+ - [clearbit : t -> t -> t ] defined as [ldiff a (1<<n)].
+ - [ones : t -> t], the number with [n] initial true bits,
+ corresponding to [2^n - 1].
+ - a logical complement [lnot]. For integer numbers it will
+ be a [t->t], doing a swap of all bits, while on natural
+ numbers, it will be a bounded complement [t->t->t], swapping
+ only the first [n] bits.
+*)
+
+(** For the moment, no shared properties about NZ here,
+ since properties and proofs for N and Z are quite different *)
diff --git a/theories/Numbers/NatInt/NZDiv.v b/theories/Numbers/NatInt/NZDiv.v
index 0807013eb..aa7ad824c 100644
--- a/theories/Numbers/NatInt/NZDiv.v
+++ b/theories/Numbers/NatInt/NZDiv.v
@@ -506,6 +506,18 @@ Proof.
apply div_mod; order.
Qed.
+Lemma mod_mul_r : forall a b c, 0<=a -> 0<b -> 0<c ->
+ a mod (b*c) == a mod b + b*((a/b) mod c).
+Proof.
+ intros a b c Ha Hb Hc.
+ apply add_cancel_l with (b*c*(a/(b*c))).
+ rewrite <- div_mod by (apply neq_mul_0; split; order).
+ rewrite <- div_div by trivial.
+ rewrite add_assoc, add_shuffle0, <- mul_assoc, <- mul_add_distr_l.
+ rewrite <- div_mod by order.
+ apply div_mod; order.
+Qed.
+
(** A last inequality: *)
Theorem div_mul_le:
diff --git a/theories/Numbers/NatInt/NZParity.v b/theories/Numbers/NatInt/NZParity.v
new file mode 100644
index 000000000..29b85724e
--- /dev/null
+++ b/theories/Numbers/NatInt/NZParity.v
@@ -0,0 +1,263 @@
+(************************************************************************)
+(* 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 NZAxioms NZMulOrder.
+
+(** Parity functions *)
+
+Module Type NZParity (Import A : NZAxiomsSig').
+ Parameter Inline even odd : t -> bool.
+ Definition Even n := exists m, n == 2*m.
+ Definition Odd n := exists m, n == 2*m+1.
+ Axiom even_spec : forall n, even n = true <-> Even n.
+ Axiom odd_spec : forall n, odd n = true <-> Odd n.
+End NZParity.
+
+Module Type NZParityProp
+ (Import A : NZOrdAxiomsSig')
+ (Import B : NZParity A)
+ (Import C : NZMulOrderProp A).
+
+(** Morphisms *)
+
+Instance Even_wd : Proper (eq==>iff) Even.
+Proof. unfold Even. solve_predicate_wd. Qed.
+
+Instance Odd_wd : Proper (eq==>iff) Odd.
+Proof. unfold Odd. solve_predicate_wd. Qed.
+
+Instance even_wd : Proper (eq==>Logic.eq) even.
+Proof.
+ intros x x' EQ. rewrite eq_iff_eq_true, 2 even_spec. now apply Even_wd.
+Qed.
+
+Instance odd_wd : Proper (eq==>Logic.eq) odd.
+Proof.
+ intros x x' EQ. rewrite eq_iff_eq_true, 2 odd_spec. now apply Odd_wd.
+Qed.
+
+(** Evenness and oddity are dual notions *)
+
+Lemma Even_or_Odd : forall x, Even x \/ Odd x.
+Proof.
+ nzinduct x.
+ left. exists 0. now nzsimpl.
+ intros x.
+ split; intros [(y,H)|(y,H)].
+ right. exists y. rewrite H. now nzsimpl.
+ left. exists (S y). rewrite H. now nzsimpl'.
+ right.
+ assert (LT : exists z, z<y).
+ destruct (lt_ge_cases 0 y) as [LT|GT]; [now exists 0 | exists x].
+ rewrite <- le_succ_l, H. nzsimpl'.
+ rewrite <- (add_0_r y) at 3. now apply add_le_mono_l.
+ destruct LT as (z,LT).
+ destruct (lt_exists_pred z y LT) as (y' & Hy' & _).
+ exists y'. rewrite <- succ_inj_wd, H, Hy'. now nzsimpl'.
+ left. exists y. rewrite <- succ_inj_wd. rewrite H. now nzsimpl.
+Qed.
+
+Lemma double_below : forall n m, n<=m -> 2*n < 2*m+1.
+Proof.
+ intros. nzsimpl'. apply lt_succ_r. now apply add_le_mono.
+Qed.
+
+Lemma double_above : forall n m, n<m -> 2*n+1 < 2*m.
+Proof.
+ intros. nzsimpl'.
+ rewrite <- le_succ_l, <- add_succ_l, <- add_succ_r.
+ apply add_le_mono; now apply le_succ_l.
+Qed.
+
+Lemma Even_Odd_False : forall x, Even x -> Odd x -> False.
+Proof.
+intros x (y,E) (z,O). rewrite O in E; clear O.
+destruct (le_gt_cases y z) as [LE|GT].
+generalize (double_below _ _ LE); order.
+generalize (double_above _ _ GT); order.
+Qed.
+
+Lemma orb_even_odd : forall n, orb (even n) (odd n) = true.
+Proof.
+ intros.
+ destruct (Even_or_Odd n) as [H|H].
+ rewrite <- even_spec in H. now rewrite H.
+ rewrite <- odd_spec in H. now rewrite H, orb_true_r.
+Qed.
+
+Lemma negb_odd : forall n, negb (odd n) = even n.
+Proof.
+ intros.
+ generalize (Even_or_Odd n) (Even_Odd_False n).
+ rewrite <- even_spec, <- odd_spec.
+ destruct (odd n), (even n); simpl; intuition.
+Qed.
+
+Lemma negb_even : forall n, negb (even n) = odd n.
+Proof.
+ intros. rewrite <- negb_odd. apply negb_involutive.
+Qed.
+
+(** Constants *)
+
+Lemma even_0 : even 0 = true.
+Proof.
+ rewrite even_spec. exists 0. now nzsimpl.
+Qed.
+
+Lemma odd_0 : odd 0 = false.
+Proof.
+ now rewrite <- negb_even, even_0.
+Qed.
+
+Lemma odd_1 : odd 1 = true.
+Proof.
+ rewrite odd_spec. exists 0. now nzsimpl'.
+Qed.
+
+Lemma even_1 : even 1 = false.
+Proof.
+ now rewrite <- negb_odd, odd_1.
+Qed.
+
+Lemma even_2 : even 2 = true.
+Proof.
+ rewrite even_spec. exists 1. now nzsimpl'.
+Qed.
+
+Lemma odd_2 : odd 2 = false.
+Proof.
+ now rewrite <- negb_even, even_2.
+Qed.
+
+(** Parity and successor *)
+
+Lemma Odd_succ : forall n, Odd (S n) <-> Even n.
+Proof.
+ split; intros (m,H).
+ exists m. apply succ_inj. now rewrite add_1_r in H.
+ exists m. rewrite add_1_r. now apply succ_wd.
+Qed.
+
+Lemma odd_succ : forall n, odd (S n) = even n.
+Proof.
+ intros. apply eq_iff_eq_true. rewrite even_spec, odd_spec.
+ apply Odd_succ.
+Qed.
+
+Lemma even_succ : forall n, even (S n) = odd n.
+Proof.
+ intros. now rewrite <- negb_odd, odd_succ, negb_even.
+Qed.
+
+Lemma Even_succ : forall n, Even (S n) <-> Odd n.
+Proof.
+ intros. now rewrite <- even_spec, even_succ, odd_spec.
+Qed.
+
+(** Parity and successor of successor *)
+
+Lemma Even_succ_succ : forall n, Even (S (S n)) <-> Even n.
+Proof.
+ intros. now rewrite Even_succ, Odd_succ.
+Qed.
+
+Lemma Odd_succ_succ : forall n, Odd (S (S n)) <-> Odd n.
+Proof.
+ intros. now rewrite Odd_succ, Even_succ.
+Qed.
+
+Lemma even_succ_succ : forall n, even (S (S n)) = even n.
+Proof.
+ intros. now rewrite even_succ, odd_succ.
+Qed.
+
+Lemma odd_succ_succ : forall n, odd (S (S n)) = odd n.
+Proof.
+ intros. now rewrite odd_succ, even_succ.
+Qed.
+
+(** Parity and addition *)
+
+Lemma even_add : forall n m, 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_add_distr_l, Hn, Hm.
+ exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm, add_assoc.
+ exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm, add_shuffle0.
+ exists (n'+m'+1). rewrite Hm,Hn. nzsimpl'. now rewrite add_shuffle1.
+Qed.
+
+Lemma odd_add : forall n m, odd (n+m) = xorb (odd n) (odd m).
+Proof.
+ intros. rewrite <- !negb_even. rewrite even_add.
+ now destruct (even n), (even m).
+Qed.
+
+(** Parity and multiplication *)
+
+Lemma even_mul : forall n m, even (mul n m) = even n || even m.
+Proof.
+ intros.
+ case_eq (even n); simpl; rewrite ?even_spec.
+ intros (n',Hn). exists (n'*m). now rewrite Hn, mul_assoc.
+ case_eq (even m); simpl; rewrite ?even_spec.
+ intros (m',Hm). exists (n*m'). now rewrite Hm, !mul_assoc, (mul_comm 2).
+ (* odd / odd *)
+ rewrite <- !negb_true_iff, !negb_even, !odd_spec.
+ intros (m',Hm) (n',Hn). exists (n'*2*m' +n'+m').
+ rewrite Hn,Hm, !mul_add_distr_l, !mul_add_distr_r, !mul_1_l, !mul_1_r.
+ now rewrite add_shuffle1, add_assoc, !mul_assoc.
+Qed.
+
+Lemma odd_mul : forall n m, odd (mul n m) = odd n && odd m.
+Proof.
+ intros. rewrite <- !negb_even. rewrite even_mul.
+ now destruct (even n), (even m).
+Qed.
+
+(** A particular case : adding by an even number *)
+
+Lemma even_add_even : forall n m, Even m -> even (n+m) = even n.
+Proof.
+ intros n m Hm. apply even_spec in Hm.
+ rewrite even_add, Hm. now destruct (even n).
+Qed.
+
+Lemma odd_add_even : forall n m, Even m -> odd (n+m) = odd n.
+Proof.
+ intros n m Hm. apply even_spec in Hm.
+ rewrite odd_add, <- (negb_even m), Hm. now destruct (odd n).
+Qed.
+
+Lemma even_add_mul_even : forall n m p, Even m -> even (n+m*p) = even n.
+Proof.
+ intros n m p Hm. apply even_spec in Hm.
+ apply even_add_even. apply even_spec. now rewrite even_mul, Hm.
+Qed.
+
+Lemma odd_add_mul_even : forall n m p, Even m -> odd (n+m*p) = odd n.
+Proof.
+ intros n m p Hm. apply even_spec in Hm.
+ apply odd_add_even. apply even_spec. now rewrite even_mul, Hm.
+Qed.
+
+Lemma even_add_mul_2 : forall n m, even (n+2*m) = even n.
+Proof.
+ intros. apply even_add_mul_even. apply even_spec, even_2.
+Qed.
+
+Lemma odd_add_mul_2 : forall n m, odd (n+2*m) = odd n.
+Proof.
+ intros. apply odd_add_mul_even. apply even_spec, even_2.
+Qed.
+
+End NZParityProp. \ No newline at end of file
diff --git a/theories/Numbers/NatInt/NZPow.v b/theories/Numbers/NatInt/NZPow.v
index 76b745bf0..23c27777a 100644
--- a/theories/Numbers/NatInt/NZPow.v
+++ b/theories/Numbers/NatInt/NZPow.v
@@ -54,6 +54,14 @@ Proof.
rewrite EQ. now nzsimpl.
Qed.
+Lemma pow_0_l' : forall a, a~=0 -> 0^a == 0.
+Proof.
+ intros a Ha.
+ destruct (lt_trichotomy a 0) as [LT|[EQ|GT]]; try order.
+ now rewrite pow_neg_r.
+ now apply pow_0_l.
+Qed.
+
Lemma pow_1_r : forall a, a^1 == a.
Proof.
intros. now nzsimpl'.
@@ -75,6 +83,36 @@ Qed.
Hint Rewrite pow_2_r : nz.
+(** Power and nullity *)
+
+Lemma pow_eq_0 : forall a b, 0<=b -> a^b == 0 -> a == 0.
+Proof.
+ intros a b Hb. apply le_ind with (4:=Hb).
+ solve_predicate_wd.
+ rewrite pow_0_r. order'.
+ clear b Hb. intros b Hb IH.
+ rewrite pow_succ_r by trivial.
+ intros H. apply eq_mul_0 in H. destruct H; trivial.
+ now apply IH.
+Qed.
+
+Lemma pow_nonzero : forall a b, a~=0 -> 0<=b -> a^b ~= 0.
+Proof.
+ intros a b Ha Hb. contradict Ha. now apply pow_eq_0 with b.
+Qed.
+
+Lemma pow_eq_0_iff : forall a b, a^b == 0 <-> b<0 \/ (0<b /\ a==0).
+Proof.
+ intros a b. split.
+ intros H.
+ destruct (lt_trichotomy b 0) as [Hb|[Hb|Hb]].
+ now left.
+ rewrite Hb, pow_0_r in H; order'.
+ right. split; trivial. apply pow_eq_0 with b; order.
+ intros [Hb|[Hb Ha]]. now rewrite pow_neg_r.
+ rewrite Ha. apply pow_0_l'. order.
+Qed.
+
(** Power and addition, multiplication *)
Lemma pow_add_r : forall a b c, 0<=b -> 0<=c ->
diff --git a/theories/Numbers/Natural/Abstract/NAxioms.v b/theories/Numbers/Natural/Abstract/NAxioms.v
index 82f072746..09438628d 100644
--- a/theories/Numbers/Natural/Abstract/NAxioms.v
+++ b/theories/Numbers/Natural/Abstract/NAxioms.v
@@ -8,7 +8,7 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-Require Export NZAxioms NZPow NZSqrt NZLog NZDiv NZGcd.
+Require Export Bool NZAxioms NZParity NZPow NZSqrt NZLog NZDiv NZGcd NZBits.
(** From [NZ], we obtain natural numbers just by stating that [pred 0] == 0 *)
@@ -19,19 +19,8 @@ End NAxiom.
Module Type NAxiomsMiniSig := NZOrdAxiomsSig <+ NAxiom.
Module Type NAxiomsMiniSig' := NZOrdAxiomsSig' <+ NAxiom.
-
(** Let's now add some more functions and their specification *)
-(** Parity functions *)
-
-Module Type Parity (Import N : NAxiomsMiniSig').
- Parameter Inline even odd : t -> bool.
- Definition Even n := exists m, n == 2*m.
- Definition Odd n := exists m, n == 2*m+1.
- Axiom even_spec : forall n, even n = true <-> Even n.
- Axiom odd_spec : forall n, odd n = true <-> Odd n.
-End Parity.
-
(** Division Function : we reuse NZDiv.DivMod and NZDiv.NZDivCommon,
and add to that a N-specific constraint. *)
@@ -39,17 +28,17 @@ 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 div mod gcd pow sqrt log2, the NZ axiomatizations are enough. *)
+(** For all other functions, the NZ axiomatizations are enough. *)
(** We now group everything together. *)
-Module Type NAxiomsSig := NAxiomsMiniSig <+ HasCompare <+ Parity
- <+ NZPow.NZPow <+ NZSqrt.NZSqrt <+ NZLog.NZLog2 <+ NZGcd.NZGcd
- <+ NZDiv.NZDiv.
+Module Type NAxiomsSig := NAxiomsMiniSig <+ HasCompare <+ HasEqBool
+ <+ NZParity.NZParity <+ NZPow.NZPow <+ NZSqrt.NZSqrt <+ NZLog.NZLog2
+ <+ NZGcd.NZGcd <+ NZDiv.NZDiv <+ NZBits.NZBits.
-Module Type NAxiomsSig' := NAxiomsMiniSig' <+ HasCompare <+ Parity
- <+ NZPow.NZPow' <+ NZSqrt.NZSqrt' <+ NZLog.NZLog2 <+ NZGcd.NZGcd'
- <+ NZDiv.NZDiv'.
+Module Type NAxiomsSig' := NAxiomsMiniSig' <+ HasCompare <+ HasEqBool
+ <+ NZParity.NZParity <+ NZPow.NZPow' <+ NZSqrt.NZSqrt' <+ NZLog.NZLog2
+ <+ NZGcd.NZGcd' <+ NZDiv.NZDiv' <+ NZBits.NZBits'.
(** It could also be interesting to have a constructive recursor function. *)
diff --git a/theories/Numbers/Natural/Abstract/NBits.v b/theories/Numbers/Natural/Abstract/NBits.v
new file mode 100644
index 000000000..2cb5bbc06
--- /dev/null
+++ b/theories/Numbers/Natural/Abstract/NBits.v
@@ -0,0 +1,1422 @@
+(************************************************************************)
+(* 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. apply pow_wd. now apply div_exact.
+ reflexivity.
+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.
+
+(** Alternative caracterisations of [testbit] *)
+
+Lemma testbit_spec' : forall a n, a.[n] == (a / 2^n) mod 2.
+Proof.
+ intros a n.
+ destruct (testbit_spec a n) as (l & h & (_,H) & EQ).
+ apply le_0_l.
+ fold (b2n a.[n]) in EQ.
+ apply mod_unique with h. destruct a.[n]; order'.
+ symmetry. apply div_unique with l; trivial.
+ now rewrite add_comm, mul_comm, (add_comm (2*h)).
+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.
+
+(** testbit is hence a morphism *)
+
+Instance testbit_wd : Proper (eq==>eq==>Logic.eq) testbit.
+Proof.
+ intros a a' Ha n n' Hn. now rewrite 2 testbit_eqb, Ha, Hn.
+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 initial specification of testbit 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_mod : forall a, a.[0] == a mod 2.
+Proof.
+ intros a. rewrite testbit_spec'. now nzsimpl.
+Qed.
+
+Lemma bit0_eqb : forall a, a.[0] = eqb (a mod 2) 1.
+Proof.
+ intros a. rewrite testbit_eqb. now nzsimpl.
+Qed.
+
+Lemma bit0_odd : forall a, a.[0] = odd a.
+Proof.
+ intros. rewrite bit0_eqb.
+ apply eq_true_iff_eq. rewrite eqb_eq, odd_spec. split.
+ intros H. exists (a/2). rewrite <- H. apply div_mod. order'.
+ intros (b,H). rewrite H, add_comm, mul_comm, mod_add, mod_small; order'.
+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_predicate_wd|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'.
+ apply add_wd; [ | now rewrite <- 2 bit0_mod, H].
+ apply mul_wd. reflexivity.
+ 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_predicate_wd.
+ 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.
+ intros a a' Ha n n' Hn. unfold setbit. now rewrite Ha, Hn.
+Qed.
+
+Instance clearbit_wd : Proper (eq==>eq==>eq) clearbit.
+Proof.
+ intros a a' Ha n n' Hn. unfold clearbit. now rewrite Ha, Hn.
+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. intros a a' Ha; unfold ones; now rewrite Ha. Qed.
+
+Instance lnot_wd : Proper (eq==>eq==>eq) lnot.
+Proof. intros a a' Ha n n' Hn; unfold lnot; now rewrite Ha, Hn. 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'.
+ apply div_wd; try easy.
+ 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.
+ apply testbit_wd; try easy.
+ 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/NDiv.v b/theories/Numbers/Natural/Abstract/NDiv.v
index 9110ec036..47f409605 100644
--- a/theories/Numbers/Natural/Abstract/NDiv.v
+++ b/theories/Numbers/Natural/Abstract/NDiv.v
@@ -219,6 +219,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:
diff --git a/theories/Numbers/Natural/Abstract/NParity.v b/theories/Numbers/Natural/Abstract/NParity.v
index bd6588686..6a1e20ce0 100644
--- a/theories/Numbers/Natural/Abstract/NParity.v
+++ b/theories/Numbers/Natural/Abstract/NParity.v
@@ -6,172 +6,31 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-Require Import Bool NSub.
+Require Import Bool NSub NZParity.
-(** Properties of [even], [odd]. *)
-
-(** NB: most parts of [NParity] and [ZParity] are common,
- but it is difficult to share them in NZ, since
- initial proofs [Even_or_Odd] and [Even_Odd_False] must
- be proved differently *)
+(** Some additionnal properties of [even], [odd]. *)
Module Type NParityProp (Import N : NAxiomsSig')(Import NP : NSubProp N).
-Instance Even_wd : Proper (eq==>iff) Even.
-Proof. unfold Even. solve_predicate_wd. Qed.
-
-Instance Odd_wd : Proper (eq==>iff) Odd.
-Proof. unfold Odd. solve_predicate_wd. Qed.
-
-Instance even_wd : Proper (eq==>Logic.eq) even.
-Proof.
- intros x x' EQ. rewrite eq_iff_eq_true, 2 even_spec. now apply Even_wd.
-Qed.
-
-Instance odd_wd : Proper (eq==>Logic.eq) odd.
-Proof.
- intros x x' EQ. rewrite eq_iff_eq_true, 2 odd_spec. now apply Odd_wd.
-Qed.
-
-Lemma Even_or_Odd : forall x, Even x \/ Odd x.
-Proof.
- induct x.
- left. exists 0. now nzsimpl.
- intros x.
- intros [(y,H)|(y,H)].
- right. exists y. rewrite H. now nzsimpl.
- left. exists (S y). rewrite H. now nzsimpl'.
-Qed.
-
-Lemma double_below : forall n m, n<=m -> 2*n < 2*m+1.
-Proof.
- intros. nzsimpl'. apply lt_succ_r. now apply add_le_mono.
-Qed.
-
-Lemma double_above : forall n m, n<m -> 2*n+1 < 2*m.
-Proof.
- intros. nzsimpl'.
- rewrite <- le_succ_l, <- add_succ_l, <- add_succ_r.
- apply add_le_mono; now apply le_succ_l.
-Qed.
-
-Lemma Even_Odd_False : forall x, Even x -> Odd x -> False.
-Proof.
-intros x (y,E) (z,O). rewrite O in E; clear O.
-destruct (le_gt_cases y z) as [LE|GT].
-generalize (double_below _ _ LE); order.
-generalize (double_above _ _ GT); order.
-Qed.
-
-Lemma orb_even_odd : forall n, orb (even n) (odd n) = true.
-Proof.
- intros.
- destruct (Even_or_Odd n) as [H|H].
- rewrite <- even_spec in H. now rewrite H.
- rewrite <- odd_spec in H. now rewrite H, orb_true_r.
-Qed.
-
-Lemma negb_odd_even : forall n, negb (odd n) = even n.
-Proof.
- intros.
- generalize (Even_or_Odd n) (Even_Odd_False n).
- rewrite <- even_spec, <- odd_spec.
- destruct (odd n), (even n); simpl; intuition.
-Qed.
-
-Lemma negb_even_odd : forall n, negb (even n) = odd n.
-Proof.
- intros. rewrite <- negb_odd_even. apply negb_involutive.
-Qed.
-
-Lemma even_0 : even 0 = true.
-Proof.
- rewrite even_spec. exists 0. now nzsimpl.
-Qed.
+Include NZParityProp N N NP.
-Lemma odd_1 : odd 1 = true.
-Proof.
- rewrite odd_spec. exists 0. now nzsimpl'.
-Qed.
-
-Lemma Odd_succ_Even : forall n, Odd (S n) <-> Even n.
-Proof.
- split; intros (m,H).
- exists m. apply succ_inj. now rewrite add_1_r in H.
- exists m. rewrite add_1_r. now apply succ_wd.
-Qed.
-
-Lemma odd_succ_even : forall n, odd (S n) = even n.
-Proof.
- intros. apply eq_iff_eq_true. rewrite even_spec, odd_spec.
- apply Odd_succ_Even.
-Qed.
-
-Lemma even_succ_odd : forall n, even (S n) = odd n.
-Proof.
- intros. now rewrite <- negb_odd_even, odd_succ_even, negb_even_odd.
-Qed.
-
-Lemma Even_succ_Odd : forall n, Even (S n) <-> Odd n.
-Proof.
- intros. now rewrite <- even_spec, even_succ_odd, odd_spec.
-Qed.
-
-Lemma odd_pred_even : forall n, n~=0 -> odd (P n) = even n.
+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_odd.
+ symmetry. apply even_succ.
Qed.
-Lemma even_pred_odd : forall n, n~=0 -> even (P n) = odd n.
+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_even.
-Qed.
-
-Lemma even_add : forall n m, 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, ?odd_spec, ?even_spec;
- intros (m',Hm) (n',Hn).
- exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm.
- exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm, add_assoc.
- exists (n'+m'). now rewrite mul_add_distr_l, Hn, Hm, add_shuffle0.
- exists (n'+m'+1). rewrite Hm,Hn. nzsimpl'. now rewrite add_shuffle1.
-Qed.
-
-Lemma odd_add : forall n m, odd (n+m) = xorb (odd n) (odd m).
-Proof.
- intros. rewrite <- !negb_even_odd. rewrite even_add.
- now destruct (even n), (even m).
-Qed.
-
-Lemma even_mul : forall n m, even (mul n m) = even n || even m.
-Proof.
- intros.
- case_eq (even n); simpl; rewrite ?even_spec.
- intros (n',Hn). exists (n'*m). now rewrite Hn, mul_assoc.
- case_eq (even m); simpl; rewrite ?even_spec.
- intros (m',Hm). exists (n*m'). now rewrite Hm, !mul_assoc, (mul_comm 2).
- (* odd / odd *)
- rewrite <- !negb_true_iff, !negb_even_odd, !odd_spec.
- intros (m',Hm) (n',Hn). exists (n'*2*m' +n'+m').
- rewrite Hn,Hm, !mul_add_distr_l, !mul_add_distr_r, !mul_1_l, !mul_1_r.
- now rewrite add_shuffle1, add_assoc, !mul_assoc.
-Qed.
-
-Lemma odd_mul : forall n m, odd (mul n m) = odd n && odd m.
-Proof.
- intros. rewrite <- !negb_even_odd. rewrite even_mul.
- now destruct (even n), (even m).
+ 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, ?odd_spec, ?even_spec;
+ 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).
@@ -197,7 +56,7 @@ Qed.
Lemma odd_sub : forall n m, m<=n -> odd (n-m) = xorb (odd n) (odd m).
Proof.
- intros. rewrite <- !negb_even_odd. rewrite even_sub by trivial.
+ intros. rewrite <- !negb_even. rewrite even_sub by trivial.
now destruct (even n), (even m).
Qed.
diff --git a/theories/Numbers/Natural/Abstract/NPow.v b/theories/Numbers/Natural/Abstract/NPow.v
index 275a5c4f5..68976624e 100644
--- a/theories/Numbers/Natural/Abstract/NPow.v
+++ b/theories/Numbers/Natural/Abstract/NPow.v
@@ -50,10 +50,21 @@ 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.
-(** Positivity *)
+(** Power and nullity *)
-Lemma pow_nonzero : forall a b, a~=0 -> a^b~=0.
-Proof. intros. rewrite neq_0_lt_0. wrap pow_pos_nonneg. Qed.
+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 *)
@@ -143,7 +154,7 @@ Qed.
Lemma odd_pow : forall a b, b~=0 -> odd (a^b) = odd a.
Proof.
- intros. now rewrite <- !negb_even_odd, even_pow.
+ 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 58e3afe78..1edb6b51f 100644
--- a/theories/Numbers/Natural/Abstract/NProperties.v
+++ b/theories/Numbers/Natural/Abstract/NProperties.v
@@ -7,10 +7,11 @@
(************************************************************************)
Require Export NAxioms.
-Require Import NMaxMin NParity NPow NSqrt NLog NDiv NGcd NLcm.
+Require Import NMaxMin NParity NPow NSqrt NLog NDiv NGcd NLcm NBits.
(** This functor summarizes all known facts about N. *)
Module Type NProp (N:NAxiomsSig) :=
NMaxMinProp N <+ NParityProp N <+ NPowProp N <+ NSqrtProp N
- <+ NLog2Prop N <+ NDivProp N <+ NGcdProp N <+ NLcmProp N.
+ <+ NLog2Prop N <+ NDivProp N <+ NGcdProp N <+ NLcmProp N
+ <+ NBitsProp N.
diff --git a/theories/Numbers/Natural/BigN/BigN.v b/theories/Numbers/Natural/BigN/BigN.v
index 209bee8c1..315876656 100644
--- a/theories/Numbers/Natural/BigN/BigN.v
+++ b/theories/Numbers/Natural/BigN/BigN.v
@@ -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,8 +21,7 @@ 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].
*)
@@ -43,6 +42,7 @@ 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.to_N [bigN_scope].
Arguments Scope BigN.succ [bigN_scope].
Arguments Scope BigN.pred [bigN_scope].
Arguments Scope BigN.square [bigN_scope].
@@ -66,8 +66,21 @@ 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].
+Arguments Scope BigN.lcm [bigN_scope bigN_scope].
Arguments Scope BigN.even [bigN_scope].
Arguments Scope BigN.odd [bigN_scope].
+Arguments Scope BigN.testbit [bigN_scope bigN_scope].
+Arguments Scope BigN.shiftl [bigN_scope bigN_scope].
+Arguments Scope BigN.shiftr [bigN_scope bigN_scope].
+Arguments Scope BigN.lor [bigN_scope bigN_scope].
+Arguments Scope BigN.land [bigN_scope bigN_scope].
+Arguments Scope BigN.ldiff [bigN_scope bigN_scope].
+Arguments Scope BigN.lxor [bigN_scope bigN_scope].
+Arguments Scope BigN.setbit [bigN_scope bigN_scope].
+Arguments Scope BigN.clearbit [bigN_scope bigN_scope].
+Arguments Scope BigN.lnot [bigN_scope bigN_scope].
+Arguments Scope BigN.div2 [bigN_scope].
+Arguments Scope BigN.ones [bigN_scope].
Local Notation "0" := BigN.zero : bigN_scope. (* temporary notation *)
Local Notation "1" := BigN.one : bigN_scope. (* temporary notation *)
diff --git a/theories/Numbers/Natural/BigN/NMake.v b/theories/Numbers/Natural/BigN/NMake.v
index 306efc19c..a55fb5900 100644
--- a/theories/Numbers/Natural/BigN/NMake.v
+++ b/theories/Numbers/Natural/BigN/NMake.v
@@ -16,7 +16,7 @@
representation. The representation-dependent (and macro-generated) part
is now in [NMake_gen]. *)
-Require Import Bool BigNumPrelude ZArith Nnat CyclicAxioms DoubleType
+Require Import Bool BigNumPrelude ZArith Nnat Ndigits CyclicAxioms DoubleType
Nbasic Wf_nat StreamMemo NSig NMake_gen.
Module Make (W0:CyclicType) <: NType.
@@ -972,6 +972,44 @@ Module Make (W0:CyclicType) <: NType.
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 pheight p :=
@@ -1212,7 +1250,7 @@ Module Make (W0:CyclicType) <: NType.
let sub_c := ZnZ.sub_c in
let add_mul_div := ZnZ.add_mul_div in
let zzero := ZnZ.zero in
- fun p x => match sub_c zdigits p with
+ fun x p => match sub_c zdigits p with
| C0 d => reduce n (add_mul_div d zzero x)
| C1 _ => zero
end).
@@ -1236,13 +1274,13 @@ Module Make (W0:CyclicType) <: NType.
rewrite Zpower_0_r; ring.
Qed.
- Theorem spec_shiftr: forall n x,
- [shiftr n x] = [x] / 2 ^ [n].
+ 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 p x. simpl.
- assert (Hx := ZnZ.spec_to_Z p).
- assert (Hy := ZnZ.spec_to_Z x).
+ 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 ] *)
@@ -1264,6 +1302,12 @@ Module Make (W0:CyclicType) <: NType.
generalize (ZnZ.spec_to_Z d); auto with zarith.
Qed.
+ Lemma spec_shiftr: forall x p, [shiftr x p] = Zshiftr [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
@@ -1273,7 +1317,7 @@ Module Make (W0:CyclicType) <: NType.
let op := dom_op n in
let add_mul_div := ZnZ.add_mul_div in
let zero := ZnZ.zero in
- fun p x => reduce n (add_mul_div p x zero)).
+ 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.
@@ -1281,20 +1325,20 @@ Module Make (W0:CyclicType) <: NType.
Lemma unsafe_shiftl_fold : unsafe_shiftl = same_level unsafe_shiftln.
Proof. red_t; reflexivity. Qed.
- Theorem spec_unsafe_shiftl_aux : forall p x K,
+ Theorem spec_unsafe_shiftl_aux : forall x p K,
0 <= K ->
[x] < 2^K ->
[p] + K <= Zpos (digits x) ->
- [unsafe_shiftl p x] = [x] * 2 ^ [p].
+ [unsafe_shiftl x p] = [x] * 2 ^ [p].
Proof.
- intros p x.
+ 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 p x.
- intros n p x K HK Hx Hp. simpl. rewrite spec_reduce.
+ 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 *).
@@ -1308,8 +1352,8 @@ Module Make (W0:CyclicType) <: NType.
apply Zpower_le_monotone2; auto with zarith.
Qed.
- Theorem spec_unsafe_shiftl: forall p x,
- [p] <= [head0 x] -> [unsafe_shiftl p x] = [x] * 2 ^ [p].
+ 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].
@@ -1436,19 +1480,19 @@ Module Make (W0:CyclicType) <: NType.
(** Finally we iterate [double_size] enough before [unsafe_shiftl]
in order to get a fully correct [shiftl]. *)
- Definition shiftl_aux_body cont n x :=
- match compare n (head0 x) with
- Gt => cont n (double_size x)
- | _ => unsafe_shiftl n x
+ 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.
+ 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.
@@ -1459,22 +1503,22 @@ Module Make (W0:CyclicType) <: NType.
rewrite Zpower_1_r; apply Zmult_le_compat_l; auto with zarith.
Qed.
- Fixpoint shiftl_aux p cont n x :=
+ 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.
@@ -1501,15 +1545,15 @@ Module Make (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.
+ 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.
@@ -1531,42 +1575,64 @@ Module Make (W0:CyclicType) <: NType.
apply Zpower_le_monotone2; auto with zarith.
Qed.
- (** * Parity test *)
+ Lemma spec_shiftl: forall x p, [shiftl x p] = Zshiftl [x] [p].
+ Proof.
+ intros.
+ now rewrite spec_shiftl_pow2, Z.shiftl_mul_pow2 by apply spec_pos.
+ Qed.
- Definition even : t -> bool := Eval red_t in
- iter_t (fun n x => ZnZ.is_even x).
+ (** Other bitwise operations *)
- Definition odd x := negb (even x).
+ Definition testbit x n := odd (shiftr x n).
- Lemma even_fold : even = iter_t (fun n x => ZnZ.is_even x).
- Proof. red_t; reflexivity. Qed.
+ Lemma spec_testbit: forall x p, testbit x p = Ztestbit [x] [p].
+ Proof.
+ intros. unfold testbit. symmetry.
+ rewrite spec_odd, spec_shiftr. apply Z.testbit_odd.
+ Qed.
- Theorem spec_even_aux: forall x,
- if even x then [x] mod 2 = 0 else [x] mod 2 = 1.
+ Definition div2 x := shiftr x one.
+
+ Lemma spec_div2: forall x, [div2 x] = Zdiv2' [x].
Proof.
- intros x. rewrite even_fold. destr_t x as (n,x).
- exact (ZnZ.spec_is_even x).
+ intros. unfold div2. symmetry.
+ rewrite spec_shiftr, spec_1. apply Zdiv2'_spec.
Qed.
- Theorem spec_even: forall x, even x = Zeven_bool [x].
+ (** TODO : provide efficient versions instead of just converting
+ from/to N (see with Laurent) *)
+
+ Definition lor x y := of_N (Nor (to_N x) (to_N y)).
+ Definition land x y := of_N (Nand (to_N x) (to_N y)).
+ Definition ldiff x y := of_N (Ndiff (to_N x) (to_N y)).
+ Definition lxor x y := of_N (Nxor (to_N x) (to_N y)).
+
+ Lemma spec_land: forall x y, [land x y] = Zand [x] [y].
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.
+ 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.
- Theorem spec_odd: forall x, odd x = Zodd_bool [x].
+ Lemma spec_lor: forall x y, [lor x y] = Zor [x] [y].
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.
+ 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] = Zdiff [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] = Zxor [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 2736013c6..8779f4be3 100644
--- a/theories/Numbers/Natural/BigN/NMake_gen.ml
+++ b/theories/Numbers/Natural/BigN/NMake_gen.ml
@@ -711,12 +711,12 @@ pr
(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 y) [x] [y] (same_level f 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 y <= fst (same_level f' x y))
+ 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).
diff --git a/theories/Numbers/Natural/BigN/Nbasic.v b/theories/Numbers/Natural/BigN/Nbasic.v
index 56dd9c8c5..552002e44 100644
--- a/theories/Numbers/Natural/BigN/Nbasic.v
+++ b/theories/Numbers/Natural/BigN/Nbasic.v
@@ -565,7 +565,7 @@ Axiom spec_same_level_dep :
(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 y) [x] [y] (same_level f 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 *)
diff --git a/theories/Numbers/Natural/Binary/NBinary.v b/theories/Numbers/Natural/Binary/NBinary.v
index 1b5d382a3..d2979bcf0 100644
--- a/theories/Numbers/Natural/Binary/NBinary.v
+++ b/theories/Numbers/Natural/Binary/NBinary.v
@@ -8,7 +8,7 @@
(* Evgeny Makarov, INRIA, 2007 *)
(************************************************************************)
-Require Import BinPos Ndiv_def Nsqrt_def Ngcd_def.
+Require Import BinPos Ndiv_def Nsqrt_def Ngcd_def Ndigits.
Require Export BinNat.
Require Import NAxioms NProperties.
@@ -178,6 +178,20 @@ Definition gcd_greatest := Ngcd_greatest.
Lemma gcd_nonneg : forall a b, 0 <= Ngcd a b.
Proof. intros. now destruct (Ngcd a b). Qed.
+(** Bitwise Operations *)
+
+Definition testbit_spec a n (_:0<=n) := Ntestbit_spec a n.
+Lemma testbit_neg_r a n (H:n<0) : Ntestbit a n = false.
+Proof. now destruct n. Qed.
+Definition shiftl_spec_low := Nshiftl_spec_low.
+Definition shiftl_spec_high a n m (_:0<=m) := Nshiftl_spec_high a n m.
+Definition shiftr_spec a n m (_:0<=m) := Nshiftr_spec a n m.
+Definition lxor_spec := Nxor_spec.
+Definition land_spec := Nand_spec.
+Definition lor_spec := Nor_spec.
+Definition ldiff_spec := Ndiff_spec.
+Definition div2_spec a : Ndiv2 a = Nshiftr a 1 := eq_refl _.
+
(** The instantiation of operations.
Placing them at the very end avoids having indirections in above lemmas. *)
@@ -207,6 +221,14 @@ Definition sqrt := Nsqrt.
Definition log2 := Nlog2.
Definition divide := Ndivide.
Definition gcd := Ngcd.
+Definition testbit := Ntestbit.
+Definition shiftl := Nshiftl.
+Definition shiftr := Nshiftr.
+Definition lxor := Nxor.
+Definition land := Nand.
+Definition lor := Nor.
+Definition ldiff := Ndiff.
+Definition div2 := Ndiv2.
Include NProp
<+ UsualMinMaxLogicalProperties <+ UsualMinMaxDecProperties.
diff --git a/theories/Numbers/Natural/Peano/NPeano.v b/theories/Numbers/Natural/Peano/NPeano.v
index 60c59b323..6f72a504c 100644
--- a/theories/Numbers/Natural/Peano/NPeano.v
+++ b/theories/Numbers/Natural/Peano/NPeano.v
@@ -9,7 +9,7 @@
(************************************************************************)
Require Import
- Bool Peano Peano_dec Compare_dec Plus Mult Minus Le Lt EqNat
+ Bool Peano Peano_dec Compare_dec Plus Mult Minus Le Lt EqNat Div2 Wf_nat
NAxioms NProperties.
(** Functions not already defined *)
@@ -69,6 +69,21 @@ Proof.
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,
@@ -332,6 +347,191 @@ Proof.
now rewrite mult_minus_distr_l, mult_assoc, Hu, Hv, 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_spec : forall a n,
+ exists l, exists h, 0<=l<2^n /\
+ a = l + ((if testbit a n then 1 else 0) + 2*h)*2^n.
+Proof.
+ intros a n. revert a. induction n; intros a; simpl testbit.
+ exists 0. exists (div2 a).
+ split. simpl. unfold lt. now split.
+ case_eq (odd a); intros EQ; simpl.
+ rewrite mult_1_r, <- plus_n_O.
+ now apply odd_double, Odd_equiv, odd_spec.
+ rewrite mult_1_r, <- plus_n_O. apply even_double.
+ destruct (Even.even_or_odd a) as [H|H]; trivial.
+ apply Odd_equiv, odd_spec in H. rewrite H in EQ; discriminate.
+ destruct (IHn (div2 a)) as (l & h & (_,H) & EQ).
+ destruct (Even.even_or_odd a) as [EV|OD].
+ exists (double l). exists h.
+ split. split. apply le_O_n.
+ unfold double; simpl. rewrite <- plus_n_O. now apply plus_lt_compat.
+ pattern a at 1. rewrite (even_double a EV).
+ pattern (div2 a) at 1. rewrite EQ.
+ rewrite !double_twice, mult_plus_distr_l. f_equal.
+ rewrite mult_assoc, (mult_comm 2), <- mult_assoc. f_equal.
+ exists (S (double l)). exists h.
+ split. split. apply le_O_n.
+ red. red in H.
+ unfold double; simpl. rewrite <- plus_n_O, plus_n_Sm, <- plus_Sn_m.
+ now apply plus_le_compat.
+ rewrite plus_Sn_m.
+ pattern a at 1. rewrite (odd_double a OD). f_equal.
+ pattern (div2 a) at 1. rewrite EQ.
+ rewrite !double_twice, mult_plus_distr_l. f_equal.
+ rewrite mult_assoc, (mult_comm 2), <- mult_assoc. f_equal.
+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] *)
@@ -525,6 +725,27 @@ 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.
+
+Definition testbit_spec a n (_:0<=n) := testbit_spec 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
diff --git a/theories/Numbers/Natural/SpecViaZ/NSig.v b/theories/Numbers/Natural/SpecViaZ/NSig.v
index dc2d27fa4..021ac29ee 100644
--- a/theories/Numbers/Natural/SpecViaZ/NSig.v
+++ b/theories/Numbers/Natural/SpecViaZ/NSig.v
@@ -56,10 +56,16 @@ Module Type NType.
Parameter div : t -> t -> t.
Parameter modulo : t -> t -> t.
Parameter gcd : t -> t -> t.
- Parameter shiftr : t -> t -> t.
- Parameter shiftl : 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 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].
@@ -84,10 +90,16 @@ Module Type NType.
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_even: forall x, even x = Zeven_bool [x].
Parameter spec_odd: forall x, odd x = Zodd_bool [x].
+ Parameter spec_testbit: forall x p, testbit x p = Ztestbit [x] [p].
+ Parameter spec_shiftr: forall x p, [shiftr x p] = Zshiftr [x] [p].
+ Parameter spec_shiftl: forall x p, [shiftl x p] = Zshiftl [x] [p].
+ Parameter spec_land: forall x y, [land x y] = Zand [x] [y].
+ Parameter spec_lor: forall x y, [lor x y] = Zor [x] [y].
+ Parameter spec_ldiff: forall x y, [ldiff x y] = Zdiff [x] [y].
+ Parameter spec_lxor: forall x y, [lxor x y] = Zxor [x] [y].
+ Parameter spec_div2: forall x, [div2 x] = Zdiv2' [x].
End NType.
diff --git a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
index 6760cfc81..a169c009d 100644
--- a/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
+++ b/theories/Numbers/Natural/SpecViaZ/NSigNAxioms.v
@@ -7,7 +7,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-Require Import ZArith Nnat NAxioms NDiv NSig.
+Require Import ZArith Nnat Ndigits NAxioms NDiv NSig.
(** * The interface [NSig.NType] implies the interface [NAxiomsSig] *)
@@ -17,7 +17,8 @@ Hint Rewrite
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_eq_bool spec_sqrt
spec_log2 spec_max spec_min spec_pow_pos spec_pow_N spec_pow
- spec_even spec_odd
+ 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, to_N; repeat red; intros; nsimpl; congruence.
@@ -219,7 +220,7 @@ Qed.
Lemma pow_N_pow : forall a b, pow_N a b == a^(of_N b).
Proof.
- intros. zify. f_equal. symmetry. apply spec_of_N.
+ intros. now zify.
Qed.
Lemma pow_pos_N : forall a p, pow_pos a p == pow_N a (Npos p).
@@ -266,7 +267,7 @@ Proof.
rewrite Zeven_bool_iff, Zeven_ex_iff.
split; intros (m,Hm).
exists (of_N (Zabs_N m)).
- zify. rewrite spec_of_N, Z_of_N_abs, Zabs_eq; auto.
+ zify. rewrite Z_of_N_abs, Zabs_eq; trivial.
generalize (spec_pos n); auto with zarith.
exists [m]. revert Hm. now zify.
Qed.
@@ -277,7 +278,7 @@ Proof.
rewrite Zodd_bool_iff, Zodd_ex_iff.
split; intros (m,Hm).
exists (of_N (Zabs_N m)).
- zify. rewrite spec_of_N, Z_of_N_abs, Zabs_eq; auto.
+ zify. rewrite Z_of_N_abs, Zabs_eq; trivial.
generalize (spec_pos n); auto with zarith.
exists [m]. revert Hm. now zify.
Qed.
@@ -308,7 +309,7 @@ Proof.
intros n m. split.
intros (p,H). exists [p]. revert H; now zify.
intros (z,H). exists (of_N (Zabs_N z)). zify.
- rewrite spec_of_N, Z_of_N_abs.
+ rewrite Z_of_N_abs.
rewrite <- (Zabs_eq [n]) by apply spec_pos.
rewrite <- Zabs_Zmult, H.
apply Zabs_eq, spec_pos.
@@ -334,6 +335,82 @@ Proof.
intros. zify. apply Zgcd_nonneg.
Qed.
+(** Bitwise operations *)
+
+Lemma testbit_spec : forall a n, 0<=n ->
+ exists l, exists h, (0<=l /\ l<2^n) /\
+ a == l + ((if testbit a n then 1 else 0) + 2*h)*2^n.
+Proof.
+ intros a n _. zify.
+ assert (Ha := spec_pos a).
+ assert (Hn := spec_pos n).
+ destruct (Ntestbit_spec (Zabs_N [a]) (Zabs_N [n])) as (l & h & (_,Hl) & EQ).
+ exists (of_N l), (of_N h).
+ zify.
+ apply Z_of_N_lt in Hl.
+ apply Z_of_N_eq in EQ.
+ revert Hl EQ.
+ rewrite <- Ztestbit_of_N.
+ rewrite Z_of_N_plus, Z_of_N_mult, <- !Zpower_Npow, Z_of_N_plus,
+ Z_of_N_mult, !Z_of_N_abs, !Zabs_eq by trivial.
+ simpl (Z_of_N 2).
+ repeat split; trivial using Z_of_N_le_0.
+ destruct Ztestbit; now zify.
+Qed.
+
+Lemma testbit_neg_r : forall a n, n<0 -> testbit a n = false.
+Proof.
+ intros a n. zify. apply Ztestbit_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 Zshiftr_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 Zmax_r by auto with zarith.
+ now apply Zshiftl_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 Zshiftl_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 Zand_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 Zor_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 Zdiff_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 Zxor_spec.
+Qed.
+
+Lemma div2_spec : forall a, div2 a == shiftr a 1.
+Proof.
+ intros a. zify. now apply Zdiv2'_spec.
+Qed.
+
(** Recursion *)
Definition recursion (A : Type) (a : A) (f : N.t -> A -> A) (n : N.t) :=
diff --git a/theories/Numbers/vo.itarget b/theories/Numbers/vo.itarget
index d077f5af7..baefbd252 100644
--- a/theories/Numbers/vo.itarget
+++ b/theories/Numbers/vo.itarget
@@ -31,6 +31,7 @@ Integer/Abstract/ZParity.vo
Integer/Abstract/ZPow.vo
Integer/Abstract/ZGcd.vo
Integer/Abstract/ZLcm.vo
+Integer/Abstract/ZBits.vo
Integer/Abstract/ZProperties.vo
Integer/BigZ/BigZ.vo
Integer/BigZ/ZMake.vo
@@ -48,11 +49,13 @@ NatInt/NZMul.vo
NatInt/NZOrder.vo
NatInt/NZProperties.vo
NatInt/NZDomain.vo
+NatInt/NZParity.vo
NatInt/NZDiv.vo
NatInt/NZPow.vo
NatInt/NZSqrt.vo
NatInt/NZLog.vo
NatInt/NZGcd.vo
+NatInt/NZBits.vo
Natural/Abstract/NAddOrder.vo
Natural/Abstract/NAdd.vo
Natural/Abstract/NAxioms.vo
@@ -72,6 +75,7 @@ Natural/Abstract/NSqrt.vo
Natural/Abstract/NLog.vo
Natural/Abstract/NGcd.vo
Natural/Abstract/NLcm.vo
+Natural/Abstract/NBits.vo
Natural/BigN/BigN.vo
Natural/BigN/Nbasic.vo
Natural/BigN/NMake_gen.vo
diff --git a/theories/PArith/BinPos.v b/theories/PArith/BinPos.v
index cb6030e26..988a9d0d3 100644
--- a/theories/PArith/BinPos.v
+++ b/theories/PArith/BinPos.v
@@ -255,6 +255,15 @@ Definition Pdiv2 (z:positive) :=
Infix "/" := Pdiv2 : positive_scope.
+(** Division by 2 rounded up *)
+
+Definition Pdiv2_up p :=
+ match p with
+ | 1 => 1
+ | p~0 => p
+ | p~1 => Psucc p
+ end.
+
(** Number of digits in a positive number *)
Fixpoint Psize (p:positive) : nat :=
@@ -1292,6 +1301,17 @@ Proof.
apply Plt_trans with (p+q); auto using Plt_plus_r.
Qed.
+Lemma Ppow_gt_1 : forall n p, 1<n -> 1<n^p.
+Proof.
+ intros n p Hn.
+ induction p using Pind.
+ now rewrite Ppow_1_r.
+ rewrite Ppow_succ_r.
+ apply Plt_trans with (n*1).
+ now rewrite Pmult_1_r.
+ now apply Pmult_lt_mono_l.
+Qed.
+
(**********************************************************************)
(** Properties of subtraction on binary positive numbers *)
diff --git a/theories/Structures/Equalities.v b/theories/Structures/Equalities.v
index 14d34f1eb..8e72bd611 100644
--- a/theories/Structures/Equalities.v
+++ b/theories/Structures/Equalities.v
@@ -7,6 +7,7 @@
(***********************************************************************)
Require Export RelationClasses.
+Require Import Bool Morphisms Setoid.
Set Implicit Arguments.
Unset Strict Implicit.
@@ -165,6 +166,18 @@ Module Dec2Bool (E:DecidableType) <: BooleanDecidableType
Module Bool2Dec (E:BooleanEqualityType) <: BooleanDecidableType
:= E <+ HasEqBool2Dec.
+(** In a BooleanEqualityType, [eqb] is compatible with [eq] *)
+
+Module BoolEqualityFacts (Import E : BooleanEqualityType).
+
+Instance eqb_compat : Proper (E.eq ==> E.eq ==> Logic.eq) eqb.
+Proof.
+intros x x' Exx' y y' Eyy'.
+apply eq_true_iff_eq.
+now rewrite 2 eqb_eq, Exx', Eyy'.
+Qed.
+
+End BoolEqualityFacts.
(** * UsualDecidableType
diff --git a/theories/Structures/EqualitiesFacts.v b/theories/Structures/EqualitiesFacts.v
index d9b1d76fd..d8a1b7581 100644
--- a/theories/Structures/EqualitiesFacts.v
+++ b/theories/Structures/EqualitiesFacts.v
@@ -8,21 +8,8 @@
Require Import Equalities Bool SetoidList RelationPairs.
-(** In a BooleanEqualityType, [eqb] is compatible with [eq] *)
-
-Module BoolEqualityFacts (Import E : BooleanEqualityType).
-
-Instance eqb_compat : Proper (E.eq ==> E.eq ==> Logic.eq) eqb.
-Proof.
-intros x x' Exx' y y' Eyy'.
-apply eq_true_iff_eq.
-rewrite 2 eqb_eq, Exx', Eyy'; auto with *.
-Qed.
-
-End BoolEqualityFacts.
-
-
(** * Keys and datas used in FMap *)
+
Module KeyDecidableType(Import D:DecidableType).
Section Elt.
diff --git a/theories/ZArith/ZArith.v b/theories/ZArith/ZArith.v
index 26d700773..e532e4ad9 100644
--- a/theories/ZArith/ZArith.v
+++ b/theories/ZArith/ZArith.v
@@ -12,7 +12,8 @@ Require Export ZArith_base.
(** Extra definitions *)
-Require Export Zpow_def Zsqrt_def Zlog_def Zgcd_def Zdiv_def.
+Require Export
+ Zpow_def Zsqrt_def Zlog_def Zgcd_def Zdiv_def Zdigits_def.
(** Extra modules using [Omega] or [Ring]. *)
diff --git a/theories/ZArith/Zdigits_def.v b/theories/ZArith/Zdigits_def.v
new file mode 100644
index 000000000..a31ef8c98
--- /dev/null
+++ b/theories/ZArith/Zdigits_def.v
@@ -0,0 +1,420 @@
+(************************************************************************)
+(* 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 *)
+(************************************************************************)
+
+(** Bitwise operations for ZArith *)
+
+Require Import Bool BinPos BinNat BinInt Ndigits Znat Zeven Zpow_def
+ Zorder Zcompare.
+
+Local Open Scope Z_scope.
+
+(** When accessing the bits of negative numbers, all functions
+ below will use the two's complement representation. For instance,
+ [-1] will correspond to an infinite stream of true bits. If this
+ isn't what you're looking for, you can use [Zabs] first and then
+ access the bits of the absolute value.
+*)
+
+(** [Ztestbit] : accessing the [n]-th bit of a number [a].
+ For negative [n], we arbitrarily answer [false]. *)
+
+Definition Ztestbit a n :=
+ match n with
+ | 0 => Zodd_bool a
+ | Zpos p => match a with
+ | 0 => false
+ | Zpos a => Ptestbit a (Npos p)
+ | Zneg a => negb (Ntestbit (Ppred_N a) (Npos p))
+ end
+ | Zneg _ => false
+ end.
+
+(** Shifts
+
+ Nota: a shift to the right by [-n] will be a shift to the left
+ by [n], and vice-versa.
+
+ For fulfilling the two's complement convention, shifting to
+ the right a negative number should correspond to a division
+ by 2 with rounding toward bottom, hence the use of [Zdiv2']
+ instead of [Zdiv2].
+*)
+
+Definition Zshiftl a n :=
+ match n with
+ | 0 => a
+ | Zpos p => iter_pos p _ (Zmult 2) a
+ | Zneg p => iter_pos p _ Zdiv2' a
+ end.
+
+Definition Zshiftr a n := Zshiftl a (-n).
+
+(** Bitwise operations Zor Zand Zdiff Zxor *)
+
+Definition Zor a b :=
+ match a, b with
+ | 0, _ => b
+ | _, 0 => a
+ | Zpos a, Zpos b => Zpos (Por a b)
+ | Zneg a, Zpos b => Zneg (Nsucc_pos (Ndiff (Ppred_N a) (Npos b)))
+ | Zpos a, Zneg b => Zneg (Nsucc_pos (Ndiff (Ppred_N b) (Npos a)))
+ | Zneg a, Zneg b => Zneg (Nsucc_pos (Nand (Ppred_N a) (Ppred_N b)))
+ end.
+
+Definition Zand a b :=
+ match a, b with
+ | 0, _ => 0
+ | _, 0 => 0
+ | Zpos a, Zpos b => Z_of_N (Pand a b)
+ | Zneg a, Zpos b => Z_of_N (Ndiff (Npos b) (Ppred_N a))
+ | Zpos a, Zneg b => Z_of_N (Ndiff (Npos a) (Ppred_N b))
+ | Zneg a, Zneg b => Zneg (Nsucc_pos (Nor (Ppred_N a) (Ppred_N b)))
+ end.
+
+Definition Zdiff a b :=
+ match a, b with
+ | 0, _ => 0
+ | _, 0 => a
+ | Zpos a, Zpos b => Z_of_N (Pdiff a b)
+ | Zneg a, Zpos b => Zneg (Nsucc_pos (Nor (Ppred_N a) (Npos b)))
+ | Zpos a, Zneg b => Z_of_N (Nand (Npos a) (Ppred_N b))
+ | Zneg a, Zneg b => Z_of_N (Ndiff (Ppred_N b) (Ppred_N a))
+ end.
+
+Definition Zxor a b :=
+ match a, b with
+ | 0, _ => b
+ | _, 0 => a
+ | Zpos a, Zpos b => Z_of_N (Pxor a b)
+ | Zneg a, Zpos b => Zneg (Nsucc_pos (Nxor (Ppred_N a) (Npos b)))
+ | Zpos a, Zneg b => Zneg (Nsucc_pos (Nxor (Npos a) (Ppred_N b)))
+ | Zneg a, Zneg b => Z_of_N (Nxor (Ppred_N a) (Ppred_N b))
+ end.
+
+(** Proofs of specifications *)
+
+Lemma Zdiv2'_spec : forall a, Zdiv2' a = Zshiftr a 1.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma Ztestbit_neg_r : forall a n, n<0 -> Ztestbit a n = false.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma Ztestbit_spec : forall a n, 0<=n ->
+ exists l, exists h, 0<=l<2^n /\
+ a = l + ((if Ztestbit a n then 1 else 0) + 2*h)*2^n.
+Proof.
+ intros a [ |n|n] Hn; (now destruct Hn) || clear Hn.
+ (* n = 0 *)
+ simpl Ztestbit.
+ exists 0. exists (Zdiv2' a). repeat split. easy.
+ now rewrite Zplus_0_l, Zmult_1_r, Zplus_comm, <- Zdiv2'_odd.
+ (* n > 0 *)
+ destruct a.
+ (* ... a = 0 *)
+ exists 0. exists 0. repeat split; try easy. now rewrite Zpower_Ppow.
+ (* ... a > 0 *)
+ simpl Ztestbit.
+ destruct (Ntestbit_spec (Npos p) (Npos n)) as (l & h & (_,Hl) & EQ).
+ exists (Z_of_N l). exists (Z_of_N h).
+ simpl Npow in *; simpl Ntestbit in *. rewrite Zpower_Ppow.
+ repeat split.
+ apply Z_of_N_le_0.
+ rewrite <-Z_of_N_pos. now apply Z_of_N_lt.
+ rewrite <-Z_of_N_pos, EQ.
+ rewrite Z_of_N_plus, Z_of_N_mult. f_equal. f_equal.
+ destruct Ptestbit; now rewrite Z_of_N_plus, Z_of_N_mult.
+ (* ... a < 0 *)
+ simpl Ztestbit.
+ destruct (Ntestbit_spec (Ppred_N p) (Npos n)) as (l & h & (_,Hl) & EQ).
+ exists (2^Zpos n - (Z_of_N l) - 1). exists (-Z_of_N h - 1).
+ simpl Npow in *. rewrite Zpower_Ppow.
+ repeat split.
+ apply Zle_minus_le_0.
+ change 1 with (Zsucc 0). apply Zle_succ_l.
+ apply Zlt_plus_swap. simpl. rewrite <-Z_of_N_pos. now apply Z_of_N_lt.
+ apply Zlt_plus_swap. unfold Zminus. rewrite Zopp_involutive.
+ fold (Zsucc (Zpos (2^n))). apply Zlt_succ_r.
+ apply Zle_plus_swap. unfold Zminus. rewrite Zopp_involutive.
+ rewrite <- (Zplus_0_r (Zpos (2^n))) at 1. apply Zplus_le_compat_l.
+ apply Z_of_N_le_0.
+ apply Zopp_inj. unfold Zminus.
+ rewrite Zopp_neg, !Zopp_plus_distr, !Zopp_involutive.
+ rewrite Zopp_mult_distr_l, Zopp_plus_distr, Zopp_mult_distr_r,
+ Zopp_plus_distr, !Zopp_involutive.
+ rewrite Ppred_N_spec in EQ at 1.
+ apply (f_equal Nsucc) in EQ. rewrite Nsucc_pred in EQ by easy.
+ rewrite <-Z_of_N_pos, EQ.
+ rewrite Z_of_N_succ, Z_of_N_plus, Z_of_N_mult, Z_of_N_pos. unfold Zsucc.
+ rewrite <- (Zplus_assoc _ 1), (Zplus_comm 1), Zplus_assoc. f_equal.
+ rewrite (Zplus_comm (- _)), <- Zplus_assoc. f_equal.
+ change (- Zpos (2^n)) with ((-1)*(Zpos (2^n))). rewrite <- Zmult_plus_distr_l.
+ f_equal.
+ rewrite Z_of_N_plus, Z_of_N_mult.
+ rewrite Zmult_plus_distr_r, Zmult_1_r, (Zplus_comm _ 2), !Zplus_assoc.
+ rewrite (Zplus_comm _ 2), Zplus_assoc. change (2+-1) with 1.
+ f_equal.
+ now destruct Ntestbit.
+Qed.
+
+(** Conversions between [Ztestbit] and [Ntestbit] *)
+
+Lemma Ztestbit_of_N : forall a n,
+ Ztestbit (Z_of_N a) (Z_of_N n) = Ntestbit a n.
+Proof.
+ intros [ |a] [ |n]; simpl; trivial. now destruct a.
+Qed.
+
+Lemma Ztestbit_of_N' : forall a n, 0<=n ->
+ Ztestbit (Z_of_N a) n = Ntestbit a (Zabs_N n).
+Proof.
+ intros. now rewrite <- Ztestbit_of_N, Z_of_N_abs, Zabs_eq.
+Qed.
+
+Lemma Ztestbit_Zpos : forall a n, 0<=n ->
+ Ztestbit (Zpos a) n = Ntestbit (Npos a) (Zabs_N n).
+Proof.
+ intros. now rewrite <- Ztestbit_of_N'.
+Qed.
+
+Lemma Ztestbit_Zneg : forall a n, 0<=n ->
+ Ztestbit (Zneg a) n = negb (Ntestbit (Ppred_N a) (Zabs_N n)).
+Proof.
+ intros a n Hn.
+ rewrite <- Ztestbit_of_N' by trivial.
+ destruct n as [ |n|n];
+ [ | simpl; now destruct (Ppred_N a) | now destruct Hn].
+ unfold Ztestbit.
+ replace (Z_of_N (Ppred_N a)) with (Zpred (Zpos a))
+ by (destruct a; trivial).
+ now rewrite Zodd_bool_pred, <- Zodd_even_bool.
+Qed.
+
+Lemma Ztestbit_0_l : forall n, Ztestbit 0 n = false.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma Ppred_div2_up : forall p,
+ Ppred_N (Pdiv2_up p) = Ndiv2 (Ppred_N p).
+Proof.
+ intros [p|p| ]; trivial.
+ simpl. rewrite Ppred_N_spec. apply (Npred_succ (Npos p)).
+ destruct p; simpl; trivial.
+Qed.
+
+Lemma Z_of_N_div2' : forall n, Z_of_N (Ndiv2 n) = Zdiv2' (Z_of_N n).
+Proof.
+ intros [|p]; trivial. now destruct p.
+Qed.
+
+Lemma Z_of_N_div2 : forall n, Z_of_N (Ndiv2 n) = Zdiv2 (Z_of_N n).
+Proof.
+ intros [|p]; trivial. now destruct p.
+Qed.
+
+(** Auxiliary results about right shift on positive numbers *)
+
+Lemma Ppred_Pshiftl_low : forall p n m, (m<n)%nat ->
+ Nbit (Ppred_N (iter_nat n _ xO p)) m = true.
+Proof.
+ induction n.
+ inversion 1.
+ intros m H. simpl.
+ destruct m.
+ now destruct (iter_nat n _ xO p).
+ apply lt_S_n in H.
+ specialize (IHn m H).
+ now destruct (iter_nat n _ xO p).
+Qed.
+
+Lemma Ppred_Pshiftl_high : forall p n m, (n<=m)%nat ->
+ Nbit (Ppred_N (iter_nat n _ xO p)) m =
+ Nbit (Nshiftl_nat (Ppred_N p) n) m.
+Proof.
+ induction n.
+ now unfold Nshiftl_nat.
+ intros m H.
+ destruct m.
+ inversion H.
+ apply le_S_n in H.
+ rewrite Nshiftl_nat_S.
+ specialize (IHn m H).
+ simpl in *.
+ unfold Nbit.
+ now destruct (Nshiftl_nat (Ppred_N p) n), (iter_nat n positive xO p).
+Qed.
+
+(** Correctness proofs about [Zshiftr] and [Zshiftl] *)
+
+Lemma Zshiftr_spec_aux : forall a n m, 0<=n -> 0<=m ->
+ Ztestbit (Zshiftr a n) m = Ztestbit a (m+n).
+Proof.
+ intros a n m Hn Hm. unfold Zshiftr.
+ destruct n as [ |n|n]; (now destruct Hn) || clear Hn; simpl.
+ now rewrite Zplus_0_r.
+ destruct a as [ |a|a].
+ (* a = 0 *)
+ replace (iter_pos n _ Zdiv2' 0) with 0
+ by (apply iter_pos_invariant; intros; subst; trivial).
+ now rewrite 2 Ztestbit_0_l.
+ (* a > 0 *)
+ rewrite <- (Z_of_N_pos a) at 1.
+ rewrite <- (iter_pos_swap_gen _ _ _ Ndiv2) by exact Z_of_N_div2'.
+ rewrite Ztestbit_Zpos, Ztestbit_of_N'; trivial.
+ rewrite Zabs_N_plus; try easy. simpl Zabs_N.
+ exact (Nshiftr_spec (Npos a) (Npos n) (Zabs_N m)).
+ now apply Zplus_le_0_compat.
+ (* a < 0 *)
+ rewrite <- (iter_pos_swap_gen _ _ _ Pdiv2_up) by trivial.
+ rewrite 2 Ztestbit_Zneg; trivial. f_equal.
+ rewrite Zabs_N_plus; try easy. simpl Zabs_N.
+ rewrite (iter_pos_swap_gen _ _ _ _ Ndiv2) by exact Ppred_div2_up.
+ exact (Nshiftr_spec (Ppred_N a) (Npos n) (Zabs_N m)).
+ now apply Zplus_le_0_compat.
+Qed.
+
+Lemma Zshiftl_spec_low : forall a n m, m<n ->
+ Ztestbit (Zshiftl a n) m = false.
+Proof.
+ intros a [ |n|n] [ |m|m] H; try easy; simpl Zshiftl.
+ rewrite iter_nat_of_P.
+ destruct (nat_of_P_is_S n) as (n' & ->).
+ simpl. now destruct (iter_nat n' _ (Zmult 2) a).
+ destruct a as [ |a|a].
+ (* a = 0 *)
+ replace (iter_pos n _ (Zmult 2) 0) with 0
+ by (apply iter_pos_invariant; intros; subst; trivial).
+ apply Ztestbit_0_l.
+ (* a > 0 *)
+ rewrite <- (iter_pos_swap_gen _ _ _ xO) by trivial.
+ rewrite Ztestbit_Zpos by easy.
+ exact (Nshiftl_spec_low (Npos a) (Npos n) (Npos m) H).
+ (* a < 0 *)
+ rewrite <- (iter_pos_swap_gen _ _ _ xO) by trivial.
+ rewrite Ztestbit_Zneg by easy.
+ simpl Zabs_N.
+ rewrite <- Nbit_Ntestbit, iter_nat_of_P. simpl nat_of_N.
+ rewrite Ppred_Pshiftl_low; trivial.
+ now apply Plt_lt.
+Qed.
+
+Lemma Zshiftl_spec_high : forall a n m, 0<=m -> n<=m ->
+ Ztestbit (Zshiftl a n) m = Ztestbit a (m-n).
+Proof.
+ intros a n m Hm H.
+ destruct n as [ |n|n]. simpl. now rewrite Zminus_0_r.
+ (* n > 0 *)
+ destruct m as [ |m|m]; try (now destruct H).
+ assert (0 <= Zpos m - Zpos n) by (now apply Zle_minus_le_0).
+ assert (EQ : Zabs_N (Zpos m - Zpos n) = (Npos m - Npos n)%N).
+ apply Z_of_N_eq_rev. rewrite Z_of_N_abs, Zabs_eq by trivial.
+ now rewrite Z_of_N_minus, !Z_of_N_pos, Zmax_r.
+ destruct a; unfold Zshiftl.
+ (* ... a = 0 *)
+ replace (iter_pos n _ (Zmult 2) 0) with 0
+ by (apply iter_pos_invariant; intros; subst; trivial).
+ now rewrite 2 Ztestbit_0_l.
+ (* ... a > 0 *)
+ rewrite <- (iter_pos_swap_gen _ _ _ xO) by trivial.
+ rewrite 2 Ztestbit_Zpos, EQ by easy.
+ exact (Nshiftl_spec_high (Npos p) (Npos n) (Npos m) H).
+ (* ... a < 0 *)
+ rewrite <- (iter_pos_swap_gen _ _ _ xO) by trivial.
+ rewrite 2 Ztestbit_Zneg, EQ by easy. f_equal.
+ simpl Zabs_N.
+ rewrite <- Nshiftl_spec_high by easy.
+ rewrite <- 2 Nbit_Ntestbit, iter_nat_of_P, <- Nshiftl_nat_equiv.
+ simpl nat_of_N.
+ apply Ppred_Pshiftl_high.
+ now apply Ple_le.
+ (* n < 0 *)
+ unfold Zminus. simpl.
+ now apply (Zshiftr_spec_aux a (Zpos n) m).
+Qed.
+
+Lemma Zshiftr_spec : forall a n m, 0<=m ->
+ Ztestbit (Zshiftr a n) m = Ztestbit a (m+n).
+Proof.
+ intros a n m Hm.
+ destruct (Zle_or_lt 0 n).
+ now apply Zshiftr_spec_aux.
+ destruct (Zle_or_lt (-n) m).
+ unfold Zshiftr.
+ rewrite (Zshiftl_spec_high a (-n) m); trivial.
+ unfold Zminus. now rewrite Zopp_involutive.
+ unfold Zshiftr.
+ rewrite (Zshiftl_spec_low a (-n) m); trivial.
+ rewrite Ztestbit_neg_r; trivial.
+ now apply Zlt_plus_swap.
+Qed.
+
+(** Correctness proofs for bitwise operations *)
+
+Lemma Zor_spec : forall a b n,
+ Ztestbit (Zor a b) n = Ztestbit a n || Ztestbit b n.
+Proof.
+ intros a b n.
+ destruct (Zle_or_lt 0 n) as [Hn|Hn]; [|now rewrite !Ztestbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?Ztestbit_0_l, ?orb_false_r; trivial; unfold Zor;
+ rewrite ?Ztestbit_Zpos, ?Ztestbit_Zneg, ?Ppred_Nsucc
+ by trivial.
+ now rewrite <- Nor_spec.
+ now rewrite Ndiff_spec, negb_andb, negb_involutive, orb_comm.
+ now rewrite Ndiff_spec, negb_andb, negb_involutive.
+ now rewrite Nand_spec, negb_andb.
+Qed.
+
+Lemma Zand_spec : forall a b n,
+ Ztestbit (Zand a b) n = Ztestbit a n && Ztestbit b n.
+Proof.
+ intros a b n.
+ destruct (Zle_or_lt 0 n) as [Hn|Hn]; [|now rewrite !Ztestbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?Ztestbit_0_l, ?andb_false_r; trivial; unfold Zand;
+ rewrite ?Ztestbit_Zpos, ?Ztestbit_Zneg, ?Ztestbit_of_N', ?Ppred_Nsucc
+ by trivial.
+ now rewrite <- Nand_spec.
+ now rewrite Ndiff_spec.
+ now rewrite Ndiff_spec, andb_comm.
+ now rewrite Nor_spec, negb_orb.
+Qed.
+
+Lemma Zdiff_spec : forall a b n,
+ Ztestbit (Zdiff a b) n = Ztestbit a n && negb (Ztestbit b n).
+Proof.
+ intros a b n.
+ destruct (Zle_or_lt 0 n) as [Hn|Hn]; [|now rewrite !Ztestbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?Ztestbit_0_l, ?andb_true_r; trivial; unfold Zdiff;
+ rewrite ?Ztestbit_Zpos, ?Ztestbit_Zneg, ?Ztestbit_of_N', ?Ppred_Nsucc
+ by trivial.
+ now rewrite <- Ndiff_spec.
+ now rewrite Nand_spec, negb_involutive.
+ now rewrite Nor_spec, negb_orb.
+ now rewrite Ndiff_spec, negb_involutive, andb_comm.
+Qed.
+
+Lemma Zxor_spec : forall a b n,
+ Ztestbit (Zxor a b) n = xorb (Ztestbit a n) (Ztestbit b n).
+Proof.
+ intros a b n.
+ destruct (Zle_or_lt 0 n) as [Hn|Hn]; [|now rewrite !Ztestbit_neg_r].
+ destruct a as [ |a|a], b as [ |b|b];
+ rewrite ?Ztestbit_0_l, ?xorb_false_l, ?xorb_false_r; trivial; unfold Zxor;
+ rewrite ?Ztestbit_Zpos, ?Ztestbit_Zneg, ?Ztestbit_of_N', ?Ppred_Nsucc
+ by trivial.
+ now rewrite <- Nxor_spec.
+ now rewrite Nxor_spec, negb_xorb_r.
+ now rewrite Nxor_spec, negb_xorb_l.
+ now rewrite Nxor_spec, xorb_negb_negb.
+Qed.
diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v
index a14f29308..c9397db8b 100644
--- a/theories/ZArith/Zdiv.v
+++ b/theories/ZArith/Zdiv.v
@@ -356,7 +356,7 @@ Proof. intros a b. zero_or_not b. intros; rewrite Z.div_opp_r_nz; auto. Qed.
(** Cancellations. *)
-Lemma Zdiv_mult_cancel_r : forall a b c:Z,
+Lemma Zdiv_mult_cancel_r : forall a b c:Z,
c <> 0 -> (a*c)/(b*c) = a/b.
Proof. intros. zero_or_not b. apply Z.div_mul_cancel_r; auto. Qed.
@@ -521,6 +521,33 @@ Proof.
split; intros (c,Hc); exists c; auto.
Qed.
+(** Particular case : dividing by 2 is related with parity *)
+
+Lemma Zdiv2'_div : forall a, Zdiv2' a = a/2.
+Proof.
+ apply Z.div2_div.
+Qed.
+
+Lemma Zmod_odd : forall a, a mod 2 = if Zodd_bool a then 1 else 0.
+Proof.
+ intros a. now rewrite <- Z.bit0_odd, <- Z.bit0_mod.
+Qed.
+
+Lemma Zmod_even : forall a, a mod 2 = if Zeven_bool a then 0 else 1.
+Proof.
+ intros a. rewrite Zmod_odd, Zodd_even_bool. now destruct Zeven_bool.
+Qed.
+
+Lemma Zodd_mod : forall a, Zodd_bool a = Zeq_bool (a mod 2) 1.
+Proof.
+ intros a. rewrite Zmod_odd. now destruct Zodd_bool.
+Qed.
+
+Lemma Zeven_mod : forall a, Zeven_bool a = Zeq_bool (a mod 2) 0.
+Proof.
+ intros a. rewrite Zmod_even. now destruct Zeven_bool.
+Qed.
+
(** * Compatibility *)
(** Weaker results kept only for compatibility *)
diff --git a/theories/ZArith/Zeven.v b/theories/ZArith/Zeven.v
index 74e2e6fbf..70ab4dacc 100644
--- a/theories/ZArith/Zeven.v
+++ b/theories/ZArith/Zeven.v
@@ -59,128 +59,183 @@ Proof.
destruct n as [|p|p]; try destruct p; simpl in *; split; easy.
Qed.
+Lemma Zodd_even_bool : forall n, Zodd_bool n = negb (Zeven_bool n).
+Proof.
+ destruct n as [|p|p]; trivial; now destruct p.
+Qed.
+
+Lemma Zeven_odd_bool : forall n, Zeven_bool n = negb (Zodd_bool n).
+Proof.
+ destruct n as [|p|p]; trivial; now destruct p.
+Qed.
+
Definition Zeven_odd_dec : forall z:Z, {Zeven z} + {Zodd z}.
Proof.
intro z. case z;
- [ left; compute in |- *; trivial
+ [ left; compute; trivial
| intro p; case p; intros;
- (right; compute in |- *; exact I) || (left; compute in |- *; exact I)
+ (right; compute; exact I) || (left; compute; exact I)
| intro p; case p; intros;
- (right; compute in |- *; exact I) || (left; compute in |- *; exact I) ].
+ (right; compute; exact I) || (left; compute; exact I) ].
Defined.
Definition Zeven_dec : forall z:Z, {Zeven z} + {~ Zeven z}.
Proof.
intro z. case z;
- [ left; compute in |- *; trivial
+ [ left; compute; trivial
| intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial)
+ (left; compute; exact I) || (right; compute; trivial)
| intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial) ].
+ (left; compute; exact I) || (right; compute; trivial) ].
Defined.
Definition Zodd_dec : forall z:Z, {Zodd z} + {~ Zodd z}.
Proof.
intro z. case z;
- [ right; compute in |- *; trivial
+ [ right; compute; trivial
| intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial)
+ (left; compute; exact I) || (right; compute; trivial)
| intro p; case p; intros;
- (left; compute in |- *; exact I) || (right; compute in |- *; trivial) ].
+ (left; compute; exact I) || (right; compute; trivial) ].
Defined.
Lemma Zeven_not_Zodd : forall n:Z, Zeven n -> ~ Zodd n.
Proof.
- intro z; destruct z; [ idtac | destruct p | destruct p ]; compute in |- *;
+ intro z; destruct z; [ idtac | destruct p | destruct p ]; compute;
trivial.
Qed.
Lemma Zodd_not_Zeven : forall n:Z, Zodd n -> ~ Zeven n.
Proof.
- intro z; destruct z; [ idtac | destruct p | destruct p ]; compute in |- *;
+ intro z; destruct z; [ idtac | destruct p | destruct p ]; compute;
trivial.
Qed.
Lemma Zeven_Sn : forall n:Z, Zodd n -> Zeven (Zsucc n).
Proof.
- intro z; destruct z; unfold Zsucc in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
+ intro z; destruct z; unfold Zsucc;
+ [ idtac | destruct p | destruct p ]; simpl;
trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ unfold Pdouble_minus_one; case p; simpl; auto.
Qed.
Lemma Zodd_Sn : forall n:Z, Zeven n -> Zodd (Zsucc n).
Proof.
- intro z; destruct z; unfold Zsucc in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
+ intro z; destruct z; unfold Zsucc;
+ [ idtac | destruct p | destruct p ]; simpl;
trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ unfold Pdouble_minus_one; case p; simpl; auto.
Qed.
Lemma Zeven_pred : forall n:Z, Zodd n -> Zeven (Zpred n).
Proof.
- intro z; destruct z; unfold Zpred in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
+ intro z; destruct z; unfold Zpred;
+ [ idtac | destruct p | destruct p ]; simpl;
trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ unfold Pdouble_minus_one; case p; simpl; auto.
Qed.
Lemma Zodd_pred : forall n:Z, Zeven n -> Zodd (Zpred n).
Proof.
- intro z; destruct z; unfold Zpred in |- *;
- [ idtac | destruct p | destruct p ]; simpl in |- *;
+ intro z; destruct z; unfold Zpred;
+ [ idtac | destruct p | destruct p ]; simpl;
trivial.
- unfold Pdouble_minus_one in |- *; case p; simpl in |- *; auto.
+ unfold Pdouble_minus_one; case p; simpl; auto.
Qed.
Hint Unfold Zeven Zodd: zarith.
+Lemma Zeven_bool_succ : forall n, Zeven_bool (Zsucc n) = Zodd_bool n.
+Proof.
+ destruct n as [ |p|p]; trivial; destruct p as [p|p| ]; trivial.
+ now destruct p.
+Qed.
+
+Lemma Zeven_bool_pred : forall n, Zeven_bool (Zpred n) = Zodd_bool n.
+Proof.
+ destruct n as [ |p|p]; trivial; destruct p as [p|p| ]; trivial.
+ now destruct p.
+Qed.
+
+Lemma Zodd_bool_succ : forall n, Zodd_bool (Zsucc n) = Zeven_bool n.
+Proof.
+ destruct n as [ |p|p]; trivial; destruct p as [p|p| ]; trivial.
+ now destruct p.
+Qed.
+
+Lemma Zodd_bool_pred : forall n, Zodd_bool (Zpred n) = Zeven_bool n.
+Proof.
+ destruct n as [ |p|p]; trivial; destruct p as [p|p| ]; trivial.
+ now destruct p.
+Qed.
(******************************************************************)
(** * Definition of [Zdiv2] and properties wrt [Zeven] and [Zodd] *)
(** [Zdiv2] is defined on all [Z], but notice that for odd negative
- integers it is not the euclidean quotient: in that case we have
- [n = 2*(n/2)-1] *)
+ integers we have [n = 2*(Zdiv2 n)-1], hence it does not
+ correspond to the usual Coq division [Zdiv], for which we would
+ have here [n = 2*(n/2)+1]. Since [Zdiv2] performs rounding
+ toward zero, it is rather a particular case of the alternative
+ division [Zquot].
+*)
Definition Zdiv2 (z:Z) :=
match z with
- | Z0 => 0
- | Zpos xH => 0
+ | 0 => 0
+ | Zpos 1 => 0
| Zpos p => Zpos (Pdiv2 p)
- | Zneg xH => 0
+ | Zneg 1 => 0
| Zneg p => Zneg (Pdiv2 p)
end.
+(** We also provide an alternative [Zdiv2'] performing round toward
+ bottom, and hence corresponding to [Zdiv]. *)
+
+Definition Zdiv2' a :=
+ match a with
+ | 0 => 0
+ | Zpos 1 => 0
+ | Zpos p => Zpos (Pdiv2 p)
+ | Zneg p => Zneg (Pdiv2_up p)
+ end.
+
+Lemma Zdiv2'_odd : forall a,
+ a = 2*(Zdiv2' a) + if Zodd_bool a then 1 else 0.
+Proof.
+ intros [ |[p|p| ]|[p|p| ]]; simpl; trivial.
+ f_equal. now rewrite xO_succ_permute, <-Ppred_minus, Ppred_succ.
+Qed.
+
Lemma Zeven_div2 : forall n:Z, Zeven n -> n = 2 * Zdiv2 n.
Proof.
intro x; destruct x.
auto with arith.
destruct p; auto with arith.
- intros. absurd (Zeven (Zpos (xI p))); red in |- *; auto with arith.
- intros. absurd (Zeven 1); red in |- *; auto with arith.
+ intros. absurd (Zeven (Zpos (xI p))); red; auto with arith.
+ intros. absurd (Zeven 1); red; auto with arith.
destruct p; auto with arith.
- intros. absurd (Zeven (Zneg (xI p))); red in |- *; auto with arith.
- intros. absurd (Zeven (-1)); red in |- *; auto with arith.
+ intros. absurd (Zeven (Zneg (xI p))); red; auto with arith.
+ intros. absurd (Zeven (-1)); red; auto with arith.
Qed.
Lemma Zodd_div2 : forall n:Z, n >= 0 -> Zodd n -> n = 2 * Zdiv2 n + 1.
Proof.
intro x; destruct x.
- intros. absurd (Zodd 0); red in |- *; auto with arith.
+ intros. absurd (Zodd 0); red; auto with arith.
destruct p; auto with arith.
- intros. absurd (Zodd (Zpos (xO p))); red in |- *; auto with arith.
- intros. absurd (Zneg p >= 0); red in |- *; auto with arith.
+ intros. absurd (Zodd (Zpos (xO p))); red; auto with arith.
+ intros. absurd (Zneg p >= 0); red; auto with arith.
Qed.
Lemma Zodd_div2_neg :
forall n:Z, n <= 0 -> Zodd n -> n = 2 * Zdiv2 n - 1.
Proof.
intro x; destruct x.
- intros. absurd (Zodd 0); red in |- *; auto with arith.
- intros. absurd (Zneg p >= 0); red in |- *; auto with arith.
+ intros. absurd (Zodd 0); red; auto with arith.
+ intros. absurd (Zneg p >= 0); red; auto with arith.
destruct p; auto with arith.
- intros. absurd (Zodd (Zneg (xO p))); red in |- *; auto with arith.
+ intros. absurd (Zodd (Zneg (xO p))); red; auto with arith.
Qed.
Lemma Z_modulo_2 :
@@ -192,10 +247,10 @@ Proof.
right. generalize b; clear b; case x.
intro b; inversion b.
intro p; split with (Zdiv2 (Zpos p)). apply (Zodd_div2 (Zpos p)); trivial.
- unfold Zge, Zcompare in |- *; simpl in |- *; discriminate.
+ unfold Zge, Zcompare; simpl; discriminate.
intro p; split with (Zdiv2 (Zpred (Zneg p))).
- pattern (Zneg p) at 1 in |- *; rewrite (Zsucc_pred (Zneg p)).
- pattern (Zpred (Zneg p)) at 1 in |- *; rewrite (Zeven_div2 (Zpred (Zneg p))).
+ pattern (Zneg p) at 1; rewrite (Zsucc_pred (Zneg p)).
+ pattern (Zpred (Zneg p)) at 1; rewrite (Zeven_div2 (Zpred (Zneg p))).
reflexivity.
apply Zeven_pred; assumption.
Qed.
diff --git a/theories/ZArith/Znat.v b/theories/ZArith/Znat.v
index 5f4a6e38d..8f4a69b1e 100644
--- a/theories/ZArith/Znat.v
+++ b/theories/ZArith/Znat.v
@@ -13,7 +13,7 @@ Require Export Arith_base.
Require Import BinPos BinInt BinNat Zcompare Zorder.
Require Import Decidable Peano_dec Min Max Compare_dec.
-Open Local Scope Z_scope.
+Local Open Scope Z_scope.
Definition neq (x y:nat) := x <> y.
@@ -341,3 +341,51 @@ Proof.
intros. unfold Zmax, Nmax. rewrite Z_of_N_compare.
case Ncompare_spec; intros; subst; trivial.
Qed.
+
+(** Results about the [Zabs_N] function, converting from Z to N *)
+
+Lemma Zabs_of_N : forall n, Zabs_N (Z_of_N n) = n.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma Zabs_N_succ_abs : forall n,
+ Zabs_N (Zsucc (Zabs n)) = Nsucc (Zabs_N n).
+Proof.
+ intros [ |n|n]; simpl; trivial; now rewrite Pplus_one_succ_r.
+Qed.
+
+Lemma Zabs_N_succ : forall n, 0<=n ->
+ Zabs_N (Zsucc n) = Nsucc (Zabs_N n).
+Proof.
+ intros n Hn. rewrite <- Zabs_N_succ_abs. repeat f_equal.
+ symmetry; now apply Zabs_eq.
+Qed.
+
+Lemma Zabs_N_plus_abs : forall n m,
+ Zabs_N (Zabs n + Zabs m) = (Zabs_N n + Zabs_N m)%N.
+Proof.
+ intros [ |n|n] [ |m|m]; simpl; trivial.
+Qed.
+
+Lemma Zabs_N_plus : forall n m, 0<=n -> 0<=m ->
+ Zabs_N (n + m) = (Zabs_N n + Zabs_N m)%N.
+Proof.
+ intros n m Hn Hm.
+ rewrite <- Zabs_N_plus_abs; repeat f_equal;
+ symmetry; now apply Zabs_eq.
+Qed.
+
+Lemma Zabs_N_mult_abs : forall n m,
+ Zabs_N (Zabs n * Zabs m) = (Zabs_N n * Zabs_N m)%N.
+Proof.
+ intros [ |n|n] [ |m|m]; simpl; trivial.
+Qed.
+
+Lemma Zabs_N_mult : forall n m, 0<=n -> 0<=m ->
+ Zabs_N (n * m) = (Zabs_N n * Zabs_N m)%N.
+Proof.
+ intros n m Hn Hm.
+ rewrite <- Zabs_N_mult_abs; repeat f_equal;
+ symmetry; now apply Zabs_eq.
+Qed.
diff --git a/theories/ZArith/Zquot.v b/theories/ZArith/Zquot.v
index 4f1c94e99..5fe105aa5 100644
--- a/theories/ZArith/Zquot.v
+++ b/theories/ZArith/Zquot.v
@@ -471,6 +471,56 @@ Proof.
rewrite Z.rem_divide; trivial. split; intros (c,Hc); exists c; auto.
Qed.
+(** Particular case : dividing by 2 is related with parity *)
+
+Lemma Zdiv2_odd_eq : forall a,
+ a = 2 * Zdiv2 a + if Zodd_bool a then Zsgn a else 0.
+Proof.
+ destruct a as [ |p|p]; try destruct p; trivial.
+Qed.
+
+Lemma Zdiv2_odd_remainder : forall a,
+ Remainder a 2 (if Zodd_bool a then Zsgn a else 0).
+Proof.
+ intros [ |p|p]. simpl.
+ left. simpl. auto with zarith.
+ left. destruct p; simpl; auto with zarith.
+ right. destruct p; simpl; split; now auto with zarith.
+Qed.
+
+Lemma Zdiv2_quot : forall a, Zdiv2 a = a÷2.
+Proof.
+ intros.
+ apply Zquot_unique_full with (if Zodd_bool a then Zsgn a else 0).
+ apply Zdiv2_odd_remainder.
+ apply Zdiv2_odd_eq.
+Qed.
+
+Lemma Zrem_odd : forall a, Zrem a 2 = if Zodd_bool a then Zsgn a else 0.
+Proof.
+ intros. symmetry.
+ apply Zrem_unique_full with (Zdiv2 a).
+ apply Zdiv2_odd_remainder.
+ apply Zdiv2_odd_eq.
+Qed.
+
+Lemma Zrem_even : forall a, Zrem a 2 = if Zeven_bool a then 0 else Zsgn a.
+Proof.
+ intros a. rewrite Zrem_odd, Zodd_even_bool. now destruct Zeven_bool.
+Qed.
+
+Lemma Zeven_rem : forall a, Zeven_bool a = Zeq_bool (Zrem a 2) 0.
+Proof.
+ intros a. rewrite Zrem_even.
+ destruct a as [ |p|p]; trivial; now destruct p.
+Qed.
+
+Lemma Zodd_rem : forall a, Zodd_bool a = negb (Zeq_bool (Zrem a 2) 0).
+Proof.
+ intros a. rewrite Zrem_odd.
+ destruct a as [ |p|p]; trivial; now destruct p.
+Qed.
+
(** * Interaction with "historic" Zdiv *)
(** They agree at least on positive numbers: *)
diff --git a/theories/ZArith/vo.itarget b/theories/ZArith/vo.itarget
index ef18d67c7..8dc8e9276 100644
--- a/theories/ZArith/vo.itarget
+++ b/theories/ZArith/vo.itarget
@@ -33,3 +33,4 @@ Zsqrt_def.vo
Zlog_def.vo
Zgcd_def.vo
Zeuclid.vo
+Zdigits_def.vo \ No newline at end of file