aboutsummaryrefslogtreecommitdiffhomepage
path: root/test-suite/bugs
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2015-04-15 09:52:13 +0200
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2015-04-15 09:52:13 +0200
commit148cf78a4d85ec56818a8ff00719a775670950b9 (patch)
treeea4540bb896b3bbedb7c41b80fcf7e0ff1cd04aa /test-suite/bugs
parent429f493997e34bfaac930c68bf6b267a5b9640ee (diff)
parent6f40831dc1d0fecfbaf9fbc8116da0e74b6e8726 (diff)
Merge branch 'v8.5'
Diffstat (limited to 'test-suite/bugs')
-rw-r--r--test-suite/bugs/closed/3199.v18
-rw-r--r--test-suite/bugs/closed/3590.v7
-rw-r--r--test-suite/bugs/closed/4120.v5
-rw-r--r--test-suite/bugs/closed/4121.v15
4 files changed, 40 insertions, 5 deletions
diff --git a/test-suite/bugs/closed/3199.v b/test-suite/bugs/closed/3199.v
new file mode 100644
index 000000000..08bf62493
--- /dev/null
+++ b/test-suite/bugs/closed/3199.v
@@ -0,0 +1,18 @@
+Axiom P : nat -> Prop.
+Axiom admit : forall n : nat, P n -> P n -> n = S n.
+Axiom foo : forall n, P n.
+
+Create HintDb bar.
+Hint Extern 3 => symmetry : bar.
+Hint Resolve admit : bar.
+Hint Immediate foo : bar.
+
+Lemma qux : forall n : nat, n = S n.
+Proof.
+intros n.
+eauto with bar.
+Defined.
+
+Goal True.
+pose (e := eq_refl (qux 0)); unfold qux in e.
+match type of e with context [eq_sym] => fail 1 | _ => idtac end.
diff --git a/test-suite/bugs/closed/3590.v b/test-suite/bugs/closed/3590.v
index 3440f0e80..3ef9270d4 100644
--- a/test-suite/bugs/closed/3590.v
+++ b/test-suite/bugs/closed/3590.v
@@ -1,13 +1,10 @@
-Require Import TestSuite.admit.
-(* Set Primitive Projections. *)
Set Implicit Arguments.
Record prod A B := pair { fst : A ; snd : B }.
Definition idS := Set.
-Goal forall x y : prod Set Set, fst x = fst y.
+Goal forall x y : prod Set Set, forall H : fst x = fst y, fst x = fst y.
intros.
change (@fst _ _ ?z) with (@fst Set idS z) at 2.
- Unshelve.
- admit.
+ apply H.
Qed.
(* Toplevel input, characters 20-58:
diff --git a/test-suite/bugs/closed/4120.v b/test-suite/bugs/closed/4120.v
new file mode 100644
index 000000000..00db8f7f3
--- /dev/null
+++ b/test-suite/bugs/closed/4120.v
@@ -0,0 +1,5 @@
+Definition id {T} (x : T) := x.
+Goal sigT (fun x => id x)%type.
+ change (fun x => ?f x) with f.
+ exists Type. exact Set.
+Defined. (* Error: Attempt to save a proof with shelved goals (in proof Unnamed_thm) *) \ No newline at end of file
diff --git a/test-suite/bugs/closed/4121.v b/test-suite/bugs/closed/4121.v
new file mode 100644
index 000000000..5f8c411ca
--- /dev/null
+++ b/test-suite/bugs/closed/4121.v
@@ -0,0 +1,15 @@
+(* -*- coq-prog-args: ("-emacs" "-nois") -*- *)
+(* File reduced by coq-bug-finder from original input, then from 830 lines to 47 lines, then from 25 lines to 11 lines *)
+(* coqc version 8.5beta1 (March 2015) compiled on Mar 11 2015 18:51:36 with OCaml 4.01.0
+ coqtop version cagnode15:/afs/csail.mit.edu/u/j/jgross/coq-8.5,v8.5 (8dbfee5c5f897af8186cb1bdfb04fd4f88eca677) *)
+
+Set Universe Polymorphism.
+Class Contr_internal (A : Type) := BuildContr { center : A }.
+Arguments center A {_}.
+Class Contr (A : Type) : Type := Contr_is_trunc : Contr_internal A.
+Hint Extern 0 => progress change Contr_internal with Contr in * : typeclass_instances.
+Definition contr_paths_contr0 {A} `{Contr A} : Contr A := {| center := center A |}.
+Instance contr_paths_contr1 {A} `{Contr A} : Contr A := {| center := center A |}.
+Check @contr_paths_contr0@{i}.
+Check @contr_paths_contr1@{i}. (* Error: Universe instance should have length 2 *)
+(** It should have length 1, just like contr_paths_contr0 *) \ No newline at end of file