1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57
|
(********************************************************************)
(* *)
(* The Why3 Verification Platform / The Why3 Development Team *)
(* Copyright 2010-2025 -- Inria - CNRS - Paris-Saclay 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.
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 {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} :
(a -> b) -> a -> b -> a -> b.
Proof.
intros 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 {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} :
forall (f:a -> b) (x:a) (v:b) (y:a),
((y = x) -> ((set f x v y) = v)) /\ (~ (y = x) -> ((set f x v y) = (f y))).
Proof.
intros f x v y.
unfold set.
case (why_decidable_eq x y).
intros <-.
now split.
intros H.
split ; intros H'.
now elim H.
easy.
Qed.
|