https://gitlab.com/nomadic-labs/mi-cho-coq
Raw File
Tip revision: 7ccb33612808dc6eba01b4d54e638e3ebade2924 authored by Raphaƫl Cauderlier on 18 March 2021, 11:08:58 UTC
Use the fuel-free WP in Dexter proofs
Tip revision: 7ccb336
dexter_verification_ep_xtzToToken.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.
Require Import tez.
Require Import syntax_type.
Require Import int64bv.

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

Definition ep_xtzToToken := Eval vm_compute in Def.ep_xtzToToken.

Lemma ep_xtzToToken_same : Def.ep_xtzToToken = ep_xtzToToken.
Proof.
  reflexivity.
Qed.

Lemma ep_xtzToToken_correct
      (env : @proto_env dexter_definition.self_type)
      (p : data dexter_definition.parameter_ep_xtzToToken_ty)
      (sto : data dexter_definition.storage_ty)
      (ret_ops : Datatypes.list (data operation))
      (ret_sto : data dexter_definition.storage_ty) :
  eval0_seq_precond env Def.ep_xtzToToken
               (fun x => x = (ret_ops, ret_sto, tt))
               (p, (sto, tt)) <->
  Spec.ep_xtzToToken env p sto ret_ops ret_sto.
Proof.
  rewrite ep_xtzToToken_same.
  destruct p as [to [min_tokens_bought deadline]].
  destruct_sto sto token_pool xtz_pool lqt_total self_is_updating_token_pool freeze_baker manager token_address lqt_address.
  simpl;
    repeat fold_mutez_goal;
    change tez.compare with (compare mutez).
  do 4 (try apply and_both_0).


  - rewrite bool_false_not. easy.
  - rewrite <- check_deadline_correct.
    rewrite comparison_to_int_opp.
    rewrite comparison_to_int_Z_compare_ltb.
    rewrite comparison_to_int_leb.
    rewrite Z.leb_gt.
    rewrite <- Zlt_is_lt_bool.
    easy.
  - fold_mutez.
    unfold Spec.get_amount_bought. unfold Spec.mutez_to_nat.
    rewrite !Z.div_1_r in * |- *.
    rewrite !Z.mod_1_r in * |- *.
    rewrite !Z2N.id in * |- *.
    split.
    + error_step.
      intro H.
      repeat (destruct_exists; destruct_conjs; try if_step; simpl in * |- *);
              try rewrite ediv_spec in * |-; destruct_conjs; subst; simplify_eqs;
              try easy; fold_amount.

      ++ lia.
      ++ intuition.
         +++ apply comparison_to_int_lt_false; auto.
         +++ rewrite N2Z.inj_le. apply Z_sub_ge_bool; auto.
         +++ unfold precond in * |-. error_step; repeat (destruct_exists; destruct_conjs).

             ++++ eexists. eexists. intuition; simplify_eqs; f_equal; try easy; eauto.
                  rewrite <- N2Z.inj_iff. rewrite Z2N.id. rewrite N2Z.inj_sub. auto.
                  apply N_sub_ge_bool; auto.
                  apply Z_geb_lt; auto.

    + intro Hspec.
      if_step. inversion Hspec.
      intuition.

      repeat (destruct_exists; destruct_conjs; try if_step; simpl in * |- *);
              try rewrite ediv_spec in * |-; destruct_conjs; subst; simplify_eqs;
              try easy; fold_amount.

      eexists. intuition. eexists. intuition. eexists. intuition.

      ++ rewrite ediv_spec.
        intuition; try easy.
        eapply not_eq_and_le_implies_lt; try easy; auto.
        fold_amount. apply N.le_0_l.
      ++ simpl.
         unfold Spec.mutez_add, add in * |-.
         fold_amount.
         intuition.
         +++ apply comparison_to_int_lt_false; auto.
         +++ if_step.
           ++++ eexists. intuition; try easy.
                simpl. unfold precond. error_step.
                fold_amount.
                congruence.
                +++++ eexists. intuition; try easy. eauto.
                ++++++ repeat f_equal; auto.
                +++++++ rewrite <- N2Z.inj_iff; auto. rewrite Z2N.id. rewrite N2Z.inj_sub. auto.
                auto. apply Z_geb_lt; auto.
                +++++++ fold_amount. congruence.
           ++++
             apply N_sub_ge_bool in H3.
             fold_amount.
             rewrite H3 in Hif0.
             congruence.
    + lia.
Qed.
back to top