aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--pretyping/detyping.ml8
-rw-r--r--test-suite/bugs/closed/bug_13732.v16
-rw-r--r--vernac/vernacentries.ml5
3 files changed, 25 insertions, 4 deletions
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 722a0a2048..c5f23482c7 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -744,9 +744,11 @@ let detype_level sigma l =
UNamed (detype_level_name sigma l)
let detype_instance sigma l =
- let l = EInstance.kind sigma l in
- if Univ.Instance.is_empty l then None
- else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l)))
+ if not !print_universes then None
+ else
+ let l = EInstance.kind sigma l in
+ if Univ.Instance.is_empty l then None
+ else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l)))
let delay (type a) (d : a delay) (f : a delay -> _ -> _ -> _ -> _ -> _ -> a glob_constr_r) flags env avoid sigma t : a glob_constr_g =
match d with
diff --git a/test-suite/bugs/closed/bug_13732.v b/test-suite/bugs/closed/bug_13732.v
new file mode 100644
index 0000000000..24840abdf6
--- /dev/null
+++ b/test-suite/bugs/closed/bug_13732.v
@@ -0,0 +1,16 @@
+Module Sort.
+ Set Printing Universes.
+
+ Implicit Types TT : Type.
+
+ Check fun TT => nat.
+End Sort.
+
+Module Ref.
+ Set Universe Polymorphism.
+
+ Axiom tele : Type.
+
+ Implicit Types TT : tele.
+ Check fun TT => nat.
+End Ref.
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 664c6b2f36..42ba63903d 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -1438,7 +1438,10 @@ let vernac_reserve bl =
let env = Global.env() in
let sigma = Evd.from_env env in
let t,ctx = Constrintern.interp_type env sigma c in
- let t = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_ctx ctx) t in
+ let t = Flags.without_option Detyping.print_universes (fun () ->
+ Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_ctx ctx) t)
+ ()
+ in
let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
Reserve.declare_reserved_type idl t)
in List.iter sb_decl bl