Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(* -*- coding: utf-8 -*- *)
(************************************************************************)
(* * The Coq Proof Assistant / The Coq Development Team *)
(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
Bit vectors interpreted as integers.
Contribution by Jean Duprat (ENS Lyon).
Require Import Bvector. Require Import ZArith. Require Export Zpower. Require Import Omega.
The evaluation of boolean vector is done both in binary and
two's complement. The computed number belongs to Z.
We hence use Omega to perform computations in Z.
Moreover, we use functions 2^n where n is a natural number
(here the vector length).
Section VALUE_OF_BOOLEAN_VECTORS.
Computations are done in the usual convention.
The values correspond either to the binary coding (nat) or
to the two's complement coding (int).
We perform the computation via Horner scheme.
The two's complement coding only makes sense on vectors whose
size is greater or equal to one (a sign bit should be present).
Definition bit_value (b:bool) : Z := match b with | true => 1%Z | false => 0%Z end.forall n : nat, Bvector n -> Zforall n : nat, Bvector n -> ZH:Bvector 0Zn:natH:Bvector n -> ZH0:Bvector (S n)Zn:natH:Bvector n -> ZH0:Bvector (S n)Zexact (bit_value h + 2 * H H2)%Z. Defined.n:natH:Bvector n -> ZH0:Bvector (S n)n0:nath:boolH2:Vector.t bool nH1:n0 = nZforall n : nat, Bvector (S n) -> Zforall n : nat, Bvector (S n) -> Zn:natH:Bvector 1Zn, n0:natH:Bvector (S n0) -> ZH0:Bvector (S (S n0))Zn:natH:Bvector 1n0:nath:boolH1:Vector.t bool 0H0:n0 = 0Zn, n0:natH:Bvector (S n0) -> ZH0:Bvector (S (S n0))Zn, n0:natH:Bvector (S n0) -> ZH0:Bvector (S (S n0))Zexact (bit_value h + 2 * H H2)%Z. Defined. End VALUE_OF_BOOLEAN_VECTORS. Section ENCODING_VALUE.n, n0:natH:Bvector (S n0) -> ZH0:Bvector (S (S n0))n1:nath:boolH2:Vector.t bool (S n0)H1:n1 = S n0Z
We compute the binary value via a Horner scheme.
Computation stops at the vector length without checks.
We define a function Zmod2 similar to Z.div2 returning the
quotient of division z=2q+r with 0<=r<=1.
The two's complement value is also computed via a Horner scheme
with Zmod2, the parameter is the size minus one.
Definition Zmod2 (z:Z) := match z with | Z0 => 0%Z | Zpos p => match p with | xI q => Zpos q | xO q => Zpos q | xH => 0%Z end | Zneg p => match p with | xI q => (Zneg q - 1)%Z | xO q => Zneg q | xH => (-1)%Z end end.forall z : Z, z = (2 * Zmod2 z + bit_value (Z.odd z))%Zforall z : Z, z = (2 * Zmod2 z + bit_value (Z.odd z))%Z0%Z = 0%Zp:positiveZ.pos p = (match match p with | (q~1)%positive | (q~0)%positive => Z.pos q | 1%positive => 0 end with | 0 => 0 | Z.pos y' => Z.pos y'~0 | Z.neg y' => Z.neg y'~0 end + bit_value match p with | (_~0)%positive => false | _ => true end)%Zp:positiveZ.neg p = (match match p with | (q~1)%positive => Z.neg (q + 1) | (q~0)%positive => Z.neg q | 1%positive => -1 end with | 0 => 0 | Z.pos y' => Z.pos y'~0 | Z.neg y' => Z.neg y'~0 end + bit_value match p with | (_~0)%positive => false | _ => true end)%Zp:positiveZ.pos p = (match match p with | (q~1)%positive | (q~0)%positive => Z.pos q | 1%positive => 0 end with | 0 => 0 | Z.pos y' => Z.pos y'~0 | Z.neg y' => Z.neg y'~0 end + bit_value match p with | (_~0)%positive => false | _ => true end)%Zp:positiveZ.neg p = (match match p with | (q~1)%positive => Z.neg (q + 1) | (q~0)%positive => Z.neg q | 1%positive => -1 end with | 0 => 0 | Z.pos y' => Z.pos y'~0 | Z.neg y' => Z.neg y'~0 end + bit_value match p with | (_~0)%positive => false | _ => true end)%Zp:positiveZ.neg p = (match match p with | (q~1)%positive => Z.neg (q + 1) | (q~0)%positive => Z.neg q | 1%positive => -1 end with | 0 => 0 | Z.pos y' => Z.pos y'~0 | Z.neg y' => Z.neg y'~0 end + bit_value match p with | (_~0)%positive => false | _ => true end)%Zp:positiveZ.neg p~1 = Z.neg (Pos.pred_double (p + 1))p:positiveZ.neg p~0 = Z.neg p~0(-1)%Z = (-1)%Zp:positiveZ.neg p~1~1 = Z.neg (Pos.pred_double (Pos.succ p))~1p:positiveZ.neg p~0~1 = Z.neg p~0~1(-3)%Z = (-3)%Zp:positiveZ.neg p~0 = Z.neg p~0(-1)%Z = (-1)%Zp:positiveZ.neg p~0~1 = Z.neg p~0~1(-3)%Z = (-3)%Zp:positiveZ.neg p~0 = Z.neg p~0(-1)%Z = (-1)%Z(-3)%Z = (-3)%Zp:positiveZ.neg p~0 = Z.neg p~0(-1)%Z = (-1)%Zp:positiveZ.neg p~0 = Z.neg p~0(-1)%Z = (-1)%Ztrivial. Qed.(-1)%Z = (-1)%Zforall n : nat, Z -> Bvector nforall n : nat, Z -> Bvector nn:natH:ZBvector 0n, n0:natH:Z -> Bvector n0H0:ZBvector (S n0)exact (Bcons (Z.odd H0) n0 (H (Z.div2 H0))). Defined.n, n0:natH:Z -> Bvector n0H0:ZBvector (S n0)forall n : nat, Z -> Bvector (S n)forall n : nat, Z -> Bvector (S n)n:natH:ZBvector 1n, n0:natH:Z -> Bvector (S n0)H0:ZBvector (S (S n0))exact (Bcons (Z.odd H0) (S n0) (H (Zmod2 H0))). Defined. End ENCODING_VALUE. Section Z_BRIC_A_BRAC.n, n0:natH:Z -> Bvector (S n0)H0:ZBvector (S (S n0))
Some auxiliary lemmas used in the next section. Large use of ZArith.
Deserve to be properly rewritten.
forall (n : nat) (b : bool) (bv : Bvector n), binary_value (S n) (b :: bv) = (bit_value b + 2 * binary_value n bv)%Zintros; auto. Qed.forall (n : nat) (b : bool) (bv : Bvector n), binary_value (S n) (b :: bv) = (bit_value b + 2 * binary_value n bv)%Zforall (n : nat) (b : bool) (z : Z), (z >= 0)%Z -> Z_to_binary (S n) (bit_value b + 2 * z) = Bcons b n (Z_to_binary n z)forall (n : nat) (b : bool) (z : Z), (z >= 0)%Z -> Z_to_binary (S n) (bit_value b + 2 * z) = Bcons b n (Z_to_binary n z)intro H; elim H; trivial. Qed.n:natp:positive(Z.neg p >= 0)%Z -> Bcons match Pos.pred_double p with | (_~0)%positive => false | _ => true end n (Z_to_binary n (Z.neg (Pos.div2_up (Pos.pred_double p)))) = Bcons true n (Z_to_binary n (Z.neg p))forall (n : nat) (bv : Bvector n), (binary_value n bv >= 0)%Zforall (n : nat) (bv : Bvector n), (binary_value n bv >= 0)%Z(0 >= 0)%Za:booln:natv:VectorDef.t bool nIHbv:(binary_value n v >= 0)%Z(bit_value a + match binary_value n v with | 0 => 0 | Z.pos y' => Z.pos y'~0 | Z.neg y' => Z.neg y'~0 end >= 0)%Za:booln:natv:VectorDef.t bool nIHbv:(binary_value n v >= 0)%Z(bit_value a + match binary_value n v with | 0 => 0 | Z.pos y' => Z.pos y'~0 | Z.neg y' => Z.neg y'~0 end >= 0)%Zauto with zarith. Qed.n:natv:VectorDef.t bool nIHbv:(0 >= 0)%Z(1 >= 0)%Zforall (n : nat) (bv : Bvector (S n)) (b : bool), two_compl_value (S n) (Bcons b (S n) bv) = (bit_value b + 2 * two_compl_value n bv)%Zintros; auto. Qed.forall (n : nat) (bv : Bvector (S n)) (b : bool), two_compl_value (S n) (Bcons b (S n) bv) = (bit_value b + 2 * two_compl_value n bv)%Zforall (n : nat) (b : bool) (z : Z), Z_to_two_compl (S n) (bit_value b + 2 * z) = Bcons b (S n) (Z_to_two_compl n z)forall (n : nat) (b : bool) (z : Z), Z_to_two_compl (S n) (bit_value b + 2 * z) = Bcons b (S n) (Z_to_two_compl n z)n:natp:positiveZ_to_two_compl (S n) (bit_value true + 2 * Z.neg p) = Bcons true (S n) (Z_to_two_compl n (Z.neg p))n:natp:positiveZ_to_two_compl (S n) (bit_value true + 2 * Z.neg p~0) = Bcons true (S n) (Z_to_two_compl n (Z.neg p~0))intros; rewrite (Pos.succ_pred_double p); trivial. Qed.n:natp:positiveBcons true (S n) (Z_to_two_compl n (Z.neg (Pos.succ (Pos.pred_double p))~0)) = Bcons true (S n) (Z_to_two_compl n (Z.neg p~0~0))forall (n : nat) (z : Z), Z_to_binary (S n) z = Bcons (Z.odd z) n (Z_to_binary n (Z.div2 z))intros; auto. Qed.forall (n : nat) (z : Z), Z_to_binary (S n) z = Bcons (Z.odd z) n (Z_to_binary n (Z.div2 z))forall z : Z, (z >= 0)%Z -> (bit_value (Z.odd z) + 2 * Z.div2 z)%Z = zforall z : Z, (z >= 0)%Z -> (bit_value (Z.odd z) + 2 * Z.div2 z)%Z = zp:positive(Z.pos p >= 0)%Z -> (bit_value (Z.odd (Z.pos p)) + 2 * Z.div2 (Z.pos p))%Z = Z.pos pp:positive(Z.neg p >= 0)%Z -> (bit_value (Z.odd (Z.neg p)) + 2 * Z.div2 (Z.neg p))%Z = Z.neg pintro H; elim H; trivial. Qed.p:positive(Z.neg p >= 0)%Z -> (bit_value (Z.odd (Z.neg p)) + 2 * Z.div2 (Z.neg p))%Z = Z.neg pforall z : Z, (z >= 0)%Z -> (Z.div2 z >= 0)%Zforall z : Z, (z >= 0)%Z -> (Z.div2 z >= 0)%Z(0 >= 0)%Z -> (Z.div2 0 >= 0)%Zp:positive(Z.pos p >= 0)%Z -> (Z.div2 (Z.pos p) >= 0)%Zp:positive(Z.neg p >= 0)%Z -> (Z.div2 (Z.neg p) >= 0)%Zp:positive(Z.pos p >= 0)%Z -> (Z.div2 (Z.pos p) >= 0)%Zp:positive(Z.neg p >= 0)%Z -> (Z.div2 (Z.neg p) >= 0)%Z(1 >= 0)%Z -> (Z.div2 1 >= 0)%Zp:positive(Z.neg p >= 0)%Z -> (Z.div2 (Z.neg p) >= 0)%Zintro H; elim H; trivial. Qed.p:positive(Z.neg p >= 0)%Z -> (Z.div2 (Z.neg p) >= 0)%Zforall (z : Z) (n : nat), (z >= 0)%Z -> (z < two_power_nat (S n))%Z -> (Z.div2 z < two_power_nat n)%Zforall (z : Z) (n : nat), (z >= 0)%Z -> (z < two_power_nat (S n))%Z -> (Z.div2 z < two_power_nat n)%Zz:Zn:natH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(Z.div2 z < two_power_nat n)%Zz:Zn:natH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(2 * Z.div2 z < 2 * two_power_nat n)%Zz:Zn:natH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(2 * Z.div2 z < two_power_nat (S n))%Zz:Zn:natH:(z >= 0)%ZH0:(z < two_power_nat (S n))%ZHeven:Zeven z(2 * Z.div2 z < two_power_nat (S n))%Zz:Zn:natH:(z >= 0)%ZH0:(z < two_power_nat (S n))%ZHodd:Zodd z(2 * Z.div2 z < two_power_nat (S n))%Zgeneralize (Zeven.Zodd_div2 z Hodd); omega. Qed.z:Zn:natH:(z >= 0)%ZH0:(z < two_power_nat (S n))%ZHodd:Zodd z(2 * Z.div2 z < two_power_nat (S n))%Zforall (n : nat) (z : Z), Z_to_two_compl (S n) z = Bcons (Z.odd z) (S n) (Z_to_two_compl n (Zmod2 z))intros; auto. Qed.forall (n : nat) (z : Z), Z_to_two_compl (S n) z = Bcons (Z.odd z) (S n) (Z_to_two_compl n (Zmod2 z))forall z : Z, Zeven z -> bit_value (Z.odd z) = 0%Zforall z : Z, Zeven z -> bit_value (Z.odd z) = 0%Zp:positiveZeven (Z.pos p) -> (if Z.odd (Z.pos p) then 1%Z else 0%Z) = 0%Zp:positiveZeven (Z.neg p) -> (if Z.odd (Z.neg p) then 1%Z else 0%Z) = 0%Zdestruct p; tauto || (intro H; elim H). Qed.p:positiveZeven (Z.neg p) -> (if Z.odd (Z.neg p) then 1%Z else 0%Z) = 0%Zforall z : Z, Zodd z -> bit_value (Z.odd z) = 1%Zforall z : Z, Zodd z -> bit_value (Z.odd z) = 1%ZZodd 0 -> (if Z.odd 0 then 1%Z else 0%Z) = 1%Zp:positiveZodd (Z.pos p) -> (if Z.odd (Z.pos p) then 1%Z else 0%Z) = 1%Zp:positiveZodd (Z.neg p) -> (if Z.odd (Z.neg p) then 1%Z else 0%Z) = 1%Zp:positiveZodd (Z.pos p) -> (if Z.odd (Z.pos p) then 1%Z else 0%Z) = 1%Zp:positiveZodd (Z.neg p) -> (if Z.odd (Z.neg p) then 1%Z else 0%Z) = 1%Zdestruct p; tauto || (intros; elim H). Qed.p:positiveZodd (Z.neg p) -> (if Z.odd (Z.neg p) then 1%Z else 0%Z) = 1%Zforall (n : nat) (z : Z), (z >= - two_power_nat (S n))%Z -> (Zmod2 z >= - two_power_nat n)%Zforall (n : nat) (z : Z), (z >= - two_power_nat (S n))%Z -> (Zmod2 z >= - two_power_nat n)%Zn:natz:Z(z >= - (2 * two_power_nat n))%Z -> (Zmod2 z >= - two_power_nat n)%Zn:natz:Zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z >= - (2 * two_power_nat n))%Z -> (Zmod2 z >= - two_power_nat n)%Zn:natz:ZH:Zeven zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z >= - (2 * two_power_nat n))%Z -> (Zmod2 z >= - two_power_nat n)%Zn:natz:ZH:Zodd zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z >= - (2 * two_power_nat n))%Z -> (Zmod2 z >= - two_power_nat n)%Zrewrite (Zodd_bit_value z H); intros; omega. Qed.n:natz:ZH:Zodd zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z >= - (2 * two_power_nat n))%Z -> (Zmod2 z >= - two_power_nat n)%Zforall (n : nat) (z : Z), (z < two_power_nat (S n))%Z -> (Zmod2 z < two_power_nat n)%Zforall (n : nat) (z : Z), (z < two_power_nat (S n))%Z -> (Zmod2 z < two_power_nat n)%Zn:natz:Z(z < 2 * two_power_nat n)%Z -> (Zmod2 z < two_power_nat n)%Zn:natz:Zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z < 2 * two_power_nat n)%Z -> (Zmod2 z < two_power_nat n)%Zn:natz:ZH:Zeven zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z < 2 * two_power_nat n)%Z -> (Zmod2 z < two_power_nat n)%Zn:natz:ZH:Zodd zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z < 2 * two_power_nat n)%Z -> (Zmod2 z < two_power_nat n)%Zrewrite (Zodd_bit_value z H); intros; omega. Qed. End Z_BRIC_A_BRAC. Section COHERENT_VALUE.n:natz:ZH:Zodd zz = (2 * Zmod2 z + bit_value (Z.odd z))%Z -> (z < 2 * two_power_nat n)%Z -> (Zmod2 z < two_power_nat n)%Z
We check that the functions are reciprocal on the definition interval.
This uses earlier library lemmas.
forall (n : nat) (bv : Bvector n), Z_to_binary n (binary_value n bv) = bvforall (n : nat) (bv : Bvector n), Z_to_binary n (binary_value n bv) = bvZ_to_binary 0 (binary_value 0 []) = []a:booln:natbv:VectorDef.t bool nIHbv:Z_to_binary n (binary_value n bv) = bvZ_to_binary (S n) (binary_value (S n) (a :: bv)) = a :: bva:booln:natbv:VectorDef.t bool nIHbv:Z_to_binary n (binary_value n bv) = bvZ_to_binary (S n) (binary_value (S n) (a :: bv)) = a :: bva:booln:natbv:VectorDef.t bool nIHbv:Z_to_binary n (binary_value n bv) = bvZ_to_binary (S n) (bit_value a + 2 * binary_value n bv) = a :: bva:booln:natbv:VectorDef.t bool nIHbv:Z_to_binary n (binary_value n bv) = bvBcons a n (Z_to_binary n (binary_value n bv)) = a :: bva:booln:natbv:VectorDef.t bool nIHbv:Z_to_binary n (binary_value n bv) = bv(binary_value n bv >= 0)%Zapply binary_value_pos. Qed.a:booln:natbv:VectorDef.t bool nIHbv:Z_to_binary n (binary_value n bv) = bv(binary_value n bv >= 0)%Zforall (n : nat) (bv : Bvector n) (b : bool), Z_to_two_compl n (two_compl_value n (Bcons b n bv)) = Bcons b n bvforall (n : nat) (bv : Bvector n) (b : bool), Z_to_two_compl n (two_compl_value n (Bcons b n bv)) = Bcons b n bvb:boolZ_to_two_compl 0 (two_compl_value 0 (Bcons b 0 [])) = Bcons b 0 []a:booln:natbv:VectorDef.t bool nIHbv:forall b0 : bool, Z_to_two_compl n (two_compl_value n (Bcons b0 n bv)) = Bcons b0 n bvb:boolZ_to_two_compl (S n) (two_compl_value (S n) (Bcons b (S n) (a :: bv))) = Bcons b (S n) (a :: bv)a:booln:natbv:VectorDef.t bool nIHbv:forall b0 : bool, Z_to_two_compl n (two_compl_value n (Bcons b0 n bv)) = Bcons b0 n bvb:boolZ_to_two_compl (S n) (two_compl_value (S n) (Bcons b (S n) (a :: bv))) = Bcons b (S n) (a :: bv)a:booln:natbv:VectorDef.t bool nIHbv:forall b0 : bool, Z_to_two_compl n (two_compl_value n (Bcons b0 n bv)) = Bcons b0 n bvb:boolZ_to_two_compl (S n) (bit_value b + 2 * two_compl_value n (a :: bv)) = Bcons b (S n) (a :: bv)rewrite IHbv; trivial. Qed.a:booln:natbv:VectorDef.t bool nIHbv:forall b0 : bool, Z_to_two_compl n (two_compl_value n (Bcons b0 n bv)) = Bcons b0 n bvb:boolBcons b (S n) (Z_to_two_compl n (two_compl_value n (a :: bv))) = Bcons b (S n) (a :: bv)forall (n : nat) (z : Z), (z >= 0)%Z -> (z < two_power_nat n)%Z -> binary_value n (Z_to_binary n z) = zforall (n : nat) (z : Z), (z >= 0)%Z -> (z < two_power_nat n)%Z -> binary_value n (Z_to_binary n z) = zforall z : Z, (z >= 0)%Z -> (z < two_power_nat 0)%Z -> binary_value 0 (Z_to_binary 0 z) = zn:natIHn:forall z : Z, (z >= 0)%Z -> (z < two_power_nat n)%Z -> binary_value n (Z_to_binary n z) = zforall z : Z, (z >= 0)%Z -> (z < two_power_nat (S n))%Z -> binary_value (S n) (Z_to_binary (S n) z) = zn:natIHn:forall z : Z, (z >= 0)%Z -> (z < two_power_nat n)%Z -> binary_value n (Z_to_binary n z) = zforall z : Z, (z >= 0)%Z -> (z < two_power_nat (S n))%Z -> binary_value (S n) (Z_to_binary (S n) z) = zn:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Zbinary_value (S n) (Bcons (Z.odd z) n (Z_to_binary n (Z.div2 z))) = zn:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(bit_value (Z.odd z) + 2 * binary_value n (Z_to_binary n (Z.div2 z)))%Z = zn:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(bit_value (Z.odd z) + 2 * Z.div2 z)%Z = zn:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(Z.div2 z >= 0)%Zn:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(Z.div2 z < two_power_nat n)%Zn:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(Z.div2 z >= 0)%Zn:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(Z.div2 z < two_power_nat n)%Zapply Zdiv2_two_power_nat; trivial. Qed.n:natIHn:forall z0 : Z, (z0 >= 0)%Z -> (z0 < two_power_nat n)%Z -> binary_value n (Z_to_binary n z0) = z0z:ZH:(z >= 0)%ZH0:(z < two_power_nat (S n))%Z(Z.div2 z < two_power_nat n)%Zforall (n : nat) (z : Z), (z >= - two_power_nat n)%Z -> (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = zforall (n : nat) (z : Z), (z >= - two_power_nat n)%Z -> (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = zforall z : Z, (z >= - two_power_nat 0)%Z -> (z < two_power_nat 0)%Z -> two_compl_value 0 (Z_to_two_compl 0 z) = zn:natIHn:forall z : Z, (z >= - two_power_nat n)%Z -> (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = zforall z : Z, (z >= - two_power_nat (S n))%Z -> (z < two_power_nat (S n))%Z -> two_compl_value (S n) (Z_to_two_compl (S n) z) = zz:ZH:(z >= -1)%ZH0:(z < 1)%Zeq_rec_r (fun n : nat => bool -> Vector.t bool n -> Z) (fun (h : bool) (_ : Vector.t bool 0) => (- bit_value h)%Z) eq_refl (Z.odd z) Bnil = zn:natIHn:forall z : Z, (z >= - two_power_nat n)%Z -> (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = zforall z : Z, (z >= - two_power_nat (S n))%Z -> (z < two_power_nat (S n))%Z -> two_compl_value (S n) (Z_to_two_compl (S n) z) = zz:ZH:(z >= -1)%ZH0:(z < 1)%Zz = (-1)%Z \/ z = 0%Zz:ZH:(z >= -1)%ZH0:(z < 1)%ZH1:z = (-1)%Z \/ z = 0%Zeq_rec_r (fun n : nat => bool -> Vector.t bool n -> Z) (fun (h : bool) (_ : Vector.t bool 0) => (- bit_value h)%Z) eq_refl (Z.odd z) Bnil = zn:natIHn:forall z : Z, (z >= - two_power_nat n)%Z -> (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = zforall z : Z, (z >= - two_power_nat (S n))%Z -> (z < two_power_nat (S n))%Z -> two_compl_value (S n) (Z_to_two_compl (S n) z) = zz:ZH:(z >= -1)%ZH0:(z < 1)%ZH1:z = (-1)%Z \/ z = 0%Zeq_rec_r (fun n : nat => bool -> Vector.t bool n -> Z) (fun (h : bool) (_ : Vector.t bool 0) => (- bit_value h)%Z) eq_refl (Z.odd z) Bnil = zn:natIHn:forall z : Z, (z >= - two_power_nat n)%Z -> (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = zforall z : Z, (z >= - two_power_nat (S n))%Z -> (z < two_power_nat (S n))%Z -> two_compl_value (S n) (Z_to_two_compl (S n) z) = zn:natIHn:forall z : Z, (z >= - two_power_nat n)%Z -> (z < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z) = zforall z : Z, (z >= - two_power_nat (S n))%Z -> (z < two_power_nat (S n))%Z -> two_compl_value (S n) (Z_to_two_compl (S n) z) = zn:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Ztwo_compl_value (S n) (Bcons (Z.odd z) (S n) (Z_to_two_compl n (Zmod2 z))) = zn:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Z(bit_value (Z.odd z) + 2 * two_compl_value n (Z_to_two_compl n (Zmod2 z)))%Z = zn:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Z(bit_value (Z.odd z) + 2 * Zmod2 z)%Z = zn:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Z(Zmod2 z >= - two_power_nat n)%Zn:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Z(Zmod2 z < two_power_nat n)%Zn:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Z(Zmod2 z >= - two_power_nat n)%Zn:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Z(Zmod2 z < two_power_nat n)%Zapply Zlt_two_power_nat_S; auto. Qed. End COHERENT_VALUE.n:natIHn:forall z0 : Z, (z0 >= - two_power_nat n)%Z -> (z0 < two_power_nat n)%Z -> two_compl_value n (Z_to_two_compl n z0) = z0z:ZH:(z >= - two_power_nat (S n))%ZH0:(z < two_power_nat (S n))%Z(Zmod2 z < two_power_nat n)%Z