diff --git a/tests/proofmode.v b/tests/proofmode.v index b2e87333f1da54e6fdf812f54c2611c040fa1ee7..981242fe35063e31a625547e3a3e329877e5493e 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -380,6 +380,11 @@ Proof. iSpecialize ("Hφ" with "[% //] HP"). done. Qed. +Lemma demo_laterN_forall {A} (Φ Ψ: A → PROP) n: (∀ x, ▷^n Φ x) -∗ ▷^n (∀ x, Φ x). +Proof. + iIntros "H" (w). iApply ("H" $! w). +Qed. + Lemma test_iNext_laterN_later P n : ▷ ▷^n P -∗ ▷^n ▷ P. Proof. iIntros "H". iNext. by iNext. Qed. Lemma test_iNext_later_laterN P n : ▷^n ▷ P -∗ ▷ ▷^n P. diff --git a/theories/proofmode/class_instances_sbi.v b/theories/proofmode/class_instances_sbi.v index 5d8aae731ec36fbf1bba8979a007462eb956d3e5..d6eaeb4965abfef957459c8e94e2d9abd375fdfc 100644 --- a/theories/proofmode/class_instances_sbi.v +++ b/theories/proofmode/class_instances_sbi.v @@ -301,6 +301,11 @@ Proof. rewrite /IntoExist=> HP. by rewrite HP plainly_exist. Qed. Global Instance into_forall_later {A} P (Φ : A → PROP) : IntoForall P Φ → IntoForall (▷ P) (λ a, ▷ (Φ a))%I. Proof. rewrite /IntoForall=> HP. by rewrite HP later_forall. Qed. + +Global Instance into_forall_laterN {A} P (Φ : A → PROP) n : + IntoForall P Φ → IntoForall (▷^n P) (λ a, ▷^n (Φ a))%I. +Proof. rewrite /IntoForall=> HP. by rewrite HP laterN_forall. Qed. + Global Instance into_forall_except_0 {A} P (Φ : A → PROP) : IntoForall P Φ → IntoForall (◇ P) (λ a, ◇ (Φ a))%I. Proof. rewrite /IntoForall=> HP. by rewrite HP except_0_forall. Qed. @@ -313,6 +318,11 @@ Proof. rewrite /IntoForall=> HP. by rewrite HP plainly_forall. Qed. Global Instance from_forall_later {A} P (Φ : A → PROP) : FromForall P Φ → FromForall (▷ P)%I (λ a, ▷ (Φ a))%I. Proof. rewrite /FromForall=> <-. by rewrite later_forall. Qed. + +Global Instance from_forall_laterN {A} P (Φ : A → PROP) n : + FromForall P Φ → FromForall (▷^n P)%I (λ a, ▷^n (Φ a))%I. +Proof. rewrite /FromForall => <-. by rewrite laterN_forall. Qed. + Global Instance from_forall_except_0 {A} P (Φ : A → PROP) : FromForall P Φ → FromForall (◇ P)%I (λ a, ◇ (Φ a))%I. Proof. rewrite /FromForall=> <-. by rewrite except_0_forall. Qed.