Skip to content

Commit 6eb3a8d

Browse files
committed
feat: BitVec.sshiftRight' in bv_decide
1 parent 5f7a40a commit 6eb3a8d

File tree

10 files changed

+318
-10
lines changed

10 files changed

+318
-10
lines changed

src/Init/Data/BitVec/Bitblast.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1092,8 +1092,8 @@ def sshiftRightRec (x : BitVec w₁) (y : BitVec w₂) (n : Nat) : BitVec w₁ :
10921092

10931093
@[simp]
10941094
theorem sshiftRightRec_zero_eq (x : BitVec w₁) (y : BitVec w₂) :
1095-
sshiftRightRec x y 0 = x.sshiftRight' (y &&& 1#w₂) := by
1096-
simp only [sshiftRightRec, twoPow_zero]
1095+
sshiftRightRec x y 0 = x.sshiftRight' (y &&& twoPow w₂ 0) := by
1096+
simp only [sshiftRightRec]
10971097

10981098
@[simp]
10991099
theorem sshiftRightRec_succ_eq (x : BitVec w₁) (y : BitVec w₂) (n : Nat) :

src/Lean/Elab/Tactic/BVDecide/Frontend/BVDecide/Reflect.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -65,6 +65,8 @@ where
6565
mkApp4 (mkConst ``BVExpr.shiftLeft) (toExpr m) (toExpr n) (go lhs) (go rhs)
6666
| .shiftRight (m := m) (n := n) lhs rhs =>
6767
mkApp4 (mkConst ``BVExpr.shiftRight) (toExpr m) (toExpr n) (go lhs) (go rhs)
68+
| .arithShiftRight (m := m) (n := n) lhs rhs =>
69+
mkApp4 (mkConst ``BVExpr.arithShiftRight) (toExpr m) (toExpr n) (go lhs) (go rhs)
6870

6971
instance : ToExpr BVBinPred where
7072
toExpr x :=

src/Lean/Elab/Tactic/BVDecide/Frontend/BVDecide/Reify.lean

Lines changed: 10 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -60,8 +60,8 @@ where
6060
``BVUnOp.shiftLeftConst
6161
``Std.Tactic.BVDecide.Reflect.BitVec.shiftLeftNat_congr
6262
else
63+
let_expr BitVec _ := β | return none
6364
shiftReflection
64-
β
6565
distanceExpr
6666
innerExpr
6767
.shiftLeft
@@ -78,8 +78,8 @@ where
7878
``BVUnOp.shiftRightConst
7979
``Std.Tactic.BVDecide.Reflect.BitVec.shiftRightNat_congr
8080
else
81+
let_expr BitVec _ := β | return none
8182
shiftReflection
82-
β
8383
distanceExpr
8484
innerExpr
8585
.shiftRight
@@ -92,6 +92,13 @@ where
9292
innerExpr
9393
.arithShiftRightConst
9494
``BVUnOp.arithShiftRightConst
95+
``Std.Tactic.BVDecide.Reflect.BitVec.arithShiftRightNat_congr
96+
| BitVec.sshiftRight' _ _ innerExpr distanceExpr =>
97+
shiftReflection
98+
distanceExpr
99+
innerExpr
100+
.arithShiftRight
101+
``BVExpr.arithShiftRight
95102
``Std.Tactic.BVDecide.Reflect.BitVec.arithShiftRight_congr
96103
| BitVec.zeroExtend _ newWidthExpr innerExpr =>
97104
let some newWidth ← getNatValue? newWidthExpr | return none
@@ -258,11 +265,10 @@ where
258265
let some distance ← ReifiedBVExpr.getNatOrBvValue? β distanceExpr | return none
259266
shiftConstLikeReflection distance innerExpr shiftOp shiftOpName congrThm
260267

261-
shiftReflection (β : Expr) (distanceExpr : Expr) (innerExpr : Expr)
268+
shiftReflection (distanceExpr : Expr) (innerExpr : Expr)
262269
(shiftOp : {m n : Nat} → BVExpr m → BVExpr n → BVExpr m) (shiftOpName : Name)
263270
(congrThm : Name) :
264271
LemmaM (Option ReifiedBVExpr) := do
265-
let_expr BitVec _ ← β | return none
266272
let some inner ← goOrAtom innerExpr | return none
267273
let some distance ← goOrAtom distanceExpr | return none
268274
let bvExpr : BVExpr inner.width := shiftOp inner.bvExpr distance.bvExpr

src/Std/Tactic/BVDecide/Bitblast/BVExpr/Basic.lean

Lines changed: 12 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -245,6 +245,10 @@ inductive BVExpr : Nat → Type where
245245
shift right by another BitVec expression. For constant shifts there exists a `BVUnop`.
246246
-/
247247
| shiftRight (lhs : BVExpr m) (rhs : BVExpr n) : BVExpr m
248+
/--
249+
shift right arithmetically by another BitVec expression. For constant shifts there exists a `BVUnop`.
250+
-/
251+
| arithShiftRight (lhs : BVExpr m) (rhs : BVExpr n) : BVExpr m
248252

249253
namespace BVExpr
250254

@@ -259,7 +263,8 @@ def toString : BVExpr w → String
259263
| .replicate n expr => s!"(replicate {n} {toString expr})"
260264
| .signExtend v expr => s!"(sext {v} {expr.toString})"
261265
| .shiftLeft lhs rhs => s!"({lhs.toString} << {rhs.toString})"
262-
| .shiftRight lhs rhs => s!"({lhs.toString} >> {rhs.toString})"
266+
| .shiftRight lhs rhs => s!"({lhs.toString} u>> {rhs.toString})"
267+
| .arithShiftRight lhs rhs => s!"({lhs.toString} s>> {rhs.toString})"
263268

264269

265270
instance : ToString (BVExpr w) := ⟨toString⟩
@@ -299,6 +304,7 @@ def eval (assign : Assignment) : BVExpr w → BitVec w
299304
| .signExtend v expr => BitVec.signExtend v (eval assign expr)
300305
| .shiftLeft lhs rhs => (eval assign lhs) <<< (eval assign rhs)
301306
| .shiftRight lhs rhs => (eval assign lhs) >>> (eval assign rhs)
307+
| .arithShiftRight lhs rhs => BitVec.sshiftRight' (eval assign lhs) (eval assign rhs)
302308

303309
@[simp]
304310
theorem eval_var : eval assign ((.var idx) : BVExpr w) = (assign.getD idx).bv.truncate _ := by
@@ -343,6 +349,11 @@ theorem eval_shiftLeft : eval assign (.shiftLeft lhs rhs) = (eval assign lhs) <<
343349
theorem eval_shiftRight : eval assign (.shiftRight lhs rhs) = (eval assign lhs) >>> (eval assign rhs) := by
344350
rfl
345351

352+
@[simp]
353+
theorem eval_arithShiftRight :
354+
eval assign (.arithShiftRight lhs rhs) = BitVec.sshiftRight' (eval assign lhs) (eval assign rhs) := by
355+
rfl
356+
346357
end BVExpr
347358

348359
/--

src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Impl/Expr.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -214,6 +214,18 @@ where
214214
dsimp only at hlaig hraig
215215
omega
216216
⟨res, this⟩
217+
| .arithShiftRight lhs rhs =>
218+
let ⟨⟨aig, lhs⟩, hlaig⟩ := go aig lhs
219+
let ⟨⟨aig, rhs⟩, hraig⟩ := go aig rhs
220+
let lhs := lhs.cast <| by
221+
dsimp only at hlaig hraig
222+
omega
223+
let res := bitblast.blastArithShiftRight aig ⟨_, lhs, rhs⟩
224+
have := by
225+
apply AIG.LawfulVecOperator.le_size_of_le_aig_size (f := bitblast.blastArithShiftRight)
226+
dsimp only at hlaig hraig
227+
omega
228+
⟨res, this⟩
217229

218230
theorem bitblast.go_decl_eq (aig : AIG BVBit) (expr : BVExpr w) :
219231
∀ (idx : Nat) (h1) (h2), (go aig expr).val.aig.decls[idx]'h2 = aig.decls[idx]'h1 := by
@@ -300,6 +312,16 @@ theorem bitblast.go_decl_eq (aig : AIG BVBit) (expr : BVExpr w) :
300312
· omega
301313
· apply Nat.lt_of_lt_of_le h1
302314
apply Nat.le_trans <;> assumption
315+
| arithShiftRight lhs rhs lih rih =>
316+
dsimp only [go]
317+
have := (bitblast.go aig lhs).property
318+
have := (bitblast.go aig lhs).property
319+
have := (go (go aig lhs).1.aig rhs).property
320+
rw [AIG.LawfulVecOperator.decl_eq (f := blastArithShiftRight)]
321+
rw [rih, lih]
322+
· omega
323+
· apply Nat.lt_of_lt_of_le h1
324+
apply Nat.le_trans <;> assumption
303325

304326
instance : AIG.LawfulVecOperator BVBit (fun _ w => BVExpr w) bitblast where
305327
le_size := by

src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Impl/Operations/ShiftRight.lean

Lines changed: 116 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -126,14 +126,14 @@ instance : AIG.LawfulVecOperator α AIG.ShiftTarget blastArithShiftRightConst wh
126126
unfold blastArithShiftRightConst
127127
simp
128128

129-
namespace blastShiftRight
130-
131129
structure TwoPowShiftTarget (aig : AIG α) (w : Nat) where
132130
n : Nat
133131
lhs : AIG.RefVec aig w
134132
rhs : AIG.RefVec aig n
135133
pow : Nat
136134

135+
namespace blastShiftRight
136+
137137
def twoPowShift (aig : AIG α) (target : TwoPowShiftTarget aig w) : AIG.RefVecEntry α w :=
138138
let ⟨n, lhs, rhs, pow⟩ := target
139139
if h : pow < n then
@@ -246,6 +246,120 @@ instance : AIG.LawfulVecOperator α AIG.ArbitraryShiftTarget blastShiftRight whe
246246
apply AIG.LawfulVecOperator.lt_size_of_lt_aig_size (f := blastShiftRight.twoPowShift)
247247
assumption
248248

249+
namespace blastArithShiftRight
250+
251+
def twoPowShift (aig : AIG α) (target : TwoPowShiftTarget aig w) : AIG.RefVecEntry α w :=
252+
let ⟨n, lhs, rhs, pow⟩ := target
253+
if h : pow < n then
254+
let res := blastArithShiftRightConst aig ⟨lhs, 2 ^ pow⟩
255+
let aig := res.aig
256+
let shifted := res.vec
257+
258+
have := AIG.LawfulVecOperator.le_size (f := blastArithShiftRightConst) ..
259+
let rhs := rhs.cast this
260+
let lhs := lhs.cast this
261+
AIG.RefVec.ite aig ⟨rhs.get pow h, shifted, lhs⟩
262+
else
263+
⟨aig, lhs⟩
264+
265+
instance : AIG.LawfulVecOperator α TwoPowShiftTarget twoPowShift where
266+
le_size := by
267+
intros
268+
unfold twoPowShift
269+
dsimp only
270+
split
271+
· apply AIG.LawfulVecOperator.le_size_of_le_aig_size (f := AIG.RefVec.ite)
272+
apply AIG.LawfulVecOperator.le_size (f := blastArithShiftRightConst)
273+
· simp
274+
decl_eq := by
275+
intros
276+
unfold twoPowShift
277+
dsimp only
278+
split
279+
· rw [AIG.LawfulVecOperator.decl_eq (f := AIG.RefVec.ite)]
280+
rw [AIG.LawfulVecOperator.decl_eq (f := blastArithShiftRightConst)]
281+
apply AIG.LawfulVecOperator.lt_size_of_lt_aig_size (f := blastArithShiftRightConst)
282+
assumption
283+
· simp
284+
285+
end blastArithShiftRight
286+
287+
def blastArithShiftRight (aig : AIG α) (target : AIG.ArbitraryShiftTarget aig w) :
288+
AIG.RefVecEntry α w :=
289+
let ⟨n, input, distance⟩ := target
290+
if n = 0 then
291+
⟨aig, input⟩
292+
else
293+
let res := blastArithShiftRight.twoPowShift aig ⟨_, input, distance, 0
294+
let aig := res.aig
295+
let acc := res.vec
296+
have := AIG.LawfulVecOperator.le_size (f := blastArithShiftRight.twoPowShift) ..
297+
let distance := distance.cast this
298+
go aig distance 0 acc
299+
where
300+
go {n : Nat} (aig : AIG α) (distance : AIG.RefVec aig n) (curr : Nat)
301+
(acc : AIG.RefVec aig w) :
302+
AIG.RefVecEntry α w :=
303+
if curr < n - 1 then
304+
let res := blastArithShiftRight.twoPowShift aig ⟨_, acc, distance, curr + 1
305+
let aig := res.aig
306+
let acc := res.vec
307+
have := AIG.LawfulVecOperator.le_size (f := blastArithShiftRight.twoPowShift) ..
308+
let distance := distance.cast this
309+
go aig distance (curr + 1) acc
310+
else
311+
⟨aig, acc⟩
312+
termination_by n - 1 - curr
313+
314+
theorem blastArithShiftRight.go_le_size (aig : AIG α) (distance : AIG.RefVec aig n) (curr : Nat)
315+
(acc : AIG.RefVec aig w) :
316+
aig.decls.size ≤ (go aig distance curr acc).aig.decls.size := by
317+
unfold go
318+
dsimp only
319+
split
320+
· refine Nat.le_trans ?_ (by apply go_le_size)
321+
apply AIG.LawfulVecOperator.le_size (f := blastArithShiftRight.twoPowShift)
322+
· simp
323+
termination_by n - 1 - curr
324+
325+
theorem blastArithShiftRight.go_decl_eq (aig : AIG α) (distance : AIG.RefVec aig n) (curr : Nat)
326+
(acc : AIG.RefVec aig w) :
327+
∀ (idx : Nat) (h1) (h2),
328+
(go aig distance curr acc).aig.decls[idx]'h2 = aig.decls[idx]'h1 := by
329+
generalize hgo : go aig distance curr acc = res
330+
unfold go at hgo
331+
dsimp only at hgo
332+
split at hgo
333+
· rw [← hgo]
334+
intros
335+
rw [blastArithShiftRight.go_decl_eq]
336+
rw [AIG.LawfulVecOperator.decl_eq (f := blastArithShiftRight.twoPowShift)]
337+
apply AIG.LawfulVecOperator.lt_size_of_lt_aig_size (f := blastArithShiftRight.twoPowShift)
338+
assumption
339+
· simp [← hgo]
340+
termination_by n - 1 - curr
341+
342+
343+
instance : AIG.LawfulVecOperator α AIG.ArbitraryShiftTarget blastArithShiftRight where
344+
le_size := by
345+
intros
346+
unfold blastArithShiftRight
347+
dsimp only
348+
split
349+
· simp
350+
· refine Nat.le_trans ?_ (by apply blastArithShiftRight.go_le_size)
351+
apply AIG.LawfulVecOperator.le_size (f := blastArithShiftRight.twoPowShift)
352+
decl_eq := by
353+
intros
354+
unfold blastArithShiftRight
355+
dsimp only
356+
split
357+
· simp
358+
· rw [blastArithShiftRight.go_decl_eq]
359+
rw [AIG.LawfulVecOperator.decl_eq (f := blastArithShiftRight.twoPowShift)]
360+
apply AIG.LawfulVecOperator.lt_size_of_lt_aig_size (f := blastArithShiftRight.twoPowShift)
361+
assumption
362+
249363
end bitblast
250364
end BVExpr
251365

src/Std/Tactic/BVDecide/Bitblast/BVExpr/Circuit/Lemmas/Expr.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -147,6 +147,18 @@ theorem go_denote_eq (aig : AIG BVBit) (expr : BVExpr w) (assign : Assignment) :
147147
· simp [Ref.hgate]
148148
· intros
149149
rw [← rih]
150+
| arithShiftRight lhs rhs lih rih =>
151+
simp only [go, eval_arithShiftRight]
152+
apply denote_blastArithShiftRight
153+
· intros
154+
dsimp only
155+
rw [go_denote_mem_prefix]
156+
rw [← lih (aig := aig)]
157+
· simp
158+
· assumption
159+
· simp [Ref.hgate]
160+
· intros
161+
rw [← rih]
150162
| bin lhs op rhs lih rih =>
151163
cases op with
152164
| and =>

0 commit comments

Comments
 (0)