From 9938aed874d3e15e5d21689ea841bdc3e6ebb40e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 11 Jul 2017 21:51:23 +0200 Subject: Safer API for Global.body_of_constant and variants. We aditionally return the abstract universe context inside the option. This is relatively painless as most uses were using the option as a boolean. --- plugins/funind/functional_principles_proofs.ml | 4 ++-- plugins/funind/functional_principles_types.ml | 2 +- plugins/funind/indfun.ml | 2 +- 3 files changed, 4 insertions(+), 4 deletions(-) (limited to 'plugins') diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index ba46f78aa..dc43ea7c4 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -957,7 +957,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num (* observe (str "rec_args_num := " ++ str (string_of_int (rec_args_num + 1) )); *) let f_def = Global.lookup_constant (fst (destConst evd f)) in let eq_lhs = mkApp(f,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i))) in - let f_body = Option.get (Global.body_of_constant_body f_def) in + let (f_body, _) = Option.get (Global.body_of_constant_body f_def) in let f_body = EConstr.of_constr f_body in let params,f_body_with_params = decompose_lam_n evd nb_params f_body in let (_,num),(_,_,bodies) = destFix evd f_body_with_params in @@ -1091,7 +1091,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam in let get_body const = match Global.body_of_constant const with - | Some body -> + | Some (body, _) -> Tacred.cbv_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) (Global.env ()) diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 8ffd15f9f..8c9b514e6 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -407,7 +407,7 @@ let get_funs_constant mp dp = function const -> let find_constant_body const = match Global.body_of_constant const with - | Some body -> + | Some (body, _) -> let body = Tacred.cbv_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) (Global.env ()) diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 2c5dae1cd..ff7667e99 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -851,7 +851,7 @@ let make_graph (f_ref:global_reference) = in (match Global.body_of_constant_body c_body with | None -> error "Cannot build a graph over an axiom!" - | Some body -> + | Some (body, _) -> let env = Global.env () in let sigma = Evd.from_env env in let extern_body,extern_type = -- cgit v1.2.3 From 71563ebb86a83bc7cdfc17f58493f59428d764b0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 12 Jul 2017 14:49:26 +0200 Subject: Safer API for constr_of_global, and getting rid of unsafe_constr_of_global. --- API/API.mli | 4 +- engine/termops.ml | 2 +- engine/universes.ml | 66 --------------------------- engine/universes.mli | 11 ----- library/global.ml | 19 ++++++++ library/global.mli | 7 +++ plugins/funind/functional_principles_types.ml | 2 +- plugins/funind/recdef.ml | 2 +- plugins/ssr/ssrvernac.ml4 | 3 +- pretyping/classops.ml | 2 +- pretyping/pretyping.ml | 4 +- printing/prettyp.ml | 6 ++- tactics/tactics.ml | 4 +- 13 files changed, 43 insertions(+), 89 deletions(-) (limited to 'plugins') diff --git a/API/API.mli b/API/API.mli index 0f3394fe9..f8fb96aa9 100644 --- a/API/API.mli +++ b/API/API.mli @@ -2684,10 +2684,8 @@ sig val fresh_inductive_instance : Environ.env -> Names.inductive -> Term.pinductive Univ.in_universe_context_set val new_Type : Names.DirPath.t -> Term.types val type_of_global : Globnames.global_reference -> Term.types Univ.in_universe_context_set - val unsafe_type_of_global : Globnames.global_reference -> Term.types val constr_of_global : Prelude.global_reference -> Term.constr val new_univ_level : Names.DirPath.t -> Univ.Level.t - val unsafe_constr_of_global : Globnames.global_reference -> Term.constr Univ.in_universe_context val new_sort_in_family : Sorts.family -> Sorts.t val pr_with_global_universes : Univ.Level.t -> Pp.std_ppcmds val pr_universe_opt_subst : universe_opt_subst -> Pp.std_ppcmds @@ -2713,6 +2711,8 @@ sig val env_of_context : Environ.named_context_val -> Environ.env val is_polymorphic : Globnames.global_reference -> bool + val constr_of_global_in_context : Environ.env -> Globnames.global_reference -> Constr.t * Univ.AUContext.t + val type_of_global_in_context : Environ.env -> Globnames.global_reference -> Constr.t * Univ.AUContext.t val type_of_global_unsafe : Globnames.global_reference -> Term.types val current_dirpath : unit -> Names.DirPath.t diff --git a/engine/termops.ml b/engine/termops.ml index fc3291df1..1aba2bbdd 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -906,7 +906,7 @@ let collect_vars sigma c = aux Id.Set.empty c let vars_of_global_reference env gr = - let c, _ = Universes.unsafe_constr_of_global gr in + let c, _ = Global.constr_of_global_in_context env gr in vars_of_global (Global.env ()) c (* Tests whether [m] is a subterm of [t]: diff --git a/engine/universes.ml b/engine/universes.ml index fc441fd0b..21854b3fa 100644 --- a/engine/universes.ml +++ b/engine/universes.ml @@ -319,9 +319,6 @@ let fresh_instance_from ctx inst = let constraints = AUContext.instantiate inst ctx in inst, (ctx', constraints) -let unsafe_instance_from ctx = - (Univ.AUContext.instance ctx, Univ.instantiate_univ_context ctx) - (** Fresh universe polymorphic construction *) let fresh_constant_instance env c inst = @@ -358,34 +355,6 @@ let fresh_constructor_instance env (ind,i) inst = let inst, ctx = fresh_instance_from (ACumulativityInfo.univ_context acumi) inst in (((ind,i),inst), ctx) -let unsafe_constant_instance env c = - let cb = lookup_constant c env in - match cb.Declarations.const_universes with - | Declarations.Monomorphic_const _ -> - ((c,Instance.empty), UContext.empty) - | Declarations.Polymorphic_const auctx -> - let inst, ctx = unsafe_instance_from auctx in ((c, inst), ctx) - -let unsafe_inductive_instance env ind = - let mib, mip = Inductive.lookup_mind_specif env ind in - match mib.Declarations.mind_universes with - | Declarations.Monomorphic_ind _ -> ((ind,Instance.empty), UContext.empty) - | Declarations.Polymorphic_ind auctx -> - let inst, ctx = unsafe_instance_from auctx in ((ind,inst), ctx) - | Declarations.Cumulative_ind acumi -> - let inst, ctx = unsafe_instance_from (ACumulativityInfo.univ_context acumi) in - ((ind,inst), ctx) - -let unsafe_constructor_instance env (ind,i) = - let mib, mip = Inductive.lookup_mind_specif env ind in - match mib.Declarations.mind_universes with - | Declarations.Monomorphic_ind _ -> (((ind, i),Instance.empty), UContext.empty) - | Declarations.Polymorphic_ind auctx -> - let inst, ctx = unsafe_instance_from auctx in (((ind, i),inst), ctx) - | Declarations.Cumulative_ind acumi -> - let inst, ctx = unsafe_instance_from (ACumulativityInfo.univ_context acumi) in - (((ind, i),inst), ctx) - open Globnames let fresh_global_instance ?names env gr = @@ -410,19 +379,6 @@ let fresh_inductive_instance env sp = let fresh_constructor_instance env sp = fresh_constructor_instance env sp None -let unsafe_global_instance env gr = - match gr with - | VarRef id -> mkVar id, UContext.empty - | ConstRef sp -> - let c, ctx = unsafe_constant_instance env sp in - mkConstU c, ctx - | ConstructRef sp -> - let c, ctx = unsafe_constructor_instance env sp in - mkConstructU c, ctx - | IndRef sp -> - let c, ctx = unsafe_inductive_instance env sp in - mkIndU c, ctx - let constr_of_global gr = let c, ctx = fresh_global_instance (Global.env ()) gr in if not (Univ.ContextSet.is_empty ctx) then @@ -437,9 +393,6 @@ let constr_of_global gr = let constr_of_reference = constr_of_global -let unsafe_constr_of_global gr = - unsafe_global_instance (Global.env ()) gr - let constr_of_global_univ (gr,u) = match gr with | VarRef id -> mkVar id @@ -513,25 +466,6 @@ let type_of_reference env r = let type_of_global t = type_of_reference (Global.env ()) t -let unsafe_type_of_reference env r = - match r with - | VarRef id -> Environ.named_type id env - | ConstRef c -> - let cb = Environ.lookup_constant c env in - Typeops.type_of_constant_type env cb.const_type - - | IndRef ind -> - let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in - let (_, inst), _ = unsafe_inductive_instance env ind in - Inductive.type_of_inductive env (specif, inst) - - | ConstructRef (ind, _ as cstr) -> - let (mib,oib as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in - let (_, inst), _ = unsafe_inductive_instance env ind in - Inductive.type_of_constructor (cstr,inst) specif - -let unsafe_type_of_global t = unsafe_type_of_reference (Global.env ()) t - let fresh_sort_in_family env = function | InProp -> prop_sort, ContextSet.empty | InSet -> set_sort, ContextSet.empty diff --git a/engine/universes.mli b/engine/universes.mli index 5f4d212b6..8d0f106de 100644 --- a/engine/universes.mli +++ b/engine/universes.mli @@ -189,22 +189,11 @@ val constr_of_global : Globnames.global_reference -> constr (** ** DEPRECATED ** synonym of [constr_of_global] *) val constr_of_reference : Globnames.global_reference -> constr -(** [unsafe_constr_of_global gr] turns [gr] into a constr, works on polymorphic - references by taking the original universe instance that is not recorded - anywhere. The constraints are forgotten as well. DO NOT USE in new code. *) -val unsafe_constr_of_global : Globnames.global_reference -> constr in_universe_context - (** Returns the type of the global reference, by creating a fresh instance of polymorphic references and computing their instantiated universe context. (side-effect on the universe counter, use with care). *) val type_of_global : Globnames.global_reference -> types in_universe_context_set -(** [unsafe_type_of_global gr] returns [gr]'s type, works on polymorphic - references by taking the original universe instance that is not recorded - anywhere. The constraints are forgotten as well. - USE with care. *) -val unsafe_type_of_global : Globnames.global_reference -> types - (** Full universes substitutions into terms *) val nf_evars_and_universes_opt_subst : (existential -> constr option) -> diff --git a/library/global.ml b/library/global.ml index 3a74f535d..5fff26566 100644 --- a/library/global.ml +++ b/library/global.ml @@ -191,6 +191,25 @@ let type_of_global_unsafe r = let inst = Univ.AUContext.instance (Declareops.inductive_polymorphic_context mib) in Inductive.type_of_constructor (cstr,inst) specif +let constr_of_global_in_context env r = + let open Constr in + match r with + | VarRef id -> mkVar id, Univ.AUContext.empty + | ConstRef c -> + let cb = Environ.lookup_constant c env in + let univs = Declareops.constant_polymorphic_context cb in + mkConstU (c, Univ.make_abstract_instance univs), univs + | IndRef ind -> + let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in + let univs = Declareops.inductive_polymorphic_context mib in + mkIndU (ind, Univ.make_abstract_instance univs), univs + | ConstructRef cstr -> + let (mib,oib as specif) = + Inductive.lookup_mind_specif env (inductive_of_constructor cstr) + in + let univs = Declareops.inductive_polymorphic_context mib in + mkConstructU (cstr, Univ.make_abstract_instance univs), univs + let type_of_global_in_context env r = match r with | VarRef id -> Environ.named_type id env, Univ.AUContext.empty diff --git a/library/global.mli b/library/global.mli index d6d2f1f85..0f1cec44a 100644 --- a/library/global.mli +++ b/library/global.mli @@ -122,6 +122,13 @@ val is_polymorphic : Globnames.global_reference -> bool val is_template_polymorphic : Globnames.global_reference -> bool val is_type_in_type : Globnames.global_reference -> bool +val constr_of_global_in_context : Environ.env -> + Globnames.global_reference -> Constr.types * Univ.AUContext.t +(** Returns the type of the constant in its local universe + context. The type should not be used without pushing it's universe + context in the environmnent of usage. For non-universe-polymorphic + constants, it does not matter. *) + val type_of_global_in_context : Environ.env -> Globnames.global_reference -> Constr.types * Univ.AUContext.t (** Returns the type of the constant in its local universe diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 8c9b514e6..d726c1a2b 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -651,7 +651,7 @@ let build_case_scheme fa = (* in *) let funs = let (_,f,_) = fa in - try fst (Universes.unsafe_constr_of_global (Smartlocate.global_with_alias f)) + try fst (Global.constr_of_global_in_context (Global.env ()) (Smartlocate.global_with_alias f)) with Not_found -> user_err ~hdr:"FunInd.build_case_scheme" (str "Cannot find " ++ Libnames.pr_reference f) in diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 1705cac78..bc550c8b9 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -90,7 +90,7 @@ let type_of_const sigma t = |_ -> assert false let constr_of_global x = - fst (Universes.unsafe_constr_of_global x) + fst (Global.constr_of_global_in_context (Global.env ()) x) let constant sl s = constr_of_global (find_reference sl s) diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4 index 4c8827bf8..cb6a2cd69 100644 --- a/plugins/ssr/ssrvernac.ml4 +++ b/plugins/ssr/ssrvernac.ml4 @@ -337,7 +337,8 @@ let coerce_search_pattern_to_sort hpat = Pattern.PApp (fp, args') in let hr, na = splay_search_pattern 0 hpat in let dc, ht = - Reductionops.splay_prod env sigma (EConstr.of_constr (Universes.unsafe_type_of_global hr)) in + let hr, _ = Global.type_of_global_in_context (Global.env ()) hr (** FIXME *) in + Reductionops.splay_prod env sigma (EConstr.of_constr hr) in let np = List.length dc in if np < na then CErrors.user_err (Pp.str "too many arguments in head search pattern") else let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 948aa26ca..078990a8c 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -403,7 +403,7 @@ type coercion = { (* Computation of the class arity *) let reference_arity_length ref = - let t = Universes.unsafe_type_of_global ref in + let t, _ = Global.type_of_global_in_context (Global.env ()) ref in List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty (EConstr.of_constr t))) (** FIXME *) let projection_arity_length p = diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index e166e0e9d..bfc6bf5cf 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -511,8 +511,8 @@ let pretype_global ?loc rigid env evd gr us = match us with | None -> evd, None | Some l -> - let _, ctx = Universes.unsafe_constr_of_global gr in - let len = Univ.UContext.size ctx in + let _, ctx = Global.constr_of_global_in_context env.ExtraEnv.env gr in + let len = Univ.AUContext.size ctx in interp_instance ?loc evd ~len l in let (sigma, c) = Evd.fresh_global ?loc ~rigid ?names:instance env.ExtraEnv.env evd gr in diff --git a/printing/prettyp.ml b/printing/prettyp.ml index d1e51c9f3..a0c88a7af 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -798,9 +798,11 @@ let print_opaque_name qid = | IndRef (sp,_) -> print_inductive sp | ConstructRef cstr as gr -> - let open EConstr in - let ty = Universes.unsafe_type_of_global gr in + let ty, ctx = Global.type_of_global_in_context env gr in + let inst = Univ.AUContext.instance ctx in + let ty = Vars.subst_instance_constr inst ty in let ty = EConstr.of_constr ty in + let open EConstr in print_typed_value (mkConstruct cstr, ty) | VarRef id -> env |> lookup_named id |> NamedDecl.set_id id |> print_named_decl diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 2e3a4e33b..177c44bcb 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -5004,7 +5004,9 @@ let cache_term_by_tactic_then ~opaque ?(goal_type=None) id gk tac tacK = in let cst = Impargs.with_implicit_protection cst () in (* let evd, lem = Evd.fresh_global (Global.env ()) evd (ConstRef cst) in *) - let lem, ctx = Universes.unsafe_constr_of_global (ConstRef cst) in + let lem, ctx = Global.constr_of_global_in_context (Global.env ()) (ConstRef cst) in + let inst = Univ.AUContext.instance ctx in + let lem = CVars.subst_instance_constr inst lem in let lem = EConstr.of_constr lem in let evd = Evd.set_universe_context evd ectx in let open Safe_typing in -- cgit v1.2.3 From 469a9b3242891b089b4a211e96b5b568277f7fc0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 12 Jul 2017 15:55:51 +0200 Subject: Remove the function Global.type_of_global_unsafe. --- API/API.mli | 1 - interp/notation.ml | 6 +++--- library/global.ml | 18 ------------------ library/global.mli | 11 ----------- library/impargs.ml | 7 ++++--- plugins/extraction/table.ml | 7 ++++--- plugins/funind/indfun_common.ml | 2 +- pretyping/typeclasses.ml | 2 +- printing/prettyp.ml | 5 +++-- tactics/hints.ml | 2 +- vernac/auto_ind_decl.ml | 2 +- vernac/class.ml | 6 ++++-- vernac/classes.ml | 2 +- vernac/search.ml | 2 +- vernac/vernacentries.ml | 2 +- 15 files changed, 25 insertions(+), 50 deletions(-) (limited to 'plugins') diff --git a/API/API.mli b/API/API.mli index f8fb96aa9..e8418552c 100644 --- a/API/API.mli +++ b/API/API.mli @@ -2713,7 +2713,6 @@ sig val constr_of_global_in_context : Environ.env -> Globnames.global_reference -> Constr.t * Univ.AUContext.t val type_of_global_in_context : Environ.env -> Globnames.global_reference -> Constr.t * Univ.AUContext.t - val type_of_global_unsafe : Globnames.global_reference -> Term.types val current_dirpath : unit -> Names.DirPath.t val body_of_constant_body : Declarations.constant_body -> (Term.constr * Univ.AUContext.t) option diff --git a/interp/notation.ml b/interp/notation.ml index 4067a6b94..c07a00943 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -718,13 +718,13 @@ let rebuild_arguments_scope (req,r,n,l,_) = match req with | ArgsScopeNoDischarge -> assert false | ArgsScopeAuto -> - let scs,cls = compute_arguments_scope_full (fst(Universes.type_of_global r)(*FIXME?*)) in + let scs,cls = compute_arguments_scope_full (fst(Global.type_of_global_in_context (Global.env ()) r)(*FIXME?*)) in (req,r,List.length scs,scs,cls) | ArgsScopeManual -> (* Add to the manually given scopes the one found automatically for the extra parameters of the section. Discard the classes of the manually given scopes to avoid further re-computations. *) - let l',cls = compute_arguments_scope_full (Global.type_of_global_unsafe r) in + let l',cls = compute_arguments_scope_full (fst (Global.type_of_global_in_context (Global.env ()) r)) in let l1 = List.firstn n l' in let cls1 = List.firstn n cls in (req,r,0,l1@l,cls1) @@ -768,7 +768,7 @@ let find_arguments_scope r = with Not_found -> [] let declare_ref_arguments_scope ref = - let t = Global.type_of_global_unsafe ref in + let t, _ = Global.type_of_global_in_context (Global.env ()) ref in let (scs,cls as o) = compute_arguments_scope_full t in declare_arguments_scope_gen ArgsScopeAuto ref (List.length scs) o diff --git a/library/global.ml b/library/global.ml index 5fff26566..5b17855dc 100644 --- a/library/global.ml +++ b/library/global.ml @@ -173,24 +173,6 @@ open Globnames (** Build a fresh instance for a given context, its associated substitution and the instantiated constraints. *) -let type_of_global_unsafe r = - let env = env() in - match r with - | VarRef id -> Environ.named_type id env - | ConstRef c -> - let cb = Environ.lookup_constant c env in - let inst = Univ.AUContext.instance (Declareops.constant_polymorphic_context cb) in - let ty = Typeops.type_of_constant_type env cb.Declarations.const_type in - Vars.subst_instance_constr inst ty - | IndRef ind -> - let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in - let inst = Univ.AUContext.instance (Declareops.inductive_polymorphic_context mib) in - Inductive.type_of_inductive env (specif, inst) - | ConstructRef cstr -> - let (mib,oib as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in - let inst = Univ.AUContext.instance (Declareops.inductive_polymorphic_context mib) in - Inductive.type_of_constructor (cstr,inst) specif - let constr_of_global_in_context env r = let open Constr in match r with diff --git a/library/global.mli b/library/global.mli index 0f1cec44a..431747c52 100644 --- a/library/global.mli +++ b/library/global.mli @@ -136,17 +136,6 @@ val type_of_global_in_context : Environ.env -> context in the environmnent of usage. For non-universe-polymorphic constants, it does not matter. *) -val type_of_global_unsafe : Globnames.global_reference -> Constr.types -(** Returns the type of the constant, forgetting its universe context if - it is polymorphic, use with care: for polymorphic constants, the - type cannot be used to produce a term used by the kernel. For safe - handling of polymorphic global references, one should look at a - particular instantiation of the reference, in some particular - universe context (part of an [env] or [evar_map]), see - e.g. [type_of_constant_in]. If you want to create a fresh instance - of the reference and get its type look at [Evd.fresh_global] or - [Evarutil.new_global] and [Retyping.get_type_of]. *) - (** Returns the universe context of the global reference (whatever its polymorphic status is). *) val universes_of_global : Globnames.global_reference -> Univ.abstract_universe_context diff --git a/library/impargs.ml b/library/impargs.ml index 351addf2f..b7125fc85 100644 --- a/library/impargs.ml +++ b/library/impargs.ml @@ -431,12 +431,13 @@ let compute_mib_implicits flags manual kn = (Array.mapi (* No need to lift, arities contain no de Bruijn *) (fun i mip -> (** No need to care about constraints here *) - Context.Rel.Declaration.LocalAssum (Name mip.mind_typename, Global.type_of_global_unsafe (IndRef (kn,i)))) + let ty, _ = Global.type_of_global_in_context env (IndRef (kn,i)) in + Context.Rel.Declaration.LocalAssum (Name mip.mind_typename, ty)) mib.mind_packets) in let env_ar = push_rel_context ar env in let imps_one_inductive i mip = let ind = (kn,i) in - let ar = Global.type_of_global_unsafe (IndRef ind) in + let ar, _ = Global.type_of_global_in_context env (IndRef ind) in ((IndRef ind,compute_semi_auto_implicits env flags manual ar), Array.mapi (fun j c -> (ConstructRef (ind,j+1),compute_semi_auto_implicits env_ar flags manual c)) @@ -674,7 +675,7 @@ let projection_implicits env p impls = let declare_manual_implicits local ref ?enriching l = let flags = !implicit_args in let env = Global.env () in - let t = Global.type_of_global_unsafe ref in + let t, _ = Global.type_of_global_in_context (Global.env ()) ref in let enriching = Option.default flags.auto enriching in let isrigid,autoimpls = compute_auto_implicits env flags enriching t in let l' = match l with diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml index 2642aeefa..dff3548fd 100644 --- a/plugins/extraction/table.ml +++ b/plugins/extraction/table.ml @@ -445,9 +445,10 @@ let error_MPfile_as_mod mp b = "Please "^s2^"use (Recursive) Extraction Library instead.\n")) let argnames_of_global r = - let typ = Global.type_of_global_unsafe r in + let env = Global.env () in + let typ, _ = Global.type_of_global_in_context env r in let rels,_ = - decompose_prod (Reduction.whd_all (Global.env ()) typ) in + decompose_prod (Reduction.whd_all env typ) in List.rev_map fst rels let msg_of_implicit = function @@ -878,7 +879,7 @@ let extract_constant_inline inline r ids s = match g with | ConstRef kn -> let env = Global.env () in - let typ = Global.type_of_global_unsafe (ConstRef kn) in + let typ, _ = Global.type_of_global_in_context env (ConstRef kn) in let typ = Reduction.whd_all env typ in if Reduction.is_arity env typ then begin diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index 6fe6888f3..61fbca23f 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -342,7 +342,7 @@ let pr_info f_info = str "function_constant_type := " ++ (try Printer.pr_lconstr - (Global.type_of_global_unsafe (ConstRef f_info.function_constant)) + (fst (Global.type_of_global_in_context (Global.env ()) (ConstRef f_info.function_constant))) with e when CErrors.noncritical e -> mt ()) ++ fnl () ++ str "equation_lemma := " ++ pr_ocst f_info.equation_lemma ++ fnl () ++ str "completeness_lemma :=" ++ pr_ocst f_info.completeness_lemma ++ fnl () ++ diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index c4418b5a6..d4fa266c0 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -400,7 +400,7 @@ let remove_instance i = remove_instance_hint i.is_impl let declare_instance info local glob = - let ty = Global.type_of_global_unsafe glob in + let ty, _ = Global.type_of_global_in_context (Global.env ()) glob in let info = Option.default {hint_priority = None; hint_pattern = None} info in match class_of_constr Evd.empty (EConstr.of_constr ty) with | Some (rels, ((tc,_), args) as _cl) -> diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 5cd79ed6d..827c0e458 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -70,7 +70,8 @@ let int_or_no n = if Int.equal n 0 then str "no" else int n let print_basename sp = pr_global (ConstRef sp) let print_ref reduce ref = - let typ = Global.type_of_global_unsafe ref in + let typ, ctx = Global.type_of_global_in_context (Global.env ()) ref in + let typ = Vars.subst_instance_constr (Univ.AUContext.instance ctx) typ in let typ = EConstr.of_constr typ in let typ = if reduce then @@ -137,7 +138,7 @@ let print_renames_list prefix l = hv 2 (prlist_with_sep pr_comma (fun x -> x) (List.map Name.print l))] let need_expansion impl ref = - let typ = Global.type_of_global_unsafe ref in + let typ, _ = Global.type_of_global_in_context (Global.env ()) ref in let ctx = prod_assum typ in let nprods = List.count is_local_assum ctx in not (List.is_empty impl) && List.length impl >= nprods && diff --git a/tactics/hints.ml b/tactics/hints.ml index c2c80e630..a572508d4 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -937,7 +937,7 @@ let make_extern pri pat tacast = let make_mode ref m = let open Term in - let ty = Global.type_of_global_unsafe ref in + let ty, _ = Global.type_of_global_in_context (Global.env ()) ref in let ctx, t = decompose_prod ty in let n = List.length ctx in let m' = Array.of_list m in diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index 248224e6b..59920742d 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -105,7 +105,7 @@ let mkFullInd (ind,u) n = else mkIndU (ind,u) let check_bool_is_defined () = - try let _ = Global.type_of_global_unsafe Coqlib.glob_bool in () + try let _ = Global.type_of_global_in_context (Global.env ()) Coqlib.glob_bool in () with e when CErrors.noncritical e -> raise (UndefinedCst "bool") let beq_scheme_kind_aux = ref (fun _ -> failwith "Undefined") diff --git a/vernac/class.ml b/vernac/class.ml index 0b510bbcf..be682977e 100644 --- a/vernac/class.ml +++ b/vernac/class.ml @@ -62,7 +62,9 @@ let explain_coercion_error g = function (* Verifications pour l'ajout d'une classe *) let check_reference_arity ref = - if not (Reductionops.is_arity (Global.env()) Evd.empty (EConstr.of_constr (Global.type_of_global_unsafe ref))) (** FIXME *) then + let env = Global.env () in + let c, _ = Global.type_of_global_in_context env ref in + if not (Reductionops.is_arity env Evd.empty (EConstr.of_constr c)) (** FIXME *) then raise (CoercionError (NotAClass ref)) let check_arity = function @@ -252,7 +254,7 @@ let warn_uniform_inheritance = let add_new_coercion_core coef stre poly source target isid = check_source source; - let t = Global.type_of_global_unsafe coef in + let t, _ = Global.type_of_global_in_context (Global.env ()) coef in if coercion_exists coef then raise (CoercionError AlreadyExists); let tg,lp = prods_of t in let llp = List.length lp in diff --git a/vernac/classes.ml b/vernac/classes.ml index d6d4b164b..ab1892a18 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -68,7 +68,7 @@ let _ = let existing_instance glob g info = let c = global g in let info = Option.default Hints.empty_hint_info info in - let instance = Global.type_of_global_unsafe c in + let instance, _ = Global.type_of_global_in_context (Global.env ()) c in let _, r = decompose_prod_assum instance in match class_of_constr Evd.empty (EConstr.of_constr r) with | Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob diff --git a/vernac/search.ml b/vernac/search.ml index 5cf6a9491..0f56f81e7 100644 --- a/vernac/search.ml +++ b/vernac/search.ml @@ -78,7 +78,7 @@ let iter_declarations (fn : global_reference -> env -> constr -> unit) = | "CONSTANT" -> let cst = Global.constant_of_delta_kn kn in let gr = ConstRef cst in - let typ = Global.type_of_global_unsafe gr in + let (typ, _) = Global.type_of_global_in_context (Global.env ()) gr in fn gr env typ | "INDUCTIVE" -> let mind = Global.mind_of_delta_kn kn in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index bbec28aff..8a647c6c1 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -986,7 +986,7 @@ let vernac_arguments locality reference args more_implicits nargs_for_red flags let sr = smart_global reference in let inf_names = - let ty = Global.type_of_global_unsafe sr in + let ty, _ = Global.type_of_global_in_context (Global.env ()) sr in Impargs.compute_implicits_names (Global.env ()) ty in let prev_names = -- cgit v1.2.3