aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2016-09-22 11:14:37 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2016-09-22 11:17:13 +0200
commit464c680b631e1ba892f2171a36002d6ca184bc4f (patch)
tree0f046047aa0e14509e1a8b7d99e832816707a0a5
parentda708e759a1518bf4304e3d0edd725ed4176941f (diff)
Fixing #5095 (non relevant too strict test in let-in abstraction).
-rw-r--r--pretyping/unification.ml9
-rw-r--r--test-suite/bugs/closed/5095.v5
2 files changed, 6 insertions, 8 deletions
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index cd0bbfa30..347bf6f9e 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -1570,14 +1570,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
let compute_dependency _ (hyp,_,_ as d) (sign,depdecls) =
match occurrences_of_hyp hyp occs with
| NoOccurrences, InHyp ->
- if indirectly_dependent c d depdecls then
- (* Told explicitly not to abstract over [d], but it is dependent *)
- let id' = indirect_dependency d depdecls in
- errorlabstrm "" (str "Cannot abstract over " ++ Nameops.pr_id id'
- ++ str " without also abstracting or erasing " ++ Nameops.pr_id hyp
- ++ str ".")
- else
- (push_named_context_val d sign,depdecls)
+ (push_named_context_val d sign,depdecls)
| AllOccurrences, InHyp as occ ->
let occ = if likefirst then LikeFirst else AtOccs occ in
let newdecl = replace_term_occ_decl_modulo occ test mkvarid d in
diff --git a/test-suite/bugs/closed/5095.v b/test-suite/bugs/closed/5095.v
new file mode 100644
index 000000000..b6f38e3e8
--- /dev/null
+++ b/test-suite/bugs/closed/5095.v
@@ -0,0 +1,5 @@
+(* Checking let-in abstraction *)
+Goal let x := Set in let y := x in True.
+ intros x y.
+ (* There used to have a too strict dependency test there *)
+ set (s := Set) in (value of x).