Skip to content
Snippets Groups Projects
Commit fd02cdaa authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

In logatom lock: Formalize that the lock is Free before acquire.

parent 6d670edc
No related branches found
No related tags found
No related merge requests found
...@@ -74,7 +74,7 @@ Section tada. ...@@ -74,7 +74,7 @@ Section tada.
tada_is_lock γ lk -∗ tada_is_lock γ lk -∗
<<< ∀∀ s, tada_lock_state γ s >>> <<< ∀∀ s, tada_lock_state γ s >>>
l.(acquire) lk @ l.(acquire) lk @
<<< tada_lock_state γ Locked, RET #() >>>. <<< s = Free tada_lock_state γ Locked, RET #() >>>.
Proof. Proof.
iIntros "#Hislock %Φ AU". iApply wp_fupd. iIntros "#Hislock %Φ AU". iApply wp_fupd.
wp_apply (l.(acquire_spec) with "Hislock"). wp_apply (l.(acquire_spec) with "Hislock").
...@@ -83,7 +83,7 @@ Section tada. ...@@ -83,7 +83,7 @@ Section tada.
iDestruct (ghost_var_agree with "Hvar1 Hvar2") as %<-. iDestruct (ghost_var_agree with "Hvar1 Hvar2") as %<-.
iMod (ghost_var_update_2 Locked with "Hvar1 Hvar2") as "[Hvar1 Hvar2]". iMod (ghost_var_update_2 Locked with "Hvar1 Hvar2") as "[Hvar1 Hvar2]".
{ rewrite Qp.quarter_three_quarter //. } { rewrite Qp.quarter_three_quarter //. }
iMod ("Hclose" with "[$Hvar2 $Hlocked $Hvar1]"). done. iMod ("Hclose" with "[$Hvar2 $Hlocked $Hvar1]"); done.
Qed. Qed.
Lemma release_tada_spec γ lk : Lemma release_tada_spec γ lk :
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment