include "ASM/Util.ma". include "ASM/Arithmetic.ma". definition nat_of_bool: bool → nat ≝ λb: bool. match b with [ true ⇒ 1 | false ⇒ 0 ]. lemma blah: ∀n: nat. ∀bv: BitVector n. ∀buffer: nat. nat_of_bitvector_aux n buffer bv = nat_of_bitvector n bv + (buffer * 2^n). #n #bv elim bv [1: #buffer elim buffer try % #buffer' #inductive_hypothesis normalize inductive_hypothesis >inductive_hypothesis [1: change with ( ? + (2 * buffer + 1) * ?) in ⊢ (??%?); change with ( ? + buffer * (2 * 2^n')) in ⊢ (???%); cases daemon ] ] cases daemon qed. lemma nat_of_bitvector_aux_hd_tl: ∀n: nat. ∀tl: BitVector n. ∀hd: bool. nat_of_bitvector (S n) (hd:::tl) = nat_of_bitvector n tl + (nat_of_bool hd * 2^n). #n #tl elim tl [1: #hd cases hd % |2: #n' #hd' #tl' #inductive_hypothesis #hd cases hd whd in ⊢ (??%?); normalize nodelta >inductive_hypothesis cases hd' normalize nodelta normalize in match (nat_of_bool ?); normalize in match (nat_of_bool ?); [4: normalize in match (2 * ?);