Skip to content

Commit 1fb75b6

Browse files
feat: add BitVec.(shiftLeft_add_distrib, shiftLeft_ushiftRight) (#5478)
Moved some Nat theorems from Mathlib --------- Co-authored-by: Tobias Grosser <github@grosser.es> Co-authored-by: Tobias Grosser <tobias@grosser.es>
1 parent 26f508d commit 1fb75b6

File tree

3 files changed

+29
-0
lines changed

3 files changed

+29
-0
lines changed

src/Init/Data/BitVec/Lemmas.lean

Lines changed: 24 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1512,6 +1512,21 @@ theorem shiftRight_add {w : Nat} (x : BitVec w) (n m : Nat) :
15121512
ext i
15131513
simp [Nat.add_assoc n m i]
15141514

1515+
theorem shiftLeft_ushiftRight {x : BitVec w} {n : Nat}:
1516+
x >>> n <<< n = x &&& BitVec.allOnes w <<< n := by
1517+
induction n generalizing x
1518+
case zero =>
1519+
ext; simp
1520+
case succ n ih =>
1521+
rw [BitVec.shiftLeft_add, Nat.add_comm, BitVec.shiftRight_add, ih,
1522+
Nat.add_comm, BitVec.shiftLeft_add, BitVec.shiftLeft_and_distrib]
1523+
ext i
1524+
by_cases hw : w = 0
1525+
· simp [hw]
1526+
· by_cases hi₂ : i.val = 0
1527+
· simp [hi₂]
1528+
· simp [Nat.lt_one_iff, hi₂, show 1 + (i.val - 1) = i by omega]
1529+
15151530
@[deprecated shiftRight_add (since := "2024-06-02")]
15161531
theorem shiftRight_shiftRight {w : Nat} (x : BitVec w) (n m : Nat) :
15171532
(x >>> n) >>> m = x >>> (n + m) := by
@@ -1737,6 +1752,15 @@ theorem ofInt_add {n} (x y : Int) : BitVec.ofInt n (x + y) =
17371752
apply eq_of_toInt_eq
17381753
simp
17391754

1755+
@[simp]
1756+
theorem shiftLeft_add_distrib {x y : BitVec w} {n : Nat} :
1757+
(x + y) <<< n = x <<< n + y <<< n := by
1758+
induction n
1759+
case zero =>
1760+
simp
1761+
case succ n ih =>
1762+
simp [ih, toNat_eq, Nat.shiftLeft_eq, ← Nat.add_mul]
1763+
17401764
/-! ### sub/neg -/
17411765

17421766
theorem sub_def {n} (x y : BitVec n) : x - y = .ofNat n ((2^n - y.toNat) + x.toNat) := by rfl

src/Init/Data/Nat/Basic.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -634,6 +634,8 @@ theorem lt_succ_of_lt (h : a < b) : a < succ b := le_succ_of_le h
634634

635635
theorem lt_add_one_of_lt (h : a < b) : a < b + 1 := le_succ_of_le h
636636

637+
theorem lt_one_iff : n < 1 ↔ n = 0 := Nat.lt_succ_iff.trans <| by rw [le_zero_eq]
638+
637639
theorem succ_pred_eq_of_ne_zero : ∀ {n}, n ≠ 0 → succ (pred n) = n
638640
| _+1, _ => rfl
639641

src/Init/Data/Nat/Lemmas.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -605,6 +605,9 @@ theorem add_mod (a b n : Nat) : (a + b) % n = ((a % n) + (b % n)) % n := by
605605
| zero => simp_all
606606
| succ k => omega
607607

608+
@[simp] theorem mod_mul_mod {a b c : Nat} : (a % c * b) % c = a * b % c := by
609+
rw [mul_mod, mod_mod, ← mul_mod]
610+
608611
/-! ### pow -/
609612

610613
theorem pow_succ' {m n : Nat} : m ^ n.succ = m * m ^ n := by

0 commit comments

Comments
 (0)