Skip to content

Commit 89686db

Browse files
mhk119eric-wieser
andcommitted
feat: delete Nat.shiftr and Nat.shiftl (#6356)
These already exists upstream (with minorly different but equal definitions) as `Nat.shiftRight` and `Nat.shiftLeft`. Co-authored-by: mhk119 <58151072+mhk119@users.noreply.github.com> Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
1 parent 41aa37b commit 89686db

File tree

8 files changed

+129
-140
lines changed

8 files changed

+129
-140
lines changed

Mathlib/Data/FP/Basic.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -16,8 +16,8 @@ import Mathlib.Data.Rat.Floor
1616

1717
@[nolint docBlame]
1818
def Int.shift2 (a b : ℕ) : ℤ → ℕ × ℕ
19-
| Int.ofNat e => (a.shiftl e, b)
20-
| Int.negSucc e => (a, b.shiftl e.succ)
19+
| Int.ofNat e => (a <<< e, b)
20+
| Int.negSucc e => (a, b <<< e.succ)
2121
#align int.shift2 Int.shift2
2222

2323
namespace FP
@@ -138,8 +138,8 @@ protected def Float.neg : Float → Float
138138

139139
@[nolint docBlame]
140140
def divNatLtTwoPow (n d : ℕ) : ℤ → Bool
141-
| Int.ofNat e => n < d.shiftl e
142-
| Int.negSucc e => n.shiftl e.succ < d
141+
| Int.ofNat e => n < d <<< e
142+
| Int.negSucc e => n <<< e.succ < d
143143
#align fp.div_nat_lt_two_pow FP.divNatLtTwoPowₓ -- Porting note: TC argument `[C : FP.FloatCfg]` no longer present
144144

145145

Mathlib/Data/Int/Bitwise.lean

Lines changed: 21 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -375,13 +375,13 @@ theorem shiftr_neg (m n : ℤ) : shiftr m (-n) = shiftl m n := by rw [← shiftl
375375

376376
-- Porting note: what's the correct new name?
377377
@[simp]
378-
theorem shiftl_coe_nat (m n : ℕ) : shiftl m n = Nat.shiftl m n :=
379-
rfl
378+
theorem shiftl_coe_nat (m n : ℕ) : shiftl m n = m <<< n :=
379+
by simp [shiftl]
380380
#align int.shiftl_coe_nat Int.shiftl_coe_nat
381381

382382
-- Porting note: what's the correct new name?
383383
@[simp]
384-
theorem shiftr_coe_nat (m n : ℕ) : shiftr m n = Nat.shiftr m n := by cases n <;> rfl
384+
theorem shiftr_coe_nat (m n : ℕ) : shiftr m n = m >>> n := by cases n <;> rfl
385385
#align int.shiftr_coe_nat Int.shiftr_coe_nat
386386

387387
@[simp]
@@ -390,63 +390,66 @@ theorem shiftl_negSucc (m n : ℕ) : shiftl -[m+1] n = -[Nat.shiftl' true m n+1]
390390
#align int.shiftl_neg_succ Int.shiftl_negSucc
391391

392392
@[simp]
393-
theorem shiftr_negSucc (m n : ℕ) : shiftr -[m+1] n = -[Nat.shiftr m n+1] := by cases n <;> rfl
393+
theorem shiftr_negSucc (m n : ℕ) : shiftr -[m+1] n = -[m >>> n+1] := by cases n <;> rfl
394394
#align int.shiftr_neg_succ Int.shiftr_negSucc
395395

396396
theorem shiftr_add : ∀ (m : ℤ) (n k : ℕ), shiftr m (n + k) = shiftr (shiftr m n) k
397397
| (m : ℕ), n, k => by
398-
rw [shiftr_coe_nat, shiftr_coe_nat, ← Int.ofNat_add, shiftr_coe_nat, Nat.shiftr_add]
398+
rw [shiftr_coe_nat, shiftr_coe_nat, ← Int.ofNat_add, shiftr_coe_nat, Nat.shiftRight_add]
399399
| -[m+1], n, k => by
400-
rw [shiftr_negSucc, shiftr_negSucc, ← Int.ofNat_add, shiftr_negSucc, Nat.shiftr_add]
400+
rw [shiftr_negSucc, shiftr_negSucc, ← Int.ofNat_add, shiftr_negSucc, Nat.shiftRight_add]
401401
#align int.shiftr_add Int.shiftr_add
402402

403403
/-! ### bitwise ops -/
404404

405405
attribute [local simp] Int.zero_div
406406

407407
theorem shiftl_add : ∀ (m : ℤ) (n : ℕ) (k : ℤ), shiftl m (n + k) = shiftl (shiftl m n) k
408-
| (m : ℕ), n, (k : ℕ) => congr_arg ofNat (Nat.shiftl_add _ _ _)
408+
| (m : ℕ), n, (k : ℕ) =>
409+
congr_arg ofNat (by simp [Nat.pow_add, mul_assoc])
409410
| -[m+1], n, (k : ℕ) => congr_arg negSucc (Nat.shiftl'_add _ _ _ _)
410411
| (m : ℕ), n, -[k+1] =>
411-
subNatNat_elim n k.succ (fun n k i => shiftl (↑m) i = Nat.shiftr (Nat.shiftl m n) k)
412+
subNatNat_elim n k.succ (fun n k i => shiftl (↑m) i = (Nat.shiftl' false m n) >>> k)
412413
(fun (i n : ℕ) =>
413-
by dsimp; rw [← Nat.shiftl_sub, add_tsub_cancel_left]; apply Nat.le_add_right)
414+
by dsimp; simp [- Nat.shiftLeft_eq, ← Nat.shiftLeft_sub _ , add_tsub_cancel_left])
414415
fun i n =>
415-
by dsimp; rw [add_assoc, Nat.shiftr_add, ← Nat.shiftl_sub, tsub_self] <;> rfl
416+
by dsimp; simp [- Nat.shiftLeft_eq, Nat.shiftLeft_zero, Nat.shiftRight_add,
417+
← Nat.shiftLeft_sub, shiftl]
416418
| -[m+1], n, -[k+1] =>
417419
subNatNat_elim n k.succ
418-
(fun n k i => shiftl -[m+1] i = -[Nat.shiftr (Nat.shiftl' true m n) k+1])
420+
(fun n k i => shiftl -[m+1] i = -[(Nat.shiftl' true m n) >>> k+1])
419421
(fun i n =>
420422
congr_arg negSucc <| by
421423
rw [← Nat.shiftl'_sub, add_tsub_cancel_left]; apply Nat.le_add_right)
422424
fun i n =>
423-
congr_arg negSucc <| by rw [add_assoc, Nat.shiftr_add, ← Nat.shiftl'_sub, tsub_self] <;> rfl
425+
congr_arg negSucc <| by rw [add_assoc, Nat.shiftRight_add, ← Nat.shiftl'_sub, tsub_self]
426+
<;> rfl
424427
#align int.shiftl_add Int.shiftl_add
425428

426429
theorem shiftl_sub (m : ℤ) (n : ℕ) (k : ℤ) : shiftl m (n - k) = shiftr (shiftl m n) k :=
427430
shiftl_add _ _ _
428431
#align int.shiftl_sub Int.shiftl_sub
429432

430433
theorem shiftl_eq_mul_pow : ∀ (m : ℤ) (n : ℕ), shiftl m n = m * ↑(2 ^ n)
431-
| (m : ℕ), _ => congr_arg ((↑) : ℕ → ℤ) (Nat.shiftl_eq_mul_pow _ _)
434+
| (m : ℕ), _ => congr_arg ((↑) : ℕ → ℤ) (by simp)
432435
| -[_+1], _ => @congr_arg ℕ ℤ _ _ (fun i => -i) (Nat.shiftl'_tt_eq_mul_pow _ _)
433436
#align int.shiftl_eq_mul_pow Int.shiftl_eq_mul_pow
434437

435438
theorem shiftr_eq_div_pow : ∀ (m : ℤ) (n : ℕ), shiftr m n = m / ↑(2 ^ n)
436-
| (m : ℕ), n => by rw [shiftr_coe_nat, Nat.shiftr_eq_div_pow _ _]; simp
439+
| (m : ℕ), n => by rw [shiftr_coe_nat, Nat.shiftRight_eq_div_pow _ _]; simp
437440
| -[m+1], n => by
438-
rw [shiftr_negSucc, negSucc_ediv, Nat.shiftr_eq_div_pow]; rfl
441+
rw [shiftr_negSucc, negSucc_ediv, Nat.shiftRight_eq_div_pow]; rfl
439442
exact ofNat_lt_ofNat_of_lt (pow_pos (by decide) _)
440443
#align int.shiftr_eq_div_pow Int.shiftr_eq_div_pow
441444

442445
theorem one_shiftl (n : ℕ) : shiftl 1 n = (2 ^ n : ℕ) :=
443-
congr_arg ((↑) : ℕ → ℤ) (Nat.one_shiftl _)
446+
congr_arg ((↑) : ℕ → ℤ) (by simp)
444447
#align int.one_shiftl Int.one_shiftl
445448

446449
@[simp]
447450
theorem zero_shiftl : ∀ n : ℤ, shiftl 0 n = 0
448-
| (n : ℕ) => congr_arg ((↑) : ℕ → ℤ) (Nat.zero_shiftl _)
449-
| -[_+1] => congr_arg ((↑) : ℕ → ℤ) (Nat.zero_shiftr _)
451+
| (n : ℕ) => congr_arg ((↑) : ℕ → ℤ) (by simp)
452+
| -[_+1] => congr_arg ((↑) : ℕ → ℤ) (by simp)
450453
#align int.zero_shiftl Int.zero_shiftl
451454

452455
@[simp]

Mathlib/Data/Nat/Bitwise.lean

Lines changed: 5 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -67,13 +67,14 @@ theorem zero_of_testBit_eq_false {n : ℕ} (h : ∀ i, testBit n i = false) : n
6767
#align nat.zero_of_test_bit_eq_ff Nat.zero_of_testBit_eq_false
6868

6969
@[simp]
70-
theorem zero_testBit (i : ℕ) : testBit 0 i = false := by simp only [testBit, shiftr_zero, bodd_zero]
70+
theorem zero_testBit (i : ℕ) : testBit 0 i = false := by
71+
simp only [testBit, zero_shiftRight, bodd_zero]
7172
#align nat.zero_test_bit Nat.zero_testBit
7273

7374
/-- The ith bit is the ith element of `n.bits`. -/
7475
theorem testBit_eq_inth (n i : ℕ) : n.testBit i = n.bits.getI i := by
7576
induction' i with i ih generalizing n
76-
· simp [testBit, shiftr, bodd_eq_bits_head, List.getI_zero_eq_headI]
77+
· simp [testBit, bodd_eq_bits_head, List.getI_zero_eq_headI]
7778
conv_lhs => rw [← bit_decomp n]
7879
rw [testBit_succ, ih n.div2, div2_bits_eq_tail]
7980
cases n.bits <;> simp
@@ -137,11 +138,11 @@ theorem lt_of_testBit {n m : ℕ} (i : ℕ) (hn : testBit n i = false) (hm : tes
137138

138139
@[simp]
139140
theorem testBit_two_pow_self (n : ℕ) : testBit (2 ^ n) n = true := by
140-
rw [testBit, shiftr_eq_div_pow, Nat.div_self (pow_pos (α := ℕ) zero_lt_two n), bodd_one]
141+
rw [testBit, shiftRight_eq_div_pow, Nat.div_self (pow_pos (α := ℕ) zero_lt_two n), bodd_one]
141142
#align nat.test_bit_two_pow_self Nat.testBit_two_pow_self
142143

143144
theorem testBit_two_pow_of_ne {n m : ℕ} (hm : n ≠ m) : testBit (2 ^ n) m = false := by
144-
rw [testBit, shiftr_eq_div_pow]
145+
rw [testBit, shiftRight_eq_div_pow]
145146
cases' hm.lt_or_lt with hm hm
146147
· rw [Nat.div_eq_zero, bodd_zero]
147148
exact Nat.pow_lt_pow_of_lt_right one_lt_two hm

Mathlib/Data/Nat/Size.lean

Lines changed: 25 additions & 34 deletions
Original file line numberDiff line numberDiff line change
@@ -17,15 +17,11 @@ namespace Nat
1717
section
1818
set_option linter.deprecated false
1919

20-
theorem shiftl_eq_mul_pow (m) : ∀ n, shiftl m n = m * 2 ^ n
21-
| 0 => (Nat.mul_one _).symm
22-
| k + 1 => by
23-
show bit0 (shiftl m k) = m * (2 ^ k * 2)
24-
rw [bit0_val, shiftl_eq_mul_pow m k, mul_comm 2, mul_assoc]
25-
#align nat.shiftl_eq_mul_pow Nat.shiftl_eq_mul_pow
20+
theorem shiftLeft_eq_mul_pow (m) : ∀ n, m <<< n = m * 2 ^ n := shiftLeft_eq _
21+
#align nat.shiftl_eq_mul_pow Nat.shiftLeft_eq_mul_pow
2622

2723
theorem shiftl'_tt_eq_mul_pow (m) : ∀ n, shiftl' true m n + 1 = (m + 1) * 2 ^ n
28-
| 0 => by simp [shiftl, shiftl', pow_zero, Nat.one_mul]
24+
| 0 => by simp [shiftl', pow_zero, Nat.one_mul]
2925
| k + 1 => by
3026
change bit1 (shiftl' true m k) + 1 = (m + 1) * (2 ^ k * 2)
3127
rw [bit1_val]
@@ -35,26 +31,17 @@ theorem shiftl'_tt_eq_mul_pow (m) : ∀ n, shiftl' true m n + 1 = (m + 1) * 2 ^
3531

3632
end
3733

38-
theorem one_shiftl (n) : shiftl 1 n = 2 ^ n :=
39-
(shiftl_eq_mul_pow _ _).trans (Nat.one_mul _)
40-
#align nat.one_shiftl Nat.one_shiftl
34+
#align nat.one_shiftl Nat.one_shiftLeft
4135

42-
@[simp]
43-
theorem zero_shiftl (n) : shiftl 0 n = 0 :=
44-
(shiftl_eq_mul_pow _ _).trans (Nat.zero_mul _)
45-
#align nat.zero_shiftl Nat.zero_shiftl
36+
theorem zero_shiftLeft (n) : 0 <<< n = 0 := by simp
37+
#align nat.zero_shiftl Nat.zero_shiftLeft
4638

47-
theorem shiftr_eq_div_pow (m) : ∀ n, shiftr m n = m / 2 ^ n
39+
theorem shiftRight_eq_div_pow (m) : ∀ n, m >>> n = m / 2 ^ n
4840
| 0 => (Nat.div_one _).symm
49-
| k + 1 =>
50-
(congr_arg div2 (shiftr_eq_div_pow m k)).trans <| by
51-
rw [div2_val, Nat.div_div_eq_div_mul, Nat.pow_succ]
52-
#align nat.shiftr_eq_div_pow Nat.shiftr_eq_div_pow
53-
54-
@[simp]
55-
theorem zero_shiftr (n) : shiftr 0 n = 0 :=
56-
(shiftr_eq_div_pow _ _).trans (Nat.zero_div _)
57-
#align nat.zero_shiftr Nat.zero_shiftr
41+
| k + 1 => by
42+
rw [shiftRight_add, shiftRight_eq_div_pow m k]
43+
simp [Nat.div_div_eq_div_mul, ← Nat.pow_succ]
44+
#align nat.shiftr_eq_div_pow Nat.shiftRight_eq_div_pow
5845

5946
theorem shiftl'_ne_zero_left (b) {m} (h : m ≠ 0) (n) : shiftl' b m n ≠ 0 := by
6047
induction n <;> simp [bit_ne_zero, shiftl', *]
@@ -119,26 +106,28 @@ theorem size_shiftl' {b m n} (h : shiftl' b m n ≠ 0) : size (shiftl' b m n) =
119106
rfl
120107
#align nat.size_shiftl' Nat.size_shiftl'
121108

122-
@[simp]
123-
theorem size_shiftl {m} (h : m ≠ 0) (n) : size (shiftl m n) = size m + n :=
124-
size_shiftl' (shiftl'_ne_zero_left _ h _)
125-
#align nat.size_shiftl Nat.size_shiftl
109+
-- TODO: decide whether `Nat.shiftLeft_eq` (which rewrites the LHS into a power) should be a simp
110+
-- lemma; it was not in mathlib3. Until then, tell the simpNF linter to ignore the issue.
111+
@[simp, nolint simpNF]
112+
theorem size_shiftLeft {m} (h : m ≠ 0) (n) : size (m <<< n) = size m + n :=
113+
by simp only [size_shiftl' (shiftl'_ne_zero_left _ h _), ← shiftl'_false]
114+
#align nat.size_shiftl Nat.size_shiftLeft
126115

127116
theorem lt_size_self (n : ℕ) : n < 2 ^ size n := by
128-
rw [← one_shiftl]
129-
have : ∀ {n}, n = 0 → n < shiftl 1 (size n) := by simp
117+
rw [← one_shiftLeft]
118+
have : ∀ {n}, n = 0 → n < 1 <<< (size n) := by simp
130119
apply binaryRec _ _ n
131120
· apply this rfl
132121
intro b n IH
133122
by_cases h : bit b n = 0
134123
· apply this h
135-
rw [size_bit h, shiftl_succ]
136-
exact bit_lt_bit0 _ IH
124+
rw [size_bit h, shiftLeft_succ, shiftLeft_eq, one_mul, ← bit0_val]
125+
exact bit_lt_bit0 _ (by simpa [shiftRight_eq_div_pow] using IH)
137126
#align nat.lt_size_self Nat.lt_size_self
138127

139128
theorem size_le {m n : ℕ} : size m ≤ n ↔ m < 2 ^ n :=
140129
fun h => lt_of_lt_of_le (lt_size_self _) (pow_le_pow_of_le_right (by decide) h), by
141-
rw [← one_shiftl]; revert n
130+
rw [← one_shiftLeft]; revert n
142131
apply binaryRec _ _ m
143132
· intro n
144133
simp
@@ -149,7 +138,9 @@ theorem size_le {m n : ℕ} : size m ≤ n ↔ m < 2 ^ n :=
149138
cases' n with n
150139
· exact e.elim (Nat.eq_zero_of_le_zero (le_of_lt_succ h))
151140
· apply succ_le_succ (IH _)
152-
apply lt_imp_lt_of_le_imp_le (fun h' => bit0_le_bit _ h') h⟩
141+
apply lt_of_mul_lt_mul_left _ (zero_le 2)
142+
simp only [← bit0_val, shiftLeft_succ] at *
143+
exact lt_of_le_of_lt (bit0_le_bit b rfl.le) h⟩
153144
#align nat.size_le Nat.size_le
154145

155146
theorem lt_size {m n : ℕ} : m < size n ↔ 2 ^ m ≤ n := by

Mathlib/Data/Num/Lemmas.lean

Lines changed: 30 additions & 35 deletions
Original file line numberDiff line numberDiff line change
@@ -955,73 +955,68 @@ theorem lxor'_to_nat : ∀ m n, (lxor m n : ℕ) = Nat.lxor' m n := by
955955
#align num.lxor_to_nat Num.lxor'_to_nat
956956

957957
@[simp, norm_cast]
958-
theorem shiftl_to_nat (m n) : (shiftl m n : ℕ) = Nat.shiftl m n := by
958+
theorem shiftl_to_nat (m n) : (shiftl m n : ℕ) = (m : ℕ) <<< (n : ℕ) := by
959959
cases m <;> dsimp only [shiftl]
960960
· symm
961-
apply Nat.zero_shiftl
961+
apply Nat.zero_shiftLeft
962962
simp only [cast_pos]
963963
induction' n with n IH
964964
· rfl
965-
simp [PosNum.shiftl_succ_eq_bit0_shiftl, Nat.shiftl_succ, IH]
965+
simp [PosNum.shiftl_succ_eq_bit0_shiftl, Nat.shiftLeft_succ, IH,
966+
Nat.bit0_val, pow_succ, ← mul_assoc, mul_comm]
966967
#align num.shiftl_to_nat Num.shiftl_to_nat
967968

968969
@[simp, norm_cast]
969-
theorem shiftr_to_nat (m n) : (shiftr m n : ℕ) = Nat.shiftr m n := by
970-
cases' m with m <;> dsimp only [shiftr]
970+
971+
theorem shiftr_to_nat (m n) : (shiftr m n : ℕ) = (m : ℕ) >>> (n : ℕ) := by
972+
cases' m with m <;> dsimp only [shiftr];
971973
· symm
972-
apply Nat.zero_shiftr
974+
apply Nat.zero_shiftRight
973975
induction' n with n IH generalizing m
974976
· cases m <;> rfl
975977
cases' m with m m <;> dsimp only [PosNum.shiftr]
976-
· rw [Nat.shiftr_eq_div_pow]
978+
· rw [Nat.shiftRight_eq_div_pow]
977979
symm
978980
apply Nat.div_eq_of_lt
979-
exact @Nat.pow_lt_pow_of_lt_right 2 (by decide) 0 (n + 1) (Nat.succ_pos _)
981+
simp [@Nat.pow_lt_pow_of_lt_right 2 (by decide) 0 (n + 1) (Nat.succ_pos _)]
980982
· trans
981983
apply IH
982-
change Nat.shiftr m n = Nat.shiftr (_root_.bit1 m) (n + 1)
983-
rw [add_comm n 1, Nat.shiftr_add]
984-
apply congr_arg fun x => Nat.shiftr x n
985-
-- Porting note: `unfold` is not repeated in Lean4.
986-
repeat unfold Nat.shiftr
987-
change (m : ℕ) = Nat.div2 (Nat.bit true m)
988-
rw [Nat.div2_bit]
984+
change Nat.shiftRight m n = Nat.shiftRight (_root_.bit1 m) (n + 1)
985+
rw [add_comm n 1, @Nat.shiftRight_eq _ (1 + n), Nat.shiftRight_add]
986+
apply congr_arg fun x => Nat.shiftRight x n
987+
simp [Nat.shiftRight_succ, Nat.shiftRight_zero, ← Nat.div2_val]
989988
· trans
990989
apply IH
991-
change Nat.shiftr m n = Nat.shiftr (_root_.bit0 m) (n + 1)
992-
rw [add_comm n 1, Nat.shiftr_add]
993-
apply congr_arg fun x => Nat.shiftr x n
994-
repeat unfold Nat.shiftr
995-
change (m : ℕ) = Nat.div2 (Nat.bit false m)
996-
rw [Nat.div2_bit]
990+
change Nat.shiftRight m n = Nat.shiftRight (_root_.bit0 m) (n + 1)
991+
rw [add_comm n 1, @Nat.shiftRight_eq _ (1 + n), Nat.shiftRight_add]
992+
apply congr_arg fun x => Nat.shiftRight x n
993+
simp [Nat.shiftRight_succ, Nat.shiftRight_zero, ← Nat.div2_val]
997994
#align num.shiftr_to_nat Num.shiftr_to_nat
998995

999996
@[simp]
1000997
theorem testBit_to_nat (m n) : testBit m n = Nat.testBit m n := by
1001998
-- Porting note: `unfold` → `dsimp only`
1002999
cases m <;> dsimp only [testBit, Nat.testBit]
10031000
case zero =>
1004-
change false = Nat.bodd (Nat.shiftr 0 n)
1005-
rw [Nat.zero_shiftr]
1001+
change false = Nat.bodd (0 >>> n)
1002+
rw [Nat.zero_shiftRight]
10061003
rfl
10071004
case pos m =>
10081005
induction' n with n IH generalizing m <;> cases' m with m m <;> dsimp only [PosNum.testBit]
10091006
· rfl
10101007
· exact (Nat.bodd_bit _ _).symm
10111008
· exact (Nat.bodd_bit _ _).symm
1012-
· change false = Nat.bodd (Nat.shiftr 1 (n + 1))
1013-
rw [add_comm, Nat.shiftr_add]
1014-
change false = Nat.bodd (Nat.shiftr 0 n)
1015-
rw [Nat.zero_shiftr]; rfl
1016-
· change PosNum.testBit m n = Nat.bodd (Nat.shiftr (Nat.bit true m) (n + 1))
1017-
rw [add_comm, Nat.shiftr_add]
1018-
dsimp only [Nat.shiftr]
1019-
rw [Nat.div2_bit]
1009+
· change false = Nat.bodd (1 >>> (n + 1))
1010+
rw [add_comm, Nat.shiftRight_add]
1011+
change false = Nat.bodd (0 >>> n)
1012+
rw [Nat.zero_shiftRight]; rfl
1013+
· change PosNum.testBit m n = Nat.bodd ((Nat.bit true m) >>> (n + 1))
1014+
rw [add_comm, Nat.shiftRight_add]
1015+
simp only [Nat.shiftRight_succ, Nat.shiftRight_zero, ← Nat.div2_val, Nat.div2_bit]
10201016
apply IH
1021-
· change PosNum.testBit m n = Nat.bodd (Nat.shiftr (Nat.bit false m) (n + 1))
1022-
rw [add_comm, Nat.shiftr_add]
1023-
dsimp only [Nat.shiftr]
1024-
rw [Nat.div2_bit]
1017+
· change PosNum.testBit m n = Nat.bodd ((Nat.bit false m) >>> (n + 1))
1018+
rw [add_comm, Nat.shiftRight_add]
1019+
simp only [Nat.shiftRight_succ, Nat.shiftRight_zero, ← Nat.div2_val, Nat.div2_bit]
10251020
apply IH
10261021
#align num.test_bit_to_nat Num.testBit_to_nat
10271022

Mathlib/Data/Ordmap/Ordnode.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -876,11 +876,11 @@ def ofAscListAux₁ : ∀ l : List α, ℕ → Ordnode α × { l' : List α // l
876876
| x :: xs => fun s =>
877877
if s = 1 then (ι x, ⟨xs, Nat.le_succ _⟩)
878878
else
879-
match ofAscListAux₁ xs (s.shiftl 1) with
879+
match ofAscListAux₁ xs (s <<< 1) with
880880
| (t, ⟨[], _⟩) => (t, ⟨[], Nat.zero_le _⟩)
881881
| (l, ⟨y :: ys, h⟩) =>
882882
have := Nat.le_succ_of_le h
883-
let (r, ⟨zs, h'⟩) := ofAscListAux₁ ys (s.shiftl 1)
883+
let (r, ⟨zs, h'⟩) := ofAscListAux₁ ys (s <<< 1)
884884
(link l y r, ⟨zs, le_trans h' (le_of_lt this)⟩)
885885
termination_by ofAscListAux₁ l => l.length
886886
#align ordnode.of_asc_list_aux₁ Ordnode.ofAscListAux₁
@@ -892,7 +892,7 @@ def ofAscListAux₂ : List α → Ordnode α → ℕ → Ordnode α
892892
match ofAscListAux₁ xs s with
893893
| (r, ⟨ys, h⟩) =>
894894
have := Nat.lt_succ_of_le h
895-
ofAscListAux₂ ys (link l x r) (s.shiftl 1)
895+
ofAscListAux₂ ys (link l x r) (s <<< 1)
896896
termination_by ofAscListAux₂ l => l.length
897897
#align ordnode.of_asc_list_aux₂ Ordnode.ofAscListAux₂
898898

Mathlib/Init/Data/Int/Bitwise.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -105,9 +105,9 @@ def lxor' : ℤ → ℤ → ℤ
105105
is obtained by left-shifting the binary representation of `m` by `n` places -/
106106
def shiftl : ℤ → ℤ → ℤ
107107
| (m : ℕ), (n : ℕ) => Nat.shiftl' false m n
108-
| (m : ℕ), -[n +1] => Nat.shiftr m (Nat.succ n)
108+
| (m : ℕ), -[n +1] => m >>> (Nat.succ n)
109109
| -[m +1], (n : ℕ) => -[Nat.shiftl' true m n +1]
110-
| -[m +1], -[n +1] => -[Nat.shiftr m (Nat.succ n) +1]
110+
| -[m +1], -[n +1] => -[m >>> (Nat.succ n) +1]
111111
#align int.shiftl Int.shiftl
112112

113113
/-- `shiftr m n` produces an integer whose binary representation

0 commit comments

Comments
 (0)