aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--test-suite/bugs/closed/HoTT_coq_123.v6
-rw-r--r--vernac/lemmas.ml8
2 files changed, 10 insertions, 4 deletions
diff --git a/test-suite/bugs/closed/HoTT_coq_123.v b/test-suite/bugs/closed/HoTT_coq_123.v
index cd9cad406..7bed956f3 100644
--- a/test-suite/bugs/closed/HoTT_coq_123.v
+++ b/test-suite/bugs/closed/HoTT_coq_123.v
@@ -104,11 +104,15 @@ Record Functor (C D : PreCategory) :=
morphism_of : forall s d, morphism C s d -> morphism D (object_of s) (object_of d)
}.
+(** Workaround to simpl losing universe constraints, see bug #5645. *)
+Ltac simpl' :=
+ simpl; match goal with [ |- ?P ] => let T := type of P in idtac end.
+
Global Instance trunc_forall `{Funext} `{P : A -> Type} `{forall a, IsTrunc n (P a)}
: IsTrunc n (forall a, P a) | 100.
Proof.
generalize dependent P.
- induction n as [ | n' IH]; (simpl; intros P ?).
+ induction n as [ | n' IH]; (simpl'; intros P ?).
- admit.
- pose (fun f g => trunc_equiv (@apD10 A P f g) ^-1); admit.
Defined.
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index cfd489dde..645320c60 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -44,12 +44,14 @@ let call_hook fix_exn hook l c =
(* Support for mutually proved theorems *)
-let retrieve_first_recthm = function
+let retrieve_first_recthm uctx = function
| VarRef id ->
(NamedDecl.get_value (Global.lookup_named id),variable_opacity id)
| ConstRef cst ->
let cb = Global.lookup_constant cst in
- let map (c, ctx) = Vars.subst_instance_constr (Univ.AUContext.instance ctx) c in
+ let (_, uctx) = UState.universe_context uctx in
+ let inst = Univ.UContext.instance uctx in
+ let map (c, ctx) = Vars.subst_instance_constr inst c in
(Option.map map (Global.body_of_constant_body cb), is_opaque cb)
| _ -> assert false
@@ -413,7 +415,7 @@ let start_proof_with_initialization kind ctx recguard thms snl hook =
let other_thms_data =
if List.is_empty other_thms then [] else
(* there are several theorems defined mutually *)
- let body,opaq = retrieve_first_recthm ref in
+ let body,opaq = retrieve_first_recthm ctx ref in
let subst = Evd.evar_universe_context_subst ctx in
let norm c = Universes.subst_opt_univs_constr subst c in
let ctx = UState.context_set (*FIXME*) ctx in