(********************************************************************) (* *) (* 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. Require HighOrd. (* This file is generated by Why3's Coq-realize driver *) (* Beware! Only edit allowed sections below *) Require Import BuiltIn. Require BuiltIn. Require Import ClassicalEpsilon. (* Why3 assumption *) Definition map (a:Type) (b:Type) := (a -> b). Global Instance map_WhyType : forall (a:Type) {a_WT:WhyType a} (b:Type) {b_WT:WhyType b}, WhyType (map a b). Proof. intros. repeat split. exact (fun _ => why_inhabitant). intros x y. apply excluded_middle_informative. Qed. (* Why3 goal *) Definition set: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b}, (a -> b) -> a -> b -> (a -> b). Proof. intros a a_WT b b_WT m x y. intros x'. destruct (why_decidable_eq x x') as [H|H]. exact y. exact (m x'). Defined. (* Why3 goal *) Lemma set_def : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b}, forall (f:(a -> b)) (x:a) (v:b), forall (y:a), ((y = x) -> (((set f x v) y) = v)) /\ ((~ (y = x)) -> (((set f x v) y) = (f y))). Proof. intros a a_WT b b_WT f x v y. unfold set. case why_decidable_eq. intros <-. easy. intros H. split ; intros H'. now elim H. easy. Qed.