Skip to content

Commit

Permalink
feat: start experimenting with 'abs' in bv_decide (#665)
Browse files Browse the repository at this point in the history
  • Loading branch information
tobiasgrosser authored Sep 28, 2024
1 parent 19b8e1b commit b71ca39
Show file tree
Hide file tree
Showing 3 changed files with 31 additions and 3 deletions.
23 changes: 23 additions & 0 deletions SSA/Projects/InstCombine/ForLean.lean
Original file line number Diff line number Diff line change
Expand Up @@ -89,6 +89,29 @@ end Nat

namespace BitVec

@[simp]
theorem replicate_one {w : Nat} : BitVec.replicate w 1#1 = cast (by simp) (allOnes w) := by
ext i
simp
omega

@[simp]
theorem replicate_zero {w : Nat} : BitVec.replicate w 0#1 = cast (by simp) (0#w) := by
ext i
simp

theorem abs_eq_add_xor {x : BitVec w} :
have y : BitVec w := BitVec.cast (by simp) (BitVec.replicate w (BitVec.ofBool x.msb))
x.abs = (x + y) ^^^ y := by
simp only [BitVec.abs, neg_eq]
by_cases h: x.msb
· simp [h, ← allOnes_sub_eq_not]
· simp [h]

theorem abs_eq_if {x : BitVec w} :
x.abs = if x.msb then -x else x := by
simp only [BitVec.abs, neg_eq]

@[simp, bv_toNat]
lemma toNat_shiftLeft' (A B : BitVec w) :
BitVec.toNat (A <<< B) = (BitVec.toNat A) * 2 ^ BitVec.toNat B % 2 ^w := by
Expand Down
6 changes: 4 additions & 2 deletions SSA/Projects/InstCombine/HackersDelight.lean
Original file line number Diff line number Diff line change
@@ -1,6 +1,8 @@
import SSA.Projects.InstCombine.ForMathlib
import SSA.Projects.InstCombine.ForStd
import SSA.Projects.InstCombine.TacticAuto
import Std.Tactic.BVDecide


set_option linter.unusedTactic false
set_option linter.unreachableTactic false
Expand All @@ -9,7 +11,7 @@ namespace HackersDelight

namespace Ch2Basics

variable {x y z : BitVec w}
variable {x y z : BitVec 32}

/- 2–3 Inequalities among Logical and Arithmetic Expressions -/

Expand Down Expand Up @@ -62,7 +64,7 @@ theorem neq_iff_neg_sub_abs :

theorem lt_iff_sub_xor_xor_and_sub_xor :
(x <ₛ y) ↔ ((x - y) ^^^ ((x ^^^ y) &&& ((x - y) ^^^ x))).msb := by
-- try alive_auto -- this leads to a heartbeat timeout
try alive_auto
all_goals sorry

theorem slt_iff_and_not_or_not_xor_and_sub :
Expand Down
5 changes: 4 additions & 1 deletion SSA/Projects/InstCombine/TacticAuto.lean
Original file line number Diff line number Diff line change
Expand Up @@ -155,7 +155,10 @@ macro "bv_auto": tactic =>
simp (config := {failIfUnchanged := false}) only [(BitVec.two_mul), ←BitVec.negOne_eq_allOnes]
bv_automata
)
| bv_decide
|
simp (config := {failIfUnchanged := false}) only [BitVec.abs_eq_if]
try split
all_goals bv_decide
| bv_distrib
)
)
Expand Down

0 comments on commit b71ca39

Please sign in to comment.