diff --git a/theories/typing/interp.v b/theories/typing/interp.v index bcc74bd006e7f214c257627b226c1452f4be2ba0..4268c8507089ae6274a1ecfc94fbc1313ee41dd1 100644 --- a/theories/typing/interp.v +++ b/theories/typing/interp.v @@ -1,6 +1,7 @@ (* ReLoC -- Relational logic for fine-grained concurrency *) (** Interpretations for System F_mu_ref types *) From iris.algebra Require Export list. +From iris.proofmode Require Import tactics. From reloc.typing Require Export types. From reloc.logic Require Import model. From Autosubst Require Import Autosubst. @@ -51,6 +52,35 @@ Section semtypes. apply I. by f_equiv. Defined. + Lemma unboxed_type_sound Ï„ Δ v v' : + UnboxedType Ï„ → + interp Ï„ Δ v v' -∗ ⌜val_is_unboxed v ∧ val_is_unboxed v'âŒ. + Proof. + induction 1; simpl; + first [iIntros "[% %]" + |iDestruct 1 as (?) "[% %]" + |iDestruct 1 as (? ?) "[% [% ?]]"]; + simplify_eq/=; eauto with iFrame. + Qed. + + Lemma eq_type_sound Ï„ Δ v v' : + EqType Ï„ → + interp Ï„ Δ v v' -∗ ⌜v = v'âŒ. + Proof. + intros HÏ„; revert v v'; induction HÏ„; iIntros (v v') "#H1 /=". + - by iDestruct "H1" as "[% %]"; subst. + - by iDestruct "H1" as (n) "[% %]"; subst. + - by iDestruct "H1" as (b) "[% %]"; subst. + - iDestruct "H1" as (?? ??) "[% [% [H1 H2]]]"; simplify_eq/=. + rewrite IHHÏ„1 IHHÏ„2. + by iDestruct "H1" as "%"; iDestruct "H2" as "%"; subst. + - iDestruct "H1" as (??) "[H1|H1]". + + iDestruct "H1" as "[% [% H1]]"; simplify_eq/=. + rewrite IHHÏ„1. by iDestruct "H1" as "%"; subst. + + iDestruct "H1" as "[% [% H1]]"; simplify_eq/=. + rewrite IHHÏ„2. by iDestruct "H1" as "%"; subst. + Qed. + Definition bin_log_related (E : coPset) (Δ : list lty2) (Γ : stringmap type) (e e' : expr) (Ï„ : type) : iProp Σ := diff --git a/theories/typing/types.v b/theories/typing/types.v index a7a08a7277fb9f0df206009e92d25ecc15446894..87da16ae383260f0555821b13b0244f70002a66b 100644 --- a/theories/typing/types.v +++ b/theories/typing/types.v @@ -26,6 +26,13 @@ Inductive EqType : type → Prop := | EqTProd Ï„ Ï„' : EqType Ï„ → EqType Ï„' → EqType (TProd Ï„ Ï„') | EqSum Ï„ Ï„' : EqType Ï„ → EqType Ï„' → EqType (TSum Ï„ Ï„'). +(** Which types are unboxed *) +Inductive UnboxedType : type → Prop := + | UnboxedTUnit : UnboxedType TUnit + | UnboxedTNat : UnboxedType TNat + | UnboxedTBool : UnboxedType TBool + | UnboxedTref Ï„ : UnboxedType (Tref Ï„). + (** Autosubst instances *) Instance Ids_type : Ids type. derive. Defined. Instance Rename_type : Rename type. derive. Defined.