File: SetAppInt.v

package info (click to toggle)
why3 1.8.2-3
  • links: PTS, VCS
  • area: main
  • in suites: forky, sid
  • size: 45,028 kB
  • sloc: xml: 185,443; ml: 111,224; ansic: 3,998; sh: 2,578; makefile: 2,568; java: 865; python: 720; javascript: 290; lisp: 205; pascal: 173
file content (59 lines) | stat: -rw-r--r-- 1,565 bytes parent folder | download
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
58
59
(********************************************************************)
(*                                                                  *)
(*  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 int.Int.
Require set.Fset.
Require set.FsetInt.
Require set.SetApp.

(* Why3 goal *)
Definition set : Type.
Proof.
exact (Fset.fset Z).
Defined.

(* Why3 goal *)
Definition to_fset : set -> set.Fset.fset Numbers.BinNums.Z.
Proof.
exact (fun x => x).
Defined.

(* Why3 goal *)
Definition mk : set.Fset.fset Numbers.BinNums.Z -> set.
Proof.
exact (fun x => x).
Defined.

(* Why3 goal *)
Lemma mk'spec :
  forall (s:set.Fset.fset Numbers.BinNums.Z), ((to_fset (mk s)) = s).
Proof.
trivial.
Qed.

(* Why3 goal *)
Definition choose : set -> Numbers.BinNums.Z.
Proof.
exact Fset.pick.
Defined.

(* Why3 goal *)
Lemma choose'spec :
  forall (s:set), ~ set.Fset.is_empty (to_fset s) ->
  set.Fset.mem (choose s) (to_fset s).
Proof.
apply Fset.pick_def.
Qed.