aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--test-suite/bugs/closed/4785.v26
-rw-r--r--theories/Compat/Coq85.v21
-rw-r--r--theories/Vectors/VectorDef.v7
3 files changed, 50 insertions, 4 deletions
diff --git a/test-suite/bugs/closed/4785.v b/test-suite/bugs/closed/4785.v
new file mode 100644
index 000000000..b844af68a
--- /dev/null
+++ b/test-suite/bugs/closed/4785.v
@@ -0,0 +1,26 @@
+Require Import Coq.Lists.List Coq.Vectors.Vector.
+Import ListNotations.
+Check [ ]%list : list _.
+Import VectorNotations ListNotations.
+Delimit Scope vector_scope with vector.
+Check [ ]%vector : Vector.t _ _.
+Check []%vector : Vector.t _ _.
+Check [ ]%list : list _.
+Check []%list : list _.
+
+Inductive mylist A := mynil | mycons (x : A) (xs : mylist A).
+Delimit Scope mylist_scope with mylist.
+Bind Scope mylist_scope with mylist.
+Arguments mynil {_}, _.
+Arguments mycons {_} _ _.
+Notation " [] " := mynil : mylist_scope.
+Notation " [ ] " := mynil (format "[ ]") : mylist_scope.
+Notation " [ x ] " := (mycons x nil) : mylist_scope.
+Notation " [ x ; y ; .. ; z ] " := (mycons x (mycons y .. (mycons z nil) ..)) : mylist_scope.
+
+Require Import Coq.Compat.Coq85.
+
+Check []%vector : Vector.t _ _.
+Check []%mylist : mylist _.
+Check [ ]%mylist : mylist _.
+Check [ ]%list : list _.
diff --git a/theories/Compat/Coq85.v b/theories/Compat/Coq85.v
index 0ddf1acde..6e495b05c 100644
--- a/theories/Compat/Coq85.v
+++ b/theories/Compat/Coq85.v
@@ -21,4 +21,23 @@ Global Unset Regular Subst Tactic.
Global Unset Structural Injection.
Global Unset Shrink Abstract.
Global Unset Shrink Obligations.
-Global Set Refolding Reduction. \ No newline at end of file
+Global Set Refolding Reduction.
+
+(** In Coq 8.5, [] meant Vector, and [ ] meant list. Restore this
+ behavior, to allow user-defined [] to not override vector
+ notations. See https://coq.inria.fr/bugs/show_bug.cgi?id=4785. *)
+
+Require Coq.Lists.List.
+Require Coq.Vectors.VectorDef.
+Module Export Coq.
+Module Export Vectors.
+Module VectorDef.
+Module VectorNotations.
+Export Coq.Vectors.VectorDef.VectorNotations.
+Notation "[]" := (VectorDef.nil _) : vector_scope.
+End VectorNotations.
+End VectorDef.
+End Vectors.
+End Coq.
+Export Vectors.VectorDef.VectorNotations.
+Close Scope vector_scope. (* vector_scope is not opened by default *)
diff --git a/theories/Vectors/VectorDef.v b/theories/Vectors/VectorDef.v
index c69223804..095cda398 100644
--- a/theories/Vectors/VectorDef.v
+++ b/theories/Vectors/VectorDef.v
@@ -30,7 +30,7 @@ Inductive t A : nat -> Type :=
|nil : t A 0
|cons : forall (h:A) (n:nat), t A n -> t A (S n).
-Local Notation "[]" := (nil _).
+Local Notation " [ ] " := (nil _) (format "[ ]").
Local Notation "h :: t" := (cons _ h _ t) (at level 60, right associativity).
Section SCHEMES.
@@ -102,7 +102,7 @@ Definition const {A} (a:A) := nat_rect _ [] (fun n x => cons _ a n x).
Computational behavior of this function should be the same as
ocaml function. *)
-Definition nth {A} :=
+Definition nth {A} :=
fix nth_fix {m} (v' : t A m) (p : Fin.t m) {struct v'} : A :=
match p in Fin.t m' return t A m' -> A with
|Fin.F1 => caseS (fun n v' => A) (fun h n t => h)
@@ -293,7 +293,8 @@ Eval cbv delta beta in fold_right (fun h H => Datatypes.cons h H) v Datatypes.ni
End VECTORLIST.
Module VectorNotations.
-Notation "[]" := [] : vector_scope.
+Delimit Scope vector_scope with vector.
+Notation " [ ] " := [] (format "[ ]") : vector_scope.
Notation "h :: t" := (h :: t) (at level 60, right associativity)
: vector_scope.
Notation " [ x ] " := (x :: []) : vector_scope.