aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/2996.v
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2014-06-28 17:13:22 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2014-06-28 18:55:31 +0200
commitcad44fcfe8a129af24d4d9d1f86c8be123707744 (patch)
tree437271cc203c9725f7f5045b6c6e032183047141 /test-suite/bugs/closed/2996.v
parent1f0e44c96872196d0051618de77c4735eb447540 (diff)
Quickly fixing bug #2996: typing functions now check when referring to
a global reference that the current (goal) env contains all the section variables that the global reference expects to be present. Note that the test for inclusion might be costly: everytime a conversion happens in a section variable copied in a goal, this conversion has to be redone when referring to a constant dependent on this section variable. It is unclear to me whether we should not instead give global names to section variables so that they exist even if they are not listed in the context of the current goal. Here are two examples which are still problematic: Section A. Let B := True : Type. Definition C := eq_refl : B = True. Theorem D : Type. clearbody B. set (x := C). unfold C in x. (* inconsistent context *) or Section A. Let B : Type. exact True. Qed. Definition C := eq_refl : B = True. (* Note that this violated the Qed. *) Theorem D : Type. set (x := C). unfold C in x. (* inconsistent context *)
Diffstat (limited to 'test-suite/bugs/closed/2996.v')
-rw-r--r--test-suite/bugs/closed/2996.v30
1 files changed, 30 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/2996.v b/test-suite/bugs/closed/2996.v
new file mode 100644
index 000000000..440cda617
--- /dev/null
+++ b/test-suite/bugs/closed/2996.v
@@ -0,0 +1,30 @@
+(* Test on definitions referring to section variables that are not any
+ longer in the current context *)
+
+Section x.
+
+ Hypothesis h : forall(n : nat), n < S n.
+
+ Definition f(n m : nat)(less : n < m) : nat := n + m.
+
+ Lemma a : forall(n : nat), f n (S n) (h n) = 1 + 2 * n.
+ Proof.
+ (* XXX *) admit.
+ Qed.
+
+ Lemma b : forall(n : nat), n < 3 + n.
+ Proof.
+ clear.
+ intros n.
+ Fail assert (H := a n).
+ Abort.
+
+ Let T := True.
+ Definition p := I : T.
+
+ Lemma paradox : False.
+ Proof.
+ clear.
+ set (T := False).
+ Fail pose proof p as H.
+ Abort.