aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--pretyping/inductiveops.ml2
-rw-r--r--pretyping/inductiveops.mli2
-rw-r--r--tactics/tacinterp.ml6
3 files changed, 6 insertions, 4 deletions
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index 74795f572c..a0b2cb80fc 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -369,3 +369,5 @@ let control_only_guard env =
| LetIn (_,c1,c2,c3) -> control_rec c1; control_rec c2; control_rec c3
in
control_rec
+
+let subst_inductive subst (kn,i) = (Mod_subst.subst_kn subst kn,i)
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 7f18a79435..aa0306d2f5 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -99,3 +99,5 @@ val make_default_case_info : env -> case_style -> inductive -> case_info
(********************)
val control_only_guard : env -> types -> unit
+
+val subst_inductive : Mod_subst.substitution -> inductive -> inductive
diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml
index 52ed344a58..87906d01ce 100644
--- a/tactics/tacinterp.ml
+++ b/tactics/tacinterp.ml
@@ -1815,7 +1815,7 @@ and interp_atomic ist gl = function
| HypArgType ->
VConstr (mkVar (interp_var ist gl (out_gen globwit_var x)))
| RefArgType ->
- VConstr (constr_of_reference
+ VConstr (constr_of_global
(pf_interp_reference ist gl (out_gen globwit_ref x)))
| SortArgType ->
VConstr (mkSort (Pretyping.interp_sort (out_gen globwit_sort x)))
@@ -1864,8 +1864,6 @@ let subst_quantified_hypothesis _ x = x
let subst_declared_or_quantified_hypothesis _ x = x
-let subst_inductive subst (kn,i) = (subst_kn subst kn,i)
-
let subst_rawconstr subst (c,e) =
assert (e=None); (* e<>None only for toplevel tactics *)
(Detyping.subst_raw subst c,None)
@@ -1905,7 +1903,7 @@ let subst_reference subst =
let subst_global_reference subst =
let subst_global ref =
let ref',t' = subst_global subst ref in
- if not (eq_constr (constr_of_reference ref') t') then
+ if not (eq_constr (constr_of_global ref') t') then
ppnl (str "Warning: the reference " ++ pr_global ref ++ str " is not " ++
str " expanded to \"" ++ prterm t' ++ str "\", but to " ++
pr_global ref') ;