aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGaëtan Gilbert2020-05-14 17:25:23 +0200
committerGaëtan Gilbert2020-05-14 17:25:23 +0200
commitcc54af3842cbf99f169f7937b0e31f737652bd3a (patch)
tree7812c281b7c0012e9dcb3e7e2f78f11950b2e9e2
parente36c415f5a9509f050246a62e6198bc5b7b193d3 (diff)
parent5c3b9f9082b2a628bbf877dffae5ac38b5923f89 (diff)
Merge PR #12296: Fixes #12234: wrong environment for Show Proof
Reviewed-by: SkySkimmer
-rw-r--r--doc/changelog/07-commands-and-options/12296-master+fix12234-show-proof-proper.rst5
-rw-r--r--test-suite/bugs/closed/bug_12234.v9
-rw-r--r--vernac/vernacentries.ml8
3 files changed, 21 insertions, 1 deletions
diff --git a/doc/changelog/07-commands-and-options/12296-master+fix12234-show-proof-proper.rst b/doc/changelog/07-commands-and-options/12296-master+fix12234-show-proof-proper.rst
new file mode 100644
index 0000000000..dc71a27eb8
--- /dev/null
+++ b/doc/changelog/07-commands-and-options/12296-master+fix12234-show-proof-proper.rst
@@ -0,0 +1,5 @@
+- **Fixed:**
+ Anomalies with :cmd:`Show Proof`
+ (`#12296 <https://github.com/coq/coq/pull/12296>`_,
+ by Hugo Herbelin; fixes
+ `#12234 <https://github.com/coq/coq/pull/12234>`_).
diff --git a/test-suite/bugs/closed/bug_12234.v b/test-suite/bugs/closed/bug_12234.v
new file mode 100644
index 0000000000..b99c5d524e
--- /dev/null
+++ b/test-suite/bugs/closed/bug_12234.v
@@ -0,0 +1,9 @@
+(* Checking a Show Proof bug *)
+Section S.
+Variable A:Prop.
+Theorem thm (a:A) : True.
+assert (b:=a).
+clear A a b.
+Show Proof.
+Abort.
+End S.
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index aac0b54ed4..6568d7bf14 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -95,8 +95,14 @@ let show_proof ~pstate =
try
let pstate = Option.get pstate in
let p = Declare.Proof.get_proof pstate in
- let sigma, env = Declare.get_current_context pstate in
+ let sigma, _ = Declare.get_current_context pstate in
let pprf = Proof.partial_proof p in
+ (* In the absence of an environment explicitly attached to the
+ proof and on top of which side effects of the proof would be pushed, ,
+ we take the global environment which in practise should be a
+ superset of the initial environment in which the proof was
+ started *)
+ let env = Global.env() in
Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf
(* We print nothing if there are no goals left *)
with