coq-club AT inria.fr
Subject: The Coq mailing list
List archive
- From: vincent rahli <vincent.rahli AT gmail.com>
- To: coq-club AT inria.fr
- Subject: [Coq-Club] Universe polymorphism in the HoTT branch of Coq
- Date: Thu, 5 Sep 2013 19:07:16 -0400
Hi,
We're trying to use the HoTT branch of Coq (downloaded today) in order to getuniverse polymorphism.
Unfortunately we keep on running into this error [Error: Universe inconsistency],
which is the very error we were trying to avoid by using the HoTT branch of Coq.
One of our problems is that [try] does not catch this error (see the code below).
- Why doesn't [try] catch that error? - Why do we get a universe inconsistency error here and not
[Error: Unable to unify.] (without the [try])?
[Error: Unable to unify.] (without the [try])?
Thanks.
Vincent & Abhishek.
-----------------------------------------------------------------------------------------
Require Export Coq.Lists.List.
Polymorphic Fixpoint LIn (A : Type) (a:A) (l:list A) : Type :=
match l with
| nil => False
| b :: m => (b = a) + LIn A a m
end.
Polymorphic Inductive NTerm : Type :=
| cterm : NTerm
| oterm : list NTerm -> NTerm.
Polymorphic Fixpoint dummy {A B} (x : list (A * B)) : list (A * B) :=
match x with
| nil => nil
| (_, _) :: _ => nil
end.
Lemma foo :
forall v t sub vars,
LIn (nat * NTerm) (v, t) (dummy sub)
->
(
LIn (nat * NTerm) (v, t) sub
*
notT (LIn nat v vars)
).
Proof.
induction sub; simpl; intros.
destruct H.
try (apply IHsub in X).
- [Coq-Club] Universe polymorphism in the HoTT branch of Coq, vincent rahli, 09/06/2013
- Re: [Coq-Club] Universe polymorphism in the HoTT branch of Coq, Jason Gross, 09/06/2013
- Re: [Coq-Club] Universe polymorphism in the HoTT branch of Coq, Matthieu Sozeau, 09/06/2013
- Re: [Coq-Club] Universe polymorphism in the HoTT branch of Coq, Jason Gross, 09/07/2013
- Re: [Coq-Club] Universe polymorphism in the HoTT branch of Coq, vincent rahli, 09/07/2013
Archive powered by MHonArc 2.6.18.