(********************************************************************) (* *) (* The Why3 Verification Platform / The Why3 Development Team *) (* Copyright 2010-2017 -- INRIA - CNRS - Paris-Sud University *) (* *) (* This software is distributed under the terms of the GNU Lesser *) (* General Public License version 2.1, with the special exception *) (* on linking described in file LICENSE. *) (* *) (********************************************************************) (* This file is generated by Why3's Coq-realize driver *) (* Beware! Only edit allowed sections below *) Require Import BuiltIn. Require BuiltIn. (* Why3 assumption *) Definition is_none {a:Type} {a_WT:WhyType a} (o:(option a)): Prop := match o with | Init.Datatypes.None => True | (Init.Datatypes.Some _) => False end. (* Why3 goal *) Lemma is_none_spec : forall {a:Type} {a_WT:WhyType a}, forall (o:(option a)), (is_none o) <-> (o = Init.Datatypes.None). Proof. intros a a_WT o. split. now destruct o. now intros ->. Qed. Global Instance option_WhyType : forall T {T_WT : WhyType T}, WhyType (option T). Proof. split. apply @None. intros [x|] [y|] ; try (now right) ; try (now left). destruct (why_decidable_eq x y) as [E|E]. left. now apply f_equal. right. contradict E. now injection E. Qed.