LogrelCC.cooperative.cooplang.rules_binary

From LogrelCC.cooperative.program_logic Require Import
     coop_ectx_language.
From LogrelCC Require Import prelude.
From iris.algebra Require Import auth frac agree gmap list.
From iris.base_logic Require Import invariants gen_heap.
From LogrelCC.cooperative.cooplang Require Export lang.
From iris.proofmode Require Import tactics.
Import uPred.

Definition specN := nroot .@ "spec".

The CMRA for the heap of the specification.
Definition tpoolUR : ucmraT := gmapUR nat (exclR coopexprC).
Definition cfgUR := prodUR tpoolUR (gen_heapUR loc coopval).
Definition CThreadUR := optionUR (exclR natC).

Fixpoint to_tpool_go (i : nat) (tp : list coopexpr) : tpoolUR :=
  match tp with
  | [] =>
  | e :: tp => <[i:=Excl e]>(to_tpool_go (S i) tp)
  end.
Definition to_tpool : list coopexpr tpoolUR := to_tpool_go 0.

The CMRA for the thread pool.
Class cfgSG Σ := CFGSG
  {
    cfg_inG :> inG Σ (authR cfgUR);
    Cthread_inG :> inG Σ (authR CThreadUR);
    cfg_name : gname;
    cth_name : gname }.

Section definitionsS.
  Context `{cfgSG Σ, invG Σ}.

  Definition heapS_mapsto (l : loc) (q : Qp) (v: coopval) : iProp Σ :=
    own cfg_name ( (, {[ l := (q, to_agree v) ]})).

  Definition tpool_mapsto (j : nat) (e: coopexpr) : iProp Σ :=
    own cfg_name ( ({[ j := Excl e ]}, )).

  Definition CurTh (n : nat) : iProp Σ :=
    own cth_name ( (Excl' n)).

  Definition spec_inv (ρ : cfg lang) : iProp Σ :=
    ( tp σ n, own cfg_name ( (to_tpool tp, to_gen_heap σ))
                  own cth_name ( Excl' n)
                  rtc step ρ (tp, n ,σ))%I.
  Definition spec_ctx (ρ : cfg lang) : iProp Σ :=
    inv specN (spec_inv ρ).

  Global Instance heapS_mapsto_timeless l q v : Timeless (heapS_mapsto l q v).
  Proof. apply _. Qed.
  Global Instance spec_ctx_persistent ρ : Persistent (spec_ctx ρ).
  Proof. apply _. Qed.
End definitionsS.
Typeclasses Opaque heapS_mapsto tpool_mapsto.

Notation "l ↦ₛ{ q } v" := (heapS_mapsto l q v)
  (at level 20, q at level 50, format "l ↦ₛ{ q } v") : bi_scope.
Notation "l ↦ₛ v" := (heapS_mapsto l 1 v) (at level 20) : bi_scope.
Notation "j ⤇ e" := (tpool_mapsto j e) (at level 20) : bi_scope.

Section conversions.
  Context `{cfgSG Σ}.

Conversion to tpools and back
  Lemma to_tpool_valid es : to_tpool es.
  Proof.
    rewrite /to_tpool. move: 0.
    induction es as [|e es]=> n //. by apply insert_valid.
  Qed.

  Lemma tpool_lookup tp j : to_tpool tp !! j = Excl <$> tp !! j.
  Proof.
    cut ( i, to_tpool_go i tp !! (i + j) = Excl <$> tp !! j).
    { intros help. apply (help 0). }
    revert j. induction tp as [|e tp IH]=> //= -[|j] i /=.
    - by rewrite Nat.add_0_r lookup_insert.
    - by rewrite -Nat.add_succ_comm lookup_insert_ne; last lia.
  Qed.
  Lemma tpool_lookup_Some tp j e : to_tpool tp !! j = Excl' e tp !! j = Some e.
  Proof. rewrite tpool_lookup fmap_Some. naive_solver. Qed.
  Hint Resolve tpool_lookup_Some.

  Lemma to_tpool_insert tp j e :
    j < length tp
    to_tpool (<[j:=e]> tp) = <[j:=Excl e]> (to_tpool tp).
  Proof.
    intros. apply: map_eq=> i. destruct (decide (i = j)) as [->|].
    - by rewrite tpool_lookup lookup_insert list_lookup_insert.
    - rewrite tpool_lookup lookup_insert_ne // list_lookup_insert_ne //.
      by rewrite tpool_lookup.
  Qed.
  Lemma to_tpool_insert' tp j e :
    is_Some (to_tpool tp !! j)
    to_tpool (<[j:=e]> tp) = <[j:=Excl e]> (to_tpool tp).
  Proof.
    rewrite tpool_lookup fmap_is_Some lookup_lt_is_Some. apply to_tpool_insert.
  Qed.

  Lemma to_tpool_snoc tp e :
    to_tpool (tp ++ [e]) = <[length tp:=Excl e]>(to_tpool tp).
  Proof.
    intros. apply: map_eq=> i.
    destruct (lt_eq_lt_dec i (length tp)) as [[?| ->]|?].
    - rewrite lookup_insert_ne; last lia. by rewrite !tpool_lookup lookup_app_l.
    - by rewrite lookup_insert tpool_lookup lookup_app_r // Nat.sub_diag.
    - rewrite lookup_insert_ne; last lia.
      rewrite !tpool_lookup ?lookup_ge_None_2 ?app_length //=;
         change (ofe_car coopexprC) with coopexpr; lia.
  Qed.

  Lemma tpool_singleton_included tp j e :
    {[j := Excl e]} to_tpool tp tp !! j = Some e.
  Proof.
    move=> /singleton_included [ex [/leibniz_equiv_iff]].
    rewrite tpool_lookup fmap_Some=> [[e' [-> ->]] /Excl_included ?]. by f_equal.
  Qed.
  Lemma tpool_singleton_included' tp j e :
    {[j := Excl e]} to_tpool tp to_tpool tp !! j = Excl' e.
  Proof. rewrite tpool_lookup. by move=> /tpool_singleton_included=> ->. Qed.

End conversions.

Section cfg.
  Context `{invG Σ, cfgSG Σ}.
  Implicit Types P Q : iProp Σ.
  Implicit Types Φ : coopval iProp Σ.
  Implicit Types σ : state.
  Implicit Types e : coopexpr.
  Implicit Types v : coopval.

  Local Hint Resolve tpool_lookup.
  Local Hint Resolve tpool_lookup_Some.
  Local Hint Resolve to_tpool_insert.
  Local Hint Resolve to_tpool_insert'.
  Local Hint Resolve tpool_singleton_included.

  Lemma step_insert_normal K tp j e σ e' σ' ef :
    let m :=
        match ef with
        | Some _ => length tp
        | None => j
        end
    in
    tp !! j = Some (fill K e) head_step K e σ e' σ' ef NormalMode
    step (tp, j, σ) (<[j:=fill K e']> tp ++ (forked_thread ef), m, σ').
  Proof.
    intros. rewrite -(take_drop_middle tp j (fill K e)) //.
    rewrite insert_app_r_alt take_length_le ?Nat.sub_diag /=;
      eauto using lookup_lt_Some, Nat.lt_le_incl.
    rewrite -(assoc_L (++)) /=.
    assert (j < length tp) by by apply lookup_lt_is_Some_1; eauto.
    eapply (step_atomic _ _ _ _ _ _ _ false); eauto.
    { rewrite take_length. lia. }
    { destruct ef; last done.
      rewrite take_length_le; auto with omega.
      rewrite drop_length. unfold m. omega. }
    by apply: Ectx_normal_step'.
  Qed.

  Lemma step_insert_no_fork_normal K tp j e σ e' σ' :
    tp !! j = Some (fill K e) head_step K e σ e' σ' None NormalMode
    step (tp, j, σ) (<[j:=fill K e']> tp, j, σ').
  Proof.
    rewrite -(right_id_L [] (++) (<[_:=_]>_)).
    by apply step_insert_normal.
  Qed.

  Lemma step_pure_normal E ρ j K e e' :
    ( σ, head_step K e σ e' σ None NormalMode)
    nclose specN E
    CurTh j spec_ctx ρ j fill K e ={E}=∗ CurTh j j fill K e'.
  Proof.
    iIntros (??) "(Hct & #Hspec & Hj)". rewrite /spec_ctx /tpool_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[?%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1,
        singleton_local_update, (exclusive_local_update _ (Excl (fill K e'))). }
    iFrame. iApply "Hclose". iNext. iExists (<[n:=fill K e']> tp), σ, n.
    rewrite to_tpool_insert'; last eauto.
    iFrame. iPureIntro. eapply rtc_r, step_insert_no_fork_normal; eauto.
  Qed.

  Lemma step_alloc E ρ j K e v : IntoVal e v
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Alloc e) ={E}=∗
      CurTh j l, j fill K (Loc l) l ↦ₛ v.
  Proof.
    iIntros (??) "(Hct & #Hspec & Hj)". rewrite /spec_ctx /tpool_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    destruct (exist_fresh (dom (gset positive) σ)) as [l Hl%not_elem_of_dom].
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[?%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1,
        singleton_local_update, (exclusive_local_update _ (Excl (fill K (Loc l)))). }
    iMod (own_update with "Hown") as "[Hown Hl]".
    { eapply auth_update_alloc, prod_local_update_2,
        (alloc_singleton_local_update _ l (1%Qp,to_agree v)); last done.
      by apply lookup_to_gen_heap_None. }
    iFrame.
    iExists l. rewrite /heapS_mapsto. iFrame "Hj Hl". iApply "Hclose". iNext.
    iExists (<[n:=fill K (Loc l)]> tp), (<[l:=v]>σ), n.
    rewrite to_gen_heap_insert to_tpool_insert'; last eauto. iFrame. iPureIntro.
    eapply rtc_r, step_insert_no_fork_normal; eauto. econstructor; eauto.
  Qed.

  Lemma step_load E ρ j K l q v:
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Load (Loc l)) l ↦ₛ{q} v
    ={E}=∗ CurTh j j fill K (of_val v) l ↦ₛ{q} v.
  Proof.
    iIntros (?) "(Hct & #Hinv & Hj & Hl)".
    rewrite /spec_ctx /tpool_mapsto /heapS_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[?%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hown Hl")
      as %[[? ?%gen_heap_singleton_included]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1, singleton_local_update,
        (exclusive_local_update _ (Excl (fill K (of_val v)))). }
    iFrame. iApply "Hclose". iNext.
    iExists (<[n:=fill K (of_val v)]> tp), σ, n.
    rewrite to_tpool_insert'; last eauto. iFrame. iPureIntro.
    eapply rtc_r, step_insert_no_fork_normal; eauto. econstructor; eauto.
  Qed.

  Lemma step_store E ρ j K l v' e v: IntoVal e v
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Store (Loc l) e) l ↦ₛ v'
    ={E}=∗ CurTh j j fill K Unit l ↦ₛ v.
  Proof.
    iIntros (??) "(Hct & #Hspec & Hj & Hl)".
    rewrite /spec_ctx /tpool_mapsto /heapS_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[?%tpool_singleton_included' _]%prod_included _]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hown Hl")
      as %[[_ Hl%gen_heap_singleton_included]%prod_included _]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1, singleton_local_update,
        (exclusive_local_update _ (Excl (fill K Unit))). }
    iMod (own_update_2 with "Hown Hl") as "[Hown Hl]".
    { eapply auth_update, prod_local_update_2, singleton_local_update,
        (exclusive_local_update _ (1%Qp, to_agree v)); last done.
      by rewrite /to_gen_heap lookup_fmap Hl. }
    iFrame. iApply "Hclose". iNext.
    iExists (<[n:=fill K Unit]> tp), (<[l:=v]>σ), n.
    rewrite to_gen_heap_insert to_tpool_insert'; last eauto. iFrame. iPureIntro.
    eapply rtc_r, step_insert_no_fork_normal; eauto. econstructor; eauto.
  Qed.

  Lemma step_rec E ρ j K e1 e2 v : IntoVal e2 v
    nclose specN E
    CurTh j spec_ctx ρ j fill K (App (Rec e1) e2)
    ={E}=∗ CurTh j j fill K (e1.[Rec e1,e2/]).
  Proof. intros ??; apply step_pure_normal => σ; eauto. econstructor; eauto. Qed.

  Lemma step_Lam E ρ j K e1 e2 v : IntoVal e2 v
    nclose specN E
    CurTh j spec_ctx ρ j fill K (App (Lam e1) e2)
    ={E}=∗ CurTh j j fill K (e1.[e2/]).
  Proof.
    intros ??. apply step_pure_normal => σ; eauto. econstructor; eauto.
  Qed.

  Lemma step_LetIn E ρ j K e1 e2 v : IntoVal e1 v
    nclose specN E
    CurTh j spec_ctx ρ j fill K (LetIn e1 e2)
    ={E}=∗ CurTh j j fill K (e2.[e1/]).
  Proof.
    intros ??. apply step_pure_normal => σ; eauto. econstructor; eauto.
  Qed.

  Lemma step_Seq E ρ j K e1 e2 : AsVal e1
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Seq e1 e2)
    ={E}=∗ CurTh j j fill K e2.
  Proof.
    intros []?. apply step_pure_normal => σ; eauto. econstructor; eauto.
  Qed.

  Lemma step_tlam E ρ j K e :
    nclose specN E
    CurTh j spec_ctx ρ j fill K (TApp (TLam e)) ={E}=∗
    CurTh j j fill K e.
  Proof. apply step_pure_normal => σ; econstructor; eauto. Qed.

  Lemma step_Fold E ρ j K e v : IntoVal e v
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Unfold (Fold e)) ={E}=∗
    CurTh j j fill K e.
  Proof. intros H1; apply step_pure_normal => σ; eauto. econstructor; eauto. Qed.

  Lemma step_fst E ρ j K e1 v1 e2 v2 : IntoVal e1 v1 IntoVal e2 v2
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Fst (Pair e1 e2)) ={E}=∗
    CurTh j j fill K e1.
  Proof. intros ??. apply step_pure_normal => σ; econstructor; eauto. Qed.

  Lemma step_snd E ρ j K e1 v1 e2 v2: IntoVal e1 v1 IntoVal e2 v2
       nclose specN E
    CurTh j spec_ctx ρ j fill K (Snd (Pair e1 e2)) ={E}=∗
    CurTh j j fill K e2.
  Proof. intros ??. apply step_pure_normal => σ; econstructor; eauto. Qed.

  Lemma step_case_inl E ρ j K e0 v0 e1 e2 : IntoVal e0 v0
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Case (InjL e0) e1 e2)
      ={E}=∗ CurTh j j fill K (e1.[e0/]).
  Proof. intros ?. apply step_pure_normal => σ; econstructor; eauto. Qed.

  Lemma step_case_inr E ρ j K e0 v0 e1 e2 : IntoVal e0 v0
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Case (InjR e0) e1 e2)
      ={E}=∗ CurTh j j fill K (e2.[e0/]).
  Proof. intros H1; apply step_pure_normal => σ; eauto. econstructor; eauto. Qed.

  Lemma step_if_false E ρ j K e1 e2 :
    nclose specN E
    CurTh j spec_ctx ρ j fill K (If (#♭ false) e1 e2) ={E}=∗
    CurTh j j fill K e2.
  Proof. apply step_pure_normal => σ; econstructor. Qed.

  Lemma step_if_true E ρ j K e1 e2 :
    nclose specN E
    CurTh j spec_ctx ρ j fill K (If (#♭ true) e1 e2) ={E}=∗
    CurTh j j fill K e1.
  Proof. apply step_pure_normal => σ; econstructor. Qed.

  Lemma step_nat_binop E ρ j K op e1 e2 v1 v2 w:
        IntoVal e1 v1 IntoVal e2 v2
    binop_eval op v1 v2 = Some w
    nclose specN E
    CurTh j spec_ctx ρ j fill K (BinOp op e1 e2)
      ={E}=∗ CurTh j j fill K (of_val w).
  Proof. intros. apply step_pure_normal => σ; eauto. econstructor; eauto. Qed.

  Lemma step_fork E ρ j K e :
    nclose specN E
    CurTh j spec_ctx ρ j fill K (CoFork e) ={E}=∗
     j', CurTh j' j fill K Unit j' e.
  Proof.
    iIntros (?) "(Hct & #Hinv & Hj)". rewrite /spec_ctx /tpool_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[?%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    assert (n < length tp) by eauto using lookup_lt_Some.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1,
        singleton_local_update, (exclusive_local_update _ (Excl (fill K Unit))). }
    iMod (own_update with "Hown") as "[Hown Hfork]".
    { eapply auth_update_alloc, prod_local_update_1,
        (alloc_singleton_local_update _ (length tp) (Excl e)); last done.
      rewrite lookup_insert_ne ?tpool_lookup; last omega.
      by rewrite lookup_ge_None_2. }
    iMod (own_update_2 with "Hth Hct") as "[Hth Hct]".
    { by eapply (auth_update _ _ (Excl' (length tp))),
      option_local_update, exclusive_local_update. }
    iExists (length tp). iFrame. iApply "Hclose". iNext.
    iExists (<[n:=fill K Unit]> tp ++ [e]), σ, (length tp).
    rewrite to_tpool_snoc insert_length to_tpool_insert //. iFrame. iPureIntro.
    eapply rtc_r, (step_insert_normal _ _ _ _ _ _ _ (Some e)); eauto.
    econstructor; eauto.
  Qed.

  Lemma step_yield_change E ρ j j' K e' :
    nclose specN E
    CurTh j spec_ctx ρ j fill K CoYield j' e' ={E}=∗
    CurTh j' j fill K Unit j' e'.
  Proof.
    iIntros (?) "(Hct & #Hinv & Hj & Hj')". rewrite /spec_ctx /tpool_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[Hij%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hown Hj'")
      as %[[Hij'%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    assert (n < length tp) by eauto using lookup_lt_Some.
    assert (j' < length tp) by eauto using lookup_lt_Some.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1,
        singleton_local_update, (exclusive_local_update _ (Excl (fill K Unit))). }
    iMod (own_update_2 with "Hth Hct") as "[Hth Hct]".
    { by eapply (auth_update _ _ (Excl' j')),
      option_local_update, exclusive_local_update. }
    iFrame. iApply "Hclose". iNext.
    iExists (<[n:=fill K Unit]> tp), σ, j'.
    rewrite to_tpool_insert //. iFrame. iPureIntro.
    eapply rtc_r; first by eauto.
    rewrite tpool_lookup in Hij.
    rewrite -(take_drop_middle tp n (fill K CoYield)); last first.
    { revert Hij; case:(tp !! n) => [? []|]; simpl; by try intros ->. }
    rewrite insert_app_r_alt take_length_le ?Nat.sub_diag /=;
      eauto using lookup_lt_Some, Nat.lt_le_incl.
    eapply (step_atomic _ _ _ _ _ _ None true _ _ n j'); simpl; eauto.
    { by rewrite app_nil_r. }
    { rewrite take_length. lia. }
    { rewrite take_length_le; auto with omega.
      rewrite drop_length. omega. }
    apply: Ectx_yield_step; eauto.
    econstructor.
  Qed.

  Lemma step_yield_dont_change E ρ j K:
    nclose specN E
    CurTh j spec_ctx ρ j fill K CoYield ={E}=∗
    CurTh j j fill K Unit.
  Proof.
    iIntros (?) "(Hct & #Hinv & Hj)". rewrite /spec_ctx /tpool_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[Hij%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    assert (n < length tp) by eauto using lookup_lt_Some.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1,
        singleton_local_update, (exclusive_local_update _ (Excl (fill K Unit))). }
    iFrame. iApply "Hclose". iNext.
    iExists (<[n:=fill K Unit]> tp), σ, n.
    rewrite to_tpool_insert //. iFrame. iPureIntro.
    eapply rtc_r; first by eauto.
    rewrite tpool_lookup in Hij.
    rewrite -(take_drop_middle tp n (fill K CoYield)); last first.
    { revert Hij; case:(tp !! n) => [? []|]; simpl; by try intros ->. }
    rewrite insert_app_r_alt take_length_le ?Nat.sub_diag /=;
      eauto using lookup_lt_Some, Nat.lt_le_incl.
    eapply (step_atomic _ _ _ _ _ _ None true _ _ n n); simpl; eauto.
    { by rewrite app_nil_r. }
    { rewrite take_length. lia. }
    { rewrite take_length_le; auto with omega. }
    apply: Ectx_yield_step; eauto.
    econstructor.
  Qed.

  Lemma step_insert_capture K tp j e σ e' σ' ef :
    tp !! j = Some (fill K e) head_step K e σ e' σ' ef CaptureMode
    step (tp, j, σ) (<[j:=fill K e']> tp, j, σ').
  Proof.
    intros ? Hstp. rewrite -(take_drop_middle tp j (fill K e)) //.
    rewrite insert_app_r_alt take_length_le ?Nat.sub_diag /=;
      eauto using lookup_lt_Some, Nat.lt_le_incl.
    assert (j < length tp) by by apply lookup_lt_is_Some_1; eauto.
    inversion Hstp; subst.
    eapply (step_atomic _ _ _ _ _ _ None false); rewrite /= ?app_nil_r; eauto.
    { rewrite take_length_le; trivial.
      by apply PeanoNat.Nat.lt_le_incl.
    (* omega bug!*) }
    by apply: Ectx_capture_step'; eauto.
  Qed.

  Lemma step_callcc E ρ j K e :
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Callcc e)
    ={E}=∗ CurTh j j fill K (e.[Cont K/]).
  Proof.
    iIntros (?) "(Hct & #Hinv & Hj)". rewrite /spec_ctx /tpool_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[?%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1,
        singleton_local_update,
        (exclusive_local_update _ (Excl (fill K (e.[Cont K/])))). }
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    iFrame. iApply "Hclose". iNext.
    iExists (<[n:=fill K (e.[Cont K/])]> tp), σ, n.
    rewrite to_tpool_insert'; last eauto.
    iFrame. iPureIntro.
    eapply rtc_r, step_insert_capture; eauto.
    econstructor.
  Qed.

  Lemma step_insert_throw K tp j e σ e' σ' ef :
    tp !! j = Some (fill K e) head_step K e σ e' σ' ef ThrowMode
    step (tp, j, σ) (<[j:= e']> tp, j, σ').
  Proof.
    intros? Hstp. rewrite -(take_drop_middle tp j (fill K e)) //.
    rewrite insert_app_r_alt take_length_le ?Nat.sub_diag /=;
      eauto using lookup_lt_Some, Nat.lt_le_incl.
    assert (j < length tp) by by apply lookup_lt_is_Some_1; eauto.
    inversion Hstp; subst.
    eapply (step_atomic _ _ _ _ _ _ None false); rewrite /= ?app_nil_r; eauto.
    { rewrite take_length_le; trivial.
      by apply PeanoNat.Nat.lt_le_incl.
    (* omega bug!*) }
    by apply: Ectx_throw_step'; eauto.
  Qed.

  Lemma step_throw E ρ j K K' e v : IntoVal e v
    nclose specN E
    CurTh j spec_ctx ρ j fill K (Throw e (Cont K'))
    ={E}=∗ CurTh j j fill K' e.
  Proof.
    iIntros (??) "(Hct & #Hinv & Hj)". rewrite /spec_ctx /tpool_mapsto.
    iInv specN as (tp σ n) ">(Hown & Hth & %)" "Hclose".
    iDestruct (own_valid_2 with "Hown Hj")
      as %[[?%tpool_singleton_included' _]%prod_included ?]%auth_valid_discrete_2.
    iDestruct (own_valid_2 with "Hth Hct") as
        %[?%Excl_included%leibniz_equiv _]%auth_valid_discrete_2; subst.
    iMod (own_update_2 with "Hown Hj") as "[Hown Hj]".
    { by eapply auth_update, prod_local_update_1,
        singleton_local_update,
        (exclusive_local_update _ (Excl (fill K' e))). }
    iFrame. iApply "Hclose". iNext.
    iExists (<[n:=fill K' e]> tp), σ, n.
    rewrite to_tpool_insert'; last eauto.
    iFrame. iPureIntro.
    eapply rtc_r, step_insert_throw; eauto.
    econstructor; eauto.
  Qed.

End cfg.