aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGaëtan Gilbert2020-08-28 15:15:00 +0200
committerGaëtan Gilbert2020-08-28 15:15:00 +0200
commit9d319872a1668775ba1d72ae856f46fd7f082bb9 (patch)
tree5b5c9b641c88c9d78cbb3633e83993c3676c4c7d
parent911f33f0a0ff648082d329841388f59e8cecf231 (diff)
par: print relevant subgoal when failing
Fix (partial) #5502
-rw-r--r--stm/stm.ml3
-rw-r--r--test-suite/output/Partac.out6
-rw-r--r--test-suite/output/Partac.v6
3 files changed, 14 insertions, 1 deletions
diff --git a/stm/stm.ml b/stm/stm.ml
index c94a6d3a5d..9999e66c45 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -1922,7 +1922,8 @@ end = struct (* {{{ *)
str" solves the goal and leaves no unresolved existential variables. The following" ++
str" existentials remain unsolved: " ++ prlist (Termops.pr_existential_key sigma) (Evar.Set.elements evars))
end) ()
- with e when CErrors.noncritical e -> RespError (CErrors.print e)
+ with e when CErrors.noncritical e ->
+ RespError (CErrors.print e ++ spc() ++ str "(for subgoal "++int (fst r_ast) ++ str ")")
let name_of_task { t_name } = t_name
let name_of_request { r_name } = r_name
diff --git a/test-suite/output/Partac.out b/test-suite/output/Partac.out
new file mode 100644
index 0000000000..889e698fa2
--- /dev/null
+++ b/test-suite/output/Partac.out
@@ -0,0 +1,6 @@
+The command has indeed failed with message:
+The term "false" has type "bool" while it is expected to have type "nat".
+(for subgoal 1)
+The command has indeed failed with message:
+The term "0" has type "nat" while it is expected to have type "bool".
+(for subgoal 2)
diff --git a/test-suite/output/Partac.v b/test-suite/output/Partac.v
new file mode 100644
index 0000000000..f579ee683b
--- /dev/null
+++ b/test-suite/output/Partac.v
@@ -0,0 +1,6 @@
+Goal nat * bool.
+Proof.
+ split.
+ Fail par: exact false.
+ Fail par: exact 0.
+Abort.