diff options
| author | Maxime Dénès | 2020-08-23 17:52:15 +0200 |
|---|---|---|
| committer | Maxime Dénès | 2020-08-24 15:39:08 +0200 |
| commit | a970072a64348a1b5af1cd45a138da819ef0a8d2 (patch) | |
| tree | 00f3a964bb3079fa1495341d11d61a5dc158617a | |
| parent | 0f5bd6a18d51a412ee3c8309b891254365e48b91 (diff) | |
Update sigma instead of erasing it in `update_global_env`
| -rw-r--r-- | engine/proofview.ml | 3 | ||||
| -rw-r--r-- | engine/proofview.mli | 3 | ||||
| -rw-r--r-- | proofs/proof.ml | 4 | ||||
| -rw-r--r-- | proofs/proof.mli | 3 | ||||
| -rw-r--r-- | vernac/declare.ml | 6 |
5 files changed, 14 insertions, 5 deletions
diff --git a/engine/proofview.ml b/engine/proofview.ml index de38104ecd..fd8512d73e 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -1032,6 +1032,9 @@ module Unsafe = struct let mark_as_unresolvables p evs = { p with solution = mark_in_evm ~goal:false p.solution evs } + let update_sigma_env pv env = + { pv with solution = Evd.update_sigma_env pv.solution env } + end module UnsafeRepr = Proof.Unsafe diff --git a/engine/proofview.mli b/engine/proofview.mli index d0a2b37a69..0f49d2f5d8 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -503,6 +503,9 @@ module Unsafe : sig val undefined : Evd.evar_map -> Proofview_monad.goal_with_state list -> Proofview_monad.goal_with_state list + (** [update_sigma_env] lifts [Evd.update_sigma_env] to the proofview *) + val update_sigma_env : proofview -> Environ.env -> proofview + end (** This module gives access to the innards of the monad. Its use is diff --git a/proofs/proof.ml b/proofs/proof.ml index a183fa7797..38fcdd6e5f 100644 --- a/proofs/proof.ml +++ b/proofs/proof.ml @@ -356,6 +356,10 @@ let compact p = let entry, proofview = Proofview.compact p.entry p.proofview in { p with proofview; entry } +let update_sigma_env p env = + let proofview = Proofview.Unsafe.update_sigma_env p.proofview env in + { p with proofview } + (*** Function manipulation proof extra informations ***) (*** Tactics ***) diff --git a/proofs/proof.mli b/proofs/proof.mli index 0e5bdaf07d..2d4966676e 100644 --- a/proofs/proof.mli +++ b/proofs/proof.mli @@ -78,6 +78,9 @@ val partial_proof : t -> EConstr.constr list val compact : t -> t +(** [update_sigma_env] lifts [Evd.update_sigma_env] to the proof *) +val update_sigma_env : t -> Environ.env -> t + (* Returns the proofs (with their type) of the initial goals. Raises [UnfinishedProof] is some goals remain to be considered. Raises [HasShelvedGoals] if some goals are left on the shelf. diff --git a/vernac/declare.ml b/vernac/declare.ml index 66537c2978..28e6f21d41 100644 --- a/vernac/declare.ml +++ b/vernac/declare.ml @@ -1735,11 +1735,7 @@ let return_proof ps = List.map (fun (((_ub, body),eff),_) -> (body,eff)) p, uctx let update_global_env = - map ~f:(fun p -> - let { Proof.sigma } = Proof.data p in - let tac = Proofview.Unsafe.tclEVARS (Evd.update_sigma_env sigma (Global.env ())) in - let p, (status,info), _ = Proof.run_tactic (Global.env ()) tac p in - p) + map ~f:(fun p -> Proof.update_sigma_env p (Global.env ())) let next = let n = ref 0 in fun () -> incr n; !n |
