RunST.future

From iris.base_logic Require Import upred derived lib.fancy_updates
 primitive lib.wsat soundness.
From iris.prelude Require Export coPset.
From iris.base_logic Require Import big_op.
From iris.proofmode Require Export classes.
From iris.algebra Require Import gmap.
From RunST Require Export naught.

Import uPred_primitive.
Import uPred.

Require Import iris.proofmode.tactics.

(* future modality *)
Definition future_def `{invG Σ}
    (E : coPset) n (P : iProp Σ) : iProp Σ :=
  (Nat.iter n (λ Q, (|={E}=> Q)) (|={E}=> P))%I.
Definition future_aux : { x | x = @future_def }. by eexists. Qed.
Definition future := proj1_sig future_aux.
Definition future_eq : @future = @future_def := proj2_sig future_aux.
Arguments future {_ _} _ _ _%I.
Instance: Params (@future) 4.

Notation "|≫{ E }=[ n ]=> Q" := (future E n Q)
  (at level 99, E at level 50, Q at level 200,
   format "|≫{ E }=[ n ]=> Q") : uPred_scope.
Notation "P ≫{ E }=[ n ]=★ Q" := (P -★ |≫{E}=[n]=> Q)%I
  (at level 99, E at level 50, Q at level 200,
   format "P ≫{ E }=[ n ]=★ Q") : uPred_scope.
Notation "P ≫{ E }=[ n ]=★ Q" := (P |≫{E}=[n]=> Q)
  (at level 99, E at level 50, Q at level 200, only parsing) : C_scope.

Lemma future_unfold `{invG Σ} (E : coPset) n (P : iProp Σ) :
  (|≫{E}=[n]=> P) ⊣⊢
  (match n with
   | O => |={E}=> P
   | S n => (Nat.iter (S n) (λ Q, (|={E}=> Q)) (|={E}=> P))
   end).
Proof.
  rewrite future_eq. destruct n; trivial.
Qed.

Lemma future_unfold_O `{invG Σ} (E : coPset) (P : iProp Σ) :
  (|≫{E}=[0]=> P) ⊣⊢ |={E}=> P.
Proof. by rewrite future_unfold. Qed.

Lemma future_unfold_S `{invG Σ} (E : coPset) n (P : iProp Σ) :
  (|≫{E}=[S n]=> P) ⊣⊢ (Nat.iter (S n) (λ Q, (|={E}=> Q)) (|={E}=> P)).
Proof. by rewrite future_unfold. Qed.

Section future_and_plain.
 Context `{invG Σ}.

 Lemma fupd_extract_plain {E1 E2} P Q `{!PlainP P} :
  (Q P) (|={E1, E2}=> Q) |={E1}=> ((|={E1, E2}=> Q) P).
 Proof.
  intros HQP. rewrite fupd_eq /fupd_def; unseal.
  split => n x Hx Hfp k x' Hkn Hxx' Hwo.
  destruct (Hfp k x' Hkn Hxx' Hwo k (le_n _)) as [z [Hz1 Hz2]];
   first by rewrite right_id.
  assert (Hz1' : ✓{k} (z)) by (by revert Hz1; rewrite right_id).
  assert (HP : ( P)%I k ).
  { rewrite -uPred_sep_eq in Hz2. apply except_0_sep' in Hz2; last done.
    rewrite uPred_sep_eq in Hz2. destruct Hz2 as (z1 & z2 & Hz & _ & Hz2).
    rewrite -uPred_sep_eq in Hz2.
    assert (Hz2vl : ✓{k} z2) by (revert Hz1'; rewrite Hz;
      eauto using cmra_validN_op_l, cmra_validN_op_r).
    apply except_0_sep' in Hz2; last done.
   rewrite uPred_sep_eq in Hz2. destruct Hz2 as (w1 & w2 & Hw & _ & Hz2).
   assert (Hw2vl : ✓{k} w2) by (revert Hz2vl; rewrite Hw;
     eauto using cmra_validN_op_l, cmra_validN_op_r).
   eapply (naught_irrel' _ _ w2); auto using ucmra_unit_validN.
   by eapply (except_0_mono _ _ HQP). }
  intros m yf Hmk Hxx'yf. exists (x x'); split; auto.
  destruct Hwo as (u1 & u2 & Hu & Hu1 & Hu2).
  rewrite -uPred_sep_eq; apply except_0_sep; first eauto using cmra_validN_le.
  rewrite uPred_sep_eq.
  exists u1, (x u2); split;
   first by rewrite (comm _ u1) -assoc (dist_le _ _ _ _ Hu) // (comm _ u1).
  split.
  { assert (✓{m} u1) by
    (revert Hxx'; rewrite Hu=> Hxx';
      eapply cmra_validN_op_l, cmra_validN_op_r, cmra_validN_le; eauto).
    apply except_0_intro; last eapply uPred_closed; eauto. }
  assert (✓{m} (x u2)) by
    (revert Hxx'; rewrite Hu=> Hxx';
      eapply cmra_validN_op_l; rewrite -assoc (comm _ u2);
      eapply cmra_validN_le; eauto).
  rewrite -uPred_sep_eq; apply except_0_sep; first done. rewrite uPred_sep_eq.
  exists u2, x; split; first by rewrite (comm _ u2).
  split.
  { assert (✓{m} u2) by eauto using cmra_validN_op_r.
    apply except_0_intro; last eapply uPred_closed; eauto. }
  rewrite -uPred_sep_eq. apply except_0_sep; first eauto using cmra_validN_op_l.
  rewrite uPred_sep_eq.
  exists x, ( : iResUR Σ); split; first by rewrite right_id.
  split.
  { apply except_0_intro; first eauto using cmra_validN_op_l.
    eapply uPred_closed; eauto; last eauto using cmra_validN_op_l. omega. }
  eapply uPred_closed; eauto using ucmra_unit_validN.
 Qed.

 Lemma later_fupd_plain {E} P `{Hpl: !PlainP P} :
   ( |={E}=> P) |={E}=> P.
 Proof.
  rewrite fupd_eq /fupd_def later_wand; unseal.
  split=> n x Hx Hfp k x' Hkn Hxx' Hwo.
  pose proof (Hwo) as Hwo'; apply later_intro in Hwo';
    last eauto using cmra_validN_op_r; rewrite uPred_later_eq in Hwo'.
  specialize (Hfp k x' Hkn Hxx' Hwo').
  assert (HP : ( P)%I k x).
  { eapply (naught_irrel' _ _ (x x')); eauto using cmra_validN_op_l.
    eapply later_mono; last (rewrite uPred_later_eq; apply Hfp); trivial.
    rewrite -uPred_bupd_eq; apply bupd_extract_plain; first typeclasses eauto.
    apply except_0_mono; rewrite -uPred_sep_eq; by iIntros "(_&_&?)".
  }
  rewrite -uPred_bupd_eq; apply bupd_intro; trivial.
  rewrite -uPred_sep_eq -uPred_later_eq. rewrite -uPred_sep_eq in Hwo.
  apply except_0_intro; trivial.
  apply sep_assoc'; trivial.
  rewrite {1}uPred_sep_eq; eexists x', x.
  repeat split; trivial; by rewrite (comm _ x).
 Qed.

 Lemma step_Sn_fupd_mono {E} P Q n :
   (P Q)
   Nat.iter n (λ Q : iProp Σ, |={E}=> Q) P
   Nat.iter n (λ Q : iProp Σ, |={E}=> Q) Q.
 Proof.
  revert P Q; induction n=> P Q HPQ; first done.
  rewrite !Nat_iter_S_r; apply IHn. iIntros "HP".
  iMod "HP"; iModIntro; iNext; by iApply HPQ.
 Qed.

 Lemma future_plain_base {E} P `{Hpl: !PlainP P} n :
   (Nat.iter (S n) (λ Q, (|={E}=> Q)) (|={E}=> P))
   |={E}=> Nat.iter (S n) (λ Q, Q) P.
 Proof.
  revert P Hpl. induction n=> P Hpl.
  - simpl. iIntros "HP"; iMod "HP"; by iApply later_fupd_plain.
  - rewrite !(Nat_iter_S_r (S n)). iIntros "Hfpn".
    iApply IHn. iApply step_Sn_fupd_mono; last trivial.
    iIntros "HP"; iMod "HP"; by iApply later_fupd_plain.
 Qed.

 Lemma later_n_mono {M} P Q n :
  (P Q) Nat.iter n (λ Q : uPred M, Q) P Nat.iter n (λ Q, Q) Q.
 Proof.
   revert P Q; induction n=> P Q HPQ; first trivial.
   rewrite !Nat_iter_S_r. by apply IHn, later_mono.
 Qed.

 Lemma later_except_0_n {M} P n :
   Nat.iter (S n) (λ Q : uPred M, Q) P Nat.iter (S n) (λ Q, Q) ( P)%I.
 Proof.
  revert P; induction n=> P; first trivial.
  rewrite !(Nat_iter_S_r (S n)). iIntros "Hfpn".
  iPoseProof (IHn with "Hfpn") as "Hfpn".
  iApply later_n_mono; last trivial.
  apply except_0_later.
 Qed.

  Lemma fupd_plain_forall_comm {E1 E2} (A :Type) (P : A iProp Σ)
   `{Hpl: x, PlainP (P x)} :
  ( x : A, (|={E1,E2}=> P x))%I |={E1}=> ( x : A, P x).
  Proof.
  rewrite fupd_eq /fupd_def; unseal.
  split=> n x Hx Hfp k x' Hkn Hxx' Hwo k' yf Hkk' Hxx'yf.
  exists (x x'); split; first trivial.
  assert (Hk'xx' : ✓{k'} (x x')) by eauto using cmra_validN_le.
  rewrite -uPred_sep_eq. rewrite -uPred_sep_eq in Hwo.
  eapply except_0_mono; first apply sep_assoc';
    last apply except_0_frame_l; trivial.
  rewrite {1}uPred_sep_eq; eexists x', x; repeat split; first by rewrite comm.
  eapply uPred_closed; eauto using cmra_validN_op_r.
  rewrite uPred_sep_eq in Hwo.
  rewrite -uPred_forall_eq; apply except_0_forall'; eauto using cmra_validN_op_l.
  rewrite uPred_forall_eq => u.
  destruct (Hfp u k x' Hkn Hxx' Hwo k (le_n _)) as [z [Hz1 Hz2]];
    first by rewrite right_id.
  assert (✓{k} z) by (by revert Hz1; rewrite right_id).
  eapply (naught_irrel' _ _ z); eauto using cmra_validN_op_l, cmra_validN_le.
  eapply except_0_mono; last (eapply uPred_closed; first eapply Hz2);
    eauto using cmra_validN_le.
  rewrite -uPred_sep_eq; iIntros "(_ & _ & ?)"; iFrame.
 Qed.

 Lemma future_S {E} P n : (|≫{E}=[S n]=> P) ⊣⊢ |={E}=> (|≫{E}=[n]=> P).
 Proof. by rewrite !future_eq. Qed.

 Lemma future_mono {E} Q P n :
  (Q P) (|≫{E}=[n]=> Q) (|≫{E}=[n]=> P).
  Proof.
    revert Q P; induction n => Q P HQP.
    - by rewrite !future_unfold_O; apply fupd_mono.
    - rewrite !future_S; auto using fupd_mono, later_mono.
  Qed.

 Lemma future_plain {E} P `{Hpl: !PlainP P} n :
   (|≫{E}=[n]=> P) |={E}=> Nat.iter n (λ Q, Q) ( P).
 Proof.
   rewrite future_unfold; destruct n; iIntros "HP".
   - by iMod "HP"; iModIntro; iApply except_0_intro.
   - iPoseProof (future_plain_base with "HP") as "HP".
     iMod "HP"; iModIntro; by iApply later_except_0_n.
 Qed.

 Lemma future_plain'_base {E} Q P `{Hpl: !PlainP P} n :
   (Q P) (|≫{E}=[n]=> Q) |={E}=> Nat.iter n (λ Q, Q) ( P).
 Proof.
  iIntros (HQP) "HQ". iApply future_plain. iApply future_mono; eauto.
 Qed.

 Lemma future_plain' {E} Q P `{Hpl: !PlainP P} n :
   (Q P)
   (|≫{E}=[n]=> Q) |={E}=> (|≫{E}=[n]=> Q) (Nat.iter n (λ Q, Q) ( P)).
 Proof.
   intros HQP.
   split => k x Hx HQ.
   pose proof (HQ) as HP; apply (future_plain'_base _ _ _ HQP) in HP; auto.
   revert HP; rewrite fupd_eq /fupd_def; unseal => HP.
   intros k' x' Hkk' Hxx' Hwo k'' yf Hk'k'' Hxx'yf.
   destruct (HP k' x' Hkk' Hxx' Hwo k'' yf Hk'k'' Hxx'yf) as [u [Hvu Hu]].
   rewrite -uPred_sep_eq in Hu.
   assert (Hvu' : ✓{k''} u) by eauto using cmra_validN_op_l.
   do 2 (apply except_0_sep' in Hu; trivial; apply sep_elim_r in Hu; trivial).
   exists (x x'); split; trivial.
   assert (Hk'xx' : ✓{k''} (x x')) by eauto using cmra_validN_le.
   rewrite -uPred_sep_eq. rewrite -uPred_sep_eq in Hwo.
   eapply except_0_mono; first apply sep_assoc';
    last apply except_0_frame_l; trivial.
   rewrite {1}uPred_sep_eq; eexists x', x; repeat split; first by rewrite comm.
   eapply uPred_closed; eauto using cmra_validN_op_r.
   rewrite uPred_sep_eq in Hwo.
   apply except_0_frame_l; eauto using cmra_validN_op_l.
   rewrite uPred_sep_eq. exists x, ( : iResUR Σ); split; first by rewrite right_id.
   split; first (eapply uPred_closed; first apply HQ);
     try omega; eauto using cmra_validN_op_l.
   rewrite -uPred_later_eq; rewrite -uPred_later_eq in Hu.
   (eapply (naught_irrel' _ _ _ ) in Hu); eauto using ucmra_unit_validN.
 Qed.

 Lemma later_n_except_0_future {E} P n :
   Nat.iter n (λ Q, Q) ( P) (|≫{E}=[n]=> P).
 Proof.
  rewrite future_eq; induction n; simpl.
  - by iIntros "HP"; iMod "HP"; iModIntro.
  - by iIntros "HP"; iModIntro; iNext; iApply IHn.
 Qed.

 Lemma future_sep {E} P Q n :
  (|≫{E}=[n]=> P) (|≫{E}=[n]=> Q) (|≫{E}=[n]=> P Q).
 Proof.
  rewrite future_eq; induction n; simpl.
  - iIntros "[HP HQ]"; iMod "HP"; iMod "HQ"; iModIntro; iFrame.
  - iIntros "[HP HQ]";iMod "HP"; iMod "HQ"; iModIntro; iNext; iApply IHn; iFrame.
 Qed.

 Lemma future_cancel_2 {E} P Q Q' m n :
  ((|≫{E}=[m - n]=> P) Q Q')
  (|≫{E}=[m]=> P) (|≫{E}=[n]=> Q) (|≫{E}=[n]=> Q').
 Proof.
  revert m; induction n => m; simpl.
  - replace (m - 0) with m by omega. intros HA.
    iIntros "[HP HQ]". rewrite !future_unfold_O.
    iMod "HQ"; iModIntro; iApply HA; iFrame.
  - destruct m.
    + replace (0 - S n) with (0 - n) by omega. intros HB; specialize (IHn _ HB).
      iIntros "[HP HQ]"; rewrite !future_S. iMod "HQ"; iModIntro; iNext.
      iApply IHn; iFrame.
    + replace (S m - S n) with (m - n) by omega.
      intros HB; specialize (IHn _ HB).
      iIntros "[HP HQ]"; rewrite !future_S. iMod "HP"; iMod "HQ"; iModIntro; iNext.
      iApply IHn; iFrame.
 Qed.

End future_and_plain.

Section future_properties.
 Context `{invG Σ}.

 Global Instance future_ne k E n :
   Proper (dist n ==> dist n) (@future Σ _ E k).
 Proof.
   rewrite future_eq /future_def; induction k; simpl; first by solve_proper.
   by intros P Q HPQ; apply fupd_ne, (contractive_ne _), IHk.
 Qed.

 Global Instance fupd_proper n E : Proper ((≡) ==> (≡)) (@future Σ _ E n).
 Proof. apply ne_proper, _. Qed.

 Lemma fupd_future_future n E P :
   (|={E}=> (|≫{E}=[n]=> P)) ⊣⊢ (|≫{E}=[n]=> P).
 Proof.
   induction n; rewrite future_unfold; simpl; iSplit;
     iIntros "H"; by iMod "H".
 Qed.

 Lemma except_0_future n E P : (|≫{E}=[n]=> P) ≫{E}=[n]=★ P.
 Proof. rewrite -fupd_future_future; apply except_0_fupd. Qed.

 Lemma future_intro n E P : P ≫{E}=[n]=★ P.
 Proof.
   induction n; [rewrite future_unfold|rewrite future_S]; simpl;
     iIntros "H"; first by iModIntro.
   by do 2 iModIntro; iApply IHn.
 Qed.

 Lemma bupd_future n E P : (|==> P) ≫{E}=[n]=★ P.
 Proof.
   induction n; [rewrite future_unfold|rewrite future_S]; simpl;
     iIntros "H"; first by iMod "H"; auto.
   by do 2 iModIntro; iApply IHn.
 Qed.

 Lemma fupd_future n E P : (|={E}=> P) ≫{E}=[n]=★ P.
 Proof.
   induction n; [rewrite future_unfold|rewrite future_S]; simpl;
     iIntros "H"; first by iMod "H"; auto.
   by do 2 iModIntro; iApply IHn.
 Qed.

 Lemma future_trans n m E P : (|≫{E}=[n]=> |≫{E}=[m]=> P) ≫{E}=[n + m]=★ P.
 Proof.
   induction n; simpl.
   - by rewrite future_unfold; rewrite fupd_future_future.
   - by rewrite !future_S; rewrite IHn.
 Qed.

 Lemma future_frame_r n E P Q : (|≫{E}=[n]=> P) Q ≫{E}=[n]=★ P Q.
 Proof.
   induction n.
   - rewrite !future_unfold; apply fupd_frame_r.
   - rewrite !future_S.
     iIntros "[>H1 H2]"; iModIntro; iNext; iApply IHn; by iFrame.
 Qed.

Derived rules

Global Instance future_mono' n E : Proper ((⊢) ==> (⊢)) (@future Σ _ E n).
Proof. intros P Q; apply future_mono. Qed.
Global Instance future_flip_mono' n E :
  Proper (flip (⊢) ==> flip (⊢)) (@future Σ _ E n).
Proof. intros P Q; apply future_mono. Qed.

Lemma fupd_except_0 n E P : (|≫{E}=[n]=> P) ≫{E}=[n]=★ P.
Proof.
  rewrite {1}(future_intro 0 E P) except_0_future future_trans.
  by replace (n + 0) with n by omega.
Qed.

Lemma future_frame_l n E P Q : P (|≫{E}=[n]=> Q) ≫{E}=[n]=★ P Q.
Proof. rewrite !(comm _ P). apply future_frame_r. Qed.
Lemma future_wand_l n E P Q : (P -★ Q) (|≫{E}=[n]=> P) ≫{E}=[n]=★ Q.
Proof. by rewrite future_frame_l wand_elim_l. Qed.
Lemma future_wand_r n E P Q : (|≫{E}=[n]=> P) (P -★ Q) ≫{E}=[n]=★ Q.
Proof. by rewrite future_frame_r wand_elim_r. Qed.

Lemma future_trans_frame n m E P Q :
  ((Q ≫{E}=[n]=★ True) |≫{E}=[m]=> (Q P)) ≫{E}=[m + n]=★ P.
Proof.
  rewrite future_frame_l assoc -(comm _ Q) wand_elim_r.
  by rewrite future_frame_r left_id future_trans.
Qed.

Lemma future_mask_frame_r n E Ef P :
  E Ef (|≫{E}=[n]=> P) ≫{E Ef}=[n]=★ P.
Proof.
  intros HE. induction n.
  - rewrite !future_unfold. iIntros "H".
    iApply fupd_mask_frame_r'; auto.
    iMod "H"; iModIntro. by iIntros.
  - rewrite !future_S. iIntros "H".
    iApply fupd_mask_frame_r'; auto.
    iMod "H"; iModIntro. iIntros; iNext; by iApply IHn.
Qed.

Lemma future_mask_mono n E1 E2 P : E1 E2 (|≫{E1}=[n]=> P) ≫{E2}=[n]=★ P.
Proof.
  intros (Ef&->&?)%subseteq_disjoint_union_L. by apply future_mask_frame_r.
Qed.

Lemma future_big_sepM `{Countable K} {A}
      n E (Φ : K A iProp Σ) (m : gmap K A) :
  ([★ map] kx m, |≫{E}=[n]=> Φ k x) ≫{E}=[n]=★ [★ map] kx m, Φ k x.
Proof.
  apply (big_opM_forall (λ P Q, P ≫{E}=[n]=★ Q)); auto using future_intro.
  intros P1 P2 HP Q1 Q2 HQ. by rewrite HP HQ -future_sep.
Qed.
Lemma future_big_sepS `{Countable A} n E (Φ : A iProp Σ) X :
  ([★ set] x X, |≫{E}=[n]=> Φ x) ≫{E}=[n]=★ [★ set] x X, Φ x.
Proof.
  apply (big_opS_forall (λ P Q, P ≫{E}=[n]=★ Q)); auto using future_intro.
  intros P1 P2 HP Q1 Q2 HQ. by rewrite HP HQ -future_sep.
Qed.

End future_properties.

Section future_proofmode_classes.
  Context `{invG Σ}.
  Implicit Types P Q : iProp Σ.

  Global Instance from_pure_future n E P φ :
    FromPure P φ FromPure (|≫{E}=[n]=> P) φ.
  Proof. rewrite /FromPure. intros <-. apply future_intro. Qed.

  Global Instance from_assumption_future n E p P Q :
    FromAssumption p P (|==> Q) FromAssumption p P (|≫{E}=[n]=> Q)%I.
  Proof. rewrite /FromAssumption=>->. apply bupd_future. Qed.

  Global Instance into_wand_future n E R P Q :
    IntoWand R P Q IntoWand R (|≫{E}=[n]=> P) (|≫{E}=[n]=> Q) | 100.
  Proof.
    rewrite /IntoWand=>->. apply wand_intro_l.
      by rewrite future_wand_r.
  Qed.

  Global Instance from_sep_future n E P Q1 Q2 :
    FromSep P Q1 Q2 FromSep (|≫{E}=[n]=> P) (|≫{E}=[n]=> Q1) (|≫{E}=[n]=> Q2).
  Proof. rewrite /FromSep=><-. apply future_sep. Qed.

  Global Instance or_split_future n E P Q1 Q2 :
    FromOr P Q1 Q2 FromOr (|≫{E}=[n]=> P) (|≫{E}=[n]=> Q1) (|≫{E}=[n]=> Q2).
  Proof.
    rewrite /FromOr=><-. apply or_elim; apply future_mono; auto with I.
  Qed.

  Global Instance exists_split_future {A} n E P (Φ : A iProp Σ) :
    FromExist P Φ FromExist (|≫{E}=[n]=> P) (λ a, |≫{E}=[n]=> Φ a)%I.
  Proof.
    rewrite /FromExist=><-. apply exist_elim=> a. by rewrite -(exist_intro a).
  Qed.

  Global Instance frame_future n E R P Q :
    Frame R P Q Frame R (|≫{E}=[n]=> P) (|≫{E}=[n]=> Q).
  Proof. rewrite /Frame=><-. by rewrite future_frame_l. Qed.

  Global Instance is_except_0_future n E P : IsExcept0 (|≫{E}=[n]=> P).
  Proof. by rewrite /IsExcept0 except_0_future. Qed.

  Global Instance into_modal_future n E P : IntoModal P (|≫{E}=[n]=> P).
  Proof. rewrite /IntoModal. apply future_intro. Qed.

  Global Instance elim_modal_bupd_future n E P Q :
    ElimModal (|==> P) P (|≫{E}=[n]=> Q) (|≫{E}=[n]=> Q).
  Proof.
    by rewrite /ElimModal (bupd_future 0 E) future_frame_r
            wand_elim_r future_trans.
  Qed.
  Global Instance elim_modal_fupd_future n E1 E2 P Q :
    ElimModal (|={E1,E2}=> P) P (|≫{E1}=[n]=> Q) (|={E2,E1}=> (|≫{E1}=[n]=> Q)).
  Proof.
    rewrite /ElimModal; iIntros "[HP HQ]".
    rewrite -{2}fupd_future_future.
    by iMod "HP"; iApply "HQ".
  Qed.

  Class CanElimFuture_by (n m : nat) : Prop := can_elim_future_by : n m.

  Global Instance elim_modal_future_future n m E P Q :
    CanElimFuture_by m n
      ElimModal (|≫{E}=[m]=> P) P (|≫{E}=[n]=> Q) ((|≫{E}=[n - m]=> Q)).
  Proof.
    rewrite /ElimModal /CanElimFuture_by;
      revert n; induction m => n; iIntros (Hnm) "[HP HQ]".
    - rewrite future_unfold_O. replace (n - 0) with n by omega.
      by iMod "HP"; iModIntro; iApply "HQ".
    - destruct n; simpl; first by inversion Hnm.
      rewrite !future_S.
      iMod "HP"; iModIntro; iNext.
      iApply IHm; first omega; iFrame.
  Qed.

End future_proofmode_classes.

Hint Extern 1 (CanElimFuture_by _ _) =>
unfold CanElimFuture_by; abstract omega : typeclass_instances.

Hint Extern 2 (coq_tactics.of_envs _ _) =>
match goal with |- _ |≫{_}=[_]=> _ => iModIntro end.