diff options
| author | Hugo Herbelin | 2019-10-15 11:12:21 +0200 |
|---|---|---|
| committer | Hugo Herbelin | 2019-10-21 14:46:50 +0200 |
| commit | 9fe7c0b44d9405d4008210d674aa742b2d343c4a (patch) | |
| tree | 1a101ddb3129412d935f8dc94cc43c542ca5f77d | |
| parent | c6056bdc54d02dc5c6a91aeacb13460e8dd365e3 (diff) | |
Fixes #10894: uconstr was not anymore typed in the Ltac-substituted environment.
More precisely: The equivalent in #7288 (4dab4fc) to the former call
to ltac_interp_name_env was not done anymore when interpreting
uconstr's.
| -rw-r--r-- | pretyping/globEnv.ml | 2 | ||||
| -rw-r--r-- | test-suite/bugs/closed/bug_10894.v | 12 |
2 files changed, 13 insertions, 1 deletions
diff --git a/pretyping/globEnv.ml b/pretyping/globEnv.ml index d49a39b547..aebe47a7a7 100644 --- a/pretyping/globEnv.ml +++ b/pretyping/globEnv.ml @@ -166,7 +166,7 @@ let interp_ltac_variable ?loc typing_fun env sigma id : Evd.evar_map * unsafe_ju here, as the call to the main pretyping function is caught inside the try but I want to avoid refactoring this function too much for now. *) - typing_fun {env with lvar} term + typing_fun {env with lvar; static_env = env.renamed_env} term with Not_found -> (* Check if [id] is a ltac variable not bound to a term *) (* and build a nice error message *) diff --git a/test-suite/bugs/closed/bug_10894.v b/test-suite/bugs/closed/bug_10894.v new file mode 100644 index 0000000000..b8c9367951 --- /dev/null +++ b/test-suite/bugs/closed/bug_10894.v @@ -0,0 +1,12 @@ +(* Check that uconstrs are interpreted in the ltac-substituted environment *) +(* Was a regression introduced in 4dab4fc (#7288) *) + +Tactic Notation "bind" hyp(x) "in" uconstr(f) "as" ident(h) := + set (h := fun x => f). + +Fact test : nat -> nat. +Proof. + intros n. + bind n in (n*n) as f. + assert (f 0 = 0) by reflexivity. +Abort. |
