From 2fc79e0d43f51993790792f7893f674ead73cae7 Mon Sep 17 00:00:00 2001 From: Ralf Jung <post@ralfj.de> Date: Sun, 18 Dec 2016 16:07:10 +0100 Subject: [PATCH] old subtyping is not used any more, comment it out --- theories/typing/perm.v | 11 ++++++----- theories/typing/typing.v | 4 ++-- 2 files changed, 8 insertions(+), 7 deletions(-) diff --git a/theories/typing/perm.v b/theories/typing/perm.v index da3c6472..1189e4b8 100644 --- a/theories/typing/perm.v +++ b/theories/typing/perm.v @@ -44,12 +44,12 @@ Section perm. Global Instance top : Top (@perm Σ) := λ _, True%I. - +(* Definition perm_incl (Ï1 Ï2 : perm) := ∀ tid, lft_ctx -∗ Ï1 tid ={⊤}=∗ Ï2 tid. Global Instance perm_equiv : Equiv perm := - λ Ï1 Ï2, perm_incl Ï1 Ï2 ∧ perm_incl Ï2 Ï1. + λ Ï1 Ï2, perm_incl Ï1 Ï2 ∧ perm_incl Ï2 Ï1. *) End perm. Delimit Scope perm_scope with P. @@ -71,14 +71,14 @@ Notation "∃ x .. y , P" := (exist (λ x, .. (exist (λ y, P)) ..)) : perm_scope. Infix "∗" := sep (at level 80, right associativity) : perm_scope. - +(* Infix "⇒" := perm_incl (at level 95, no associativity) : C_scope. Notation "(⇒)" := perm_incl (only parsing) : C_scope. Notation "Ï1 ⇔ Ï2" := (equiv (A:=perm) Ï1%P Ï2%P) (at level 95, no associativity) : C_scope. Notation "(⇔)" := (equiv (A:=perm)) (only parsing) : C_scope. - +*) Section has_type. Context `{typeG Σ}. @@ -106,7 +106,7 @@ Section has_type. iApply ("IH" with "[] [HΦ]"). done. simpl. wp_op. inversion EQν. eauto. Qed. End has_type. - +(* Section perm_incl. Context `{typeG Σ}. @@ -177,3 +177,4 @@ Section perm_incl. Lemma perm_lftincl_trans κ1 κ2 κ3 : κ1 ⊑ κ2 ∗ κ2 ⊑ κ3 ⇒ κ1 ⊑ κ3. Proof. iIntros (tid) "_ [#?#?]!>". iApply (lft_incl_trans with "[] []"); auto. Qed. End perm_incl. +*) \ No newline at end of file diff --git a/theories/typing/typing.v b/theories/typing/typing.v index 3ea3ed1d..a9f97dd8 100644 --- a/theories/typing/typing.v +++ b/theories/typing/typing.v @@ -112,14 +112,14 @@ Section typing. iIntros "[H↦ H↦']". iMod "Hcons" as "[Hown' Hcons]". iMod ("Hcons" with "H↦'") as "[$$]". iApply "Hupd". by iFrame. Qed. - +(* Lemma typed_weaken Ï1 Ï2 e: typed_program Ï2 e → (Ï1 ⇒ Ï2) → typed_program Ï1 e. Proof. iIntros (HÏ2 HÏ12 tid) "!#(#HEAP & #LFT & HÏ1 & Htl)". iApply (HÏ2 with ">"). iFrame "∗#". iApply (HÏ12 with "LFT HÏ1"). Qed. - +*) Lemma typed_program_exists {A} Ï Î¸ e: (∀ x:A, typed_program (Ï âˆ— θ x) e) → typed_program (Ï âˆ— ∃ x, θ x) e. -- GitLab