summaryrefslogtreecommitdiff
path: root/test-suite/bugs/closed/5500.v
diff options
context:
space:
mode:
authorGravatar Benjamin Barenblat <bbaren@debian.org>2019-02-02 19:29:28 -0500
committerGravatar Benjamin Barenblat <bbaren@debian.org>2019-02-02 19:29:28 -0500
commit1ef7f1c0c6897535a86daa77799714e25638f5e9 (patch)
tree5bcca733632ecc84d2c6b1ee48cb2e557a7adba5 /test-suite/bugs/closed/5500.v
parent3a2fac7bcee36fd9dcb4f39a615c8ac0349abcc9 (diff)
parent9ebf44d84754adc5b64fcf612c6816c02c80462d (diff)
Updated version 8.9.0 from 'upstream/8.9.0'
Diffstat (limited to 'test-suite/bugs/closed/5500.v')
-rw-r--r--test-suite/bugs/closed/5500.v35
1 files changed, 35 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/5500.v b/test-suite/bugs/closed/5500.v
new file mode 100644
index 00000000..aa63e2ab
--- /dev/null
+++ b/test-suite/bugs/closed/5500.v
@@ -0,0 +1,35 @@
+(* Too weak check on the correctness of return clause was leading to an anomaly *)
+
+Inductive Vector A: nat -> Type :=
+ nil: Vector A O
+| cons: forall n, A -> Vector A n -> Vector A (S n).
+
+(* This could be made working with a better inference of inner return
+ predicates from the return predicate at the higher level of the
+ nested matching. Currently, we only check that it does not raise an
+ anomaly, but eventually, the "Fail" could be removed. *)
+
+Fail Definition hd_fst A x n (v: A * Vector A (S n)) :=
+ match v as v0 return match v0 with
+ (l, r) =>
+ match r in Vector _ n return match n with 0 => Type | S _ => Type end with
+ nil _ => A
+ | cons _ _ _ _ => A
+ end
+ end with
+ (_, nil _) => x
+ | (_, cons _ n hd tl) => hd
+ end.
+
+(* This is another example of failure but involving beta-reduction and
+ not iota-reduction. Thus, for this one, I don't see how it could be
+ solved by small inversion, whatever smart is small inversion. *)
+
+Inductive A : (Type->Type) -> Type := J : A (fun x => x).
+
+Fail Check fun x : nat * A (fun x => x) =>
+ match x return match x with
+ (y,z) => match z in A f return f Type with J => bool end
+ end with
+ (y,J) => true
+ end.