@@ -9,6 +9,7 @@ import Init.Data.Bool
9
9
import Init.Data.BitVec.Basic
10
10
import Init.Data.Fin.Lemmas
11
11
import Init.Data.Nat.Lemmas
12
+ import Init.Data.Nat.Div.Lemmas
12
13
import Init.Data.Nat.Mod
13
14
import Init.Data.Int.Bitwise.Lemmas
14
15
import Init.Data.Int.Pow
@@ -442,6 +443,10 @@ theorem toInt_eq_toNat_cond (x : BitVec n) :
442
443
(x.toNat : Int) - (2 ^n : Nat) :=
443
444
rfl
444
445
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
+
445
450
theorem msb_eq_false_iff_two_mul_lt {x : BitVec w} : x.msb = false ↔ 2 * x.toNat < 2 ^w := by
446
451
cases w <;> simp [Nat.pow_succ, Nat.mul_comm _ 2 , msb_eq_decide, toNat_of_zero_length]
447
452
@@ -454,6 +459,9 @@ theorem toInt_eq_msb_cond (x : BitVec w) :
454
459
simp only [BitVec.toInt, ← msb_eq_false_iff_two_mul_lt]
455
460
cases x.msb <;> rfl
456
461
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]
457
465
458
466
theorem toInt_eq_toNat_bmod (x : BitVec n) : x.toInt = Int.bmod x.toNat (2 ^n) := by
459
467
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 =
2300
2308
@[simp, bv_toNat] theorem toNat_neg (x : BitVec n) : (- x).toNat = (2 ^n - x.toNat) % 2 ^n := by
2301
2309
simp [Neg.neg, BitVec.neg]
2302
2310
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
+
2303
2317
theorem toInt_neg {x : BitVec w} :
2304
2318
(-x).toInt = (-x.toInt).bmod (2 ^ w) := by
2305
2319
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) :
2586
2600
rw [← udiv_eq]
2587
2601
simp [udiv, bv_toNat, h, Nat.mod_eq_of_lt]
2588
2602
2603
+ @[simp]
2604
+ theorem toFin_udiv {x y : BitVec n} : (x / y).toFin = x.toFin / y.toFin := by
2605
+ rfl
2606
+
2589
2607
@[simp, bv_toNat]
2590
2608
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
2596
2610
2597
2611
@[simp]
2598
2612
theorem zero_udiv {x : BitVec w} : (0 #w) / x = 0 #w := by
@@ -2628,6 +2642,45 @@ theorem udiv_self {x : BitVec w} :
2628
2642
↓reduceIte, toNat_udiv]
2629
2643
rw [Nat.div_self (by omega), Nat.mod_eq_of_lt (by omega)]
2630
2644
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
+
2631
2684
/-! ### umod -/
2632
2685
2633
2686
theorem umod_def {x y : BitVec n} :
0 commit comments