aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--interp/notation.ml3
-rw-r--r--interp/notation.mli2
-rw-r--r--interp/notation_ops.ml9
-rw-r--r--test-suite/output/Notations.v11
-rw-r--r--toplevel/metasyntax.ml4
5 files changed, 24 insertions, 5 deletions
diff --git a/interp/notation.ml b/interp/notation.ml
index 389a1c9df..ba80cf1d2 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -561,11 +561,12 @@ let interpretation_eq (vars1, t1) (vars2, t2) =
List.equal var_attributes_eq vars1 vars2 &&
Notation_ops.eq_notation_constr (List.map fst vars1, List.map fst vars2) t1 t2
-let exists_notation_in_scope scopt ntn r =
+let exists_notation_in_scope scopt ntn onlyprint r =
let scope = match scopt with Some s -> s | None -> default_scope in
try
let sc = String.Map.find scope !scope_map in
let n = String.Map.find ntn sc.notations in
+ onlyprint = n.not_onlyprinting &&
interpretation_eq n.not_interp r
with Not_found -> false
diff --git a/interp/notation.mli b/interp/notation.mli
index 2e92a00a8..303fa8c7a 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -142,7 +142,7 @@ val interp_notation_as_global_reference : Loc.t -> (global_reference -> bool) ->
(** Checks for already existing notations *)
val exists_notation_in_scope : scope_name option -> notation ->
- interpretation -> bool
+ bool -> interpretation -> bool
(** Declares and looks for scopes associated to arguments of a global ref *)
val declare_arguments_scope :
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml
index 564882153..26e61d13a 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -47,9 +47,16 @@ let compare_glob_constr f add t1 t2 = match t1,t2 with
| GHole _ | GSort _ | GLetIn _), _
-> false
+(* helper for NVar, NVar case in eq_notation_constr *)
+let get_var_ndx id vs = try Some (List.index Id.equal id vs) with Not_found -> None
+
let rec eq_notation_constr (vars1,vars2 as vars) t1 t2 = match t1, t2 with
| NRef gr1, NRef gr2 -> eq_gr gr1 gr2
-| NVar id1, NVar id2 -> Int.equal (List.index Id.equal id1 vars1) (List.index Id.equal id2 vars2)
+| NVar id1, NVar id2 -> (
+ match (get_var_ndx id1 vars1,get_var_ndx id2 vars2) with
+ | Some n,Some m -> Int.equal n m
+ | None ,None -> Id.equal id1 id2
+ | _ -> false)
| NApp (t1, a1), NApp (t2, a2) ->
(eq_notation_constr vars) t1 t2 && List.equal (eq_notation_constr vars) a1 a2
| NHole (_, _, _), NHole (_, _, _) -> true (** FIXME? *)
diff --git a/test-suite/output/Notations.v b/test-suite/output/Notations.v
index 2ccca829d..b9985a594 100644
--- a/test-suite/output/Notations.v
+++ b/test-suite/output/Notations.v
@@ -1,3 +1,14 @@
+(* Bug 5568, don't warn for notations in repeated module import *)
+
+Module foo.
+Notation compose := (fun g f => g f).
+Notation "g & f" := (compose g f) (at level 10).
+End foo.
+
+Import foo.
+Import foo.
+Import foo.
+
(**********************************************************************)
(* Notations for if and let (submitted by Roland Zumkeller) *)
diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml
index 9618cf1b2..628a829e2 100644
--- a/toplevel/metasyntax.ml
+++ b/toplevel/metasyntax.ml
@@ -1073,11 +1073,11 @@ let open_notation i (_, nobj) =
let scope = nobj.notobj_scope in
let (ntn, df) = nobj.notobj_notation in
let pat = nobj.notobj_interp in
- let fresh = not (Notation.exists_notation_in_scope scope ntn pat) in
+ let onlyprint = nobj.notobj_onlyprint in
+ let fresh = not (Notation.exists_notation_in_scope scope ntn onlyprint pat) in
let active = is_active_compat nobj.notobj_compat in
if Int.equal i 1 && fresh && active then begin
(* Declare the interpretation *)
- let onlyprint = nobj.notobj_onlyprint in
let () = Notation.declare_notation_interpretation ntn scope pat df ~onlyprint in
(* Declare the uninterpretation *)
if not nobj.notobj_onlyparse then