Skip to content

Commit d2c4471

Browse files
alexkeizerkim-em
andauthored
feat: BitVec.{toInt, toFin, msb}_udiv (#6402)
This PR adds a `toFin` and `msb` lemma for unsigned bitvector division. We *don't* have `toInt_udiv`, since the only truly general statement we can make does no better than unfolding the definition, and it's not uncontroversially clear how to unfold `toInt` (see `toInt_eq_msb_cond`/`toInt_eq_toNat_cond`/`toInt_eq_toNat_bmod` for a few options currently provided). Instead, we do have `toInt_udiv_of_msb` that's able to provide a more meaningful rewrite given an extra side-condition (that `x.msb = false`). This PR also upstreams a minor `Nat` theorem (`Nat.div_le_div_left`) needed for the above from Mathlib. --------- Co-authored-by: Kim Morrison <[email protected]>
1 parent c07948a commit d2c4471

File tree

2 files changed

+66
-5
lines changed

2 files changed

+66
-5
lines changed

src/Init/Data/BitVec/Lemmas.lean

+58-5
Original file line numberDiff line numberDiff line change
@@ -9,6 +9,7 @@ import Init.Data.Bool
99
import Init.Data.BitVec.Basic
1010
import Init.Data.Fin.Lemmas
1111
import Init.Data.Nat.Lemmas
12+
import Init.Data.Nat.Div.Lemmas
1213
import Init.Data.Nat.Mod
1314
import Init.Data.Int.Bitwise.Lemmas
1415
import Init.Data.Int.Pow
@@ -442,6 +443,10 @@ theorem toInt_eq_toNat_cond (x : BitVec n) :
442443
(x.toNat : Int) - (2^n : Nat) :=
443444
rfl
444445

446+
theorem toInt_eq_toNat_of_lt {x : BitVec n} (h : 2 * x.toNat < 2^n) :
447+
x.toInt = x.toNat := by
448+
simp [toInt_eq_toNat_cond, h]
449+
445450
theorem msb_eq_false_iff_two_mul_lt {x : BitVec w} : x.msb = false2 * x.toNat < 2^w := by
446451
cases w <;> simp [Nat.pow_succ, Nat.mul_comm _ 2, msb_eq_decide, toNat_of_zero_length]
447452

@@ -454,6 +459,9 @@ theorem toInt_eq_msb_cond (x : BitVec w) :
454459
simp only [BitVec.toInt, ← msb_eq_false_iff_two_mul_lt]
455460
cases x.msb <;> rfl
456461

462+
theorem toInt_eq_toNat_of_msb {x : BitVec w} (h : x.msb = false) :
463+
x.toInt = x.toNat := by
464+
simp [toInt_eq_msb_cond, h]
457465

458466
theorem toInt_eq_toNat_bmod (x : BitVec n) : x.toInt = Int.bmod x.toNat (2^n) := by
459467
simp only [toInt_eq_toNat_cond]
@@ -2300,6 +2308,12 @@ theorem ofNat_sub_ofNat {n} (x y : Nat) : BitVec.ofNat n x - BitVec.ofNat n y =
23002308
@[simp, bv_toNat] theorem toNat_neg (x : BitVec n) : (- x).toNat = (2^n - x.toNat) % 2^n := by
23012309
simp [Neg.neg, BitVec.neg]
23022310

2311+
theorem toNat_neg_of_pos {x : BitVec n} (h : 0#n < x) :
2312+
(- x).toNat = 2^n - x.toNat := by
2313+
change 0 < x.toNat at h
2314+
rw [toNat_neg, Nat.mod_eq_of_lt]
2315+
omega
2316+
23032317
theorem toInt_neg {x : BitVec w} :
23042318
(-x).toInt = (-x.toInt).bmod (2 ^ w) := by
23052319
rw [← BitVec.zero_sub, toInt_sub]
@@ -2586,13 +2600,13 @@ theorem udiv_def {x y : BitVec n} : x / y = BitVec.ofNat n (x.toNat / y.toNat) :
25862600
rw [← udiv_eq]
25872601
simp [udiv, bv_toNat, h, Nat.mod_eq_of_lt]
25882602

2603+
@[simp]
2604+
theorem toFin_udiv {x y : BitVec n} : (x / y).toFin = x.toFin / y.toFin := by
2605+
rfl
2606+
25892607
@[simp, bv_toNat]
25902608
theorem toNat_udiv {x y : BitVec n} : (x / y).toNat = x.toNat / y.toNat := by
2591-
rw [udiv_def]
2592-
by_cases h : y = 0
2593-
· simp [h]
2594-
· rw [toNat_ofNat, Nat.mod_eq_of_lt]
2595-
exact Nat.lt_of_le_of_lt (Nat.div_le_self ..) (by omega)
2609+
rfl
25962610

25972611
@[simp]
25982612
theorem zero_udiv {x : BitVec w} : (0#w) / x = 0#w := by
@@ -2628,6 +2642,45 @@ theorem udiv_self {x : BitVec w} :
26282642
↓reduceIte, toNat_udiv]
26292643
rw [Nat.div_self (by omega), Nat.mod_eq_of_lt (by omega)]
26302644

2645+
theorem msb_udiv (x y : BitVec w) :
2646+
(x / y).msb = (x.msb && y == 1#w) := by
2647+
cases msb_x : x.msb
2648+
· suffices x.toNat / y.toNat < 2 ^ (w - 1) by simpa [msb_eq_decide]
2649+
calc
2650+
x.toNat / y.toNat ≤ x.toNat := by apply Nat.div_le_self
2651+
_ < 2 ^ (w - 1) := by simpa [msb_eq_decide] using msb_x
2652+
. rcases w with _|w
2653+
· contradiction
2654+
· have : (y == 1#_) = decide (y.toNat = 1) := by
2655+
simp [(· == ·), toNat_eq]
2656+
simp only [this, Bool.true_and]
2657+
match hy : y.toNat with
2658+
| 0 =>
2659+
obtain rfl : y = 0#_ := eq_of_toNat_eq hy
2660+
simp
2661+
| 1 =>
2662+
obtain rfl : y = 1#_ := eq_of_toNat_eq (by simp [hy])
2663+
simpa using msb_x
2664+
| y + 2 =>
2665+
suffices x.toNat / (y + 2) < 2 ^ w by
2666+
simp_all [msb_eq_decide, hy]
2667+
calc
2668+
x.toNat / (y + 2)
2669+
≤ x.toNat / 2 := by apply Nat.div_add_le_right (by omega)
2670+
_ < 2 ^ w := by omega
2671+
2672+
theorem msb_udiv_eq_false_of {x : BitVec w} (h : x.msb = false) (y : BitVec w) :
2673+
(x / y).msb = false := by
2674+
simp [msb_udiv, h]
2675+
2676+
/--
2677+
If `x` is nonnegative (i.e., does not have its msb set),
2678+
then `x / y` is nonnegative, thus `toInt` and `toNat` coincide.
2679+
-/
2680+
theorem toInt_udiv_of_msb {x : BitVec w} (h : x.msb = false) (y : BitVec w) :
2681+
(x / y).toInt = x.toNat / y.toNat := by
2682+
simp [toInt_eq_msb_cond, msb_udiv_eq_false_of h]
2683+
26312684
/-! ### umod -/
26322685

26332686
theorem umod_def {x y : BitVec n} :

src/Init/Data/Nat/Div/Lemmas.lean

+8
Original file line numberDiff line numberDiff line change
@@ -49,4 +49,12 @@ theorem lt_div_mul_self (h : 0 < k) (w : k ≤ x) : x - k < x / k * k := by
4949
have : x % k < k := mod_lt x h
5050
omega
5151

52+
theorem div_le_div_left (hcb : c ≤ b) (hc : 0 < c) : a / b ≤ a / c :=
53+
(Nat.le_div_iff_mul_le hc).2 <|
54+
Nat.le_trans (Nat.mul_le_mul_left _ hcb) (Nat.div_mul_le_self a b)
55+
56+
theorem div_add_le_right {z : Nat} (h : 0 < z) (x y : Nat) :
57+
x / (y + z) ≤ x / z :=
58+
div_le_div_left (Nat.le_add_left z y) h
59+
5260
end Nat

0 commit comments

Comments
 (0)