mirror of
https://github.com/ethereum/solidity
synced 2023-10-03 13:03:40 +00:00
Merge pull request #10445 from ethereum/modifiers
[Sol->Yul] Implement function modifiers.
This commit is contained in:
commit
c4ade1753e
@ -34,6 +34,27 @@ Consequently, if the padding space within a struct is used to store data (e.g. i
|
||||
|
||||
We have the same behavior for implicit delete, for example when array of structs is shortened.
|
||||
|
||||
* Function modifiers are implemented in a slightly different way regarding function parameters.
|
||||
This especially has an effect if the placeholder ``_;`` is evaluated multiple times in a modifier.
|
||||
In the old code generator, each function parameter has a fixed slot on the stack. If the function
|
||||
is run multiple times because ``_;`` is used multiple times or used in a loop, then a change to the
|
||||
function parameter's value is visible in the next execution of the function.
|
||||
The new code generator implements modifiers using actual functions and passes function parameters on.
|
||||
This means that multiple executions of a function will get the same values for the parameters.
|
||||
|
||||
::
|
||||
// SPDX-License-Identifier: GPL-3.0
|
||||
pragma solidity >=0.7.0;
|
||||
contract C {
|
||||
function f(uint a) public pure mod() returns (uint r) {
|
||||
r = a++;
|
||||
}
|
||||
modifier mod() { _; _; }
|
||||
}
|
||||
|
||||
If you execute ``f(0)`` in the old code generator, it will return ``2``, while
|
||||
it will return ``1`` when using the new code generator.
|
||||
|
||||
* The order of contract initialization has changed in case of inheritance.
|
||||
|
||||
The order used to be:
|
||||
|
@ -32,10 +32,9 @@ YulArity YulArity::fromType(FunctionType const& _functionType)
|
||||
TupleType(_functionType.returnParameterTypes()).sizeOnStack()
|
||||
};
|
||||
}
|
||||
|
||||
string IRNames::function(FunctionDefinition const& _function)
|
||||
{
|
||||
// @TODO previously, we had to distinguish creation context and runtime context,
|
||||
// but since we do not work with jump positions anymore, this should not be a problem, right?
|
||||
return "fun_" + _function.name() + "_" + to_string(_function.id());
|
||||
}
|
||||
|
||||
@ -44,6 +43,21 @@ string IRNames::function(VariableDeclaration const& _varDecl)
|
||||
return "getter_fun_" + _varDecl.name() + "_" + to_string(_varDecl.id());
|
||||
}
|
||||
|
||||
string IRNames::modifierInvocation(ModifierInvocation const& _modifierInvocation)
|
||||
{
|
||||
// This uses the ID of the modifier invocation because it has to be unique
|
||||
// for each invocation.
|
||||
solAssert(!_modifierInvocation.name().path().empty(), "");
|
||||
string const& modifierName = _modifierInvocation.name().path().back();
|
||||
solAssert(!modifierName.empty(), "");
|
||||
return "modifier_" + modifierName + "_" + to_string(_modifierInvocation.id());
|
||||
}
|
||||
|
||||
string IRNames::functionWithModifierInner(FunctionDefinition const& _function)
|
||||
{
|
||||
return "fun_" + _function.name() + "_" + to_string(_function.id()) + "_inner";
|
||||
}
|
||||
|
||||
string IRNames::creationObject(ContractDefinition const& _contract)
|
||||
{
|
||||
return _contract.name() + "_" + toString(_contract.id());
|
||||
|
@ -49,6 +49,8 @@ struct IRNames
|
||||
{
|
||||
static std::string function(FunctionDefinition const& _function);
|
||||
static std::string function(VariableDeclaration const& _varDecl);
|
||||
static std::string modifierInvocation(ModifierInvocation const& _modifierInvocation);
|
||||
static std::string functionWithModifierInner(FunctionDefinition const& _function);
|
||||
static std::string creationObject(ContractDefinition const& _contract);
|
||||
static std::string runtimeObject(ContractDefinition const& _contract);
|
||||
static std::string internalDispatch(YulArity const& _arity);
|
||||
|
@ -80,6 +80,11 @@ IRVariable const& IRGenerationContext::localVariable(VariableDeclaration const&
|
||||
return m_localVariables.at(&_varDecl);
|
||||
}
|
||||
|
||||
void IRGenerationContext::resetLocalVariables()
|
||||
{
|
||||
m_localVariables.clear();
|
||||
}
|
||||
|
||||
void IRGenerationContext::registerImmutableVariable(VariableDeclaration const& _variable)
|
||||
{
|
||||
solAssert(_variable.immutable(), "Attempted to register a non-immutable variable as immutable.");
|
||||
|
@ -97,6 +97,7 @@ public:
|
||||
IRVariable const& addLocalVariable(VariableDeclaration const& _varDecl);
|
||||
bool isLocalVariable(VariableDeclaration const& _varDecl) const { return m_localVariables.count(&_varDecl); }
|
||||
IRVariable const& localVariable(VariableDeclaration const& _varDecl);
|
||||
void resetLocalVariables();
|
||||
|
||||
/// Registers an immutable variable of the contract.
|
||||
/// Should only be called at construction time.
|
||||
|
@ -249,10 +249,10 @@ string IRGenerator::generateFunction(FunctionDefinition const& _function)
|
||||
{
|
||||
string functionName = IRNames::function(_function);
|
||||
return m_context.functionCollector().createFunction(functionName, [&]() {
|
||||
solUnimplementedAssert(_function.modifiers().empty(), "Modifiers not implemented yet.");
|
||||
m_context.resetLocalVariables();
|
||||
Whiskers t(R"(
|
||||
function <functionName>(<params>)<?+retParams> -> <retParams></+retParams> {
|
||||
<initReturnVariables>
|
||||
<retInit>
|
||||
<body>
|
||||
}
|
||||
)");
|
||||
@ -268,8 +268,137 @@ string IRGenerator::generateFunction(FunctionDefinition const& _function)
|
||||
retParams += m_context.addLocalVariable(*varDecl).stackSlots();
|
||||
retInit += generateInitialAssignment(*varDecl);
|
||||
}
|
||||
|
||||
t("retParams", joinHumanReadable(retParams));
|
||||
t("initReturnVariables", retInit);
|
||||
t("retInit", retInit);
|
||||
|
||||
if (_function.modifiers().empty())
|
||||
t("body", generate(_function.body()));
|
||||
else
|
||||
{
|
||||
for (size_t i = 0; i < _function.modifiers().size(); ++i)
|
||||
{
|
||||
ModifierInvocation const& modifier = *_function.modifiers().at(i);
|
||||
string next =
|
||||
i + 1 < _function.modifiers().size() ?
|
||||
IRNames::modifierInvocation(*_function.modifiers().at(i + 1)) :
|
||||
IRNames::functionWithModifierInner(_function);
|
||||
generateModifier(modifier, _function, next);
|
||||
}
|
||||
t("body",
|
||||
(retParams.empty() ? string{} : joinHumanReadable(retParams) + " := ") +
|
||||
IRNames::modifierInvocation(*_function.modifiers().at(0)) +
|
||||
"(" +
|
||||
joinHumanReadable(retParams + params) +
|
||||
")"
|
||||
);
|
||||
// Now generate the actual inner function.
|
||||
generateFunctionWithModifierInner(_function);
|
||||
}
|
||||
return t.render();
|
||||
});
|
||||
}
|
||||
|
||||
string IRGenerator::generateModifier(
|
||||
ModifierInvocation const& _modifierInvocation,
|
||||
FunctionDefinition const& _function,
|
||||
string const& _nextFunction
|
||||
)
|
||||
{
|
||||
string functionName = IRNames::modifierInvocation(_modifierInvocation);
|
||||
return m_context.functionCollector().createFunction(functionName, [&]() {
|
||||
m_context.resetLocalVariables();
|
||||
Whiskers t(R"(
|
||||
function <functionName>(<params>)<?+retParams> -> <retParams></+retParams> {
|
||||
<assignRetParams>
|
||||
<evalArgs>
|
||||
<body>
|
||||
}
|
||||
)");
|
||||
t("functionName", functionName);
|
||||
vector<string> retParamsIn;
|
||||
for (auto const& varDecl: _function.returnParameters())
|
||||
retParamsIn += IRVariable(*varDecl).stackSlots();
|
||||
vector<string> params = retParamsIn;
|
||||
for (auto const& varDecl: _function.parameters())
|
||||
params += m_context.addLocalVariable(*varDecl).stackSlots();
|
||||
t("params", joinHumanReadable(params));
|
||||
vector<string> retParams;
|
||||
string assignRetParams;
|
||||
for (size_t i = 0; i < retParamsIn.size(); ++i)
|
||||
{
|
||||
retParams.emplace_back(m_context.newYulVariable());
|
||||
assignRetParams += retParams.back() + " := " + retParamsIn[i] + "\n";
|
||||
}
|
||||
t("retParams", joinHumanReadable(retParams));
|
||||
t("assignRetParams", assignRetParams);
|
||||
|
||||
solAssert(*_modifierInvocation.name().annotation().requiredLookup == VirtualLookup::Virtual, "");
|
||||
|
||||
ModifierDefinition const& modifier = dynamic_cast<ModifierDefinition const&>(
|
||||
*_modifierInvocation.name().annotation().referencedDeclaration
|
||||
).resolveVirtual(m_context.mostDerivedContract());
|
||||
|
||||
solAssert(
|
||||
modifier.parameters().empty() ==
|
||||
(!_modifierInvocation.arguments() || _modifierInvocation.arguments()->empty()),
|
||||
""
|
||||
);
|
||||
IRGeneratorForStatements expressionEvaluator(m_context, m_utils);
|
||||
if (_modifierInvocation.arguments())
|
||||
for (size_t i = 0; i < _modifierInvocation.arguments()->size(); i++)
|
||||
{
|
||||
IRVariable argument = expressionEvaluator.evaluateExpression(
|
||||
*_modifierInvocation.arguments()->at(i),
|
||||
*modifier.parameters()[i]->annotation().type
|
||||
);
|
||||
expressionEvaluator.define(
|
||||
m_context.addLocalVariable(*modifier.parameters()[i]),
|
||||
argument
|
||||
);
|
||||
}
|
||||
|
||||
t("evalArgs", expressionEvaluator.code());
|
||||
IRGeneratorForStatements generator(m_context, m_utils, [&]() {
|
||||
string ret = joinHumanReadable(retParams);
|
||||
return
|
||||
(ret.empty() ? "" : ret + " := ") +
|
||||
_nextFunction + "(" + joinHumanReadable(params) + ")\n";
|
||||
});
|
||||
generator.generate(modifier.body());
|
||||
t("body", generator.code());
|
||||
return t.render();
|
||||
});
|
||||
}
|
||||
|
||||
string IRGenerator::generateFunctionWithModifierInner(FunctionDefinition const& _function)
|
||||
{
|
||||
string functionName = IRNames::functionWithModifierInner(_function);
|
||||
return m_context.functionCollector().createFunction(functionName, [&]() {
|
||||
m_context.resetLocalVariables();
|
||||
Whiskers t(R"(
|
||||
function <functionName>(<params>)<?+retParams> -> <retParams></+retParams> {
|
||||
<assignRetParams>
|
||||
<body>
|
||||
}
|
||||
)");
|
||||
t("functionName", functionName);
|
||||
vector<string> retParams;
|
||||
vector<string> retParamsIn;
|
||||
for (auto const& varDecl: _function.returnParameters())
|
||||
retParams += m_context.addLocalVariable(*varDecl).stackSlots();
|
||||
string assignRetParams;
|
||||
for (size_t i = 0; i < retParams.size(); ++i)
|
||||
{
|
||||
retParamsIn.emplace_back(m_context.newYulVariable());
|
||||
assignRetParams += retParams.back() + " := " + retParamsIn[i] + "\n";
|
||||
}
|
||||
vector<string> params = retParamsIn;
|
||||
for (auto const& varDecl: _function.parameters())
|
||||
params += m_context.addLocalVariable(*varDecl).stackSlots();
|
||||
t("params", joinHumanReadable(params));
|
||||
t("retParams", joinHumanReadable(retParams));
|
||||
t("assignRetParams", assignRetParams);
|
||||
t("body", generate(_function.body()));
|
||||
return t.render();
|
||||
});
|
||||
@ -526,6 +655,7 @@ void IRGenerator::generateImplicitConstructors(ContractDefinition const& _contra
|
||||
ContractDefinition const* contract = _contract.annotation().linearizedBaseContracts[i];
|
||||
baseConstructorParams.erase(contract);
|
||||
|
||||
m_context.resetLocalVariables();
|
||||
m_context.functionCollector().createFunction(IRNames::implicitConstructor(*contract), [&]() {
|
||||
Whiskers t(R"(
|
||||
function <functionName>(<params><comma><baseParams>) {
|
||||
@ -537,16 +667,8 @@ void IRGenerator::generateImplicitConstructors(ContractDefinition const& _contra
|
||||
)");
|
||||
vector<string> params;
|
||||
if (contract->constructor())
|
||||
{
|
||||
for (auto const& modifierInvocation: contract->constructor()->modifiers())
|
||||
// This can be ContractDefinition too for super arguments. That is supported.
|
||||
solUnimplementedAssert(
|
||||
!dynamic_cast<ModifierDefinition const*>(modifierInvocation->name().annotation().referencedDeclaration),
|
||||
"Modifiers not implemented yet."
|
||||
);
|
||||
for (ASTPointer<VariableDeclaration> const& varDecl: contract->constructor()->parameters())
|
||||
params += m_context.addLocalVariable(*varDecl).stackSlots();
|
||||
}
|
||||
t("params", joinHumanReadable(params));
|
||||
vector<string> baseParams = listAllParams(baseConstructorParams);
|
||||
t("baseParams", joinHumanReadable(baseParams));
|
||||
@ -565,7 +687,37 @@ void IRGenerator::generateImplicitConstructors(ContractDefinition const& _contra
|
||||
else
|
||||
t("hasNextConstructor", false);
|
||||
t("initStateVariables", initStateVariables(*contract));
|
||||
t("userDefinedConstructorBody", contract->constructor() ? generate(contract->constructor()->body()) : "");
|
||||
string body;
|
||||
if (FunctionDefinition const* constructor = contract->constructor())
|
||||
{
|
||||
vector<ModifierInvocation*> realModifiers;
|
||||
for (auto const& modifierInvocation: constructor->modifiers())
|
||||
// Filter out the base constructor calls
|
||||
if (dynamic_cast<ModifierDefinition const*>(modifierInvocation->name().annotation().referencedDeclaration))
|
||||
realModifiers.emplace_back(modifierInvocation.get());
|
||||
if (realModifiers.empty())
|
||||
body = generate(constructor->body());
|
||||
else
|
||||
{
|
||||
for (size_t i = 0; i < realModifiers.size(); ++i)
|
||||
{
|
||||
ModifierInvocation const& modifier = *realModifiers.at(i);
|
||||
string next =
|
||||
i + 1 < realModifiers.size() ?
|
||||
IRNames::modifierInvocation(*realModifiers.at(i + 1)) :
|
||||
IRNames::functionWithModifierInner(*constructor);
|
||||
generateModifier(modifier, *constructor, next);
|
||||
}
|
||||
body =
|
||||
IRNames::modifierInvocation(*constructor->modifiers().at(0)) +
|
||||
"(" +
|
||||
joinHumanReadable(params) +
|
||||
")";
|
||||
// Now generate the actual inner function.
|
||||
generateFunctionWithModifierInner(*constructor);
|
||||
}
|
||||
}
|
||||
t("userDefinedConstructorBody", move(body));
|
||||
|
||||
return t.render();
|
||||
});
|
||||
|
@ -73,6 +73,12 @@ private:
|
||||
InternalDispatchMap generateInternalDispatchFunctions();
|
||||
/// Generates code for and returns the name of the function.
|
||||
std::string generateFunction(FunctionDefinition const& _function);
|
||||
std::string generateModifier(
|
||||
ModifierInvocation const& _modifierInvocation,
|
||||
FunctionDefinition const& _function,
|
||||
std::string const& _nextFunction
|
||||
);
|
||||
std::string generateFunctionWithModifierInner(FunctionDefinition const& _function);
|
||||
/// Generates a getter for the given declaration and returns its name
|
||||
std::string generateGetter(VariableDeclaration const& _varDecl);
|
||||
|
||||
|
@ -581,6 +581,12 @@ bool IRGeneratorForStatements::visit(IfStatement const& _ifStatement)
|
||||
return false;
|
||||
}
|
||||
|
||||
void IRGeneratorForStatements::endVisit(PlaceholderStatement const&)
|
||||
{
|
||||
solAssert(m_placeholderCallback, "");
|
||||
m_code << m_placeholderCallback();
|
||||
}
|
||||
|
||||
bool IRGeneratorForStatements::visit(ForStatement const& _forStatement)
|
||||
{
|
||||
setLocation(_forStatement);
|
||||
|
@ -40,8 +40,13 @@ class YulUtilFunctions;
|
||||
class IRGeneratorForStatements: public ASTConstVisitor
|
||||
{
|
||||
public:
|
||||
IRGeneratorForStatements(IRGenerationContext& _context, YulUtilFunctions& _utils):
|
||||
IRGeneratorForStatements(
|
||||
IRGenerationContext& _context,
|
||||
YulUtilFunctions& _utils,
|
||||
std::function<std::string()> _placeholderCallback = {}
|
||||
):
|
||||
m_context(_context),
|
||||
m_placeholderCallback(std::move(_placeholderCallback)),
|
||||
m_utils(_utils)
|
||||
{}
|
||||
|
||||
@ -58,6 +63,9 @@ public:
|
||||
/// Calculates expression's value and returns variable where it was stored
|
||||
IRVariable evaluateExpression(Expression const& _expression, Type const& _to);
|
||||
|
||||
/// Defines @a _var using the value of @a _value while performing type conversions, if required.
|
||||
void define(IRVariable const& _var, IRVariable const& _value) { declareAssign(_var, _value, true); }
|
||||
|
||||
/// @returns the name of a function that computes the value of the given constant
|
||||
/// and also generates the function.
|
||||
std::string constantValueFunction(VariableDeclaration const& _constant);
|
||||
@ -66,6 +74,7 @@ public:
|
||||
bool visit(Conditional const& _conditional) override;
|
||||
bool visit(Assignment const& _assignment) override;
|
||||
bool visit(TupleExpression const& _tuple) override;
|
||||
void endVisit(PlaceholderStatement const& _placeholder) override;
|
||||
bool visit(Block const& _block) override;
|
||||
void endVisit(Block const& _block) override;
|
||||
bool visit(IfStatement const& _ifStatement) override;
|
||||
@ -135,8 +144,7 @@ private:
|
||||
/// @returns an output stream that can be used to define @a _var using a function call or
|
||||
/// single stack slot expression.
|
||||
std::ostream& define(IRVariable const& _var);
|
||||
/// Defines @a _var using the value of @a _value while performing type conversions, if required.
|
||||
void define(IRVariable const& _var, IRVariable const& _value) { declareAssign(_var, _value, true); }
|
||||
|
||||
/// Assigns @a _var to the value of @a _value while performing type conversions, if required.
|
||||
void assign(IRVariable const& _var, IRVariable const& _value) { declareAssign(_var, _value, false); }
|
||||
/// Declares variable @a _var.
|
||||
@ -189,6 +197,7 @@ private:
|
||||
|
||||
std::ostringstream m_code;
|
||||
IRGenerationContext& m_context;
|
||||
std::function<std::string()> m_placeholderCallback;
|
||||
YulUtilFunctions& m_utils;
|
||||
std::optional<IRLValue> m_currentLValue;
|
||||
langutil::SourceLocation m_currentLocation;
|
||||
|
@ -34,5 +34,7 @@ contract C is B {
|
||||
return (x, y);
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// test() -> 5, 10
|
||||
|
@ -8,6 +8,8 @@ contract C {
|
||||
return b + c;
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f(uint16,uint16,uint16): 0xe000, 0xe500, 2 -> 58626
|
||||
// f(uint16,uint16,uint16): 0x1000, 0xe500, 0xe000 -> FAILURE, hex"4e487b71", 0x11
|
||||
|
@ -28,6 +28,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f() -> true
|
||||
// g() -> FAILURE
|
||||
|
@ -14,6 +14,8 @@ contract C {
|
||||
x = t;
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// x() -> 0
|
||||
// f() ->
|
||||
|
@ -14,6 +14,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// x() -> 0
|
||||
// f() ->
|
||||
|
@ -0,0 +1,22 @@
|
||||
contract A { constructor(uint) {} }
|
||||
contract B { constructor(uint) {} }
|
||||
contract C { constructor(uint) {} }
|
||||
|
||||
contract D is A, B, C {
|
||||
uint[] x;
|
||||
constructor() m2(f(1)) B(f(2)) m1(f(3)) C(f(4)) m3(f(5)) A(f(6)) {
|
||||
f(7);
|
||||
}
|
||||
|
||||
function query() public view returns (uint[] memory) { return x; }
|
||||
|
||||
modifier m1(uint) { _; }
|
||||
modifier m2(uint) { _; }
|
||||
modifier m3(uint) { _; }
|
||||
|
||||
function f(uint y) internal returns (uint) { x.push(y); return 0; }
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// query() -> 0x20, 7, 4, 2, 6, 1, 3, 5, 7
|
@ -8,6 +8,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// getOne() -> 0
|
||||
// getOne(), 1 wei -> 1
|
||||
|
@ -45,5 +45,7 @@ contract C is A {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// getData() -> 0x4300
|
||||
|
@ -13,5 +13,7 @@ contract C is A {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f() -> false
|
||||
|
@ -22,6 +22,7 @@ contract C is A {
|
||||
_;
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// getData() -> 6
|
||||
|
@ -24,5 +24,7 @@ contract Test {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f() -> 0x202
|
||||
|
@ -30,5 +30,7 @@ contract Test {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f() -> 0x202
|
||||
|
@ -14,6 +14,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f(bool): true -> 0
|
||||
// f(bool): false -> 3
|
||||
|
@ -8,6 +8,7 @@ contract C {
|
||||
r += 1;
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: false
|
||||
// ----
|
||||
// f() -> 10
|
||||
|
@ -0,0 +1,14 @@
|
||||
contract C {
|
||||
modifier repeat(uint256 count) {
|
||||
uint256 i;
|
||||
for (i = 0; i < count; ++i) _;
|
||||
}
|
||||
|
||||
function f() public repeat(10) returns (uint256 r) {
|
||||
r += 1;
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: true
|
||||
// ----
|
||||
// f() -> 1
|
@ -8,7 +8,8 @@ contract C {
|
||||
r += 1;
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: false
|
||||
// ----
|
||||
// f(bool): false -> 1
|
||||
// f(bool): true -> 2
|
||||
|
@ -0,0 +1,15 @@
|
||||
contract C {
|
||||
modifier repeat(bool twice) {
|
||||
if (twice) _;
|
||||
_;
|
||||
}
|
||||
|
||||
function f(bool twice) public repeat(twice) returns (uint256 r) {
|
||||
r += 1;
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: true
|
||||
// ----
|
||||
// f(bool): false -> 1
|
||||
// f(bool): true -> 1
|
@ -11,7 +11,8 @@ contract C {
|
||||
return r;
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: false
|
||||
// ----
|
||||
// f(bool): false -> 1
|
||||
// f(bool): true -> 2
|
||||
|
@ -10,6 +10,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f(uint256): 3 -> 10
|
||||
// a() -> 10
|
||||
|
@ -13,6 +13,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f(uint256): 3 -> 10
|
||||
// a() -> 0
|
||||
|
@ -15,5 +15,7 @@ contract C is A {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f() -> false
|
||||
|
@ -0,0 +1,15 @@
|
||||
contract C {
|
||||
modifier m(bool condition) {
|
||||
if (condition) _;
|
||||
}
|
||||
|
||||
function f(uint x) public m(x >= 10) returns (uint[5] memory r) {
|
||||
r[2] = 3;
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f(uint256): 9 -> 0x00, 0x00, 0x00, 0x00, 0x00
|
||||
// f(uint256): 10 -> 0x00, 0x00, 3, 0x00, 0x00
|
@ -10,6 +10,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// x() -> 0
|
||||
// f() -> 2
|
||||
|
@ -14,6 +14,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// x() -> 0
|
||||
// f() ->
|
||||
|
@ -15,6 +15,8 @@ contract C {
|
||||
}
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// x() -> 0
|
||||
// f() -> 42
|
||||
|
@ -13,6 +13,8 @@ contract C {
|
||||
return address(this).balance;
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f(), 27 wei -> FAILURE
|
||||
// balance() -> 0
|
||||
|
@ -16,6 +16,8 @@ contract C {
|
||||
x += 3;
|
||||
}
|
||||
}
|
||||
// ====
|
||||
// compileViaYul: also
|
||||
// ----
|
||||
// f1() ->
|
||||
// x() -> 0x08
|
||||
|
Loading…
Reference in New Issue
Block a user