diff --git a/theories/heap_lang/lib/coin_flip.v b/theories/heap_lang/lib/coin_flip.v index 2bdbab5cee53f5a5a47ed09c7009f172c95446a0..804f25440dbf1b362132fd1ded5e5d9052e6e46d 100644 --- a/theories/heap_lang/lib/coin_flip.v +++ b/theories/heap_lang/lib/coin_flip.v @@ -76,7 +76,7 @@ Section coinflip. Proof using N. iApply wp_atomic_intro. iIntros (Φ) "AU". wp_lam. wp_apply wp_new_proph; first done. - iIntros (p) "Hp". iDestruct "Hp" as (v) "Hp". + iIntros (v p) "Hp". wp_let. wp_bind (_ <- _)%E. iMod "AU" as "[Hl [_ Hclose]]". diff --git a/theories/heap_lang/lifting.v b/theories/heap_lang/lifting.v index a51849155adcf3ede97dc06cfe3fd23f29230316..835ee9bc8f9a7e04d00bdc22b042840fcd76e17d 100644 --- a/theories/heap_lang/lifting.v +++ b/theories/heap_lang/lifting.v @@ -289,7 +289,7 @@ Qed. (** Lifting lemmas for creating and resolving prophecy variables *) Lemma wp_new_proph : - {{{ True }}} NewProph {{{ (p : proph), RET (LitV (LitProphecy p)); p ⥱ - }}}. + {{{ True }}} NewProph {{{ v (p : proph), RET (LitV (LitProphecy p)); p ⥱ v }}}. Proof. iIntros (Φ) "_ HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1 κs) "[Hσ HR] !>". iDestruct "HR" as (R [Hfr Hdom]) "HR". iSplit. @@ -304,7 +304,7 @@ Proof. iPureIntro. split. + apply first_resolve_insert; auto. + rewrite dom_insert_L. by apply union_mono_l. - - iApply "HΦ". by iExists _. + - iApply "HΦ". done. Qed. Lemma wp_resolve_proph e1 e2 p v w: @@ -320,7 +320,7 @@ Proof. unfold cons_obs. simpl. iNext; iIntros (κ κs' v2 σ2 efs [Hstep ->]); inv_head_step. iApply fupd_frame_l. iSplit=> //. iFrame. - iMod ((@proph_map_remove _ _ _ _ _ _ _ p0) with "HR Hp") as "Hp". iModIntro. + iMod (@proph_map_remove with "HR Hp") as "Hp". iModIntro. iSplitR "HΦ". - iExists _. iFrame. iPureIntro. split; first by eapply first_resolve_delete. rewrite dom_delete. rewrite <- difference_empty_L. by eapply difference_mono. diff --git a/theories/heap_lang/proph_map.v b/theories/heap_lang/proph_map.v index fab888a471158782617c124d1d0eb2595e4f4a24..3a32f3487a54b9fd602486d02429a2715a5db968 100644 --- a/theories/heap_lang/proph_map.v +++ b/theories/heap_lang/proph_map.v @@ -20,7 +20,7 @@ Section first_resolve. (map_of_list pvs : gmap P V) !! p. Definition first_resolve_in_list R pvs := - forall p v, p ∈ dom (gset _) R → + ∀ p v, p ∈ dom (gset _) R → first_resolve pvs p = Some v → R !! p = Some (Some v). @@ -156,13 +156,11 @@ Section proph_map. p ∉ dom (gset _) R → proph_map_auth R ==∗ proph_map_auth (<[p := v]> R) ∗ p ⥱ v. Proof. - iIntros (?) "HR". rewrite /proph_map_ctx p_mapsto_eq /p_mapsto_def. + iIntros (Hp) "HR". rewrite /proph_map_ctx p_mapsto_eq /p_mapsto_def. iMod (own_update with "HR") as "[HR Hl]". - { - eapply auth_update_alloc, + { eapply auth_update_alloc, (alloc_singleton_local_update _ _ (Excl $ (v : option (leibnizC _))))=> //. - apply lookup_to_proph_map_None. apply (iffLR (not_elem_of_dom _ _) H1). - } + apply lookup_to_proph_map_None. apply (iffLR (not_elem_of_dom _ _) Hp). } iModIntro. rewrite /proph_map_auth to_proph_map_insert. iFrame. Qed. diff --git a/theories/program_logic/language.v b/theories/program_logic/language.v index 7b284d5b2e97a20aefa738c70243d8bd7c798c98..102c381fce69b6cb81fe38642226738f0459eb0f 100644 --- a/theories/program_logic/language.v +++ b/theories/program_logic/language.v @@ -103,6 +103,7 @@ Section language. Ï2 = (t1 ++ e2 :: t2 ++ efs, σ2) → prim_step e1 σ1 κ e2 σ2 efs → step Ï1 κ Ï2. + Hint Constructors step. (* TODO (MR) introduce notation ::? for cons_obs and suggest for inclusion to stdpp? *) Definition cons_obs {A} (x : option A) (xs : list A) := @@ -115,10 +116,9 @@ Section language. step Ï1 κ Ï2 → nsteps n Ï2 κs Ï3 → nsteps (S n) Ï1 (cons_obs κ κs) Ï3. + Hint Constructors nsteps. - Definition erased_step (Ï1 Ï2 : cfg Λ) := exists κ, step Ï1 κ Ï2. - - Hint Constructors step nsteps. + Definition erased_step (Ï1 Ï2 : cfg Λ) := ∃ κ, step Ï1 κ Ï2. Lemma erased_steps_nsteps Ï1 Ï2 : rtc erased_step Ï1 Ï2 → diff --git a/theories/program_logic/ownp.v b/theories/program_logic/ownp.v index a1bc373db0c7b826699ae81ebbee265d26a49766..0038df660701d6effaf44774d9fc66d958ace0ce 100644 --- a/theories/program_logic/ownp.v +++ b/theories/program_logic/ownp.v @@ -134,8 +134,8 @@ Section lifting. iDestruct "Hσκs" as "[Hσ Hκs]". rewrite /ownP_state /ownP_obs. iMod (own_update_2 with "Hσ Hσf") as "[Hσ Hσf]". - { apply auth_update. apply: option_local_update. - by apply: (exclusive_local_update _ (Excl σ2)). } + { apply auth_update. apply: option_local_update. + by apply: (exclusive_local_update _ (Excl σ2)). } iMod (own_update_2 with "Hκs Hκsf") as "[Hκs Hκsf]". { apply auth_update. apply: option_local_update. by apply: (exclusive_local_update _ (Excl (κs'':leibnizC _))). } diff --git a/theories/program_logic/weakestpre.v b/theories/program_logic/weakestpre.v index 7370591dba8ec647a52825c4856d19c8da07fffb..dd29e8dbdff982134ee51d0cb5850271fbf4ef52 100644 --- a/theories/program_logic/weakestpre.v +++ b/theories/program_logic/weakestpre.v @@ -58,7 +58,7 @@ Proof. (* FIXME: reflexivity, as being called many times by f_equiv and f_contractive is very slow here *) do 24 (f_contractive || f_equiv). apply IH; first lia. - intros v. eapply dist_le; eauto with omega. + intros v. eapply dist_le; eauto with lia. Qed. Global Instance wp_proper s E e : Proper (pointwise_relation _ (≡) ==> (≡)) (wp (PROP:=iProp Σ) s E e).