diff --git a/tests/proofmode.v b/tests/proofmode.v index 1c51da0645d7f0229fd5cd15ee1199b982a4b0b3..f657edfe7c92084191dc6dc285e6b1e0c55b7972 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -5,6 +5,10 @@ Section tests. Context {PROP : sbi}. Implicit Types P Q R : PROP. +Lemma type_sbi_internal_eq_annot P : + P ≡@{PROP} P ∧ (≡@{PROP}) P P ∧ (P ≡.) P ∧ (.≡ P) P. +Proof. done. Qed. + Lemma test_eauto_emp_isplit_biwand P : emp ⊢ P ∗-∗ P. Proof. eauto 6. Qed. diff --git a/theories/bi/interface.v b/theories/bi/interface.v index 1d0895ae68d40a0570b0be3258807d8ec68a541f..41313692c789b9978168d660fa159f949ca1f059 100644 --- a/theories/bi/interface.v +++ b/theories/bi/interface.v @@ -290,6 +290,12 @@ Notation "∃ x .. y , P" := Notation "'<pers>' P" := (bi_persistently P) : bi_scope. Infix "≡" := sbi_internal_eq : bi_scope. +Infix "≡@{ A }" := (sbi_internal_eq (A := A)) (only parsing) : bi_scope. + +Notation "( X ≡.)" := (sbi_internal_eq X) (only parsing) : bi_scope. +Notation "(.≡ X )" := (λ Y, Y ≡ X)%I (only parsing) : bi_scope. +Notation "(≡@{ A } )" := (sbi_internal_eq (A := A)) (only parsing) : bi_scope. + Notation "▷ P" := (sbi_later P) : bi_scope. Definition bi_emp_valid {PROP : bi} (P : PROP) : Prop := emp ⊢ P.