From iris.heap_lang Require Export proofmode notation. From iris.heap_lang Require Import spin_lock. From iris.base_logic.lib Require Import cancelable_invariants auth saved_prop. From iris.algebra Require Import auth agree excl frac gmap gset. From iris.bi.lib Require Import fractional. Inductive lockstate := | Locked | Unlocked. Canonical Structure lockstateC := leibnizC lockstate. Instance lockstate_inhabited : Inhabited lockstate := populate Unlocked. (** We need to be able to allocate resources into the lock even when it is in use -- for that we need to have /active/ and /pending/ propositions. If the lock is unlocked (not acquired), then all the propositions are active (can be acquired); if the lock is locked (acquired), then some thread has the ownership of the active propositions, but any other thread can add a pending proposition. Pending propositions are thus always held by the lock. When a thread acquires the lock, it gets all the active propositions. For the whole construction to make sense, the thread should return exactly the same propositions once it releases the lock. For this to work we need to ensure that the set of active proposition has not changed since the lock has been acquired. Thus, we use ghost state to track the exact content of the set of active propositions. If we have a full access to the resource (with fraction 1), we may want to cancel it and get the proposition back. In order to do that we need to make sure that if we have the full fraction for accessing the resource, then this resource cannot be active. For that purpose we also need to store the information about the fractional permissions for the active propositions. Finally, the actual propositional content of the resources is shared via cancellable invariants. Each proposition (active or pending) is stored in a cancellable invariant associated with it: C(X) = X.prop ∗ X.token₁ ∨ X.token₂ The exclusive tokens allow the thread to take out the proposition for the whole duration of the critical section. *) Record flock_name := { (** ghost name for the actual spin lock invariant *) flock_lock_name : gname; (** -//- for keeping the track of the state of the flock *) flock_state_name : gname; (** -//- for keeping track of all propositions in the flock *) flock_props_name : gname; (** -//- for keeping track of the "active" propositions. Once you relase the lock you need to know that the set of the active propositions has not changed since you have acquired it *) flock_props_active_name : gname }. (** The data about pending and active propositions will be stored in the maps `lock_res_gname -> frac` *) (** One of the tricks in this lock is to store only ghost names associated with a 'flock resource' in the flock invariant. If we start storing propositions, then we end up in the LATER HELL. The `lock_res_name` datastructure contains all the ghosts names needed to manage an individual proposition. *) Record lock_res_gname := { flock_cinv_name : gname; flock_token1_name : gname; flock_token2_name : gname; }. Canonical Structure lock_res_gnameC := leibnizC lock_res_gname. Instance lock_res_gname_eq_dec : EqDecision lock_res_gname. Proof. solve_decision. Defined. Instance lock_res_gname_countable : Countable lock_res_gname. Proof. apply inj_countable' with (f:=(λ x, (flock_cinv_name x, flock_token1_name x, flock_token2_name x))) (g:=(λ '(γ1,γ2,γ3), Build_lock_res_gname γ1 γ2 γ3)). intros []. reflexivity. Defined. (** Before we define the type of flock resources, we need to have access to the appropriate CMRAs. The totality of propositions , and the pending propositions will be represented as a map: `lock_res_gname → frac`, so a standart ghost heap. Active propositions will be modelled by a map `excl (lock_res_name → frac)`. In other words, this map will contain *exact* information of which propositions are active and with what fractions were they activated. *) Class flockG Σ := FlockG { flock_stateG :> inG Σ (authR (optionUR (exclR lockstateC))); flock_lockG :> lockG Σ; flock_cinvG :> cinvG Σ; (* note the difference between the two RAs here ! *) flock_props_active :> inG Σ (authR (optionUR (exclR (gmapC lock_res_gname fracR)))); flock_props :> inG Σ (authR (gmapUR lock_res_gname fracR)); flock_tokens :> inG Σ (exclR unitC); }. Definition flockΣ : gFunctors := #[GFunctor (authR (optionUR (exclR lockstateC))) ;lockΣ ;cinvΣ ;GFunctor (authR (optionUR (exclR (gmapC lock_res_gname fracC)))) ;GFunctor (authR (gmapUR lock_res_gname fracR))%CF]. Instance subG_flockΣ Σ : subG flockΣ Σ → flockG Σ. Proof. solve_inG. Qed. Section flock. Context `{heapG Σ, flockG Σ}. Variable N : namespace. Definition invN := N.@"flock_inv". Definition lockN := N.@"flock_lock". Definition resN := N.@"flock_res". (** * Definitions **) Definition token₁ (γ : lock_res_gname) : iProp Σ := own (flock_token1_name γ) (Excl ()). Definition token₂ (γ : lock_res_gname) : iProp Σ := own (flock_token2_name γ) (Excl ()). Definition C (R : iProp Σ) (ρ : lock_res_gname) : iProp Σ := (R ∗ token₁ ρ ∨ token₂ ρ)%I. (** Tokens that allow you to get all the propositions from `A` out of the invariants *) Definition all_tokens (P : gmap lock_res_gname ()) : iProp Σ := ([∗ map] ρ ↦ _ ∈ P, token₂ ρ)%I. (** For active propositions we also need to know the fraction which was used to access it *) Definition from_active (A : gmap lock_res_gname frac) := fmap (const ()) A. (** A flock resource `(ρ,X,π)` contains: - Knowledge that the `ρ ↦ π` is in the set of propositions managed by the flock; - There is a cancellable invariant `C X ρ`; - We have π- permissiones to access that invariant. *) Definition N_of_ρ (ρ : lock_res_gname) := (flock_token1_name ρ, flock_token2_name ρ). (* When creating a new invariant we need to pick a name that depends _only_ on the ghost names of the tokens. We cannot depend on the ghost name for the cancellable invariant, because that would be circular. I mean we potentially can but that would be annoying. *) Definition flock_res (γ : flock_name) (ρ : lock_res_gname) (π : frac) (P : iProp Σ) : iProp Σ := (own (flock_props_name γ) (◯ {[ρ := π]}) ∗ cinv (resN.@(N_of_ρ ρ)) (flock_cinv_name ρ) (C P ρ) ∗ cinv_own (flock_cinv_name ρ) π)%I. (** Basically says that the propositions from the map `A` are active. This predicate is designed to contain enough information to deactivate those propositions. When we are deactivating the propositions we need to make sure that the set of the active propositions has not changed. That's why we need a CMRA to keep track of the exact contents of that set. *) Definition flocked (γ : flock_name) (A : gmap lock_res_gname (frac*iProp Σ)) : iProp Σ := (let fa := fmap fst A : gmap lock_res_gname frac in (* Information we retain: the flock is locked .. *) own (flock_state_name γ) (◯ (Excl' Locked)) (* What are the exact propositions that we have activated.. *) ∗ own (flock_props_active_name γ) (◯ (Excl' fa)) (* Tokens and permissions for closing the active propositions .. *) ∗ ([∗ map] ρ ↦ πR ∈ A, token₂ ρ ∗ cinv (resN.@(N_of_ρ ρ)) (flock_cinv_name ρ) (C πR.2 ρ) ∗ cinv_own (flock_cinv_name ρ) πR.1))%I. Definition to_props_map (f : gmap lock_res_gname ()) : gmapUR lock_res_gname fracR := fmap (const 1%Qp) f. Definition flock_inv (γ : flock_name) : iProp Σ := (∃ (s : lockstate) (** fa -- active propositions, fp -- pending propositions *) (fa : gmap lock_res_gname frac) (fp : gmap lock_res_gname unit), (** aka mapset lock_res_name *) (** ... and those sets are disjoint *) ⌜fp ##ₘ from_active fa⌝ ∗ own (flock_state_name γ) (● (Excl' s)) ∗ (** the totality of resources: fp ∪ fa *) own (flock_props_name γ) (● to_props_map (fp ∪ from_active fa)) ∗ (** but we also need to know the exact contents of fa *) own (flock_props_active_name γ) (● Excl' fa) ∗ (** the invariant keeps track of the tokens (token₂) for the pending resource; when we activate a resource, we transfer this token to the client, which uses it to open the cinvariant and get the actual proposition. *) all_tokens fp ∗ match s with | Locked => (** If the lock is acuired, then we hold on to the token from the "old" specification *) locked (flock_lock_name γ) ∗ (** .. and we take hold of all the ghost pointsto `s ↦{π} ρ`. This is necessary to have if we want to free the resources. If we wish to free the resource `ρ ↦ 1`, then we need to know that it is not currently active (has not been acquired). The predicate below ensures that. *) own (flock_props_name γ) (◯ fa) | Unlocked => (** If the lock is NOT acquired, then there are no active proposition *) own (flock_props_active_name γ) (◯ Excl' ∅) end)%I. Definition is_flock (γ : flock_name) (lk : val) : iProp Σ := (inv invN (flock_inv γ) ∗ is_lock lockN (flock_lock_name γ) lk (own (flock_state_name γ) (◯ (Excl' Unlocked))))%I. (** * Lemmata **) (** ** Basic properties of the CAPs *) Lemma flock_res_op (γ : flock_name) (ρ : lock_res_gname) (R : iProp Σ) (π1 π2 : frac) : flock_res γ ρ (π1+π2) R ⊣⊢ flock_res γ ρ π1 R ∗ flock_res γ ρ π2 R. Proof. rewrite /flock_res. iSplit. - iDestruct 1 as "(Hs & #Hinv & Hρ)". iDestruct "Hρ" as "[Hρ1 Hρ2]". iDestruct "Hs" as "[Hs1 Hs2]". iSplitL "Hρ1 Hs1"; by iFrame. - iIntros "[H1 H2]". iDestruct "H1" as "(Hs1 & #Hinv1 & Hρ1)". iDestruct "H2" as "(Hs2 & #Hinv2 & Hρ2)". iCombine "Hs1 Hs2" as "Hs". iFrame. rewrite /C /=. iFrame "Hinv1". Qed. Global Instance flock_res_fractional γ i R : Fractional (λ π, flock_res γ i π R). Proof. intros p q. apply flock_res_op. Qed. Global Instance flock_res_as_fractional γ i R π : AsFractional (flock_res γ i π R) (λ π, flock_res γ i π R) π. Proof. split. done. apply _. Qed. Lemma to_props_map_singleton_included fp q ρ: {[ρ := q]} ≼ to_props_map fp → fp !! ρ = Some (). Proof. rewrite singleton_included=> -[[q' av] []]. rewrite /to_props_map lookup_fmap fmap_Some_equiv. move=> -[[] [Hρ _]]//. Qed. Lemma to_props_map_delete fp i: delete i (to_props_map fp) = to_props_map (delete i fp). Proof. by rewrite /to_props_map fmap_delete. Qed. (** ** Spectral operations *) Lemma flock_res_alloc_strong (X : gset lock_res_gname) γ lk R E : ↑N ⊆ E → is_flock γ lk -∗ ▷ R ={E}=∗ ∃ ρ, ⌜ρ ∉ X⌝ ∧ flock_res γ ρ 1 R. Proof. iIntros (?) "Hl HR". rewrite /is_flock. iDestruct "Hl" as "(#Hfl & #Hlk)". (* Pick a fresh name *) iInv invN as (s fa fp) "(>% & >Hstate & >Haprops & >Haactive & >Htokens & Hst)" "Hcl". (* Alloc all the data for the resource *) (* We need to guarantee that the new name - is not in the set X, - does not clash with the existing names in the invariant *) set (X₁ := collection_map flock_token1_name X : gset gname). set (X₂ := collection_map flock_token2_name (dom (gset lock_res_gname) (fp ∪ from_active fa)) : gset gname). iMod (own_alloc_strong (Excl ()) X₁) as (γt₁ Ht₁) "T1"; first done. iMod (own_alloc_strong (Excl ()) X₂) as (γt₂ Ht₂) "T2"; first done. iMod (cinv_alloc _ (resN.@(γt₁,γt₂)) (R ∗ own γt₁ (Excl ()) ∨ own γt₂ (Excl ()))%I with "[HR T1]") as (γc) "[#HR Hρ]". { iNext. iLeft. by iFrame. } set (ρ := {| flock_cinv_name := γc; flock_token1_name := γt₁; flock_token2_name := γt₂ |}). (* Put it in the map of flock resources *) iMod (own_update with "Haprops") as "Haprops". { eapply (auth_update_alloc _ (to_props_map (<[ρ := ()]> fp ∪ from_active fa)) {[ ρ := 1%Qp ]}). rewrite -insert_union_l. rewrite /to_props_map /= fmap_insert. apply alloc_local_update; last done. apply (not_elem_of_dom (to_props_map (fp ∪ from_active fa)) _ (D:=gset lock_res_gname)). rewrite /to_props_map dom_fmap. intros Hρ. apply Ht₂. by apply (elem_of_map_2 flock_token2_name _ ρ). } iDestruct "Haprops" as "[Haprops Hi]". iMod ("Hcl" with "[-Hi Hρ]") as "_". { iNext. iExists s, fa,(<[ρ:=()]>fp). iFrame. rewrite /all_tokens big_sepM_insert; last first. { apply (not_elem_of_dom _ ρ (D:=gset lock_res_gname)). intros Hρ. apply Ht₂. apply (elem_of_map_2 flock_token2_name _ ρ). rewrite dom_union_L elem_of_union //. eauto. } iFrame. iPureIntro. apply map_disjoint_insert_l_2; eauto. eapply (not_elem_of_dom (D:=gset lock_res_gname)). intros Hi. apply Ht₂. apply (elem_of_map_2 flock_token2_name _ ρ). rewrite dom_union_L elem_of_union. eauto. } iModIntro; iExists ρ; iSplit; last by iFrame. iPureIntro. intros Hρ. apply Ht₁. unfold X₁. by apply (elem_of_map_2 flock_token1_name X ρ). Qed. Lemma flock_res_dealloc γ lk ρ R E : ↑N ⊆ E → is_flock γ lk -∗ flock_res γ ρ 1 R ={E}=∗ ▷ R. Proof. iIntros (?) "Hl HR". rewrite /is_flock. iDestruct "Hl" as "(#Hfl & #Hlk)". iDestruct "HR" as "(Hρ & #Hiinv & Hcinv)". iInv invN as ([|] fa fp) "(>% & >Hstate & >Haprops & >Haactive & >Htokens & Hst)" "Hcl". (* Locked *) - iDestruct "Hst" as ">[Hlocked Hactives]". (* We can now show that the proposition `ρ` is *not* active *) iDestruct (own_valid_2 with "Haprops Hρ") as %[Hfoo%to_props_map_singleton_included _]%auth_valid_discrete_2. iAssert (⌜fa !! ρ = None⌝)%I with "[-]" as %Hbar. { remember (fa !! ρ) as faρ. destruct faρ as [π|]; last done. symmetry in Heqfaρ. iCombine "Hactives Hρ" as "H". iDestruct (own_valid with "H") as %Hbaz. exfalso. apply auth_own_valid in Hbaz. simpl in *. specialize (Hbaz ρ). revert Hbaz. rewrite lookup_op Heqfaρ lookup_singleton. rewrite -Some_op Some_valid. eapply exclusive_r ; eauto. apply _. } assert (fp !! ρ = Some ()) as Hbaz. { apply lookup_union_Some in Hfoo; last done. destruct Hfoo as [? | Hfoo]; first done. exfalso. revert Hfoo. by rewrite /from_active lookup_fmap Hbar. } iMod (own_update_2 with "Haprops Hρ") as "Haprops". { apply auth_update_dealloc, (delete_singleton_local_update _ ρ _). } rewrite /all_tokens (big_sepM_delete _ fp ρ) //. iDestruct "Htokens" as "[T2 Htokens]". iMod (cinv_cancel with "Hiinv Hcinv") as "HC". solve_ndisj. rewrite /C /=. iDestruct "HC" as "[[HR >T1] | >T2']"; last first. { iDestruct (own_valid_2 with "T2 T2'") as %?. done. } iFrame "HR". (* Now that the resource is cancelled we close the flock invariant *) iApply "Hcl". iNext. iExists Locked,fa,(delete ρ fp). iFrame. iSplit. + iPureIntro. solve_map_disjoint. + rewrite to_props_map_delete delete_union. rewrite (delete_notin (from_active fa)) //. rewrite /from_active lookup_fmap Hbar //. (* Unlocked *) - iDestruct "Hst" as ">Hfactive". iAssert (⌜fa = ∅⌝)%I with "[-]" as %->. { iDestruct (own_valid_2 with "Haactive Hfactive") as %[Hfoo%Excl_included _]%auth_valid_discrete_2. iPureIntro. by unfold_leibniz. } rewrite /from_active fmap_empty /= right_id. iDestruct (own_valid_2 with "Haprops Hρ") as %[Hfoo%to_props_map_singleton_included _]%auth_valid_discrete_2. iMod (own_update_2 with "Haprops Hρ") as "Haprops". { apply auth_update_dealloc, (delete_singleton_local_update _ ρ _). } rewrite /all_tokens (big_sepM_delete _ fp ρ _) //. iDestruct "Htokens" as "[T2 Htokens]". iMod (cinv_cancel with "Hiinv Hcinv") as "HC". solve_ndisj. rewrite /C /=. iDestruct "HC" as "[[HR >T1] | >T2']"; last first. { iDestruct (own_valid_2 with "T2 T2'") as %?. done. } iFrame "HR". (* Now that the resource is cancelled we close the flock invariant *) iApply "Hcl". iNext. iExists Unlocked,∅,(delete ρ fp). iFrame. iSplit. + iPureIntro. solve_map_disjoint. + rewrite /from_active fmap_empty /= right_id. by rewrite to_props_map_delete. Qed. (** `flocked` supports invariant access just like regular invariants *) Lemma flocked_inv_open_single E γ ρ πR I : ↑resN ⊆ E → I !! ρ = Some πR → flocked γ I ={E}=∗ ▷ πR.2 ∗ (▷ πR.2 ={E}=∗ flocked γ I). Proof. destruct πR as [π R]. iIntros (? Hρ). rewrite {1}/flocked. iIntros "(Hst & Hfa & Htokens)". rewrite (big_sepM_lookup_acc _ I ρ (π,R)) //. iDestruct "Htokens" as "[(T2 & #Hinv & Hρ) Htokens]". iMod (cinv_open with "Hinv Hρ") as "(HC & Hρ & Hcl)"; first solve_ndisj. rewrite /C /=. iDestruct "HC" as "[[HR >T1] | >T2']"; last first. { iDestruct (own_valid_2 with "T2 T2'") as %?. done. } iMod ("Hcl" with "[T2]") as "_". { iNext. iRight. done. } iModIntro. iFrame "HR". iIntros "HR". iMod (cinv_open with "Hinv Hρ") as "(HC & Hρ & Hcl)"; first solve_ndisj. iDestruct "HC" as "[[? >T1'] | >T2]". { iDestruct (own_valid_2 with "T1 T1'") as %?. done. } iMod ("Hcl" with "[T1 HR]") as "_". { iNext. iLeft. iFrame. } iModIntro. rewrite /flocked. iFrame "Hst Hfa". iApply "Htokens". by iFrame. Qed. Lemma flocked_inv_open E γ I : ↑resN ⊆ E → flocked γ I ={E}=∗ ▷ ([∗ map] ρ ↦ πR ∈ I, πR.2) ∗ (▷ ([∗ map] ρ ↦ πR ∈ I, πR.2) ={E}=∗ flocked γ I). Proof. iIntros (?). rewrite {1}/flocked. iIntros "(Hst & Hfa & Htokens)". rewrite !big_sepM_sepM. iDestruct "Htokens" as "(T2s & Hinvs & Hρs)". rewrite /flocked. iFrame "Hst Hfa". iInduction I as [|ρ πR I Hs] "IH" using map_ind. - rewrite /flocked !big_sepM_empty. eauto with iFrame. - rewrite !big_sepM_insert //. iDestruct "T2s" as "[T2 T2s]". iDestruct "Hinvs" as "[#Hinv Hinvs]". iDestruct "Hρs" as "[Hρ Hρs]". iMod ("IH" with "T2s Hinvs Hρs") as "[$ IH']". iClear "IH". (* The main induction step *) iMod (cinv_open with "Hinv Hρ") as "(HC & Hρ & Hcl)"; first solve_ndisj. rewrite /C /=. iDestruct "HC" as "[[HR >T1] | >T2']"; last first. { iDestruct (own_valid_2 with "T2 T2'") as %?. done. } iMod ("Hcl" with "[T2]") as "_". { iNext. iRight. done. } iModIntro. iFrame "HR". iIntros "[HR HRs]". iMod ("IH'" with "HRs") as "$". iMod (cinv_open with "Hinv Hρ") as "(HC & Hρ & Hcl)"; first solve_ndisj. iDestruct "HC" as "[[? >T1'] | >$]". { iDestruct (own_valid_2 with "T1 T1'") as %?. done. } iMod ("Hcl" with "[T1 HR]") as "_". { iNext. iLeft. iFrame. } iModIntro. by iFrame. Qed. (** ** Physical operations *) Lemma newflock_spec : {{{ True }}} newlock #() {{{ lk γ, RET lk; is_flock γ lk }}}. Proof. iIntros (Φ) "_ HΦ". rewrite -wp_fupd. iMod (own_alloc (● (Excl' Unlocked) ⋅ ◯ (Excl' Unlocked))) as (γ_state) "[Hauth Hfrag]"; first done. iMod (own_alloc (● to_props_map ∅)) as (γ_props) "Hprops". { apply auth_valid_discrete. simpl. split; last done. apply ucmra_unit_least. } iMod (own_alloc ((● Excl' ∅) ⋅ (◯ Excl' ∅))) as (γ_ac_props) "[Hacprops1 Hacprops2]". { apply auth_valid_discrete. simpl. split; last done. by rewrite left_id. } iApply (newlock_spec lockN (own γ_state (◯ (Excl' Unlocked))) with "Hfrag"). iNext. iIntros (lk γ_lock) "#Hlock". pose (γ := (Build_flock_name γ_lock γ_state γ_props γ_ac_props)). iMod (inv_alloc invN _ (flock_inv γ) with "[-HΦ]") as "#Hinv". { iNext. iExists Unlocked, ∅, ∅. rewrite /from_active fmap_empty right_id. iFrame. iSplit; eauto. by rewrite /all_tokens big_sepM_empty. } iModIntro. iApply ("HΦ" $! lk γ with "[-]"). rewrite /is_flock. by iFrame "Hlock". Qed. (** We need two additional lemmas, and they are annoying! *) Lemma big_sepM_own_frag {A B : Type} {C} `{EqDecision A, Countable A} `{inG Σ (authR (gmapUR A C))} (f : B → C) (m : gmap A B) (γ : gname) : own γ (◯ ∅) -∗ own γ (◯ (f <$> m)) ∗-∗ [∗ map] i↦x ∈ m, own γ (◯ {[ i := f x ]}). Proof. simple refine (map_ind (λ m, _)%I _ _ m); simpl. - iIntros "He". rewrite fmap_empty big_sepM_empty. iSplit; eauto. - iIntros (i x m' Hi IH) "He". rewrite fmap_insert insert_union_singleton_l. assert (({[i := f x]} ∪ (f <$> m')) = {[i := f x]} ⋅ (f <$> m')) as ->. { rewrite /op /cmra_op /= /gmap_op. apply map_eq. intros j. destruct (decide (i = j)) as [->|?]. - etransitivity. eapply lookup_union_Some_l. apply lookup_insert. symmetry. rewrite lookup_merge lookup_insert. rewrite lookup_fmap Hi /=. done. - remember (m' !! j) as mj. destruct mj; simplify_eq/=. + etransitivity. apply lookup_union_Some_raw. right. split; first by rewrite lookup_insert_ne. by rewrite lookup_fmap -Heqmj. symmetry. rewrite lookup_merge lookup_singleton_ne; eauto. rewrite lookup_fmap -Heqmj. done. + etransitivity. apply lookup_union_None. split; first by rewrite lookup_singleton_ne. rewrite lookup_fmap -Heqmj //. symmetry. rewrite lookup_merge lookup_singleton_ne // lookup_fmap -Heqmj //. } rewrite auth_frag_op own_op IH big_sepM_insert; last eauto. iSplit; iIntros "[$ Hm']"; by iApply "He". Qed. Lemma own_frag_empty γ X : own (flock_props_name γ) (● X) ==∗ own (flock_props_name γ) (● X) ∗ own (flock_props_name γ) (◯ ∅). Proof. iIntros "H". rewrite -own_op. iApply (own_update with "H"). by apply auth_update_alloc. Qed. Lemma release_cancel_spec γ lk I : {{{ is_flock γ lk ∗ flocked γ I }}} release lk {{{ RET #(); [∗ map] ρ ↦ πR ∈ I, flock_res γ ρ πR.1 πR.2 }}}. Proof. iIntros (Φ) "(#Hl & H) HΦ". rewrite -fupd_wp. rewrite /is_flock. iDestruct "Hl" as "(#Hinv & #Hlk)". rewrite {1}/flocked /=. iDestruct "H" as "(Hflkd & Hfactive & Hfa)". do 2 rewrite big_sepM_sepM. iDestruct "Hfa" as "(HT2s & #Hinvs & Hρs)". iInv invN as ([|] fa fp) "(>% & >Hstate & >Haprops & >Haactive & >Htokens & Hst)" "Hcl"; last first. - iDestruct (own_valid_2 with "Hstate Hflkd") as %[Hfoo%Excl_included _]%auth_valid_discrete_2. fold_leibniz. inversion Hfoo. - iDestruct "Hst" as ">[Hlocked Hactives]". iDestruct (own_valid_2 with "Haactive Hfactive") as %[Hfoo%Excl_included _]%auth_valid_discrete_2. fold_leibniz. simplify_eq/=. (* Locked ~~> Unlocked *) iMod (own_update_2 with "Hstate Hflkd") as "Hstate". { apply (auth_update _ _ (Excl' Unlocked) (Excl' Unlocked)). apply option_local_update. by apply exclusive_local_update. } iDestruct "Hstate" as "[Hstate Hunflkd]". (* Empty up the set of active propositions *) iMod (own_update_2 with "Haactive Hfactive") as "[Haactive Hfactive]". { apply (auth_update _ _ (Excl' ∅) (Excl' ∅)). by apply option_local_update, exclusive_local_update. } pose (fa := fst <$> I). iMod (own_frag_empty with "Haprops") as "[Haprops Hemp]". (* You may wonder, why the hell do we need this useless Hemp proposition? .. Well you will find out towards the end of the proof *) iMod ("Hcl" with "[-HΦ Hlocked Hactives Hunflkd Hρs Hemp]") as "_". { iNext. iExists Unlocked,∅,(fp ∪ from_active fa). iSplitR; eauto. - rewrite /from_active fmap_empty. iPureIntro. solve_map_disjoint. - iFrame. rewrite /from_active fmap_empty right_id /=. iFrame "Haprops". rewrite /all_tokens. rewrite big_sepM_union // -map_fmap_compose. rewrite big_sepM_fmap. by iFrame. } iApply (release_spec with "[$Hlk $Hlocked $Hunflkd]"). iModIntro. iNext. iIntros "_". iApply "HΦ". rewrite /flock_res !big_sepM_sepM. iFrame "Hρs Hinvs". by iApply (big_sepM_own_frag with "Hemp"). Qed. Lemma acquire_flock_spec γ lk (I : gmap lock_res_gname (frac*iProp Σ)) : {{{ is_flock γ lk ∗ [∗ map] ρ ↦ πR ∈ I, flock_res γ ρ πR.1 πR.2 }}} acquire lk {{{ RET #(); flocked γ I }}}. Proof. iIntros (Φ) "(Hl & HRres) HΦ". rewrite /is_flock. iDestruct "Hl" as "(#Hfl & #Hlk)". iApply wp_fupd. iApply (acquire_spec with "Hlk"). iNext. iIntros "[Hlocked Hunlk]". iInv invN as ([|] fa fp) "(>% & >Hstate & >Haprops & >Haactive & >Htokens & Hst)" "Hcl". - iDestruct "Hst" as "(>Hlocked2 & ?)". iExFalso. iApply (locked_exclusive with "Hlocked Hlocked2"). - iDestruct "Hst" as ">Hfactive". iAssert (⌜fa = ∅⌝)%I with "[-]" as %->. { iDestruct (own_valid_2 with "Haactive Hfactive") as %[Hfoo%Excl_included _]%auth_valid_discrete_2. iPureIntro. by unfold_leibniz. } rewrite /from_active fmap_empty /= right_id. (* Unlocked ~~> Locked *) iMod (own_update_2 with "Hstate Hunlk") as "Hstate". { apply (auth_update _ _ (Excl' Locked) (Excl' Locked)). apply option_local_update. by apply exclusive_local_update. } iDestruct "Hstate" as "[Hstate Hflkd]". iApply "HΦ". rewrite /flocked. iFrame "Hflkd". (* Designate the propositions from I as active *) pose (fa := fst <$> I). iMod (own_update_2 with "Haactive Hfactive") as "[Haactive Hfactive]". { apply (auth_update _ _ (Excl' fa) (Excl' fa)). by apply option_local_update, exclusive_local_update. } iFrame "Hfactive". rewrite /flock_res. rewrite !big_sepM_sepM. iDestruct "HRres" as "(HI & #Hinvs & Hρs)". iFrame "Hinvs Hρs". (* this is going to be annoying .. *) (* show that I ⊆ fp, or, better fp = fp' ∪ I *) (* first obtain the empty fragment for big_sepM_own_frag *) iMod (own_frag_empty with "Haprops") as "[Haprops #Hemp]". iAssert (own (flock_props_name γ) (◯ fa)) with "[HI Hemp]" as "HI". { by iApply (big_sepM_own_frag with "Hemp"). } (* I' ≼ fp *) iDestruct (own_valid_2 with "Haprops HI") as %[Hfoo _]%auth_valid_discrete_2. (* TODO: RK, can this proof script be simplified? The idea is not complicated. *) (* We are going to separate the active resources I out of the fp map *) pose (I'' := fmap (const ()) I). assert (I'' ⊆ fp). { subst I'' fa. unfold to_props_map in *. pose (f := fst : frac*(iProp Σ) → frac). pose (g := const 1%Qp : () → frac). assert (const () <$> (f <$> I) ≼ (const () <$> (g <$> fp) : gmapUR lock_res_gname unitR)) as Hbar. { by apply gmap_fmap_mono. } move: Hbar. rewrite -!map_fmap_compose=>Hbar. (* TODO: WHY can't I do apply (lookup_included (fst <$> I) (const 1%Qp <$> fp)) in Hbar *) assert (∀ x, (const () ∘ f <$> I) !! x ≼ (const () ∘ g <$> fp) !! x) as Hbaz. { by apply lookup_included. } apply map_subseteq_spec. intros ρ []. specialize (Hbaz ρ). revert Hbaz. rewrite !lookup_fmap. intros Hbaz Hρ. rewrite Hρ in Hbaz. destruct (fp !! ρ) as [[]|] eqn:Hfp; auto. simpl in Hbaz. exfalso. destruct Hbaz as [[[]|] Hbaz]; fold_leibniz; inversion Hbaz. } pose (P' := fp ∖ I''). assert (I'' ##ₘ P'). { by apply map_disjoint_difference_r. } assert (fp = P' ∪ I'') as HP'. { symmetry. rewrite map_union_comm//. by apply map_difference_union. } rewrite HP'. rewrite /all_tokens big_sepM_union //. iDestruct "Htokens" as "[Htokens H]". rewrite /I'' big_sepM_fmap. iFrame. iApply "Hcl". iNext. iExists Locked, fa, P'. assert (I'' = from_active fa) as <-. { by rewrite /from_active -map_fmap_compose. } eauto with iFrame. Qed. End flock.