https://gitlab.com/nomadic-labs/mi-cho-coq
Raw File
Tip revision: d7b21f4f74735ee5a48d9a8c09af765c38a30d44 authored by Guillaume Claret on 12 March 2021, 11:19:43 UTC
[mi-cho-coq] Replace String.string_dec by String.eqb
Tip revision: d7b21f4
dexter_verification_ep_addLiquidity.v
(* Open Source License *)
(* Copyright (c) 2019 Nomadic Labs. <contact@nomadic-labs.com> *)

(* 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. *)

(* Coq *)
Require Import String.
Require Import Michocoq.macros.
Import syntax.
Import comparable.
Require Import NArith.
Require Import ZArith.
Require Import Coq.Setoids.Setoid.
Require Import Coq.micromega.Lia.
From Coq Require Import Program.Tactics.

(* Mi-Cho-Coq *)
Require Import semantics.
Require Import util.
Import error.
Require List.

(* Dexter *)
Require dexter_definition.
Require dexter_spec.
Require Import dexter_util.
Require Import dexter_spec_util.

Module Spec := dexter_spec.

Module Def := dexter_definition.

Import Def.Storage.

Import Tactics.

Require Import contract_semantics.

(* It's better not to unfold compare in this proof,
 * this makes it easier to use util.eqb_eq.
 *)
Opaque N.mul.
Opaque N.add.
Opaque Z.modulo.
Opaque compare.

Lemma ep_addLiquidity_correct
        (env : @proto_env dexter_definition.self_type)
        (p : data dexter_definition.parameter_ep_addLiquidity_ty)
        (sto : data dexter_definition.storage_ty)
        (ret_ops : Datatypes.list (data operation))
        (ret_sto : data dexter_definition.storage_ty)
        (fuel : Datatypes.nat) :
    (* TODO: fuel may have to be adapted to suit the entrypoint *)
    400 <= fuel ->

    eval_seq_precond fuel env Def.ep_addLiquidity
                 (fun x => x = (ret_ops, ret_sto, tt))
                 (p, (sto, tt)) <->
    Spec.ep_addLiquidity env p sto ret_ops ret_sto.
  Proof.
    intros Hfuel.
    do 6 (more_fuel; simpl).
    destruct_sto sto tokenPool xtzPool lqtTotal selfIsUpdatingTokenPool freezeBaker manager tokenAddress lqtAddress.
    destruct p as (owner, (min_lqt_minted, (max_tokens_deposited, deadline))).
    simpl.

    (* update token *)
    apply and_both_0; [apply bool_false_not|].

    (* deadline *)
    apply and_both_0; [apply check_deadline_correct_alt|].

    rewrite mutez_to_nat.

    rewrite ex_eq_simpl_some.

    rewrite mutez_to_nat. rewrite ex_eq_simpl_some.

    rewrite <- !ediv_N_some_iff.

    rewrite and_idempotent.

    apply and_both_0; [apply tez_to_Z_to_N_neq_zero_iff|].

    unfold Spec.mutez_to_nat.

    rewrite !(N.mul_comm _ (Z.to_N (tez.to_Z (amount env)))).

    rewrite !(N.add_comm 1).

    rewrite comparison_to_int_compare_nat_eqb.

    fold_mutez.

    if_step.

    all: rewrite precond_exists; unfold precond_ex.
    all: rewrite !ex_and_comm; rewrite <- ex_order; rewrite ex_order3; rewrite ex_order.
    all: apply forall_ex; intro xtzPool'.
    all: rewrite !ex_and_comm; rewrite ex_order; apply forall_ex; intro lqt_transfer.
    all: rewrite !ex_and_comm; apply forall_ex; intro lqt_mintOrBurn.
    all: rewrite !Z_gtb_false_leb_iff, !Z_ltb_false_geb_iff,
      !comparison_to_int_compare_nat_le, !comparison_to_int_compare_nat_ge.

    1: rewrite Spec_Nceiling_mod by now apply N.eqb_eq.
    2: rewrite Spec_Nceiling_rem by now apply N.eqb_neq.

    all: unfold Spec.lqt.mint, Spec.lqt.mintOrBurn.
    all: rewrite N.ge_le_iff; repeat rewrite shim_pair_equal_spec; intuition.
Qed.
back to top