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.
(************************************************************************)
(*         *   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)         *)
(************************************************************************)
This library has been deprecated since Coq version 8.10.

Int31 numbers defines indeed a cyclic structure : Z/(2^31)Z

Author: Arnaud Spiwack (+ Pierre Letouzey)
Require Import List.
Require Import Min.
Require Export Int31.
Require Import Znumtheory.
Require Import Zgcd_alt.
Require Import Zpow_facts.
Require Import CyclicAxioms.
Require Import Lia.

Local Open Scope nat_scope.
Local Open Scope int31_scope.

Local Hint Resolve Z.lt_gt Z.div_pos : zarith.

Section Basics.

Basic results about iszero, shiftl, shiftr

 

forall x : int31, iszero x = true -> x = 0

forall x : int31, iszero x = true -> x = 0
d, d0, d1, d2, d3, d4, d5, d6, d7, d8, d9, d10, d11, d12, d13, d14, d15, d16, d17, d18, d19, d20, d21, d22, d23, d24, d25, d26, d27, d28, d29:digits
H:match d with | D0 => match d0 with | D0 => match d1 with | D0 => match d2 with | D0 => match d3 with | D0 => match d4 with | D0 => match d5 with | D0 => match d6 with | D0 => match d7 with | D0 => match d8 with | D0 => match d9 with | D0 => match d10 with | D0 => match ... with | ... ... | ... false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end | D1 => false end = true

I31 d d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29 = 0
H:true = true

0 = 0
reflexivity. Qed.

forall x : int31, iszero x = false -> x <> 0

forall x : int31, iszero x = false -> x <> 0
intros x H Eq; rewrite Eq in H; simpl in *; discriminate. Qed.

forall x : int31, x = sneakl (firstr x) (shiftr x)

forall x : int31, x = sneakl (firstr x) (shiftr x)
destruct x; simpl; auto. Qed.

forall x : int31, x = sneakr (firstl x) (shiftl x)

forall x : int31, x = sneakr (firstl x) (shiftl x)
destruct x; simpl; auto. Qed.

forall x : int31, twice x = 0 <-> twice_plus_one x = 1

forall x : int31, twice x = 0 <-> twice_plus_one x = 1
destruct x; simpl in *; split; intro H; injection H; intros; subst; auto. Qed.

forall x : int31, x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)

forall x : int31, x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)
x:int31
H:firstr x = D0

x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)
x:int31
H:firstr x = D1
x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)
x:int31
H:firstr x = D1

x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)
destruct x; simpl in *; rewrite H; auto. Qed.

Iterated shift to the right

 Definition nshiftr x := nat_rect _ x (fun _ => shiftr).

 

forall (n : nat) (x : int31), nshiftr x (S n) = shiftr (nshiftr x n)

forall (n : nat) (x : int31), nshiftr x (S n) = shiftr (nshiftr x n)
reflexivity. Qed.

forall (n : nat) (x : int31), nshiftr x (S n) = nshiftr (shiftr x) n

forall (n : nat) (x : int31), nshiftr x (S n) = nshiftr (shiftr x) n
n:nat

forall n0 : nat, (forall x : int31, shiftr (nshiftr x n0) = nshiftr (shiftr x) n0) -> forall x : int31, shiftr (shiftr (nshiftr x n0)) = shiftr (nshiftr (shiftr x) n0)
intros; now f_equal. Qed.

forall n : nat, nshiftr 0 n = 0

forall n : nat, nshiftr 0 n = 0
n:nat
IHn:nshiftr 0 n = 0

shiftr (nshiftr 0 n) = 0
rewrite IHn; auto. Qed.

forall x : int31, nshiftr x size = 0

forall x : int31, nshiftr x size = 0
destruct x; simpl; auto. Qed.

forall (k : nat) (x : int31), size <= k -> nshiftr x k = 0

forall (k : nat) (x : int31), size <= k -> nshiftr x k = 0
k:nat
x:int31
H:size <= k

nshiftr x k = 0
k:nat
x:int31
H:size <= k

nshiftr x (k - size + size) = 0
k:nat
x:int31
H:size <= k

nshiftr x (0 + size) = 0
k:nat
x:int31
H:size <= k
n:nat
IHn:nshiftr x (n + size) = 0
nshiftr x (S n + size) = 0
k:nat
x:int31
H:size <= k
n:nat
IHn:nshiftr x (n + size) = 0

nshiftr x (S n + size) = 0
simpl; rewrite IHn; auto. Qed.

Iterated shift to the left

 Definition nshiftl x := nat_rect _ x (fun _ => shiftl).

 

forall (n : nat) (x : int31), nshiftl x (S n) = shiftl (nshiftl x n)

forall (n : nat) (x : int31), nshiftl x (S n) = shiftl (nshiftl x n)
reflexivity. Qed.

forall (n : nat) (x : int31), nshiftl x (S n) = nshiftl (shiftl x) n

forall (n : nat) (x : int31), nshiftl x (S n) = nshiftl (shiftl x) n
intros n; elim n; simpl; intros; now f_equal. Qed.

forall n : nat, nshiftl 0 n = 0

forall n : nat, nshiftl 0 n = 0
n:nat
IHn:nshiftl 0 n = 0

shiftl (nshiftl 0 n) = 0
rewrite IHn; auto. Qed.

forall x : int31, nshiftl x size = 0

forall x : int31, nshiftl x size = 0
destruct x; simpl; auto. Qed.

forall (k : nat) (x : int31), size <= k -> nshiftl x k = 0

forall (k : nat) (x : int31), size <= k -> nshiftl x k = 0
k:nat
x:int31
H:size <= k

nshiftl x k = 0
k:nat
x:int31
H:size <= k

nshiftl x (k - size + size) = 0
k:nat
x:int31
H:size <= k

nshiftl x (0 + size) = 0
k:nat
x:int31
H:size <= k
n:nat
IHn:nshiftl x (n + size) = 0
nshiftl x (S n + size) = 0
k:nat
x:int31
H:size <= k
n:nat
IHn:nshiftl x (n + size) = 0

nshiftl x (S n + size) = 0
simpl; rewrite IHn; auto. Qed.

forall x : int31, firstr x = firstl (nshiftl x (Init.Nat.pred size))

forall x : int31, firstr x = firstl (nshiftl x (Init.Nat.pred size))
destruct x; simpl; auto. Qed.

forall x : int31, firstl x = firstr (nshiftr x (Init.Nat.pred size))

forall x : int31, firstl x = firstr (nshiftr x (Init.Nat.pred size))
destruct x; simpl; auto. Qed.
More advanced results about nshiftr
 

forall x : int31, nshiftr x (Init.Nat.pred size) = 0 -> firstl x = D0

forall x : int31, nshiftr x (Init.Nat.pred size) = 0 -> firstl x = D0
destruct x; compute; intros H; injection H; intros; subst; auto. Qed.

forall (n p : nat) (x : int31), n <= p -> nshiftr x n = 0 -> nshiftr x p = 0

forall (n p : nat) (x : int31), n <= p -> nshiftr x n = 0 -> nshiftr x p = 0
n, p:nat
x:int31
H:n <= p
H0:nshiftr x n = 0

nshiftr x p = 0
n, p:nat
x:int31
H:n <= p
H0:nshiftr x n = 0

nshiftr x (p - n + n) = 0
n, p:nat
x:int31
H:n <= p
H0:nshiftr x n = 0

nshiftr x (0 + n) = 0
n, p:nat
x:int31
H:n <= p
H0:nshiftr x n = 0
n0:nat
IHn0:nshiftr x (n0 + n) = 0
nshiftr x (S n0 + n) = 0
n, p:nat
x:int31
H:n <= p
H0:nshiftr x n = 0
n0:nat
IHn0:nshiftr x (n0 + n) = 0

nshiftr x (S n0 + n) = 0
simpl; rewrite IHn0; auto. Qed.

forall (n : nat) (x : int31), n < size -> nshiftr x n = 0 -> firstl x = D0

forall (n : nat) (x : int31), n < size -> nshiftr x n = 0 -> firstl x = D0
n:nat
x:int31
H:n < size
H0:nshiftr x n = 0

firstl x = D0
n:nat
x:int31
H:n < size
H0:nshiftr x n = 0

nshiftr x (Init.Nat.pred size) = 0
apply nshiftr_0_propagates with n; auto; omega. Qed.

Some induction principles over int31

 
Not used for the moment. Are they really useful ?
 

forall P : int31 -> Prop, P 0 -> (forall (x : int31) (d : digits), P x -> P (sneakl d x)) -> forall x : int31, P x

forall P : int31 -> Prop, P 0 -> (forall (x : int31) (d : digits), P x -> P (sneakl d x)) -> forall x : int31, P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31

P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31

forall n : nat, n <= size -> P (nshiftr x (size - n))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))
P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:0 <= size

P (nshiftr x (size - 0))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
n:nat
IHn:n <= size -> P (nshiftr x (size - n))
H1:S n <= size
P (nshiftr x (size - S n))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))
P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
n:nat
IHn:n <= size -> P (nshiftr x (size - n))
H1:S n <= size

P (nshiftr x (size - S n))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))
P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
n:nat
IHn:n <= size -> P (nshiftr x (size - n))
H1:S n <= size

P (sneakl (firstr (nshiftr x (size - S n))) (shiftr (nshiftr x (size - S n))))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))
P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
n:nat
IHn:n <= size -> P (nshiftr x (size - n))
H1:S n <= size

P (shiftr (nshiftr x (size - S n)))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))
P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
n:nat
IHn:n <= size -> P (nshiftr x (size - n))
H1:S n <= size

P (nshiftr x (S (size - S n)))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))
P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
n:nat
IHn:n <= size -> P (nshiftr x (size - n))
H1:S n <= size

P (nshiftr x (size - n))
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))
P x
P:int31 -> Prop
H:P 0
H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)
x:int31
H1:forall n : nat, n <= size -> P (nshiftr x (size - n))

P x
change x with (nshiftr x (size-size)); auto. Qed.

forall P : int31 -> Prop, P 0 -> (forall x : int31, P x -> P (twice x)) -> (forall x : int31, P x -> P (twice_plus_one x)) -> forall x : int31, P x

forall P : int31 -> Prop, P 0 -> (forall x : int31, P x -> P (twice x)) -> (forall x : int31, P x -> P (twice_plus_one x)) -> forall x : int31, P x
P:int31 -> Prop
H:P 0
H0:forall x0 : int31, P x0 -> P (twice x0)
H1:forall x0 : int31, P x0 -> P (twice_plus_one x0)
x:int31
d:digits
IHx:P x

P (sneakl d x)
destruct d; auto. Qed.

Some generic results about recr

 Section Recr.
recr satisfies the fixpoint equation used for its definition.
 Variable (A:Type)(case0:A)(caserec:digits->int31->A->A).

 
A:Type
case0:A
caserec:digits -> int31 -> A -> A

forall (n : nat) (x : int31), iszero x = false -> recr_aux (S n) A case0 caserec x = caserec (firstr x) (shiftr x) (recr_aux n A case0 caserec (shiftr x))
A:Type
case0:A
caserec:digits -> int31 -> A -> A

forall (n : nat) (x : int31), iszero x = false -> recr_aux (S n) A case0 caserec x = caserec (firstr x) (shiftr x) (recr_aux n A case0 caserec (shiftr x))
intros; simpl; rewrite H; auto. Qed.
A:Type
case0:A
caserec:digits -> int31 -> A -> A

forall (n p : nat) (x : int31), n <= size -> n <= p -> recr_aux n A case0 caserec (nshiftr x (size - n)) = recr_aux p A case0 caserec (nshiftr x (size - n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A

forall (n p : nat) (x : int31), n <= size -> n <= p -> recr_aux n A case0 caserec (nshiftr x (size - n)) = recr_aux p A case0 caserec (nshiftr x (size - n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A

forall (p : nat) (x : int31), 0 <= size -> 0 <= p -> recr_aux 0 A case0 caserec (nshiftr x (size - 0)) = recr_aux p A case0 caserec (nshiftr x (size - 0))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p : nat) (x : int31), n <= size -> n <= p -> recr_aux n A case0 caserec (nshiftr x (size - n)) = recr_aux p A case0 caserec (nshiftr x (size - n))
forall (p : nat) (x : int31), S n <= size -> S n <= p -> recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux p A case0 caserec (nshiftr x (size - S n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
p:nat
x:int31
H:0 <= size
H0:0 <= p

recr_aux 0 A case0 caserec (nshiftr x 31) = recr_aux p A case0 caserec (nshiftr x 31)
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p : nat) (x : int31), n <= size -> n <= p -> recr_aux n A case0 caserec (nshiftr x (size - n)) = recr_aux p A case0 caserec (nshiftr x (size - n))
forall (p : nat) (x : int31), S n <= size -> S n <= p -> recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux p A case0 caserec (nshiftr x (size - S n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p : nat) (x : int31), n <= size -> n <= p -> recr_aux n A case0 caserec (nshiftr x (size - n)) = recr_aux p A case0 caserec (nshiftr x (size - n))

forall (p : nat) (x : int31), S n <= size -> S n <= p -> recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux p A case0 caserec (nshiftr x (size - S n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= p

recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux p A case0 caserec (nshiftr x (size - S n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p : nat) (x0 : int31), n <= size -> n <= p -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p A case0 caserec (nshiftr x0 (size - n))
x:int31
H:S n <= size
H0:S n <= 0

recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux 0 A case0 caserec (nshiftr x (size - S n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= S p
recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux (S p) A case0 caserec (nshiftr x (size - S n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= S p

recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux (S p) A case0 caserec (nshiftr x (size - S n))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= S p

(if iszero (nshiftr x (size - S n)) then case0 else caserec (firstr (nshiftr x (size - S n))) (shiftr (nshiftr x (size - S n))) (recr_aux n A case0 caserec (shiftr (nshiftr x (size - S n))))) = (if iszero (nshiftr x (size - S n)) then case0 else caserec (firstr (nshiftr x (size - S n))) (shiftr (nshiftr x (size - S n))) (recr_aux p A case0 caserec (shiftr (nshiftr x (size - S n)))))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= S p

caserec (firstr (nshiftr x (size - S n))) (shiftr (nshiftr x (size - S n))) (recr_aux n A case0 caserec (shiftr (nshiftr x (size - S n)))) = caserec (firstr (nshiftr x (size - S n))) (shiftr (nshiftr x (size - S n))) (recr_aux p A case0 caserec (shiftr (nshiftr x (size - S n))))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= S p

recr_aux n A case0 caserec (shiftr (nshiftr x (size - S n))) = recr_aux p A case0 caserec (shiftr (nshiftr x (size - S n)))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= S p

recr_aux n A case0 caserec (nshiftr x (S (size - S n))) = recr_aux p A case0 caserec (nshiftr x (S (size - S n)))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
n:nat
IHn:forall (p0 : nat) (x0 : int31), n <= size -> n <= p0 -> recr_aux n A case0 caserec (nshiftr x0 (size - n)) = recr_aux p0 A case0 caserec (nshiftr x0 (size - n))
p:nat
x:int31
H:S n <= size
H0:S n <= S p

recr_aux n A case0 caserec (nshiftr x (size - n)) = recr_aux p A case0 caserec (nshiftr x (size - n))
apply IHn; auto with arith. Qed.
A:Type
case0:A
caserec:digits -> int31 -> A -> A

forall x : int31, iszero x = false -> recr A case0 caserec x = caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x))
A:Type
case0:A
caserec:digits -> int31 -> A -> A

forall x : int31, iszero x = false -> recr A case0 caserec x = caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
x:int31
H:iszero x = false

recr A case0 caserec x = caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
x:int31
H:iszero x = false

recr_aux size A case0 caserec x = caserec (firstr x) (shiftr x) (recr_aux size A case0 caserec (shiftr x))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
x:int31
H:iszero x = false

recr_aux size A case0 caserec (nshiftr x (size - size)) = caserec (firstr (nshiftr x (size - size))) (shiftr (nshiftr x (size - size))) (recr_aux size A case0 caserec (shiftr (nshiftr x (size - size))))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
x:int31
H:iszero x = false

recr_aux (S size) A case0 caserec (nshiftr x (size - size)) = caserec (firstr (nshiftr x (size - size))) (shiftr (nshiftr x (size - size))) (recr_aux size A case0 caserec (shiftr (nshiftr x (size - size))))
rewrite recr_aux_eqn; auto. Qed.
recr is usually equivalent to a variant recrbis written without iszero check.
 Fixpoint recrbis_aux (n:nat)(A:Type)(case0:A)(caserec:digits->int31->A->A)
 (i:int31) : A :=
  match n with
  | O => case0
  | S next =>
      let si := shiftr i in
      caserec (firstr i) si (recrbis_aux next A case0 caserec si)
  end.

 Definition recrbis := recrbis_aux size.

 Hypothesis case0_caserec : caserec D0 0 case0 = case0.

 
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0

forall (n : nat) (x : int31), recrbis_aux n A case0 caserec x = recr_aux n A case0 caserec x
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0

forall (n : nat) (x : int31), recrbis_aux n A case0 caserec x = recr_aux n A case0 caserec x
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0
n:nat
IHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0
x:int31

caserec (firstr x) (shiftr x) (recrbis_aux n A case0 caserec (shiftr x)) = (if iszero x then case0 else caserec (firstr x) (shiftr x) (recr_aux n A case0 caserec (shiftr x)))
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0
n:nat
IHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0
x:int31
H:iszero x = true

caserec (firstr x) (shiftr x) (recrbis_aux n A case0 caserec (shiftr x)) = case0
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0
n:nat
IHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0
x:int31
H:iszero x = true

caserec D0 0 (recrbis_aux n A case0 caserec 0) = case0
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0
n:nat
IHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0
x:int31
H:iszero x = true

case0 = recrbis_aux n A case0 caserec 0
clear H IHn; induction n; simpl; congruence. Qed.
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0

forall x : int31, recrbis A case0 caserec x = recr A case0 caserec x
A:Type
case0:A
caserec:digits -> int31 -> A -> A
case0_caserec:caserec D0 0 case0 = case0

forall x : int31, recrbis A case0 caserec x = recr A case0 caserec x
intros; apply recrbis_aux_equiv; auto. Qed. End Recr.

Incrementation

 Section Incr.
Variant of incr via recrbis
 Let Incr (b : digits) (si rec : int31) :=
  match b with
   | D0 => sneakl D1 si
   | D1 => sneakl D0 rec
  end.

 Definition incrbis_aux n x := recrbis_aux n _ In Incr x.

 
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, incrbis_aux size x = incr x
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, incrbis_aux size x = incr x
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31

recrbis_aux size int31 In Incr x = recr_aux size int31 In Incr x
apply recrbis_aux_equiv; auto. Qed.
Recursive equations satisfied by incr
 
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, firstr x = D0 -> incr x = twice_plus_one (shiftr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, firstr x = D0 -> incr x = twice_plus_one (shiftr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D0

incr x = twice_plus_one (shiftr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D0
H0:iszero x = true

incr x = twice_plus_one (shiftr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D0
H0:iszero x = false
incr x = twice_plus_one (shiftr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D0
H0:iszero x = false

incr x = twice_plus_one (shiftr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D0
H0:iszero x = false

match firstr x with | D0 => sneakl D1 (shiftr x) | D1 => sneakl D0 (incr (shiftr x)) end = twice_plus_one (shiftr x)
rewrite H; auto. Qed.
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, firstr x = D1 -> incr x = twice (incr (shiftr x))
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, firstr x = D1 -> incr x = twice (incr (shiftr x))
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D1

incr x = twice (incr (shiftr x))
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D1
H0:iszero x = true

incr x = twice (incr (shiftr x))
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D1
H0:iszero x = false
incr x = twice (incr (shiftr x))
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D1
H0:iszero x = false

incr x = twice (incr (shiftr x))
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstr x = D1
H0:iszero x = false

match firstr x with | D0 => sneakl D1 (shiftr x) | D1 => sneakl D0 (incr (shiftr x)) end = twice (incr (shiftr x))
rewrite H; auto. Qed.
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, incr (twice x) = twice_plus_one x
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, incr (twice x) = twice_plus_one x
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31

incr (twice x) = twice_plus_one x
rewrite incr_eqn1; destruct x; simpl; auto. Qed.
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, firstl x = D0 -> incr (twice_plus_one x) = twice (incr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31

forall x : int31, firstl x = D0 -> incr (twice_plus_one x) = twice (incr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstl x = D0

incr (twice_plus_one x) = twice (incr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstl x = D0

twice (incr (shiftr (twice_plus_one x))) = twice (incr x)
Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31
x:int31
H:firstl x = D0

shiftr (twice_plus_one x) = x
destruct x; simpl in *; rewrite H; auto. Qed.
The previous result is actually true even without the constraint on firstl, but this is harder to prove (see later).
 End Incr.

Conversion to Z : the phi function

 Section Phi.
Variant of phi via recrbis
 Let Phi := fun b (_:int31) =>
       match b with D0 => Z.double | D1 => Z.succ_double end.

 Definition phibis_aux n x := recrbis_aux n _ Z0 Phi x.

 
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, phibis_aux size x = phi x
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, phibis_aux size x = phi x
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31

recrbis_aux size Z 0%Z Phi x = recr_aux size Z 0%Z Phi x
apply recrbis_aux_equiv; auto. Qed.
Recursive equations satisfied by phi
 
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstr x = D0 -> phi x = Z.double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstr x = D0 -> phi x = Z.double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D0

phi x = Z.double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D0
H0:iszero x = true

phi x = Z.double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D0
H0:iszero x = false
phi x = Z.double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D0
H0:iszero x = false

phi x = Z.double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D0
H0:iszero x = false

match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phi (shiftr x)) = Z.double (phi (shiftr x))
rewrite H; auto. Qed.
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstr x = D1 -> phi x = Z.succ_double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstr x = D1 -> phi x = Z.succ_double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D1

phi x = Z.succ_double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D1
H0:iszero x = true

phi x = Z.succ_double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D1
H0:iszero x = false
phi x = Z.succ_double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D1
H0:iszero x = false

phi x = Z.succ_double (phi (shiftr x))
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstr x = D1
H0:iszero x = false

match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phi (shiftr x)) = Z.succ_double (phi (shiftr x))
rewrite H; auto. Qed.
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstl x = D0 -> phi (twice x) = Z.double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstl x = D0 -> phi (twice x) = Z.double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstl x = D0

phi (twice x) = Z.double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstl x = D0

Z.double (phi (shiftr (twice x))) = Z.double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstl x = D0

shiftr (twice x) = x
destruct x; simpl in *; rewrite H; auto. Qed.
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstl x = D0 -> phi (twice_plus_one x) = Z.succ_double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z

forall x : int31, firstl x = D0 -> phi (twice_plus_one x) = Z.succ_double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstl x = D0

phi (twice_plus_one x) = Z.succ_double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstl x = D0

Z.succ_double (phi (shiftr (twice_plus_one x))) = Z.succ_double (phi x)
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Z
x:int31
H:firstl x = D0

shiftr (twice_plus_one x) = x
destruct x; simpl in *; rewrite H; auto. Qed. End Phi.
phi x is positive and lower than 2^31
 

forall (n : nat) (x : int31), (0 <= phibis_aux n x)%Z

forall (n : nat) (x : int31), (0 <= phibis_aux n x)%Z

forall x : int31, (0 <= phibis_aux 0 x)%Z
n:nat
IHn:forall x : int31, (0 <= phibis_aux n x)%Z
forall x : int31, (0 <= phibis_aux (S n) x)%Z
n:nat
IHn:forall x : int31, (0 <= phibis_aux n x)%Z

forall x : int31, (0 <= phibis_aux (S n) x)%Z
n:nat
IHn:forall x0 : int31, (0 <= phibis_aux n x0)%Z
x:int31

(0 <= phibis_aux (S n) x)%Z
n:nat
IHn:forall x0 : int31, (0 <= phibis_aux n x0)%Z
x:int31

(0 <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, (0 <= phibis_aux n x0)%Z
x:int31

(0 <= Z.double (phibis_aux n (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, (0 <= phibis_aux n x0)%Z
x:int31
(0 <= Z.succ_double (phibis_aux n (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, (0 <= phibis_aux n x0)%Z
x:int31

(0 <= Z.succ_double (phibis_aux n (shiftr x)))%Z
specialize IHn with (shiftr x); rewrite Z.succ_double_spec; omega. Qed.

forall (n : nat) (x : int31), n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Z

forall (n : nat) (x : int31), n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Z

forall x : int31, 0 <= size -> (phibis_aux 0 (nshiftr x (size - 0)) < 2 ^ Z.of_nat 0)%Z
n:nat
IHn:forall x : int31, n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Z
forall x : int31, S n <= size -> (phibis_aux (S n) (nshiftr x (size - S n)) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x : int31, n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Z

forall x : int31, S n <= size -> (phibis_aux (S n) (nshiftr x (size - S n)) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size

(phibis_aux (S n) (nshiftr x (size - S n)) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size

(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr (nshiftr x (size - S n)))) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size

shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr (nshiftr x (size - S n)))) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size

shiftr (nshiftr x (size - S n)) = nshiftr x (S (size - S n))
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr (nshiftr x (size - S n)))) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)

(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr (nshiftr x (size - S n)))) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)

(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (nshiftr x (size - n))) < 2 ^ Z.of_nat (S n))%Z
n:nat
IHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Z
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:n <= size

(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (nshiftr x (size - n))) < 2 ^ Z.of_nat (S n))%Z
n:nat
x:int31
IHn:(phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Z
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:n <= size

(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (nshiftr x (size - n))) < 2 ^ Z.of_nat (S n))%Z
n:nat
x:int31
y:=phibis_aux n (nshiftr x (size - n)):Z
IHn:(y < 2 ^ Z.of_nat n)%Z
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:n <= size

(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end y < 2 ^ Z.of_nat (S n))%Z
n:nat
x:int31
y:=phibis_aux n (nshiftr x (size - n)):Z
IHn:(y < 2 ^ Z.of_nat n)%Z
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:n <= size

(match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end y < 2 * 2 ^ Z.of_nat n)%Z
n:nat
x:int31
y:=phibis_aux n (nshiftr x (size - n)):Z
IHn:(y < 2 ^ Z.of_nat n)%Z
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:n <= size
H2:firstr (nshiftr x (size - S n)) = D0

(Z.double y < 2 * 2 ^ Z.of_nat n)%Z
n:nat
x:int31
y:=phibis_aux n (nshiftr x (size - n)):Z
IHn:(y < 2 ^ Z.of_nat n)%Z
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:n <= size
H2:firstr (nshiftr x (size - S n)) = D1
(Z.succ_double y < 2 * 2 ^ Z.of_nat n)%Z
n:nat
x:int31
y:=phibis_aux n (nshiftr x (size - n)):Z
IHn:(y < 2 ^ Z.of_nat n)%Z
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:n <= size
H2:firstr (nshiftr x (size - S n)) = D1

(Z.succ_double y < 2 * 2 ^ Z.of_nat n)%Z
rewrite Z.succ_double_spec; auto with zarith. Qed.

forall x : int31, (0 <= phi x)%Z

forall x : int31, (0 <= phi x)%Z
x:int31

(0 <= phi x)%Z
x:int31

(0 <= phibis_aux size x)%Z
apply phibis_aux_pos. Qed. Hint Resolve phi_nonneg : zarith.

forall x : int31, (0 <= phi x < 2 ^ Z.of_nat size)%Z

forall x : int31, (0 <= phi x < 2 ^ Z.of_nat size)%Z
x:int31

(0 <= phi x < 2 ^ Z.of_nat size)%Z
x:int31

(phi x < 2 ^ Z.of_nat size)%Z
x:int31

(phibis_aux size x < 2 ^ Z.of_nat size)%Z
x:int31

(phibis_aux size (nshiftr x (size - size)) < 2 ^ Z.of_nat size)%Z
apply phibis_aux_bounded; auto. Qed.

forall (n : nat) (x : int31), firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z

forall (n : nat) (x : int31), firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z

forall x : int31, firstr (nshiftr x 0) = D1 -> (2 ^ Z.of_nat 0 <= phibis_aux 1 x)%Z
n:nat
IHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z
forall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Z
x:int31
H:firstr (nshiftr x 0) = D1

(2 ^ Z.of_nat 0 <= phibis_aux 1 x)%Z
n:nat
IHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z
forall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Z
x:int31
H:firstr x = D1

(1 <= phibis_aux 1 x)%Z
n:nat
IHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z
forall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Z
x:int31
H:firstr x = D1

(1 <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end 0)%Z
n:nat
IHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z
forall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Z
n:nat
IHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Z

forall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1

(2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Z
n, m:nat
Heqm:m = S n
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux m x0)%Z
x:int31
H:firstr (nshiftr x m) = D1

(2 ^ Z.of_nat m <= phibis_aux (S m) x)%Z
n, m:nat
Heqm:m = S n
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux m x0)%Z
x:int31
H:firstr (nshiftr x m) = D1

(2 ^ Z.of_nat m <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux m (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1

(2 ^ Z.of_nat (S n) <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1

(2 * 2 ^ Z.of_nat n <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1

(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z
(2 * 2 ^ Z.of_nat n <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1

firstr (nshiftr (shiftr x) n) = D1
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z
(2 * 2 ^ Z.of_nat n <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z

(2 * 2 ^ Z.of_nat n <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z

(2 * 2 ^ Z.of_nat n <= Z.double (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z
(2 * 2 ^ Z.of_nat n <= Z.succ_double (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z

(2 * 2 ^ Z.of_nat n <= 2 * phibis_aux (S n) (shiftr x))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z
(2 * 2 ^ Z.of_nat n <= Z.succ_double (phibis_aux (S n) (shiftr x)))%Z
n:nat
IHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Z
x:int31
H:firstr (nshiftr x (S n)) = D1
H0:(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Z

(2 * 2 ^ Z.of_nat n <= Z.succ_double (phibis_aux (S n) (shiftr x)))%Z
rewrite Z.succ_double_spec; omega. Qed.

forall x : int31, firstl x = D1 -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Z

forall x : int31, firstl x = D1 -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Z
x:int31
H:firstl x = D1

(2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Z
x:int31
H:firstl x = D1

(firstr (nshiftr x (Init.Nat.pred size)) = D1 -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phibis_aux (S (Init.Nat.pred size)) x)%Z) -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Z
x:int31
H:firstl x = D1

(firstl x = D1 -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phibis_aux (S (Init.Nat.pred size)) x)%Z) -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Z
x:int31
H:firstl x = D1

(firstl x = D1 -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phibis_aux size x)%Z) -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Z
rewrite phibis_aux_equiv; auto. Qed.

Equivalence modulo 2^n

 Section EqShiftL.
After killing n bits at the left, are the numbers equal ?
 Definition EqShiftL n x y :=
  nshiftl x n = nshiftl y n.

 

forall x y : int31, EqShiftL 0 x y <-> x = y

forall x y : int31, EqShiftL 0 x y <-> x = y
unfold EqShiftL; intros; unfold nshiftl; simpl; split; auto. Qed.

forall (k : nat) (x y : int31), size <= k -> EqShiftL k x y

forall (k : nat) (x y : int31), size <= k -> EqShiftL k x y
red; intros; rewrite 2 nshiftl_above_size; auto. Qed.

forall (k k' : nat) (x y : int31), k <= k' -> EqShiftL k x y -> EqShiftL k' x y

forall (k k' : nat) (x y : int31), k <= k' -> EqShiftL k x y -> EqShiftL k' x y
k, k':nat
x, y:int31
H:k <= k'
H0:nshiftl x k = nshiftl y k

nshiftl x k' = nshiftl y k'
k, k':nat
x, y:int31
H:k <= k'
H0:nshiftl x k = nshiftl y k

nshiftl x (k' - k + k) = nshiftl y (k' - k + k)
k, k':nat
x, y:int31
H:k <= k'
H0:nshiftl x k = nshiftl y k
n:nat
Heqn:n = (k' - k)%nat

nshiftl x (n + k) = nshiftl y (n + k)
k:nat
x, y:int31
H0:nshiftl x k = nshiftl y k
n:nat

nshiftl x (n + k) = nshiftl y (n + k)
k:nat
x, y:int31
H0:nshiftl x k = nshiftl y k
n:nat
IHn:nshiftl x (n + k) = nshiftl y (n + k)

shiftl (nshiftl x (n + k)) = shiftl (nshiftl y (n + k))
f_equal; auto. Qed.

forall (k : nat) (x y : int31), k < size -> EqShiftL k x y -> firstr x = firstr y

forall (k : nat) (x y : int31), k < size -> EqShiftL k x y -> firstr x = firstr y
k:nat
x, y:int31
H:k < size
H0:EqShiftL k x y

firstr x = firstr y
k:nat
x, y:int31
H:k < size
H0:EqShiftL k x y

firstl (nshiftl x (Init.Nat.pred size)) = firstl (nshiftl y (Init.Nat.pred size))
k:nat
x, y:int31
H:k < size
H0:EqShiftL k x y

nshiftl x (Init.Nat.pred size) = nshiftl y (Init.Nat.pred size)
k:nat
x, y:int31
H:k < size
H0:EqShiftL k x y

k <= Init.Nat.pred size
k:nat
x, y:int31
H:k < size
H0:EqShiftL k x y

k <= Init.Nat.pred 31
auto with arith. Qed.

forall (k : nat) (x y : int31), EqShiftL k (twice x) (twice y) <-> EqShiftL (S k) x y

forall (k : nat) (x y : int31), EqShiftL k (twice x) (twice y) <-> EqShiftL (S k) x y
k:nat
x, y:int31

nshiftl (twice x) k = nshiftl (twice y) k <-> nshiftl x (S k) = nshiftl y (S k)
rewrite 2 nshiftl_S_tail; split; auto. Qed.

From int31 to list of digits.

 
Lower (=rightmost) bits comes first.
 Definition i2l := recrbis _ nil (fun d _ rec => d::rec).

 

forall x : int31, length (i2l x) = size

forall x : int31, length (i2l x) = size
intros; reflexivity. Qed. Fixpoint lshiftl l x := match l with | nil => x | d::l => sneakl d (lshiftl l x) end. Definition l2i l := lshiftl l On.

forall x : int31, l2i (i2l x) = x

forall x : int31, l2i (i2l x) = x
destruct x; compute; auto. Qed.

forall (x : int31) (d : digits), i2l (sneakr d x) = tl (i2l x) ++ d :: nil

forall (x : int31) (d : digits), i2l (sneakr d x) = tl (i2l x) ++ d :: nil
destruct x; compute; auto. Qed.

forall (x : int31) (d : digits), i2l (sneakl d x) = d :: removelast (i2l x)

forall (x : int31) (d : digits), i2l (sneakl d x) = d :: removelast (i2l x)
destruct x; compute; auto. Qed.

forall l : list digits, length l = size -> i2l (l2i l) = l

forall l : list digits, length l = size -> i2l (l2i l) = l
d, d0, d1, d2, d3, d4, d5, d6, d7, d8, d9, d10, d11, d12, d13, d14, d15, d16, d17, d18, d19, d20, d21, d22, d23, d24, d25, d26, d27, d28, d29:digits
l:list digits

length (d :: d0 :: d1 :: d2 :: d3 :: d4 :: d5 :: d6 :: d7 :: d8 :: d9 :: d10 :: d11 :: d12 :: d13 :: d14 :: d15 :: d16 :: d17 :: d18 :: d19 :: d20 :: d21 :: ...) = size -> i2l (l2i (d :: d0 :: d1 :: d2 :: d3 :: d4 :: d5 :: d6 :: d7 :: d8 :: d9 :: d10 :: d11 :: d12 :: d13 :: d14 :: d15 :: d16 :: d17 :: d18 :: d19 :: ...)) = d :: d0 :: d1 :: d2 :: d3 :: d4 :: d5 :: d6 :: d7 :: d8 :: d9 :: d10 :: d11 :: d12 :: d13 :: d14 :: d15 :: d16 :: d17 :: d18 :: d19 :: d20 :: d21 :: d22 :: d23 :: ...
d, d0, d1, d2, d3, d4, d5, d6, d7, d8, d9, d10, d11, d12, d13, d14, d15, d16, d17, d18, d19, d20, d21, d22, d23, d24, d25, d26, d27, d28, d29:digits

length (d :: d0 :: d1 :: d2 :: d3 :: d4 :: d5 :: d6 :: d7 :: d8 :: d9 :: d10 :: d11 :: d12 :: d13 :: d14 :: d15 :: d16 :: d17 :: d18 :: d19 :: d20 :: d21 :: ...) = size -> i2l (l2i (d :: d0 :: d1 :: d2 :: d3 :: d4 :: d5 :: d6 :: d7 :: d8 :: d9 :: d10 :: d11 :: d12 :: d13 :: d14 :: d15 :: d16 :: d17 :: d18 :: d19 :: ...)) = d :: d0 :: d1 :: d2 :: d3 :: d4 :: d5 :: d6 :: d7 :: d8 :: d9 :: d10 :: d11 :: d12 :: d13 :: d14 :: d15 :: d16 :: d17 :: d18 :: d19 :: d20 :: d21 :: d22 :: d23 :: ...
intros _; compute; auto. Qed. Fixpoint cstlist (A:Type)(a:A) n := match n with | O => nil | S n => a::cstlist _ a n end.

forall (n : nat) (x : int31), n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)

forall (n : nat) (x : int31), n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)

forall x : int31, 0 <= size -> i2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)
n:nat
IHn:forall x : int31, n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)
forall x : int31, S n <= size -> i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
x:int31
H:0 <= size

i2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)
n:nat
IHn:forall x : int31, n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)
forall x : int31, S n <= size -> i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
x:int31
H:0 <= size

firstn (size - 0) (i2l x) = i2l x
x:int31
H:0 <= size
H0:firstn (size - 0) (i2l x) = i2l x
i2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)
n:nat
IHn:forall x : int31, n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)
forall x : int31, S n <= size -> i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
x:int31
H:0 <= size

firstn (length (i2l x)) (i2l x) = i2l x
x:int31
H:0 <= size
H0:firstn (size - 0) (i2l x) = i2l x
i2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)
n:nat
IHn:forall x : int31, n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)
forall x : int31, S n <= size -> i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
x:int31
H:0 <= size
H0:firstn (size - 0) (i2l x) = i2l x

i2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)
n:nat
IHn:forall x : int31, n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)
forall x : int31, S n <= size -> i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
x:int31
H:0 <= size

i2l (nshiftl x 0) = cstlist digits D0 0 ++ i2l x
n:nat
IHn:forall x : int31, n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)
forall x : int31, S n <= size -> i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x : int31, n <= size -> i2l (nshiftl x n) = cstlist digits D0 n ++ firstn (size - n) (i2l x)

forall x : int31, S n <= size -> i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

i2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

i2l (shiftl (nshiftl x n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

D0 :: removelast (i2l (nshiftl x n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

D0 :: removelast (i2l (nshiftl x n)) = (D0 :: cstlist digits D0 n) ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

removelast (i2l (nshiftl x n)) = cstlist digits D0 n ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

removelast (cstlist digits D0 n ++ firstn (size - n) (i2l x)) = cstlist digits D0 n ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

cstlist digits D0 n ++ removelast (firstn (size - n) (i2l x)) = cstlist digits D0 n ++ firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
firstn (size - n) (i2l x) <> nil
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

removelast (firstn (size - n) (i2l x)) = firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
firstn (size - n) (i2l x) <> nil
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

removelast (firstn (S (size - S n)) (i2l x)) = firstn (size - S n) (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
firstn (size - n) (i2l x) <> nil
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

size - S n < length (i2l x)
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
firstn (size - n) (i2l x) <> nil
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

firstn (size - n) (i2l x) <> nil
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

length (firstn (size - n) (i2l x)) = Init.Nat.min (size - n) (length (i2l x)) -> firstn (size - n) (i2l x) <> nil
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size

length (firstn (size - n) (i2l x)) = Init.Nat.min (size - n) size -> firstn (size - n) (i2l x) <> nil
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
H0:length (firstn (size - n) (i2l x)) = Init.Nat.min (size - n) size
H1:firstn (size - n) (i2l x) = nil

False
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
H0:length nil = Init.Nat.min (size - n) size
H1:firstn (size - n) (i2l x) = nil

False
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
H0:length nil = (size - n)%nat
H1:firstn (size - n) (i2l x) = nil

False
n:nat
IHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)
x:int31
H:S n <= size
H0:0%nat = (size - n)%nat
H1:firstn (size - n) (i2l x) = nil

False
omega. Qed.
i2l can be used to define a relation equivalent to EqShiftL
 

forall (k : nat) (x y : int31), EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)

forall (k : nat) (x y : int31), EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31

EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hle:size <= k

EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hle:size <= k
H:EqShiftL k x y

firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hle:size <= k
H:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
EqShiftL k x y
k:nat
x, y:int31
Hlt:k < size
EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hle:size <= k
H:EqShiftL k x y

firstn 0 (i2l x) = firstn 0 (i2l y)
k:nat
x, y:int31
Hle:size <= k
H:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
EqShiftL k x y
k:nat
x, y:int31
Hlt:k < size
EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hle:size <= k
H:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)

EqShiftL k x y
k:nat
x, y:int31
Hlt:k < size
EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size

EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size

nshiftl x k = nshiftl y k <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
H:k <= size

nshiftl x k = nshiftl y k <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:nshiftl x k = nshiftl y k

firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
nshiftl x k = nshiftl y k
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:nshiftl x k = nshiftl y k
H1:i2l (nshiftl x k) = i2l (nshiftl y k)

firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
nshiftl x k = nshiftl y k
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:nshiftl x k = nshiftl y k
H1:cstlist digits D0 k ++ firstn (size - k) (i2l x) = cstlist digits D0 k ++ firstn (size - k) (i2l y)

firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
nshiftl x k = nshiftl y k
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)

nshiftl x k = nshiftl y k
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)

i2l (nshiftl x k) = i2l (nshiftl y k)
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
H1:i2l (nshiftl x k) = i2l (nshiftl y k)
nshiftl x k = nshiftl y k
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)

cstlist digits D0 k ++ firstn (size - k) (i2l x) = cstlist digits D0 k ++ firstn (size - k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
H1:i2l (nshiftl x k) = i2l (nshiftl y k)
nshiftl x k = nshiftl y k
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
H1:i2l (nshiftl x k) = i2l (nshiftl y k)

nshiftl x k = nshiftl y k
k:nat
x, y:int31
Hlt:k < size
H:k <= size
H0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)
H1:i2l (nshiftl x k) = i2l (nshiftl y k)

l2i (i2l (nshiftl x k)) = l2i (i2l (nshiftl y k))
f_equal; auto. Qed.
This equivalence allows proving easily the following delicate result
 

forall (k : nat) (x y : int31), EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y

forall (k : nat) (x y : int31), EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y
k:nat
x, y:int31

EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y
k:nat
x, y:int31
Hle:size <= k

EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y
k:nat
x, y:int31
Hlt:k < size
EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y
k:nat
x, y:int31
Hlt:k < size

EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x y
k:nat
x, y:int31
Hlt:k < size

firstn (size - k) (i2l (twice_plus_one x)) = firstn (size - k) (i2l (twice_plus_one y)) <-> firstn (size - S k) (i2l x) = firstn (size - S k) (i2l y)
k:nat
x, y:int31
Hlt:k < size

firstn (size - k) (i2l (sneakl D1 x)) = firstn (size - k) (i2l (sneakl D1 y)) <-> firstn (size - S k) (i2l x) = firstn (size - S k) (i2l y)
k:nat
x, y:int31
Hlt:k < size

firstn (size - k) (D1 :: removelast (i2l x)) = firstn (size - k) (D1 :: removelast (i2l y)) <-> firstn (size - S k) (i2l x) = firstn (size - S k) (i2l y)
k:nat
x, y:int31
Hlt:k < size

firstn (S (size - S k)) (D1 :: removelast (i2l x)) = firstn (S (size - S k)) (D1 :: removelast (i2l y)) <-> firstn (size - S k) (i2l x) = firstn (size - S k) (i2l y)
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat

firstn (S n) (D1 :: removelast (i2l x)) = firstn (S n) (D1 :: removelast (i2l y)) <-> firstn n (i2l x) = firstn n (i2l y)
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x

firstn (S n) (D1 :: removelast lx) = firstn (S n) (D1 :: removelast (i2l y)) <-> firstn n lx = firstn n (i2l y)
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y

firstn (S n) (D1 :: removelast lx) = firstn (S n) (D1 :: removelast ly) <-> firstn n lx = firstn n ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y

D1 :: firstn n (removelast lx) = D1 :: firstn n (removelast ly) <-> firstn n lx = firstn n ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y

D1 :: firstn n lx = D1 :: firstn n ly <-> firstn n lx = firstn n ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
n < length ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
n < length lx
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
H:D1 :: firstn n lx = D1 :: firstn n ly

firstn n lx = firstn n ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
H:firstn n lx = firstn n ly
D1 :: firstn n lx = D1 :: firstn n ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
n < length ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
n < length lx
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
H:firstn n lx = firstn n ly

D1 :: firstn n lx = D1 :: firstn n ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
n < length ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
n < length lx
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y

n < length ly
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y
n < length lx
k:nat
x, y:int31
Hlt:k < size
n:nat
Heqn:n = (size - S k)%nat
lx:list digits
Heqlx:lx = i2l x
ly:list digits
Heqly:ly = i2l y

n < length lx
subst lx n; rewrite i2l_length; omega. Qed.

forall (k : nat) (x y : int31), EqShiftL k x y -> EqShiftL (S k) (shiftr x) (shiftr y)

forall (k : nat) (x y : int31), EqShiftL k x y -> EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y

EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hle:size <= S k

EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size

EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D0

EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1
EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D0

EqShiftL k (twice (shiftr x)) (twice (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1
EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D0

EqShiftL k (sneakl (firstr x) (shiftr x)) (sneakl (firstr x) (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1
EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D0

EqShiftL k x (sneakl (firstr x) (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1
EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D0

EqShiftL k x (sneakl (firstr y) (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D0
k < size
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1
EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D0

k < size
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1
EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1

EqShiftL (S k) (shiftr x) (shiftr y)
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1

EqShiftL k (twice_plus_one (shiftr x)) (twice_plus_one (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1

EqShiftL k (sneakl (firstr x) (shiftr x)) (sneakl (firstr x) (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1

EqShiftL k x (sneakl (firstr x) (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1

EqShiftL k x (sneakl (firstr y) (shiftr y))
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1
k < size
k:nat
x, y:int31
H:EqShiftL k x y
Hlt:S k < size
H0:firstr x = D1

k < size
omega. Qed.

forall (n k : nat) (x y : int31), n <= size -> (n + k)%nat = S size -> EqShiftL k x y -> EqShiftL k (incrbis_aux n x) (incrbis_aux n y)

forall (n k : nat) (x y : int31), n <= size -> (n + k)%nat = S size -> EqShiftL k x y -> EqShiftL k (incrbis_aux n x) (incrbis_aux n y)
k:nat
x, y:int31
H:0 <= size
H0:k = S size
H1:EqShiftL k x y

EqShiftL k (incrbis_aux 0 x) (incrbis_aux 0 y)
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
EqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y

EqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
e:k = size

EqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
EqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size

EqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size

EqShiftL k match firstr x with | D0 => sneakl D1 (shiftr x) | D1 => sneakl D0 (incrbis_aux n (shiftr x)) end match firstr y with | D0 => sneakl D1 (shiftr y) | D1 => sneakl D0 (incrbis_aux n (shiftr y)) end
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size

EqShiftL k match firstr y with | D0 => sneakl D1 (shiftr x) | D1 => sneakl D0 (incrbis_aux n (shiftr x)) end match firstr y with | D0 => sneakl D1 (shiftr y) | D1 => sneakl D0 (incrbis_aux n (shiftr y)) end
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
H2:firstr y = D0

EqShiftL k (sneakl D1 (shiftr x)) (sneakl D1 (shiftr y))
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
H2:firstr y = D1
EqShiftL k (sneakl D0 (incrbis_aux n (shiftr x))) (sneakl D0 (incrbis_aux n (shiftr y)))
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
H2:firstr y = D0

EqShiftL (S k) (shiftr x) (shiftr y)
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
H2:firstr y = D1
EqShiftL k (sneakl D0 (incrbis_aux n (shiftr x))) (sneakl D0 (incrbis_aux n (shiftr y)))
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
H2:firstr y = D1

EqShiftL k (sneakl D0 (incrbis_aux n (shiftr x))) (sneakl D0 (incrbis_aux n (shiftr y)))
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
H2:firstr y = D1

EqShiftL (S k) (incrbis_aux n (shiftr x)) (incrbis_aux n (shiftr y))
n:nat
IHn:forall (k0 : nat) (x0 y0 : int31), n <= size -> (n + k0)%nat = S size -> EqShiftL k0 x0 y0 -> EqShiftL k0 (incrbis_aux n x0) (incrbis_aux n y0)
k:nat
x, y:int31
H:S n <= size
H0:S (n + k) = S size
H1:EqShiftL k x y
n0:k <> size
H2:firstr y = D1

EqShiftL (S k) (shiftr x) (shiftr y)
apply EqShiftL_shiftr; auto. Qed.

forall x y : int31, EqShiftL 1 x y -> EqShiftL 1 (incr x) (incr y)

forall x y : int31, EqShiftL 1 x y -> EqShiftL 1 (incr x) (incr y)
x, y:int31
H:EqShiftL 1 x y

EqShiftL 1 (incr x) (incr y)
x, y:int31
H:EqShiftL 1 x y

EqShiftL 1 (incrbis_aux size x) (incrbis_aux size y)
apply EqShiftL_incrbis; auto. Qed. End EqShiftL.

More equations about incr

 

forall x : int31, incr (twice_plus_one x) = twice (incr x)

forall x : int31, incr (twice_plus_one x) = twice (incr x)
x:int31

incr (twice_plus_one x) = twice (incr x)
x:int31

twice (incr (shiftr (twice_plus_one x))) = twice (incr x)
x:int31

EqShiftL 1 (shiftr (twice_plus_one x)) x
red; destruct x; simpl; auto. Qed.

forall x : int31, firstr (incr x) <> firstr x

forall x : int31, firstr (incr x) <> firstr x
x:int31

firstr (incr x) <> firstr x
x:int31
H:firstr x = D0

firstr (incr x) <> D0
x:int31
H:firstr x = D1
firstr (incr x) <> D1
x:int31
H:firstr x = D0

firstr (twice_plus_one (shiftr x)) <> D0
x:int31
H:firstr x = D1
firstr (incr x) <> D1
x:int31
H:firstr x = D1

firstr (incr x) <> D1
x:int31
H:firstr x = D1

firstr (twice (incr (shiftr x))) <> D1
destruct (incr (shiftr x)); simpl; discriminate. Qed.

forall x y : int31, incr x = twice_plus_one y -> x = twice y

forall x y : int31, incr x = twice_plus_one y -> x = twice y
x, y:int31
H:incr x = twice_plus_one y

x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = true

x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
x = twice y
x, y:int31
H:incr 0 = twice_plus_one y
H0:iszero x = true

0 = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
x = twice y
x, y:int31
H:1 = twice_plus_one y
H0:iszero x = true

0 = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false

x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
H1:firstr x = D0

x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
H1:firstr x = D1
x = twice y
x, y:int31
H:twice_plus_one (shiftr x) = twice_plus_one y
H0:iszero x = false
H1:firstr x = D0

x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
H1:firstr x = D1
x = twice y
d, d0, d1, d2, d3, d4, d5, d6, d7, d8, d9, d10, d11, d12, d13, d14, d15, d16, d17, d18, d19, d20, d21, d22, d23, d24, d25, d26, d27, d28, d29, d30, d31, d32, d33, d34, d35, d36, d37, d38, d39, d40, d41, d42, d43, d44, d45, d46, d47, d48, d49, d50, d51, d52, d53, d54, d55, d56, d57, d58, d59, d60:digits
H:I31 d d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 D1 = I31 d31 d32 d33 d34 d35 d36 d37 d38 d39 d40 d41 d42 d43 d44 d45 d46 d47 d48 d49 d50 d51 d52 d53 d54 d55 d56 d57 d58 d59 d60 D1
H1:d29 = D0

I31 d d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29 = I31 d31 d32 d33 d34 d35 d36 d37 d38 d39 d40 d41 d42 d43 d44 d45 d46 d47 d48 d49 d50 d51 d52 d53 d54 d55 d56 d57 d58 d59 d60 D0
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
H1:firstr x = D1
x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
H1:firstr x = D1

x = twice y
x, y:int31
H:incr x = twice_plus_one y
H0:iszero x = false
H1:firstr x = D1

firstr (incr x) = firstr x
rewrite H1, H; destruct y; simpl; auto. Qed.

Conversion from Z : the phi_inv function

 
First, recursive equations
 

forall z : Z, phi_inv (Z.succ_double z) = twice_plus_one (phi_inv z)

forall z : Z, phi_inv (Z.succ_double z) = twice_plus_one (phi_inv z)
p:positive

incr (complement_negative (Pos.pred_double p)) = twice_plus_one (incr (complement_negative p))
p:positive
IHp:incr (complement_negative (Pos.pred_double p)) = twice_plus_one (incr (complement_negative p))

incr (twice (twice_plus_one (complement_negative p))) = twice_plus_one (incr (twice (complement_negative p)))
p:positive
IHp:incr (complement_negative (Pos.pred_double p)) = twice_plus_one (incr (complement_negative p))
incr (twice (complement_negative (Pos.pred_double p))) = twice_plus_one (incr (twice_plus_one (complement_negative p)))

incr 2147483646 = 2147483647
p:positive
IHp:incr (complement_negative (Pos.pred_double p)) = twice_plus_one (incr (complement_negative p))

incr (twice (complement_negative (Pos.pred_double p))) = twice_plus_one (incr (twice_plus_one (complement_negative p)))

incr 2147483646 = 2147483647
p:positive
IHp:incr (complement_negative (Pos.pred_double p)) = twice_plus_one (incr (complement_negative p))

twice_plus_one (complement_negative (Pos.pred_double p)) = twice_plus_one (twice (incr (complement_negative p)))

incr 2147483646 = 2147483647
p:positive
IHp:incr (complement_negative (Pos.pred_double p)) = twice_plus_one (incr (complement_negative p))

complement_negative (Pos.pred_double p) = twice (incr (complement_negative p))

incr 2147483646 = 2147483647

incr 2147483646 = 2147483647
auto. Qed.

forall z : Z, phi_inv (Z.double z) = twice (phi_inv z)

forall z : Z, phi_inv (Z.double z) = twice (phi_inv z)
p:positive

incr (twice_plus_one (complement_negative p)) = twice (incr (complement_negative p))
rewrite incr_twice_plus_one; auto. Qed.

forall z : Z, phi_inv (Z.succ z) = incr (phi_inv z)

forall z : Z, phi_inv (Z.succ z) = incr (phi_inv z)

phi_inv (Z.succ 0) = incr (phi_inv 0)
p:positive
phi_inv (Z.succ (Z.pos p)) = incr (phi_inv (Z.pos p))
p:positive
phi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))
p:positive

phi_inv (Z.succ (Z.pos p)) = incr (phi_inv (Z.pos p))
p:positive
phi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))
p:positive

phi_inv_positive (p + 1) = incr (phi_inv_positive p)
p:positive
phi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))
p:positive
IHp:phi_inv_positive (p + 1) = incr (phi_inv_positive p)

twice (phi_inv_positive (Pos.succ p)) = incr (twice_plus_one (phi_inv_positive p))
p:positive
IHp:phi_inv_positive (p + 1) = incr (phi_inv_positive p)
twice_plus_one (phi_inv_positive p) = incr (twice (phi_inv_positive p))
p:positive
phi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))
p:positive
IHp:phi_inv_positive (p + 1) = incr (phi_inv_positive p)

twice_plus_one (phi_inv_positive p) = incr (twice (phi_inv_positive p))
p:positive
phi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))
p:positive

phi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))
p:positive

phi_inv (Z.pos_sub 1 p) = incr (incr (complement_negative p))
p:positive

incr (twice_plus_one (complement_negative p)) = incr (incr (twice (complement_negative p)))
p:positive
incr (complement_negative (Pos.pred_double p)) = incr (incr (twice_plus_one (complement_negative p)))
p:positive

incr (complement_negative (Pos.pred_double p)) = incr (incr (twice_plus_one (complement_negative p)))
p:positive

complement_negative (Pos.pred_double p) = incr (twice_plus_one (complement_negative p))
p:positive

complement_negative (Pos.pred_double p) = twice (incr (complement_negative p))
p:positive
IHp:complement_negative (Pos.pred_double p) = twice (incr (complement_negative p))

twice (twice_plus_one (complement_negative p)) = twice (incr (twice (complement_negative p)))
p:positive
IHp:complement_negative (Pos.pred_double p) = twice (incr (complement_negative p))
twice (complement_negative (Pos.pred_double p)) = twice (incr (twice_plus_one (complement_negative p)))
p:positive
IHp:complement_negative (Pos.pred_double p) = twice (incr (complement_negative p))

twice (complement_negative (Pos.pred_double p)) = twice (incr (twice_plus_one (complement_negative p)))
p:positive
IHp:complement_negative (Pos.pred_double p) = twice (incr (complement_negative p))

complement_negative (Pos.pred_double p) = incr (twice_plus_one (complement_negative p))
rewrite incr_twice_plus_one; auto. Qed.
phi_inv o inv, the always-exact and easy-to-prove trip : from int31 to Z and then back to int31.
 

forall (n : nat) (x : int31), n <= size -> phi_inv (phibis_aux n (nshiftr x (size - n))) = nshiftr x (size - n)

forall (n : nat) (x : int31), n <= size -> phi_inv (phibis_aux n (nshiftr x (size - n))) = nshiftr x (size - n)

forall x : int31, 0 <= size -> phi_inv (phibis_aux 0 (nshiftr x (size - 0))) = nshiftr x (size - 0)
n:nat
IHn:forall x : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x (size - n))) = nshiftr x (size - n)
forall x : int31, S n <= size -> phi_inv (phibis_aux (S n) (nshiftr x (size - S n))) = nshiftr x (size - S n)
x:int31
H:0 <= size

phi_inv (phibis_aux 0 (nshiftr x 31)) = nshiftr x 31
n:nat
IHn:forall x : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x (size - n))) = nshiftr x (size - n)
forall x : int31, S n <= size -> phi_inv (phibis_aux (S n) (nshiftr x (size - S n))) = nshiftr x (size - S n)
n:nat
IHn:forall x : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x (size - n))) = nshiftr x (size - n)

forall x : int31, S n <= size -> phi_inv (phibis_aux (S n) (nshiftr x (size - S n))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size

phi_inv (phibis_aux (S n) (nshiftr x (size - S n))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size

phi_inv (match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr (nshiftr x (size - S n))))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size

shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
phi_inv (match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr (nshiftr x (size - S n))))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)

phi_inv (match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr (nshiftr x (size - S n))))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)

phi_inv (match firstr (nshiftr x (size - S n)) with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D0

phi_inv (Z.double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1
phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D0

twice (phi_inv (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1
phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D0

twice (nshiftr x (size - n)) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1
phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D0

twice (shiftr (nshiftr x (size - S n))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1
phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
y:(fun _ : nat => int31) (size - S n)%nat
Heqy:y = nshiftr x (size - S n)
H0:shiftr y = nshiftr x (size - n)
H1:firstr y = D0

twice (shiftr y) = y
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1
phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1

phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1

twice_plus_one (phi_inv (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1

twice_plus_one (nshiftr x (size - n)) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
H0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)
H1:firstr (nshiftr x (size - S n)) = D1

twice_plus_one (shiftr (nshiftr x (size - S n))) = nshiftr x (size - S n)
n:nat
IHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)
x:int31
H:S n <= size
y:(fun _ : nat => int31) (size - S n)%nat
Heqy:y = nshiftr x (size - S n)
H0:shiftr y = nshiftr x (size - n)
H1:firstr y = D1

twice_plus_one (shiftr y) = y
destruct y; simpl in H1; rewrite H1; auto. Qed.

forall x : int31, phi_inv (phi x) = x

forall x : int31, phi_inv (phi x) = x
x:int31

phi_inv (phi x) = x
x:int31

phi_inv (phibis_aux size x) = x
x:int31

phi_inv (phibis_aux size (nshiftr x (size - size))) = nshiftr x (size - size)
apply phi_inv_phi_aux; auto. Qed.
The other composition phi o phi_inv is harder to prove correct. In particular, an overflow can happen, so a modulo is needed. For the moment, we proceed via several steps, the first one being a detour to positive_to_in31.
 

positive_to_int31

 
A variant of p2i with twice and twice_plus_one instead of 2*i and 2*i+1
 Fixpoint p2ibis n p : (N*int31)%type :=
  match n with
    | O => (Npos p, On)
    | S n => match p with
               | xO p => let (r,i) := p2ibis n p in (r, twice i)
               | xI p => let (r,i) := p2ibis n p in (r, twice_plus_one i)
               | xH => (N0, In)
             end
  end.

 

forall (n : nat) (p : positive), nshiftr (snd (p2ibis n p)) n = 0

forall (n : nat) (p : positive), nshiftr (snd (p2ibis n p)) n = 0

forall p : positive, nshiftr (snd (p2ibis 0 p)) 0 = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
forall p : positive, nshiftr (snd (p2ibis (S n) p)) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0

forall p : positive, nshiftr (snd (p2ibis (S n) p)) (S n) = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive

nshiftr (snd match p with | (p0~1)%positive => let (r, i) := p2ibis n p0 in (r, twice_plus_one i) | (p0~0)%positive => let (r, i) := p2ibis n p0 in (r, twice i) | 1%positive => (0%N, In) end) (S n) = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive

nshiftr (snd (let (r, i) := p2ibis n p in (r, twice_plus_one i))) (S n) = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
IHn:nshiftr (snd (p2ibis n p)) n = 0

nshiftr (snd (let (r, i) := p2ibis n p in (r, twice_plus_one i))) (S n) = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr (snd (n0, i)) n = 0

nshiftr (snd (n0, twice_plus_one i)) (S n) = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0

nshiftr (twice_plus_one i) (S n) = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0

nshiftr (shiftr (twice_plus_one i)) n = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hle:size <= n

nshiftr (shiftr (twice_plus_one i)) n = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size
nshiftr (shiftr (twice_plus_one i)) n = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size

nshiftr (shiftr (twice_plus_one i)) n = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size
H:firstl i = D0

nshiftr (shiftr (twice_plus_one i)) n = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size
H:firstl i = D0

i = shiftr (twice_plus_one i)
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
d, d0, d1, d2, d3, d4, d5, d6, d7, d8, d9, d10, d11, d12, d13, d14, d15, d16, d17, d18, d19, d20, d21, d22, d23, d24, d25, d26, d27, d28, d29:digits
IHn:nshiftr (I31 d d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29) n = 0
Hlt:n < size
H:d = D0

I31 d d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29 = I31 D0 d0 d1 d2 d3 d4 d5 d6 d7 d8 d9 d10 d11 d12 d13 d14 d15 d16 d17 d18 d19 d20 d21 d22 d23 d24 d25 d26 d27 d28 d29
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive
nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
IHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0
p:positive

nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
IHn:nshiftr (snd (p2ibis n p)) n = 0

nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0

nshiftr (twice i) (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0

nshiftr (shiftr (twice i)) n = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hle:size <= n

nshiftr (shiftr (twice i)) n = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size
nshiftr (shiftr (twice i)) n = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size

nshiftr (shiftr (twice i)) n = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size
H:firstl i = D0

nshiftr (shiftr (twice i)) n = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
p:positive
n0:N
i:int31
IHn:nshiftr i n = 0
Hlt:n < size
H:firstl i = D0

i = shiftr (twice i)
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0
nshiftr In (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0

nshiftr In (S n) = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0

nshiftr (shiftr In) n = 0
n:nat
IHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0

nshiftr 0 n = 0
apply nshiftr_n_0. Qed. Local Open Scope Z_scope.

forall (n : nat) (p : positive), (n <= size)%nat -> Z.pos p = Z.of_N (fst (p2ibis n p)) * 2 ^ Z.of_nat n + phi (snd (p2ibis n p))

forall (n : nat) (p : positive), (n <= size)%nat -> Z.pos p = Z.of_N (fst (p2ibis n p)) * 2 ^ Z.of_nat n + phi (snd (p2ibis n p))
p:positive
H:(0 <= size)%nat

Z.pos p = Z.of_N (fst (p2ibis 0 p)) * 2 ^ Z.of_nat 0 + phi (snd (p2ibis 0 p))
n:nat
IHn:forall p0 : positive, (n <= size)%nat -> Z.pos p0 = Z.of_N (fst (p2ibis n p0)) * 2 ^ Z.of_nat n + phi (snd (p2ibis n p0))
p:positive
H:(S n <= size)%nat
Z.pos p = Z.of_N (fst (p2ibis (S n) p)) * 2 ^ Z.of_nat (S n) + phi (snd (p2ibis (S n) p))
n:nat
IHn:forall p0 : positive, (n <= size)%nat -> Z.pos p0 = Z.of_N (fst (p2ibis n p0)) * 2 ^ Z.of_nat n + phi (snd (p2ibis n p0))
p:positive
H:(S n <= size)%nat

Z.pos p = Z.of_N (fst (p2ibis (S n) p)) * 2 ^ Z.of_nat (S n) + phi (snd (p2ibis (S n) p))
n:nat
IHn:forall p0 : positive, (n <= size)%nat -> Z.pos p0 = Z.of_N (fst (p2ibis n p0)) * 2 ^ Z.of_nat n + phi (snd (p2ibis n p0))
p:positive
H:(S n <= size)%nat

Z.pos p = Z.of_N (fst (p2ibis (S n) p)) * (2 * 2 ^ Z.of_nat n) + phi (snd (p2ibis (S n) p))
n:nat
IHn:forall p0 : positive, (n <= size)%nat -> Z.pos p0 = Z.of_N (fst (p2ibis n p0)) * 2 ^ Z.of_nat n + phi (snd (p2ibis n p0))
p:positive
H:(S n <= size)%nat

Z.pos p = Z.of_N (fst (p2ibis (S n) p)) * (2 ^ Z.of_nat n * 2) + phi (snd (p2ibis (S n) p))
n:nat
IHn:forall p0 : positive, (n <= size)%nat -> Z.pos p0 = Z.of_N (fst (p2ibis n p0)) * 2 ^ Z.of_nat n + phi (snd (p2ibis n p0))
p:positive
H:(S n <= size)%nat
H0:(n <= size)%nat

Z.pos p = Z.of_N (fst (p2ibis (S n) p)) * (2 ^ Z.of_nat n * 2) + phi (snd (p2ibis (S n) p))
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

Z.pos p~1 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice_plus_one i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31
Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

2 * Z.pos p + 1 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice_plus_one i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31
Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

2 * Z.pos p + 1 = Z.of_N r * (2 ^ Z.of_nat n * 2) + (2 * phi i + 1)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31
firstl i = D0
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31
Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

firstl i = D0
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31
Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

2 * Z.pos p = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

2 * Z.pos p = Z.of_N r * (2 ^ Z.of_nat n * 2) + Z.double (phi i)
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31
firstl i = D0
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

2 * Z.pos p = Z.of_N r * (2 ^ Z.of_nat n * 2) + 2 * phi i
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31
firstl i = D0
n:nat
p:positive
r:N
i:int31
IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi i
H:(S n <= size)%nat
H0:(n <= size)%nat
H1:nshiftr i n = 0%int31

firstl i = D0
apply (nshiftr_0_firstl n); auto; try omega. Qed.
We now prove that this p2ibis is related to phi_inv_positive
 

forall (n : nat) (p : positive), (n <= size)%nat -> EqShiftL (size - n) (phi_inv_positive p) (snd (p2ibis n p))

forall (n : nat) (p : positive), (n <= size)%nat -> EqShiftL (size - n) (phi_inv_positive p) (snd (p2ibis n p))

forall p : positive, (0 <= size)%nat -> EqShiftL (size - 0) (phi_inv_positive p) (snd (p2ibis 0 p))
n:nat
IHn:forall p : positive, (n <= size)%nat -> EqShiftL (size - n) (phi_inv_positive p) (snd (p2ibis n p))
forall p : positive, (S n <= size)%nat -> EqShiftL (size - S n) (phi_inv_positive p) (snd (p2ibis (S n) p))
p:positive
H:(0 <= size)%nat

EqShiftL (size - 0) (phi_inv_positive p) (snd (p2ibis 0 p))
n:nat
IHn:forall p : positive, (n <= size)%nat -> EqShiftL (size - n) (phi_inv_positive p) (snd (p2ibis n p))
forall p : positive, (S n <= size)%nat -> EqShiftL (size - S n) (phi_inv_positive p) (snd (p2ibis (S n) p))
n:nat
IHn:forall p : positive, (n <= size)%nat -> EqShiftL (size - n) (phi_inv_positive p) (snd (p2ibis n p))

forall p : positive, (S n <= size)%nat -> EqShiftL (size - S n) (phi_inv_positive p) (snd (p2ibis (S n) p))
n:nat
IHn:forall p0 : positive, (n <= size)%nat -> EqShiftL (size - n) (phi_inv_positive p0) (snd (p2ibis n p0))
p:positive
H:(S n <= size)%nat

EqShiftL (size - S n) (phi_inv_positive p) (snd (p2ibis (S n) p))
simpl p2ibis; destruct p; [ | | red; auto]; specialize IHn with p; destruct (p2ibis n p); simpl @snd in *; simpl phi_inv_positive; rewrite ?EqShiftL_twice_plus_one, ?EqShiftL_twice; replace (S (size - S n))%nat with (size - n)%nat by omega; apply IHn; omega. Qed.
This gives the expected result about phi o phi_inv, at least for the positive case.
 

forall p : positive, phi (phi_inv_positive p) = Z.pos p mod 2 ^ Z.of_nat size

forall p : positive, phi (phi_inv_positive p) = Z.pos p mod 2 ^ Z.of_nat size
p:positive

phi (phi_inv_positive p) = Z.pos p mod 2 ^ Z.of_nat size
p:positive

phi (snd (p2ibis size p)) = Z.pos p mod 2 ^ Z.of_nat size
p:positive
snd (p2ibis size p) = phi_inv_positive p
p:positive

phi (snd (p2ibis size p)) = (Z.of_N (fst (p2ibis size p)) * 2 ^ Z.of_nat size + phi (snd (p2ibis size p))) mod 2 ^ Z.of_nat size
p:positive
snd (p2ibis size p) = phi_inv_positive p
p:positive

phi (snd (p2ibis size p)) = phi (snd (p2ibis size p)) mod 2 ^ Z.of_nat size
p:positive
2 ^ Z.of_nat size > 0
p:positive
snd (p2ibis size p) = phi_inv_positive p
p:positive

0 <= phi (snd (p2ibis size p)) < 2 ^ Z.of_nat size
p:positive
2 ^ Z.of_nat size > 0
p:positive
snd (p2ibis size p) = phi_inv_positive p
p:positive

2 ^ Z.of_nat size > 0
p:positive
snd (p2ibis size p) = phi_inv_positive p
p:positive

snd (p2ibis size p) = phi_inv_positive p
p:positive

phi_inv_positive p = snd (p2ibis size p)
p:positive

EqShiftL 0 (phi_inv_positive p) (snd (p2ibis size p))
apply (phi_inv_positive_p2ibis size p); auto. Qed.
Moreover, p2ibis is also related with p2i and hence with positive_to_int31.
 

forall x : int31, firstl x = D0 -> (Twon * x)%int31 = twice x

forall x : int31, firstl x = D0 -> (Twon * x)%int31 = twice x
x:int31
H:firstl x = D0

(Twon * x)%int31 = twice x
x:int31
H:firstl x = D0

phi_inv (phi Twon * phi x) = twice x
rewrite <- Z.double_spec, <- phi_twice_firstl, phi_inv_phi; auto. Qed.

forall x : int31, firstl x = D0 -> (Twon * x + In)%int31 = twice_plus_one x

forall x : int31, firstl x = D0 -> (Twon * x + In)%int31 = twice_plus_one x
x:int31
H:firstl x = D0

(Twon * x + In)%int31 = twice_plus_one x
x:int31
H:firstl x = D0

(twice x + In)%int31 = twice_plus_one x
x:int31
H:firstl x = D0

phi_inv (phi (twice x) + phi In) = twice_plus_one x
rewrite phi_twice_firstl, <- Z.succ_double_spec, <- phi_twice_plus_one_firstl, phi_inv_phi; auto. Qed.

forall (n : nat) (p : positive), (n <= size)%nat -> p2i n p = p2ibis n p

forall (n : nat) (p : positive), (n <= size)%nat -> p2i n p = p2ibis n p
n:nat
IHn:forall p0 : positive, (n <= size)%nat -> p2i n p0 = p2ibis n p0
p:positive
H:(S n <= size)%nat

match p with | (p0~1)%positive => let (r, i) := p2i n p0 in (r, (Twon * i + In)%int31) | (p0~0)%positive => let (r, i) := p2i n p0 in (r, (Twon * i)%int31) | 1%positive => (0%N, In) end = match p with | (p0~1)%positive => let (r, i) := p2ibis n p0 in (r, twice_plus_one i) | (p0~0)%positive => let (r, i) := p2ibis n p0 in (r, twice i) | 1%positive => (0%N, In) end
n:nat
p:positive
n0:N
i:int31
IHn:(n <= size)%nat -> p2i n p = (n0, i)
H:(S n <= size)%nat
H0:nshiftr i n = 0%int31

(Twon * i + In)%int31 = twice_plus_one i
n:nat
p:positive
n0:N
i:int31
IHn:(n <= size)%nat -> p2i n p = (n0, i)
H:(S n <= size)%nat
H0:nshiftr i n = 0%int31
(Twon * i)%int31 = twice i
n:nat
p:positive
n0:N
i:int31
IHn:(n <= size)%nat -> p2i n p = (n0, i)
H:(S n <= size)%nat
H0:nshiftr i n = 0%int31

firstl i = D0
n:nat
p:positive
n0:N
i:int31
IHn:(n <= size)%nat -> p2i n p = (n0, i)
H:(S n <= size)%nat
H0:nshiftr i n = 0%int31
(Twon * i)%int31 = twice i
n:nat
p:positive
n0:N
i:int31
IHn:(n <= size)%nat -> p2i n p = (n0, i)
H:(S n <= size)%nat
H0:nshiftr i n = 0%int31

(Twon * i)%int31 = twice i
n:nat
p:positive
n0:N
i:int31
IHn:(n <= size)%nat -> p2i n p = (n0, i)
H:(S n <= size)%nat
H0:nshiftr i n = 0%int31

firstl i = D0
apply (nshiftr_0_firstl n); auto; omega. Qed.

forall p : positive, snd (positive_to_int31 p) = phi_inv_positive p

forall p : positive, snd (positive_to_int31 p) = phi_inv_positive p
p:positive

snd (p2i size p) = phi_inv_positive p
p:positive

snd (p2ibis size p) = phi_inv_positive p
p:positive

phi_inv_positive p = snd (p2ibis size p)
p:positive

EqShiftL 0 (phi_inv_positive p) (snd (p2ibis size p))
apply (phi_inv_positive_p2ibis size); auto. Qed.

forall p : positive, Z.pos p = Z.of_N (fst (positive_to_int31 p)) * 2 ^ Z.of_nat size + phi (snd (positive_to_int31 p))

forall p : positive, Z.pos p = Z.of_N (fst (positive_to_int31 p)) * 2 ^ Z.of_nat size + phi (snd (positive_to_int31 p))

forall p : positive, Z.pos p = Z.of_N (fst (p2i size p)) * 2 ^ Z.of_nat size + phi (snd (p2i size p))
p:positive

Z.pos p = Z.of_N (fst (p2ibis size p)) * 2 ^ Z.of_nat size + phi (snd (p2ibis size p))
apply p2ibis_spec; auto. Qed.
Thanks to the result about phi o phi_inv_positive, we can now establish easily the most general results about phi o twice and so one.
 

forall x : int31, phi (twice x) = Z.double (phi x) mod 2 ^ Z.of_nat size

forall x : int31, phi (twice x) = Z.double (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (twice x) = Z.double (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (twice (phi_inv (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (phi_inv (Z.double (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat size
x:int31

0 <= Z.double (phi x)
x:int31
H:0 <= Z.double (phi x)
phi (phi_inv (Z.double (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat size
x:int31
H:0 <= Z.double (phi x)

phi (phi_inv (Z.double (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat size
x:int31
H:0 <= 0

phi (phi_inv 0) = 0 mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.pos p
phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.pos p

phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p

phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
compute in H; elim H; auto. Qed.

forall x : int31, phi (twice_plus_one x) = Z.succ_double (phi x) mod 2 ^ Z.of_nat size

forall x : int31, phi (twice_plus_one x) = Z.succ_double (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (twice_plus_one x) = Z.succ_double (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (twice_plus_one (phi_inv (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (phi_inv (Z.succ_double (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat size
x:int31

0 <= Z.succ_double (phi x)
x:int31
H:0 <= Z.succ_double (phi x)
phi (phi_inv (Z.succ_double (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat size
x:int31
H:0 <= Z.succ_double (phi x)

phi (phi_inv (Z.succ_double (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat size
x:int31
H:0 <= 0

phi (phi_inv 0) = 0 mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.pos p
phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.pos p

phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p

phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
compute in H; elim H; auto. Qed.

forall x : int31, phi (incr x) = Z.succ (phi x) mod 2 ^ Z.of_nat size

forall x : int31, phi (incr x) = Z.succ (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (incr x) = Z.succ (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (incr (phi_inv (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat size
x:int31

phi (phi_inv (Z.succ (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat size
x:int31

0 <= Z.succ (phi x)
x:int31
H:0 <= Z.succ (phi x)
phi (phi_inv (Z.succ (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat size
x:int31
H:0 <= Z.succ (phi x)

phi (phi_inv (Z.succ (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat size
x:int31
H:0 <= 0

phi (phi_inv 0) = 0 mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.pos p
phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.pos p

phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
x:int31
p:positive
H:0 <= Z.neg p

phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
compute in H; elim H; auto. Qed.
With the previous results, we can deal with phi o phi_inv even in the negative case
 

forall p : positive, phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size

forall p : positive, phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size

phi (incr (complement_negative p~1)) = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size

phi (incr (twice (complement_negative p))) = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
IHp:Z.succ (phi (complement_negative p)) mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat size

phi (incr (twice (complement_negative p))) = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
IHp:Z.succ (phi (complement_negative p)) mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat size

Z.succ_double (phi (complement_negative p)) mod 2 ^ Z.of_nat size = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
q:Z
Heqq:q = phi (complement_negative p)
IHp:Z.succ q mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat size

Z.succ_double q mod 2 ^ Z.of_nat size = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
q:Z
Heqq:q = phi (complement_negative p)
IHp:Z.succ q mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat size

(2 * q + 1) mod 2 ^ Z.of_nat size = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
q:Z
Heqq:q = phi (complement_negative p)
IHp:Z.succ q mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat size

(2 * Z.succ q - 1) mod 2 ^ Z.of_nat size = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
q:Z
Heqq:q = phi (complement_negative p)
IHp:Z.succ q mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat size

((2 * (Z.neg p mod 2 ^ Z.of_nat size)) mod 2 ^ Z.of_nat size - 1) mod 2 ^ Z.of_nat size = Z.neg p~1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size
phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size

phi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size

phi (incr (twice_plus_one (complement_negative p))) = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
IHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat size

Z.double (phi (incr (complement_negative p))) mod 2 ^ Z.of_nat size = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
p:positive
q:Z
Heqq:q = phi (incr (complement_negative p))
IHp:q = Z.neg p mod 2 ^ Z.of_nat size

Z.double q mod 2 ^ Z.of_nat size = Z.neg p~0 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size

phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat size
simpl; auto. Qed.

forall z : Z, phi (phi_inv z) = z mod 2 ^ Z.of_nat size

forall z : Z, phi (phi_inv z) = z mod 2 ^ Z.of_nat size

phi (phi_inv 0) = 0 mod 2 ^ Z.of_nat size
p:positive
phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
p:positive
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
p:positive

phi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat size
p:positive
phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
p:positive

phi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
apply phi_phi_inv_negative. Qed. End Basics. Instance int31_ops : ZnZ.Ops int31 := { digits := 31%positive; (* number of digits *) zdigits := 31; (* number of digits *) to_Z := phi; (* conversion to Z *) of_pos := positive_to_int31; (* positive -> N*int31 : p => N,i where p = N*2^31+phi i *) head0 := head031; (* number of head 0 *) tail0 := tail031; (* number of tail 0 *) zero := 0; one := 1; minus_one := Tn; (* 2^31 - 1 *) compare := compare31; eq0 := fun i => match i ?= 0 with Eq => true | _ => false end; opp_c := fun i => 0 -c i; opp := opp31; opp_carry := fun i => 0-i-1; succ_c := fun i => i +c 1; add_c := add31c; add_carry_c := add31carryc; succ := fun i => i + 1; add := add31; add_carry := fun i j => i + j + 1; pred_c := fun i => i -c 1; sub_c := sub31c; sub_carry_c := sub31carryc; pred := fun i => i - 1; sub := sub31; sub_carry := fun i j => i - j - 1; mul_c := mul31c; mul := mul31; square_c := fun x => x *c x; div21 := div3121; div_gt := div31; (* this is supposed to be the special case of division a/b where a > b *) div := div31; modulo_gt := fun i j => let (_,r) := i/j in r; modulo := fun i j => let (_,r) := i/j in r; gcd_gt := gcd31; gcd := gcd31; add_mul_div := addmuldiv31; pos_mod := (* modulo 2^p *) fun p i => match p ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31-p) i 0) | _ => i end; is_even := fun i => let (_,r) := i/2 in match r ?= 0 with Eq => true | _ => false end; sqrt2 := sqrt312; sqrt := sqrt31; lor := lor31; land := land31; lxor := lxor31 }. Section Int31_Specs. Local Open Scope Z_scope. Notation "[| x |]" := (phi x) (at level 0, x at level 99). Notation wB := (2 ^ (Z.of_nat size)).

wB > 0

wB > 0
auto with zarith. Qed. Notation "[+| c |]" := (interp_carry 1 wB phi c) (at level 0, c at level 99). Notation "[-| c |]" := (interp_carry (-1) wB phi c) (at level 0, c at level 99). Notation "[|| x ||]" := (zn2z_to_Z wB phi x) (at level 0, x at level 99).

[|31|] = 31

[|31|] = 31
reflexivity. Qed.

1 < 31

1 < 31
auto with zarith. Qed.

[|0|] = 0

[|0|] = 0
reflexivity. Qed.

[|1|] = 1

[|1|] = 1
reflexivity. Qed.

[|Tn|] = wB - 1

[|Tn|] = wB - 1
reflexivity. Qed.

forall x y : int31, (x ?= y)%int31 = ([|x|] ?= [|y|])

forall x y : int31, (x ?= y)%int31 = ([|x|] ?= [|y|])
reflexivity. Qed.
Addition
 

forall x y : int31, [+|x +c y|] = [|x|] + [|y|]

forall x y : int31, [+|x +c y|] = [|x|] + [|y|]
x, y:int31

match match ([|x|] + [|y|]) mod wB ?= [|x|] + [|y|] with | Eq => C0 (phi_inv ([|x|] + [|y|])) | _ => C1 (phi_inv ([|x|] + [|y|])) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = [|x|] + [|y|]
x, y:int31
H:0 <= [|x|] < wB
H0:0 <= [|y|] < wB

match match ([|x|] + [|y|]) mod wB ?= [|x|] + [|y|] with | Eq => C0 (phi_inv ([|x|] + [|y|])) | _ => C1 (phi_inv ([|x|] + [|y|])) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = [|x|] + [|y|]
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y
match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y) mod wB <> X + Y

1 * wB + (X + Y) mod wB = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y
match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y) mod wB <> X + Y
l:X + Y < wB

1 * wB + (X + Y) mod wB = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y) mod wB <> X + Y
l:wB <= X + Y
1 * wB + (X + Y) mod wB = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y
match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y) mod wB <> X + Y
l:wB <= X + Y

1 * wB + (X + Y) mod wB = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y
match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y) mod wB <> X + Y
l:wB <= X + Y

1 * wB + (X + Y + -1 * wB) mod wB = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y
match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y

match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Y
Heq:((X + Y) mod wB ?= X + Y) = Eq -> (X + Y) mod wB = X + Y

match match (X + Y) mod wB ?= X + Y with | Eq => C0 (phi_inv (X + Y)) | _ => C1 (phi_inv (X + Y)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y
destruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.

forall x : int31, [+|x +c 1|] = [|x|] + 1

forall x : int31, [+|x +c 1|] = [|x|] + 1
intros; apply spec_add_c. Qed.

forall x y : int31, [+|add31carryc x y|] = [|x|] + [|y|] + 1

forall x y : int31, [+|add31carryc x y|] = [|x|] + [|y|] + 1
x, y:int31

[+|add31carryc x y|] = [|x|] + [|y|] + 1
x, y:int31

match match ([|x|] + [|y|] + 1) mod wB ?= [|x|] + [|y|] + 1 with | Eq => C0 (phi_inv ([|x|] + [|y|] + 1)) | _ => C1 (phi_inv ([|x|] + [|y|] + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = [|x|] + [|y|] + 1
x, y:int31
H:0 <= [|x|] < wB
H0:0 <= [|y|] < wB

match match ([|x|] + [|y|] + 1) mod wB ?= [|x|] + [|y|] + 1 with | Eq => C0 (phi_inv ([|x|] + [|y|] + 1)) | _ => C1 (phi_inv ([|x|] + [|y|] + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = [|x|] + [|y|] + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1
match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y + 1) mod wB <> X + Y + 1

1 * wB + (X + Y + 1) mod wB = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1
match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y + 1) mod wB <> X + Y + 1
l:X + Y + 1 < wB

1 * wB + (X + Y + 1) mod wB = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y + 1) mod wB <> X + Y + 1
l:wB <= X + Y + 1
1 * wB + (X + Y + 1) mod wB = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1
match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y + 1) mod wB <> X + Y + 1
l:wB <= X + Y + 1

1 * wB + (X + Y + 1) mod wB = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1
match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X + Y + 1) mod wB <> X + Y + 1
l:wB <= X + Y + 1

1 * wB + (X + Y + 1 + -1 * wB) mod wB = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1
match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1

match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1
Heq:((X + Y + 1) mod wB ?= X + Y + 1) = Eq -> (X + Y + 1) mod wB = X + Y + 1

match match (X + Y + 1) mod wB ?= X + Y + 1 with | Eq => C0 (phi_inv (X + Y + 1)) | _ => C1 (phi_inv (X + Y + 1)) end with | C0 x0 => [|x0|] | C1 x0 => 1 * wB + [|x0|] end = X + Y + 1
destruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.

forall x y : int31, [|x + y|] = ([|x|] + [|y|]) mod wB

forall x y : int31, [|x + y|] = ([|x|] + [|y|]) mod wB
intros; apply phi_phi_inv. Qed.

forall x y : int31, [|x + y + 1|] = ([|x|] + [|y|] + 1) mod wB

forall x y : int31, [|x + y + 1|] = ([|x|] + [|y|] + 1) mod wB
x, y:int31

[|phi_inv ([|phi_inv ([|x|] + [|y|])|] + [|1|])|] = ([|x|] + [|y|] + 1) mod wB
x, y:int31

(([|x|] + [|y|]) mod wB + [|1|]) mod wB = ([|x|] + [|y|] + 1) mod wB
apply Zplus_mod_idemp_l. Qed.

forall x : int31, [|x + 1|] = ([|x|] + 1) mod wB

forall x : int31, [|x + 1|] = ([|x|] + 1) mod wB
intros; rewrite <- spec_1; apply spec_add. Qed.
Subtraction
 

forall x y : int31, [-|x -c y|] = [|x|] - [|y|]

forall x y : int31, [-|x -c y|] = [|x|] - [|y|]
x, y:int31

match match [|phi_inv ([|x|] - [|y|])|] ?= [|x|] - [|y|] with | Eq => C0 (phi_inv ([|x|] - [|y|])) | _ => C1 (phi_inv ([|x|] - [|y|])) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = [|x|] - [|y|]
x, y:int31

match match ([|x|] - [|y|]) mod wB ?= [|x|] - [|y|] with | Eq => C0 (phi_inv ([|x|] - [|y|])) | _ => C1 (phi_inv ([|x|] - [|y|])) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = [|x|] - [|y|]
x, y:int31
H:0 <= [|x|] < wB
H0:0 <= [|y|] < wB

match match ([|x|] - [|y|]) mod wB ?= [|x|] - [|y|] with | Eq => C0 (phi_inv ([|x|] - [|y|])) | _ => C1 (phi_inv ([|x|] - [|y|])) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = [|x|] - [|y|]
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y
match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y) mod wB <> X - Y

-1 * wB + (X - Y) mod wB = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y
match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y) mod wB <> X - Y
l:X - Y < 0

-1 * wB + (X - Y) mod wB = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y) mod wB <> X - Y
l:0 <= X - Y
-1 * wB + (X - Y) mod wB = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y
match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y) mod wB <> X - Y
l:X - Y < 0

-1 * wB + (X - Y + 1 * wB) mod wB = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y) mod wB <> X - Y
l:0 <= X - Y
-1 * wB + (X - Y) mod wB = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y
match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y) mod wB <> X - Y
l:0 <= X - Y

-1 * wB + (X - Y) mod wB = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y
match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y

match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Y
Heq:((X - Y) mod wB ?= X - Y) = Eq -> (X - Y) mod wB = X - Y

match match (X - Y) mod wB ?= X - Y with | Eq => C0 (phi_inv (X - Y)) | _ => C1 (phi_inv (X - Y)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y
destruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.

forall x y : int31, [-|sub31carryc x y|] = [|x|] - [|y|] - 1

forall x y : int31, [-|sub31carryc x y|] = [|x|] - [|y|] - 1
x, y:int31

match match [|phi_inv ([|x|] - [|y|] - 1)|] ?= [|x|] - [|y|] - 1 with | Eq => C0 (phi_inv ([|x|] - [|y|] - 1)) | _ => C1 (phi_inv ([|x|] - [|y|] - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = [|x|] - [|y|] - 1
x, y:int31

match match ([|x|] - [|y|] - 1) mod wB ?= [|x|] - [|y|] - 1 with | Eq => C0 (phi_inv ([|x|] - [|y|] - 1)) | _ => C1 (phi_inv ([|x|] - [|y|] - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = [|x|] - [|y|] - 1
x, y:int31
H:0 <= [|x|] < wB
H0:0 <= [|y|] < wB

match match ([|x|] - [|y|] - 1) mod wB ?= [|x|] - [|y|] - 1 with | Eq => C0 (phi_inv ([|x|] - [|y|] - 1)) | _ => C1 (phi_inv ([|x|] - [|y|] - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = [|x|] - [|y|] - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB

((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1
match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y - 1) mod wB <> X - Y - 1

-1 * wB + (X - Y - 1) mod wB = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1
match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y - 1) mod wB <> X - Y - 1
l:X - Y - 1 < 0

-1 * wB + (X - Y - 1) mod wB = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y - 1) mod wB <> X - Y - 1
l:0 <= X - Y - 1
-1 * wB + (X - Y - 1) mod wB = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1
match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y - 1) mod wB <> X - Y - 1
l:X - Y - 1 < 0

-1 * wB + (X - Y - 1 + 1 * wB) mod wB = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y - 1) mod wB <> X - Y - 1
l:0 <= X - Y - 1
-1 * wB + (X - Y - 1) mod wB = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1
match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:(X - Y - 1) mod wB <> X - Y - 1
l:0 <= X - Y - 1

-1 * wB + (X - Y - 1) mod wB = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1
match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1

match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
x, y:int31
X:Z
H:0 <= X < wB
Y:Z
H0:0 <= Y < wB
H1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1
Heq:((X - Y - 1) mod wB ?= X - Y - 1) = Eq -> (X - Y - 1) mod wB = X - Y - 1

match match (X - Y - 1) mod wB ?= X - Y - 1 with | Eq => C0 (phi_inv (X - Y - 1)) | _ => C1 (phi_inv (X - Y - 1)) end with | C0 x0 => [|x0|] | C1 x0 => -1 * wB + [|x0|] end = X - Y - 1
destruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.

forall x y : int31, [|x - y|] = ([|x|] - [|y|]) mod wB

forall x y : int31, [|x - y|] = ([|x|] - [|y|]) mod wB
intros; apply phi_phi_inv. Qed.

forall x y : int31, [|x - y - 1|] = ([|x|] - [|y|] - 1) mod wB

forall x y : int31, [|x - y - 1|] = ([|x|] - [|y|] - 1) mod wB
x, y:int31

[|phi_inv ([|phi_inv ([|x|] - [|y|])|] - [|1|])|] = ([|x|] - [|y|] - 1) mod wB
x, y:int31

(([|x|] - [|y|]) mod wB - [|1|]) mod wB = ([|x|] - [|y|] - 1) mod wB
apply Zminus_mod_idemp_l. Qed.

forall x : int31, [-|0 -c x|] = - [|x|]

forall x : int31, [-|0 -c x|] = - [|x|]
intros; apply spec_sub_c. Qed.

forall x : int31, [|0 - x|] = - [|x|] mod wB

forall x : int31, [|0 - x|] = - [|x|] mod wB
intros; apply phi_phi_inv. Qed.

forall x : int31, [|0 - x - 1|] = wB - [|x|] - 1

forall x : int31, [|0 - x - 1|] = wB - [|x|] - 1
x:int31

[|phi_inv ([|phi_inv ([|0|] - [|x|])|] - [|1|])|] = wB - [|x|] - 1
x:int31

(([|0|] - [|x|]) mod wB - [|1|]) mod wB = wB - [|x|] - 1
x:int31

((0 - [|x|]) mod wB - 1) mod wB = wB - [|x|] - 1
x:int31

((0 - [|x|] + 1 * wB) mod wB - 1) mod wB = wB - [|x|] - 1
x:int31

(0 - [|x|] + 1 * wB - 1) mod wB = wB - [|x|] - 1
rewrite Zmod_small; generalize (phi_bounded x); lia. Qed.

forall x : int31, [-|x -c 1|] = [|x|] - 1

forall x : int31, [-|x -c 1|] = [|x|] - 1
intros; apply spec_sub_c. Qed.

forall x : int31, [|x - 1|] = ([|x|] - 1) mod wB

forall x : int31, [|x - 1|] = ([|x|] - 1) mod wB
intros; apply spec_sub. Qed.
Multiplication
 

forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2

forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2

forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z

(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z

(z / wB) mod wB = z / wB - z / wB / wB * wB
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB

(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB

z mod wB = z - z / wB * wB
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
H0:z mod wB = z - z / wB * wB
(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
H0:z mod wB = z - z / wB * wB

(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
H0:z mod wB = z - z / wB * wB

(z / wB - z / wB / wB * wB) * wB + z mod wB = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
H0:z mod wB = z - z / wB * wB

(z / wB - z / wB / wB * wB) * wB + (z - z / wB * wB) = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
H0:z mod wB = z - z / wB * wB

- (z / wB / wB) * wB ^ 2 + z = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
H0:z mod wB = z - z / wB * wB

- (z / (wB * wB)) * wB ^ 2 + z = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
z:Z
H:(z / wB) mod wB = z / wB - z / wB / wB * wB
H0:z mod wB = z - z / wB * wB

- (z / (wB * wB)) * wB ^ 2 + (wB * wB * (z / (wB * wB)) + z mod (wB * wB)) = z mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2

forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2
H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2

forall x : Z, [||match x with | 0 => W0 | _ => WW (phi_inv (x / base)) (phi_inv x) end||] = x mod wB ^ 2
destruct x; unfold zn2z_to_Z; rewrite ?phi_phi_inv; change base with wB; auto. Qed.

forall x y : int31, [||x *c y||] = [|x|] * [|y|]

forall x y : int31, [||x *c y||] = [|x|] * [|y|]
x, y:int31

[||phi_inv2 ([|x|] * [|y|])||] = [|x|] * [|y|]
x, y:int31

([|x|] * [|y|]) mod wB ^ 2 = [|x|] * [|y|]
x, y:int31

0 <= [|x|] * [|y|] < wB ^ 2
x, y:int31
H:0 <= [|x|] < wB
H0:0 <= [|y|] < wB

0 <= [|x|] * [|y|] < wB ^ 2
x, y:int31
H:0 <= [|x|] < wB
H0:0 <= [|y|] < wB

0 <= [|x|] * [|y|] < wB * wB
auto using Z.mul_lt_mono_nonneg with zarith. Qed.

forall x y : int31, [|x * y|] = ([|x|] * [|y|]) mod wB

forall x y : int31, [|x * y|] = ([|x|] * [|y|]) mod wB
intros; apply phi_phi_inv. Qed.

forall x : int31, [||x *c x||] = [|x|] * [|x|]

forall x : int31, [||x *c x||] = [|x|] * [|x|]
intros; apply spec_mul_c. Qed.
Division
 

forall a1 a2 b : int31, wB / 2 <= [|b|] -> [|a1|] < [|b|] -> let (q, r) := div3121 a1 a2 b in [|a1|] * wB + [|a2|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]

forall a1 a2 b : int31, wB / 2 <= [|b|] -> [|a1|] < [|b|] -> let (q, r) := div3121 a1 a2 b in [|a1|] * wB + [|a2|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]

let (q, r) := let (q, r) := Z.div_eucl (phi2 a1 a2) [|b|] in (phi_inv q, phi_inv r) in [|a1|] * wB + [|a2|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB

let (q, r) := let (q, r) := Z.div_eucl (phi2 a1 a2) [|b|] in (phi_inv q, phi_inv r) in [|a1|] * wB + [|a2|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0

let (q, r) := let (q, r) := Z.div_eucl (phi2 a1 a2) [|b|] in (phi_inv q, phi_inv r) in [|a1|] * wB + [|a2|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0

(let (q, r) := Z.div_eucl (phi2 a1 a2) [|b|] in phi2 a1 a2 = [|b|] * q + r /\ 0 <= r < [|b|]) -> (0 <= phi2 a1 a2 -> 0 <= phi2 a1 a2 / [|b|]) -> let (q, r) := let (q, r) := Z.div_eucl (phi2 a1 a2) [|b|] in (phi_inv q, phi_inv r) in [|a1|] * wB + [|a2|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z

phi2 a1 a2 = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> (0 <= phi2 a1 a2 -> 0 <= z) -> [|a1|] * wB + [|a2|] = [|phi_inv z|] * [|b|] + [|phi_inv z0|] /\ 0 <= [|phi_inv z0|] < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z

phi2 a1 a2 = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> (0 <= phi2 a1 a2 -> 0 <= z) -> [|a1|] * wB + [|a2|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:phi2 a1 a2 = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= phi2 a1 a2 -> 0 <= z

[|a1|] * wB + [|a2|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * base + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

[|a1|] * wB + [|a2|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

[|a1|] * wB + [|a2|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

[|a1|] * wB + [|a2|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

z * [|b|] + z0 = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

z * [|b|] + z0 = z mod wB * [|b|] + z0 /\ 0 <= z0 < [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

z = z mod wB
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

0 <= z < wB
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

0 <= z
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z
z < wB
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

z < wB
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

z * [|b|] < wB * [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

[|b|] * z < wB * [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

[|b|] * z + z0 < wB * [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

[|a1|] * wB + [|a2|] < wB * [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

[|a1|] * wB + (wB - 1) < wB * [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

wB * ([|a1|] + 1) - 1 < wB * [|b|]
a1, a2, b:int31
H:wB / 2 <= [|b|]
H0:[|a1|] < [|b|]
H1:0 <= [|a1|] < wB
H2:0 <= [|a2|] < wB
H3:0 <= [|b|] < wB
H4:[|b|] > 0
z, z0:Z
H5:[|a1|] * wB + [|a2|] = [|b|] * z + z0
H6:0 <= z0 < [|b|]
H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z

wB * ([|a1|] + 1) <= wB * [|b|]
apply Z.mul_le_mono_nonneg; omega. Qed.

forall a b : int31, 0 < [|b|] -> let (q, r) := (a / b)%int31 in [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]

forall a b : int31, 0 < [|b|] -> let (q, r) := (a / b)%int31 in [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a, b:int31
H:0 < [|b|]

let (q, r) := let (q, r) := Z.div_eucl [|a|] [|b|] in (phi_inv q, phi_inv r) in [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0

let (q, r) := let (q, r) := Z.div_eucl [|a|] [|b|] in (phi_inv q, phi_inv r) in [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0

(let (q, r) := Z.div_eucl [|a|] [|b|] in [|a|] = [|b|] * q + r /\ 0 <= r < [|b|]) -> (0 <= [|a|] -> 0 <= [|a|] / [|b|]) -> let (q, r) := let (q, r) := Z.div_eucl [|a|] [|b|] in (phi_inv q, phi_inv r) in [|a|] = [|q|] * [|b|] + [|r|] /\ 0 <= [|r|] < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z

[|a|] = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> (0 <= [|a|] -> 0 <= z) -> [|a|] = [|phi_inv z|] * [|b|] + [|phi_inv z0|] /\ 0 <= [|phi_inv z0|] < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z

[|a|] = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> (0 <= [|a|] -> 0 <= z) -> [|a|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z

[|a|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z

z * [|b|] + z0 = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

z * [|b|] + z0 = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

z * [|b|] + z0 = z mod wB * [|b|] + z0 /\ 0 <= z0 < [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

z = z mod wB
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

0 <= z < wB
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

z < wB
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

z <= [|a|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

z <= [|b|] * z + z0
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

z <= [|b|] * z
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|a|] -> 0 <= z
H4:0 <= [|a|] < wB
H5:0 <= [|b|] < wB

1 * z <= [|b|] * z
apply Z.mul_le_mono_nonneg; auto with zarith. Qed.

forall a b : int31, 0 < [|b|] -> [|(let (_, r) := (a / b)%int31 in r)|] = [|a|] mod [|b|]

forall a b : int31, 0 < [|b|] -> [|(let (_, r) := (a / b)%int31 in r)|] = [|a|] mod [|b|]
a, b:int31
H:0 < [|b|]

[|(let (_, r) := let (q, r) := Z.div_eucl [|a|] [|b|] in (phi_inv q, phi_inv r) in r)|] = [|a|] mod [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0

[|(let (_, r) := let (q, r) := Z.div_eucl [|a|] [|b|] in (phi_inv q, phi_inv r) in r)|] = [|a|] mod [|b|]
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0

[|(let (_, r) := let (q, r) := Z.div_eucl [|a|] [|b|] in (phi_inv q, phi_inv r) in r)|] = (let (_, r) := Z.div_eucl [|a|] [|b|] in r)
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0

(let (q, r) := Z.div_eucl [|a|] [|b|] in [|a|] = [|b|] * q + r /\ 0 <= r < [|b|]) -> [|(let (_, r) := let (q, r) := Z.div_eucl [|a|] [|b|] in (phi_inv q, phi_inv r) in r)|] = (let (_, r) := Z.div_eucl [|a|] [|b|] in r)
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z

[|a|] = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> [|phi_inv z0|] = z0
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z

[|a|] = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> z0 mod wB = z0
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]

z0 mod wB = z0
a, b:int31
H:0 < [|b|]
H0:[|b|] > 0
z, z0:Z
H1:[|a|] = [|b|] * z + z0
H2:0 <= z0 < [|b|]
H3:0 <= [|b|] < wB

z0 mod wB = z0
apply Zmod_small; omega. Qed.

forall i j : int31, [|gcd31 i j|] = Zgcdn (2 * size) [|j|] [|i|]

forall i j : int31, [|gcd31 i j|] = Zgcdn (2 * size) [|j|] [|i|]

forall i j : int31, [|euler (2 * size) i j|] = Zgcdn (2 * size) [|j|] [|i|]
i, j:int31

[|euler 0 i j|] = Zgcdn 0 [|j|] [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
[|euler (S n) i j|] = Zgcdn (S n) [|j|] [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31

[|euler (S n) i j|] = Zgcdn (S n) [|j|] [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31

[|match (j ?= On)%int31 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) [|j|] [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31

[|match [|j|] ?= [|On|] with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) [|j|] [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31

[|match [|j|] ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) [|j|] [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB

[|match [|j|] ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) [|j|] [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
H1:[|j|] = 0

[|match 0 ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) 0 [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.pos p
[|match Z.pos p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.pos p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.neg p
[|match Z.neg p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.neg p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
H1:[|j|] = 0

[|i|] = Z.abs [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.pos p
[|match Z.pos p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.pos p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.neg p
[|match Z.neg p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.neg p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.pos p

[|match Z.pos p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.pos p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.neg p
[|match Z.neg p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.neg p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.pos p

[|euler n j (let (_, r) := (i / j)%int31 in r)|] = Zgcdn n ([|i|] mod Z.pos p) (Z.pos p)
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.neg p
[|match Z.neg p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.neg p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.pos p

[|(let (_, r) := (i / j)%int31 in r)|] = [|i|] mod Z.pos p
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.neg p
[|match Z.neg p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.neg p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.pos p

0 < [|j|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.neg p
[|match Z.neg p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.neg p) [|i|]
n:nat
IHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]
i, j:int31
H:0 <= [|j|] < wB
H0:0 <= [|i|] < wB
p:positive
H1:[|j|] = Z.neg p

[|match Z.neg p ?= 0 with | Eq => i | _ => euler n j (let (_, r) := (i / j)%int31 in r) end|] = Zgcdn (S n) (Z.neg p) [|i|]
rewrite H1 in H; destruct H as [H _]; compute in H; elim H; auto. Qed.

forall a b : int31, Zis_gcd [|a|] [|b|] [|gcd31 a b|]

forall a b : int31, Zis_gcd [|a|] [|b|] [|gcd31 a b|]
a, b:int31

Zis_gcd [|a|] [|b|] [|gcd31 a b|]
a, b:int31

Zis_gcd [|a|] [|b|] (Zgcdn (2 * size) [|b|] [|a|])
a, b:int31

Zis_gcd [|b|] [|a|] (Zgcdn (2 * size) [|b|] [|a|])
a, b:int31

(Zgcd_bound [|b|] <= 2 * size)%nat
a, b:int31

(match [|b|] with | 0%Z => 1 | Z.pos p | Z.neg p => Pos.size_nat p + Pos.size_nat p end <= 2 * size)%nat
a, b:int31

0 <= [|b|] < wB -> (match [|b|] with | 0%Z => 1 | Z.pos p | Z.neg p => Pos.size_nat p + Pos.size_nat p end <= 2 * size)%nat
a, b:int31

0 <= 0 < wB -> (1 <= 2 * size)%nat
a, b:int31
p:positive
0 <= Z.pos p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nat
a, b:int31
p:positive
0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nat
a, b:int31
p:positive

0 <= Z.pos p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nat
a, b:int31
p:positive
0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nat
a, b:int31
p:positive
H:Z.pos p < wB

(Pos.size_nat p + Pos.size_nat p <= 2 * size)%nat
a, b:int31
p:positive
0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nat
a, b:int31
p:positive

0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nat
intros (H,_); compute in H; elim H; auto. Qed.

forall (A : Type) (f : A -> A) (i : int31) (a : A), iter_int31 i A f a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat [|i|])

forall (A : Type) (f : A -> A) (i : int31) (a : A), iter_int31 i A f a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat [|i|])
A:Type
f:A -> A
i:int31
a:A

iter_int31 i A f a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat [|i|])
A:Type
f:A -> A
i:int31
a:A

recr (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat [|i|])
A:Type
f:A -> A
i:int31
a:A

recrbis_aux size (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat [|i|])
A:Type
f:A -> A
i:int31
a:A

recrbis_aux size (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (phibis_aux size i))
A:Type
f:A -> A

forall (i : int31) (a : A), recrbis_aux 0 (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (phibis_aux 0 i))
A:Type
f:A -> A
n:nat
IHn:forall (i : int31) (a : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (phibis_aux n i))
forall (i : int31) (a : A), recrbis_aux (S n) (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (phibis_aux (S n) i))
A:Type
f:A -> A
n:nat
IHn:forall (i : int31) (a : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (phibis_aux n i))

forall (i : int31) (a : A), recrbis_aux (S n) (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (phibis_aux (S n) i))
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A

match firstr i with | D0 => fun x : A => recrbis_aux n (A -> A) (fun x0 : A => x0) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x0 : A => rec (rec x0) | D1 => fun x0 : A => f (rec (rec x0)) end) (shiftr i) (recrbis_aux n (A -> A) (fun x0 : A => x0) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x0 : A => rec (rec x0) | D1 => fun x0 : A => f (rec (rec x0)) end) (shiftr i) x) | D1 => fun x : A => f (recrbis_aux n (A -> A) (fun x0 : A => x0) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x0 : A => rec (rec x0) | D1 => fun x0 : A => f (rec (rec x0)) end) (shiftr i) (recrbis_aux n (A -> A) (fun x0 : A => x0) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x0 : A => rec (rec x0) | D1 => fun x0 : A => f (rec (rec x0)) end) (shiftr i) x)) end a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (phibis_aux (S n) i))
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D0
z:Z
H0:0 <= z

nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat z + Z.abs_nat z) = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (Z.double z))
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z
f (nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat z + Z.abs_nat z)) = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (Z.succ_double z))
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D0
z:Z
H0:0 <= z

(Z.abs_nat z + Z.abs_nat z)%nat = Z.abs_nat (Z.double z)
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z
f (nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat z + Z.abs_nat z)) = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (Z.succ_double z))
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D0
z:Z
H0:0 <= z

(Z.abs_nat z + Z.abs_nat z)%nat = Z.abs_nat (z + z)
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z
f (nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat z + Z.abs_nat z)) = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (Z.succ_double z))
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z

f (nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat z + Z.abs_nat z)) = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat (Z.succ_double z))
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z

(S (Z.abs_nat z) + Z.abs_nat z)%nat = Z.abs_nat (Z.succ_double z)
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z

(S (Z.abs_nat z) + Z.abs_nat z)%nat = Z.abs_nat (z + z + 1)
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z

(S (Z.abs_nat z) + Z.abs_nat z)%nat = (Z.abs_nat (z + z) + Z.abs_nat 1)%nat
A:Type
f:A -> A
n:nat
IHn:forall (i0 : int31) (a0 : A), recrbis_aux n (A -> A) (fun x : A => x) (fun (b : digits) (_ : int31) (rec : A -> A) => match b with | D0 => fun x : A => rec (rec x) | D1 => fun x : A => f (rec (rec x)) end) i0 a0 = nat_rect (fun _ : nat => A) a0 (fun _ : nat => f) (Z.abs_nat (phibis_aux n i0))
i:int31
a:A
H:firstr i = D1
z:Z
H0:0 <= z

(S (Z.abs_nat z) + Z.abs_nat z)%nat = (Z.abs_nat z + Z.abs_nat z + Z.abs_nat 1)%nat
change (Z.abs_nat 1) with 1%nat; omega. Qed. Fixpoint addmuldiv31_alt n i j := match n with | O => i | S n => addmuldiv31_alt n (sneakl (firstl j) i) (shiftl j) end.

forall p x y : int31, addmuldiv31 p x y = addmuldiv31_alt (Z.abs_nat [|p|]) x y

forall p x y : int31, addmuldiv31 p x y = addmuldiv31_alt (Z.abs_nat [|p|]) x y
p, x, y:int31

addmuldiv31 p x y = addmuldiv31_alt (Z.abs_nat [|p|]) x y
p, x, y:int31

(let (res, _) := iter_int31 p (int31 * int31) (fun ij : int31 * int31 => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) (x, y) in res) = addmuldiv31_alt (Z.abs_nat [|p|]) x y
p, x, y:int31

(let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) (Z.abs_nat [|p|]) in res) = addmuldiv31_alt (Z.abs_nat [|p|]) x y
x, y:int31
n:nat

(let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) n in res) = addmuldiv31_alt n x y

forall x y : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) 0 in res) = addmuldiv31_alt 0 x y
n:nat
IHn:forall x y : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) n in res) = addmuldiv31_alt n x y
forall x y : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) (S n) in res) = addmuldiv31_alt (S n) x y
n:nat
IHn:forall x y : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) n in res) = addmuldiv31_alt n x y

forall x y : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) (S n) in res) = addmuldiv31_alt (S n) x y
n:nat
IHn:forall x0 y0 : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x0, y0) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) n in res) = addmuldiv31_alt n x0 y0
x, y:int31

(let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) (S n) in res) = addmuldiv31_alt (S n) x y
n:nat
IHn:forall x0 y0 : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x0, y0) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) n in res) = addmuldiv31_alt n x0 y0
x, y:int31

(let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) (S n) in res) = addmuldiv31_alt n (sneakl (firstl y) x) (shiftl y)
n:nat
IHn:forall x0 y0 : int31, (let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x0, y0) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) n in res) = addmuldiv31_alt n x0 y0
x, y:int31

(let (res, _) := nat_rect (fun _ : nat => (int31 * int31)%type) (x, y) (fun (_ : nat) (ij : int31 * int31) => let (i, j) := ij in (sneakl (firstl j) i, shiftl j)) (n + 1) in res) = addmuldiv31_alt n (sneakl (firstl y) x) (shiftl y)
rewrite nat_rect_plus; simpl; auto. Qed.

forall x y p : int31, [|p|] <= 31 -> [|addmuldiv31 p x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wB

forall x y p : int31, [|p|] <= 31 -> [|addmuldiv31 p x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wB
x, y, p:int31
H:[|p|] <= 31

[|addmuldiv31 p x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wB
x, y, p:int31
H:[|p|] <= 31

[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wB
x, y, p:int31
H:[|p|] <= 31

[|p|] = Z.of_nat (Z.abs_nat [|p|])
x, y, p:int31
H:[|p|] <= 31
H0:[|p|] = Z.of_nat (Z.abs_nat [|p|])
[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wB
x, y, p:int31
H:[|p|] <= 31

0 <= [|p|]
x, y, p:int31
H:[|p|] <= 31
H0:[|p|] = Z.of_nat (Z.abs_nat [|p|])
[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wB
x, y, p:int31
H:[|p|] <= 31
H0:[|p|] = Z.of_nat (Z.abs_nat [|p|])

[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wB
x, y, p:int31
H:Z.of_nat (Z.abs_nat [|p|]) <= 31

[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ Z.of_nat (Z.abs_nat [|p|]) + [|y|] / 2 ^ (31 - Z.of_nat (Z.abs_nat [|p|]))) mod wB
x, y, p:int31
n:nat
H:Z.of_nat n <= 31

[|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y, p:int31
n:nat
H:Z.of_nat n <= 31

(n <= 31)%nat
x, y, p:int31
n:nat
H:Z.of_nat n <= 31
H0:(n <= 31)%nat
[|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y, p:int31
n:nat
H:Z.of_nat n <= 31
H0:(n <= 31)%nat

[|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
n:nat
H0:(n <= 31)%nat

forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
H0:(0 <= 31)%nat

forall x y : int31, [|addmuldiv31_alt 0 x y|] = ([|x|] * 2 ^ Z.of_nat 0 + [|y|] / 2 ^ (31 - Z.of_nat 0)) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
forall x y : int31, [|addmuldiv31_alt (S n) x y|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
H0:(0 <= 31)%nat
x, y:int31

[|addmuldiv31_alt 0 x y|] = ([|x|] * 2 ^ 0 + [|y|] / 2 ^ (31 - 0)) mod 2 ^ 31
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
forall x y : int31, [|addmuldiv31_alt (S n) x y|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
H0:(0 <= 31)%nat
x, y:int31

[|addmuldiv31_alt 0 x y|] = ([|x|] + [|y|] / 2 ^ (31 - 0)) mod 2 ^ 31
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
forall x y : int31, [|addmuldiv31_alt (S n) x y|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
H0:(0 <= 31)%nat
x, y:int31

[|addmuldiv31_alt 0 x y|] = ([|x|] + 0) mod 2 ^ 31
H0:(0 <= 31)%nat
x, y:int31
0 = [|y|] / 2 ^ (31 - 0)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
forall x y : int31, [|addmuldiv31_alt (S n) x y|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
H0:(0 <= 31)%nat
x, y:int31

[|addmuldiv31_alt 0 x y|] = [|x|] mod 2 ^ 31
H0:(0 <= 31)%nat
x, y:int31
0 = [|y|] / 2 ^ (31 - 0)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
forall x y : int31, [|addmuldiv31_alt (S n) x y|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
H0:(0 <= 31)%nat
x, y:int31

0 = [|y|] / 2 ^ (31 - 0)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB
forall x y : int31, [|addmuldiv31_alt (S n) x y|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB

forall x y : int31, [|addmuldiv31_alt (S n) x y|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31

[|addmuldiv31_alt n (sneakl (firstl y) x) (shiftl y)|] = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31

([|sneakl (firstl y) x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

([|sneakl D0 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

((2 * [|x|]) mod wB * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

((2 * [|x|]) mod wB * 2 ^ Z.of_nat n + Z.double [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

((2 * [|x|]) mod wB * 2 ^ Z.of_nat n + 2 * [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

((2 * [|x|]) mod wB * 2 ^ Z.of_nat n + 2 * [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * (2 * 2 ^ Z.of_nat n) + [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

(2 * [|x|] * 2 ^ Z.of_nat n + 2 * [|y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * (2 * 2 ^ Z.of_nat n) + [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

2 * [|x|] * 2 ^ Z.of_nat n + 2 * [|y|] / 2 ^ (31 - Z.of_nat n) = [|x|] * (2 * 2 ^ Z.of_nat n) + [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

2 * [|x|] * 2 ^ Z.of_nat n = [|x|] * (2 * 2 ^ Z.of_nat n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0
2 * [|y|] / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

2 * [|y|] / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

2 * [|y|] / 2 ^ Z.succ (31 - Z.succ (Z.of_nat n)) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D0

2 * [|y|] / 2 / 2 ^ (31 - Z.succ (Z.of_nat n)) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

([|sneakl D1 x|] * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

((2 * [|x|] + 1) mod wB * 2 ^ Z.of_nat n + [|shiftl y|] / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

((2 * [|x|] + 1) mod wB * 2 ^ Z.of_nat n + Z.double [|y|] mod wB / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

((2 * [|x|] + 1) mod wB * 2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * 2 ^ Z.of_nat (S n) + [|y|] / 2 ^ (31 - Z.of_nat (S n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

((2 * [|x|] + 1) mod wB * 2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * (2 * 2 ^ Z.of_nat n) + [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

((2 * [|x|] + 1) * 2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n)) mod wB = ([|x|] * (2 * 2 ^ Z.of_nat n) + [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

(2 * [|x|] * 2 ^ Z.of_nat n + (2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n))) mod wB = ([|x|] * (2 * 2 ^ Z.of_nat n) + [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

2 * [|x|] * 2 ^ Z.of_nat n + (2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n)) = [|x|] * (2 * 2 ^ Z.of_nat n) + [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

2 * [|x|] * 2 ^ Z.of_nat n = [|x|] * (2 * 2 ^ Z.of_nat n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1

(2 * [|y|]) mod wB = 2 * [|y|] - wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1

(2 * [|y|]) mod wB = 2 * [|y|] - wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1

2 * [|y|] - wB = (2 * [|y|]) mod wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1

0 <= 2 * [|y|] - wB < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1

2 ^ Z.of_nat (Init.Nat.pred size) <= [|y|] -> 0 <= [|y|] < wB -> 0 <= 2 * [|y|] - wB < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1
wB':=2 ^ Z.of_nat (Init.Nat.pred size):Z

wB' <= [|y|] -> 0 <= [|y|] < wB -> 0 <= 2 * [|y|] - wB < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1
wB':=2 ^ Z.of_nat (Init.Nat.pred size):Z

2 * wB' = wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1
wB':=2 ^ Z.of_nat (Init.Nat.pred size):Z

2 * 2 ^ Z.of_nat (Init.Nat.pred size) = wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
y:int31
H:firstl y = D1
wB':=2 ^ Z.of_nat (Init.Nat.pred size):Z

2 ^ Z.of_nat (S (Init.Nat.pred size)) = wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB
2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 ^ Z.of_nat n + (2 * [|y|] - wB) / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 ^ Z.of_nat n + (2 * [|y|] - 2 ^ Z.of_nat n * 2 ^ (31 - Z.of_nat n)) / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 ^ Z.of_nat n + (2 * [|y|] + - 2 ^ Z.of_nat n * 2 ^ (31 + - Z.of_nat n)) / 2 ^ (31 + - Z.of_nat n) = [|y|] / 2 ^ (31 + - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 ^ Z.of_nat n + (2 * [|y|] / 2 ^ (31 + - Z.of_nat n) + - 2 ^ Z.of_nat n) = [|y|] / 2 ^ (31 + - Z.succ (Z.of_nat n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 * [|y|] / 2 ^ (31 + - Z.of_nat n) = [|y|] / 2 ^ (31 + - (Z.of_nat n + 1))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 * [|y|] / 2 ^ Z.succ (31 - Z.succ (Z.of_nat n)) = [|y|] / 2 ^ (31 + - (Z.of_nat n + 1))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 y0 : int31, [|addmuldiv31_alt n x0 y0|] = ([|x0|] * 2 ^ Z.of_nat n + [|y0|] / 2 ^ (31 - Z.of_nat n)) mod wB
x, y:int31
H:firstl y = D1
H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB

2 * [|y|] / 2 / 2 ^ (31 - Z.succ (Z.of_nat n)) = [|y|] / 2 ^ (31 + - (Z.of_nat n + 1))
rewrite Z.mul_comm, Z_div_mult; auto with zarith. Qed.

forall n p a : Z, 0 <= p <= n -> ((a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p)) mod 2 ^ n = a mod 2 ^ p

forall n p a : Z, 0 <= p <= n -> ((a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p)) mod 2 ^ n = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n

((a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p)) mod 2 ^ n = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n

(a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n

(a * 2 ^ (n - p) - a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n

(a * 2 ^ (n - p) + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n)) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n

a + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n

2 ^ n = 2 ^ (n - p) * 2 ^ p
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p
a + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n

2 ^ n = 2 ^ (n - p + p)
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p
a + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p

a + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p

a + - (a * 2 ^ (n - p) / (2 ^ (n - p) * 2 ^ p) * (2 ^ (n - p) * 2 ^ p)) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p

a + - (a / 2 ^ p * (2 ^ (n - p) * 2 ^ p)) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p

a + - (a / 2 ^ p * 2 ^ p * 2 ^ (n - p)) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p

a + - (a / 2 ^ p * 2 ^ p) * 2 ^ (n - p) / 2 ^ (n - p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
H0:2 ^ n = 2 ^ (n - p) * 2 ^ p

a + - (a / 2 ^ p * 2 ^ p) = a mod 2 ^ p
n, p, a:Z
H:0 <= p <= n
0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n

0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)

0 <= b mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n

0 <= b mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n

0 <= b mod 2 ^ n / 2 ^ (n - p)
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n
b mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n

b mod 2 ^ n / 2 ^ (n - p) < 2 ^ n
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n

b mod 2 ^ n < 2 ^ n * 2 ^ (n - p)
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n

2 ^ n <= 2 ^ n * 2 ^ (n - p)
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n

2 ^ n * 1 <= 2 ^ n * 2 ^ (n - p)
n, p, a:Z
H:0 <= p <= n
b:Z
Heqb:b = a * 2 ^ (n - p)
H0:0 <= b mod 2 ^ n
H1:b mod 2 ^ n < 2 ^ n

1 <= 2 ^ (n - p)
cut (0 < 2 ^ (n-p)); auto with zarith. Qed.

forall w p : int31, [|ZnZ.pos_mod p w|] = [|w|] mod 2 ^ [|p|]

forall w p : int31, [|ZnZ.pos_mod p w|] = [|w|] mod 2 ^ [|p|]

forall w p : int31, [|match [|p|] ?= [|31|] with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]

forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]

forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p
H:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p
forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
w:int31
p:Z
H:31 <= p

[|w|] = [|w|] mod 2 ^ p
H:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p
forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
w:int31
p:Z
H:31 <= p

0 <= [|w|] < wB -> [|w|] = [|w|] mod 2 ^ p
H:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p
forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
w:int31
p:Z
H:31 <= p
H0:0 <= [|w|] < wB

0 <= [|w|] < 2 ^ p
H:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p
forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
w:int31
p:Z
H:31 <= p
H0:0 <= [|w|] < wB

[|w|] < 2 ^ p
H:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p
forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
w:int31
p:Z
H:31 <= p
H0:0 <= [|w|] < wB

wB <= 2 ^ p
H:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p
forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
H:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ p

forall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31

[|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:([|p|] ?= 31) = Lt

[|addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0)|] = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31

[|addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0)|] = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31

([|0|] * 2 ^ [|p|] + [|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31

([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB

([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB

31 - [|p|] < wB
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB

31 < wB
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB

([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB

[|31 - p|] = 31 - [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
H4:[|31 - p|] = 31 - [|p|]
([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB

([|31|] - [|p|]) mod wB = 31 - [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
H4:[|31 - p|] = 31 - [|p|]
([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB

(31 - [|p|]) mod wB = 31 - [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
H4:[|31 - p|] = 31 - [|p|]
([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
H4:[|31 - p|] = 31 - [|p|]

([|addmuldiv31 (31 - p) w 0|] / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
H4:[|31 - p|] = 31 - [|p|]

(([|w|] * 2 ^ [|31 - p|] + [|0|] / 2 ^ (31 - [|31 - p|])) mod wB / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
H4:[|31 - p|] = 31 - [|p|]

(([|w|] * 2 ^ [|31 - p|]) mod wB / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0
w, p:int31
H0:[|p|] < 31
H1:0 <= [|p|] < wB
H2:0 <= [|w|] < wB
H3:31 - [|p|] < wB
H4:[|31 - p|] = 31 - [|p|]

(([|w|] * 2 ^ (31 - [|p|])) mod wB / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
apply shift_unshift_mod_2; simpl; auto with zarith. Qed.
Shift operations
 

forall x : int31, [|x|] = 0 -> [|head031 x|] = 31

forall x : int31, [|x|] = 0 -> [|head031 x|] = 31
x:int31
H:[|x|] = 0

[|head031 x|] = 31
x:int31
H:[|x|] = 0

phi_inv [|x|] = x -> [|head031 x|] = 31
x:int31
H:[|x|] = 0

On = x -> [|head031 x|] = 31
x:int31
H:[|x|] = 0
H':On = x

[|head031 On|] = 31
simpl; auto. Qed. Fixpoint head031_alt n x := match n with | O => 0%nat | S n => match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0%nat end end.

forall x : int31, [|head031 x|] = Z.of_nat (head031_alt size x)

forall x : int31, [|head031 x|] = Z.of_nat (head031_alt size x)
x:int31

[|head031 x|] = Z.of_nat (head031_alt size x)
x:int31
H:iszero x = true

[|head031 x|] = Z.of_nat (head031_alt size x)
x:int31
H:iszero x = false
[|head031 x|] = Z.of_nat (head031_alt size x)
x:int31
H:iszero x = true

[|head031 0|] = Z.of_nat (head031_alt size 0)
x:int31
H:iszero x = false
[|head031 x|] = Z.of_nat (head031_alt size x)
x:int31
H:iszero x = false

[|head031 x|] = Z.of_nat (head031_alt size x)
x:int31
H:iszero x = false

[|recl_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x On|] = Z.of_nat (head031_alt size x)
x:int31
H:iszero x = false

[|recl_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - size)))|] = Z.of_nat (head031_alt size x)
x:int31
H:iszero x = false

[|recl_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - size)))|] = Z.of_nat (head031_alt size x + (31 - size))
x:int31
H:iszero x = false
H0:(size <= 31)%nat

[|recl_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - size)))|] = Z.of_nat (head031_alt size x + (31 - size))
H0:(0 <= 31)%nat
x:int31
H:iszero x = false

[|recl_aux 0 (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - 0)))|] = Z.of_nat (head031_alt 0 x + (31 - 0))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
[|recl_aux (S n) (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (head031_alt (S n) x + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|recl_aux (S n) (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (head031_alt (S n) x + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|(if iszero x then fun _ : int31 => T31 else fun n0 : int31 => match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n1 : int31) => match b with | D0 => rec (n1 + In)%int31 | D1 => n1 end) (shiftl x) (n0 + In)%int31 | D1 => n0 end) (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (head031_alt (S n) x + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|(if iszero x then fun _ : int31 => T31 else fun n0 : int31 => match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n1 : int31) => match b with | D0 => rec (n1 + In)%int31 | D1 => n1 end) (shiftl x) (n0 + In)%int31 | D1 => n0 end) (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat (31 - S n) mod wB = Z.of_nat (31 - S n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

0 <= Z.of_nat (31 - S n) < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

0 <= Z.of_nat (31 - S n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
Z.of_nat (31 - S n) < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat (31 - S n) < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat (31 - S n) <= Z.of_nat 31
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
Z.of_nat 31 < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat 31 < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)

[|match firstl x with | D0 => recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

[|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = Z.of_nat (S (head031_alt n (shiftl x)) + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

[|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = Z.of_nat (head031_alt n (shiftl x) + S (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

[|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = Z.of_nat (head031_alt n (shiftl x) + (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

[|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftl x) (phi_inv (Z.of_nat (31 - n)))|]
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0
iszero (shiftl x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

(phi_inv (Z.of_nat (31 - S n)) + In)%int31 = phi_inv (Z.of_nat (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0
iszero (shiftl x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

phi_inv ([|phi_inv (Z.of_nat (31 - S n))|] + [|In|]) = phi_inv (Z.of_nat (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0
iszero (shiftl x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

phi_inv (Z.of_nat (31 - S n) + [|In|]) = phi_inv (Z.of_nat (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0
iszero (shiftl x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

Z.of_nat (31 - S n) + [|In|] = Z.of_nat (31 - n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0
iszero (shiftl x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

Z.of_nat (31 - S n) + 1 = Z.of_nat (31 - n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0
iszero (shiftl x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

Z.of_nat (31 - S n) + 1 = Z.of_nat (S (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0
iszero (shiftl x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recl_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (head031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstl x = D0

iszero (shiftl x) = false
x:int31
H:iszero x = false
H2:firstl x = D0

iszero (shiftl x) = false
x:int31
H:iszero (sneakr (firstl x) (shiftl x)) = false
H2:firstl x = D0

iszero (shiftl x) = false
x:int31
H:iszero (sneakr D0 (shiftl x)) = false
H2:firstl x = D0

iszero (shiftl x) = false
x:int31
H:iszero (sneakr D0 (shiftl x)) = false
H2:firstl x = D0
H0:iszero (shiftl x) = true

true = false
rewrite (iszero_eq0 _ H0) in H; discriminate. Qed.

forall x : int31, 0 < [|x|] <-> x <> 0%int31

forall x : int31, 0 < [|x|] <-> x <> 0%int31
x:int31
H:0 < [|x|]

x <> 0%int31
x:int31
H:x <> 0%int31
0 < [|x|]
x:int31
H:x <> 0%int31

0 < [|x|]
x:int31
H:x <> 0%int31

[|x|] <> 0
x:int31
H:x <> 0%int31
H0:[|x|] <> 0
0 < [|x|]
x:int31
H:[|x|] = 0

x = 0%int31
x:int31
H:x <> 0%int31
H0:[|x|] <> 0
0 < [|x|]
x:int31
H:x <> 0%int31
H0:[|x|] <> 0

0 < [|x|]
generalize (phi_bounded x); auto with zarith. Qed.

forall x : int31, 0 < [|x|] -> wB / 2 <= 2 ^ [|head031 x|] * [|x|] < wB

forall x : int31, 0 < [|x|] -> wB / 2 <= 2 ^ [|head031 x|] * [|x|] < wB
x:int31
H:0 < [|x|]

wB / 2 <= 2 ^ [|head031 x|] * [|x|] < wB
x:int31
H:0 < [|x|]

wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wB
x:int31
H:0 < [|x|]

nshiftl x size = 0%int31
x:int31
H:0 < [|x|]
H0:nshiftl x size = 0%int31
wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wB
x:int31
H:0 < [|x|]
H0:nshiftl x size = 0%int31

wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wB

forall x : int31, 0 < [|x|] -> nshiftl x size = 0%int31 -> wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wB

forall x : int31, 0 < [|x|] -> nshiftl x size = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < 2 ^ Z.of_nat 31

forall x : int31, 0 < [|x|] -> nshiftl x 0 = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt 0 x) * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftl x n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x) * [|x|] < 2 ^ Z.of_nat 31
forall x : int31, 0 < [|x|] -> nshiftl x (S n) = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt (S n) x) * [|x|] < 2 ^ Z.of_nat 31

forall x : int31, 0 < [|x|] -> nshiftl x 0 = 0%int31 -> 2 ^ 31 / 2 <= 2 ^ 0 * [|x|] < 2 ^ 31
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftl x n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x) * [|x|] < 2 ^ Z.of_nat 31
forall x : int31, 0 < [|x|] -> nshiftl x (S n) = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt (S n) x) * [|x|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x 0 = 0%int31

2 ^ 31 / 2 <= 2 ^ 0 * [|x|] < 2 ^ 31
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftl x n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x) * [|x|] < 2 ^ Z.of_nat 31
forall x : int31, 0 < [|x|] -> nshiftl x (S n) = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt (S n) x) * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftl x n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x) * [|x|] < 2 ^ Z.of_nat 31

forall x : int31, 0 < [|x|] -> nshiftl x (S n) = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt (S n) x) * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31

2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt (S n) x) * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31

2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat match firstl x with | D0 => S (head031_alt n (shiftl x)) | D1 => 0 end * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0

2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (S (head031_alt n (shiftl x))) * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0

2 ^ Z.of_nat 31 / 2 <= 2 * 2 ^ Z.of_nat (head031_alt n (shiftl x)) * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0

2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n (shiftl x)) * (2 * [|x|]) < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0

2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n (shiftl x)) * [|twice x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0

0 < [|twice x|]
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0
nshiftl (twice x) n = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0
H:twice x = 0%int31

x = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0
nshiftl (twice x) n = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0
H:shiftl x = 0%int31

x = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0
nshiftl (twice x) n = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D0

nshiftl (twice x) n = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1

2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1

2 ^ Z.of_nat 31 / 2 <= [|x|] < 2 ^ Z.of_nat 31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
H2:0 <= [|x|] < 2 ^ Z.of_nat 31

2 ^ Z.of_nat 31 / 2 <= [|x|]
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftl x0 n = 0%int31 -> 2 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n x0) * [|x0|] < 2 ^ Z.of_nat 31
x:int31
H:0 < [|x|]
H0:nshiftl x (S n) = 0%int31
H1:firstl x = D1
H2:0 <= [|x|] < 2 ^ Z.of_nat 31

2 ^ Z.of_nat (Init.Nat.pred size) <= [|x|]
apply phi_lowerbound; auto. Qed.

forall x : int31, [|x|] = 0 -> [|tail031 x|] = 31

forall x : int31, [|x|] = 0 -> [|tail031 x|] = 31
x:int31
H:[|x|] = 0

[|tail031 x|] = 31
x:int31
H:[|x|] = 0

phi_inv [|x|] = x -> [|tail031 x|] = 31
x:int31
H:[|x|] = 0

On = x -> [|tail031 x|] = 31
x:int31
H:[|x|] = 0
H':On = x

[|tail031 On|] = 31
simpl; auto. Qed. Fixpoint tail031_alt n x := match n with | O => 0%nat | S n => match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0%nat end end.

forall x : int31, [|tail031 x|] = Z.of_nat (tail031_alt size x)

forall x : int31, [|tail031 x|] = Z.of_nat (tail031_alt size x)
x:int31

[|tail031 x|] = Z.of_nat (tail031_alt size x)
x:int31
H:iszero x = true

[|tail031 x|] = Z.of_nat (tail031_alt size x)
x:int31
H:iszero x = false
[|tail031 x|] = Z.of_nat (tail031_alt size x)
x:int31
H:iszero x = true

[|tail031 0|] = Z.of_nat (tail031_alt size 0)
x:int31
H:iszero x = false
[|tail031 x|] = Z.of_nat (tail031_alt size x)
x:int31
H:iszero x = false

[|tail031 x|] = Z.of_nat (tail031_alt size x)
x:int31
H:iszero x = false

[|recr_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x On|] = Z.of_nat (tail031_alt size x)
x:int31
H:iszero x = false

[|recr_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - size)))|] = Z.of_nat (tail031_alt size x)
x:int31
H:iszero x = false

[|recr_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - size)))|] = Z.of_nat (tail031_alt size x + (31 - size))
x:int31
H:iszero x = false
H0:(size <= 31)%nat

[|recr_aux size (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - size)))|] = Z.of_nat (tail031_alt size x + (31 - size))
H0:(0 <= 31)%nat
x:int31
H:iszero x = false

[|recr_aux 0 (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n : int31) => match b with | D0 => rec (n + In)%int31 | D1 => n end) x (phi_inv (Z.of_nat (31 - 0)))|] = Z.of_nat (tail031_alt 0 x + (31 - 0))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
[|recr_aux (S n) (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (tail031_alt (S n) x + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|recr_aux (S n) (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (tail031_alt (S n) x + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|(if iszero x then fun _ : int31 => T31 else fun n0 : int31 => match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n1 : int31) => match b with | D0 => rec (n1 + In)%int31 | D1 => n1 end) (shiftr x) (n0 + In)%int31 | D1 => n0 end) (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (tail031_alt (S n) x + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|(if iszero x then fun _ : int31 => T31 else fun n0 : int31 => match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n1 : int31) => match b with | D0 => rec (n1 + In)%int31 | D1 => n1 end) (shiftr x) (n0 + In)%int31 | D1 => n0 end) (phi_inv (Z.of_nat (31 - S n)))|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat (31 - S n) mod wB = Z.of_nat (31 - S n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

0 <= Z.of_nat (31 - S n) < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

0 <= Z.of_nat (31 - S n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
Z.of_nat (31 - S n) < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat (31 - S n) < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat (31 - S n) <= Z.of_nat 31
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
Z.of_nat 31 < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false

Z.of_nat 31 < wB
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)

[|match firstr x with | D0 => recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31 | D1 => phi_inv (Z.of_nat (31 - S n)) end|] = Z.of_nat (match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

[|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = Z.of_nat (S (tail031_alt n (shiftr x)) + (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

[|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = Z.of_nat (tail031_alt n (shiftr x) + S (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

[|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = Z.of_nat (tail031_alt n (shiftr x) + (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

[|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - S n)) + In)%int31|] = [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) (shiftr x) (phi_inv (Z.of_nat (31 - n)))|]
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0
iszero (shiftr x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

(phi_inv (Z.of_nat (31 - S n)) + In)%int31 = phi_inv (Z.of_nat (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0
iszero (shiftr x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

phi_inv ([|phi_inv (Z.of_nat (31 - S n))|] + [|In|]) = phi_inv (Z.of_nat (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0
iszero (shiftr x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

phi_inv (Z.of_nat (31 - S n) + [|In|]) = phi_inv (Z.of_nat (31 - n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0
iszero (shiftr x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

Z.of_nat (31 - S n) + [|In|] = Z.of_nat (31 - n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0
iszero (shiftr x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

Z.of_nat (31 - S n) + 1 = Z.of_nat (31 - n)
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0
iszero (shiftr x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

Z.of_nat (31 - S n) + 1 = Z.of_nat (S (31 - S n))
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0
iszero (shiftr x) = false
n:nat
H0:(S n <= 31)%nat
IHn:(n <= 31)%nat -> forall x0 : int31, iszero x0 = false -> [|recr_aux n (int31 -> int31) (fun _ : int31 => T31) (fun (b : digits) (_ : int31) (rec : int31 -> int31) (n0 : int31) => match b with | D0 => rec (n0 + In)%int31 | D1 => n0 end) x0 (phi_inv (Z.of_nat (31 - n)))|] = Z.of_nat (tail031_alt n x0 + (31 - n))
x:int31
H:iszero x = false
H1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)
H2:firstr x = D0

iszero (shiftr x) = false
x:int31
H:iszero x = false
H2:firstr x = D0

iszero (shiftr x) = false
x:int31
H:iszero (sneakl (firstr x) (shiftr x)) = false
H2:firstr x = D0

iszero (shiftr x) = false
x:int31
H:iszero (sneakl D0 (shiftr x)) = false
H2:firstr x = D0

iszero (shiftr x) = false
x:int31
H:iszero (sneakl D0 (shiftr x)) = false
H2:firstr x = D0
H0:iszero (shiftr x) = true

true = false
rewrite (iszero_eq0 _ H0) in H; discriminate. Qed.

forall x : int31, 0 < [|x|] -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail031 x|]

forall x : int31, 0 < [|x|] -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail031 x|]
x:int31
H:0 < [|x|]

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail031 x|]
x:int31
H:0 < [|x|]

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt size x)
x:int31
H:0 < [|x|]

nshiftr x size = 0%int31
x:int31
H:0 < [|x|]
H0:nshiftr x size = 0%int31
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt size x)
x:int31
H:0 < [|x|]
H0:nshiftr x size = 0%int31

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt size x)

forall x : int31, 0 < [|x|] -> nshiftr x size = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt size x)

forall x : int31, 0 < [|x|] -> nshiftr x 0 = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt 0 x)
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftr x n = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x)
forall x : int31, 0 < [|x|] -> nshiftr x (S n) = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt (S n) x)

forall x : int31, 0 < [|x|] -> nshiftr x 0 = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ 0
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftr x n = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x)
forall x : int31, 0 < [|x|] -> nshiftr x (S n) = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt (S n) x)
x:int31
H:0 < [|x|]
H0:nshiftr x 0 = 0%int31

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ 0
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftr x n = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x)
forall x : int31, 0 < [|x|] -> nshiftr x (S n) = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt (S n) x)
n:nat
IHn:forall x : int31, 0 < [|x|] -> nshiftr x n = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x)

forall x : int31, 0 < [|x|] -> nshiftr x (S n) = 0%int31 -> exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt (S n) x)
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt (S n) x)
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat match firstr x with | D0 => S (tail031_alt n (shiftr x)) | D1 => 0 end
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (S (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0

0 < [|shiftr x|]
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
nshiftr (shiftr x) n = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y0 : Z, 0 <= y0 /\ [|x0|] = (2 * y0 + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
y:Z
Hy1:0 <= y
Hy2:[|shiftr x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n (shiftr x))
exists y0 : Z, 0 <= y0 /\ [|x|] = (2 * y0 + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
H:shiftr x = 0%int31

x = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
nshiftr (shiftr x) n = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y0 : Z, 0 <= y0 /\ [|x0|] = (2 * y0 + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
y:Z
Hy1:0 <= y
Hy2:[|shiftr x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n (shiftr x))
exists y0 : Z, 0 <= y0 /\ [|x|] = (2 * y0 + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0

nshiftr (shiftr x) n = 0%int31
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y0 : Z, 0 <= y0 /\ [|x0|] = (2 * y0 + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
y:Z
Hy1:0 <= y
Hy2:[|shiftr x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n (shiftr x))
exists y0 : Z, 0 <= y0 /\ [|x|] = (2 * y0 + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y0 : Z, 0 <= y0 /\ [|x0|] = (2 * y0 + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
y:Z
Hy1:0 <= y
Hy2:[|shiftr x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n (shiftr x))

exists y0 : Z, 0 <= y0 /\ [|x|] = (2 * y0 + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y0 : Z, 0 <= y0 /\ [|x0|] = (2 * y0 + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
y:Z
Hy1:0 <= y
Hy2:[|shiftr x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n (shiftr x))

[|x|] = (2 * y + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y0 : Z, 0 <= y0 /\ [|x0|] = (2 * y0 + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D0
y:Z
Hy1:0 <= y
Hy2:[|shiftr x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n (shiftr x))

Z.double [|shiftr x|] = (2 * y + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1

exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1

0 <= [|shiftr x|] /\ [|x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1

0 <= [|shiftr x|]
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1
[|x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1

[|x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0
n:nat
IHn:forall x0 : int31, 0 < [|x0|] -> nshiftr x0 n = 0%int31 -> exists y : Z, 0 <= y /\ [|x0|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt n x0)
x:int31
H:0 < [|x|]
H0:nshiftr x (S n) = 0%int31
H1:firstr x = D1

Z.succ_double [|shiftr x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0
rewrite Z.succ_double_spec; simpl; ring. Qed. (* Sqrt *) (* Direct transcription of an old proof of a fortran program in boyer-moore *)
a:Z

a - 1 <= a / 2 + a / 2
a:Z

a - 1 <= a / 2 + a / 2
a:Z

0 <= a mod 2 -> a mod 2 < 2 -> a - 1 <= a / 2 + a / 2
intros H1; rewrite Zmod_eq_full; auto with zarith. Qed.
j, k:Z

0 <= j -> 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2
j, k:Z

0 <= j -> 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2

0 <= 0 -> forall k : Z, 0 <= k -> 0 * k + 0 <= ((0 + k) / 2 + 1) ^ 2

forall x : Z, 0 <= x -> (0 <= x -> forall k : Z, 0 <= k -> x * k + x <= ((x + k) / 2 + 1) ^ 2) -> 0 <= Z.succ x -> forall k : Z, 0 <= k -> Z.succ x * k + Z.succ x <= ((Z.succ x + k) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

0 <= (k / 2 + 1) ^ 2

forall x : Z, 0 <= x -> (0 <= x -> forall k : Z, 0 <= k -> x * k + x <= ((x + k) / 2 + 1) ^ 2) -> 0 <= Z.succ x -> forall k : Z, 0 <= k -> Z.succ x * k + Z.succ x <= ((Z.succ x + k) / 2 + 1) ^ 2

forall x : Z, 0 <= x -> (0 <= x -> forall k : Z, 0 <= k -> x * k + x <= ((x + k) / 2 + 1) ^ 2) -> 0 <= Z.succ x -> forall k : Z, 0 <= k -> Z.succ x * k + Z.succ x <= ((Z.succ x + k) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2

Z.succ j * 0 + Z.succ j <= ((Z.succ j + 0) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2
forall x : Z, 0 <= x -> Z.succ j * x + Z.succ j <= ((Z.succ j + x) / 2 + 1) ^ 2 -> Z.succ j * Z.succ x + Z.succ j <= ((Z.succ j + Z.succ x) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2

Z.succ j <= (Z.succ j / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2
forall x : Z, 0 <= x -> Z.succ j * x + Z.succ j <= ((Z.succ j + x) / 2 + 1) ^ 2 -> Z.succ j * Z.succ x + Z.succ j <= ((Z.succ j + Z.succ x) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2

(j + 1) / 2 * ((j + 1) / 2) >= 0 -> j + 1 - 1 <= (j + 1) / 2 + (j + 1) / 2 -> j + 1 <= ((j + 1) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2
forall x : Z, 0 <= x -> Z.succ j * x + Z.succ j <= ((Z.succ j + x) / 2 + 1) ^ 2 -> Z.succ j * Z.succ x + Z.succ j <= ((Z.succ j + Z.succ x) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2

(j + 1) / 2 * ((j + 1) / 2) >= 0 -> j + 1 - 1 <= (j + 1) / 2 + (j + 1) / 2 -> j + 1 <= (j + 1) / 2 * ((j + 1) / 2) + (j + 1) / 2 * 1 + (1 * ((j + 1) / 2) + 1 * 1)
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2
forall x : Z, 0 <= x -> Z.succ j * x + Z.succ j <= ((Z.succ j + x) / 2 + 1) ^ 2 -> Z.succ j * Z.succ x + Z.succ j <= ((Z.succ j + Z.succ x) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2

forall x : Z, 0 <= x -> Z.succ j * x + Z.succ j <= ((Z.succ j + x) / 2 + 1) ^ 2 -> Z.succ j * Z.succ x + Z.succ j <= ((Z.succ j + Z.succ x) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

Z.succ j * Z.succ k + Z.succ j <= ((Z.succ j + Z.succ k) / 2 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

Z.succ j * Z.succ k + Z.succ j <= ((j + k) / 2 + 1 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k
(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

j * k + j <= ((j + k) / 2 + 1) ^ 2 -> j + k - 1 <= (j + k) / 2 + (j + k) / 2 -> Z.succ j * Z.succ k + Z.succ j <= ((j + k) / 2 + 1 + 1) ^ 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k
(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

j * k + j <= (j + k) / 2 * ((j + k) / 2) + (j + k) / 2 * 1 + (1 * ((j + k) / 2) + 1 * 1) -> j + k - 1 <= (j + k) / 2 + (j + k) / 2 -> j * k + j * 1 + (1 * k + 1 * 1) + (j + 1) <= (j + k) / 2 * ((j + k) / 2) + (j + k) / 2 * 1 + (j + k) / 2 * 1 + (1 * ((j + k) / 2) + 1 * 1 + 1 * 1) + (1 * ((j + k) / 2) + 1 * 1 + 1 * 1)
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k
(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

j * k + j <= (j + k) / 2 * ((j + k) / 2) + (j + k) / 2 + ((j + k) / 2 + 1) -> j + k - 1 <= (j + k) / 2 + (j + k) / 2 -> j * k + j + (k + 1) + (j + 1) <= (j + k) / 2 * ((j + k) / 2) + (j + k) / 2 + (j + k) / 2 + ((j + k) / 2 + 1 + 1) + ((j + k) / 2 + 1 + 1)
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k
(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2
j:Z
Hj:0 <= j
Hrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2
k:Z
Hk:0 <= k

(1 * 2 + (j + k)) / 2 = (Z.succ j + Z.succ k) / 2
apply f_equal2 with (f := Z.div); auto with zarith. Qed.
i, j:Z

0 <= i -> 0 < j -> i < ((j + i / j) / 2 + 1) ^ 2
i, j:Z

0 <= i -> 0 < j -> i < ((j + i / j) / 2 + 1) ^ 2
i, j:Z
Hi:0 <= i
Hj:0 < j

i < ((j + i / j) / 2 + 1) ^ 2
i, j:Z
Hi:0 <= i
Hj:0 < j
Hij:0 <= i / j

i < ((j + i / j) / 2 + 1) ^ 2
i, j:Z
Hi:0 <= i
Hj:0 < j
Hij:0 <= i / j

i < j * (i / j) + j
pattern i at 1; rewrite (Z_div_mod_eq i j); case (Z_mod_lt i j); auto with zarith. Qed.
i:Z

1 < i -> i < (i / 2 + 1) ^ 2
i:Z

1 < i -> i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i

i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2

i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2

1 <= (i / 2) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= (i / 2) ^ 2
i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2

1 <= ((1 * 2 + (i - 2)) / 2) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= (i / 2) ^ 2
i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2

1 <= (1 + (i - 2) / 2) * (1 + (i - 2) / 2)
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= (i / 2) ^ 2
i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2

(i - 2) / 2 * ((i - 2) / 2) >= 0 -> (2 > 0 -> 0 <= i - 2 -> 0 <= (i - 2) / 2) -> 1 <= (1 + (i - 2) / 2) * (1 + (i - 2) / 2)
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= (i / 2) ^ 2
i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2

(i - 2) / 2 * ((i - 2) / 2) >= 0 -> (2 > 0 -> 0 <= i - 2 -> 0 <= (i - 2) / 2) -> 1 <= 1 * 1 + 1 * ((i - 2) / 2) + ((i - 2) / 2 * 1 + (i - 2) / 2 * ((i - 2) / 2))
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= (i / 2) ^ 2
i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= (i / 2) ^ 2

i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= (i / 2) ^ 2

i - 1 <= i / 2 + i / 2 -> i < (i / 2 + 1) ^ 2
i:Z
Hi:1 < i
H1:0 <= i - 2
H2:1 <= i / 2 * (i / 2)

i - 1 <= i / 2 + i / 2 -> i < i / 2 * (i / 2) + i / 2 + (i / 2 + 1)
auto with zarith. Qed.
i, j:Z

0 <= i -> 0 < j -> i / j >= j -> j ^ 2 <= i
i, j:Z

0 <= i -> 0 < j -> i / j >= j -> j ^ 2 <= i
i, j:Z
Hi:0 <= i
Hj:0 < j
Hd:i / j >= j

j * j <= i
i, j:Z
Hi:0 <= i
Hj:0 < j
Hd:i / j >= j

j * (i / j) <= i
apply Z_mult_div_ge; auto with zarith. Qed.
i, j:Z

0 <= i -> 0 < j -> i / j < j -> (j + i / j) / 2 < j
i, j:Z

0 <= i -> 0 < j -> i / j < j -> (j + i / j) / 2 < j
i, j:Z
Hi:0 <= i
Hj:0 < j
H:i / j < j

j <= (j + i / j) / 2 -> (j + i / j) / 2 < j
i, j:Z
Hi:0 <= i
Hj:0 < j
H1:j <= (j + i / j) / 2

j <= i / j
i, j:Z
Hi:0 <= i
Hj:0 < j
H1:j <= (j + i / j) / 2

2 * j <= j + i / j
i, j:Z
Hi:0 <= i
Hj:0 < j
H1:j <= (j + i / j) / 2

2 * ((j + i / j) / 2) <= j + i / j
apply Z_mult_div_ge; auto with zarith. Qed.
rec:int31 -> int31 -> int31
i, j:int31

sqrt31_step rec i j = match (fst (i / j) ?= j)%int31 with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j end
rec:int31 -> int31 -> int31
i, j:int31

sqrt31_step rec i j = match (fst (i / j) ?= j)%int31 with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j end
rec:int31 -> int31 -> int31
i, j, i0, i1:int31

match (i0 ?= j)%int31 with | Lt => rec i (fst ((j + i0) / 2)%int31) | _ => j end = match (fst (i0, i1) ?= j)%int31 with | Lt => rec i (fst ((j + fst (i0, i1)) / 2)%int31) | _ => j end
simpl; case compare31; auto. Qed.
i, j:int31

0 < [|j|] -> [|fst (i / j)%int31|] = [|i|] / [|j|]
i, j:int31
Hj:0 < [|j|]

(let (q, r) := (i / j)%int31 in [|i|] = [|q|] * [|j|] + [|r|] /\ 0 <= [|r|] < [|j|]) -> [|fst (i / j)%int31|] = [|i|] / [|j|]
i, j:int31
Hj:0 < [|j|]
q, r:int31

[|i|] = [|q|] * [|j|] + [|r|] /\ 0 <= [|r|] < [|j|] -> [|q|] = [|i|] / [|j|]
i, j:int31
Hj:0 < [|j|]
q, r:int31
H1:[|i|] = [|q|] * [|j|] + [|r|]
H2:0 <= [|r|] < [|j|]

[|i|] = [|j|] * [|q|] + [|r|]
rewrite H1; ring. Qed.
rec:int31 -> int31 -> int31
i, j:int31

0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31

0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]

0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2

[|match (fst (i / j) ?= j)%int31 with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j end|] ^ 2 <= [|i|] < ([|match (fst (i / j) ?= j)%int31 with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2

[|match [|i|] / [|j|] ?= [|j|] with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j end|] ^ 2 <= [|i|] < ([|match [|i|] / [|j|] ?= [|j|] with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]

[|rec i (fst ((j + fst (i / j)) / 2)%int31)|] ^ 2 <= [|i|] < ([|rec i (fst ((j + fst (i / j)) / 2)%int31)|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]

[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]
[|rec i (fst ((j + fst (i / j)) / 2)%int31)|] ^ 2 <= [|i|] < ([|rec i (fst ((j + fst (i / j)) / 2)%int31)|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]

[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]
rewrite spec_add, div31_phi; auto using Z.mod_small with zarith.
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

[|rec i (fst ((j + fst (i / j)) / 2)%int31)|] ^ 2 <= [|i|] < ([|rec i (fst ((j + fst (i / j)) / 2)%int31)|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < ([|j|] + [|i|] / [|j|]) / [|2|] < [|j|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < ([|j|] + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:Z.succ 0 <= [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < ([|j|] + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 <= [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < ([|j|] + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < ([|j|] + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 = [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]
0 < ([|j|] + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < ([|j|] + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < (1 * 2 + ([|j|] - 2 + [|i|] / [|j|])) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < 1 + ([|j|] - 2 + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]
H:0 <= [|i|] / [|j|]

0 < 1 + ([|j|] - 2 + [|i|] / [|j|]) / [|2|]
assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / [|2|]); auto with zarith.
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 = [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < ([|j|] + [|i|] / [|j|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 = [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < (1 + [|i|]) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 = [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < (1 * 2 + ([|i|] - 1)) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 = [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]

0 < 1 + ([|i|] - 1) / [|2|]
rec:int31 -> int31 -> int31
i, j:int31
Hp2:0 < [|2|]
Hi:0 < [|i|]
Hj:1 = [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2
Hc:[|i|] / [|j|] < [|j|]
E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]
H:0 <= ([|i|] - 1) / 2

0 < 1 + ([|i|] - 1) / [|2|]
change ([|2|]) with 2; auto with zarith. Qed.
n:nat
rec:int31 -> int31 -> int31
i, j:int31

0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2
n:nat
rec:int31 -> int31 -> int31
i, j:int31

0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2

forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 0 + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step rec i j|] ^ 2 <= [|i|] < ([|sqrt31_step rec i j|] + 1) ^ 2

forall n : nat, (forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] ^ 2 <= [|i|] < ([|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 0 + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2

forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2

forall n : nat, (forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] ^ 2 <= [|i|] < ([|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
Hrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat 0 + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
H:0 < [|j1|] < [|j|]
H0:[|i|] < ([|j1|] + 1) ^ 2

2 ^ Z.of_nat 0 + [|j1|] <= [|j|]

forall n : nat, (forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] ^ 2 <= [|i|] < ([|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] + 1) ^ 2

forall n : nat, (forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|iter31_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31) (i j : int31), 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> [|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] ^ 2 <= [|i|] < ([|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j0|] -> [|i0|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec0 i0 j1|] ^ 2 <= [|i0|] < ([|rec0 i0 j1|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2

[|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] ^ 2 <= [|i|] < ([|sqrt31_step (iter31_sqrt n (iter31_sqrt n rec)) i j|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j0|] -> [|i0|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec0 i0 j1|] ^ 2 <= [|i0|] < ([|rec0 i0 j1|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2

forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|iter31_sqrt n (iter31_sqrt n rec) i j1|] ^ 2 <= [|i|] < ([|iter31_sqrt n (iter31_sqrt n rec) i j1|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j2 : int31, 0 < [|j2|] -> 2 ^ Z.of_nat n + [|j2|] <= [|j0|] -> [|i0|] < ([|j2|] + 1) ^ 2 -> 2 * [|j2|] < wB -> [|rec0 i0 j2|] ^ 2 <= [|i0|] < ([|rec0 i0 j2|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:[|i|] < ([|j1|] + 1) ^ 2

forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat n + [|j0|] <= [|j1|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|iter31_sqrt n rec i j0|] ^ 2 <= [|i|] < ([|iter31_sqrt n rec i j0|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j3 : int31, 0 < [|j3|] -> 2 ^ Z.of_nat n + [|j3|] <= [|j0|] -> [|i0|] < ([|j3|] + 1) ^ 2 -> 2 * [|j3|] < wB -> [|rec0 i0 j3|] ^ 2 <= [|i0|] < ([|rec0 i0 j3|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:[|i|] < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:[|i|] < ([|j2|] + 1) ^ 2
Hj31:2 * [|j2|] < wB

forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat n + [|j0|] <= [|j2|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> [|i0|] < ([|j4|] + 1) ^ 2 -> 2 * [|j4|] < wB -> [|rec0 i0 j4|] ^ 2 <= [|i0|] < ([|rec0 i0 j4|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:[|i|] < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:[|i|] < ([|j2|] + 1) ^ 2
Hj31:2 * [|j2|] < wB
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

[|i|] < ([|j3|] + 1) ^ 2 -> 2 * [|j3|] < wB -> [|rec i j3|] ^ 2 <= [|i|] < ([|rec i j3|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> [|i0|] < ([|j4|] + 1) ^ 2 -> 2 * [|j4|] < wB -> [|rec0 i0 j4|] ^ 2 <= [|i0|] < ([|rec0 i0 j4|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:[|i|] < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:[|i|] < ([|j2|] + 1) ^ 2
Hj31:2 * [|j2|] < wB
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

2 ^ Z.of_nat (S n) + [|j3|] <= [|j|]
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> [|i0|] < ([|j4|] + 1) ^ 2 -> 2 * [|j4|] < wB -> [|rec0 i0 j4|] ^ 2 <= [|i0|] < ([|rec0 i0 j4|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:[|i|] < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:[|i|] < ([|j2|] + 1) ^ 2
Hj31:2 * [|j2|] < wB
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

2 * 2 ^ Z.of_nat n + [|j3|] <= [|j|]
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> [|i0|] < ([|j4|] + 1) ^ 2 -> 2 * [|j4|] < wB -> [|rec0 i0 j4|] ^ 2 <= [|i0|] < ([|rec0 i0 j4|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:[|i|] < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:[|i|] < ([|j2|] + 1) ^ 2
Hj31:2 * [|j2|] < wB
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]
0 <= Z.of_nat n
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31) (i0 j0 : int31), 0 < [|i0|] -> 0 < [|j0|] -> [|i0|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> [|i0|] < ([|j4|] + 1) ^ 2 -> 2 * [|j4|] < wB -> [|rec0 i0 j4|] ^ 2 <= [|i0|] < ([|rec0 i0 j4|] + 1) ^ 2) -> [|iter31_sqrt n rec0 i0 j0|] ^ 2 <= [|i0|] < ([|iter31_sqrt n rec0 i0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31
i, j:int31
Hi:0 < [|i|]
Hj:0 < [|j|]
Hij:[|i|] < ([|j|] + 1) ^ 2
H31:2 * [|j|] < wB
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> [|i|] < ([|j0|] + 1) ^ 2 -> 2 * [|j0|] < wB -> [|rec i j0|] ^ 2 <= [|i|] < ([|rec i j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:[|i|] < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:[|i|] < ([|j2|] + 1) ^ 2
Hj31:2 * [|j2|] < wB
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

0 <= Z.of_nat n
apply Nat2Z.is_nonneg. Qed.

forall x : int31, [|sqrt31 x|] ^ 2 <= [|x|] < ([|sqrt31 x|] + 1) ^ 2

forall x : int31, [|sqrt31 x|] ^ 2 <= [|x|] < ([|sqrt31 x|] + 1) ^ 2
i:int31

[|match (1 ?= i)%int31 with | Eq => 1 | Lt => iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31) | Gt => 0 end|] ^ 2 <= [|i|] < ([|match (1 ?= i)%int31 with | Eq => 1 | Lt => iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31) | Gt => 0 end|] + 1) ^ 2
i:int31

[|match [|1|] ?= [|i|] with | Eq => 1 | Lt => iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31) | Gt => 0 end|] ^ 2 <= [|i|] < ([|match [|1|] ?= [|i|] with | Eq => 1 | Lt => iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31) | Gt => 0 end|] + 1) ^ 2
i:int31
Hi:1 = [|i|]

1 ^ 2 <= [|i|] < (1 + 1) ^ 2
i:int31
Hi:1 < [|i|]
[|iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31)|] ^ 2 <= [|i|] < ([|iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31)|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

[|iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31)|] ^ 2 <= [|i|] < ([|iter31_sqrt 31 (fun _ j : int31 => j) i (fst (i / 2)%int31)|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

0 < [|fst (i / 2)%int31|]
i:int31
Hi:1 < [|i|]
[|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
2 * [|fst (i / 2)%int31|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

0 < [|i|] / 2
i:int31
Hi:1 < [|i|]
[|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
2 * [|fst (i / 2)%int31|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

0 < (1 * 2 + ([|i|] - 2)) / 2
i:int31
Hi:1 < [|i|]
[|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
2 * [|fst (i / 2)%int31|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
H:0 <= ([|i|] - 2) / 2

0 < (1 * 2 + ([|i|] - 2)) / 2
i:int31
Hi:1 < [|i|]
[|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
2 * [|fst (i / 2)%int31|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

[|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
2 * [|fst (i / 2)%int31|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

[|i|] < ([|i|] / 2 + 1) ^ 2
i:int31
Hi:1 < [|i|]
2 * [|fst (i / 2)%int31|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

2 * [|fst (i / 2)%int31|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

2 * ([|i|] / 2) < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

2 * ([|i|] / 2) <= [|i|]
i:int31
Hi:1 < [|i|]
[|i|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

[|i|] < wB
i:int31
Hi:1 < [|i|]
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]

forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|fst (i / 2)%int31|] -> [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> [|j1|] ^ 2 <= [|i|] < ([|j1|] + 1) ^ 2
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]

[|fst (i / 2)%int31|] < 2 ^ Z.of_nat 31 + [|j2|]
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]

[|i|] / 2 < 2 ^ Z.of_nat 31 + [|j2|]
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]

[|i|] / 2 <= [|i|]
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]
[|i|] < 2 ^ Z.of_nat 31 + [|j2|]
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]
H:0 <= [|i|] / 2

[|i|] / 2 <= [|i|]
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]
[|i|] < 2 ^ Z.of_nat 31 + [|j2|]
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]
H:0 <= [|i|] / 2

2 * ([|i|] / 2) <= [|i|]
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]
[|i|] < 2 ^ Z.of_nat 31 + [|j2|]
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:1 < [|i|]
j2:int31
H1:0 < [|j2|]

[|i|] < 2 ^ Z.of_nat 31 + [|j2|]
i:int31
Hi:[|i|] < 1
[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:[|i|] < 1

[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2
i:int31
Hi:[|i|] < 1

0 ^ 2 <= [|i|] < (0 + 1) ^ 2
case (phi_bounded i); repeat rewrite Z.pow_2_r; auto with zarith. Qed.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31

sqrt312_step rec ih il j = match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31

sqrt312_step rec ih il j = match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end
rec:int31 -> int31 -> int31 -> int31
ih, il, j, i, i0:int31

match (ih ?= j)%int31 with | Lt => match (i ?= j)%int31 with | Lt => rec ih il match j +c i with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + 1073741824)%int31 end | _ => j end | _ => j end = match (ih ?= j)%int31 with | Lt => match (fst (i, i0) ?= j)%int31 with | Lt => rec ih il match j +c fst (i, i0) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end | _ => j end | _ => j end
simpl; case compare31; auto. Qed.
ih, il, j:int31

phi2 ih il < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|]
ih, il, j:int31

phi2 ih il < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|]
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2

[|ih|] <= [|j|]
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]

[|ih|] <= [|j|]
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]

[|ih|] <= [|j|]
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] <= [|j|]
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] < [|j|] + 1
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] * [|ih|] < ([|j|] + 1) * ([|j|] + 1)
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] ^ 2 <= phi2 ih il
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] ^ 2 <= [|ih|] * wB
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB
[|ih|] * wB <= phi2 ih il
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] ^ 2 <= [|ih|] * wB
rewrite ? Z.pow_2_r; auto with zarith.
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] * wB <= phi2 ih il
ih, il, j:int31
H1:phi2 ih il < ([|j|] + 1) ^ 2
Hbj:0 <= [|j|]
Hbil:0 <= [|il|]
Hbih:0 <= [|ih|]
Hbih1:[|ih|] < wB

[|ih|] * wB <= [|ih|] * base + [|il|]
change base with wB; auto with zarith. Qed.
ih, il, j:int31

2 ^ 30 <= [|j|] -> [|ih|] < [|j|] -> [|fst (div3121 ih il j)|] = phi2 ih il / [|j|]
ih, il, j:int31

2 ^ 30 <= [|j|] -> [|ih|] < [|j|] -> [|fst (div3121 ih il j)|] = phi2 ih il / [|j|]
ih, il, j:int31
Hj:2 ^ 30 <= [|j|]
Hj1:[|ih|] < [|j|]

[|fst (div3121 ih il j)|] = phi2 ih il / [|j|]
ih, il, j:int31
Hj:2 ^ 30 <= [|j|]
Hj1:[|ih|] < [|j|]

(let (q, r) := div3121 ih il j in [|ih|] * wB + [|il|] = [|q|] * [|j|] + [|r|] /\ 0 <= [|r|] < [|j|]) -> [|fst (div3121 ih il j)|] = phi2 ih il / [|j|]
ih, il, j:int31
Hj:2 ^ 30 <= [|j|]
Hj1:[|ih|] < [|j|]
q, r:int31
Hq:[|ih|] * wB + [|il|] = [|q|] * [|j|] + [|r|]
Hr:0 <= [|r|] < [|j|]

[|fst (q, r)|] = phi2 ih il / [|j|]
ih, il, j:int31
Hj:2 ^ 30 <= [|j|]
Hj1:[|ih|] < [|j|]
q, r:int31
Hq:[|ih|] * wB + [|il|] = [|q|] * [|j|] + [|r|]
Hr:0 <= [|r|] < [|j|]

phi2 ih il = [|j|] * [|fst (q, r)|] + [|r|]
simpl @fst; apply eq_trans with (1 := Hq); ring. Qed.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31

2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step rec ih il j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31

2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step rec ih il j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]

2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step rec ih il j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2

[|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]

[|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]

[|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]

[|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB

[|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB

0 < phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
[|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB

0 < phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB

0 < [|ih|] * base
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB

0 < [|ih|]
apply Z.lt_le_trans with (2:= Hih); auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il

[|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match (ih ?= j)%int31 with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il

[|match [|ih|] ?= [|j|] with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] ^ 2 <= phi2 ih il < ([|match [|ih|] ?= [|j|] with | Lt => match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]

[|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
[|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] ^ 2 <= phi2 ih il < ([|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]
[|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]

[|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]

[|j|] ^ 2 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]

0 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]
phi2 ih il / [|j|] >= [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]

0 <= phi2 ih il
unfold phi2, base; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]

phi2 ih il / [|j|] >= [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]

([|j|] * base + [|il|]) / [|j|] >= [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]
H:0 <= [|il|] / [|j|]

([|j|] * base + [|il|]) / [|j|] >= [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]
H:0 <= [|il|] / [|j|]

base + [|il|] / [|j|] >= [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] = [|j|]
H:0 <= [|il|] / [|j|]

wB + [|il|] / [|j|] >= [|j|]
auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]

[|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] ^ 2 <= phi2 ih il < ([|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]

[|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] ^ 2 <= phi2 ih il < ([|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
[|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] ^ 2 <= phi2 ih il < ([|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]

[|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] ^ 2 <= phi2 ih il < ([|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]

[|(let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m)|] ^ 2 <= phi2 ih il < ([|(let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m)|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
phi2 ih il < ([|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:Z.succ 0 <= [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 <= [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < ([|j|] + phi2 ih il / [|j|]) / 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < ([|j|] + phi2 ih il / [|j|]) / 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < (1 * 2 + ([|j|] - 2 + phi2 ih il / [|j|])) / 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]

0 < 1 + ([|j|] - 2 + phi2 ih il / [|j|]) / 2
assert (0 <= ([|j|] - 2 + phi2 ih il / [|j|]) / 2) ; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2

([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2

([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
apply sqrt_test_false; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]

0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]

[+|j +c fst (div3121 ih il j)|] = [|j|] + [|fst (div3121 ih il j)|] -> 0 < [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31

[|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31
1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31 + v30|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31

[|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31

[|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|r|] / 2 < [|j|]
intros HH; rewrite HH; clear HH; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31

1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31 + v30|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31

1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31 + v30|] < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31

1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < ([|r|] / 2 + [|v30|]) mod wB < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31
HH:wB + [|r|] = [|j|] + phi2 ih il / [|j|]

0 < ([|r|] / 2 + [|v30|]) mod wB < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31
HH:wB + [|r|] = [|j|] + phi2 ih il / [|j|]

0 < (([|v30|] * 2 + [|r|]) / 2) mod wB < [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:1 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
Hf1:0 <= phi2 ih il / [|j|]
Hf3:0 < ([|j|] + phi2 ih il / [|j|]) / 2
Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]
r:int31
HH:wB + [|r|] = [|j|] + phi2 ih il / [|j|]

0 < ((wB + [|r|]) / 2) mod wB < [|j|]
rewrite HH, Zmod_small; auto with zarith. }
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]

phi2 ih il < ([|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]

([|j|] + phi2 ih il / [|j|]) / 2 = [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]

[+|j +c fst (div3121 ih il j)|] = [|j|] + [|fst (div3121 ih il j)|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => fst (m1 / 2)%int31 + v30 end|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31

[|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31 + v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31

[|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31

[|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|r|] / [|2|]
intros HH; rewrite HH; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31

1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31 + v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31

1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31 + v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

(1 * wB + [|r|]) / 2 = [|fst (r / 2)%int31 + v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

([|v30|] * 2 + [|r|]) / 2 = [|fst (r / 2)%int31 + v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|v30|] + [|r|] / 2 = [|fst (r / 2)%int31 + v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 + [|v30|] = [|fst (r / 2)%int31 + v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 + [|v30|] = [|fst (r / 2)%int31|] + [|v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]
0 <= [|fst (r / 2)%int31|] + [|v30|] < wB
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 + [|v30|] = [|fst (r / 2)%int31|] + [|v30|]
rewrite div31_phi; auto.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

0 <= [|fst (r / 2)%int31|] + [|v30|] < wB
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

0 <= [|fst (r / 2)%int31|] + [|v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]
[|fst (r / 2)%int31|] + [|v30|] < wB
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

0 <= [|fst (r / 2)%int31|] + [|v30|]
case (phi_bounded (fst (r/2)%int31)); case (phi_bounded v30); auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|fst (r / 2)%int31|] + [|v30|] < wB
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 + [|v30|] < wB
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 + [|v30|] < base / 2 + [|v30|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 < base / 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 * 2 < base / 2 * 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 * 2 < base
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 * 2 <= [|r|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]
[|r|] < base
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] / 2 * 2 <= [|r|]
rewrite Z.mul_comm; apply Z_mult_div_ge; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:2 ^ 30 <= [|j|]
Hc:phi2 ih il / [|j|] < [|j|]
r:int31
HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]

[|r|] < base
case (phi_bounded r); auto with zarith. }
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30

[|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] ^ 2 <= phi2 ih il < ([|match (fst (div3121 ih il j) ?= j)%int31 with | Lt => let m := match j +c fst (div3121 ih il j) with | C0 m1 => fst (m1 / 2)%int31 | C1 m1 => (fst (m1 / 2) + v30)%int31 end in rec ih il m | _ => j end|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30

([|j|] + 1) ^ 2 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30

([|j|] + 1) ^ 2 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30

([|j|] + 1) ^ 2 <= 2 ^ 30 * 2 ^ 30
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30
2 ^ 30 * 2 ^ 30 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30

([|j|] + 1) ^ 2 <= 2 ^ 30 * 2 ^ 30
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30
H0:0 <= 1 + [|j|]

([|j|] + 1) ^ 2 <= 2 ^ 30 * 2 ^ 30
apply Z.mul_le_mono_nonneg; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30

2 ^ 30 * 2 ^ 30 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30

2 ^ 29 * base <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < base
Hp3:0 < phi2 ih il
Hc1:[|ih|] < [|j|]
Hjj:[|j|] < 2 ^ 30
H:1 + [|j|] <= 2 ^ 30

[|ih|] * base <= phi2 ih il
unfold phi2, base; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

[|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

[|j|] ^ 2 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

0 <= phi2 ih il
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]
phi2 ih il / [|j|] >= [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

0 <= phi2 ih il
unfold phi2, base; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

phi2 ih il / [|j|] >= [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

[|j|] <= [|j|] * base / [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]
[|j|] * base / [|j|] <= phi2 ih il / [|j|]
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

[|j|] <= [|j|] * base / [|j|]
rewrite Z.mul_comm, Z_div_mult; auto with zarith.
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hp2:0 < [|2|]
Hih:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2
H1:[|ih|] <= [|j|]
Hih1:0 <= [|ih|]
Hil1:0 <= [|il|]
Hj1:[|j|] < wB
Hp3:0 < phi2 ih il
Hc1:[|j|] < [|ih|]

[|j|] * base / [|j|] <= phi2 ih il / [|j|]
apply Z.ge_le; apply Z_div_ge; auto with zarith. Qed.
n:nat
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31

2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j|] + 1) ^ 2
n:nat
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31

2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j|] + 1) ^ 2

forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 0 + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step rec ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step rec ih il j|] + 1) ^ 2

forall n : nat, (forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 0 + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2

forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2

forall n : nat, (forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
Hrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat 0 + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
H:0 < [|j1|] < [|j|]
H0:phi2 ih il < ([|j1|] + 1) ^ 2

2 ^ Z.of_nat 0 + [|j1|] <= [|j|]

forall n : nat, (forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] + 1) ^ 2

forall n : nat, (forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|iter312_sqrt n rec ih il j|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j|] + 1) ^ 2) -> forall (rec : int31 -> int31 -> int31 -> int31) (ih il j : int31), 2 ^ 29 <= [|ih|] -> 0 < [|j|] -> phi2 ih il < ([|j|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2) -> [|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j0|] -> phi2 ih0 il0 < ([|j1|] + 1) ^ 2 -> [|rec0 ih0 il0 j1|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j1|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2

[|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] ^ 2 <= phi2 ih il < ([|sqrt312_step (iter312_sqrt n (iter312_sqrt n rec)) ih il j|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat n + [|j1|] <= [|j0|] -> phi2 ih0 il0 < ([|j1|] + 1) ^ 2 -> [|rec0 ih0 il0 j1|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j1|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat (S n) + [|j1|] <= [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2

forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|iter312_sqrt n (iter312_sqrt n rec) ih il j1|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n (iter312_sqrt n rec) ih il j1|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j2 : int31, 0 < [|j2|] -> 2 ^ Z.of_nat n + [|j2|] <= [|j0|] -> phi2 ih0 il0 < ([|j2|] + 1) ^ 2 -> [|rec0 ih0 il0 j2|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j2|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2

forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat n + [|j0|] <= [|j1|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|iter312_sqrt n rec ih il j0|] ^ 2 <= phi2 ih il < ([|iter312_sqrt n rec ih il j0|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j3 : int31, 0 < [|j3|] -> 2 ^ Z.of_nat n + [|j3|] <= [|j0|] -> phi2 ih0 il0 < ([|j3|] + 1) ^ 2 -> [|rec0 ih0 il0 j3|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j3|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2

forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat n + [|j0|] <= [|j2|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> phi2 ih0 il0 < ([|j4|] + 1) ^ 2 -> [|rec0 ih0 il0 j4|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j4|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

phi2 ih il < ([|j3|] + 1) ^ 2 -> [|rec ih il j3|] ^ 2 <= phi2 ih il < ([|rec ih il j3|] + 1) ^ 2
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> phi2 ih0 il0 < ([|j4|] + 1) ^ 2 -> [|rec0 ih0 il0 j4|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j4|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

2 ^ Z.of_nat (S n) + [|j3|] <= [|j|]
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> phi2 ih0 il0 < ([|j4|] + 1) ^ 2 -> [|rec0 ih0 il0 j4|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j4|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

2 * 2 ^ Z.of_nat n + [|j3|] <= [|j|]
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> phi2 ih0 il0 < ([|j4|] + 1) ^ 2 -> [|rec0 ih0 il0 j4|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j4|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]
0 <= Z.of_nat n
n:nat
Hrec:forall (rec0 : int31 -> int31 -> int31 -> int31) (ih0 il0 j0 : int31), 2 ^ 29 <= [|ih0|] -> 0 < [|j0|] -> phi2 ih0 il0 < ([|j0|] + 1) ^ 2 -> (forall j4 : int31, 0 < [|j4|] -> 2 ^ Z.of_nat n + [|j4|] <= [|j0|] -> phi2 ih0 il0 < ([|j4|] + 1) ^ 2 -> [|rec0 ih0 il0 j4|] ^ 2 <= phi2 ih0 il0 < ([|rec0 ih0 il0 j4|] + 1) ^ 2) -> [|iter312_sqrt n rec0 ih0 il0 j0|] ^ 2 <= phi2 ih0 il0 < ([|iter312_sqrt n rec0 ih0 il0 j0|] + 1) ^ 2
rec:int31 -> int31 -> int31 -> int31
ih, il, j:int31
Hi:2 ^ 29 <= [|ih|]
Hj:0 < [|j|]
Hij:phi2 ih il < ([|j|] + 1) ^ 2
HHrec:forall j0 : int31, 0 < [|j0|] -> 2 ^ Z.of_nat (S n) + [|j0|] <= [|j|] -> phi2 ih il < ([|j0|] + 1) ^ 2 -> [|rec ih il j0|] ^ 2 <= phi2 ih il < ([|rec ih il j0|] + 1) ^ 2
j1:int31
Hj1:0 < [|j1|] < [|j|]
Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2
j2:int31
Hj2:0 < [|j2|]
H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]
Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2
j3:int31
Hj3:0 < [|j3|]
Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]

0 <= Z.of_nat n
apply Nat2Z.is_nonneg. Qed. (* Avoid expanding [iter312_sqrt] before variables in the context. *) Strategy 1 [iter312_sqrt].

forall x y : int31, wB / 4 <= [|x|] -> let (s, r) := sqrt312 x y in [||WW x y||] = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]

forall x y : int31, wB / 4 <= [|x|] -> let (s, r) := sqrt312 x y in [||WW x y||] = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]

let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in [||WW ih il||] = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]

let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2

let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base

let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base

phi2 ih il < ([|Tn|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base

phi2 ih il < ([|Tn|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base

phi2 ih il < 2 ^ 62
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base

phi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
(2 ^ 31 - 1) * base + (2 ^ 31 - 1) < 2 ^ 62
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base

phi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
H1:0 <= [|il|]
H2:[|il|] < wB
H3:0 <= [|ih|]
H4:[|ih|] < wB

phi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
H1:0 <= [|il|]
H2:[|il|] < 2147483648
H3:0 <= [|ih|]
H4:[|ih|] < 2147483648

phi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
H1:0 <= [|il|]
H2:[|il|] < 2147483648
H3:0 <= [|ih|]
H4:[|ih|] < 2147483648

[|ih|] * base + [|il|] <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
H1:0 <= [|il|]
H2:[|il|] < 2147483648
H3:0 <= [|ih|]
H4:[|ih|] < 2147483648

[|ih|] * base + [|il|] <= (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (...))))))))))) - 1) * base + (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * (2 * ...))))))))))) - 1)
auto with zarith.
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2

let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2

0 < [|Tn|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|Tn|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|j1|] ^ 2 <= phi2 ih il < ([|j1|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
[|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] + 1) ^ 2 -> let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2

forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|Tn|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|j1|] ^ 2 <= phi2 ih il < ([|j1|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
[|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] + 1) ^ 2 -> let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
j1:int31

~ 2 ^ Z.of_nat 31 + [|j1|] <= [|Tn|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
[|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] + 1) ^ 2 -> let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
j1:int31

[|Tn|] < 2 ^ Z.of_nat 31 + [|j1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
[|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] + 1) ^ 2 -> let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
j1:int31

2147483647 < 2 ^ Z.of_nat 31 + [|j1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
[|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] + 1) ^ 2 -> let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
j1:int31

2147483647 < 2147483648 + [|j1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
[|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] + 1) ^ 2 -> let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2

[|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn|] + 1) ^ 2 -> let (s, r) := match iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn *c iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C1 il2) | _ => (iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn, C0 il2) end end end in phi2 ih il = [|s|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31

[|s|] ^ 2 <= phi2 ih il -> phi2 ih il < ([|s|] + 1) ^ 2 -> let (s0, r) := match s *c s with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2

let (s0, r) := match s *c s with | W0 => (0%int31, C0 0%int31) | WW ih1 il1 => match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2

[||W0||] = [|s|] * [|s|] -> phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
HH:0 = [|s|] * [|s|]

phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
HH:0 = [|s|] * [|s|]

[|s|] = 0
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
HH:0 = [|s|] * [|s|]
H:[|s|] = 0
phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
HH:0 = [|s|] * [|s|]

[|s|] = 0
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
HH:[|s|] * [|s|] = 0

[|s|] = 0
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
HH:[|s|] = 0 \/ [|s|] = 0

[|s|] = 0
destruct HH; auto.
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
HH:0 = [|s|] * [|s|]
H:[|s|] = 0

phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
HH:0 = [|s|] * [|s|]
H:[|s|] = 0

(0 + 1) ^ 2 <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
HH:0 = [|s|] * [|s|]
H:[|s|] = 0

1 <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
HH:0 = [|s|] * [|s|]
H:[|s|] = 0

1 <= wB / 4 * base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
HH:0 = [|s|] * [|s|]
H:[|s|] = 0
wB / 4 * base <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
HH:0 = [|s|] * [|s|]
H:[|s|] = 0

wB / 4 * base <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
HH:0 = [|s|] * [|s|]
H:[|s|] = 0

[|ih|] * base <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2

forall i i0 : int31, [||WW i i0||] = [|s|] * [|s|] -> let (s0, r) := match il -c i0 with | C0 il2 => match (ih ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= i)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31

[||WW ih1 il1||] = [|s|] * [|s|] -> let (s0, r) := match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31

phi2 ih1 il1 = [|s|] * [|s|] -> let (s0, r) := match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]

let (s0, r) := match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]

[-|il -c il1|] = [|il|] - [|il1|] -> let (s0, r) := match il -c il1 with | C0 il2 => match (ih ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end | C1 il2 => match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]

let (s0, r) := match (ih ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]

[|ih|] = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]

[|ih|] = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]

phi2 ih il = [|s|] ^ 2 + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]

phi2 ih il = phi2 ih1 il1 + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]

[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]

phi2 ih il - phi2 ih1 il1 <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]
phi2 ih il - phi2 ih1 il1 = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]

phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]
phi2 ih il - phi2 ih1 il1 = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih|] = [|ih1|]

phi2 ih il - phi2 ih1 il1 = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]

[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]

[|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]

~ [|s|] ^ 2 <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]

phi2 ih il < phi2 ih1 il1
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]

[|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]
H2:[|il|] < wB

[|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]
H2:[|il|] < wB

[|ih|] * base + [|il|] < ([|ih|] + 1) * base + 0
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]
H2:[|il|] < wB
([|ih|] + 1) * base + 0 <= [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]
H2:[|il|] < wB

([|ih|] + 1) * base + 0 <= [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
H1:[|ih|] < [|ih1|]
H2:[|il|] < wB
H3:0 <= [|il1|]

([|ih|] + 1) * base + 0 <= [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]
[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C0 il2|] = [|il|] - [|il1|]

[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]

[|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]

[|ih1|] < [|ih|] -> phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] < [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 <= [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]

([|s|] + 1) ^ 2 <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]

phi2 ih1 il1 + 2 * [|s|] + 1 <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]

[|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]

[|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]

[|il1|] <= [|il|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]
Hl1l:[|il1|] <= [|il|]
[|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]

[|il1|] <= [|il|]
case (phi_bounded il2); rewrite Hil2; auto with zarith.
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]
Hl1l:[|il1|] <= [|il|]

[|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]
Hl1l:[|il1|] <= [|il|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]
Hl1l:[|il1|] <= [|il|]
Hps:[|s|] < base

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]
Hl1l:[|il1|] <= [|il|]
Hps:[|s|] < base
Hpih1:0 <= [|ih1|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]
Hl1l:[|il1|] <= [|il|]
Hps:[|s|] < base
Hpih1:0 <= [|ih1|]

[|ih1|] * base + 2 * [|s|] + 1 <= ([|ih1|] + 2) * base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
Hpil:0 <= [|il|]
Hl1l:[|il1|] <= [|il|]
Hps:[|s|] < base
Hpih1:0 <= [|ih1|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih1|] * base + 2 * base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 < [|ih|]

phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|]))
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]

base + ([|il|] - [|il1|]) <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]

phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]
phi2 ih il - [|s|] * [|s|] = base + ([|il|] - [|il1|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[|il2|] = [|il|] - [|il1|]
H1:[|ih1|] + 1 = [|ih|]

phi2 ih il - [|s|] * [|s|] = base + ([|il|] - [|il1|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:[-|C1 il2|] = [|il|] - [|il1|]

let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + [+|r|] /\ [+|r|] <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]

let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + match r with | C0 x => [|x|] | C1 x => 1 * wB + [|x|] end /\ match r with | C0 x => [|x|] | C1 x => 1 * wB + [|x|] end <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]

let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + match r with | C0 x => [|x|] | C1 x => base + [|x|] end /\ match r with | C0 x => [|x|] | C1 x => base + [|x|] end <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]

[|ih - 1|] = [|ih|] - 1
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + match r with | C0 x => [|x|] | C1 x => base + [|x|] end /\ match r with | C0 x => [|x|] | C1 x => base + [|x|] end <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]

[|ih - 1|] = [|ih|] - 1
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]

0 <= [|ih|] - 1 < wB
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
H1:0 <= [|ih|]
H2:[|ih|] < wB

0 <= [|ih|] - 1 < wB
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
H1:0 <= [|ih|]
H2:[|ih|] < wB

536870912 <= [|ih|] -> 0 <= [|ih|] - 1 < wB
split; auto with zarith.
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1

let (s0, r) := match (ih - 1 ?= ih1)%int31 with | Gt => (s, C1 il2) | _ => (s, C0 il2) end in phi2 ih il = [|s0|] ^ 2 + match r with | C0 x => [|x|] | C1 x => base + [|x|] end /\ match r with | C0 x => [|x|] | C1 x => base + [|x|] end <= 2 * [|s0|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1

[|ih - 1|] = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1

[|ih|] - 1 = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

phi2 ih il = [|s|] ^ 2 + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

phi2 ih il = phi2 ih1 il1 + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

[|ih|] * base + [|il|] = ([|ih|] - 1) * base + [|il1|] + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

[|ih|] * base + [|il|] = [|ih|] * base + [|il1|] + ([|il|] - [|il1|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
[|ih|] * base + [|il1|] + ([|il|] - [|il1|]) = ([|ih|] - 1) * base + [|il1|] + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

[|ih|] * base + [|il1|] + ([|il|] - [|il1|]) = ([|ih|] - 1) * base + [|il1|] + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

[|ih|] * base + [|il1|] + (-1 * wB + [|il2|]) = ([|ih|] - 1) * base + [|il1|] + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

[|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

phi2 ih il - phi2 ih1 il1 <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
phi2 ih il - phi2 ih1 il1 = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
phi2 ih il - phi2 ih1 il1 = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

phi2 ih il - phi2 ih1 il1 = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

[|ih|] * base + [|il|] - ([|ih1|] * base + [|il1|]) = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

[|ih|] * base + [|il|] - (([|ih|] - 1) * base + [|il1|]) = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

base + [|il|] - [|il1|] = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

base + [|il|] - [|il1|] = base + ([|il|] - [|il1|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]
base + ([|il|] - [|il1|]) = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

base + ([|il|] - [|il1|]) = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 = [|ih1|]

base + (-1 * wB + [|il2|]) = [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1

[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]

phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]

[|ih|] = [|ih1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
He:[|ih|] = [|ih1|]
phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]

[|ih|] = [|ih1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]

[|ih1|] <= [|ih|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]

[|ih1|] <= [|ih|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]

phi2 ih il < phi2 ih1 il1
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]

[|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]
Hpil1:[|il|] < base

[|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]
Hpil1:[|il|] < base

[|ih|] * base + [|il|] < ([|ih|] + 1) * base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]
Hpil1:[|il|] < base
([|ih|] + 1) * base <= [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]
Hpil1:[|il|] < base

([|ih|] + 1) * base <= [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
H2:[|ih|] < [|ih1|]
Hpil1:[|il|] < base
Hpil2:0 <= [|il1|]

([|ih|] + 1) * base <= [|ih1|] * base + [|il1|]
apply Z.le_trans with (([|ih1|]) * base); auto with zarith.
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
He:[|ih|] = [|ih1|]

phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
He:[|ih|] = [|ih1|]

[|ih|] * base + [|il|] = [|ih|] * base + [|il1|] + [|il2|] /\ [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
He:[|ih|] = [|ih1|]

phi2 ih il < phi2 ih1 il1
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
He:[|ih|] = [|ih1|]

[|ih1|] * base + [|il|] < [|ih1|] * base + [|il1|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
He:[|ih|] = [|ih1|]

[|il|] - [|il1|] < 0
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih|] - 1 < [|ih1|]
He:[|ih|] = [|ih1|]

-1 * wB + [|il2|] < 0
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1

[|ih1|] < [|ih - 1|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]

phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]

phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 <= [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

([|s|] + 1) ^ 2 <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

phi2 ih1 il1 + 2 * [|s|] + 1 <= phi2 ih il
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

[|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + ([|il|] - [|il1|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + (-1 * wB + [|il2|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + (- base + [|il2|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
Hpil2:0 <= [|il2|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + (- base + [|il2|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
Hpil2:0 <= [|il2|]

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + - base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
Hpil2:0 <= [|il2|]
Hps:[|s|] < base

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + - base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
Hpil2:0 <= [|il2|]
Hps:[|s|] < base
H:2 * [|s|] + 1 <= 2 * base

[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + - base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
Hpil2:0 <= [|il2|]
Hps:[|s|] < base
H:2 * [|s|] + 1 <= 2 * base

[|ih1|] * base + 2 * base <= [|ih|] * base + - base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
Hpil2:0 <= [|il2|]
Hps:[|s|] < base
H:2 * [|s|] + 1 <= 2 * base
Hi:([|ih1|] + 3) * base <= [|ih|] * base

[|ih1|] * base + 2 * base <= [|ih|] * base + - base
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]
phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 < [|ih|]

phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

([|ih1|] + 2) * base + [|il|] = [|ih1|] * base + [|il1|] + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

([|ih1|] + 2) * base + [|il|] = [|ih1|] * base + [|il1|] + (base + [|il2|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

([|ih1|] + 2) * base + ([|il|] - [|il1|] + [|il1|]) = [|ih1|] * base + [|il1|] + (base + [|il2|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

([|ih1|] + 2) * base + (-1 * wB + [|il2|] + [|il1|]) = [|ih1|] * base + [|il1|] + (base + [|il2|])
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

base + [|il2|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

phi2 ih il - phi2 ih1 il1 <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il - phi2 ih1 il1 = base + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]
phi2 ih il - phi2 ih1 il1 = base + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

phi2 ih il - phi2 ih1 il1 = base + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

([|ih1|] + 2) * base + [|il|] - ([|ih1|] * base + [|il1|]) = base + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

([|ih1|] + 2) * base + ([|il|] - [|il1|] + [|il1|]) - ([|ih1|] * base + [|il1|]) = base + [|il2|]
ih, il:int31
Hih:wB / 4 <= [|ih|]
Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2
Hb:0 <= base
Hi2:phi2 ih il < ([|Tn|] + 1) ^ 2
s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31
Hs1:[|s|] ^ 2 <= phi2 ih il
Hs2:phi2 ih il < ([|s|] + 1) ^ 2
ih1, il1:int31
Hihl1:phi2 ih1 il1 = [|s|] * [|s|]
il2:int31
Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]
Hsih:[|ih - 1|] = [|ih|] - 1
H1:[|ih1|] < [|ih - 1|]
H2:[|ih1|] + 2 = [|ih|]

([|ih1|] + 2) * base + (-1 * wB + [|il2|] + [|il1|]) - ([|ih1|] * base + [|il1|]) = base + [|il2|]
change (-1 * 2 ^ Z.of_nat size) with (-base); ring. Qed.
iszero
 

forall x : int31, ZnZ.eq0 x = true -> [|x|] = 0

forall x : int31, ZnZ.eq0 x = true -> [|x|] = 0

forall x : int31, match (x ?= 0)%int31 with | Eq => true | _ => false end = true -> [|x|] = 0
x:int31
H:match [|x|] ?= [|0|] with | Eq => true | _ => false end = true

[|x|] = 0
x:int31
H:match [|x|] ?= 0 with | Eq => true | _ => false end = true

[|x|] = 0
x:int31
H:match [|x|] ?= 0 with | Eq => true | _ => false end = true

([|x|] ?= 0) = Eq
now destruct ([|x|] ?= 0). Qed. (* Even *)

forall x : int31, if ZnZ.is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1

forall x : int31, if ZnZ.is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x:int31

if let (_, r) := (x / 2)%int31 in match (r ?= 0)%int31 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x:int31

(0 < [|2|] -> let (q, r) := (x / 2)%int31 in [|x|] = [|q|] * [|2|] + [|r|] /\ 0 <= [|r|] < [|2|]) -> if let (_, r) := (x / 2)%int31 in match (r ?= 0)%int31 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x, q, r:int31
H:0 < [|2|] -> [|x|] = [|q|] * [|2|] + [|r|] /\ 0 <= [|r|] < [|2|]

if match (r ?= 0)%int31 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x, q, r:int31
H:0 < [|2|] -> [|x|] = [|q|] * [|2|] + [|r|] /\ 0 <= [|r|] < [|2|]

if match [|r|] ?= [|0|] with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x, q, r:int31
H:0 < 2 -> [|x|] = [|q|] * 2 + [|r|] /\ 0 <= [|r|] < 2

if match [|r|] ?= [|0|] with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x, q, r:int31
H:0 < 2 -> [|x|] = [|q|] * 2 + [|r|] /\ 0 <= [|r|] < 2

if match [|r|] ?= 0 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x, q, r:int31
H:[|x|] = [|q|] * 2 + [|r|]
H0:0 <= [|r|] < 2

if match [|r|] ?= 0 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1
x, q, r:int31
H:[|x|] = [|q|] * 2 + [|r|]
H0:0 <= [|r|] < 2

if match [|r|] ?= 0 with | Eq => true | _ => false end then [|r|] = 0 else [|r|] = 1
x, q, r:int31
H:[|x|] = [|q|] * 2 + [|r|]
H0:0 <= [|r|] < 2
[|r|] = [|x|] mod 2
x, q, r:int31
H0:0 <= [|r|] < 2

if match [|r|] ?= 0 with | Eq => true | _ => false end then [|r|] = 0 else [|r|] = 1
x, q, r:int31
H:[|x|] = [|q|] * 2 + [|r|]
H0:0 <= [|r|] < 2
[|r|] = [|x|] mod 2
x, q, r:int31
H:[|x|] = [|q|] * 2 + [|r|]
H0:0 <= [|r|] < 2

[|r|] = [|x|] mod 2
apply Zmod_unique with [|q|]; auto with zarith. Qed. (* Bitwise *)
x:int31

Z.log2 [|x|] < Z.of_nat size
x:int31

Z.log2 [|x|] < Z.of_nat size
x:int31
H:0 <= [|x|]
H':[|x|] < wB

Z.log2 [|x|] < Z.of_nat size
x:int31
H:0 < [|x|]
H':[|x|] < wB

Z.log2 [|x|] < Z.of_nat size
x:int31
H:0 = [|x|]
H':[|x|] < wB
Z.log2 [|x|] < Z.of_nat size
x:int31
H:0 < [|x|]
H':[|x|] < wB

Z.log2 [|x|] < Z.of_nat size
now apply Z.log2_lt_pow2.
x:int31
H:0 = [|x|]
H':[|x|] < wB

Z.log2 [|x|] < Z.of_nat size
now rewrite <- H. Qed.
x, y:int31

[|ZnZ.lor x y|] = Z.lor [|x|] [|y|]
x, y:int31

[|ZnZ.lor x y|] = Z.lor [|x|] [|y|]
x, y:int31

[|lor31 x y|] = Z.lor [|x|] [|y|]
x, y:int31

[|phi_inv (Z.lor [|x|] [|y|])|] = Z.lor [|x|] [|y|]
x, y:int31

Z.lor [|x|] [|y|] mod wB = Z.lor [|x|] [|y|]
x, y:int31

0 <= Z.lor [|x|] [|y|]
x, y:int31
Z.lor [|x|] [|y|] < wB
x, y:int31

0 <= Z.lor [|x|] [|y|]
apply Z.lor_nonneg; split; apply phi_bounded.
x, y:int31

Z.lor [|x|] [|y|] < wB
x, y:int31

Z.log2 (Z.lor [|x|] [|y|]) < Z.log2 wB
x, y:int31

Z.log2 (Z.lor [|x|] [|y|]) < Z.of_nat size
x, y:int31

Z.max (Z.log2 [|x|]) (Z.log2 [|y|]) < Z.of_nat size
apply Z.max_lub_lt; apply log2_phi_bounded. Qed.
x, y:int31

[|ZnZ.land x y|] = Z.land [|x|] [|y|]
x, y:int31

[|ZnZ.land x y|] = Z.land [|x|] [|y|]
x, y:int31

[|land31 x y|] = Z.land [|x|] [|y|]
x, y:int31

[|phi_inv (Z.land [|x|] [|y|])|] = Z.land [|x|] [|y|]
x, y:int31

Z.land [|x|] [|y|] mod wB = Z.land [|x|] [|y|]
x, y:int31

0 <= Z.land [|x|] [|y|]
x, y:int31
Z.land [|x|] [|y|] < wB
x, y:int31

0 <= Z.land [|x|] [|y|]
apply Z.land_nonneg; left; apply phi_bounded.
x, y:int31

Z.land [|x|] [|y|] < wB
x, y:int31

Z.log2 (Z.land [|x|] [|y|]) < Z.log2 wB
x, y:int31

Z.log2 (Z.land [|x|] [|y|]) < Z.of_nat size
x, y:int31

Z.log2 (Z.land [|x|] [|y|]) <= ?m
x, y:int31
?m < Z.of_nat size
x, y:int31

Z.min (Z.log2 [|x|]) (Z.log2 [|y|]) < Z.of_nat size
apply Z.min_lt_iff; left; apply log2_phi_bounded. Qed.
x, y:int31

[|ZnZ.lxor x y|] = Z.lxor [|x|] [|y|]
x, y:int31

[|ZnZ.lxor x y|] = Z.lxor [|x|] [|y|]
x, y:int31

[|lxor31 x y|] = Z.lxor [|x|] [|y|]
x, y:int31

[|phi_inv (Z.lxor [|x|] [|y|])|] = Z.lxor [|x|] [|y|]
x, y:int31

Z.lxor [|x|] [|y|] mod wB = Z.lxor [|x|] [|y|]
x, y:int31

0 <= Z.lxor [|x|] [|y|]
x, y:int31
Z.lxor [|x|] [|y|] < wB
x, y:int31

0 <= Z.lxor [|x|] [|y|]
apply Z.lxor_nonneg; split; intros; apply phi_bounded.
x, y:int31

Z.lxor [|x|] [|y|] < wB
x, y:int31

Z.log2 (Z.lxor [|x|] [|y|]) < Z.log2 wB
x, y:int31

Z.log2 (Z.lxor [|x|] [|y|]) < Z.of_nat size
x, y:int31

Z.log2 (Z.lxor [|x|] [|y|]) <= ?m
x, y:int31
?m < Z.of_nat size
x, y:int31

Z.max (Z.log2 [|x|]) (Z.log2 [|y|]) < Z.of_nat size
apply Z.max_lub_lt; apply log2_phi_bounded. Qed. Instance int31_specs : ZnZ.Specs int31_ops := { spec_to_Z := phi_bounded; spec_of_pos := positive_to_int31_spec; spec_zdigits := spec_zdigits; spec_more_than_1_digit := spec_more_than_1_digit; spec_0 := spec_0; spec_1 := spec_1; spec_m1 := spec_m1; spec_compare := spec_compare; spec_eq0 := spec_eq0; spec_opp_c := spec_opp_c; spec_opp := spec_opp; spec_opp_carry := spec_opp_carry; spec_succ_c := spec_succ_c; spec_add_c := spec_add_c; spec_add_carry_c := spec_add_carry_c; spec_succ := spec_succ; spec_add := spec_add; spec_add_carry := spec_add_carry; spec_pred_c := spec_pred_c; spec_sub_c := spec_sub_c; spec_sub_carry_c := spec_sub_carry_c; spec_pred := spec_pred; spec_sub := spec_sub; spec_sub_carry := spec_sub_carry; spec_mul_c := spec_mul_c; spec_mul := spec_mul; spec_square_c := spec_square_c; spec_div21 := spec_div21; spec_div_gt := fun a b _ => spec_div a b; spec_div := spec_div; spec_modulo_gt := fun a b _ => spec_mod a b; spec_modulo := spec_mod; spec_gcd_gt := fun a b _ => spec_gcd a b; spec_gcd := spec_gcd; spec_head00 := spec_head00; spec_head0 := spec_head0; spec_tail00 := spec_tail00; spec_tail0 := spec_tail0; spec_add_mul_div := spec_add_mul_div; spec_pos_mod := spec_pos_mod; spec_is_even := spec_is_even; spec_sqrt2 := spec_sqrt2; spec_sqrt := spec_sqrt; spec_lor := spec_lor; spec_land := spec_land; spec_lxor := spec_lxor }. End Int31_Specs. Module Int31Cyclic <: CyclicType. Definition t := int31. Definition ops := int31_ops. Definition specs := int31_specs. End Int31Cyclic.