aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorEmilio Jesus Gallego Arias2020-04-22 13:17:37 +0200
committerEmilio Jesus Gallego Arias2020-04-22 13:17:37 +0200
commit0c4775fcb48207c96752b81c76f67c17e5fd3c99 (patch)
treed29cfaf78ff92638b655d7fb148eae47f14559ca
parent3b7f9271609bc4a008a1c754fb0cc00d932ec82f (diff)
parent33aaa5d64771c3644de4771338c371ce6b865647 (diff)
Merge PR #11694: Support printing argument-free abbreviations in custom entries with a global rule
Reviewed-by: ejgallego Ack-by: gares
-rw-r--r--interp/constrextern.ml28
-rw-r--r--test-suite/output/Notations4.out16
-rw-r--r--test-suite/output/Notations4.v7
3 files changed, 34 insertions, 17 deletions
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 7a14ca3e48..a37bac3275 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -282,9 +282,9 @@ let insert_pat_alias ?loc p = function
| Anonymous -> p
| Name _ as na -> CAst.make ?loc @@ CPatAlias (p,(CAst.make ?loc na))
-let rec insert_coercion ?loc l c = match l with
+let rec insert_entry_coercion ?loc l c = match l with
| [] -> c
- | (inscope,ntn)::l -> CAst.make ?loc @@ CNotation (Some inscope,ntn,([insert_coercion ?loc l c],[],[],[]))
+ | (inscope,ntn)::l -> CAst.make ?loc @@ CNotation (Some inscope,ntn,([insert_entry_coercion ?loc l c],[],[],[]))
let rec insert_pat_coercion ?loc l c = match l with
| [] -> c
@@ -453,7 +453,8 @@ let rec extern_cases_pattern_in_scope (custom,scopes as allscopes) vars pat =
with No_match ->
let loc = pat.CAst.loc in
match DAst.get pat with
- | PatVar (Name id) when entry_has_ident custom -> CAst.make ?loc (CPatAtom (Some (qualid_of_ident ?loc id)))
+ | PatVar (Name id) when entry_has_global custom || entry_has_ident custom ->
+ CAst.make ?loc (CPatAtom (Some (qualid_of_ident ?loc id)))
| pat ->
match availability_of_entry_coercion custom InConstrEntrySomeLevel with
| None -> raise No_match
@@ -615,6 +616,10 @@ let is_projection nargs r =
let is_hole = function CHole _ | CEvar _ -> true | _ -> false
+let isCRef_no_univ = function
+ | CRef (_,None) -> true
+ | _ -> false
+
let is_significant_implicit a =
not (is_hole (a.CAst.v))
@@ -849,7 +854,7 @@ let extern_possible_prim_token (custom,scopes) r =
| Some coercion ->
match availability_of_prim_token n sc scopes with
| None -> raise No_match
- | Some key -> insert_coercion coercion (insert_delimiters (CAst.make ?loc:(loc_of_glob_constr r) @@ CPrim n) key)
+ | Some key -> insert_entry_coercion coercion (insert_delimiters (CAst.make ?loc:(loc_of_glob_constr r) @@ CPrim n) key)
let filter_enough_applied nargs l =
match nargs with
@@ -931,7 +936,8 @@ let rec extern inctx ?impargs scopes vars r =
match DAst.get r with
| GRef (ref,us) when entry_has_global (fst scopes) -> CAst.make ?loc (extern_ref vars ref us)
- | GVar id when entry_has_ident (fst scopes) -> CAst.make ?loc (extern_var ?loc id)
+ | GVar id when entry_has_global (fst scopes) || entry_has_ident (fst scopes) ->
+ CAst.make ?loc (extern_var ?loc id)
| c ->
@@ -1081,7 +1087,7 @@ let rec extern inctx ?impargs scopes vars r =
| GFloat f -> extern_float f (snd scopes)
- in insert_coercion coercion (CAst.make ?loc c)
+ in insert_entry_coercion coercion (CAst.make ?loc c)
and extern_typ ?impargs (subentry,(_,scopes)) =
extern true ?impargs (subentry,(Notation.current_type_scope_name (),scopes))
@@ -1279,14 +1285,11 @@ and extern_notation (custom,scopes as allscopes) vars t rules =
pi3 (extern_local_binder (subentry,(scopt,scl@scopes')) vars bl))
binderlists in
let c = make_notation loc specific_ntn (l,ll,bl,bll) in
- let c = insert_coercion coercion (insert_delimiters c key) in
+ let c = insert_entry_coercion coercion (insert_delimiters c key) in
let args = fill_arg_scopes args argsscopes allscopes in
let args = extern_args (extern true) vars args in
CAst.make ?loc @@ extern_applied_notation nallargs argsimpls c args)
| SynDefRule kn ->
- match availability_of_entry_coercion custom InConstrEntrySomeLevel with
- | None -> raise No_match
- | Some coercion ->
let l =
List.map (fun (c,(subentry,(scopt,scl))) ->
extern true (subentry,(scopt,scl@snd scopes)) vars c)
@@ -1296,7 +1299,10 @@ and extern_notation (custom,scopes as allscopes) vars t rules =
let args = fill_arg_scopes args argsscopes allscopes in
let args = extern_args (extern true) vars args in
let c = CAst.make ?loc @@ extern_applied_syntactic_definition nallargs argsimpls (a,cf) l args in
- insert_coercion coercion c
+ if isCRef_no_univ c.CAst.v && entry_has_global custom then c
+ else match availability_of_entry_coercion custom InConstrEntrySomeLevel with
+ | None -> raise No_match
+ | Some coercion -> insert_entry_coercion coercion c
with
No_match -> extern_notation allscopes vars t rules
diff --git a/test-suite/output/Notations4.out b/test-suite/output/Notations4.out
index e121b5e86c..f48eaac4c9 100644
--- a/test-suite/output/Notations4.out
+++ b/test-suite/output/Notations4.out
@@ -14,6 +14,10 @@ Entry constr:myconstr is
: nat
[<< # 0 >>]
: option nat
+[b + c]
+ : nat
+fun a : nat => [a + a]
+ : nat -> nat
[1 {f 1}]
: Expr
fun (x : nat) (y z : Expr) => [1 + y z + {f x}]
@@ -81,18 +85,18 @@ fun x : nat => [x]
: nat -> nat
∀ x : nat, x = x
: Prop
-File "stdin", line 219, characters 0-160:
+File "stdin", line 226, characters 0-160:
Warning: Notation "∀ _ .. _ , _" was already defined with a different format
in scope type_scope. [notation-incompatible-format,parsing]
∀x : nat,x = x
: Prop
-File "stdin", line 232, characters 0-60:
+File "stdin", line 239, characters 0-60:
Warning: Notation "_ %%% _" was already defined with a different format.
[notation-incompatible-format,parsing]
-File "stdin", line 236, characters 0-64:
+File "stdin", line 243, characters 0-64:
Warning: Notation "_ %%% _" was already defined with a different format.
[notation-incompatible-format,parsing]
-File "stdin", line 241, characters 0-62:
+File "stdin", line 248, characters 0-62:
Warning: Lonely notation "_ %%%% _" was already defined with a different
format. [notation-incompatible-format,parsing]
3 %% 4
@@ -101,9 +105,9 @@ format. [notation-incompatible-format,parsing]
: nat
3 %% 4
: nat
-File "stdin", line 269, characters 0-61:
+File "stdin", line 276, characters 0-61:
Warning: The format modifier is irrelevant for only parsing rules.
[irrelevant-format-only-parsing,parsing]
-File "stdin", line 273, characters 0-63:
+File "stdin", line 280, characters 0-63:
Warning: The only parsing modifier has no effect in Reserved Notation.
[irrelevant-reserved-notation-only-parsing,parsing]
diff --git a/test-suite/output/Notations4.v b/test-suite/output/Notations4.v
index 1cf0d919b1..4d4b37a8b2 100644
--- a/test-suite/output/Notations4.v
+++ b/test-suite/output/Notations4.v
@@ -22,6 +22,13 @@ Notation "<< x >>" := x (in custom myconstr at level 3, x custom anotherconstr a
Notation "# x" := (Some x) (in custom anotherconstr at level 8, x constr at level 9).
Check [ << # 0 >> ].
+(* Now check with global *)
+
+Axiom c : nat.
+Notation "x" := x (in custom myconstr at level 0, x global).
+Check [ b + c ].
+Check fun a => [ a + a ].
+
End A.
Module B.