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_tokenToToken.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_tokenToToken_correct
      (env : @proto_env dexter_definition.self_type)
      (p : data dexter_definition.parameter_ep_tokenToToken_ty)
      (sto : data dexter_definition.storage_ty)
      (ret_ops : Datatypes.list (data operation))
      (ret_sto : data dexter_definition.storage_ty)
      (fuel : Datatypes.nat) :
  8 <= fuel ->
  eval_seq_precond fuel env Def.ep_tokenToToken
                (fun x => x = (ret_ops, ret_sto, tt))
                (p, (sto, tt)) <->
  Spec.ep_tokenToToken env p sto ret_ops ret_sto.
Proof.
  intros Hfuel.
  simpl in p.
  destruct p
    as (outputDexterContract & minTokensBought & to & tokensSold & deadline).
  destruct_sto sto
    tokenPool
    xtzPool
    lqtTotal
    selfIsUpdatingTokenPool
    freezeBaker
    manager
    tokenAddress
    lqtAddress.
  do 8 more_fuel; simpl.

  repeat fold_mutez_goal.

  unfold Spec.get_amount_bought, Spec.mutez_to_nat; simpl.

  Ltac ex_false_crush := simpl; try (rewrite ex_false; intuition; destruct_conjs; intuition (try congruence); fail).

  fold dexter_definition.parameter_ep_xtzToToken_ty.
  destruct (contract_ (Some "%xtzToToken"%string)
    dexter_definition.parameter_ep_xtzToToken_ty outputDexterContract).
  rewrite ex_eq_some_simpl.

  rewrite bool_false_not; apply and_both.
  rewrite and_comm_3.
  rewrite Z_gtb_false_leb_iff, Z_geb_false_ltb_iff,
    comparison_to_int_tez_compare_le_zero, check_deadline_correct.
  do 2 apply and_both.
  rewrite Zmod_1_r; simpl.
  rewrite Zdiv_1_r.
  rewrite ex_eq_some_simpl.
  rewrite <- ediv_N_some_iff.
  destruct (N.eq_dec (tokenPool * 1000 + tokensSold * 997) 0) as [H_divisor | H_divisor]; try now intuition.
  
  rewrite <- Zred_factor0.
  destruct (Spec.nat_to_mutez _) as [|xtz_bought_mutez] eqn:Hxtz_bought_mutez;
    unfold Spec.nat_to_mutez in Hxtz_bought_mutez;
    rewrite Hxtz_bought_mutez; simpl; try easy.

  unfold Spec.mutez_sub; simpl.
  destruct (tez.of_Z (tez.to_Z xtzPool - tez.to_Z xtz_bought_mutez));
    try ex_false_crush.
  apply and_left; trivial.
  apply forall_ex; intro transfer_contract.

  split; intro H.
  { eexists.
    destruct_conjs.
    repeat rewrite shim_pair_equal_spec in *.
    destruct_conjs.
    intuition; congruence.
  }
  { destruct_conjs.
    intuition.
    repeat rewrite shim_pair_equal_spec.
    intuition.
    match goal with
    | H : Return _ = Return _ |- _ => injection H
    end.
    congruence.
  }
  { split; intro H.
    { destruct_conjs.
      congruence.
    }
    { destruct (N.eq_dec _ 0); try ex_false_crush.
      destruct (Spec.nat_to_mutez _); ex_false_crush.
    }
  }
Qed.
back to top