Documentation

Std.Tactic.BVDecide.Bitblast.BVExpr.Circuit.Lemmas.Const

This module contains the verification of the BitVec constant bitblaster from Impl.Const.

@[irreducible]
theorem Std.Tactic.BVDecide.BVExpr.bitblast.blastConst.go_get_aux {α : Type} [Hashable α] [DecidableEq α] {w : Nat} (aig : Sat.AIG α) (c : BitVec w) (curr : Nat) (hcurr : curr w) (s : aig.RefVec curr) (idx : Nat) (hidx : idx < curr) (hfoo : aig.decls.size aig.decls.size) :
(go aig c curr s hcurr).get idx = (s.get idx hidx).cast hfoo
theorem Std.Tactic.BVDecide.BVExpr.bitblast.blastConst.go_get {α : Type} [Hashable α] [DecidableEq α] {w : Nat} (aig : Sat.AIG α) (c : BitVec w) (curr : Nat) (hcurr : curr w) (s : aig.RefVec curr) (idx : Nat) (hidx : idx < curr) :
(go aig c curr s hcurr).get idx = s.get idx hidx
@[irreducible]
theorem Std.Tactic.BVDecide.BVExpr.bitblast.blastConst.go_denote_eq {α : Type} [Hashable α] [DecidableEq α] {w : Nat} (aig : Sat.AIG α) (c : BitVec w) (assign : αBool) (curr : Nat) (hcurr : curr w) (s : aig.RefVec curr) (idx : Nat) (hidx1 : idx < w) :
curr idxassign, { aig := aig, ref := (go aig c curr s hcurr).get idx hidx1 } = c.getLsbD idx
@[simp]
theorem Std.Tactic.BVDecide.BVExpr.bitblast.denote_blastConst {α : Type} [Hashable α] [DecidableEq α] {w : Nat} (aig : Sat.AIG α) (c : BitVec w) (assign : αBool) (idx : Nat) (hidx : idx < w) :
assign, { aig := aig, ref := (blastConst aig c).get idx hidx } = c.getLsbD idx