Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

Require Export BinNums BinPos Pnat.
Require Import BinNat Bool Equalities GenericMinMax
 OrdersFacts ZAxioms ZProperties.
Require BinIntDef.

(***********************************************************)

Binary Integers

(***********************************************************)
Initial author: Pierre Crégut, CNET, Lannion, France
The type Z and its constructors Z0 and Zpos and Zneg are now defined in BinNums.v
Local Open Scope Z_scope.
Every definitions and early properties about binary integers are placed in a module Z for qualification purpose.
Module Z
 <: ZAxiomsSig
 <: UsualOrderedTypeFull
 <: UsualDecidableTypeFull
 <: TotalOrder.

Definitions of operations, now in a separate file

Include BinIntDef.Z.

Register add as num.Z.add.
Register opp as num.Z.opp.
Register succ as num.Z.succ.
Register pred as num.Z.pred.
Register sub as num.Z.sub.
Register mul as num.Z.mul.
Register of_nat as num.Z.of_nat.
When including property functors, only inline t eq zero one two
Set Inline Level 30.

Logic Predicates

Definition eq := @Logic.eq Z.
Definition eq_equiv := @eq_equivalence Z.

Definition lt x y := (x ?= y) = Lt.
Definition gt x y := (x ?= y) = Gt.
Definition le x y := (x ?= y) <> Gt.
Definition ge x y := (x ?= y) <> Lt.

Infix "<=" := le : Z_scope.
Infix "<" := lt : Z_scope.
Infix ">=" := ge : Z_scope.
Infix ">" := gt : Z_scope.

Notation "x <= y <= z" := (x <= y /\ y <= z) : Z_scope.
Notation "x <= y < z" := (x <= y /\ y < z) : Z_scope.
Notation "x < y < z" := (x < y /\ y < z) : Z_scope.
Notation "x < y <= z" := (x < y /\ y <= z) : Z_scope.

Definition divide x y := exists z, y = z*x.
Notation "( x | y )" := (divide x y) (at level 0).

Definition Even a := exists b, a = 2*b.
Definition Odd a := exists b, a = 2*b+1.

Register le as num.Z.le.
Register lt as num.Z.lt.
Register ge as num.Z.ge.
Register gt as num.Z.gt.

Decidability of equality.

x, y:Z

{x = y} + {x <> y}
x, y:Z

{x = y} + {x <> y}
decide equality; apply Pos.eq_dec. Defined.

Proofs of morphisms, obvious since eq is Leibniz

Local Obligation Tactic := simpl_relation.
Definition succ_wd : Proper (eq==>eq) succ := _.
Definition pred_wd : Proper (eq==>eq) pred := _.
Definition opp_wd : Proper (eq==>eq) opp := _.
Definition add_wd : Proper (eq==>eq==>eq) add := _.
Definition sub_wd : Proper (eq==>eq==>eq) sub := _.
Definition mul_wd : Proper (eq==>eq==>eq) mul := _.
Definition lt_wd : Proper (eq==>eq==>iff) lt := _.
Definition div_wd : Proper (eq==>eq==>eq) div := _.
Definition mod_wd : Proper (eq==>eq==>eq) modulo := _.
Definition quot_wd : Proper (eq==>eq==>eq) quot := _.
Definition rem_wd : Proper (eq==>eq==>eq) rem := _.
Definition pow_wd : Proper (eq==>eq==>eq) pow := _.
Definition testbit_wd : Proper (eq==>eq==>Logic.eq) testbit := _.

Properties of pos_sub

pos_sub can be written in term of positive comparison and subtraction (cf. earlier definition of addition of Z)
p, q:positive

pos_sub p q = match (p ?= q)%positive with | Eq => 0 | Lt => neg (q - p) | Gt => pos (p - q) end
p, q:positive

pos_sub p q = match (p ?= q)%positive with | Eq => 0 | Lt => neg (q - p) | Gt => pos (p - q) end
p:positive

forall q : positive, pos_sub p q = match (p ?= q)%positive with | Eq => 0 | Lt => neg (q - p) | Gt => pos (p - q) end
induction p; destruct q; simpl; trivial; rewrite ?Pos.compare_xI_xI, ?Pos.compare_xO_xI, ?Pos.compare_xI_xO, ?Pos.compare_xO_xO, IHp; simpl; case Pos.compare_spec; intros; simpl; trivial; (now rewrite Pos.sub_xI_xI) || (now rewrite Pos.sub_xO_xO) || (now rewrite Pos.sub_xO_xI) || (now rewrite Pos.sub_xI_xO) || subst; unfold Pos.sub; simpl; now rewrite Pos.sub_mask_diag. Qed.
p, q:positive

match pos_sub p q with | 0 => p = q | pos k => p = (q + k)%positive | neg k => q = (p + k)%positive end
p, q:positive

match pos_sub p q with | 0 => p = q | pos k => p = (q + k)%positive | neg k => q = (p + k)%positive end
p, q:positive

match match (p ?= q)%positive with | Eq => 0 | Lt => neg (q - p) | Gt => pos (p - q) end with | 0 => p = q | pos k => p = (q + k)%positive | neg k => q = (p + k)%positive end
case Pos.compare_spec; auto; intros; now rewrite Pos.add_comm, Pos.sub_add. Qed.
Particular cases of the previous result
p:positive

pos_sub p p = 0
p:positive

pos_sub p p = 0
now rewrite pos_sub_spec, Pos.compare_refl. Qed.
p, q:positive

(p < q)%positive -> pos_sub p q = neg (q - p)
p, q:positive

(p < q)%positive -> pos_sub p q = neg (q - p)
p, q:positive
H:(p < q)%positive

pos_sub p q = neg (q - p)
now rewrite pos_sub_spec, H. Qed.
p, q:positive

(q < p)%positive -> pos_sub p q = pos (p - q)
p, q:positive

(q < p)%positive -> pos_sub p q = pos (p - q)
p, q:positive
H:(q < p)%positive

pos_sub p q = pos (p - q)
now rewrite pos_sub_spec, Pos.compare_antisym, H. Qed.
The opposite of pos_sub is pos_sub with reversed arguments
p, q:positive

- pos_sub p q = pos_sub q p
p, q:positive

- pos_sub p q = pos_sub q p
revert q; induction p; destruct q; simpl; trivial; rewrite <- IHp; now destruct pos_sub. Qed.
In the following module, we group results that are needed now to prove specifications of operations, but will also be provided later by the generic functor of properties.
Module Import Private_BootStrap.

Operations and constants

n:Z

n + 0 = n
n:Z

n + 0 = n
now destruct n. Qed.
n:Z

n * 0 = 0
n:Z

n * 0 = 0
now destruct n. Qed.
n:Z

1 * n = n
n:Z

1 * n = n
now destruct n. Qed.

Addition is commutative

n, m:Z

n + m = m + n
n, m:Z

n + m = m + n
destruct n, m; simpl; trivial; now rewrite Pos.add_comm. Qed.

Opposite distributes over addition

n, m:Z

- (n + m) = - n + - m
n, m:Z

- (n + m) = - n + - m
destruct n, m; simpl; trivial using pos_sub_opp. Qed.

Opposite is injective

n, m:Z

- n = - m -> n = m
n, m:Z

- n = - m -> n = m
destruct n, m; simpl; intros H; destr_eq H; now f_equal. Qed.

Addition is associative

p, q, r:positive

pos_sub (p + q) r = pos p + pos_sub q r
p, q, r:positive

pos_sub (p + q) r = pos p + pos_sub q r
p, q, r:positive

pos_sub (p + q) r = match pos_sub q r with | 0 => pos p | pos y' => pos (p + y') | neg y' => pos_sub p y' end
p, q, r:positive

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = match match (q ?= r)%positive with | Eq => 0 | Lt => neg (r - q) | Gt => pos (q - r) end with | 0 => pos p | pos y' => pos (p + y') | neg y' => pos_sub p y' end
p, q, r:positive
E0:q = r

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos p
p, q, r:positive
E0:(q < r)%positive
match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos_sub p (r - q)
p, q, r:positive
E0:(r < q)%positive
match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos (p + (q - r))
p, q, r:positive
E0:q = r

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos p
p, r:positive

match (p + r ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + r)) | Gt => pos (p + r - r) end = pos p
p, r:positive
H:(r < r + p)%positive

match (p + r ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + r)) | Gt => pos (p + r - r) end = pos p
p, r:positive
H:(r < p + r)%positive

match (p + r ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + r)) | Gt => pos (p + r - r) end = pos p
p, r:positive
H:(p + r > r)%positive

match (p + r ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + r)) | Gt => pos (p + r - r) end = pos p
now rewrite H, Pos.add_sub.
p, q, r:positive
E0:(q < r)%positive

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos_sub p (r - q)
p, q, r:positive
E0:(q < r)%positive

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = match (p ?= r - q)%positive with | Eq => 0 | Lt => neg (r - q - p) | Gt => pos (p - (r - q)) end
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = match (p ?= r - q)%positive with | Eq => 0 | Lt => neg (r - q - p) | Gt => pos (p - (r - q)) end
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive

match (p + q ?= r - q + q)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = match (p ?= r - q)%positive with | Eq => 0 | Lt => neg (r - q - p) | Gt => pos (p - (r - q)) end
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive

match (p ?= r - q)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = match (p ?= r - q)%positive with | Eq => 0 | Lt => neg (r - q - p) | Gt => pos (p - (r - q)) end
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(p < r - q)%positive

(r - (p + q))%positive = (r - q - p)%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(r - q < p)%positive
(p + q - r)%positive = (p - (r - q))%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(p < r - q)%positive

(r - (q + p))%positive = (r - q - p)%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(r - q < p)%positive
(p + q - r)%positive = (p - (r - q))%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(p < r - q)%positive

(q + p < r)%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(r - q < p)%positive
(p + q - r)%positive = (p - (r - q))%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(p < r - q)%positive

(p + q < r - q + q)%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(r - q < p)%positive
(p + q - r)%positive = (p - (r - q))%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(r - q < p)%positive

(p + q - r)%positive = (p - (r - q))%positive
p, q, r:positive
E0:(q < r)%positive
Hr:r = (r - q + q)%positive
E1:(r - q < p)%positive

(p - (r - q))%positive = (p + q - r)%positive
apply Pos.sub_sub_distr; trivial.
p, q, r:positive
E0:(r < q)%positive

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos (p + (q - r))
p, q, r:positive
E0:(r < q)%positive

(r < p + q)%positive
p, q, r:positive
E0:(r < q)%positive
LT:(r < p + q)%positive
match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos (p + (q - r))
p, q, r:positive
E0:(r < q)%positive

(r < p + q)%positive
p, q, r:positive
E0:(r < q)%positive

(q < p + q)%positive
p, q, r:positive
E0:(r < q)%positive

(q < q + p)%positive
apply Pos.lt_add_r.
p, q, r:positive
E0:(r < q)%positive
LT:(r < p + q)%positive

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos (p + (q - r))
p, q, r:positive
E0:(r < q)%positive
LT:(p + q > r)%positive

match (p + q ?= r)%positive with | Eq => 0 | Lt => neg (r - (p + q)) | Gt => pos (p + q - r) end = pos (p + (q - r))
p, q, r:positive
E0:(r < q)%positive
LT:(p + q > r)%positive

pos (p + q - r) = pos (p + (q - r))
p, q, r:positive
E0:(r < q)%positive
LT:(p + q > r)%positive

(p + q - r)%positive = (p + (q - r))%positive
p, q, r:positive
E0:(r < q)%positive
LT:(p + q > r)%positive

(p + (q - r))%positive = (p + q - r)%positive
now apply Pos.add_sub_assoc. Qed. Local Arguments add !x !y.
p:positive
n, m:Z

pos p + (n + m) = pos p + n + m
p:positive
n, m:Z

pos p + (n + m) = pos p + n + m
p, n, m:positive

pos (p + (n + m)) = pos (p + n + m)
p, n, m:positive
pos p + pos_sub n m = pos_sub (p + n) m
p, n:positive
pos_sub p n = pos_sub p n + 0
p, n, m:positive
pos p + pos_sub m n = pos_sub p n + pos m
p, n, m:positive
pos_sub p (n + m) = pos_sub p n + neg m
p, n, m:positive

pos (p + (n + m)) = pos (p + n + m)
now rewrite Pos.add_assoc.
p, n, m:positive

pos p + pos_sub n m = pos_sub (p + n) m
p, n, m:positive

pos_sub (p + n) m = pos p + pos_sub n m
apply pos_sub_add.
p, n:positive

pos_sub p n = pos_sub p n + 0
p, n:positive

pos_sub p n + 0 = pos_sub p n
apply add_0_r.
p, n, m:positive

pos p + pos_sub m n = pos_sub p n + pos m
now rewrite <- pos_sub_add, add_comm, <- pos_sub_add, Pos.add_comm.
p, n, m:positive

pos_sub p (n + m) = pos_sub p n + neg m
p, n, m:positive

- pos_sub p (n + m) = - (pos_sub p n + neg m)
p, n, m:positive

pos_sub (n + m) p = pos_sub n p + - neg m
p, n, m:positive

pos_sub (m + n) p = - neg m + pos_sub n p
apply pos_sub_add. Qed.
n, m, p:Z

n + (m + p) = n + m + p
n, m, p:Z

n + (m + p) = n + m + p
m, p:Z

0 + (m + p) = 0 + m + p
p0:positive
m, p:Z
pos p0 + (m + p) = pos p0 + m + p
p0:positive
m, p:Z
neg p0 + (m + p) = neg p0 + m + p
m, p:Z

0 + (m + p) = 0 + m + p
trivial.
p0:positive
m, p:Z

pos p0 + (m + p) = pos p0 + m + p
apply add_assoc_pos.
p0:positive
m, p:Z

neg p0 + (m + p) = neg p0 + m + p
p0:positive
m, p:Z

- (neg p0 + (m + p)) = - (neg p0 + m + p)
p0:positive
m, p:Z

- neg p0 + (- m + - p) = - neg p0 + - m + - p
p0:positive
m, p:Z

pos p0 + (- m + - p) = pos p0 + - m + - p
apply add_assoc_pos. Qed.

Opposite is inverse for addition

n:Z

n + - n = 0
n:Z

n + - n = 0
destruct n; simpl; trivial; now rewrite pos_sub_diag. Qed.

Multiplication and Opposite

n, m:Z

n * - m = - (n * m)
n, m:Z

n * - m = - (n * m)
now destruct n, m. Qed.

Distributivity of multiplication over addition

p:positive
n, m:Z

(n + m) * pos p = n * pos p + m * pos p
p:positive
n, m:Z

(n + m) * pos p = n * pos p + m * pos p
p, n, m:positive

pos ((n + m) * p) = pos (n * p + m * p)
p, n, m:positive
pos_sub n m * pos p = pos_sub (n * p) (m * p)
p, n, m:positive
pos_sub m n * pos p = pos_sub (m * p) (n * p)
p, n, m:positive
neg ((n + m) * p) = neg (n * p + m * p)
p, n, m:positive

pos ((n + m) * p) = pos (n * p + m * p)
now rewrite Pos.mul_add_distr_r.
p, n, m:positive

pos_sub n m * pos p = pos_sub (n * p) (m * p)
rewrite ?pos_sub_spec, ?Pos.mul_compare_mono_r; case Pos.compare_spec; simpl; trivial; intros; now rewrite Pos.mul_sub_distr_r.
p, n, m:positive

pos_sub m n * pos p = pos_sub (m * p) (n * p)
rewrite ?pos_sub_spec, ?Pos.mul_compare_mono_r; case Pos.compare_spec; simpl; trivial; intros; now rewrite Pos.mul_sub_distr_r.
p, n, m:positive

neg ((n + m) * p) = neg (n * p + m * p)
now rewrite Pos.mul_add_distr_r. Qed.
n, m, p:Z

(n + m) * p = n * p + m * p
n, m, p:Z

(n + m) * p = n * p + m * p
n, m:Z

(n + m) * 0 = n * 0 + m * 0
n, m:Z
p:positive
(n + m) * pos p = n * pos p + m * pos p
n, m:Z
p:positive
(n + m) * neg p = n * neg p + m * neg p
n, m:Z

(n + m) * 0 = n * 0 + m * 0
now rewrite !mul_0_r.
n, m:Z
p:positive

(n + m) * pos p = n * pos p + m * pos p
apply mul_add_distr_pos.
n, m:Z
p:positive

(n + m) * neg p = n * neg p + m * neg p
n, m:Z
p:positive

- ((n + m) * neg p) = - (n * neg p + m * neg p)
n, m:Z
p:positive

(n + m) * - neg p = n * - neg p + m * - neg p
apply mul_add_distr_pos. Qed. End Private_BootStrap.

Proofs of specifications

Specification of constants


1 = succ 0

1 = succ 0
reflexivity. Qed.

2 = succ 1

2 = succ 1
reflexivity. Qed.

Specification of addition

n:Z

0 + n = n
n:Z

0 + n = n
now destruct n. Qed.
n, m:Z

succ n + m = succ (n + m)
n, m:Z

succ n + m = succ (n + m)
n, m:Z

n + 1 + m = n + m + 1
now rewrite 2 (add_comm _ 1), add_assoc. Qed.

Specification of opposite


- 0 = 0

- 0 = 0
reflexivity. Qed.
n:Z

- succ n = pred (- n)
n:Z

- succ n = pred (- n)
n:Z

- (n + 1) = - n + -1
apply opp_add_distr. Qed.

Specification of successor and predecessor

Local Arguments pos_sub : simpl nomatch.

n:Z

succ (pred n) = n
n:Z

succ (pred n) = n
n:Z

n + -1 + 1 = n
now rewrite <- add_assoc, add_opp_diag_r, add_0_r. Qed.
n:Z

pred (succ n) = n
n:Z

pred (succ n) = n
n:Z

n + 1 + -1 = n
now rewrite <- add_assoc, add_opp_diag_r, add_0_r. Qed.

Specification of subtraction

n:Z

n - 0 = n
n:Z

n - 0 = n
apply add_0_r. Qed.
n, m:Z

n - succ m = pred (n - m)
n, m:Z

n - succ m = pred (n - m)
n, m:Z

n + - (m + 1) = n + - m + -1
now rewrite opp_add_distr, add_assoc. Qed.

Specification of multiplication

n:Z

0 * n = 0
n:Z

0 * n = 0
reflexivity. Qed.
n, m:Z

succ n * m = n * m + m
n, m:Z

succ n * m = n * m + m
n, m:Z

(n + 1) * m = n * m + m
now rewrite mul_add_distr_r, mul_1_l. Qed.

Specification of comparisons and order

n, m:Z

(n =? m) = true <-> n = m
n, m:Z

(n =? m) = true <-> n = m
destruct n, m; simpl; try (now split); rewrite Pos.eqb_eq; split; (now injection 1) || (intros; now f_equal). Qed.
n, m:Z

(n <? m) = true <-> n < m
n, m:Z

(n <? m) = true <-> n < m
n, m:Z

match n ?= m with | Lt => true | _ => false end = true <-> (n ?= m) = Lt
destruct compare; easy'. Qed.
n, m:Z

(n <=? m) = true <-> n <= m
n, m:Z

(n <=? m) = true <-> n <= m
n, m:Z

match n ?= m with | Gt => false | _ => true end = true <-> (n ?= m) <> Gt
destruct compare; easy'. Qed.
n, m:Z

(n ?= m) = Eq <-> n = m
n, m:Z

(n ?= m) = Eq <-> n = m
destruct n, m; simpl; rewrite ?CompOpp_iff, ?Pos.compare_eq_iff; split; congruence. Qed.
n, m:Z

(n ?= m) = (n - m ?= 0)
n, m:Z

(n ?= m) = (n - m ?= 0)
destruct n as [|n|n], m as [|m|m]; simpl; trivial; rewrite <- ? Pos.compare_antisym, ?pos_sub_spec; case Pos.compare_spec; trivial. Qed.
n, m:Z

(m ?= n) = CompOpp (n ?= m)
n, m:Z

(m ?= n) = CompOpp (n ?= m)
destruct n, m; simpl; trivial; now rewrite <- ?Pos.compare_antisym. Qed.
n, m:Z

(n ?= m) = Lt <-> n < m
n, m:Z

(n ?= m) = Lt <-> n < m
reflexivity. Qed.
n, m:Z

(n ?= m) <> Gt <-> n <= m
n, m:Z

(n ?= m) <> Gt <-> n <= m
reflexivity. Qed.
Some more advanced properties of comparison and orders, including compare_spec and lt_irrefl and lt_eq_cases.
Include BoolOrderFacts.
Remaining specification of lt and le
n, m:Z

n < succ m <-> n <= m
n, m:Z

n < succ m <-> n <= m
n, m:Z

(n ?= succ m) = Lt <-> (n ?= m) <> Gt
n, m:Z

(pred (n - m) ?= 0) = Lt <-> (n ?= m) <> Gt
n, m:Z

(pred (n - m) ?= 0) = Lt <-> (n - m ?= 0) <> Gt
destruct (n-m) as [|[ | | ]|]; easy'. Qed.

Specification of minimum and maximum

n, m:Z

m <= n -> max n m = n
n, m:Z

m <= n -> max n m = n
n, m:Z

(m ?= n) <> Gt -> match n ?= m with | Lt => m | _ => n end = n
n, m:Z

CompOpp (n ?= m) <> Gt -> match n ?= m with | Lt => m | _ => n end = n
case compare; intuition. Qed.
n, m:Z

n <= m -> max n m = m
n, m:Z

n <= m -> max n m = m
n, m:Z

(n ?= m) <> Gt -> match n ?= m with | Lt => m | _ => n end = m
case compare_spec; intuition. Qed.
n, m:Z

n <= m -> min n m = n
n, m:Z

n <= m -> min n m = n
n, m:Z

(n ?= m) <> Gt -> match n ?= m with | Gt => m | _ => n end = n
case compare_spec; intuition. Qed.
n, m:Z

m <= n -> min n m = m
n, m:Z

m <= n -> min n m = m
n, m:Z

(m ?= n) <> Gt -> match n ?= m with | Gt => m | _ => n end = m
n, m:Z

CompOpp (n ?= m) <> Gt -> match n ?= m with | Gt => m | _ => n end = m
case compare_spec; intuition. Qed.

Induction principles based on successor / predecessor

P:Z -> Prop

P 0 -> (forall x : Z, P x -> P (succ x)) -> (forall x : Z, P x -> P (pred x)) -> forall z : Z, P z
P:Z -> Prop

P 0 -> (forall x : Z, P x -> P (succ x)) -> (forall x : Z, P x -> P (pred x)) -> forall z : Z, P z
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)

P 0
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
P (pos p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
P (neg p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive

P (pos p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
P (neg p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)

P 1
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
IHp:P (pos p)
P (pos (Pos.succ p))
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
P (neg p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
IHp:P (pos p)

P (pos (Pos.succ p))
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
P (neg p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
IHp:P (pos p)

P (pos (p + 1))
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
P (neg p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive

P (neg p)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)

P (-1)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
IHp:P (neg p)
P (neg (Pos.succ p))
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
IHp:P (neg p)

P (neg (Pos.succ p))
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x -> P (succ x)
Hp:forall x : Z, P x -> P (pred x)
p:positive
IHp:P (neg p)

P (neg (p + 1))
now apply (Hp (neg p)). Qed.
P:Z -> Prop

Proper (eq ==> iff) P -> P 0 -> (forall x : Z, P x <-> P (succ x)) -> forall z : Z, P z
P:Z -> Prop

Proper (eq ==> iff) P -> P 0 -> (forall x : Z, P x <-> P (succ x)) -> forall z : Z, P z
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)

forall z : Z, P z
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)

P 0
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)
z:Z
IHz:P z
P (succ z)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)
z:Z
IHz:P z
P (pred z)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)
z:Z
IHz:P z

P (succ z)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)
z:Z
IHz:P z
P (pred z)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)
z:Z
IHz:P z

P (pred z)
P:Z -> Prop
H0:P 0
Hs:forall x : Z, P x <-> P (succ x)
z:Z
IHz:P z

P (succ (pred z))
now rewrite succ_pred. Qed.
We can now derive all properties of basic functions and orders, and use these properties for proving the specs of more advanced functions.
Include ZBasicProp <+ UsualMinMaxLogicalProperties <+ UsualMinMaxDecProperties.

Register eq_decidable as num.Z.eq_decidable.
Register le_decidable as num.Z.le_decidable.
Register lt_decidable as num.Z.lt_decidable.

Specification of absolute value

n:Z

0 <= n -> abs n = n
n:Z

0 <= n -> abs n = n
p:positive

0 <= neg p -> abs (neg p) = neg p
now destruct 1. Qed.
n:Z

n <= 0 -> abs n = - n
n:Z

n <= 0 -> abs n = - n
p:positive

pos p <= 0 -> abs (pos p) = - pos p
now destruct 1. Qed.

Specification of sign

n:Z

n = 0 -> sgn n = 0
n:Z

n = 0 -> sgn n = 0
n:Z
H:n = 0

sgn n = 0
now subst. Qed.
n:Z

0 < n -> sgn n = 1
n:Z

0 < n -> sgn n = 1
now destruct n. Qed.
n:Z

n < 0 -> sgn n = -1
n:Z

n < 0 -> sgn n = -1
now destruct n. Qed.

Specification of power

n:Z

n ^ 0 = 1
n:Z

n ^ 0 = 1
reflexivity. Qed.
n, m:Z

0 <= m -> n ^ succ m = n * n ^ m
n, m:Z

0 <= m -> n ^ succ m = n * n ^ m
n:Z
m:positive

pow_pos n (m + 1) = n * pow_pos n m
n:Z
m:positive

Pos.iter (mul n) 1 (m + 1) = n * Pos.iter (mul n) 1 m
now rewrite Pos.add_comm, Pos.iter_add. Qed.
n, m:Z

m < 0 -> n ^ m = 0
n, m:Z

m < 0 -> n ^ m = 0
now destruct m. Qed.
For folding back a pow_pos into a pow
n:Z
p:positive

pow_pos n p = n ^ pos p
n:Z
p:positive

pow_pos n p = n ^ pos p
reflexivity. Qed.

Specification of square

n:Z

square n = n * n
n:Z

square n = n * n
destruct n; trivial; simpl; f_equal; apply Pos.square_spec. Qed.

Specification of square root

n:Z

0 <= n -> let (s, r) := sqrtrem n in n = s * s + r /\ 0 <= r <= 2 * s
n:Z

0 <= n -> let (s, r) := sqrtrem n in n = s * s + r /\ 0 <= r <= 2 * s

0 <= 0 -> let (s, r) := sqrtrem 0 in 0 = s * s + r /\ 0 <= r <= 2 * s
p:positive
0 <= pos p -> let (s, r) := sqrtrem (pos p) in pos p = s * s + r /\ 0 <= r <= 2 * s
p:positive
0 <= neg p -> let (s, r) := sqrtrem (neg p) in neg p = s * s + r /\ 0 <= r <= 2 * s
p:positive

0 <= pos p -> let (s, r) := sqrtrem (pos p) in pos p = s * s + r /\ 0 <= r <= 2 * s
p:positive
0 <= neg p -> let (s, r) := sqrtrem (neg p) in neg p = s * s + r /\ 0 <= r <= 2 * s
p:positive

Pos.SqrtSpec (Pos.sqrtrem p) p -> 0 <= pos p -> let (s, r) := sqrtrem (pos p) in pos p = s * s + r /\ 0 <= r <= 2 * s
p:positive
0 <= neg p -> let (s, r) := sqrtrem (neg p) in neg p = s * s + r /\ 0 <= r <= 2 * s
p:positive

Pos.SqrtSpec (Pos.sqrtrem p) p -> 0 <= pos p -> let (s, r) := let (s, m) := Pos.sqrtrem p in match m with | Pos.IsPos r => (pos s, pos r) | _ => (pos s, 0) end in pos p = s * s + r /\ 0 <= r <= match s with | 0 => 0 | pos y' => pos y'~0 | neg y' => neg y'~0 end
p:positive
0 <= neg p -> let (s, r) := sqrtrem (neg p) in neg p = s * s + r /\ 0 <= r <= 2 * s
p:positive

0 <= neg p -> let (s, r) := sqrtrem (neg p) in neg p = s * s + r /\ 0 <= r <= 2 * s
now destruct 1. Qed.
n:Z

0 <= n -> let s := sqrt n in s * s <= n < succ s * succ s
n:Z

0 <= n -> let s := sqrt n in s * s <= n < succ s * succ s

0 <= 0 -> let s := sqrt 0 in s * s <= 0 < succ s * succ s
p:positive
0 <= pos p -> let s := sqrt (pos p) in s * s <= pos p < succ s * succ s
p:positive
0 <= neg p -> let s := sqrt (neg p) in s * s <= neg p < succ s * succ s
p:positive

0 <= pos p -> let s := sqrt (pos p) in s * s <= pos p < succ s * succ s
p:positive
0 <= neg p -> let s := sqrt (neg p) in s * s <= neg p < succ s * succ s
p:positive

0 <= pos p -> pos (Pos.sqrt p) * pos (Pos.sqrt p) <= pos p < succ (pos (Pos.sqrt p)) * succ (pos (Pos.sqrt p))
p:positive
0 <= neg p -> let s := sqrt (neg p) in s * s <= neg p < succ s * succ s
p:positive

pos (Pos.sqrt p) * pos (Pos.sqrt p) <= pos p < succ (pos (Pos.sqrt p)) * succ (pos (Pos.sqrt p))
p:positive
0 <= neg p -> let s := sqrt (neg p) in s * s <= neg p < succ s * succ s
p:positive

pos (Pos.sqrt p) * pos (Pos.sqrt p) <= pos p < pos (Pos.sqrt p + 1) * pos (Pos.sqrt p + 1)
p:positive
0 <= neg p -> let s := sqrt (neg p) in s * s <= neg p < succ s * succ s
p:positive

pos (Pos.sqrt p) * pos (Pos.sqrt p) <= pos p < pos (Pos.succ (Pos.sqrt p)) * pos (Pos.succ (Pos.sqrt p))
p:positive
0 <= neg p -> let s := sqrt (neg p) in s * s <= neg p < succ s * succ s
p:positive

0 <= neg p -> let s := sqrt (neg p) in s * s <= neg p < succ s * succ s
now destruct 1. Qed.
n:Z

n < 0 -> sqrt n = 0
n:Z

n < 0 -> sqrt n = 0
now destruct n. Qed.
n:Z

fst (sqrtrem n) = sqrt n
n:Z

fst (sqrtrem n) = sqrt n
p:positive

fst (sqrtrem (pos p)) = sqrt (pos p)
p:positive

fst (let (s, m) := Pos.sqrtrem p in match m with | Pos.IsPos r => (pos s, pos r) | _ => (pos s, 0) end) = pos (fst (Pos.sqrtrem p))
p, s:positive
r:Pos.mask

fst match r with | Pos.IsPos r0 => (pos s, pos r0) | _ => (pos s, 0) end = pos (fst (s, r))
now destruct r. Qed.

Specification of logarithm

n:Z

0 < n -> 2 ^ log2 n <= n < 2 ^ succ (log2 n)
n:Z

0 < n -> 2 ^ log2 n <= n < 2 ^ succ (log2 n)
n:Z

forall p q : positive, pos (p ^ q) = pos p ^ pos q
n:Z
Pow:forall p q : positive, pos (p ^ q) = pos p ^ pos q
0 < n -> 2 ^ log2 n <= n < 2 ^ succ (log2 n)
n:Z

forall p q : positive, pos (p ^ q) = pos p ^ pos q
n:Z
p, q:positive

pos (p ^ q) = pos p ^ pos q
now apply Pos.iter_swap_gen.
n:Z
Pow:forall p q : positive, pos (p ^ q) = pos p ^ pos q

0 < n -> 2 ^ log2 n <= n < 2 ^ succ (log2 n)
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~1

pos (2 ^ Pos.size p) <= pos p~1
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~1
pos p~1 < pos (2 ^ Pos.succ (Pos.size p))
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0
pos (2 ^ Pos.size p) <= pos p~0
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0
pos p~0 < pos (2 ^ Pos.succ (Pos.size p))
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~1

(2 ^ Pos.size p <= Pos.succ p~0)%positive
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~1
pos p~1 < pos (2 ^ Pos.succ (Pos.size p))
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0
pos (2 ^ Pos.size p) <= pos p~0
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0
pos p~0 < pos (2 ^ Pos.succ (Pos.size p))
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~1

pos p~1 < pos (2 ^ Pos.succ (Pos.size p))
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0
pos (2 ^ Pos.size p) <= pos p~0
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0
pos p~0 < pos (2 ^ Pos.succ (Pos.size p))
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0

pos (2 ^ Pos.size p) <= pos p~0
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0
pos p~0 < pos (2 ^ Pos.succ (Pos.size p))
p:positive
Pow:forall p0 q : positive, pos (p0 ^ q) = pos p0 ^ pos q
Hn:0 < pos p~0

pos p~0 < pos (2 ^ Pos.succ (Pos.size p))
apply Pos.size_gt. Qed.
n:Z

n <= 0 -> log2 n = 0
n:Z

n <= 0 -> log2 n = 0
destruct n as [|p|p]; trivial; now destruct 1. Qed.
Specification of parity functions
n:Z

even n = true <-> Even n
n:Z

even n = true <-> Even n
n:Z

even n = true -> Even n
n:Z
Even n -> even n = true
n:Z
H:even n = true

n = 2 * div2 n
n:Z
Even n -> even n = true
n:Z

Even n -> even n = true
m:Z

even (2 * m) = true
now destruct m. Qed.
n:Z

odd n = true <-> Odd n
n:Z

odd n = true <-> Odd n
n:Z

odd n = true -> Odd n
n:Z
Odd n -> odd n = true
n:Z
H:odd n = true

n = 2 * div2 n + 1
n:Z
Odd n -> odd n = true
p:positive
H:odd (neg p~1) = true

neg p~1 = neg (Pos.pred_double (Pos.succ p))
n:Z
Odd n -> odd n = true
n:Z

Odd n -> odd n = true
m:Z

odd (2 * m + 1) = true
now destruct m as [|[ | | ]|[ | | ]]. Qed.

Multiplication and Doubling

n:Z

double n = 2 * n
n:Z

double n = 2 * n
reflexivity. Qed.
n:Z

succ_double n = 2 * n + 1
n:Z

succ_double n = 2 * n + 1
now destruct n. Qed.
n:Z

pred_double n = 2 * n - 1
n:Z

pred_double n = 2 * n - 1
now destruct n. Qed.

Correctness proofs for Trunc division

a:positive
b:Z

0 < b -> let (q, r) := pos_div_eucl a b in pos a = q * b + r
a:positive
b:Z

0 < b -> let (q, r) := pos_div_eucl a b in pos a = q * b + r
a:positive
b:Z
Hb:0 < b

let (q, r) := pos_div_eucl a b in pos a = q * b + r
a:positive
b:Z
Hb:0 < b
IHa:let (q, r) := pos_div_eucl a b in pos a = q * b + r

let (q, r) := let (q, r) := pos_div_eucl a b in if 2 * r + 1 <? b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - b) in pos a~1 = q * b + r
a:positive
b:Z
Hb:0 < b
IHa:let (q, r) := pos_div_eucl a b in pos a = q * b + r
let (q, r) := let (q, r) := pos_div_eucl a b in if 2 * r <? b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - b) in pos a~0 = q * b + r
b:Z
Hb:0 < b
let (q, r) := if 2 <=? b then (0, 1) else (1, 0) in 1 = q * b + r
a:positive
b:Z
Hb:0 < b
IHa:let (q, r) := pos_div_eucl a b in pos a = q * b + r

let (q, r) := let (q, r) := pos_div_eucl a b in if 2 * r + 1 <? b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - b) in pos a~1 = q * b + r
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

let (q0, r0) := if 2 * r + 1 <? b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - b) in pos a~1 = q0 * b + r0
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

let (q0, r0) := if 2 * r + 1 <? b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - b) in 2 * pos a + 1 = q0 * b + r0
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

let (q0, r0) := if 2 * r + 1 <? b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - b) in 2 * q * b + 2 * r + 1 = q0 * b + r0
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * q * b + 2 * r + 1 = 2 * q * b + (2 * r + 1)
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r
2 * q * b + 2 * r + 1 = (2 * q + 1) * b + (2 * r + 1 - b)
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * q * b + 2 * r + 1 = (2 * q + 1) * b + (2 * r + 1 - b)
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * q * b + (2 * r + 1) = 2 * q * b + (b + (2 * r + 1 - b))
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * r + 1 = b + (2 * r + 1 - b)
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * r + 1 = b + (2 * r + 1 + - b)
now rewrite (add_comm _ (-b)), add_assoc, add_opp_diag_r.
a:positive
b:Z
Hb:0 < b
IHa:let (q, r) := pos_div_eucl a b in pos a = q * b + r

let (q, r) := let (q, r) := pos_div_eucl a b in if 2 * r <? b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - b) in pos a~0 = q * b + r
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

let (q0, r0) := if 2 * r <? b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - b) in pos a~0 = q0 * b + r0
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

let (q0, r0) := if 2 * r <? b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - b) in 2 * pos a = q0 * b + r0
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

let (q0, r0) := if 2 * r <? b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - b) in 2 * q * b + 2 * r = q0 * b + r0
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * q * b + 2 * r = 2 * q * b + 2 * r
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r
2 * q * b + 2 * r = (2 * q + 1) * b + (2 * r - b)
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * q * b + 2 * r = (2 * q + 1) * b + (2 * r - b)
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * q * b + 2 * r = 2 * q * b + (b + (2 * r - b))
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * r = b + (2 * r - b)
a:positive
b:Z
Hb:0 < b
q, r:Z
IHa:pos a = q * b + r

2 * r = b + (2 * r + - b)
now rewrite (add_comm _ (-b)), add_assoc, add_opp_diag_r.
b:Z
Hb:0 < b

let (q, r) := if 2 <=? b then (0, 1) else (1, 0) in 1 = q * b + r
b:Z
Hb:0 < b

b < 2 -> 1 = 1 * b + 0
b:Z
Hb:0 < b
Hb':b < 2

1 = 1 * b + 0
b:positive
Hb':pos b < 2

1 = 1 * pos b + 0
b:positive
Hb':pos b < 2

1%positive = b
b:positive
Hb':pos b < 2

(1 <= b)%positive
b:positive
Hb':pos b < 2
(b <= 1)%positive
b:positive
Hb':pos b < 2

(b <= 1)%positive
now apply Pos.lt_succ_r. Qed.
a, b:Z

b <> 0 -> let (q, r) := div_eucl a b in a = b * q + r
a, b:Z

b <> 0 -> let (q, r) := div_eucl a b in a = b * q + r
a, b:positive
q, r:Z

pos a = pos b * q + r -> pos a = pos b * q + r
a, b:positive
q, r:Z
pos a = pos b * q + r -> let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), neg b + r) end in pos a = neg b * q0 + r0
a, b:positive
q, r:Z
pos a = pos b * q + r -> let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in neg a = pos b * q0 + r0
a, b:positive
q, r:Z
pos a = pos b * q + r -> neg a = neg b * q + - r
a, b:positive
q, r:Z

pos a = pos b * q + r -> pos a = pos b * q + r
trivial.
a, b:positive
q, r:Z

pos a = pos b * q + r -> let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), neg b + r) end in pos a = neg b * q0 + r0
a, b:positive
q, r:Z

let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), neg b + r) end in pos b * q + r = neg b * q0 + r0
destruct r as [ |r|r]; rewrite <- !mul_opp_comm; trivial; rewrite mul_add_distr_l, mul_1_r, <- add_assoc; f_equal; now rewrite add_assoc, add_opp_diag_r.
a, b:positive
q, r:Z

pos a = pos b * q + r -> let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in neg a = pos b * q0 + r0
a, b:positive
q, r:Z

pos a = pos b * q + r -> let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in - pos a = pos b * q0 + r0
a, b:positive
q, r:Z

let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in - (pos b * q + r) = pos b * q0 + r0
a, b:positive
q, r:Z

let (q0, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in pos b * - q + - r = pos b * q0 + r0
destruct r as [ |r|r]; trivial; rewrite opp_add_distr, mul_add_distr_l, <- add_assoc; f_equal; unfold sub; now rewrite add_assoc, mul_opp_r, mul_1_r, add_opp_diag_l.
a, b:positive
q, r:Z

pos a = pos b * q + r -> neg a = neg b * q + - r
a, b:positive
q, r:Z

pos a = pos b * q + r -> - pos a = neg b * q + - r
a, b:positive
q, r:Z

- (pos b * q + r) = neg b * q + - r
now rewrite opp_add_distr, <- mul_opp_l. Qed.
a, b:Z

b <> 0 -> a = b * (a / b) + a mod b
a, b:Z

b <> 0 -> a = b * (a / b) + a mod b
a, b:Z
Hb:b <> 0

a = b * (a / b) + a mod b
a, b:Z
Hb:b <> 0

(let (q, r) := div_eucl a b in a = b * q + r) -> a = b * (a / b) + a mod b
a, b:Z
Hb:b <> 0

(let (q, r) := div_eucl a b in a = b * q + r) -> a = b * (let (q, _) := div_eucl a b in q) + (let (_, r) := div_eucl a b in r)
now destruct div_eucl. Qed.
a:positive
b:Z

0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b:Z

0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b:Z

forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b, m:Z
p:positive

m < pos p~0 -> m - pos p < pos p
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b, m:Z
p:positive

(m ?= pos p~0) = Lt -> (m - pos p ?= pos p) = Lt
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b, m:Z
p:positive

(m - pos p~0 ?= 0) = Lt -> (m - pos p - pos p ?= 0) = Lt
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b, m:Z
p:positive

(m + - pos p~0 ?= 0) = Lt -> (m + - pos p + - pos p ?= 0) = Lt
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b, m:Z
p:positive

(m + - pos p~0 ?= 0) = Lt -> (m + (- pos p + - pos p) ?= 0) = Lt
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b, m:Z
p:positive

(m + neg p~0 ?= 0) = Lt -> (m + neg (p + p) ?= 0) = Lt
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p

0 < b -> 0 <= snd (pos_div_eucl a b) < b
a:positive
b:Z
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
Hb:0 < b

0 <= snd (pos_div_eucl a b) < b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p

0 <= snd (pos_div_eucl a (pos b)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b

0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r + 1 <? pos b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - pos b)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
(* ~1 *)
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
IHa:0 <= snd (q, r) < pos b

0 <= snd (if 2 * r + 1 <? pos b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - pos b)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b

0 <= snd (if 2 * r + 1 <? pos b then (2 * q, 2 * r + 1) else (2 * q + 1, 2 * r + 1 - pos b)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:2 * r + 1 < pos b

0 <= 2 * r + 1 < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1
0 <= 2 * r + 1 - pos b < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:2 * r + 1 < pos b

0 <= 2 * r + 1
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1
0 <= 2 * r + 1 - pos b < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1

0 <= 2 * r + 1 - pos b < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1

0 <= 2 * r + 1 - pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1
2 * r + 1 - pos b < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1

(0 ?= 2 * r + 1 - pos b) <> Gt
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1
2 * r + 1 - pos b < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1

2 * r + 1 - pos b < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1

2 * r + 1 < pos b~0
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r + 1

succ_double r < pos b~0
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p0 : positive), m < pos p0~0 -> m - pos p0 < pos p0
q:Z
p:positive
Hr:0 <= pos p
Hr':pos p < pos b
H:pos b <= 2 * pos p + 1

succ_double (pos p) < pos b~0
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p0 : positive), (m ?= pos p0~0) = Lt -> (m - pos p0 ?= pos p0) = Lt
q:Z
p:positive
Hr:0 <= pos p
Hr':(p ?= b)%positive = Lt
H:pos b <= pos p~1

(p~1 ?= b~0)%positive = Lt
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b
0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
IHa:0 <= snd (pos_div_eucl a (pos b)) < pos b

0 <= snd (let (q, r) := pos_div_eucl a (pos b) in if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
(* ~0 *)
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
IHa:0 <= snd (q, r) < pos b

0 <= snd (if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b

0 <= snd (if 2 * r <? pos b then (2 * q, 2 * r) else (2 * q + 1, 2 * r - pos b)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:2 * r < pos b

0 <= 2 * r < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r
0 <= 2 * r - pos b < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:2 * r < pos b

0 <= 2 * r
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r
0 <= 2 * r - pos b < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r

0 <= 2 * r - pos b < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r

0 <= 2 * r - pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r
2 * r - pos b < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r

(0 ?= 2 * r - pos b) <> Gt
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r
2 * r - pos b < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r

2 * r - pos b < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
a, b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
q, r:Z
Hr:0 <= r
Hr':r < pos b
H:pos b <= 2 * r

2 * r < pos b~0
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p

0 <= snd (if 2 <=? pos b then (0, 1) else (1, 0)) < pos b
(* 1 *)
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
H:2 <= pos b

1 < pos b
b:positive
AUX:forall (m : Z) (p : positive), m < pos p~0 -> m - pos p < pos p
H:2 <= pos b

(1 ?= b)%positive = Lt
now apply Pos.le_succ_l. Qed.
a, b:Z

0 < b -> 0 <= a mod b < b
a, b:Z

0 < b -> 0 <= a mod b < b
a:Z
b:positive

0 <= a mod pos b < pos b
b:positive

0 <= 0 < pos b
a, b:positive
0 <= (let (_, r) := pos_div_eucl a (pos b) in r) < pos b
a, b:positive
0 <= (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in r) < pos b
a, b:positive

0 <= (let (_, r) := pos_div_eucl a (pos b) in r) < pos b
a, b:positive
0 <= (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in r) < pos b
a, b:positive

0 <= (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in r) < pos b
a, b:positive

0 <= snd (pos_div_eucl a (pos b)) < pos b -> 0 <= (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in r) < pos b
a, b:positive
q, r:Z
Hr:0 <= r
Hr':r < pos b

0 <= (let (_, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), pos b - r) end in r0) < pos b
a, b:positive
q:Z
Hr':0 < pos b

0 <= 0 < pos b
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
0 <= pos b - pos r < pos b
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

0 <= pos b - pos r < pos b
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

0 <= pos b - pos r
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
pos b - pos r < pos b
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

(0 ?= pos b - pos r) <> Gt
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
pos b - pos r < pos b
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

pos b - pos r < pos b
a, b:positive
q:Z
r:positive
Hr':(r ?= b)%positive = Lt

(pos_sub b r ?= pos b) = Lt
a, b:positive
q:Z
r:positive
Hr':(r ?= b)%positive = Lt

(pos (b - r) ?= pos b) = Lt
a, b:positive
q:Z
r:positive
Hr':(r ?= b)%positive = Lt

(b - r ?= b)%positive = Lt
now apply Pos.sub_decr. Qed. Definition mod_bound_pos a b (_:0<=a) := mod_pos_bound a b.
a, b:Z

b < 0 -> b < a mod b <= 0
a, b:Z

b < 0 -> b < a mod b <= 0
a:Z
b:positive

neg b < a mod neg b <= 0
b:positive

neg b < 0 <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in match r with | 0 => (- q, 0) | _ => (- (q + 1), neg b + r) end in r) <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive

neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in match r with | 0 => (- q, 0) | _ => (- (q + 1), neg b + r) end in r) <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive

0 <= snd (pos_div_eucl a (pos b)) < pos b -> neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in match r with | 0 => (- q, 0) | _ => (- (q + 1), neg b + r) end in r) <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q, r:Z
Hr:0 <= r
Hr':r < pos b

neg b < (let (_, r0) := match r with | 0 => (- q, 0) | _ => (- (q + 1), neg b + r) end in r0) <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
Hr':0 < pos b

neg b < 0 <= 0
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
neg b < neg b + pos r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

neg b < neg b + pos r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

neg b < neg b + pos r
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
neg b + pos r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':(r ?= b)%positive = Lt

match pos_sub r b with | neg y' => CompOpp (b ?= y')%positive | _ => Lt end = Lt
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
neg b + pos r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':(r ?= b)%positive = Lt

CompOpp (b ?= b - r)%positive = Lt
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
neg b + pos r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':(r ?= b)%positive = Lt

(b - r ?= b)%positive = Lt
a, b:positive
q:Z
r:positive
Hr':pos r < pos b
neg b + pos r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

neg b + pos r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':pos r < pos b

neg b - neg r <= 0
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':(pos r ?= pos b) = Lt

(neg b - neg r ?= 0) <> Gt
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':(pos r ?= pos b) = Lt

(neg b ?= neg r) <> Gt
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q:Z
r:positive
Hr':(r ?= b)%positive = Lt

CompOpp (b ?= r)%positive <> Gt
a, b:positive
neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive

neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive

0 <= snd (pos_div_eucl a (pos b)) < pos b -> neg b < (let (_, r) := let (q, r) := pos_div_eucl a (pos b) in (q, - r) in r) <= 0
a, b:positive
q, r:Z
Hr:0 <= r
Hr':r < pos b

neg b < - r <= 0
a, b:positive
q:Z
p:positive
Hr:0 <= pos p
Hr':pos p < pos b

neg b < - pos p
red; simpl; now rewrite <- Pos.compare_antisym. Qed.

Correctness proofs for Floor division

a, b:Z

let (q, r) := quotrem a b in a = q * b + r
a, b:Z

let (q, r) := quotrem a b in a = q * b + r
destruct a as [|a|a], b as [|b|b]; simpl; trivial; generalize (N.pos_div_eucl_spec a (N.pos b)); case N.pos_div_eucl; trivial; intros q r; try change (neg a) with (-pos a); change (pos a) with (of_N (N.pos a)); intros ->; now destruct q, r. Qed.
a, b:Z

a = b * (a ÷ b) + rem a b
a, b:Z

a = b * (a ÷ b) + rem a b
a, b:Z

a = a ÷ b * b + rem a b
a, b:Z

(let (q, r) := quotrem a b in a = q * b + r) -> a = a ÷ b * b + rem a b
a, b:Z

(let (q, r) := quotrem a b in a = q * b + r) -> a = fst (quotrem a b) * b + snd (quotrem a b)
now destruct quotrem. Qed.
a, b:Z

b <> 0 -> a = b * (a ÷ b) + rem a b
a, b:Z

b <> 0 -> a = b * (a ÷ b) + rem a b
a, b:Z

a = b * (a ÷ b) + rem a b
apply quot_rem'. Qed.
a, b:Z

0 <= a -> 0 < b -> 0 <= rem a b < b
a, b:Z

0 <= a -> 0 < b -> 0 <= rem a b < b
a, b:Z
Ha:0 <= a
Hb:0 < b

0 <= rem a b < b
b:positive

0 <= rem 0 (pos b) < pos b
a, b:positive
0 <= rem (pos a) (pos b) < pos b
b:positive

(Eq = Gt -> False) /\ Lt = Lt
a, b:positive
0 <= rem (pos a) (pos b) < pos b
a, b:positive

0 <= rem (pos a) (pos b) < pos b
a, b:positive

0 <= snd (let (q, r) := N.pos_div_eucl a (N.pos b) in (of_N q, of_N r)) < pos b
a, b:positive
H:N.pos b <> 0%N -> (snd (N.pos_div_eucl a (N.pos b)) < N.pos b)%N

0 <= snd (let (q, r) := N.pos_div_eucl a (N.pos b) in (of_N q, of_N r)) < pos b
a, b:positive
q:N
r:positive
H:N.pos b <> 0%N -> (snd (q, N.pos r) < N.pos b)%N

pos r < pos b
now apply H. Qed.
a, b:Z

rem (- a) b = - rem a b
a, b:Z

rem (- a) b = - rem a b
destruct a, b; trivial; unfold rem; simpl; now destruct N.pos_div_eucl as (q,[|r]). Qed.
a, b:Z

rem a (- b) = rem a b
a, b:Z

rem a (- b) = rem a b
destruct a, b; trivial; unfold rem; simpl; now destruct N.pos_div_eucl as (q,[|r]). Qed.
a, b:Z

b <> 0 -> rem (- a) b = - rem a b
a, b:Z

b <> 0 -> rem (- a) b = - rem a b
a, b:Z

rem (- a) b = - rem a b
apply rem_opp_l'. Qed.
a, b:Z

b <> 0 -> rem a (- b) = rem a b
a, b:Z

b <> 0 -> rem a (- b) = rem a b
a, b:Z

rem a (- b) = rem a b
apply rem_opp_r'. Qed.

Extra properties about divide

p, q:positive

(pos p | pos q) <-> (p | q)%positive
p, q:positive

(pos p | pos q) <-> (p | q)%positive
p, q:positive

(pos p | pos q) -> (p | q)%positive
p, q:positive
(p | q)%positive -> (pos p | pos q)
p, q, r:positive
H:q = (r * p)%positive

(p | q)%positive
p, q:positive
(p | q)%positive -> (pos p | pos q)
p, q:positive

(p | q)%positive -> (pos p | pos q)
p, q, r:positive
H:q = (r * p)%positive

(pos p | pos q)
exists (pos r); simpl; now f_equal. Qed.
n:Z
p:positive

(n | pos p) <-> (n | neg p)
n:Z
p:positive

(n | pos p) <-> (n | neg p)
split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- H. Qed.
n:Z
p:positive

(pos p | n) <-> (neg p | n)
n:Z
p:positive

(pos p | n) <-> (neg p | n)
split; intros (m,H); exists (-m); now rewrite mul_opp_l, <- mul_opp_r. Qed.

Correctness proofs for gcd

a, b:Z

fst (ggcd a b) = gcd a b
a, b:Z

fst (ggcd a b) = gcd a b
destruct a as [ |p|p], b as [ |q|q]; simpl; auto; generalize (Pos.ggcd_gcd p q); destruct Pos.ggcd as (g,(aa,bb)); simpl; congruence. Qed.
a, b:Z

let '(g, (aa, bb)) := ggcd a b in a = g * aa /\ b = g * bb
a, b:Z

let '(g, (aa, bb)) := ggcd a b in a = g * aa /\ b = g * bb
destruct a as [ |p|p], b as [ |q|q]; simpl; rewrite ?Pos.mul_1_r; auto; generalize (Pos.ggcd_correct_divisors p q); destruct Pos.ggcd as (g,(aa,bb)); simpl; destruct 1; now subst. Qed.
a, b:Z

(gcd a b | a)
a, b:Z

(gcd a b | a)
a, b:Z

(fst (ggcd a b) | a)
a, b:Z

(let '(g, (aa, bb)) := ggcd a b in a = g * aa /\ b = g * bb) -> (fst (ggcd a b) | a)
a, b, g, aa, bb:Z

a = g * aa /\ b = g * bb -> (g | a)
a, b, g, aa, bb:Z
H:a = g * aa

(g | a)
a, b, g, aa, bb:Z
H:a = g * aa

a = aa * g
now rewrite mul_comm. Qed.
a, b:Z

(gcd a b | b)
a, b:Z

(gcd a b | b)
a, b:Z

(fst (ggcd a b) | b)
a, b:Z

(let '(g, (aa, bb)) := ggcd a b in a = g * aa /\ b = g * bb) -> (fst (ggcd a b) | b)
a, b, g, aa, bb:Z

a = g * aa /\ b = g * bb -> (g | b)
a, b, g, aa, bb:Z
H:b = g * bb

(g | b)
a, b, g, aa, bb:Z
H:b = g * bb

b = bb * g
now rewrite mul_comm. Qed.
a, b, c:Z

(c | a) -> (c | b) -> (c | gcd a b)
a, b, c:Z

(c | a) -> (c | b) -> (c | gcd a b)
a, b, c:Z

forall (p q : positive) (r : Z), (r | pos p) -> (r | pos q) -> (r | pos (Pos.gcd p q))
a, b, c:Z
H:forall (p q : positive) (r : Z), (r | pos p) -> (r | pos q) -> (r | pos (Pos.gcd p q))
(c | a) -> (c | b) -> (c | gcd a b)
a, b, c:Z

forall (p q : positive) (r : Z), (r | pos p) -> (r | pos q) -> (r | pos (Pos.gcd p q))
a, b, c:Z
p, q:positive
H:(0 | pos p)
H':(0 | pos q)

(0 | pos (Pos.gcd p q))
a, b, c:Z
p, q, r:positive
H:(pos r | pos p)
H':(pos r | pos q)
(pos r | pos (Pos.gcd p q))
a, b, c:Z
p, q, r:positive
H:(neg r | pos p)
H':(neg r | pos q)
(neg r | pos (Pos.gcd p q))
a, b, c:Z
p, q, r:positive
H:(pos r | pos p)
H':(pos r | pos q)

(pos r | pos (Pos.gcd p q))
a, b, c:Z
p, q, r:positive
H:(neg r | pos p)
H':(neg r | pos q)
(neg r | pos (Pos.gcd p q))
a, b, c:Z
p, q, r:positive
H:(neg r | pos p)
H':(neg r | pos q)

(neg r | pos (Pos.gcd p q))
apply divide_Zpos_Zneg_l, divide_Zpos, Pos.gcd_greatest; now apply divide_Zpos, divide_Zpos_Zneg_l.
a, b, c:Z
H:forall (p q : positive) (r : Z), (r | pos p) -> (r | pos q) -> (r | pos (Pos.gcd p q))

(c | a) -> (c | b) -> (c | gcd a b)
destruct a, b; simpl; auto; intros; try apply H; trivial; now apply divide_Zpos_Zneg_r. Qed.
a, b:Z

0 <= gcd a b
a, b:Z

0 <= gcd a b
now destruct a, b. Qed.
ggcd and opp : an auxiliary result used in QArith
a, b:Z

ggcd (- a) b = (let '(g, (aa, bb)) := ggcd a b in (g, (- aa, bb)))
a, b:Z

ggcd (- a) b = (let '(g, (aa, bb)) := ggcd a b in (g, (- aa, bb)))
destruct a as [|a|a], b as [|b|b]; unfold ggcd, opp; auto; destruct (Pos.ggcd a b) as (g,(aa,bb)); auto. Qed.

Extra properties about testbit

a, n:N

testbit (of_N a) (of_N n) = N.testbit a n
a, n:N

testbit (of_N a) (of_N n) = N.testbit a n
a:positive

match a with | (_~0)%positive => false | _ => true end = Pos.testbit a 0
now destruct a. Qed.
a:N
n:Z

0 <= n -> testbit (of_N a) n = N.testbit a (to_N n)
a:N
n:Z

0 <= n -> testbit (of_N a) n = N.testbit a (to_N n)
a:N
n:Z
Hn:0 <= n

testbit (of_N a) n = N.testbit a (to_N n)
a:N
n:Z
Hn:0 <= n

testbit (of_N a) n = testbit (of_N a) (of_N (to_N n))
a:N
n:Z
Hn:0 <= n

n = of_N (to_N n)
destruct n; trivial; now destruct Hn. Qed.
a:positive
n:Z

0 <= n -> testbit (pos a) n = N.testbit (N.pos a) (to_N n)
a:positive
n:Z

0 <= n -> testbit (pos a) n = N.testbit (N.pos a) (to_N n)
a:positive
n:Z
Hn:0 <= n

testbit (pos a) n = N.testbit (N.pos a) (to_N n)
now rewrite <- testbit_of_N'. Qed.
a:positive
n:Z

0 <= n -> testbit (neg a) n = negb (N.testbit (Pos.pred_N a) (to_N n))
a:positive
n:Z

0 <= n -> testbit (neg a) n = negb (N.testbit (Pos.pred_N a) (to_N n))
a:positive
n:Z
Hn:0 <= n

testbit (neg a) n = negb (N.testbit (Pos.pred_N a) (to_N n))
a:positive
n:Z
Hn:0 <= n

testbit (neg a) n = negb (testbit (of_N (Pos.pred_N a)) n)
a:positive
Hn:0 <= 0

testbit (neg a) 0 = negb (testbit (of_N (Pos.pred_N a)) 0)
a:positive
Hn:0 <= 0

odd (neg a) = negb (odd (of_N (Pos.pred_N a)))
now destruct a as [|[ | | ]| ]. Qed.

Proofs of specifications for bitwise operations

a:Z

div2 a = shiftr a 1
a:Z

div2 a = shiftr a 1
reflexivity. Qed.
n:Z

testbit 0 n = false
n:Z

testbit 0 n = false
now destruct n. Qed.
a, n:Z

n < 0 -> testbit a n = false
a, n:Z

n < 0 -> testbit a n = false
now destruct n. Qed.
a:Z

testbit (2 * a + 1) 0 = true
a:Z

testbit (2 * a + 1) 0 = true
now destruct a as [|a|[a|a|]]. Qed.
a:Z

testbit (2 * a) 0 = false
a:Z

testbit (2 * a) 0 = false
now destruct a. Qed.
a, n:Z

0 <= n -> testbit (2 * a + 1) (succ n) = testbit a n
a, n:Z

0 <= n -> testbit (2 * a + 1) (succ n) = testbit a n
a:Z

testbit (2 * a + 1) (succ 0) = testbit a 0
a:Z
n:positive
testbit (2 * a + 1) (succ (pos n)) = testbit a (pos n)
a:positive

negb (Pos.testbit (Pos.pred_double a) 0) = false
a:Z
n:positive
testbit (2 * a + 1) (succ (pos n)) = testbit a (pos n)
a:Z
n:positive

testbit (2 * a + 1) (succ (pos n)) = testbit a (pos n)
a:Z
n:positive

match match a with | 0 => 0 | pos y' => pos y'~0 | neg y' => neg y'~0 end + 1 with | 0 => false | pos a0 => Pos.testbit a0 (N.pos (n + 1)) | neg a0 => negb (N.testbit (Pos.pred_N a0) (N.pos (n + 1))) end = match a with | 0 => false | pos a0 => Pos.testbit a0 (N.pos n) | neg a0 => negb (N.testbit (Pos.pred_N a0) (N.pos n)) end
destruct a as [|a|[a|a|]]; simpl; trivial; rewrite ?Pos.add_1_r, ?Pos.pred_N_succ; now destruct n. Qed.
a, n:Z

0 <= n -> testbit (2 * a) (succ n) = testbit a n
a, n:Z

0 <= n -> testbit (2 * a) (succ n) = testbit a n
a:Z

testbit (2 * a) (succ 0) = testbit a 0
a:Z
n:positive
testbit (2 * a) (succ (pos n)) = testbit a (pos n)
a:positive

negb (Pos.testbit (Pos.pred_double a) 0) = false
a:Z
n:positive
testbit (2 * a) (succ (pos n)) = testbit a (pos n)
a:Z
n:positive

testbit (2 * a) (succ (pos n)) = testbit a (pos n)
a:Z
n:positive

match match a with | 0 => 0 | pos y' => pos y'~0 | neg y' => neg y'~0 end with | 0 => false | pos a0 => Pos.testbit a0 (N.pos (n + 1)) | neg a0 => negb (N.testbit (Pos.pred_N a0) (N.pos (n + 1))) end = match a with | 0 => false | pos a0 => Pos.testbit a0 (N.pos n) | neg a0 => negb (N.testbit (Pos.pred_N a0) (N.pos n)) end
destruct a as [|a|[a|a|]]; simpl; trivial; rewrite ?Pos.add_1_r, ?Pos.pred_N_succ; now destruct n. Qed.
Correctness proofs about Z.shiftr and Z.shiftl
a, n, m:Z

0 <= n -> 0 <= m -> testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z

0 <= n -> 0 <= m -> testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hn:0 <= n
Hm:0 <= m

testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hn:0 <= n
Hm:0 <= m

testbit (shiftl a (- n)) m = testbit a (m + n)
a, m:Z
Hm:0 <= m

testbit a m = testbit a (m + 0)
a:Z
n:positive
m:Z
Hm:0 <= m
testbit (Pos.iter div2 a n) m = testbit a (m + pos n)
a:Z
n:positive
m:Z
Hm:0 <= m

testbit (Pos.iter div2 a n) m = testbit a (m + pos n)
a:Z
n:positive
m:Z
Hm:0 <= m

forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
a:Z
n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
testbit (Pos.iter div2 a n) m = testbit a (m + pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N

testbit (Pos.iter div2 a n) m = testbit a (m + pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N

forall p : positive, 0 <= m + pos p
a:Z
n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 a n) m = testbit a (m + pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit (Pos.iter div2 a n) m = testbit a (m + pos n)
n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit (Pos.iter div2 0 n) m = testbit 0 (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (pos a) n) m = testbit (pos a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (neg a) n) m = testbit (neg a) (m + pos n)
(* a = 0 *)
n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit 0 m = testbit 0 (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (pos a) n) m = testbit (pos a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (neg a) n) m = testbit (neg a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit (Pos.iter div2 (pos a) n) m = testbit (pos a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (neg a) n) m = testbit (neg a) (m + pos n)
(* a > 0 *)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit (Pos.iter div2 (of_N (N.pos a)) n) m = testbit (pos a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (neg a) n) m = testbit (neg a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit (of_N (Pos.iter N.div2 (N.pos a) n)) m = testbit (pos a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (neg a) n) m = testbit (neg a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

N.testbit (Pos.iter N.div2 (N.pos a) n) (to_N m) = N.testbit (N.pos a) (to_N m + N.pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p
testbit (Pos.iter div2 (neg a) n) m = testbit (neg a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit (Pos.iter div2 (neg a) n) m = testbit (neg a) (m + pos n)
(* a < 0 *)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

testbit (neg (Pos.iter Pos.div2_up a n)) m = testbit (neg a) (m + pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

negb (N.testbit (Pos.pred_N (Pos.iter Pos.div2_up a n)) (to_N m)) = negb (N.testbit (Pos.pred_N a) (to_N m + N.pos n))
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

N.testbit (Pos.pred_N (Pos.iter Pos.div2_up a n)) (to_N m) = N.testbit (Pos.pred_N a) (to_N m + N.pos n)
a, n:positive
m:Z
Hm:0 <= m
H:forall p : positive, to_N (m + pos p) = (to_N m + N.pos p)%N
H0:forall p : positive, 0 <= m + pos p

N.testbit (Pos.iter N.div2 (Pos.pred_N a) n) (to_N m) = N.testbit (Pos.pred_N a) (to_N m + N.pos n)
exact (N.shiftr_spec' (Pos.pred_N a) (N.pos n) (to_N m)). Qed.
a, n, m:Z

m < n -> testbit (shiftl a n) m = false
a, n, m:Z

m < n -> testbit (shiftl a n) m = false
a, n, m:Z
H:m < n

testbit (shiftl a n) m = false
a:Z
n:positive
H:0 < pos n

testbit (Pos.iter (mul 2) a n) 0 = false
a:Z
n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) a n) (pos m) = false
a:Z
n, m:positive
H:pos m < pos n

testbit (Pos.iter (mul 2) a n) (pos m) = false
n, m:positive
H:pos m < pos n

testbit (Pos.iter (mul 2) 0 n) (pos m) = false
a, n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) (pos a) n) (pos m) = false
a, n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) (neg a) n) (pos m) = false
(* a = 0 *)
n, m:positive
H:pos m < pos n

testbit 0 (pos m) = false
a, n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) (pos a) n) (pos m) = false
a, n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) (neg a) n) (pos m) = false
a, n, m:positive
H:pos m < pos n

testbit (Pos.iter (mul 2) (pos a) n) (pos m) = false
a, n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) (neg a) n) (pos m) = false
(* a > 0 *)
a, n, m:positive
H:pos m < pos n

testbit (pos (Pos.iter xO a n)) (pos m) = false
a, n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) (neg a) n) (pos m) = false
a, n, m:positive
H:pos m < pos n

N.testbit (N.pos (Pos.iter xO a n)) (to_N (pos m)) = false
a, n, m:positive
H:pos m < pos n
testbit (Pos.iter (mul 2) (neg a) n) (pos m) = false
a, n, m:positive
H:pos m < pos n

testbit (Pos.iter (mul 2) (neg a) n) (pos m) = false
(* a < 0 *)
a, n, m:positive
H:pos m < pos n

testbit (neg (Pos.iter xO a n)) (pos m) = false
a, n, m:positive
H:pos m < pos n

negb (N.testbit (Pos.pred_N (Pos.iter xO a n)) (to_N (pos m))) = false
now rewrite (N.pos_pred_shiftl_low a (N.pos n)). Qed.
a, n, m:Z

0 <= m -> n <= m -> testbit (shiftl a n) m = testbit a (m - n)
a, n, m:Z

0 <= m -> n <= m -> testbit (shiftl a n) m = testbit a (m - n)
a, n, m:Z
Hm:0 <= m
H:n <= m

testbit (shiftl a n) m = testbit a (m - n)
a, m:Z
Hm, H:0 <= m

testbit (shiftl a 0) m = testbit a (m - 0)
a:Z
n:positive
m:Z
Hm:0 <= m
H:pos n <= m
testbit (shiftl a (pos n)) m = testbit a (m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a, m:Z
Hm, H:0 <= m

testbit a m = testbit a (m - 0)
a:Z
n:positive
m:Z
Hm:0 <= m
H:pos n <= m
testbit (shiftl a (pos n)) m = testbit a (m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:pos n <= m

testbit (shiftl a (pos n)) m = testbit a (m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
(* n > 0 *)
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m

testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m

0 <= pos m - pos n
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m

(0 ?= pos m - pos n) <> Gt
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n

testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n

to_N (pos m - pos n) = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:(pos n ?= pos m) <> Gt
H0:0 <= pos m - pos n

to_N (pos m - pos n) = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:(n ?= m)%positive <> Gt
H0:0 <= pos m - pos n

to_N (pos m - pos n) = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:(n ?= m)%positive <> Gt
H0:0 <= pos m - pos n

to_N (pos_sub m n) = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:(n ?= m)%positive <> Gt
H0:0 <= pos m - pos n

to_N match CompOpp (n ?= m)%positive with | Eq => 0 | Lt => neg (n - m) | Gt => pos (m - n) end = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:Eq <> Gt
H0:0 <= pos m - pos n
H':n = m

to_N match CompOpp Eq with | Eq => 0 | Lt => neg (n - m) | Gt => pos (m - n) end = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:Lt <> Gt
H0:0 <= pos m - pos n
H':(n < m)%positive
to_N match CompOpp Lt with | Eq => 0 | Lt => neg (n - m) | Gt => pos (m - n) end = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
m:positive
Hm:0 <= pos m
H:Eq <> Gt
H0:0 <= pos m - pos m

to_N match CompOpp Eq with | Eq => 0 | Lt => neg (m - m) | Gt => pos (m - m) end = (N.pos m - N.pos m)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:Lt <> Gt
H0:0 <= pos m - pos n
H':(n < m)%positive
to_N match CompOpp Lt with | Eq => 0 | Lt => neg (n - m) | Gt => pos (m - n) end = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:Lt <> Gt
H0:0 <= pos m - pos n
H':(n < m)%positive

to_N match CompOpp Lt with | Eq => 0 | Lt => neg (n - m) | Gt => pos (m - n) end = (N.pos m - N.pos n)%N
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:Lt <> Gt
H0:0 <= pos m - pos n
H':(n < m)%positive

N.pos (m - n) = match Pos.sub_mask m n with | Pos.IsPos p => N.pos p | _ => 0%N end
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, p:positive
Hm:0 <= pos (n + p)
H:Lt <> Gt
H':(n < n + p)%positive
H0:0 <= pos (n + p) - pos n

N.pos (n + p - n) = N.pos p
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, p:positive
Hm:0 <= pos (n + p)
H:Lt <> Gt
H':(n < n + p)%positive
H0:0 <= pos (n + p) - pos n

(n + p - n)%positive = p
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

testbit (shiftl a (pos n)) (pos m) = testbit a (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

testbit (Pos.iter (mul 2) 0 n) (pos m) = testbit 0 (pos m - pos n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (Pos.iter (mul 2) (pos p) n) (pos m) = testbit (pos p) (pos m - pos n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (Pos.iter (mul 2) (neg p) n) (pos m) = testbit (neg p) (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
(* ... a = 0 *)
n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

testbit 0 (pos m) = testbit 0 (pos m - pos n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (Pos.iter (mul 2) (pos p) n) (pos m) = testbit (pos p) (pos m - pos n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (Pos.iter (mul 2) (neg p) n) (pos m) = testbit (neg p) (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

testbit (Pos.iter (mul 2) (pos p) n) (pos m) = testbit (pos p) (pos m - pos n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (Pos.iter (mul 2) (neg p) n) (pos m) = testbit (neg p) (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
(* ... a > 0 *)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

testbit (pos (Pos.iter xO p n)) (pos m) = testbit (pos p) (pos m - pos n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (Pos.iter (mul 2) (neg p) n) (pos m) = testbit (neg p) (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

N.testbit (N.pos (Pos.iter xO p n)) (to_N (pos m)) = N.testbit (N.pos p) (N.pos m - N.pos n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N
testbit (Pos.iter (mul 2) (neg p) n) (pos m) = testbit (neg p) (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

testbit (Pos.iter (mul 2) (neg p) n) (pos m) = testbit (neg p) (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
(* ... a < 0 *)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

testbit (neg (Pos.iter xO p n)) (pos m) = testbit (neg p) (pos m - pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

negb (N.testbit (Pos.pred_N (Pos.iter xO p n)) (to_N (pos m))) = negb (N.testbit (Pos.pred_N p) (N.pos m - N.pos n))
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

N.testbit (Pos.pred_N (Pos.iter xO p n)) (to_N (pos m)) = N.testbit (Pos.pred_N p) (N.pos m - N.pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

N.testbit (Pos.pred_N (Pos.iter xO p n)) (N.pos m) = N.testbit (Pos.pred_N p) (N.pos m - N.pos n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
p, n, m:positive
Hm:0 <= pos m
H:pos n <= pos m
H0:0 <= pos m - pos n
EQ:to_N (pos m - pos n) = (N.pos m - N.pos n)%N

N.testbit (Pos.pred_N (Pos.iter xO p n)) (N.pos m) = N.testbit (N.shiftl (Pos.pred_N p) (N.pos n)) (N.pos m)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m
testbit (shiftl a (neg n)) m = testbit a (m - neg n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m

testbit (shiftl a (neg n)) m = testbit a (m - neg n)
(* n < 0 *)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m

testbit (shiftl a (neg n)) m = testbit a (m + - neg n)
a:Z
n:positive
m:Z
Hm:0 <= m
H:neg n <= m

testbit (Pos.iter div2 a n) m = testbit a (m + pos n)
now apply (shiftr_spec_aux a (pos n) m). Qed.
a, n, m:Z

0 <= m -> testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z

0 <= m -> testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m

testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:0 <= n

testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0

testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
LE:- n <= m

testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:m < - n
testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
LE:- n <= m

testbit (shiftl a (- n)) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:m < - n
testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
LE:- n <= m

testbit a (m - - n) = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:m < - n
testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:m < - n

testbit (shiftr a n) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:m < - n

testbit (shiftl a (- n)) m = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:m < - n

false = testbit a (m + n)
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:m < - n

m + n < 0
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:(m ?= - n) = Lt

m + n < 0
a, n, m:Z
Hm:0 <= m
H:n < 0
GT:(m - - n ?= 0) = Lt

m + n < 0
now destruct n. Qed.
Correctness proofs for bitwise operations
a, b, n:Z

testbit (lor a b) n = testbit a n || testbit b n
a, b, n:Z

testbit (lor a b) n = testbit a n || testbit b n
a, b, n:Z
Hn:0 <= n

testbit (lor a b) n = testbit a n || testbit b n
a, b:positive
n:Z
Hn:0 <= n

N.testbit (N.pos (Pos.lor a b)) (to_N n) = N.testbit (N.pos a) (to_N n) || N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.ldiff (Pos.pred_N b) (N.pos a)) (to_N n)) = N.testbit (N.pos a) (to_N n) || negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.ldiff (Pos.pred_N a) (N.pos b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) || N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.land (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) || negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n

negb (N.testbit (N.ldiff (Pos.pred_N b) (N.pos a)) (to_N n)) = N.testbit (N.pos a) (to_N n) || negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.ldiff (Pos.pred_N a) (N.pos b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) || N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.land (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) || negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n

negb (N.testbit (N.ldiff (Pos.pred_N a) (N.pos b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) || N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.land (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) || negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n

negb (N.testbit (N.land (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) || negb (N.testbit (Pos.pred_N b) (to_N n))
now rewrite N.land_spec, negb_andb. Qed.
a, b, n:Z

testbit (land a b) n = testbit a n && testbit b n
a, b, n:Z

testbit (land a b) n = testbit a n && testbit b n
a, b, n:Z
Hn:0 <= n

testbit (land a b) n = testbit a n && testbit b n
a, b:positive
n:Z
Hn:0 <= n

N.testbit (Pos.land a b) (to_N n) = N.testbit (N.pos a) (to_N n) && N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.ldiff (N.pos a) (Pos.pred_N b)) (to_N n) = N.testbit (N.pos a) (to_N n) && negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.ldiff (N.pos b) (Pos.pred_N a)) (to_N n) = negb (N.testbit (Pos.pred_N a) (to_N n)) && N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lor (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n

N.testbit (N.ldiff (N.pos a) (Pos.pred_N b)) (to_N n) = N.testbit (N.pos a) (to_N n) && negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.ldiff (N.pos b) (Pos.pred_N a)) (to_N n) = negb (N.testbit (Pos.pred_N a) (to_N n)) && N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lor (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n

N.testbit (N.ldiff (N.pos b) (Pos.pred_N a)) (to_N n) = negb (N.testbit (Pos.pred_N a) (to_N n)) && N.testbit (N.pos b) (to_N n)
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lor (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (N.testbit (Pos.pred_N b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n

negb (N.testbit (N.lor (Pos.pred_N a) (Pos.pred_N b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (N.testbit (Pos.pred_N b) (to_N n))
now rewrite N.lor_spec, negb_orb. Qed.
a, b, n:Z

testbit (ldiff a b) n = testbit a n && negb (testbit b n)
a, b, n:Z

testbit (ldiff a b) n = testbit a n && negb (testbit b n)
a, b, n:Z
Hn:0 <= n

testbit (ldiff a b) n = testbit a n && negb (testbit b n)
a, b:positive
n:Z
Hn:0 <= n

N.testbit (Pos.ldiff a b) (to_N n) = N.testbit (N.pos a) (to_N n) && negb (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.land (N.pos a) (Pos.pred_N b)) (to_N n) = N.testbit (N.pos a) (to_N n) && negb (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lor (Pos.pred_N a) (N.pos b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.ldiff (Pos.pred_N b) (Pos.pred_N a)) (to_N n) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n

N.testbit (N.land (N.pos a) (Pos.pred_N b)) (to_N n) = N.testbit (N.pos a) (to_N n) && negb (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lor (Pos.pred_N a) (N.pos b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.ldiff (Pos.pred_N b) (Pos.pred_N a)) (to_N n) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n

negb (N.testbit (N.lor (Pos.pred_N a) (N.pos b)) (to_N n)) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.ldiff (Pos.pred_N b) (Pos.pred_N a)) (to_N n) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n

N.testbit (N.ldiff (Pos.pred_N b) (Pos.pred_N a)) (to_N n) = negb (N.testbit (Pos.pred_N a) (to_N n)) && negb (negb (N.testbit (Pos.pred_N b) (to_N n)))
now rewrite N.ldiff_spec, negb_involutive, andb_comm. Qed.
a, b, n:Z

testbit (lxor a b) n = xorb (testbit a n) (testbit b n)
a, b, n:Z

testbit (lxor a b) n = xorb (testbit a n) (testbit b n)
a, b, n:Z
Hn:0 <= n

testbit (lxor a b) n = xorb (testbit a n) (testbit b n)
a, b:positive
n:Z
Hn:0 <= n

N.testbit (Pos.lxor a b) (to_N n) = xorb (N.testbit (N.pos a) (to_N n)) (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lxor (N.pos a) (Pos.pred_N b)) (to_N n)) = xorb (N.testbit (N.pos a) (to_N n)) (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lxor (Pos.pred_N a) (N.pos b)) (to_N n)) = xorb (negb (N.testbit (Pos.pred_N a) (to_N n))) (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.lxor (Pos.pred_N a) (Pos.pred_N b)) (to_N n) = xorb (negb (N.testbit (Pos.pred_N a) (to_N n))) (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n

negb (N.testbit (N.lxor (N.pos a) (Pos.pred_N b)) (to_N n)) = xorb (N.testbit (N.pos a) (to_N n)) (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n
negb (N.testbit (N.lxor (Pos.pred_N a) (N.pos b)) (to_N n)) = xorb (negb (N.testbit (Pos.pred_N a) (to_N n))) (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.lxor (Pos.pred_N a) (Pos.pred_N b)) (to_N n) = xorb (negb (N.testbit (Pos.pred_N a) (to_N n))) (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n

negb (N.testbit (N.lxor (Pos.pred_N a) (N.pos b)) (to_N n)) = xorb (negb (N.testbit (Pos.pred_N a) (to_N n))) (N.testbit (N.pos b) (to_N n))
a, b:positive
n:Z
Hn:0 <= n
N.testbit (N.lxor (Pos.pred_N a) (Pos.pred_N b)) (to_N n) = xorb (negb (N.testbit (Pos.pred_N a) (to_N n))) (negb (N.testbit (Pos.pred_N b) (to_N n)))
a, b:positive
n:Z
Hn:0 <= n

N.testbit (N.lxor (Pos.pred_N a) (Pos.pred_N b)) (to_N n) = xorb (negb (N.testbit (Pos.pred_N a) (to_N n))) (negb (N.testbit (Pos.pred_N b) (to_N n)))
now rewrite N.lxor_spec, xorb_negb_negb. Qed.
Generic properties of advanced functions.
Include ZExtraProp.
In generic statements, the predicates lt and le have been favored, whereas gt and ge don't even exist in the abstract layers. The use of gt and ge is hence not recommended. We provide here the bare minimal results to related them with lt and le.
n, m:Z

n > m <-> m < n
n, m:Z

n > m <-> m < n
n, m:Z

(n ?= m) = Gt <-> (m ?= n) = Lt
now rewrite compare_antisym, CompOpp_iff. Qed.
n, m:Z

n > m -> m < n
n, m:Z

n > m -> m < n
apply gt_lt_iff. Qed.
n, m:Z

n < m -> m > n
n, m:Z

n < m -> m > n
apply gt_lt_iff. Qed.
n, m:Z

n >= m <-> m <= n
n, m:Z

n >= m <-> m <= n
n, m:Z

(n ?= m) <> Lt <-> (m ?= n) <> Gt
now rewrite compare_antisym, CompOpp_iff. Qed.
n, m:Z

n >= m -> m <= n
n, m:Z

n >= m -> m <= n
apply ge_le_iff. Qed.
n, m:Z

n <= m -> m >= n
n, m:Z

n <= m -> m >= n
apply ge_le_iff. Qed.
We provide a tactic converting from one style to the other.
Ltac swap_greater := rewrite ?gt_lt_iff in *; rewrite ?ge_le_iff in *.
Similarly, the boolean comparisons ltb and leb are favored over their dual gtb and geb. We prove here the equivalence and a few minimal results.
n, m:Z

(n >? m) = (m <? n)
n, m:Z

(n >? m) = (m <? n)
n, m:Z

match n ?= m with | Gt => true | _ => false end = match m ?= n with | Lt => true | _ => false end
n, m:Z

match CompOpp (m ?= n) with | Gt => true | _ => false end = match m ?= n with | Lt => true | _ => false end
now case compare. Qed.
n, m:Z

(n >=? m) = (m <=? n)
n, m:Z

(n >=? m) = (m <=? n)
n, m:Z

match n ?= m with | Lt => false | _ => true end = match m ?= n with | Gt => false | _ => true end
n, m:Z

match CompOpp (m ?= n) with | Lt => false | _ => true end = match m ?= n with | Gt => false | _ => true end
now case compare. Qed.
n, m:Z

(n >? m) = true <-> m < n
n, m:Z

(n >? m) = true <-> m < n
n, m:Z

(m <? n) = true <-> m < n
apply ltb_lt. Qed.
n, m:Z

(n >=? m) = true <-> m <= n
n, m:Z

(n >=? m) = true <-> m <= n
n, m:Z

(m <=? n) = true <-> m <= n
apply leb_le. Qed.
n, m:Z

BoolSpec (m < n) (n <= m) (n >? m)
n, m:Z

BoolSpec (m < n) (n <= m) (n >? m)
n, m:Z

BoolSpec (m < n) (n <= m) (m <? n)
apply ltb_spec. Qed.
n, m:Z

BoolSpec (m <= n) (n < m) (n >=? m)
n, m:Z

BoolSpec (m <= n) (n < m) (n >=? m)
n, m:Z

BoolSpec (m <= n) (n < m) (m <=? n)
apply leb_spec. Qed.
TODO : to add in Numbers ?
n, m, p:Z

n + m = n + p -> m = p
n, m, p:Z

n + m = n + p -> m = p
exact (proj1 (add_cancel_l m p n)). Qed.
n, m, p:Z

p <> 0 -> p * n = p * m -> n = m
n, m, p:Z

p <> 0 -> p * n = p * m -> n = m
exact (fun Hp => proj1 (mul_cancel_l n m p Hp)). Qed.
n, m, p:Z

p <> 0 -> n * p = m * p -> n = m
n, m, p:Z

p <> 0 -> n * p = m * p -> n = m
exact (fun Hp => proj1 (mul_cancel_r n m p Hp)). Qed.
n:Z

- n = n * -1
n:Z

- n = n * -1
n:Z

- n = -1 * n
now destruct n. Qed.
n:Z

n + n = 2 * n
n:Z

n + n = 2 * n
n:Z

n + n = (1 + 1) * n
now rewrite mul_add_distr_r, !mul_1_l. Qed.

Comparison and opposite

n, m:Z

(- n ?= - m) = (m ?= n)
n, m:Z

(- n ?= - m) = (m ?= n)
destruct n, m; simpl; trivial; intros; now rewrite <- Pos.compare_antisym. Qed.

Comparison and addition

n, m, p:Z

(n + m ?= n + p) = (m ?= p)
n, m, p:Z

(n + m ?= n + p) = (m ?= p)
n, m, p:Z

(n + m - (n + p) ?= 0) = (m - p ?= 0)
n, m, p:Z

n + m - (n + p) = m - p
n, m, p:Z

n + m + - (n + p) = m + - p
n, m, p:Z

m + n + - n + - p = m + - p
n, m, p:Z

m + n + - n = m
now rewrite <- add_assoc, add_opp_diag_r, add_0_r. Qed.

testbit in terms of comparison.

a, n, i:Z
H:0 <= n

testbit (a mod 2 ^ n) i = (i <? n) && testbit a i
a, n, i:Z
H:0 <= n

testbit (a mod 2 ^ n) i = (i <? n) && testbit a i
destruct (ltb_spec i n); rewrite ?mod_pow2_bits_low, ?mod_pow2_bits_high by auto; auto. Qed.
n, i:Z
H:0 <= n

testbit (ones n) i = (0 <=? i) && (i <? n)
n, i:Z
H:0 <= n

testbit (ones n) i = (0 <=? i) && (i <? n)
destruct (leb_spec 0 i), (ltb_spec i n); cbn; rewrite ?testbit_neg_r, ?ones_spec_low, ?ones_spec_high by auto; trivial. Qed.
n, i:Z
Hn:0 <= n
Hi:0 <= i

testbit (ones n) i = (i <? n)
n, i:Z
Hn:0 <= n
Hi:0 <= i

testbit (ones n) i = (i <? n)
n, i:Z
Hn:0 <= n
Hi:0 <= i

(0 <=? i) && (i <? n) = (i <? n)
destruct (leb_spec 0 i); cbn; solve [ trivial | destruct (proj1 (Z.le_ngt _ _) Hi ltac:(eassumption)) ]. Qed. End Z. Bind Scope Z_scope with Z.t Z.
Re-export Notations
Numeral Notation Z Z.of_int Z.to_int : Z_scope.

Infix "+" := Z.add : Z_scope.
Notation "- x" := (Z.opp x) : Z_scope.
Infix "-" := Z.sub : Z_scope.
Infix "*" := Z.mul : Z_scope.
Infix "^" := Z.pow : Z_scope.
Infix "/" := Z.div : Z_scope.
Infix "mod" := Z.modulo (at level 40, no associativity) : Z_scope.
Infix "÷" := Z.quot (at level 40, left associativity) : Z_scope.
Infix "?=" := Z.compare (at level 70, no associativity) : Z_scope.
Infix "=?" := Z.eqb (at level 70, no associativity) : Z_scope.
Infix "<=?" := Z.leb (at level 70, no associativity) : Z_scope.
Infix "<?" := Z.ltb (at level 70, no associativity) : Z_scope.
Infix ">=?" := Z.geb (at level 70, no associativity) : Z_scope.
Infix ">?" := Z.gtb (at level 70, no associativity) : Z_scope.
Notation "( x | y )" := (Z.divide x y) (at level 0) : Z_scope.
Infix "<=" := Z.le : Z_scope.
Infix "<" := Z.lt : Z_scope.
Infix ">=" := Z.ge : Z_scope.
Infix ">" := Z.gt : Z_scope.
Notation "x <= y <= z" := (x <= y /\ y <= z) : Z_scope.
Notation "x <= y < z" := (x <= y /\ y < z) : Z_scope.
Notation "x < y < z" := (x < y /\ y < z) : Z_scope.
Notation "x < y <= z" := (x < y /\ y <= z) : Z_scope.
Conversions from / to positive numbers
Module Pos2Z.

p:positive

Z.to_pos (Zpos p) = p
p:positive

Z.to_pos (Zpos p) = p
reflexivity. Qed.
p, q:positive

Zpos p = Zpos q -> p = q
p, q:positive

Zpos p = Zpos q -> p = q
now injection 1. Qed.
p, q:positive

Zpos p = Zpos q <-> p = q
p, q:positive

Zpos p = Zpos q <-> p = q
p, q:positive

Zpos p = Zpos q -> p = q
p, q:positive
p = q -> Zpos p = Zpos q
p, q:positive

p = q -> Zpos p = Zpos q
intros; now f_equal. Qed.
p:positive

0 < Zpos p
p:positive

0 < Zpos p
reflexivity. Qed.
p:positive

0 <= Zpos p
p:positive

0 <= Zpos p
easy. Qed.

1 = 1

1 = 1
reflexivity. Qed.
p:positive

Zpos p~0 = 2 * Zpos p
p:positive

Zpos p~0 = 2 * Zpos p
reflexivity. Qed.
p:positive

Zpos p~1 = 2 * Zpos p + 1
p:positive

Zpos p~1 = 2 * Zpos p + 1
reflexivity. Qed.
p:positive

Zpos (Pos.succ p) = Z.succ (Zpos p)
p:positive

Zpos (Pos.succ p) = Z.succ (Zpos p)
p:positive

Zpos (Pos.succ p) = Zpos (p + 1)
now rewrite Pos.add_1_r. Qed.
p, q:positive

Zpos (p + q) = Zpos p + Zpos q
p, q:positive

Zpos (p + q) = Zpos p + Zpos q
reflexivity. Qed.
p, q:positive

(p < q)%positive -> Zpos (q - p) = Zpos q - Zpos p
p, q:positive

(p < q)%positive -> Zpos (q - p) = Zpos q - Zpos p
p, q:positive
H:(p < q)%positive

Zpos (q - p) = Zpos q - Zpos p
p, q:positive
H:(p < q)%positive

Zpos (q - p) = Z.pos_sub q p
now rewrite Z.pos_sub_gt. Qed.
p, q:positive

Zpos (p - q) = Z.max 1 (Zpos p - Zpos q)
p, q:positive

Zpos (p - q) = Z.max 1 (Zpos p - Zpos q)
p, q:positive

Zpos (p - q) = Z.max 1 (Z.pos_sub p q)
p, q:positive

Zpos (p - q) = Z.max 1 match (p ?= q)%positive with | Eq => 0 | Lt => Zneg (q - p) | Gt => Zpos (p - q) end
p, q:positive
H:p = q

Zpos (p - q) = Z.max 1 0
p, q:positive
H:(p < q)%positive
Zpos (p - q) = Z.max 1 (Zneg (q - p))
p, q:positive
H:(q < p)%positive
Zpos (p - q) = Z.max 1 (Zpos (p - q))
p, q:positive
H:p = q

Zpos (p - q) = Z.max 1 0
subst; now rewrite Pos.sub_diag.
p, q:positive
H:(p < q)%positive

Zpos (p - q) = Z.max 1 (Zneg (q - p))
now rewrite Pos.sub_lt.
p, q:positive
H:(q < p)%positive

Zpos (p - q) = Z.max 1 (Zpos (p - q))
now destruct (p-q)%positive. Qed.
p:positive

p <> 1%positive -> Zpos (Pos.pred p) = Z.pred (Zpos p)
p:positive

p <> 1%positive -> Zpos (Pos.pred p) = Z.pred (Zpos p)
destruct p; easy || now destruct 1. Qed.
p, q:positive

Zpos (p * q) = Zpos p * Zpos q
p, q:positive

Zpos (p * q) = Zpos p * Zpos q
reflexivity. Qed.
p, q:positive

Zpos (p ^ q) = Z.pow_pos (Zpos p) q
p, q:positive

Zpos (p ^ q) = Z.pow_pos (Zpos p) q
now apply Pos.iter_swap_gen. Qed.
p, q:positive

Zpos (p ^ q) = Zpos p ^ Zpos q
p, q:positive

Zpos (p ^ q) = Zpos p ^ Zpos q
apply inj_pow_pos. Qed.
p:positive

Zpos (Pos.square p) = Z.square (Zpos p)
p:positive

Zpos (Pos.square p) = Z.square (Zpos p)
reflexivity. Qed.
p, q:positive

(p ?= q)%positive = (Zpos p ?= Zpos q)
p, q:positive

(p ?= q)%positive = (Zpos p ?= Zpos q)
reflexivity. Qed.
p, q:positive

(p <=? q)%positive = (Zpos p <=? Zpos q)
p, q:positive

(p <=? q)%positive = (Zpos p <=? Zpos q)
reflexivity. Qed.
p, q:positive

(p <? q)%positive = (Zpos p <? Zpos q)
p, q:positive

(p <? q)%positive = (Zpos p <? Zpos q)
reflexivity. Qed.
p, q:positive

(p =? q)%positive = (Zpos p =? Zpos q)
p, q:positive

(p =? q)%positive = (Zpos p =? Zpos q)
reflexivity. Qed.
p, q:positive

Zpos (Pos.max p q) = Z.max (Zpos p) (Zpos q)
p, q:positive

Zpos (Pos.max p q) = Z.max (Zpos p) (Zpos q)
p, q:positive

Zpos match (p ?= q)%positive with | Gt => p | _ => q end = match Zpos p ?= Zpos q with | Lt => Zpos q | _ => Zpos p end
p, q:positive

Zpos match Zpos p ?= Zpos q with | Gt => p | _ => q end = match Zpos p ?= Zpos q with | Lt => Zpos q | _ => Zpos p end
now case Z.compare_spec. Qed.
p, q:positive

Zpos (Pos.min p q) = Z.min (Zpos p) (Zpos q)
p, q:positive

Zpos (Pos.min p q) = Z.min (Zpos p) (Zpos q)
p, q:positive

Zpos match (p ?= q)%positive with | Gt => q | _ => p end = match Zpos p ?= Zpos q with | Gt => Zpos q | _ => Zpos p end
p, q:positive

Zpos match Zpos p ?= Zpos q with | Gt => q | _ => p end = match Zpos p ?= Zpos q with | Gt => Zpos q | _ => Zpos p end
now case Z.compare_spec. Qed.
p:positive

Zpos (Pos.sqrt p) = Z.sqrt (Zpos p)
p:positive

Zpos (Pos.sqrt p) = Z.sqrt (Zpos p)
reflexivity. Qed.
p, q:positive

Zpos (Pos.gcd p q) = Z.gcd (Zpos p) (Zpos q)
p, q:positive

Zpos (Pos.gcd p q) = Z.gcd (Zpos p) (Zpos q)
reflexivity. Qed.
p, q:positive

(Zpos p | Zpos q) <-> (p | q)%positive
p, q:positive

(Zpos p | Zpos q) <-> (p | q)%positive
apply Z.divide_Zpos. Qed.
a:positive
n:Z

0 <= n -> Z.testbit (Zpos a) n = N.testbit (N.pos a) (Z.to_N n)
a:positive
n:Z

0 <= n -> Z.testbit (Zpos a) n = N.testbit (N.pos a) (Z.to_N n)
apply Z.testbit_Zpos. Qed.
Some results concerning Z.neg and Z.pos
p, q:positive

Zneg p = Zneg q -> p = q
p, q:positive

Zneg p = Zneg q -> p = q
now injection 1. Qed.
p, q:positive

Zneg p = Zneg q <-> p = q
p, q:positive

Zneg p = Zneg q <-> p = q
p, q:positive

Zneg p = Zneg q -> p = q
p, q:positive
p = q -> Zneg p = Zneg q
p, q:positive

p = q -> Zneg p = Zneg q
intros; now f_equal. Qed.
p, q:positive

Zpos p = Zpos q -> p = q
p, q:positive

Zpos p = Zpos q -> p = q
now injection 1. Qed.
p, q:positive

Zpos p = Zpos q <-> p = q
p, q:positive

Zpos p = Zpos q <-> p = q
p, q:positive

Zpos p = Zpos q -> p = q
p, q:positive
p = q -> Zpos p = Zpos q
p, q:positive

p = q -> Zpos p = Zpos q
intros; now f_equal. Qed.
p:positive

Zneg p < 0
p:positive

Zneg p < 0
reflexivity. Qed.
p:positive

Zneg p <= 0
p:positive

Zneg p <= 0
easy. Qed.
p:positive

0 < Zpos p
p:positive

0 < Zpos p
reflexivity. Qed.
p:positive

0 <= Zpos p
p:positive

0 <= Zpos p
easy. Qed.
p, q:positive

Zneg p <= Zpos q
p, q:positive

Zneg p <= Zpos q
easy. Qed.
p, q:positive

Zneg p < Zpos q
p, q:positive

Zneg p < Zpos q
easy. Qed.
p, q:positive

(q <= p)%positive -> Zneg p <= Zneg q
p, q:positive

(q <= p)%positive -> Zneg p <= Zneg q
p, q:positive
H:(q <= p)%positive

CompOpp (p ?= q)%positive <> Gt
now rewrite <- Pos.compare_antisym. Qed.
p, q:positive

(q < p)%positive -> Zneg p < Zneg q
p, q:positive

(q < p)%positive -> Zneg p < Zneg q
p, q:positive
H:(q < p)%positive

CompOpp (p ?= q)%positive = Lt
now rewrite <- Pos.compare_antisym. Qed.
p, q:positive

(p <= q)%positive -> Zpos p <= Zpos q
p, q:positive

(p <= q)%positive -> Zpos p <= Zpos q
easy. Qed.
p, q:positive

(p < q)%positive -> Zpos p < Zpos q
p, q:positive

(p < q)%positive -> Zpos p < Zpos q
easy. Qed.
p:positive

Zneg p~0 = 2 * Zneg p
p:positive

Zneg p~0 = 2 * Zneg p
reflexivity. Qed.
p:positive

Zneg p~1 = 2 * Zneg p - 1
p:positive

Zneg p~1 = 2 * Zneg p - 1
reflexivity. Qed.
p:positive

Zpos p~0 = 2 * Zpos p
p:positive

Zpos p~0 = 2 * Zpos p
reflexivity. Qed.
p:positive

Zpos p~1 = 2 * Zpos p + 1
p:positive

Zpos p~1 = 2 * Zpos p + 1
reflexivity. Qed.
p:positive

- Zneg p = Zpos p
p:positive

- Zneg p = Zpos p
reflexivity. Qed.
p:positive

- Zpos p = Zneg p
p:positive

- Zpos p = Zneg p
reflexivity. Qed.
p, q:positive

Zneg p + Zneg q = Zneg (p + q)
p, q:positive

Zneg p + Zneg q = Zneg (p + q)
reflexivity. Qed.
p, q:positive

Zpos p + Zneg q = Z.pos_sub p q
p, q:positive

Zpos p + Zneg q = Z.pos_sub p q
reflexivity. Qed.
p, q:positive

Zneg p + Zpos q = Z.pos_sub q p
p, q:positive

Zneg p + Zpos q = Z.pos_sub q p
reflexivity. Qed.
p, q:positive

Zpos p + Zpos q = Zpos (p + q)
p, q:positive

Zpos p + Zpos q = Zpos (p + q)
reflexivity. Qed.
n:Z
p:positive

(n | Zpos p) <-> (n | Zneg p)
n:Z
p:positive

(n | Zpos p) <-> (n | Zneg p)
apply Z.divide_Zpos_Zneg_r. Qed.
n:Z
p:positive

(Zpos p | n) <-> (Zneg p | n)
n:Z
p:positive

(Zpos p | n) <-> (Zneg p | n)
apply Z.divide_Zpos_Zneg_l. Qed.
a:positive
n:Z

0 <= n -> Z.testbit (Zneg a) n = negb (N.testbit (Pos.pred_N a) (Z.to_N n))
a:positive
n:Z

0 <= n -> Z.testbit (Zneg a) n = negb (N.testbit (Pos.pred_N a) (Z.to_N n))
apply Z.testbit_Zneg. Qed.
a:positive
n:Z

0 <= n -> Z.testbit (Zpos a) n = N.testbit (N.pos a) (Z.to_N n)
a:positive
n:Z

0 <= n -> Z.testbit (Zpos a) n = N.testbit (N.pos a) (Z.to_N n)
apply Z.testbit_Zpos. Qed. End Pos2Z. Module Z2Pos.
x:Z

0 < x -> Zpos (Z.to_pos x) = x
x:Z

0 < x -> Zpos (Z.to_pos x) = x
now destruct x. Qed.
x, y:Z

0 < x -> 0 < y -> Z.to_pos x = Z.to_pos y -> x = y
x, y:Z

0 < x -> 0 < y -> Z.to_pos x = Z.to_pos y -> x = y
p:positive
y:Z

0 < Zpos p -> 0 < y -> p = Z.to_pos y -> Zpos p = y
y:Z
H:0 < y

Zpos (Z.to_pos y) = y
now apply id. Qed.
x, y:Z

0 < x -> 0 < y -> Z.to_pos x = Z.to_pos y <-> x = y
x, y:Z

0 < x -> 0 < y -> Z.to_pos x = Z.to_pos y <-> x = y
x, y:Z
H:0 < x
H0:0 < y

Z.to_pos x = Z.to_pos y -> x = y
x, y:Z
H:0 < x
H0:0 < y
x = y -> Z.to_pos x = Z.to_pos y
x, y:Z
H:0 < x
H0:0 < y

x = y -> Z.to_pos x = Z.to_pos y
intros; now f_equal. Qed.
x:Z

x <= 0 -> Z.to_pos x = 1%positive
x:Z

x <= 0 -> Z.to_pos x = 1%positive
p:positive

Zpos p <= 0 -> Z.to_pos (Zpos p) = 1%positive
now destruct 1. Qed.

Z.to_pos 1 = 1%positive

Z.to_pos 1 = 1%positive
reflexivity. Qed.
x:Z

0 < x -> Z.to_pos (Z.double x) = ((Z.to_pos x)~0)%positive
x:Z

0 < x -> Z.to_pos (Z.double x) = ((Z.to_pos x)~0)%positive
now destruct x. Qed.
x:Z

0 < x -> Z.to_pos (Z.succ_double x) = ((Z.to_pos x)~1)%positive
x:Z

0 < x -> Z.to_pos (Z.succ_double x) = ((Z.to_pos x)~1)%positive
now destruct x. Qed.
x:Z

0 < x -> Z.to_pos (Z.succ x) = Pos.succ (Z.to_pos x)
x:Z

0 < x -> Z.to_pos (Z.succ x) = Pos.succ (Z.to_pos x)
p:positive

0 < Zpos p -> Z.to_pos (Z.succ (Zpos p)) = Pos.succ (Z.to_pos (Zpos p))
p:positive

0 < Zpos p -> (p + 1)%positive = Pos.succ p
now rewrite Pos.add_1_r. Qed.
x, y:Z

0 < x -> 0 < y -> Z.to_pos (x + y) = (Z.to_pos x + Z.to_pos y)%positive
x, y:Z

0 < x -> 0 < y -> Z.to_pos (x + y) = (Z.to_pos x + Z.to_pos y)%positive
destruct x; easy || now destruct y. Qed.
x, y:Z

0 < x < y -> Z.to_pos (y - x) = (Z.to_pos y - Z.to_pos x)%positive
x, y:Z

0 < x < y -> Z.to_pos (y - x) = (Z.to_pos y - Z.to_pos x)%positive
p:positive
y:Z

0 < Zpos p < y -> Z.to_pos (y - Zpos p) = (Z.to_pos y - Z.to_pos (Zpos p))%positive
p, p0:positive

0 < Zpos p < Zpos p0 -> Z.to_pos (Zpos p0 - Zpos p) = (Z.to_pos (Zpos p0) - Z.to_pos (Zpos p))%positive
p, p0:positive

0 < Zpos p < Zpos p0 -> Z.to_pos (Z.pos_sub p0 p) = (p0 - p)%positive
p, p0:positive
H:0 < Zpos p < Zpos p0

Z.to_pos (Z.pos_sub p0 p) = (p0 - p)%positive
now rewrite Z.pos_sub_gt. Qed.
x:Z

1 < x -> Z.to_pos (Z.pred x) = Pos.pred (Z.to_pos x)
x:Z

1 < x -> Z.to_pos (Z.pred x) = Pos.pred (Z.to_pos x)
now destruct x as [|[x|x|]|]. Qed.
x, y:Z

0 < x -> 0 < y -> Z.to_pos (x * y) = (Z.to_pos x * Z.to_pos y)%positive
x, y:Z

0 < x -> 0 < y -> Z.to_pos (x * y) = (Z.to_pos x * Z.to_pos y)%positive
destruct x; easy || now destruct y. Qed.
x, y:Z

0 < x -> 0 < y -> Z.to_pos (x ^ y) = (Z.to_pos x ^ Z.to_pos y)%positive
x, y:Z

0 < x -> 0 < y -> Z.to_pos (x ^ y) = (Z.to_pos x ^ Z.to_pos y)%positive
x, y:Z
H:0 < x
H0:0 < y

Z.to_pos (x ^ y) = (Z.to_pos x ^ Z.to_pos y)%positive
x, y:Z
H:0 < x
H0:0 < y

Zpos (Z.to_pos (x ^ y)) = Zpos (Z.to_pos x ^ Z.to_pos y)
x, y:Z
H:0 < x
H0:0 < y

0 < x ^ y
x, y:Z
H:0 < x
H0:0 < y

0 < x
x, y:Z
H:0 < x
H0:0 < y
0 <= y
x, y:Z
H:0 < x
H0:0 < y

0 <= y
now apply Z.lt_le_incl. Qed.
x:Z
p:positive

0 < x -> Z.to_pos (Z.pow_pos x p) = (Z.to_pos x ^ p)%positive
x:Z
p:positive

0 < x -> Z.to_pos (Z.pow_pos x p) = (Z.to_pos x ^ p)%positive
x:Z
p:positive
H:0 < x

Z.to_pos (Z.pow_pos x p) = (Z.to_pos x ^ p)%positive
now apply (inj_pow x (Z.pos p)). Qed.
x, y:Z

0 < x -> 0 < y -> (x ?= y) = (Z.to_pos x ?= Z.to_pos y)%positive
x, y:Z

0 < x -> 0 < y -> (x ?= y) = (Z.to_pos x ?= Z.to_pos y)%positive
destruct x; easy || now destruct y. Qed.
x, y:Z

0 < x -> 0 < y -> (x <=? y) = (Z.to_pos x <=? Z.to_pos y)%positive
x, y:Z

0 < x -> 0 < y -> (x <=? y) = (Z.to_pos x <=? Z.to_pos y)%positive
destruct x; easy || now destruct y. Qed.
x, y:Z

0 < x -> 0 < y -> (x <? y) = (Z.to_pos x <? Z.to_pos y)%positive
x, y:Z

0 < x -> 0 < y -> (x <? y) = (Z.to_pos x <? Z.to_pos y)%positive
destruct x; easy || now destruct y. Qed.
x, y:Z

0 < x -> 0 < y -> (x =? y) = (Z.to_pos x =? Z.to_pos y)%positive
x, y:Z

0 < x -> 0 < y -> (x =? y) = (Z.to_pos x =? Z.to_pos y)%positive
destruct x; easy || now destruct y. Qed.
x, y:Z

Z.to_pos (Z.max x y) = Pos.max (Z.to_pos x) (Z.to_pos y)
x, y:Z

Z.to_pos (Z.max x y) = Pos.max (Z.to_pos x) (Z.to_pos y)
y:Z

Z.to_pos (Z.max 0 y) = Z.to_pos y
p:positive
y:Z
Z.to_pos (Z.max (Zpos p) y) = Pos.max p (Z.to_pos y)
p:positive
y:Z
Z.to_pos (Z.max (Zneg p) y) = Z.to_pos y
y:Z

Z.to_pos (Z.max 0 y) = Z.to_pos y
now destruct y.
p:positive
y:Z

Z.to_pos (Z.max (Zpos p) y) = Pos.max p (Z.to_pos y)
destruct y; simpl; now rewrite ?Pos.max_1_r, <- ?Pos2Z.inj_max.
p:positive
y:Z

Z.to_pos (Z.max (Zneg p) y) = Z.to_pos y
p, p0:positive

Z.to_pos (Z.max (Zneg p) (Zneg p0)) = 1%positive
p, p0:positive

Z.max (Zneg p) (Zneg p0) <= 0
now apply Z.max_lub. Qed.
x, y:Z

Z.to_pos (Z.min x y) = Pos.min (Z.to_pos x) (Z.to_pos y)
x, y:Z

Z.to_pos (Z.min x y) = Pos.min (Z.to_pos x) (Z.to_pos y)
y:Z

Z.to_pos (Z.min 0 y) = 1%positive
p:positive
y:Z
Z.to_pos (Z.min (Zpos p) y) = Pos.min p (Z.to_pos y)
p:positive
y:Z
Z.to_pos (Z.min (Zneg p) y) = 1%positive
y:Z

Z.to_pos (Z.min 0 y) = 1%positive
now destruct y.
p:positive
y:Z

Z.to_pos (Z.min (Zpos p) y) = Pos.min p (Z.to_pos y)
destruct y; simpl; now rewrite ?Pos.min_1_r, <- ?Pos2Z.inj_min.
p:positive
y:Z

Z.to_pos (Z.min (Zneg p) y) = 1%positive
p, p0:positive

Z.to_pos (Z.min (Zneg p) (Zneg p0)) = 1%positive
p, p0:positive

Z.min (Zneg p) (Zneg p0) <= 0
p, p0:positive

Zneg p <= 0 \/ Zneg p0 <= 0
now left. Qed.
x:Z

Z.to_pos (Z.sqrt x) = Pos.sqrt (Z.to_pos x)
x:Z

Z.to_pos (Z.sqrt x) = Pos.sqrt (Z.to_pos x)
now destruct x. Qed.
x, y:Z

0 < x -> 0 < y -> Z.to_pos (Z.gcd x y) = Pos.gcd (Z.to_pos x) (Z.to_pos y)
x, y:Z

0 < x -> 0 < y -> Z.to_pos (Z.gcd x y) = Pos.gcd (Z.to_pos x) (Z.to_pos y)
destruct x; easy || now destruct y. Qed. End Z2Pos.
Compatibility Notations
Notation Zdouble_plus_one := Z.succ_double (only parsing).
Notation Zdouble_minus_one := Z.pred_double (only parsing).
Notation ZPminus := Z.pos_sub (only parsing).
Notation Zplus := Z.add (only parsing). (* Slightly incompatible *)
Notation Zminus := Z.sub (only parsing).
Notation Zmult := Z.mul (only parsing).
Notation Z_of_nat := Z.of_nat (only parsing).
Notation Z_of_N := Z.of_N (only parsing).

Notation Zind := Z.peano_ind (only parsing).
Notation Zplus_0_l := Z.add_0_l (only parsing).
Notation Zplus_0_r := Z.add_0_r (only parsing).
Notation Zplus_comm := Z.add_comm (only parsing).
Notation Zopp_plus_distr := Z.opp_add_distr (only parsing).
Notation Zplus_opp_r := Z.add_opp_diag_r (only parsing).
Notation Zplus_opp_l := Z.add_opp_diag_l (only parsing).
Notation Zplus_assoc := Z.add_assoc (only parsing).
Notation Zplus_permute := Z.add_shuffle3 (only parsing).
Notation Zplus_reg_l := Z.add_reg_l (only parsing).
Notation Zplus_succ_l := Z.add_succ_l (only parsing).
Notation Zplus_succ_comm := Z.add_succ_comm (only parsing).
Notation Zsucc_discr := Z.neq_succ_diag_r (only parsing).
Notation Zsucc'_discr := Z.neq_succ_diag_r (only parsing).
Notation Zminus_0_r := Z.sub_0_r (only parsing).
Notation Zminus_diag := Z.sub_diag (only parsing).
Notation Zminus_plus_distr := Z.sub_add_distr (only parsing).
Notation Zminus_succ_r := Z.sub_succ_r (only parsing).
Notation Zminus_plus := Z.add_simpl_l (only parsing).
Notation Zmult_0_l := Z.mul_0_l (only parsing).
Notation Zmult_0_r := Z.mul_0_r (only parsing).
Notation Zmult_1_l := Z.mul_1_l (only parsing).
Notation Zmult_1_r := Z.mul_1_r (only parsing).
Notation Zmult_comm := Z.mul_comm (only parsing).
Notation Zmult_assoc := Z.mul_assoc (only parsing).
Notation Zmult_permute := Z.mul_shuffle3 (only parsing).
Notation Zmult_1_inversion_l := Z.mul_eq_1 (only parsing).
Notation Zdouble_mult := Z.double_spec (only parsing).
Notation Zdouble_plus_one_mult := Z.succ_double_spec (only parsing).
Notation Zopp_mult_distr_l_reverse := Z.mul_opp_l (only parsing).
Notation Zmult_opp_opp := Z.mul_opp_opp (only parsing).
Notation Zmult_opp_comm := Z.mul_opp_comm (only parsing).
Notation Zopp_eq_mult_neg_1 := Z.opp_eq_mul_m1 (only parsing).
Notation Zmult_plus_distr_r := Z.mul_add_distr_l (only parsing).
Notation Zmult_plus_distr_l := Z.mul_add_distr_r (only parsing).
Notation Zmult_minus_distr_r := Z.mul_sub_distr_r (only parsing).
Notation Zmult_reg_l := Z.mul_reg_l (only parsing).
Notation Zmult_reg_r := Z.mul_reg_r (only parsing).
Notation Zmult_succ_l := Z.mul_succ_l (only parsing).
Notation Zmult_succ_r := Z.mul_succ_r (only parsing).

Notation Zpos_xI := Pos2Z.inj_xI (only parsing).
Notation Zpos_xO := Pos2Z.inj_xO (only parsing).
Notation Zneg_xI := Pos2Z.neg_xI (only parsing).
Notation Zneg_xO := Pos2Z.neg_xO (only parsing).
Notation Zopp_neg := Pos2Z.opp_neg (only parsing).
Notation Zpos_succ_morphism := Pos2Z.inj_succ (only parsing).
Notation Zpos_mult_morphism := Pos2Z.inj_mul (only parsing).
Notation Zpos_minus_morphism := Pos2Z.inj_sub (only parsing).
Notation Zpos_eq_rev := Pos2Z.inj (only parsing).
Notation Zpos_plus_distr := Pos2Z.inj_add (only parsing).
Notation Zneg_plus_distr := Pos2Z.add_neg_neg (only parsing).

Notation Z := Z (only parsing).
Notation Z_rect := Z_rect (only parsing).
Notation Z_rec := Z_rec (only parsing).
Notation Z_ind := Z_ind (only parsing).
Notation Z0 := Z0 (only parsing).
Notation Zpos := Zpos (only parsing).
Notation Zneg := Zneg (only parsing).
Compatibility lemmas. These could be notations, but scope information would be lost.
Notation SYM1 lem := (fun n => eq_sym (lem n)).
Notation SYM2 lem := (fun n m => eq_sym (lem n m)).
Notation SYM3 lem := (fun n m p => eq_sym (lem n m p)).


forall n m p : BinNums.Z, n + m + p = n + (m + p)
Proof (SYM3 Z.add_assoc).

forall n m : BinNums.Z, Z.succ (n + m) = n + Z.succ m
Proof (SYM2 Z.add_succ_r). Notation Zplus_succ_r := Zplus_succ_r_reverse (only parsing).

forall n : BinNums.Z, n = n + 0
Proof (SYM1 Z.add_0_r).

forall n m p q : BinNums.Z, n = m -> p = q -> n + p = m + q
Proof (f_equal2 Z.add).

forall n : BinNums.Z, n = Z.succ (Z.pred n)
Proof (SYM1 Z.succ_pred).

forall n : BinNums.Z, n = Z.pred (Z.succ n)
Proof (SYM1 Z.pred_succ).

forall n m : BinNums.Z, n = m -> Z.succ n = Z.succ m
Proof (f_equal Z.succ).

forall n : BinNums.Z, n = n - 0
Proof (SYM1 Z.sub_0_r).

forall n : BinNums.Z, 0 = n - n
Proof (SYM1 Z.sub_diag).

forall n m : BinNums.Z, Z.succ (n - m) = Z.succ n - m
Proof (SYM2 Z.sub_succ_l).

forall n m p : BinNums.Z, n = m + p -> p = n - m

forall n m p : BinNums.Z, n = m + p -> p = n - m
n, m, p:BinNums.Z
H:n = m + p

p = n - m
now apply Z.add_move_l. Qed.

forall n m : BinNums.Z, n + (m - n) = m
Proof (fun n m => eq_trans (Z.add_comm n (m-n)) (Z.sub_add n m)).

forall n m p : BinNums.Z, p + n - (p + m) = n - m
Proof (fun n m p => Z.add_add_simpl_l_l p n m).

forall n m p : BinNums.Z, n - m = p + n - (p + m)
Proof (SYM3 Zminus_plus_simpl_l).

forall n m p : BinNums.Z, n + p - (m + p) = n - m
Proof (fun n m p => Z.add_add_simpl_r_r n p m).

forall n m : BinNums.Z, n = m -> n - m = 0
Proof (fun n m => proj2 (Z.sub_move_0_r n m)).

forall n m : BinNums.Z, n - m = 0 -> n = m
Proof (fun n m => proj1 (Z.sub_move_0_r n m)).

forall n : BinNums.Z, 0 = n * 0
Proof (SYM1 Z.mul_0_r).

forall n m p : BinNums.Z, n * m * p = n * (m * p)
Proof (SYM3 Z.mul_assoc).

forall n m : BinNums.Z, n * m = 0 -> n = 0 \/ m = 0
Proof (fun n m => proj1 (Z.mul_eq_0 n m)).

forall n m : BinNums.Z, n <> 0 -> m * n = 0 -> m = 0
Proof (fun n m H H' => Z.mul_eq_0_l m n H' H).

forall n m : BinNums.Z, - (n * m) = - n * m
Proof (SYM2 Z.mul_opp_l).

forall n m : BinNums.Z, - (n * m) = n * - m
Proof (SYM2 Z.mul_opp_r).

forall n m p : BinNums.Z, p * (n - m) = p * n - p * m
Proof (fun n m p => Z.mul_sub_distr_l p n m).

forall n m : BinNums.Z, n * m + n = n * Z.succ m
Proof (SYM2 Z.mul_succ_r).

forall n m : BinNums.Z, n * m + m = Z.succ n * m
Proof (SYM2 Z.mul_succ_l).

forall p q : positive, p = q -> BinNums.Zpos p = BinNums.Zpos q

forall p q : positive, p = q -> BinNums.Zpos p = BinNums.Zpos q
congruence. Qed.

forall p q : positive, p = q <-> BinNums.Zpos p = BinNums.Zpos q
Proof (fun p q => iff_sym (Pos2Z.inj_iff p q)). Hint Immediate Zsucc_pred: zarith. (* Not kept : Zplus_0_simpl_l Zplus_0_simpl_l_reverse Zplus_opp_expand Zsucc_inj_contrapositive Zsucc_succ' Zpred_pred' *) (* No compat notation for : weak_assoc (now Z.add_assoc_pos) weak_Zmult_plus_distr_r (now Z.mul_add_distr_pos) *)
Obsolete stuff
Definition Zne (x y:Z) := x <> y. (* TODO : to remove someday ? *)

Register Zne as plugins.omega.Zne.

Ltac elim_compare com1 com2 :=
  case (Dcompare (com1 ?= com2)%Z);
    [ idtac | let x := fresh "H" in
      (intro x; case x; clear x) ].


2%nat = (1 + 1)%nat

2%nat = (1 + 1)%nat
reflexivity. Qed.
n:BinNums.Z

n + n = n * 2
n:BinNums.Z

n + n = n * 2
n:BinNums.Z

n + n = 2 * n
apply Z.add_diag. Qed.
n, m:BinNums.Z

m = 0 -> m * n = 0
n, m:BinNums.Z

m = 0 -> m * n = 0
intros; now subst. Qed.