diff options
| author | Emilio Jesus Gallego Arias | 2018-03-31 01:06:18 +0200 |
|---|---|---|
| committer | Emilio Jesus Gallego Arias | 2018-03-31 01:06:18 +0200 |
| commit | 7b9967a1a145ab70181d362ddafeba7ecb0916bd (patch) | |
| tree | 2d74732a4e05aa094cc05fcc8a6469aad550b6fc | |
| parent | 238e074cd1767b5ce12a733d8e835dd7d715c518 (diff) | |
| parent | 820442433915e4176861e558b95413e4b832b3fa (diff) | |
Merge PR #7111: Fix #6631: Derive Plugin gives "Anomaly: more than one statement".
| -rw-r--r-- | tactics/tactics.ml | 16 | ||||
| -rw-r--r-- | test-suite/bugs/closed/6631.v | 7 |
2 files changed, 15 insertions, 8 deletions
diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 834d73bdda..0d9f3d8216 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4987,15 +4987,15 @@ let anon_id = Id.of_string "anonymous" let name_op_to_name name_op object_kind suffix = let open Proof_global in let default_gk = (Global, false, object_kind) in + let name, gk = match Proof_global.V82.get_current_initial_conclusions () with + | (id, (_, gk)) -> Some id, gk + | exception NoCurrentProof -> None, default_gk + in match name_op with - | Some s -> - (try let _, gk, _ = Pfedit.current_proof_statement () in s, gk - with NoCurrentProof -> s, default_gk) - | None -> - let name, gk = - try let name, gk, _ = Pfedit.current_proof_statement () in name, gk - with NoCurrentProof -> anon_id, default_gk in - add_suffix name suffix, gk + | Some s -> s, gk + | None -> + let name = Option.default anon_id name in + add_suffix name suffix, gk let tclABSTRACT ?(opaque=true) name_op tac = let s, gk = if opaque diff --git a/test-suite/bugs/closed/6631.v b/test-suite/bugs/closed/6631.v new file mode 100644 index 0000000000..100dc13fc8 --- /dev/null +++ b/test-suite/bugs/closed/6631.v @@ -0,0 +1,7 @@ +Require Import Coq.derive.Derive. + +Derive f SuchThat (f = 1 + 1) As feq. +Proof. + transitivity 2; [refine (eq_refl 2)|]. + transitivity 2. + 2:abstract exact (eq_refl 2). |
