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.
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.
forall x : int31, iszero x = true -> x = 0forall x : int31, iszero x = true -> x = 0d, 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:digitsH: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 = trueI31 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 = 0reflexivity. Qed.H:true = true0 = 0forall x : int31, iszero x = false -> x <> 0intros x H Eq; rewrite Eq in H; simpl in *; discriminate. Qed.forall x : int31, iszero x = false -> x <> 0forall x : int31, x = sneakl (firstr x) (shiftr x)destruct x; simpl; auto. Qed.forall x : int31, x = sneakl (firstr x) (shiftr x)forall x : int31, x = sneakr (firstl x) (shiftl x)destruct x; simpl; auto. Qed.forall x : int31, x = sneakr (firstl x) (shiftl x)forall x : int31, twice x = 0 <-> twice_plus_one x = 1destruct x; simpl in *; split; intro H; injection H; intros; subst; auto. Qed.forall x : int31, twice x = 0 <-> twice_plus_one x = 1forall 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:int31H:firstr x = D0x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)x:int31H:firstr x = D1x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)destruct x; simpl in *; rewrite H; auto. Qed.x:int31H:firstr x = D1x = twice (shiftr x) \/ x = twice_plus_one (shiftr x)
Definition nshiftr x := nat_rect _ x (fun _ => shiftr).forall (n : nat) (x : int31), nshiftr x (S n) = shiftr (nshiftr x n)reflexivity. Qed.forall (n : nat) (x : int31), nshiftr x (S n) = shiftr (nshiftr x n)forall (n : nat) (x : int31), nshiftr x (S n) = nshiftr (shiftr x) nforall (n : nat) (x : int31), nshiftr x (S n) = nshiftr (shiftr x) nintros; now f_equal. Qed.n:natforall 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)forall n : nat, nshiftr 0 n = 0forall n : nat, nshiftr 0 n = 0rewrite IHn; auto. Qed.n:natIHn:nshiftr 0 n = 0shiftr (nshiftr 0 n) = 0forall x : int31, nshiftr x size = 0destruct x; simpl; auto. Qed.forall x : int31, nshiftr x size = 0forall (k : nat) (x : int31), size <= k -> nshiftr x k = 0forall (k : nat) (x : int31), size <= k -> nshiftr x k = 0k:natx:int31H:size <= knshiftr x k = 0k:natx:int31H:size <= knshiftr x (k - size + size) = 0k:natx:int31H:size <= knshiftr x (0 + size) = 0k:natx:int31H:size <= kn:natIHn:nshiftr x (n + size) = 0nshiftr x (S n + size) = 0simpl; rewrite IHn; auto. Qed.k:natx:int31H:size <= kn:natIHn:nshiftr x (n + size) = 0nshiftr x (S n + size) = 0
Definition nshiftl x := nat_rect _ x (fun _ => shiftl).forall (n : nat) (x : int31), nshiftl x (S n) = shiftl (nshiftl x n)reflexivity. Qed.forall (n : nat) (x : int31), nshiftl x (S n) = shiftl (nshiftl x n)forall (n : nat) (x : int31), nshiftl x (S n) = nshiftl (shiftl x) nintros n; elim n; simpl; intros; now f_equal. Qed.forall (n : nat) (x : int31), nshiftl x (S n) = nshiftl (shiftl x) nforall n : nat, nshiftl 0 n = 0forall n : nat, nshiftl 0 n = 0rewrite IHn; auto. Qed.n:natIHn:nshiftl 0 n = 0shiftl (nshiftl 0 n) = 0forall x : int31, nshiftl x size = 0destruct x; simpl; auto. Qed.forall x : int31, nshiftl x size = 0forall (k : nat) (x : int31), size <= k -> nshiftl x k = 0forall (k : nat) (x : int31), size <= k -> nshiftl x k = 0k:natx:int31H:size <= knshiftl x k = 0k:natx:int31H:size <= knshiftl x (k - size + size) = 0k:natx:int31H:size <= knshiftl x (0 + size) = 0k:natx:int31H:size <= kn:natIHn:nshiftl x (n + size) = 0nshiftl x (S n + size) = 0simpl; rewrite IHn; auto. Qed.k:natx:int31H:size <= kn:natIHn:nshiftl x (n + size) = 0nshiftl x (S n + size) = 0forall x : int31, firstr x = firstl (nshiftl x (Init.Nat.pred size))destruct x; simpl; auto. Qed.forall x : int31, firstr x = firstl (nshiftl x (Init.Nat.pred size))forall x : int31, firstl x = firstr (nshiftr x (Init.Nat.pred size))destruct x; simpl; auto. Qed.forall x : int31, firstl x = firstr (nshiftr x (Init.Nat.pred size))
More advanced results about nshiftr
forall x : int31, nshiftr x (Init.Nat.pred size) = 0 -> firstl x = D0destruct x; compute; intros H; injection H; intros; subst; auto. Qed.forall x : int31, nshiftr x (Init.Nat.pred size) = 0 -> firstl x = D0forall (n p : nat) (x : int31), n <= p -> nshiftr x n = 0 -> nshiftr x p = 0forall (n p : nat) (x : int31), n <= p -> nshiftr x n = 0 -> nshiftr x p = 0n, p:natx:int31H:n <= pH0:nshiftr x n = 0nshiftr x p = 0n, p:natx:int31H:n <= pH0:nshiftr x n = 0nshiftr x (p - n + n) = 0n, p:natx:int31H:n <= pH0:nshiftr x n = 0nshiftr x (0 + n) = 0n, p:natx:int31H:n <= pH0:nshiftr x n = 0n0:natIHn0:nshiftr x (n0 + n) = 0nshiftr x (S n0 + n) = 0simpl; rewrite IHn0; auto. Qed.n, p:natx:int31H:n <= pH0:nshiftr x n = 0n0:natIHn0:nshiftr x (n0 + n) = 0nshiftr x (S n0 + n) = 0forall (n : nat) (x : int31), n < size -> nshiftr x n = 0 -> firstl x = D0forall (n : nat) (x : int31), n < size -> nshiftr x n = 0 -> firstl x = D0n:natx:int31H:n < sizeH0:nshiftr x n = 0firstl x = D0apply nshiftr_0_propagates with n; auto; omega. Qed.n:natx:int31H:n < sizeH0:nshiftr x n = 0nshiftr x (Init.Nat.pred size) = 0
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 xforall P : int31 -> Prop, P 0 -> (forall (x : int31) (d : digits), P x -> P (sneakl d x)) -> forall x : int31, P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31forall n : nat, n <= size -> P (nshiftr x (size - n))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:0 <= sizeP (nshiftr x (size - 0))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31n:natIHn:n <= size -> P (nshiftr x (size - n))H1:S n <= sizeP (nshiftr x (size - S n))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31n:natIHn:n <= size -> P (nshiftr x (size - n))H1:S n <= sizeP (nshiftr x (size - S n))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31n:natIHn:n <= size -> P (nshiftr x (size - n))H1:S n <= sizeP (sneakl (firstr (nshiftr x (size - S n))) (shiftr (nshiftr x (size - S n))))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31n:natIHn:n <= size -> P (nshiftr x (size - n))H1:S n <= sizeP (shiftr (nshiftr x (size - S n)))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31n:natIHn:n <= size -> P (nshiftr x (size - n))H1:S n <= sizeP (nshiftr x (S (size - S n)))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xP:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31n:natIHn:n <= size -> P (nshiftr x (size - n))H1:S n <= sizeP (nshiftr x (size - n))P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xchange x with (nshiftr x (size-size)); auto. Qed.P:int31 -> PropH:P 0H0:forall (x0 : int31) (d : digits), P x0 -> P (sneakl d x0)x:int31H1:forall n : nat, n <= size -> P (nshiftr x (size - n))P xforall 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 xforall 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 xdestruct d; auto. Qed.P:int31 -> PropH:P 0H0:forall x0 : int31, P x0 -> P (twice x0)H1:forall x0 : int31, P x0 -> P (twice_plus_one x0)x:int31d:digitsIHx:P xP (sneakl d x)
Section Recr.
recr satisfies the fixpoint equation used for its definition.
Variable (A:Type)(case0:A)(caserec:digits->int31->A->A).A:Typecase0:Acaserec:digits -> int31 -> A -> Aforall (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:Typecase0:Acaserec:digits -> int31 -> A -> Aforall (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:Typecase0:Acaserec:digits -> int31 -> A -> Aforall (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:Typecase0:Acaserec:digits -> int31 -> A -> Aforall (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:Typecase0:Acaserec:digits -> int31 -> A -> Aforall (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:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:Typecase0:Acaserec:digits -> int31 -> A -> Ap:natx:int31H:0 <= sizeH0:0 <= precr_aux 0 A case0 caserec (nshiftr x 31) = recr_aux p A case0 caserec (nshiftr x 31)A:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0:S n <= precr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux p A case0 caserec (nshiftr x (size - S n))A:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:int31H:S n <= sizeH0:S n <= 0recr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux 0 A case0 caserec (nshiftr x (size - S n))A:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0:S n <= S precr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux (S p) A case0 caserec (nshiftr x (size - S n))A:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0:S n <= S precr_aux (S n) A case0 caserec (nshiftr x (size - S n)) = recr_aux (S p) A case0 caserec (nshiftr x (size - S n))A:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0: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:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0:S n <= S pcaserec (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:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0:S n <= S precr_aux n A case0 caserec (shiftr (nshiftr x (size - S n))) = recr_aux p A case0 caserec (shiftr (nshiftr x (size - S n)))A:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0:S n <= S precr_aux n A case0 caserec (nshiftr x (S (size - S n))) = recr_aux p A case0 caserec (nshiftr x (S (size - S n)))apply IHn; auto with arith. Qed.A:Typecase0:Acaserec:digits -> int31 -> A -> An:natIHn: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:natx:int31H:S n <= sizeH0:S n <= S precr_aux n A case0 caserec (nshiftr x (size - n)) = recr_aux p A case0 caserec (nshiftr x (size - n))A:Typecase0:Acaserec:digits -> int31 -> A -> Aforall x : int31, iszero x = false -> recr A case0 caserec x = caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x))A:Typecase0:Acaserec:digits -> int31 -> A -> Aforall x : int31, iszero x = false -> recr A case0 caserec x = caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x))A:Typecase0:Acaserec:digits -> int31 -> A -> Ax:int31H:iszero x = falserecr A case0 caserec x = caserec (firstr x) (shiftr x) (recr A case0 caserec (shiftr x))A:Typecase0:Acaserec:digits -> int31 -> A -> Ax:int31H:iszero x = falserecr_aux size A case0 caserec x = caserec (firstr x) (shiftr x) (recr_aux size A case0 caserec (shiftr x))A:Typecase0:Acaserec:digits -> int31 -> A -> Ax:int31H:iszero x = falserecr_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))))rewrite recr_aux_eqn; auto. Qed.A:Typecase0:Acaserec:digits -> int31 -> A -> Ax:int31H:iszero x = falserecr_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))))
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:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0forall (n : nat) (x : int31), recrbis_aux n A case0 caserec x = recr_aux n A case0 caserec xA:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0forall (n : nat) (x : int31), recrbis_aux n A case0 caserec x = recr_aux n A case0 caserec xA:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0n:natIHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0x:int31caserec (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:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0n:natIHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0x:int31H:iszero x = truecaserec (firstr x) (shiftr x) (recrbis_aux n A case0 caserec (shiftr x)) = case0A:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0n:natIHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0x:int31H:iszero x = truecaserec D0 0 (recrbis_aux n A case0 caserec 0) = case0clear H IHn; induction n; simpl; congruence. Qed.A:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0n:natIHn:forall x0 : int31, recrbis_aux n A case0 caserec x0 = recr_aux n A case0 caserec x0x:int31H:iszero x = truecase0 = recrbis_aux n A case0 caserec 0A:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0forall x : int31, recrbis A case0 caserec x = recr A case0 caserec xintros; apply recrbis_aux_equiv; auto. Qed. End Recr.A:Typecase0:Acaserec:digits -> int31 -> A -> Acase0_caserec:caserec D0 0 case0 = case0forall x : int31, recrbis A case0 caserec x = recr A case0 caserec x
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 -> int31forall x : int31, incrbis_aux size x = incr xIncr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31forall x : int31, incrbis_aux size x = incr xapply recrbis_aux_equiv; auto. Qed.Incr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31x:int31recrbis_aux size int31 In Incr x = recr_aux size int31 In Incr x
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 -> int31forall 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 -> int31forall 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 -> int31x:int31H:firstr x = D0incr 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 -> int31x:int31H:firstr x = D0H0:iszero x = trueincr 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 -> int31x:int31H:firstr x = D0H0:iszero x = falseincr 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 -> int31x:int31H:firstr x = D0H0:iszero x = falseincr x = 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 -> int31x:int31H:firstr x = D0H0:iszero x = falsematch firstr x with | D0 => sneakl D1 (shiftr x) | D1 => sneakl D0 (incr (shiftr x)) end = 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 -> int31forall 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 -> int31forall 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 -> int31x:int31H:firstr x = D1incr 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 -> int31x:int31H:firstr x = D1H0:iszero x = trueincr 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 -> int31x:int31H:firstr x = D1H0:iszero x = falseincr 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 -> int31x:int31H:firstr x = D1H0:iszero x = falseincr x = 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 -> int31x:int31H:firstr x = D1H0:iszero x = falsematch firstr x with | D0 => sneakl D1 (shiftr x) | D1 => sneakl D0 (incr (shiftr x)) end = 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 -> int31forall x : int31, incr (twice x) = twice_plus_one xIncr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31forall x : int31, incr (twice x) = twice_plus_one xrewrite 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 -> int31x:int31incr (twice x) = twice_plus_one xIncr:=fun (b : digits) (si rec : int31) => match b with | D0 => sneakl D1 si | D1 => sneakl D0 rec end:digits -> int31 -> int31 -> int31forall 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 -> int31forall 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 -> int31x:int31H:firstl x = D0incr (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 -> int31x:int31H:firstl x = D0twice (incr (shiftr (twice_plus_one x))) = twice (incr x)destruct x; simpl in *; 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 -> int31x:int31H:firstl x = D0shiftr (twice_plus_one x) = x
The previous result is actually true even without the
constraint on firstl, but this is harder to prove
(see later).
End Incr.
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 -> Zforall x : int31, phibis_aux size x = phi xPhi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Zforall x : int31, phibis_aux size x = phi xapply recrbis_aux_equiv; auto. Qed.Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Zx:int31recrbis_aux size Z 0%Z Phi x = recr_aux size Z 0%Z Phi x
Recursive equations satisfied by phi
Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Zforall 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 -> Zforall 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 -> Zx:int31H:firstr x = D0phi 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 -> Zx:int31H:firstr x = D0H0:iszero x = truephi 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 -> Zx:int31H:firstr x = D0H0:iszero x = falsephi 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 -> Zx:int31H:firstr x = D0H0:iszero x = falsephi 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 -> Zx:int31H:firstr x = D0H0:iszero x = falsematch firstr x with | D0 => Z.double | D1 => Z.succ_double end (phi (shiftr 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 -> Zforall 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 -> Zforall 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 -> Zx:int31H:firstr x = D1phi 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 -> Zx:int31H:firstr x = D1H0:iszero x = truephi 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 -> Zx:int31H:firstr x = D1H0:iszero x = falsephi 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 -> Zx:int31H:firstr x = D1H0:iszero x = falsephi 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 -> Zx:int31H:firstr x = D1H0:iszero x = falsematch firstr x with | D0 => Z.double | D1 => Z.succ_double end (phi (shiftr 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 -> Zforall 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 -> Zforall 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 -> Zx:int31H:firstl x = D0phi (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 -> Zx:int31H:firstl x = D0Z.double (phi (shiftr (twice x))) = Z.double (phi 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 -> Zx:int31H:firstl x = D0shiftr (twice x) = xPhi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Zforall 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 -> Zforall 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 -> Zx:int31H:firstl x = D0phi (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 -> Zx:int31H:firstl x = D0Z.succ_double (phi (shiftr (twice_plus_one x))) = Z.succ_double (phi x)destruct x; simpl in *; rewrite H; auto. Qed. End Phi.Phi:=fun (b : digits) (_ : int31) => match b with | D0 => Z.double | D1 => Z.succ_double end:digits -> int31 -> Z -> Zx:int31H:firstl x = D0shiftr (twice_plus_one x) = x
phi x is positive and lower than 2^31
forall (n : nat) (x : int31), (0 <= phibis_aux n x)%Zforall (n : nat) (x : int31), (0 <= phibis_aux n x)%Zforall x : int31, (0 <= phibis_aux 0 x)%Zn:natIHn:forall x : int31, (0 <= phibis_aux n x)%Zforall x : int31, (0 <= phibis_aux (S n) x)%Zn:natIHn:forall x : int31, (0 <= phibis_aux n x)%Zforall x : int31, (0 <= phibis_aux (S n) x)%Zn:natIHn:forall x0 : int31, (0 <= phibis_aux n x0)%Zx:int31(0 <= phibis_aux (S n) x)%Zn:natIHn:forall x0 : int31, (0 <= phibis_aux n x0)%Zx:int31(0 <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end (phibis_aux n (shiftr x)))%Zn:natIHn:forall x0 : int31, (0 <= phibis_aux n x0)%Zx:int31(0 <= Z.double (phibis_aux n (shiftr x)))%Zn:natIHn:forall x0 : int31, (0 <= phibis_aux n x0)%Zx:int31(0 <= Z.succ_double (phibis_aux n (shiftr x)))%Zspecialize IHn with (shiftr x); rewrite Z.succ_double_spec; omega. Qed.n:natIHn:forall x0 : int31, (0 <= phibis_aux n x0)%Zx:int31(0 <= Z.succ_double (phibis_aux n (shiftr x)))%Zforall (n : nat) (x : int31), n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Zforall (n : nat) (x : int31), n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Zforall x : int31, 0 <= size -> (phibis_aux 0 (nshiftr x (size - 0)) < 2 ^ Z.of_nat 0)%Zn:natIHn:forall x : int31, n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Zforall x : int31, S n <= size -> (phibis_aux (S n) (nshiftr x (size - S n)) < 2 ^ Z.of_nat (S n))%Zn:natIHn:forall x : int31, n <= size -> (phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%Zforall x : int31, S n <= size -> (phibis_aux (S n) (nshiftr x (size - S n)) < 2 ^ Z.of_nat (S n))%Zn:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= size(phibis_aux (S n) (nshiftr x (size - S n)) < 2 ^ Z.of_nat (S n))%Zn:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H: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))%Zn:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= sizeshiftr (nshiftr x (size - S n)) = nshiftr x (size - n)n:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= sizeH0: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))%Zn:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= sizeshiftr (nshiftr x (size - S n)) = nshiftr x (S (size - S n))n:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= sizeH0: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))%Zn:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= sizeH0: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))%Zn:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= sizeH0: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))%Zn:natIHn:forall x0 : int31, n <= size -> (phibis_aux n (nshiftr x0 (size - n)) < 2 ^ Z.of_nat n)%Zx:int31H:S n <= sizeH0: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))%Zn:natx:int31IHn:(phibis_aux n (nshiftr x (size - n)) < 2 ^ Z.of_nat n)%ZH:S n <= sizeH0: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))%Zn:natx:int31y:=phibis_aux n (nshiftr x (size - n)):ZIHn:(y < 2 ^ Z.of_nat n)%ZH:S n <= sizeH0: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))%Zn:natx:int31y:=phibis_aux n (nshiftr x (size - n)):ZIHn:(y < 2 ^ Z.of_nat n)%ZH:S n <= sizeH0: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)%Zn:natx:int31y:=phibis_aux n (nshiftr x (size - n)):ZIHn:(y < 2 ^ Z.of_nat n)%ZH:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:n <= sizeH2:firstr (nshiftr x (size - S n)) = D0(Z.double y < 2 * 2 ^ Z.of_nat n)%Zn:natx:int31y:=phibis_aux n (nshiftr x (size - n)):ZIHn:(y < 2 ^ Z.of_nat n)%ZH:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:n <= sizeH2:firstr (nshiftr x (size - S n)) = D1(Z.succ_double y < 2 * 2 ^ Z.of_nat n)%Zrewrite Z.succ_double_spec; auto with zarith. Qed.n:natx:int31y:=phibis_aux n (nshiftr x (size - n)):ZIHn:(y < 2 ^ Z.of_nat n)%ZH:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:n <= sizeH2:firstr (nshiftr x (size - S n)) = D1(Z.succ_double y < 2 * 2 ^ Z.of_nat n)%Zforall x : int31, (0 <= phi x)%Zforall x : int31, (0 <= phi x)%Zx:int31(0 <= phi x)%Zapply phibis_aux_pos. Qed. Hint Resolve phi_nonneg : zarith.x:int31(0 <= phibis_aux size x)%Zforall x : int31, (0 <= phi x < 2 ^ Z.of_nat size)%Zforall x : int31, (0 <= phi x < 2 ^ Z.of_nat size)%Zx:int31(0 <= phi x < 2 ^ Z.of_nat size)%Zx:int31(phi x < 2 ^ Z.of_nat size)%Zx:int31(phibis_aux size x < 2 ^ Z.of_nat size)%Zapply phibis_aux_bounded; auto. Qed.x:int31(phibis_aux size (nshiftr x (size - size)) < 2 ^ Z.of_nat size)%Zforall (n : nat) (x : int31), firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Zforall (n : nat) (x : int31), firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Zforall x : int31, firstr (nshiftr x 0) = D1 -> (2 ^ Z.of_nat 0 <= phibis_aux 1 x)%Zn:natIHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Zforall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Zx:int31H:firstr (nshiftr x 0) = D1(2 ^ Z.of_nat 0 <= phibis_aux 1 x)%Zn:natIHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Zforall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Zx:int31H:firstr x = D1(1 <= phibis_aux 1 x)%Zn:natIHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Zforall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Zx:int31H:firstr x = D1(1 <= match firstr x with | D0 => Z.double | D1 => Z.succ_double end 0)%Zn:natIHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Zforall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Zn:natIHn:forall x : int31, firstr (nshiftr x n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x)%Zforall x : int31, firstr (nshiftr x (S n)) = D1 -> (2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1(2 ^ Z.of_nat (S n) <= phibis_aux (S (S n)) x)%Zn, m:natHeqm:m = S nIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux m x0)%Zx:int31H:firstr (nshiftr x m) = D1(2 ^ Z.of_nat m <= phibis_aux (S m) x)%Zn, m:natHeqm:m = S nIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux m x0)%Zx:int31H: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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H: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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H: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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1(2 ^ Z.of_nat n <= phibis_aux (S n) (shiftr x))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1firstr (nshiftr (shiftr x) n) = D1n:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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)))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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))%Zn:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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)))%Zrewrite Z.succ_double_spec; omega. Qed.n:natIHn:forall x0 : int31, firstr (nshiftr x0 n) = D1 -> (2 ^ Z.of_nat n <= phibis_aux (S n) x0)%Zx:int31H:firstr (nshiftr x (S n)) = D1H0:(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)))%Zforall x : int31, firstl x = D1 -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Zforall x : int31, firstl x = D1 -> (2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Zx:int31H:firstl x = D1(2 ^ Z.of_nat (Init.Nat.pred size) <= phi x)%Zx:int31H: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)%Zx:int31H: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)%Zrewrite phibis_aux_equiv; auto. Qed.x:int31H: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
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 = yunfold EqShiftL; intros; unfold nshiftl; simpl; split; auto. Qed.forall x y : int31, EqShiftL 0 x y <-> x = yforall (k : nat) (x y : int31), size <= k -> EqShiftL k x yred; intros; rewrite 2 nshiftl_above_size; auto. Qed.forall (k : nat) (x y : int31), size <= k -> EqShiftL k x yforall (k k' : nat) (x y : int31), k <= k' -> EqShiftL k x y -> EqShiftL k' x yforall (k k' : nat) (x y : int31), k <= k' -> EqShiftL k x y -> EqShiftL k' x yk, k':natx, y:int31H:k <= k'H0:nshiftl x k = nshiftl y knshiftl x k' = nshiftl y k'k, k':natx, y:int31H:k <= k'H0:nshiftl x k = nshiftl y knshiftl x (k' - k + k) = nshiftl y (k' - k + k)k, k':natx, y:int31H:k <= k'H0:nshiftl x k = nshiftl y kn:natHeqn:n = (k' - k)%natnshiftl x (n + k) = nshiftl y (n + k)k:natx, y:int31H0:nshiftl x k = nshiftl y kn:natnshiftl x (n + k) = nshiftl y (n + k)f_equal; auto. Qed.k:natx, y:int31H0:nshiftl x k = nshiftl y kn:natIHn:nshiftl x (n + k) = nshiftl y (n + k)shiftl (nshiftl x (n + k)) = shiftl (nshiftl y (n + k))forall (k : nat) (x y : int31), k < size -> EqShiftL k x y -> firstr x = firstr yforall (k : nat) (x y : int31), k < size -> EqShiftL k x y -> firstr x = firstr yk:natx, y:int31H:k < sizeH0:EqShiftL k x yfirstr x = firstr yk:natx, y:int31H:k < sizeH0:EqShiftL k x yfirstl (nshiftl x (Init.Nat.pred size)) = firstl (nshiftl y (Init.Nat.pred size))k:natx, y:int31H:k < sizeH0:EqShiftL k x ynshiftl x (Init.Nat.pred size) = nshiftl y (Init.Nat.pred size)k:natx, y:int31H:k < sizeH0:EqShiftL k x yk <= Init.Nat.pred sizeauto with arith. Qed.k:natx, y:int31H:k < sizeH0:EqShiftL k x yk <= Init.Nat.pred 31forall (k : nat) (x y : int31), EqShiftL k (twice x) (twice y) <-> EqShiftL (S k) x yforall (k : nat) (x y : int31), EqShiftL k (twice x) (twice y) <-> EqShiftL (S k) x yrewrite 2 nshiftl_S_tail; split; auto. Qed.k:natx, y:int31nshiftl (twice x) k = nshiftl (twice y) k <-> nshiftl x (S k) = nshiftl y (S k)
Lower (=rightmost) bits comes first.
Definition i2l := recrbis _ nil (fun d _ rec => d::rec).forall x : int31, length (i2l x) = sizeintros; 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, length (i2l x) = sizeforall x : int31, l2i (i2l x) = xdestruct x; compute; auto. Qed.forall x : int31, l2i (i2l x) = xforall (x : int31) (d : digits), i2l (sneakr d x) = tl (i2l x) ++ d :: nildestruct x; compute; auto. Qed.forall (x : int31) (d : digits), i2l (sneakr d x) = tl (i2l x) ++ d :: nilforall (x : int31) (d : digits), i2l (sneakl d x) = d :: removelast (i2l x)destruct x; compute; auto. Qed.forall (x : int31) (d : digits), i2l (sneakl d x) = d :: removelast (i2l x)forall l : list digits, length l = size -> i2l (l2i l) = lforall l : list digits, length l = size -> i2l (l2i l) = ld, 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:digitsl:list digitslength (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.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:digitslength (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 :: ...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:natIHn: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:int31H:0 <= sizei2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)n:natIHn: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:int31H:0 <= sizefirstn (size - 0) (i2l x) = i2l xx:int31H:0 <= sizeH0:firstn (size - 0) (i2l x) = i2l xi2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)n:natIHn: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:int31H:0 <= sizefirstn (length (i2l x)) (i2l x) = i2l xx:int31H:0 <= sizeH0:firstn (size - 0) (i2l x) = i2l xi2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)n:natIHn: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:int31H:0 <= sizeH0:firstn (size - 0) (i2l x) = i2l xi2l (nshiftl x 0) = cstlist digits D0 0 ++ firstn (size - 0) (i2l x)n:natIHn: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:int31H:0 <= sizei2l (nshiftl x 0) = cstlist digits D0 0 ++ i2l xn:natIHn: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:natIHn: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:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizei2l (nshiftl x (S n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizei2l (shiftl (nshiftl x n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeD0 :: removelast (i2l (nshiftl x n)) = cstlist digits D0 (S n) ++ firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeD0 :: removelast (i2l (nshiftl x n)) = (D0 :: cstlist digits D0 n) ++ firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeremovelast (i2l (nshiftl x n)) = cstlist digits D0 n ++ firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeremovelast (cstlist digits D0 n ++ firstn (size - n) (i2l x)) = cstlist digits D0 n ++ firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizecstlist digits D0 n ++ removelast (firstn (size - n) (i2l x)) = cstlist digits D0 n ++ firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizefirstn (size - n) (i2l x) <> niln:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeremovelast (firstn (size - n) (i2l x)) = firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizefirstn (size - n) (i2l x) <> niln:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeremovelast (firstn (S (size - S n)) (i2l x)) = firstn (size - S n) (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizefirstn (size - n) (i2l x) <> niln:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizesize - S n < length (i2l x)n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizefirstn (size - n) (i2l x) <> niln:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizefirstn (size - n) (i2l x) <> niln:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizelength (firstn (size - n) (i2l x)) = Init.Nat.min (size - n) (length (i2l x)) -> firstn (size - n) (i2l x) <> niln:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizelength (firstn (size - n) (i2l x)) = Init.Nat.min (size - n) size -> firstn (size - n) (i2l x) <> niln:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeH0:length (firstn (size - n) (i2l x)) = Init.Nat.min (size - n) sizeH1:firstn (size - n) (i2l x) = nilFalsen:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeH0:length nil = Init.Nat.min (size - n) sizeH1:firstn (size - n) (i2l x) = nilFalsen:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeH0:length nil = (size - n)%natH1:firstn (size - n) (i2l x) = nilFalseomega. Qed.n:natIHn:forall x0 : int31, n <= size -> i2l (nshiftl x0 n) = cstlist digits D0 n ++ firstn (size - n) (i2l x0)x:int31H:S n <= sizeH0:0%nat = (size - n)%natH1:firstn (size - n) (i2l x) = nilFalse
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:natx, y:int31EqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hle:size <= kEqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hlt:k < sizeEqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hle:size <= kH:EqShiftL k x yfirstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hle:size <= kH:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)EqShiftL k x yk:natx, y:int31Hlt:k < sizeEqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hle:size <= kH:EqShiftL k x yfirstn 0 (i2l x) = firstn 0 (i2l y)k:natx, y:int31Hle:size <= kH:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)EqShiftL k x yk:natx, y:int31Hlt:k < sizeEqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hle:size <= kH:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)EqShiftL k x yk:natx, y:int31Hlt:k < sizeEqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hlt:k < sizeEqShiftL k x y <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hlt:k < sizenshiftl x k = nshiftl y k <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hlt:k < sizeH:k <= sizenshiftl x k = nshiftl y k <-> firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hlt:k < sizeH:k <= sizeH0:nshiftl x k = nshiftl y kfirstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)nshiftl x k = nshiftl y kk:natx, y:int31Hlt:k < sizeH:k <= sizeH0:nshiftl x k = nshiftl y kH1:i2l (nshiftl x k) = i2l (nshiftl y k)firstn (size - k) (i2l x) = firstn (size - k) (i2l y)k:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)nshiftl x k = nshiftl y kk:natx, y:int31Hlt:k < sizeH:k <= sizeH0:nshiftl x k = nshiftl y kH1: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:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)nshiftl x k = nshiftl y kk:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)nshiftl x k = nshiftl y kk:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)i2l (nshiftl x k) = i2l (nshiftl y k)k:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)H1:i2l (nshiftl x k) = i2l (nshiftl y k)nshiftl x k = nshiftl y kk:natx, y:int31Hlt:k < sizeH:k <= sizeH0: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:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)H1:i2l (nshiftl x k) = i2l (nshiftl y k)nshiftl x k = nshiftl y kk:natx, y:int31Hlt:k < sizeH:k <= sizeH0:firstn (size - k) (i2l x) = firstn (size - k) (i2l y)H1:i2l (nshiftl x k) = i2l (nshiftl y k)nshiftl x k = nshiftl y kf_equal; auto. Qed.k:natx, y:int31Hlt:k < sizeH:k <= sizeH0: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))
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 yforall (k : nat) (x y : int31), EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x yk:natx, y:int31EqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x yk:natx, y:int31Hle:size <= kEqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x yk:natx, y:int31Hlt:k < sizeEqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x yk:natx, y:int31Hlt:k < sizeEqShiftL k (twice_plus_one x) (twice_plus_one y) <-> EqShiftL (S k) x yk:natx, y:int31Hlt:k < sizefirstn (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:natx, y:int31Hlt:k < sizefirstn (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:natx, y:int31Hlt:k < sizefirstn (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:natx, y:int31Hlt:k < sizefirstn (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:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natfirstn (S n) (D1 :: removelast (i2l x)) = firstn (S n) (D1 :: removelast (i2l y)) <-> firstn n (i2l x) = firstn n (i2l y)k:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xfirstn (S n) (D1 :: removelast lx) = firstn (S n) (D1 :: removelast (i2l y)) <-> firstn n lx = firstn n (i2l y)k:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yfirstn (S n) (D1 :: removelast lx) = firstn (S n) (D1 :: removelast ly) <-> firstn n lx = firstn n lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yD1 :: firstn n (removelast lx) = D1 :: firstn n (removelast ly) <-> firstn n lx = firstn n lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yD1 :: firstn n lx = D1 :: firstn n ly <-> firstn n lx = firstn n lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lxk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yH:D1 :: firstn n lx = D1 :: firstn n lyfirstn n lx = firstn n lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yH:firstn n lx = firstn n lyD1 :: firstn n lx = D1 :: firstn n lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lxk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yH:firstn n lx = firstn n lyD1 :: firstn n lx = D1 :: firstn n lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lxk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lyk:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lxsubst lx n; rewrite i2l_length; omega. Qed.k:natx, y:int31Hlt:k < sizen:natHeqn:n = (size - S k)%natlx:list digitsHeqlx:lx = i2l xly:list digitsHeqly:ly = i2l yn < length lxforall (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:natx, y:int31H:EqShiftL k x yEqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHle:size <= S kEqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeEqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeEqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D0EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D0EqShiftL k (twice (shiftr x)) (twice (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D0EqShiftL k (sneakl (firstr x) (shiftr x)) (sneakl (firstr x) (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D0EqShiftL k x (sneakl (firstr x) (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D0EqShiftL k x (sneakl (firstr y) (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D0k < sizek:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D0k < sizek:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL (S k) (shiftr x) (shiftr y)k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL k (twice_plus_one (shiftr x)) (twice_plus_one (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL k (sneakl (firstr x) (shiftr x)) (sneakl (firstr x) (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL k x (sneakl (firstr x) (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1EqShiftL k x (sneakl (firstr y) (shiftr y))k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1k < sizeomega. Qed.k:natx, y:int31H:EqShiftL k x yHlt:S k < sizeH0:firstr x = D1k < sizeforall (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:natx, y:int31H:0 <= sizeH0:k = S sizeH1:EqShiftL k x yEqShiftL k (incrbis_aux 0 x) (incrbis_aux 0 y)n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yEqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yEqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x ye:k = sizeEqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeEqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeEqShiftL k (incrbis_aux (S n) x) (incrbis_aux (S n) y)n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeEqShiftL 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)) endn:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeEqShiftL 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)) endn:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeH2:firstr y = D0EqShiftL k (sneakl D1 (shiftr x)) (sneakl D1 (shiftr y))n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeH2:firstr y = D1EqShiftL k (sneakl D0 (incrbis_aux n (shiftr x))) (sneakl D0 (incrbis_aux n (shiftr y)))n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeH2:firstr y = D0EqShiftL (S k) (shiftr x) (shiftr y)n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeH2:firstr y = D1EqShiftL k (sneakl D0 (incrbis_aux n (shiftr x))) (sneakl D0 (incrbis_aux n (shiftr y)))n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeH2:firstr y = D1EqShiftL k (sneakl D0 (incrbis_aux n (shiftr x))) (sneakl D0 (incrbis_aux n (shiftr y)))n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeH2:firstr y = D1EqShiftL (S k) (incrbis_aux n (shiftr x)) (incrbis_aux n (shiftr y))apply EqShiftL_shiftr; auto. Qed.n:natIHn: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:natx, y:int31H:S n <= sizeH0:S (n + k) = S sizeH1:EqShiftL k x yn0:k <> sizeH2:firstr y = D1EqShiftL (S k) (shiftr x) (shiftr y)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:int31H:EqShiftL 1 x yEqShiftL 1 (incr x) (incr y)apply EqShiftL_incrbis; auto. Qed. End EqShiftL.x, y:int31H:EqShiftL 1 x yEqShiftL 1 (incrbis_aux size x) (incrbis_aux size y)
forall x : int31, incr (twice_plus_one x) = twice (incr x)forall x : int31, incr (twice_plus_one x) = twice (incr x)x:int31incr (twice_plus_one x) = twice (incr x)x:int31twice (incr (shiftr (twice_plus_one x))) = twice (incr x)red; destruct x; simpl; auto. Qed.x:int31EqShiftL 1 (shiftr (twice_plus_one x)) xforall x : int31, firstr (incr x) <> firstr xforall x : int31, firstr (incr x) <> firstr xx:int31firstr (incr x) <> firstr xx:int31H:firstr x = D0firstr (incr x) <> D0x:int31H:firstr x = D1firstr (incr x) <> D1x:int31H:firstr x = D0firstr (twice_plus_one (shiftr x)) <> D0x:int31H:firstr x = D1firstr (incr x) <> D1x:int31H:firstr x = D1firstr (incr x) <> D1destruct (incr (shiftr x)); simpl; discriminate. Qed.x:int31H:firstr x = D1firstr (twice (incr (shiftr x))) <> D1forall x y : int31, incr x = twice_plus_one y -> x = twice yforall x y : int31, incr x = twice_plus_one y -> x = twice yx, y:int31H:incr x = twice_plus_one yx = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = truex = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falsex = twice yx, y:int31H:incr 0 = twice_plus_one yH0:iszero x = true0 = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falsex = twice yx, y:int31H:1 = twice_plus_one yH0:iszero x = true0 = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falsex = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falsex = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falseH1:firstr x = D0x = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falseH1:firstr x = D1x = twice yx, y:int31H:twice_plus_one (shiftr x) = twice_plus_one yH0:iszero x = falseH1:firstr x = D0x = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falseH1:firstr x = D1x = twice yd, 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:digitsH: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 D1H1:d29 = D0I31 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 D0x, y:int31H:incr x = twice_plus_one yH0:iszero x = falseH1:firstr x = D1x = twice yx, y:int31H:incr x = twice_plus_one yH0:iszero x = falseH1:firstr x = D1x = twice yrewrite H1, H; destruct y; simpl; auto. Qed.x, y:int31H:incr x = twice_plus_one yH0:iszero x = falseH1:firstr x = D1firstr (incr x) = firstr x
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:positiveincr (complement_negative (Pos.pred_double p)) = twice_plus_one (incr (complement_negative p))p:positiveIHp: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:positiveIHp: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 = 2147483647p:positiveIHp: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 = 2147483647p:positiveIHp: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 = 2147483647p:positiveIHp: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 = 2147483647auto. Qed.incr 2147483646 = 2147483647forall z : Z, phi_inv (Z.double z) = twice (phi_inv z)forall z : Z, phi_inv (Z.double z) = twice (phi_inv z)rewrite incr_twice_plus_one; auto. Qed.p:positiveincr (twice_plus_one (complement_negative p)) = twice (incr (complement_negative p))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:positivephi_inv (Z.succ (Z.pos p)) = incr (phi_inv (Z.pos p))p:positivephi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))p:positivephi_inv (Z.succ (Z.pos p)) = incr (phi_inv (Z.pos p))p:positivephi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))p:positivephi_inv_positive (p + 1) = incr (phi_inv_positive p)p:positivephi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))p:positiveIHp: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:positiveIHp:phi_inv_positive (p + 1) = incr (phi_inv_positive p)twice_plus_one (phi_inv_positive p) = incr (twice (phi_inv_positive p))p:positivephi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))p:positiveIHp:phi_inv_positive (p + 1) = incr (phi_inv_positive p)twice_plus_one (phi_inv_positive p) = incr (twice (phi_inv_positive p))p:positivephi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))p:positivephi_inv (Z.succ (Z.neg p)) = incr (phi_inv (Z.neg p))p:positivephi_inv (Z.pos_sub 1 p) = incr (incr (complement_negative p))p:positiveincr (twice_plus_one (complement_negative p)) = incr (incr (twice (complement_negative p)))p:positiveincr (complement_negative (Pos.pred_double p)) = incr (incr (twice_plus_one (complement_negative p)))p:positiveincr (complement_negative (Pos.pred_double p)) = incr (incr (twice_plus_one (complement_negative p)))p:positivecomplement_negative (Pos.pred_double p) = incr (twice_plus_one (complement_negative p))p:positivecomplement_negative (Pos.pred_double p) = twice (incr (complement_negative p))p:positiveIHp: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:positiveIHp: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:positiveIHp: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)))rewrite incr_twice_plus_one; auto. Qed.p:positiveIHp:complement_negative (Pos.pred_double p) = twice (incr (complement_negative p))complement_negative (Pos.pred_double p) = incr (twice_plus_one (complement_negative p))
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:natIHn: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:int31H:0 <= sizephi_inv (phibis_aux 0 (nshiftr x 31)) = nshiftr x 31n:natIHn: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:natIHn: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:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizephi_inv (phibis_aux (S n) (nshiftr x (size - S n))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizephi_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:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeshiftr (nshiftr x (size - S n)) = nshiftr x (size - n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0: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:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0: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:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0: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:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D0phi_inv (Z.double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D0twice (phi_inv (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D0twice (nshiftr x (size - n)) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D0twice (shiftr (nshiftr x (size - S n))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizey:(fun _ : nat => int31) (size - S n)%natHeqy:y = nshiftr x (size - S n)H0:shiftr y = nshiftr x (size - n)H1:firstr y = D0twice (shiftr y) = yn:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1phi_inv (Z.succ_double (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1twice_plus_one (phi_inv (phibis_aux n (nshiftr x (size - n)))) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1twice_plus_one (nshiftr x (size - n)) = nshiftr x (size - S n)n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizeH0:shiftr (nshiftr x (size - S n)) = nshiftr x (size - n)H1:firstr (nshiftr x (size - S n)) = D1twice_plus_one (shiftr (nshiftr x (size - S n))) = nshiftr x (size - S n)destruct y; simpl in H1; rewrite H1; auto. Qed.n:natIHn:forall x0 : int31, n <= size -> phi_inv (phibis_aux n (nshiftr x0 (size - n))) = nshiftr x0 (size - n)x:int31H:S n <= sizey:(fun _ : nat => int31) (size - S n)%natHeqy:y = nshiftr x (size - S n)H0:shiftr y = nshiftr x (size - n)H1:firstr y = D1twice_plus_one (shiftr y) = yforall x : int31, phi_inv (phi x) = xforall x : int31, phi_inv (phi x) = xx:int31phi_inv (phi x) = xx:int31phi_inv (phibis_aux size x) = xapply phi_inv_phi_aux; auto. Qed.x:int31phi_inv (phibis_aux size (nshiftr x (size - size))) = nshiftr x (size - size)
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.
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 = 0forall (n : nat) (p : positive), nshiftr (snd (p2ibis n p)) n = 0forall p : positive, nshiftr (snd (p2ibis 0 p)) 0 = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0forall p : positive, nshiftr (snd (p2ibis (S n) p)) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0forall p : positive, nshiftr (snd (p2ibis (S n) p)) (S n) = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (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) = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice_plus_one i))) (S n) = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiveIHn:nshiftr (snd (p2ibis n p)) n = 0nshiftr (snd (let (r, i) := p2ibis n p in (r, twice_plus_one i))) (S n) = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr (snd (n0, i)) n = 0nshiftr (snd (n0, twice_plus_one i)) (S n) = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0nshiftr (twice_plus_one i) (S n) = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0nshiftr (shiftr (twice_plus_one i)) n = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hle:size <= nnshiftr (shiftr (twice_plus_one i)) n = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizenshiftr (shiftr (twice_plus_one i)) n = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizenshiftr (shiftr (twice_plus_one i)) n = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizeH:firstl i = D0nshiftr (shiftr (twice_plus_one i)) n = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizeH:firstl i = D0i = shiftr (twice_plus_one i)n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Nd, 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:digitsIHn: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 = 0Hlt:n < sizeH:d = D0I31 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 d29n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natIHn:forall p0 : positive, nshiftr (snd (p2ibis n p0)) n = 0p:positivenshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiveIHn:nshiftr (snd (p2ibis n p)) n = 0nshiftr (snd (let (r, i) := p2ibis n p in (r, twice i))) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0nshiftr (twice i) (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0nshiftr (shiftr (twice i)) n = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hle:size <= nnshiftr (shiftr (twice i)) n = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizenshiftr (shiftr (twice i)) n = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizenshiftr (shiftr (twice i)) n = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizeH:firstl i = D0nshiftr (shiftr (twice i)) n = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natp:positiven0:Ni:int31IHn:nshiftr i n = 0Hlt:n < sizeH:firstl i = D0i = shiftr (twice i)n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr In (S n) = 0n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr (shiftr In) n = 0apply nshiftr_n_0. Qed. Local Open Scope Z_scope.n:natIHn:forall p : positive, nshiftr (snd (p2ibis n p)) n = 0nshiftr 0 n = 0forall (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:positiveH:(0 <= size)%natZ.pos p = Z.of_N (fst (p2ibis 0 p)) * 2 ^ Z.of_nat 0 + phi (snd (p2ibis 0 p))n:natIHn: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:positiveH:(S n <= size)%natZ.pos p = Z.of_N (fst (p2ibis (S n) p)) * 2 ^ Z.of_nat (S n) + phi (snd (p2ibis (S n) p))n:natIHn: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:positiveH:(S n <= size)%natZ.pos p = Z.of_N (fst (p2ibis (S n) p)) * 2 ^ Z.of_nat (S n) + phi (snd (p2ibis (S n) p))n:natIHn: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:positiveH:(S n <= size)%natZ.pos p = Z.of_N (fst (p2ibis (S n) p)) * (2 * 2 ^ Z.of_nat n) + phi (snd (p2ibis (S n) p))n:natIHn: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:positiveH:(S n <= size)%natZ.pos p = Z.of_N (fst (p2ibis (S n) p)) * (2 ^ Z.of_nat n * 2) + phi (snd (p2ibis (S n) p))n:natIHn: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:positiveH:(S n <= size)%natH0:(n <= size)%natZ.pos p = Z.of_N (fst (p2ibis (S n) p)) * (2 ^ Z.of_nat n * 2) + phi (snd (p2ibis (S n) p))n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31Z.pos p~1 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice_plus_one i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int312 * Z.pos p + 1 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice_plus_one i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int312 * Z.pos p + 1 = Z.of_N r * (2 ^ Z.of_nat n * 2) + (2 * phi i + 1)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31firstl i = D0n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31firstl i = D0n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31Z.pos p~0 = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int312 * Z.pos p = Z.of_N r * (2 ^ Z.of_nat n * 2) + phi (twice i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int312 * Z.pos p = Z.of_N r * (2 ^ Z.of_nat n * 2) + Z.double (phi i)n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31firstl i = D0n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int312 * Z.pos p = Z.of_N r * (2 ^ Z.of_nat n * 2) + 2 * phi in:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31firstl i = D0apply (nshiftr_0_firstl n); auto; try omega. Qed.n:natp:positiver:Ni:int31IHn:Z.pos p = Z.of_N r * 2 ^ Z.of_nat n + phi iH:(S n <= size)%natH0:(n <= size)%natH1:nshiftr i n = 0%int31firstl i = D0
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:natIHn: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:positiveH:(0 <= size)%natEqShiftL (size - 0) (phi_inv_positive p) (snd (p2ibis 0 p))n:natIHn: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:natIHn: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))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.n:natIHn:forall p0 : positive, (n <= size)%nat -> EqShiftL (size - n) (phi_inv_positive p0) (snd (p2ibis n p0))p:positiveH:(S n <= size)%natEqShiftL (size - S n) (phi_inv_positive p) (snd (p2ibis (S n) p))
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 sizeforall p : positive, phi (phi_inv_positive p) = Z.pos p mod 2 ^ Z.of_nat sizep:positivephi (phi_inv_positive p) = Z.pos p mod 2 ^ Z.of_nat sizep:positivephi (snd (p2ibis size p)) = Z.pos p mod 2 ^ Z.of_nat sizep:positivesnd (p2ibis size p) = phi_inv_positive pp:positivephi (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 sizep:positivesnd (p2ibis size p) = phi_inv_positive pp:positivephi (snd (p2ibis size p)) = phi (snd (p2ibis size p)) mod 2 ^ Z.of_nat sizep:positive2 ^ Z.of_nat size > 0p:positivesnd (p2ibis size p) = phi_inv_positive pp:positive0 <= phi (snd (p2ibis size p)) < 2 ^ Z.of_nat sizep:positive2 ^ Z.of_nat size > 0p:positivesnd (p2ibis size p) = phi_inv_positive pp:positive2 ^ Z.of_nat size > 0p:positivesnd (p2ibis size p) = phi_inv_positive pp:positivesnd (p2ibis size p) = phi_inv_positive pp:positivephi_inv_positive p = snd (p2ibis size p)apply (phi_inv_positive_p2ibis size p); auto. Qed.p:positiveEqShiftL 0 (phi_inv_positive p) (snd (p2ibis size p))
Moreover, p2ibis is also related with p2i and hence with
positive_to_int31.
forall x : int31, firstl x = D0 -> (Twon * x)%int31 = twice xforall x : int31, firstl x = D0 -> (Twon * x)%int31 = twice xx:int31H:firstl x = D0(Twon * x)%int31 = twice xrewrite <- Z.double_spec, <- phi_twice_firstl, phi_inv_phi; auto. Qed.x:int31H:firstl x = D0phi_inv (phi Twon * phi x) = twice xforall x : int31, firstl x = D0 -> (Twon * x + In)%int31 = twice_plus_one xforall x : int31, firstl x = D0 -> (Twon * x + In)%int31 = twice_plus_one xx:int31H:firstl x = D0(Twon * x + In)%int31 = twice_plus_one xx:int31H:firstl x = D0(twice x + In)%int31 = twice_plus_one xrewrite phi_twice_firstl, <- Z.succ_double_spec, <- phi_twice_plus_one_firstl, phi_inv_phi; auto. Qed.x:int31H:firstl x = D0phi_inv (phi (twice x) + phi In) = twice_plus_one xforall (n : nat) (p : positive), (n <= size)%nat -> p2i n p = p2ibis n pforall (n : nat) (p : positive), (n <= size)%nat -> p2i n p = p2ibis n pn:natIHn:forall p0 : positive, (n <= size)%nat -> p2i n p0 = p2ibis n p0p:positiveH:(S n <= size)%natmatch 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) endn:natp:positiven0:Ni:int31IHn:(n <= size)%nat -> p2i n p = (n0, i)H:(S n <= size)%natH0:nshiftr i n = 0%int31(Twon * i + In)%int31 = twice_plus_one in:natp:positiven0:Ni:int31IHn:(n <= size)%nat -> p2i n p = (n0, i)H:(S n <= size)%natH0:nshiftr i n = 0%int31(Twon * i)%int31 = twice in:natp:positiven0:Ni:int31IHn:(n <= size)%nat -> p2i n p = (n0, i)H:(S n <= size)%natH0:nshiftr i n = 0%int31firstl i = D0n:natp:positiven0:Ni:int31IHn:(n <= size)%nat -> p2i n p = (n0, i)H:(S n <= size)%natH0:nshiftr i n = 0%int31(Twon * i)%int31 = twice in:natp:positiven0:Ni:int31IHn:(n <= size)%nat -> p2i n p = (n0, i)H:(S n <= size)%natH0:nshiftr i n = 0%int31(Twon * i)%int31 = twice iapply (nshiftr_0_firstl n); auto; omega. Qed.n:natp:positiven0:Ni:int31IHn:(n <= size)%nat -> p2i n p = (n0, i)H:(S n <= size)%natH0:nshiftr i n = 0%int31firstl i = D0forall p : positive, snd (positive_to_int31 p) = phi_inv_positive pforall p : positive, snd (positive_to_int31 p) = phi_inv_positive pp:positivesnd (p2i size p) = phi_inv_positive pp:positivesnd (p2ibis size p) = phi_inv_positive pp:positivephi_inv_positive p = snd (p2ibis size p)apply (phi_inv_positive_p2ibis size); auto. Qed.p:positiveEqShiftL 0 (phi_inv_positive p) (snd (p2ibis size 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 (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))apply p2ibis_spec; auto. Qed.p:positiveZ.pos p = Z.of_N (fst (p2ibis size p)) * 2 ^ Z.of_nat size + phi (snd (p2ibis size p))
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 sizeforall x : int31, phi (twice x) = Z.double (phi x) mod 2 ^ Z.of_nat sizex:int31phi (twice x) = Z.double (phi x) mod 2 ^ Z.of_nat sizex:int31phi (twice (phi_inv (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat sizex:int31phi (phi_inv (Z.double (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat sizex:int310 <= Z.double (phi x)x:int31H:0 <= Z.double (phi x)phi (phi_inv (Z.double (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat sizex:int31H:0 <= Z.double (phi x)phi (phi_inv (Z.double (phi x))) = Z.double (phi x) mod 2 ^ Z.of_nat sizex:int31H:0 <= 0phi (phi_inv 0) = 0 mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.pos pphi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.pos pphi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizecompute in H; elim H; auto. Qed.x:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizeforall x : int31, phi (twice_plus_one x) = Z.succ_double (phi x) mod 2 ^ Z.of_nat sizeforall x : int31, phi (twice_plus_one x) = Z.succ_double (phi x) mod 2 ^ Z.of_nat sizex:int31phi (twice_plus_one x) = Z.succ_double (phi x) mod 2 ^ Z.of_nat sizex:int31phi (twice_plus_one (phi_inv (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat sizex:int31phi (phi_inv (Z.succ_double (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat sizex:int310 <= Z.succ_double (phi x)x:int31H:0 <= Z.succ_double (phi x)phi (phi_inv (Z.succ_double (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat sizex:int31H:0 <= Z.succ_double (phi x)phi (phi_inv (Z.succ_double (phi x))) = Z.succ_double (phi x) mod 2 ^ Z.of_nat sizex:int31H:0 <= 0phi (phi_inv 0) = 0 mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.pos pphi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.pos pphi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizecompute in H; elim H; auto. Qed.x:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizeforall x : int31, phi (incr x) = Z.succ (phi x) mod 2 ^ Z.of_nat sizeforall x : int31, phi (incr x) = Z.succ (phi x) mod 2 ^ Z.of_nat sizex:int31phi (incr x) = Z.succ (phi x) mod 2 ^ Z.of_nat sizex:int31phi (incr (phi_inv (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat sizex:int31phi (phi_inv (Z.succ (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat sizex:int310 <= Z.succ (phi x)x:int31H:0 <= Z.succ (phi x)phi (phi_inv (Z.succ (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat sizex:int31H:0 <= Z.succ (phi x)phi (phi_inv (Z.succ (phi x))) = Z.succ (phi x) mod 2 ^ Z.of_nat sizex:int31H:0 <= 0phi (phi_inv 0) = 0 mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.pos pphi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.pos pphi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizex:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizecompute in H; elim H; auto. Qed.x:int31p:positiveH:0 <= Z.neg pphi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat size
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 sizeforall p : positive, phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~1)) = Z.neg p~1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (twice (complement_negative p))) = Z.neg p~1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveIHp:Z.succ (phi (complement_negative p)) mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (twice (complement_negative p))) = Z.neg p~1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveIHp:Z.succ (phi (complement_negative p)) mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat sizeZ.succ_double (phi (complement_negative p)) mod 2 ^ Z.of_nat size = Z.neg p~1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveq:ZHeqq:q = phi (complement_negative p)IHp:Z.succ q mod 2 ^ Z.of_nat size = Z.neg p mod 2 ^ Z.of_nat sizeZ.succ_double q mod 2 ^ Z.of_nat size = Z.neg p~1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveq:ZHeqq: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 sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveq:ZHeqq: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 sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveq:ZHeqq: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 sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (complement_negative p~0)) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizephi (incr (twice_plus_one (complement_negative p))) = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveIHp:phi (incr (complement_negative p)) = Z.neg p mod 2 ^ Z.of_nat sizeZ.double (phi (incr (complement_negative p))) mod 2 ^ Z.of_nat size = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizep:positiveq:ZHeqq:q = phi (incr (complement_negative p))IHp:q = Z.neg p mod 2 ^ Z.of_nat sizeZ.double q mod 2 ^ Z.of_nat size = Z.neg p~0 mod 2 ^ Z.of_nat sizephi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizesimpl; auto. Qed.phi (incr (complement_negative 1)) = -1 mod 2 ^ Z.of_nat sizeforall z : Z, phi (phi_inv z) = z mod 2 ^ Z.of_nat sizeforall z : Z, phi (phi_inv z) = z mod 2 ^ Z.of_nat sizephi (phi_inv 0) = 0 mod 2 ^ Z.of_nat sizep:positivephi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizep:positivephi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizep:positivephi (phi_inv (Z.pos p)) = Z.pos p mod 2 ^ Z.of_nat sizep:positivephi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizeapply 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)).p:positivephi (phi_inv (Z.neg p)) = Z.neg p mod 2 ^ Z.of_nat sizewB > 0auto 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).wB > 0[|31|] = 31reflexivity. Qed.[|31|] = 311 < 31auto with zarith. Qed.1 < 31[|0|] = 0reflexivity. Qed.[|0|] = 0[|1|] = 1reflexivity. Qed.[|1|] = 1[|Tn|] = wB - 1reflexivity. Qed.[|Tn|] = wB - 1forall x y : int31, (x ?= y)%int31 = ([|x|] ?= [|y|])reflexivity. Qed.forall x y : int31, (x ?= y)%int31 = ([|x|] ?= [|y|])
Addition
forall x y : int31, [+|x +c y|] = [|x|] + [|y|]forall x y : int31, [+|x +c y|] = [|x|] + [|y|]x, y:int31match 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:int31H:0 <= [|x|] < wBH0:0 <= [|y|] < wBmatch 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:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBmatch 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 + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wB((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Ymatch 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 + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y) mod wB <> X + Y1 * wB + (X + Y) mod wB = X + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Ymatch 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 + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y) mod wB <> X + Yl:X + Y < wB1 * wB + (X + Y) mod wB = X + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y) mod wB <> X + Yl:wB <= X + Y1 * wB + (X + Y) mod wB = X + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Ymatch 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 + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y) mod wB <> X + Yl:wB <= X + Y1 * wB + (X + Y) mod wB = X + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Ymatch 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 + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y) mod wB <> X + Yl:wB <= X + Y1 * wB + (X + Y + -1 * wB) mod wB = X + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Ymatch 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 + Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + Ymatch 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 + Ydestruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y) mod wB ?= X + Y) <> Eq -> [+|C1 (phi_inv (X + Y))|] = X + YHeq:((X + Y) mod wB ?= X + Y) = Eq -> (X + Y) mod wB = X + Ymatch 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 + Yforall x : int31, [+|x +c 1|] = [|x|] + 1intros; apply spec_add_c. Qed.forall x : int31, [+|x +c 1|] = [|x|] + 1forall x y : int31, [+|add31carryc x y|] = [|x|] + [|y|] + 1forall x y : int31, [+|add31carryc x y|] = [|x|] + [|y|] + 1x, y:int31[+|add31carryc x y|] = [|x|] + [|y|] + 1x, y:int31match 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|] + 1x, y:int31H:0 <= [|x|] < wBH0:0 <= [|y|] < wBmatch 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|] + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBmatch 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 + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wB((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1match 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 + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y + 1) mod wB <> X + Y + 11 * wB + (X + Y + 1) mod wB = X + Y + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1match 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 + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y + 1) mod wB <> X + Y + 1l:X + Y + 1 < wB1 * wB + (X + Y + 1) mod wB = X + Y + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y + 1) mod wB <> X + Y + 1l:wB <= X + Y + 11 * wB + (X + Y + 1) mod wB = X + Y + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1match 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 + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y + 1) mod wB <> X + Y + 1l:wB <= X + Y + 11 * wB + (X + Y + 1) mod wB = X + Y + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1match 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 + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X + Y + 1) mod wB <> X + Y + 1l:wB <= X + Y + 11 * wB + (X + Y + 1 + -1 * wB) mod wB = X + Y + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1match 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 + 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1match 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 + 1destruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X + Y + 1) mod wB ?= X + Y + 1) <> Eq -> [+|C1 (phi_inv (X + Y + 1))|] = X + Y + 1Heq:((X + Y + 1) mod wB ?= X + Y + 1) = Eq -> (X + Y + 1) mod wB = X + Y + 1match 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 + 1forall x y : int31, [|x + y|] = ([|x|] + [|y|]) mod wBintros; apply phi_phi_inv. Qed.forall x y : int31, [|x + y|] = ([|x|] + [|y|]) mod wBforall x y : int31, [|x + y + 1|] = ([|x|] + [|y|] + 1) mod wBforall x y : int31, [|x + y + 1|] = ([|x|] + [|y|] + 1) mod wBx, y:int31[|phi_inv ([|phi_inv ([|x|] + [|y|])|] + [|1|])|] = ([|x|] + [|y|] + 1) mod wBapply Zplus_mod_idemp_l. Qed.x, y:int31(([|x|] + [|y|]) mod wB + [|1|]) mod wB = ([|x|] + [|y|] + 1) mod wBforall x : int31, [|x + 1|] = ([|x|] + 1) mod wBintros; rewrite <- spec_1; apply spec_add. Qed.forall x : int31, [|x + 1|] = ([|x|] + 1) mod wB
Subtraction
forall x y : int31, [-|x -c y|] = [|x|] - [|y|]forall x y : int31, [-|x -c y|] = [|x|] - [|y|]x, y:int31match 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:int31match 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:int31H:0 <= [|x|] < wBH0:0 <= [|y|] < wBmatch 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:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBmatch 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 - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wB((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Ymatch 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 - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y) mod wB <> X - Y-1 * wB + (X - Y) mod wB = X - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Ymatch 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 - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y) mod wB <> X - Yl:X - Y < 0-1 * wB + (X - Y) mod wB = X - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y) mod wB <> X - Yl:0 <= X - Y-1 * wB + (X - Y) mod wB = X - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Ymatch 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 - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y) mod wB <> X - Yl:X - Y < 0-1 * wB + (X - Y + 1 * wB) mod wB = X - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y) mod wB <> X - Yl:0 <= X - Y-1 * wB + (X - Y) mod wB = X - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Ymatch 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 - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y) mod wB <> X - Yl:0 <= X - Y-1 * wB + (X - Y) mod wB = X - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Ymatch 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 - Yx, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - Ymatch 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 - Ydestruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y) mod wB ?= X - Y) <> Eq -> [-|C1 (phi_inv (X - Y))|] = X - YHeq:((X - Y) mod wB ?= X - Y) = Eq -> (X - Y) mod wB = X - Ymatch 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 - Yforall x y : int31, [-|sub31carryc x y|] = [|x|] - [|y|] - 1forall x y : int31, [-|sub31carryc x y|] = [|x|] - [|y|] - 1x, y:int31match 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|] - 1x, y:int31match 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|] - 1x, y:int31H:0 <= [|x|] < wBH0:0 <= [|y|] < wBmatch 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|] - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBmatch 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 - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wB((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1match 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 - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y - 1) mod wB <> X - Y - 1-1 * wB + (X - Y - 1) mod wB = X - Y - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1match 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 - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y - 1) mod wB <> X - Y - 1l:X - Y - 1 < 0-1 * wB + (X - Y - 1) mod wB = X - Y - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y - 1) mod wB <> X - Y - 1l:0 <= X - Y - 1-1 * wB + (X - Y - 1) mod wB = X - Y - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1match 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 - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y - 1) mod wB <> X - Y - 1l:X - Y - 1 < 0-1 * wB + (X - Y - 1 + 1 * wB) mod wB = X - Y - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y - 1) mod wB <> X - Y - 1l:0 <= X - Y - 1-1 * wB + (X - Y - 1) mod wB = X - Y - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1match 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 - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:(X - Y - 1) mod wB <> X - Y - 1l:0 <= X - Y - 1-1 * wB + (X - Y - 1) mod wB = X - Y - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1match 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 - 1x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1match 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 - 1destruct Z.compare; intros; [ rewrite phi_phi_inv; auto | now apply H1 | now apply H1]. Qed.x, y:int31X:ZH:0 <= X < wBY:ZH0:0 <= Y < wBH1:((X - Y - 1) mod wB ?= X - Y - 1) <> Eq -> [-|C1 (phi_inv (X - Y - 1))|] = X - Y - 1Heq:((X - Y - 1) mod wB ?= X - Y - 1) = Eq -> (X - Y - 1) mod wB = X - Y - 1match 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 - 1forall x y : int31, [|x - y|] = ([|x|] - [|y|]) mod wBintros; apply phi_phi_inv. Qed.forall x y : int31, [|x - y|] = ([|x|] - [|y|]) mod wBforall x y : int31, [|x - y - 1|] = ([|x|] - [|y|] - 1) mod wBforall x y : int31, [|x - y - 1|] = ([|x|] - [|y|] - 1) mod wBx, y:int31[|phi_inv ([|phi_inv ([|x|] - [|y|])|] - [|1|])|] = ([|x|] - [|y|] - 1) mod wBapply Zminus_mod_idemp_l. Qed.x, y:int31(([|x|] - [|y|]) mod wB - [|1|]) mod wB = ([|x|] - [|y|] - 1) mod wBforall x : int31, [-|0 -c x|] = - [|x|]intros; apply spec_sub_c. Qed.forall x : int31, [-|0 -c x|] = - [|x|]forall x : int31, [|0 - x|] = - [|x|] mod wBintros; apply phi_phi_inv. Qed.forall x : int31, [|0 - x|] = - [|x|] mod wBforall x : int31, [|0 - x - 1|] = wB - [|x|] - 1forall x : int31, [|0 - x - 1|] = wB - [|x|] - 1x:int31[|phi_inv ([|phi_inv ([|0|] - [|x|])|] - [|1|])|] = wB - [|x|] - 1x:int31(([|0|] - [|x|]) mod wB - [|1|]) mod wB = wB - [|x|] - 1x:int31((0 - [|x|]) mod wB - 1) mod wB = wB - [|x|] - 1x:int31((0 - [|x|] + 1 * wB) mod wB - 1) mod wB = wB - [|x|] - 1rewrite Zmod_small; generalize (phi_bounded x); lia. Qed.x:int31(0 - [|x|] + 1 * wB - 1) mod wB = wB - [|x|] - 1forall x : int31, [-|x -c 1|] = [|x|] - 1intros; apply spec_sub_c. Qed.forall x : int31, [-|x -c 1|] = [|x|] - 1forall x : int31, [|x - 1|] = ([|x|] - 1) mod wBintros; apply spec_sub. Qed.forall x : int31, [|x - 1|] = ([|x|] - 1) mod wB
Multiplication
forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:Z(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:Z(z / wB) mod wB = z / wB - z / wB / wB * wBz:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wB(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wB(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBz mod wB = z - z / wB * wBz:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBH0:z mod wB = z - z / wB * wB(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBH0:z mod wB = z - z / wB * wB(z / wB) mod wB * wB + z mod wB = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBH0:z mod wB = z - z / wB * wB(z / wB - z / wB / wB * wB) * wB + z mod wB = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBH0:z mod wB = z - z / wB * wB(z / wB - z / wB / wB * wB) * wB + (z - z / wB * wB) = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBH0:z mod wB = z - z / wB * wB- (z / wB / wB) * wB ^ 2 + z = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBH0:z mod wB = z - z / wB * wB- (z / (wB * wB)) * wB ^ 2 + z = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2z:ZH:(z / wB) mod wB = z / wB - z / wB / wB * wBH0:z mod wB = z - z / wB * wB- (z / (wB * wB)) * wB ^ 2 + (wB * wB * (z / (wB * wB)) + z mod (wB * wB)) = z mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||phi_inv2 x||] = x mod wB ^ 2destruct x; unfold zn2z_to_Z; rewrite ?phi_phi_inv; change base with wB; auto. Qed.H:forall z : Z, (z / wB) mod wB * wB + z mod wB = z mod wB ^ 2forall x : Z, [||match x with | 0 => W0 | _ => WW (phi_inv (x / base)) (phi_inv x) end||] = x mod wB ^ 2forall 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:int310 <= [|x|] * [|y|] < wB ^ 2x, y:int31H:0 <= [|x|] < wBH0:0 <= [|y|] < wB0 <= [|x|] * [|y|] < wB ^ 2auto using Z.mul_lt_mono_nonneg with zarith. Qed.x, y:int31H:0 <= [|x|] < wBH0:0 <= [|y|] < wB0 <= [|x|] * [|y|] < wB * wBforall x y : int31, [|x * y|] = ([|x|] * [|y|]) mod wBintros; apply phi_phi_inv. Qed.forall x y : int31, [|x * y|] = ([|x|] * [|y|]) mod wBforall x : int31, [||x *c x||] = [|x|] * [|x|]intros; apply spec_mul_c. Qed.forall x : int31, [||x *c x||] = [|x|] * [|x|]
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:int31H: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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBlet (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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0let (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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:Zphi2 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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:Zphi2 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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:phi2 a1 a2 = [|b|] * z + z0H6: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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * base + [|a2|] = [|b|] * z + z0H6: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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6: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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6: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:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zz * [|b|] + z0 = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zz * [|b|] + z0 = z mod wB * [|b|] + z0 /\ 0 <= z0 < [|b|]a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zz = z mod wBa1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z0 <= z < wBa1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z0 <= za1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zz < wBa1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zz < wBa1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zz * [|b|] < wB * [|b|]a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z[|b|] * z < wB * [|b|]a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z[|b|] * z + z0 < wB * [|b|]a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z[|a1|] * wB + [|a2|] < wB * [|b|]a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= z[|a1|] * wB + (wB - 1) < wB * [|b|]a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zwB * ([|a1|] + 1) - 1 < wB * [|b|]apply Z.mul_le_mono_nonneg; omega. Qed.a1, a2, b:int31H:wB / 2 <= [|b|]H0:[|a1|] < [|b|]H1:0 <= [|a1|] < wBH2:0 <= [|a2|] < wBH3:0 <= [|b|] < wBH4:[|b|] > 0z, z0:ZH5:[|a1|] * wB + [|a2|] = [|b|] * z + z0H6:0 <= z0 < [|b|]H7:0 <= [|a1|] * base + [|a2|] -> 0 <= zwB * ([|a1|] + 1) <= wB * [|b|]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:int31H: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:int31H:0 < [|b|]H0:[|b|] > 0let (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:int31H: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:int31H:0 < [|b|]H0:[|b|] > 0z, 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:int31H:0 < [|b|]H0:[|b|] > 0z, 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:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= z[|a|] = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zz * [|b|] + z0 = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wBz * [|b|] + z0 = z mod wB * [|b|] + z0 mod wB /\ 0 <= z0 mod wB < [|b|]a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wBz * [|b|] + z0 = z mod wB * [|b|] + z0 /\ 0 <= z0 < [|b|]a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wBz = z mod wBa, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wB0 <= z < wBa, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wBz < wBa, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wBz <= [|a|]a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wBz <= [|b|] * z + z0a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wBz <= [|b|] * zapply Z.mul_le_mono_nonneg; auto with zarith. Qed.a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|a|] -> 0 <= zH4:0 <= [|a|] < wBH5:0 <= [|b|] < wB1 * z <= [|b|] * zforall 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:int31H: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:int31H: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:int31H: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:int31H: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:int31H:0 < [|b|]H0:[|b|] > 0z, z0:Z[|a|] = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> [|phi_inv z0|] = z0a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:Z[|a|] = [|b|] * z + z0 /\ 0 <= z0 < [|b|] -> z0 mod wB = z0a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]z0 mod wB = z0apply Zmod_small; omega. Qed.a, b:int31H:0 < [|b|]H0:[|b|] > 0z, z0:ZH1:[|a|] = [|b|] * z + z0H2:0 <= z0 < [|b|]H3:0 <= [|b|] < wBz0 mod wB = z0forall 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:natIHn: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:natIHn: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:natIHn: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:natIHn: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:natIHn: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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0: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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBH1:[|j|] = 0[|i|] = Z.abs [|i|]n:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|j|] = Z.pos p[|(let (_, r) := (i / j)%int31 in r)|] = [|i|] mod Z.pos pn:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|j|] = Z.pos p0 < [|j|]n:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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.n:natIHn:forall i0 j0 : int31, [|euler n i0 j0|] = Zgcdn n [|j0|] [|i0|]i, j:int31H:0 <= [|j|] < wBH0:0 <= [|i|] < wBp:positiveH1:[|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|]forall a b : int31, Zis_gcd [|a|] [|b|] [|gcd31 a b|]forall a b : int31, Zis_gcd [|a|] [|b|] [|gcd31 a b|]a, b:int31Zis_gcd [|a|] [|b|] [|gcd31 a b|]a, b:int31Zis_gcd [|a|] [|b|] (Zgcdn (2 * size) [|b|] [|a|])a, b:int31Zis_gcd [|b|] [|a|] (Zgcdn (2 * size) [|b|] [|a|])a, b:int31(Zgcd_bound [|b|] <= 2 * size)%nata, 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)%nata, b:int310 <= [|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)%nata, b:int310 <= 0 < wB -> (1 <= 2 * size)%nata, b:int31p:positive0 <= Z.pos p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nata, b:int31p:positive0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nata, b:int31p:positive0 <= Z.pos p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nata, b:int31p:positive0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%nata, b:int31p:positiveH:Z.pos p < wB(Pos.size_nat p + Pos.size_nat p <= 2 * size)%nata, b:int31p:positive0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%natintros (H,_); compute in H; elim H; auto. Qed.a, b:int31p:positive0 <= Z.neg p < wB -> (Pos.size_nat p + Pos.size_nat p <= 2 * size)%natforall (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:Typef:A -> Ai:int31a:Aiter_int31 i A f a = nat_rect (fun _ : nat => A) a (fun _ : nat => f) (Z.abs_nat [|i|])A:Typef:A -> Ai:int31a:Arecr (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:Typef:A -> Ai:int31a:Arecrbis_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:Typef:A -> Ai:int31a:Arecrbis_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:Typef:A -> Aforall (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:Typef:A -> An:natIHn: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:Typef:A -> An:natIHn: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:Typef:A -> An:natIHn: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:int31a:Amatch 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:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D0z:ZH0:0 <= znat_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:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= zf (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:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D0z:ZH0:0 <= z(Z.abs_nat z + Z.abs_nat z)%nat = Z.abs_nat (Z.double z)A:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= zf (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:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D0z:ZH0:0 <= z(Z.abs_nat z + Z.abs_nat z)%nat = Z.abs_nat (z + z)A:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= zf (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:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= zf (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:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= z(S (Z.abs_nat z) + Z.abs_nat z)%nat = Z.abs_nat (Z.succ_double z)A:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= z(S (Z.abs_nat z) + Z.abs_nat z)%nat = Z.abs_nat (z + z + 1)A:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= z(S (Z.abs_nat z) + Z.abs_nat z)%nat = (Z.abs_nat (z + z) + Z.abs_nat 1)%natchange (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.A:Typef:A -> An:natIHn: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:int31a:AH:firstr i = D1z:ZH0:0 <= z(S (Z.abs_nat z) + Z.abs_nat z)%nat = (Z.abs_nat z + Z.abs_nat z + Z.abs_nat 1)%natforall p x y : int31, addmuldiv31 p x y = addmuldiv31_alt (Z.abs_nat [|p|]) x yforall p x y : int31, addmuldiv31 p x y = addmuldiv31_alt (Z.abs_nat [|p|]) x yp, x, y:int31addmuldiv31 p x y = addmuldiv31_alt (Z.abs_nat [|p|]) x yp, 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 yp, 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 yx, y:int31n: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 yforall 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 yn:natIHn: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 yforall 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 yn:natIHn: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 yforall 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 yn:natIHn: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 y0x, 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 yn:natIHn: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 y0x, 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)rewrite nat_rect_plus; simpl; auto. Qed.n:natIHn: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 y0x, 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)forall x y p : int31, [|p|] <= 31 -> [|addmuldiv31 p x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wBforall x y p : int31, [|p|] <= 31 -> [|addmuldiv31 p x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wBx, y, p:int31H:[|p|] <= 31[|addmuldiv31 p x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wBx, y, p:int31H:[|p|] <= 31[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wBx, y, p:int31H:[|p|] <= 31[|p|] = Z.of_nat (Z.abs_nat [|p|])x, y, p:int31H:[|p|] <= 31H0:[|p|] = Z.of_nat (Z.abs_nat [|p|])[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wBx, y, p:int31H:[|p|] <= 310 <= [|p|]x, y, p:int31H:[|p|] <= 31H0:[|p|] = Z.of_nat (Z.abs_nat [|p|])[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wBx, y, p:int31H:[|p|] <= 31H0:[|p|] = Z.of_nat (Z.abs_nat [|p|])[|addmuldiv31_alt (Z.abs_nat [|p|]) x y|] = ([|x|] * 2 ^ [|p|] + [|y|] / 2 ^ (31 - [|p|])) mod wBx, y, p:int31H: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 wBx, y, p:int31n:natH:Z.of_nat n <= 31[|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wBx, y, p:int31n:natH:Z.of_nat n <= 31(n <= 31)%natx, y, p:int31n:natH:Z.of_nat n <= 31H0:(n <= 31)%nat[|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wBx, y, p:int31n:natH:Z.of_nat n <= 31H0:(n <= 31)%nat[|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wBn:natH0:(n <= 31)%natforall x y : int31, [|addmuldiv31_alt n x y|] = ([|x|] * 2 ^ Z.of_nat n + [|y|] / 2 ^ (31 - Z.of_nat n)) mod wBH0:(0 <= 31)%natforall x y : int31, [|addmuldiv31_alt 0 x y|] = ([|x|] * 2 ^ Z.of_nat 0 + [|y|] / 2 ^ (31 - Z.of_nat 0)) mod wBn:natH0:(S n <= 31)%natIHn:(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 wBforall 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 wBH0:(0 <= 31)%natx, y:int31[|addmuldiv31_alt 0 x y|] = ([|x|] * 2 ^ 0 + [|y|] / 2 ^ (31 - 0)) mod 2 ^ 31n:natH0:(S n <= 31)%natIHn:(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 wBforall 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 wBH0:(0 <= 31)%natx, y:int31[|addmuldiv31_alt 0 x y|] = ([|x|] + [|y|] / 2 ^ (31 - 0)) mod 2 ^ 31n:natH0:(S n <= 31)%natIHn:(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 wBforall 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 wBH0:(0 <= 31)%natx, y:int31[|addmuldiv31_alt 0 x y|] = ([|x|] + 0) mod 2 ^ 31H0:(0 <= 31)%natx, y:int310 = [|y|] / 2 ^ (31 - 0)n:natH0:(S n <= 31)%natIHn:(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 wBforall 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 wBH0:(0 <= 31)%natx, y:int31[|addmuldiv31_alt 0 x y|] = [|x|] mod 2 ^ 31H0:(0 <= 31)%natx, y:int310 = [|y|] / 2 ^ (31 - 0)n:natH0:(S n <= 31)%natIHn:(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 wBforall 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 wBH0:(0 <= 31)%natx, y:int310 = [|y|] / 2 ^ (31 - 0)n:natH0:(S n <= 31)%natIHn:(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 wBforall 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 wBn:natH0:(S n <= 31)%natIHn:(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 wBforall 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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, 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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, 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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D02 * [|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:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D02 * [|x|] * 2 ^ Z.of_nat n = [|x|] * (2 * 2 ^ Z.of_nat n)n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D02 * [|y|] / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D02 * [|y|] / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D02 * [|y|] / 2 ^ Z.succ (31 - Z.succ (Z.of_nat n)) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D02 * [|y|] / 2 / 2 ^ (31 - Z.succ (Z.of_nat n)) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H: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 wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D12 * [|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:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D12 * [|x|] * 2 ^ Z.of_nat n = [|x|] * (2 * 2 ^ Z.of_nat n)n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D12 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D12 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1(2 * [|y|]) mod wB = 2 * [|y|] - wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D1(2 * [|y|]) mod wB = 2 * [|y|] - wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D12 * [|y|] - wB = (2 * [|y|]) mod wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D10 <= 2 * [|y|] - wB < wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D12 ^ Z.of_nat (Init.Nat.pred size) <= [|y|] -> 0 <= [|y|] < wB -> 0 <= 2 * [|y|] - wB < wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D1wB':=2 ^ Z.of_nat (Init.Nat.pred size):ZwB' <= [|y|] -> 0 <= [|y|] < wB -> 0 <= 2 * [|y|] - wB < wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D1wB':=2 ^ Z.of_nat (Init.Nat.pred size):Z2 * wB' = wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D1wB':=2 ^ Z.of_nat (Init.Nat.pred size):Z2 * 2 ^ Z.of_nat (Init.Nat.pred size) = wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))y:int31H:firstl y = D1wB':=2 ^ Z.of_nat (Init.Nat.pred size):Z2 ^ Z.of_nat (S (Init.Nat.pred size)) = wBn:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|]) mod wB / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ Z.of_nat n + (2 * [|y|] - wB) / 2 ^ (31 - Z.of_nat n) = [|y|] / 2 ^ (31 - Z.succ (Z.of_nat n))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ 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:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ 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:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 ^ 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:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 * [|y|] / 2 ^ (31 + - Z.of_nat n) = [|y|] / 2 ^ (31 + - (Z.of_nat n + 1))n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 * [|y|] / 2 ^ Z.succ (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.n:natH0:(S n <= 31)%natIHn:(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 wBx, y:int31H:firstl y = D1H1:(2 * [|y|]) mod wB = 2 * [|y|] - wB2 * [|y|] / 2 / 2 ^ (31 - Z.succ (Z.of_nat n)) = [|y|] / 2 ^ (31 + - (Z.of_nat n + 1))forall n p a : Z, 0 <= p <= n -> ((a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p)) mod 2 ^ n = a mod 2 ^ pforall n p a : Z, 0 <= p <= n -> ((a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p)) mod 2 ^ n = a mod 2 ^ pn, p, a:ZH:0 <= p <= n((a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p)) mod 2 ^ n = a mod 2 ^ pn, p, a:ZH:0 <= p <= n(a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= n(a * 2 ^ (n - p) - a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= n(a * 2 ^ (n - p) + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n)) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= na + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= n2 ^ n = 2 ^ (n - p) * 2 ^ pn, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= n2 ^ n = 2 ^ (n - p + p)n, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a * 2 ^ (n - p) / 2 ^ n * 2 ^ n) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a * 2 ^ (n - p) / (2 ^ (n - p) * 2 ^ p) * (2 ^ (n - p) * 2 ^ p)) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a / 2 ^ p * (2 ^ (n - p) * 2 ^ p)) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a / 2 ^ p * 2 ^ p * 2 ^ (n - p)) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a / 2 ^ p * 2 ^ p) * 2 ^ (n - p) / 2 ^ (n - p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nH0:2 ^ n = 2 ^ (n - p) * 2 ^ pa + - (a / 2 ^ p * 2 ^ p) = a mod 2 ^ pn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= n0 <= (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)0 <= b mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ n0 <= b mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ n0 <= b mod 2 ^ n / 2 ^ (n - p)n, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ nb mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ nb mod 2 ^ n / 2 ^ (n - p) < 2 ^ nn, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ nb mod 2 ^ n < 2 ^ n * 2 ^ (n - p)n, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ n2 ^ n <= 2 ^ n * 2 ^ (n - p)n, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ n2 ^ n * 1 <= 2 ^ n * 2 ^ (n - p)cut (0 < 2 ^ (n-p)); auto with zarith. Qed.n, p, a:ZH:0 <= p <= nb:ZHeqb:b = a * 2 ^ (n - p)H0:0 <= b mod 2 ^ nH1:b mod 2 ^ n < 2 ^ n1 <= 2 ^ (n - p)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 ^ pH:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ pforall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]w:int31p:ZH:31 <= p[|w|] = [|w|] mod 2 ^ pH:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ pforall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]w:int31p:ZH:31 <= p0 <= [|w|] < wB -> [|w|] = [|w|] mod 2 ^ pH:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ pforall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]w:int31p:ZH:31 <= pH0:0 <= [|w|] < wB0 <= [|w|] < 2 ^ pH:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ pforall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]w:int31p:ZH:31 <= pH0:0 <= [|w|] < wB[|w|] < 2 ^ pH:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ pforall w p : int31, [|match [|p|] ?= 31 with | Lt => addmuldiv31 p 0 (addmuldiv31 (31 - p) w 0) | _ => w end|] = [|w|] mod 2 ^ [|p|]w:int31p:ZH:31 <= pH0:0 <= [|w|] < wBwB <= 2 ^ pH:forall (w : int31) (p : Z), 31 <= p -> [|w|] = [|w|] mod 2 ^ pforall 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 ^ pforall 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 ^ p0w, 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 ^ p0w, p:int31H0:([|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 ^ p0w, p:int31H0:[|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 ^ p0w, p:int31H0:[|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 ^ p0w, p:int31H0:[|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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2: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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wB31 - [|p|] < wBH:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3: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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wB31 < wBH:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3: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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3: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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wB[|31 - p|] = 31 - [|p|]H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wBH4:[|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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wB([|31|] - [|p|]) mod wB = 31 - [|p|]H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wBH4:[|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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wB(31 - [|p|]) mod wB = 31 - [|p|]H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wBH4:[|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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wBH4:[|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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wBH4:[|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 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wBH4:[|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.H:forall (w0 : int31) (p0 : Z), 31 <= p0 -> [|w0|] = [|w0|] mod 2 ^ p0w, p:int31H0:[|p|] < 31H1:0 <= [|p|] < wBH2:0 <= [|w|] < wBH3:31 - [|p|] < wBH4:[|31 - p|] = 31 - [|p|](([|w|] * 2 ^ (31 - [|p|])) mod wB / 2 ^ (31 - [|p|])) mod wB = [|w|] mod 2 ^ [|p|]
Shift operations
forall x : int31, [|x|] = 0 -> [|head031 x|] = 31forall x : int31, [|x|] = 0 -> [|head031 x|] = 31x:int31H:[|x|] = 0[|head031 x|] = 31x:int31H:[|x|] = 0phi_inv [|x|] = x -> [|head031 x|] = 31x:int31H:[|x|] = 0On = x -> [|head031 x|] = 31simpl; 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.x:int31H:[|x|] = 0H':On = x[|head031 On|] = 31forall 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:int31H:iszero x = true[|head031 x|] = Z.of_nat (head031_alt size x)x:int31H:iszero x = false[|head031 x|] = Z.of_nat (head031_alt size x)x:int31H:iszero x = true[|head031 0|] = Z.of_nat (head031_alt size 0)x:int31H:iszero x = false[|head031 x|] = Z.of_nat (head031_alt size x)x:int31H:iszero x = false[|head031 x|] = Z.of_nat (head031_alt size x)x:int31H: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:int31H: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:int31H: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:int31H:iszero x = falseH0:(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)%natx:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = false[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) mod wB = Z.of_nat (31 - S n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = false0 <= Z.of_nat (31 - S n) < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = false0 <= Z.of_nat (31 - S n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) <= Z.of_nat 31n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat 31 < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat 31 < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0phi_inv ([|phi_inv (Z.of_nat (31 - S n))|] + [|In|]) = phi_inv (Z.of_nat (31 - n))n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0phi_inv (Z.of_nat (31 - S n) + [|In|]) = phi_inv (Z.of_nat (31 - n))n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0Z.of_nat (31 - S n) + [|In|] = Z.of_nat (31 - n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0Z.of_nat (31 - S n) + 1 = Z.of_nat (31 - n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0Z.of_nat (31 - S n) + 1 = Z.of_nat (S (31 - S n))n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstl x = D0iszero (shiftl x) = falsex:int31H:iszero x = falseH2:firstl x = D0iszero (shiftl x) = falsex:int31H:iszero (sneakr (firstl x) (shiftl x)) = falseH2:firstl x = D0iszero (shiftl x) = falsex:int31H:iszero (sneakr D0 (shiftl x)) = falseH2:firstl x = D0iszero (shiftl x) = falserewrite (iszero_eq0 _ H0) in H; discriminate. Qed.x:int31H:iszero (sneakr D0 (shiftl x)) = falseH2:firstl x = D0H0:iszero (shiftl x) = truetrue = falseforall x : int31, 0 < [|x|] <-> x <> 0%int31forall x : int31, 0 < [|x|] <-> x <> 0%int31x:int31H:0 < [|x|]x <> 0%int31x:int31H:x <> 0%int310 < [|x|]x:int31H:x <> 0%int310 < [|x|]x:int31H:x <> 0%int31[|x|] <> 0x:int31H:x <> 0%int31H0:[|x|] <> 00 < [|x|]x:int31H:[|x|] = 0x = 0%int31x:int31H:x <> 0%int31H0:[|x|] <> 00 < [|x|]generalize (phi_bounded x); auto with zarith. Qed.x:int31H:x <> 0%int31H0:[|x|] <> 00 < [|x|]forall x : int31, 0 < [|x|] -> wB / 2 <= 2 ^ [|head031 x|] * [|x|] < wBforall x : int31, 0 < [|x|] -> wB / 2 <= 2 ^ [|head031 x|] * [|x|] < wBx:int31H:0 < [|x|]wB / 2 <= 2 ^ [|head031 x|] * [|x|] < wBx:int31H:0 < [|x|]wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wBx:int31H:0 < [|x|]nshiftl x size = 0%int31x:int31H:0 < [|x|]H0:nshiftl x size = 0%int31wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wBx:int31H:0 < [|x|]H0:nshiftl x size = 0%int31wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wBforall x : int31, 0 < [|x|] -> nshiftl x size = 0%int31 -> wB / 2 <= 2 ^ Z.of_nat (head031_alt size x) * [|x|] < wBforall 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 31forall 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 31n:natIHn: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 31forall 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 31forall x : int31, 0 < [|x|] -> nshiftl x 0 = 0%int31 -> 2 ^ 31 / 2 <= 2 ^ 0 * [|x|] < 2 ^ 31n:natIHn: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 31forall 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 31x:int31H:0 < [|x|]H0:nshiftl x 0 = 0%int312 ^ 31 / 2 <= 2 ^ 0 * [|x|] < 2 ^ 31n:natIHn: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 31forall 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 31n:natIHn: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 31forall 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 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int312 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt (S n) x) * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int312 ^ 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 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D02 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (S (head031_alt n (shiftl x))) * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D02 ^ Z.of_nat 31 / 2 <= 2 * 2 ^ Z.of_nat (head031_alt n (shiftl x)) * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D02 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n (shiftl x)) * (2 * [|x|]) < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D02 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat (head031_alt n (shiftl x)) * [|twice x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D00 < [|twice x|]n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D0nshiftl (twice x) n = 0%int31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H0:nshiftl x (S n) = 0%int31H1:firstl x = D0H:twice x = 0%int31x = 0%int31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D0nshiftl (twice x) n = 0%int31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H0:nshiftl x (S n) = 0%int31H1:firstl x = D0H:shiftl x = 0%int31x = 0%int31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D0nshiftl (twice x) n = 0%int31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D0nshiftl (twice x) n = 0%int31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= 2 ^ Z.of_nat 0 * [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D12 ^ Z.of_nat 31 / 2 <= [|x|] < 2 ^ Z.of_nat 31n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D1H2:0 <= [|x|] < 2 ^ Z.of_nat 312 ^ Z.of_nat 31 / 2 <= [|x|]apply phi_lowerbound; auto. Qed.n:natIHn: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 31x:int31H:0 < [|x|]H0:nshiftl x (S n) = 0%int31H1:firstl x = D1H2:0 <= [|x|] < 2 ^ Z.of_nat 312 ^ Z.of_nat (Init.Nat.pred size) <= [|x|]forall x : int31, [|x|] = 0 -> [|tail031 x|] = 31forall x : int31, [|x|] = 0 -> [|tail031 x|] = 31x:int31H:[|x|] = 0[|tail031 x|] = 31x:int31H:[|x|] = 0phi_inv [|x|] = x -> [|tail031 x|] = 31x:int31H:[|x|] = 0On = x -> [|tail031 x|] = 31simpl; 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.x:int31H:[|x|] = 0H':On = x[|tail031 On|] = 31forall 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:int31H:iszero x = true[|tail031 x|] = Z.of_nat (tail031_alt size x)x:int31H:iszero x = false[|tail031 x|] = Z.of_nat (tail031_alt size x)x:int31H:iszero x = true[|tail031 0|] = Z.of_nat (tail031_alt size 0)x:int31H:iszero x = false[|tail031 x|] = Z.of_nat (tail031_alt size x)x:int31H:iszero x = false[|tail031 x|] = Z.of_nat (tail031_alt size x)x:int31H: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:int31H: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:int31H: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:int31H:iszero x = falseH0:(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)%natx:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H: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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = false[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) mod wB = Z.of_nat (31 - S n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = false0 <= Z.of_nat (31 - S n) < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = false0 <= Z.of_nat (31 - S n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat (31 - S n) <= Z.of_nat 31n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat 31 < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseZ.of_nat 31 < wBn:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|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:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0phi_inv ([|phi_inv (Z.of_nat (31 - S n))|] + [|In|]) = phi_inv (Z.of_nat (31 - n))n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0phi_inv (Z.of_nat (31 - S n) + [|In|]) = phi_inv (Z.of_nat (31 - n))n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0Z.of_nat (31 - S n) + [|In|] = Z.of_nat (31 - n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0Z.of_nat (31 - S n) + 1 = Z.of_nat (31 - n)n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0Z.of_nat (31 - S n) + 1 = Z.of_nat (S (31 - S n))n:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsen:natH0:(S n <= 31)%natIHn:(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:int31H:iszero x = falseH1:[|phi_inv (Z.of_nat (31 - S n))|] = Z.of_nat (31 - S n)H2:firstr x = D0iszero (shiftr x) = falsex:int31H:iszero x = falseH2:firstr x = D0iszero (shiftr x) = falsex:int31H:iszero (sneakl (firstr x) (shiftr x)) = falseH2:firstr x = D0iszero (shiftr x) = falsex:int31H:iszero (sneakl D0 (shiftr x)) = falseH2:firstr x = D0iszero (shiftr x) = falserewrite (iszero_eq0 _ H0) in H; discriminate. Qed.x:int31H:iszero (sneakl D0 (shiftr x)) = falseH2:firstr x = D0H0:iszero (shiftr x) = truetrue = falseforall 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:int31H:0 < [|x|]exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail031 x|]x:int31H:0 < [|x|]exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt size x)x:int31H:0 < [|x|]nshiftr x size = 0%int31x:int31H:0 < [|x|]H0:nshiftr x size = 0%int31exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt size x)x:int31H:0 < [|x|]H0:nshiftr x size = 0%int31exists 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:natIHn: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 ^ 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x 0 = 0%int31exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ 0n:natIHn: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:natIHn: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:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (tail031_alt (S n) x)n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31exists 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 endn:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat (S (tail031_alt n (shiftr x)))n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * (2 * 2 ^ Z.of_nat (tail031_alt n (shiftr x)))n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D00 < [|shiftr x|]n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0nshiftr (shiftr x) n = 0%int31n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0y:ZHy1:0 <= yHy2:[|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:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H0:nshiftr x (S n) = 0%int31H1:firstr x = D0H:shiftr x = 0%int31x = 0%int31n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0nshiftr (shiftr x) n = 0%int31n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0y:ZHy1:0 <= yHy2:[|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:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0nshiftr (shiftr x) n = 0%int31n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0y:ZHy1:0 <= yHy2:[|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:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0y:ZHy1:0 <= yHy2:[|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:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0y:ZHy1:0 <= yHy2:[|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:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D0y:ZHy1:0 <= yHy2:[|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:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1exists y : Z, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D10 <= [|shiftr x|] /\ [|x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D10 <= [|shiftr x|]n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1[|x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1[|x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0rewrite Z.succ_double_spec; simpl; ring. Qed. (* Sqrt *) (* Direct transcription of an old proof of a fortran program in boyer-moore *)n:natIHn: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:int31H:0 < [|x|]H0:nshiftr x (S n) = 0%int31H1:firstr x = D1Z.succ_double [|shiftr x|] = (2 * [|shiftr x|] + 1) * 2 ^ Z.of_nat 0a:Za - 1 <= a / 2 + a / 2a:Za - 1 <= a / 2 + a / 2intros H1; rewrite Zmod_eq_full; auto with zarith. Qed.a:Z0 <= a mod 2 -> a mod 2 < 2 -> a - 1 <= a / 2 + a / 2j, k:Z0 <= j -> 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2j, k:Z0 <= j -> 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 20 <= 0 -> forall k : Z, 0 <= k -> 0 * k + 0 <= ((0 + k) / 2 + 1) ^ 2forall 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) ^ 2k:ZHk:0 <= k0 <= (k / 2 + 1) ^ 2forall 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) ^ 2forall 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) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2Z.succ j * 0 + Z.succ j <= ((Z.succ j + 0) / 2 + 1) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2forall 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) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2Z.succ j <= (Z.succ j / 2 + 1) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2forall 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) ^ 2j:ZHj:0 <= jHrec: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) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2forall 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) ^ 2j:ZHj:0 <= jHrec: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:ZHj:0 <= jHrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2forall 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) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k : Z, 0 <= k -> j * k + j <= ((j + k) / 2 + 1) ^ 2forall 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) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= kZ.succ j * Z.succ k + Z.succ j <= ((Z.succ j + Z.succ k) / 2 + 1) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= kZ.succ j * Z.succ k + Z.succ j <= ((j + k) / 2 + 1 + 1) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= k(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= kj * 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) ^ 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= k(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= kj * 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:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= k(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= kj * 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:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= k(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= k(j + k) / 2 + 1 = (Z.succ j + Z.succ k) / 2apply f_equal2 with (f := Z.div); auto with zarith. Qed.j:ZHj:0 <= jHrec:0 <= j -> forall k0 : Z, 0 <= k0 -> j * k0 + j <= ((j + k0) / 2 + 1) ^ 2k:ZHk:0 <= k(1 * 2 + (j + k)) / 2 = (Z.succ j + Z.succ k) / 2i, j:Z0 <= i -> 0 < j -> i < ((j + i / j) / 2 + 1) ^ 2i, j:Z0 <= i -> 0 < j -> i < ((j + i / j) / 2 + 1) ^ 2i, j:ZHi:0 <= iHj:0 < ji < ((j + i / j) / 2 + 1) ^ 2i, j:ZHi:0 <= iHj:0 < jHij:0 <= i / ji < ((j + i / j) / 2 + 1) ^ 2pattern i at 1; rewrite (Z_div_mod_eq i j); case (Z_mod_lt i j); auto with zarith. Qed.i, j:ZHi:0 <= iHj:0 < jHij:0 <= i / ji < j * (i / j) + ji:Z1 < i -> i < (i / 2 + 1) ^ 2i:Z1 < i -> i < (i / 2 + 1) ^ 2i:ZHi:1 < ii < (i / 2 + 1) ^ 2i:ZHi:1 < iH1:0 <= i - 2i < (i / 2 + 1) ^ 2i:ZHi:1 < iH1:0 <= i - 21 <= (i / 2) ^ 2i:ZHi:1 < iH1:0 <= i - 2H2:1 <= (i / 2) ^ 2i < (i / 2 + 1) ^ 2i:ZHi:1 < iH1:0 <= i - 21 <= ((1 * 2 + (i - 2)) / 2) ^ 2i:ZHi:1 < iH1:0 <= i - 2H2:1 <= (i / 2) ^ 2i < (i / 2 + 1) ^ 2i:ZHi:1 < iH1:0 <= i - 21 <= (1 + (i - 2) / 2) * (1 + (i - 2) / 2)i:ZHi:1 < iH1:0 <= i - 2H2:1 <= (i / 2) ^ 2i < (i / 2 + 1) ^ 2i:ZHi:1 < iH1: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:ZHi:1 < iH1:0 <= i - 2H2:1 <= (i / 2) ^ 2i < (i / 2 + 1) ^ 2i:ZHi:1 < iH1: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:ZHi:1 < iH1:0 <= i - 2H2:1 <= (i / 2) ^ 2i < (i / 2 + 1) ^ 2i:ZHi:1 < iH1:0 <= i - 2H2:1 <= (i / 2) ^ 2i < (i / 2 + 1) ^ 2i:ZHi:1 < iH1:0 <= i - 2H2:1 <= (i / 2) ^ 2i - 1 <= i / 2 + i / 2 -> i < (i / 2 + 1) ^ 2auto with zarith. Qed.i:ZHi:1 < iH1:0 <= i - 2H2:1 <= i / 2 * (i / 2)i - 1 <= i / 2 + i / 2 -> i < i / 2 * (i / 2) + i / 2 + (i / 2 + 1)i, j:Z0 <= i -> 0 < j -> i / j >= j -> j ^ 2 <= ii, j:Z0 <= i -> 0 < j -> i / j >= j -> j ^ 2 <= ii, j:ZHi:0 <= iHj:0 < jHd:i / j >= jj * j <= iapply Z_mult_div_ge; auto with zarith. Qed.i, j:ZHi:0 <= iHj:0 < jHd:i / j >= jj * (i / j) <= ii, j:Z0 <= i -> 0 < j -> i / j < j -> (j + i / j) / 2 < ji, j:Z0 <= i -> 0 < j -> i / j < j -> (j + i / j) / 2 < ji, j:ZHi:0 <= iHj:0 < jH:i / j < jj <= (j + i / j) / 2 -> (j + i / j) / 2 < ji, j:ZHi:0 <= iHj:0 < jH1:j <= (j + i / j) / 2j <= i / ji, j:ZHi:0 <= iHj:0 < jH1:j <= (j + i / j) / 22 * j <= j + i / japply Z_mult_div_ge; auto with zarith. Qed.i, j:ZHi:0 <= iHj:0 < jH1:j <= (j + i / j) / 22 * ((j + i / j) / 2) <= j + i / jrec:int31 -> int31 -> int31i, j:int31sqrt31_step rec i j = match (fst (i / j) ?= j)%int31 with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j endrec:int31 -> int31 -> int31i, j:int31sqrt31_step rec i j = match (fst (i / j) ?= j)%int31 with | Lt => rec i (fst ((j + fst (i / j)) / 2)%int31) | _ => j endsimpl; case compare31; auto. Qed.rec:int31 -> int31 -> int31i, j, i0, i1:int31match (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 endi, j:int310 < [|j|] -> [|fst (i / j)%int31|] = [|i|] / [|j|]i, j:int31Hj:0 < [|j|](let (q, r) := (i / j)%int31 in [|i|] = [|q|] * [|j|] + [|r|] /\ 0 <= [|r|] < [|j|]) -> [|fst (i / j)%int31|] = [|i|] / [|j|]i, j:int31Hj:0 < [|j|]q, r:int31[|i|] = [|q|] * [|j|] + [|r|] /\ 0 <= [|r|] < [|j|] -> [|q|] = [|i|] / [|j|]rewrite H1; ring. Qed.i, j:int31Hj:0 < [|j|]q, r:int31H1:[|i|] = [|q|] * [|j|] + [|r|]H2:0 <= [|r|] < [|j|][|i|] = [|j|] * [|q|] + [|r|]rec:int31 -> int31 -> int31i, j:int310 < [|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) ^ 2rec:int31 -> int31 -> int31i, j:int310 < [|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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hp2: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|][|rec i (fst ((j + fst (i / j)) / 2)%int31)|] ^ 2 <= [|i|] < ([|rec i (fst ((j + fst (i / j)) / 2)%int31)|] + 1) ^ 2rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|][|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|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) ^ 2rewrite spec_add, div31_phi; auto using Z.mod_small with zarith.rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|][|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|] < [|j|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:Z.succ 0 <= [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 <= [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 = [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < (1 * 2 + ([|j|] - 2 + [|i|] / [|j|])) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < 1 + ([|j|] - 2 + [|i|] / [|j|]) / [|2|]assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / [|2|]); auto with zarith.rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]H:0 <= [|i|] / [|j|]0 < 1 + ([|j|] - 2 + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 = [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < ([|j|] + [|i|] / [|j|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 = [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < (1 + [|i|]) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 = [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < (1 * 2 + ([|i|] - 1)) / [|2|]rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 = [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]0 < 1 + ([|i|] - 1) / [|2|]change ([|2|]) with 2; auto with zarith. Qed.rec:int31 -> int31 -> int31i, j:int31Hp2:0 < [|2|]Hi:0 < [|i|]Hj:1 = [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec:forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2Hc:[|i|] / [|j|] < [|j|]E:[|j + fst (i / j)%int31|] = [|j|] + [|i|] / [|j|]H:0 <= ([|i|] - 1) / 20 < 1 + ([|i|] - 1) / [|2|]n:natrec:int31 -> int31 -> int31i, j:int310 < [|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) ^ 2n:natrec:int31 -> int31 -> int31i, j:int310 < [|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) ^ 2forall (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) ^ 2forall 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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec: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) ^ 2forall j1 : int31, 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2forall 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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHrec: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) ^ 2j1:int31H:0 < [|j1|] < [|j|]H0:[|i|] < ([|j1|] + 1) ^ 22 ^ 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) ^ 2forall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2forall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:[|i|] < ([|j1|] + 1) ^ 2forall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:[|i|] < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:[|i|] < ([|j2|] + 1) ^ 2Hj31:2 * [|j2|] < wBforall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:[|i|] < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:[|i|] < ([|j2|] + 1) ^ 2Hj31:2 * [|j2|] < wBj3:int31Hj3: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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:[|i|] < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:[|i|] < ([|j2|] + 1) ^ 2Hj31:2 * [|j2|] < wBj3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]2 ^ Z.of_nat (S n) + [|j3|] <= [|j|]n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:[|i|] < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:[|i|] < ([|j2|] + 1) ^ 2Hj31:2 * [|j2|] < wBj3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]2 * 2 ^ Z.of_nat n + [|j3|] <= [|j|]n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:[|i|] < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:[|i|] < ([|j2|] + 1) ^ 2Hj31:2 * [|j2|] < wBj3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]0 <= Z.of_nat napply Nat2Z.is_nonneg. Qed.n:natHrec: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) ^ 2rec:int31 -> int31 -> int31i, j:int31Hi:0 < [|i|]Hj:0 < [|j|]Hij:[|i|] < ([|j|] + 1) ^ 2H31:2 * [|j|] < wBHHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:[|i|] < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:[|i|] < ([|j2|] + 1) ^ 2Hj31:2 * [|j2|] < wBj3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]0 <= Z.of_nat nforall x : int31, [|sqrt31 x|] ^ 2 <= [|x|] < ([|sqrt31 x|] + 1) ^ 2forall x : int31, [|sqrt31 x|] ^ 2 <= [|x|] < ([|sqrt31 x|] + 1) ^ 2i: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) ^ 2i: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) ^ 2i:int31Hi:1 = [|i|]1 ^ 2 <= [|i|] < (1 + 1) ^ 2i:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]0 < [|fst (i / 2)%int31|]i:int31Hi:1 < [|i|][|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * [|fst (i / 2)%int31|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]0 < [|i|] / 2i:int31Hi:1 < [|i|][|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * [|fst (i / 2)%int31|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]0 < (1 * 2 + ([|i|] - 2)) / 2i:int31Hi:1 < [|i|][|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * [|fst (i / 2)%int31|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]H:0 <= ([|i|] - 2) / 20 < (1 * 2 + ([|i|] - 2)) / 2i:int31Hi:1 < [|i|][|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * [|fst (i / 2)%int31|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|][|i|] < ([|fst (i / 2)%int31|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * [|fst (i / 2)%int31|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|][|i|] < ([|i|] / 2 + 1) ^ 2i:int31Hi:1 < [|i|]2 * [|fst (i / 2)%int31|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * [|fst (i / 2)%int31|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * ([|i|] / 2) < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]2 * ([|i|] / 2) <= [|i|]i:int31Hi:1 < [|i|][|i|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|][|i|] < wBi:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi: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) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|][|fst (i / 2)%int31|] < 2 ^ Z.of_nat 31 + [|j2|]i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|][|i|] / 2 < 2 ^ Z.of_nat 31 + [|j2|]i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|][|i|] / 2 <= [|i|]i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|][|i|] < 2 ^ Z.of_nat 31 + [|j2|]i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|]H:0 <= [|i|] / 2[|i|] / 2 <= [|i|]i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|][|i|] < 2 ^ Z.of_nat 31 + [|j2|]i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|]H:0 <= [|i|] / 22 * ([|i|] / 2) <= [|i|]i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|][|i|] < 2 ^ Z.of_nat 31 + [|j2|]i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:1 < [|i|]j2:int31H1:0 < [|j2|][|i|] < 2 ^ Z.of_nat 31 + [|j2|]i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2i:int31Hi:[|i|] < 1[|0|] ^ 2 <= [|i|] < ([|0|] + 1) ^ 2case (phi_bounded i); repeat rewrite Z.pow_2_r; auto with zarith. Qed.i:int31Hi:[|i|] < 10 ^ 2 <= [|i|] < (0 + 1) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31sqrt312_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 endrec:int31 -> int31 -> int31 -> int31ih, il, j:int31sqrt312_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 endsimpl; case compare31; auto. Qed.rec:int31 -> int31 -> int31 -> int31ih, il, j, i, i0:int31match (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 endih, il, j:int31phi2 ih il < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|]ih, il, j:int31phi2 ih il < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|]ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2[|ih|] <= [|j|]ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|][|ih|] <= [|j|]ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|][|ih|] <= [|j|]ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] <= [|j|]ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] < [|j|] + 1ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] * [|ih|] < ([|j|] + 1) * ([|j|] + 1)ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] ^ 2 <= phi2 ih ilih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] ^ 2 <= [|ih|] * wBih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] * wB <= phi2 ih ilrewrite ? Z.pow_2_r; auto with zarith.ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] ^ 2 <= [|ih|] * wBih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] * wB <= phi2 ih ilchange base with wB; auto with zarith. Qed.ih, il, j:int31H1:phi2 ih il < ([|j|] + 1) ^ 2Hbj:0 <= [|j|]Hbil:0 <= [|il|]Hbih:0 <= [|ih|]Hbih1:[|ih|] < wB[|ih|] * wB <= [|ih|] * base + [|il|]ih, il, j:int312 ^ 30 <= [|j|] -> [|ih|] < [|j|] -> [|fst (div3121 ih il j)|] = phi2 ih il / [|j|]ih, il, j:int312 ^ 30 <= [|j|] -> [|ih|] < [|j|] -> [|fst (div3121 ih il j)|] = phi2 ih il / [|j|]ih, il, j:int31Hj:2 ^ 30 <= [|j|]Hj1:[|ih|] < [|j|][|fst (div3121 ih il j)|] = phi2 ih il / [|j|]ih, il, j:int31Hj: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:int31Hj:2 ^ 30 <= [|j|]Hj1:[|ih|] < [|j|]q, r:int31Hq:[|ih|] * wB + [|il|] = [|q|] * [|j|] + [|r|]Hr:0 <= [|r|] < [|j|][|fst (q, r)|] = phi2 ih il / [|j|]simpl @fst; apply eq_trans with (1 := Hq); ring. Qed.ih, il, j:int31Hj:2 ^ 30 <= [|j|]Hj1:[|ih|] < [|j|]q, r:int31Hq:[|ih|] * wB + [|il|] = [|q|] * [|j|] + [|r|]Hr:0 <= [|r|] < [|j|]phi2 ih il = [|j|] * [|fst (q, r)|] + [|r|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int312 ^ 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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int312 ^ 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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wB0 < phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wB0 < phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wB0 < [|ih|] * baseapply Z.lt_le_trans with (2:= Hih); auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wB0 < [|ih|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|][|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|][|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|][|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|][|j|] ^ 2 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]0 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]phi2 ih il / [|j|] >= [|j|]unfold phi2, base; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]0 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]phi2 ih il / [|j|] >= [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]([|j|] * base + [|il|]) / [|j|] >= [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]H:0 <= [|il|] / [|j|]([|j|] * base + [|il|]) / [|j|] >= [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]H:0 <= [|il|] / [|j|]base + [|il|] / [|j|] >= [|j|]auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] = [|j|]H:0 <= [|il|] / [|j|]wB + [|il|] / [|j|] >= [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:Z.succ 0 <= [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 <= [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]Hf1:0 <= phi2 ih il / [|j|]0 < ([|j|] + phi2 ih il / [|j|]) / 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 20 < [|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]Hf1:0 <= phi2 ih il / [|j|]0 < ([|j|] + phi2 ih il / [|j|]) / 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|])) / 2assert (0 <= ([|j|] - 2 + phi2 ih il / [|j|]) / 2) ; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 20 < [|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|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|]apply sqrt_test_false; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int311 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31 + v30|] < [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int31[|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31|] < [|j|]intros HH; rewrite HH; clear HH; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int31[|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|r|] / 2 < [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int311 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31 + v30|] < [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int311 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < [|fst (r / 2)%int31 + v30|] < [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int311 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> 0 < ([|r|] / 2 + [|v30|]) mod wB < [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int31HH:wB + [|r|] = [|j|] + phi2 ih il / [|j|]0 < ([|r|] / 2 + [|v30|]) mod wB < [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int31HH:wB + [|r|] = [|j|] + phi2 ih il / [|j|]0 < (([|v30|] * 2 + [|r|]) / 2) mod wB < [|j|]rewrite HH, Zmod_small; auto with zarith. }rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:1 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]) / 2Hf4:([|j|] + phi2 ih il / [|j|]) / 2 < [|j|]r:int31HH:wB + [|r|] = [|j|] + phi2 ih il / [|j|]0 < ((wB + [|r|]) / 2) mod wB < [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int311 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31 + v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]intros HH; rewrite HH; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int311 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31 + v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int311 * wB + [|r|] = [|j|] + phi2 ih il / [|j|] -> ([|j|] + phi2 ih il / [|j|]) / 2 = [|fst (r / 2)%int31 + v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|](1 * wB + [|r|]) / 2 = [|fst (r / 2)%int31 + v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]([|v30|] * 2 + [|r|]) / 2 = [|fst (r / 2)%int31 + v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|v30|] + [|r|] / 2 = [|fst (r / 2)%int31 + v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 + [|v30|] = [|fst (r / 2)%int31 + v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 + [|v30|] = [|fst (r / 2)%int31|] + [|v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]0 <= [|fst (r / 2)%int31|] + [|v30|] < wBrewrite div31_phi; auto.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 + [|v30|] = [|fst (r / 2)%int31|] + [|v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]0 <= [|fst (r / 2)%int31|] + [|v30|] < wBrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]0 <= [|fst (r / 2)%int31|] + [|v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|fst (r / 2)%int31|] + [|v30|] < wBcase (phi_bounded (fst (r/2)%int31)); case (phi_bounded v30); auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|]0 <= [|fst (r / 2)%int31|] + [|v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|fst (r / 2)%int31|] + [|v30|] < wBrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 + [|v30|] < wBrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 + [|v30|] < base / 2 + [|v30|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 < base / 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 * 2 < base / 2 * 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 * 2 < baserec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 * 2 <= [|r|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] < baserewrite Z.mul_comm; apply Z_mult_div_ge; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] / 2 * 2 <= [|r|]case (phi_bounded r); auto with zarith. }rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:2 ^ 30 <= [|j|]Hc:phi2 ih il / [|j|] < [|j|]r:int31HH:1 * wB + [|r|] = [|j|] + phi2 ih il / [|j|][|r|] < baserec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30([|j|] + 1) ^ 2 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 30([|j|] + 1) ^ 2 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 30([|j|] + 1) ^ 2 <= 2 ^ 30 * 2 ^ 30rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 302 ^ 30 * 2 ^ 30 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 30([|j|] + 1) ^ 2 <= 2 ^ 30 * 2 ^ 30apply Z.mul_le_mono_nonneg; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 30H0:0 <= 1 + [|j|]([|j|] + 1) ^ 2 <= 2 ^ 30 * 2 ^ 30rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 302 ^ 30 * 2 ^ 30 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 302 ^ 29 * base <= phi2 ih ilunfold phi2, base; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < baseHp3:0 < phi2 ih ilHc1:[|ih|] < [|j|]Hjj:[|j|] < 2 ^ 30H:1 + [|j|] <= 2 ^ 30[|ih|] * base <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|][|j|] ^ 2 <= phi2 ih il < ([|j|] + 1) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|][|j|] ^ 2 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|]0 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|]phi2 ih il / [|j|] >= [|j|]unfold phi2, base; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|]0 <= phi2 ih ilrec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|]phi2 ih il / [|j|] >= [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|][|j|] <= [|j|] * base / [|j|]rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|][|j|] * base / [|j|] <= phi2 ih il / [|j|]rewrite Z.mul_comm, Z_div_mult; auto with zarith.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|][|j|] <= [|j|] * base / [|j|]apply Z.ge_le; apply Z_div_ge; auto with zarith. Qed.rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hp2:0 < [|2|]Hih:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2H1:[|ih|] <= [|j|]Hih1:0 <= [|ih|]Hil1:0 <= [|il|]Hj1:[|j|] < wBHp3:0 < phi2 ih ilHc1:[|j|] < [|ih|][|j|] * base / [|j|] <= phi2 ih il / [|j|]n:natrec:int31 -> int31 -> int31 -> int31ih, il, j:int312 ^ 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) ^ 2n:natrec:int31 -> int31 -> int31 -> int31ih, il, j:int312 ^ 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) ^ 2forall (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) ^ 2forall 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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2forall j1 : int31, 0 < [|j1|] < [|j|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|rec ih il j1|] ^ 2 <= phi2 ih il < ([|rec ih il j1|] + 1) ^ 2forall 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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2Hrec: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) ^ 2j1:int31H:0 < [|j1|] < [|j|]H0:phi2 ih il < ([|j1|] + 1) ^ 22 ^ 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) ^ 2forall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2forall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2forall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2forall 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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2j3:int31Hj3: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) ^ 2n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2j3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]2 ^ Z.of_nat (S n) + [|j3|] <= [|j|]n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2j3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]2 * 2 ^ Z.of_nat n + [|j3|] <= [|j|]n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2j3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]0 <= Z.of_nat napply Nat2Z.is_nonneg. Qed. (* Avoid expanding [iter312_sqrt] before variables in the context. *) Strategy 1 [iter312_sqrt].n:natHrec: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) ^ 2rec:int31 -> int31 -> int31 -> int31ih, il, j:int31Hi:2 ^ 29 <= [|ih|]Hj:0 < [|j|]Hij:phi2 ih il < ([|j|] + 1) ^ 2HHrec: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) ^ 2j1:int31Hj1:0 < [|j1|] < [|j|]Hjp1:phi2 ih il < ([|j1|] + 1) ^ 2j2:int31Hj2:0 < [|j2|]H2j2:2 ^ Z.of_nat n + [|j2|] <= [|j1|]Hjp2:phi2 ih il < ([|j2|] + 1) ^ 2j3:int31Hj3:0 < [|j3|]Hpj3:2 ^ Z.of_nat n + [|j3|] <= [|j2|]0 <= Z.of_nat nforall 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:int31Hih: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:int31Hih: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2let (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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baselet (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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= basephi2 ih il < ([|Tn|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2let (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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= basephi2 ih il < ([|Tn|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= basephi2 ih il < 2 ^ 62ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= basephi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= base(2 ^ 31 - 1) * base + (2 ^ 31 - 1) < 2 ^ 62ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= basephi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseH1:0 <= [|il|]H2:[|il|] < wBH3:0 <= [|ih|]H4:[|ih|] < wBphi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseH1:0 <= [|il|]H2:[|il|] < 2147483648H3:0 <= [|ih|]H4:[|ih|] < 2147483648phi2 ih il <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseH1:0 <= [|il|]H2:[|il|] < 2147483648H3:0 <= [|ih|]H4:[|ih|] < 2147483648[|ih|] * base + [|il|] <= (2 ^ 31 - 1) * base + (2 ^ 31 - 1)auto with zarith.ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseH1:0 <= [|il|]H2:[|il|] < 2147483648H3: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)ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2let (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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 20 < [|Tn|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|Tn|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|j1|] ^ 2 <= phi2 ih il < ([|j1|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2forall j1 : int31, 0 < [|j1|] -> 2 ^ Z.of_nat 31 + [|j1|] <= [|Tn|] -> phi2 ih il < ([|j1|] + 1) ^ 2 -> [|j1|] ^ 2 <= phi2 ih il < ([|j1|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2j1:int31~ 2 ^ Z.of_nat 31 + [|j1|] <= [|Tn|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2j1:int31[|Tn|] < 2 ^ Z.of_nat 31 + [|j1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2j1:int312147483647 < 2 ^ Z.of_nat 31 + [|j1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2j1:int312147483647 < 2147483648 + [|j1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s : Z, s * s + 2 * s + 1 = (s + 1) ^ 2Hb:0 <= baseHi2: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2let (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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2[||W0||] = [|s|] * [|s|] -> phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2HH:0 = [|s|] * [|s|]phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2HH:0 = [|s|] * [|s|][|s|] = 0ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2HH:0 = [|s|] * [|s|]H:[|s|] = 0phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2HH:0 = [|s|] * [|s|][|s|] = 0ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2HH:[|s|] * [|s|] = 0[|s|] = 0destruct HH; auto.ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2HH:[|s|] = 0 \/ [|s|] = 0[|s|] = 0ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2HH:0 = [|s|] * [|s|]H:[|s|] = 0phi2 ih il = [|0|] ^ 2 + [+|C0 0%int31|] /\ [+|C0 0%int31|] <= 2 * [|0|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHH:0 = [|s|] * [|s|]H:[|s|] = 0(0 + 1) ^ 2 <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHH:0 = [|s|] * [|s|]H:[|s|] = 01 <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHH:0 = [|s|] * [|s|]H:[|s|] = 01 <= wB / 4 * baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHH:0 = [|s|] * [|s|]H:[|s|] = 0wB / 4 * base <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHH:0 = [|s|] * [|s|]H:[|s|] = 0wB / 4 * base <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHH:0 = [|s|] * [|s|]H:[|s|] = 0[|ih|] * base <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2forall 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31phi2 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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1: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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|][|ih|] = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|]phi2 ih il = [|s|] ^ 2 + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|]phi2 ih il = phi2 ih1 il1 + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|]phi2 ih il - phi2 ih1 il1 <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|]phi2 ih il - phi2 ih1 il1 = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|]phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|]phi2 ih il - phi2 ih1 il1 = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih|] = [|ih1|]phi2 ih il - phi2 ih1 il1 = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [+|C0 il2|] /\ [+|C0 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|]~ [|s|] ^ 2 <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|]phi2 ih il < phi2 ih1 il1ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|][|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|]H2:[|il|] < wB[|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|]H2:[|il|] < wB[|ih|] * base + [|il|] < ([|ih|] + 1) * base + 0ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|]H2:[|il|] < wB([|ih|] + 1) * base + 0 <= [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|]H2:[|il|] < wB([|ih|] + 1) * base + 0 <= [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|]H1:[|ih|] < [|ih1|]H2:[|il|] < wBH3:0 <= [|il1|]([|ih|] + 1) * base + 0 <= [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|C0 il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + [+|C1 il2|] /\ [+|C1 il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|][|ih1|] < [|ih|] -> phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] < [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 <= [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]([|s|] + 1) ^ 2 <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|][|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|][|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|][|il1|] <= [|il|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|]Hl1l:[|il1|] <= [|il|][|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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|]case (phi_bounded il2); rewrite Hil2; auto with zarith.ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|][|il1|] <= [|il|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|]Hl1l:[|il1|] <= [|il|][|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|]Hl1l:[|il1|] <= [|il|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|]Hl1l:[|il1|] <= [|il|]Hps:[|s|] < base[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|]Hl1l:[|il1|] <= [|il|]Hps:[|s|] < baseHpih1:0 <= [|ih1|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|]Hl1l:[|il1|] <= [|il|]Hps:[|s|] < baseHpih1:0 <= [|ih1|][|ih1|] * base + 2 * [|s|] + 1 <= ([|ih1|] + 2) * baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]Hpil:0 <= [|il|]Hl1l:[|il1|] <= [|il|]Hps:[|s|] < baseHpih1:0 <= [|ih1|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih1|] * base + 2 * baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|])) /\ base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + ([|il|] - [|il1|]))ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]base + ([|il|] - [|il1|]) <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il - [|s|] * [|s|] = base + ([|il|] - [|il1|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[|il2|] = [|il|] - [|il1|]H1:[|ih1|] + 1 = [|ih|]phi2 ih il - [|s|] * [|s|] = base + ([|il|] - [|il1|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:[-|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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|][|ih - 1|] = [|ih|] - 1ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1let (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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|][|ih - 1|] = [|ih|] - 1ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]0 <= [|ih|] - 1 < wBih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]H1:0 <= [|ih|]H2:[|ih|] < wB0 <= [|ih|] - 1 < wBsplit; auto with zarith.ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]H1:0 <= [|ih|]H2:[|ih|] < wB536870912 <= [|ih|] -> 0 <= [|ih|] - 1 < wBih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1let (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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih|] - 1 = [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]phi2 ih il = [|s|] ^ 2 + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]phi2 ih il = phi2 ih1 il1 + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|ih|] * base + [|il|] = ([|ih|] - 1) * base + [|il1|] + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|ih|] * base + [|il|] = [|ih|] * base + [|il1|] + ([|il|] - [|il1|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|ih|] * base + [|il1|] + ([|il|] - [|il1|]) = ([|ih|] - 1) * base + [|il1|] + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|ih|] * base + [|il1|] + ([|il|] - [|il1|]) = ([|ih|] - 1) * base + [|il1|] + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|ih|] * base + [|il1|] + (-1 * wB + [|il2|]) = ([|ih|] - 1) * base + [|il1|] + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]phi2 ih il - phi2 ih1 il1 <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]phi2 ih il - phi2 ih1 il1 = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]phi2 ih il - phi2 ih1 il1 = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]phi2 ih il - phi2 ih1 il1 = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|ih|] * base + [|il|] - ([|ih1|] * base + [|il1|]) = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|][|ih|] * base + [|il|] - (([|ih|] - 1) * base + [|il1|]) = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]base + [|il|] - [|il1|] = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]base + [|il|] - [|il1|] = base + ([|il|] - [|il1|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]base + ([|il|] - [|il1|]) = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]base + ([|il|] - [|il1|]) = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 = [|ih1|]base + (-1 * wB + [|il2|]) = [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1[|ih - 1|] < [|ih1|] -> phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|][|ih|] = [|ih1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]He:[|ih|] = [|ih1|]phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|][|ih|] = [|ih1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|][|ih1|] <= [|ih|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|][|ih1|] <= [|ih|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|]phi2 ih il < phi2 ih1 il1ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|][|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|]Hpil1:[|il|] < base[|ih|] * base + [|il|] < [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|]Hpil1:[|il|] < base[|ih|] * base + [|il|] < ([|ih|] + 1) * baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|]Hpil1:[|il|] < base([|ih|] + 1) * base <= [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|]Hpil1:[|il|] < base([|ih|] + 1) * base <= [|ih1|] * base + [|il1|]apply Z.le_trans with (([|ih1|]) * base); auto with zarith.ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]H2:[|ih|] < [|ih1|]Hpil1:[|il|] < baseHpil2:0 <= [|il1|]([|ih|] + 1) * base <= [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]He:[|ih|] = [|ih1|]phi2 ih il = [|s|] ^ 2 + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]He:[|ih|] = [|ih1|][|ih|] * base + [|il|] = [|ih|] * base + [|il1|] + [|il2|] /\ [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]He:[|ih|] = [|ih1|]phi2 ih il < phi2 ih1 il1ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]He:[|ih|] = [|ih1|][|ih1|] * base + [|il|] < [|ih1|] * base + [|il1|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]He:[|ih|] = [|ih1|][|il|] - [|il1|] < 0ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih|] - 1 < [|ih1|]He:[|ih|] = [|ih1|]-1 * wB + [|il2|] < 0ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-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:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]phi2 ih il = [|s|] ^ 2 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 <= [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]([|s|] + 1) ^ 2 <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 <= phi2 ih ilih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|][|ih1|] * base + [|il1|] + 2 * [|s|] + 1 <= [|ih|] * base + [|il|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + ([|il|] - [|il1|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + (-1 * wB + [|il2|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + (- base + [|il2|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]Hpil2:0 <= [|il2|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + (- base + [|il2|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]Hpil2:0 <= [|il2|][|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + - baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]Hpil2:0 <= [|il2|]Hps:[|s|] < base[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + - baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]Hpil2:0 <= [|il2|]Hps:[|s|] < baseH:2 * [|s|] + 1 <= 2 * base[|ih1|] * base + 2 * [|s|] + 1 <= [|ih|] * base + - baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]Hpil2:0 <= [|il2|]Hps:[|s|] < baseH:2 * [|s|] + 1 <= 2 * base[|ih1|] * base + 2 * base <= [|ih|] * base + - baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]Hpil2:0 <= [|il2|]Hps:[|s|] < baseH:2 * [|s|] + 1 <= 2 * baseHi:([|ih1|] + 3) * base <= [|ih|] * base[|ih1|] * base + 2 * base <= [|ih|] * base + - baseih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 < [|ih|]phi2 ih1 il1 + 2 * [|s|] + 1 = ([|s|] + 1) ^ 2ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il = phi2 ih1 il1 + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]([|ih1|] + 2) * base + [|il|] = [|ih1|] * base + [|il1|] + (base + [|il2|]) /\ base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]([|ih1|] + 2) * base + [|il|] = [|ih1|] * base + [|il1|] + (base + [|il2|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]([|ih1|] + 2) * base + ([|il|] - [|il1|] + [|il1|]) = [|ih1|] * base + [|il1|] + (base + [|il2|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]([|ih1|] + 2) * base + (-1 * wB + [|il2|] + [|il1|]) = [|ih1|] * base + [|il1|] + (base + [|il2|])ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]base + [|il2|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il - phi2 ih1 il1 <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il - phi2 ih1 il1 = base + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il - [|s|] * [|s|] <= 2 * [|s|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il - phi2 ih1 il1 = base + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]phi2 ih il - phi2 ih1 il1 = base + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]([|ih1|] + 2) * base + [|il|] - ([|ih1|] * base + [|il1|]) = base + [|il2|]ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]([|ih1|] + 2) * base + ([|il|] - [|il1|] + [|il1|]) - ([|ih1|] * base + [|il1|]) = base + [|il2|]change (-1 * 2 ^ Z.of_nat size) with (-base); ring. Qed.ih, il:int31Hih:wB / 4 <= [|ih|]Hbin:forall s0 : Z, s0 * s0 + 2 * s0 + 1 = (s0 + 1) ^ 2Hb:0 <= baseHi2:phi2 ih il < ([|Tn|] + 1) ^ 2s:=iter312_sqrt 31 (fun _ _ j : int31 => j) ih il Tn:int31Hs1:[|s|] ^ 2 <= phi2 ih ilHs2:phi2 ih il < ([|s|] + 1) ^ 2ih1, il1:int31Hihl1:phi2 ih1 il1 = [|s|] * [|s|]il2:int31Hil2:-1 * wB + [|il2|] = [|il|] - [|il1|]Hsih:[|ih - 1|] = [|ih|] - 1H1:[|ih1|] < [|ih - 1|]H2:[|ih1|] + 2 = [|ih|]([|ih1|] + 2) * base + (-1 * wB + [|il2|] + [|il1|]) - ([|ih1|] * base + [|il1|]) = base + [|il2|]
iszero
forall x : int31, ZnZ.eq0 x = true -> [|x|] = 0forall x : int31, ZnZ.eq0 x = true -> [|x|] = 0forall x : int31, match (x ?= 0)%int31 with | Eq => true | _ => false end = true -> [|x|] = 0x:int31H:match [|x|] ?= [|0|] with | Eq => true | _ => false end = true[|x|] = 0x:int31H:match [|x|] ?= 0 with | Eq => true | _ => false end = true[|x|] = 0now destruct ([|x|] ?= 0). Qed. (* Even *)x:int31H:match [|x|] ?= 0 with | Eq => true | _ => false end = true([|x|] ?= 0) = Eqforall x : int31, if ZnZ.is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1forall x : int31, if ZnZ.is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1x:int31if let (_, r) := (x / 2)%int31 in match (r ?= 0)%int31 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1x: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 = 1x, q, r:int31H: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 = 1x, q, r:int31H: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 = 1x, q, r:int31H:0 < 2 -> [|x|] = [|q|] * 2 + [|r|] /\ 0 <= [|r|] < 2if match [|r|] ?= [|0|] with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1x, q, r:int31H:0 < 2 -> [|x|] = [|q|] * 2 + [|r|] /\ 0 <= [|r|] < 2if match [|r|] ?= 0 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1x, q, r:int31H:[|x|] = [|q|] * 2 + [|r|]H0:0 <= [|r|] < 2if match [|r|] ?= 0 with | Eq => true | _ => false end then [|x|] mod 2 = 0 else [|x|] mod 2 = 1x, q, r:int31H:[|x|] = [|q|] * 2 + [|r|]H0:0 <= [|r|] < 2if match [|r|] ?= 0 with | Eq => true | _ => false end then [|r|] = 0 else [|r|] = 1x, q, r:int31H:[|x|] = [|q|] * 2 + [|r|]H0:0 <= [|r|] < 2[|r|] = [|x|] mod 2x, q, r:int31H0:0 <= [|r|] < 2if match [|r|] ?= 0 with | Eq => true | _ => false end then [|r|] = 0 else [|r|] = 1x, q, r:int31H:[|x|] = [|q|] * 2 + [|r|]H0:0 <= [|r|] < 2[|r|] = [|x|] mod 2apply Zmod_unique with [|q|]; auto with zarith. Qed. (* Bitwise *)x, q, r:int31H:[|x|] = [|q|] * 2 + [|r|]H0:0 <= [|r|] < 2[|r|] = [|x|] mod 2x:int31Z.log2 [|x|] < Z.of_nat sizex:int31Z.log2 [|x|] < Z.of_nat sizex:int31H:0 <= [|x|]H':[|x|] < wBZ.log2 [|x|] < Z.of_nat sizex:int31H:0 < [|x|]H':[|x|] < wBZ.log2 [|x|] < Z.of_nat sizex:int31H:0 = [|x|]H':[|x|] < wBZ.log2 [|x|] < Z.of_nat sizenow apply Z.log2_lt_pow2.x:int31H:0 < [|x|]H':[|x|] < wBZ.log2 [|x|] < Z.of_nat sizenow rewrite <- H. Qed.x:int31H:0 = [|x|]H':[|x|] < wBZ.log2 [|x|] < Z.of_nat sizex, 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:int31Z.lor [|x|] [|y|] mod wB = Z.lor [|x|] [|y|]x, y:int310 <= Z.lor [|x|] [|y|]x, y:int31Z.lor [|x|] [|y|] < wBapply Z.lor_nonneg; split; apply phi_bounded.x, y:int310 <= Z.lor [|x|] [|y|]x, y:int31Z.lor [|x|] [|y|] < wBx, y:int31Z.log2 (Z.lor [|x|] [|y|]) < Z.log2 wBx, y:int31Z.log2 (Z.lor [|x|] [|y|]) < Z.of_nat sizeapply Z.max_lub_lt; apply log2_phi_bounded. Qed.x, y:int31Z.max (Z.log2 [|x|]) (Z.log2 [|y|]) < Z.of_nat sizex, 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:int31Z.land [|x|] [|y|] mod wB = Z.land [|x|] [|y|]x, y:int310 <= Z.land [|x|] [|y|]x, y:int31Z.land [|x|] [|y|] < wBapply Z.land_nonneg; left; apply phi_bounded.x, y:int310 <= Z.land [|x|] [|y|]x, y:int31Z.land [|x|] [|y|] < wBx, y:int31Z.log2 (Z.land [|x|] [|y|]) < Z.log2 wBx, y:int31Z.log2 (Z.land [|x|] [|y|]) < Z.of_nat sizex, y:int31Z.log2 (Z.land [|x|] [|y|]) <= ?mx, y:int31?m < Z.of_nat sizeapply Z.min_lt_iff; left; apply log2_phi_bounded. Qed.x, y:int31Z.min (Z.log2 [|x|]) (Z.log2 [|y|]) < Z.of_nat sizex, 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:int31Z.lxor [|x|] [|y|] mod wB = Z.lxor [|x|] [|y|]x, y:int310 <= Z.lxor [|x|] [|y|]x, y:int31Z.lxor [|x|] [|y|] < wBapply Z.lxor_nonneg; split; intros; apply phi_bounded.x, y:int310 <= Z.lxor [|x|] [|y|]x, y:int31Z.lxor [|x|] [|y|] < wBx, y:int31Z.log2 (Z.lxor [|x|] [|y|]) < Z.log2 wBx, y:int31Z.log2 (Z.lxor [|x|] [|y|]) < Z.of_nat sizex, y:int31Z.log2 (Z.lxor [|x|] [|y|]) <= ?mx, y:int31?m < Z.of_nat sizeapply 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.x, y:int31Z.max (Z.log2 [|x|]) (Z.log2 [|y|]) < Z.of_nat size