https://gitlab.com/nomadic-labs/mi-cho-coq
Raw File
Tip revision: 0ff13fc3a670383ce13b7772d9df85724cade8b3 authored by Arvid Jakobsson on 02 March 2021, 17:52:07 UTC
Dexter2/Proof: addLiquidity
Tip revision: 0ff13fc
dexter_verification_ep_default.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_default_correct
      (env : @proto_env dexter_definition.self_type)
      (p : data dexter_definition.parameter_ep_default_ty)
      (sto : data dexter_definition.storage_ty)
      (ret_ops : Datatypes.list (data operation))
      (ret_sto : data dexter_definition.storage_ty)
      (fuel : Datatypes.nat) :
  400 <= fuel ->
  eval_precond fuel env Def.ep_default
               (fun x => x = (ret_ops, ret_sto, tt))
               (p, (sto, tt)) <->
  Spec.ep_default env p sto ret_ops ret_sto.
Proof.
(*
 * intro Hfuel.
 * auto_fuel.
 * destruct_sto sto
 *              accounts self_is_updating_token_pool freeze_baker lq_total manager token_address token_pool xtz_pool.
 * rewrite precond_exists.
 * unfold precond_ex.
 * apply and_both_0.
 * unfold Spec.Storage.sto_selfIsUpdatingTokenPool.
 * - (* Check self_updating_token_pool is false *)
 *   apply bool_false_not.
 * - (* Check xtz_pool is set to xtz_pool + amount *)
 *   unfold Spec.ep_default, Spec.mutez_add, add.
 *   unfold Spec.Storage.sto_xtzPool.
 *   unfold Spec.Storage.sto_set_xtzPool.
 *   simpl.
 *   intuition.
 *   + destruct H.
 *     destruct H.
 *     exists x.
 *     rewrite <- H.
 *     rewrite Zplus_comm.
 *     injection H0.
 *     clear H0.
 *     intros.
 *     rewrite H0, H1.
 *     intuition reflexivity.
 *   + destruct H.
 *     destruct H.
 *     exists x.
 *     rewrite <- H.
 *     rewrite Zplus_comm.
 *     destruct H0.
 *     subst.
 *     intuition reflexivity.
 *)
Admitted.
back to top