summaryrefslogtreecommitdiff
path: root/test-suite/bugs/closed/4161.v
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2015-11-13 11:43:34 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2015-11-13 11:43:34 +0100
commit4e76c4f01b69b77f40686e06c4544aa156efaa5a (patch)
treeaefad2e3de35f75c46729f9310d33b56d3821961 /test-suite/bugs/closed/4161.v
parent64fa31c7ee53e79b112507fb2eea27dc7648328d (diff)
parent91dbeab8eef959c3f64960909ca69d4e68c8198d (diff)
Imported Upstream version 8.5~beta3+dfsg
Diffstat (limited to 'test-suite/bugs/closed/4161.v')
-rw-r--r--test-suite/bugs/closed/4161.v27
1 files changed, 27 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4161.v b/test-suite/bugs/closed/4161.v
new file mode 100644
index 00000000..aa2b189b
--- /dev/null
+++ b/test-suite/bugs/closed/4161.v
@@ -0,0 +1,27 @@
+
+ (* Inductive t : Type -> Type := *)
+ (* | Just : forall (A : Type), t A -> t A. *)
+
+ (* Fixpoint test {A : Type} (x : t A) : t (A + unit) := *)
+ (* match x in t A return t (A + unit) with *)
+ (* | Just T x => @test T x *)
+ (* end. *)
+
+
+ Definition Type1 := Type.
+Definition Type2 := Type.
+Definition cast (x:Type2) := x:Type1.
+Axiom f: Type2 -> Prop.
+Definition A :=
+ let T := fun A:Type1 => _ in
+ fun A':Type2 =>
+ eq_refl : T A' = f A' :> Prop.
+(* Type2 <= Type1... f A -> Type1 <= Type2 *)
+
+Inductive t : Type -> Type :=
+ | Just : forall (A : Type), t A -> t A.
+
+Fixpoint test {A : Type} (x : t A) : t (A + unit) :=
+ match x in t A with
+ | Just B x => @test B x
+ end. \ No newline at end of file