diff options
author | Stephane Glondu <steph@glondu.net> | 2010-07-21 09:46:51 +0200 |
---|---|---|
committer | Stephane Glondu <steph@glondu.net> | 2010-07-21 09:46:51 +0200 |
commit | 5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 (patch) | |
tree | 631ad791a7685edafeb1fb2e8faeedc8379318ae /test-suite/bugs/closed/shouldsucceed/2137.v | |
parent | da178a880e3ace820b41d38b191d3785b82991f5 (diff) |
Imported Upstream snapshot 8.3~beta0+13298
Diffstat (limited to 'test-suite/bugs/closed/shouldsucceed/2137.v')
-rw-r--r-- | test-suite/bugs/closed/shouldsucceed/2137.v | 52 |
1 files changed, 52 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/shouldsucceed/2137.v b/test-suite/bugs/closed/shouldsucceed/2137.v new file mode 100644 index 00000000..6c2023ab --- /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 |