aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs/closed
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <mattam@mattam.org>2014-10-14 14:53:42 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2014-10-14 15:03:17 +0200
commita0f3e1fe045d069cb28af21e88ea60d6b3b79c74 (patch)
tree095437a81f9982dfbcf6340a2da28caead9f9697 /test-suite/bugs/closed
parent4c330191042c7f395bd5754a6b56cf9cac4b4514 (diff)
Fix bug #3698: stack overflow due to eta+canonical structures in
unification.
Diffstat (limited to 'test-suite/bugs/closed')
-rw-r--r--test-suite/bugs/closed/3698.v25
1 files changed, 25 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/3698.v b/test-suite/bugs/closed/3698.v
new file mode 100644
index 000000000..3c53d243e
--- /dev/null
+++ b/test-suite/bugs/closed/3698.v
@@ -0,0 +1,25 @@
+(* File reduced by coq-bug-finder from original input, then from 5479 lines to 4682 lines, then from 4214 lines to 86 lines, then from 60 lines to 25 lines *)
+(* coqc version trunk (October 2014) compiled on Oct 1 2014 18:13:54 with OCaml 4.01.0
+ coqtop version cagnode16:/afs/csail.mit.edu/u/j/jgross/coq-trunk,trunk (68846802a7be637ec805a5e374655a426b5723a5) *)
+Set Primitive Projections.
+Notation "( x ; y )" := (existT _ x y) : fibration_scope.
+Open Scope fibration_scope.
+Notation pr1 := projT1.
+Notation "x .1" := (pr1 x) (at level 3, format "x '.1'") : fibration_scope.
+Axiom ap : forall {A B:Type} (f:A -> B) {x y:A} (p:x = y), f x = f y.
+Class IsEquiv {A B : Type} (f : A -> B) := { equiv_inv : B -> A }.
+Record Equiv A B := { equiv_fun :> A -> B ; equiv_isequiv :> IsEquiv equiv_fun }.
+Global Existing Instance equiv_isequiv.
+Notation "A <~> B" := (Equiv A B) (at level 85) : equiv_scope.
+Notation "f ^-1" := (@equiv_inv _ _ f _) (at level 3, format "f '^-1'") : equiv_scope.
+Axiom IsHSet : Type -> Type.
+Local Open Scope equiv_scope.
+Record hSet := BuildhSet {setT:> Type; iss :> IsHSet setT}.
+Canonical Structure default_HSet:= fun T P => (@BuildhSet T P).
+Axiom issig_hSet: (sigT IsHSet) <~> hSet.
+Definition isequiv_ap_setT X Y : IsEquiv (@ap _ _ setT X Y).
+Proof.
+ assert (H'' : forall g : X = Y -> (issig_hSet^-1 X).1 = (issig_hSet^-1 Y).1,
+ g = g -> IsEquiv g) by admit.
+ Eval compute in (@projT1 Type IsHSet (@equiv_inv _ _ _ (equiv_isequiv _ _ issig_hSet) X)).
+ Fail apply H''. (* stack overflow *) \ No newline at end of file