diff --git a/theories/typing/contextual_refinement.v b/theories/typing/contextual_refinement.v index 5f47dae793af6b4e45ad197a217d274208bd1842..bbed1775e30441f8eb77558d09ffebb7b28a186b 100644 --- a/theories/typing/contextual_refinement.v +++ b/theories/typing/contextual_refinement.v @@ -239,7 +239,7 @@ Inductive typed_ctx: ctx → stringmap type → type → stringmap type → type typed_ctx (k :: K) Γ1 Ï„1 Γ3 Ï„3. (* Observable types are, at the moment, exactly the types which support equality. *) -Definition ObsType : type → Prop := EqType. +Definition ObsType : type → Prop := λ Ï„, EqType Ï„ ∧ UnboxedType Ï„. Definition ctx_refines (Γ : stringmap type) (e e' : expr) (Ï„ : type) : Prop := ∀ K thp σ₀ σ₠v Ï„', @@ -311,20 +311,6 @@ Definition ctx_refines_alt (Γ : stringmap type) rtc erased_step ([fill_ctx K e], σ₀) (of_val v1 :: thp, σâ‚) → ∃ thp' σâ‚' v2, rtc erased_step ([fill_ctx K e'], σ₀) (of_val v2 :: thp', σâ‚'). -(* Lemma erased_step_ectx (e e' : expr) tp' σ σ' K : *) -(* erased_step ([e], σ) (e' :: tp', σ') → *) -(* erased_step ([fill K e], σ) (fill K e' :: tp', σ'). *) -(* Proof. *) -(* intros [κ Hst]. inversion Hst; simplify_eq/=. *) -(* symmetry in H. apply app_singleton in H. *) -(* assert (t1 = [] ∧ e1 = e ∧ t2 = []) as (->&->&->). *) -(* { naive_solver. } *) -(* assert (e2 = e' ∧ tp' = efs) as [-> ->]. *) -(* { naive_solver. } *) -(* eapply fill_prim_step in H1. simpl in H1. *) -(* econstructor. eapply step_atomic with (t1 := []); eauto. *) -(* Qed. *) - (* Helper lemmas *) Lemma erased_step_ectx (e e' : expr) tp tp' σ σ' K : erased_step (e :: tp, σ) (e' :: tp', σ') → @@ -340,14 +326,6 @@ Proof. + rewrite app_comm_cons. reflexivity. Qed. - -Local Definition ffill (K : list ectx_item) : - (list expr * state) → (list expr * state) := - fun x => match x with - | (e :: tp, σ) => (fill K e :: tp, σ) - | ([], σ) => ([], σ) - end. - Lemma erased_step_nonempty (tp : list expr) σ tp' σ' : erased_step (tp, σ) (tp', σ') → tp' ≠[]. Proof. @@ -366,6 +344,14 @@ Proof. intros ? ?%erased_step_nonempty. naive_solver. Qed. + +Local Definition ffill (K : list ectx_item) : + (list expr * state) → (list expr * state) := + fun x => match x with + | (e :: tp, σ) => (fill K e :: tp, σ) + | ([], σ) => ([], σ) + end. + Lemma rtc_erased_step_ectx (e e' : expr) tp tp' σ σ' K : rtc erased_step (e :: tp, σ) (e' :: tp', σ') → rtc erased_step (fill K e :: tp, σ) (fill K e' :: tp', σ'). @@ -434,6 +420,7 @@ Proof. fold Ï1 Ï2. intros HÏ. change σ1 with Ï1.2. eapply nice_ctx_lemma; eauto. Qed. +(*/Helper lemmas*) Lemma ctx_refines_impl_alt Γ e1 e2 Ï„ : (Γ ⊨ e1 ≤ctx≤ e2 : Ï„) → diff --git a/theories/typing/types.v b/theories/typing/types.v index 4c76bdb3ef867f7810b098c0b5cfbced6b6c728a..574390092bd740874387119b6c55a14a3140bf3b 100644 --- a/theories/typing/types.v +++ b/theories/typing/types.v @@ -18,7 +18,7 @@ Inductive type := | TExists : {bind 1 of type} → type | TRef : type → type. -(** Which types support equality testing *) +(** Types which support direct equality test (which coincides with ctx equiv) *) Inductive EqType : type → Prop := | EqTUnit : EqType TUnit | EqTNat : EqType TNat