File: bug_2136.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (61 lines) | stat: -rw-r--r-- 1,495 bytes parent folder | download | duplicates (6)
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
60
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.