From 3a0b543af4ac99b29efdebe27b1d204d044a7bf0 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Sat, 31 Mar 2018 17:43:18 +0200 Subject: Evar maps contain econstrs. We bootstrap the circular evar_map <-> econstr dependency by moving the internal EConstr.API module to Evd.MiniEConstr. Then we make the Evd functions use econstr. --- dev/doc/changes.md | 9 ++ dev/top_printers.ml | 2 +- engine/eConstr.ml | 148 +++---------------------- engine/eConstr.mli | 5 +- engine/evarutil.ml | 51 +++++---- engine/evarutil.mli | 2 +- engine/evd.ml | 105 ++++++++++++++++++ engine/evd.mli | 133 +++++++++++++++------- engine/proofview.ml | 9 +- engine/termops.ml | 17 +-- plugins/firstorder/sequent.ml | 2 +- plugins/funind/functional_principles_proofs.ml | 3 +- plugins/funind/functional_principles_types.ml | 15 ++- plugins/funind/glob_termops.ml | 7 +- plugins/funind/indfun.ml | 11 +- plugins/funind/invfun.ml | 8 +- plugins/ltac/extratactics.ml4 | 2 - plugins/ltac/rewrite.ml | 8 +- plugins/ssr/ssrcommon.ml | 16 +-- plugins/ssr/ssrelim.ml | 2 +- plugins/ssr/ssrequality.ml | 6 +- plugins/ssr/ssripats.ml | 2 +- plugins/ssr/ssrview.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 24 ++-- pretyping/cases.ml | 2 +- pretyping/coercion.ml | 2 +- pretyping/evarconv.ml | 36 +++--- pretyping/evardefine.ml | 12 +- pretyping/evarsolve.ml | 55 ++++----- pretyping/nativenorm.ml | 6 +- pretyping/patternops.ml | 2 +- pretyping/pretyping.ml | 8 +- pretyping/reductionops.ml | 15 +-- pretyping/retyping.ml | 8 +- pretyping/tacred.ml | 4 +- pretyping/typeclasses.ml | 2 +- pretyping/typing.ml | 3 +- pretyping/unification.ml | 30 ++--- pretyping/vnorm.ml | 2 +- printing/printer.ml | 20 ++-- proofs/clenv.ml | 23 ++-- proofs/evar_refiner.ml | 8 +- proofs/goal.ml | 9 +- proofs/pfedit.ml | 2 +- proofs/proof_global.ml | 2 +- proofs/refine.ml | 16 +-- stm/proofBlockDelimiter.ml | 2 +- stm/stm.ml | 3 +- tactics/auto.ml | 5 +- tactics/class_tactics.ml | 6 +- tactics/equality.ml | 4 - tactics/inv.ml | 2 - tactics/tacticals.ml | 7 +- tactics/tactics.ml | 11 +- vernac/class.ml | 1 + vernac/himsg.ml | 20 ++-- vernac/obligations.ml | 11 +- vernac/vernacentries.ml | 2 +- 58 files changed, 462 insertions(+), 468 deletions(-) diff --git a/dev/doc/changes.md b/dev/doc/changes.md index ab78b0956..1a24f23e5 100644 --- a/dev/doc/changes.md +++ b/dev/doc/changes.md @@ -1,3 +1,12 @@ +## Changes between Coq 8.8 and Coq 8.9 + +### ML API + +Proof engine + + More functions have been changed to use `EConstr`, notably the + functions in `Evd`. + ## Changes between Coq 8.7 and Coq 8.8 ### Bug tracker diff --git a/dev/top_printers.ml b/dev/top_printers.ml index ba0c54407..f9b402586 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -181,7 +181,7 @@ let ppproofview p = pp(pr_enum Goal.pr_goal gls ++ fnl () ++ Termops.pr_evar_map (Some 1) sigma) let ppopenconstr (x : Evd.open_constr) = - let (evd,c) = x in pp (Termops.pr_evar_map (Some 2) evd ++ envpp pr_constr_env c) + let (evd,c) = x in pp (Termops.pr_evar_map (Some 2) evd ++ envpp pr_econstr_env c) (* spiwack: deactivated until a replacement is found let pppftreestate p = pp(print_pftreestate p) *) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 01b4fe851..a72bdee12 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -13,138 +13,8 @@ open Util open Names open Constr open Context -open Evd - -module API : -sig -module ESorts : -sig -type t -val make : Sorts.t -> t -val kind : Evd.evar_map -> t -> Sorts.t -val unsafe_to_sorts : t -> Sorts.t -end -module EInstance : -sig -type t -val make : Univ.Instance.t -> t -val kind : Evd.evar_map -> t -> Univ.Instance.t -val empty : t -val is_empty : t -> bool -val unsafe_to_instance : t -> Univ.Instance.t -end -type t -val kind : Evd.evar_map -> t -> (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -val kind_upto : Evd.evar_map -> constr -> (constr, types, Sorts.t, Univ.Instance.t) Constr.kind_of_term -val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type -val whd_evar : Evd.evar_map -> t -> t -val of_kind : (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -> t -val of_constr : Constr.t -> t -val to_constr : ?abort_on_undefined_evars:bool -> evar_map -> t -> Constr.t -val unsafe_to_constr : t -> Constr.t -val unsafe_eq : (t, Constr.t) eq -val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (t, t) Context.Named.Declaration.pt -val unsafe_to_named_decl : (t, t) Context.Named.Declaration.pt -> (Constr.t, Constr.types) Context.Named.Declaration.pt -val unsafe_to_rel_decl : (t, t) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt -val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (t, t) Context.Rel.Declaration.pt -val to_rel_decl : Evd.evar_map -> (t, t) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt -end = -struct - -module ESorts = -struct - type t = Sorts.t - let make s = s - let kind sigma = function - | Sorts.Type u -> Sorts.sort_of_univ (Evd.normalize_universe sigma u) - | s -> s - let unsafe_to_sorts s = s -end -module EInstance = -struct - type t = Univ.Instance.t - let make i = i - let kind sigma i = - if Univ.Instance.is_empty i then i - else Evd.normalize_universe_instance sigma i - let empty = Univ.Instance.empty - let is_empty = Univ.Instance.is_empty - let unsafe_to_instance t = t -end - -type t = Constr.t - -let safe_evar_value sigma ev = - try Some (Evd.existential_value sigma ev) - with NotInstantiatedEvar | Not_found -> None - -let rec whd_evar sigma c = - match Constr.kind c with - | Evar ev -> - begin match safe_evar_value sigma ev with - | Some c -> whd_evar sigma c - | None -> c - end - | App (f, args) when isEvar f -> - (** Enforce smart constructor invariant on applications *) - let ev = destEvar f in - begin match safe_evar_value sigma ev with - | None -> c - | Some f -> whd_evar sigma (mkApp (f, args)) - end - | Cast (c0, k, t) when isEvar c0 -> - (** Enforce smart constructor invariant on casts. *) - let ev = destEvar c0 in - begin match safe_evar_value sigma ev with - | None -> c - | Some c -> whd_evar sigma (mkCast (c, k, t)) - end - | _ -> c - -let kind sigma c = Constr.kind (whd_evar sigma c) -let kind_upto = kind -let kind_of_type sigma c = Term.kind_of_type (whd_evar sigma c) -let of_kind = Constr.of_kind -let of_constr c = c -let unsafe_to_constr c = c -let unsafe_eq = Refl - -let to_constr ?(abort_on_undefined_evars=true) sigma c = -let rec to_constr c = match Constr.kind c with -| Evar ev -> - begin match safe_evar_value sigma ev with - | Some c -> to_constr c - | None -> - if abort_on_undefined_evars then - anomaly ~label:"econstr" Pp.(str "grounding a non evar-free term") - else - Constr.map (fun c -> to_constr c) c - end -| Sort (Sorts.Type u) -> - let u' = Evd.normalize_universe sigma u in - if u' == u then c else mkSort (Sorts.sort_of_univ u') -| Const (c', u) when not (Univ.Instance.is_empty u) -> - let u' = Evd.normalize_universe_instance sigma u in - if u' == u then c else mkConstU (c', u') -| Ind (i, u) when not (Univ.Instance.is_empty u) -> - let u' = Evd.normalize_universe_instance sigma u in - if u' == u then c else mkIndU (i, u') -| Construct (co, u) when not (Univ.Instance.is_empty u) -> - let u' = Evd.normalize_universe_instance sigma u in - if u' == u then c else mkConstructU (co, u') -| _ -> Constr.map (fun c -> to_constr c) c -in to_constr c - -let of_named_decl d = d -let unsafe_to_named_decl d = d -let of_rel_decl d = d -let unsafe_to_rel_decl d = d -let to_rel_decl sigma d = Context.Rel.Declaration.map_constr (to_constr sigma) d - -end - -include API +include Evd.MiniEConstr type types = t type constr = t @@ -387,8 +257,7 @@ let decompose_prod_n_assum sigma n c = in prodec_rec Context.Rel.empty n c -let existential_type sigma (evk, args) = - of_constr (existential_type sigma (evk, Array.map unsafe_to_constr args)) +let existential_type = Evd.existential_type let map sigma f c = match kind sigma c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ @@ -749,7 +618,7 @@ let universes_of_constr env sigma c = LSet.fold LSet.add (Universe.levels u) s | Evar (k, args) -> let concl = Evd.evar_concl (Evd.find sigma k) in - fold sigma aux (aux s (of_constr concl)) c + fold sigma aux (aux s concl) c | _ -> fold sigma aux s c in aux LSet.empty c @@ -907,6 +776,10 @@ let named_context e = cast_named_context (sym unsafe_eq) (named_context e) let val_of_named_context e = val_of_named_context (cast_named_context unsafe_eq e) let named_context_of_val e = cast_named_context (sym unsafe_eq) (named_context_of_val e) +let of_existential : Constr.existential -> existential = + let gen : type a b. (a,b) eq -> 'c * b array -> 'c * a array = fun Refl x -> x in + gen unsafe_eq + let lookup_rel i e = cast_rel_decl (sym unsafe_eq) (lookup_rel i e) let lookup_named n e = cast_named_decl (sym unsafe_eq) (lookup_named n e) let lookup_named_val n e = cast_named_decl (sym unsafe_eq) (lookup_named_val n e) @@ -922,7 +795,7 @@ let map_rel_context_in_env f env sign = let fresh_global ?loc ?rigid ?names env sigma reference = let (evd,t) = Evd.fresh_global ?loc ?rigid ?names env sigma reference in - evd, of_constr t + evd, t let is_global sigma gr c = Globnames.is_global gr (to_constr sigma c) @@ -934,5 +807,10 @@ let to_instance = EInstance.unsafe_to_instance let to_constr = unsafe_to_constr let to_rel_decl = unsafe_to_rel_decl let to_named_decl = unsafe_to_named_decl +let to_named_context = + let gen : type a b. (a, b) eq -> (a,a) Context.Named.pt -> (b,b) Context.Named.pt + = fun Refl x -> x + in + gen unsafe_eq let eq = unsafe_eq end diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 9cc9bf680..9a5b5ec3a 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -13,7 +13,7 @@ open Names open Constr open Environ -type t +type t = Evd.econstr (** Type of incomplete terms. Essentially a wrapper around {!Constr.t} ensuring that {!Constr.kind} does not observe defined evars. *) @@ -290,6 +290,7 @@ val is_global : Evd.evar_map -> Globnames.global_reference -> t -> bool (** {5 Extra} *) +val of_existential : Constr.existential -> existential val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (t, types) Context.Named.Declaration.pt val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (t, types) Context.Rel.Declaration.pt @@ -308,6 +309,8 @@ sig val to_named_decl : (t, types) Context.Named.Declaration.pt -> (Constr.t, Constr.types) Context.Named.Declaration.pt (** Physical identity. Does not care for defined evars. *) + val to_named_context : (t, types) Context.Named.pt -> Context.Named.t + val to_sorts : ESorts.t -> Sorts.t (** Physical identity. Does not care for normalization. *) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index c707b37b6..065b42bf6 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -23,7 +23,8 @@ module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration let safe_evar_value sigma ev = - try Some (Evd.existential_value sigma ev) + let ev = EConstr.of_existential ev in + try Some (EConstr.Unsafe.to_constr @@ Evd.existential_value sigma ev) with NotInstantiatedEvar | Not_found -> None (** Combinators *) @@ -44,11 +45,11 @@ let evd_comb2 f evdref x y = z let e_new_global evdref x = - EConstr.of_constr (evd_comb1 (Evd.fresh_global (Global.env())) evdref x) + evd_comb1 (Evd.fresh_global (Global.env())) evdref x let new_global evd x = let (evd, c) = Evd.fresh_global (Global.env()) evd x in - (evd, EConstr.of_constr c) + (evd, c) (****************************************************) (* Expanding/testing/exposing existential variables *) @@ -61,7 +62,7 @@ exception Uninstantiated_evar of Evar.t let rec flush_and_check_evars sigma c = match kind c with | Evar (evk,_ as ev) -> - (match existential_opt_value sigma ev with + (match existential_opt_value0 sigma ev with | None -> raise (Uninstantiated_evar evk) | Some c -> flush_and_check_evars sigma c) | _ -> Constr.map (flush_and_check_evars sigma) c @@ -102,7 +103,8 @@ let nf_evar_map_universes evm = if Univ.LMap.is_empty subst then evm, nf_evar0 evm else let f = nf_evars_universes evm in - Evd.raw_map (fun _ -> map_evar_info f) evm, f + let f' c = EConstr.of_constr (f (EConstr.Unsafe.to_constr c)) in + Evd.raw_map (fun _ -> map_evar_info f') evm, f let nf_named_context_evar sigma ctx = Context.Named.map (nf_evar0 sigma) ctx @@ -115,7 +117,7 @@ let nf_env_evar sigma env = let rel' = nf_rel_context_evar sigma (EConstr.rel_context env) in EConstr.push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env) -let nf_evar_info evc info = map_evar_info (nf_evar0 evc) info +let nf_evar_info evc info = map_evar_info (nf_evar evc) info let nf_evar_map evm = Evd.raw_map (fun _ evi -> nf_evar_info evm evi) evm @@ -414,7 +416,6 @@ let push_rel_context_to_named_context env sigma typ = let default_source = Loc.tag @@ Evar_kinds.InternalHole let restrict_evar evd evk filter ?src candidates = - let candidates = Option.map (fun l -> List.map EConstr.Unsafe.to_constr l) candidates in let evd, evk' = Evd.restrict evk filter ?candidates ?src evd in Evd.declare_future_goal evk' evd, evk' @@ -424,8 +425,6 @@ let new_pure_evar_full evd evi = (evd, evk) let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?candidates ?(store = Store.empty) ?naming ?(principal=false) typ = - let typ = EConstr.Unsafe.to_constr typ in - let candidates = Option.map (fun l -> List.map EConstr.Unsafe.to_constr l) candidates in let default_naming = Misctypes.IntroAnonymous in let naming = Option.default default_naming naming in let name = match naming with @@ -513,7 +512,7 @@ let generalize_evar_over_rels sigma (ev,args) = List.fold_left2 (fun (c,inst as x) a d -> if isRel sigma a then (mkNamedProd_or_LetIn d c,a::inst) else x) - (EConstr.of_constr evi.evar_concl,[]) (Array.to_list args) sign + (evi.evar_concl,[]) (Array.to_list args) sign (************************************) (* Removing a dependency in an evar *) @@ -549,7 +548,8 @@ let rec check_and_clear_in_constr env evdref err ids global c = | Evar (evk,l as ev) -> if Evd.is_defined !evdref evk then (* If evk is already defined we replace it by its definition *) - let nc = Evd.existential_value !evdref ev in + let nc = Evd.existential_value !evdref (EConstr.of_existential ev) in + let nc = EConstr.Unsafe.to_constr nc in (check_and_clear_in_constr env evdref err ids global nc) else (* We check for dependencies to elements of ids in the @@ -559,8 +559,7 @@ let rec check_and_clear_in_constr env evdref err ids global c = removed *) let evi = Evd.find_undefined !evdref evk in let ctxt = Evd.evar_filtered_context evi in - let ctxt = List.map (fun d -> map_named_decl EConstr.of_constr d) ctxt in - let (rids,filter) = + let (rids,filter) = List.fold_right2 (fun h a (ri,filter) -> try @@ -586,7 +585,8 @@ let rec check_and_clear_in_constr env evdref err ids global c = try let nids = Id.Map.domain rids in let global = Id.Set.exists is_section_variable nids in - check_and_clear_in_constr env evdref (EvarTypingBreak ev) nids global (evar_concl evi) + let concl = EConstr.Unsafe.to_constr (evar_concl evi) in + check_and_clear_in_constr env evdref (EvarTypingBreak ev) nids global concl with ClearDependencyError (rid,err) -> raise (ClearDependencyError (Id.Map.find rid rids,err)) in @@ -597,7 +597,7 @@ let rec check_and_clear_in_constr env evdref err ids global c = let evd = !evdref in let (evd,_) = restrict_evar evd evk filter None in evdref := evd; - Evd.existential_value !evdref ev + Evd.existential_value0 !evdref ev | _ -> Constr.map (check_and_clear_in_constr env evdref err ids global) c @@ -643,7 +643,7 @@ let clear_hyps2_in_evi env evdref hyps t concl ids = let queue_set q is_dependent set = Evar.Set.iter (fun a -> Queue.push (is_dependent,a) q) set let queue_term q is_dependent c = - queue_set q is_dependent (evars_of_term c) + queue_set q is_dependent (evars_of_term (EConstr.Unsafe.to_constr c)) let process_dependent_evar q acc evm is_dependent e = let evi = Evd.find evm e in @@ -656,12 +656,12 @@ let process_dependent_evar q acc evm is_dependent e = match decl with | LocalAssum _ -> () | LocalDef (_,b,_) -> queue_term q true b - end (Environ.named_context_of_val evi.evar_hyps); + end (EConstr.named_context_of_val evi.evar_hyps); match evi.evar_body with | Evar_empty -> if is_dependent then Evar.Map.add e None acc else acc | Evar_defined b -> - let subevars = evars_of_term b in + let subevars = evars_of_term (EConstr.Unsafe.to_constr b) in (* evars appearing in the definition of an evar [e] are marked as dependent when [e] is dependent itself: if [e] is a non-dependent goal, then, unless they are reach from another @@ -729,11 +729,11 @@ let undefined_evars_of_named_context evd nc = ~init:Evar.Set.empty let undefined_evars_of_evar_info evd evi = - Evar.Set.union (undefined_evars_of_term evd (EConstr.of_constr evi.evar_concl)) + Evar.Set.union (undefined_evars_of_term evd evi.evar_concl) (Evar.Set.union (match evi.evar_body with | Evar_empty -> Evar.Set.empty - | Evar_defined b -> undefined_evars_of_term evd (EConstr.of_constr b)) + | Evar_defined b -> undefined_evars_of_term evd b) (undefined_evars_of_named_context evd (named_context_of_val evi.evar_hyps))) @@ -781,10 +781,11 @@ let filtered_undefined_evars_of_evar_info ?cache sigma evi = in let accu = match evi.evar_body with | Evar_empty -> Evar.Set.empty - | Evar_defined b -> evars_of_term b + | Evar_defined b -> evars_of_term (EConstr.Unsafe.to_constr b) in - let accu = Evar.Set.union (undefined_evars_of_term sigma (EConstr.of_constr evi.evar_concl)) accu in - evars_of_named_context cache accu (evar_filtered_context evi) + let accu = Evar.Set.union (undefined_evars_of_term sigma evi.evar_concl) accu in + let ctxt = EConstr.Unsafe.to_named_context (evar_filtered_context evi) in + evars_of_named_context cache accu ctxt (* spiwack: this is a more complete version of {!Termops.occur_evar}. The latter does not look recursively into an @@ -794,7 +795,7 @@ let occur_evar_upto sigma n c = let c = EConstr.Unsafe.to_constr c in let rec occur_rec c = match kind c with | Evar (sp,_) when Evar.equal sp n -> raise Occur - | Evar e -> Option.iter occur_rec (existential_opt_value sigma e) + | Evar e -> Option.iter occur_rec (existential_opt_value0 sigma e) | _ -> Constr.iter occur_rec c in try occur_rec c; false with Occur -> true @@ -849,6 +850,8 @@ let compare_constructor_instances evd u u' = let eq_constr_univs_test sigma1 sigma2 t u = (* spiwack: mild code duplication with {!Evd.eq_constr_univs}. *) let open Evd in + let t = EConstr.Unsafe.to_constr t + and u = EConstr.Unsafe.to_constr u in let fold cstr sigma = try Some (add_universe_constraints sigma cstr) with Univ.UniverseInconsistency _ | UniversesDiffer -> None diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 972b0b9e1..40c1ee082 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -201,7 +201,7 @@ val kind_of_term_upto : evar_map -> Constr.constr -> universes. The term [t] is interpreted in [sigma1] while [u] is interpreted in [sigma2]. The universe constraints in [sigma2] are assumed to be an extention of those in [sigma1]. *) -val eq_constr_univs_test : evar_map -> evar_map -> Constr.constr -> Constr.constr -> bool +val eq_constr_univs_test : evar_map -> evar_map -> constr -> constr -> bool (** [compare_cumulative_instances cv_pb variance u1 u2 sigma] Returns [Inl sigma'] where [sigma'] is [sigma] augmented with universe diff --git a/engine/evd.ml b/engine/evd.ml index f6e13e1f4..6dcec2760 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -21,6 +21,9 @@ open Environ (* module RelDecl = Context.Rel.Declaration *) module NamedDecl = Context.Named.Declaration +type econstr = constr +type etypes = types + (** Generic filters *) module Filter : sig @@ -537,10 +540,14 @@ let existential_value d (n, args) = | Evar_empty -> raise NotInstantiatedEvar +let existential_value0 = existential_value + let existential_opt_value d ev = try Some (existential_value d ev) with NotInstantiatedEvar -> None +let existential_opt_value0 = existential_opt_value + let existential_type d (n, args) = let info = try find d n @@ -548,6 +555,8 @@ let existential_type d (n, args) = anomaly (str "Evar " ++ str (string_of_existential n) ++ str " was not declared.") in instantiate_evar_array info info.evar_concl args +let existential_type0 = existential_type + let add_constraints d c = { d with universes = UState.add_constraints d.universes c } @@ -1065,6 +1074,7 @@ let meta_ftype evd mv = | Clval(_,_,b) -> b let meta_type evd mv = (meta_ftype evd mv).rebus +let meta_type0 = meta_type let meta_declare mv v ?(name=Anonymous) evd = let metas = Metamap.add mv (Cltyp(name,mk_freelisted v)) evd.metas in @@ -1217,3 +1227,98 @@ let normalize_evar_universe_context_variables = UState.normalize_variables let abstract_undefined_variables = UState.abstract_undefined_variables let normalize_evar_universe_context = UState.minimize let nf_constraints = minimize_universes + +module MiniEConstr = struct + + module ESorts = + struct + type t = Sorts.t + let make s = s + let kind sigma = function + | Sorts.Type u -> Sorts.sort_of_univ (normalize_universe sigma u) + | s -> s + let unsafe_to_sorts s = s + end + + module EInstance = + struct + type t = Univ.Instance.t + let make i = i + let kind sigma i = + if Univ.Instance.is_empty i then i + else normalize_universe_instance sigma i + let empty = Univ.Instance.empty + let is_empty = Univ.Instance.is_empty + let unsafe_to_instance t = t + end + + type t = econstr + + let safe_evar_value sigma ev = + try Some (existential_value sigma ev) + with NotInstantiatedEvar | Not_found -> None + + let rec whd_evar sigma c = + match Constr.kind c with + | Evar ev -> + begin match safe_evar_value sigma ev with + | Some c -> whd_evar sigma c + | None -> c + end + | App (f, args) when isEvar f -> + (** Enforce smart constructor invariant on applications *) + let ev = destEvar f in + begin match safe_evar_value sigma ev with + | None -> c + | Some f -> whd_evar sigma (mkApp (f, args)) + end + | Cast (c0, k, t) when isEvar c0 -> + (** Enforce smart constructor invariant on casts. *) + let ev = destEvar c0 in + begin match safe_evar_value sigma ev with + | None -> c + | Some c -> whd_evar sigma (mkCast (c, k, t)) + end + | _ -> c + + let kind sigma c = Constr.kind (whd_evar sigma c) + let kind_upto = kind + let kind_of_type sigma c = Term.kind_of_type (whd_evar sigma c) + let of_kind = Constr.of_kind + let of_constr c = c + let unsafe_to_constr c = c + let unsafe_eq = Refl + + let to_constr ?(abort_on_undefined_evars=true) sigma c = + let rec to_constr c = match Constr.kind c with + | Evar ev -> + begin match safe_evar_value sigma ev with + | Some c -> to_constr c + | None -> + if abort_on_undefined_evars then + anomaly ~label:"econstr" Pp.(str "grounding a non evar-free term") + else + Constr.map (fun c -> to_constr c) c + end + | Sort (Sorts.Type u) -> + let u' = normalize_universe sigma u in + if u' == u then c else mkSort (Sorts.sort_of_univ u') + | Const (c', u) when not (Univ.Instance.is_empty u) -> + let u' = normalize_universe_instance sigma u in + if u' == u then c else mkConstU (c', u') + | Ind (i, u) when not (Univ.Instance.is_empty u) -> + let u' = normalize_universe_instance sigma u in + if u' == u then c else mkIndU (i, u') + | Construct (co, u) when not (Univ.Instance.is_empty u) -> + let u' = normalize_universe_instance sigma u in + if u' == u then c else mkConstructU (co, u') + | _ -> Constr.map (fun c -> to_constr c) c + in to_constr c + + let of_named_decl d = d + let unsafe_to_named_decl d = d + let of_rel_decl d = d + let unsafe_to_rel_decl d = d + let to_rel_decl sigma d = Context.Rel.Declaration.map_constr (to_constr sigma) d + +end diff --git a/engine/evd.mli b/engine/evd.mli index 911799c44..5ce16459c 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -28,6 +28,9 @@ open Environ It also contains conversion constraints, debugging information and information about meta variables. *) +type econstr +type etypes = econstr + (** {5 Existential variables and unification states} *) type evar = Evar.t @@ -86,16 +89,16 @@ end type evar_body = | Evar_empty - | Evar_defined of constr + | Evar_defined of econstr module Store : Store.S (** Datatype used to store additional information in evar maps. *) type evar_info = { - evar_concl : constr; + evar_concl : econstr; (** Type of the evar. *) - evar_hyps : named_context_val; + evar_hyps : named_context_val; (** TODO econstr? *) (** Context of the evar. *) evar_body : evar_body; (** Optional content of the evar. *) @@ -105,16 +108,16 @@ type evar_info = { in the solution *) evar_source : Evar_kinds.t located; (** Information about the evar. *) - evar_candidates : constr list option; + evar_candidates : econstr list option; (** List of possible solutions when known that it is a finite list *) evar_extra : Store.t (** Extra store, used for clever hacks. *) } -val make_evar : named_context_val -> types -> evar_info -val evar_concl : evar_info -> constr -val evar_context : evar_info -> Context.Named.t -val evar_filtered_context : evar_info -> Context.Named.t +val make_evar : named_context_val -> etypes -> evar_info +val evar_concl : evar_info -> econstr +val evar_context : evar_info -> (econstr, etypes) Context.Named.pt +val evar_filtered_context : evar_info -> (econstr, etypes) Context.Named.pt val evar_hyps : evar_info -> named_context_val val evar_filtered_hyps : evar_info -> named_context_val val evar_body : evar_info -> evar_body @@ -122,8 +125,8 @@ val evar_filter : evar_info -> Filter.t val evar_env : evar_info -> env val evar_filtered_env : evar_info -> env -val map_evar_body : (constr -> constr) -> evar_body -> evar_body -val map_evar_info : (constr -> constr) -> evar_info -> evar_info +val map_evar_body : (econstr -> econstr) -> evar_body -> evar_body +val map_evar_info : (econstr -> econstr) -> evar_info -> evar_info (** {6 Unification state} **) @@ -190,7 +193,7 @@ val raw_map_undefined : (Evar.t -> evar_info -> evar_info) -> evar_map -> evar_m (** Same as {!raw_map}, but restricted to undefined evars. For efficiency reasons. *) -val define : Evar.t-> constr -> evar_map -> evar_map +val define : Evar.t-> econstr -> evar_map -> evar_map (** Set the body of an evar to the given constr. It is expected that: {ul {- The evar is already present in the evarmap.} @@ -198,7 +201,7 @@ val define : Evar.t-> constr -> evar_map -> evar_map {- All the evars present in the constr should be present in the evar map.} } *) -val cmap : (constr -> constr) -> evar_map -> evar_map +val cmap : (econstr -> econstr) -> evar_map -> evar_map (** Map the function on all terms in the evar map. *) val is_evar : evar_map -> Evar.t-> bool @@ -222,20 +225,26 @@ val drop_all_defined : evar_map -> evar_map exception NotInstantiatedEvar -val existential_value : evar_map -> existential -> constr +val existential_value : evar_map -> econstr pexistential -> econstr (** [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has no body and [Not_found] if it does not exist in [sigma] *) -val existential_type : evar_map -> existential -> types +val existential_value0 : evar_map -> existential -> constr + +val existential_type : evar_map -> econstr pexistential -> etypes -val existential_opt_value : evar_map -> existential -> constr option +val existential_type0 : evar_map -> existential -> types + +val existential_opt_value : evar_map -> econstr pexistential -> econstr option (** Same as {!existential_value} but returns an option instead of raising an exception. *) +val existential_opt_value0 : evar_map -> existential -> constr option + val evar_instance_array : (Context.Named.Declaration.t -> 'a -> bool) -> evar_info -> 'a array -> (Id.t * 'a) list -val instantiate_evar_array : evar_info -> constr -> constr array -> constr +val instantiate_evar_array : evar_info -> econstr -> econstr array -> econstr val evars_reset_evd : ?with_conv_pbs:bool -> ?with_univs:bool -> evar_map -> evar_map -> evar_map @@ -243,7 +252,7 @@ val evars_reset_evd : ?with_conv_pbs:bool -> ?with_univs:bool -> (** {6 Misc} *) -val restrict : Evar.t-> Filter.t -> ?candidates:constr list -> +val restrict : Evar.t-> Filter.t -> ?candidates:econstr list -> ?src:Evar_kinds.t located -> evar_map -> evar_map * Evar.t (** Restrict an undefined evar into a new evar by filtering context and possibly limiting the instances to a set of candidates *) @@ -251,7 +260,7 @@ val restrict : Evar.t-> Filter.t -> ?candidates:constr list -> val is_restricted_evar : evar_info -> Evar.t option (** Tell if an evar comes from restriction of another evar, and if yes, which *) -val downcast : Evar.t-> types -> evar_map -> evar_map +val downcast : Evar.t-> etypes -> evar_map -> evar_map (** Change the type of an undefined evar to a new type assumed to be a subtype of its current type; subtyping must be ensured by caller *) @@ -341,7 +350,7 @@ val shelve_on_future_goals : Evar.t list -> future_goals -> future_goals Evar maps also keep track of the universe constraints defined at a given point. This section defines the relevant manipulation functions. *) -val whd_sort_variable : evar_map -> constr -> constr +val whd_sort_variable : evar_map -> econstr -> econstr exception UniversesDiffer @@ -397,8 +406,8 @@ type 'a freelisted = { rebus : 'a; freemetas : Metaset.t } -val metavars_of : constr -> Metaset.t -val mk_freelisted : constr -> constr freelisted +val metavars_of : econstr -> Metaset.t +val mk_freelisted : econstr -> econstr freelisted val map_fl : ('a -> 'b) -> 'a freelisted -> 'b freelisted (** Status of an instance found by unification wrt to the meta it solves: @@ -436,12 +445,12 @@ type instance_status = instance_constraint * instance_typing_status (** Clausal environments *) type clbinding = - | Cltyp of Name.t * constr freelisted - | Clval of Name.t * (constr freelisted * instance_status) * constr freelisted + | Cltyp of Name.t * econstr freelisted + | Clval of Name.t * (econstr freelisted * instance_status) * econstr freelisted (** Unification constraints *) type conv_pb = Reduction.conv_pb -type evar_constraint = conv_pb * env * constr * constr +type evar_constraint = conv_pb * env * econstr * econstr val add_conv_pb : ?tail:bool -> evar_constraint -> evar_map -> evar_map val extract_changed_conv_pbs : evar_map -> @@ -457,7 +466,7 @@ val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t option val evars_of_term : constr -> Evar.Set.t (** including evars in instances of evars *) -val evars_of_named_context : Context.Named.t -> Evar.Set.t +val evars_of_named_context : (econstr, etypes) Context.Named.pt -> Evar.Set.t val evars_of_filtered_evar_info : evar_info -> Evar.Set.t @@ -465,19 +474,20 @@ val evars_of_filtered_evar_info : evar_info -> Evar.Set.t val meta_list : evar_map -> (metavariable * clbinding) list val meta_defined : evar_map -> metavariable -> bool -val meta_value : evar_map -> metavariable -> constr +val meta_value : evar_map -> metavariable -> econstr (** [meta_fvalue] raises [Not_found] if meta not in map or [Anomaly] if meta has no value *) -val meta_fvalue : evar_map -> metavariable -> constr freelisted * instance_status -val meta_opt_fvalue : evar_map -> metavariable -> (constr freelisted * instance_status) option -val meta_type : evar_map -> metavariable -> types -val meta_ftype : evar_map -> metavariable -> types freelisted +val meta_fvalue : evar_map -> metavariable -> econstr freelisted * instance_status +val meta_opt_fvalue : evar_map -> metavariable -> (econstr freelisted * instance_status) option +val meta_type : evar_map -> metavariable -> etypes +val meta_type0 : evar_map -> metavariable -> types +val meta_ftype : evar_map -> metavariable -> etypes freelisted val meta_name : evar_map -> metavariable -> Name.t val meta_declare : - metavariable -> types -> ?name:Name.t -> evar_map -> evar_map -val meta_assign : metavariable -> constr * instance_status -> evar_map -> evar_map -val meta_reassign : metavariable -> constr * instance_status -> evar_map -> evar_map + metavariable -> etypes -> ?name:Name.t -> evar_map -> evar_map +val meta_assign : metavariable -> econstr * instance_status -> evar_map -> evar_map +val meta_reassign : metavariable -> econstr * instance_status -> evar_map -> evar_map val clear_metas : evar_map -> evar_map @@ -485,10 +495,10 @@ val clear_metas : evar_map -> evar_map val meta_merge : ?with_univs:bool -> evar_map -> evar_map -> evar_map val undefined_metas : evar_map -> metavariable list -val map_metas_fvalue : (constr -> constr) -> evar_map -> evar_map -val map_metas : (constr -> constr) -> evar_map -> evar_map +val map_metas_fvalue : (econstr -> econstr) -> evar_map -> evar_map +val map_metas : (econstr -> econstr) -> evar_map -> evar_map -type metabinding = metavariable * constr * instance_status +type metabinding = metavariable * econstr * instance_status val retract_coercible_metas : evar_map -> metabinding list * evar_map @@ -639,13 +649,13 @@ val fresh_inductive_instance : ?loc:Loc.t -> env -> evar_map -> inductive -> eva val fresh_constructor_instance : ?loc:Loc.t -> env -> evar_map -> constructor -> evar_map * pconstructor val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> env -> - evar_map -> Globnames.global_reference -> evar_map * constr + evar_map -> Globnames.global_reference -> evar_map * econstr (********************************************************************) (* constr with holes and pending resolution of classes, conversion *) (* problems, candidates, etc. *) -type open_constr = evar_map * constr (* Special case when before is empty *) +type open_constr = evar_map * econstr (* Special case when before is empty *) (** Partially constructed constrs. *) @@ -665,3 +675,50 @@ val create_evar_defs : evar_map -> evar_map (** Create an [evar_map] with empty meta map: *) +(** Use this module only to bootstrap EConstr *) +module MiniEConstr : sig + module ESorts : sig + type t + val make : Sorts.t -> t + val kind : evar_map -> t -> Sorts.t + val unsafe_to_sorts : t -> Sorts.t + end + + module EInstance : sig + type t + val make : Univ.Instance.t -> t + val kind : evar_map -> t -> Univ.Instance.t + val empty : t + val is_empty : t -> bool + val unsafe_to_instance : t -> Univ.Instance.t + end + + type t = econstr + + val kind : evar_map -> t -> (t, t, ESorts.t, EInstance.t) Constr.kind_of_term + val kind_upto : evar_map -> constr -> (constr, types, Sorts.t, Univ.Instance.t) Constr.kind_of_term + val kind_of_type : evar_map -> t -> (t, t) Term.kind_of_type + + val whd_evar : evar_map -> t -> t + + val of_kind : (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -> t + + val of_constr : Constr.t -> t + + val to_constr : ?abort_on_undefined_evars:bool -> evar_map -> t -> Constr.t + + val unsafe_to_constr : t -> Constr.t + + val unsafe_eq : (t, Constr.t) eq + + val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> + (t, t) Context.Named.Declaration.pt + val unsafe_to_named_decl : (t, t) Context.Named.Declaration.pt -> + (Constr.t, Constr.types) Context.Named.Declaration.pt + val unsafe_to_rel_decl : (t, t) Context.Rel.Declaration.pt -> + (Constr.t, Constr.types) Context.Rel.Declaration.pt + val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> + (t, t) Context.Rel.Declaration.pt + val to_rel_decl : evar_map -> (t, t) Context.Rel.Declaration.pt -> + (Constr.t, Constr.types) Context.Rel.Declaration.pt +end diff --git a/engine/proofview.ml b/engine/proofview.ml index 639f48e77..54237ceb4 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -45,9 +45,9 @@ let compact el ({ solution } as pv) = let pruned_solution = Evd.drop_all_defined solution in let apply_subst_einfo _ ei = Evd.({ ei with - evar_concl = nf0 ei.evar_concl; + evar_concl = nf ei.evar_concl; evar_hyps = Environ.map_named_val nf0 ei.evar_hyps; - evar_candidates = Option.map (List.map nf0) ei.evar_candidates }) in + evar_candidates = Option.map (List.map nf) ei.evar_candidates }) in let new_solution = Evd.raw_map_undefined apply_subst_einfo pruned_solution in let new_size = Evd.fold (fun _ _ i -> i+1) new_solution 0 in Feedback.msg_info (Pp.str (Printf.sprintf "Evars: %d -> %d\n" size new_size)); @@ -875,8 +875,7 @@ module Progress = struct (** equality function on hypothesis contexts *) let eq_named_context_val sigma1 sigma2 ctx1 ctx2 = - let open Environ in - let c1 = named_context_of_val ctx1 and c2 = named_context_of_val ctx2 in + let c1 = EConstr.named_context_of_val ctx1 and c2 = EConstr.named_context_of_val ctx2 in let eq_named_declaration d1 d2 = match d1, d2 with | LocalAssum (i1,t1), LocalAssum (i2,t2) -> @@ -1101,7 +1100,7 @@ module Goal = struct let gmake_with info env sigma goal state = { env = Environ.reset_with_named_context (Evd.evar_filtered_hyps info) env ; sigma = sigma ; - concl = EConstr.of_constr (Evd.evar_concl info); + concl = Evd.evar_concl info; state = state ; self = goal } diff --git a/engine/termops.ml b/engine/termops.ml index b7531f6fc..df43be28e 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -115,7 +115,7 @@ let pr_evar_suggested_name evk sigma = | _,Evar_kinds.GoalEvar -> Id.of_string "Goal" | _ -> let env = reset_with_named_context evi.evar_hyps (Global.env()) in - Namegen.id_of_name_using_hdchar env sigma (EConstr.of_constr evi.evar_concl) Anonymous + Namegen.id_of_name_using_hdchar env sigma evi.evar_concl Anonymous in let names = EvMap.mapi base_id (undefined_map sigma) in let id = EvMap.find evk names in @@ -154,7 +154,7 @@ let protect f x = with e -> str "EXCEPTION: " ++ str (Printexc.to_string e) let print_kconstr a = - protect (fun c -> print_constr (EConstr.of_constr c)) a + protect (fun c -> print_constr c) a let pr_meta_map evd = let open Evd in @@ -197,11 +197,11 @@ let pr_evar_source = function let print_constr = print_kconstr in let id = Option.get ido in str "parameter " ++ Id.print id ++ spc () ++ str "of" ++ - spc () ++ print_constr (printable_constr_of_global c) + spc () ++ print_constr (EConstr.of_constr @@ printable_constr_of_global c) | Evar_kinds.InternalHole -> str "internal placeholder" | Evar_kinds.TomatchTypeParameter (ind,n) -> let print_constr = print_kconstr in - pr_nth n ++ str " argument of type " ++ print_constr (mkInd ind) + pr_nth n ++ str " argument of type " ++ print_constr (EConstr.mkInd ind) | Evar_kinds.GoalEvar -> str "goal evar" | Evar_kinds.ImpossibleCase -> str "type of impossible pattern-matching clause" | Evar_kinds.MatchingVar _ -> str "matching variable" @@ -256,7 +256,7 @@ let compute_evar_dependency_graph sigma = in match evar_body evi with | Evar_empty -> acc - | Evar_defined c -> Evar.Set.fold fold_ev (evars_of_term c) acc + | Evar_defined c -> Evar.Set.fold fold_ev (evars_of_term (EConstr.Unsafe.to_constr c)) acc in Evd.fold fold sigma EvMap.empty @@ -314,7 +314,8 @@ let print_env_short env = let print_constr = print_kconstr in let pr_rel_decl = function | RelDecl.LocalAssum (n,_) -> Name.print n - | RelDecl.LocalDef (n,b,_) -> str "(" ++ Name.print n ++ str " := " ++ print_constr b ++ str ")" + | RelDecl.LocalDef (n,b,_) -> str "(" ++ Name.print n ++ str " := " + ++ print_constr (EConstr.of_constr b) ++ str ")" in let pr_named_decl = NamedDecl.to_rel_decl %> pr_rel_decl in let nc = List.rev (named_context env) in @@ -335,11 +336,11 @@ let pr_evar_constraints sigma pbs = Namegen.make_all_name_different env sigma in print_env_short env ++ spc () ++ str "|-" ++ spc () ++ - protect (print_constr_env env sigma) (EConstr.of_constr t1) ++ spc () ++ + protect (print_constr_env env sigma) t1 ++ spc () ++ str (match pbty with | Reduction.CONV -> "==" | Reduction.CUMUL -> "<=") ++ - spc () ++ protect (print_constr_env env Evd.empty) (EConstr.of_constr t2) + spc () ++ protect (print_constr_env env Evd.empty) t2 in prlist_with_sep fnl pr_evconstr pbs diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index 83fbc2d5d..7c612c0d8 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -197,7 +197,7 @@ let extend_with_ref_list env sigma l seq = let l = expand_constructor_hints l in let f gr (seq, sigma) = let sigma, c = Evd.fresh_global env sigma gr in - let sigma, typ= Typing.type_of env sigma (EConstr.of_constr c) in + let sigma, typ= Typing.type_of env sigma c in (add_formula env sigma Hyp gr typ seq, sigma) in List.fold_right f l (seq, sigma) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index d04887a48..8da0e1c4f 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -1050,8 +1050,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a (Global.env ()) !evd (Constrintern.locate_reference (qualid_of_ident equation_lemma_id)) in - let res = EConstr.of_constr res in - evd:=evd'; + evd:=evd'; let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd res in res in diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 804548ce5..04a23cdb9 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -266,7 +266,7 @@ let change_property_sort evd toSort princ princName = (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident princName)) in let init = let nargs = (princ_info.nparams + (List.length princ_info.predicates)) in - mkApp(princName_as_constr, + mkApp(EConstr.Unsafe.to_constr princName_as_constr, Array.init nargs (fun i -> mkRel (nargs - i ))) in @@ -630,11 +630,14 @@ let build_scheme fas = in let evd',f = Evd.fresh_global (Global.env ()) !evd f_as_constant in let _ = evd := evd' in - let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr f) in - if isConst f - then (destConst f,sort) - else user_err Pp.(pr_constr_env (Global.env ()) !evd f ++spc () ++ str "should be the named of a globally defined function") - ) + let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd f in + let c, u = + try EConstr.destConst !evd f + with DestKO -> + user_err Pp.(pr_econstr_env (Global.env ()) !evd f ++spc () ++ str "should be the named of a globally defined function") + in + (c, EConstr.EInstance.kind !evd u), sort + ) fas ) in let bodies_types = diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 40ea40b6b..845104c3c 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -564,6 +564,7 @@ let resolve_and_replace_implicits ?(flags=Pretyping.all_and_fail_flags) ?(expect If someone knows how to prevent solved existantial removal in understand, please do not hesitate to change the computation of [ctx] here *) let ctx,_,_ = Pretyping.ise_pretype_gen flags env sigma Glob_ops.empty_lvar expected_type rt in let ctx, f = Evarutil.nf_evars_and_universes ctx in + let f c = EConstr.of_constr (f (EConstr.Unsafe.to_constr c)) in (* then we map [rt] to replace the implicit holes by their values *) let rec change rt = @@ -586,8 +587,8 @@ If someone knows how to prevent solved existantial removal in understand, pleas with Found evi -> (* we found the evar corresponding to this hole *) match evi.evar_body with | Evar_defined c -> - (* we just have to lift the solution in glob_term *) - Detyping.detype Detyping.Now false Id.Set.empty env ctx (EConstr.of_constr (f c)) + (* we just have to lift the solution in glob_term *) + Detyping.detype Detyping.Now false Id.Set.empty env ctx (f c) | Evar_empty -> rt (* the hole was not solved : we do nothing *) ) | (GHole(BinderType na,_,_)) -> (* we only want to deal with implicit arguments *) @@ -609,7 +610,7 @@ If someone knows how to prevent solved existantial removal in understand, pleas match evi.evar_body with | Evar_defined c -> (* we just have to lift the solution in glob_term *) - Detyping.detype Detyping.Now false Id.Set.empty env ctx (EConstr.of_constr (f c)) + Detyping.detype Detyping.Now false Id.Set.empty env ctx (f c) | Evar_empty -> rt (* the hole was not solved : we d when falseo nothing *) in res diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index d395e3601..748d8add2 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -77,8 +77,7 @@ let functional_induction with_clean c princl pat = user_err (str "Cannot find induction principle for " ++ Printer.pr_leconstr_env (Tacmach.pf_env g) sigma (mkConst c') ) in - let princ = EConstr.of_constr princ in - (princ,NoBindings,Tacmach.pf_unsafe_type_of g' princ,g') + (princ,NoBindings,Tacmach.pf_unsafe_type_of g' princ,g') | _ -> raise (UserError(None,str "functional induction must be used with a function" )) end | Some ((princ,binding)) -> @@ -261,7 +260,6 @@ let derive_inversion fix_names = let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident id)) in - let c = EConstr.of_constr c in let (cst, u) = destConst evd c in evd, (cst, EInstance.kind evd u) :: l ) @@ -283,8 +281,7 @@ let derive_inversion fix_names = (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident (mk_rel_id id))) in - let id = EConstr.of_constr id in - evd,(fst (destInd evd id))::l + evd,(fst (destInd evd id))::l ) fix_names (evd',[]) @@ -388,7 +385,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error let evd = ref (Evd.from_env env) in let evd',uprinc = Evd.fresh_global env !evd princ in let _ = evd := evd' in - let princ_type = Typing.e_type_of ~refresh:true env evd (EConstr.of_constr uprinc) in + let princ_type = Typing.e_type_of ~refresh:true env evd uprinc in let princ_type = EConstr.Unsafe.to_constr princ_type in Functional_principles_types.generate_functional_principle evd @@ -425,7 +422,6 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in - let c = EConstr.of_constr c in let (cst, u) = destConst evd c in let u = EInstance.kind evd u in evd,((cst, u) :: l) @@ -442,7 +438,6 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in - let c = EConstr.of_constr c in let (cst, u) = destConst evd c in let u = EInstance.kind evd u in evd,((cst, u) :: l) diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index ae84eaa93..28e85268a 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -102,7 +102,6 @@ let generate_type evd g_to_f f graph i = let evd',graph = Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd !evd graph))) in - let graph = EConstr.of_constr graph in evd:=evd'; let graph_arity = Typing.e_type_of (Global.env ()) evd graph in let ctxt,_ = decompose_prod_assum !evd graph_arity in @@ -172,7 +171,6 @@ let find_induction_principle evd f = | None -> raise Not_found | Some rect_lemma -> let evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (Globnames.ConstRef rect_lemma) in - let rect_lemma = EConstr.of_constr rect_lemma in let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' rect_lemma in evd:=evd'; rect_lemma,typ @@ -823,8 +821,7 @@ let derive_correctness make_scheme (funs: pconstant list) (graphs:inductive list (* let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in *) let _,lem_cst_constr = Evd.fresh_global (Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in - let lem_cst_constr = EConstr.of_constr lem_cst_constr in - let (lem_cst,_) = destConst !evd lem_cst_constr in + let (lem_cst,_) = destConst !evd lem_cst_constr in update_Function {finfo with correctness_lemma = Some lem_cst}; ) @@ -884,8 +881,7 @@ let derive_correctness make_scheme (funs: pconstant list) (graphs:inductive list let finfo = find_Function_infos (fst f_as_constant) in let _,lem_cst_constr = Evd.fresh_global (Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in - let lem_cst_constr = EConstr.of_constr lem_cst_constr in - let (lem_cst,_) = destConst !evd lem_cst_constr in + let (lem_cst,_) = destConst !evd lem_cst_constr in update_Function {finfo with completeness_lemma = Some lem_cst} ) funs) diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 2e90ce90c..a5f8060ae 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -296,7 +296,6 @@ let project_hint ~poly pri l2r r = let env = Global.env() in let sigma = Evd.from_env env in let sigma, c = Evd.fresh_global env sigma gr in - let c = EConstr.of_constr c in let t = Retyping.get_type_of env sigma c in let t = Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) t in @@ -307,7 +306,6 @@ let project_hint ~poly pri l2r r = let p = if l2r then build_coq_iff_left_proj () else build_coq_iff_right_proj () in let sigma, p = Evd.fresh_global env sigma p in - let p = EConstr.of_constr p in let c = Reductionops.whd_beta sigma (mkApp (c, Context.Rel.to_extended_vect mkRel 0 sign)) in let c = it_mkLambda_or_LetIn (mkApp (p,[|mkArrow a (lift 1 b);mkArrow b (lift 1 a);c|])) sign in diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index d32a2faef..9eb55aa5e 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -428,7 +428,8 @@ let split_head = function | [] -> assert(false) let eq_pb (ty, env, x, y as pb) (ty', env', x', y' as pb') = - pb == pb' || (ty == ty' && Constr.equal x x' && Constr.equal y y') + let equal x y = Constr.equal (EConstr.Unsafe.to_constr x) (EConstr.Unsafe.to_constr y) in + pb == pb' || (ty == ty' && equal x x' && equal y y') let problem_inclusion x y = List.for_all (fun pb -> List.exists (fun pb' -> eq_pb pb pb') y) x @@ -626,9 +627,9 @@ let solve_remaining_by env sigma holes by = (** Evar should not be defined, but just in case *) | Some evi -> let env = Environ.reset_with_named_context evi.evar_hyps env in - let ty = EConstr.of_constr evi.evar_concl in + let ty = evi.evar_concl in let c, sigma = Pfedit.refine_by_tactic env sigma ty solve_tac in - Evd.define evk c sigma + Evd.define evk (EConstr.of_constr c) sigma in List.fold_left solve sigma indep @@ -1862,7 +1863,6 @@ let declare_projection n instance_id r = let env = Global.env () in let sigma = Evd.from_env env in let sigma,c = Evd.fresh_global env sigma r in - let c = EConstr.of_constr c in let ty = Retyping.get_type_of env sigma c in let term = proper_projection sigma c ty in let sigma, typ = Typing.type_of env sigma term in diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml index e07bc48a4..06c26edbe 100644 --- a/plugins/ssr/ssrcommon.ml +++ b/plugins/ssr/ssrcommon.ml @@ -509,11 +509,12 @@ let pf_abs_evars2 gl rigid (sigma, c0) = let nenv = env_size (pf_env gl) in let abs_evar n k = let evi = Evd.find sigma k in - let dc = CList.firstn n (evar_filtered_context evi) in + let concl = EConstr.Unsafe.to_constr evi.evar_concl in + let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in let abs_dc c = function | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c) | NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in - let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in + let t = Context.Named.fold_inside abs_dc ~init:concl dc in nf_evar sigma t in let rec put evlist c = match Constr.kind c with | Evar (k, a) -> @@ -569,11 +570,12 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let nenv = env_size (pf_env gl) in let abs_evar n k = let evi = Evd.find sigma k in - let dc = CList.firstn n (evar_filtered_context evi) in + let concl = EConstr.Unsafe.to_constr evi.evar_concl in + let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in let abs_dc c = function | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c) | NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in - let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in + let t = Context.Named.fold_inside abs_dc ~init:concl dc in nf_evar sigma0 (nf_evar sigma t) in let rec put evlist c = match Constr.kind c with | Evar (k, a) -> @@ -581,7 +583,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let n = max 0 (Array.length a - nenv) in let k_ty = Retyping.get_sort_family_of - (pf_env gl) sigma (EConstr.of_constr (Evd.evar_concl (Evd.find sigma k))) in + (pf_env gl) sigma (Evd.evar_concl (Evd.find sigma k)) in let is_prop = k_ty = InProp in let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t | _ -> Constr.fold put evlist c in @@ -746,7 +748,7 @@ let pf_mkSsrConst name gl = let pf_fresh_global name gl = let sigma, env, it = project gl, pf_env gl, sig_it gl in let sigma,t = Evd.fresh_global env sigma name in - t, re_sig it sigma + EConstr.Unsafe.to_constr t, re_sig it sigma let mkProt t c gl = let prot, gl = pf_mkSsrConst "protect_term" gl in @@ -980,7 +982,7 @@ let applyn ~with_evars ?beta ?(with_shelve=false) n t gl = if not (EConstr.Vars.closed0 sigma ty) then raise dependent_apply_error; let m = Evarutil.new_meta () in - loop (meta_declare m (EConstr.Unsafe.to_constr ty) sigma) bo ((EConstr.mkMeta m)::args) (n-1) + loop (meta_declare m ty sigma) bo ((EConstr.mkMeta m)::args) (n-1) | _ -> assert false in loop sigma t [] n in pp(lazy(str"Refiner.refiner " ++ Printer.pr_econstr_env (pf_env gl) (project gl) t)); diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml index 717657a24..de8ffb976 100644 --- a/plugins/ssr/ssrelim.ml +++ b/plugins/ssr/ssrelim.ml @@ -356,7 +356,7 @@ let ssrelim ?(ind=ref None) ?(is_case=false) deps what ?elim eqid elim_intro_tac let ev = List.fold_left Evar.Set.union Evar.Set.empty patterns_ev in let ty_ev = Evar.Set.fold (fun i e -> let ex = i in - let i_ty = EConstr.of_constr (Evd.evar_concl (Evd.find (project gl) ex)) in + let i_ty = Evd.evar_concl (Evd.find (project gl) ex) in Evar.Set.union e (evars_of_term i_ty)) ev Evar.Set.empty in let inter = Evar.Set.inter ev ty_ev in diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml index 2f0433b64..7748ba2b0 100644 --- a/plugins/ssr/ssrequality.ml +++ b/plugins/ssr/ssrequality.ml @@ -359,7 +359,7 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl = let evs = Evar.Set.elements (Evarutil.undefined_evars_of_term sigma t) in let open_evs = List.filter (fun k -> Sorts.InProp <> Retyping.get_sort_family_of - env sigma (EConstr.of_constr (Evd.evar_concl (Evd.find sigma k)))) + env sigma (Evd.evar_concl (Evd.find sigma k))) evs in if open_evs <> [] then Some name else None) (List.combine (Array.to_list args) names) @@ -478,10 +478,10 @@ let rwprocess_rule dir rule gl = | _ -> let ra = Array.append a [|r|] in function 1 -> let sigma, pi1 = Evd.fresh_global env sigma coq_prod.Coqlib.proj1 in - EConstr.mkApp (EConstr.of_constr pi1, ra), sigma + EConstr.mkApp (pi1, ra), sigma | _ -> let sigma, pi2 = Evd.fresh_global env sigma coq_prod.Coqlib.proj2 in - EConstr.mkApp (EConstr.of_constr pi2, ra), sigma in + EConstr.mkApp (pi2, ra), sigma in if EConstr.eq_constr sigma a.(0) (EConstr.of_constr (Universes.constr_of_global @@ Coqlib.build_coq_True ())) then let s, sigma = sr sigma 2 in loop (converse_dir d) sigma s a.(1) rs 0 diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml index af78bf36a..6b7a96deb 100644 --- a/plugins/ssr/ssripats.ml +++ b/plugins/ssr/ssripats.ml @@ -623,7 +623,7 @@ let tacFIND_ABSTRACT_PROOF check_lock abstract_n = Goal.enter_one ~__LOC__ begin fun g -> let sigma, env = Goal.(sigma g, env g) in let l = Evd.fold_undefined (fun e ei l -> - match EConstr.kind sigma (EConstr.of_constr ei.Evd.evar_concl) with + match EConstr.kind sigma ei.Evd.evar_concl with | Term.App(hd, [|ty; n; lock|]) when (not check_lock || (occur_existential_or_casted_meta sigma ty && diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml index afe03533d..fc50b24a6 100644 --- a/plugins/ssr/ssrview.ml +++ b/plugins/ssr/ssrview.ml @@ -265,7 +265,7 @@ Goal.enter_one ~__LOC__ begin fun g -> List.fold_left (fun l k -> if Evd.is_defined sigma k then let bo = get_body Evd.(evar_body (find sigma k)) in - k :: l @ Evar.Set.elements (evars_of_econstr sigma bo) + k :: l @ Evar.Set.elements (evars_of_econstr sigma (EConstr.Unsafe.to_constr bo)) else l ) [] s in let und0 = (* Unassigned evars in the initial goal *) diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 2ba6acc03..a10437a63 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -283,7 +283,7 @@ exception NoProgress (* comparison can be much faster than the HO one. *) let unif_EQ env sigma p c = - let evars = existential_opt_value sigma, Evd.universes sigma in + let evars = existential_opt_value0 sigma, Evd.universes sigma in try let _ = Reduction.conv env p ~evars c in true with _ -> false let unif_EQ_args env sigma pa a = @@ -337,7 +337,7 @@ let nf_open_term sigma0 ise c = let s = ise and s' = ref sigma0 in let rec nf c' = match kind c' with | Evar ex -> - begin try nf (existential_value s ex) with _ -> + begin try nf (existential_value0 s ex) with _ -> let k, a = ex in let a' = Array.map nf a in if not (Evd.mem !s' k) then s' := Evd.add !s' k (Evarutil.nf_evar_info s (Evd.find s k)); @@ -347,7 +347,9 @@ let nf_open_term sigma0 ise c = let copy_def k evi () = if evar_body evi != Evd.Evar_empty then () else match Evd.evar_body (Evd.find s k) with - | Evar_defined c' -> s' := Evd.define k (nf c') !s' + | Evar_defined c' -> + let c' = EConstr.of_constr (nf (EConstr.Unsafe.to_constr c')) in + s' := Evd.define k c' !s' | _ -> () in let c' = nf c in let _ = Evd.fold copy_def sigma0 () in !s', Evd.evar_universe_context s, EConstr.of_constr c' @@ -446,7 +448,7 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 = let nenv = env_size env + if hack then 1 else 0 in let rec put c = match kind c with | Evar (k, a as ex) -> - begin try put (existential_value !sigma ex) + begin try put (existential_value0 !sigma ex) with NotInstantiatedEvar -> if Evd.mem sigma0 k then map put c else let evi = Evd.find !sigma k in @@ -457,11 +459,13 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 = | Context.Named.Declaration.LocalAssum (x, t) -> mkVar x :: d, mkNamedProd x (put t) c in let a, t = - Context.Named.fold_inside abs_dc ~init:([], (put evi.evar_concl)) dc in + Context.Named.fold_inside abs_dc + ~init:([], (put @@ EConstr.Unsafe.to_constr evi.evar_concl)) + (EConstr.Unsafe.to_named_context dc) in let m = Evarutil.new_meta () in - ise := meta_declare m t !ise; - sigma := Evd.define k (applistc (mkMeta m) a) !sigma; - put (existential_value !sigma ex) + ise := meta_declare m (EConstr.of_constr t) !ise; + sigma := Evd.define k (EConstr.of_constr (applistc (mkMeta m) a)) !sigma; + put (existential_value0 !sigma ex) end | _ -> map put c in let c1 = put c0 in !ise, c1 @@ -541,7 +545,7 @@ let splay_app ise = | App (f, a') -> loop f (Array.append a' a) | Cast (c', _, _) -> loop c' a | Evar ex -> - (try loop (existential_value ise ex) a with _ -> c, a) + (try loop (existential_value0 ise ex) a with _ -> c, a) | _ -> c, a in fun c -> match kind c with | App (f, a) -> loop f a @@ -1255,7 +1259,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let fs sigma x = nf_evar sigma x in let pop_evar sigma e p = let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in - let e_body = match e_body with Evar_defined c -> c + let e_body = match e_body with Evar_defined c -> EConstr.Unsafe.to_constr c | _ -> errorstrm (str "Matching the pattern " ++ pr_constr_env env0 sigma0 p ++ str " did not instantiate ?" ++ int (Evar.repr e) ++ spc () ++ str "Does the variable bound by the \"in\" construct occur "++ diff --git a/pretyping/cases.ml b/pretyping/cases.ml index b16f1a9ed..4c87b4e7e 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -1015,7 +1015,7 @@ let adjust_impossible_cases pb pred tomatch submat = | Evar (evk,_) when snd (evar_source evk !(pb.evdref)) == Evar_kinds.ImpossibleCase -> if not (Evd.is_defined !(pb.evdref) evk) then begin let evd, default = use_unit_judge !(pb.evdref) in - pb.evdref := Evd.define evk (EConstr.Unsafe.to_constr default.uj_type) evd + pb.evdref := Evd.define evk default.uj_type evd end; add_assert_false_case pb tomatch | _ -> diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 04cb6a59f..6dc3687a0 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -251,7 +251,7 @@ and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr) let (n, dom, rng) = destLambda !evdref t in if isEvar !evdref dom then let (domk, args) = destEvar !evdref dom in - evdref := define domk (EConstr.Unsafe.to_constr a) !evdref; + evdref := define domk a !evdref; else (); t, rng | _ -> raise NoSubtacCoercion diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index d37090a65..144166a34 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -114,9 +114,6 @@ let flex_kind_of_term ts env evd c sk = | Fix _ -> Rigid (* happens when the fixpoint is partially applied *) | Cast _ | App _ | Case _ -> assert false -let add_conv_pb (pb, env, x, y) sigma = - Evd.add_conv_pb (pb, env, EConstr.Unsafe.to_constr x, EConstr.Unsafe.to_constr y) sigma - let apprec_nohdbeta ts env evd c = let (t,sk as appr) = Reductionops.whd_nored_state evd (c, []) in if Stack.not_purely_applicative sk @@ -1045,7 +1042,7 @@ let choose_less_dependent_instance evk evd term args = let subst' = List.filter (fun (id,c) -> EConstr.eq_constr evd c term) subst in match subst' with | [] -> None - | (id, _) :: _ -> Some (Evd.define evk (Constr.mkVar id) evd) + | (id, _) :: _ -> Some (Evd.define evk (mkVar id) evd) let apply_on_subterm env evdref f c t = let rec applyrec (env,(k,c) as acc) t = @@ -1085,7 +1082,7 @@ let filter_possible_projections evd c ty ctxt args = let a = Array.unsafe_get args i in (match decl with | NamedDecl.LocalAssum _ -> false - | NamedDecl.LocalDef (_,c,_) -> not (isRel evd (EConstr.of_constr c) || isVar evd (EConstr.of_constr c))) || + | NamedDecl.LocalDef (_,c,_) -> not (isRel evd c || isVar evd c)) || a == c || (* Here we make an approximation, for instance, we could also be *) (* interested in finding a term u convertible to c such that a occurs *) @@ -1135,7 +1132,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = end | decl'::ctxt', c::l, occs::occsl -> let id = NamedDecl.get_id decl' in - let t = EConstr.of_constr (NamedDecl.get_type decl') in + let t = NamedDecl.get_type decl' in let evs = ref [] in let ty = Retyping.get_type_of env_rhs evd c in let filter' = filter_possible_projections evd c ty ctxt args in @@ -1183,7 +1180,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = (* We force abstraction over this unconstrained occurrence *) (* and we use typing to propagate this instantiation *) (* This is an arbitrary choice *) - let evd = Evd.define evk (Constr.mkVar id) evd in + let evd = Evd.define evk (mkVar id) evd in match evar_conv_x ts env_evar evd CUMUL idty evty with | UnifFailure _ -> user_err Pp.(str "Cannot find an instance") | Success evd -> @@ -1205,14 +1202,11 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = (evar_conv_x full_transparent_state) with IllTypedInstance _ -> raise (TypingFailed evd) in - Evd.define evk (EConstr.Unsafe.to_constr rhs) evd + Evd.define evk rhs evd in abstract_free_holes evd subst, true with TypingFailed evd -> evd, false -let to_pb (pb, env, t1, t2) = - (pb, env, EConstr.Unsafe.to_constr t1, EConstr.Unsafe.to_constr t2) - let second_order_matching_with_args ts env evd pbty ev l t = (* let evd,ev = evar_absorb_arguments env evd ev l in @@ -1222,7 +1216,7 @@ let second_order_matching_with_args ts env evd pbty ev l t = else UnifFailure (evd, ConversionFailed (env,mkApp(mkEvar ev,l),t)) if b then Success evd else *) - let pb = to_pb (pbty,env,mkApp(mkEvar ev,l),t) in + let pb = (pbty,env,mkApp(mkEvar ev,l),t) in UnifFailure (evd, CannotSolveConstraint (pb,ProblemBeyondCapabilities)) let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = @@ -1245,7 +1239,7 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = | Some evd -> Success evd | None -> let reason = ProblemBeyondCapabilities in - UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason))) + UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason))) | (Rel _|Var _), Evar (evk2,args2) when app_empty && List.for_all (fun a -> EConstr.eq_constr evd a term1 || isEvar evd a) (remove_instance_local_defs evd evk2 args2) -> @@ -1255,7 +1249,7 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = | Some evd -> Success evd | None -> let reason = ProblemBeyondCapabilities in - UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason))) + UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason))) | Evar (evk1,args1), Evar (evk2,args2) when Evar.equal evk1 evk2 -> let f env evd pbty x y = is_fconv ~reds:ts pbty env evd x y in Success (solve_refl ~can_drop:true f env evd @@ -1295,10 +1289,10 @@ let error_cannot_unify env evd pb ?reason t1 t2 = let check_problems_are_solved env evd = match snd (extract_all_conv_pbs evd) with - | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb (EConstr.of_constr t1) (EConstr.of_constr t2) + | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb t1 t2 | _ -> () -exception MaxUndefined of (Evar.t * evar_info * Constr.t list) +exception MaxUndefined of (Evar.t * evar_info * EConstr.t list) let max_undefined_with_candidates evd = let fold evk evi () = match evi.evar_candidates with @@ -1326,7 +1320,7 @@ let rec solve_unconstrained_evars_with_candidates ts evd = | a::l -> try let conv_algo = evar_conv_x ts in - let evd = check_evar_instance evd evk (EConstr.of_constr a) conv_algo in + let evd = check_evar_instance evd evk a conv_algo in let evd = Evd.define evk a evd in match reconsider_unif_constraints conv_algo evd with | Success evd -> solve_unconstrained_evars_with_candidates ts evd @@ -1348,7 +1342,7 @@ let solve_unconstrained_impossible_cases env evd = let ty = j_type j in let conv_algo = evar_conv_x full_transparent_state in let evd' = check_evar_instance evd' evk ty conv_algo in - Evd.define evk (EConstr.Unsafe.to_constr ty) evd' + Evd.define evk ty evd' | _ -> evd') evd evd let solve_unif_constraints_with_heuristics env @@ -1357,8 +1351,6 @@ let solve_unif_constraints_with_heuristics env let rec aux evd pbs progress stuck = match pbs with | (pbty,env,t1,t2 as pb) :: pbs -> - let t1 = EConstr.of_constr t1 in - let t2 = EConstr.of_constr t2 in (match apply_conversion_problem_heuristic ts env evd pbty t1 t2 with | Success evd' -> let (evd', rest) = extract_all_conv_pbs evd' in @@ -1375,9 +1367,7 @@ let solve_unif_constraints_with_heuristics env match stuck with | [] -> (* We're finished *) evd | (pbty,env,t1,t2 as pb) :: _ -> - let t1 = EConstr.of_constr t1 in - let t2 = EConstr.of_constr t2 in - (* There remains stuck problems *) + (* There remains stuck problems *) error_cannot_unify env evd pb t1 t2 in let (evd,pbs) = extract_all_conv_pbs evd in diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 4cffbbb83..b452755b1 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -77,7 +77,7 @@ let define_pure_evar_as_product evd evk = let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in let id = next_ident_away idx (Environ.ids_of_named_context_val evi.evar_hyps) in - let concl = Reductionops.whd_all evenv evd (EConstr.of_constr evi.evar_concl) in + let concl = Reductionops.whd_all evenv evd evi.evar_concl in let s = destSort evd concl in let evksrc = evar_source evk evd in let src = subterm_source evk ~where:Domain evksrc in @@ -101,7 +101,7 @@ let define_pure_evar_as_product evd evk = evd3, rng in let prod = mkProd (Name id, dom, subst_var id rng) in - let evd3 = Evd.define evk (EConstr.Unsafe.to_constr prod) evd2 in + let evd3 = Evd.define evk prod evd2 in evd3,prod (* Refine an applied evar to a product and returns its instantiation *) @@ -128,7 +128,7 @@ let define_pure_evar_as_lambda env evd evk = let open Context.Named.Declaration in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in - let typ = Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi)) in + let typ = Reductionops.whd_all evenv evd (evar_concl evi) in let evd1,(na,dom,rng) = match EConstr.kind evd typ with | Prod (na,dom,rng) -> (evd,(na,dom,rng)) | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd evd typ @@ -141,7 +141,7 @@ let define_pure_evar_as_lambda env evd evk = let src = subterm_source evk ~where:Body (evar_source evk evd1) in let evd2,body = new_evar newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in let lam = mkLambda (Name id, dom, subst_var id body) in - Evd.define evk (EConstr.Unsafe.to_constr lam) evd2, lam + Evd.define evk lam evd2, lam let define_evar_as_lambda env evd (evk,args) = let evd,lam = define_pure_evar_as_lambda env evd evk in @@ -166,9 +166,9 @@ let define_evar_as_sort env evd (ev,args) = let evd, u = new_univ_variable univ_rigid evd in let evi = Evd.find_undefined evd ev in let s = Type u in - let concl = Reductionops.whd_all (evar_env evi) evd (EConstr.of_constr evi.evar_concl) in + let concl = Reductionops.whd_all (evar_env evi) evd evi.evar_concl in let sort = destSort evd concl in - let evd' = Evd.define ev (Constr.mkSort s) evd in + let evd' = Evd.define ev (mkSort s) evd in Evd.set_leq_sort env evd' (Type (Univ.super u)) (ESorts.kind evd' sort), s (* Propagation of constraints through application and abstraction: diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 0c109b026..b7eaff078 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -89,9 +89,9 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) Array.iter (refresh_term_evars onevars false) args | Evar (ev, a) when onevars -> let evi = Evd.find !evdref ev in - let ty' = refresh ~onlyalg univ_flexible ~direction:true (EConstr.of_constr evi.evar_concl) in + let ty' = refresh ~onlyalg univ_flexible ~direction:true evi.evar_concl in if !modified then - evdref := Evd.add !evdref ev {evi with evar_concl = EConstr.Unsafe.to_constr ty'} + evdref := Evd.add !evdref ev {evi with evar_concl = ty'} else () | _ -> EConstr.iter !evdref (refresh_term_evars onevars false) t and refresh_polymorphic_positions args pos = @@ -137,8 +137,6 @@ let test_success conv_algo env evd c c' rhs = is_success (conv_algo env evd c c' rhs) let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd = - let t1 = EConstr.Unsafe.to_constr t1 in - let t2 = EConstr.Unsafe.to_constr t2 in match pbty with | Some true -> add_conv_pb ~tail (Reduction.CUMUL,env,t1,t2) evd | Some false -> add_conv_pb ~tail (Reduction.CUMUL,env,t2,t1) evd @@ -197,7 +195,7 @@ let restrict_evar_key evd evk filter candidates = | None -> evar_filter evi | Some filter -> filter in let candidates = match candidates with - | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates + | NoUpdate -> evi.evar_candidates | UpdateWith c -> Some c in restrict_evar evd evk filter candidates end @@ -600,7 +598,6 @@ let solve_pattern_eqn env sigma l c = let make_projectable_subst aliases sigma evi args = let sign = evar_filtered_context evi in - let sign = List.map (fun d -> map_named_decl EConstr.of_constr d) sign in let evar_aliases = compute_var_aliases sign sigma in let (_,full_subst,cstr_subst) = List.fold_right @@ -877,7 +874,7 @@ let choose_projection evi sols = let rec do_projection_effects define_fun env ty evd = function | ProjectVar -> evd | ProjectEvar ((evk,argsv),evi,id,p) -> - let evd = Evd.define evk (Constr.mkVar id) evd in + let evd = Evd.define evk (mkVar id) evd in (* TODO: simplify constraints involving evk *) let evd = do_projection_effects define_fun env ty evd p in let ty = whd_all env evd (Lazy.force ty) in @@ -887,7 +884,7 @@ let rec do_projection_effects define_fun env ty evd = function one (however, regarding coercions, because t is obtained by unif, we know that no coercion can be inserted) *) let subst = make_pure_subst evi argsv in - let ty' = replace_vars subst (EConstr.of_constr evi.evar_concl) in + let ty' = replace_vars subst evi.evar_concl in if isEvar evd ty' then define_fun env evd (Some false) (destEvar evd ty') ty else evd else evd @@ -1004,7 +1001,7 @@ let filter_effective_candidates evd evi filter candidates = let filter_candidates evd evk filter candidates_update = let evi = Evd.find_undefined evd evk in let candidates = match candidates_update with - | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates + | NoUpdate -> evi.evar_candidates | UpdateWith c -> Some c in match candidates with @@ -1023,13 +1020,12 @@ let closure_of_filter evd evk = function | None -> None | Some filter -> let evi = Evd.find_undefined evd evk in - let vars = collect_vars evd (EConstr.of_constr (evar_concl evi)) in + let vars = collect_vars evd (evar_concl evi) in let test b decl = b || Id.Set.mem (get_id decl) vars || match decl with | LocalAssum _ -> false | LocalDef (_,c,_) -> - let c = EConstr.of_constr c in not (isRel evd c || isVar evd c) in let newfilter = Filter.map_along test filter (evar_context evi) in @@ -1062,7 +1058,7 @@ let do_restrict_hyps evd (evk,args as ev) filter candidates = match candidates,filter with | UpdateWith [], _ -> user_err Pp.(str "Not solvable.") | UpdateWith [nc],_ -> - let evd = Evd.define evk (EConstr.Unsafe.to_constr nc) evd in + let evd = Evd.define evk nc evd in raise (EvarSolvedWhileRestricting (evd,mkEvar ev)) | NoUpdate, None -> evd,ev | _ -> restrict_applied_evar evd ev filter candidates @@ -1113,9 +1109,6 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs = * Note: argument f is the function used to instantiate evars. *) -let instantiate_evar_array evi c args = - EConstr.of_constr (instantiate_evar_array evi (EConstr.Unsafe.to_constr c) (Array.map EConstr.Unsafe.to_constr args)) - let filter_compatible_candidates conv_algo env evd evi args rhs c = let c' = instantiate_evar_array evi c args in match conv_algo env evd Reduction.CONV rhs c' with @@ -1135,8 +1128,6 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) = | _, None -> filter_candidates evd evk1 filter1 NoUpdate | None, Some _ -> raise DoesNotPreserveCandidateRestriction | Some l1, Some l2 -> - let l1 = List.map EConstr.of_constr l1 in - let l2 = List.map EConstr.of_constr l2 in let l1 = filter_effective_candidates evd evi1 filter1 l1 in let l1' = List.filter (fun c1 -> let c1' = instantiate_evar_array evi1 c1 argsv1 in @@ -1242,9 +1233,9 @@ let check_evar_instance evd evk1 body conv_algo = try Retyping.get_type_of ~lax:true evenv evd body with Retyping.RetypeError _ -> user_err Pp.(str "Ill-typed evar instance") in - match conv_algo evenv evd Reduction.CUMUL ty (EConstr.of_constr evi.evar_concl) with + match conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl with | Success evd -> evd - | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,EConstr.of_constr evi.evar_concl)) + | UnifFailure _ -> raise (IllTypedInstance (evenv,ty, evi.evar_concl)) let update_evar_source ev1 ev2 evd = let loc, evs2 = evar_source ev2 evd in @@ -1257,7 +1248,7 @@ let update_evar_source ev1 ev2 evd = let solve_evar_evar_l2r force f g env evd aliases pbty ev1 (evk2,_ as ev2) = try let evd,body = project_evar_on_evar force g env evd aliases 0 pbty ev1 ev2 in - let evd' = Evd.define evk2 (EConstr.Unsafe.to_constr body) evd in + let evd' = Evd.define evk2 body evd in let evd' = update_evar_source (fst (destEvar evd body)) evk2 evd' in check_evar_instance evd' evk2 body g with EvarSolvedOnTheFly (evd,c) -> @@ -1292,17 +1283,19 @@ let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 a let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) = let pbty = if force then None else pbty in let evi = Evd.find evd evk1 in - let downcast evk t evd = downcast evk (EConstr.Unsafe.to_constr t) evd in + let downcast evk t evd = downcast evk t evd in let evd = try (* ?X : Π Δ. Type i = ?Y : Π Δ'. Type j. The body of ?X and ?Y just has to be of type Π Δ. Type k for some k <= i, j. *) let evienv = Evd.evar_env evi in - let ctx1, i = Reduction.dest_arity evienv evi.evar_concl in + let concl1 = EConstr.Unsafe.to_constr evi.evar_concl in + let ctx1, i = Reduction.dest_arity evienv concl1 in let ctx1 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx1 in let evi2 = Evd.find evd evk2 in let evi2env = Evd.evar_env evi2 in - let ctx2, j = Reduction.dest_arity evi2env evi2.evar_concl in + let concl2 = EConstr.Unsafe.to_constr evi2.evar_concl in + let ctx2, j = Reduction.dest_arity evi2env concl2 in let ctx2 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx2 in let ui, uj = univ_of_sort i, univ_of_sort j in if i == j || Evd.check_eq evd ui uj @@ -1375,14 +1368,14 @@ let solve_candidates conv_algo env evd (evk,argsv) rhs = | Some l -> let l' = List.map_filter - (fun c -> filter_compatible_candidates conv_algo env evd evi argsv rhs (EConstr.of_constr c)) l in + (fun c -> filter_compatible_candidates conv_algo env evd evi argsv rhs c) l in match l' with | [] -> raise IncompatibleCandidates | [c,evd] -> (* solve_candidates might have been called recursively in the mean *) (* time and the evar been solved by the filtering process *) if Evd.is_undefined evd evk then - let evd' = Evd.define evk (EConstr.Unsafe.to_constr c) evd in + let evd' = Evd.define evk c evd in check_evar_instance evd' evk c conv_algo else evd | l when List.length l < List.length l' -> @@ -1401,8 +1394,8 @@ let occur_evar_upto_types sigma n c = Array.iter occur_rec args else ( seen := Evar.Set.add sp !seen; - Option.iter occur_rec (existential_opt_value sigma e); - occur_rec (Evd.existential_type sigma e)) + Option.iter occur_rec (existential_opt_value0 sigma e); + occur_rec (Evd.existential_type0 sigma e)) | _ -> Constr.iter occur_rec c in try occur_rec c; false with Occur -> true @@ -1529,7 +1522,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = (* Try to project (a restriction of) the left evar ... *) try let evd,body = project_evar_on_evar false conv_algo env' evd aliases 0 None ev'' ev' in - let evd = Evd.define evk' (EConstr.Unsafe.to_constr body) evd in + let evd = Evd.define evk' body evd in check_evar_instance evd evk' body conv_algo with | EvarSolvedOnTheFly _ -> assert false (* ev has no candidates *) @@ -1644,7 +1637,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = print_constr body); raise e in*) let evd' = check_evar_instance evd' evk body conv_algo in - Evd.define evk (EConstr.Unsafe.to_constr body) evd' + Evd.define evk body evd' with | NotEnoughInformationToProgress sols -> postpone_non_unique_projection env evd pbty ev sols rhs @@ -1691,8 +1684,6 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = *) let status_changed evd lev (pbty,_,t1,t2) = - let t1 = EConstr.of_constr t1 in - let t2 = EConstr.of_constr t2 in (try Evar.Set.mem (head_evar evd t1) lev with NoHeadEvar -> false) || (try Evar.Set.mem (head_evar evd t2) lev with NoHeadEvar -> false) @@ -1702,7 +1693,7 @@ let reconsider_unif_constraints conv_algo evd = (fun p (pbty,env,t1,t2 as x) -> match p with | Success evd -> - (match conv_algo env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) with + (match conv_algo env evd pbty t1 t2 with | Success _ as x -> x | UnifFailure (i,e) -> UnifFailure (i,CannotSolveConstraint (x,e))) | UnifFailure _ as x -> x) diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index fcbf50fea..85911394f 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -401,9 +401,9 @@ and nf_evar env sigma evk ty args = mkEvar (evk, Array.of_list args), ty let evars_of_evar_map sigma = - { Nativelambda.evars_val = Evd.existential_opt_value sigma; - Nativelambda.evars_typ = Evd.existential_type sigma; - Nativelambda.evars_metas = Evd.meta_type sigma } + { Nativelambda.evars_val = Evd.existential_opt_value0 sigma; + Nativelambda.evars_typ = Evd.existential_type0 sigma; + Nativelambda.evars_metas = Evd.meta_type0 sigma } (* fork perf process, return profiler's process id *) let start_profiler_linux profile_fn = diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index e52112fda..27e457e3b 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -184,7 +184,7 @@ let pattern_of_constr env sigma t = | Evar_kinds.GoalEvar | Evar_kinds.VarInstance _ -> (* These are the two evar kinds used for existing goals *) (* see Proofview.mark_in_evm *) - if Evd.is_defined sigma evk then pattern_of_constr env (Evd.existential_value sigma ev) + if Evd.is_defined sigma evk then pattern_of_constr env (Evd.existential_value0 sigma ev) else PEvar (evk,Array.map (pattern_of_constr env) ctxt) | Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar ido) -> assert false | _ -> diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 2c371d5cf..947469ca0 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -315,7 +315,7 @@ let apply_inference_hook hook evdref frozen = match frozen with then try let sigma, c = hook sigma evk in - Evd.define evk (EConstr.Unsafe.to_constr c) sigma + Evd.define evk c sigma with Exit -> sigma else @@ -532,7 +532,7 @@ let pretype_global ?loc rigid env evd gr us = interp_instance ?loc evd ~len l in let (sigma, c) = Evd.fresh_global ?loc ~rigid ?names:instance env.ExtraEnv.env evd gr in - (sigma, EConstr.of_constr c) + (sigma, c) let pretype_ref ?loc evdref env ref us = match ref with @@ -1109,7 +1109,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = let f decl (subst,update) = let id = NamedDecl.get_id decl in - let t = replace_vars subst (EConstr.of_constr (NamedDecl.get_type decl)) in + let t = replace_vars subst (NamedDecl.get_type decl) in let c, update = try let c = List.assoc id update in @@ -1150,7 +1150,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar c = match D (* Correction of bug #5315 : we need to define an evar for *all* holes *) let evkt = e_new_evar env evdref ~src:(loc, knd) ~naming (mkSort s) in let ev,_ = destEvar !evdref evkt in - evdref := Evd.define ev (to_constr ~abort_on_undefined_evars:false !evdref v) !evdref; + evdref := Evd.define ev (nf_evar !evdref v) !evdref; (* End of correction of bug #5315 *) { utj_val = v; utj_type = s } diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 360c6e86e..244b8e60b 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -871,7 +871,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Evar ev -> fold () | Meta ev -> (match safe_meta_value sigma ev with - | Some body -> whrec cst_l (EConstr.of_constr body, stack) + | Some body -> whrec cst_l (body, stack) | None -> fold ()) | Const (c,u as const) -> reduction_effect_hook env sigma (EConstr.to_constr sigma x) @@ -1106,7 +1106,7 @@ let local_whd_state_gen flags sigma = | Evar ev -> s | Meta ev -> (match safe_meta_value sigma ev with - Some c -> whrec (EConstr.of_constr c,stack) + Some c -> whrec (c,stack) | None -> s) | Construct ((ind,c),u) -> @@ -1392,7 +1392,7 @@ let vm_infer_conv ?(pb=Reduction.CUMUL) env t1 t2 = (********************************************************************) let whd_meta sigma c = match EConstr.kind sigma c with - | Meta p -> (try EConstr.of_constr (meta_value sigma p) with Not_found -> c) + | Meta p -> (try meta_value sigma p with Not_found -> c) | _ -> c let default_plain_instance_ident = Id.of_string "H" @@ -1612,7 +1612,7 @@ let meta_value evd mv = match meta_opt_fvalue evd mv with | Some (b,_) -> let metas = Metamap.bind valrec b.freemetas in - instance evd metas (EConstr.of_constr b.rebus) + instance evd metas b.rebus | None -> mkMeta mv in valrec mv @@ -1625,9 +1625,8 @@ let meta_instance sigma b = instance sigma c_sigma b.rebus let nf_meta sigma c = - let c = EConstr.Unsafe.to_constr c in let cl = mk_freelisted c in - meta_instance sigma { cl with rebus = EConstr.of_constr cl.rebus } + meta_instance sigma { cl with rebus = cl.rebus } (* Instantiate metas that create beta/iota redexes *) @@ -1648,7 +1647,6 @@ let meta_reducible_instance evd b = (match try let g, s = Metamap.find m metas in - let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in if isConstruct evd g || not is_coerce then Some g else None with Not_found -> None @@ -1660,7 +1658,6 @@ let meta_reducible_instance evd b = (match try let g, s = Metamap.find m metas in - let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in if isLambda evd g || not is_coerce then Some g else None with Not_found -> None @@ -1669,7 +1666,6 @@ let meta_reducible_instance evd b = | None -> mkApp (f,Array.map irec l)) | Meta m -> (try let g, s = Metamap.find m metas in - let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in if not is_coerce then irec g else u with Not_found -> u) @@ -1678,7 +1674,6 @@ let meta_reducible_instance evd b = (match try let g, s = Metamap.find m metas in - let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in if isConstruct evd g || not is_coerce then Some g else None with Not_found -> None diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 95aa5ebef..746a68b21 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -57,8 +57,8 @@ let get_type_from_constraints env sigma t = if isEvar sigma (fst (decompose_app_vect sigma t)) then match List.map_filter (fun (pbty,env,t1,t2) -> - if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t1) then Some t2 - else if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t2) then Some t1 + if is_fconv Reduction.CONV env sigma t t1 then Some t2 + else if is_fconv Reduction.CONV env sigma t t2 then Some t1 else None) (snd (Evd.extract_all_conv_pbs sigma)) with @@ -99,7 +99,7 @@ let retype ?(polyprop=true) sigma = let rec type_of env cstr = match EConstr.kind sigma cstr with | Meta n -> - (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus) + (try strip_outer_cast sigma (Evd.meta_ftype sigma n).Evd.rebus with Not_found -> retype_error (BadMeta n)) | Rel n -> let ty = RelDecl.get_type (lookup_rel n env) in @@ -115,7 +115,7 @@ let retype ?(polyprop=true) sigma = try Inductiveops.find_rectype env sigma t with Not_found -> try - let t = EConstr.of_constr (get_type_from_constraints env sigma t) in + let t = get_type_from_constraints env sigma t in Inductiveops.find_rectype env sigma t with Not_found -> retype_error BadRecursiveType in diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 977713fd8..696001ab7 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -416,7 +416,7 @@ exception Partial reduction is solved by the expanded fix term. *) let solve_arity_problem env sigma fxminargs c = let evm = ref sigma in - let set_fix i = evm := Evd.define i (Constr.mkVar vfx) !evm in + let set_fix i = evm := Evd.define i (mkVar vfx) !evm in let rec check strict c = let c' = whd_betaiotazeta sigma c in let (h,rcargs) = decompose_app_vect sigma c' in @@ -641,7 +641,7 @@ let whd_nothing_for_iota env sigma s = | _ -> s) | Evar ev -> s | Meta ev -> - (try whrec (EConstr.of_constr (Evd.meta_value sigma ev), stack) + (try whrec (Evd.meta_value sigma ev, stack) with Not_found -> s) | Const (const, u) when is_transparent_constant full_transparent_state const -> let u = EInstance.kind sigma u in diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 08051fd3a..30ddeffa6 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -158,7 +158,7 @@ let rec is_class_type evd c = | _ -> is_class_constr evd c let is_class_evar evd evi = - is_class_type evd (EConstr.of_constr evi.Evd.evar_concl) + is_class_type evd evi.Evd.evar_concl (* * classes persistent object diff --git a/pretyping/typing.ml b/pretyping/typing.ml index c4eb6af89..aaec73f04 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -29,7 +29,6 @@ let meta_type evd mv = let ty = try Evd.meta_ftype evd mv with Not_found -> anomaly (str "unknown meta ?" ++ str (Nameops.string_of_meta mv) ++ str ".") in - let ty = Evd.map_fl EConstr.of_constr ty in meta_instance evd ty let inductive_type_knowing_parameters env sigma (ind,u) jl = @@ -129,7 +128,7 @@ let e_is_correct_arity env evdref c pj ind specif params = then error () | Evar (ev,_), [] -> let evd, s = Evd.fresh_sort_in_family env !evdref (max_sort allowed_sorts) in - evdref := Evd.define ev (Constr.mkSort s) evd + evdref := Evd.define ev (mkSort s) evd | _, (LocalDef _ as d)::ar' -> srec (push_rel d env) (lift 1 pt') ar' | _ -> diff --git a/pretyping/unification.ml b/pretyping/unification.ml index ca03b96ab..d98ce9aba 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -84,7 +84,7 @@ let occur_meta_or_undefined_evar evd c = | Evar (ev,args) -> (match evar_body (Evd.find evd ev) with | Evar_defined c -> - occrec c; Array.iter occrec args + occrec (EConstr.Unsafe.to_constr c); Array.iter occrec args | Evar_empty -> raise Occur) | _ -> Constr.iter occrec c in try occrec c; false with Occur | Not_found -> true @@ -189,10 +189,9 @@ let pose_all_metas_as_evars env evd t = let rec aux t = match EConstr.kind !evdref t with | Meta mv -> (match Evd.meta_opt_fvalue !evdref mv with - | Some ({rebus=c},_) -> EConstr.of_constr c + | Some ({rebus=c},_) -> c | None -> let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in - let ty = EConstr.of_constr ty in let ty = if Evd.Metaset.is_empty mvs then ty else aux ty in let ty = if Flags.version_strictly_greater Flags.V8_6 @@ -200,7 +199,7 @@ let pose_all_metas_as_evars env evd t = else ty (* some beta-iota-normalization "regression" in 8.5 and 8.6 *) in let src = Evd.evar_source_of_meta mv !evdref in let ev = Evarutil.e_new_evar env evdref ~src ty in - evdref := meta_assign mv (EConstr.Unsafe.to_constr ev,(Conv,TypeNotProcessed)) !evdref; + evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref; ev) | _ -> EConstr.map !evdref aux t in @@ -1060,7 +1059,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e (evd,t2::ks, m-1) else let mv = new_meta () in - let evd' = meta_declare mv (EConstr.Unsafe.to_constr (substl ks b)) evd in + let evd' = meta_declare mv (substl ks b) evd in (evd', mkMeta mv :: ks, m - 1)) (sigma,[],List.length bs) bs in @@ -1247,7 +1246,7 @@ let try_to_coerce env evd c cty tycon = let j = make_judge c cty in let (evd',j') = inh_conv_coerce_rigid_to true env evd j tycon in let evd' = Evarconv.solve_unif_constraints_with_heuristics env evd' in - let evd' = Evd.map_metas_fvalue (fun c -> EConstr.Unsafe.to_constr (nf_evar evd' (EConstr.of_constr c))) evd' in + let evd' = Evd.map_metas_fvalue (fun c -> nf_evar evd' c) evd' in (evd',j'.uj_val) let w_coerce_to_type env evd c cty mvty = @@ -1359,11 +1358,11 @@ let w_merge env with_types flags (evd,metas,evars : subst0) = if meta_defined evd mv then let {rebus=c'},(status',_) = meta_fvalue evd mv in let (take_left,st,(evd,metas',evars')) = - merge_instances env evd flags status' status (EConstr.of_constr c') c + merge_instances env evd flags status' status c' c in let evd' = if take_left then evd - else meta_reassign mv (EConstr.Unsafe.to_constr c,(st,TypeProcessed)) evd + else meta_reassign mv (c,(st,TypeProcessed)) evd in w_merge_rec evd' (metas'@metas@metas'') (evars'@evars'') eqns else @@ -1372,7 +1371,7 @@ let w_merge env with_types flags (evd,metas,evars : subst0) = if isMetaOf evd mv (whd_all env evd c) then evd else error_cannot_unify env evd (mkMeta mv,c) else - meta_assign mv (EConstr.Unsafe.to_constr c,(status,TypeProcessed)) evd in + meta_assign mv (c,(status,TypeProcessed)) evd in w_merge_rec evd' (metas''@metas) evars'' eqns | [] -> (* Process type eqns *) @@ -1396,17 +1395,17 @@ let w_merge env with_types flags (evd,metas,evars : subst0) = let (evd', c) = applyHead sp_env evd nargs hdc in let (evd'',mc,ec) = unify_0 sp_env evd' CUMUL flags - (get_type_of sp_env evd' c) (EConstr.of_constr ev.evar_concl) in + (get_type_of sp_env evd' c) ev.evar_concl in let evd''' = w_merge_rec evd'' mc ec [] in if evd' == evd''' - then Evd.define sp (EConstr.Unsafe.to_constr c) evd''' - else Evd.define sp (EConstr.Unsafe.to_constr (Evarutil.nf_evar evd''' c)) evd''' in + then Evd.define sp c evd''' + else Evd.define sp (Evarutil.nf_evar evd''' c) evd''' in let check_types evd = let metas = Evd.meta_list evd in let eqns = List.fold_left (fun acc (mv, b) -> match b with - | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, EConstr.of_constr t.rebus) :: acc + | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, t.rebus) :: acc | _ -> acc) [] metas in w_merge_rec evd [] [] eqns in @@ -1417,11 +1416,6 @@ let w_merge env with_types flags (evd,metas,evars : subst0) = in if with_types then check_types res else res -let retract_coercible_metas evd = - let (metas, evd) = retract_coercible_metas evd in - let map (mv, c, st) = (mv, EConstr.of_constr c, st) in - (List.map map metas, evd) - let w_unify_meta_types env ?(flags=default_unify_flags ()) evd = let metas,evd = retract_coercible_metas evd in w_merge env true flags.merge_unify_flags (evd,metas,[]) diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index f314ae0d6..049c3aff5 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -205,7 +205,7 @@ and nf_univ_args ~nb_univs mk env sigma stk = and nf_evar env sigma evk stk = let evi = try Evd.find sigma evk with Not_found -> assert false in let hyps = Environ.named_context_of_val (Evd.evar_filtered_hyps evi) in - let concl = Evd.evar_concl evi in + let concl = EConstr.Unsafe.to_constr @@ Evd.evar_concl evi in if List.is_empty hyps then nf_stk env sigma (mkEvar (evk, [||])) concl stk else match stk with diff --git a/printing/printer.ml b/printing/printer.ml index 199aa79c6..edcce874d 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -93,13 +93,13 @@ let _ = Hook.set Refine.pr_constr pr_constr_env let pr_lconstr_goal_style_env env sigma c = pr_leconstr_core true env sigma (EConstr.of_constr c) let pr_constr_goal_style_env env sigma c = pr_econstr_core true env sigma (EConstr.of_constr c) -let pr_open_lconstr_env env sigma (_,c) = pr_lconstr_env env sigma c -let pr_open_constr_env env sigma (_,c) = pr_constr_env env sigma c - let pr_econstr_n_env env sigma c = pr_econstr_n_core false env sigma c let pr_leconstr_env env sigma c = pr_leconstr_core false env sigma c let pr_econstr_env env sigma c = pr_econstr_core false env sigma c +let pr_open_lconstr_env env sigma (_,c) = pr_leconstr_env env sigma c +let pr_open_constr_env env sigma (_,c) = pr_econstr_env env sigma c + (* NB do not remove the eta-redexes! Global.env() has side-effects... *) let pr_lconstr t = let (sigma, env) = Pfedit.get_current_context () in @@ -108,12 +108,12 @@ let pr_constr t = let (sigma, env) = Pfedit.get_current_context () in pr_constr_env env sigma t -let pr_open_lconstr (_,c) = pr_lconstr c -let pr_open_constr (_,c) = pr_constr c - let pr_leconstr c = pr_lconstr (EConstr.Unsafe.to_constr c) let pr_econstr c = pr_constr (EConstr.Unsafe.to_constr c) +let pr_open_lconstr (_,c) = pr_leconstr c +let pr_open_constr (_,c) = pr_econstr c + let pr_constr_under_binders_env_gen pr env sigma (ids,c) = (* Warning: clashes can occur with variables of same name in env but *) (* we also need to preserve the actual names of the patterns *) @@ -541,12 +541,12 @@ let pr_evgl_sign sigma evi = if List.is_empty ids then mt () else (str " (" ++ prlist_with_sep pr_comma pr_id ids ++ str " cannot be used)") in - let pc = pr_lconstr_env env sigma evi.evar_concl in + let pc = pr_leconstr_env env sigma evi.evar_concl in let candidates = match evi.evar_body, evi.evar_candidates with | Evar_empty, Some l -> spc () ++ str "= {" ++ - prlist_with_sep (fun () -> str "|") (pr_lconstr_env env sigma) l ++ str "}" + prlist_with_sep (fun () -> str "|") (pr_leconstr_env env sigma) l ++ str "}" | _ -> mt () in @@ -622,8 +622,8 @@ let print_evar_constraints gl sigma = end in let pr_evconstr (pbty,env,t1,t2) = - let t1 = Evarutil.nf_evar sigma (EConstr.of_constr t1) - and t2 = Evarutil.nf_evar sigma (EConstr.of_constr t2) in + let t1 = Evarutil.nf_evar sigma t1 + and t2 = Evarutil.nf_evar sigma t2 in let env = (** We currently allow evar instances to refer to anonymous de Bruijn indices, so we protect the error printing code in this case by giving diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 03ff580ad..aeaf16723 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -62,9 +62,6 @@ let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) c exception NotExtensibleClause -let mk_freelisted c = - map_fl EConstr.of_constr (mk_freelisted (EConstr.Unsafe.to_constr c)) - let clenv_push_prod cl = let typ = whd_all (cl_env cl) (cl_sigma cl) (clenv_type cl) in let rec clrec typ = match EConstr.kind cl.evd typ with @@ -73,7 +70,7 @@ let clenv_push_prod cl = let mv = new_meta () in let dep = not (noccurn (cl_sigma cl) 1 u) in let na' = if dep then na else Anonymous in - let e' = meta_declare mv (EConstr.Unsafe.to_constr t) ~name:na' cl.evd in + let e' = meta_declare mv t ~name:na' cl.evd in let concl = if dep then subst1 (mkMeta mv) u else u in let def = applist (cl.templval.rebus,[mkMeta mv]) in { templval = mk_freelisted def; @@ -107,8 +104,7 @@ let clenv_environments evd bound t = let mv = new_meta () in let dep = not (noccurn evd 1 t2) in let na' = if dep then na else Anonymous in - let t1 = EConstr.Unsafe.to_constr t1 in - let e' = meta_declare mv t1 ~name:na' e in + let e' = meta_declare mv t1 ~name:na' e in clrec (e', (mkMeta mv)::metas) (Option.map ((+) (-1)) n) (if dep then (subst1 (mkMeta mv) t2) else t2) | (n, LetIn (na,b,_,t)) -> clrec (e,metas) n (subst1 b t) @@ -167,13 +163,13 @@ let clenv_assign mv rhs clenv = user_err Pp.(str "clenv_assign: circularity in unification"); try if meta_defined clenv.evd mv then - if not (EConstr.eq_constr clenv.evd (EConstr.of_constr (fst (meta_fvalue clenv.evd mv)).rebus) rhs) then + if not (EConstr.eq_constr clenv.evd (fst (meta_fvalue clenv.evd mv)).rebus rhs) then error_incompatible_inst clenv mv else clenv else let st = (Conv,TypeNotProcessed) in - {clenv with evd = meta_assign mv (EConstr.Unsafe.to_constr rhs_fls.rebus,st) clenv.evd} + {clenv with evd = meta_assign mv (rhs_fls.rebus,st) clenv.evd} with Not_found -> user_err Pp.(str "clenv_assign: undefined meta") @@ -218,7 +214,7 @@ let clenv_assign mv rhs clenv = *) let clenv_metas_in_type_of_meta evd mv = - (mk_freelisted (meta_instance evd (map_fl EConstr.of_constr (meta_ftype evd mv)))).freemetas + (mk_freelisted (meta_instance evd (meta_ftype evd mv))).freemetas let dependent_in_type_of_metas clenv mvs = List.fold_right @@ -288,11 +284,11 @@ let adjust_meta_source evd mv = function in situations like "ex_intro (fun x => P) ?ev p" *) let f = function (mv',(Cltyp (_,t) | Clval (_,_,t))) -> if Metaset.mem mv t.freemetas then - let f,l = decompose_app evd (EConstr.of_constr t.rebus) in + let f,l = decompose_app evd t.rebus in match EConstr.kind evd f with | Meta mv'' -> (match meta_opt_fvalue evd mv'' with - | Some (c,_) -> match_name (EConstr.of_constr c.rebus) l + | Some (c,_) -> match_name c.rebus l | None -> None) | _ -> None else None in @@ -502,7 +498,6 @@ let clenv_assign_binding clenv k c = let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in let c_typ = nf_betaiota clenv.env clenv.evd (clenv_get_type_of clenv c) in let status,clenv',c = clenv_unify_binding_type clenv c c_typ k_typ in - let c = EConstr.Unsafe.to_constr c in { clenv' with evd = meta_assign k (c,(Conv,status)) clenv'.evd } let clenv_match_args bl clenv = @@ -515,7 +510,7 @@ let clenv_match_args bl clenv = (fun clenv {CAst.loc;v=(b,c)} -> let k = meta_of_binder clenv loc mvs b in if meta_defined clenv.evd k then - if EConstr.eq_constr clenv.evd (EConstr.of_constr (fst (meta_fvalue clenv.evd k)).rebus) c then clenv + if EConstr.eq_constr clenv.evd (fst (meta_fvalue clenv.evd k)).rebus c then clenv else error_already_defined b else clenv_assign_binding clenv k c) @@ -677,7 +672,7 @@ let define_with_type sigma env ev c = let j = Environ.make_judge c ty in let (sigma, j) = Coercion.inh_conv_coerce_to true env sigma j t in let (ev, _) = destEvar sigma ev in - let sigma = Evd.define ev (EConstr.Unsafe.to_constr j.Environ.uj_val) sigma in + let sigma = Evd.define ev j.Environ.uj_val sigma in sigma let solve_evar_clause env sigma hyp_only clause = function diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index 0d197c92c..c80f370fd 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -25,8 +25,6 @@ open Ltac_pretype type glob_constr_ltac_closure = ltac_var_map * glob_constr let depends_on_evar sigma evk _ (pbty,_,t1,t2) = - let t1 = EConstr.of_constr t1 in - let t2 = EConstr.of_constr t2 in try Evar.equal (head_evar sigma t1) evk with NoHeadEvar -> try Evar.equal (head_evar sigma t2) evk @@ -35,12 +33,12 @@ let depends_on_evar sigma evk _ (pbty,_,t1,t2) = let define_and_solve_constraints evk c env evd = if Termops.occur_evar evd evk c then Pretype_errors.error_occur_check env evd evk c; - let evd = define evk (EConstr.Unsafe.to_constr c) evd in + let evd = define evk c evd in let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evd evk) in match List.fold_left (fun p (pbty,env,t1,t2) -> match p with - | Success evd -> Evarconv.evar_conv_x full_transparent_state env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) + | Success evd -> Evarconv.evar_conv_x full_transparent_state env evd pbty t1 t2 | UnifFailure _ as x -> x) (Success evd) pbs with @@ -59,7 +57,7 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma = Pretyping.fail_evar = false; Pretyping.expand_evars = true } in try Pretyping.understand_ltac flags - env sigma ltac_var (Pretyping.OfType (EConstr.of_constr evi.evar_concl)) rawc + env sigma ltac_var (Pretyping.OfType evi.evar_concl) rawc with e when CErrors.noncritical e -> let loc = Glob_ops.loc_of_glob_constr rawc in user_err ?loc diff --git a/proofs/goal.ml b/proofs/goal.ml index 6912db364..1440d1636 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -48,7 +48,7 @@ module V82 = struct (* Access to ".evar_concl" *) let concl evars gl = let evi = Evd.find evars gl in - EConstr.of_constr evi.Evd.evar_concl + evi.Evd.evar_concl (* Access to ".evar_extra" *) let extra evars gl = @@ -61,7 +61,6 @@ module V82 = struct be shelved. It must not appear as a future_goal, so the future goals are restored to their initial value after the evar is created. *) - let concl = EConstr.Unsafe.to_constr concl in let prev_future_goals = Evd.save_future_goals evars in let evi = { Evd.evar_hyps = hyps; Evd.evar_concl = concl; @@ -86,7 +85,7 @@ module V82 = struct if not (Evarutil.occur_evar_upto sigma evk c) then () else Pretype_errors.error_occur_check Environ.empty_env sigma evk c in - Evd.define evk (EConstr.Unsafe.to_constr c) sigma + Evd.define evk c sigma (* Instantiates a goal with an open term, using name of goal for evk' *) let partial_solution_to sigma evk evk' c = @@ -100,7 +99,9 @@ module V82 = struct let same_goal evars1 gl1 evars2 gl2 = let evi1 = Evd.find evars1 gl1 in let evi2 = Evd.find evars2 gl2 in - Constr.equal evi1.Evd.evar_concl evi2.Evd.evar_concl && + let c1 = EConstr.Unsafe.to_constr evi1.Evd.evar_concl in + let c2 = EConstr.Unsafe.to_constr evi2.Evd.evar_concl in + Constr.equal c1 c2 && Environ.eq_named_context_val evi1.Evd.evar_hyps evi2.Evd.evar_hyps let weak_progress glss gls = diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index 8725f51cd..abda04ff1 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -233,7 +233,7 @@ let apply_implicit_tactic tac = (); fun env sigma evk -> (Environ.named_context env) -> let tac = Proofview.tclTHEN tac (Proofview.tclEXTEND [] (Proofview.tclZERO (CErrors.UserError (None,Pp.str"Proof is not complete."))) []) in (try - let c = Evarutil.nf_evars_universes sigma evi.evar_concl in + let c = Evarutil.nf_evars_universes sigma (EConstr.Unsafe.to_constr evi.evar_concl) in let c = EConstr.of_constr c in if Evarutil.has_undefined_evars sigma c then raise Exit; let (ans, _, ctx) = diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index d6c0e3341..fc7c437e6 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -340,7 +340,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now have existential variables in the initial types of goals, we need to normalise them for the kernel. *) let subst_evar k = - Proof.in_proof proof (fun m -> Evd.existential_opt_value m k) in + Proof.in_proof proof (fun m -> Evd.existential_opt_value0 m k) in let nf = Universes.nf_evars_and_universes_opt_subst subst_evar (UState.subst universes) in let make_body = diff --git a/proofs/refine.ml b/proofs/refine.ml index 909556b1e..5a2d82977 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -15,7 +15,7 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration let extract_prefix env info = - let ctx1 = List.rev (Environ.named_context env) in + let ctx1 = List.rev (EConstr.named_context env) in let ctx2 = List.rev (Evd.evar_context info) in let rec share l1 l2 accu = match l1, l2 with | d1 :: l1, d2 :: l2 -> @@ -29,21 +29,21 @@ let typecheck_evar ev env sigma = let info = Evd.find sigma ev in (** Typecheck the hypotheses. *) let type_hyp (sigma, env) decl = - let t = EConstr.of_constr (NamedDecl.get_type decl) in + let t = NamedDecl.get_type decl in let evdref = ref sigma in let _ = Typing.e_sort_of env evdref t in let () = match decl with | LocalAssum _ -> () - | LocalDef (_,body,_) -> Typing.e_check env evdref (EConstr.of_constr body) t + | LocalDef (_,body,_) -> Typing.e_check env evdref body t in - (!evdref, Environ.push_named decl env) + (!evdref, EConstr.push_named decl env) in let (common, changed) = extract_prefix env info in - let env = Environ.reset_with_named_context (Environ.val_of_named_context common) env in + let env = Environ.reset_with_named_context (EConstr.val_of_named_context common) env in let (sigma, env) = List.fold_left type_hyp (sigma, env) changed in (** Typecheck the conclusion *) let evdref = ref sigma in - let _ = Typing.e_sort_of env evdref (EConstr.of_constr (Evd.evar_concl info)) in + let _ = Typing.e_sort_of env evdref (Evd.evar_concl info) in !evdref let typecheck_proof c concl env sigma = @@ -106,7 +106,6 @@ let generic_refine ~typecheck f gl = let evs = Evd.map_filter_future_goals (Proofview.Unsafe.advance sigma) evs in let comb,shelf,given_up,evkmain = Evd.dispatch_future_goals evs in (** Proceed to the refinement *) - let c = EConstr.Unsafe.to_constr c in let sigma = match Proofview.Unsafe.advance sigma self with | None -> (** Nothing to do, the goal has been solved by side-effect *) @@ -124,7 +123,8 @@ let generic_refine ~typecheck f gl = (** Mark goals *) let sigma = CList.fold_left Proofview.Unsafe.mark_as_goal sigma comb in let comb = CList.map (fun x -> Proofview.goal_with_state x state) comb in - let trace () = Pp.(hov 2 (str"simple refine"++spc()++ Hook.get pr_constrv env sigma c)) in + let trace () = Pp.(hov 2 (str"simple refine"++spc()++ + Hook.get pr_constrv env sigma (EConstr.Unsafe.to_constr c))) in Proofview.Trace.name_tactic trace (Proofview.tclUNIT v) >>= fun v -> Proofview.Unsafe.tclSETENV (Environ.reset_context env) <*> Proofview.Unsafe.tclEVARS sigma <*> diff --git a/stm/proofBlockDelimiter.ml b/stm/proofBlockDelimiter.ml index 23f976120..0af766219 100644 --- a/stm/proofBlockDelimiter.ml +++ b/stm/proofBlockDelimiter.ml @@ -41,7 +41,7 @@ let simple_goal sigma g gs = let open Evd in let open Evarutil in let evi = Evd.find sigma g in - Set.is_empty (evars_of_term evi.evar_concl) && + Set.is_empty (evars_of_term (EConstr.Unsafe.to_constr evi.evar_concl)) && Set.is_empty (evars_of_filtered_evar_info (nf_evar_info sigma evi)) && not (List.exists (Proofview.depends_on sigma g) gs) diff --git a/stm/stm.ml b/stm/stm.ml index 30aa9ea06..326b6d1c2 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1896,7 +1896,7 @@ end = struct (* {{{ *) stm_purify (fun () -> let _,_,_,_,sigma0 = Proof.proof (Proof_global.give_me_the_proof ()) in let g = Evd.find sigma0 r_goal in - let is_ground c = Evarutil.is_ground_term sigma0 (EConstr.of_constr c) in + let is_ground c = Evarutil.is_ground_term sigma0 c in if not ( is_ground Evd.(evar_concl g) && List.for_all (Context.Named.Declaration.for_all is_ground) @@ -1919,7 +1919,6 @@ end = struct (* {{{ *) match Evd.(evar_body (find sigma r_goal)) with | Evd.Evar_empty -> RespNoProgress | Evd.Evar_defined t -> - let t = EConstr.of_constr t in let t = Evarutil.nf_evar sigma t in if Evarutil.is_ground_term sigma t then let t = EConstr.Unsafe.to_constr t in diff --git a/tactics/auto.ml b/tactics/auto.ml index 0c0d9bcfc..15a24fb37 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -8,8 +8,6 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -module CVars = Vars - open Pp open Util open Names @@ -82,14 +80,13 @@ let connect_hint_clenv poly (c, _, ctx) clenv gl = if poly then (** Refresh the instance of the hint *) let (subst, ctx) = Universes.fresh_universe_context_set_instance ctx in - let map c = CVars.subst_univs_level_constr subst c in let emap c = Vars.subst_univs_level_constr subst c in let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in (** Only metas are mentioning the old universes. *) let clenv = { templval = Evd.map_fl emap clenv.templval; templtyp = Evd.map_fl emap clenv.templtyp; - evd = Evd.map_metas map evd; + evd = Evd.map_metas emap evd; env = Proofview.Goal.env gl; } in clenv, emap c diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 0260460e6..b11e36bce 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1030,8 +1030,8 @@ module Intpart = Unionfind.Make(Evar.Set)(Evar.Map) let deps_of_constraints cstrs evm p = List.iter (fun (_, _, x, y) -> - let evx = Evarutil.undefined_evars_of_term evm (EConstr.of_constr x) in - let evy = Evarutil.undefined_evars_of_term evm (EConstr.of_constr y) in + let evx = Evarutil.undefined_evars_of_term evm x in + let evy = Evarutil.undefined_evars_of_term evm y in Intpart.union_set (Evar.Set.union evx evy) p) cstrs @@ -1076,7 +1076,7 @@ let error_unresolvable env comp evd = | Some s -> Evar.Set.mem ev s in let fold ev evi (found, accu) = - let ev_class = class_of_constr evd (EConstr.of_constr evi.evar_concl) in + let ev_class = class_of_constr evd evi.evar_concl in if not (Option.is_empty ev_class) && is_part ev then (* focus on one instance if only one was searched for *) if not found then (true, Some ev) diff --git a/tactics/equality.ml b/tactics/equality.ml index 98f627f21..d7a3e9470 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1108,8 +1108,6 @@ let make_tuple env sigma (rterm,rty) lind = let p = mkLambda (na, a, rty) in let sigma, exist_term = Evd.fresh_global env sigma sigdata.intro in let sigma, sig_term = Evd.fresh_global env sigma sigdata.typ in - let exist_term = EConstr.of_constr exist_term in - let sig_term = EConstr.of_constr sig_term in sigma, (applist(exist_term,[a;p;(mkRel lind);rterm]), applist(sig_term,[a;p])) @@ -1203,7 +1201,6 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = let w_type = unsafe_type_of env !evdref w in if Evarconv.e_cumul env evdref w_type a then let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in - let exist_term = EConstr.of_constr exist_term in applist(exist_term,[a;p_i_minus_1;w;tuple_tail]) else user_err Pp.(str "Cannot solve a unification problem.") @@ -1372,7 +1369,6 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = let sigma, (injbody,resty) = build_injector e_env !evdref t1' (mkVar e) cpath in let injfun = mkNamedLambda e t injbody in let sigma,congr = Evd.fresh_global env sigma eq.congr in - let congr = EConstr.of_constr congr in let pf = applist(congr,[t;resty;injfun;t1;t2]) in let sigma, pf_typ = Typing.type_of env sigma pf in let inj_clause = apply_on_clause (pf,pf_typ) eq_clause in diff --git a/tactics/inv.ml b/tactics/inv.ml index d76c9a697..412954989 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -124,12 +124,10 @@ let make_inv_predicate env evd indf realargs id status concl = in let eq_term = eqdata.Coqlib.eq in let eq = Evarutil.evd_comb1 (Evd.fresh_global env) evd eq_term in - let eq = EConstr.of_constr eq in let eqn = applist (eq,[eqnty;lhs;rhs]) in let eqns = (Anonymous, lift n eqn) :: eqns in let refl_term = eqdata.Coqlib.refl in let refl_term = Evarutil.evd_comb1 (Evd.fresh_global env) evd refl_term in - let refl_term = EConstr.of_constr refl_term in let refl = mkApp (refl_term, [|eqnty; rhs|]) in let _ = Evarutil.evd_comb1 (Typing.type_of env) evd refl in let args = refl :: args in diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index a97ae8f65..5e81e2d4b 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -263,7 +263,7 @@ let pf_with_evars glsev k gls = tclTHEN (Refiner.tclEVARS evd) (k a) gls let pf_constr_of_global gr k = - pf_with_evars (fun gls -> on_snd EConstr.of_constr (pf_apply Evd.fresh_global gls gr)) k + pf_with_evars (fun gls -> pf_apply Evd.fresh_global gls gr) k (** Tacticals of Ltac defined directly in term of Proofview *) module New = struct @@ -506,7 +506,7 @@ module New = struct let evi = Evd.find sigma evk in match Evd.evar_body evi with | Evd.Evar_empty -> Some (evk,evi) - | Evd.Evar_defined c -> match Constr.kind c with + | Evd.Evar_defined c -> match Constr.kind (EConstr.Unsafe.to_constr c) with | Term.Evar (evk,l) -> is_undefined_up_to_restriction sigma evk | _ -> (* We make the assumption that there is no way to refine an @@ -709,7 +709,7 @@ module New = struct let gl_make_elim ind = begin fun gl -> let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in let (sigma, c) = pf_apply Evd.fresh_global gl gr in - (sigma, EConstr.of_constr c) + (sigma, c) end let gl_make_case_dep (ind, u) = begin fun gl -> @@ -769,7 +769,6 @@ module New = struct Proofview.tclEVARMAP >>= fun sigma -> Proofview.tclENV >>= fun env -> let (sigma, c) = Evd.fresh_global env sigma ref in - let c = EConstr.of_constr c in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclUNIT c end diff --git a/tactics/tactics.ml b/tactics/tactics.ml index d0ec3358a..efb46e7d9 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1258,7 +1258,6 @@ let cut c = end let error_uninstantiated_metas t clenv = - let t = EConstr.Unsafe.to_constr t in let na = meta_name clenv.evd (List.hd (Metaset.elements (metavars_of t))) in let id = match na with Name id -> id | _ -> anomaly (Pp.str "unnamed dependent meta.") in user_err (str "Cannot find an instance for " ++ Id.print id ++ str".") @@ -1268,7 +1267,7 @@ let check_unresolved_evars_of_metas sigma clenv = (* Refiner.pose_all_metas_as_evars are resolved *) List.iter (fun (mv,b) -> match b with | Clval (_,(c,_),_) -> - (match Constr.kind c.rebus with + (match Constr.kind (EConstr.Unsafe.to_constr c.rebus) with | Evar (evk,_) when Evd.is_undefined clenv.evd evk && not (Evd.mem sigma evk) -> error_uninstantiated_metas (mkMeta mv) clenv @@ -1445,9 +1444,7 @@ let is_nonrec mind = (Global.lookup_mind (fst mind)).mind_finite == Declarations let find_ind_eliminator ind s gl = let gr = lookup_eliminator ind s in - let evd, c = Tacmach.New.pf_apply Evd.fresh_global gl gr in - let c = EConstr.of_constr c in - evd, c + Tacmach.New.pf_apply Evd.fresh_global gl gr let find_eliminator c gl = let ((ind,u),t) = Tacmach.New.pf_reduce_to_quantified_ind gl (Tacmach.New.pf_unsafe_type_of gl c) in @@ -2612,9 +2609,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let eqdata = build_coq_eq_data () in let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in let (sigma, eq) = Evd.fresh_global env sigma eqdata.eq in - let eq = EConstr.of_constr eq in let (sigma, refl) = Evd.fresh_global env sigma eqdata.refl in - let refl = EConstr.of_constr refl in let eq = applist (eq,args) in let refl = applist (refl, [t;mkVar id]) in let term = mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)) in @@ -2668,9 +2663,7 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let eqdata = build_coq_eq_data () in let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in let (sigma, eq) = Evd.fresh_global env sigma eqdata.eq in - let eq = EConstr.of_constr eq in let (sigma, refl) = Evd.fresh_global env sigma eqdata.refl in - let refl = EConstr.of_constr refl in let eq = applist (eq,args) in let refl = applist (refl, [t;mkVar id]) in let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in diff --git a/vernac/class.ml b/vernac/class.ml index 59d933108..f0b01061b 100644 --- a/vernac/class.ml +++ b/vernac/class.ml @@ -181,6 +181,7 @@ let build_id_coercion idf_opt source poly = let sigma, vs = match source with | CL_CONST sp -> Evd.fresh_global env sigma (ConstRef sp) | _ -> error_not_transparent source in + let vs = EConstr.Unsafe.to_constr vs in let c = match constant_opt_value_in env (destConst vs) with | Some c -> c | None -> error_not_transparent source in diff --git a/vernac/himsg.ml b/vernac/himsg.ml index 23c863cab..acb461cac 100644 --- a/vernac/himsg.ml +++ b/vernac/himsg.ml @@ -75,11 +75,7 @@ let rec contract3' env sigma a b c = function | MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities | UnifUnivInconsistency _ as x -> contract3 env sigma a b c, x | CannotSolveConstraint ((pb,env',t,u),x) -> - let t = EConstr.of_constr t in - let u = EConstr.of_constr u in let env',t,u = contract2 env' sigma t u in - let t = EConstr.Unsafe.to_constr t in - let u = EConstr.Unsafe.to_constr u in let y,x = contract3' env sigma a b c x in y,CannotSolveConstraint ((pb,env',t,u),x) @@ -322,8 +318,6 @@ let explain_unification_error env sigma p1 p2 = function else [str "universe inconsistency"] | CannotSolveConstraint ((pb,env,t,u),e) -> - let t = EConstr.of_constr t in - let u = EConstr.of_constr u in let env = make_all_name_different env sigma in (strbrk "cannot satisfy constraint " ++ pr_leconstr_env env sigma t ++ str " == " ++ pr_leconstr_env env sigma u) @@ -562,9 +556,9 @@ let rec explain_evar_kind env sigma evk ty = function | Evar_kinds.SubEvar (where,evk') -> let evi = Evd.find sigma evk' in let pc = match evi.evar_body with - | Evar_defined c -> pr_leconstr_env env sigma (EConstr.of_constr c) + | Evar_defined c -> pr_leconstr_env env sigma c | Evar_empty -> assert false in - let ty' = EConstr.of_constr evi.evar_concl in + let ty' = evi.evar_concl in (match where with | Some Evar_kinds.Body -> str "the body of " | Some Evar_kinds.Domain -> str "the domain of " @@ -577,11 +571,11 @@ let rec explain_evar_kind env sigma evk ty = function (pr_leconstr_env env sigma ty') (snd evi.evar_source) let explain_typeclass_resolution env sigma evi k = - match Typeclasses.class_of_constr sigma (EConstr.of_constr evi.evar_concl) with + match Typeclasses.class_of_constr sigma evi.evar_concl with | Some _ -> let env = Evd.evar_filtered_env evi in fnl () ++ str "Could not find an instance for " ++ - pr_lconstr_env env sigma evi.evar_concl ++ + pr_leconstr_env env sigma evi.evar_concl ++ pr_trailing_ne_context_of env sigma | _ -> mt() @@ -590,14 +584,14 @@ let explain_placeholder_kind env sigma c e = | Some (SeveralInstancesFound n) -> strbrk " (several distinct possible type class instances found)" | None -> - match Typeclasses.class_of_constr sigma (EConstr.of_constr c) with + match Typeclasses.class_of_constr sigma c with | Some _ -> strbrk " (no type class instance found)" | _ -> mt () let explain_unsolvable_implicit env sigma evk explain = let evi = Evarutil.nf_evar_info sigma (Evd.find_undefined sigma evk) in let env = Evd.evar_filtered_env evi in - let type_of_hole = pr_lconstr_env env sigma evi.evar_concl in + let type_of_hole = pr_leconstr_env env sigma evi.evar_concl in let pe = pr_trailing_ne_context_of env sigma in strbrk "Cannot infer " ++ explain_evar_kind env sigma evk type_of_hole (snd evi.evar_source) ++ @@ -765,7 +759,7 @@ let pr_constraints printenv env sigma evars cstrs = let evs = prlist (fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++ - str " : " ++ pr_lconstr_env env' sigma evi.evar_concl ++ fnl ()) l + str " : " ++ pr_leconstr_env env' sigma evi.evar_concl ++ fnl ()) l in h 0 (pe ++ evs ++ pr_evar_constraints sigma cstrs) else diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 064e40b9b..3f2792518 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -209,8 +209,10 @@ let eterm_obligations env name evm fs ?status t ty = List.fold_right (fun (id, (n, nstr), ev) l -> let hyps = Evd.evar_filtered_context ev in - let hyps = trunc_named_context nc_len hyps in - let evtyp, deps, transp = etype_of_evar l hyps ev.evar_concl in + let hyps = trunc_named_context nc_len hyps in + let hyps = EConstr.Unsafe.to_named_context hyps in + let concl = EConstr.Unsafe.to_constr ev.evar_concl in + let evtyp, deps, transp = etype_of_evar l hyps concl in let evtyp, hyps, chop = match chop_product fs evtyp with | Some t -> t, trunc_named_context fs hyps, fs @@ -356,7 +358,7 @@ let _ = optread = get_shrink_obligations; optwrite = set_shrink_obligations; } -let evar_of_obligation o = make_evar (Global.named_context_val ()) o.obl_type +let evar_of_obligation o = make_evar (Global.named_context_val ()) (EConstr.of_constr o.obl_type) let get_obligation_body expand obl = match obl.obl_body with @@ -813,10 +815,9 @@ let rec string_of_list sep f = function 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 - id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in + id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps evi.evar_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 diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index b44c7cccb..a9d1631ba 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -449,7 +449,7 @@ let start_proof_and_print k l hook = let evi = Evarutil.nf_evar_info sigma evi in let env = Evd.evar_filtered_env evi in try - let concl = EConstr.of_constr evi.Evd.evar_concl in + let concl = evi.Evd.evar_concl in if not (Evarutil.is_ground_env sigma env && Evarutil.is_ground_term sigma concl) then raise Exit; -- cgit v1.2.3