summaryrefslogtreecommitdiff
path: root/test-suite/bugs/closed/4622.v
diff options
context:
space:
mode:
Diffstat (limited to 'test-suite/bugs/closed/4622.v')
-rw-r--r--test-suite/bugs/closed/4622.v24
1 files changed, 24 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4622.v b/test-suite/bugs/closed/4622.v
new file mode 100644
index 00000000..ffa478cb
--- /dev/null
+++ b/test-suite/bugs/closed/4622.v
@@ -0,0 +1,24 @@
+Set Primitive Projections.
+
+Record foo : Type := bar { x : unit }.
+
+Goal forall t u, bar t = bar u -> t = u.
+Proof.
+ intros.
+ injection H.
+ trivial.
+Qed.
+(* Was: Error: Pattern-matching expression on an object of inductive type foo has invalid information. *)
+
+(** Dependent pattern-matching is ok on this one as it has eta *)
+Definition baz (x : foo) :=
+ match x as x' return x' = x' with
+ | bar u => eq_refl
+ end.
+
+Inductive foo' : Type := bar' {x' : unit; y: foo'}.
+(** Dependent pattern-matching is not ok on this one *)
+Fail Definition baz' (x : foo') :=
+ match x as x' return x' = x' with
+ | bar' u y => eq_refl
+ end.