Skip to content

Commit

Permalink
feat: break theorems up into smaller lemmas
Browse files Browse the repository at this point in the history
  • Loading branch information
bollu committed Nov 11, 2024
1 parent 6878a2d commit 7d2bc58
Show file tree
Hide file tree
Showing 3 changed files with 33 additions and 15 deletions.
41 changes: 26 additions & 15 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1240,27 +1240,37 @@ theorem toNat_ushiftRight_lt (x : BitVec w) (n : Nat) (hn : n ≤ w) :
· apply hn
· apply Nat.pow_pos (by decide)

theorem toNat_ushiftRight_eq_zero (x : BitVec w) (n : Nat) (hn : w ≤ n) :
(x >>> n).toNat = 0 := by
rw [toNat_ushiftRight, Nat.shiftRight_eq_div_pow, Nat.div_eq_of_lt]
exact Nat.lt_of_lt_of_le x.isLt (Nat.pow_le_pow_of_le Nat.one_lt_two hn)

@[simp]
theorem toInt_ushiftRight_zero (x : BitVec w) :
(x >>> 0).toInt = x.toInt := by simp
/-- Shifting right by `n`, which is larger than the bitwidth `w` produces `0. -/
theorem ushiftRight_eq_zero (x : BitVec w) (n : Nat) (hn : w ≤ n) :
x >>> n = 0#w := by
simp [bv_toNat]
have : 2^w ≤ 2^n := Nat.pow_le_pow_of_le Nat.one_lt_two hn
rw [Nat.shiftRight_eq_div_pow, Nat.div_eq_of_lt (by omega)]


@[simp]
theorem toInt_ushiftRight_pos (x : BitVec w) (n : Nat) (hn : 0 < n) :
/--
Unsigned shift right by at least one bit makes the value of the bitvector less than or equal to `2^(w-1)`,
makes the interpretation of the bitvector `Int` and `Nat` agree.
-/
theorem toInt_ushiftRight_pos_eq_toNat_shiftRight (x : BitVec w) (n : Nat) (hn : 0 < n) :
(x >>> n).toInt = x.toNat >>> n := by
rw [toInt_eq_toNat_cond]
simp only [toNat_ushiftRight, ite_eq_left_iff, Nat.not_lt]
intros h
by_cases hn: n ≤ w
· have h1 := Nat.mul_lt_mul_of_pos_left (toNat_ushiftRight_lt x n hn) Nat.two_pos
rw [← Nat.pow_succ'] at h1
simp only [Nat.lt_of_lt_of_le h1 (Nat.pow_le_pow_of_le (Nat.one_lt_two) (show _ ≤ w by omega))]
simp
· rw [← toNat_ushiftRight]
simp [toNat_ushiftRight_eq_zero x n (by omega), Nat.two_pow_pos w]
simp only [toNat_ushiftRight, Nat.zero_lt_succ, Nat.mul_lt_mul_left] at h1
have : 2 ^ (w - n).succ ≤ 2^ w := Nat.pow_le_pow_of_le (by decide) (by omega)
have := show 2 * x.toNat >>> n < 2 ^ w by
omega
omega
· have : x.toNat >>> n = 0 := by
apply Nat.shiftRight_eq_zero
have : 2^w ≤ 2^n := Nat.pow_le_pow_of_le (by decide) (by omega)
omega
simp [this] at h
omega

@[simp]
theorem toFin_uShiftRight (x : BitVec w) (n : Nat) :
Expand All @@ -1269,7 +1279,8 @@ theorem toFin_uShiftRight (x : BitVec w) (n : Nat) :
by_cases hn : n < w
· simp [Nat.shiftRight_eq_div_pow, Nat.mod_eq_of_lt (Nat.pow_lt_pow_of_lt Nat.one_lt_two hn)]
· simp only [Nat.not_lt] at hn
simp [toNat_ushiftRight_eq_zero x n hn, Nat.dvd_iff_mod_eq_zero.mp (Nat.pow_dvd_pow 2 hn)]
rw [ushiftRight_eq_zero _ _ (by omega)]
simp [Nat.dvd_iff_mod_eq_zero.mp (Nat.pow_dvd_pow 2 hn)]

@[simp]
theorem getMsbD_ushiftRight {x : BitVec w} {i n : Nat} :
Expand Down
4 changes: 4 additions & 0 deletions src/Init/Data/Int/Bitwise/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -34,4 +34,8 @@ theorem shiftRight_eq_div_pow (m : Int) (n : Nat) :
theorem zero_shiftRight (n : Nat) : (0 : Int) >>> n = 0 := by
simp [Int.shiftRight_eq_div_pow]

@[simp]
theorem shiftRight_zero (n : Int) : n >>> 0 = n := by
simp [Int.shiftRight_eq_div_pow]

end Int
3 changes: 3 additions & 0 deletions src/Init/Data/Nat/Bitwise/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -71,6 +71,9 @@ theorem shiftRight_eq_div_pow (m : Nat) : ∀ n, m >>> n = m / 2 ^ n
rw [shiftRight_add, shiftRight_eq_div_pow m k]
simp [Nat.div_div_eq_div_mul, ← Nat.pow_succ, shiftRight_succ]

theorem shiftRight_eq_zero (m n : Nat) (hn : m < 2 ^n): m >>> n = 0 := by
simp [Nat.shiftRight_eq_div_pow, Nat.div_eq_of_lt hn]

/-!
### testBit
We define an operation for testing individual bits in the binary representation
Expand Down

0 comments on commit 7d2bc58

Please sign in to comment.