From 886a9c2fb25e32bd87b3fce38023b3e701134d23 Mon Sep 17 00:00:00 2001 From: Emilio Jesus Gallego Arias Date: Wed, 27 Dec 2017 20:22:23 +0100 Subject: [econstr] Continue consolidation of EConstr API under `interp`. This commit was motivated by true spurious conversions arising in my `to_constr` debug branch. The changes here need careful review as the tradeoffs are subtle and still a lot of clean up remains to be done in `vernac/*`. We have opted for penalize [minimally] the few users coming from true `Constr`-land, but I am sure we can tweak code in a much better way. In particular, it is not clear if internalization should take an `evar_map` even in the cases where it is not triggered, see the changes under `plugins` for a good example. Also, the new return type of `Pretyping.understand` should undergo careful review. We don't touch `Impargs` as it is not clear how to proceed, however, the current type of `compute_implicits_gen` looks very suspicious as it is called often with free evars. Some TODOs are: - impargs was calling whd_all, the Econstr equivalent can be either + Reductionops.whd_all [which does refolding and no sharing] + Reductionops.clos_whd_flags with all as a flag. --- dev/base_include | 5 +- .../06511-ejgallego-econstr+more_fix.sh | 7 ++ dev/top_printers.ml | 8 +- engine/termops.ml | 4 +- engine/termops.mli | 7 +- ide/ide_slave.ml | 2 +- interp/constrintern.ml | 46 ++++---- interp/constrintern.mli | 39 +++---- interp/declare.ml | 8 +- interp/impargs.ml | 128 +++++++++++---------- interp/impargs.mli | 6 +- interp/modintern.ml | 7 +- interp/notation.ml | 59 +++++----- interp/notation.mli | 6 +- plugins/funind/glob_term_to_relation.ml | 47 ++++---- plugins/funind/indfun.ml | 7 +- plugins/funind/recdef.ml | 9 +- plugins/ltac/extratactics.ml4 | 9 +- plugins/ltac/rewrite.ml | 1 - plugins/ltac/tacintern.ml | 6 +- plugins/ltac/tacinterp.ml | 2 +- plugins/setoid_ring/newring.ml | 2 +- plugins/ssr/ssrcommon.ml | 2 +- plugins/ssr/ssrvernac.ml4 | 8 +- plugins/ssr/ssrview.ml | 3 +- plugins/ssrmatching/ssrmatching.ml4 | 6 +- pretyping/inductiveops.ml | 7 +- pretyping/inductiveops.mli | 2 +- pretyping/pretyping.ml | 3 +- pretyping/pretyping.mli | 6 +- printing/ppconstr.ml | 9 +- proofs/proof.ml | 2 +- tactics/class_tactics.ml | 2 +- tactics/hints.ml | 8 +- tactics/tactics.ml | 1 - vernac/classes.ml | 2 +- vernac/comAssumption.ml | 8 +- vernac/comFixpoint.ml | 10 +- vernac/comInductive.ml | 22 ++-- vernac/comProgramFixpoint.ml | 4 +- vernac/obligations.ml | 6 +- vernac/record.ml | 4 +- vernac/vernacentries.ml | 21 ++-- 43 files changed, 286 insertions(+), 265 deletions(-) create mode 100644 dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh diff --git a/dev/base_include b/dev/base_include index 350ccaa10..39e3ef2bd 100644 --- a/dev/base_include +++ b/dev/base_include @@ -196,7 +196,10 @@ let parse_tac = Pcoq.parse_string Ltac_plugin.Pltac.tactic;; (* build a term of type glob_constr without type-checking or resolution of implicit syntax *) -let e s = Constrintern.intern_constr (Global.env()) (parse_constr s);; +let e s = + let env = Global.env () in + let sigma = Evd.from_env env in + Constrintern.intern_constr env sigma (parse_constr s);; (* build a term of type constr with type-checking and resolution of implicit syntax *) diff --git a/dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh b/dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh new file mode 100644 index 000000000..4b681909d --- /dev/null +++ b/dev/ci/user-overlays/06511-ejgallego-econstr+more_fix.sh @@ -0,0 +1,7 @@ + if [ "$CI_PULL_REQUEST" = "6511" ] || [ "$CI_BRANCH" = "econstr+more_fix" ]; then + ltac2_CI_BRANCH=econstr+more_fix + ltac2_CI_GITURL=https://github.com/ejgallego/ltac2 + + Equations_CI_BRANCH=econstr+more_fix + Equations_CI_GITURL=https://github.com/ejgallego/Coq-Equations +fi diff --git a/dev/top_printers.ml b/dev/top_printers.ml index f99e2593d..2c983fd8d 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -313,6 +313,8 @@ let constr_display csr = in pp (str (term_display csr) ++fnl ()) +let econstr_display c = constr_display EConstr.Unsafe.(to_constr c) ;; + open Format;; let print_pure_constr csr = @@ -452,6 +454,8 @@ let print_pure_constr csr = print_string (Printexc.to_string e);print_flush (); raise e +let print_pure_econstr c = print_pure_constr EConstr.Unsafe.(to_constr c) ;; + let pploc x = let (l,r) = Loc.unloc x in print_string"(";print_int l;print_string",";print_int r;print_string")" @@ -505,7 +509,7 @@ let _ = (function [c] when genarg_tag c = unquote (topwit wit_constr) && true -> let c = out_gen (rawwit wit_constr) c in - (fun ~atts ~st -> in_current_context constr_display c; st) + (fun ~atts ~st -> in_current_context econstr_display c; st) | _ -> failwith "Vernac extension: cannot occur") with e -> pp (CErrors.print e) @@ -521,7 +525,7 @@ let _ = (function [c] when genarg_tag c = unquote (topwit wit_constr) && true -> let c = out_gen (rawwit wit_constr) c in - (fun ~atts ~st -> in_current_context print_pure_constr c; st) + (fun ~atts ~st -> in_current_context print_pure_econstr c; st) | _ -> failwith "Vernac extension: cannot occur") with e -> pp (CErrors.print e) diff --git a/engine/termops.ml b/engine/termops.ml index 40b3d0d8b..668ae8777 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -797,9 +797,9 @@ let fold_constr_with_binders sigma g f n acc c = each binder traversal; it is not recursive and the order with which subterms are processed is not specified *) -let iter_constr_with_full_binders g f l c = +let iter_constr_with_full_binders sigma g f l c = let open RelDecl in - match kind c with + match EConstr.kind sigma c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _) -> () | Cast (c,_, t) -> f l c; f l t diff --git a/engine/termops.mli b/engine/termops.mli index a3559a693..5aa6235f5 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -76,9 +76,10 @@ val fold_constr_with_full_binders : Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b -val iter_constr_with_full_binders : - (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> Constr.constr -> unit) -> 'a -> - Constr.constr -> unit +val iter_constr_with_full_binders : Evd.evar_map -> + (rel_declaration -> 'a -> 'a) -> + ('a -> constr -> unit) -> 'a -> + constr -> unit (**********************************************************************) diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index fe86df084..8f6ae9760 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -281,7 +281,7 @@ let pattern_of_string ?env s = | Some e -> e in let constr = Pcoq.parse_string Pcoq.Constr.lconstr_pattern s in - let (_, pat) = Constrintern.intern_constr_pattern env constr in + let (_, pat) = Constrintern.intern_constr_pattern env Evd.empty constr in pat let dirpath_of_string_list s = diff --git a/interp/constrintern.ml b/interp/constrintern.ml index d03aa3552..9053cdc24 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -184,14 +184,14 @@ let compute_explicitable_implicit imps = function (* Unable to know in advance what the implicit arguments will be *) [] -let compute_internalization_data env ty typ impl = - let impl = compute_implicits_with_manual env typ (is_implicit_args()) impl in +let compute_internalization_data env sigma ty typ impl = + let impl = compute_implicits_with_manual env sigma typ (is_implicit_args()) impl in let expls_impl = compute_explicitable_implicit impl ty in - (ty, expls_impl, impl, compute_arguments_scope typ) + (ty, expls_impl, impl, compute_arguments_scope sigma typ) -let compute_internalization_env env ?(impls=empty_internalization_env) ty = +let compute_internalization_env env sigma ?(impls=empty_internalization_env) ty = List.fold_left3 - (fun map id typ impl -> Id.Map.add id (compute_internalization_data env ty typ impl) map) + (fun map id typ impl -> Id.Map.add id (compute_internalization_data env sigma ty typ impl) map) impls (**********************************************************************) @@ -2207,9 +2207,9 @@ let extract_ids env = (Termops.ids_of_rel_context (Environ.rel_context env)) Id.Set.empty -let scope_of_type_kind = function +let scope_of_type_kind sigma = function | IsType -> Notation.current_type_scope_name () - | OfType typ -> compute_type_scope (EConstr.Unsafe.to_constr typ) + | OfType typ -> compute_type_scope sigma typ | WithoutTypeConstraint -> None let empty_ltac_sign = { @@ -2218,19 +2218,17 @@ let empty_ltac_sign = { ltac_extra = Genintern.Store.empty; } -let intern_gen kind env +let intern_gen kind env sigma ?(impls=empty_internalization_env) ?(pattern_mode=false) ?(ltacvars=empty_ltac_sign) c = - let tmp_scope = scope_of_type_kind kind in + let tmp_scope = scope_of_type_kind sigma kind in internalize env {ids = extract_ids env; unb = false; tmp_scope = tmp_scope; scopes = []; impls = impls} pattern_mode (ltacvars, Id.Map.empty) c -let intern_constr env c = intern_gen WithoutTypeConstraint env c - -let intern_type env c = intern_gen IsType env c - +let intern_constr env sigma c = intern_gen WithoutTypeConstraint env sigma c +let intern_type env sigma c = intern_gen IsType env sigma c let intern_pattern globalenv patt = try intern_cases_pattern globalenv Id.Map.empty (None,[]) empty_alias patt @@ -2245,7 +2243,7 @@ let intern_pattern globalenv patt = (* All evars resolved *) let interp_gen kind env sigma ?(impls=empty_internalization_env) c = - let c = intern_gen kind ~impls env c in + let c = intern_gen kind ~impls env sigma c in understand ~expected_type:kind env sigma c let interp_constr env sigma ?(impls=empty_internalization_env) c = @@ -2260,13 +2258,13 @@ let interp_casted_constr env sigma ?(impls=empty_internalization_env) c typ = (* Not all evars expected to be resolved *) let interp_open_constr env sigma c = - understand_tcc env sigma (intern_constr env c) + understand_tcc env sigma (intern_constr env sigma c) (* Not all evars expected to be resolved and computation of implicit args *) let interp_constr_evars_gen_impls env sigma ?(impls=empty_internalization_env) expected_type c = - let c = intern_gen expected_type ~impls env c in + let c = intern_gen expected_type ~impls env sigma c in let imps = Implicit_quantifiers.implicits_of_glob_constr ~with_products:(expected_type == IsType) c in let sigma, c = understand_tcc env sigma ~expected_type c in sigma, (c, imps) @@ -2283,7 +2281,7 @@ let interp_type_evars_impls env sigma ?(impls=empty_internalization_env) c = (* Not all evars expected to be resolved, with side-effect on evars *) let interp_constr_evars_gen env sigma ?(impls=empty_internalization_env) expected_type c = - let c = intern_gen expected_type ~impls env c in + let c = intern_gen expected_type ~impls env sigma c in understand_tcc env sigma ~expected_type c let interp_constr_evars env evdref ?(impls=empty_internalization_env) c = @@ -2297,9 +2295,9 @@ let interp_type_evars env sigma ?(impls=empty_internalization_env) c = (* Miscellaneous *) -let intern_constr_pattern env ?(as_type=false) ?(ltacvars=empty_ltac_sign) c = +let intern_constr_pattern env sigma ?(as_type=false) ?(ltacvars=empty_ltac_sign) c = let c = intern_gen (if as_type then IsType else WithoutTypeConstraint) - ~pattern_mode:true ~ltacvars env c in + ~pattern_mode:true ~ltacvars env sigma c in pattern_of_glob_constr c let interp_notation_constr env ?(impls=empty_internalization_env) nenv a = @@ -2323,16 +2321,14 @@ let interp_notation_constr env ?(impls=empty_internalization_env) nenv a = (* Interpret binders and contexts *) let interp_binder env sigma na t = - let t = intern_gen IsType env t in + let t = intern_gen IsType env sigma t in let t' = locate_if_hole ?loc:(loc_of_glob_constr t) na t in understand ~expected_type:IsType env sigma t' -let interp_binder_evars env evdref na t = - let t = intern_gen IsType env t in +let interp_binder_evars env sigma na t = + let t = intern_gen IsType env sigma t in let t' = locate_if_hole ?loc:(loc_of_glob_constr t) na t in - let evd, c = understand_tcc env !evdref ~expected_type:IsType t' in - evdref := evd; - c + understand_tcc env sigma ~expected_type:IsType t' let my_intern_constr env lvar acc c = internalize env acc false lvar c diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 7411fb84b..563230b59 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -7,17 +7,17 @@ (************************************************************************) open Names -open Constr open Evd open Environ +open Misctypes open Libnames open Globnames open Glob_term open Pattern +open EConstr open Constrexpr open Notation_term open Pretyping -open Misctypes (** Translation from front abstract syntax of term to untyped terms (glob_constr) *) @@ -58,10 +58,10 @@ type internalization_env = var_internalization_data Id.Map.t val empty_internalization_env : internalization_env -val compute_internalization_data : env -> var_internalization_type -> +val compute_internalization_data : env -> evar_map -> var_internalization_type -> types -> Impargs.manual_explicitation list -> var_internalization_data -val compute_internalization_env : env -> ?impls:internalization_env -> var_internalization_type -> +val compute_internalization_env : env -> evar_map -> ?impls:internalization_env -> var_internalization_type -> Id.t list -> types list -> Impargs.manual_explicitation list list -> internalization_env @@ -78,11 +78,10 @@ val empty_ltac_sign : ltac_sign (** {6 Internalization performs interpretation of global names and notations } *) -val intern_constr : env -> constr_expr -> glob_constr - -val intern_type : env -> constr_expr -> glob_constr +val intern_constr : env -> evar_map -> constr_expr -> glob_constr +val intern_type : env -> evar_map -> constr_expr -> glob_constr -val intern_gen : typing_constraint -> env -> +val intern_gen : typing_constraint -> env -> evar_map -> ?impls:internalization_env -> ?pattern_mode:bool -> ?ltacvars:ltac_sign -> constr_expr -> glob_constr @@ -100,7 +99,7 @@ val interp_constr : env -> evar_map -> ?impls:internalization_env -> constr_expr -> constr Evd.in_evar_universe_context val interp_casted_constr : env -> evar_map -> ?impls:internalization_env -> - constr_expr -> EConstr.types -> constr Evd.in_evar_universe_context + constr_expr -> types -> constr Evd.in_evar_universe_context val interp_type : env -> evar_map -> ?impls:internalization_env -> constr_expr -> types Evd.in_evar_universe_context @@ -108,37 +107,37 @@ val interp_type : env -> evar_map -> ?impls:internalization_env -> (** Main interpretation function expecting all postponed problems to be resolved, but possibly leaving evars. *) -val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * EConstr.constr +val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * constr (** Accepting unresolved evars *) val interp_constr_evars : env -> evar_map -> - ?impls:internalization_env -> constr_expr -> evar_map * EConstr.constr + ?impls:internalization_env -> constr_expr -> evar_map * constr val interp_casted_constr_evars : env -> evar_map -> - ?impls:internalization_env -> constr_expr -> EConstr.types -> evar_map * EConstr.constr + ?impls:internalization_env -> constr_expr -> types -> evar_map * constr val interp_type_evars : env -> evar_map -> - ?impls:internalization_env -> constr_expr -> evar_map * EConstr.types + ?impls:internalization_env -> constr_expr -> evar_map * types (** Accepting unresolved evars and giving back the manual implicit arguments *) val interp_constr_evars_impls : env -> evar_map -> ?impls:internalization_env -> constr_expr -> - evar_map * (EConstr.constr * Impargs.manual_implicits) + evar_map * (constr * Impargs.manual_implicits) val interp_casted_constr_evars_impls : env -> evar_map -> - ?impls:internalization_env -> constr_expr -> EConstr.types -> - evar_map * (EConstr.constr * Impargs.manual_implicits) + ?impls:internalization_env -> constr_expr -> types -> + evar_map * (constr * Impargs.manual_implicits) val interp_type_evars_impls : env -> evar_map -> ?impls:internalization_env -> constr_expr -> - evar_map * (EConstr.types * Impargs.manual_implicits) + evar_map * (types * Impargs.manual_implicits) (** Interprets constr patterns *) val intern_constr_pattern : - env -> ?as_type:bool -> ?ltacvars:ltac_sign -> + env -> evar_map -> ?as_type:bool -> ?ltacvars:ltac_sign -> constr_pattern_expr -> patvar list * constr_pattern (** Raise Not_found if syndef not bound to a name and error if unexisting ref *) @@ -152,14 +151,14 @@ val interp_reference : ltac_sign -> reference -> glob_constr val interp_binder : env -> evar_map -> Name.t -> constr_expr -> types Evd.in_evar_universe_context -val interp_binder_evars : env -> evar_map ref -> Name.t -> constr_expr -> EConstr.types +val interp_binder_evars : env -> evar_map -> Name.t -> constr_expr -> evar_map * types (** Interpret contexts: returns extended env and context *) val interp_context_evars : ?global_level:bool -> ?impl_env:internalization_env -> ?shift:int -> env -> evar_map -> local_binder_expr list -> - evar_map * (internalization_env * ((env * EConstr.rel_context) * Impargs.manual_implicits)) + evar_map * (internalization_env * ((env * rel_context) * Impargs.manual_implicits)) (* val interp_context_gen : (env -> glob_constr -> unsafe_type_judgment Evd.in_evar_universe_context) -> *) (* (env -> Evarutil.type_constraint -> glob_constr -> unsafe_judgment Evd.in_evar_universe_context) -> *) diff --git a/interp/declare.ml b/interp/declare.ml index dfa84f278..f6d7b45c3 100644 --- a/interp/declare.ml +++ b/interp/declare.ml @@ -135,7 +135,7 @@ let set_declare_scheme f = declare_scheme := f let update_tables c = declare_constant_implicits c; Heads.declare_head (EvalConstRef c); - Notation.declare_ref_arguments_scope (ConstRef c) + Notation.declare_ref_arguments_scope Evd.empty (ConstRef c) let register_side_effect (c, role) = let o = inConstant { @@ -275,7 +275,7 @@ let inVariable : variable_obj -> obj = let declare_variable id obj = let oname = add_leaf id (inVariable (Inr (id,obj))) in declare_var_implicits id; - Notation.declare_ref_arguments_scope (VarRef id); + Notation.declare_ref_arguments_scope Evd.empty (VarRef id); Heads.declare_head (EvalVarRef id); oname @@ -283,9 +283,9 @@ let declare_variable id obj = let declare_inductive_argument_scopes kn mie = List.iteri (fun i {mind_entry_consnames=lc} -> - Notation.declare_ref_arguments_scope (IndRef (kn,i)); + Notation.declare_ref_arguments_scope Evd.empty (IndRef (kn,i)); for j=1 to List.length lc do - Notation.declare_ref_arguments_scope (ConstructRef ((kn,i),j)); + Notation.declare_ref_arguments_scope Evd.empty (ConstructRef ((kn,i),j)); done) mie.mind_entry_inds let inductive_names sp kn mie = diff --git a/interp/impargs.ml b/interp/impargs.ml index ed1cd5276..6767af6d8 100644 --- a/interp/impargs.ml +++ b/interp/impargs.ml @@ -7,21 +7,20 @@ (************************************************************************) open CErrors +open Pp open Util open Names -open Globnames open Constr -open Reduction +open Globnames open Declarations -open Environ -open Libobject +open Decl_kinds open Lib -open Pp -open Constrexpr +open Libobject +open EConstr open Termops +open Reductionops +open Constrexpr open Namegen -open Decl_kinds -open Context.Named.Declaration module NamedDecl = Context.Named.Declaration @@ -165,8 +164,8 @@ let update pos rig (na,st) = in na, Some e (* modified is_rigid_reference with a truncated env *) -let is_flexible_reference env bound depth f = - match kind f with +let is_flexible_reference env sigma bound depth f = + match kind sigma f with | Rel n when n >= bound+depth -> (* inductive type *) false | Rel n when n >= depth -> (* previous argument *) true | Rel n -> (* since local definitions have been expanded *) false @@ -174,102 +173,101 @@ let is_flexible_reference env bound depth f = let cb = Environ.lookup_constant kn env in (match cb.const_body with Def _ -> true | _ -> false) | Var id -> - env |> Environ.lookup_named id |> is_local_def + env |> Environ.lookup_named id |> NamedDecl.is_local_def | Ind _ | Construct _ -> false | _ -> true let push_lift d (e,n) = (push_rel d e,n+1) -let is_reversible_pattern bound depth f l = - isRel f && let n = destRel f in (n < bound+depth) && (n >= depth) && - Array.for_all (fun c -> isRel c && destRel c < depth) l && +let is_reversible_pattern sigma bound depth f l = + isRel sigma f && let n = destRel sigma f in (n < bound+depth) && (n >= depth) && + Array.for_all (fun c -> isRel sigma c && destRel sigma c < depth) l && Array.distinct l (* Precondition: rels in env are for inductive types only *) -let add_free_rels_until strict strongly_strict revpat bound env m pos acc = +let add_free_rels_until strict strongly_strict revpat bound env sigma m pos acc = let rec frec rig (env,depth as ed) c = - let hd = if strict then whd_all env c else c in + let hd = if strict then whd_all env sigma c else c in let c = if strongly_strict then hd else c in - match kind hd with + match kind sigma hd with | Rel n when (n < bound+depth) && (n >= depth) -> let i = bound + depth - n - 1 in acc.(i) <- update pos rig acc.(i) - | App (f,l) when revpat && is_reversible_pattern bound depth f l -> - let i = bound + depth - destRel f - 1 in + | App (f,l) when revpat && is_reversible_pattern sigma bound depth f l -> + let i = bound + depth - EConstr.destRel sigma f - 1 in acc.(i) <- update pos rig acc.(i) - | App (f,_) when rig && is_flexible_reference env bound depth f -> + | App (f,_) when rig && is_flexible_reference env sigma bound depth f -> if strict then () else - iter_constr_with_full_binders push_lift (frec false) ed c + iter_constr_with_full_binders sigma push_lift (frec false) ed c | Proj (p,c) when rig -> if strict then () else - iter_constr_with_full_binders push_lift (frec false) ed c + iter_constr_with_full_binders sigma push_lift (frec false) ed c | Case _ when rig -> if strict then () else - iter_constr_with_full_binders push_lift (frec false) ed c + iter_constr_with_full_binders sigma push_lift (frec false) ed c | Evar _ -> () | _ -> - iter_constr_with_full_binders push_lift (frec rig) ed c + iter_constr_with_full_binders sigma push_lift (frec rig) ed c in - let () = if not (Vars.noccur_between 1 bound m) then frec true (env,1) m in + let () = if not (Vars.noccur_between sigma 1 bound m) then frec true (env,1) m in acc -let rec is_rigid_head t = match kind t with +let rec is_rigid_head sigma t = match kind sigma t with | Rel _ | Evar _ -> false | Ind _ | Const _ | Var _ | Sort _ -> true - | Case (_,_,f,_) -> is_rigid_head f + | Case (_,_,f,_) -> is_rigid_head sigma f | Proj (p,c) -> true | App (f,args) -> - (match kind f with - | Fix ((fi,i),_) -> is_rigid_head (args.(fi.(i))) - | _ -> is_rigid_head f) + (match kind sigma f with + | Fix ((fi,i),_) -> is_rigid_head sigma (args.(fi.(i))) + | _ -> is_rigid_head sigma f) | Lambda _ | LetIn _ | Construct _ | CoFix _ | Fix _ | Prod _ | Meta _ | Cast _ -> assert false (* calcule la liste des arguments implicites *) let find_displayed_name_in all avoid na (env, b) = - let b = EConstr.of_constr b in let envnames_b = (env, b) in let flag = RenamingElsewhereFor envnames_b in if all then compute_and_force_displayed_name_in Evd.empty flag avoid na b else compute_displayed_name_in Evd.empty flag avoid na b -let compute_implicits_gen strict strongly_strict revpat contextual all env t = +let compute_implicits_gen strict strongly_strict revpat contextual all env sigma (t : EConstr.t) = let rigid = ref true in let open Context.Rel.Declaration in - let rec aux env avoid n names t = - let t = whd_all env t in - match kind t with + let rec aux env avoid n names (t : EConstr.t) = + let t = whd_all env sigma t in + match kind sigma t with | Prod (na,a,b) -> let na',avoid' = find_displayed_name_in all avoid na (names,b) in - add_free_rels_until strict strongly_strict revpat n env a (Hyp (n+1)) + add_free_rels_until strict strongly_strict revpat n env sigma a (Hyp (n+1)) (aux (push_rel (LocalAssum (na',a)) env) avoid' (n+1) (na'::names) b) | _ -> - rigid := is_rigid_head t; + rigid := is_rigid_head sigma t; let names = List.rev names in let v = Array.map (fun na -> na,None) (Array.of_list names) in if contextual then - add_free_rels_until strict strongly_strict revpat n env t Conclusion v + add_free_rels_until strict strongly_strict revpat n env sigma t Conclusion v else v in - match kind (whd_all env t) with + match kind sigma (whd_all env sigma t) with | Prod (na,a,b) -> let na',avoid = find_displayed_name_in all Id.Set.empty na ([],b) in let v = aux (push_rel (LocalAssum (na',a)) env) avoid 1 [na'] b in !rigid, Array.to_list v | _ -> true, [] -let compute_implicits_flags env f all t = +let compute_implicits_flags env sigma f all t = compute_implicits_gen (f.strict || f.strongly_strict) f.strongly_strict - f.reversible_pattern f.contextual all env t + f.reversible_pattern f.contextual all env sigma t -let compute_auto_implicits env flags enriching t = - if enriching then compute_implicits_flags env flags true t - else compute_implicits_gen false false false true true env t +let compute_auto_implicits env sigma flags enriching t = + if enriching then compute_implicits_flags env sigma flags true t + else compute_implicits_gen false false false true true env sigma t -let compute_implicits_names env t = - let _, impls = compute_implicits_gen false false false false true env t in +let compute_implicits_names env sigma t = + let _, impls = compute_implicits_gen false false false false true env sigma t in List.map fst impls (* Extra information about implicit arguments *) @@ -398,24 +396,25 @@ let set_manual_implicits env flags enriching autoimps l = in merge 1 l autoimps -let compute_semi_auto_implicits env f manual t = +let compute_semi_auto_implicits env sigma f manual t = match manual with | [] -> if not f.auto then [DefaultImpArgs, []] - else let _,l = compute_implicits_flags env f false t in + else let _,l = compute_implicits_flags env sigma f false t in [DefaultImpArgs, prepare_implicits f l] | _ -> - let _,autoimpls = compute_auto_implicits env f f.auto t in + let _,autoimpls = compute_auto_implicits env sigma f f.auto t in [DefaultImpArgs, set_manual_implicits env f f.auto autoimpls manual] (*s Constants. *) let compute_constant_implicits flags manual cst = let env = Global.env () in + let sigma = Evd.from_env env in let cb = Environ.lookup_constant cst env in - let ty = cb.const_type in - let impls = compute_semi_auto_implicits env flags manual ty in - impls + let ty = of_constr cb.const_type in + let impls = compute_semi_auto_implicits env sigma flags manual ty in + impls (*s Inductives and constructors. Their implicit arguments are stored in an array, indexed by the inductive number, of pairs $(i,v)$ where @@ -424,7 +423,8 @@ let compute_constant_implicits flags manual cst = let compute_mib_implicits flags manual kn = let env = Global.env () in - let mib = lookup_mind kn env in + let sigma = Evd.from_env env in + let mib = Environ.lookup_mind kn env in let ar = Array.to_list (Array.mapi (* No need to lift, arities contain no de Bruijn *) @@ -433,14 +433,14 @@ let compute_mib_implicits flags manual kn = 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 env_ar = Environ.push_rel_context ar env in let imps_one_inductive i mip = let ind = (kn,i) in let ar, _ = Global.type_of_global_in_context env (IndRef ind) in - ((IndRef ind,compute_semi_auto_implicits env flags manual ar), + ((IndRef ind,compute_semi_auto_implicits env sigma flags manual (of_constr ar)), Array.mapi (fun j c -> - (ConstructRef (ind,j+1),compute_semi_auto_implicits env_ar flags manual c)) - mip.mind_nf_lc) + (ConstructRef (ind,j+1),compute_semi_auto_implicits env_ar sigma flags manual c)) + (Array.map of_constr mip.mind_nf_lc)) in Array.mapi imps_one_inductive mib.mind_packets @@ -453,7 +453,8 @@ let compute_all_mib_implicits flags manual kn = let compute_var_implicits flags manual id = let env = Global.env () in - compute_semi_auto_implicits env flags manual (NamedDecl.get_type (lookup_named id env)) + let sigma = Evd.from_env env in + compute_semi_auto_implicits env sigma flags manual (NamedDecl.get_type (lookup_named id env)) (* Implicits of a global reference. *) @@ -524,7 +525,7 @@ let impls_of_context ctx = | Implicit -> Some (NamedDecl.get_id decl, Manual, (true, true)) | _ -> None in - List.rev_map map (List.filter (fst %> is_local_assum) ctx) + List.rev_map map (List.filter (fst %> NamedDecl.is_local_assum) ctx) let adjust_side_condition p = function | LessArgsThan n -> LessArgsThan (n+p) @@ -649,8 +650,8 @@ type manual_explicitation = Constrexpr.explicitation * (bool * bool * bool) type manual_implicits = manual_explicitation list -let compute_implicits_with_manual env typ enriching l = - let _,autoimpls = compute_auto_implicits env !implicit_args enriching typ in +let compute_implicits_with_manual env sigma typ enriching l = + let _,autoimpls = compute_auto_implicits env sigma !implicit_args enriching typ in set_manual_implicits env !implicit_args enriching autoimpls l let check_inclusion l = @@ -674,9 +675,10 @@ 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_in_context (Global.env ()) ref in + let sigma = Evd.from_env env in + let t, _ = Global.type_of_global_in_context env ref in let enriching = Option.default flags.auto enriching in - let isrigid,autoimpls = compute_auto_implicits env flags enriching t in + let isrigid,autoimpls = compute_auto_implicits env sigma flags enriching (of_constr t) in let l' = match l with | [] -> assert false | [l] -> diff --git a/interp/impargs.mli b/interp/impargs.mli index 40fa4cb26..dcfe527b0 100644 --- a/interp/impargs.mli +++ b/interp/impargs.mli @@ -7,7 +7,7 @@ (************************************************************************) open Names -open Constr +open EConstr open Globnames open Environ @@ -90,10 +90,10 @@ type manual_explicitation = Constrexpr.explicitation * type manual_implicits = manual_explicitation list -val compute_implicits_with_manual : env -> types -> bool -> +val compute_implicits_with_manual : env -> Evd.evar_map -> types -> bool -> manual_implicits -> implicit_status list -val compute_implicits_names : env -> types -> Name.t list +val compute_implicits_names : env -> Evd.evar_map -> types -> Name.t list (** {6 Computation of implicits (done using the global environment). } *) diff --git a/interp/modintern.ml b/interp/modintern.ml index e631b3ea4..128152cc2 100644 --- a/interp/modintern.ml +++ b/interp/modintern.ml @@ -61,13 +61,16 @@ let transl_with_decl env = function | CWith_Module ((_,fqid),qid) -> WithMod (fqid,lookup_module qid), Univ.ContextSet.empty | CWith_Definition ((_,fqid),c) -> - let c, ectx = interp_constr env (Evd.from_env env) c in + let sigma = Evd.from_env env in + let c, ectx = interp_constr env sigma c in if Flags.is_universe_polymorphism () then let ctx = UState.context ectx in let inst, ctx = Univ.abstract_universes ctx in - let c = Vars.subst_univs_level_constr (Univ.make_instance_subst inst) c in + let c = EConstr.Vars.subst_univs_level_constr (Univ.make_instance_subst inst) c in + let c = EConstr.to_constr sigma c in WithDef (fqid,(c, Some ctx)), Univ.ContextSet.empty else + let c = EConstr.to_constr sigma c in WithDef (fqid,(c, None)), UState.context_set ectx let loc_of_module l = l.CAst.loc diff --git a/interp/notation.ml b/interp/notation.ml index ea7ef21b1..bf39c726a 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -668,8 +668,8 @@ type scope_class = cl_typ let scope_class_compare : scope_class -> scope_class -> int = cl_typ_ord -let compute_scope_class t = - let (cl,_,_) = find_class_type Evd.empty (EConstr.of_constr t) in +let compute_scope_class sigma t = + let (cl,_,_) = find_class_type sigma t in cl module ScopeClassOrd = @@ -698,22 +698,22 @@ let find_scope_class_opt = function (**********************************************************************) (* Special scopes associated to arguments of a global reference *) -let rec compute_arguments_classes t = - match Constr.kind (EConstr.Unsafe.to_constr (Reductionops.whd_betaiotazeta Evd.empty (EConstr.of_constr t))) with +let rec compute_arguments_classes sigma t = + match EConstr.kind sigma (Reductionops.whd_betaiotazeta sigma t) with | Prod (_,t,u) -> - let cl = try Some (compute_scope_class t) with Not_found -> None in - cl :: compute_arguments_classes u + let cl = try Some (compute_scope_class sigma t) with Not_found -> None in + cl :: compute_arguments_classes sigma u | _ -> [] -let compute_arguments_scope_full t = - let cls = compute_arguments_classes t in +let compute_arguments_scope_full sigma t = + let cls = compute_arguments_classes sigma t in let scs = List.map find_scope_class_opt cls in scs, cls -let compute_arguments_scope t = fst (compute_arguments_scope_full t) +let compute_arguments_scope sigma t = fst (compute_arguments_scope_full sigma t) -let compute_type_scope t = - find_scope_class_opt (try Some (compute_scope_class t) with Not_found -> None) +let compute_type_scope sigma t = + find_scope_class_opt (try Some (compute_scope_class sigma t) with Not_found -> None) let current_type_scope_name () = find_scope_class_opt (Some CL_SORT) @@ -779,20 +779,24 @@ let discharge_arguments_scope (_,(req,r,n,l,_)) = let classify_arguments_scope (req,_,_,_,_ as obj) = if req == ArgsScopeNoDischarge then Dispose else Substitute obj -let rebuild_arguments_scope (req,r,n,l,_) = +let rebuild_arguments_scope sigma (req,r,n,l,_) = match req with | ArgsScopeNoDischarge -> assert false | ArgsScopeAuto -> - 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) + let env = Global.env () in (*FIXME?*) + let typ = EConstr.of_constr @@ fst (Global.type_of_global_in_context env r) in + let scs,cls = compute_arguments_scope_full sigma typ 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 (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) + (* 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 env = Global.env () in (*FIXME?*) + let typ = EConstr.of_constr @@ fst (Global.type_of_global_in_context env r) in + let l',cls = compute_arguments_scope_full sigma typ in + let l1 = List.firstn n l' in + let cls1 = List.firstn n cls in + (req,r,0,l1@l,cls1) type arguments_scope_obj = arguments_scope_discharge_request * global_reference * @@ -807,7 +811,8 @@ let inArgumentsScope : arguments_scope_obj -> obj = subst_function = subst_arguments_scope; classify_function = classify_arguments_scope; discharge_function = discharge_arguments_scope; - rebuild_function = rebuild_arguments_scope } + (* XXX: Should we pass the sigma here or not, see @herbelin's comment in 6511 *) + rebuild_function = rebuild_arguments_scope Evd.empty } let is_local local ref = local || isVarRef ref && Lib.is_in_section ref @@ -819,7 +824,7 @@ let declare_arguments_scope local r scl = (* We empty the list of argument classes to disable further scope re-computations and keep these manually given scopes. *) declare_arguments_scope_gen req r 0 (scl,[]) - + let find_arguments_scope r = try let (scl,cls,stamp) = Refmap.find r !arguments_scope in @@ -832,12 +837,12 @@ let find_arguments_scope r = scl' with Not_found -> [] -let declare_ref_arguments_scope ref = - let t, _ = Global.type_of_global_in_context (Global.env ()) ref in - let (scs,cls as o) = compute_arguments_scope_full t in +let declare_ref_arguments_scope sigma ref = + let env = Global.env () in (* FIXME? *) + let typ = EConstr.of_constr @@ fst @@ Global.type_of_global_in_context env ref in + let (scs,cls as o) = compute_arguments_scope_full sigma typ in declare_arguments_scope_gen ArgsScopeAuto ref (List.length scs) o - (********************************) (* Encoding notations as string *) diff --git a/interp/notation.mli b/interp/notation.mli index a4c79d6d3..79b56bd41 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -163,10 +163,10 @@ val subst_scope_class : Mod_subst.substitution -> scope_class -> scope_class option val declare_scope_class : scope_name -> scope_class -> unit -val declare_ref_arguments_scope : global_reference -> unit +val declare_ref_arguments_scope : Evd.evar_map -> global_reference -> unit -val compute_arguments_scope : Constr.types -> scope_name option list -val compute_type_scope : Constr.types -> scope_name option +val compute_arguments_scope : Evd.evar_map -> EConstr.types -> scope_name option list +val compute_type_scope : Evd.evar_map -> EConstr.types -> scope_name option (** Get the current scope bound to Sortclass, if it exists *) val current_type_scope_name : unit -> scope_name option diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 22881c32c..7159614d9 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -352,9 +352,9 @@ let raw_push_named (na,raw_value,raw_typ) env = let typ,_ = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in (match raw_value with | None -> - Environ.push_named (NamedDecl.LocalAssum (id,typ)) env + EConstr.push_named (NamedDecl.LocalAssum (id,typ)) env | Some value -> - Environ.push_named (NamedDecl.LocalDef (id, value, typ)) env) + EConstr.push_named (NamedDecl.LocalDef (id, value, typ)) env) let add_pat_variables pat typ env : Environ.env = @@ -519,7 +519,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = The "value" of this branch is then simply [res] *) let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in - let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr rt_as_constr) in + let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) rt_as_constr in let res_raw_type = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) rt_typ in let res = fresh_id args_res.to_avoid "_res" in let new_avoid = res::args_res.to_avoid in @@ -631,12 +631,11 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let v = match typ with None -> v | Some t -> DAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in let v_res = build_entry_lc env funnames avoid v in let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in - let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in - let v_type = EConstr.Unsafe.to_constr v_type in + let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in let new_env = match n with Anonymous -> env - | Name id -> Environ.push_named (NamedDecl.LocalDef (id,v_as_constr,v_type)) env + | Name id -> EConstr.push_named (NamedDecl.LocalDef (id,v_as_constr,v_type)) env in let b_res = build_entry_lc new_env funnames avoid b in combine_results (combine_letin n) v_res b_res @@ -648,7 +647,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = build_entry_lc_from_case env funnames make_discr el brl avoid | GIf(b,(na,e_option),lhs,rhs) -> let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in - let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in + let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in let (ind,_) = try Inductiveops.find_inductive env (Evd.from_env env) b_typ with Not_found -> @@ -680,7 +679,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = nal in let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in - let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in + let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in let (ind,_) = try Inductiveops.find_inductive env (Evd.from_env env) b_typ with Not_found -> @@ -726,7 +725,7 @@ and build_entry_lc_from_case env funname make_discr let types = List.map (fun (case_arg,_) -> let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in - EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr case_arg_as_constr)) + EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) case_arg_as_constr) ) el in (****** The next works only if the match is not dependent ****) @@ -948,7 +947,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = mkGApp(mkGVar(mk_rel_id this_relname),List.tl args'@[res_rt]) in let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in - let new_env = Environ.push_rel (LocalAssum (n,t')) env in + let new_env = EConstr.push_rel (LocalAssum (n,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname @@ -983,7 +982,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let subst_b = if is_in_b then b else replace_var_by_term id rt b in - let new_env = Environ.push_rel (LocalAssum (n,t')) env in + let new_env = EConstr.push_rel (LocalAssum (n,t')) env in let new_b,id_to_exclude = rebuild_cons new_env @@ -995,7 +994,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = with Continue -> let jmeq = Globnames.IndRef (fst (EConstr.destInd Evd.empty (jmeq ()))) in let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in - let ind,args' = Inductive.find_inductive env ty' in + let ind,args' = Inductiveops.find_inductive env Evd.(from_env env) ty' in let mib,_ = Global.lookup_inductive (fst ind) in let nparam = mib.Declarations.mind_nparams in let params,arg' = @@ -1017,14 +1016,14 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = observe (str "computing new type for jmeq : " ++ pr_glob_constr_env env eq'); let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in observe (str " computing new type for jmeq : done") ; + let sigma = Evd.(from_env env) in let new_args = - match Constr.kind eq'_as_constr with + match EConstr.kind sigma eq'_as_constr with | App(_,[|_;_;ty;_|]) -> - let ty = Array.to_list (snd (destApp ty)) in + let ty = Array.to_list (snd (EConstr.destApp sigma ty)) in let ty' = snd (Util.List.chop nparam ty) in List.fold_left2 (fun acc var_as_constr arg -> - let arg = EConstr.of_constr arg in if isRel var_as_constr then let na = RelDecl.get_name (Environ.lookup_rel (destRel var_as_constr) env) in @@ -1065,7 +1064,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = in let new_env = let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in - Environ.push_rel (LocalAssum (n,t')) env + EConstr.push_rel (LocalAssum (n,t')) env in let new_b,id_to_exclude = rebuild_cons @@ -1103,7 +1102,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = with Continue -> observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt); let t',ctx = Pretyping.understand env (Evd.from_env env) t in - let new_env = Environ.push_rel (LocalAssum (n,t')) env in + let new_env = EConstr.push_rel (LocalAssum (n,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname @@ -1119,7 +1118,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = | _ -> observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt); let t',ctx = Pretyping.understand env (Evd.from_env env) t in - let new_env = Environ.push_rel (LocalAssum (n,t')) env in + let new_env = EConstr.push_rel (LocalAssum (n,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname @@ -1140,7 +1139,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let t',ctx = Pretyping.understand env (Evd.from_env env) t in match n with | Name id -> - let new_env = Environ.push_rel (LocalAssum (n,t')) env in + let new_env = EConstr.push_rel (LocalAssum (n,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname @@ -1163,7 +1162,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let evd = (Evd.from_env env) in let t',ctx = Pretyping.understand env evd t in let evd = Evd.from_ctx ctx in - let type_t' = Typing.unsafe_type_of env evd (EConstr.of_constr t') in + let type_t' = Typing.unsafe_type_of env evd t' in + let t' = EConstr.Unsafe.to_constr t' in let type_t' = EConstr.Unsafe.to_constr type_t' in let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in let new_b,id_to_exclude = @@ -1189,7 +1189,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = depth t in let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in - let new_env = Environ.push_rel (LocalAssum (na,t')) env in + let new_env = EConstr.push_rel (LocalAssum (na,t')) env in let new_b,id_to_exclude = rebuild_cons new_env nb_args relname @@ -1369,8 +1369,9 @@ let do_build_inductive *) let rel_arities = Array.mapi rel_arity funsargs in Util.Array.fold_left2 (fun env rel_name rel_ar -> - Environ.push_named (LocalAssum (rel_name, - fst (with_full_print (Constrintern.interp_constr env evd) rel_ar))) env) env relnames rel_arities + let rex = fst (with_full_print (Constrintern.interp_constr env evd) rel_ar) in + let rex = EConstr.Unsafe.to_constr rex in + Environ.push_named (LocalAssum (rel_name,rex)) env) env relnames rel_arities in (* and of the real constructors*) let constr i res = diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index e19fc9b62..13eda3952 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -141,8 +141,7 @@ let rec abstract_glob_constr c = function | Constrexpr.CLocalPattern _::bl -> assert false let interp_casted_constr_with_implicits env sigma impls c = - Constrintern.intern_gen Pretyping.WithoutTypeConstraint env ~impls - c + Constrintern.intern_gen Pretyping.WithoutTypeConstraint env sigma ~impls c (* Construct a fixpoint as a Glob_term @@ -160,9 +159,9 @@ let build_newrecursive let arity,ctx = Constrintern.interp_type env0 sigma arityc in let evd = Evd.from_env env0 in let evd, (_, (_, impls')) = Constrintern.interp_context_evars env evd bl in - let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity impls' in + let impl = Constrintern.compute_internalization_data env0 evd Constrintern.Recursive arity impls' in let open Context.Named.Declaration in - (Environ.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls)) + (EConstr.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls)) (env0,Constrintern.empty_internalization_env) lnameargsardef in let recdef = (* Declare local notations *) diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 8fe05b497..e2d7de6d5 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -210,6 +210,7 @@ let (value_f: Constr.t list -> global_reference -> Constr.t) = DAst.make @@ GVar v_id)]) in let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in + let body = EConstr.Unsafe.to_constr body in it_mkLambda_or_LetIn body context let (declare_f : Id.t -> logical_kind -> Constr.t list -> global_reference -> global_reference) = @@ -1400,7 +1401,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp (fun c -> Proofview.V82.of_tactic (Tacticals.New.tclTHENLIST [intros; - Simple.apply (EConstr.of_constr (fst (interp_constr (Global.env()) Evd.empty c))) (*FIXME*); + Simple.apply (fst (interp_constr (Global.env()) Evd.empty c)) (*FIXME*); Tacticals.New.tclCOMPLETE Auto.default_auto ]) ) @@ -1599,7 +1600,9 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num and functional_ref = destConst (constr_of_global functional_ref) and eq_ref = destConst (constr_of_global eq_ref) in generate_induction_principle f_ref tcc_lemma_constr - functional_ref eq_ref rec_arg_num (EConstr.of_constr rec_arg_type) (nb_prod evd (EConstr.of_constr res)) (EConstr.of_constr relation); + functional_ref eq_ref rec_arg_num + (EConstr.of_constr rec_arg_type) + (nb_prod evd (EConstr.of_constr res)) relation; Flags.if_verbose msgnl (h 1 (Ppconstr.pr_id function_name ++ spc () ++ str"is defined" )++ fnl () ++ @@ -1614,7 +1617,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num tcc_lemma_constr is_mes functional_ref (EConstr.of_constr rec_arg_type) - (EConstr.of_constr relation) rec_arg_num + relation rec_arg_num term_id using_lemmas (List.length res_vars) diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 10be8a842..6a758ed1f 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -253,6 +253,7 @@ let add_rewrite_hint ~poly bases ort t lcsr = let sigma = Evd.from_env env in let f ce = let c, ctx = Constrintern.interp_constr env sigma ce in + let c = EConstr.to_constr sigma c in let ctx = let ctx = UState.context_set ctx in if poly then ctx @@ -554,6 +555,7 @@ let add_transitivity_lemma left lem = let env = Global.env () in let sigma = Evd.from_env env in let lem',ctx (*FIXME*) = Constrintern.interp_constr env sigma lem in + let lem' = EConstr.to_constr sigma lem' in add_anonymous_leaf (inTransitivity (left,lem')) (* Vernacular syntax *) @@ -611,8 +613,10 @@ END VERNAC COMMAND EXTEND RetroknowledgeRegister CLASSIFIED AS SIDEFF | [ "Register" constr(c) "as" retroknowledge_field(f) "by" constr(b)] -> - [ let tc,ctx = Constrintern.interp_constr (Global.env ()) Evd.empty c in - let tb,ctx(*FIXME*) = Constrintern.interp_constr (Global.env ()) Evd.empty b in + [ let tc,_ctx = Constrintern.interp_constr (Global.env ()) Evd.empty c in + let tb,_ctx(*FIXME*) = Constrintern.interp_constr (Global.env ()) Evd.empty b in + let tc = EConstr.to_constr Evd.empty tc in + let tb = EConstr.to_constr Evd.empty tb in Global.register f tc tb ] END @@ -705,7 +709,6 @@ let hResolve id c occ t = resolve_hole (subst_hole_with_term loc_begin c_raw t_hole) in let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in - let t_constr = EConstr.of_constr t_constr in let sigma = Evd.merge_universe_context sigma ctx in let t_constr_type = Retyping.get_type_of env sigma t_constr in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index e73a18b79..3433987e3 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1896,7 +1896,6 @@ let declare_projection n instance_id r = let build_morphism_signature env sigma m = let m,ctx = Constrintern.interp_constr env sigma m in - let m = EConstr.of_constr m in let sigma = Evd.from_ctx ctx in let t = Typing.unsafe_type_of env sigma m in let cstrs = diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 22ec6c5b1..9a10fb805 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -199,7 +199,7 @@ let intern_constr_gen pattern_mode isarity {ltacvars=lfun; genv=env; extra} c = ltac_extra = extra; } in let c' = - warn (Constrintern.intern_gen scope ~pattern_mode ~ltacvars env) c + warn (Constrintern.intern_gen scope ~pattern_mode ~ltacvars env Evd.(from_env env)) c in (c',if !strict_check then None else Some c) @@ -316,7 +316,7 @@ let intern_constr_pattern ist ~as_type ~ltacvars pc = ltac_extra = ist.extra; } in let metas,pat = Constrintern.intern_constr_pattern - ist.genv ~as_type ~ltacvars pc + ist.genv Evd.(from_env ist.genv) ~as_type ~ltacvars pc in let (glob,_ as c) = intern_constr_gen true false ist pc in let bound_names = Glob_ops.bound_glob_vars glob in @@ -335,7 +335,7 @@ let intern_typed_pattern ist ~as_type ~ltacvars p = ltac_bound = Id.Set.empty; ltac_extra = ist.extra; } in - Constrintern.intern_constr_pattern ist.genv ~as_type ~ltacvars p + Constrintern.intern_constr_pattern ist.genv Evd.(from_env ist.genv) ~as_type ~ltacvars p else [], dummy_pat in let (glob,_ as c) = intern_constr_gen true false ist p in diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 79b5c1622..6a04badf3 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -584,7 +584,7 @@ let interp_glob_closure ist env sigma ?(kind=WithoutTypeConstraint) ?(pattern_mo ltac_bound = Id.Map.domain ist.lfun; ltac_extra = Genintern.Store.empty; } in - { closure ; term = intern_gen kind ~pattern_mode ~ltacvars env term_expr } + { closure ; term = intern_gen kind ~pattern_mode ~ltacvars env sigma term_expr } let interp_uconstr ist env sigma c = interp_glob_closure ist env sigma c diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 125afb1a0..0f0dbfff3 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -148,7 +148,7 @@ let ic c = let ic_unsafe c = (*FIXME remove *) let env = Global.env() in let sigma = Evd.from_env env in - EConstr.of_constr (fst (Constrintern.interp_constr env sigma c)) + fst (Constrintern.interp_constr env sigma c) let decl_constant na univs c = let open Constr in diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml index 4ae746288..1a02fb91c 100644 --- a/plugins/ssr/ssrcommon.ml +++ b/plugins/ssr/ssrcommon.ml @@ -203,7 +203,7 @@ let glob_constr ist genv = function let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.Tacinterp.lfun Id.Set.empty in let ltacvars = { Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in - Constrintern.intern_gen Pretyping.WithoutTypeConstraint ~ltacvars genv ce + Constrintern.intern_gen Pretyping.WithoutTypeConstraint ~ltacvars genv Evd.(from_env genv) ce | rc, None -> rc let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4 index 8e6e0347f..96ded5abf 100644 --- a/plugins/ssr/ssrvernac.ml4 +++ b/plugins/ssr/ssrvernac.ml4 @@ -335,7 +335,8 @@ let push_rels_assum l e = push_rels_assum l e let coerce_search_pattern_to_sort hpat = - let env = Global.env () and sigma = Evd.empty in + let env = Global.env () in + let sigma = Evd.(from_env env) in let mkPApp fp n_imps args = let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in Pattern.PApp (fp, args') in @@ -393,8 +394,9 @@ let interp_search_arg arg = interp_search_notation ~loc s key | RGlobSearchSubPattern p -> try - let intern = Constrintern.intern_constr_pattern in - Search.GlobSearchSubPattern (snd (intern (Global.env()) p)) + let env = Global.env () in + let _, p = Constrintern.intern_constr_pattern env (Evd.from_env env) p in + Search.GlobSearchSubPattern p with e -> let e = CErrors.push e in iraise (ExplainErr.process_vernac_interp_error e)) arg in let hpat, a1 = match arg with | (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a' diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml index 61b65e347..f2990070b 100644 --- a/plugins/ssr/ssrview.ml +++ b/plugins/ssr/ssrview.ml @@ -55,7 +55,8 @@ let in_viewhint = Libobject.classify_function = classify_viewhint } let glob_view_hints lvh = - List.map (Constrintern.intern_constr (Global.env ())) lvh + let env = Global.env () in + List.map (Constrintern.intern_constr env Evd.(from_env env)) lvh let add_view_hints lvh i = Lib.add_anonymous_leaf (in_viewhint (i, lvh)) diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 73e212365..d05f50e6d 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -73,11 +73,11 @@ let env_size env = List.length (Environ.named_context env) let safeDestApp c = match kind c with App (f, a) -> f, a | _ -> c, [| |] (* Toplevel constr must be globalized twice ! *) -let glob_constr ist genv = function +let glob_constr ist genv sigma = function | _, Some ce -> let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.lfun Id.Set.empty in let ltacvars = { Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in - Constrintern.intern_gen WithoutTypeConstraint ~ltacvars:ltacvars genv ce + Constrintern.intern_gen WithoutTypeConstraint ~ltacvars:ltacvars genv sigma ce | rc, None -> rc (* Term printing utilities functions for deciding bracketing. *) @@ -1009,7 +1009,7 @@ let interp_wit wit ist gl x = sigma, Value.cast (topwit wit) arg let interp_open_constr ist gl gc = interp_wit wit_open_constr ist gl gc -let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c +let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) gl.sigma c let interp_term ist gl (_, c) = on_snd EConstr.Unsafe.to_constr (interp_open_constr ist gl c) let pr_ssrterm _ _ _ = pr_term let input_ssrtermkind strm = match stream_nth 0 strm with diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 275a079d5..33d674ff5 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -643,8 +643,9 @@ let type_of_projection_knowing_arg env sigma p c ty = (* A function which checks that a term well typed verifies both syntactic conditions *) -let control_only_guard env c = - let check_fix_cofix e c = match kind c with +let control_only_guard env sigma c = + let check_fix_cofix e c = + match kind (EConstr.to_constr sigma c) with | CoFix (_,(_,_,_) as cofix) -> Inductive.check_cofix e cofix | Fix (_,(_,_,_) as fix) -> @@ -653,6 +654,6 @@ let control_only_guard env c = in let rec iter env c = check_fix_cofix env c; - iter_constr_with_full_binders push_rel iter env c + iter_constr_with_full_binders sigma EConstr.push_rel iter env c in iter env c diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index 55149552a..8efa087a7 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -198,4 +198,4 @@ val type_of_inductive_knowing_conclusion : env -> evar_map -> Inductive.mind_specif Univ.puniverses -> EConstr.types -> evar_map * EConstr.types (********************) -val control_only_guard : env -> types -> unit +val control_only_guard : env -> Evd.evar_map -> EConstr.types -> unit diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 6700748eb..79d255b0b 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -1205,8 +1205,7 @@ let all_no_fail_flags = default_inference_flags false let ise_pretype_gen_ctx flags env sigma lvar kind c = let evd, c, _ = ise_pretype_gen flags env sigma lvar kind c in - let evd, f = Evarutil.nf_evars_and_universes evd in - f (EConstr.Unsafe.to_constr c), Evd.evar_universe_context evd + c, Evd.evar_universe_context evd (** Entry points of the high-level type synthesis algorithm *) diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index 864768fe5..ee568fbc4 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -12,7 +12,6 @@ into elementary ones, insertion of coercions and resolution of implicit arguments. *) -open Constr open Environ open Evd open EConstr @@ -26,7 +25,7 @@ val interp_known_glob_level : ?loc:Loc.t -> Evd.evar_map -> (** An auxiliary function for searching for fixpoint guard indexes *) val search_guard : - ?loc:Loc.t -> env -> int list list -> rec_declaration -> int array + ?loc:Loc.t -> env -> int list list -> Constr.rec_declaration -> int array type typing_constraint = OfType of types | IsType | WithoutTypeConstraint @@ -85,9 +84,8 @@ val understand_ltac : inference_flags -> heuristics (but no external tactic solver hook), as well as to ensure that conversion problems are all solved and that no unresolved evar remains, expanding evars. *) - val understand : ?flags:inference_flags -> ?expected_type:typing_constraint -> - env -> evar_map -> glob_constr -> Constr.constr Evd.in_evar_universe_context + env -> evar_map -> glob_constr -> constr Evd.in_evar_universe_context (** Trying to solve remaining evars and remaining conversion problems possibly using type classes, heuristics, external tactic solver diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 3c7095505..dedad1990 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -701,13 +701,16 @@ let tag_var = tag Tag.variable | { CAst.v = CAppExpl ((None,f,us),[]) } -> str "@" ++ pr_cref f us | c -> pr prec c - let transf env c = + let transf env sigma c = if !Flags.beautify_file then - let r = Constrintern.for_grammar (Constrintern.intern_constr env) c in + let r = Constrintern.for_grammar (Constrintern.intern_constr env sigma) c in Constrextern.extern_glob_constr (Termops.vars_of_env env) r else c - let pr_expr prec c = pr prec (transf (Global.env()) c) + let pr_expr prec c = + let env = Global.env () in + let sigma = Evd.from_env env in + pr prec (transf env sigma c) let pr_simpleconstr = pr_expr lsimpleconstr diff --git a/proofs/proof.ml b/proofs/proof.ml index 04e707cd6..f5ff35eeb 100644 --- a/proofs/proof.ml +++ b/proofs/proof.ml @@ -432,7 +432,7 @@ module V82 = struct CList.nth evl (n-1) in let env = Evd.evar_filtered_env evi in - let rawc = Constrintern.intern_constr env com in + let rawc = Constrintern.intern_constr env sigma com in let ltac_vars = Glob_ops.empty_lvar in let sigma = Evar_refiner.w_refine (evk, evi) (ltac_vars, rawc) sigma in Proofview.Unsafe.tclEVARS sigma diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index a95e6b941..fe163cabc 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -581,7 +581,7 @@ let make_resolve_hyp env sigma st flags only_classes pri decl = (fun (path,info,c) -> let info = { info with Vernacexpr.hint_pattern = - Option.map (Constrintern.intern_constr_pattern env) + Option.map (Constrintern.intern_constr_pattern env sigma) info.Vernacexpr.hint_pattern } in make_resolves env sigma ~name:(PathHints path) diff --git a/tactics/hints.ml b/tactics/hints.ml index 7f9b5ef34..10cd7e1f6 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1270,7 +1270,7 @@ let prepare_hint check (poly,local) env init (sigma,c) = let interp_hints poly = fun h -> - let env = (Global.env()) in + let env = Global.env () in let sigma = Evd.from_env env in let f poly c = let evd,c = Constrintern.interp_open_constr env sigma c in @@ -1279,9 +1279,7 @@ let interp_hints poly = let gr = global_with_alias r in Dumpglob.add_glob ?loc:(loc_of_reference r) gr; gr in - let fr r = - evaluable_of_global_reference (Global.env()) (fref r) - in + let fr r = evaluable_of_global_reference env (fref r) in let fi c = match c with | HintsReference c -> @@ -1289,7 +1287,7 @@ let interp_hints poly = (PathHints [gr], poly, IsGlobRef gr) | HintsConstr c -> (PathAny, poly, f poly c) in - let fp = Constrintern.intern_constr_pattern (Global.env()) in + let fp = Constrintern.intern_constr_pattern env sigma in let fres (info, b, r) = let path, poly, gr = fi r in let info = { info with hint_pattern = Option.map fp info.hint_pattern } in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 7e281e2fe..945b178a1 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1991,7 +1991,6 @@ let exact_proof c = Proofview.Goal.enter begin fun gl -> Refine.refine ~typecheck:false begin fun sigma -> let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in - let c = EConstr.of_constr c in let sigma = Evd.merge_universe_context sigma ctx in (sigma, c) end diff --git a/vernac/classes.ml b/vernac/classes.ml index 695be74bb..93af94ad6 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -52,7 +52,7 @@ let _ = let open Vernacexpr in { info with hint_pattern = Option.map - (Constrintern.intern_constr_pattern (Global.env())) + (Constrintern.intern_constr_pattern (Global.env()) Evd.(from_env Global.(env()))) info.hint_pattern } in Flags.silently (fun () -> Hints.add_hints local [typeclasses_db] diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml index 7e5b941ad..9fd2153cb 100644 --- a/vernac/comAssumption.ml +++ b/vernac/comAssumption.ml @@ -10,7 +10,6 @@ open Pp open CErrors open Util open Vars -open Environ open Declare open Names open Globnames @@ -87,7 +86,6 @@ match local with let interp_assumption sigma env impls bl c = let c = mkCProdN ?loc:(local_binders_loc bl) bl c in let sigma, (ty, impls) = interp_type_evars_impls env sigma ~impls c in - let ty = EConstr.Unsafe.to_constr ty in sigma, (ty, impls) (* When monomorphic the universe constraints are declared with the first declaration only. *) @@ -151,9 +149,9 @@ let do_assumptions kind nl l = let (sigma,_,_),l = List.fold_left_map (fun (sigma,env,ienv) (is_coe,(idl,c)) -> let sigma,(t,imps) = interp_assumption sigma env ienv [] c in let env = - push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in + EConstr.push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in let ienv = List.fold_right (fun {CAst.v=id} ienv -> - let impls = compute_internalization_data env Variable t imps in + let impls = compute_internalization_data env sigma Variable t imps in Id.Map.add id impls ienv) idl ienv in ((sigma,env,ienv),((is_coe,idl),t,imps))) (sigma,env,empty_internalization_env) l @@ -161,7 +159,7 @@ let do_assumptions kind nl l = let sigma = solve_remaining_evars all_and_fail_flags env sigma Evd.empty in (* The universe constraints come from the whole telescope. *) let sigma = Evd.nf_constraints sigma in - let nf_evar c = EConstr.to_constr sigma (EConstr.of_constr c) in + let nf_evar c = EConstr.to_constr sigma c in let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) -> let t = nf_evar t in let uvars = Univ.LSet.union uvars (Univops.universes_of_constr env t) in diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml index 489f299a2..edfe7aa81 100644 --- a/vernac/comFixpoint.ml +++ b/vernac/comFixpoint.ml @@ -212,8 +212,7 @@ let interp_recursive ~program_mode ~cofix fixl notations = let env_rec = push_named_context rec_sign env in (* Get interpretation metadatas *) - let fixtypes = List.map EConstr.Unsafe.to_constr fixtypes in - let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in + let impls = compute_internalization_env env sigma Recursive fixnames fixtypes fiximps in (* Interp bodies with rollback because temp use of notations/implicit *) let sigma, fixdefs = @@ -226,10 +225,9 @@ let interp_recursive ~program_mode ~cofix fixl notations = (* Instantiate evars and check all are resolved *) let sigma = solve_unif_constraints_with_heuristics env_rec sigma in - let sigma, nf = nf_evars_and_universes sigma in - let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in - let fixdefs = List.map (Option.map nf) fixdefs in - let fixtypes = List.map nf fixtypes in + let sigma, _ = nf_evars_and_universes sigma in + let fixdefs = List.map (fun c -> Option.map EConstr.(to_constr sigma) c) fixdefs in + let fixtypes = List.map EConstr.(to_constr sigma) fixtypes in let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in (* Build the fix declaration block *) diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml index c650e9e40..cef5546c6 100644 --- a/vernac/comInductive.ml +++ b/vernac/comInductive.ml @@ -11,7 +11,6 @@ open CErrors open Sorts open Util open Constr -open Termops open Environ open Declare open Names @@ -51,7 +50,7 @@ let rec complete_conclusion a cs = CAst.map_with_loc (fun ?loc -> function ) let push_types env idl tl = - List.fold_left2 (fun env id t -> Environ.push_rel (LocalAssum (Name id,t)) env) + List.fold_left2 (fun env id t -> EConstr.push_rel (LocalAssum (Name id,t)) env) env idl tl type structured_one_inductive_expr = { @@ -90,7 +89,7 @@ let check_all_names_different indl = | _ -> raise (InductiveError (SameNamesOverlap l)) let mk_mltype_data sigma env assums arity indname = - let is_ml_type = is_sort env sigma (EConstr.of_constr arity) in + let is_ml_type = is_sort env sigma arity in (is_ml_type,indname,assums) let prepare_param = function @@ -130,14 +129,13 @@ let is_impredicative env u = u = Prop Null || (is_impredicative_set env && u = Prop Pos) let interp_ind_arity env sigma ind = - let c = intern_gen IsType env ind.ind_arity in + let c = intern_gen IsType env sigma ind.ind_arity in let impls = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in let sigma,t = understand_tcc env sigma ~expected_type:IsType c in let pseudo_poly = check_anonymous_type c in let () = if not (Reductionops.is_arity env sigma t) then user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity") in - let t = EConstr.Unsafe.to_constr t in sigma, (t, pseudo_poly, impls) let interp_cstrs env sigma impls mldata arity ind = @@ -272,7 +270,6 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = let sigma, (impls, ((env_params, ctx_params), userimpls)) = interp_context_evars env0 sigma paramsl in - let ctx_params = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx_params in let indnames = List.map (fun ind -> ind.ind_name) indl in (* Names of parameters as arguments of the inductive type (defs removed) *) @@ -282,16 +279,16 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = (* Interpret the arities *) let sigma, arities = List.fold_left_map (fun sigma -> interp_ind_arity env_params sigma) sigma indl in - let fullarities = List.map (fun (c, _, _) -> Term.it_mkProd_or_LetIn c ctx_params) arities in + let fullarities = List.map (fun (c, _, _) -> EConstr.it_mkProd_or_LetIn c ctx_params) arities in let env_ar = push_types env0 indnames fullarities in - let env_ar_params = push_rel_context ctx_params env_ar in + let env_ar_params = EConstr.push_rel_context ctx_params env_ar in (* Compute interpretation metadatas *) let indimpls = List.map (fun (_, _, impls) -> userimpls @ lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in let arities = List.map pi1 arities and aritypoly = List.map pi2 arities in - let impls = compute_internalization_env env0 ~impls (Inductive (params,true)) indnames fullarities indimpls in - let ntn_impls = compute_internalization_env env0 (Inductive (params,true)) indnames fullarities indimpls in + let impls = compute_internalization_env env0 sigma ~impls (Inductive (params,true)) indnames fullarities indimpls in + let ntn_impls = compute_internalization_env env0 sigma (Inductive (params,true)) indnames fullarities indimpls in let mldatas = List.map2 (mk_mltype_data sigma env_params params) arities indnames in let sigma, constructors = @@ -306,15 +303,14 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = let sigma = solve_remaining_evars all_and_fail_flags env_params sigma Evd.empty in (* Compute renewed arities *) let sigma, nf = nf_evars_and_universes sigma in - let arities = List.map nf arities in let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf cl,impsl)) constructors in + let arities = List.map EConstr.(to_constr sigma) arities in let sigma = List.fold_left2 (fun sigma ty poly -> make_conclusion_flexible sigma ty poly) sigma arities aritypoly in let sigma, arities = inductive_levels env_ar_params sigma poly arities constructors in let sigma, nf' = nf_evars_and_universes sigma in - let nf x = nf' (nf x) in let arities = List.map nf' arities in let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in - let ctx_params = Context.Rel.map nf ctx_params in + let ctx_params = List.map Termops.(map_rel_decl (EConstr.to_constr sigma)) ctx_params in let uctx = Evd.check_univ_decl ~poly sigma decl in List.iter (fun c -> check_evars env_params Evd.empty sigma (EConstr.of_constr c)) arities; Context.Rel.iter (fun c -> check_evars env0 Evd.empty sigma (EConstr.of_constr c)) ctx_params; diff --git a/vernac/comProgramFixpoint.ml b/vernac/comProgramFixpoint.ml index af34f8b29..bd7ee0978 100644 --- a/vernac/comProgramFixpoint.ml +++ b/vernac/comProgramFixpoint.ml @@ -171,8 +171,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let sigma, intern_body = let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in let (r, l, impls, scopes) = - Constrintern.compute_internalization_data env - Constrintern.Recursive (EConstr.Unsafe.to_constr full_arity) impls + Constrintern.compute_internalization_data env sigma + Constrintern.Recursive full_arity impls in let newimpls = Id.Map.singleton recname (r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))], diff --git a/vernac/obligations.ml b/vernac/obligations.ml index e4bcbc4bb..6447fc350 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -816,13 +816,13 @@ let solve_by_tac name evi t poly ctx = let id = name in let concl = EConstr.of_constr evi.evar_concl in (* spiwack: the status is dropped. *) - let (entry,_,ctx') = Pfedit.build_constant_by_tactic + let (entry,_,ctx') = Pfedit.build_constant_by_tactic id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in let env = Global.env () in let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in let body, () = Future.force entry.const_entry_body in let ctx' = Evd.merge_context_set ~sideff:true Evd.univ_rigid (Evd.from_ctx ctx') (snd body) in - Inductiveops.control_only_guard (Global.env ()) (fst body); + Inductiveops.control_only_guard env ctx' (EConstr.of_constr (fst body)); (fst body), entry.const_entry_type, Evd.evar_universe_context ctx' let obligation_terminator name num guard hook auto pf = @@ -838,7 +838,7 @@ let obligation_terminator name num guard hook auto pf = let (body, cstr), () = Future.force entry.Entries.const_entry_body in let sigma = Evd.from_ctx uctx in let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in - Inductiveops.control_only_guard (Global.env ()) body; + Inductiveops.control_only_guard (Global.env ()) sigma (EConstr.of_constr body); (** Declare the obligation ourselves and drop the hook *) let prg = get_info (ProgMap.find name !from_prg) in (** Ensure universes are substituted properly in body and type *) diff --git a/vernac/record.ml b/vernac/record.ml index 1140e3d37..44113bfad 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -70,7 +70,7 @@ let interp_fields_evars env sigma impls_env nots l = let impls = match i with | Anonymous -> impls - | Name id -> Id.Map.add id (compute_internalization_data env Constrintern.Method (EConstr.to_constr sigma t') impl) impls + | Name id -> Id.Map.add id (compute_internalization_data env sigma Constrintern.Method t' impl) impls in let d = match b' with | None -> LocalAssum (i,t') @@ -145,7 +145,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs = let assums = List.filter is_local_assum newps in let params = List.map (RelDecl.get_name %> Name.get_id) assums in let ty = Inductive (params,(finite != Declarations.BiFinite)) in - let impls_env = compute_internalization_env env0 ~impls:impls_env ty [id] [EConstr.to_constr sigma arity] [imps] in + let impls_env = compute_internalization_env env0 sigma ~impls:impls_env ty [id] [arity] [imps] in let env2,sigma,impls,newfs,data = interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs) in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 4613100fc..2aee2065c 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1025,7 +1025,9 @@ let vernac_arguments ~atts reference args more_implicits nargs_for_red flags = let sr = smart_global reference in let inf_names = let ty, _ = Global.type_of_global_in_context (Global.env ()) sr in - Impargs.compute_implicits_names (Global.env ()) ty + let env = Global.env () in + let sigma = Evd.from_env env in + Impargs.compute_implicits_names env sigma (EConstr.of_constr ty) in let prev_names = try Arguments_renaming.arguments_names sr with Not_found -> inf_names @@ -1253,7 +1255,7 @@ 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) (EConstr.of_constr t) in + let t = 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 @@ -1627,6 +1629,7 @@ let vernac_global_check c = let senv = Global.safe_env() in let uctx = UState.context_set uctx in let senv = Safe_typing.push_context_set false uctx senv in + let c = EConstr.to_constr sigma c in let j = Safe_typing.typing senv c in let env = Safe_typing.env_of_safe_env senv in print_safe_judgment env sigma j ++ @@ -1747,10 +1750,10 @@ let interp_search_restriction = function open Search -let interp_search_about_item env = +let interp_search_about_item env sigma = function | SearchSubPattern pat -> - let _,pat = intern_constr_pattern env pat in + let _,pat = intern_constr_pattern env sigma pat in GlobSearchSubPattern pat | SearchString (s,None) when Id.is_valid s -> GlobSearchString s @@ -1796,13 +1799,13 @@ let vernac_search ~atts s gopt r = (* if goal selector is given and wrong, then let exceptions be raised. *) | Some g -> snd (Pfedit.get_goal_context g) , Some g in - let get_pattern c = snd (intern_constr_pattern env c) in + let get_pattern c = snd (intern_constr_pattern env Evd.(from_env env) c) in let pr_search ref env c = let pr = pr_global ref in let pp = if !search_output_name_only then pr else begin - let pc = pr_lconstr_env env Evd.empty c in + let pc = pr_lconstr_env env Evd.(from_env env) c in hov 2 (pr ++ str":" ++ spc () ++ pc) end in Feedback.msg_notice pp @@ -1815,7 +1818,8 @@ let vernac_search ~atts s gopt r = | SearchHead c -> (Search.search_by_head gopt (get_pattern c) r |> Search.prioritize_search) pr_search | SearchAbout sl -> - (Search.search_about gopt (List.map (on_snd (interp_search_about_item env)) sl) r |> Search.prioritize_search) pr_search + (Search.search_about gopt (List.map (on_snd (interp_search_about_item env Evd.(from_env env))) sl) r |> + Search.prioritize_search) pr_search let vernac_locate = function | LocateAny (AN qid) -> print_located_qualid qid @@ -1910,8 +1914,7 @@ let vernac_check_guard () = let message = try let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in - Inductiveops.control_only_guard (Goal.V82.env sigma gl) - (EConstr.Unsafe.to_constr pfterm); + Inductiveops.control_only_guard (Goal.V82.env sigma gl) sigma pfterm; (str "The condition holds up to here") with UserError(_,s) -> (str ("Condition violated: ") ++s) -- cgit v1.2.3