aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-05-19 12:20:19 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2017-05-19 12:26:55 +0200
commit5c68b57c37c23a3b7b2afe4f6ff073568c7b8db9 (patch)
treecd6070f0374179932649b7c9f3c9d0ff40256cce
parent9f11adda4bff194a3c6a66d573ce7001d4399886 (diff)
Fixing an extra bug with pattern_of_constr.
Ensure in type constr_pattern that those preexisting existential variables of the goal which do not contribute as pattern variables are expanded: constr_pattern is not observed up to evar expansion (like EConstr does), so we need to pre-normalize defined evars in patterns to that matching against an EConstr works.
-rw-r--r--pretyping/patternops.ml5
-rw-r--r--test-suite/success/change.v9
2 files changed, 12 insertions, 2 deletions
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index a22db1407..638e12d7c 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -155,14 +155,15 @@ let pattern_of_constr env sigma t =
| Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp))
| Proj (p, c) ->
pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) []))
- | Evar (evk,ctxt) ->
+ | Evar (evk,ctxt as ev) ->
(match snd (Evd.evar_source evk sigma) with
| Evar_kinds.MatchingVar (b,id) ->
assert (not b); PMeta (Some id)
| Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ ->
(* These are the two evar kinds used for existing goals *)
(* see Proofview.mark_in_evm *)
- PEvar (evk,Array.map (pattern_of_constr env) ctxt)
+ if Evd.is_defined sigma evk then pattern_of_constr env (Evd.existential_value sigma ev)
+ else PEvar (evk,Array.map (pattern_of_constr env) ctxt)
| _ ->
PMeta None)
| Case (ci,p,a,br) ->
diff --git a/test-suite/success/change.v b/test-suite/success/change.v
index 1f0b7d38a..a9821b027 100644
--- a/test-suite/success/change.v
+++ b/test-suite/success/change.v
@@ -59,3 +59,12 @@ unfold x.
(* check that n in 0+n is not interpreted as the n from "fun n" *)
change n with (0+n).
Abort.
+
+(* Check non-collision of non-normalized defined evars with pattern variables *)
+
+Goal exists x, 1=1 -> x=1/\x=1.
+eexists ?[n]; intros; split.
+eassumption.
+match goal with |- ?x=1 => change (x=1) with (0+x=1) end.
+match goal with |- 0+1=1 => trivial end.
+Qed.