aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed/4375.v
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2015-10-28 11:02:09 -0400
committerGravatar Matthieu Sozeau <matthieu.sozeau@inria.fr>2015-10-28 11:02:52 -0400
commitb5a0e384b405f64fd0854d5e88b55e8c2a159c02 (patch)
tree6f5ae6fc34acb395543c91dfd3e172d30998be7c /test-suite/bugs/closed/4375.v
parente3ec13976d39909ac6f1a82bf1b243ba8a895190 (diff)
Univs: fix bug #4375, accept universe binders on (co)-fixpoints
Diffstat (limited to 'test-suite/bugs/closed/4375.v')
-rw-r--r--test-suite/bugs/closed/4375.v106
1 files changed, 106 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4375.v b/test-suite/bugs/closed/4375.v
new file mode 100644
index 000000000..03af16535
--- /dev/null
+++ b/test-suite/bugs/closed/4375.v
@@ -0,0 +1,106 @@
+
+
+Polymorphic Fixpoint g@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ t.
+
+
+Module A.
+Polymorphic Fixpoint foo@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ match n with
+ | 0 => t
+ | S n => bar t n
+ end
+
+with bar (t : Type@{i}) (n : nat) : Type@{i} :=
+ match n with
+ | 0 => t
+ | S n => foo t n
+ end.
+End A.
+
+Module B.
+Polymorphic Fixpoint foo@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ match n with
+ | 0 => t
+ | S n => bar t n
+ end
+
+with bar@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ match n with
+ | 0 => t
+ | S n => foo t n
+ end.
+End B.
+
+Module C.
+Fail Polymorphic Fixpoint foo@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ match n with
+ | 0 => t
+ | S n => bar t n
+ end
+
+with bar@{j} (t : Type@{j}) (n : nat) : Type@{j} :=
+ match n with
+ | 0 => t
+ | S n => foo t n
+ end.
+End C.
+
+Module D.
+Polymorphic Fixpoint foo@{i j} (t : Type@{i}) (n : nat) : Type@{i} :=
+ match n with
+ | 0 => t
+ | S n => bar t n
+ end
+
+with bar@{i j} (t : Type@{j}) (n : nat) : Type@{j} :=
+ match n with
+ | 0 => t
+ | S n => foo t n
+ end.
+End D.
+
+Module E.
+Fail Polymorphic Fixpoint foo@{i j} (t : Type@{i}) (n : nat) : Type@{i} :=
+ match n with
+ | 0 => t
+ | S n => bar t n
+ end
+
+with bar@{j i} (t : Type@{j}) (n : nat) : Type@{j} :=
+ match n with
+ | 0 => t
+ | S n => foo t n
+ end.
+End E.
+
+(*
+Polymorphic Fixpoint g@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ t.
+
+Print g.
+
+Polymorphic Fixpoint a@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ t
+with b@{i} (t : Type@{i}) (n : nat) : Type@{i} :=
+ t.
+
+Print a.
+Print b.
+*)
+
+Polymorphic CoInductive foo@{i} (T : Type@{i}) : Type@{i} :=
+| A : foo T -> foo T.
+
+Polymorphic CoFixpoint cg@{i} (t : Type@{i}) : foo@{i} t :=
+ @A@{i} t (cg@{i} t).
+
+Print cg.
+
+Polymorphic CoFixpoint ca@{i} (t : Type@{i}) : foo@{i} t :=
+ @A@{i} t (@cb@{i} t)
+with cb@{i} (t : Type@{i}) : foo@{i} t :=
+ @A@{i} t (@ca@{i} t).
+
+Print ca.
+Print cb. \ No newline at end of file