RunST.regions

From iris.base_logic Require Import auth big_op.
From iris.proofmode Require Import tactics classes.
From iris.algebra Require Import dec_agree auth gmap gset excl.
From RunST Require Import lang rules saved_pred.
Import uPred.

Definition bijective (g : gset (loc * loc)) :=
   x y, (x, y) g ( z, (x, z) g z = y) ( z, (z, y) g z = x).

Lemma empty_bijective : bijective .
Proof. by intros x y Hxy; apply not_elem_of_empty in Hxy. Qed.

Lemma bijective_extend g l l' :
  bijective g ( y, (l, y) g) ( y, (y, l') g)
  bijective (({[(l, l')]} : gset (loc * loc)) g).
Proof.
  intros bij Hl Hl' x y Hxy.
  apply elem_of_union in Hxy; destruct Hxy as [Hxy|Hxy].
  - apply elem_of_singleton_1 in Hxy; inversion Hxy; subst.
    split; intros z Hz; apply elem_of_union in Hz; destruct Hz as [Hz|Hz];
      try (apply elem_of_singleton_1 in Hz; inversion Hz; subst); trivial;
        try (by apply Hl in Hz); try (by apply Hl' in Hz).
  - assert (x l) by (by intros Heql; subst; apply Hl in Hxy).
    assert (y l') by (by intros Heql'; subst; apply Hl' in Hxy).
    apply bij in Hxy; destruct Hxy as [Hxy1 Hx2].
    split; intros z Hz; apply elem_of_union in Hz; destruct Hz as [Hz|Hz];
      try (apply elem_of_singleton_1 in Hz; inversion Hz; subst); trivial;
        try match goal with H : ?A = ?A |- _ => by contradict H end;
        auto.
Qed.

Lemma bij_eq_iff g l1 l2 l1' l2' :
  bijective g ((l1, l1') g) ((l2, l2') g) l1 = l2 l1' = l2'.
Proof.
  intros bij Hl1 Hl2; split => Hl1eq; subst;
    by pose proof (bij _ _ Hl1) as Hb1; apply Hb1 in Hl2.
Qed.

Definition RegN := (gname * gname)%type.

Definition regsUR := gmapUR positive (dec_agreeR RegN).

Definition bijUR := gsetUR (loc * loc).

Definition predsUR := gmapUR (loc * loc) (dec_agreeR gname).

Class LogRelG (Σ : gFunctors) := Build_LogRelG {
  regionsG :> authG Σ regsUR;
  regions_name : gname;
  bijG :> authG Σ bijUR;
  predsG :> authG Σ predsUR;
  spredG :> @savedPredG (prodC valC valC) Σ
}.

Class PreLogRelG (Σ : gFunctors) := BuildPreLogRelG {
  PreregionsG :> authG Σ regsUR;
  PrebijG :> authG Σ bijUR;
  PrepredsG :> authG Σ predsUR;
  PrespredG :> @savedPredG (prodC valC valC) Σ
}.

Definition PreLogRelG_LogRelG `{PreLogRelG Σ} γ : LogRelG Σ :=
  {|
    regionsG := PreregionsG;
    regions_name := γ;
    bijG := PrebijG;
    predsG := PrepredsG;
    spredG := PrespredG
  |}.

Definition PreLogRelΣ :=
  #[authΣ regsUR; authΣ bijUR;
      authΣ predsUR; @savedPredΣ_ex (prodC valC valC)].

Global Instance subG_PreLogRelΣ Σ : subG PreLogRelΣ Σ PreLogRelG Σ.
Proof.
  intros ?.
  repeat match goal with
    H : subG _ _ |- _ => apply subG_inv in H; destruct H as [? ?]
  end.
  econstructor; try apply subG_savedPropΣ_ex; apply _.
Qed.

Section Regions.
  Context `{!LogRelG Σ}.

  Definition REGS_def (R : regsUR) : iProp Σ := own regions_name ( R).
  Definition REGS_aux : { x | x = @REGS_def }. by eexists. Qed.
  Definition REGS := proj1_sig REGS_aux.
  Definition REGS_eq : @REGS = @REGS_def := proj2_sig REGS_aux.

  Global Instance REGS_timeless R : TimelessP (REGS R).
  Proof. rewrite REGS_eq /REGS_def. apply _. Qed.

  Definition REG_def n (reg : RegN) : iProp Σ :=
    own regions_name ( {[ n := DecAgree reg]}).
  Definition REG_aux : { x | x = @REG_def }. by eexists. Qed.
  Definition REG := proj1_sig REG_aux.
  Definition REG_eq : @REG = @REG_def := proj2_sig REG_aux.

  Global Instance REG_persistent n r : PersistentP (REG n r).
  Proof. rewrite REG_eq /REG_def. apply _. Qed.

  Global Instance REG_timeless n r : TimelessP (REG n r).
  Proof. rewrite REG_eq /REG_def. apply _. Qed.

  Lemma fresh_reg_alloc R r :
    REGS R ==★ n, (R !! n = None) REGS (<[n := DecAgree r]>R) REG n r.
  Proof.
    iIntros "H"; rewrite REGS_eq /REGS_def REG_eq /REG_def.
    iMod (own_update with "H") as "H".
    - apply auth_update_alloc.
      eapply (alloc_local_update _ _ (fresh (dom (gset _) R)) (DecAgree r));
        last done.
      eapply (@not_elem_of_dom _ _ (gset positive)); try typeclasses eauto.
      apply is_fresh.
    - rewrite own_op. iModIntro; iExists _; iFrame. iPureIntro.
      rewrite -(@not_elem_of_dom _ _ (gset positive)). apply is_fresh.
  Qed.

  Lemma reg_agree n r r' :
    REG n r REG n r' r = r'.
  Proof.
    iIntros "H"; rewrite REG_eq /REG_def -own_op.
    rewrite -auth_frag_op op_singleton.
    iDestruct (own_valid with "H") as %Hv.
    by apply singleton_valid, dec_agree_op_inv in Hv; inversion Hv.
  Qed.

  Lemma reg_in R n r:
    REGS R REG n r (R = <[n := DecAgree r]>(delete n R)).
  Proof.
    iIntros "[H1 H2]"; rewrite REG_eq /REG_def REGS_eq /REGS_def.
    iDestruct (own_valid_2 with "[$H1 $H2]") as %Hv.
    iPureIntro.
    apply auth_valid_discrete_2 in Hv; destruct Hv as [Hv1 Hv2].
    specialize (Hv2 n).
    apply singleton_included in Hv1.
    destruct Hv1 as (y & Heq & Hi).
    revert Hv2; rewrite Heq => Hv2.
    revert Hi; rewrite Some_included_total => Hi.
    destruct y as [y|]; last inversion Hv2.
    apply DecAgree_included in Hi; subst.
    by rewrite insert_delete insert_id ?(leibniz_equiv _ _ Heq).
  Qed.

  Definition BIJ_def γbij (L : bijUR) : iProp Σ :=
    (own γbij ( L) ( bijective L))%I.
  Definition BIJ_aux : { x | x = @BIJ_def }. by eexists. Qed.
  Definition BIJ := proj1_sig BIJ_aux.
  Definition BIJ_eq : @BIJ = @BIJ_def := proj2_sig BIJ_aux.

  Global Instance BIJ_timeless γbij L : TimelessP (BIJ γbij L).
  Proof. rewrite BIJ_eq /BIJ_def. apply _. Qed.

  Lemma alloc_empty_bij : True ==★ γ, BIJ γ .
  Proof.
    iIntros "". rewrite BIJ_eq /BIJ_def.
    iMod (own_alloc ( ( : bijUR))) as (γ) "H"; first done.
    iModIntro; iExists _; iFrame. iPureIntro. apply empty_bijective.
  Qed.

  Definition inBij_def γbij (l l' : loc) :=
    own γbij ( ({[ (l, l') ]} : gset (loc * loc))).
  Definition inBij_aux : { x | x = @inBij_def }. by eexists. Qed.
  Definition inBij := proj1_sig inBij_aux.
  Definition inBij_eq : @inBij = @inBij_def := proj2_sig inBij_aux.

  Global Instance inBij_timeless γbij l l' : TimelessP (inBij γbij l l').
  Proof. rewrite inBij_eq /inBij_def. apply _. Qed.

  Global Instance inBij_persistent γbij l l' : PersistentP (inBij γbij l l').
  Proof. rewrite inBij_eq /inBij_def. apply _. Qed.

  Lemma bij_alloc γbij (L : gset (loc * loc)) l l' :
    ( y, (l, y) L) ( y, (y, l') L)
    BIJ γbij L ==★
      BIJ γbij (({[(l, l')]} : gset (loc * loc)) L) inBij γbij l l'.
  Proof.
    iIntros (H1 H2) "H"; rewrite BIJ_eq /BIJ_def inBij_eq /inBij_def.
    iDestruct "H" as "[H1 H2]"; iDestruct "H2" as %Hbij.
    iMod (own_update with "H1") as "H1".
    - apply auth_update_alloc.
      apply (gset_local_update _ _ (({[(l, l')]} : gset (loc * loc)) L)).
      apply union_subseteq_r.
    - rewrite -{2}gset_op_union. rewrite own_op.
      change (own γbij ( ({[l, l']} L))) with
      (auth.auth_own γbij (({[l, l']} L))). rewrite auth_own_op /auth.auth_own.
      iDestruct "H1" as "[H1 [H2 _]]"; iFrame.
      iModIntro; iPureIntro. apply bijective_extend; auto.
  Qed.

  Lemma bij_agree γbij L l1 l2 l1' l2' :
    BIJ γbij L inBij γbij l1 l1' inBij γbij l2 l2'
         (l1 = l2 l1' = l2').
  Proof.
    iIntros "(H1 & H2 & H3)". rewrite BIJ_eq /BIJ_def inBij_eq /inBij_def.
    iDestruct "H1" as "[H1 HL]"; iDestruct "HL" as %HL.
    iDestruct (own_valid_2 with "[$H1 $H2]") as %Hv1%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "[$H1 $H3]") as %Hv2%auth_valid_discrete_2.
    iPureIntro.
    destruct Hv1 as [Hv1 _]; destruct Hv2 as [Hv2 _].
    apply gset_included, elem_of_subseteq_singleton in Hv1;
      apply gset_included, elem_of_subseteq_singleton in Hv2.
    eapply bij_eq_iff; eauto.
  Qed.

  Definition PREDS_def γrel (g : gset (loc * loc)) : iProp Σ :=
    ( P, ( (dom (gset _) P = g)) own γrel ( P))%I.
  Definition PREDS_aux : { x | x = @PREDS_def }. by eexists. Qed.
  Definition PREDS := proj1_sig PREDS_aux.
  Definition PREDS_eq : @PREDS = @PREDS_def := proj2_sig PREDS_aux.

  Global Instance PREDS_timeless γrel g : TimelessP (PREDS γrel g).
  Proof. rewrite PREDS_eq /PREDS_def. apply _. Qed.

  Lemma alloc_empty_preds : True ==★ γ, PREDS γ .
  Proof.
    iIntros "". rewrite PREDS_eq /PREDS_def.
    iMod (own_alloc ( ( : gmapUR (loc * loc) (dec_agreeR gname))))
      as (γ) "H"; first done.
    iModIntro; iExists γ, ; iFrame.
    iPureIntro. by rewrite dom_empty_L.
  Qed.

  Definition PREDS_include_def γrel (g : gset (loc * loc)) : iProp Σ :=
    ( P, ( (dom (gset _) P = g)) own γrel ( P))%I.
  Definition PREDS_include_aux : { x | x = @PREDS_include_def }. by eexists. Qed.
  Definition PREDS_include := proj1_sig PREDS_include_aux.
  Definition PREDS_include_eq : @PREDS_include = @PREDS_include_def :=
    proj2_sig PREDS_include_aux.

  Global Instance PREDS_include_timeless γrel g :
    TimelessP (PREDS_include γrel g).
  Proof. rewrite PREDS_include_eq /PREDS_include_def. apply _. Qed.

  Global Instance PREDS_include_persistent γrel g :
    PersistentP (PREDS_include γrel g).
  Proof. rewrite PREDS_include_eq /PREDS_include_def. apply _. Qed.

  Lemma PREDS_PREDS_include γrel g :
    PREDS γrel g ==★ PREDS γrel g PREDS_include γrel g.
  Proof.
    iIntros "H".
    rewrite PREDS_eq /PREDS_def PREDS_include_eq /PREDS_include_def.
    iDestruct "H" as (P) "[Hd HP]"; iDestruct "Hd" as %Hd.
    iMod (own_update with "HP") as "HP".
    - eapply (auth_update_alloc _ P P).
      intros n mz HP Heq; simpl in *; split; trivial.
      destruct mz as [mz|]; simpl in *; trivial.
      intros i; specialize (Heq i); revert Heq.
      rewrite !lookup_op lookup_empty left_id_L =>->.
      case: (mz !! i) => [[z|]|]; rewrite -?Some_op; auto.
      by rewrite dec_agree_idemp.
    - rewrite own_op; iDestruct "HP" as "[HP1 HP2]".
      by iModIntro; iSplitL "HP1"; iExists _; iFrame.
  Qed.

  Lemma PREDS_include_union γrel g g' :
    PREDS_include γrel g PREDS_include γrel g' PREDS_include γrel (g g').
  Proof.
    iIntros "[H1 H2]". rewrite PREDS_include_eq /PREDS_include_def.
    iDestruct "H1" as (P1) "[Hd1 HP1]"; iDestruct "Hd1" as %Hd1; subst.
    iDestruct "H2" as (P2) "[Hd2 HP2]"; iDestruct "Hd2" as %Hd2; subst.
    iCombine "HP1" "HP2" as "HP". rewrite -auth_frag_op.
    iExists _; iFrame. iPureIntro. apply dom_op.
  Qed.

  Lemma PREDS_include_correct γrel g g' :
    PREDS γrel g PREDS_include γrel g' (g' g).
  Proof.
    iIntros "[H1 H2]".
    rewrite PREDS_eq /PREDS_def PREDS_include_eq /PREDS_include_def.
    iDestruct "H1" as (P1) "[Hd1 HP1]"; iDestruct "Hd1" as %Hd1; subst.
    iDestruct "H2" as (P2) "[Hd2 HP2]"; iDestruct "Hd2" as %Hd2; subst.
    iDestruct (own_valid_2 with "[$HP1 $HP2]") as %Hv%auth_valid_discrete_2.
    destruct Hv as [Hv _].
    by iPureIntro; apply dom_included.
  Qed.

  Definition Rel_name_def γrel l l' γ :=
    own γrel ( {[(l, l') := DecAgree γ ]}).
  Definition Rel_name_aux : { x | x = @Rel_name_def }. by eexists. Qed.
  Definition Rel_name := proj1_sig Rel_name_aux.
  Definition Rel_name_eq : @Rel_name = @Rel_name_def := proj2_sig Rel_name_aux.

  Global Instance Rel_name_timeless γrel l l' γ :
    TimelessP (Rel_name γrel l l' γ).
  Proof. rewrite Rel_name_eq /Rel_name_def. apply _. Qed.

  Global Instance Rel_name_persistent γrel l l' γ :
    PersistentP (Rel_name γrel l l' γ).
  Proof. rewrite Rel_name_eq /Rel_name_def. apply _. Qed.

  Lemma Rel_name_alloc γrel g l l' γ :
    (l, l') g
    PREDS γrel g
      ==★ PREDS γrel ({[(l, l')]} g) Rel_name γrel l l' γ.
  Proof.
    iIntros (Hni) "Hpr". rewrite PREDS_eq /PREDS_def Rel_name_eq /Rel_name_def.
    iDestruct "Hpr" as (P) "[H1 Hpr]"; iDestruct "H1" as %Hdm.
    iMod (own_update with "Hpr") as "Hpr".
    - apply auth_update_alloc.
      eapply (alloc_singleton_local_update _ _ (DecAgree γ)); eauto.
      by apply not_elem_of_dom; rewrite Hdm.
      done.
    - rewrite own_op; iDestruct "Hpr" as "[Hp Hγ]"; iModIntro; iFrame.
      iExists _; iFrame; iPureIntro.
      by rewrite dom_insert_L Hdm.
  Qed.

  Lemma Rel_name_agree γrel l l' γ γ' :
    Rel_name γrel l l' γ Rel_name γrel l l' γ' γ = γ'.
  Proof.
    iIntros "H". rewrite Rel_name_eq /Rel_name_def.
    repeat match goal with
      |- context [own ?A ( ?B)] =>
      change (own A ( B)) with (auth.auth_own A B)
    end.
    rewrite -auth_own_op /auth.auth_own.
    rewrite op_singleton.
    iDestruct (own_valid with "H") as %Hv.
    apply singleton_valid in Hv.
    by apply dec_agree_op_inv in Hv; inversion Hv.
  Qed.

  Lemma Rel_name_inrels γ γrel r l l' :
    Rel_name γrel l l' γ PREDS γrel r ((l, l') r).
  Proof.
    iIntros "[H1 H2]". rewrite Rel_name_eq /Rel_name_def PREDS_eq /PREDS_def /=.
    iDestruct "H2" as (P) "[Heq HP]"; iDestruct "Heq" as %Heq; subst.
    iDestruct (own_valid_2 with "[$HP $H1]") as %Hv. iPureIntro.
    apply auth_valid_discrete_2 in Hv; destruct Hv as [Hv _].
    apply dom_included in Hv. revert Hv; rewrite dom_singleton =>Hv.
    by apply elem_of_subseteq_singleton.
  Qed.

  Program Definition Rel_for_def γrel (l l' : loc) :
    ((prodC valC valC) -n> iProp Σ) -n> iProp Σ :=
    λne Φ, ( γ, Rel_name γrel l l' γ saved_pred_own γ Φ)%I.
  Next Obligation.
  Proof. by intros ?????? H; apply exist_ne => ?; rewrite H. Qed.

  Definition Rel_for_aux : { x | x = @Rel_for_def }. by eexists. Qed.
  Definition Rel_for := proj1_sig Rel_for_aux.
  Definition Rel_for_eq : @Rel_for = @Rel_for_def := proj2_sig Rel_for_aux.

  Lemma Rel_for_inrels γrel r l l' Φ :
    Rel_for γrel l l' Φ PREDS γrel r ((l, l') r).
  Proof.
    iIntros "[H1 H2]". rewrite Rel_for_eq /Rel_for_def /=.
    iDestruct "H1" as (γ) "[H1 _]".
    iApply (Rel_name_inrels with "[$H1 $H2]").
  Qed.

  Global Instance Rel_for_persistent γrel l l' Φ :
    PersistentP (Rel_for γrel l l' Φ).
  Proof. rewrite Rel_for_eq /Rel_for_def. apply _. Qed.

  Program Definition In_Rel_for_def γrel (l l' : loc) :
    (prodC valC valC) -n> iProp Σ :=
    λne vs,
    ( γ (Φ : (prodC valC valC) -n> iProp Σ), ( ( vs, PersistentP (Φ vs)))
               Rel_name γrel l l' γ saved_pred_own γ Φ Φ vs)%I.
  Next Obligation.
  Proof. by intros ?????? H; repeat apply exist_ne => ?; rewrite H. Qed.

  Definition In_Rel_for_aux : { x | x = @In_Rel_for_def }. by eexists. Qed.
  Definition In_Rel_for := proj1_sig In_Rel_for_aux.
  Definition In_Rel_for_eq : @In_Rel_for = @In_Rel_for_def :=
    proj2_sig In_Rel_for_aux.

  Global Instance In_Rel_for_persistent
         γrel l l' (Φ : (prodC valC valC) -n> iProp Σ) vs :
    PersistentP (In_Rel_for γrel l l' vs).
  Proof.
    rewrite In_Rel_for_eq /In_Rel_for_def.
    apply exist_persistent => ?.
    apply exist_persistent => Φ'.
    unfold PersistentP. iIntros "(#H1 & #H2 & #H3 & H4)". iDestruct "H1" as %HP.
    iAssert ( Φ' vs)%I with "[H4]" as "#H4".
    { rewrite always_later; iNext; by iApply HP. }
    iAlways. by repeat iSplit.
  Qed.

  Lemma In_Rel_for_correct γrel l l' Φ vs :
    Rel_for γrel l l' Φ In_Rel_for γrel l l' vs (Φ vs).
  Proof.
    rewrite Rel_for_eq /Rel_for_def In_Rel_for_eq /In_Rel_for_def /=.
    iIntros "[H1 H2]".
    iDestruct "H1" as (γ1) "[#H11 #H12]".
    iDestruct "H2" as (γ2 Φ') "(#H21 & #H22 & #H23 & H24)".
    iPoseProof (saved_pred_impl with "[H24]") as "H"; last iExact "H".
    iFrame "H23".
    iDestruct (Rel_name_agree with "[]") as %Heq; first iFrame "H11 H22"; subst.
    iFrame "#"; iFrame.
  Qed.

  Lemma Rel_alloc γrel g l l' (Φ : (prodC valC valC) -n> iProp Σ) vs :
    ( vs, PersistentP (Φ vs))
    (l, l') g
    PREDS γrel g (Φ vs)
      ==★ PREDS γrel ({[(l, l')]} g) PREDS_include γrel {[(l, l')]}
          Rel_for γrel l l' Φ In_Rel_for γrel l l' vs.
  Proof.
    iIntros (Hps Hni) "[HR #HΦ]".
    iMod (saved_pred_alloc_strong Φ ) as (γ) "[_ #Hpred]".
    iMod ((Rel_name_alloc _ _ _ _ γ) with "HR") as "(HR & #Hrn)"; first eauto.
    iModIntro. iFrame.
    rewrite Rel_for_eq /Rel_for_def In_Rel_for_eq /In_Rel_for_def
            PREDS_include_eq /PREDS_include_def /=.
    repeat iSplit; try by repeat iExists _; iFrame "#".
    rewrite Rel_name_eq /Rel_name_def.
    iExists {[(l, l') := DecAgree γ]}; iFrame "#".
    iPureIntro. by rewrite dom_singleton_L.
  Qed.

  Lemma is_In_Rel_for γrel l l' (Φ : (prodC valC valC) -n> iProp Σ) vs :
    ( vs, PersistentP (Φ vs))
     Rel_for γrel l l' Φ Φ vs In_Rel_for γrel l l' vs.
  Proof.
    iIntros (Hps) "[#HR #HΦ]".
    rewrite Rel_for_eq /Rel_for_def In_Rel_for_eq /In_Rel_for_def /=.
    iDestruct "HR" as (γ) "[Hrn Hsp]".
    iExists _, _; iFrame "#". by iPureIntro.
  Qed.

  Program Definition STRefREL_def rn :
    ((prodC valC valC) -n> iProp Σ) -n> ((prodC valC valC) -n> iProp Σ) :=
    λne Φ vs,
      ( l l' R, vs = (LitV (LitLoc l), LitV (LitLoc l'))
         REG rn R inBij (fst R) l l' Rel_for (snd R) l l' Φ)%I.
  Next Obligation.
  Proof.
    intros ? ? ? [? ?] [? ?] [He1 He2]; repeat apply exist_ne =>?.
    by cbv in He1, He2; subst.
  Qed.
  Next Obligation.
  Proof. by intros ? ? ? ? Heq ?; repeat apply exist_ne =>?; rewrite Heq. Qed.

  Definition STRefREL_aux : { x | x = @STRefREL_def }. by eexists. Qed.
  Definition STRefREL := proj1_sig STRefREL_aux.
  Definition STRefREL_eq : @STRefREL = @STRefREL_def := proj2_sig STRefREL_aux.

  Global Instance STRefREL_persistent rn (Φ : (prodC valC valC) -n> iProp Σ) vs:
    ( vs, PersistentP (Φ vs)) PersistentP (STRefREL rn Φ vs).
  Proof. rewrite STRefREL_eq /STRefREL_def; apply _. Qed.

  Lemma STRefREL_locs rn Φ v v' :
    STRefREL rn Φ (v, v') ( l l', v = LitV $ LitLoc l v' = LitV $ LitLoc l').
  Proof.
    rewrite STRefREL_eq /STRefREL_def /=.
    iDestruct 1 as (l l' R) "(H1 & H2 & H3)"; iDestruct "H1" as %H1.
    iPureIntro. inversion H1; subst; eauto.
  Qed.

  Context `{heapG Σ}.

  Definition regIN := (nroot .@ "RegI").

  Definition _reg_inv : iProp Σ :=
    inv regIN
      ( Rs, REGS Rs
        [★ map] n dR Rs, R, dR = DecAgree R
          ( g, BIJ (fst R) g PREDS_include (snd R) g))%I.

  Definition reg_inv_def := _reg_inv.
  Definition reg_inv_aux : { x | x = @reg_inv_def }. by eexists. Qed.
  Definition reg_inv := proj1_sig reg_inv_aux.
  Definition reg_inv_eq : @reg_inv = @reg_inv_def := proj2_sig reg_inv_aux.

  Global Instance Persistent_reg_inv : PersistentP reg_inv.
  Proof. rewrite reg_inv_eq /reg_inv_def; apply _. Qed.

  Definition Region_def n γh γh' : iProp Σ :=
    ( r R, REG n R PREDS (snd R) r
       [★ set] ll' r,
          l l' v v', ll' = (l, l') l ↦{γh} v l' ↦{γh'} v'
           In_Rel_for (snd R) l l' (v, v'))%I.
  Definition Region_aux : { x | x = @Region_def }. by eexists. Qed.
  Definition Region := proj1_sig Region_aux.
  Definition Region_eq : @Region = @Region_def := proj2_sig Region_aux.

  Lemma region_alloc E γh γh' :
    nclose regIN E reg_inv ={E}=★ n, Region n γh γh'.
  Proof.
    iIntros (HE) "#HI". rewrite reg_inv_eq /reg_inv_def.
    iInv regIN as (Rs) "(>HRs & >HI')" "Hclose".
    iMod (alloc_empty_preds) as (γrel) "Hrel".
    iMod (alloc_empty_bij) as (γbij) "Hbij".
    iMod ((fresh_reg_alloc _ (γbij, γrel)) with "HRs") as
        (rn) "(Hni & HRs & Hreg)"; iDestruct "Hni" as %Hni.
    iMod (PREDS_PREDS_include with "Hrel") as "[Hrel Hri]".
    iMod ("Hclose" with "[HRs HI' Hbij Hri]") as "_".
    { iNext. iExists _; iFrame "HRs".
      rewrite big_sepM_insert; last trivial; iSplitR "HI'"; iFrame.
      iExists _; iSplit; trivial. iExists _; iFrame. }
    iModIntro. iExists rn; rewrite Region_eq /Region_def.
    iExists _, (γbij, γrel); iFrame.
    by rewrite big_sepS_empty.
  Qed.

  Lemma related_pairs_eq_iff E rn Φ Φ' l1 l1' l2 l2' :
    nclose regIN E
    STRefREL rn Φ (LitV (LitLoc l1), LitV (LitLoc l1'))
       STRefREL rn Φ' (LitV (LitLoc l2), LitV (LitLoc l2')) reg_inv
      ={E}=★ (l1 = l2 l1' = l2').
  Proof.
    rewrite STRefREL_eq /STRefREL_def; simpl.
    iIntros (HE) "(Hl1 & Hl2 & #HI)". rewrite reg_inv_eq /reg_inv_def.
    iDestruct "Hl1" as (k1 k1' R1) "(Heq1 & HRg1 & Hbij1 & HRL1)".
    iDestruct "Hl2" as (k2 k2' R2) "(Heq2 & HRg2 & Hbij2 & HRL2)".
    iDestruct "Heq1" as %Heq1; iDestruct "Heq2" as %Heq2;
      inversion Heq1; inversion Heq2; subst.
    iDestruct (reg_agree with "[$HRg1 $HRg2]") as %HReq; subst.
    iInv regIN as (Rs) "(>HRs & >HI')" "Hclose".
    iDestruct (reg_in with "[$HRs $HRg1]") as %->.
    rewrite big_sepM_insert; last by rewrite lookup_delete.
    iDestruct "HI'" as "[Hrn HI']".
    iDestruct "Hrn" as (R) "(HeqR & Hrn)".
    iDestruct "HeqR" as %HeqR; inversion HeqR; subst.
    iDestruct "Hrn" as (g) "(HRbij & HRrel)".
    iDestruct (bij_agree with "[$HRbij $Hbij1 $Hbij2]") as %Hiff.
    iMod ("Hclose" with "[-]"); last by iModIntro; iPureIntro.
    iNext; simpl.
    iExists (<[rn:=DecAgree R]> (delete rn Rs)); iFrame.
    rewrite big_sepM_insert; last by rewrite lookup_delete.
    iFrame.
    iExists _; iSplit; first done. iExists _; iFrame.
  Qed.

  Lemma not_in_region_r γrel γh γh' r l1 v1 :
      l1 ↦{γh} v1
         ([★ set] ll' r, (l l' : loc) (v v' : val),
             ll' = (l, l') l ↦{γh} v l' ↦{γh'} v'
                            (In_Rel_for γrel l l') (v, v'))
          ( l2, (l1, l2) r).
  Proof.
    iIntros "[Hl1 Hr]".
    rewrite pure_forall. iIntros (l2). rewrite pure_impl. iIntros (Hin).
    iDestruct ((big_sepS_elem_of _ r (l1, l2)) with "[Hr]") as "Hl12"; trivial.
    simpl.
    iDestruct "Hl12" as (l1' l2' w w') "(Heq & Hl1' & Hrest)".
    iDestruct "Heq" as %Heq; inversion Heq; subst.
    iApply (mapsto_no_dup with "[$Hl1 $Hl1']").
  Qed.

  Lemma not_in_region_l γrel γh γh' r l2 v2 :
      l2 ↦{γh'} v2
         ([★ set] ll' r, (l l' : loc) (v v' : val),
             ll' = (l, l') l ↦{γh} v l' ↦{γh'} v'
                            (In_Rel_for γrel l l') (v, v'))
          ( l1, (l1, l2) r).
  Proof.
    iIntros "[Hl2 Hr]".
    rewrite pure_forall. iIntros (l1). rewrite pure_impl. iIntros (Hin).
    iDestruct ((big_sepS_elem_of _ r (l1, l2)) with "[Hr]") as "Hl12"; trivial.
    simpl.
    iDestruct "Hl12" as (l1' l2' w w') "(Heq & Hl1 & Hl2' & Hrest)".
    iDestruct "Heq" as %Heq; inversion Heq; subst.
    iApply (mapsto_no_dup with "[$Hl2 $Hl2']").
  Qed.

  Lemma Region_extend E rn γh γh' (Φ : (prodC valC valC) -n> iProp Σ)
        l1 l2 v1 v2 :
    ( vs, PersistentP (Φ vs))
    nclose regIN E
    reg_inv Region rn γh γh' l1 ↦{γh} v1 l2 ↦{γh'} v2 Φ (v1 ,v2)
      ={E}=★ Region rn γh γh'
        STRefREL rn Φ (LitV (LitLoc l1), LitV (LitLoc l2)).
  Proof.
    rewrite STRefREL_eq /STRefREL_def Region_eq
            /Region_def reg_inv_eq /reg_inv_def /=.
    iIntros (Hps HE) "(#HI & HRg & Hl1 & Hl2 & #HPhi)".
    iDestruct "HRg" as (r R) "(HR & Hpr & Hpts)".
    iInv regIN as (Rs) "(>HRs & >HI')" "Hclose".
    iDestruct (reg_in with "[$HRs $HR]") as %->.
    rewrite big_sepM_insert; last by rewrite lookup_delete.
    iDestruct "HI'" as "[Hrn HI']".
    iDestruct "Hrn" as (R') "(HeqR & Hrn)".
    iDestruct "HeqR" as %HeqR; inversion HeqR; subst.
    iDestruct "Hrn" as (g) "(HRbij & HRrel)".
    iDestruct (not_in_region_r with "[$Hl1 $Hpts]") as %Hni1.
    iDestruct (not_in_region_l with "[$Hl2 $Hpts]") as %Hni2.
    iDestruct (PREDS_include_correct with "[$Hpr $HRrel]") as %Hincl.
    iMod ((Rel_alloc _ _ l1 l2) with "[$Hpr]") as
        "(Hpr & HRincl & Hrel & Hirel)"; auto.
    iDestruct (PREDS_include_union with "[$HRincl $HRrel]") as "HRrel".
    iMod (bij_alloc with "HRbij") as "[HRbij Hinbij]".
    { intros y Hy; eapply Hni1, Hincl; eauto. }
    { intros y Hy; eapply Hni2, Hincl; eauto. }
    iMod ("Hclose" with "[HRs HI' HRbij HRrel]") as "_".
    { iNext; simpl.
      iExists (<[rn:=DecAgree R']> (delete rn Rs)); iFrame.
      rewrite big_sepM_insert; last by rewrite lookup_delete.
      iFrame. iExists _; iSplit; trivial. iExists _; iFrame. }
    iModIntro.
    iSplit.
    - iExists _, _; iFrame.
      rewrite big_sepS_insert; last apply Hni1; iFrame.
      by iExists _, _, _, _; iFrame.
    - by iExists _, _, _; iFrame.
  Qed.

  Definition open_Region_def n γh γh' l l' : iProp Σ :=
    ( r R, REG n R PREDS (snd R) r
       [★ set] ll' (r {[ (l, l') ]}),
          l l' v v', ll' = (l, l') l ↦{γh} v l' ↦{γh'} v'
           In_Rel_for (snd R) l l' (v, v'))%I.
  Definition open_Region_aux : { x | x = @open_Region_def }. by eexists. Qed.
  Definition open_Region := proj1_sig open_Region_aux.
  Definition open_Region_eq : @open_Region = @open_Region_def :=
    proj2_sig open_Region_aux.

  Lemma region_open rn γh γh' l l' Φ :
    STRefREL rn Φ (LitV (LitLoc l), LitV (LitLoc l')) Region rn γh γh'
      v v', open_Region rn γh γh' l l'
          l ↦{γh} v l' ↦{γh'} v' Φ (v ,v').
  Proof.
    iIntros "[Hrf Hreg]".
    rewrite STRefREL_eq /STRefREL_def Region_eq /Region_def
            open_Region_eq /open_Region_def /=.
    iDestruct "Hrf" as (l1 l2 R) "(Hleq & HR & Hbij & #Hrelf)".
    iDestruct "Hleq" as %Hleq; inversion Hleq; subst.
    iDestruct "Hreg" as (r R') "(HR' & HPreds & Hpts)".
    iDestruct (reg_agree with "[$HR $HR']") as %HReq; subst.
    iDestruct (Rel_for_inrels with "[$Hrelf $HPreds]") as %Hin.
    rewrite big_sepS_delete; eauto.
    iDestruct "Hpts" as "[Hll Hpts]".
    iDestruct "Hll" as (l l' v v') "(Heq & Hl & Hl' & Hirel)".
    iDestruct "Heq" as %Heq; inversion Heq; subst.
    iExists _, _; iFrame.
    iSplitR "Hirel".
    - iExists _, _; iFrame.
    - iApply In_Rel_for_correct; iFrame "#"; iFrame.
  Qed.

  Lemma region_close rn γh γh' l l' v v' (Φ : (prodC valC valC) -n> iProp Σ) :
    ( vs, PersistentP (Φ vs))
    STRefREL rn Φ (LitV (LitLoc l), LitV (LitLoc l'))
      open_Region rn γh γh' l l' l ↦{γh} v l' ↦{γh'} v' Φ (v ,v')
       Region rn γh γh'.
  Proof.
    iIntros (HPs) "(Hrf & Hreg & Hl & Hl' & HP)".
    rewrite STRefREL_eq /STRefREL_def Region_eq /Region_def
            open_Region_eq /open_Region_def /=.
    iDestruct "Hrf" as (l1 l2 R) "(Hleq & HR & Hbij & #Hrelf)".
    iDestruct "Hleq" as %Hleq; inversion Hleq; subst.
    iDestruct "Hreg" as (r R') "(HR' & HPreds & Hpts)".
    iDestruct (reg_agree with "[$HR $HR']") as %HReq; subst.
    iDestruct (Rel_for_inrels with "[$Hrelf $HPreds]") as %Hin.
    iExists _, _; iFrame "HPreds"; iFrame.
    rewrite (big_sepS_delete _ r (l1, l2)); eauto.
    iSplitL "Hl Hl' HP"; last by iFrame.
    iDestruct (is_In_Rel_for with "[$HP]") as "HP"; first iFrame "#".
    by iExists _, _, _, _; iFrame.
  Qed.

End Regions.

Lemma reg_inv_alloc E `{heapG Σ} `{PreLogRelG Σ} :
  True ={E}=★ _ : LogRelG Σ, reg_inv.
Proof.
  iIntros "".
  iMod (own_alloc ( ( : regsUR))) as (γ) "H"; first done.
  set (PLRG := PreLogRelG_LogRelG γ).
  iMod (inv_alloc regIN E
                  ( Rs, REGS Rs
                     [★ map] n dR Rs, R, dR = DecAgree R
                     ( g, BIJ (fst R) g PREDS_include (snd R) g))%I
        with "[H]") as "Hinv".
  - rewrite REGS_eq /REGS_def.
    iNext; iExists _; iFrame.
    by rewrite big_sepM_empty.
  - iModIntro. iExists _. by rewrite reg_inv_eq /reg_inv_def /_reg_inv.
Qed.