Michael.W基于Foundry精读Openzeppelin

2024年08月13日更新 127 人订阅
专栏简介 Michael.W基于Foundry精读Openzeppelin第8期——Context.sol Michael.W基于Foundry精读Openzeppelin第1期——Address.sol Michael.W基于Foundry精读Openzeppelin第2期——StorageSlot.sol Michael.W基于Foundry精读Openzeppelin第3期——Arrays.sol Michael.W基于Foundry精读Openzeppelin第4期——Base64.sol Michael.W基于Foundry精读Openzeppelin第5期——Counters.sol Michael.W基于Foundry精读Openzeppelin第6期——Strings.sol Michael.W基于Foundry精读Openzeppelin第7期——Timers.sol Michael.W基于Foundry精读Openzeppelin第9期——Multicall.sol Michael.W基于Foundry精读Openzeppelin第10期——Create2.sol Michael.W基于Foundry精读Openzeppelin第11期——Math.sol Michael.W基于Foundry精读Openzeppelin第12期——SafeCast.sol Michael.W基于Foundry精读Openzeppelin第13期——Checkpoints.sol Michael.W基于Foundry精读Openzeppelin第14期——SafeMath.sol Michael.W基于Foundry精读Openzeppelin第15期——SignedMath.sol Michael.W基于Foundry精读Openzeppelin第16期——SignedSafeMath.sol Michael.W基于Foundry精读Openzeppelin第17期——BitMaps.sol Michael.W基于Foundry精读Openzeppelin第18期——DoubleEndedQueue.sol Michael.W基于Foundry精读Openzeppelin第19期——EnumerableSet.sol Michael.W基于Foundry精读Openzeppelin第20期——EnumerableMap.sol Michael.W基于Foundry精读Openzeppelin第21期——ERC165.sol (番外篇)Michael.W基于Foundry精读Openzeppelin第22期——内联汇编staticcall Michael.W基于Foundry精读Openzeppelin第23期——ERC165Checker.sol Michael.W基于Foundry精读Openzeppelin第24期——ERC165Storage.sol Michael.W基于Foundry精读Openzeppelin第25期——IERC1820Registry.sol Michael.W基于Foundry精读Openzeppelin第26期——ERC1820Implementer.sol Michael.W基于Foundry精读Openzeppelin第27期——Escrow.sol Michael.W基于Foundry精读Openzeppelin第28期——ConditionalEscrow.sol Michael.W基于Foundry精读Openzeppelin第29期——RefundEscrow.sol Michael.W基于Foundry精读Openzeppelin第30期——ECDSA.sol Michael.W基于Foundry精读Openzeppelin第31期——IERC1271.sol Michael.W基于Foundry精读Openzeppelin第32期——SignatureChecker.sol Michael.W基于Foundry精读Openzeppelin第33期——EIP712.sol Michael.W基于Foundry精读Openzeppelin第34期——MerkleProof.sol Michael.W基于Foundry精读Openzeppelin第35期——Ownable.sol Michael.W基于Foundry精读Openzeppelin第36期——Ownable2Step.sol Michael.W基于Foundry精读Openzeppelin第37期——AccessControl.sol Michael.W基于Foundry精读Openzeppelin第38期——AccessControlEnumerable.sol Michael.W基于Foundry精读Openzeppelin第39期——ERC20.sol Michael.W基于Foundry精读Openzeppelin第40期——ERC20Burnable.sol Michael.W基于Foundry精读Openzeppelin第41期——ERC20Capped.sol Michael.W基于Foundry精读Openzeppelin第42期——draft-ERC20Permit.sol Michael.W基于Foundry精读Openzeppelin第43期——Pausable.sol Michael.W基于Foundry精读Openzeppelin第44期——ERC20Pausable.sol Michael.W基于Foundry精读Openzeppelin第45期——ERC20FlashMint.sol Michael.W基于Foundry精读Openzeppelin第46期——ERC20Snapshot.sol Michael.W基于Foundry精读Openzeppelin第47期——SafeERC20.sol Michael.W基于Foundry精读Openzeppelin第48期——TokenTimelock.sol Michael.W基于Foundry精读Openzeppelin第49期——ERC20Wrapper.sol Michael.W基于Foundry精读Openzeppelin第50期——ERC20Votes.sol Michael.W基于Foundry精读Openzeppelin第51期——ERC20VotesComp.sol Michael.W基于Foundry精读Openzeppelin第52期——ERC4626.sol Michael.W基于Foundry精读Openzeppelin第53期——ERC20PresetFixedSupply.sol Michael.W基于Foundry精读Openzeppelin第54期——ERC20PresetMinterPauser.sol Michael.W基于Foundry精读Openzeppelin第55期——PaymentSplitter.sol Michael.W基于Foundry精读Openzeppelin第56期——VestingWallet.sol Michael.W基于Foundry精读Openzeppelin第57期——ReentrancyGuard.sol Michael.W基于Foundry精读Openzeppelin第58期——PullPayment.sol Michael.W基于Foundry精读Openzeppelin第59期——Proxy.sol Michael.W基于Foundry精读Openzeppelin第60期——Clones.sol Michael.W基于Foundry精读Openzeppelin第61期——ERC1967Upgrade.sol Michael.W基于Foundry精读Openzeppelin第62期——ERC1967Proxy.sol Michael.W基于Foundry精读Openzeppelin第63期——Initializable.sol Michael.W基于Foundry精读Openzeppelin第64期——UUPSUpgradeable.sol Michael.W基于Foundry精读Openzeppelin第65期——TransparentUpgradeableProxy.sol Michael.W基于Foundry精读Openzeppelin第66期——ProxyAdmin.sol Michael.W基于Foundry精读Openzeppelin第67期——BeaconProxy.sol Michael.W基于Foundry精读Openzeppelin第68期——UpgradeableBeacon.sol

Michael.W基于Foundry精读Openzeppelin第66期——ProxyAdmin.sol

  • Michael.W
  • 发布于 2024-07-21 14:59
  • 阅读 2263

ProxyAdmin库是指定用于做透明代理TransparentUpgradeableProxy库admin的管理员合约。

0. 版本

[openzeppelin]:v4.8.3,[forge-std]:v1.5.6

0.1 ProxyAdmin.sol

Github: https://github.com/OpenZeppelin/openzeppelin-contracts/blob/v4.8.3/contracts/proxy/transparent/ProxyAdmin.sol

ProxyAdmin库是指定用于做透明代理TransparentUpgradeableProxy库admin的管理员合约。

注:用一个额外合约来做透明代理合约的admin的原因及透明代理合约详解参见:https://learnblockchain.cn/article/8770

1. 目标合约

ProxyAdmin合约可直接部署。

全部foundry测试合约:

Github: https://github.com/RevelationOfTuring/foundry-openzeppelin-contracts/blob/master/test/proxy/transparent/ProxyAdmin/ProxyAdmin.t.sol

测试使用的物料合约:

Github: https://github.com/RevelationOfTuring/foundry-openzeppelin-contracts/blob/master/test/proxy/transparent/ProxyAdmin/Implementation.sol

// SPDX-License-Identifier: UNLICENSED
pragma solidity ^0.8.0;

interface IImplementation {
    event ChangeStorageUint(uint, uint);
}

contract Implementation is IImplementation {
    // storage
    uint public i;

    function __Implementation_init(uint i_) external {
        i = i_;
    }
}

contract ImplementationNew is Implementation {
    // add a function
    function addI(uint i_) external payable {
        i += i_;
        emit ChangeStorageUint(i, msg.value);
    }
}

2. 代码精读

注:以下所有方法只有当本ProxyAdmin合约为透明代理合约proxy的admin时才能调用成功。

2.1 getProxyImplementation(ITransparentUpgradeableProxy proxy)

查询透明代理合约proxy背后的逻辑合约地址。

    function getProxyImplementation(ITransparentUpgradeableProxy proxy) public view virtual returns (address) {
        // staticcall到透明代理合约proxy(calldata为0x5c60da1b),如同去调用proxy合约的implementation()方法
        (bool success, bytes memory returndata) = address(proxy).staticcall(hex"5c60da1b");
        // 要求调用必须成功
        require(success);
        // 将staticcall返回的bytes解码成address类型并返回
        return abi.decode(returndata, (address));
    }

foundry代码验证:

contract ProxyAdminTest is Test, IERC1967, IImplementation {
    ProxyAdmin private _testing = new ProxyAdmin();
    Implementation private _implementation1 = new Implementation();
    Implementation private _implementation2 = new Implementation();
    TransparentUpgradeableProxy private _transparentUpgradeableProxy1 = new TransparentUpgradeableProxy(
        address(_implementation1),
        address(_testing),
        abi.encodeCall(
            _implementation1.__Implementation_init,
            (1024)
        )
    );
    TransparentUpgradeableProxy private _transparentUpgradeableProxy2 = new TransparentUpgradeableProxy(
        address(_implementation2),
        address(_testing),
        abi.encodeCall(
            _implementation2.__Implementation_init,
            (2048)
        )
    );

    function test_GetProxyImplementation() external {
        assertEq(
            _testing.getProxyImplementation(
                ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1))
            ),
            address(_implementation1)
        );
        assertEq(
            _testing.getProxyImplementation(
                ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2))
            ),
            address(_implementation2)
        );
    }
}

2.2 getProxyAdmin(ITransparentUpgradeableProxy proxy) && changeProxyAdmin(ITransparentUpgradeableProxy proxy, address newAdmin)

  • getProxyAdmin(ITransparentUpgradeableProxy proxy):查询透明代理合约proxy的admin;
  • changeProxyAdmin(ITransparentUpgradeableProxy proxy, address newAdmin):本合约的owner修改透明代理合约proxy的admin为newAdmin。
    function getProxyAdmin(ITransparentUpgradeableProxy proxy) public view virtual returns (address) {
        // staticcall到透明代理合约proxy(calldata为f851a440),如同去调用proxy合约的admin()方法
        (bool success, bytes memory returndata) = address(proxy).staticcall(hex"f851a440");
        // 要求调用必须成功
        require(success);
        // 将staticcall返回的bytes解码成address类型并返回
        return abi.decode(returndata, (address));
    }

    function changeProxyAdmin(ITransparentUpgradeableProxy proxy, address newAdmin) public virtual onlyOwner {
        // 如同调用proxy合约的changeAdmin(address)方法
        proxy.changeAdmin(newAdmin);
    }

foundry代码验证:

contract ProxyAdminTest is Test, IERC1967, IImplementation {
    ProxyAdmin private _testing = new ProxyAdmin();
    Implementation private _implementation1 = new Implementation();
    Implementation private _implementation2 = new Implementation();
    TransparentUpgradeableProxy private _transparentUpgradeableProxy1 = new TransparentUpgradeableProxy(
        address(_implementation1),
        address(_testing),
        abi.encodeCall(
            _implementation1.__Implementation_init,
            (1024)
        )
    );
    TransparentUpgradeableProxy private _transparentUpgradeableProxy2 = new TransparentUpgradeableProxy(
        address(_implementation2),
        address(_testing),
        abi.encodeCall(
            _implementation2.__Implementation_init,
            (2048)
        )
    );

    function test_GetProxyAdminAndChangeProxyAdmin() external {
        assertEq(
            _testing.getProxyAdmin(
                ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1))
            ),
            address(_testing)
        );
        assertEq(
            _testing.getProxyAdmin(
                ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2))
            ),
            address(_testing)
        );

        // deploy another ProxyAdmin
        ProxyAdmin newProxyAdmin = new ProxyAdmin();
        // test changeProxyAdmin(ITransparentUpgradeableProxy,address)
        vm.expectEmit(address(_transparentUpgradeableProxy1));
        emit IERC1967.AdminChanged(address(_testing), address(newProxyAdmin));

        _testing.changeProxyAdmin(ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1)), address(newProxyAdmin));

        vm.expectEmit(address(_transparentUpgradeableProxy2));
        emit IERC1967.AdminChanged(address(_testing), address(newProxyAdmin));

        _testing.changeProxyAdmin(ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2)), address(newProxyAdmin));

        assertEq(
            newProxyAdmin.getProxyAdmin(
                ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1))
            ),
            address(newProxyAdmin)
        );
        assertEq(
            newProxyAdmin.getProxyAdmin(
                ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2))
            ),
            address(newProxyAdmin)
        );

        // revert if not owner calls
        assertEq(newProxyAdmin.owner(), address(this));
        vm.prank(address(1));
        vm.expectRevert("Ownable: caller is not the owner");
        newProxyAdmin.changeProxyAdmin(
            ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1)),
            address(1)
        );
    }
}

2.3 upgrade(ITransparentUpgradeableProxy proxy, address implementation)

本合约的owner升级透明代理合约proxy的逻辑合约地址为implementation。

    function upgrade(ITransparentUpgradeableProxy proxy, address implementation) public virtual onlyOwner {
        // 如同调用proxy合约的upgradeTo(address)方法
        proxy.upgradeTo(implementation);
    }

foundry代码验证:

contract ProxyAdminTest is Test, IERC1967, IImplementation {
    ProxyAdmin private _testing = new ProxyAdmin();
    Implementation private _implementation1 = new Implementation();
    Implementation private _implementation2 = new Implementation();
    TransparentUpgradeableProxy private _transparentUpgradeableProxy1 = new TransparentUpgradeableProxy(
        address(_implementation1),
        address(_testing),
        abi.encodeCall(
            _implementation1.__Implementation_init,
            (1024)
        )
    );
    TransparentUpgradeableProxy private _transparentUpgradeableProxy2 = new TransparentUpgradeableProxy(
        address(_implementation2),
        address(_testing),
        abi.encodeCall(
            _implementation2.__Implementation_init,
            (2048)
        )
    );
    ImplementationNew private _implementationNew = new ImplementationNew();

    function test_Upgrade() external {
        // upgrade one transparent upgradeable proxy
        vm.expectEmit(address(_transparentUpgradeableProxy1));
        emit IERC1967.Upgraded(address(_implementationNew));

        _testing.upgrade(
            ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1)),
            address(_implementationNew)
        );

        // check the result of upgrade
        assertEq(
            _testing.getProxyImplementation(ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1))),
            address(_implementationNew)
        );

        ImplementationNew transparentUpgradeableProxy1AsNew = ImplementationNew(address(_transparentUpgradeableProxy1));
        assertEq(transparentUpgradeableProxy1AsNew.i(), 1024);
        vm.expectEmit(address(transparentUpgradeableProxy1AsNew));
        emit IImplementation.ChangeStorageUint(1024 + 1, 0);

        transparentUpgradeableProxy1AsNew.addI(1);
        assertEq(transparentUpgradeableProxy1AsNew.i(), 1024 + 1);

        // upgrade another transparent upgradeable proxy
        vm.expectEmit(address(_transparentUpgradeableProxy2));
        emit IERC1967.Upgraded(address(_implementationNew));

        _testing.upgrade(
            ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2)),
            address(_implementationNew)
        );

        // check the result of upgrade
        assertEq(
            _testing.getProxyImplementation(ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2))),
            address(_implementationNew)
        );

        ImplementationNew transparentUpgradeableProxy2AsNew = ImplementationNew(address(_transparentUpgradeableProxy2));
        assertEq(transparentUpgradeableProxy2AsNew.i(), 2048);
        vm.expectEmit(address(transparentUpgradeableProxy2AsNew));
        emit IImplementation.ChangeStorageUint(2048 + 2, 0);

        transparentUpgradeableProxy2AsNew.addI(2);
        assertEq(transparentUpgradeableProxy2AsNew.i(), 2048 + 2);

        // revert if not owner calls
        assertEq(_testing.owner(), address(this));
        vm.prank(address(1));
        vm.expectRevert("Ownable: caller is not the owner");
        _testing.upgrade(
            ITransparentUpgradeableProxy(address(transparentUpgradeableProxy1AsNew)),
            address(_implementationNew)
        );
    }
}

2.4 upgradeAndCall(ITransparentUpgradeableProxy proxy, address implementation, bytes memory data)

本合约的owner升级透明代理合约proxy的逻辑合约地址为implementation并随后以data为calldata执行一次delegatecall。

    function upgradeAndCall(
        ITransparentUpgradeableProxy proxy,
        address implementation,
        bytes memory data
    ) public payable virtual onlyOwner {
        // 如同携带本次call的全部msg.value去调用proxy合约的upgradeToAndCall(address,bytes)方法
        proxy.upgradeToAndCall{value: msg.value}(implementation, data);
    }

foundry代码验证:

contract ProxyAdminTest is Test, IERC1967, IImplementation {
    ProxyAdmin private _testing = new ProxyAdmin();
    Implementation private _implementation1 = new Implementation();
    Implementation private _implementation2 = new Implementation();
    TransparentUpgradeableProxy private _transparentUpgradeableProxy1 = new TransparentUpgradeableProxy(
        address(_implementation1),
        address(_testing),
        abi.encodeCall(
            _implementation1.__Implementation_init,
            (1024)
        )
    );
    TransparentUpgradeableProxy private _transparentUpgradeableProxy2 = new TransparentUpgradeableProxy(
        address(_implementation2),
        address(_testing),
        abi.encodeCall(
            _implementation2.__Implementation_init,
            (2048)
        )
    );
    ImplementationNew private _implementationNew = new ImplementationNew();

    function test_UpgradeAndCall() external {
        // upgrade one transparent upgradeable proxy and delegatecall the added function in new implementation
        uint ethValue = 1024;
        vm.expectEmit(address(_transparentUpgradeableProxy1));
        emit IERC1967.Upgraded(address(_implementationNew));
        vm.expectEmit(address(_transparentUpgradeableProxy1));
        emit IImplementation.ChangeStorageUint(1024 + 1, ethValue);

        _testing.upgradeAndCall{value: ethValue}(
            ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1)),
            address(_implementationNew),
            abi.encodeCall(
                _implementationNew.addI,
                (1)
            )
        );

        // check the result of upgrade
        assertEq(
            _testing.getProxyImplementation(ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1))),
            address(_implementationNew)
        );

        assertEq(ImplementationNew(address(_transparentUpgradeableProxy1)).i(), 1024 + 1);
        assertEq(address(_transparentUpgradeableProxy1).balance, ethValue);

        // upgrade another transparent upgradeable proxy and delegatecall the added function in new implementation
        vm.expectEmit(address(_transparentUpgradeableProxy2));
        emit IERC1967.Upgraded(address(_implementationNew));
        vm.expectEmit(address(_transparentUpgradeableProxy2));
        emit IImplementation.ChangeStorageUint(2048 + 2, ethValue);

        _testing.upgradeAndCall{value: ethValue}(
            ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2)),
            address(_implementationNew),
            abi.encodeCall(
                _implementationNew.addI,
                (2)
            )
        );

        // check the result of upgrade
        assertEq(
            _testing.getProxyImplementation(ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy2))),
            address(_implementationNew)
        );

        assertEq(ImplementationNew(address(_transparentUpgradeableProxy2)).i(), 2048 + 2);
        assertEq(address(_transparentUpgradeableProxy2).balance, ethValue);

        // revert if not owner calls
        assertEq(_testing.owner(), address(this));
        address nonOwner = address(1);
        vm.deal(nonOwner, ethValue);
        vm.prank(nonOwner);
        vm.expectRevert("Ownable: caller is not the owner");

        _testing.upgradeAndCall{value: ethValue}(
            ITransparentUpgradeableProxy(address(_transparentUpgradeableProxy1)),
            address(_implementationNew),
            abi.encodeCall(
                _implementationNew.addI,
                (1)
            )
        );
    }
}

ps: 本人热爱图灵,热爱中本聪,热爱V神。 以下是我个人的公众号,如果有技术问题可以关注我的公众号来跟我交流。 同时我也会在这个公众号上每周更新我的原创文章,喜欢的小伙伴或者老伙计可以支持一下! 如果需要转发,麻烦注明作者。十分感谢!

1.jpeg

公众号名称:后现代泼痞浪漫主义奠基人

点赞 0
收藏 0
分享
本文参与登链社区写作激励计划 ,好文好收益,欢迎正在阅读的你也加入。

0 条评论

请先 登录 后评论