mirror of
https://github.com/ethereum/solidity
synced 2023-10-03 13:03:40 +00:00
Control flow side effects on non-disambiguated source.
This commit is contained in:
parent
25c41546ee
commit
776ae466bc
@ -20,6 +20,7 @@
|
|||||||
|
|
||||||
#include <libyul/AST.h>
|
#include <libyul/AST.h>
|
||||||
#include <libyul/Dialect.h>
|
#include <libyul/Dialect.h>
|
||||||
|
#include <libyul/FunctionReferenceResolver.h>
|
||||||
|
|
||||||
#include <libsolutil/Common.h>
|
#include <libsolutil/Common.h>
|
||||||
#include <libsolutil/CommonData.h>
|
#include <libsolutil/CommonData.h>
|
||||||
@ -35,16 +36,15 @@ using namespace solidity::yul;
|
|||||||
|
|
||||||
ControlFlowBuilder::ControlFlowBuilder(Block const& _ast)
|
ControlFlowBuilder::ControlFlowBuilder(Block const& _ast)
|
||||||
{
|
{
|
||||||
for (auto const& statement: _ast.statements)
|
m_currentNode = newNode();
|
||||||
if (auto const* function = get_if<FunctionDefinition>(&statement))
|
(*this)(_ast);
|
||||||
(*this)(*function);
|
|
||||||
}
|
}
|
||||||
|
|
||||||
void ControlFlowBuilder::operator()(FunctionCall const& _functionCall)
|
void ControlFlowBuilder::operator()(FunctionCall const& _functionCall)
|
||||||
{
|
{
|
||||||
walkVector(_functionCall.arguments | ranges::views::reverse);
|
walkVector(_functionCall.arguments | ranges::views::reverse);
|
||||||
newConnectedNode();
|
newConnectedNode();
|
||||||
m_currentNode->functionCall = _functionCall.functionName.name;
|
m_currentNode->functionCall = &_functionCall;
|
||||||
}
|
}
|
||||||
|
|
||||||
void ControlFlowBuilder::operator()(If const& _if)
|
void ControlFlowBuilder::operator()(If const& _if)
|
||||||
@ -78,7 +78,9 @@ void ControlFlowBuilder::operator()(Switch const& _switch)
|
|||||||
void ControlFlowBuilder::operator()(FunctionDefinition const& _function)
|
void ControlFlowBuilder::operator()(FunctionDefinition const& _function)
|
||||||
{
|
{
|
||||||
ScopedSaveAndRestore currentNode(m_currentNode, nullptr);
|
ScopedSaveAndRestore currentNode(m_currentNode, nullptr);
|
||||||
yulAssert(!m_leave && !m_break && !m_continue, "Function hoister has not been used.");
|
ScopedSaveAndRestore leave(m_leave, nullptr);
|
||||||
|
ScopedSaveAndRestore _break(m_break, nullptr);
|
||||||
|
ScopedSaveAndRestore _continue(m_continue, nullptr);
|
||||||
|
|
||||||
FunctionFlow flow;
|
FunctionFlow flow;
|
||||||
flow.exit = newNode();
|
flow.exit = newNode();
|
||||||
@ -90,7 +92,7 @@ void ControlFlowBuilder::operator()(FunctionDefinition const& _function)
|
|||||||
|
|
||||||
m_currentNode->successors.emplace_back(flow.exit);
|
m_currentNode->successors.emplace_back(flow.exit);
|
||||||
|
|
||||||
m_functionFlows[_function.name] = move(flow);
|
m_functionFlows[&_function] = move(flow);
|
||||||
|
|
||||||
m_leave = nullptr;
|
m_leave = nullptr;
|
||||||
}
|
}
|
||||||
@ -164,14 +166,17 @@ ControlFlowSideEffectsCollector::ControlFlowSideEffectsCollector(
|
|||||||
Block const& _ast
|
Block const& _ast
|
||||||
):
|
):
|
||||||
m_dialect(_dialect),
|
m_dialect(_dialect),
|
||||||
m_cfgBuilder(_ast)
|
m_cfgBuilder(_ast),
|
||||||
|
m_functionReferences(FunctionReferenceResolver{_ast}.references())
|
||||||
{
|
{
|
||||||
for (auto&& [name, flow]: m_cfgBuilder.functionFlows())
|
for (auto&& [function, flow]: m_cfgBuilder.functionFlows())
|
||||||
{
|
{
|
||||||
yulAssert(!flow.entry->functionCall);
|
yulAssert(!flow.entry->functionCall);
|
||||||
m_processedNodes[name] = {};
|
yulAssert(function);
|
||||||
m_pendingNodes[name].push_front(flow.entry);
|
m_processedNodes[function] = {};
|
||||||
m_functionSideEffects[name] = {false, false, false};
|
m_pendingNodes[function].push_front(flow.entry);
|
||||||
|
m_functionSideEffects[function] = {false, false, false};
|
||||||
|
m_functionCalls[function] = {};
|
||||||
}
|
}
|
||||||
|
|
||||||
// Process functions while we have progress. For now, we are only interested
|
// Process functions while we have progress. For now, we are only interested
|
||||||
@ -180,8 +185,8 @@ ControlFlowSideEffectsCollector::ControlFlowSideEffectsCollector(
|
|||||||
while (progress)
|
while (progress)
|
||||||
{
|
{
|
||||||
progress = false;
|
progress = false;
|
||||||
for (auto const& functionName: m_pendingNodes | ranges::views::keys)
|
for (FunctionDefinition const* function: m_pendingNodes | ranges::views::keys)
|
||||||
if (processFunction(functionName))
|
if (processFunction(*function))
|
||||||
progress = true;
|
progress = true;
|
||||||
}
|
}
|
||||||
|
|
||||||
@ -190,57 +195,64 @@ ControlFlowSideEffectsCollector::ControlFlowSideEffectsCollector(
|
|||||||
// If we have not set `canContinue` by now, the function's exit
|
// If we have not set `canContinue` by now, the function's exit
|
||||||
// is not reachable.
|
// is not reachable.
|
||||||
|
|
||||||
for (auto&& [functionName, calls]: m_functionCalls)
|
// Now it is sufficient to handle the reachable function calls (`m_functionCalls`),
|
||||||
|
// we do not have to consider the control-flow graph anymore.
|
||||||
|
for (auto&& [function, calls]: m_functionCalls)
|
||||||
{
|
{
|
||||||
ControlFlowSideEffects& sideEffects = m_functionSideEffects[functionName];
|
yulAssert(function);
|
||||||
auto _visit = [&, visited = std::set<YulString>{}](YulString _function, auto&& _recurse) mutable {
|
ControlFlowSideEffects& functionSideEffects = m_functionSideEffects[function];
|
||||||
if (sideEffects.canTerminate && sideEffects.canRevert)
|
auto _visit = [&, visited = std::set<FunctionDefinition const*>{}](FunctionDefinition const& _function, auto&& _recurse) mutable {
|
||||||
|
// Worst side-effects already, stop searching.
|
||||||
|
if (functionSideEffects.canTerminate && functionSideEffects.canRevert)
|
||||||
return;
|
return;
|
||||||
if (!visited.insert(_function).second)
|
if (!visited.insert(&_function).second)
|
||||||
return;
|
return;
|
||||||
|
|
||||||
ControlFlowSideEffects const* calledSideEffects = nullptr;
|
for (FunctionCall const* call: m_functionCalls.at(&_function))
|
||||||
if (BuiltinFunction const* f = _dialect.builtin(_function))
|
{
|
||||||
calledSideEffects = &f->controlFlowSideEffects;
|
ControlFlowSideEffects const& calledSideEffects = sideEffects(*call);
|
||||||
else
|
if (calledSideEffects.canTerminate)
|
||||||
calledSideEffects = &m_functionSideEffects.at(_function);
|
functionSideEffects.canTerminate = true;
|
||||||
|
if (calledSideEffects.canRevert)
|
||||||
|
functionSideEffects.canRevert = true;
|
||||||
|
|
||||||
if (calledSideEffects->canTerminate)
|
if (m_functionReferences.count(call))
|
||||||
sideEffects.canTerminate = true;
|
_recurse(*m_functionReferences.at(call), _recurse);
|
||||||
if (calledSideEffects->canRevert)
|
}
|
||||||
sideEffects.canRevert = true;
|
|
||||||
|
|
||||||
set<YulString> emptySet;
|
|
||||||
for (YulString callee: util::valueOrDefault(m_functionCalls, _function, emptySet))
|
|
||||||
_recurse(callee, _recurse);
|
|
||||||
};
|
};
|
||||||
for (auto const& call: calls)
|
_visit(*function, _visit);
|
||||||
_visit(call, _visit);
|
|
||||||
}
|
}
|
||||||
|
|
||||||
}
|
}
|
||||||
|
|
||||||
bool ControlFlowSideEffectsCollector::processFunction(YulString _name)
|
map<YulString, ControlFlowSideEffects> ControlFlowSideEffectsCollector::functionSideEffectsNamed() const
|
||||||
|
{
|
||||||
|
map<YulString, ControlFlowSideEffects> result;
|
||||||
|
for (auto&& [function, sideEffects]: m_functionSideEffects)
|
||||||
|
yulAssert(result.insert({function->name, sideEffects}).second);
|
||||||
|
return result;
|
||||||
|
}
|
||||||
|
|
||||||
|
bool ControlFlowSideEffectsCollector::processFunction(FunctionDefinition const& _function)
|
||||||
{
|
{
|
||||||
bool progress = false;
|
bool progress = false;
|
||||||
while (ControlFlowNode const* node = nextProcessableNode(_name))
|
while (ControlFlowNode const* node = nextProcessableNode(_function))
|
||||||
{
|
{
|
||||||
if (node == m_cfgBuilder.functionFlows().at(_name).exit)
|
if (node == m_cfgBuilder.functionFlows().at(&_function).exit)
|
||||||
{
|
{
|
||||||
m_functionSideEffects[_name].canContinue = true;
|
m_functionSideEffects[&_function].canContinue = true;
|
||||||
return true;
|
return true;
|
||||||
}
|
}
|
||||||
for (ControlFlowNode const* s: node->successors)
|
for (ControlFlowNode const* s: node->successors)
|
||||||
recordReachabilityAndQueue(_name, s);
|
recordReachabilityAndQueue(_function, s);
|
||||||
|
|
||||||
progress = true;
|
progress = true;
|
||||||
}
|
}
|
||||||
return progress;
|
return progress;
|
||||||
}
|
}
|
||||||
|
|
||||||
ControlFlowNode const* ControlFlowSideEffectsCollector::nextProcessableNode(YulString _functionName)
|
ControlFlowNode const* ControlFlowSideEffectsCollector::nextProcessableNode(FunctionDefinition const& _function)
|
||||||
{
|
{
|
||||||
std::list<ControlFlowNode const*>& nodes = m_pendingNodes[_functionName];
|
std::list<ControlFlowNode const*>& nodes = m_pendingNodes[&_function];
|
||||||
auto it = ranges::find_if(nodes, [this](ControlFlowNode const* _node) {
|
auto it = ranges::find_if(nodes, [this](ControlFlowNode const* _node) {
|
||||||
return !_node->functionCall || sideEffects(*_node->functionCall).canContinue;
|
return !_node->functionCall || sideEffects(*_node->functionCall).canContinue;
|
||||||
});
|
});
|
||||||
@ -252,22 +264,22 @@ ControlFlowNode const* ControlFlowSideEffectsCollector::nextProcessableNode(YulS
|
|||||||
return node;
|
return node;
|
||||||
}
|
}
|
||||||
|
|
||||||
ControlFlowSideEffects const& ControlFlowSideEffectsCollector::sideEffects(YulString _functionName) const
|
ControlFlowSideEffects const& ControlFlowSideEffectsCollector::sideEffects(FunctionCall const& _call) const
|
||||||
{
|
{
|
||||||
if (auto const* builtin = m_dialect.builtin(_functionName))
|
if (auto const* builtin = m_dialect.builtin(_call.functionName.name))
|
||||||
return builtin->controlFlowSideEffects;
|
return builtin->controlFlowSideEffects;
|
||||||
else
|
else
|
||||||
return m_functionSideEffects.at(_functionName);
|
return m_functionSideEffects.at(m_functionReferences.at(&_call));
|
||||||
}
|
}
|
||||||
|
|
||||||
void ControlFlowSideEffectsCollector::recordReachabilityAndQueue(
|
void ControlFlowSideEffectsCollector::recordReachabilityAndQueue(
|
||||||
YulString _functionName,
|
FunctionDefinition const& _function,
|
||||||
ControlFlowNode const* _node
|
ControlFlowNode const* _node
|
||||||
)
|
)
|
||||||
{
|
{
|
||||||
if (_node->functionCall)
|
if (_node->functionCall)
|
||||||
m_functionCalls[_functionName].insert(*_node->functionCall);
|
m_functionCalls[&_function].insert(_node->functionCall);
|
||||||
if (m_processedNodes[_functionName].insert(_node).second)
|
if (m_processedNodes[&_function].insert(_node).second)
|
||||||
m_pendingNodes.at(_functionName).push_front(_node);
|
m_pendingNodes.at(&_function).push_front(_node);
|
||||||
}
|
}
|
||||||
|
|
||||||
|
@ -34,8 +34,8 @@ struct Dialect;
|
|||||||
struct ControlFlowNode
|
struct ControlFlowNode
|
||||||
{
|
{
|
||||||
std::vector<ControlFlowNode const*> successors;
|
std::vector<ControlFlowNode const*> successors;
|
||||||
/// Name of the called function if the node calls a function.
|
/// Function call AST node, if present.
|
||||||
std::optional<YulString> functionCall;
|
FunctionCall const* functionCall = nullptr;
|
||||||
};
|
};
|
||||||
|
|
||||||
/**
|
/**
|
||||||
@ -56,7 +56,7 @@ public:
|
|||||||
/// Computes the control-flows of all function defined in the block.
|
/// Computes the control-flows of all function defined in the block.
|
||||||
/// Assumes the functions are hoisted to the topmost block.
|
/// Assumes the functions are hoisted to the topmost block.
|
||||||
explicit ControlFlowBuilder(Block const& _ast);
|
explicit ControlFlowBuilder(Block const& _ast);
|
||||||
std::map<YulString, FunctionFlow> const& functionFlows() const { return m_functionFlows; }
|
std::map<FunctionDefinition const*, FunctionFlow> const& functionFlows() const { return m_functionFlows; }
|
||||||
|
|
||||||
private:
|
private:
|
||||||
using ASTWalker::operator();
|
using ASTWalker::operator();
|
||||||
@ -79,12 +79,14 @@ private:
|
|||||||
ControlFlowNode const* m_break = nullptr;
|
ControlFlowNode const* m_break = nullptr;
|
||||||
ControlFlowNode const* m_continue = nullptr;
|
ControlFlowNode const* m_continue = nullptr;
|
||||||
|
|
||||||
std::map<YulString, FunctionFlow> m_functionFlows;
|
std::map<FunctionDefinition const*, FunctionFlow> m_functionFlows;
|
||||||
};
|
};
|
||||||
|
|
||||||
|
|
||||||
/**
|
/**
|
||||||
* Requires: Disambiguator, Function Hoister.
|
* Computes control-flow side-effects for user-defined functions.
|
||||||
|
* Source does not have to be disambiguated, unless you want the side-effects
|
||||||
|
* based on function names.
|
||||||
*/
|
*/
|
||||||
class ControlFlowSideEffectsCollector
|
class ControlFlowSideEffectsCollector
|
||||||
{
|
{
|
||||||
@ -94,36 +96,43 @@ public:
|
|||||||
Block const& _ast
|
Block const& _ast
|
||||||
);
|
);
|
||||||
|
|
||||||
std::map<YulString, ControlFlowSideEffects> const& functionSideEffects() const
|
std::map<FunctionDefinition const*, ControlFlowSideEffects> const& functionSideEffects() const
|
||||||
{
|
{
|
||||||
return m_functionSideEffects;
|
return m_functionSideEffects;
|
||||||
}
|
}
|
||||||
|
/// Returns the side effects by function name, requires unique function names.
|
||||||
|
std::map<YulString, ControlFlowSideEffects> functionSideEffectsNamed() const;
|
||||||
private:
|
private:
|
||||||
|
|
||||||
/// @returns false if nothing could be processed.
|
/// @returns false if nothing could be processed.
|
||||||
bool processFunction(YulString _name);
|
bool processFunction(FunctionDefinition const& _function);
|
||||||
|
|
||||||
/// @returns the next pending node of the function that is not
|
/// @returns the next pending node of the function that is not
|
||||||
/// a function call to a function that might not continue.
|
/// a function call to a function that might not continue.
|
||||||
/// De-queues the node or returns nullptr if no such node is found.
|
/// De-queues the node or returns nullptr if no such node is found.
|
||||||
ControlFlowNode const* nextProcessableNode(YulString _functionName);
|
ControlFlowNode const* nextProcessableNode(FunctionDefinition const& _function);
|
||||||
|
|
||||||
/// @returns the side-effects of either a builtin call or a user defined function
|
/// @returns the side-effects of either a builtin call or a user defined function
|
||||||
/// call (as far as already computed).
|
/// call (as far as already computed).
|
||||||
ControlFlowSideEffects const& sideEffects(YulString _functionName) const;
|
ControlFlowSideEffects const& sideEffects(FunctionCall const& _call) const;
|
||||||
|
|
||||||
/// Queues the given node to be processed (if not already visited)
|
/// Queues the given node to be processed (if not already visited)
|
||||||
/// and if it is a function call, records that `_functionName` calls
|
/// and if it is a function call, records that `_functionName` calls
|
||||||
/// `*_node->functionCall`.
|
/// `*_node->functionCall`.
|
||||||
void recordReachabilityAndQueue(YulString _functionName, ControlFlowNode const* _node);
|
void recordReachabilityAndQueue(FunctionDefinition const& _function, ControlFlowNode const* _node);
|
||||||
|
|
||||||
Dialect const& m_dialect;
|
Dialect const& m_dialect;
|
||||||
ControlFlowBuilder m_cfgBuilder;
|
ControlFlowBuilder m_cfgBuilder;
|
||||||
std::map<YulString, ControlFlowSideEffects> m_functionSideEffects;
|
/// Function references, but only for calls to user-defined functions.
|
||||||
std::map<YulString, std::list<ControlFlowNode const*>> m_pendingNodes;
|
std::map<FunctionCall const*, FunctionDefinition const*> m_functionReferences;
|
||||||
std::map<YulString, std::set<ControlFlowNode const*>> m_processedNodes;
|
/// Side effects of user-defined functions, is being constructod.
|
||||||
/// `x` is in `m_functionCalls[y]` if a direct call to `x` is reachable inside `y`
|
std::map<FunctionDefinition const*, ControlFlowSideEffects> m_functionSideEffects;
|
||||||
std::map<YulString, std::set<YulString>> m_functionCalls;
|
/// Control flow nodes still to process, per function.
|
||||||
|
std::map<FunctionDefinition const*, std::list<ControlFlowNode const*>> m_pendingNodes;
|
||||||
|
/// Control flow nodes already processed, per function.
|
||||||
|
std::map<FunctionDefinition const*, std::set<ControlFlowNode const*>> m_processedNodes;
|
||||||
|
/// Set of reachable function calls nodes in each function (including calls to builtins).
|
||||||
|
std::map<FunctionDefinition const*, std::set<FunctionCall const*>> m_functionCalls;
|
||||||
};
|
};
|
||||||
|
|
||||||
|
|
||||||
|
@ -29,8 +29,10 @@ using namespace solidity::util;
|
|||||||
|
|
||||||
void ConditionalSimplifier::run(OptimiserStepContext& _context, Block& _ast)
|
void ConditionalSimplifier::run(OptimiserStepContext& _context, Block& _ast)
|
||||||
{
|
{
|
||||||
ControlFlowSideEffectsCollector sideEffects(_context.dialect, _ast);
|
ConditionalSimplifier{
|
||||||
ConditionalSimplifier{_context.dialect, sideEffects.functionSideEffects()}(_ast);
|
_context.dialect,
|
||||||
|
ControlFlowSideEffectsCollector{_context.dialect, _ast}.functionSideEffectsNamed()
|
||||||
|
}(_ast);
|
||||||
}
|
}
|
||||||
|
|
||||||
void ConditionalSimplifier::operator()(Switch& _switch)
|
void ConditionalSimplifier::operator()(Switch& _switch)
|
||||||
|
@ -62,12 +62,12 @@ public:
|
|||||||
private:
|
private:
|
||||||
explicit ConditionalSimplifier(
|
explicit ConditionalSimplifier(
|
||||||
Dialect const& _dialect,
|
Dialect const& _dialect,
|
||||||
std::map<YulString, ControlFlowSideEffects> const& _sideEffects
|
std::map<YulString, ControlFlowSideEffects> _sideEffects
|
||||||
):
|
):
|
||||||
m_dialect(_dialect), m_functionSideEffects(_sideEffects)
|
m_dialect(_dialect), m_functionSideEffects(move(_sideEffects))
|
||||||
{}
|
{}
|
||||||
Dialect const& m_dialect;
|
Dialect const& m_dialect;
|
||||||
std::map<YulString, ControlFlowSideEffects> const& m_functionSideEffects;
|
std::map<YulString, ControlFlowSideEffects> m_functionSideEffects;
|
||||||
};
|
};
|
||||||
|
|
||||||
}
|
}
|
||||||
|
@ -30,8 +30,10 @@ using namespace solidity::util;
|
|||||||
|
|
||||||
void ConditionalUnsimplifier::run(OptimiserStepContext& _context, Block& _ast)
|
void ConditionalUnsimplifier::run(OptimiserStepContext& _context, Block& _ast)
|
||||||
{
|
{
|
||||||
ControlFlowSideEffectsCollector sideEffects(_context.dialect, _ast);
|
ConditionalUnsimplifier{
|
||||||
ConditionalUnsimplifier{_context.dialect, sideEffects.functionSideEffects()}(_ast);
|
_context.dialect,
|
||||||
|
ControlFlowSideEffectsCollector{_context.dialect, _ast}.functionSideEffectsNamed()
|
||||||
|
}(_ast);
|
||||||
}
|
}
|
||||||
|
|
||||||
void ConditionalUnsimplifier::operator()(Switch& _switch)
|
void ConditionalUnsimplifier::operator()(Switch& _switch)
|
||||||
|
@ -40,7 +40,7 @@ void DeadCodeEliminator::run(OptimiserStepContext& _context, Block& _ast)
|
|||||||
ControlFlowSideEffectsCollector sideEffects(_context.dialect, _ast);
|
ControlFlowSideEffectsCollector sideEffects(_context.dialect, _ast);
|
||||||
DeadCodeEliminator{
|
DeadCodeEliminator{
|
||||||
_context.dialect,
|
_context.dialect,
|
||||||
sideEffects.functionSideEffects()
|
sideEffects.functionSideEffectsNamed()
|
||||||
}(_ast);
|
}(_ast);
|
||||||
}
|
}
|
||||||
|
|
||||||
|
@ -23,6 +23,7 @@
|
|||||||
|
|
||||||
#include <libyul/optimiser/ASTWalker.h>
|
#include <libyul/optimiser/ASTWalker.h>
|
||||||
#include <libyul/YulString.h>
|
#include <libyul/YulString.h>
|
||||||
|
#include <libyul/ControlFlowSideEffects.h>
|
||||||
|
|
||||||
#include <map>
|
#include <map>
|
||||||
#include <set>
|
#include <set>
|
||||||
@ -31,7 +32,6 @@ namespace solidity::yul
|
|||||||
{
|
{
|
||||||
struct Dialect;
|
struct Dialect;
|
||||||
struct OptimiserStepContext;
|
struct OptimiserStepContext;
|
||||||
struct ControlFlowSideEffects;
|
|
||||||
|
|
||||||
/**
|
/**
|
||||||
* Optimisation stage that removes unreachable code
|
* Optimisation stage that removes unreachable code
|
||||||
@ -62,11 +62,11 @@ public:
|
|||||||
private:
|
private:
|
||||||
DeadCodeEliminator(
|
DeadCodeEliminator(
|
||||||
Dialect const& _dialect,
|
Dialect const& _dialect,
|
||||||
std::map<YulString, ControlFlowSideEffects> const& _sideEffects
|
std::map<YulString, ControlFlowSideEffects> _sideEffects
|
||||||
): m_dialect(_dialect), m_functionSideEffects(_sideEffects) {}
|
): m_dialect(_dialect), m_functionSideEffects(move(_sideEffects)) {}
|
||||||
|
|
||||||
Dialect const& m_dialect;
|
Dialect const& m_dialect;
|
||||||
std::map<YulString, ControlFlowSideEffects> const& m_functionSideEffects;
|
std::map<YulString, ControlFlowSideEffects> m_functionSideEffects;
|
||||||
};
|
};
|
||||||
|
|
||||||
}
|
}
|
||||||
|
@ -22,6 +22,7 @@
|
|||||||
#include <test/libyul/Common.h>
|
#include <test/libyul/Common.h>
|
||||||
|
|
||||||
#include <libyul/Object.h>
|
#include <libyul/Object.h>
|
||||||
|
#include <libyul/AST.h>
|
||||||
#include <libyul/ControlFlowSideEffects.h>
|
#include <libyul/ControlFlowSideEffects.h>
|
||||||
#include <libyul/ControlFlowSideEffectsCollector.h>
|
#include <libyul/ControlFlowSideEffectsCollector.h>
|
||||||
#include <libyul/backends/evm/EVMDialect.h>
|
#include <libyul/backends/evm/EVMDialect.h>
|
||||||
@ -61,19 +62,15 @@ TestCase::TestResult ControlFlowSideEffectsTest::run(ostream& _stream, string co
|
|||||||
if (!obj.code)
|
if (!obj.code)
|
||||||
BOOST_THROW_EXCEPTION(runtime_error("Parsing input failed."));
|
BOOST_THROW_EXCEPTION(runtime_error("Parsing input failed."));
|
||||||
|
|
||||||
std::map<YulString, ControlFlowSideEffects> sideEffects =
|
ControlFlowSideEffectsCollector sideEffects(
|
||||||
ControlFlowSideEffectsCollector(
|
EVMDialect::strictAssemblyForEVMObjects(langutil::EVMVersion()),
|
||||||
EVMDialect::strictAssemblyForEVMObjects(langutil::EVMVersion()),
|
*obj.code
|
||||||
*obj.code
|
);
|
||||||
).functionSideEffects();
|
|
||||||
|
|
||||||
std::map<std::string, std::string> controlFlowSideEffectsStr;
|
|
||||||
for (auto&& [fun, effects]: sideEffects)
|
|
||||||
controlFlowSideEffectsStr[fun.str()] = toString(effects);
|
|
||||||
|
|
||||||
m_obtainedResult.clear();
|
m_obtainedResult.clear();
|
||||||
for (auto&& [functionName, effect]: controlFlowSideEffectsStr)
|
forEach<FunctionDefinition const>(*obj.code, [&](FunctionDefinition const& _fun) {
|
||||||
m_obtainedResult += functionName + (effect.empty() ? ":" : ": " + effect) + "\n";
|
string effectStr = toString(sideEffects.functionSideEffects().at(&_fun));
|
||||||
|
m_obtainedResult += _fun.name.str() + (effectStr.empty() ? ":" : ": " + effectStr) + "\n";
|
||||||
|
});
|
||||||
|
|
||||||
return checkResult(_stream, _linePrefix, _formatted);
|
return checkResult(_stream, _linePrefix, _formatted);
|
||||||
}
|
}
|
||||||
|
19
test/libyul/controlFlowSideEffects/nondisambiguated.yul
Normal file
19
test/libyul/controlFlowSideEffects/nondisambiguated.yul
Normal file
@ -0,0 +1,19 @@
|
|||||||
|
{
|
||||||
|
function a() {
|
||||||
|
{
|
||||||
|
function b() { if calldataloda(0) { return(0, 0) } }
|
||||||
|
b()
|
||||||
|
}
|
||||||
|
{
|
||||||
|
function b() { revert(0, 0) }
|
||||||
|
b()
|
||||||
|
}
|
||||||
|
}
|
||||||
|
function b() {
|
||||||
|
leave
|
||||||
|
revert(0, 0)
|
||||||
|
}
|
||||||
|
}
|
||||||
|
// ----
|
||||||
|
// a: can revert
|
||||||
|
// b: can continue
|
Loading…
Reference in New Issue
Block a user