Skip to content
Snippets Groups Projects
core.v 1.54 KiB
Newer Older
From iris.base_logic Require Import base_logic.
From iris.proofmode Require Import tactics.
Set Default Proof Using "Type".
Import uPred.

(** The "core" of an assertion is its maximal persistent part.
    It can be defined entirely within the logic... at least
    in the shallow embedding.
    WARNING: The function "coreP" is NOT NON-EXPANSIVE.
    This is because the turnstile is not non-expansive as a function
    from iProp to (discreteC Prop).
    To obtain a core that's non-expansive, we would have to add another
    modality to the logic: a box that removes access to *all* resources,
    not just restricts access to the core.
*)

Definition coreP {M : ucmraT} (P : uPred M) : uPred M :=
  ( `(!PersistentP Q), P  Q  Q)%I.
Robbert Krebbers's avatar
Robbert Krebbers committed
Instance: Params (@coreP) 1.
Typeclasses Opaque coreP.

Section core.
  Context {M : ucmraT}.
  Implicit Types P Q : uPred M.

Ralf Jung's avatar
Ralf Jung committed
  Lemma coreP_intro P : P -∗ coreP P.
Robbert Krebbers's avatar
Robbert Krebbers committed
  Proof. iIntros "HP". by iIntros (Q HQ ->). Qed.
Ralf Jung's avatar
Ralf Jung committed
  Global Instance coreP_persistent P : PersistentP (coreP P).
Robbert Krebbers's avatar
Robbert Krebbers committed
  Proof. rewrite /coreP. apply _. Qed.
Jacques-Henri Jourdan's avatar
Jacques-Henri Jourdan committed
  Global Instance coreP_mono : Proper (() ==> ()) (@coreP M).
  Proof.
Robbert Krebbers's avatar
Robbert Krebbers committed
    rewrite /coreP. iIntros (P P' ?) "H"; iIntros (Q ??).
    iApply ("H" $! Q with "[%]"). by etrans.
Jacques-Henri Jourdan's avatar
Jacques-Henri Jourdan committed
  Qed.

  Global Instance coreP_proper : Proper ((⊣⊢) ==> (⊣⊢)) (@coreP M).
Robbert Krebbers's avatar
Robbert Krebbers committed
  Proof. intros P Q. rewrite !equiv_spec=>-[??]. by split; apply coreP_mono. Qed.
Ralf Jung's avatar
Ralf Jung committed
  Lemma coreP_elim P : PersistentP P  coreP P -∗ P.
Robbert Krebbers's avatar
Robbert Krebbers committed
  Proof. rewrite /coreP. iIntros (?) "HCP". unshelve iApply ("HCP" $! P); auto. Qed.