This module contains the implementation of a bitblaster for BitVec.shiftLeft
.
It distinguishes two cases:
- Shifting by a constant distance (trivial)
- Shifting by a symbolic
BitVec
distance (requires symbolic branches over the distance).
def
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeftConst
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(target : aig.ShiftTarget w)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[irreducible]
def
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeftConst.go
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(input : aig.RefVec w)
(distance curr : Nat)
(hcurr : curr ≤ w)
(s : aig.RefVec curr)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[irreducible]
theorem
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeftConst.go_le_size
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(distance : Nat)
(input : aig.RefVec w)
(curr : Nat)
(hcurr : curr ≤ w)
(s : aig.RefVec curr)
:
@[irreducible]
theorem
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeftConst.go_decl_eq
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(distance : Nat)
(input : aig.RefVec w)
(curr : Nat)
(hcurr : curr ≤ w)
(s : aig.RefVec curr)
(idx : Nat)
(h1 : idx < aig.decls.size)
(h2 : idx < (go aig input distance curr hcurr s).aig.decls.size)
:
instance
Std.Tactic.BVDecide.BVExpr.bitblast.instLawfulVecOperatorShiftTargetBlastShiftLeftConst
{α : Type}
[Hashable α]
[DecidableEq α]
:
Sat.AIG.LawfulVecOperator α Sat.AIG.ShiftTarget fun {len : Nat} => blastShiftLeftConst
structure
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeft.TwoPowShiftTarget
{α : Type}
[Hashable α]
[DecidableEq α]
(aig : Sat.AIG α)
(w : Nat)
:
Instances For
def
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeft.twoPowShift
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(target : TwoPowShiftTarget aig w)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
instance
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeft.instLawfulVecOperatorTwoPowShiftTargetTwoPowShift
{α : Type}
[Hashable α]
[DecidableEq α]
:
Sat.AIG.LawfulVecOperator α TwoPowShiftTarget fun {len : Nat} => twoPowShift
def
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeft
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(target : aig.ArbitraryShiftTarget w)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[irreducible]
def
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeft.go
{α : Type}
[Hashable α]
[DecidableEq α]
{w n : Nat}
(aig : Sat.AIG α)
(distance : aig.RefVec n)
(curr : Nat)
(acc : aig.RefVec w)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[irreducible]
theorem
Std.Tactic.BVDecide.BVExpr.bitblast.blastShiftLeft.go_le_size
{α : Type}
[Hashable α]
[DecidableEq α]
{n w : Nat}
(aig : Sat.AIG α)
(distance : aig.RefVec n)
(curr : Nat)
(acc : aig.RefVec w)
:
@[irreducible]
instance
Std.Tactic.BVDecide.BVExpr.bitblast.instLawfulVecOperatorArbitraryShiftTargetBlastShiftLeft
{α : Type}
[Hashable α]
[DecidableEq α]
:
Sat.AIG.LawfulVecOperator α Sat.AIG.ArbitraryShiftTarget fun {len : Nat} => blastShiftLeft