From fd68e4d619d13e3a9dde19a62c65752421a1dc42 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 25 Sep 2024 10:48:27 +1000 Subject: [PATCH] chore: missing BitVec lemmas --- src/Init/Data/BitVec/Lemmas.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index e26803b3fe39..497c06b8674d 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -927,6 +927,10 @@ theorem not_def {x : BitVec v} : ~~~x = allOnes v ^^^ x := rfl ext simp +@[simp] theorem not_allOnes : ~~~ allOnes w = 0#w := by + ext + simp + @[simp] theorem xor_allOnes {x : BitVec w} : x ^^^ allOnes w = ~~~ x := by ext i simp @@ -1410,6 +1414,10 @@ theorem msb_append {x : BitVec w} {y : BitVec v} : rw [getLsbD_append] simpa using lt_of_getLsbD +@[simp] theorem zero_append_zero : 0#v ++ 0#w = 0#(v + w) := by + ext + simp only [getLsbD_append, getLsbD_zero, Bool.cond_self] + @[simp] theorem cast_append_right (h : w + v = w + v') (x : BitVec w) (y : BitVec v) : cast h (x ++ y) = x ++ cast (by omega) y := by ext @@ -1655,6 +1663,10 @@ theorem getElem_concat (x : BitVec w) (b : Bool) (i : Nat) (h : i < w + 1) : (concat x a) ^^^ (concat y b) = concat (x ^^^ y) (a ^^ b) := by ext i; cases i using Fin.succRecOn <;> simp +@[simp] theorem zero_concat_false : concat 0#w false = 0#(w + 1) := by + ext + simp [getLsbD_concat] + /-! ### add -/ theorem add_def {n} (x y : BitVec n) : x + y = .ofNat n (x.toNat + y.toNat) := rfl @@ -2166,6 +2178,20 @@ theorem twoPow_zero {w : Nat} : twoPow w 0 = 1#w := by theorem getLsbD_one {w i : Nat} : (1#w).getLsbD i = (decide (0 < w) && decide (0 = i)) := by rw [← twoPow_zero, getLsbD_twoPow] +@[simp] theorem true_cons_zero : cons true 0#w = twoPow (w + 1) w := by + ext + simp [getLsbD_cons] + omega + +@[simp] theorem false_cons_zero : cons false 0#w = 0#(w + 1) := by + ext + simp [getLsbD_cons] + +@[simp] theorem zero_concat_true : concat 0#w true = 1#(w + 1) := by + ext + simp [getLsbD_concat] + omega + /- ### setWidth, setWidth, and bitwise operations -/ /--