Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat: add BitVec.ofFn and lemmas #1078

Open
wants to merge 21 commits into
base: main
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from 15 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
1 change: 1 addition & 0 deletions Batteries.lean
Original file line number Diff line number Diff line change
Expand Up @@ -16,6 +16,7 @@ import Batteries.Data.Array
import Batteries.Data.AssocList
import Batteries.Data.BinaryHeap
import Batteries.Data.BinomialHeap
import Batteries.Data.BitVec
import Batteries.Data.ByteArray
import Batteries.Data.ByteSubarray
import Batteries.Data.Char
Expand Down
88 changes: 88 additions & 0 deletions Batteries/Data/BitVec.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,88 @@
/-
Copyright (c) 2024 François G. Dorais. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: François G. Dorais
-/
fgdorais marked this conversation as resolved.
Show resolved Hide resolved

import Batteries.Data.Fin.OfBits

namespace BitVec

theorem getElem_shiftConcat (v : BitVec n) (b : Bool) (i) (h : i < n) :
(v.shiftConcat b)[i] = if i = 0 then b else v[i-1] := by
rw [← getLsbD_eq_getElem, getLsbD_shiftConcat, getLsbD_eq_getElem, decide_eq_true h,
Bool.true_and]

@[simp] theorem getElem_shiftConcat_zero (v : BitVec n) (b : Bool) (h : 0 < n) :
(v.shiftConcat b)[0] = b := by simp [getElem_shiftConcat]

@[simp] theorem getElem_shiftConcat_succ (v : BitVec n) (b : Bool) (h : i + 1 < n) :
(v.shiftConcat b)[i+1] = v[i] := by simp [getElem_shiftConcat]

/-- `ofFnLEAux f` returns the `BitVec m` whose `i`th bit is `f i` when `i < m`, little endian. -/
@[inline] def ofFnLEAux (m : Nat) (f : Fin n → Bool) : BitVec m :=
Fin.foldr n (fun i v => v.shiftConcat (f i)) 0

@[simp] theorem toNat_ofFnLEAux (m : Nat) (f : Fin n → Bool) :
(ofFnLEAux m f).toNat = Nat.ofBits f % 2 ^ m := by
simp only [ofFnLEAux]
induction n with
| zero => rfl
| succ n ih =>
rw [Fin.foldr_succ, toNat_shiftConcat, Nat.shiftLeft_eq, Nat.pow_one, Nat.ofBits_succ, ih,
← Nat.mod_add_div (Nat.ofBits (f ∘ Fin.succ)) (2 ^ m), Nat.mul_add 2, Nat.add_right_comm,
Nat.mul_left_comm, Nat.add_mul_mod_self_left, Nat.mul_comm 2]
rfl

@[simp] theorem toFin_ofFnLEAux (m : Nat) (f : Fin n → Bool) :
(ofFnLEAux m f).toFin = Fin.ofNat' (2 ^ m) (Nat.ofBits f) := by
ext; simp

/-- `ofFnLE f` returns the `BitVec n` whose `i`th bit is `f i` with little endian ordering. -/
@[inline] def ofFnLE (f : Fin n → Bool) : BitVec n := ofFnLEAux n f

@[simp] theorem toNat_ofFnLE (f : Fin n → Bool) : (ofFnLE f).toNat = Nat.ofBits f := by
rw [ofFnLE, toNat_ofFnLEAux, Nat.mod_eq_of_lt (Nat.ofBits_lt_two_pow f)]

@[simp] theorem toFin_ofFnLE (f : Fin n → Bool) : (ofFnLE f).toFin = Fin.ofBits f := by
ext; simp

theorem getElem_ofFnLEAux (f : Fin n → Bool) (i) (h : i < n) (h' : i < m) :
(ofFnLEAux m f)[i] = f ⟨i, h⟩ := by
simp only [ofFnLEAux]
induction n generalizing i m with
| zero => contradiction
| succ n ih =>
simp only [Fin.foldr_succ, getElem_shiftConcat]
cases i with
| zero =>
simp
| succ i =>
rw [ih (fun i => f i.succ)] <;> try omega
simp

@[simp] theorem getElem_ofFnLE (f : Fin n → Bool) (i) (h : i < n) : (ofFnLE f)[i] = f ⟨i, h⟩ :=
getElem_ofFnLEAux ..

theorem getLsb'_ofFnLE (f : Fin n → Bool) (i) : (ofFnLE f).getLsb' i = f i := by simp

@[simp] theorem getMsb'_ofFnLE (f : Fin n → Bool) (i) : (ofFnLE f).getMsb' i = f i.rev := by
simp [getMsb'_eq_getLsb', Fin.rev]; congr 2; omega

/-- `ofFnBE f` returns the `BitVec n` whose `i`th bit is `f i` with big endian ordering. -/
@[inline] def ofFnBE (f : Fin n → Bool) : BitVec n := ofFnLE fun i => f i.rev

@[simp] theorem toNat_ofFnBE (f : Fin n → Bool) : (ofFnBE f).toNat = Nat.ofBits (f ∘ Fin.rev) := by
simp [ofFnBE]; rfl

@[simp] theorem toFin_ofFnBE (f : Fin n → Bool) : (ofFnBE f).toFin = Fin.ofBits (f ∘ Fin.rev) := by
ext; simp

@[simp] theorem getElem_ofFnBE (f : Fin n → Bool) (i) (h : i < n) :
(ofFnBE f)[i] = f (Fin.rev ⟨i, h⟩) := by simp [ofFnBE]

theorem getLsb'_ofFnBE (f : Fin n → Bool) (i) : (ofFnBE f).getLsb' i = f i.rev := by
simp

@[simp] theorem getMsb'_ofFnBE (f : Fin n → Bool) (i) : (ofFnBE f).getMsb' i = f i := by
simp [ofFnBE]
1 change: 1 addition & 0 deletions Batteries/Data/Fin.lean
Original file line number Diff line number Diff line change
@@ -1,3 +1,4 @@
import Batteries.Data.Fin.Basic
import Batteries.Data.Fin.Fold
import Batteries.Data.Fin.Lemmas
import Batteries.Data.Fin.OfBits
16 changes: 16 additions & 0 deletions Batteries/Data/Fin/OfBits.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,16 @@
/-
Copyright (c) 2025 François G. Dorais. All rights reserved.
Released under Apache 2. license as described in the file LICENSE.
Authors: François G. Dorais
-/

import Batteries.Data.Nat.Lemmas

namespace Fin

/--
Construct an element of `Fin (2 ^ n)` from a sequence of bits (little endian).
-/
abbrev ofBits (f : Fin n → Bool) : Fin (2 ^ n) := ⟨Nat.ofBits f, Nat.ofBits_lt_two_pow f⟩

@[simp] theorem val_ofBits (f : Fin n → Bool) : (ofBits f).val = Nat.ofBits f := rfl
7 changes: 6 additions & 1 deletion Batteries/Data/Nat/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,6 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro

namespace Nat


/--
Recursor identical to `Nat.recOn` but uses notations `0` for `Nat.zero` and `·+1` for `Nat.succ`
-/
Expand Down Expand Up @@ -103,3 +102,9 @@ where
else
guess
termination_by guess

/--
Construct a natural number from a sequence of bits using little endian convention.
-/
@[inline] def ofBits (f : Fin n → Bool) : Nat :=
Fin.foldr n (fun i v => 2 * v + (f i).toNat) 0
42 changes: 42 additions & 0 deletions Batteries/Data/Nat/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -162,3 +162,45 @@ protected def sum_trichotomy (a b : Nat) : a < b ⊕' a = b ⊕' b < a :=

@[simp] theorem sum_append {l₁ l₂ : List Nat}: (l₁ ++ l₂).sum = l₁.sum + l₂.sum := by
induction l₁ <;> simp [*, Nat.add_assoc]

/-! ### ofBits -/

@[simp] theorem ofBits_zero (f : Fin 0 → Bool) : ofBits f = 0 := rfl

theorem ofBits_succ (f : Fin (n+1) → Bool) : ofBits f = 2 * ofBits (f ∘ Fin.succ) + (f 0).toNat :=
Fin.foldr_succ ..

theorem ofBits_lt_two_pow (f : Fin n → Bool) : ofBits f < 2 ^ n := by
induction n with
| zero => simp
| succ n ih =>
calc ofBits f
= 2 * ofBits (f ∘ Fin.succ) + (f 0).toNat := ofBits_succ ..
_ < 2 * (ofBits (f ∘ Fin.succ) + 1) := Nat.add_lt_add_left (Bool.toNat_lt _) ..
_ ≤ 2 * 2 ^ n := Nat.mul_le_mul_left 2 (ih ..)
_ = 2 ^ (n + 1) := Nat.pow_add_one' .. |>.symm

@[simp] theorem testBit_ofBits_lt (f : Fin n → Bool) (i : Nat) (h : i < n) :
(ofBits f).testBit i = f ⟨i, h⟩ := by
induction n generalizing i with
| zero => contradiction
| succ n ih =>
simp only [ofBits_succ]
match i with
| 0 => simp [mod_eq_of_lt (Bool.toNat_lt _)]
| i+1 =>
rw [testBit_add_one, mul_add_div Nat.zero_lt_two, Nat.div_eq_of_lt (Bool.toNat_lt _)]
exact ih (f ∘ Fin.succ) i (Nat.lt_of_succ_lt_succ h)

@[simp] theorem testBit_ofBits_ge (f : Fin n → Bool) (i : Nat) (h : n ≤ i) :
(ofBits f).testBit i = false := by
apply testBit_lt_two_pow
apply Nat.lt_of_lt_of_le
· exact ofBits_lt_two_pow f
· exact pow_le_pow_of_le_right Nat.zero_lt_two h

theorem testBit_ofBits (f : Fin n → Bool) :
(ofBits f).testBit i = if h : i < n then f ⟨i, h⟩ else false := by
cases Nat.lt_or_ge i n with
| inl h => simp [h]
| inr h => simp [h, Nat.not_lt_of_ge h]
Loading