Skip to content

Commit cdab0d7

Browse files
tydeuluisacicolini
authored andcommitted
feat: Nat.shiftRight_bitwise_distrib (leanprover#6334)
This PR adds `Nat` theorems for distributing `>>>` over bitwise operations, paralleling those of `BitVec`. This enables closing goals like the following using `simp`: ```lean example (n : Nat) : (n <<< 2 ||| 3) >>> 2 = n := by simp [Nat.shiftRight_or_distrib] ``` It might be nice for these theorems to be `simp` lemmas, but they are not currently in order to be consistent with the existing `BitVec` and `div_two` theorems.
1 parent 3f79af2 commit cdab0d7

File tree

1 file changed

+48
-12
lines changed

1 file changed

+48
-12
lines changed

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

+48-12
Original file line numberDiff line numberDiff line change
@@ -106,9 +106,21 @@ theorem testBit_add_one (x i : Nat) : testBit x (i + 1) = testBit (x/2) i := by
106106
unfold testBit
107107
simp [shiftRight_succ_inside]
108108

109+
theorem testBit_add (x i n : Nat) : testBit x (i + n) = testBit (x / 2 ^ n) i := by
110+
revert x
111+
induction n with
112+
| zero => simp
113+
| succ n ih =>
114+
intro x
115+
rw [← Nat.add_assoc, testBit_add_one, ih (x / 2),
116+
Nat.pow_succ, Nat.div_div_eq_div_mul, Nat.mul_comm]
117+
109118
theorem testBit_div_two (x i : Nat) : testBit (x / 2) i = testBit x (i + 1) := by
110119
simp
111120

121+
theorem testBit_div_two_pow (x i : Nat) : testBit (x / 2 ^ n) i = testBit x (i + n) :=
122+
testBit_add .. |>.symm
123+
112124
theorem testBit_to_div_mod {x : Nat} : testBit x i = decide (x / 2^i % 2 = 1) := by
113125
induction i generalizing x with
114126
| zero =>
@@ -365,20 +377,20 @@ theorem testBit_two_pow_of_ne {n m : Nat} (hm : n ≠ m) : testBit (2 ^ n) m = f
365377

366378
/-! ### bitwise -/
367379

368-
theorem testBit_bitwise (false_false_axiom : f false false = false) (x y i : Nat) :
380+
theorem testBit_bitwise (of_false_false : f false false = false) (x y i : Nat) :
369381
(bitwise f x y).testBit i = f (x.testBit i) (y.testBit i) := by
370382
induction i using Nat.strongRecOn generalizing x y with
371383
| ind i hyp =>
372384
unfold bitwise
373385
if x_zero : x = 0 then
374386
cases p : f false true <;>
375387
cases yi : testBit y i <;>
376-
simp [x_zero, p, yi, false_false_axiom]
388+
simp [x_zero, p, yi, of_false_false]
377389
else if y_zero : y = 0 then
378390
simp [x_zero, y_zero]
379391
cases p : f true false <;>
380392
cases xi : testBit x i <;>
381-
simp [p, xi, false_false_axiom]
393+
simp [p, xi, of_false_false]
382394
else
383395
simp only [x_zero, y_zero, ←Nat.two_mul]
384396
cases i with
@@ -440,6 +452,11 @@ theorem bitwise_lt_two_pow (left : x < 2^n) (right : y < 2^n) : (Nat.bitwise f x
440452
case neg =>
441453
apply Nat.add_lt_add <;> exact hyp1
442454

455+
theorem bitwise_div_two_pow (of_false_false : f false false = false := by rfl) :
456+
(bitwise f x y) / 2 ^ n = bitwise f (x / 2 ^ n) (y / 2 ^ n) := by
457+
apply Nat.eq_of_testBit_eq
458+
simp [testBit_bitwise of_false_false, testBit_div_two_pow]
459+
443460
/-! ### and -/
444461

445462
@[simp] theorem testBit_and (x y i : Nat) : (x &&& y).testBit i = (x.testBit i && y.testBit i) := by
@@ -495,9 +512,11 @@ theorem and_pow_two_sub_one_of_lt_two_pow {x : Nat} (lt : x < 2^n) : x &&& 2^n -
495512
rw [testBit_and]
496513
simp
497514

498-
theorem and_div_two : (a &&& b) / 2 = a / 2 &&& b / 2 := by
499-
apply Nat.eq_of_testBit_eq
500-
simp [testBit_and, ← testBit_add_one]
515+
theorem and_div_two_pow : (a &&& b) / 2 ^ n = a / 2 ^ n &&& b / 2 ^ n :=
516+
bitwise_div_two_pow
517+
518+
theorem and_div_two : (a &&& b) / 2 = a / 2 &&& b / 2 :=
519+
and_div_two_pow (n := 1)
501520

502521
/-! ### lor -/
503522

@@ -563,9 +582,11 @@ theorem or_lt_two_pow {x y n : Nat} (left : x < 2^n) (right : y < 2^n) : x ||| y
563582
rw [testBit_or]
564583
simp
565584

566-
theorem or_div_two : (a ||| b) / 2 = a / 2 ||| b / 2 := by
567-
apply Nat.eq_of_testBit_eq
568-
simp [testBit_or, ← testBit_add_one]
585+
theorem or_div_two_pow : (a ||| b) / 2 ^ n = a / 2 ^ n ||| b / 2 ^ n :=
586+
bitwise_div_two_pow
587+
588+
theorem or_div_two : (a ||| b) / 2 = a / 2 ||| b / 2 :=
589+
or_div_two_pow (n := 1)
569590

570591
/-! ### xor -/
571592

@@ -619,9 +640,11 @@ theorem and_xor_distrib_left {a b c : Nat} : a &&& (b ^^^ c) = (a &&& b) ^^^ (a
619640
rw [testBit_xor]
620641
simp
621642

622-
theorem xor_div_two : (a ^^^ b) / 2 = a / 2 ^^^ b / 2 := by
623-
apply Nat.eq_of_testBit_eq
624-
simp [testBit_xor, ← testBit_add_one]
643+
theorem xor_div_two_pow : (a ^^^ b) / 2 ^ n = a / 2 ^ n ^^^ b / 2 ^ n :=
644+
bitwise_div_two_pow
645+
646+
theorem xor_div_two : (a ^^^ b) / 2 = a / 2 ^^^ b / 2 :=
647+
xor_div_two_pow (n := 1)
625648

626649
/-! ### Arithmetic -/
627650

@@ -693,6 +716,19 @@ theorem mul_add_lt_is_or {b : Nat} (b_lt : b < 2^i) (a : Nat) : 2^i * a + b = 2^
693716
simp only [testBit, one_and_eq_mod_two, mod_two_bne_zero]
694717
exact (Bool.beq_eq_decide_eq _ _).symm
695718

719+
theorem shiftRight_bitwise_distrib {a b : Nat} (of_false_false : f false false = false := by rfl) :
720+
(bitwise f a b) >>> i = bitwise f (a >>> i) (b >>> i) := by
721+
simp [shiftRight_eq_div_pow, bitwise_div_two_pow of_false_false]
722+
723+
theorem shiftRight_and_distrib {a b : Nat} : (a &&& b) >>> i = a >>> i &&& b >>> i :=
724+
shiftRight_bitwise_distrib
725+
726+
theorem shiftRight_or_distrib {a b : Nat} : (a ||| b) >>> i = a >>> i ||| b >>> i :=
727+
shiftRight_bitwise_distrib
728+
729+
theorem shiftRight_xor_distrib {a b : Nat} : (a ^^^ b) >>> i = a >>> i ^^^ b >>> i :=
730+
shiftRight_bitwise_distrib
731+
696732
/-! ### le -/
697733

698734
theorem le_of_testBit {n m : Nat} (h : ∀ i, n.testBit i = true → m.testBit i = true) : n ≤ m := by

0 commit comments

Comments
 (0)