aboutsummaryrefslogtreecommitdiff
path: root/stm/proofBlockDelimiter.ml
diff options
context:
space:
mode:
authorMaxime Dénès2020-08-31 10:36:32 +0200
committerMaxime Dénès2020-09-01 16:20:36 +0200
commit636fe1deaf220f1c30821846343b3a70ef7a078c (patch)
treebe0d256f863dc7ecf1cd2dbbb510c7622282ad24 /stm/proofBlockDelimiter.ml
parent0d30f79268fea18ef99c040a859956f61c3d978a (diff)
Unify the shelves
Before this patch, the proof engine had three notions of shelves: - A local shelf in `proofview` - A global shelf in `Proof.t` - A future shelf in `evar_map` This has lead to a lot of confusion and limitations or bugs, because some components have only a partial view of the shelf: the pretyper can see only the future shelf, tactics can see only the local and future shelves. In particular, this refactoring is needed for #7825. The solution we choose is to move shelf information to the evar map, as a shelf stack (for nested `unshelve` tacticals). Closes #8770. Closes #6292. Co-authored-by: Gaëtan Gilbert <gaetan.gilbert@skyskimmer.net>
Diffstat (limited to 'stm/proofBlockDelimiter.ml')
-rw-r--r--stm/proofBlockDelimiter.ml4
1 files changed, 2 insertions, 2 deletions
diff --git a/stm/proofBlockDelimiter.ml b/stm/proofBlockDelimiter.ml
index e41f62361d..f367167d48 100644
--- a/stm/proofBlockDelimiter.ml
+++ b/stm/proofBlockDelimiter.ml
@@ -51,8 +51,8 @@ let is_focused_goal_simple ~doc id =
| `Valid (Some { Vernacstate.lemmas }) ->
Option.cata (Vernacstate.LemmaStack.with_top ~f:(fun proof ->
let proof = Declare.Proof.get proof in
- let Proof.{ goals=focused; stack=r1; shelf=r2; sigma } = Proof.data proof in
- let rest = List.(flatten (map (fun (x,y) -> x @ y) r1)) @ r2 @ (Evar.Set.elements @@ Evd.given_up sigma) in
+ let Proof.{ goals=focused; stack=r1; sigma } = Proof.data proof in
+ let rest = List.(flatten (map (fun (x,y) -> x @ y) r1)) @ (Evd.shelf sigma) @ (Evar.Set.elements @@ Evd.given_up sigma) in
if List.for_all (fun x -> simple_goal sigma x rest) focused
then `Simple focused
else `Not)) `Not lemmas