diff --git a/theories/base_logic/lib/proph_map.v b/theories/base_logic/lib/proph_map.v index f4ead7f01b8ea0e92e3045a9f5a72efaf796eef9..1158cd502259be76be9d0af2917ae52c5ac74b7c 100644 --- a/theories/base_logic/lib/proph_map.v +++ b/theories/base_logic/lib/proph_map.v @@ -134,10 +134,8 @@ Section proph_map. rewrite proph_eq /proph_def. iIntros "Hp1 Hp2". iCombine "Hp1 Hp2" as "Hp". iDestruct (own_valid with "Hp") as %Hp. - (* FIXME: [apply auth_frag_valid in Hp] and - [rewrite ->auth_frag_valid in Hp] both should work but do not. *) - move:Hp. rewrite auth_frag_valid singleton_valid=>Hp. - done. + (* FIXME: FIXME(Coq #6294): needs new unification *) + move:Hp. rewrite auth_frag_valid singleton_valid //. Qed. Lemma proph_map_new_proph p ps pvs : @@ -169,7 +167,8 @@ Section proph_map. assert (vs = v :: proph_list_resolves pvs p) as ->. { rewrite (Hres p vs HR). simpl. by rewrite decide_True. } iMod (own_update_2 with "H◠Hp") as "[H◠H◯]". - { eapply auth_update. apply: singleton_local_update. + { (* FIXME: FIXME(Coq #6294): needs new unification *) + eapply auth_update. apply: singleton_local_update. - by rewrite /to_proph_map lookup_fmap HR. - by apply (exclusive_local_update _ (Excl (proph_list_resolves pvs p : list (leibnizC V)))). } rewrite /to_proph_map -fmap_insert. diff --git a/theories/heap_lang/lib/counter.v b/theories/heap_lang/lib/counter.v index cde2c0b41a6e2354cebc9610193d42884d74db07..3d46fbd6e5d1a2457353e6c340eb03feaebd9c3a 100644 --- a/theories/heap_lang/lib/counter.v +++ b/theories/heap_lang/lib/counter.v @@ -60,8 +60,9 @@ Section mono_proof. wp_cas_suc. iModIntro. iSplitL "Hl Hγ". { iNext. iExists (S c). rewrite Nat2Z.inj_succ Z.add_1_l. by iFrame. } wp_if. iApply "HΦ"; iExists γ; repeat iSplit; eauto. - iApply (own_mono with "Hγf"). apply: auth_frag_mono. - by apply mnat_included, le_n_S. + iApply (own_mono with "Hγf"). + (* FIXME: FIXME(Coq #6294): needs new unification *) + apply: auth_frag_mono. by apply mnat_included, le_n_S. - wp_cas_fail; first (by intros [= ?%Nat2Z.inj]). iModIntro. iSplitL "Hl Hγ"; [iNext; iExists c'; by iFrame|]. wp_if. iApply ("IH" with "[Hγf] [HΦ]"); last by auto. diff --git a/theories/program_logic/ownp.v b/theories/program_logic/ownp.v index 4b4b4a9a20292fc6b99b4da7e912cb5f0dee4222..c57af7bba3ca6202ef6cfddec5a0b6d710803990 100644 --- a/theories/program_logic/ownp.v +++ b/theories/program_logic/ownp.v @@ -118,8 +118,8 @@ Section lifting. iModIntro; iSplit; [by destruct s|]; iNext; iIntros (e2 σ2 efs Hstep). iDestruct "Hσκs" as "Hσ". rewrite /ownP. 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)). } iFrame "Hσ". iApply ("H" with "[]"); eauto with iFrame. Qed.