LogrelCC.coop_logrel.correctness

From iris.algebra Require Import auth frac agree.
From iris.proofmode Require Import tactics.
From LogrelCC.program_logic Require Import adequacy.
From LogrelCC.program_logic Require Import cl_weakestpre.
From LogrelCC.cooperative.program_logic Require coop_language.
From LogrelCC.coop_logrel Require Import logrel compilation
     translation light_weight_threads logical_correctness.
From LogrelCC.F_mu_ref_cc Require Import lang rules_unary.

Notation CompiledStep := (@language.step F_mu_ref_cc.lang.lang).
Import coop_notations.
Import coop_type_notations.

Class heapPreIG Σ := HeapPreIG {
  heap_preG_iris :> invPreG Σ;
  heap_preG_heap :> gen_heapPreG loc val Σ
}.

Opaque sim_fork sim_yield.

Lemma basic_soundness Σ `{heapPreIG Σ, inG Σ (authR cfgUR),
                          inG Σ (authR (optionUR (exclR natC))),
                          na_invG Σ}
    e τ v th hp :
  (coop_typed [] e τ)
  rtc CompiledStep ([compile e], ) (of_val v :: th, hp)
  ( thp' hp' v' n, rtc step ([e], 0, ) (thp', n, hp')
                    (thp' !! n) = Some (coop_of_val v')).
Proof.
  intros Htp Hsteps.
  cut (adequate NotStuck (compile e) (λ _, ( thp' hp' v' n, rtc step ([e], 0, ) (thp', n, hp')
                    (thp' !! n) = Some (coop_of_val v')))).
  { destruct 1; naive_solver. }
  eapply (wp_adequacy Σ _); iIntros (Hinv).
  iMod (own_alloc ( to_gen_heap )) as (γ) "Hh".
  { apply (auth_auth_valid _ (to_gen_heap_valid _ _ )). }
  iMod (own_alloc ( (to_tpool [e], )
     ((to_tpool [e] : tpoolUR, ) : cfgUR))) as (γc) "[Hcfg1 Hcfg2]".
  { apply auth_valid_discrete_2. split=>//. split=>//. apply to_tpool_valid. }
  iMod (own_alloc ( (Excl' 0) (Excl' 0 : optionUR (exclR natC)))) as (γth) "[Hth1 Hth2]".
  { apply auth_valid_discrete_2. split=>//. }
  set (Hcfg := CFGSG _ _ _ γc γth).
  iMod (inv_alloc specN _ (spec_inv ([e], 0, )) with "[Hcfg1 Hth1]")
    as "#Hcfginv".
  { iNext. iExists [e], , 0. rewrite /to_gen_heap fin_maps.map_fmap_empty.
    iFrame. auto. }
  set (HeapΣ := (HeapG Σ Hinv (GenHeapG _ _ Σ _ _ _ γ))).
  iMod (na_alloc) as (p) "Hna".
  set (LR_na := {| logrel_nais := p |}).
  iExists (λ σ, own γ ( to_gen_heap σ)); iFrame.
  iModIntro.
  unfold compile.
  iApply (clwp_cl [LetInCtx _]);
    first by iApply (clwp_LTH_Lib_basic).
  iIntros (w); iDestruct 1 as (l) "[Hl %]"; subst w; simpl.
  pose proof (typed_Translate_n_closed _ _ _ Htp) as Hcl; simpl in *.
  iApply (wp_LetIn []); eauto; simpl.
  iNext.
  rewrite Hcl.
  iApply (wp_snd [AppRCtx (LamV _); AppLCtx _]); eauto; simpl.
  iNext.
  iApply (wp_Lam [AppLCtx _]); eauto; simpl.
  iNext.
  iApply (wp_fst [AppRCtx (LamV _)]); eauto; simpl.
  iNext.
  iApply (wp_Lam []); eauto.
  iNext; simpl. asimpl.
  set (coop_LBI := {| queue_loc := l |}).
  iAssert (|={}=> coop_inv)%I with "[Hl]" as "Hinv".
  { iMod (na_inv_alloc with "[Hl]"); last done.
    iNext. iExists []; iFrame.
    by rewrite big_sepL_nil. }
  iApply fupd_wp.
  iMod "Hinv" as "#Hinv'". iModIntro.
  iApply wp_fupd. iApply wp_wand_r.
  iSplitL.
  - iPoseProof ((@binary_fundamental Σ _ _ _ _ _ _ _ Htp [] [] ([e], 0, ) _)
                  with "[$Hcfginv]") as "Hrel"; iFrame "#".
    { by iApply interp_env_nil. }
    iSpecialize ("Hrel" $! ([], []) 0 with "[$Hna $Hth2 Hcfg2]"); simpl.
    { rewrite coop_empty_env_subst insert_empty /tpool_mapsto. iFrame.
      iAlways. iIntros (vv j) "(Hna & HCT & Hj & #Hrl)".
      iApply wp_value; eauto.
      iExists _, _; iFrame. }
    rewrite -(n_closed_subst_head_simpl_2 2 _ _ _ []); auto.
  - iIntros (_). iDestruct 1 as (v2 j') "(Hna & HCT & Hj')".
    iInv specN as (tp σ n) ">(Hcfg' & Hth & Hsteps)" "Hcl".
    iDestruct "Hsteps" as %Hsteps'.
    rewrite /tpool_mapsto /=.
    iDestruct (own_valid_2 with "Hcfg' Hj'") as %Hvalid.
    move: Hvalid=> /auth_valid_discrete_2
                   [/prod_included [/tpool_singleton_included Hv2 _] _].
    iDestruct (own_valid_2 with "Hth HCT") as
        %[Hvalid%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    iMod ("Hcl" with "[-]") as "_"; [iExists tp, σ, _; iFrame; eauto|].
    iIntros "!> !%"; eauto.
   eexists _, _, _, _; split; eauto.
Qed.

Instance : subG #[GFunctor (authR cfgUR)] Σ inG Σ (authR cfgUR).
Proof. solve_inG. Qed.

Instance : subG #[GFunctor (authR (optionUR (exclR natC)))] Σ
           inG Σ (authR (optionUR (exclR natC))).
Proof. solve_inG. Qed.

Lemma soundness e τ v th hp :
  (coop_typed [] e τ)
  rtc CompiledStep ([compile e], ) (of_val v :: th, hp)
  ( thp' hp' v' n, rtc step ([e], 0, ) (thp', n, hp')
                    (thp' !! n) = Some (coop_of_val v')).
Proof.
  intros Htp Hsteps.
  set (Σ := #[invΣ; na_invΣ; gen_heapΣ loc val; GFunctor (authR cfgUR);
              GFunctor (authR (optionUR (exclR natC)))]).
  set (HG := HeapPreIG Σ _ _).
  eapply (@basic_soundness Σ); eauto; typeclasses eauto.
Qed.