aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-06-18 16:19:49 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2010-06-18 16:19:49 +0000
commit5818c354fff1671fac26345ed9d6300ac25efcb2 (patch)
tree1834ffa1bf71bb9c5c06a49f85d2a87971812b61 /test-suite/bugs
parent726c669428f1cd1f69fa8f73e0c97b34169ff665 (diff)
add in test-suite the scripts about fsetdec bugs
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@13165 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'test-suite/bugs')
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2136.v61
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2137.v52
-rw-r--r--test-suite/bugs/closed/shouldsucceed/2281.v50
3 files changed, 163 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/shouldsucceed/2136.v b/test-suite/bugs/closed/shouldsucceed/2136.v
new file mode 100644
index 000000000..d2b926f37
--- /dev/null
+++ b/test-suite/bugs/closed/shouldsucceed/2136.v
@@ -0,0 +1,61 @@
+(* Bug #2136
+
+The fsetdec tactic seems to get confused by hypotheses like
+ HeqH1 : H1 = MkEquality s0 s1 b
+If I clear them then it is able to solve my goal; otherwise it is not.
+I would expect it to be able to solve the goal even without this hypothesis
+being cleared. A small, self-contained example is below.
+
+I have coq r12238.
+
+
+Thanks
+Ian
+*)
+
+
+Require Import FSets.
+Require Import Arith.
+Require Import FSetWeakList.
+
+Module DecidableNat.
+Definition t := nat.
+Definition eq := @eq nat.
+Definition eq_refl := @refl_equal nat.
+Definition eq_sym := @sym_eq nat.
+Definition eq_trans := @trans_eq nat.
+Definition eq_dec := eq_nat_dec.
+End DecidableNat.
+
+Module NatSet := Make(DecidableNat).
+
+Module Export Dec := WDecide (NatSet).
+Import FSetDecideAuxiliary.
+
+Parameter MkEquality : forall ( s0 s1 : NatSet.t )
+ ( x : nat ),
+ NatSet.Equal s1 (NatSet.add x s0).
+
+Lemma ThisLemmaWorks : forall ( s0 s1 : NatSet.t )
+ ( a b : nat ),
+ NatSet.In a s0
+ -> NatSet.In a s1.
+Proof.
+intros.
+remember (MkEquality s0 s1 b) as H1.
+clear HeqH1.
+fsetdec.
+Qed.
+
+Lemma ThisLemmaWasFailing : forall ( s0 s1 : NatSet.t )
+ ( a b : nat ),
+ NatSet.In a s0
+ -> NatSet.In a s1.
+Proof.
+intros.
+remember (MkEquality s0 s1 b) as H1.
+fsetdec.
+(*
+Error: Tactic failure: because the goal is beyond the scope of this tactic.
+*)
+Qed. \ No newline at end of file
diff --git a/test-suite/bugs/closed/shouldsucceed/2137.v b/test-suite/bugs/closed/shouldsucceed/2137.v
new file mode 100644
index 000000000..6c2023ab7
--- /dev/null
+++ b/test-suite/bugs/closed/shouldsucceed/2137.v
@@ -0,0 +1,52 @@
+(* Bug #2137
+
+The fsetdec tactic is sensitive to which way round the arguments to <> are.
+In the small, self-contained example below, it is able to solve the goal
+if it knows that "b <> a", but not if it knows that "a <> b". I would expect
+it to be able to solve hte goal in either case.
+
+I have coq r12238.
+
+
+Thanks
+Ian
+
+*)
+
+Require Import Arith FSets FSetWeakList.
+
+Module DecidableNat.
+Definition t := nat.
+Definition eq := @eq nat.
+Definition eq_refl := @refl_equal nat.
+Definition eq_sym := @sym_eq nat.
+Definition eq_trans := @trans_eq nat.
+Definition eq_dec := eq_nat_dec.
+End DecidableNat.
+
+Module NatSet := Make(DecidableNat).
+
+Module Export NameSetDec := WDecide (NatSet).
+
+Lemma ThisLemmaWorks : forall ( s0 : NatSet.t )
+ ( a b : nat ),
+ b <> a
+ -> ~(NatSet.In a s0)
+ -> ~(NatSet.In a (NatSet.add b s0)).
+Proof.
+intros.
+fsetdec.
+Qed.
+
+Lemma ThisLemmaWasFailing : forall ( s0 : NatSet.t )
+ ( a b : nat ),
+ a <> b
+ -> ~(NatSet.In a s0)
+ -> ~(NatSet.In a (NatSet.add b s0)).
+Proof.
+intros.
+fsetdec.
+(*
+Error: Tactic failure: because the goal is beyond the scope of this tactic.
+*)
+Qed. \ No newline at end of file
diff --git a/test-suite/bugs/closed/shouldsucceed/2281.v b/test-suite/bugs/closed/shouldsucceed/2281.v
new file mode 100644
index 000000000..40948d905
--- /dev/null
+++ b/test-suite/bugs/closed/shouldsucceed/2281.v
@@ -0,0 +1,50 @@
+(** Bug #2281
+
+In the code below, coq is confused by an equality unless it is first 'subst'ed
+away, yet http://coq.inria.fr/stdlib/Coq.FSets.FSetDecide.html says
+
+ fsetdec will first perform any necessary zeta and beta reductions and will
+invoke subst to eliminate any Coq equalities between finite sets or their
+elements.
+
+I have coq r12851.
+
+*)
+
+Require Import Arith.
+Require Import FSets.
+Require Import FSetWeakList.
+
+Module DecidableNat.
+Definition t := nat.
+Definition eq := @eq nat.
+Definition eq_refl := @refl_equal nat.
+Definition eq_sym := @sym_eq nat.
+Definition eq_trans := @trans_eq nat.
+Definition eq_dec := eq_nat_dec.
+End DecidableNat.
+
+Module NatSet := Make(DecidableNat).
+
+Module Export NameSetDec := WDecide (NatSet).
+
+Lemma ThisLemmaWorks : forall ( s1 s2 : NatSet.t )
+ ( H : s1 = s2 ),
+ NatSet.Equal s1 s2.
+Proof.
+intros.
+subst.
+fsetdec.
+Qed.
+
+Import FSetDecideAuxiliary.
+
+Lemma ThisLemmaWasFailing : forall ( s1 s2 : NatSet.t )
+ ( H : s1 = s2 ),
+ NatSet.Equal s1 s2.
+Proof.
+intros.
+fsetdec.
+(* Error: Tactic failure: because the goal is beyond the scope of this tactic.
+*)
+Qed. \ No newline at end of file