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_removeLiquidity.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_removeLiquidity_correct
      (env : dexter_spec.env_ty)
      (p : data dexter_definition.parameter_ep_removeLiquidity_ty)
      (sto : data dexter_definition.storage_ty)
      (ret_ops : Datatypes.list (data operation))
      (ret_sto : data dexter_definition.storage_ty)
      (fuel : Datatypes.nat) :
  10 <= fuel ->
  eval_seq_precond fuel env Def.ep_removeLiquidity
               (fun x => x = (ret_ops, ret_sto, tt))
               (p, (sto, tt)) <->
  Spec.ep_removeLiquidity env p sto ret_ops ret_sto.
Proof.
  intros Hfuel.
  destruct_sto sto tokenPool xtzPool lqtTotal selfIsUpdatingTokenPool freezeBaker manager tokenAddress lqtAddress.
  destruct p as (to, (lqt_burned, (min_xtz_withdrawn, (min_tokens_withdrawn, deadline)))).
  do 7 (more_fuel; simpl); repeat fold_mutez_goal.


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

  (* amount *)
  rewrite and_comm_3.
  apply and_both_0.

  rewrite Z_gtb_false_leb_iff.
  now rewrite comparison_to_int_tez_compare_le_zero.

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

  (* goal manip *)
  rewrite mutez_to_nat.
  rewrite ex_eq_simpl_some.
  rewrite <- !ediv_N_some_iff.

  unfold Spec.mutez_to_nat, Spec.nat_to_mutez.
  unfold precond.

  change (tez.to_Z (1 ~Mutez)) with 1%Z.
  rewrite <- Zred_factor0.

  rewrite N.mul_comm.
  (* TODO: why this does intuition need explicit exfalso? *)
  destruct (tez.of_Z (Z.of_N (Z.to_N (tez.to_Z xtzPool) * lqt_burned / lqtTotal)))
    as [|xtz_withdrawn] eqn:Hxtz_withdrawn;
    [ intuition exfalso; assumption | ] .

  rewrite <- !ediv_N_some_iff.

  destruct (Z.of_N lqtTotal - Z.of_N lqt_burned >=? 0)%Z eqn:HsufficientLqt.

  - rewrite ex_eq_simpl_some; simpl.
    destruct (Z.of_N tokenPool - Z.of_N (lqt_burned * tokenPool / lqtTotal) >=? 0)%Z eqn:Htoken_withdrawn.
    + rewrite ex_eq_simpl_some; simpl.

      unfold Spec.mutez_sub. simpl sub. rewrite <- ex_match_error.

      repeat rewrite ex_and_comm.

      do 4 rewrite <- underex_and_assoc.
      rewrite underex_and_comm2.

      destruct (tez.of_Z (tez.to_Z xtzPool - tez.to_Z xtz_withdrawn)) as [|xtzPool'] eqn:HxtzPool'.
      * apply iff_both_false; intros; destruct_conjs; intuition.
      * rewrite ex_match_option.

        repeat rewrite ex_and_comm.
        repeat rewrite <- underex_and_assoc.
        rewrite underex_and_comm2.

        rewrite ex_order.
        symmetry. rewrite ex_order. symmetry.
        apply forall_ex. intro contract_token_transfer.

        rewrite underex_and_comm2.
        rewrite ex_order.
        apply forall_ex. intro contract_lqtLedger_mintOrBurn.

        repeat rewrite ex_and_comm.
        apply forall_ex. intro contract_to.

        (* start proof *)

        rewrite !Z_ltb_false_geb_iff.
        rewrite !comparison_to_int_compare_nat_ge.
        rewrite !comparison_to_int_compare_mutez_ge.

        rewrite Z.geb_le in HsufficientLqt.
        rewrite Z.geb_le in Htoken_withdrawn.
        rewrite <- tez.of_Z_to_Z_eqv in Hxtz_withdrawn.

        rewrite <- Z_of_N_ge_Z_to_N_compat.
        rewrite !(N.mul_comm tokenPool).

        rewrite !Z_to_N_of_N_distrib_sub.

        repeat rewrite shim_pair_equal_spec.

        intuition; lia.
    + apply iff_both_false.
      rewrite ex_eq_simpl_none. intuition.
      rewrite N.mul_comm in Htoken_withdrawn.
      rewrite Z.geb_leb in Htoken_withdrawn.
      rewrite Z.leb_gt in Htoken_withdrawn.
      lia.
  - apply iff_both_false.
    rewrite ex_eq_simpl_none. intuition.
    rewrite Z.geb_leb in HsufficientLqt.
    rewrite Z.leb_gt in HsufficientLqt.
    lia.
Qed.
back to top