Merge pull request #10310 from blishko/issue-10306

[SMTChecker] Adding a dummy frame to the call stack for the implicit constructor
This commit is contained in:
Leonardo 2020-11-17 09:34:49 -01:00 committed by GitHub
commit 65c2089b43
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
4 changed files with 62 additions and 23 deletions

View File

@ -120,7 +120,10 @@ void BMC::endVisit(ContractDefinition const& _contract)
constructor->accept(*this); constructor->accept(*this);
else else
{ {
/// Visiting implicit constructor - we need a dummy callstack frame
pushCallStack({nullptr, nullptr});
inlineConstructorHierarchy(_contract); inlineConstructorHierarchy(_contract);
popCallStack();
/// Check targets created by state variable initialization. /// Check targets created by state variable initialization.
smtutil::Expression constraints = m_context.assertions(); smtutil::Expression constraints = m_context.assertions();
checkVerificationTargets(constraints); checkVerificationTargets(constraints);
@ -831,32 +834,28 @@ void BMC::checkCondition(
{ {
case smtutil::CheckResult::SATISFIABLE: case smtutil::CheckResult::SATISFIABLE:
{ {
solAssert(!_callStack.empty(), "");
std::ostringstream message; std::ostringstream message;
message << "BMC: " << _description << " happens here."; message << "BMC: " << _description << " happens here.";
if (_callStack.size()) std::ostringstream modelMessage;
{ modelMessage << "Counterexample:\n";
std::ostringstream modelMessage; solAssert(values.size() == expressionNames.size(), "");
modelMessage << "Counterexample:\n"; map<string, string> sortedModel;
solAssert(values.size() == expressionNames.size(), ""); for (size_t i = 0; i < values.size(); ++i)
map<string, string> sortedModel; if (expressionsToEvaluate.at(i).name != values.at(i))
for (size_t i = 0; i < values.size(); ++i) sortedModel[expressionNames.at(i)] = values.at(i);
if (expressionsToEvaluate.at(i).name != values.at(i))
sortedModel[expressionNames.at(i)] = values.at(i);
for (auto const& eval: sortedModel) for (auto const& eval: sortedModel)
modelMessage << " " << eval.first << " = " << eval.second << "\n"; modelMessage << " " << eval.first << " = " << eval.second << "\n";
m_errorReporter.warning( m_errorReporter.warning(
_errorHappens, _errorHappens,
_location, _location,
message.str(), message.str(),
SecondarySourceLocation().append(modelMessage.str(), SourceLocation{}) SecondarySourceLocation().append(modelMessage.str(), SourceLocation{})
.append(SMTEncoder::callStackMessage(_callStack)) .append(SMTEncoder::callStackMessage(_callStack))
.append(move(secondaryLocation)) .append(move(secondaryLocation))
); );
}
else
m_errorReporter.warning(6084_error, _location, message.str(), secondaryLocation);
break; break;
} }
case smtutil::CheckResult::UNSATISFIABLE: case smtutil::CheckResult::UNSATISFIABLE:

View File

@ -0,0 +1,21 @@
pragma experimental SMTChecker;
contract C {
uint x = initX();
function initX() internal pure returns (uint) {
return 42;
}
}
contract D is C {
uint y;
constructor() {
assert(x == 42);
y = x;
}
}
// ====
// SMTEngine: bmc
// ----

View File

@ -0,0 +1,19 @@
pragma experimental SMTChecker;
contract C {
uint x = initX();
uint y = initY();
function initX() internal pure returns (uint) {
return 42;
}
function initY() internal view returns (uint) {
assert(x == 42);
return x;
}
}
// ====
// SMTEngine: bmc
// ----
// Warning 4661: (205-220): BMC: Assertion violation happens here.

View File

@ -22,4 +22,4 @@ contract C {
// Warning 4144: (217-222): BMC: Underflow (resulting value less than 0) happens here. // Warning 4144: (217-222): BMC: Underflow (resulting value less than 0) happens here.
// Warning 2661: (293-298): BMC: Overflow (resulting value larger than 2**256 - 1) happens here. // Warning 2661: (293-298): BMC: Overflow (resulting value larger than 2**256 - 1) happens here.
// Warning 3046: (369-374): BMC: Division by zero happens here. // Warning 3046: (369-374): BMC: Division by zero happens here.
// Warning 6084: (68-73): BMC: Underflow (resulting value less than 0) happens here. // Warning 4144: (68-73): BMC: Underflow (resulting value less than 0) happens here.