aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorHugo Herbelin2018-10-23 16:33:39 +0200
committerHugo Herbelin2020-05-13 22:37:01 +0200
commit466e6737de8772f46f08ea8e38fda196993597c0 (patch)
tree93b5bf81ba3be2f0094a67546366b9ddd98c333f
parent91b5990e724acc863a5dba66acc33fd698ac26f0 (diff)
Extending support for mixing binders and terms in abbreviations.
-rw-r--r--interp/constrintern.ml11
-rw-r--r--interp/syntax_def.ml16
-rw-r--r--interp/syntax_def.mli8
-rw-r--r--test-suite/bugs/closed/bug_7903.v2
-rw-r--r--test-suite/output/Notations4.out8
-rw-r--r--test-suite/output/Notations4.v10
-rw-r--r--vernac/metasyntax.ml23
7 files changed, 41 insertions, 37 deletions
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index f82783f47d..9d0552817f 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -976,10 +976,6 @@ let split_by_type_pat ?loc ids subst =
assert (terms = [] && termlists = []);
subst
-let make_subst ids l =
- let fold accu (id, scopes) a = Id.Map.add id (a, scopes) accu in
- List.fold_left2 fold Id.Map.empty ids l
-
let intern_notation intern env ntnvars loc ntn fullargs =
(* Adjust to parsing of { } *)
let ntn,fullargs = contract_curly_brackets ntn fullargs in
@@ -1113,8 +1109,7 @@ let intern_qualid ?(no_secvar=false) qid intern env ntnvars us args =
if List.length args < nids then error_not_enough_arguments ?loc;
let args1,args2 = List.chop nids args in
check_no_explicitation args1;
- let terms = make_subst ids (List.map fst args1) in
- let subst = (terms, Id.Map.empty, Id.Map.empty, Id.Map.empty) in
+ let subst = split_by_type ids (List.map fst args1,[],[],[]) in
let infos = (Id.Map.empty, env) in
let c = instantiate_notation_constr loc intern intern_cases_pattern_as_binder ntnvars subst infos c in
let loc = c.loc in
@@ -1624,8 +1619,8 @@ let drop_notations_pattern looked_for genv =
let nvars = List.length vars in
if List.length pats < nvars then error_not_enough_arguments ?loc:qid.loc;
let pats1,pats2 = List.chop nvars pats in
- let subst = make_subst vars pats1 in
- let idspl1 = List.map (in_not false qid.loc scopes (subst, Id.Map.empty) []) args in
+ let subst = split_by_type_pat vars (pats1,[]) in
+ let idspl1 = List.map (in_not false qid.loc scopes subst []) args in
let (_,argscs) = find_remaining_scopes pats1 pats2 g in
Some (g, idspl1, List.map2 (in_pat_sc scopes) argscs pats2)
| _ -> raise Not_found
diff --git a/interp/syntax_def.ml b/interp/syntax_def.ml
index 7184f5ea29..bd3e234a91 100644
--- a/interp/syntax_def.ml
+++ b/interp/syntax_def.ml
@@ -8,7 +8,6 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-open Util
open Pp
open CErrors
open Names
@@ -82,16 +81,9 @@ let in_syntax_constant : (bool * syndef) -> obj =
subst_function = subst_syntax_constant;
classify_function = classify_syntax_constant }
-type syndef_interpretation = (Id.t * subscopes) list * notation_constr
-
-(* Coercions to the general format of notation that also supports
- variables bound to list of expressions *)
-let in_pat (ids,ac) = (List.map (fun (id,sc) -> (id,((Constrexpr.InConstrEntrySomeLevel,sc),NtnTypeConstr))) ids,ac)
-let out_pat (ids,ac) = (List.map (fun (id,((_,sc),typ)) -> (id,sc)) ids,ac)
-
let declare_syntactic_definition ~local deprecation id ~onlyparsing pat =
let syndef =
- { syndef_pattern = in_pat pat;
+ { syndef_pattern = pat;
syndef_onlyparsing = onlyparsing;
syndef_deprecation = deprecation;
}
@@ -106,14 +98,12 @@ let warn_deprecated_syntactic_definition =
let search_syntactic_definition ?loc kn =
let syndef = KNmap.find kn !syntax_table in
- let def = out_pat syndef.syndef_pattern in
Option.iter (fun d -> warn_deprecated_syntactic_definition ?loc (kn,d)) syndef.syndef_deprecation;
- def
+ syndef.syndef_pattern
let search_filtered_syntactic_definition ?loc filter kn =
let syndef = KNmap.find kn !syntax_table in
- let def = out_pat syndef.syndef_pattern in
- let res = filter def in
+ let res = filter syndef.syndef_pattern in
if Option.has_some res then
Option.iter (fun d -> warn_deprecated_syntactic_definition ?loc (kn,d)) syndef.syndef_deprecation;
res
diff --git a/interp/syntax_def.mli b/interp/syntax_def.mli
index 8b323462a1..66a3132f2a 100644
--- a/interp/syntax_def.mli
+++ b/interp/syntax_def.mli
@@ -13,12 +13,10 @@ open Notation_term
(** Syntactic definitions. *)
-type syndef_interpretation = (Id.t * subscopes) list * notation_constr
-
val declare_syntactic_definition : local:bool -> Deprecation.t option -> Id.t ->
- onlyparsing:bool -> syndef_interpretation -> unit
+ onlyparsing:bool -> interpretation -> unit
-val search_syntactic_definition : ?loc:Loc.t -> KerName.t -> syndef_interpretation
+val search_syntactic_definition : ?loc:Loc.t -> KerName.t -> interpretation
val search_filtered_syntactic_definition : ?loc:Loc.t ->
- (syndef_interpretation -> 'a option) -> KerName.t -> 'a option
+ (interpretation -> 'a option) -> KerName.t -> 'a option
diff --git a/test-suite/bugs/closed/bug_7903.v b/test-suite/bugs/closed/bug_7903.v
index 55c7ee99a7..18e1884ca7 100644
--- a/test-suite/bugs/closed/bug_7903.v
+++ b/test-suite/bugs/closed/bug_7903.v
@@ -1,4 +1,4 @@
(* Slightly improving interpretation of Ltac subterms in notations *)
Notation bar x f := (let z := ltac:(exact 1) in (fun x : nat => f)).
-Check bar x (x + x).
+Check fun x => bar x (x + x).
diff --git a/test-suite/output/Notations4.out b/test-suite/output/Notations4.out
index f48eaac4c9..9cb019ca56 100644
--- a/test-suite/output/Notations4.out
+++ b/test-suite/output/Notations4.out
@@ -111,3 +111,11 @@ Warning: The format modifier is irrelevant for only parsing rules.
File "stdin", line 280, characters 0-63:
Warning: The only parsing modifier has no effect in Reserved Notation.
[irrelevant-reserved-notation-only-parsing,parsing]
+fun x : nat => U (S x)
+ : nat -> nat
+V tt
+ : unit * (unit -> unit)
+fun x : nat => V x
+ : forall x : nat, nat * (?T -> ?T)
+where
+?T : [x : nat x0 : ?T |- Type] (x0 cannot be used)
diff --git a/test-suite/output/Notations4.v b/test-suite/output/Notations4.v
index 4d4b37a8b2..b3270d4f92 100644
--- a/test-suite/output/Notations4.v
+++ b/test-suite/output/Notations4.v
@@ -280,3 +280,13 @@ Notation "###" := 0 (at level 0, only parsing, format "###").
Reserved Notation "##" (at level 0, only parsing, format "##").
End N.
+
+Module O.
+
+Notation U t := (match t with 0 => 0 | S t => t | _ => 0 end).
+Check fun x => U (S x).
+Notation V t := (t,fun t => t).
+Check V tt.
+Check fun x : nat => V x.
+
+End O.
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index 3b9c771b93..32affd562f 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -1073,12 +1073,12 @@ let make_internalization_vars recvars mainvars typs =
let extratyps = List.map (fun (x,y) -> (y,List.assoc x maintyps)) recvars in
maintyps @ extratyps
-let make_interpretation_type isrec isonlybinding = function
+let make_interpretation_type isrec isonlybinding default_if_binding = function
(* Parsed as constr list *)
| ETConstr (_,None,_) when isrec -> NtnTypeConstrList
- (* Parsed as constr, but interpreted as a binder: default is to parse it as an ident only *)
+ (* Parsed as constr, but interpreted as a binder *)
| ETConstr (_,Some bk,_) -> NtnTypeBinder (NtnBinderParsedAsConstr bk)
- | ETConstr (_,None,_) when isonlybinding -> NtnTypeBinder (NtnBinderParsedAsConstr AsIdent)
+ | ETConstr (_,None,_) when isonlybinding -> NtnTypeBinder (NtnBinderParsedAsConstr default_if_binding)
(* Parsed as constr, interpreted as constr *)
| ETConstr (_,None,_) -> NtnTypeConstr
(* Others *)
@@ -1096,7 +1096,9 @@ let subentry_of_constr_prod_entry = function
| ETConstr (InConstrEntry,_,_) -> InConstrEntrySomeLevel
| _ -> InConstrEntrySomeLevel
-let make_interpretation_vars recvars allvars typs =
+let make_interpretation_vars
+ (* For binders, default is to parse only as an ident *) ?(default_if_binding=AsIdent)
+ recvars allvars typs =
let eq_subscope (sc1, l1) (sc2, l2) =
Option.equal String.equal sc1 sc2 &&
List.equal String.equal l1 l2
@@ -1113,7 +1115,7 @@ let make_interpretation_vars recvars allvars typs =
Id.Map.mapi (fun x (isonlybinding, sc) ->
let typ = Id.List.assoc x typs in
((subentry_of_constr_prod_entry typ,sc),
- make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding typ)) mainvars
+ make_interpretation_type (Id.List.mem_assoc x recvars) isonlybinding default_if_binding typ)) mainvars
let check_rule_productivity l =
if List.for_all (function NonTerminal _ | Break _ -> true | _ -> false) l then
@@ -1792,8 +1794,8 @@ let try_interp_name_alias = function
| _ -> raise Not_found
let add_syntactic_definition ~local deprecation env ident (vars,c) { onlyparsing } =
- let vars,reversibility,pat =
- try [], APrioriReversible, NRef (try_interp_name_alias (vars,c))
+ let acvars,pat,reversibility =
+ try Id.Map.empty, NRef (try_interp_name_alias (vars,c)), APrioriReversible
with Not_found ->
let fold accu id = Id.Map.add id NtnInternTypeAny accu in
let i_vars = List.fold_left fold Id.Map.empty vars in
@@ -1801,10 +1803,11 @@ let add_syntactic_definition ~local deprecation env ident (vars,c) { onlyparsing
ninterp_var_type = i_vars;
ninterp_rec_vars = Id.Map.empty;
} in
- let nvars, pat, reversibility = interp_notation_constr env nenv c in
- let map id = let (_,sc) = Id.Map.find id nvars in (id, sc) in
- List.map map vars, reversibility, pat
+ interp_notation_constr env nenv c
in
+ let in_pat id = (id,ETConstr (Constrexpr.InConstrEntry,None,(NextLevel,0))) in
+ let interp = make_interpretation_vars ~default_if_binding:AsIdentOrPattern [] acvars (List.map in_pat vars) in
+ let vars = List.map (fun x -> (x, Id.Map.find x interp)) vars in
let onlyparsing = onlyparsing || fst (printability None false reversibility pat) in
Syntax_def.declare_syntactic_definition ~local deprecation ident ~onlyparsing (vars,pat)