This module contains the implementation of a bitblaster for BitVec.add
. The implemented
circuit is a ripple carry adder.
structure
Std.Tactic.BVDecide.BVExpr.bitblast.FullAdderInput
{α : Type}
[Hashable α]
[DecidableEq α]
(aig : Sat.AIG α)
:
- lhs : aig.Ref
- rhs : aig.Ref
- cin : aig.Ref
Instances For
def
Std.Tactic.BVDecide.BVExpr.bitblast.FullAdderInput.cast
{α : Type}
[Hashable α]
[DecidableEq α]
{aig1 aig2 : Sat.AIG α}
(val : FullAdderInput aig1)
(h : aig1.decls.size ≤ aig2.decls.size)
:
FullAdderInput aig2
Equations
- { lhs := lhs, rhs := rhs, cin := cin }.cast h = { lhs := lhs.cast h, rhs := rhs.cast h, cin := cin.cast h }
Instances For
@[simp]
theorem
Std.Tactic.BVDecide.BVExpr.bitblast.FullAdderInput.lhs_cast
{α : Type}
[Hashable α]
[DecidableEq α]
{aig1 aig2 : Sat.AIG α}
(s : FullAdderInput aig1)
(hcast : aig1.decls.size ≤ aig2.decls.size)
:
(s.cast hcast).lhs = s.lhs.cast hcast
@[simp]
theorem
Std.Tactic.BVDecide.BVExpr.bitblast.FullAdderInput.rhs_cast
{α : Type}
[Hashable α]
[DecidableEq α]
{aig1 aig2 : Sat.AIG α}
(s : FullAdderInput aig1)
(hcast : aig1.decls.size ≤ aig2.decls.size)
:
(s.cast hcast).rhs = s.rhs.cast hcast
@[simp]
theorem
Std.Tactic.BVDecide.BVExpr.bitblast.FullAdderInput.cin_cast
{α : Type}
[Hashable α]
[DecidableEq α]
{aig1 aig2 : Sat.AIG α}
(s : FullAdderInput aig1)
(hcast : aig1.decls.size ≤ aig2.decls.size)
:
(s.cast hcast).cin = s.cin.cast hcast
def
Std.Tactic.BVDecide.BVExpr.bitblast.mkFullAdderOut
{α : Type}
[Hashable α]
[DecidableEq α]
(aig : Sat.AIG α)
(input : FullAdderInput aig)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
def
Std.Tactic.BVDecide.BVExpr.bitblast.mkFullAdderCarry
{α : Type}
[Hashable α]
[DecidableEq α]
(aig : Sat.AIG α)
(input : FullAdderInput aig)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
structure
Std.Tactic.BVDecide.BVExpr.bitblast.FullAdderOutput
{α : Type}
[Hashable α]
[DecidableEq α]
(old : Sat.AIG α)
:
Instances For
def
Std.Tactic.BVDecide.BVExpr.bitblast.mkFullAdder
{α : Type}
[Hashable α]
[DecidableEq α]
(aig : Sat.AIG α)
(input : FullAdderInput aig)
:
FullAdderOutput aig
Equations
- One or more equations did not get rendered due to their size.
Instances For
def
Std.Tactic.BVDecide.BVExpr.bitblast.blastAdd
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(input : aig.BinaryRefVec w)
:
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[irreducible]
def
Std.Tactic.BVDecide.BVExpr.bitblast.blastAdd.go
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(lhs rhs : aig.RefVec w)
(curr : Nat)
(hcurr : curr ≤ w)
(cin : aig.Ref)
(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.blastAdd.go_le_size
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(curr : Nat)
(hcurr : curr ≤ w)
(cin : aig.Ref)
(s : aig.RefVec curr)
(lhs rhs : aig.RefVec w)
:
@[irreducible]
theorem
Std.Tactic.BVDecide.BVExpr.bitblast.blastAdd.go_decl_eq
{α : Type}
[Hashable α]
[DecidableEq α]
{w : Nat}
(aig : Sat.AIG α)
(curr : Nat)
(hcurr : curr ≤ w)
(cin : aig.Ref)
(s : aig.RefVec curr)
(lhs rhs : aig.RefVec w)
(idx : Nat)
(h1 : idx < aig.decls.size)
(h2 : idx < (go aig lhs rhs curr hcurr cin s).aig.decls.size)
:
instance
Std.Tactic.BVDecide.BVExpr.bitblast.blastAdd.instLawfulVecOperatorBinaryRefVec
{α : Type}
[Hashable α]
[DecidableEq α]
:
Sat.AIG.LawfulVecOperator α Sat.AIG.BinaryRefVec fun {len : Nat} => blastAdd