From 5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 Mon Sep 17 00:00:00 2001 From: Stephane Glondu Date: Wed, 21 Jul 2010 09:46:51 +0200 Subject: Imported Upstream snapshot 8.3~beta0+13298 --- test-suite/bugs/closed/shouldsucceed/1951.v | 63 +++++++++++++++++++++++++++++ 1 file changed, 63 insertions(+) create mode 100644 test-suite/bugs/closed/shouldsucceed/1951.v (limited to 'test-suite/bugs/closed/shouldsucceed/1951.v') diff --git a/test-suite/bugs/closed/shouldsucceed/1951.v b/test-suite/bugs/closed/shouldsucceed/1951.v new file mode 100644 index 00000000..12c0ef9b --- /dev/null +++ b/test-suite/bugs/closed/shouldsucceed/1951.v @@ -0,0 +1,63 @@ + +(* First a simplification of the bug *) + +Set Printing Universes. + +Inductive enc (A:Type (*1*)) (* : Type.1 *) := C : A -> enc A. + +Definition id (X:Type(*5*)) (x:X) := x. + +Lemma test : let S := Type(*6 : 7*) in enc S -> S. +simpl; intros. +apply enc. +apply id. +apply Prop. +Defined. + +(* Then the original bug *) + +Require Import List. + +Inductive a : Set := (* some dummy inductive *) +b : (list a) -> a. (* i don't know if this *) + (* happens for smaller *) + (* ones *) + +Inductive sg : Type := Sg. (* single *) + +Definition ipl2 (P : a -> Type) := (* in Prop, that means P is true forall *) +fold_right (fun x => prod (P x)) sg. (* the elements of a given list *) + +Definition ind + : forall S : a -> Type, + (forall ls : list a, ipl2 S ls -> S (b ls)) -> forall s : a, S s := +fun (S : a -> Type) + (X : forall ls : list a, ipl2 S ls -> S (b ls)) => +fix ind2 (s : a) := +match s as a return (S a) with +| b l => + X l + (list_rect (fun l0 : list a => ipl2 S l0) Sg + (fun (a0 : a) (l0 : list a) (IHl : ipl2 S l0) => + pair (ind2 a0) IHl) l) +end. (* some induction principle *) + +Implicit Arguments ind [S]. + +Lemma k : a -> Type. (* some ininteresting lemma *) +intro;pattern H;apply ind;intros. + assert (K : Type). + induction ls. + exact sg. + exact sg. + exact (prod K sg). +Defined. + +Lemma k' : a -> Type. (* same lemma but with our bug *) +intro;pattern H;apply ind;intros. + apply prod. + induction ls. + exact sg. + exact sg. + exact sg. (* Proof complete *) +Defined. (* bug *) -- cgit v1.2.3