mirror of
https://github.com/ethereum/solidity
synced 2023-10-03 13:03:40 +00:00
Merge pull request #10389 from ethereum/smt_constructors_nonlinear_clauses
[SMTChecker] Use nonlinear clauses instead of inlining base constructors in CHC
This commit is contained in:
commit
4e4aaaf028
@ -149,6 +149,9 @@ bool BMC::visit(FunctionDefinition const& _function)
|
||||
resetStateVariables();
|
||||
}
|
||||
|
||||
if (_function.isConstructor())
|
||||
inlineConstructorHierarchy(dynamic_cast<ContractDefinition const&>(*_function.scope()));
|
||||
|
||||
/// Already visits the children.
|
||||
SMTEncoder::visit(_function);
|
||||
|
||||
|
@ -112,20 +112,11 @@ vector<string> CHC::unhandledQueries() const
|
||||
bool CHC::visit(ContractDefinition const& _contract)
|
||||
{
|
||||
resetContractAnalysis();
|
||||
|
||||
initContract(_contract);
|
||||
|
||||
m_stateVariables = SMTEncoder::stateVariablesIncludingInheritedAndPrivate(_contract);
|
||||
|
||||
clearIndices(&_contract);
|
||||
|
||||
m_stateVariables = SMTEncoder::stateVariablesIncludingInheritedAndPrivate(_contract);
|
||||
solAssert(m_currentContract, "");
|
||||
m_constructorSummaryPredicate = createSymbolicBlock(
|
||||
constructorSort(*m_currentContract, state()),
|
||||
"summary_constructor_" + contractSuffix(_contract),
|
||||
PredicateType::ConstructorSummary,
|
||||
&_contract
|
||||
);
|
||||
|
||||
SMTEncoder::visit(_contract);
|
||||
return false;
|
||||
@ -133,27 +124,64 @@ bool CHC::visit(ContractDefinition const& _contract)
|
||||
|
||||
void CHC::endVisit(ContractDefinition const& _contract)
|
||||
{
|
||||
auto implicitConstructorPredicate = createSymbolicBlock(
|
||||
implicitConstructorSort(state()),
|
||||
"implicit_constructor_" + contractSuffix(_contract),
|
||||
PredicateType::ImplicitConstructor,
|
||||
&_contract
|
||||
);
|
||||
addRule(
|
||||
(*implicitConstructorPredicate)({0, state().thisAddress(), state().crypto(), state().tx(), state().state()}),
|
||||
implicitConstructorPredicate->functor().name
|
||||
);
|
||||
setCurrentBlock(*implicitConstructorPredicate);
|
||||
|
||||
if (auto constructor = _contract.constructor())
|
||||
constructor->accept(*this);
|
||||
else
|
||||
inlineConstructorHierarchy(_contract);
|
||||
|
||||
defineContractInitializer(_contract);
|
||||
|
||||
auto const& entry = *createConstructorBlock(_contract, "implicit_constructor_entry");
|
||||
|
||||
// In case constructors use uninitialized state variables,
|
||||
// they need to be zeroed.
|
||||
// This is not part of `initialConstraints` because it's only true here,
|
||||
// at the beginning of the deployment routine.
|
||||
smtutil::Expression zeroes(true);
|
||||
for (auto var: stateVariablesIncludingInheritedAndPrivate(_contract))
|
||||
zeroes = zeroes && currentValue(*var) == smt::zeroValue(var->type());
|
||||
addRule(smtutil::Expression::implies(initialConstraints(_contract) && zeroes, predicate(entry)), entry.functor().name);
|
||||
setCurrentBlock(entry);
|
||||
|
||||
solAssert(!m_errorDest, "");
|
||||
m_errorDest = m_constructorSummaries.at(&_contract);
|
||||
// We need to evaluate the base constructor calls (arguments) from derived -> base
|
||||
auto baseArgs = baseArguments(_contract);
|
||||
for (auto base: _contract.annotation().linearizedBaseContracts)
|
||||
{
|
||||
if (base != &_contract)
|
||||
{
|
||||
m_callGraph[&_contract].insert(base);
|
||||
vector<ASTPointer<Expression>> const& args = baseArgs.count(base) ? baseArgs.at(base) : decltype(args){};
|
||||
|
||||
auto baseConstructor = base->constructor();
|
||||
if (baseConstructor && !args.empty())
|
||||
{
|
||||
auto const& params = baseConstructor->parameters();
|
||||
solAssert(params.size() == args.size(), "");
|
||||
for (unsigned i = 0; i < params.size(); ++i)
|
||||
{
|
||||
args.at(i)->accept(*this);
|
||||
if (params.at(i))
|
||||
{
|
||||
solAssert(m_context.knownVariable(*params.at(i)), "");
|
||||
m_context.addAssertion(currentValue(*params.at(i)) == expr(*args.at(i), params.at(i)->type()));
|
||||
}
|
||||
}
|
||||
}
|
||||
}
|
||||
}
|
||||
m_errorDest = nullptr;
|
||||
// Then call initializer_Base from base -> derived
|
||||
for (auto base: _contract.annotation().linearizedBaseContracts | boost::adaptors::reversed)
|
||||
{
|
||||
errorFlag().increaseIndex();
|
||||
m_context.addAssertion(smt::constructorCall(*m_contractInitializers.at(base), m_context));
|
||||
connectBlocks(m_currentBlock, summary(_contract), errorFlag().currentValue() > 0);
|
||||
m_context.addAssertion(errorFlag().currentValue() == 0);
|
||||
}
|
||||
|
||||
connectBlocks(m_currentBlock, summary(_contract));
|
||||
|
||||
setCurrentBlock(*m_constructorSummaryPredicate);
|
||||
|
||||
setCurrentBlock(*m_constructorSummaries.at(&_contract));
|
||||
m_queryPlaceholders[&_contract].push_back({smtutil::Expression(true), errorFlag().currentValue(), m_currentBlock});
|
||||
connectBlocks(m_currentBlock, interface(), errorFlag().currentValue() == 0);
|
||||
|
||||
@ -168,16 +196,7 @@ bool CHC::visit(FunctionDefinition const& _function)
|
||||
return false;
|
||||
}
|
||||
|
||||
// This is the case for base constructor inlining.
|
||||
if (m_currentFunction)
|
||||
{
|
||||
solAssert(m_currentFunction->isConstructor(), "");
|
||||
solAssert(_function.isConstructor(), "");
|
||||
solAssert(_function.scope() != m_currentContract, "");
|
||||
SMTEncoder::visit(_function);
|
||||
return false;
|
||||
}
|
||||
|
||||
// No inlining.
|
||||
solAssert(!m_currentFunction, "Function inlining should not happen in CHC.");
|
||||
m_currentFunction = &_function;
|
||||
|
||||
@ -189,23 +208,19 @@ bool CHC::visit(FunctionDefinition const& _function)
|
||||
auto functionPred = predicate(*functionEntryBlock);
|
||||
auto bodyPred = predicate(*bodyBlock);
|
||||
|
||||
if (_function.isConstructor())
|
||||
connectBlocks(m_currentBlock, functionPred);
|
||||
else
|
||||
addRule(functionPred, functionPred.name);
|
||||
addRule(functionPred, functionPred.name);
|
||||
|
||||
m_context.addAssertion(errorFlag().currentValue() == 0);
|
||||
for (auto const* var: m_stateVariables)
|
||||
m_context.addAssertion(m_context.variable(*var)->valueAtIndex(0) == currentValue(*var));
|
||||
for (auto const& var: _function.parameters())
|
||||
m_context.addAssertion(m_context.variable(*var)->valueAtIndex(0) == currentValue(*var));
|
||||
m_context.addAssertion(state().state(0) == state().state());
|
||||
solAssert(m_currentContract, "");
|
||||
m_context.addAssertion(initialConstraints(*m_currentContract, &_function));
|
||||
|
||||
connectBlocks(functionPred, bodyPred);
|
||||
|
||||
setCurrentBlock(*bodyBlock);
|
||||
|
||||
solAssert(!m_errorDest, "");
|
||||
m_errorDest = m_summaries.at(m_currentContract).at(&_function);
|
||||
SMTEncoder::visit(*m_currentFunction);
|
||||
m_errorDest = nullptr;
|
||||
|
||||
return false;
|
||||
}
|
||||
@ -216,55 +231,29 @@ void CHC::endVisit(FunctionDefinition const& _function)
|
||||
return;
|
||||
|
||||
solAssert(m_currentFunction && m_currentContract, "");
|
||||
// No inlining.
|
||||
solAssert(m_currentFunction == &_function, "");
|
||||
|
||||
// This is the case for base constructor inlining.
|
||||
if (m_currentFunction != &_function)
|
||||
connectBlocks(m_currentBlock, summary(_function));
|
||||
setCurrentBlock(*m_summaries.at(m_currentContract).at(&_function));
|
||||
|
||||
// Query placeholders for constructors are not created here because
|
||||
// of contracts without constructors.
|
||||
// Instead, those are created in endVisit(ContractDefinition).
|
||||
if (!_function.isConstructor())
|
||||
{
|
||||
solAssert(m_currentFunction && m_currentFunction->isConstructor(), "");
|
||||
solAssert(_function.isConstructor(), "");
|
||||
solAssert(_function.scope() != m_currentContract, "");
|
||||
}
|
||||
else
|
||||
{
|
||||
// We create an extra exit block for constructors that simply
|
||||
// connects to the interface in case an explicit constructor
|
||||
// exists in the hierarchy.
|
||||
// It is not connected directly here, as normal functions are,
|
||||
// because of the case where there are only implicit constructors.
|
||||
// This is done in endVisit(ContractDefinition).
|
||||
if (_function.isConstructor())
|
||||
auto sum = summary(_function);
|
||||
auto ifacePre = smt::interfacePre(*m_interfaces.at(m_currentContract), *m_currentContract, m_context);
|
||||
if (_function.isPublic())
|
||||
{
|
||||
string suffix = m_currentContract->name() + "_" + to_string(m_currentContract->id());
|
||||
solAssert(m_currentContract, "");
|
||||
auto constructorExit = createSymbolicBlock(
|
||||
constructorSort(*m_currentContract, state()),
|
||||
"constructor_exit_" + suffix,
|
||||
PredicateType::ConstructorSummary,
|
||||
m_currentContract
|
||||
);
|
||||
connectBlocks(m_currentBlock, predicate(*constructorExit));
|
||||
|
||||
setCurrentBlock(*constructorExit);
|
||||
auto txConstraints = m_context.state().txConstraints(_function);
|
||||
m_queryPlaceholders[&_function].push_back({txConstraints && sum, errorFlag().currentValue(), ifacePre});
|
||||
connectBlocks(ifacePre, interface(), txConstraints && sum && errorFlag().currentValue() == 0);
|
||||
}
|
||||
else
|
||||
{
|
||||
auto assertionError = errorFlag().currentValue();
|
||||
auto sum = summary(_function);
|
||||
connectBlocks(m_currentBlock, sum);
|
||||
auto iface = interface();
|
||||
setCurrentBlock(*m_interfaces.at(m_currentContract));
|
||||
|
||||
auto ifacePre = smt::interfacePre(*m_interfaces.at(m_currentContract), *m_currentContract, m_context);
|
||||
if (_function.isPublic())
|
||||
{
|
||||
auto txConstraints = m_context.state().txConstraints(_function);
|
||||
m_queryPlaceholders[&_function].push_back({txConstraints && sum, assertionError, ifacePre});
|
||||
connectBlocks(ifacePre, iface, txConstraints && sum && (assertionError == 0));
|
||||
}
|
||||
}
|
||||
m_currentFunction = nullptr;
|
||||
}
|
||||
|
||||
m_currentFunction = nullptr;
|
||||
|
||||
SMTEncoder::endVisit(_function);
|
||||
}
|
||||
|
||||
@ -564,10 +553,11 @@ void CHC::internalFunctionCall(FunctionCall const& _funCall)
|
||||
|
||||
m_context.addAssertion(predicate(_funCall));
|
||||
|
||||
solAssert(m_errorDest, "");
|
||||
connectBlocks(
|
||||
m_currentBlock,
|
||||
(m_currentFunction && !m_currentFunction->isConstructor()) ? summary(*m_currentFunction) : summary(*m_currentContract),
|
||||
(errorFlag().currentValue() > 0)
|
||||
predicate(*m_errorDest),
|
||||
errorFlag().currentValue() > 0
|
||||
);
|
||||
m_context.addAssertion(errorFlag().currentValue() == 0);
|
||||
}
|
||||
@ -644,9 +634,10 @@ void CHC::externalFunctionCallToTrustedCode(FunctionCall const& _funCall)
|
||||
state().newTx();
|
||||
m_context.addAssertion(originalTx == state().tx());
|
||||
|
||||
solAssert(m_errorDest, "");
|
||||
connectBlocks(
|
||||
m_currentBlock,
|
||||
(m_currentFunction && !m_currentFunction->isConstructor()) ? summary(*m_currentFunction) : summary(*m_currentContract),
|
||||
predicate(*m_errorDest),
|
||||
(errorFlag().currentValue() > 0)
|
||||
);
|
||||
m_context.addAssertion(errorFlag().currentValue() == 0);
|
||||
@ -728,6 +719,8 @@ void CHC::resetSourceAnalysis()
|
||||
m_summaries.clear();
|
||||
m_interfaces.clear();
|
||||
m_nondetInterfaces.clear();
|
||||
m_constructorSummaries.clear();
|
||||
m_contractInitializers.clear();
|
||||
Predicate::reset();
|
||||
ArraySlicePredicate::reset();
|
||||
m_blockCounter = 0;
|
||||
@ -840,6 +833,8 @@ void CHC::defineInterfacesAndSummaries(SourceUnit const& _source)
|
||||
string suffix = contract->name() + "_" + to_string(contract->id());
|
||||
m_interfaces[contract] = createSymbolicBlock(interfaceSort(*contract, state()), "interface_" + suffix, PredicateType::Interface, contract);
|
||||
m_nondetInterfaces[contract] = createSymbolicBlock(nondetInterfaceSort(*contract, state()), "nondet_interface_" + suffix, PredicateType::NondetInterface, contract);
|
||||
m_constructorSummaries[contract] = createConstructorBlock(*contract, "summary_constructor");
|
||||
m_contractInitializers[contract] = createConstructorBlock(*contract, "contract_initializer");
|
||||
|
||||
for (auto const* var: stateVariablesIncludingInheritedAndPrivate(*contract))
|
||||
if (!m_context.knownVariable(*var))
|
||||
@ -890,6 +885,39 @@ void CHC::defineInterfacesAndSummaries(SourceUnit const& _source)
|
||||
}
|
||||
}
|
||||
|
||||
void CHC::defineContractInitializer(ContractDefinition const& _contract)
|
||||
{
|
||||
auto const& implicitConstructorPredicate = *createConstructorBlock(_contract, "contract_initializer_entry");
|
||||
|
||||
auto implicitFact = smt::constructor(implicitConstructorPredicate, m_context);
|
||||
addRule(smtutil::Expression::implies(initialConstraints(_contract), implicitFact), implicitFact.name);
|
||||
setCurrentBlock(implicitConstructorPredicate);
|
||||
|
||||
solAssert(!m_errorDest, "");
|
||||
m_errorDest = m_contractInitializers.at(&_contract);
|
||||
for (auto var: _contract.stateVariables())
|
||||
if (var->value())
|
||||
{
|
||||
var->value()->accept(*this);
|
||||
assignment(*var, *var->value());
|
||||
}
|
||||
m_errorDest = nullptr;
|
||||
|
||||
auto const& afterInit = *createConstructorBlock(_contract, "contract_initializer_after_init");
|
||||
connectBlocks(m_currentBlock, predicate(afterInit));
|
||||
setCurrentBlock(afterInit);
|
||||
|
||||
if (auto constructor = _contract.constructor())
|
||||
{
|
||||
errorFlag().increaseIndex();
|
||||
m_context.addAssertion(smt::functionCall(*m_summaries.at(&_contract).at(constructor), &_contract, m_context));
|
||||
connectBlocks(m_currentBlock, initializer(_contract), errorFlag().currentValue() > 0);
|
||||
m_context.addAssertion(errorFlag().currentValue() == 0);
|
||||
}
|
||||
|
||||
connectBlocks(m_currentBlock, initializer(_contract));
|
||||
}
|
||||
|
||||
smtutil::Expression CHC::interface()
|
||||
{
|
||||
solAssert(m_currentContract, "");
|
||||
@ -911,14 +939,19 @@ smtutil::Expression CHC::error(unsigned _idx)
|
||||
return m_errorPredicate->functor(_idx)({});
|
||||
}
|
||||
|
||||
smtutil::Expression CHC::initializer(ContractDefinition const& _contract)
|
||||
{
|
||||
return predicate(*m_contractInitializers.at(&_contract));
|
||||
}
|
||||
|
||||
smtutil::Expression CHC::summary(ContractDefinition const& _contract)
|
||||
{
|
||||
return constructor(*m_constructorSummaryPredicate, _contract, m_context);
|
||||
return predicate(*m_constructorSummaries.at(&_contract));
|
||||
}
|
||||
|
||||
smtutil::Expression CHC::summary(FunctionDefinition const& _function, ContractDefinition const& _contract)
|
||||
{
|
||||
return smt::function(*m_summaries.at(&_contract).at(&_function), _function, &_contract, m_context);
|
||||
return smt::function(*m_summaries.at(&_contract).at(&_function), &_contract, m_context);
|
||||
}
|
||||
|
||||
smtutil::Expression CHC::summary(FunctionDefinition const& _function)
|
||||
@ -942,14 +975,22 @@ Predicate const* CHC::createBlock(ASTNode const* _node, PredicateType _predType,
|
||||
|
||||
Predicate const* CHC::createSummaryBlock(FunctionDefinition const& _function, ContractDefinition const& _contract)
|
||||
{
|
||||
auto block = createSymbolicBlock(
|
||||
return createSymbolicBlock(
|
||||
functionSort(_function, &_contract, state()),
|
||||
"summary_" + uniquePrefix() + "_" + predicateName(&_function, &_contract),
|
||||
PredicateType::FunctionSummary,
|
||||
&_function
|
||||
);
|
||||
}
|
||||
|
||||
return block;
|
||||
Predicate const* CHC::createConstructorBlock(ContractDefinition const& _contract, string const& _prefix)
|
||||
{
|
||||
return createSymbolicBlock(
|
||||
constructorSort(_contract, state()),
|
||||
_prefix + "_" + contractSuffix(_contract) + "_" + uniquePrefix(),
|
||||
PredicateType::ConstructorSummary,
|
||||
&_contract
|
||||
);
|
||||
}
|
||||
|
||||
void CHC::createErrorBlock()
|
||||
@ -967,6 +1008,21 @@ void CHC::connectBlocks(smtutil::Expression const& _from, smtutil::Expression co
|
||||
addRule(edge, _from.name + "_to_" + _to.name);
|
||||
}
|
||||
|
||||
smtutil::Expression CHC::initialConstraints(ContractDefinition const& _contract, FunctionDefinition const* _function)
|
||||
{
|
||||
smtutil::Expression conj = state().state() == state().state(0);
|
||||
conj = conj && errorFlag().currentValue() == 0;
|
||||
for (auto var: stateVariablesIncludingInheritedAndPrivate(_contract))
|
||||
conj = conj && m_context.variable(*var)->valueAtIndex(0) == currentValue(*var);
|
||||
|
||||
FunctionDefinition const* function = _function ? _function : _contract.constructor();
|
||||
if (function)
|
||||
for (auto var: function->parameters())
|
||||
conj = conj && m_context.variable(*var)->valueAtIndex(0) == currentValue(*var);
|
||||
|
||||
return conj;
|
||||
}
|
||||
|
||||
vector<smtutil::Expression> CHC::initialStateVariables()
|
||||
{
|
||||
return stateVariablesAtIndex(0);
|
||||
@ -1021,16 +1077,11 @@ smtutil::Expression CHC::predicate(Predicate const& _block)
|
||||
case PredicateType::Interface:
|
||||
solAssert(m_currentContract, "");
|
||||
return ::interface(_block, *m_currentContract, m_context);
|
||||
case PredicateType::ImplicitConstructor:
|
||||
solAssert(m_currentContract, "");
|
||||
return implicitConstructor(_block, *m_currentContract, m_context);
|
||||
case PredicateType::ConstructorSummary:
|
||||
solAssert(m_currentContract, "");
|
||||
return constructor(_block, *m_currentContract, m_context);
|
||||
return constructor(_block, m_context);
|
||||
case PredicateType::FunctionEntry:
|
||||
case PredicateType::FunctionSummary:
|
||||
solAssert(m_currentFunction, "");
|
||||
return smt::function(_block, *m_currentFunction, m_currentContract, m_context);
|
||||
return smt::function(_block, m_currentContract, m_context);
|
||||
case PredicateType::FunctionBlock:
|
||||
solAssert(m_currentFunction, "");
|
||||
return functionBlock(_block, *m_currentFunction, m_currentContract, m_context);
|
||||
@ -1173,9 +1224,10 @@ void CHC::verificationTargetEncountered(
|
||||
errorFlag().increaseIndex();
|
||||
|
||||
// create an error edge to the summary
|
||||
solAssert(m_errorDest, "");
|
||||
connectBlocks(
|
||||
m_currentBlock,
|
||||
scopeIsFunction ? summary(*m_currentFunction) : summary(*m_currentContract),
|
||||
predicate(*m_errorDest),
|
||||
currentPathConditions() && _errorCondition && errorFlag().currentValue() == errorId
|
||||
);
|
||||
|
||||
@ -1416,16 +1468,12 @@ optional<string> CHC::generateCounterexample(CHCSolverInterface::CexGraph const&
|
||||
string txCex = summaryPredicate->formatSummaryCall(summaryArgs);
|
||||
path.emplace_back(txCex);
|
||||
|
||||
/// Recurse on the next interface node which represents the previous transaction
|
||||
/// or stop.
|
||||
if (interfaceId)
|
||||
{
|
||||
Predicate const* interfacePredicate = Predicate::predicate(_graph.nodes.at(*interfaceId).name);
|
||||
solAssert(interfacePredicate && interfacePredicate->isInterface(), "");
|
||||
node = *interfaceId;
|
||||
}
|
||||
else
|
||||
/// Stop when we reach the summary of the analyzed constructor.
|
||||
if (summaryPredicate->type() == PredicateType::ConstructorSummary)
|
||||
break;
|
||||
|
||||
/// Recurse on the next interface node which represents the previous transaction.
|
||||
node = *interfaceId;
|
||||
}
|
||||
|
||||
return localState + "\nTransaction trace:\n" + boost::algorithm::join(boost::adaptors::reverse(path), "\n");
|
||||
|
@ -126,6 +126,11 @@ private:
|
||||
/// in a given _source.
|
||||
void defineInterfacesAndSummaries(SourceUnit const& _source);
|
||||
|
||||
/// Creates a CHC system that, for a given contract,
|
||||
/// - initializes its state variables (as 0 or given value, if any).
|
||||
/// - "calls" the explicit constructor function of the contract, if any.
|
||||
void defineContractInitializer(ContractDefinition const& _contract);
|
||||
|
||||
/// Interface predicate over current variables.
|
||||
smtutil::Expression interface();
|
||||
smtutil::Expression interface(ContractDefinition const& _contract);
|
||||
@ -139,12 +144,18 @@ private:
|
||||
/// The contract is needed here because of inheritance.
|
||||
Predicate const* createSummaryBlock(FunctionDefinition const& _function, ContractDefinition const& _contract);
|
||||
|
||||
/// @returns a block related to @a _contract's constructor.
|
||||
Predicate const* createConstructorBlock(ContractDefinition const& _contract, std::string const& _prefix);
|
||||
|
||||
/// Creates a new error block to be used by an assertion.
|
||||
/// Also registers the predicate.
|
||||
void createErrorBlock();
|
||||
|
||||
void connectBlocks(smtutil::Expression const& _from, smtutil::Expression const& _to, smtutil::Expression const& _constraints = smtutil::Expression(true));
|
||||
|
||||
/// @returns The initial constraints that set up the beginning of a function.
|
||||
smtutil::Expression initialConstraints(ContractDefinition const& _contract, FunctionDefinition const* _function = nullptr);
|
||||
|
||||
/// @returns the symbolic values of the state variables at the beginning
|
||||
/// of the current transaction.
|
||||
std::vector<smtutil::Expression> initialStateVariables();
|
||||
@ -160,6 +171,8 @@ private:
|
||||
smtutil::Expression predicate(Predicate const& _block);
|
||||
/// @returns the summary predicate for the called function.
|
||||
smtutil::Expression predicate(FunctionCall const& _funCall);
|
||||
/// @returns a predicate that defines a contract initializer.
|
||||
smtutil::Expression initializer(ContractDefinition const& _contract);
|
||||
/// @returns a predicate that defines a constructor summary.
|
||||
smtutil::Expression summary(ContractDefinition const& _contract);
|
||||
/// @returns a predicate that defines a function summary.
|
||||
@ -231,10 +244,6 @@ private:
|
||||
|
||||
/// Predicates.
|
||||
//@{
|
||||
/// Constructor summary predicate, exists after the constructor
|
||||
/// (implicit or explicit) and before the interface.
|
||||
Predicate const* m_constructorSummaryPredicate = nullptr;
|
||||
|
||||
/// Artificial Interface predicate.
|
||||
/// Single entry block for all functions.
|
||||
std::map<ContractDefinition const*, Predicate const*> m_interfaces;
|
||||
@ -245,6 +254,9 @@ private:
|
||||
/// nondeterministically.
|
||||
std::map<ContractDefinition const*, Predicate const*> m_nondetInterfaces;
|
||||
|
||||
std::map<ContractDefinition const*, Predicate const*> m_constructorSummaries;
|
||||
std::map<ContractDefinition const*, Predicate const*> m_contractInitializers;
|
||||
|
||||
/// Artificial Error predicate.
|
||||
/// Single error block for all assertions.
|
||||
Predicate const* m_errorPredicate = nullptr;
|
||||
@ -311,9 +323,16 @@ private:
|
||||
bool m_unknownFunctionCallSeen = false;
|
||||
|
||||
/// Block where a loop break should go to.
|
||||
Predicate const* m_breakDest;
|
||||
Predicate const* m_breakDest = nullptr;
|
||||
/// Block where a loop continue should go to.
|
||||
Predicate const* m_continueDest;
|
||||
Predicate const* m_continueDest = nullptr;
|
||||
|
||||
/// Block where an error condition should go to.
|
||||
/// This can be:
|
||||
/// 1) Constructor initializer summary, if error happens while evaluating initial values of state variables.
|
||||
/// 2) Constructor summary, if error happens while evaluating base constructor arguments.
|
||||
/// 3) Function summary, if error happens inside a function.
|
||||
Predicate const* m_errorDest = nullptr;
|
||||
//@}
|
||||
|
||||
/// CHC solver.
|
||||
|
@ -141,12 +141,12 @@ optional<vector<VariableDeclaration const*>> Predicate::stateVariables() const
|
||||
|
||||
bool Predicate::isSummary() const
|
||||
{
|
||||
return functor().name.rfind("summary", 0) == 0;
|
||||
return m_type == PredicateType::ConstructorSummary || m_type == PredicateType::FunctionSummary;
|
||||
}
|
||||
|
||||
bool Predicate::isInterface() const
|
||||
{
|
||||
return functor().name.rfind("interface", 0) == 0;
|
||||
return m_type == PredicateType::Interface;
|
||||
}
|
||||
|
||||
string Predicate::formatSummaryCall(vector<smtutil::Expression> const& _args) const
|
||||
@ -190,7 +190,7 @@ string Predicate::formatSummaryCall(vector<smtutil::Expression> const& _args) co
|
||||
vector<optional<string>> Predicate::summaryStateValues(vector<smtutil::Expression> const& _args) const
|
||||
{
|
||||
/// The signature of a function summary predicate is: summary(error, this, cryptoFunctions, txData, preBlockchainState, preStateVars, preInputVars, postBlockchainState, postStateVars, postInputVars, outputVars).
|
||||
/// The signature of an implicit constructor summary predicate is: summary(error, this, cryptoFunctions, txData, preBlockchainState, postBlockchainState, postStateVars).
|
||||
/// The signature of the summary predicate of a contract without constructor is: summary(error, this, cryptoFunctions, txData, preBlockchainState, postBlockchainState, preStateVars, postStateVars).
|
||||
/// Here we are interested in postStateVars.
|
||||
auto stateVars = stateVariables();
|
||||
solAssert(stateVars.has_value(), "");
|
||||
@ -204,7 +204,7 @@ vector<optional<string>> Predicate::summaryStateValues(vector<smtutil::Expressio
|
||||
}
|
||||
else if (programContract())
|
||||
{
|
||||
stateFirst = _args.begin() + 6;
|
||||
stateFirst = _args.begin() + 6 + static_cast<int>(stateVars->size());
|
||||
stateLast = stateFirst + static_cast<int>(stateVars->size());
|
||||
}
|
||||
else
|
||||
|
@ -34,7 +34,6 @@ enum class PredicateType
|
||||
{
|
||||
Interface,
|
||||
NondetInterface,
|
||||
ImplicitConstructor,
|
||||
ConstructorSummary,
|
||||
FunctionEntry,
|
||||
FunctionSummary,
|
||||
|
@ -51,31 +51,50 @@ smtutil::Expression nondetInterface(Predicate const& _pred, ContractDefinition c
|
||||
);
|
||||
}
|
||||
|
||||
smtutil::Expression implicitConstructor(Predicate const& _pred, ContractDefinition const&, EncodingContext& _context)
|
||||
smtutil::Expression constructor(Predicate const& _pred, EncodingContext& _context)
|
||||
{
|
||||
auto& state = _context.state();
|
||||
vector<smtutil::Expression> stateExprs{state.errorFlag().currentValue(), state.thisAddress(0), state.crypto(0), state.tx(0), state.state(0)};
|
||||
return _pred(stateExprs);
|
||||
}
|
||||
|
||||
smtutil::Expression constructor(Predicate const& _pred, ContractDefinition const& _contract, EncodingContext& _context)
|
||||
{
|
||||
if (auto const* constructor = _contract.constructor())
|
||||
return _pred(currentFunctionVariables(*constructor, &_contract, _context));
|
||||
auto const& contract = dynamic_cast<ContractDefinition const&>(*_pred.programNode());
|
||||
if (auto const* constructor = contract.constructor())
|
||||
return _pred(currentFunctionVariablesForDefinition(*constructor, &contract, _context));
|
||||
|
||||
auto& state = _context.state();
|
||||
vector<smtutil::Expression> stateExprs{state.errorFlag().currentValue(), state.thisAddress(0), state.crypto(0), state.tx(0), state.state(0), state.state()};
|
||||
return _pred(stateExprs + currentStateVariables(_contract, _context));
|
||||
return _pred(stateExprs + initialStateVariables(contract, _context) + currentStateVariables(contract, _context));
|
||||
}
|
||||
|
||||
smtutil::Expression constructorCall(Predicate const& _pred, EncodingContext& _context)
|
||||
{
|
||||
auto const& contract = dynamic_cast<ContractDefinition const&>(*_pred.programNode());
|
||||
if (auto const* constructor = contract.constructor())
|
||||
return _pred(currentFunctionVariablesForCall(*constructor, &contract, _context));
|
||||
|
||||
auto& state = _context.state();
|
||||
vector<smtutil::Expression> stateExprs{state.errorFlag().currentValue(), state.thisAddress(0), state.crypto(0), state.tx(0), state.state()};
|
||||
state.newState();
|
||||
stateExprs += vector<smtutil::Expression>{state.state()};
|
||||
stateExprs += currentStateVariables(contract, _context);
|
||||
stateExprs += newStateVariables(contract, _context);
|
||||
return _pred(stateExprs);
|
||||
}
|
||||
|
||||
smtutil::Expression function(
|
||||
Predicate const& _pred,
|
||||
FunctionDefinition const& _function,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
)
|
||||
{
|
||||
return _pred(currentFunctionVariables(_function, _contract, _context));
|
||||
auto const& function = dynamic_cast<FunctionDefinition const&>(*_pred.programNode());
|
||||
return _pred(currentFunctionVariablesForDefinition(function, _contract, _context));
|
||||
}
|
||||
|
||||
smtutil::Expression functionCall(
|
||||
Predicate const& _pred,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
)
|
||||
{
|
||||
auto const& function = dynamic_cast<FunctionDefinition const&>(*_pred.programNode());
|
||||
return _pred(currentFunctionVariablesForCall(function, _contract, _context));
|
||||
}
|
||||
|
||||
smtutil::Expression functionBlock(
|
||||
@ -111,14 +130,22 @@ vector<smtutil::Expression> currentStateVariables(ContractDefinition const& _con
|
||||
);
|
||||
}
|
||||
|
||||
vector<smtutil::Expression> currentFunctionVariables(
|
||||
vector<smtutil::Expression> newStateVariables(ContractDefinition const& _contract, EncodingContext& _context)
|
||||
{
|
||||
return applyMap(
|
||||
SMTEncoder::stateVariablesIncludingInheritedAndPrivate(_contract),
|
||||
[&](auto _var) { return _context.variable(*_var)->increaseIndex(); }
|
||||
);
|
||||
}
|
||||
|
||||
vector<smtutil::Expression> currentFunctionVariablesForDefinition(
|
||||
FunctionDefinition const& _function,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
)
|
||||
{
|
||||
auto& state = _context.state();
|
||||
vector<smtutil::Expression> exprs{_context.state().errorFlag().currentValue(), state.thisAddress(0), state.crypto(0), state.tx(0), state.state(0)};
|
||||
vector<smtutil::Expression> exprs{state.errorFlag().currentValue(), state.thisAddress(0), state.crypto(0), state.tx(0), state.state(0)};
|
||||
exprs += _contract ? initialStateVariables(*_contract, _context) : vector<smtutil::Expression>{};
|
||||
exprs += applyMap(_function.parameters(), [&](auto _var) { return _context.variable(*_var)->valueAtIndex(0); });
|
||||
exprs += vector<smtutil::Expression>{state.state()};
|
||||
@ -128,9 +155,29 @@ vector<smtutil::Expression> currentFunctionVariables(
|
||||
return exprs;
|
||||
}
|
||||
|
||||
vector<smtutil::Expression> currentFunctionVariablesForCall(
|
||||
FunctionDefinition const& _function,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
)
|
||||
{
|
||||
auto& state = _context.state();
|
||||
vector<smtutil::Expression> exprs{state.errorFlag().currentValue(), state.thisAddress(0), state.crypto(0), state.tx(0), state.state()};
|
||||
exprs += _contract ? currentStateVariables(*_contract, _context) : vector<smtutil::Expression>{};
|
||||
exprs += applyMap(_function.parameters(), [&](auto _var) { return _context.variable(*_var)->currentValue(); });
|
||||
|
||||
state.newState();
|
||||
|
||||
exprs += vector<smtutil::Expression>{state.state()};
|
||||
exprs += _contract ? newStateVariables(*_contract, _context) : vector<smtutil::Expression>{};
|
||||
exprs += applyMap(_function.parameters(), [&](auto _var) { return _context.variable(*_var)->increaseIndex(); });
|
||||
exprs += applyMap(_function.returnParameters(), [&](auto _var) { return _context.variable(*_var)->currentValue(); });
|
||||
return exprs;
|
||||
}
|
||||
|
||||
vector<smtutil::Expression> currentBlockVariables(FunctionDefinition const& _function, ContractDefinition const* _contract, EncodingContext& _context)
|
||||
{
|
||||
return currentFunctionVariables(_function, _contract, _context) +
|
||||
return currentFunctionVariablesForDefinition(_function, _contract, _context) +
|
||||
applyMap(
|
||||
SMTEncoder::localVariablesIncludingModifiers(_function),
|
||||
[&](auto _var) { return _context.variable(*_var)->currentValue(); }
|
||||
|
@ -36,13 +36,17 @@ smtutil::Expression interface(Predicate const& _pred, ContractDefinition const&
|
||||
|
||||
smtutil::Expression nondetInterface(Predicate const& _pred, ContractDefinition const& _contract, EncodingContext& _context, unsigned _preIdx, unsigned _postIdx);
|
||||
|
||||
smtutil::Expression constructor(Predicate const& _pred, ContractDefinition const& _contract, EncodingContext& _context);
|
||||
|
||||
smtutil::Expression implicitConstructor(Predicate const& _pred, ContractDefinition const& _contract, EncodingContext& _context);
|
||||
smtutil::Expression constructor(Predicate const& _pred, EncodingContext& _context);
|
||||
smtutil::Expression constructorCall(Predicate const& _pred, EncodingContext& _context);
|
||||
|
||||
smtutil::Expression function(
|
||||
Predicate const& _pred,
|
||||
FunctionDefinition const& _function,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
);
|
||||
|
||||
smtutil::Expression functionCall(
|
||||
Predicate const& _pred,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
);
|
||||
@ -62,7 +66,15 @@ std::vector<smtutil::Expression> stateVariablesAtIndex(unsigned _index, Contract
|
||||
|
||||
std::vector<smtutil::Expression> currentStateVariables(ContractDefinition const& _contract, EncodingContext& _context);
|
||||
|
||||
std::vector<smtutil::Expression> currentFunctionVariables(
|
||||
std::vector<smtutil::Expression> newStateVariables(ContractDefinition const& _contract, EncodingContext& _context);
|
||||
|
||||
std::vector<smtutil::Expression> currentFunctionVariablesForDefinition(
|
||||
FunctionDefinition const& _function,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
);
|
||||
|
||||
std::vector<smtutil::Expression> currentFunctionVariablesForCall(
|
||||
FunctionDefinition const& _function,
|
||||
ContractDefinition const* _contract,
|
||||
EncodingContext& _context
|
||||
|
@ -46,21 +46,15 @@ SortPointer nondetInterfaceSort(ContractDefinition const& _contract, SymbolicSta
|
||||
);
|
||||
}
|
||||
|
||||
SortPointer implicitConstructorSort(SymbolicState& _state)
|
||||
{
|
||||
return make_shared<FunctionSort>(
|
||||
vector<SortPointer>{_state.errorFlagSort(), _state.thisAddressSort(), _state.cryptoSort(), _state.txSort(), _state.stateSort()},
|
||||
SortProvider::boolSort
|
||||
);
|
||||
}
|
||||
|
||||
SortPointer constructorSort(ContractDefinition const& _contract, SymbolicState& _state)
|
||||
{
|
||||
if (auto const* constructor = _contract.constructor())
|
||||
return functionSort(*constructor, &_contract, _state);
|
||||
|
||||
auto varSorts = stateSorts(_contract);
|
||||
vector<SortPointer> stateSort{_state.stateSort()};
|
||||
return make_shared<FunctionSort>(
|
||||
vector<SortPointer>{_state.errorFlagSort(), _state.thisAddressSort(), _state.cryptoSort(), _state.txSort(), _state.stateSort(), _state.stateSort()} + stateSorts(_contract),
|
||||
vector<SortPointer>{_state.errorFlagSort(), _state.thisAddressSort(), _state.cryptoSort(), _state.txSort(), _state.stateSort(), _state.stateSort()} + varSorts + varSorts,
|
||||
SortProvider::boolSort
|
||||
);
|
||||
}
|
||||
|
@ -39,19 +39,17 @@ namespace solidity::frontend::smt
|
||||
* The nondeterminism behavior of a contract. Signature:
|
||||
* nondet_interface(blockchainState, stateVariables, blockchainState', stateVariables').
|
||||
*
|
||||
* 3. Implicit constructor
|
||||
* The implicit constructor of a contract, that is, without input parameters. Signature:
|
||||
* implicit_constructor(error, this, cryptoFunctions, txData, blockchainState).
|
||||
* 3. Constructor entry/summary
|
||||
* The summary of a contract's deployment procedure.
|
||||
* Signature:
|
||||
* If the contract has a constructor function, this is the same as the summary of that function. Otherwise:
|
||||
* constructor_summary(error, this, cryptoFunctions, txData, blockchainState, blockchainState', stateVariables, stateVariables').
|
||||
*
|
||||
* 4. Constructor entry/summary
|
||||
* The summary of an implicit constructor. Signature:
|
||||
* constructor_summary(error, this, cryptoFunctions, txData, blockchainState, blockchainState', stateVariables').
|
||||
*
|
||||
* 5. Function entry/summary
|
||||
* 4. Function entry/summary
|
||||
* The entry point of a function definition. Signature:
|
||||
* function_entry(error, this, cryptoFunctions, txData, blockchainState, stateVariables, inputVariables, blockchainState', stateVariables', inputVariables', outputVariables').
|
||||
*
|
||||
* 6. Function body
|
||||
* 5. Function body
|
||||
* Use for any predicate within a function. Signature:
|
||||
* function_body(error, this, txData, blockchainState, stateVariables, inputVariables, blockchainState', stateVariables', inputVariables', outputVariables', localVariables).
|
||||
*/
|
||||
@ -62,9 +60,6 @@ smtutil::SortPointer interfaceSort(ContractDefinition const& _contract, Symbolic
|
||||
/// @returns the nondeterminisc interface predicate sort for _contract.
|
||||
smtutil::SortPointer nondetInterfaceSort(ContractDefinition const& _contract, SymbolicState& _state);
|
||||
|
||||
/// @returns the implicit constructor predicate sort.
|
||||
smtutil::SortPointer implicitConstructorSort(SymbolicState& _state);
|
||||
|
||||
/// @returns the constructor entry/summary predicate sort for _contract.
|
||||
smtutil::SortPointer constructorSort(ContractDefinition const& _contract, SymbolicState& _state);
|
||||
|
||||
|
@ -131,9 +131,6 @@ bool SMTEncoder::visit(FunctionDefinition const& _function)
|
||||
{
|
||||
m_modifierDepthStack.push_back(-1);
|
||||
|
||||
if (_function.isConstructor())
|
||||
inlineConstructorHierarchy(dynamic_cast<ContractDefinition const&>(*_function.scope()));
|
||||
|
||||
initializeLocalVariables(_function);
|
||||
|
||||
_function.parameterList().accept(*this);
|
||||
@ -2563,6 +2560,37 @@ SourceUnit const* SMTEncoder::sourceUnitContaining(Scopable const& _scopable)
|
||||
solAssert(false, "");
|
||||
}
|
||||
|
||||
map<ContractDefinition const*, vector<ASTPointer<frontend::Expression>>> SMTEncoder::baseArguments(ContractDefinition const& _contract)
|
||||
{
|
||||
map<ContractDefinition const*, vector<ASTPointer<Expression>>> baseArgs;
|
||||
|
||||
for (auto contract: _contract.annotation().linearizedBaseContracts)
|
||||
{
|
||||
/// Collect base contracts and potential constructor arguments.
|
||||
for (auto specifier: contract->baseContracts())
|
||||
{
|
||||
solAssert(specifier, "");
|
||||
auto const& base = dynamic_cast<ContractDefinition const&>(*specifier->name().annotation().referencedDeclaration);
|
||||
if (auto args = specifier->arguments())
|
||||
baseArgs[&base] = *args;
|
||||
}
|
||||
/// Collect base constructor arguments given as constructor modifiers.
|
||||
if (auto constructor = contract->constructor())
|
||||
for (auto mod: constructor->modifiers())
|
||||
{
|
||||
auto decl = mod->name()->annotation().referencedDeclaration;
|
||||
if (auto base = dynamic_cast<ContractDefinition const*>(decl))
|
||||
{
|
||||
solAssert(!baseArgs.count(base), "");
|
||||
if (auto args = mod->arguments())
|
||||
baseArgs[base] = *args;
|
||||
}
|
||||
}
|
||||
}
|
||||
|
||||
return baseArgs;
|
||||
}
|
||||
|
||||
void SMTEncoder::createReturnedExpressions(FunctionCall const& _funCall)
|
||||
{
|
||||
FunctionDefinition const* funDef = functionCallToDefinition(_funCall);
|
||||
|
@ -77,6 +77,9 @@ public:
|
||||
/// @returns the SourceUnit that contains _scopable.
|
||||
static SourceUnit const* sourceUnitContaining(Scopable const& _scopable);
|
||||
|
||||
/// @returns the arguments for each base constructor call in the hierarchy of @a _contract.
|
||||
std::map<ContractDefinition const*, std::vector<ASTPointer<frontend::Expression>>> baseArguments(ContractDefinition const& _contract);
|
||||
|
||||
protected:
|
||||
// TODO: Check that we do not have concurrent reads and writes to a variable,
|
||||
// because the order of expression evaluation is undefined
|
||||
|
@ -19,6 +19,6 @@ contract A is B {
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (203-208): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (244-249): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (232-250): CHC: Assertion violation happens here.
|
||||
// Warning 4984: (203-208): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
|
@ -18,6 +18,6 @@ contract A is B {
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (230-235): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (207-212): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (198-203): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (230-235): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
|
@ -28,6 +28,6 @@ contract A is B2, B1 {
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (160-165): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (225-230): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (241-246): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (225-230): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (334-350): CHC: Assertion violation happens here.
|
||||
|
@ -19,8 +19,8 @@ contract C {
|
||||
function f() public view {
|
||||
uint y = this.m(0,1,2,3);
|
||||
assert(y == m[0][1][2][3]); // should hold
|
||||
assert(y == 1); // should fail
|
||||
// Disabled because of Spacer seg fault
|
||||
//assert(y == 1); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (401-415): CHC: Assertion violation happens here.
|
||||
|
@ -14,8 +14,8 @@ contract C {
|
||||
function f() public view {
|
||||
uint y = this.m(0,1,2,3);
|
||||
assert(y == m[0][1][2][3]); // should hold
|
||||
assert(y == 1); // should fail
|
||||
// Disabled because Spacer seg faults
|
||||
//assert(y == 1); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (349-363): CHC: Assertion violation happens here.
|
||||
|
@ -0,0 +1,36 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
|
||||
function g() internal returns (uint) {
|
||||
x = 42;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
contract Z is B {
|
||||
constructor(uint z) B(z + f()) {
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z(5) {
|
||||
constructor() {
|
||||
assert(x == 6);
|
||||
assert(x > 9); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (325-332): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (400-413): CHC: Assertion violation happens here.
|
@ -0,0 +1,38 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A(9) {
|
||||
constructor(uint b) {
|
||||
x += b;
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
|
||||
function g() internal returns (uint) {
|
||||
x = 42;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
contract Z is B {
|
||||
constructor(uint z) B(z + f()) {
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z(5) {
|
||||
constructor() {
|
||||
assert(x == 15);
|
||||
assert(x > 90); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (143-149): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (333-340): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (409-423): CHC: Assertion violation happens here.
|
@ -0,0 +1,34 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b + f()) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
abstract contract Z is A {
|
||||
uint k;
|
||||
constructor(uint z) {
|
||||
k = z;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z, B {
|
||||
constructor() B(x) Z(x) {
|
||||
assert(x == 1);
|
||||
assert(k == 0);
|
||||
assert(x == k); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (138-145): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (384-398): CHC: Assertion violation happens here.
|
@ -0,0 +1,33 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
abstract contract Z is A {
|
||||
uint k;
|
||||
constructor(uint z) {
|
||||
k = z;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z, B {
|
||||
constructor() B(f()) Z(f()) {
|
||||
assert(x == 1);
|
||||
assert(k == 2);
|
||||
assert(x == k); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (382-396): CHC: Assertion violation happens here.
|
@ -0,0 +1,35 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b + f()) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
abstract contract Z is A {
|
||||
uint k;
|
||||
constructor(uint z) {
|
||||
k = z;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z, B {
|
||||
constructor(uint c) B(c) Z(x) {
|
||||
assert(x == c + 1);
|
||||
assert(k == 0);
|
||||
assert(x == k); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (138-145): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (366-371): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (394-408): CHC: Assertion violation happens here.
|
@ -0,0 +1,35 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b + f()) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
abstract contract Z is A {
|
||||
uint k;
|
||||
constructor(uint z) {
|
||||
k = z;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z, B {
|
||||
constructor(uint c) Z(x) B(c) {
|
||||
assert(x == c + 1);
|
||||
assert(k == 0);
|
||||
assert(x == k); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (138-145): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (366-371): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (394-408): CHC: Assertion violation happens here.
|
@ -0,0 +1,39 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b + f()) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
|
||||
function g() internal returns (uint) {
|
||||
x = 42;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
abstract contract Z is A {
|
||||
uint k;
|
||||
constructor(uint z) {
|
||||
k = z;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z, B {
|
||||
constructor() Z(g()) B(f()) {
|
||||
assert(x == 44);
|
||||
assert(k == 42);
|
||||
assert(x == k); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 4984: (138-145): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (456-470): CHC: Assertion violation happens here.
|
@ -0,0 +1,38 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
|
||||
function g() internal returns (uint) {
|
||||
x = 42;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
abstract contract Z is A {
|
||||
uint k;
|
||||
constructor(uint z) {
|
||||
k = z;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z, B {
|
||||
constructor() Z(g()) B(f()) {
|
||||
assert(x == 1);
|
||||
assert(k == 42);
|
||||
assert(x == k); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (449-463): CHC: Assertion violation happens here.
|
@ -0,0 +1,35 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint a) { x = a; }
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor(uint b) A(b) {
|
||||
}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
|
||||
function g() internal returns (uint) {
|
||||
x = 42;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
contract Z is B {
|
||||
constructor() B(f()) {
|
||||
}
|
||||
}
|
||||
|
||||
contract C is Z {
|
||||
constructor() {
|
||||
assert(x == 1);
|
||||
assert(x > 2); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (387-400): CHC: Assertion violation happens here.
|
@ -0,0 +1,22 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x;
|
||||
constructor(uint) {}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is A {
|
||||
constructor() A(f()) {
|
||||
assert(x == 1);
|
||||
assert(x == 0); // should fail
|
||||
assert(x > 2000); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (218-232): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (251-267): CHC: Assertion violation happens here.
|
@ -0,0 +1,24 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint public x = 42;
|
||||
constructor(uint) {}
|
||||
|
||||
function f() internal returns (uint) {
|
||||
x = x + 1;
|
||||
return x;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is A {
|
||||
constructor() A(f()) {
|
||||
assert(x == 42);
|
||||
assert(x == 0); // should fail
|
||||
assert(x == 1); // should fail
|
||||
assert(x > 2000); // should fail
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (224-238): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (257-271): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (290-306): CHC: Assertion violation happens here.
|
@ -0,0 +1,39 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
int x;
|
||||
constructor (int a) { x = a;}
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
int y;
|
||||
constructor(int a) A(-a) {
|
||||
if (a > 0) {
|
||||
y = 2;
|
||||
}
|
||||
else {
|
||||
y = 4;
|
||||
}
|
||||
}
|
||||
}
|
||||
|
||||
contract C is B {
|
||||
constructor(int a) B(a) {
|
||||
assert(y != 3); // should hold
|
||||
assert(y == 4); // should fail
|
||||
if (a > 0) {
|
||||
assert(x < 0 && y == 2); // should hold
|
||||
assert(x < 0 && y == 4); // should fail
|
||||
}
|
||||
else {
|
||||
assert(x >= 0 && y == 4); // should hold
|
||||
assert(x >= 0 && y == 2); // should fail
|
||||
assert(x > 0); // should fail
|
||||
}
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (280-294): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (372-395): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (472-496): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (516-529): CHC: Assertion violation happens here.
|
@ -23,9 +23,9 @@ contract A is B {
|
||||
|
||||
// ----
|
||||
// Warning 4984: (157-162): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (216-221): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (239-244): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (261-266): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (261-270): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (287-292): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (275-293): CHC: Assertion violation happens here.
|
||||
// Warning 4984: (216-221): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
|
@ -23,8 +23,8 @@ contract A is B {
|
||||
|
||||
// ----
|
||||
// Warning 4984: (157-163): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (217-222): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (240-245): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (262-268): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 4984: (285-290): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
// Warning 6328: (273-291): CHC: Assertion violation happens here.
|
||||
// Warning 4984: (217-222): CHC: Overflow (resulting value larger than 2**256 - 1) happens here.
|
||||
|
@ -0,0 +1,43 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
int x;
|
||||
constructor (int a) { x = a; }
|
||||
}
|
||||
|
||||
contract Z {
|
||||
int z;
|
||||
constructor(int _z) {
|
||||
z = _z;
|
||||
}
|
||||
}
|
||||
|
||||
contract B is A, Z {
|
||||
constructor(int b) A(b) Z(x) {
|
||||
assert(x == b);
|
||||
assert(z == 0);
|
||||
}
|
||||
}
|
||||
|
||||
contract F is Z, A {
|
||||
constructor(int b) Z(x) A(b) {
|
||||
assert(x == b);
|
||||
assert(z == 0);
|
||||
}
|
||||
}
|
||||
|
||||
contract C is B {
|
||||
constructor(int c) B(-c) {
|
||||
if (x > 0) {
|
||||
assert(c < 0); // should hold
|
||||
assert(c >= 0); // should fail
|
||||
}
|
||||
else {
|
||||
assert(c < 0); // should fail
|
||||
assert(c >= 0); // should hold
|
||||
}
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (436-450): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (483-496): CHC: Assertion violation happens here.
|
@ -5,11 +5,17 @@ contract A {
|
||||
}
|
||||
|
||||
contract B is A {
|
||||
constructor() { x = 2; }
|
||||
constructor() {
|
||||
assert(x == 1);
|
||||
x = 2;
|
||||
}
|
||||
}
|
||||
|
||||
contract C is A {
|
||||
constructor() { x = 3; }
|
||||
constructor() {
|
||||
assert(x == 1);
|
||||
x = 3;
|
||||
}
|
||||
}
|
||||
|
||||
contract D is B, C {
|
||||
@ -19,4 +25,5 @@ contract D is B, C {
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (214-228): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (167-181): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (256-270): CHC: Assertion violation happens here.
|
||||
|
@ -0,0 +1,19 @@
|
||||
pragma experimental SMTChecker;
|
||||
|
||||
contract A {
|
||||
uint x;
|
||||
constructor() {
|
||||
x = 42;
|
||||
}
|
||||
function f() public view returns(uint256) {
|
||||
return x;
|
||||
}
|
||||
}
|
||||
contract B is A {
|
||||
uint y = f();
|
||||
}
|
||||
contract C is B {
|
||||
function g() public view {
|
||||
assert(y == 42);
|
||||
}
|
||||
}
|
@ -15,10 +15,10 @@ contract C
|
||||
// Erasing knowledge about memory references should not
|
||||
// erase knowledge about storage references.
|
||||
assert(c[0] == 42);
|
||||
assert(a[0] == 2);
|
||||
// Removed because current Spacer seg faults in cex generation.
|
||||
//assert(a[0] == 2);
|
||||
// Removed because current Spacer seg faults in cex generation.
|
||||
//assert(b[0] == 1);
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (476-493): CHC: Assertion violation happens here.
|
||||
|
@ -21,7 +21,8 @@ contract C {
|
||||
// Removed because current Spacer seg faults in cex generation.
|
||||
//assert(s1.t.y == s2.t.y);
|
||||
s1.a[2] = 4;
|
||||
assert(s1.a[2] == s2.a[2]);
|
||||
// Removed because current Spacer seg faults in cex generation.
|
||||
//assert(s1.a[2] == s2.a[2]);
|
||||
s1.ts[3].y = 5;
|
||||
// Removed because current Spacer seg faults in cex generation.
|
||||
//assert(s1.ts[3].y == s2.ts[3].y);
|
||||
@ -30,5 +31,4 @@ contract C {
|
||||
}
|
||||
}
|
||||
// ----
|
||||
// Warning 6328: (456-482): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (629-667): CHC: Assertion violation happens here.
|
||||
// Warning 6328: (697-735): CHC: Assertion violation happens here.
|
||||
|
Loading…
Reference in New Issue
Block a user