diff options
author | Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr> | 2015-10-02 16:27:58 +0200 |
---|---|---|
committer | Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr> | 2015-10-02 16:33:15 +0200 |
commit | 944c8de0bfe4048e0733a487e6388db4dfc9075a (patch) | |
tree | af037ad2d990da53529356fec44860ad9ca87577 /test-suite/bugs/closed/4287.v | |
parent | 16c88c9be5c37ee2e4fe04f7342365964031e7dd (diff) | |
parent | 8860362de4a26286b0cb20cf4e02edc5209bdbd1 (diff) |
Merge branch 'v8.5'
Diffstat (limited to 'test-suite/bugs/closed/4287.v')
-rw-r--r-- | test-suite/bugs/closed/4287.v | 124 |
1 files changed, 124 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/4287.v b/test-suite/bugs/closed/4287.v new file mode 100644 index 000000000..e139c5b6c --- /dev/null +++ b/test-suite/bugs/closed/4287.v @@ -0,0 +1,124 @@ + +Universe b. + +Universe c. + +Definition U : Type@{b} := Type@{c}. + +Module Type MT. + +Definition T := Prop. +End MT. + +Module M : MT. + Definition T := Type@{b}. + +Print Universes. +Fail End M. + +Set Universe Polymorphism. + +(* This is a modified version of Hurkens with all universes floating *) +Section Hurkens. + +Variable down : Type -> Type. +Variable up : Type -> Type. + +Hypothesis back : forall A, up (down A) -> A. + +Hypothesis forth : forall A, A -> up (down A). + +Hypothesis backforth : forall (A:Type) (P:A->Type) (a:A), + P (back A (forth A a)) -> P a. + +Hypothesis backforth_r : forall (A:Type) (P:A->Type) (a:A), + P a -> P (back A (forth A a)). + +(** Proof *) +Definition V : Type := forall A:Type, ((up A -> Type) -> up A -> Type) -> up A -> Type. +Definition U : Type := V -> Type. + +Definition sb (z:V) : V := fun A r a => r (z A r) a. +Definition le (i:U -> Type) (x:U) : Type := x (fun A r a => i (fun v => sb v A r a)). +Definition le' (i:up (down U) -> Type) (x:up (down U)) : Type := le (fun a:U => i (forth _ a)) (back _ x). +Definition induct (i:U -> Type) : Type := forall x:U, up (le i x) -> up (i x). +Definition WF : U := fun z => down (induct (fun a => z (down U) le' (forth _ a))). +Definition I (x:U) : Type := + (forall i:U -> Type, up (le i x) -> up (i (fun v => sb v (down U) le' (forth _ x)))) -> False. + +Lemma Omega : forall i:U -> Type, induct i -> up (i WF). +Proof. +intros i y. +apply y. +unfold le, WF, induct. +apply forth. +intros x H0. +apply y. +unfold sb, le', le. +compute. +apply backforth_r. +exact H0. +Qed. + +Lemma lemma1 : induct (fun u => down (I u)). +Proof. +unfold induct. +intros x p. +apply forth. +intro q. +generalize (q (fun u => down (I u)) p). +intro r. +apply back in r. +apply r. +intros i j. +unfold le, sb, le', le in j |-. +apply backforth in j. +specialize q with (i := fun y => i (fun v:V => sb v (down U) le' (forth _ y))). +apply q. +exact j. +Qed. + +Lemma lemma2 : (forall i:U -> Type, induct i -> up (i WF)) -> False. +Proof. +intro x. +generalize (x (fun u => down (I u)) lemma1). +intro r; apply back in r. +apply r. +intros i H0. +apply (x (fun y => i (fun v => sb v (down U) le' (forth _ y)))). +unfold le, WF in H0. +apply back in H0. +exact H0. +Qed. + +Theorem paradox : False. +Proof. +exact (lemma2 Omega). +Qed. + +End Hurkens. + +Polymorphic Record box (T : Type) := wrap {unwrap : T}. + +(* Here we instantiate to Set *) + +Fail Definition down (x : Type) : Prop := box x. +Definition up (x : Prop) : Type := x. + +Fail Definition back A : up (down A) -> A := unwrap A. + +Fail Definition forth A : A -> up (down A) := wrap A. + +Definition id {A : Type} (a : A) := a. +Definition setlt (A : Type@{i}) := + let foo := Type@{i} : Type@{j} in True. + +Definition setle (B : Type@{i}) := + let foo (A : Type@{j}) := A in foo B. + +Fail Check @setlt@{j Prop}. +Check @setlt@{Prop j}. +Check @setle@{Prop j}. + +Fail Definition foo := @setle@{j Prop}. +Definition foo := @setle@{Prop j}. |