diff --git a/CHANGELOG.md b/CHANGELOG.md index e5a006f462aad96148893a528d5d342848236947..6b056c827590fea750ad05505f39f9c2d4e52ac0 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -34,6 +34,7 @@ Changes in Coq: - Camera elements such that `core x = x`: `Persistent` -> `CoreId` - Persistent propositions: `PersistentP` -> `Persistent` - The persistent modality: `always` -> `persistently` + - Adequacy for non-stuck weakestpre: `adequate_safe` -> `adequate_not_stuck` - Consistently SnakeCase identifiers: + `CMRAMixin` -> `CmraMixin` + `CMRAT` -> `CmraT` diff --git a/theories/program_logic/adequacy.v b/theories/program_logic/adequacy.v index 6192bd377000010addebb5ccdf7fbf5cc6489a8c..ada5bbd34cea8eec72687747d0d61c79e7256e5c 100644 --- a/theories/program_logic/adequacy.v +++ b/theories/program_logic/adequacy.v @@ -37,7 +37,7 @@ Qed. Record adequate {Λ} (s : stuckness) (e1 : expr Λ) (σ1 : state Λ) (φ : val Λ → Prop) := { adequate_result t2 σ2 v2 : rtc step ([e1], σ1) (of_val v2 :: t2, σ2) → φ v2; - adequate_safe t2 σ2 e2 : + adequate_not_stuck t2 σ2 e2 : s = NotStuck → rtc step ([e1], σ1) (t2, σ2) → e2 ∈ t2 → (is_Some (to_val e2) ∨ reducible e2 σ2) @@ -51,7 +51,7 @@ Proof. intros Had ?. destruct (decide (Forall (λ e, is_Some (to_val e)) t2)) as [|Ht2]; [by left|]. apply (not_Forall_Exists _), Exists_exists in Ht2; destruct Ht2 as (e2&?&He2). - destruct (adequate_safe NotStuck e1 σ1 φ Had t2 σ2 e2) as [?|(e3&σ3&efs&?)]; + destruct (adequate_not_stuck NotStuck e1 σ1 φ Had t2 σ2 e2) as [?|(e3&σ3&efs&?)]; rewrite ?eq_None_not_Some; auto. { exfalso. eauto. } destruct (elem_of_list_split t2 e2) as (t2'&t2''&->); auto.