Skip to content

Commit ffac974

Browse files
tydeukim-em
authored andcommitted
feat: more UInt bitwise theorems (#6188)
This PR completes the `toNat` theorems for the bitwise operations (`and`, `or`, `xor`, `shiftLeft`, `shiftRight`) of the UInt types and adds `toBitVec` theorems as well. It also renames `and_toNat` to `toNat_and` to fit with the current naming convention.
1 parent e8de522 commit ffac974

File tree

2 files changed

+27
-21
lines changed

2 files changed

+27
-21
lines changed

src/Init/Data/UInt/Bitwise.lean

Lines changed: 23 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -1,25 +1,39 @@
11
/-
22
Copyright (c) 2024 Lean FRO, LLC. All Rights Reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Markus Himmel
4+
Authors: Markus Himmel, Mac Malone
55
-/
66
prelude
7-
import Init.Data.UInt.Basic
7+
import Init.Data.UInt.Lemmas
88
import Init.Data.Fin.Bitwise
99
import Init.Data.BitVec.Lemmas
1010

1111
set_option hygiene false in
12-
macro "declare_bitwise_uint_theorems" typeName:ident : command =>
12+
macro "declare_bitwise_uint_theorems" typeName:ident bits:term:arg : command =>
1313
`(
1414
namespace $typeName
1515

16-
@[simp] protected theorem and_toNat (a b : $typeName) : (a &&& b).toNat = a.toNat &&& b.toNat := BitVec.toNat_and ..
16+
@[simp] protected theorem toBitVec_and (a b : $typeName) : (a &&& b).toBitVec = a.toBitVec &&& b.toBitVec := rfl
17+
@[simp] protected theorem toBitVec_or (a b : $typeName) : (a ||| b).toBitVec = a.toBitVec ||| b.toBitVec := rfl
18+
@[simp] protected theorem toBitVec_xor (a b : $typeName) : (a ^^^ b).toBitVec = a.toBitVec ^^^ b.toBitVec := rfl
19+
@[simp] protected theorem toBitVec_shiftLeft (a b : $typeName) : (a <<< b).toBitVec = a.toBitVec <<< (b.toBitVec % $bits) := rfl
20+
@[simp] protected theorem toBitVec_shiftRight (a b : $typeName) : (a >>> b).toBitVec = a.toBitVec >>> (b.toBitVec % $bits) := rfl
21+
22+
@[simp] protected theorem toNat_and (a b : $typeName) : (a &&& b).toNat = a.toNat &&& b.toNat := by simp [toNat]
23+
@[simp] protected theorem toNat_or (a b : $typeName) : (a ||| b).toNat = a.toNat ||| b.toNat := by simp [toNat]
24+
@[simp] protected theorem toNat_xor (a b : $typeName) : (a ^^^ b).toNat = a.toNat ^^^ b.toNat := by simp [toNat]
25+
@[simp] protected theorem toNat_shiftLeft (a b : $typeName) : (a <<< b).toNat = a.toNat <<< (b.toNat % $bits) % 2 ^ $bits := by simp [toNat]
26+
@[simp] protected theorem toNat_shiftRight (a b : $typeName) : (a >>> b).toNat = a.toNat >>> (b.toNat % $bits) := by simp [toNat]
27+
28+
open $typeName (toNat_and) in
29+
@[deprecated toNat_and (since := "2024-11-28")]
30+
protected theorem and_toNat (a b : $typeName) : (a &&& b).toNat = a.toNat &&& b.toNat := BitVec.toNat_and ..
1731

1832
end $typeName
1933
)
2034

21-
declare_bitwise_uint_theorems UInt8
22-
declare_bitwise_uint_theorems UInt16
23-
declare_bitwise_uint_theorems UInt32
24-
declare_bitwise_uint_theorems UInt64
25-
declare_bitwise_uint_theorems USize
35+
declare_bitwise_uint_theorems UInt8 8
36+
declare_bitwise_uint_theorems UInt16 16
37+
declare_bitwise_uint_theorems UInt32 32
38+
declare_bitwise_uint_theorems UInt64 64
39+
declare_bitwise_uint_theorems USize System.Platform.numBits

src/Std/Data/DHashMap/Internal/Index.lean

Lines changed: 4 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -46,19 +46,11 @@ cf. https://github.com/leanprover/lean4/issues/4157
4646
@[irreducible, inline] def mkIdx (sz : Nat) (h : 0 < sz) (hash : UInt64) :
4747
{ u : USize // u.toNat < sz } :=
4848
⟨(scrambleHash hash).toUSize &&& (sz.toUSize - 1), by
49-
-- Making this proof significantly less painful will be a good test for our USize API
49+
-- This proof is a good test for our USize API
5050
by_cases h' : sz < USize.size
51-
· rw [USize.and_toNat, ← USize.ofNat_one, USize.toNat_sub_of_le, USize.toNat_ofNat_of_lt]
52-
· refine Nat.lt_of_le_of_lt Nat.and_le_right (Nat.sub_lt h ?_)
53-
rw [USize.toNat_ofNat_of_lt]
54-
· exact Nat.one_pos
55-
· exact Nat.lt_of_le_of_lt h h'
56-
· exact h'
57-
· rw [USize.le_def, BitVec.le_def]
58-
change _ ≤ (_ % _)
59-
rw [Nat.mod_eq_of_lt h', USize.ofNat, BitVec.toNat_ofNat, Nat.mod_eq_of_lt]
60-
· exact h
61-
· exact Nat.lt_of_le_of_lt h h'
51+
· rw [USize.toNat_and, USize.toNat_sub_of_le, USize.toNat_ofNat_of_lt h']
52+
· exact Nat.lt_of_le_of_lt Nat.and_le_right (Nat.sub_lt h (by simp))
53+
· simp [USize.le_iff_toNat_le, Nat.mod_eq_of_lt h', Nat.succ_le_of_lt h]
6254
· exact Nat.lt_of_lt_of_le (USize.toNat_lt_size _) (Nat.le_of_not_lt h')⟩
6355

6456
end Std.DHashMap.Internal

0 commit comments

Comments
 (0)