diff --git a/CHANGELOG.md b/CHANGELOG.md index ac57e8efe25ad5571ebddce44166be2eaed1dd03..9c8723765fbc998c9be8ce2702685112a1c4b2bf 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -108,6 +108,7 @@ API-breaking change is listed. - Rename `insert_delete` → `insert_delete_insert`; add new `insert_delete` that is consistent with `delete_insert`. - Fix statement of `sum_inhabited_r`. (by Paolo G. Giarrusso) +- Make `done` work on goals of the form `is_Some`. The following `sed` script should perform most of the renaming (on macOS, replace `sed` by `gsed`, installed via e.g. `brew install gnu-sed`). diff --git a/theories/option.v b/theories/option.v index a3513f5cb6668f37b73c3cfdcc935e622336b2df..0fd7dfaae60935620c14e3f221f9b46fb23ca92a 100644 --- a/theories/option.v +++ b/theories/option.v @@ -40,12 +40,14 @@ Proof. congruence. Qed. Definition is_Some {A} (mx : option A) := ∃ x, mx = Some x. Global Instance: Params (@is_Some) 1 := {}. +Global Hint Extern 0 (is_Some _) => by eexists : core. + Lemma is_Some_alt {A} (mx : option A) : is_Some mx ↔ match mx with Some _ => True | None => False end. Proof. unfold is_Some. destruct mx; naive_solver. Qed. Lemma mk_is_Some {A} (mx : option A) x : mx = Some x → is_Some mx. -Proof. intros; red; subst; eauto. Qed. +Proof. by intros ->. Qed. Global Hint Resolve mk_is_Some : core. Lemma is_Some_None {A} : ¬is_Some (@None A). Proof. by destruct 1. Qed. @@ -135,7 +137,7 @@ Section setoids. Proof. split; [inversion 1; naive_solver|naive_solver (by constructor)]. Qed. Global Instance is_Some_proper : Proper ((≡@{option A}) ==> iff) is_Some. - Proof. inversion_clear 1; split; eauto. Qed. + Proof. by inversion_clear 1. Qed. Global Instance from_option_proper {B} (R : relation B) : Proper (((≡@{A}) ==> R) ==> R ==> (≡) ==> R) from_option. Proof. destruct 3; simpl; auto. Qed.