1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64
|
(* Former open goals in nested proofs were lost *)
(* This used to fail with "Incorrect number of goals (expected 1 tactic)." *)
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.
(* This used to leave the goal on the shelf and fails at reflexivity *)
Goal (True/\0=0 -> True) -> True.
intro f.
refine
(f ltac:(split; only 1:exact I)).
reflexivity.
Qed.
(* The "Unshelve" used to not see the explicitly "shelved" goal *)
Lemma f (b:comparison) : b=b.
refine (match b with
Eq => ltac:(shelve)
| Lt => ltac:(give_up)
| Gt => _
end).
exact (eq_refl Gt).
Unshelve.
exact (eq_refl Eq).
Fail auto. (* Check that there are no more regular subgoals *)
Admitted.
(* The "Unshelve" used to not see the explicitly "shelved" goal *)
Lemma f2 (b:comparison) : b=b.
refine (match b with
Eq => ltac:(shelve)
| Lt => ltac:(give_up)
| Gt => _
end).
Unshelve. (* Note: Unshelve puts goals at the end *)
exact (eq_refl Gt).
exact (eq_refl Eq).
Fail auto. (* Check that there are no more regular subgoals *)
Admitted.
(* The "unshelve" used to not see the explicitly "shelved" goal *)
Lemma f3 (b:comparison) : b=b.
unshelve refine (match b with
Eq => ltac:(shelve)
| Lt => ltac:(give_up)
| Gt => _
end).
(* Note: unshelve puts goals at the beginning *)
exact (eq_refl Eq).
exact (eq_refl Gt).
Fail auto. (* Check that there are no more regular subgoals *)
Admitted.
|