aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorHugo Herbelin2018-12-16 20:33:09 +0100
committerHugo Herbelin2018-12-18 18:24:31 +0100
commit4e529454022b7d2dc0c57d29c813c5801dfd438c (patch)
treed84962c9f9dd70977f3f151610b9c4769ac2023e
parent4c733a9282bf2a272eb0ff48811b528aebbfb5a0 (diff)
Fixes #9229 (Infix not robust wrt choice of variable names).
-rw-r--r--interp/constrexpr_ops.ml8
-rw-r--r--interp/constrexpr_ops.mli3
-rw-r--r--test-suite/bugs/closed/bug_9229.v6
-rw-r--r--vernac/metasyntax.ml9
4 files changed, 23 insertions, 3 deletions
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 3a5af1dd5f..7bc5d090b4 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -366,6 +366,14 @@ let free_vars_of_constr_expr c =
| c -> fold_constr_expr_with_binders (fun a l -> a::l) aux bdvars l c
in aux [] Id.Set.empty c
+let names_of_constr_expr c =
+ let vars = ref Id.Set.empty in
+ let rec aux () () = function
+ | { CAst.v = CRef (qid, _) } when qualid_is_ident qid ->
+ let id = qualid_basename qid in vars := Id.Set.add id !vars
+ | c -> fold_constr_expr_with_binders (fun a () -> vars := Id.Set.add a !vars) aux () () c
+ in aux () () c; !vars
+
let occur_var_constr_expr id c = Id.Set.mem id (free_vars_of_constr_expr c)
(* Used in correctness and interface *)
diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli
index 7f14eb4583..8c735edfc9 100644
--- a/interp/constrexpr_ops.mli
+++ b/interp/constrexpr_ops.mli
@@ -119,6 +119,9 @@ val ids_of_cases_indtype : cases_pattern_expr -> Id.Set.t
val free_vars_of_constr_expr : constr_expr -> Id.Set.t
val occur_var_constr_expr : Id.t -> constr_expr -> bool
+(** Return all (non-qualified) names treating binders as names *)
+val names_of_constr_expr : constr_expr -> Id.Set.t
+
val split_at_annot : local_binder_expr list -> lident option -> local_binder_expr list * local_binder_expr list
val ntn_loc : ?loc:Loc.t -> constr_notation_substitution -> notation -> (int * int) list
diff --git a/test-suite/bugs/closed/bug_9229.v b/test-suite/bugs/closed/bug_9229.v
new file mode 100644
index 0000000000..7514741af4
--- /dev/null
+++ b/test-suite/bugs/closed/bug_9229.v
@@ -0,0 +1,6 @@
+(* There was a problem of freshness with Infix choice of vars *)
+
+(* In particular, x and y were special... *)
+
+Infix "#" := (fun x y => x + y) (at level 50, left associativity).
+Check (3 # 5).
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index 4e79b50b79..3da12e7714 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -1563,14 +1563,17 @@ let add_notation_extra_printing_rule df k v =
(* Infix notations *)
-let inject_var x = CAst.make @@ CRef (qualid_of_ident (Id.of_string x),None)
+let inject_var x = CAst.make @@ CRef (qualid_of_ident x,None)
let add_infix local env ({CAst.loc;v=inf},modifiers) pr sc =
check_infix_modifiers modifiers;
(* check the precedence *)
- let metas = [inject_var "x"; inject_var "y"] in
+ let vars = names_of_constr_expr pr in
+ let x = Namegen.next_ident_away (Id.of_string "x") vars in
+ let y = Namegen.next_ident_away (Id.of_string "y") vars in
+ let metas = [inject_var x; inject_var y] in
let c = mkAppC (pr,metas) in
- let df = CAst.make ?loc @@ "x "^(quote_notation_token inf)^" y" in
+ let df = CAst.make ?loc @@ Id.to_string x ^" "^(quote_notation_token inf)^" "^Id.to_string y in
add_notation local env c (df,modifiers) sc
(**********************************************************************)