This module contains the definition of the BitVec
fragment that bv_decide
internally operates
on as BVLogicalExpr
. The preprocessing steps of bv_decide
reduce all supported BitVec
operations to the ones provided in this file. For verification purposes BVLogicalExpr.Sat
and
BVLogicalExpr.Unsat
are provided.
All supported binary operations on BVExpr
.
- and: Std.Tactic.BVDecide.BVBinOp
Bitwise and.
- or: Std.Tactic.BVDecide.BVBinOp
Bitwise or.
- xor: Std.Tactic.BVDecide.BVBinOp
Bitwise xor.
- add: Std.Tactic.BVDecide.BVBinOp
Addition.
- mul: Std.Tactic.BVDecide.BVBinOp
Multiplication.
- udiv: Std.Tactic.BVDecide.BVBinOp
Unsigned division.
- umod: Std.Tactic.BVDecide.BVBinOp
Unsigned modulo.
Instances For
The semantics for BVBinOp
.
Instances For
All supported unary operators on BVExpr
.
- not: Std.Tactic.BVDecide.BVUnOp
Bitwise not.
- shiftLeftConst: Nat → Std.Tactic.BVDecide.BVUnOp
- shiftRightConst: Nat → Std.Tactic.BVDecide.BVUnOp
- rotateLeft: Nat → Std.Tactic.BVDecide.BVUnOp
Rotating left by a constant value.
- rotateRight: Nat → Std.Tactic.BVDecide.BVUnOp
Rotating right by a constant value.
- arithShiftRightConst: Nat → Std.Tactic.BVDecide.BVUnOp
Instances For
Equations
- Std.Tactic.BVDecide.BVUnOp.not.toString = "~"
- (Std.Tactic.BVDecide.BVUnOp.shiftLeftConst n).toString = toString "<< " ++ toString n ++ toString ""
- (Std.Tactic.BVDecide.BVUnOp.shiftRightConst n).toString = toString ">> " ++ toString n ++ toString ""
- (Std.Tactic.BVDecide.BVUnOp.rotateLeft n).toString = toString "rotL " ++ toString n ++ toString ""
- (Std.Tactic.BVDecide.BVUnOp.rotateRight n).toString = toString "rotR " ++ toString n ++ toString ""
- (Std.Tactic.BVDecide.BVUnOp.arithShiftRightConst n).toString = toString ">>a " ++ toString n ++ toString ""
Instances For
The semantics for BVUnOp
.
Instances For
All supported expressions involving BitVec
and operations on them.
- var: {w : Nat} → Nat → Std.Tactic.BVDecide.BVExpr w
A
BitVec
variable, referred to through an index. - const: {w : Nat} → BitVec w → Std.Tactic.BVDecide.BVExpr w
A constant
BitVec
value. - zeroExtend: {w : Nat} → (v : Nat) → Std.Tactic.BVDecide.BVExpr w → Std.Tactic.BVDecide.BVExpr v
zero extend a
BitVec
by some constant amount. - extract: {w : Nat} → Nat → (len : Nat) → Std.Tactic.BVDecide.BVExpr w → Std.Tactic.BVDecide.BVExpr len
Extract a slice from a
BitVec
. - bin: {w : Nat} →
Std.Tactic.BVDecide.BVExpr w →
Std.Tactic.BVDecide.BVBinOp → Std.Tactic.BVDecide.BVExpr w → Std.Tactic.BVDecide.BVExpr w
A binary operation on two
BVExpr
. - un: {w : Nat} → Std.Tactic.BVDecide.BVUnOp → Std.Tactic.BVDecide.BVExpr w → Std.Tactic.BVDecide.BVExpr w
A unary operation on two
BVExpr
. - append: {l r : Nat} → Std.Tactic.BVDecide.BVExpr l → Std.Tactic.BVDecide.BVExpr r → Std.Tactic.BVDecide.BVExpr (l + r)
Concatenate two bit vectors.
- replicate: {w : Nat} → (n : Nat) → Std.Tactic.BVDecide.BVExpr w → Std.Tactic.BVDecide.BVExpr (w * n)
Concatenate a bit vector with itself
n
times. - signExtend: {w : Nat} → (v : Nat) → Std.Tactic.BVDecide.BVExpr w → Std.Tactic.BVDecide.BVExpr v
sign extend a
BitVec
by some constant amount. - shiftLeft: {m n : Nat} → Std.Tactic.BVDecide.BVExpr m → Std.Tactic.BVDecide.BVExpr n → Std.Tactic.BVDecide.BVExpr m
shift left by another BitVec expression. For constant shifts there exists a
BVUnop
. - shiftRight: {m n : Nat} → Std.Tactic.BVDecide.BVExpr m → Std.Tactic.BVDecide.BVExpr n → Std.Tactic.BVDecide.BVExpr m
shift right by another BitVec expression. For constant shifts there exists a
BVUnop
. - arithShiftRight: {m n : Nat} → Std.Tactic.BVDecide.BVExpr m → Std.Tactic.BVDecide.BVExpr n → Std.Tactic.BVDecide.BVExpr m
shift right arithmetically by another BitVec expression. For constant shifts there exists a
BVUnop
.
Instances For
Instances For
Equations
- Std.Tactic.BVDecide.BVExpr.instToString = { toString := Std.Tactic.BVDecide.BVExpr.toString }
Get the value of a BVExpr.var
from an Assignment
.
Equations
- assign.getD idx = List.getD assign idx (Std.Tactic.BVDecide.BVExpr.PackedBitVec.mk (BitVec.zero 0))
Instances For
The semantics for BVExpr
.
Instances For
Supported binary predicates on BVExpr
.
- eq: Std.Tactic.BVDecide.BVBinPred
Equality.
- ult: Std.Tactic.BVDecide.BVBinPred
Unsigned Less Than
Instances For
Instances For
The semantics for BVBinPred
.
Instances For
Supported predicates on BVExpr
.
- bin: {w : Nat} →
Std.Tactic.BVDecide.BVExpr w →
Std.Tactic.BVDecide.BVBinPred → Std.Tactic.BVDecide.BVExpr w → Std.Tactic.BVDecide.BVPred
A binary predicate on
BVExpr
. - getLsbD: {w : Nat} → Std.Tactic.BVDecide.BVExpr w → Nat → Std.Tactic.BVDecide.BVPred
Getting a constant LSB from a
BitVec
.
Instances For
Pack two BVExpr
of equivalent width into one parameter-less structure.
- w : Nat
- lhs : Std.Tactic.BVDecide.BVExpr self.w
- rhs : Std.Tactic.BVDecide.BVExpr self.w
Instances For
Equations
The semantics for BVPred
.
Equations
- Std.Tactic.BVDecide.BVPred.eval assign (Std.Tactic.BVDecide.BVPred.bin lhs op rhs) = op.eval (Std.Tactic.BVDecide.BVExpr.eval assign lhs) (Std.Tactic.BVDecide.BVExpr.eval assign rhs)
- Std.Tactic.BVDecide.BVPred.eval assign (Std.Tactic.BVDecide.BVPred.getLsbD expr idx) = (Std.Tactic.BVDecide.BVExpr.eval assign expr).getLsbD idx
Instances For
Boolean substructure of problems involving predicates on BitVec as atoms.
Equations
Instances For
The semantics of boolean problems involving BitVec predicates as atoms.
Instances For
Equations
- x.Sat assign = (Std.Tactic.BVDecide.BVLogicalExpr.eval assign x = true)
Instances For
Equations
- x.Unsat = ∀ (f : Std.Tactic.BVDecide.BVExpr.Assignment), Std.Tactic.BVDecide.BVLogicalExpr.eval f x = false