diff options
author | Hugo Herbelin <Hugo.Herbelin@inria.fr> | 2017-12-06 09:32:51 +0100 |
---|---|---|
committer | Maxime Dénès <mail@maximedenes.fr> | 2018-03-08 21:58:32 +0100 |
commit | 6f30f76def8f6cf1abe7859f482b68c91b4c8709 (patch) | |
tree | fbee213ff8e429ae89c58d60f55242df9a256235 /test-suite/bugs/closed | |
parent | b731022c022cfeae9203f6b10b4e1f68b85d9071 (diff) |
Proof engine: support for nesting tactic-in-term within other tactics.
Tactic-in-term can be called from within a tactic itself. We have to
preserve the preexisting future_goals (if called from pretyping) and
we have to inform of the existence of pending goals, using
future_goals which is the only way to tell it in the absence of being
part of an encapsulating proofview.
This fixes #6313.
Conversely, future goals, created by pretyping, can call ltac:(giveup) or
ltac:(shelve), and this has to be remembered. So, we do it.
Diffstat (limited to 'test-suite/bugs/closed')
-rw-r--r-- | test-suite/bugs/closed/6313.v | 20 |
1 files changed, 20 insertions, 0 deletions
diff --git a/test-suite/bugs/closed/6313.v b/test-suite/bugs/closed/6313.v new file mode 100644 index 000000000..1715e7f90 --- /dev/null +++ b/test-suite/bugs/closed/6313.v @@ -0,0 +1,20 @@ +(* Former open goals in nested proofs were lost *) + +Inductive foo := a | b | c. +Goal foo -> foo. + intro x. + simple refine (match x with + | a => _ + | b => ltac:(exact b) + | c => _ + end); [exact a|exact c]. +Abort. + +(* Another example *) + +Goal (True/\0=0 -> True) -> True. + intro f. + refine + (f ltac:(split; only 1:exact I)). + reflexivity. +Qed. |