(* Open Source License *) (* Copyright (c) 2021 Nomadic Labs. *) (* Permission is hereby granted, free of charge, to any person obtaining a *) (* copy of this software and associated documentation files (the "Software"), *) (* to deal in the Software without restriction, including without limitation *) (* the rights to use, copy, modify, merge, publish, distribute, sublicense, *) (* and/or sell copies of the Software, and to permit persons to whom the *) (* Software is furnished to do so, subject to the following conditions: *) (* The above copyright notice and this permission notice shall be included *) (* in all copies or substantial portions of the Software. *) (* THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR *) (* IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, *) (* FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL *) (* THE AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER *) (* LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING *) (* FROM, OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER *) (* DEALINGS IN THE SOFTWARE. *) Require Import NArith. Require Import ZArith. Require Import syntax. Require Import semantics. Require Import comparable. Require Import error. Require Import utils. Require Import fa12_interface. Require Import fa12_camlcase. Import List.ListNotations. Open Scope list_scope. Module FA12CamlCaseSpec (C : ContractContext). Module fa12AxiomsImpl := FA12AxiomsImpl C. Import fa12AxiomsImpl. Import fa12TypesImpl fa12DataImpl fa12Notations. Import self semantics. (** Entry point specifications *) Section EntryPoints. Variable (env : @proto_env self_type). Notation sender := (sender env). (** Entry point: ep_approve *) Definition ep_approve (p : data parameter_ep_approve_ty) (sto : data storage_ty) (ret_ops : data (list operation)) (ret_sto : data storage_ty) := let '(spender, allowance') := p in if (compare_equal sender spender) then ret_sto = sto /\ ret_ops = nil else let allowance := getAllowance sto sender spender in (allowance = 0%N \/ allowance' = 0%N) /\ ret_sto = setAllowance sto sender spender allowance' /\ ret_ops = nil. (** Entry point: ep_transfer *) Definition ep_transfer (p : data parameter_ep_transfer_ty) (sto : data storage_ty) (ret_ops : data (list operation)) (ret_sto : data storage_ty) := let '(from, (to, amount)) := p in isOwner sto from /\ let from_balance := getBalance sto from in (amount <= from_balance)%N /\ let from_balance' := Z.abs_N (Z.of_N from_balance - Z.of_N amount) in let sto' := setBalance sto from from_balance' in let to_balance := getBalance sto' to in let sto'' := setBalance sto' to (amount + to_balance)%N in if (compare_equal sender from) then ret_sto = sto'' /\ ret_ops = nil else isSpender sto from sender /\ let allowance := getAllowance sto from sender in (amount <= allowance)%N /\ let allowance' := Z.abs_N (Z.of_N allowance - Z.of_N amount) in ret_sto = setAllowance sto'' from sender allowance' /\ ret_ops = nil. (** Entry point: ep_getTotalSupply *) Definition ep_getTotalSupply (p : data parameter_ep_getTotalSupply_ty) (sto : data storage_ty) (ret_ops : data (list operation)) (ret_sto : data storage_ty) := let '(_, contr) := p in let supply := getTotalSupply sto in let op := transfer_tokens env nat supply (amount env) contr in ret_sto = sto /\ ret_ops = [op]. (** Entry point: ep_getBalance *) Definition ep_getBalance (p : data parameter_ep_getBalance_ty) (sto : data storage_ty) (ret_ops : data (list operation)) (ret_sto : data storage_ty) := let '(owner, contr) := p in let balance := getBalance sto owner in let op := transfer_tokens env nat balance (amount env) contr in isOwner sto owner /\ ret_sto = sto /\ ret_ops = [op]. (** Entry point ep_getAllowance *) Definition ep_getAllowance (p : data parameter_ep_getAllowance_ty) (sto : data storage_ty) (ret_ops : data (list operation)) (ret_sto : data storage_ty) := let '((owner, spender), contr) := p in let allowance := getAllowance sto owner spender in let op := transfer_tokens env nat allowance (amount env) contr in isOwner sto owner /\ isSpender sto owner spender /\ ret_sto = sto /\ ret_ops = [op]. End EntryPoints. (** Full spec *) Definition spec (env : @proto_env self_type) (p : data parameter_ty) (sto : data storage_ty) (ret_ops : data (list operation)) (ret_sto : data storage_ty) := match p with | inl (inl p) => ep_approve env p sto ret_ops ret_sto | inl (inr p) => ep_transfer env p sto ret_ops ret_sto | inr (inl p) => ep_getAllowance env p sto ret_ops ret_sto | inr (inr (inl p)) => ep_getBalance env p sto ret_ops ret_sto | inr (inr (inr p)) => ep_getTotalSupply env p sto ret_ops ret_sto end. End FA12CamlCaseSpec.