https://gitlab.com/nomadic-labs/mi-cho-coq
Raw File
Tip revision: 10366a22de97e7dbababf7ca5acb9fd648bd9b70 authored by Arvid Jakobsson on 30 April 2021, 10:24:31 UTC
Merge branch 'cpmm2-verification--xtzToToken' into 'arvid@cpmm2-verification'
Tip revision: 10366a2
dexter_definition_helpers.v
(* Open Source License *)
(* Copyright (c) 2020 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. *)

Require Import syntax.
Require error.

(*+ Extract entrypoint type by name +*)
Definition is_annotation an (oan' : Datatypes.option String.string) :=
  match oan' with
    | None => false
    | Some an' => if String.string_dec an an' then true else false
  end.

Fixpoint get_ep_ty_by_name (ty: type) (ep : String.string) : Datatypes.list type :=
  match ty with
    | (or a an1 b an2) =>
      if is_annotation ep an1 then
        cons a nil
      else
        if is_annotation ep an2 then
        cons b nil
        else
          (get_ep_ty_by_name a ep) ++ (get_ep_ty_by_name b ep)
    | _ => nil
  end.

Fixpoint get_ep_ty_by_name1 (ty: type)  (ep : String.string) : Datatypes.option type :=
  match get_ep_ty_by_name ty ep with
    | cons ty nil  => Some ty
    | _ => None
  end.
back to top