Skip to content
Snippets Groups Projects
pviewshifts.v 7.73 KiB
Newer Older
From iris.program_logic Require Export iris.
From iris.program_logic Require Import wsat.
From iris.algebra Require Import upred_big_op gmap.
From iris.proofmode Require Import tactics classes.
Ralf Jung's avatar
Ralf Jung committed

Program Definition pvs_def `{irisG Λ Σ}
    (E1 E2 : coPset) (P : iProp Σ) : iProp Σ :=
  (wsat  ownE E1 =r=★  (wsat  ownE E2  P))%I.
Ralf Jung's avatar
Ralf Jung committed
Definition pvs_aux : { x | x = @pvs_def }. by eexists. Qed.
Definition pvs := proj1_sig pvs_aux.
Definition pvs_eq : @pvs = @pvs_def := proj2_sig pvs_aux.
Arguments pvs {_ _ _} _ _ _%I.
Instance: Params (@pvs) 5.
Notation "|={ E1 , E2 }=> Q" := (pvs E1 E2 Q)
  (at level 99, E1, E2 at level 50, Q at level 200,
   format "|={ E1 , E2 }=>  Q") : uPred_scope.
Notation "P ={ E1 , E2 }=★ Q" := (P -★ |={E1,E2}=> Q)%I
  (at level 99, E1,E2 at level 50, Q at level 200,
   format "P  ={ E1 , E2 }=★  Q") : uPred_scope.
Notation "P ={ E1 , E2 }=> Q" := (P  |={E1,E2}=> Q)
  (at level 99, E1, E2 at level 50, Q at level 200, only parsing) : C_scope.

Notation "|={ E }=> Q" := (pvs E E Q)
  (at level 99, E at level 50, Q at level 200,
   format "|={ E }=>  Q") : uPred_scope.
Notation "P ={ E }=★ Q" := (P -★ |={E}=> Q)%I
  (at level 99, E at level 50, Q at level 200,
   format "P  ={ E }=★  Q") : uPred_scope.
Notation "P ={ E }=> Q" := (P  |={E}=> Q)
  (at level 99, E at level 50, Q at level 200, only parsing) : C_scope.
Notation "|={ E1 , E2 }▷=> Q" := (|={E1,E2}=> ( |={E2,E1}=> Q))%I
  (at level 99, E1, E2 at level 50, Q at level 200,
   format "|={ E1 , E2 }▷=>  Q") : uPred_scope.
Notation "|={ E }▷=> Q" := (|={E,E}▷=> Q)%I
  (at level 99, E at level 50, Q at level 200,
   format "|={ E }▷=>  Q") : uPred_scope.
Section pvs.
Context `{irisG Λ Σ}.
Implicit Types P Q : iProp Σ.
Global Instance pvs_ne E1 E2 n : Proper (dist n ==> dist n) (@pvs Λ Σ _ E1 E2).
Proof. rewrite pvs_eq. solve_proper. Qed.
Global Instance pvs_proper E1 E2 : Proper (() ==> ()) (@pvs Λ Σ _ E1 E2).
Proof. apply ne_proper, _. Qed.

Lemma pvs_intro_mask E1 E2 P : E2  E1  P  |={E1,E2}=> |={E2,E1}=> P.
Proof.
  intros (E1''&->&?)%subseteq_disjoint_union_L.
  rewrite pvs_eq /pvs_def ownE_op //. iIntros "H ($ & $ & HE) !==>".
  iApply except_last_intro. iIntros "[$ $] !==>" . iApply except_last_intro.
  by iFrame.
Qed.
Lemma except_last_pvs E1 E2 P :  (|={E1,E2}=> P) ={E1,E2}=> P.
Proof.
  rewrite pvs_eq. iIntros "H [Hw HE]". iTimeless "H". iApply "H"; by iFrame.
Qed.
Lemma rvs_pvs E P : (|=r=> P) ={E}=> P.
Proof.
  rewrite pvs_eq /pvs_def. iIntros "H [$ $]"; iVs "H".
  iVsIntro. by iApply except_last_intro.
Qed.
Lemma pvs_mono E1 E2 P Q : (P  Q)  (|={E1,E2}=> P) ={E1,E2}=> Q.
Proof.
  rewrite pvs_eq /pvs_def. iIntros (HPQ) "HP HwE".
  rewrite -HPQ. by iApply "HP".
Qed.
Lemma pvs_trans E1 E2 E3 P : (|={E1,E2}=> |={E2,E3}=> P) ={E1,E3}=> P.
Proof.
  rewrite pvs_eq /pvs_def. iIntros "HP HwE".
  iVs ("HP" with "HwE") as ">(Hw & HE & HP)". iApply "HP"; by iFrame.
Qed.
Lemma pvs_mask_frame_r' E1 E2 Ef P :
  E1  Ef  (|={E1,E2}=> E2  Ef  P) ={E1  Ef,E2  Ef}=> P.
Proof.
  intros. rewrite pvs_eq /pvs_def ownE_op //. iIntros "Hvs (Hw & HE1 &HEf)".
  iVs ("Hvs" with "[Hw HE1]") as ">($ & HE2 & HP)"; first by iFrame.
  iDestruct (ownE_op' with "[HE2 HEf]") as "[? $]"; first by iFrame.
  iVsIntro; iApply except_last_intro. by iApply "HP".
Qed.
Lemma pvs_frame_r E1 E2 P Q : (|={E1,E2}=> P)  Q ={E1,E2}=> P  Q.
Proof. rewrite pvs_eq /pvs_def. by iIntros "[HwP $]". Qed.
(** * Derived rules *)
Global Instance pvs_mono' E1 E2 : Proper (() ==> ()) (@pvs Λ Σ _ E1 E2).
Proof. intros P Q; apply pvs_mono. Qed.
Global Instance pvs_flip_mono' E1 E2 :
  Proper (flip () ==> flip ()) (@pvs Λ Σ _ E1 E2).
Proof. intros P Q; apply pvs_mono. Qed.

Lemma pvs_intro E P : P ={E}=> P.
Proof. iIntros "HP". by iApply rvs_pvs. Qed.
Lemma pvs_intro_mask' E1 E2 : E2  E1  True  |={E1,E2}=> |={E2,E1}=> True.
Proof. exact: pvs_intro_mask. Qed.
Lemma pvs_except_last E1 E2 P : (|={E1,E2}=>  P) ={E1,E2}=> P.
Proof. by rewrite {1}(pvs_intro E2 P) except_last_pvs pvs_trans. Qed.
Lemma pvs_frame_l E1 E2 P Q : (P  |={E1,E2}=> Q) ={E1,E2}=> P  Q.
Proof. rewrite !(comm _ P); apply pvs_frame_r. Qed.
Lemma pvs_wand_l E1 E2 P Q : (P -★ Q)  (|={E1,E2}=> P) ={E1,E2}=> Q.
Proof. by rewrite pvs_frame_l wand_elim_l. Qed.
Lemma pvs_wand_r E1 E2 P Q : (|={E1,E2}=> P)  (P -★ Q) ={E1,E2}=> Q.
Proof. by rewrite pvs_frame_r wand_elim_r. Qed.

Lemma pvs_trans_frame E1 E2 E3 P Q :
  ((Q ={E2,E3}=★ True)  |={E1,E2}=> (Q  P)) ={E1,E3}=> P.
Proof.
  rewrite pvs_frame_l assoc -(comm _ Q) wand_elim_r.
  by rewrite pvs_frame_r left_id pvs_trans.
Qed.

Lemma pvs_mask_frame_r E1 E2 Ef P :
  E1  Ef  (|={E1,E2}=> P) ={E1  Ef,E2  Ef}=> P.
Proof.
  iIntros (?) "H". iApply pvs_mask_frame_r'; auto.
  iApply pvs_wand_r; iFrame "H"; eauto.
Qed.
Lemma pvs_mask_mono E1 E2 P : E1  E2  (|={E1}=> P) ={E2}=> P.
Proof.
  intros (Ef&->&?)%subseteq_disjoint_union_L. by apply pvs_mask_frame_r.
Qed.

Lemma pvs_sep E P Q : (|={E}=> P)  (|={E}=> Q) ={E}=> P  Q.
Proof. by rewrite pvs_frame_r pvs_frame_l pvs_trans. Qed.
Lemma pvs_big_sepM `{Countable K} {A} E (Φ : K  A  iProp Σ) (m : gmap K A) :
  ([ map] kx  m, |={E}=> Φ k x) ={E}=> [ map] kx  m, Φ k x.
Proof.
  apply (big_opM_forall (λ P Q, P ={E}=> Q)); auto using pvs_intro.
  intros P1 P2 HP Q1 Q2 HQ. by rewrite HP HQ -pvs_sep.
Lemma pvs_big_sepS `{Countable A} E (Φ : A  iProp Σ) X :
  ([ set] x  X, |={E}=> Φ x) ={E}=> [ set] x  X, Φ x.
Proof.
  apply (big_opS_forall (λ P Q, P ={E}=> Q)); auto using pvs_intro.
  intros P1 P2 HP Q1 Q2 HQ. by rewrite HP HQ -pvs_sep.
End pvs.

(** Proofmode class instances *)
Section proofmode_classes.
  Context `{irisG Λ Σ}.
  Implicit Types P Q : iProp Σ.

  Global Instance from_pure_pvs E P φ : FromPure P φ  FromPure (|={E}=> P) φ.
  Proof. rewrite /FromPure. intros <-. apply pvs_intro. Qed.

  Global Instance from_assumption_pvs E p P Q :
    FromAssumption p P (|=r=> Q)  FromAssumption p P (|={E}=> Q)%I.
  Proof. rewrite /FromAssumption=>->. apply rvs_pvs. Qed.

  Global Instance into_wand_pvs E1 E2 R P Q :
    IntoWand R P Q  IntoWand R (|={E1,E2}=> P) (|={E1,E2}=> Q) | 100.
  Proof. rewrite /IntoWand=>->. apply wand_intro_l. by rewrite pvs_wand_r. Qed.

  Global Instance from_sep_pvs E P Q1 Q2 :
    FromSep P Q1 Q2  FromSep (|={E}=> P) (|={E}=> Q1) (|={E}=> Q2).
  Proof. rewrite /FromSep=><-. apply pvs_sep. Qed.

  Global Instance or_split_pvs E1 E2 P Q1 Q2 :
    FromOr P Q1 Q2  FromOr (|={E1,E2}=> P) (|={E1,E2}=> Q1) (|={E1,E2}=> Q2).
  Proof. rewrite /FromOr=><-. apply or_elim; apply pvs_mono; auto with I. Qed.

  Global Instance exists_split_pvs {A} E1 E2 P (Φ : A  iProp Σ) :
    FromExist P Φ  FromExist (|={E1,E2}=> P) (λ a, |={E1,E2}=> Φ a)%I.
  Proof.
    rewrite /FromExist=><-. apply exist_elim=> a. by rewrite -(exist_intro a).
  Qed.

  Global Instance frame_pvs E1 E2 R P Q :
    Frame R P Q  Frame R (|={E1,E2}=> P) (|={E1,E2}=> Q).
  Proof. rewrite /Frame=><-. by rewrite pvs_frame_l. Qed.

  Global Instance is_except_last_pvs E1 E2 P : IsExceptLast (|={E1,E2}=> P).
  Proof. by rewrite /IsExceptLast except_last_pvs. Qed.

  Global Instance from_vs_pvs E P : FromVs (|={E}=> P) P.
  Proof. by rewrite /FromVs -rvs_pvs. Qed.

  Global Instance elim_vs_rvs_pvs E1 E2 P Q :
    ElimVs (|=r=> P) P (|={E1,E2}=> Q) (|={E1,E2}=> Q).
  Proof. by rewrite /ElimVs (rvs_pvs E1) pvs_frame_r wand_elim_r pvs_trans. Qed.
  Global Instance elim_vs_pvs_pvs E1 E2 E3 P Q :
    ElimVs (|={E1,E2}=> P) P (|={E1,E3}=> Q) (|={E2,E3}=> Q).
  Proof. by rewrite /ElimVs pvs_frame_r wand_elim_r pvs_trans. Qed.
End proofmode_classes.

Hint Extern 2 (coq_tactics.of_envs _  _) =>
  match goal with |- _  |={_}=> _ => iVsIntro end.