diff --git a/theories/logrel/session_types.v b/theories/logrel/session_types.v index 9d8bc21a7934ce402477ce9b03e4a64480c15baa..a9a15d6e7f64bcc2a993ac014b2351823e618922 100644 --- a/theories/logrel/session_types.v +++ b/theories/logrel/session_types.v @@ -106,7 +106,11 @@ Section Propers. Global Instance lsty_app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@lsty_app Σ). Proof. apply ne_proper_2, _. Qed. Global Instance lsty_app_assoc : Assoc (≡) (@lsty_app Σ). - Proof. intros s1 s2 s3. rewrite /lsty_app. by rewrite iProto_app_assoc. Qed. + Proof. intros S1 S2 S3. rewrite /lsty_app. by rewrite iProto_app_assoc. Qed. + Global Instance lsty_app_end_l : LeftId (≡) lsty_end (@lsty_app Σ). + Proof. intros [S]. rewrite /lsty_app. by rewrite left_id. Qed. + Global Instance lsty_app_end_r : RightId (≡) lsty_end (@lsty_app Σ). + Proof. intros [S]. rewrite /lsty_app. by rewrite right_id. Qed. End Propers. Notation "'END'" := lsty_end : lsty_scope. diff --git a/theories/logrel/subtyping.v b/theories/logrel/subtyping.v index 9aa01087c66f9112c901824a3f51a5e600bbd264..11e4b5d829a9737d537dfa438f27197930d3ed6e 100644 --- a/theories/logrel/subtyping.v +++ b/theories/logrel/subtyping.v @@ -300,6 +300,22 @@ Section subtype. ⊢ ((S1 <++> S2) <++> S3) <s: (S1 <++> (S2 <++> S3)). Proof. rewrite lsty_app_assoc. iApply lsty_le_refl. Qed. + Lemma lsty_le_app_id_l_l S : + ⊢ (END <++> S) <s: S. + Proof. rewrite left_id. iApply lsty_le_refl. Qed. + + Lemma lsty_le_app_id_l_r S : + ⊢ (S <++> END) <s: S. + Proof. rewrite right_id. iApply lsty_le_refl. Qed. + + Lemma lsty_le_app_id_r_l S : + ⊢ S <s: (END <++> S). + Proof. rewrite left_id. iApply lsty_le_refl. Qed. + + Lemma lsty_le_app_id_r_r S : + ⊢ S <s: (S <++> END). + Proof. rewrite right_id. iApply lsty_le_refl. Qed. + Lemma lsty_le_dual S1 S2 : S2 <s: S1 -∗ lsty_dual S1 <s: lsty_dual S2. Proof. iIntros "#H !>". by iApply iProto_le_dual. Qed.