diff --git a/theories/typing/contextual_refinement.v b/theories/typing/contextual_refinement.v index 7003acef7a029435eb99d98b2916851001d45a07..cd25e7cdabd01c8a6c36a19f41904dd2b995af9d 100644 --- a/theories/typing/contextual_refinement.v +++ b/theories/typing/contextual_refinement.v @@ -303,6 +303,33 @@ Proof. eapply typed_ctx_compose; eauto. Qed. +(* Alternative formulation of contextual refinement +without restricting to contexts of the ground type *) +Definition ctx_refines_alt (Γ : stringmap type) + (e e' : expr) (Ï„ : type) : Prop := ∀ K thp σ₀ σ₠v1 Ï„', + typed_ctx K Γ Ï„ ∅ Ï„' → + 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 ctx_refines_impl_alt Γ e1 e2 Ï„ : + (Γ ⊨ e1 ≤ctx≤ e2 : Ï„) → + ctx_refines_alt Γ e1 e2 Ï„. +Proof. + intros H C thp σ0 σ1 v1 Ï„' HC Hstep. + pose (C' := (CTX_AppR (λ: <>, #true)%E)::C). + cut (∃ (thp' : list expr) σâ‚', + rtc erased_step ([fill_ctx C' e2], σ0) + (of_val #true :: thp', σâ‚')). + - unfold C'; simpl. + destruct 1 as (thp' & σ1' & Hstep'). + exists thp', σ1'. + admit. + - eapply (H C' thp _ σ1 #true TBool). + + repeat econstructor; eauto. + + repeat econstructor; eauto. + + unfold C'. simpl. + admit. +Admitted. Definition ctx_equiv Γ e1 e2 Ï„ := (Γ ⊨ e1 ≤ctx≤ e2 : Ï„) ∧ (Γ ⊨ e2 ≤ctx≤ e1 : Ï„).