Skip to content
Snippets Groups Projects
invariants.v 2.83 KiB
Newer Older
From algebra Require Export base.
From program_logic Require Import ownership.
From program_logic Require Export namespaces pviewshifts weakestpre.
Import uPred.
Local Hint Extern 100 (@eq coPset _ _) => solve_elem_of.
Local Hint Extern 100 (@subseteq coPset _ _) => solve_elem_of.
Local Hint Extern 100 (_  _) => solve_elem_of.
Local Hint Extern 99 ({[ _ ]}  _) => apply elem_of_subseteq_singleton.
Robbert Krebbers's avatar
Robbert Krebbers committed

(** Derived forms and lemmas about them. *)
Definition inv {Λ Σ} (N : namespace) (P : iProp Λ Σ) : iProp Λ Σ :=
  ( i,  (i  nclose N)  ownI i P)%I.
Instance: Params (@inv) 3.
Typeclasses Opaque inv.

Section inv.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types i : positive.
Implicit Types N : namespace.
Implicit Types P Q R : iProp Λ Σ.

Global Instance inv_contractive N : Contractive (@inv Λ Σ N).
Proof. intros n ???. apply exists_ne=>i. by apply and_ne, ownI_contractive. Qed.
Global Instance inv_always_stable N P : AlwaysStable (inv N P).
Proof. rewrite /inv; apply _. Qed.

Lemma always_inv N P : ( inv N P)%I  inv N P.
Proof. by rewrite always_always. Qed.

(** Invariants can be opened around any frame-shifting assertion. *)
Lemma inv_fsa {A} (fsa : FSA Λ Σ A) `{!FrameShiftAssertion fsaV fsa}
    E N P (Q : A  iProp Λ Σ) R :
  fsaV 
  nclose N  E 
  R  ( P -★ fsa (E  nclose N) (λ a,  P  Q a)) 
  R  fsa E Q.
  intros ? HN Hinv Hinner.
  rewrite -[R](idemp ()%I) {1}Hinv Hinner =>{Hinv Hinner R}.
  rewrite always_and_sep_l /inv sep_exist_r. apply exist_elim=>i.
  rewrite always_and_sep_l -assoc. apply const_elim_sep_l=>HiN.
  rewrite -(fsa_open_close E (E  {[encode i]})) //; last by solve_elem_of+.
  (* Add this to the local context, so that solve_elem_of finds it. *)
  assert ({[encode i]}  nclose N) by eauto.
  rewrite (always_sep_dup (ownI _ _)).
  rewrite {1}pvs_openI !pvs_frame_r.
  apply pvs_mask_frame_mono; [solve_elem_of..|].
  rewrite (comm _ (▷_)%I) -assoc wand_elim_r fsa_frame_l.
  apply fsa_mask_frame_mono; [solve_elem_of..|]. intros a.
  rewrite assoc -always_and_sep_l pvs_closeI pvs_frame_r left_id.
  apply pvs_mask_frame'; solve_elem_of.
Qed.

(* Derive the concrete forms for pvs and wp, because they are useful. *)

  R  inv N P 
  R  (P -★ pvs (E  nclose N) (E  nclose N) (P  Q)) 
  R  pvs E E Q.
Ralf Jung's avatar
Ralf Jung committed
Proof. intros. by apply: (inv_fsa pvs_fsa). Qed.
Lemma wp_open_close E e N P (Q : val Λ  iProp Λ Σ) R :
Ralf Jung's avatar
Ralf Jung committed
  atomic e  nclose N  E 
  R  ( P -★ wp (E  nclose N) e (λ v, P  Q v)) 
Ralf Jung's avatar
Ralf Jung committed
Proof. intros. by apply: (inv_fsa (wp_fsa e)). Qed.
Lemma inv_alloc N P :  P  pvs N N (inv N P).
Proof. by rewrite /inv (pvs_allocI N); last apply coPset_suffixes_infinite. Qed.