diff options
author | Hugo Herbelin <Hugo.Herbelin@inria.fr> | 2017-05-01 16:56:25 +0200 |
---|---|---|
committer | Hugo Herbelin <Hugo.Herbelin@inria.fr> | 2017-05-01 16:58:04 +0200 |
commit | 12f1c409daf2cdbd7d0323f0d61723819532b362 (patch) | |
tree | 3fff6be70bc0f76472e0c40569205a3e398a877d | |
parent | 8511d1d9d903e419543e39eca83c64171da2663b (diff) |
Really fixing #2602 which was wrongly working because of #5487 hiding the cause.
The cause was a missing evar/evar clause in ltac pattern-matching
function (constr_matching.ml).
-rw-r--r-- | pretyping/constr_matching.ml | 3 | ||||
-rw-r--r-- | test-suite/success/ltac.v | 13 |
2 files changed, 16 insertions, 0 deletions
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 3c47cfdc4..afdf601c2 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -347,6 +347,9 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | PFix c1, Fix _ when eq_constr (mkFix c1) cT -> subst | PCoFix c1, CoFix _ when eq_constr (mkCoFix c1) cT -> subst + | PEvar (c1,args1), Evar (c2,args2) when c1 = c2 -> + Array.fold_left2 (sorec ctx env) subst args1 args2 + | _ -> raise PatternMatchingFailure in diff --git a/test-suite/success/ltac.v b/test-suite/success/ltac.v index ce9099059..d7ec092d7 100644 --- a/test-suite/success/ltac.v +++ b/test-suite/success/ltac.v @@ -317,3 +317,16 @@ let T := constr:(fun a b : nat => a) in end. exact (eq_refl n). Qed. + +(* A variant of #2602 which was wrongly succeeding because "a", bound to + "?m", was then internally turned into a "_" in the second matching *) + +Goal exists m, S m > 0. +eexists. +Fail match goal with + | |- context [ S ?a ] => + match goal with + | |- S a > a => idtac + end +end. +Abort. |