diff --git a/theories/heap_lang/lang.v b/theories/heap_lang/lang.v index e02449f2115159a6e11bb0de00bd0fa0afbc9692..1d6861cdb40701df1cb9b112df7116613066c535 100644 --- a/theories/heap_lang/lang.v +++ b/theories/heap_lang/lang.v @@ -153,7 +153,8 @@ read and written. Also notice that the sets of boxed and unboxed values are disjoint. *) Definition lit_is_unboxed (l: base_lit) : Prop := match l with - (* disallow comparing (erased) prophecies with (erased) prophecies, by considering them boxed *) + (** Disallow comparing (erased) prophecies with (erased) prophecies, by + considering them boxed. *) | LitProphecy _ | LitErased => False | _ => True end. diff --git a/theories/heap_lang/lifting.v b/theories/heap_lang/lifting.v index a006612fa082eb8ea75816c0bbf3551df684b4e1..8dda1c8f27c9f92c557b1f97fe0131b2bbf4f96d 100644 --- a/theories/heap_lang/lifting.v +++ b/theories/heap_lang/lifting.v @@ -593,7 +593,7 @@ Proof. Qed. Lemma wp_load_offset_vec s E l sz (off : fin sz) (vs : vec val sz) : - {{{ ▷ l ↦∗ vs }}} ! #(l +ₗ off) @ s; E {{{ RET (vs !!! off); l ↦∗ vs }}}. + {{{ ▷ l ↦∗ vs }}} ! #(l +ₗ off) @ s; E {{{ RET vs !!! off; l ↦∗ vs }}}. Proof. apply wp_load_offset. by apply vlookup_lookup. Qed. Lemma wp_store_offset s E l off vs v : @@ -628,7 +628,7 @@ Proof. Qed. Lemma wp_cmpxchg_suc_offset_vec s E l sz (off : fin sz) (vs : vec val sz) v1 v2 : - (vs !!! off) = v1 → + vs !!! off = v1 → vals_compare_safe (vs !!! off) v1 → {{{ ▷ l ↦∗ vs }}} CmpXchg #(l +ₗ off) v1 v2 @ s; E @@ -655,7 +655,7 @@ Proof. Qed. Lemma wp_cmpxchg_fail_offset_vec s E l sz (off : fin sz) (vs : vec val sz) v1 v2 : - (vs !!! off) ≠v1 → + vs !!! off ≠v1 → vals_compare_safe (vs !!! off) v1 → {{{ ▷ l ↦∗ vs }}} CmpXchg #(l +ₗ off) v1 v2 @ s; E