Mentions légales du service

Skip to content
Snippets Groups Projects
Commit e58d9baf authored by Sylvain Dailler's avatar Sylvain Dailler
Browse files

Merge branch 'issue_290' #fix 290

parents f5daea84 23cf299e
Branches
Tags
1 merge request!116Issue 290
......@@ -30,6 +30,10 @@ IDE
* display of counterexamples in the Task view has been improved
* auto jumping to next unproved goal can now be disabled in the preferences
Realizations
* Added experimental realizations for new Set theories in both Isabelle and
Coq
Version 1.2.0, February 11, 2019
--------------------------------
......
......@@ -974,7 +974,7 @@ COQLIBS_REAL = $(addprefix lib/coq/real/, $(COQLIBS_REAL_FILES))
COQLIBS_NUMBER_FILES = Divisibility Gcd Parity Prime Coprime
COQLIBS_NUMBER = $(addprefix lib/coq/number/, $(COQLIBS_NUMBER_FILES))
COQLIBS_SET_FILES = Set
COQLIBS_SET_FILES = Set Cardinal Fset FsetInduction FsetInt FsetSum SetApp SetAppInt SetImp SetImpInt
COQLIBS_SET = $(addprefix lib/coq/set/, $(COQLIBS_SET_FILES))
COQLIBS_MAP_FILES = Map Const Occ MapPermut MapInjection
......
This diff is collapsed.
(* 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.
(* Why3 goal *)
Definition fset : forall (a:Type), Type.
Proof.
intros.
(* "apply (sig Cardinal.is_finite)." is not possible: a is not Why3Type *)
apply (sig (fun (f: a -> bool) => exists l: List.list a, List.NoDup l /\ forall e, List.In e l <-> f e = true)).
Defined.
(* Why3 goal *)
Definition mem {a:Type} {a_WT:WhyType a} : a -> fset a -> Prop.
Proof.
intros. destruct X0 as (f, P).
(* TODO remove this *)
Require set.Set.
apply (set.Set.mem X f).
Defined.
(* Why3 assumption *)
Definition infix_eqeq {a:Type} {a_WT:WhyType a} (s1:fset a) (s2:fset a) :
Prop :=
forall (x:a), mem x s1 <-> mem x s2.
(* Why3 goal *)
Lemma extensionality {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), infix_eqeq s1 s2 -> (s1 = s2).
Proof.
intros s1 s2 h1.
unfold infix_eqeq in h1. unfold mem in h1.
destruct s1, s2.
assert (x = x0).
eapply set.Set.extensionality. intro. eauto.
subst.
assert (e = e0).
(* TODO maybe provable on such property ? *)
Require Logic.ProofIrrelevance.
apply Logic.ProofIrrelevance.proof_irrelevance.
subst. reflexivity.
Qed.
(* Why3 assumption *)
Definition subset {a:Type} {a_WT:WhyType a} (s1:fset a) (s2:fset a) : Prop :=
forall (x:a), mem x s1 -> mem x s2.
(* Why3 goal *)
Lemma subset_refl {a:Type} {a_WT:WhyType a} : forall (s:fset a), subset s s.
Proof.
intros s.
destruct s. eapply set.Set.subset_refl.
Qed.
(* Why3 goal *)
Lemma subset_trans {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a) (s3:fset a), subset s1 s2 -> subset s2 s3 ->
subset s1 s3.
Proof.
intros s1 s2 s3 h1 h2.
destruct s1, s2, s3.
eapply set.Set.subset_trans; eauto.
Qed.
(* Why3 assumption *)
Definition is_empty {a:Type} {a_WT:WhyType a} (s:fset a) : Prop :=
forall (x:a), ~ mem x s.
(* Why3 goal *)
Definition empty {a:Type} {a_WT:WhyType a} : fset a.
Proof.
exists (fun x => false).
(* TODO remove this *)
Require Cardinal.
apply Cardinal.is_finite_empty. unfold set.Set.is_empty.
unfold set.Set.mem. intuition.
Defined.
(* Why3 goal *)
Lemma is_empty_empty {a:Type} {a_WT:WhyType a} : is_empty (empty : fset a).
Proof.
unfold empty, is_empty, mem, set.Set.mem. intuition.
Qed.
(* Why3 goal *)
Lemma empty_is_empty {a:Type} {a_WT:WhyType a} :
forall (s:fset a), is_empty s -> (s = (empty : fset a)).
Proof.
intros s h1.
eapply extensionality. intro. unfold empty, is_empty, mem, set.Set.mem in *.
destruct s. intuition. destruct (h1 _ H).
Qed.
(* Why3 goal *)
Definition add {a:Type} {a_WT:WhyType a} : a -> fset a -> fset a.
Proof.
intros e f.
destruct f as (f, H).
exists (map.Map.set f e true).
apply Cardinal.is_finite_add. assumption.
Defined.
(* Why3 goal *)
Lemma add_def {a:Type} {a_WT:WhyType a} :
forall (x:a) (s:fset a) (y:a), mem y (add x s) <-> mem y s \/ (y = x).
Proof.
intros x s y.
unfold mem. unfold add. destruct s.
unfold Map.set, set.Set.mem. destruct why_decidable_eq; intuition.
Qed.
(* Why3 goal *)
Lemma mem_singleton {a:Type} {a_WT:WhyType a} :
forall (x:a) (y:a), mem y (add x (empty : fset a)) -> (y = x).
Proof.
intros x y h1.
unfold empty, mem, add, Map.set, set.Set.mem in *.
destruct why_decidable_eq; inversion h1; eauto.
Qed.
(* Why3 goal *)
Definition remove {a:Type} {a_WT:WhyType a} : a -> fset a -> fset a.
Proof.
intros e f.
destruct f as (f, H).
exists (Map.set f e false). eapply Cardinal.is_finite_remove.
assumption.
Defined.
(* Why3 goal *)
Lemma remove_def {a:Type} {a_WT:WhyType a} :
forall (x:a) (s:fset a) (y:a), mem y (remove x s) <-> mem y s /\ ~ (y = x).
Proof.
intros x s y.
unfold mem, remove, set.Set.mem, Map.set. destruct s.
destruct why_decidable_eq; intuition.
Qed.
(* Why3 goal *)
Lemma add_remove {a:Type} {a_WT:WhyType a} :
forall (x:a) (s:fset a), mem x s -> ((add x (remove x s)) = s).
Proof.
intros x s h1.
apply extensionality.
intro.
unfold mem, add, remove, mem in *.
destruct s.
rewrite set.Set.add_remove; eauto.
reflexivity.
Qed.
(* Why3 goal *)
Lemma remove_add {a:Type} {a_WT:WhyType a} :
forall (x:a) (s:fset a), ((remove x (add x s)) = (remove x s)).
Proof.
intros x s.
apply extensionality.
intro.
unfold mem, add, remove in *.
destruct s.
rewrite set.Set.remove_add; eauto.
reflexivity.
Qed.
(* Why3 goal *)
Lemma subset_remove {a:Type} {a_WT:WhyType a} :
forall (x:a) (s:fset a), subset (remove x s) s.
Proof.
intros x s.
unfold mem, remove in *.
destruct s.
apply set.Set.subset_remove; eauto.
Qed.
(* Why3 goal *)
Definition union {a:Type} {a_WT:WhyType a} : fset a -> fset a -> fset a.
Proof.
intros f1 f2.
destruct f1 as (f1, H1).
destruct f2 as (f2, H2).
exists (set.Set.union f1 f2). eapply Cardinal.is_finite_union; eauto.
Defined.
(* Why3 goal *)
Lemma union_def {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a) (x:a),
mem x (union s1 s2) <-> mem x s1 \/ mem x s2.
Proof.
intros s1 s2 x.
unfold mem, union.
destruct s1, s2.
eapply set.Set.union_def.
Qed.
(* Why3 goal *)
Lemma subset_union_1 {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), subset s1 (union s1 s2).
Proof.
intros s1 s2.
unfold mem, union.
destruct s1, s2.
eapply set.Set.subset_union_1.
Qed.
(* Why3 goal *)
Lemma subset_union_2 {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), subset s2 (union s1 s2).
Proof.
intros s1 s2.
unfold mem, union.
destruct s1, s2.
eapply set.Set.subset_union_2.
Qed.
(* Why3 goal *)
Definition inter {a:Type} {a_WT:WhyType a} : fset a -> fset a -> fset a.
Proof.
intros f1 f2.
destruct f1 as (f1, H1).
destruct f2 as (f2, H2).
exists (set.Set.inter f1 f2). eapply Cardinal.is_finite_inter_left; eauto.
Defined.
(* Why3 goal *)
Lemma inter_def {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a) (x:a),
mem x (inter s1 s2) <-> mem x s1 /\ mem x s2.
Proof.
intros s1 s2 x.
unfold mem, inter.
destruct s1, s2.
eapply set.Set.inter_def.
Qed.
(* Why3 goal *)
Lemma subset_inter_1 {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), subset (inter s1 s2) s1.
Proof.
intros s1 s2.
unfold mem, inter.
destruct s1, s2.
eapply set.Set.subset_inter_1.
Qed.
(* Why3 goal *)
Lemma subset_inter_2 {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), subset (inter s1 s2) s2.
Proof.
intros s1 s2.
unfold mem, inter.
destruct s1, s2.
eapply set.Set.subset_inter_2.
Qed.
(* Why3 goal *)
Definition diff {a:Type} {a_WT:WhyType a} : fset a -> fset a -> fset a.
Proof.
intros f1 f2.
destruct f1 as (f1, H1).
destruct f2 as (f2, H2).
exists (set.Set.diff f1 f2). eapply Cardinal.is_finite_diff; eauto.
Defined.
(* Why3 goal *)
Lemma diff_def {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a) (x:a),
mem x (diff s1 s2) <-> mem x s1 /\ ~ mem x s2.
Proof.
intros s1 s2 x.
unfold mem, diff.
destruct s1, s2.
eapply set.Set.diff_def.
Qed.
(* Why3 goal *)
Lemma subset_diff {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), subset (diff s1 s2) s1.
Proof.
intros s1 s2.
unfold mem, diff.
destruct s1, s2.
eapply set.Set.subset_diff.
Qed.
(* Why3 goal *)
Definition pick {a:Type} {a_WT:WhyType a} : fset a -> a.
Proof.
intros f.
destruct f as (f, H).
apply (set.Set.pick f).
Defined.
(* Why3 goal *)
Lemma pick_def {a:Type} {a_WT:WhyType a} :
forall (s:fset a), ~ is_empty s -> mem (pick s) s.
Proof.
intros s h1.
unfold mem, pick.
destruct s.
eapply set.Set.pick_def.
intuition.
Qed.
(* Why3 assumption *)
Definition disjoint {a:Type} {a_WT:WhyType a} (s1:fset a) (s2:fset a) : Prop :=
forall (x:a), ~ mem x s1 \/ ~ mem x s2.
(* Why3 goal *)
Lemma disjoint_inter_empty {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), disjoint s1 s2 <-> is_empty (inter s1 s2).
Proof.
intros s1 s2.
unfold disjoint, mem, is_empty, inter.
destruct s1, s2.
simpl. eapply set.Set.disjoint_inter_empty.
Qed.
(* Why3 goal *)
Lemma disjoint_diff_eq {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), disjoint s1 s2 <-> ((diff s1 s2) = s1).
Proof.
intros s1 s2.
split; intros.
- apply extensionality. intro.
unfold diff, disjoint, mem, is_empty, inter in *.
destruct s1, s2.
eapply set.Set.disjoint_diff_eq in H. rewrite H. reflexivity.
- assert (forall e, mem e (diff s1 s2) <-> mem e s1).
rewrite H. intuition.
clear H.
unfold diff, disjoint, mem, is_empty, inter in *.
destruct s1, s2.
intros. apply set.Set.disjoint_diff_eq.
apply set.Set.extensionality. intro. eapply H0.
Qed.
(* Why3 goal *)
Lemma disjoint_diff_s2 {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), disjoint (diff s1 s2) s2.
Proof.
intros s1 s2.
unfold disjoint, mem, is_empty, inter.
destruct s1, s2.
eapply set.Set.disjoint_diff_s2.
Qed.
Lemma filter_NoDup: forall {A} (l: list A) f, List.NoDup l -> List.NoDup (List.filter f l).
Proof.
induction l; intros.
- constructor.
- simpl. destruct (f a); eauto. econstructor; eauto.
rewrite List.filter_In. intro. inversion H. intuition.
eapply IHl; eauto. inversion H; eauto.
eapply IHl. inversion H; eauto.
Qed.
(* Why3 goal *)
Definition filter {a:Type} {a_WT:WhyType a} :
fset a -> (a -> Init.Datatypes.bool) -> fset a.
Proof.
intros s filter.
destruct s as (f, H).
exists (fun x => filter x && f x)%bool.
destruct H as (l, Hl).
exists (List.filter filter l).
split.
apply filter_NoDup; eauto. apply Hl.
intros.
rewrite List.filter_In. rewrite Bool.andb_true_iff. destruct Hl. rewrite H0. intuition.
Defined.
(* Why3 goal *)
Lemma filter_def {a:Type} {a_WT:WhyType a} :
forall (s:fset a) (p:a -> Init.Datatypes.bool) (x:a),
mem x (filter s p) <-> mem x s /\ ((p x) = Init.Datatypes.true).
Proof.
intros s p x.
unfold mem, filter.
destruct s. unfold set.Set.mem.
rewrite Bool.andb_true_iff. intuition.
Qed.
(* Why3 goal *)
Lemma subset_filter {a:Type} {a_WT:WhyType a} :
forall (s:fset a) (p:a -> Init.Datatypes.bool), subset (filter s p) s.
Proof.
intros s p.
unfold subset, filter, mem, set.Set.mem.
destruct s.
intros.
rewrite Bool.andb_true_iff in H. intuition.
Qed.
(* Why3 goal *)
Definition map {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} :
(a -> b) -> fset a -> fset b.
Proof.
intros map fs.
destruct fs as (fs, H).
exists (set.Set.map map fs).
eapply Cardinal.is_finite_map.
assumption.
Defined.
(* Why3 goal *)
Lemma map_def {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} :
forall (f:a -> b) (u:fset a) (y:b),
mem y (map f u) <-> (exists x:a, mem x u /\ (y = (f x))).
Proof.
intros f u y.
unfold map, mem.
destruct u.
eapply set.Set.map_def.
Qed.
(* Why3 goal *)
Lemma mem_map {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} :
forall (f:a -> b) (u:fset a), forall (x:a), mem x u -> mem (f x) (map f u).
Proof.
intros f u x h1.
unfold map, mem.
destruct u.
eapply set.Set.mem_map.
assumption.
Qed.
(* Why3 goal *)
Definition cardinal {a:Type} {a_WT:WhyType a} : fset a -> Numbers.BinNums.Z.
Proof.
intros fs.
destruct fs as (fs, H).
eapply (Cardinal.cardinal fs).
Defined.
(* Why3 goal *)
Lemma cardinal_nonneg {a:Type} {a_WT:WhyType a} :
forall (s:fset a), (0%Z <= (cardinal s))%Z.
Proof.
intros s.
unfold cardinal. destruct s.
eapply Cardinal.cardinal_nonneg.
Qed.
(* Why3 goal *)
Lemma cardinal_empty {a:Type} {a_WT:WhyType a} :
forall (s:fset a), is_empty s <-> ((cardinal s) = 0%Z).
Proof.
intros s.
unfold cardinal. destruct s.
eapply Cardinal.cardinal_empty.
assumption.
Qed.
(* Why3 goal *)
Lemma cardinal_add {a:Type} {a_WT:WhyType a} :
forall (x:a), forall (s:fset a),
(mem x s -> ((cardinal (add x s)) = (cardinal s))) /\
(~ mem x s -> ((cardinal (add x s)) = ((cardinal s) + 1%Z)%Z)).
Proof.
intros x s.
unfold cardinal. destruct s.
eapply Cardinal.cardinal_add.
assumption.
Qed.
(* Why3 goal *)
Lemma cardinal_remove {a:Type} {a_WT:WhyType a} :
forall (x:a), forall (s:fset a),
(mem x s -> ((cardinal (remove x s)) = ((cardinal s) - 1%Z)%Z)) /\
(~ mem x s -> ((cardinal (remove x s)) = (cardinal s))).
Proof.
intros x s.
unfold cardinal. destruct s.
eapply Cardinal.cardinal_remove.
assumption.
Qed.
(* Why3 goal *)
Lemma cardinal_subset {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), subset s1 s2 ->
((cardinal s1) <= (cardinal s2))%Z.
Proof.
intros s1 s2 h1.
unfold cardinal. destruct s1, s2.
eapply Cardinal.cardinal_subset; assumption.
Qed.
(* Why3 goal *)
Lemma subset_eq {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), subset s1 s2 ->
((cardinal s1) = (cardinal s2)) -> (s1 = s2).
Proof.
intros s1 s2 h1 h2.
apply extensionality. intro.
unfold cardinal, subset, mem in *.
destruct s1, s2.
eapply Cardinal.subset_eq in h2; eauto.
rewrite h2. reflexivity.
Qed.
(* Why3 goal *)
Lemma cardinal1 {a:Type} {a_WT:WhyType a} :
forall (s:fset a), ((cardinal s) = 1%Z) -> forall (x:a), mem x s ->
(x = (pick s)).
Proof.
intros s h1 x h2.
unfold mem, pick in *.
destruct s.
eapply Cardinal.cardinal1; eauto.
Qed.
(* Why3 goal *)
Lemma cardinal_union {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a),
((cardinal (union s1 s2)) =
(((cardinal s1) + (cardinal s2))%Z - (cardinal (inter s1 s2)))%Z).
Proof.
intros s1 s2.
unfold cardinal, union, inter in *.
destruct s1, s2.
eapply Cardinal.cardinal_union; eauto.
Qed.
(* Why3 goal *)
Lemma cardinal_inter_disjoint {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a), disjoint s1 s2 ->
((cardinal (inter s1 s2)) = 0%Z).
Proof.
intros s1 s2 h1.
unfold cardinal, inter, disjoint, mem in *.
destruct s1, s2.
eapply Cardinal.cardinal_inter_disjoint; eauto.
Qed.
(* Why3 goal *)
Lemma cardinal_diff {a:Type} {a_WT:WhyType a} :
forall (s1:fset a) (s2:fset a),
((cardinal (diff s1 s2)) = ((cardinal s1) - (cardinal (inter s1 s2)))%Z).
Proof.
intros s1 s2.
unfold cardinal, inter, disjoint, mem in *.
destruct s1, s2.
eapply Cardinal.cardinal_diff; eauto.
Qed.
(* Why3 goal *)
Lemma cardinal_filter {a:Type} {a_WT:WhyType a} :
forall (s:fset a) (p:a -> Init.Datatypes.bool),
((cardinal (filter s p)) <= (cardinal s))%Z.
Proof.
intros s p.
unfold cardinal, filter in *.
destruct s.
unfold Cardinal.cardinal.
destruct ClassicalEpsilon.excluded_middle_informative; [| intuition].
- destruct ClassicalEpsilon.classical_indefinite_description.
destruct ClassicalEpsilon.excluded_middle_informative; [| intuition].
destruct ClassicalEpsilon.classical_indefinite_description.
specialize (a1 i0). specialize (a0 i).
destruct a0, a1.
assert (length x0 <= length x1).
{
eapply List.NoDup_incl_length; eauto.
intro. intros. rewrite H2. rewrite H0 in H3. rewrite Bool.andb_true_iff in H3.
apply H3.
}
omega.
- destruct ClassicalEpsilon.classical_indefinite_description.
destruct ClassicalEpsilon.excluded_middle_informative; [| intuition].
unfold Z.zero. omega.
Qed.
(* Why3 goal *)
Lemma cardinal_map {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} :
forall (f:a -> b) (s:fset a), ((cardinal (map f s)) <= (cardinal s))%Z.
Proof.
intros f s.
unfold cardinal, inter, disjoint, mem in *.
destruct s.
eapply Cardinal.cardinal_map; eauto.
Qed.
(* 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.
(* Why3 goal *)
Definition t : Type.
Proof.
(* Example bool *)
apply bool.
Defined.
(* Why3 goal *)
Definition p : set.Fset.fset t -> Prop.
Proof.
intros f.
apply True.
Defined.
(* Why3 goal *)
Lemma Induction :
(forall (s:set.Fset.fset t), set.Fset.is_empty s -> p s) ->
(forall (s:set.Fset.fset t), p s -> (forall (t1:t), p (set.Fset.add t1 s))) ->
forall (s:set.Fset.fset t), p s.
Proof.
intros h1 h2 s.
(* TODO make something interesting *)
unfold p. constructor.
Qed.
(* 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.
(* Why3 goal *)
Definition min_elt : set.Fset.fset Numbers.BinNums.Z -> Numbers.BinNums.Z.
Proof.
intros.
destruct X as (fs, H).
(* We use the list to define the algorithm *)
eapply ClassicalEpsilon.constructive_indefinite_description in H.
destruct H as (l, Hl).
destruct l.
- apply Z.zero. (* If the list is empty the result is unspecified *)
- apply (List.fold_left (fun acc x => if Z_le_dec acc x then acc else x) l z).
Defined.
(* Why3 goal *)
Lemma min_elt_def :
forall (s:set.Fset.fset Numbers.BinNums.Z), ~ set.Fset.is_empty s ->
set.Fset.mem (min_elt s) s /\
(forall (x:Numbers.BinNums.Z), set.Fset.mem x s -> ((min_elt s) <= x)%Z).
Proof.
intros s h1.
unfold min_elt, Fset.is_empty, Fset.mem, set.Set.mem in *.
destruct s. simpl.
destruct ClassicalEpsilon.constructive_indefinite_description as [l [Hdupl Heql]].
destruct l.
{ destruct h1. intros. intro. apply Heql in H. destruct H. }
assert (forall l z,
forall x, List.In x l -> List.fold_left (fun x1 acc : int => if Z_le_dec x1 acc then x1 else acc) l z <= x)%Z.
{
induction l0; intros.
{ destruct H. }
assert (forall l z, List.fold_left (fun x1 acc : int => if Z_le_dec x1 acc then x1 else acc) l z <= z)%Z.
{
induction l1; intros; simpl; eauto. omega.
simpl. destruct Z_le_dec. eauto. eapply Z.le_trans with a0; eauto. omega.
}
simpl. destruct Z_le_dec.
destruct (Z_le_dec z0 x0).
eapply Z.le_trans with z0. eapply H0. assumption.
simpl in H. destruct H. subst. omega.
eapply IHl0; eauto.
destruct (Z_le_dec a x0).
eapply Z.le_trans with a. eapply H0. assumption.
simpl in H. destruct H. subst. omega.
eapply IHl0; eauto.
}
assert (forall l z, List.In (List.fold_left (fun x acc => if Z_le_dec x acc then x else acc) l z) l \/
List.fold_left (fun x acc => if Z_le_dec x acc then x else acc) l z = z).
{
induction l0; intros.
+ simpl. right. reflexivity.
+ simpl. destruct Z_le_dec.
- specialize (IHl0 z0). intuition.
- specialize (IHl0 a). intuition.
}
split.
rewrite <- Heql. simpl. specialize (H0 l z). intuition.
intros. eapply Heql in H1. eapply (H (List.cons z l) z) in H1; eauto. simpl in H1.
destruct Z_le_dec. assumption.
omega.
Qed.
(* Why3 goal *)
Definition max_elt : set.Fset.fset Numbers.BinNums.Z -> Numbers.BinNums.Z.
Proof.
intros.
apply (- min_elt (Fset.map (fun x => - x) X))%Z.
Defined.
(* Why3 goal *)
Lemma max_elt_def :
forall (s:set.Fset.fset Numbers.BinNums.Z), ~ set.Fset.is_empty s ->
set.Fset.mem (max_elt s) s /\
(forall (x:Numbers.BinNums.Z), set.Fset.mem x s -> (x <= (max_elt s))%Z).
Proof.
intros s h1.
unfold max_elt.
assert (~ Fset.is_empty (Fset.map (fun x : int => (- x)%Z) s))%Z.
{ unfold Fset.is_empty in *.
intro. destruct h1. intros. intro. specialize (H (-x)%Z).
apply H. eapply Fset.mem_map. assumption.
}
specialize (min_elt_def (Fset.map (fun x => -x)%Z s) H).
intros. destruct H0.
split.
clear H1 H h1.
rewrite Fset.map_def in H0. destruct H0.
destruct H. rewrite H0. replace ( - - x)%Z with x. assumption.
ring.
intros. clear H0. clear h1 H.
assert (min_elt (Fset.map (fun x => - x) s) <= -x)%Z.
{
eapply H1; eauto. apply Fset.mem_map. assumption.
}
omega.
Qed.
Fixpoint seqZ l len : list Numbers.BinNums.Z :=
match len with
| S n => List.cons l (seqZ (l + 1)%Z n)
| O => List.nil
end.
Lemma seqZ_le: forall len x l, List.In x (seqZ l len) -> (l <= x)%Z.
Proof.
induction len; simpl; intuition.
eapply IHlen in H0; intuition.
Qed.
Lemma seqZ_le2: forall len x l, List.In x (seqZ l len) -> (x < l + Z.of_nat len)%Z.
Proof.
induction len; simpl; intuition.
- subst. zify. omega.
- eapply IHlen in H0. zify. omega.
Qed.
Lemma seqZ_rev: forall len x l, (l <= x < l + Z.of_nat len)%Z -> List.In x (seqZ l len).
Proof.
induction len; intros; simpl in *.
+ omega.
+ destruct (Z_eq_dec l x); eauto.
right. eapply IHlen; eauto. zify; omega.
Qed.
Lemma seqZ_In_iff: forall l len x, List.In x (seqZ l len) <-> (l <= x < l + Z.of_nat len)%Z.
Proof.
split.
intros. eauto using seqZ_le, seqZ_le2.
eapply seqZ_rev.
Qed.
Lemma seqZ_NoDup: forall len l, List.NoDup (seqZ l len).
Proof.
induction len; intros.
+ constructor.
+ simpl. constructor; eauto.
intro Habs. eapply seqZ_le in Habs. omega.
Qed.
Lemma seqZ_length: forall len l, List.length (seqZ l len) = len.
Proof.
induction len; eauto.
intros. simpl. rewrite IHlen. reflexivity.
Qed.
(* Why3 goal *)
Definition interval :
Numbers.BinNums.Z -> Numbers.BinNums.Z -> set.Fset.fset Numbers.BinNums.Z.
Proof.
intros l r.
exists (fun x => if Z_le_dec l x then
if Z_lt_dec x r then true
else false
else false).
destruct (Z_le_dec l r).
+ exists (seqZ l (Z.to_nat (r - l))%Z).
split.
- eapply seqZ_NoDup.
- intros.
rewrite seqZ_In_iff.
rewrite Z2Nat.id; [|omega].
destruct Z_le_dec.
* destruct Z_lt_dec. split; intros; [reflexivity|].
intuition.
intuition. inversion H.
* intuition. inversion H.
+ exists List.nil.
split.
- constructor.
- intros.
destruct Z_le_dec; try destruct Z_lt_dec; intuition.
omega.
inversion H.
inversion H.
Defined.
(* Why3 goal *)
Lemma interval_def :
forall (l:Numbers.BinNums.Z) (r:Numbers.BinNums.Z) (x:Numbers.BinNums.Z),
set.Fset.mem x (interval l r) <-> (l <= x)%Z /\ (x < r)%Z.
Proof.
intros l r x.
unfold interval, Fset.mem, set.Set.mem.
destruct Z_le_dec; try destruct Z_lt_dec; intuition; try inversion H.
Qed.
Lemma interval_is_finite: forall l r,
Cardinal.is_finite (fun x : int =>
if Z_le_dec l x then if Z_lt_dec x r then true
else false else false).
Proof.
intros.
unfold Cardinal.is_finite.
destruct (Z_le_dec l r).
+ exists (seqZ l (Z.to_nat (r - l)%Z)).
split. apply seqZ_NoDup.
intros. rewrite seqZ_In_iff.
rewrite Z2Nat.id; [|omega].
destruct Z_le_dec; try destruct Z_lt_dec; intuition; try inversion H.
+ exists nil. split. constructor.
simpl. intros. destruct Z_le_dec; try destruct Z_lt_dec; intuition.
Qed.
(* Why3 goal *)
Lemma cardinal_interval :
forall (l:Numbers.BinNums.Z) (r:Numbers.BinNums.Z),
((l <= r)%Z -> ((set.Fset.cardinal (interval l r)) = (r - l)%Z)) /\
(~ (l <= r)%Z -> ((set.Fset.cardinal (interval l r)) = 0%Z)).
Proof.
intros l r.
unfold interval, Fset.mem, set.Set.mem, Fset.cardinal, Cardinal.cardinal.
assert (H := interval_is_finite l r).
destruct ClassicalEpsilon.excluded_middle_informative; [| intuition].
destruct ClassicalEpsilon.classical_indefinite_description.
specialize (a i).
split.
+ intros.
destruct a.
assert (length (seqZ l (Z.to_nat (r - l)%Z)) = length x).
{
eapply Nat.le_antisymm.
+ eapply List.NoDup_incl_length. eapply seqZ_NoDup. intro. rewrite H2.
rewrite seqZ_In_iff. destruct Z_le_dec; try destruct Z_lt_dec; intuition.
rewrite Z2Nat.id in H5; omega.
+ eapply List.NoDup_incl_length. assumption. intro. rewrite H2.
rewrite seqZ_In_iff. destruct Z_le_dec; try destruct Z_lt_dec; intuition.
rewrite Z2Nat.id; omega.
inversion H3.
inversion H3.
}
rewrite <- H3. rewrite seqZ_length. rewrite Z2Nat.id; omega.
+ intros. destruct a.
destruct x. reflexivity.
specialize (H2 z). contradict H2. destruct Z_le_dec.
destruct Z_lt_dec. omega. intuition. intuition.
Qed.
(* 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.
(* Why3 goal *)
Definition sum {a:Type} {a_WT:WhyType a} :
set.Fset.fset a -> (a -> Numbers.BinNums.Z) -> Numbers.BinNums.Z.
Proof.
intros fs conv.
destruct fs as (f, H).
(* We use the list to define the algorithm *)
eapply ClassicalEpsilon.constructive_indefinite_description in H.
destruct H as (l, H).
apply (List.fold_left (fun acc x => conv x + acc)%Z l 0%Z).
Defined.
(* Why3 goal *)
Lemma sum_def_empty {a:Type} {a_WT:WhyType a} :
forall (s:set.Fset.fset a) (f:a -> Numbers.BinNums.Z),
set.Fset.is_empty s -> ((sum s f) = 0%Z).
Proof.
intros s f h1.
unfold sum, Fset.is_empty, Fset.mem, set.Set.mem in *. destruct s.
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct x0.
+ reflexivity.
+ destruct a0. destruct (h1 a1).
apply H0. left. reflexivity.
Qed.
Lemma fold_left_symm: forall {A} {B} l i s (conv: B -> A)
(Hsymm: forall a (b: A), s a b = s b a)
(Hassoc: forall a b c, s (s a b) c = s a (s b c)) v,
s (List.fold_left (fun acc x => s (conv x) acc) l i) v =
List.fold_left (fun acc x => s (conv x) acc) l (s i v).
Proof.
induction l; intros.
+ simpl. reflexivity.
+ simpl. rewrite IHl; eauto.
rewrite Hassoc. reflexivity.
Qed.
Lemma fold_left_iff_symm: forall {A} {B} l l1 i (s: A -> A -> A) (conv: B -> A)
(Hsymm: forall a b, s a b = s b a)
(Hassoc: forall a b c, s (s a b) c = s a (s b c))
(Heq: forall e, List.In e l <-> List.In e l1)
(Hdup: List.NoDup l) (Hdup': List.NoDup l1),
List.fold_left (fun acc x => s (conv x) acc) l i =
List.fold_left (fun acc x => s (conv x) acc) l1 i.
Proof.
induction l; intros.
+ destruct l1. reflexivity. specialize (Heq b). destruct Heq. destruct H0.
left. reflexivity.
+ simpl.
destruct (List.in_split a l1) as [l1' [l1'' Hsplit]].
- eapply Heq. left. reflexivity.
- rewrite (IHl (List.app l1' l1'')); eauto.
* rewrite Hsplit. rewrite List.fold_left_app.
rewrite List.fold_left_app. simpl. rewrite (Hsymm (conv a) (List.fold_left _ _ _)).
rewrite fold_left_symm; eauto. rewrite Hsymm. reflexivity.
* intros. rewrite List.in_app_iff. rewrite Hsplit in Heq.
specialize (Heq e). rewrite List.in_app_iff in Heq. simpl in Heq.
split; intros.
++ intuition. subst. inversion Hdup; intuition. subst. inversion Hdup; intuition.
++ destruct H. destruct Heq. assert (a = e \/ List.In e l). apply H1. auto. destruct H2; eauto.
subst. eapply List.NoDup_remove_2 in Hdup'. rewrite List.in_app_iff in Hdup'. destruct Hdup'; eauto.
destruct Heq. assert (a = e \/ List.In e l). apply H1. auto. destruct H2; eauto.
subst. eapply List.NoDup_remove_2 in Hdup'. rewrite List.in_app_iff in Hdup'. destruct Hdup'; eauto.
* inversion Hdup; auto.
* rewrite Hsplit in Hdup'. eapply List.NoDup_remove_1; eauto.
Qed.
(* Why3 goal *)
Lemma sum_add {a:Type} {a_WT:WhyType a} :
forall (s:set.Fset.fset a) (f:a -> Numbers.BinNums.Z) (x:a),
(set.Fset.mem x s -> ((sum (set.Fset.add x s) f) = (sum s f))) /\
(~ set.Fset.mem x s ->
((sum (set.Fset.add x s) f) = ((sum s f) + (f x))%Z)).
Proof.
intros s f x.
unfold sum, Fset.mem, set.Set.mem, Fset.add.
destruct s as (sf, H).
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct a0 as (Hx0dup, Hx0eq). destruct a1 as (Hx1dup, Hx1eq).
split. intros.
+ eapply fold_left_iff_symm; eauto.
* intuition.
* intuition.
* intros. rewrite Hx0eq. rewrite Hx1eq. unfold Map.set.
destruct why_decidable_eq; try subst; intuition.
+ intros.
assert (List.In x x0). { eapply Hx0eq. unfold Map.set. destruct why_decidable_eq; eauto. }
destruct (List.in_split x x0 H1) as [x0' [x0'' Hx0]].
rewrite Hx0. rewrite List.fold_left_app. simpl.
rewrite (Zplus_comm (f x)). symmetry.
erewrite <- (fold_left_iff_symm (List.app x0' x0'')); eauto.
* rewrite List.fold_left_app. rewrite fold_left_symm.
++ auto.
++ intuition.
++ intuition.
* intuition.
* intuition.
* intros. rewrite List.in_app_iff. rewrite Hx0 in Hx0eq.
specialize (Hx0eq e). rewrite List.in_app_iff in Hx0eq. simpl in Hx0eq.
unfold Map.set in *. split; intros.
++ destruct H2.
** apply Hx1eq. destruct why_decidable_eq.
-- subst. eapply List.NoDup_remove_2 in Hx0dup. intuition.
-- intuition.
** eapply Hx1eq. destruct why_decidable_eq.
-- subst. eapply List.NoDup_remove_2 in Hx0dup. intuition.
-- intuition.
++ eapply Hx1eq in H2.
destruct why_decidable_eq.
-- subst. destruct H0. assumption.
-- intuition.
* rewrite Hx0 in Hx0dup. eapply List.NoDup_remove_1 in Hx0dup; eauto.
Qed.
(* Why3 goal *)
Lemma sum_remove {a:Type} {a_WT:WhyType a} :
forall (s:set.Fset.fset a) (f:a -> Numbers.BinNums.Z) (x:a),
(set.Fset.mem x s ->
((sum (set.Fset.remove x s) f) = ((sum s f) - (f x))%Z)) /\
(~ set.Fset.mem x s -> ((sum (set.Fset.remove x s) f) = (sum s f))).
Proof.
intros s f x.
split; intros.
+ assert (s = Fset.add x (Fset.remove x s)).
{
apply Fset.extensionality. intro. rewrite Fset.add_remove; eauto. reflexivity.
}
rewrite H0 at 2. destruct (sum_add (Fset.remove x s) f x).
rewrite H2. ring.
rewrite Fset.remove_def. intuition.
+ assert (s = Fset.remove x s).
{
apply Fset.extensionality. intro.
rewrite Fset.remove_def. intuition. subst. destruct H; assumption.
}
rewrite <- H0. reflexivity.
Qed.
Lemma sum_diff {a:Type} {a_WT:WhyType a} :
forall (s1:set.Fset.fset a) (s2:set.Fset.fset a),
forall (f:a -> Numbers.BinNums.Z),
Fset.subset s1 s2 ->
(sum s2 f - sum s1 f = sum (set.Fset.diff s2 s1) f)%Z.
Proof.
intros.
assert (sum s2 f = sum (Fset.diff s2 s1) f + sum s1 f)%Z.
{
unfold sum, Fset.diff, Fset.union, Fset.subset, Fset.disjoint,
Fset.inter, Fset.mem, set.Set.mem, Fset.add in *.
destruct s1 as (sf1, H1).
destruct s2 as (sf2, H2).
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct a1 as (Hdidup, Hdieq).
destruct a0 as (Hx0dup, Hx0eq).
destruct a2 as (Hx1dup, Hx1eq).
rewrite fold_left_symm; try now intuition.
rewrite Z.add_0_l. rewrite <- List.fold_left_app.
eapply fold_left_iff_symm; eauto.
+ intuition.
+ intuition.
+ intros. rewrite List.in_app_iff. rewrite Hx1eq. rewrite Hx0eq.
rewrite Hdieq. rewrite set.Set.diff_def. unfold set.Set.mem.
split; intros.
- destruct (Bool.bool_dec (sf1 e) true); intuition.
- intuition.
+ eapply Cardinal.NoDup_app; eauto.
- intros. rewrite Hx1eq in H0. rewrite Hdieq. rewrite set.Set.diff_def.
intuition.
- intros. rewrite Hx1eq. rewrite Hdieq in H0. rewrite set.Set.diff_def in H0.
intuition.
}
omega.
Qed.
Lemma sum_union_disj {a:Type} {a_WT:WhyType a} :
forall (s1:set.Fset.fset a) (s2:set.Fset.fset a),
forall (f:a -> Numbers.BinNums.Z)
(Hdisj: set.Fset.disjoint s1 s2),
(sum (set.Fset.union s1 s2) f) = ((sum s1 f) + (sum s2 f))%Z.
Proof.
intros s1 s2 f Hdisj.
unfold sum, Fset.union, Fset.disjoint, Fset.inter, Fset.mem, set.Set.mem, Fset.add in *.
destruct s1 as (sf1, H1).
destruct s2 as (sf2, H2).
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct ClassicalEpsilon.constructive_indefinite_description.
destruct a0 as (Hundup, Huneq).
destruct a1 as (Hx0dup, Hx0eq).
destruct a2 as (Hx1dup, Hx1eq).
rewrite fold_left_symm; try now intuition.
rewrite Z.add_0_l. rewrite <- List.fold_left_app.
eapply fold_left_iff_symm; eauto.
+ intuition.
+ intuition.
+ intros. rewrite List.in_app_iff. rewrite Hx1eq. rewrite Hx0eq. rewrite Huneq.
rewrite set.Set.union_def. clear - e. intuition.
+ eapply Cardinal.NoDup_app; eauto.
* intros. intro. apply Hx0eq in H0. apply Hx1eq in H.
specialize (Hdisj e). intuition.
* intros. intro. apply Hx1eq in H0. apply Hx0eq in H.
specialize (Hdisj e). intuition.
Qed.
(* Why3 goal *)
Lemma sum_union {a:Type} {a_WT:WhyType a} :
forall (s1:set.Fset.fset a) (s2:set.Fset.fset a),
forall (f:a -> Numbers.BinNums.Z),
((sum (set.Fset.union s1 s2) f) =
(((sum s1 f) + (sum s2 f))%Z - (sum (set.Fset.inter s1 s2) f))%Z).
Proof.
intros s1 s2 f.
assert (sum (Fset.union s1 s2) f - sum (Fset.inter s1 s2) f =
(sum s1 f - sum (Fset.inter s1 s2) f) + (sum s2 f - sum (Fset.inter s1 s2) f))%Z.
{
rewrite sum_diff; [rewrite sum_diff; [rewrite sum_diff |] |].
+ assert (Fset.diff (Fset.union s1 s2) (Fset.inter s1 s2) =
Fset.union (Fset.diff s1 (Fset.inter s1 s2)) (Fset.diff s2 (Fset.inter s1 s2))).
{
eapply Fset.extensionality. intro e.
rewrite Fset.union_def. rewrite Fset.diff_def. rewrite Fset.diff_def. rewrite Fset.diff_def.
rewrite Fset.union_def. rewrite Fset.inter_def. intuition.
}
rewrite H.
pose (s1' := Fset.diff s1 (Fset.inter s1 s2)). fold s1'.
pose (s2' := Fset.diff s2 (Fset.inter s1 s2)). fold s2'.
eapply sum_union_disj.
unfold s1', s2'. intro. rewrite Fset.diff_def. rewrite Fset.diff_def.
rewrite Fset.inter_def. tauto.
+ eapply Fset.subset_inter_2.
+ eapply Fset.subset_inter_1.
+ eapply Fset.subset_trans with s1. eapply Fset.subset_inter_1.
eapply Fset.subset_union_1.
}
omega.
Qed.
(* Why3 goal *)
Lemma sum_eq {a:Type} {a_WT:WhyType a} :
forall (s:set.Fset.fset a),
forall (f:a -> Numbers.BinNums.Z) (g:a -> Numbers.BinNums.Z),
(forall (x:a), set.Fset.mem x s -> ((f x) = (g x))) ->
((sum s f) = (sum s g)).
Proof.
intros s f g h1.
unfold sum, Fset.mem, set.Set.mem in *.
destruct s as (sf, H).
destruct ClassicalEpsilon.constructive_indefinite_description.
clear H.
assert (forall e, List.In e x -> f e = g e).
{
intros. rewrite h1. reflexivity. apply a0. assumption.
}
assert (forall l (f: int -> a -> int) g a
(Heq: forall e acc, List.In e l -> f acc e = g acc e),
List.fold_left f l a = List.fold_left g l a).
{
induction l; simpl; intros; eauto.
rewrite Heq; eauto.
}
erewrite H0; eauto.
intros. simpl. rewrite H; eauto.
Qed.
(* Why3 goal *)
Lemma cardinal_is_sum {a:Type} {a_WT:WhyType a} :
forall (s:set.Fset.fset a),
((set.Fset.cardinal s) = (sum s (fun (us:a) => 1%Z))).
Proof.
intros s.
unfold sum, Fset.mem, set.Set.mem, Fset.cardinal in *.
destruct s as (sf, H).
destruct ClassicalEpsilon.constructive_indefinite_description.
unfold Cardinal.cardinal.
destruct ClassicalEpsilon.excluded_middle_informative; [|intuition].
destruct ClassicalEpsilon.classical_indefinite_description.
specialize (a1 H).
assert (Z.of_nat (length x0) = Z.of_nat (length x)).
{
erewrite (Cardinal.length_prop why_decidable_eq x nil x0); eauto.
+ simpl. rewrite Nat.sub_0_r. reflexivity.
+ apply a0.
+ constructor.
+ apply a1.
+ intros e Habs. inversion Habs.
+ intros. simpl. destruct a1, a0. rewrite H1, H3. intuition.
}
rewrite H0.
clear -x.
induction x; eauto.
assert (Z.of_nat (length (a0 :: x)) = Z.of_nat (length x) + 1)%Z.
simpl. rewrite Zpos_P_of_succ_nat. ring.
rewrite H.
rewrite IHx. rewrite @fold_left_symm; eauto.
intuition.
intuition.
Qed.
......@@ -385,7 +385,7 @@ set (P := fun (x:a) => mem x s /\ y = f x).
assert (inhabited a).
destruct a_WT.
exact (inhabits why_inhabitant).
set (x := epsilon H P).
set (x := epsilon H P).
destruct b_WT.
destruct (why_decidable_eq y (f x)).
exact (s x).
......
(* 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.
(* Why3 goal *)
Definition elt : Type.
Proof.
(* TODO find something more interesting *)
apply bool.
Defined.
(* Why3 goal *)
Definition set : Type.
Proof.
(* TODO find something more interesting *)
apply bool.
Defined.
(* Why3 goal *)
Definition to_fset : set -> set.Fset.fset elt.
Proof.
(* TODO find something more interesting *)
intros. exists (fun _ => false).
exists nil. split. constructor. intros. split.
intros. destruct H. intro. inversion H.
Defined.
(* Why3 goal *)
Definition choose : set -> elt.
Proof.
intros.
apply true.
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.
intros s h1.
destruct h1. unfold set.Fset.is_empty. intros.
unfold to_fset. simpl. unfold set.Set.mem.
intuition.
Qed.
(* 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.
(* TODO find something more interesting. *)
apply bool.
Defined.
(* Why3 goal *)
Definition to_fset : set -> set.Fset.fset Numbers.BinNums.Z.
Proof.
(* TODO find something more interesting. *)
intros. exists (fun _ => false).
exists nil. intuition.
constructor. inversion H.
Defined.
(* Why3 goal *)
Definition choose : set -> Numbers.BinNums.Z.
Proof.
(* TODO find something more interesting. *)
intros. apply Z.zero.
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.
intros s h1.
destruct h1. unfold to_fset, Fset.is_empty, Fset.mem, set.Set.mem.
intuition.
Qed.
(* 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.
(* Why3 goal *)
Definition elt : Type.
Proof.
(* TODO find something more interesting. *)
apply bool.
Defined.
(* Why3 goal *)
Definition set : Type.
Proof.
(* TODO find something more interesting. *)
apply bool.
Defined.
(* Why3 goal *)
Definition to_fset : set -> set.Fset.fset elt.
Proof.
(* TODO find something more interesting. *)
intros. exists (fun _ => false).
exists nil. intuition.
constructor. inversion H.
Defined.
(* Why3 goal *)
Definition choose : set -> elt.
Proof.
(* TODO find something more interesting. *)
intros. apply true.
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.
intros s h1.
destruct h1. unfold to_fset, Fset.is_empty, Fset.mem, set.Set.mem.
intuition.
Qed.
(* 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.SetImp.
(* Why3 goal *)
Definition set : Type.
Proof.
(* TODO find something more interesting. *)
apply bool.
Defined.
(* Why3 goal *)
Definition to_fset : set -> set.Fset.fset Numbers.BinNums.Z.
Proof.
(* TODO find something more interesting. *)
intros. exists (fun _ => false).
exists nil. intuition.
constructor. inversion H.
Defined.
(* Why3 goal *)
Definition choose : set -> Numbers.BinNums.Z.
Proof.
(* TODO find something more interesting. *)
intros. apply Z.zero.
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.
intros s h1.
destruct h1. unfold to_fset, Fset.is_empty, Fset.mem, set.Set.mem.
intuition.
Qed.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Please register or to comment