aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--parsing/g_prim.ml45
-rw-r--r--parsing/g_vernac.ml42
-rw-r--r--parsing/pcoq.ml1
-rw-r--r--parsing/pcoq.mli1
-rw-r--r--test-suite/success/Notations.v4
-rw-r--r--toplevel/metasyntax.ml13
6 files changed, 18 insertions, 8 deletions
diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4
index b90e06cd3..a38043d0c 100644
--- a/parsing/g_prim.ml4
+++ b/parsing/g_prim.ml4
@@ -34,7 +34,7 @@ GEXTEND Gram
GLOBAL:
bigint natural integer identref name ident var preident
fullyqualid qualid reference dirpath ne_lstring
- ne_string string pattern_ident pattern_identref by_notation smart_global;
+ ne_string string lstring pattern_ident pattern_identref by_notation smart_global;
preident:
[ [ s = IDENT -> s ] ]
;
@@ -106,6 +106,9 @@ GEXTEND Gram
string:
[ [ s = STRING -> s ] ]
;
+ lstring:
+ [ [ s = string -> (!@loc, s) ] ]
+ ;
integer:
[ [ i = INT -> my_int_of_string (!@loc) i
| "-"; i = INT -> - my_int_of_string (!@loc) i ] ]
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index e61be53a9..0c4dbcc8d 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -1122,7 +1122,7 @@ GEXTEND Gram
idl = LIST0 ident; ":="; c = constr; b = only_parsing ->
VernacSyntacticDefinition
(id,(idl,c),local,b)
- | IDENT "Notation"; local = obsolete_locality; s = ne_lstring; ":=";
+ | IDENT "Notation"; local = obsolete_locality; s = lstring; ":=";
c = constr;
modl = [ "("; l = LIST1 syntax_modifier SEP ","; ")" -> l | -> [] ];
sc = OPT [ ":"; sc = IDENT -> sc ] ->
diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml
index 7dc02190e..007a6c767 100644
--- a/parsing/pcoq.ml
+++ b/parsing/pcoq.ml
@@ -267,6 +267,7 @@ module Prim =
let integer = gec_gen "integer"
let bigint = Gram.entry_create "Prim.bigint"
let string = gec_gen "string"
+ let lstring = Gram.entry_create "Prim.lstring"
let reference = make_gen_entry uprim "reference"
let by_notation = Gram.entry_create "by_notation"
let smart_global = Gram.entry_create "smart_global"
diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli
index 37165f6ce..fc1727fc1 100644
--- a/parsing/pcoq.mli
+++ b/parsing/pcoq.mli
@@ -137,6 +137,7 @@ module Prim :
val bigint : Bigint.bigint Gram.entry
val integer : int Gram.entry
val string : string Gram.entry
+ val lstring : string located Gram.entry
val qualid : qualid located Gram.entry
val fullyqualid : Id.t list located Gram.entry
val reference : reference Gram.entry
diff --git a/test-suite/success/Notations.v b/test-suite/success/Notations.v
index 07bbb60c4..32baeaa57 100644
--- a/test-suite/success/Notations.v
+++ b/test-suite/success/Notations.v
@@ -128,3 +128,7 @@ Notation " |- {{ a }} b" := (a=b) (no associativity, at level 10).
Goal True.
{{ exact I. }}
Qed.
+
+(* Check that we can have notations without any symbol iff they are "only printing". *)
+Fail Notation "" := (@nil).
+Notation "" := (@nil) (only printing).
diff --git a/toplevel/metasyntax.ml b/toplevel/metasyntax.ml
index 008d5cf9f..ccb9c99d7 100644
--- a/toplevel/metasyntax.ml
+++ b/toplevel/metasyntax.ml
@@ -903,8 +903,8 @@ let find_precedence lev etyps symbols =
let first_symbol =
let rec aux = function
| Break _ :: t -> aux t
- | h :: t -> h
- | [] -> assert false (* rule is known to be productive *) in
+ | h :: t -> Some h
+ | [] -> None in
aux symbols in
let last_is_terminal () =
let rec aux b = function
@@ -914,7 +914,8 @@ let find_precedence lev etyps symbols =
| [] -> b in
aux false symbols in
match first_symbol with
- | NonTerminal x ->
+ | None -> [],0
+ | Some (NonTerminal x) ->
(try match List.assoc x etyps with
| ETConstr _ ->
error "The level of the leftmost non-terminal cannot be changed."
@@ -937,11 +938,11 @@ let find_precedence lev etyps symbols =
if Option.is_empty lev then
error "A left-recursive notation must have an explicit level."
else [],Option.get lev)
- | Terminal _ when last_is_terminal () ->
+ | Some (Terminal _) when last_is_terminal () ->
if Option.is_empty lev then
([Feedback.msg_info ?loc:None ,strbrk "Setting notation at level 0."], 0)
else [],Option.get lev
- | _ ->
+ | Some _ ->
if Option.is_empty lev then error "Cannot determine the level.";
[],Option.get lev
@@ -991,7 +992,7 @@ let compute_syntax_data df modifiers =
let symbols' = remove_curly_brackets symbols in
let need_squash = not (List.equal Notation.symbol_eq symbols symbols') in
let ntn_for_grammar = make_notation_key symbols' in
- check_rule_productivity symbols';
+ if not onlyprint then check_rule_productivity symbols';
let msgs,n = find_precedence n etyps symbols' in
let innerlevel = NumLevel 200 in
let typs =