summaryrefslogtreecommitdiff
path: root/test-suite/bugs/closed/3469.v
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
commit7cfc4e5146be5666419451bdd516f1f3f264d24a (patch)
treee4197645da03dc3c7cc84e434cc31d0a0cca7056 /test-suite/bugs/closed/3469.v
parent420f78b2caeaaddc6fe484565b2d0e49c66888e5 (diff)
Imported Upstream version 8.5~beta1+dfsg
Diffstat (limited to 'test-suite/bugs/closed/3469.v')
-rw-r--r--test-suite/bugs/closed/3469.v29
1 files changed, 29 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/3469.v b/test-suite/bugs/closed/3469.v
new file mode 100644
index 00000000..b09edc65
--- /dev/null
+++ b/test-suite/bugs/closed/3469.v
@@ -0,0 +1,29 @@
+(* File reduced by coq-bug-finder from original input, then from 538 lines to 31 lines *)
+Open Scope type_scope.
+Global Set Primitive Projections.
+Set Implicit Arguments.
+Record sig (A : Type) (P : A -> Type) := exist { proj1_sig : A ; proj2_sig : P proj1_sig }.
+Notation sigT := sig (only parsing).
+Notation "{ x : A & P }" := (sigT (fun x:A => P)) : type_scope.
+Notation projT1 := proj1_sig (only parsing).
+Notation projT2 := proj2_sig (only parsing).
+Variables X : Type.
+Variable R : X -> X -> Type.
+Lemma dependent_choice :
+ (forall x:X, {y : _ & R x y}) ->
+ forall x0, {f : nat -> X & (f O = x0) * (forall n, R (f n) (f (S n)))}.
+Proof.
+ intros H x0.
+ set (f:=fix f n := match n with O => x0 | S n' => projT1 (H (f n')) end).
+ exists f.
+ split.
+ reflexivity.
+ induction n; simpl in *.
+ clear.
+ apply (proj2_sig (H x0)).
+ Undo.
+ apply @proj2_sig.
+
+
+(* Toplevel input, characters 21-31:
+Error: Found no subterm matching "proj1_sig ?206" in the current *)