From 71292adc1bcbfddde5cea5b250d0a3743a15e789 Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Thu, 27 Nov 2025 18:28:33 +0800 Subject: [PATCH 01/11] Port lang.v to use Binding --- ctx-equiv-ixfree/examples.v | 17 +- ctx-equiv-ixfree/lang.v | 3268 +++++++++++++++++++---------------- 2 files changed, 1760 insertions(+), 1525 deletions(-) diff --git a/ctx-equiv-ixfree/examples.v b/ctx-equiv-ixfree/examples.v index fa6a434..6f60ad3 100644 --- a/ctx-equiv-ixfree/examples.v +++ b/ctx-equiv-ixfree/examples.v @@ -90,25 +90,11 @@ Abort. (* cannot be proven *) Definition l_rel e1 e2 : Prop := ∀ n, n ⊨ L_rel e1 e2. -Definition nbigstep n e v := ∃ e', nsteps contextual_step n e e' ∧ to_val e' = Some v. -Definition nterminates n e := ∃ v, nbigstep n e v. - Lemma to_val_eq_Some e v : to_val e = Some v → e = ret v. Proof. destruct e; simpl; congruence. Qed. -Lemma nterminates_zero e : - nterminates O e → - ∃ (v : val), e = v. -Proof. - unfold nterminates, nbigstep. - intros (v & e' & Hnsteps & Heq). - inversion Hnsteps. subst. - apply to_val_eq_Some in Heq. - exists v. exact Heq. -Qed. - Lemma nterminates_succ e n : nterminates (S n) e → @@ -210,8 +196,7 @@ Proof. * exact He1'. Qed. -Lemma terminates_impl_nterminates e : - terminates e → ∃ n, nterminates n e. + Proof. unfold terminates, nterminates. unfold bigstep, nbigstep. diff --git a/ctx-equiv-ixfree/lang.v b/ctx-equiv-ixfree/lang.v index 457b952..0d817aa 100644 --- a/ctx-equiv-ixfree/lang.v +++ b/ctx-equiv-ixfree/lang.v @@ -1,133 +1,130 @@ - +From stdpp Require Export relations. +(** Note: both stdpp and Binding define `fmap`, but with different + signature. What we want in this file is Binding's `fmap`, thus + we Require Import Binding after stdpp *) +From Stdlib Require Import Utf8. +From Binding Require Import Lib Auto. +Require Import Binding.Set. From IxFree Require Import Lib Nat. From CtxEquivIxFree Require Import ixfree_tactics. From CtxEquivIxFree Require Import tactics. -(* From stdpp Require Export binders. *) -From stdpp Require Export gmap. -From stdpp Require Export strings. + +Local Close Scope stdpp_scope. Definition loc : Set := nat. -Inductive expr := - (* | eint (z : Z) *) - (* | eunit *) - | ret (v:val) - (* | ens (H : hprop) *) - (* | ensp (P : prop) *) - (* | var (x : binder) *) - | var (x : string) - (* | bind (e1 : expr) (x : binder) (e2 : expr) *) - | app (e1 e2: expr) - (* | abs (x : binder) (e : expr) *) - (* | eplus (e1 e2: expr) *) - -with val := - (* TODO reenable context tests when this is restored *) - (*| vint (z : Z) *) - | vlambda (x : string) (e: expr) - (* | vbool (b : bool) *) - (* | vloc (l : loc) *) - (* | vunit *). - -Definition to_val (e : expr) : option val := - match e with - | ret v => Some v - (* | abs x e => Some (vlambda x e) *) - | _ => None - end. +Inductive expr (V : Set) : Set := +| ret (v : val V) +| app (e1 e2 : expr V) -Inductive hprop := - | emp - | pts (l : loc) (v : val) - | sep (h1 h2: hprop). +with val (V : Set) : Set := +| v_var (x : V) +| v_lambda (e : expr (inc V)). -Fixpoint subst_val (x : string) (es : expr) (v : val) : val := - match v with - (*| vunit => vunit - | vint n => vint n*) - | vlambda y e => - vlambda y $ if decide (x = y) then e else - (subst x es e) - end +Arguments ret {V} v. +Arguments app {V} e1 e2. +Arguments v_var {V} x. +Arguments v_lambda {V} e. -with subst (x : string) (es : expr) (e : expr) : expr := +#[global] +Instance SetPureCore_value : SetPureCore val := + { set_pure := @v_var }. + +Fixpoint emap {A B} (f : A [→] B) (e : expr A) : expr B := match e with - | ret v => ret (subst_val x es v) - (* The function [decide] can be used to decide propositions. - [decide P] is of type {P} + {¬ P}. - It can only be applied to propositions for which, by type class inference, - it can be determined that the proposition is decidable. *) - | var y => if decide (x = y) then es else var y - (* | abs y e => - abs y $ if decide (BNamed x = y) then e else subst x es e *) - | app e1 e2 => app (subst x es e1) (subst x es e2) - (* | eplus e1 e2 => eplus (subst x es e1) (subst x es e2) *) + | ret v => ret (vmap f v) + | app e1 e2 => app (emap f e1) (emap f e2) + end + +with vmap {A B} (f : A [→] B) (v : val A) : val B := + match v with + | v_var x => v_var (f x) + | v_lambda e => v_lambda (emap (lift f) e) end. -(* Definition subst' (mx : binder) (es : expr) : expr → expr := - match mx with BNamed x => subst x es | BAnon => id end. *) +#[global] +Instance FunctorCore_emap : FunctorCore expr := @emap. + +#[global] +Instance FunctorCore_vmap : FunctorCore val := @vmap. -Definition is_val (e : expr) : Prop := +Fixpoint ebind {A B} (f : A [⇒] B) (e : expr A) : expr B := match e with - | ret v => True - (* | abs x e => True *) - | _ => False - end. + | ret v => ret (vbind f v) + | app e1 e2 => app (ebind f e1) (ebind f e2) + end -(* Definition of_val (v : val) : expr := +with vbind {A B} (f : A [⇒] B) (v : val A) : val B := match v with - | vunit => eunit - | vint n => eint n - | vlambda x e => abs x e - end. *) + | v_var x => f x + | v_lambda e => v_lambda (ebind (lift f) e) + end. + +#[global] +Instance BindCore_ebind : BindCore expr := @ebind. + +#[global] +Instance BindCore_vbind : BindCore val := @vbind. -Notation of_val := ret. Coercion ret : val >-> expr. -Inductive base_step : expr → expr → Prop := - | BetaS x e1 e2 e' : - is_val e2 → - e' = subst x e2 e1 → - base_step (app (ret (vlambda x e1)) e2) e' - (* | PlusS e1 e2 (n1 n2 n3 : Z): - e1 = (eint n1) → - e2 = (eint n2) → - (n1 + n2)%Z = n3 → - base_step (Plus e1 e2) (eint n3) *) -. - -(* inside-out contexts, similar to a "reversed" list *) -Inductive ectx := - | ectx_hole : ectx - | ectx_app1 : ectx → expr → ectx - | ectx_app2 : val → ectx → ectx. +(** Inside-out contexts, similar to a "reversed" list *) (* Imagine the list is from left-to-right, with the following structure: ectx_hole ... ectx_app1 e ... ectx_app1 e ... ectx_app2 v. The actual structure is zig-zag, but let's linearize it so that we can implement and reason about this data-structure just like a list *) +Inductive ectx (V : Set) := + | ectx_hole + | ectx_app1 (E : ectx V) (e : expr V) + | ectx_app2 (v : val V) (E : ectx V). + +Arguments ectx_hole {V}. +Arguments ectx_app1 {V} E e. +Arguments ectx_app2 {V} v E. + +Fixpoint ectx_map {A B} (f : A [→] B) (E : ectx A) : ectx B := + match E with + | ectx_hole => ectx_hole + | ectx_app1 E' e => ectx_app1 (ectx_map f E') (fmap f e) + | ectx_app2 v E' => ectx_app2 (fmap f v) (ectx_map f E') + end. + +#[global] +Instance FunctorCore_ectx_map : FunctorCore ectx := @ectx_map. + +Fixpoint ectx_bind {A B} (f : A [⇒] B) (E : ectx A) : ectx B := + match E with + | ectx_hole => ectx_hole + | ectx_app1 E' e => ectx_app1 (ectx_bind f E') (bind f e) + | ectx_app2 v E' => ectx_app2 (bind f v) (ectx_bind f E') + end. + +#[global] +Instance BindCore_ectx_bind : BindCore ectx := @ectx_bind. (* similar to foldr of a "reversed" list (foldl of a normal list) *) -(* ectx_hole -----> *) -Fixpoint plug (E : ectx) (e : expr) : expr := +Fixpoint plug {V} (E : ectx V) (e : expr V) : expr V := match E with | ectx_hole => e | ectx_app1 E' e' => plug E' (app e e') | ectx_app2 v E' => plug E' (app v e) end. +Lemma fold_unfold_plug_ectx_hole {V} (e : expr V) : + plug ectx_hole e = e. +Proof. auto. Qed. + (* similar to "prepend" of a "reversed" list ("append" of a normal list) *) -Fixpoint ectx_comp (E1 E2 : ectx) : ectx := +Fixpoint ectx_comp {V} (E1 E2 : ectx V) : ectx V := match E2 with | ectx_hole => E1 | ectx_app1 E2' e => ectx_app1 (ectx_comp E1 E2') e | ectx_app2 v E2' => ectx_app2 v (ectx_comp E1 E2') end. -Notation fill := plug. - -Lemma ectx_comp_assoc E1 E2 E3 : +Lemma ectx_comp_assoc {V} (E1 E2 E3 : ectx V) : ectx_comp E1 (ectx_comp E2 E3) = ectx_comp (ectx_comp E1 E2) E3. Proof. induction E3; simpl. @@ -136,123 +133,159 @@ Proof. - rewrite -> IHE3. reflexivity. Qed. -Lemma ectx_comp_correct E1 E2 e : +Lemma ectx_comp_correct {V} (E1 E2 : ectx V) (e : expr V) : plug (ectx_comp E1 E2) e = plug E1 (plug E2 e). Proof. revert e. induction E2; intros e'. - - simpl in *. reflexivity. - - simpl in *. rewrite (IHE2 (app e' e)). reflexivity. - - simpl in *. rewrite (IHE2 (app v e')). reflexivity. + - simpl. reflexivity. + - simpl. rewrite -> (IHE2 (app e' e)). reflexivity. + - simpl. rewrite -> (IHE2 (app v e')). reflexivity. Qed. -(** Outside-in evaluation contexts *) -(* similar to a normal list *) -Inductive rctx := - | rctx_hole : rctx - | rctx_app1 : rctx → expr → rctx - | rctx_app2 : val → rctx → rctx. +(** Outside-in evaluation contexts, similar to a normal list *) + +Inductive rctx (V : Set) := + | rctx_hole + | rctx_app1 (R : rctx V) (e : expr V) + | rctx_app2 (v : val V) (R : rctx V). + +Arguments rctx_hole {V}. +Arguments rctx_app1 {V} R e. +Arguments rctx_app2 {V} v R. + +Fixpoint rctx_map {A B} (f : A [→] B) (R : rctx A) : rctx B := + match R with + | rctx_hole => rctx_hole + | rctx_app1 R' e => rctx_app1 (rctx_map f R') (fmap f e) + | rctx_app2 v R' => rctx_app2 (fmap f v) (rctx_map f R') + end. + +#[global] +Instance FunctorCore_rctx_map : FunctorCore rctx := @rctx_map. + +Fixpoint rctx_bind {A B} (f : A [⇒] B) (R : rctx A) : rctx B := + match R with + | rctx_hole => rctx_hole + | rctx_app1 R' e => rctx_app1 (rctx_bind f R') (bind f e) + | rctx_app2 v R' => rctx_app2 (bind f v) (rctx_bind f R') + end. + +#[global] +Instance BindCore_rctx_bind : BindCore rctx := @rctx_bind. (* similar to foldr of a normal list *) -Fixpoint rplug (E : rctx) (e : expr) : expr := - match E with +Fixpoint rplug {V} (R : rctx V) (e : expr V) : expr V := + match R with | rctx_hole => e - | rctx_app1 E' e1 => app (rplug E' e) e1 - | rctx_app2 v E' => app v (rplug E' e) + | rctx_app1 R' e' => app (rplug R' e) e' + | rctx_app2 v R' => app v (rplug R' e) end. -Fixpoint rctx_comp (E1 E2 : rctx) : rctx := - match E1 with - | rctx_hole => E2 - | rctx_app1 E1' e => rctx_app1 (rctx_comp E1' E2) e - | rctx_app2 v E1' => rctx_app2 v (rctx_comp E1' E2) +(* similar to append of a normal list *) +Fixpoint rctx_comp {V} (R1 R2 : rctx V) : rctx V := + match R1 with + | rctx_hole => R2 + | rctx_app1 R1' e => rctx_app1 (rctx_comp R1' R2) e + | rctx_app2 v R1' => rctx_app2 v (rctx_comp R1' R2) end. -Lemma rctx_comp_assoc (E1 E2 E3 : rctx) : - rctx_comp (rctx_comp E1 E2) E3 = rctx_comp E1 (rctx_comp E2 E3). +Lemma rctx_comp_assoc {V} (R1 R2 R3 : rctx V) : + rctx_comp (rctx_comp R1 R2) R3 = rctx_comp R1 (rctx_comp R2 R3). Proof. - induction E1. + induction R1. - simpl. reflexivity. - - simpl. rewrite -> IHE1. reflexivity. - - simpl. rewrite -> IHE1. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. Qed. -(* similar to reverse_prepend : reverse E, and then prepend F to it *) -Fixpoint ectx_comp_rctx1 (F : ectx) (E : rctx) : ectx := - match E with - | rctx_hole => F - | rctx_app1 E e => ectx_comp_rctx1 (ectx_app1 F e) E - | rctx_app2 v E => ectx_comp_rctx1 (ectx_app2 v F) E +Lemma rctx_comp_correct {V} (R1 R2 : rctx V) (e : expr V) : + rplug (rctx_comp R1 R2) e = rplug R1 (rplug R2 e). +Proof. + induction R1. + - simpl. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. +Qed. + +(** Equivalence of ectx and rctx *) + +(* similar to reverse_prepend : reverse R, and then prepend E to it *) +Fixpoint ectx_comp_rctx1 {V} (E : ectx V) (R : rctx V) : ectx V := + match R with + | rctx_hole => E + | rctx_app1 R' e => ectx_comp_rctx1 (ectx_app1 E e) R' + | rctx_app2 v R' => ectx_comp_rctx1 (ectx_app2 v E) R' end. (* similar to reverse *) -Definition rctx_to_ectx : rctx -> ectx := ectx_comp_rctx1 ectx_hole. +Definition rctx_to_ectx {V} : rctx V -> ectx V := + ectx_comp_rctx1 ectx_hole. -Lemma ectx_comp_rctx1_correct (F : ectx) (E : rctx) (e : expr) : - plug (ectx_comp_rctx1 F E) e = plug F (rplug E e). +Lemma ectx_comp_rctx1_correct {V} (E : ectx V) (R : rctx V) (e : expr V) : + plug (ectx_comp_rctx1 E R) e = plug E (rplug R e). Proof. - revert F. - induction E; intros F. + revert E. + induction R; intros E. - simpl. reflexivity. - - simpl. rewrite -> (IHE (ectx_app1 F e0)). simpl. reflexivity. - - simpl. rewrite -> (IHE (ectx_app2 v F)). simpl. reflexivity. + - simpl. rewrite -> (IHR (ectx_app1 E _)). simpl. reflexivity. + - simpl. rewrite -> (IHR (ectx_app2 _ E)). simpl. reflexivity. Qed. -(* similar to reverse_append : reverse E, and then append to F *) -(* E1 ... En | F1 ... Fn ~> En ... E1 F1 ... Fn *) -Fixpoint ectx_comp_rctx2 (E : ectx) (F : rctx) : rctx := +(* similar to reverse_append : reverse E, and then append to R *) +Fixpoint ectx_comp_rctx2 {V} (E : ectx V) (R : rctx V) : rctx V := match E with - | ectx_hole => F - | ectx_app1 E e => ectx_comp_rctx2 E (rctx_app1 F e) - | ectx_app2 v E => ectx_comp_rctx2 E (rctx_app2 v F) + | ectx_hole => R + | ectx_app1 E e => ectx_comp_rctx2 E (rctx_app1 R e) + | ectx_app2 v E => ectx_comp_rctx2 E (rctx_app2 v R) end. -Definition ectx_to_rctx (E : ectx) : rctx := +Definition ectx_to_rctx {V} (E : ectx V) : rctx V := ectx_comp_rctx2 E rctx_hole. -Lemma ectx_comp_rctx2_correct (E : ectx) (F : rctx) (e : expr) : - rplug (ectx_comp_rctx2 E F) e = plug E (rplug F e). +Lemma ectx_comp_rctx2_correct {V} (E : ectx V) (R : rctx V) (e : expr V) : + rplug (ectx_comp_rctx2 E R) e = plug E (rplug R e). Proof. - revert F. - induction E; intros F. + revert R. + induction E; intros R. - simpl. reflexivity. - - simpl. rewrite -> (IHE (rctx_app1 F e0)). simpl. reflexivity. - - simpl. rewrite -> (IHE (rctx_app2 v F)). simpl. reflexivity. + - simpl. rewrite -> (IHE (rctx_app1 R _)). simpl. reflexivity. + - simpl. rewrite -> (IHE (rctx_app2 _ R)). simpl. reflexivity. Qed. -Lemma ectx_comp_rctx1_reset (F : ectx) (E : rctx) : - ectx_comp_rctx1 F E = ectx_comp F (ectx_comp_rctx1 ectx_hole E). +Lemma ectx_comp_rctx1_reset {V} (E : ectx V) (R : rctx V) : + ectx_comp_rctx1 E R = ectx_comp E (ectx_comp_rctx1 ectx_hole R). Proof. - revert F. - induction E; intros F; simpl. + revert E. + induction R; intros E; simpl. - reflexivity. - - rewrite -> (IHE (ectx_app1 F e)). - rewrite -> (IHE (ectx_app1 ectx_hole e)). + - rewrite -> (IHR (ectx_app1 E e)). + rewrite -> (IHR (ectx_app1 ectx_hole e)). rewrite -> ectx_comp_assoc. simpl. reflexivity. - - rewrite -> (IHE (ectx_app2 v F)). - rewrite -> (IHE (ectx_app2 v ectx_hole)). + - rewrite -> (IHR (ectx_app2 v E)). + rewrite -> (IHR (ectx_app2 v ectx_hole)). rewrite -> ectx_comp_assoc. simpl. reflexivity. Qed. -Lemma ectx_comp_rctx2_reset (E : ectx) (F : rctx) : - ectx_comp_rctx2 E F = rctx_comp (ectx_comp_rctx2 E rctx_hole) F. +Lemma ectx_comp_rctx2_reset {V} (E : ectx V) (R : rctx V) : + ectx_comp_rctx2 E R = rctx_comp (ectx_comp_rctx2 E rctx_hole) R. Proof. - revert F. - induction E; intros F; simpl. + revert R. + induction E; intros R; simpl. - reflexivity. - - rewrite -> (IHE (rctx_app1 F e)). + - rewrite -> (IHE (rctx_app1 R e)). rewrite -> (IHE (rctx_app1 rctx_hole e)). rewrite -> rctx_comp_assoc. simpl. reflexivity. - - rewrite -> (IHE (rctx_app2 v F)). + - rewrite -> (IHE (rctx_app2 v R)). rewrite -> (IHE (rctx_app2 v rctx_hole)). rewrite -> rctx_comp_assoc. simpl. reflexivity. Qed. -Lemma ectx_rctx_bijection1_aux (E : ectx) (R : rctx) : +Lemma ectx_rctx_bijection_aux {V} (E : ectx V) (R : rctx V) : ectx_comp_rctx1 ectx_hole (ectx_comp_rctx2 E R) = ectx_comp_rctx1 E R. Proof. revert R. @@ -262,15 +295,14 @@ Proof. - simpl. rewrite -> (IHE (rctx_app2 v R)). simpl. reflexivity. Qed. -Lemma ectx_rctx_bijection1 E : +Lemma ectx_rctx_bijection {V} (E : ectx V) : rctx_to_ectx (ectx_to_rctx E) = E. Proof. unfold rctx_to_ectx, ectx_to_rctx. - rewrite -> (ectx_rctx_bijection1_aux E rctx_hole). - simpl. reflexivity. + exact (ectx_rctx_bijection_aux E rctx_hole). Qed. -Lemma ectx_rctx_bijection2_aux (E : ectx) (R : rctx) : +Lemma rctx_ectx_bijection_aux {V} (E : ectx V) (R : rctx V) : ectx_comp_rctx2 (ectx_comp_rctx1 E R) rctx_hole = ectx_comp_rctx2 E R. Proof. revert E. @@ -280,623 +312,617 @@ Proof. - simpl. rewrite -> (IHR (ectx_app2 v E)). simpl. reflexivity. Qed. -Lemma ectx_rctx_bijection2 R : +Lemma rctx_ectx_bijection {V} (R : rctx V) : ectx_to_rctx (rctx_to_ectx R) = R. Proof. unfold ectx_to_rctx, rctx_to_ectx. - rewrite -> (ectx_rctx_bijection2_aux ectx_hole R). - simpl. reflexivity. + exact (rctx_ectx_bijection_aux ectx_hole R). Qed. -Lemma plug_rplug_equiv E e : +Lemma plug_rplug_equiv {V} (E : ectx V) (e : expr V) : plug E e = rplug (ectx_to_rctx E) e. Proof. unfold ectx_to_rctx. - rewrite -> (ectx_comp_rctx2_correct E rctx_hole e). - simpl. reflexivity. + rewrite -> ectx_comp_rctx2_correct. simpl. reflexivity. Qed. -Inductive contextual_step (e1 : expr) (e2 : expr) : Prop := - Ectx_step K e1' e2' : - e1 = fill K e1' → - e2 = fill K e2' → - base_step e1' e2' → - contextual_step e1 e2. +Lemma plug_eq_val_inv {V} E e (v : val V) : + plug E e = v → + E = ectx_hole ∧ e = v. +Proof. + revert e. + induction E; intros e' H_eq; simpl in *. + - auto. + - specialize (IHE _ H_eq) as (_ & H_absurd). discriminate. + - specialize (IHE _ H_eq) as (_ & H_absurd). discriminate. +Qed. -Definition contextual_reducible (e : expr) := - ∃ e', contextual_step e e'. +Lemma rplug_eq_val_inv {V} R e (v : val V) : + rplug R e = v → + R = rctx_hole ∧ e = v. +Proof. + intros H_eq. + destruct R; simpl in *. + - auto. + - discriminate. + - discriminate. +Qed. -Definition bigstep e1 (v:val) := - ∃ e2, rtc contextual_step e1 e2 ∧ to_val e2 = Some v. +(** Properties of syntax *) -Definition terminates e := ∃ v, bigstep e v. +Lemma fmap_plug {A B} (f : A [→] B) (E : ectx A) (e : expr A) : + fmap f (plug E e) = plug (fmap f E) (fmap f e). +Proof. + revert e. + induction E; intros e'. + - term_simpl. reflexivity. + - term_simpl. rewrite -> (IHE (app e' e)). term_simpl. reflexivity. + - term_simpl. rewrite -> (IHE (app v e')). term_simpl. reflexivity. +Qed. -Lemma terminates_val v : - terminates (ret v). +Lemma bind_plug {A B} (f : A [⇒] B) (E : ectx A) (e : expr A) : + bind f (plug E e) = plug (bind f E) (bind f e). Proof. - exists v. - exists v. - split; done. + revert e. + induction E; intros e'. + - term_simpl. reflexivity. + - term_simpl. rewrite -> (IHE (app e' e)). term_simpl. reflexivity. + - term_simpl. rewrite -> (IHE (app v e')). term_simpl. reflexivity. Qed. -Lemma contextual_step_comp : - ∀ K e1 e2, - contextual_step e1 e2 → - contextual_step (fill K e1) (fill K e2). +Lemma subst_plug {V} (E : ectx (inc V)) e v : + subst (plug E e) v = plug (subst E v) (subst e v). Proof. - intros K e1 e2 H_step. - inversion H_step. subst. econstructor. - - rewrite ectx_comp_correct. reflexivity. - - rewrite ectx_comp_correct. reflexivity. - - assumption. + revert e. + induction E as [| E IHE e' | v' E IHE]; intros e. + - term_simpl. reflexivity. + - term_simpl. rewrite -> (IHE (app e e')). term_simpl. reflexivity. + - term_simpl. rewrite -> (IHE (app v' e)). term_simpl. reflexivity. Qed. -Notation name := string. -Definition sub : Set := gmap name val. -Definition scope : Set := gset name. +#[global] Hint Rewrite @fmap_plug : term_simpl. +#[global] Hint Rewrite @bind_plug : term_simpl. +#[global] Hint Rewrite @subst_plug : term_simpl. -Fixpoint subst_map_val (γ : sub) (v : val) : val := - match v with - (*| vunit => vunit - | vint n => vint n*) - | vlambda x e => vlambda x (subst_map (delete x γ) e) - end -with subst_map (γ : sub) (e : expr) : expr := - match e with - | ret v => ret (subst_map_val γ v) - (* | eunit => eunit *) - | var x => match γ !! x with Some v => ret v | _ => var x end - | app e1 e2 => app (subst_map γ e1) (subst_map γ e2) - (* | abs x e => abs x (subst_map (binder_delete x xs) e) *) - (* | eplus e1 e2 => eplus (subst_map xs e1) (subst_map xs e2) *) - end. +#[global] +Instance SetPure_val : SetPure val. +Proof. + split. + - simpl. unfold SetPureCore_value. term_simpl. reflexivity. + - simpl. unfold SetPureCore_value. term_simpl. reflexivity. +Qed. -Lemma fold_unfold_subst_map_val_vlambda γ x e : - subst_map_val γ (vlambda x e) = - vlambda x (subst_map (delete x γ) e). -Proof. auto. Qed. +(** Functor instances *) -Lemma fold_unfold_subst_map_ret γ v : - subst_map γ (ret v) = - ret (subst_map_val γ v). -Proof. auto. Qed. +Fixpoint emap_id {A} (f : A [→] A) (e : expr A) : + equal f (arrow_id A) → fmap f e = e +with vmap_id {A} (f : A [→] A) (v : val A) : + equal f (arrow_id A) → fmap f v = v. +Proof. + - auto_map_id. + - auto_map_id. +Qed. -Lemma fold_unfold_subst_map_var γ x : - subst_map γ (var x) = - match γ !! x with - | Some v => ret v - | _ => var x - end. -Proof. auto. Qed. +Fixpoint emap_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (e : expr A) : + equal (arrow_comp f g) h → fmap f (fmap g e) = fmap h e +with vmap_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (v : val A) : + equal (arrow_comp f g) h → fmap f (fmap g v) = fmap h v. +Proof. + - auto_map_comp. + - auto_map_comp. +Qed. -Lemma fold_unfold_subst_map_app γ e1 e2 : - subst_map γ (app e1 e2) = - app (subst_map γ e1) (subst_map γ e2). -Proof. auto. Qed. +#[global] +Instance Functor_expr : Functor expr. +Proof. + constructor. + - exact @emap_id. + - exact @emap_comp. +Qed. -Fixpoint is_closed (X : scope) (e : expr) : bool := - match e with - | var x => bool_decide (x ∈ X) - (*| ret vunit | ret (vint _) => true*) - | ret (vlambda x e) => is_closed (X ∪ {[x]}) e - | app e1 e2 - (* | eplus e1 e2 *) - => is_closed X e1 && is_closed X e2 - end. +#[global] +Instance Functor_val : Functor val. +Proof. + constructor. + - exact @vmap_id. + - exact @vmap_comp. +Qed. -(* aka a scoping judgment, the untyped equivalent of a typing judgment *) -Definition closed (X : scope) (e : expr) : Prop := Is_true (is_closed X e). +Fixpoint ectx_map_id {A} (f : A [→] A) (E : ectx A) : + equal f (arrow_id A) → fmap f E = E. +Proof. auto_map_id. Qed. -Lemma closed_weaken e X Y: - closed X e → X ⊆ Y → closed Y e -with closed_weaken_val (v:val) X Y: - closed X v → X ⊆ Y → closed Y v. +Fixpoint ectx_map_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (E : ectx A) : + equal (arrow_comp f g) h → fmap f (fmap g E) = fmap h E. +Proof. auto_map_comp. Qed. + +#[global] +Instance Functor_ectx : Functor ectx. Proof. - { revert X Y. - induction e; intros. - - apply (closed_weaken_val _ _ _ H H0). - - unfold closed, is_closed in *. - apply bool_decide_unpack in H. - apply bool_decide_pack. - set_solver. - - unfold closed in *. simpl in *. - apply andb_prop_intro. - apply andb_prop_elim in H. - destruct H. - split. - apply (IHe1 _ _ H H0). - apply (IHe2 _ _ H1 H0). } - { revert X Y. - induction v; intros. - (*- constructor.*) - - unfold closed in *. - simpl in *. - apply (closed_weaken e _ _ H). - set_solver. - (*- constructor.*) } + constructor. + - exact @ectx_map_id. + - exact @ectx_map_comp. Qed. -Lemma closed_subst : - ∀ Γ x e1 e2, - closed (Γ ∪ {[x]}) e1 → - closed ∅ e2 → - closed Γ (subst x e2 e1) -with closed_subst_val : - ∀ Γ x (v : val) e, - closed (Γ ∪ {[x]}) v → - closed ∅ e → - closed Γ (subst_val x e v). +(** Bind-Map_Pure instances *) + +Fixpoint ebind_map_pure {A B} (f : A [→] B) g (e : expr A) : + equal (subst_of_arr f) g → fmap f e = bind g e +with vbind_map_pure {A B} (f : A [→] B) g (v : val A) : + equal (subst_of_arr f) g → fmap f v = bind g v. Proof. - { - unfold closed in *. - intros Γ x e1 e2 H_closed1 H_closed2. - induction e1. - - simpl in *. auto. - - simpl in *. - rewrite -> decide_bool_decide. - destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. - + eapply closed_weaken. - exact H_closed2. - set_solver. - + simpl in *. - apply bool_decide_unpack in H_closed1. - apply bool_decide_pack. - set_solver. - - simpl in *. - apply andb_prop_elim in H_closed1 as [Hc1 Hc2]. - apply andb_prop_intro. split. - + exact (IHe1_1 Hc1). - + exact (IHe1_2 Hc2). - } - { - unfold closed in *. - intros Γ x v e H_closed1 H_closed2. - induction v. - (* - simpl in *. auto. *) - - simpl in *. - rewrite -> decide_bool_decide. - destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. - + rewrite -> H_eq in H_closed1. - eapply closed_weaken. - exact H_closed1. set_solver. - + simpl in *. - apply closed_subst. - eapply closed_weaken. - exact H_closed1. set_solver. - exact H_closed2. - (*- simpl in *. auto.*) - } + - auto_map_bind_pure. + - auto_map_bind_pure. Qed. -Lemma closed_app xs e1 e2: - closed xs (app e1 e2) ↔ - closed xs e1 ∧ closed xs e2. -Proof. unfold closed. simpl. by rewrite andb_True. Qed. +#[global] +Instance BindMapPure_expr : BindMapPure expr. +Proof. constructor. exact @ebind_map_pure. Qed. -Lemma closed_lambda e Γ x : closed Γ (vlambda x e) ↔ closed (Γ ∪ {[x]}) e. -Proof. split. auto. auto. Qed. +#[global] +Instance BindMapPure_val : BindMapPure val. +Proof. constructor. exact @vbind_map_pure. Qed. -Lemma closed_var Γ x : closed Γ (var x) ↔ x ∈ Γ. -Proof. unfold closed. simpl. by rewrite bool_decide_spec. Qed. +Fixpoint ectx_bind_map_pure {A B} (f : A [→] B) g (E : ectx A) : + equal (subst_of_arr f) g → fmap f E = bind g E. +Proof. auto_map_bind_pure. Qed. -Lemma base_step_preserve_closedness : - ∀ e1 e1', - base_step e1 e1' → - closed ∅ e1 → - closed ∅ e1'. +#[global] +Instance BindMapPure_ectx : BindMapPure ectx. +Proof. constructor. exact @ectx_bind_map_pure. Qed. + +(** Bind-Map_Comm instances *) + +Fixpoint ebind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (e : expr A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 e) = fmap f1 (bind g1 e) +with vbind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (v : val A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 v) = fmap f1 (bind g1 v). Proof. - unfold closed. - intros e1 e1' Hred H_closed. - inversion Hred. subst. simpl in *. - apply andb_prop_elim in H_closed as [Hc1 Hc2]. - exact (closed_subst ∅ x e0 e2 Hc1 Hc2). + - auto_map_bind_comm. + - auto_map_bind_comm. Qed. -Definition ectx_is_closed (X : scope) (E : ectx) := - ∀ e, closed ∅ e → closed X (fill E e). +#[global] +Instance BindMapComm_expr : BindMapComm expr. +Proof. constructor. exact @ebind_map_comm. Qed. -Lemma closed_decompose : - ∀ E e, - closed ∅ (fill E e) → - ectx_is_closed ∅ E ∧ closed ∅ e. +#[global] +Instance BindMapComm_val : BindMapComm val. +Proof. constructor. exact @vbind_map_comm. Qed. + +Fixpoint ectx_bind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (E : ectx A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 E) = fmap f1 (bind g1 E). +Proof. auto_map_bind_comm. Qed. + +#[global] +Instance BindMapComm_ectx : BindMapComm ectx. +Proof. constructor. exact @ectx_bind_map_comm. Qed. + +(** Bind instances *) + +Fixpoint ebind_id {A} (f : A [⇒] A) (e : expr A) : + equal f (arrow_id A) → bind f e = e +with vbind_id {A} (f : A [⇒] A) (v : val A) : + equal f (arrow_id A) → bind f v = v. Proof. - unfold ectx_is_closed. - intros E. - induction E; intros e' H_closed. - - simpl in *. auto. - - simpl in *. - apply IHE in H_closed as [H_fill H_closed]. - apply closed_app in H_closed as [H_closed1 H_closed2]. - split. - + intros e'' H_closed3. - apply H_fill. - apply closed_app. auto. - + auto. - - simpl in *. - apply IHE in H_closed as [H_fill H_closed]. - apply closed_app in H_closed as [H_closed1 H_closed2]. - split. - + intros e'' H_closed3. - apply H_fill. - apply closed_app. auto. - + auto. + - auto_bind_id. + - auto_bind_id. Qed. -Lemma closed_compose : - ∀ E e, - ectx_is_closed ∅ E → - closed ∅ e → - closed ∅ (fill E e). -Proof. unfold ectx_is_closed. auto. Qed. - -Lemma contextual_step_preserve_closedness : - ∀ e1 e1', - contextual_step e1 e1' → - closed ∅ e1 → - closed ∅ e1'. +Fixpoint ebind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (e : expr A) : + equal (arrow_comp f g) h → bind f (bind g e) = bind h e +with vbind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (v : val A) : + equal (arrow_comp f g) h → bind f (bind g v) = bind h v. Proof. - unfold closed. - intros e1 e1' Hred H_closed. - inversion Hred. subst. simpl in *. clear Hred. - apply closed_decompose in H_closed as [H_closed1 H_closed2]. - apply (base_step_preserve_closedness _ _ H1) in H_closed2. - apply closed_compose; auto. + - auto_bind_comp. + - auto_bind_comp. Qed. -(** subscoped from the Erlang paper *) -Definition subst_is_closed (Γ free : scope) (sub : sub) := - Γ = dom sub ∧ - ∀ x, x ∈ Γ → - ∀ v, sub !! x = Some v → closed free (ret v). - -Lemma subst_is_closed_subseteq: ∀ (Γ1 Γ2 X : scope) (γ1 γ2: sub), - Γ1 = dom γ1 → - γ1 ⊆ γ2 → Γ1 ⊆ Γ2 → subst_is_closed Γ2 X γ2 → subst_is_closed Γ1 X γ1. +#[global] +Instance Bind_expr : Bind expr. Proof. - intros * Hd Hγ HΓ Hclosed2. - destruct Hclosed2 as [Hd2 Hc2]. - split. - assumption. - intros x Hl v Hs. - rewrite (map_subseteq_spec γ1 γ2) in Hγ. - (* specialize (Hγ _ _ Hs). *) - specialise Hγ Hs. - apply (Hc2 x ltac:(set_solver) _ Hγ). + constructor. + - exact @ebind_id. + - exact @ebind_comp. Qed. -Lemma closed_ectx_app1 : - ∀ E e, - ectx_is_closed ∅ E → - closed ∅ e → - ectx_is_closed ∅ (ectx_app1 E e). +#[global] +Instance Bind_val : Bind val. Proof. - intros E e Hc1 Hc2. - unfold ectx_is_closed in *. - intros e' Hc3. simpl. - apply Hc1. apply closed_app. auto. + constructor. + - exact @vbind_id. + - exact @vbind_comp. Qed. -Lemma closed_ectx_app2 : - ∀ (v : val) E, - closed ∅ v → - ectx_is_closed ∅ E → - ectx_is_closed ∅ (ectx_app2 v E). +Fixpoint ectx_bind_id {A} (f : A [⇒] A) (E : ectx A) : + equal f (arrow_id A) → bind f E = E. +Proof. auto_bind_id. Qed. + +Fixpoint ectx_bind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (E : ectx A) : + equal (arrow_comp f g) h → bind f (bind g E) = bind h E. +Proof. auto_bind_comp. Qed. + +#[global] +Instance Bind_ectx : Bind ectx. Proof. - intros v E Hc1 Hc2. - unfold ectx_is_closed in *. - intros e' Hc3. simpl. - apply Hc2. apply closed_app. auto. + constructor. + - exact @ectx_bind_id. + - exact @ectx_bind_comp. Qed. -Lemma subst_is_closed_elim_closed Γ (γ:sub) x X (v:val): - subst_is_closed Γ X γ → - γ !! x = Some v → - closed X v. +(** Reduction *) + +Inductive base_step {V} : expr V → expr V → Prop := +| Beta_step (e : expr (inc V)) (v : val V) : + base_step (app (v_lambda e) v) (subst (Inc:=inc) e v). + +Inductive contextual_step {V} : expr V → expr V → Prop := +| Ectx_step E e1 e2 : + base_step e1 e2 → + contextual_step (plug E e1) (plug E e2). + +Definition big_step {V} e (v : val V) := + rtc contextual_step e v. + +Definition terminates {V} (e : expr V) := + ∃ v, big_step e v. + +Lemma not_base_step_val {V} (v : val V) e : ¬ base_step v e. +Proof. inversion_clear 1. Qed. + +Lemma not_contextual_step_val {V} (v : val V) e : ¬ contextual_step v e. Proof. - intros [Hdom Hsc] He. - assert (H := elem_of_dom_2 _ _ _ He). - (* have: elem_of_dom_2 He. *) - (* pose proof (elem_of_dom_2 _ _ _ He). *) - assert (x ∈ Γ). set_solver. - apply (Hsc x H0 v He). + intros Hstep. + inversion Hstep as [E' e1' e2' H_step' Hv He]. + apply plug_eq_val_inv in Hv as [_ ->]. + by eapply not_base_step_val. Qed. -(* if e is closed under Y, we can split the variables in Y between X and γ *) -Lemma subst_map_closed' (e : expr) (X Y : scope) (γ:sub) : - closed Y e → - (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → - closed X (subst_map γ e) -with subst_map_closed'_val (v : val) (X Y : scope) (γ:sub) : - closed Y (ret v) → - (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → - closed X (subst_map_val γ v). +Lemma big_step_val {V} (v : val V) : big_step v v. +Proof. unfold big_step. done. Qed. + +Lemma terminates_val {V} (v : val V) : terminates v. +Proof. unfold terminates. exists v. apply big_step_val. Qed. + +Lemma contextual_step_comp {V} (E : ectx V) e1 e2 : + contextual_step e1 e2 → + contextual_step (plug E e1) (plug E e2). Proof. - { - revert X Y γ. induction e. - { eapply subst_map_closed'_val; eauto. } - { intros * Hc H. - (* e is a variable x *) - unfold closed in Hc; simpl in Hc; apply bool_decide_unpack in Hc. - specialize (H x Hc). simpl. - destruct (γ !! x) eqn:He. - - assumption. - - unfold closed; simpl; apply bool_decide_pack. assumption. - } - { intros *. - unfold closed. simpl. - rewrite !andb_True. intros [? ?] **. - split. - by eapply IHe1. - by eapply IHe2. } - } - { revert X Y γ. induction v. - (*{ intros. assumption. }*) - { unfold closed. simpl. - intros * Hce H. - eapply subst_map_closed'. eassumption. - intros y [|]%elem_of_union. - { destruct (decide (x = y)). - { by subst; rewrite lookup_delete_eq with (m:=γ); set_solver. } - rewrite lookup_delete_ne with (m:=γ). 2: { assumption. } - eapply H in H0. - destruct lookup; last set_solver. - eapply closed_weaken; eauto with set_solver. } - { rewrite elem_of_singleton in H0. - subst. rewrite lookup_delete_eq with (m:=γ). set_solver. } - } - (*{ intros. assumption. }*) } + intros H_step. + inversion_clear H_step as [E' e1' e2' H_step']. + rewrite <- ectx_comp_correct. + rewrite <- ectx_comp_correct. + constructor. exact H_step'. Qed. -(* simple corollary of [subst_map_closed'], - where we have split Y into X and dom γ upfront *) -Lemma subst_map_closed'_2 Γ X γ e : - closed (X ∪ (dom γ)) e -> - subst_is_closed Γ X γ -> - closed X (subst_map γ e). +Lemma contextual_step_terminates {V} (e e' : expr V) : + contextual_step e e' → + terminates e' → + terminates e. Proof. - intros Hcl Hsubst. - eapply subst_map_closed'; first eassumption. - intros x Hx. - destruct (γ !! x) as [e'|] eqn:Heq. - - apply (subst_is_closed_elim_closed _ _ _ _ _ Hsubst Heq). - - apply not_elem_of_dom in Heq. - set_solver. + unfold terminates, big_step. + intros H_step [v H_steps]. + exists v. econstructor; eauto. Qed. -(* given a value and a substitution closed under the same scope, - applying the substitution gives us a completely closed value *) -Lemma subst_map_closed'_3 (v:val) Γ γ: - closed Γ v -> - subst_is_closed Γ ∅ γ -> - closed ∅ (subst_map γ v). +Lemma base_step_is_deterministic {V} (e1 e2 e3 : expr V) : + base_step e1 e2 → + base_step e1 e3 → + e2 = e3. Proof. - pose proof (subst_map_closed'_2 Γ ∅ γ v). - simpl in H. - intros. - apply H. 2: { assumption. } - destruct H1 as [? _]. - rewrite <- H1. - replace (∅ ∪ Γ) with Γ. assumption. - set_solver. + intros Hstep2 Hstep3. + inversion Hstep2. + inversion Hstep3. + congruence. Qed. -Lemma subst_map_val_closed_val_aux Γ γ (v : val) : - closed Γ v → - Γ ∩ dom γ = ∅ → - subst_map_val γ v = v -with subst_map_closed_aux Γ γ (e : expr) : - closed Γ e → - Γ ∩ dom γ = ∅ → - subst_map γ e = e. +Inductive potential_redex {V} : expr V -> Prop := +| pr_app (v1 v2 : val V) : potential_redex (app v1 v2). + +Lemma not_potential_redex_val {V} (v : val V) : ¬ potential_redex v. +Proof. inversion_clear 1. Qed. + +Lemma potential_redex_app_inv {V} e1 e2 : + potential_redex (app e1 e2) → + ∃ (v1 v2 : val V), e1 = v1 ∧ e2 = v2. +Proof. inversion_clear 1. eauto. Qed. + +Lemma unique_rdecomposition {V} (R1 R2 : rctx V) e1 e2 : + potential_redex e1 → + potential_redex e2 → + rplug R1 e1 = rplug R2 e2 → + R1 = R2 ∧ e1 = e2. Proof. - { intros Hc Hdom. - induction v. - - rewrite -> closed_lambda in Hc. - rewrite -> fold_unfold_subst_map_val_vlambda. - rewrite -> (subst_map_closed_aux (Γ ∪ {[x]}) (delete x γ) e Hc ltac:(set_solver)). - reflexivity. } - { intros Hc Hdom. - induction e. - - rewrite -> fold_unfold_subst_map_ret. - f_equal. by eapply subst_map_val_closed_val_aux. - - unfold closed in Hc. - simpl in Hc. - rewrite -> bool_decide_spec in Hc. - rewrite -> fold_unfold_subst_map_var. - assert (H_not_in : x ∉ dom γ) by set_solver. - rewrite -> (not_elem_of_dom γ x) in H_not_in. - setoid_rewrite -> H_not_in. - reflexivity. - - apply closed_app in Hc as [Hc1 Hc2]. - rewrite -> fold_unfold_subst_map_app. - rewrite -> (IHe1 Hc1). - rewrite -> (IHe2 Hc2). - reflexivity. } + intros He1 He2. + revert R2. + induction R1; intros R2 Heq. + - destruct R2; simpl in *. + + auto. + + rewrite -> Heq in He1. + apply potential_redex_app_inv in He1 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv1 as [_ ->]. + contradict (not_potential_redex_val _ He2). + + rewrite -> Heq in He1. + apply potential_redex_app_inv in He1 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv2 as [_ ->]. + contradict (not_potential_redex_val _ He2). + - destruct R2; simpl in *. + + rewrite <- Heq in He2. + apply potential_redex_app_inv in He2 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv1 as [_ ->]. + contradict (not_potential_redex_val _ He1). + + injection Heq as Heq1 Heq2. + apply IHR1 in Heq1 as [Heq11 Heq12]. + split; congruence. + + injection Heq as Heq1 Heq2. + apply rplug_eq_val_inv in Heq1 as [_ ->]. + contradict (not_potential_redex_val _ He1). + - destruct R2; simpl in *. + + rewrite <- Heq in He2. + apply potential_redex_app_inv in He2 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv2 as [_ ->]. + contradict (not_potential_redex_val _ He1). + + injection Heq as Heq1 Heq2. symmetry in Heq1. + apply rplug_eq_val_inv in Heq1 as [_ ->]. + contradict (not_potential_redex_val _ He2). + + injection Heq as Heq1 Heq2. + apply IHR1 in Heq2 as [Heq11 Heq12]. + split; congruence. +Qed. + +Lemma unique_decomposition {V} (E1 E2 : ectx V) e1 e2 : + potential_redex e1 → + potential_redex e2 → + plug E1 e1 = plug E2 e2 → + E1 = E2 ∧ e1 = e2. +Proof. + intros He1 He2 Heq. + rewrite -> plug_rplug_equiv in Heq. + rewrite -> plug_rplug_equiv in Heq. + destruct (unique_rdecomposition _ _ _ _ He1 He2 Heq) as [Heq1 Heq2]. + split. + - rewrite <- (ectx_rctx_bijection E1). + rewrite <- (ectx_rctx_bijection E2). + f_equal. exact Heq1. + - exact Heq2. Qed. -Lemma subst_map_val_closed_val γ (v : val) : - closed ∅ v → - subst_map_val γ v = v. +Lemma base_step_potential_redex {V} (e e' : expr V) : + base_step e e' → + potential_redex e. Proof. - intros Hc. - eapply subst_map_val_closed_val_aux. - - exact Hc. - - set_solver. + inversion_clear 1. + constructor. Qed. -Lemma subst_map_closed γ (e : expr) : - closed ∅ e → - subst_map γ e = e. +Lemma contextual_step_is_deterministic {V} (e1 e2 e3 : expr V) : + contextual_step e1 e2 → + contextual_step e1 e3 → + e2 = e3. Proof. - intros Hc. - eapply subst_map_closed_aux. - - exact Hc. - - set_solver. + intros Hstep2 Hstep3. + inversion Hstep2 as [E2 e12 e2' Hstep2' He12 He2']. + inversion Hstep3 as [E3 e13 e3' Hstep3' He13 He3']. + assert (Hpr2 := base_step_potential_redex _ _ Hstep2'). + assert (Hpr3 := base_step_potential_redex _ _ Hstep3'). + destruct (unique_decomposition E2 E3 e12 e13 Hpr2 Hpr3) as [HE_eq He_eq]. + { congruence. } + rewrite -> He_eq in Hstep2'. + assert (He_eq' := base_step_is_deterministic e13 e2' e3' Hstep2' Hstep3'). + congruence. Qed. -(** Relations *) +(** Relations for closed term *) -Definition expr_rel := expr ⇒ᵢ expr ⇒ᵢ IRel. -Definition val_rel := val ⇒ᵢ val ⇒ᵢ IRel. -Definition sub_rel := sub ⇒ᵢ sub ⇒ᵢ IRel. -Definition ctx_rel := ectx ⇒ᵢ ectx ⇒ᵢ IRel. +Definition expr_rel := expr ∅ ⇒ᵢ expr ∅ ⇒ᵢ IRel. +Definition val_rel := val ∅ ⇒ᵢ val ∅ ⇒ᵢ IRel. +Definition ectx_rel := ectx ∅ ⇒ᵢ ectx ∅ ⇒ᵢ IRel. Definition L_rel_pre (L_rel : expr_rel) : expr_rel := λ e1 e2, - (∀ v1 : val, e1 = v1 → terminates e2)ᵢ ∧ᵢ - ∀ᵢ e1' : expr, (contextual_step e1 e1')ᵢ →ᵢ ▷ L_rel e1' e2. + (∀ v1 : val ∅, e1 = v1 → terminates e2)ᵢ ∧ᵢ + (∀ᵢ e1' : expr ∅, (contextual_step e1 e1')ᵢ →ᵢ ▷ L_rel e1' e2). Definition L_rel_fix := I_fix L_rel_pre. Definition L_rel := L_rel_pre L_rel_fix. -Definition O_rel e1 e2 := L_rel e1 e2 ∧ᵢ L_rel e2 e1. -Definition K_rel_pre (V_rel : val_rel) := +Definition O_rel : expr_rel := + λ e1 e2, L_rel e1 e2 ∧ᵢ L_rel e2 e1. + +Definition K_rel_pre (V_rel : val_rel) : ectx_rel := λ E1 E2, - (∀ᵢ (v1 v2 : val), + ∀ᵢ (v1 v2 : val ∅), V_rel v1 v2 →ᵢ - O_rel (fill E1 v1) (fill E2 v2)). + O_rel (plug E1 v1) (plug E2 v2). -Definition R_rel_pre (V_rel : val_rel) (e1 e2 : expr) := - ∀ᵢ E1 E2, - ▷ K_rel_pre V_rel E1 E2 →ᵢ O_rel (fill E1 e1) (fill E2 e2). +Definition R_rel_pre (V_rel : val_rel) : expr_rel := + λ e1 e2, + ∀ᵢ E1 E2, + ▷ K_rel_pre V_rel E1 E2 →ᵢ + O_rel (plug E1 e1) (plug E2 e2). Definition V_rel_pre (V_rel : val_rel) : val_rel := λ v1 v2, - (closed ∅ v1)ᵢ ∧ᵢ - (closed ∅ v2)ᵢ ∧ᵢ - ∀ᵢ (u1 u2 : val), + ∀ᵢ u1 u2, ▷ V_rel u1 u2 →ᵢ R_rel_pre V_rel (app v1 u1) (app v2 u2). Definition V_rel_fix := I_fix V_rel_pre. - Definition V_rel := V_rel_pre V_rel_fix. Definition R_rel := R_rel_pre V_rel_fix. Definition K_rel := K_rel_pre V_rel_fix. -Definition E_rel : expr_rel := - λ e1 e2, - ∀ᵢ E1 E2 : ectx, K_rel E1 E2 →ᵢ O_rel (fill E1 e1) (fill E2 e2). +Definition E_rel (e1 e2 : expr ∅) := + ∀ᵢ E1 E2, + K_rel E1 E2 →ᵢ + O_rel (plug E1 e1) (plug E2 e2). (** Relations for open terms *) -Definition G_rel (Γ: scope) (γ1 γ2 : sub) : IProp := - (subst_is_closed Γ ∅ γ1)ᵢ ∧ᵢ - (subst_is_closed Γ ∅ γ2)ᵢ ∧ᵢ - ∀ᵢ x v1 v2, - (γ1 !! x = Some v1)ᵢ →ᵢ - (γ2 !! x = Some v2)ᵢ →ᵢ - V_rel v1 v2. +Definition G_rel {V} (γ1 γ2 : V [⇒] ∅) : IProp := + ∀ᵢ x, V_rel (γ1 x) (γ2 x). -Definition E_rel_o (Γ: scope) (e1 e2 : expr) : IProp := - ∀ᵢ γ1 γ2, G_rel Γ γ1 γ2 →ᵢ E_rel (subst_map γ1 e1) (subst_map γ2 e2). +Definition E_rel_o {V} (e1 e2 : expr V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ E_rel (bind γ1 e1) (bind γ2 e2). -Definition V_rel_o (Γ: scope) (v1 v2 : val) : IProp := - ∀ᵢ γ1 γ2, G_rel Γ γ1 γ2 →ᵢ V_rel (subst_map_val γ1 v1) (subst_map_val γ2 v2). +Definition V_rel_o {V} (v1 v2 : val V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ V_rel (bind γ1 v1) (bind γ2 v2). -Definition P_rel_o (Γ: scope) (e1 e2 : expr) : IProp := - ∀ᵢ γ1 γ2, G_rel Γ γ1 γ2 →ᵢ O_rel (subst_map γ1 e1) (subst_map γ2 e2). +Definition O_rel_o {V} (e1 e2 : expr V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ O_rel (bind γ1 e1) (bind γ2 e2). (** Contractiveness and unrolling fixpoint *) Lemma L_rel_pre_contractive : contractive L_rel_pre. -Proof. - intro n; iintros; unfold L_rel_pre; auto_contr. -Qed. +Proof. intro n; iintros; unfold L_rel_pre; auto_contr. Qed. Lemma L_rel_roll p1 p2 n : - n ⊨ L_rel p1 p2 → n ⊨ L_rel_fix p1 p2. + n ⊨ L_rel p1 p2 → + n ⊨ L_rel_fix p1 p2. Proof. - intro H; iapply (I_fix_roll expr_rel); [ | exact H ]. + intro H; iapply (I_fix_roll expr_rel); [| exact H]. apply L_rel_pre_contractive. Qed. Lemma L_rel_unroll p1 p2 n : - n ⊨ L_rel_fix p1 p2 → n ⊨ L_rel p1 p2. + n ⊨ L_rel_fix p1 p2 → + n ⊨ L_rel p1 p2. Proof. - intro H; iapply (I_fix_unroll expr_rel); [ | exact H ]. + intro H; iapply (I_fix_unroll expr_rel); [| exact H]. apply L_rel_pre_contractive. Qed. Lemma V_rel_pre_contractive : contractive V_rel_pre. -Proof. - intro n; iintros; unfold V_rel_pre, R_rel_pre, K_rel_pre; auto_contr. -Qed. +Proof. intro n; iintros; unfold V_rel_pre, R_rel_pre, K_rel_pre; auto_contr. Qed. Lemma V_rel_roll v1 v2 n : - n ⊨ V_rel v1 v2 → n ⊨ V_rel_fix v1 v2. + n ⊨ V_rel v1 v2 → + n ⊨ V_rel_fix v1 v2. Proof. - intro H; iapply (I_fix_roll val_rel); [ | exact H ]. + intro H; iapply (I_fix_roll val_rel); [| exact H]. apply V_rel_pre_contractive. Qed. Lemma V_rel_unroll v1 v2 n : - n ⊨ V_rel_fix v1 v2 → n ⊨ V_rel v1 v2. + n ⊨ V_rel_fix v1 v2 → + n ⊨ V_rel v1 v2. Proof. - intro H; iapply (I_fix_unroll val_rel); [ | exact H ]. + intro H; iapply (I_fix_unroll val_rel); [| exact H]. apply V_rel_pre_contractive. Qed. (** Introduction and elimination lemmas *) -Lemma V_rel_intro (v1 v2 : val) n : - closed ∅ v1 → - closed ∅ v2 → - (n ⊨ ∀ᵢ (u1 u2:val), - ▷ V_rel u1 u2 →ᵢ - R_rel (app v1 u1) (app v2 u2)) → - n ⊨ V_rel v1 v2. +Lemma L_rel_intro (e1 e2 : expr ∅) n : + (∀ v1 : val ∅, e1 = v1 → terminates e2) → + n ⊨ (∀ᵢ e1' : expr ∅, (contextual_step e1 e1')ᵢ →ᵢ ▷ L_rel e1' e2) → + n ⊨ L_rel e1 e2. Proof. - intros H_closed1 H_closed2 Hv. - isplit; [| isplit]. - - apply I_prop_intro. assumption. - - apply I_prop_intro. assumption. - - iintros u1 u2 Hv_later. - ispecialize Hv u1. - ispecialize Hv u2. - iapply Hv. - later_shift. - apply V_rel_unroll. - assumption. + intros H_val H_expr. + unfold L_rel, L_rel_pre. + isplit. + - iintro. exact H_val. + - iintros e1' H_step. + ispec H_expr e1' H_step. + later_shift. apply L_rel_roll. exact H_expr. Qed. -Lemma V_rel_elim (v1 v2 : val) n : - n ⊨ V_rel v1 v2 → - closed ∅ v1 ∧ - closed ∅ v2 ∧ - (n ⊨ (∀ᵢ (u1 u2 : val), - ▷ V_rel u1 u2 →ᵢ - R_rel (app v1 u1) (app v2 u2))). +Lemma L_rel_elim (e1 e2 : expr ∅) n : + n ⊨ L_rel e1 e2 → + (∀ v1 : val ∅, e1 = v1 → terminates e2) ∧ + (n ⊨ ∀ᵢ e1' : expr ∅, (contextual_step e1 e1')ᵢ →ᵢ ▷ L_rel e1' e2). Proof. - intros Hv. - unfold V_rel in Hv. - unfold V_rel_pre in Hv. - idestruct Hv as H_closed1 Hv. idestruct H_closed1. - idestruct Hv as H_closed2 Hv. idestruct H_closed2. - split; [| split]. - - assumption. - - assumption. - - iintros u1 u2 Hv_later. - ispecialize Hv u1. ispecialize Hv u2. - iapply Hv. - later_shift. - apply V_rel_roll. - assumption. + intros He. + unfold L_rel, L_rel_pre in He. + idestruct He as He1 He2. + split. + - idestruct He1. exact He1. + - iintros e1' H_step. + ispec He2 e1' H_step. + later_shift. apply L_rel_unroll. exact He2. Qed. -Lemma K_rel_intro (E1 E2 : ectx) n : - n ⊨ (∀ᵢ v1 v2, V_rel v1 v2 →ᵢ O_rel (fill E1 v1) (fill E2 v2)) → - n ⊨ K_rel E1 E2. +Lemma O_rel_intro (e1 e2 : expr ∅) n : + n ⊨ L_rel e1 e2 → + n ⊨ L_rel e2 e1 → + n ⊨ O_rel e1 e2. Proof. - intros HE. - unfold K_rel, K_rel_pre. + intros He1 He2. + unfold O_rel. isplit; assumption. +Qed. + +Lemma O_rel_elim (e1 e2 : expr ∅) n : + n ⊨ O_rel e1 e2 → + (n ⊨ L_rel e1 e2) ∧ + (n ⊨ L_rel e2 e1). +Proof. + unfold O_rel. + intros He. idestruct He as He1 He2. + split; assumption. +Qed. + +Lemma V_rel_intro (v1 v2 : val ∅) n : + (n ⊨ ∀ᵢ u1 u2, + ▷ V_rel u1 u2 →ᵢ + R_rel (app v1 u1) (app v2 u2)) → + n ⊨ V_rel v1 v2. +Proof. + intros Hv. + unfold V_rel, V_rel_pre. + iintros u1 u2 Hu. + ispecialize Hv u1. + ispecialize Hv u2. + iapply Hv. later_shift. + apply V_rel_unroll. exact Hu. +Qed. + +Lemma V_rel_elim (v1 v2 : val ∅) n : + n ⊨ V_rel v1 v2 → + n ⊨ ∀ᵢ u1 u2, + ▷ V_rel u1 u2 →ᵢ + R_rel (app v1 u1) (app v2 u2). +Proof. + intros Hv. + unfold V_rel, V_rel_pre in Hv. + iintros u1 u2 Hu. + ispecialize Hv u1. + ispecialize Hv u2. + iapply Hv. later_shift. + apply V_rel_roll. exact Hu. +Qed. + +Lemma V_rel_elimR (v1 v2 u1 u2 : val ∅) n : + n ⊨ V_rel v1 v2 → + n ⊨ ▷ V_rel u1 u2 → + n ⊨ R_rel (app v1 u1) (app v2 u2). +Proof. + intros Hv Hu. + apply V_rel_elim in Hv. + iapply Hv. exact Hu. +Qed. + +Lemma K_rel_intro (E1 E2 : ectx ∅) n : + n ⊨ (∀ᵢ v1 v2, V_rel v1 v2 →ᵢ O_rel (plug E1 v1) (plug E2 v2)) → + n ⊨ K_rel E1 E2. +Proof. + intros HE. + unfold K_rel, K_rel_pre. iintros v1 v2 Hv. iapply HE. apply V_rel_unroll. exact Hv. Qed. -Lemma K_rel_elim (E1 E2 : ectx) n : +Lemma K_rel_elim (E1 E2 : ectx ∅) n : n ⊨ K_rel E1 E2 → - (n ⊨ ∀ᵢ v1 v2, V_rel v1 v2 →ᵢ O_rel (fill E1 v1) (fill E2 v2)). + n ⊨ ∀ᵢ v1 v2, V_rel v1 v2 →ᵢ O_rel (plug E1 v1) (plug E2 v2). Proof. unfold K_rel, K_rel_pre. intros HE. @@ -907,201 +933,183 @@ Qed. Lemma K_rel_elimO E1 E2 v1 v2 n : n ⊨ K_rel E1 E2 → n ⊨ V_rel v1 v2 → - n ⊨ O_rel (fill E1 v1) (fill E2 v2). + n ⊨ O_rel (plug E1 v1) (plug E2 v2). Proof. intros HE Hv. apply K_rel_elim in HE. iapply HE. exact Hv. Qed. -Lemma R_rel_intro (e1 e2 : expr) n : - n ⊨ (∀ᵢ E1 E2, ▷ K_rel E1 E2 →ᵢ O_rel (fill E1 e1) (fill E2 e2)) -> +Lemma R_rel_intro (e1 e2 : expr ∅) n : + n ⊨ (∀ᵢ E1 E2, ▷ K_rel E1 E2 →ᵢ O_rel (plug E1 e1) (plug E2 e2)) -> n ⊨ R_rel e1 e2. Proof. auto. Qed. -Lemma R_rel_elim (e1 e2 : expr) n : +Lemma R_rel_elim (e1 e2 : expr ∅) n : n ⊨ R_rel e1 e2 → - n ⊨ (∀ᵢ E1 E2, ▷ K_rel E1 E2 →ᵢ O_rel (fill E1 e1) (fill E2 e2)). + n ⊨ (∀ᵢ E1 E2, ▷ K_rel E1 E2 →ᵢ O_rel (plug E1 e1) (plug E2 e2)). Proof. auto. Qed. -Lemma R_rel_elimO (e1 e2 : expr) E1 E2 n : +Lemma R_rel_elimO (e1 e2 : expr ∅) E1 E2 n : n ⊨ R_rel e1 e2 → n ⊨ ▷ K_rel E1 E2 → - n ⊨ O_rel (fill E1 e1) (fill E2 e2). + n ⊨ O_rel (plug E1 e1) (plug E2 e2). Proof. intros He HE. apply R_rel_elim in He. iapply He. exact HE. Qed. -Lemma E_rel_intro (e1 e2 : expr) n : - (n ⊨ ∀ᵢ E1 E2, K_rel E1 E2 →ᵢ O_rel (fill E1 e1) (fill E2 e2)) -> +Lemma E_rel_intro (e1 e2 : expr ∅) n : + (n ⊨ ∀ᵢ E1 E2, K_rel E1 E2 →ᵢ O_rel (plug E1 e1) (plug E2 e2)) → n ⊨ E_rel e1 e2. -Proof. - intros HE. - unfold E_rel. - exact HE. -Qed. +Proof. auto. Qed. -Lemma E_rel_elim (e1 e2 : expr) n : +Lemma E_rel_elim (e1 e2 : expr ∅) n : n ⊨ E_rel e1 e2 → - (n ⊨ ∀ᵢ E1 E2, K_rel E1 E2 →ᵢ O_rel (fill E1 e1) (fill E2 e2)). -Proof. - intros He. - unfold E_rel in He. - assumption. -Qed. + n ⊨ ∀ᵢ E1 E2, K_rel E1 E2 →ᵢ O_rel (plug E1 e1) (plug E2 e2). +Proof. auto. Qed. (** Bind lemma *) Lemma E_rel_elimO e1 e2 E1 E2 n : n ⊨ E_rel e1 e2 → n ⊨ K_rel E1 E2 → - n ⊨ O_rel (fill E1 e1) (fill E2 e2). + n ⊨ O_rel (plug E1 e1) (plug E2 e2). Proof. intros He HE. apply E_rel_elim in He. iapply He. exact HE. Qed. -Lemma V_rel_elimE (v1 v2 u1 u2 : val) n : +Lemma V_rel_elimE (v1 v2 u1 u2 : val ∅) n : n ⊨ V_rel v1 v2 → n ⊨ V_rel u1 u2 → n ⊨ E_rel (app v1 u1) (app v2 u2). Proof. intros Hv Hu. - destruct (V_rel_elim _ _ _ Hv) as (Hv1_closed & Hv2_closed & Hv'). - destruct (V_rel_elim _ _ _ Hu) as (Hu1_closed & Hu2_closed & _). apply E_rel_intro. - iintros E1 E2 HE. simpl. + iintros E1 E2 HE. apply R_rel_elimO. - - iapply Hv'. - later_shift. exact Hu. + - apply V_rel_elimR. + + exact Hv. + + later_shift. exact Hu. - later_shift. exact HE. Qed. -Lemma G_rel_intro Γ γ1 γ2 n : - subst_is_closed Γ ∅ γ1 → - subst_is_closed Γ ∅ γ2 → - n ⊨ - (∀ᵢ x v1 v2, - (γ1 !! x = Some v1)ᵢ →ᵢ - (γ2 !! x = Some v2)ᵢ →ᵢ - V_rel v1 v2) → - n ⊨ G_rel Γ γ1 γ2. -Proof. - intros H_closed1 H_closed2 Hγ. - unfold G_rel. - isplit; [| isplit]. - - iintro. exact H_closed1. - - iintro. exact H_closed2. - - exact Hγ. -Qed. +Lemma G_rel_intro {V} (γ1 γ2 : V [⇒] ∅) n : + (n ⊨ ∀ᵢ x, V_rel (γ1 x) (γ2 x)) → + n ⊨ G_rel γ1 γ2. +Proof. auto. Qed. -Lemma G_rel_elim Γ γ1 γ2 n : - n ⊨ G_rel Γ γ1 γ2 → - subst_is_closed Γ ∅ γ1 ∧ - subst_is_closed Γ ∅ γ2 ∧ - (n ⊨ - ∀ᵢ x v1 v2, - (γ1 !! x = Some v1)ᵢ →ᵢ - (γ2 !! x = Some v2)ᵢ →ᵢ - V_rel v1 v2). -Proof. - unfold G_rel. - intros Hγ. - idestruct Hγ as H_closed1 Hγ. idestruct H_closed1. - idestruct Hγ as H_closed2 Hγ. idestruct H_closed2. - auto. -Qed. +Lemma G_rel_elim {V} (γ1 γ2 : V [⇒] ∅) n : + n ⊨ G_rel γ1 γ2 → + n ⊨ ∀ᵢ x, V_rel (γ1 x) (γ2 x). +Proof. auto. Qed. -Lemma E_rel_o_intro Γ e1 e2 n : - (n ⊨ ∀ᵢ γ1 γ2, - G_rel Γ γ1 γ2 →ᵢ - E_rel (subst_map γ1 e1) (subst_map γ2 e2)) → - n ⊨ E_rel_o Γ e1 e2. +Lemma G_rel_elimV {V} (γ1 γ2 : V [⇒] ∅) (x : V) n : + n ⊨ G_rel γ1 γ2 → + n ⊨ V_rel (γ1 x) (γ2 x). Proof. - intros He. - unfold E_rel_o. - exact He. + intros Hγ. + apply G_rel_elim in Hγ. + iapply Hγ. Qed. -Lemma E_rel_o_elim Γ e1 e2 n : - n ⊨ E_rel_o Γ e1 e2 → - (n ⊨ ∀ᵢ γ1 γ2, G_rel Γ γ1 γ2 →ᵢ E_rel (subst_map γ1 e1) (subst_map γ2 e2)). +Lemma E_rel_o_intro {V} (e1 e2 : expr V) n : + (n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ E_rel (bind γ1 e1) (bind γ2 e2)) → + n ⊨ E_rel_o e1 e2. +Proof. auto. Qed. + +Lemma E_rel_o_elim {V} (e1 e2 : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ E_rel (bind γ1 e1) (bind γ2 e2). +Proof. auto. Qed. + +Lemma E_rel_o_elimE {V} (e1 e2 : expr V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ E_rel (bind γ1 e1) (bind γ2 e2). Proof. - unfold E_rel_o. - intros He. - auto. + intros He Hγ. + apply E_rel_o_elim in He. + iapply He. exact Hγ. Qed. -Lemma V_rel_o_intro Γ (v1 v2 : val) n : - (n ⊨ ∀ᵢ γ1 γ2, - G_rel Γ γ1 γ2 →ᵢ - V_rel (subst_map_val γ1 v1) (subst_map_val γ2 v2)) → - n ⊨ V_rel_o Γ v1 v2. +Lemma V_rel_o_intro {V} (v1 v2 : val V) n : + (n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ V_rel (bind γ1 v1) (bind γ2 v2)) → + n ⊨ V_rel_o v1 v2. +Proof. auto. Qed. + +Lemma V_rel_o_elim {V} (v1 v2 : val V) n : + n ⊨ V_rel_o v1 v2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ V_rel (bind γ1 v1) (bind γ2 v2). +Proof. auto. Qed. + +Lemma V_rel_o_elimV {V} (v1 v2 : val V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ V_rel_o v1 v2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ V_rel (bind γ1 v1) (bind γ2 v2). Proof. - intros Hv. - unfold V_rel_o. - exact Hv. + intros Hv Hγ. + apply V_rel_o_elim in Hv. + iapply Hv. exact Hγ. Qed. -Lemma V_rel_o_elim Γ (v1 v2 : val) n : - n ⊨ V_rel_o Γ v1 v2 → - (n ⊨ ∀ᵢ γ1 γ2, - G_rel Γ γ1 γ2 →ᵢ - V_rel (subst_map_val γ1 v1) (subst_map_val γ2 v2)). +Lemma O_rel_o_intro {V} (e1 e2 : expr V) n : + (n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ O_rel (bind γ1 e1) (bind γ2 e2)) → + n ⊨ O_rel_o e1 e2. +Proof. auto. Qed. + +Lemma O_rel_o_elim {V} (e1 e2 : expr V) n : + n ⊨ O_rel_o e1 e2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ O_rel (bind γ1 e1) (bind γ2 e2). +Proof. auto. Qed. + +Lemma O_rel_o_elimO {V} (e1 e2 : expr V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ O_rel_o e1 e2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ O_rel (bind γ1 e1) (bind γ2 e2). Proof. - unfold V_rel_o. - intros Hv. exact Hv. + intros He Hγ. + apply O_rel_o_elim in He. + iapply He. exact Hγ. Qed. (** Compatibility lemmas *) (* aka val inclusion *) -Lemma compat_val (Γ : scope) (v1 v2 : val) n : - n ⊨ V_rel_o Γ v1 v2 → - n ⊨ E_rel_o Γ v1 v2. +Lemma compat_val {V} (v1 v2 : val V) n : + n ⊨ V_rel_o v1 v2 → + n ⊨ E_rel_o v1 v2. Proof. intros Hv. apply V_rel_o_elim in Hv. - apply E_rel_o_intro. - iintros γ1 γ2 Hγ. - ispecialize Hv γ1. ispecialize Hv γ2. ispec Hv Hγ. - apply E_rel_intro. iintros E1 E2 HE. simpl. - apply (K_rel_elimO E1 E2 _ _ _ HE). - apply V_rel_elim in Hv as (H_closed1 & H_closed2 & Hv). - apply V_rel_intro. - { exact H_closed1. } - { exact H_closed2. } - { exact Hv. } + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + apply E_rel_intro. iintros E1 E2 HE. + apply (K_rel_elimO E1 E2 (bind γ1 v1) (bind γ2 v2) _ HE). + iapply Hv. exact Hγ. Qed. -Lemma compat_app (Γ:scope) (e1 e2 e1' e2' : expr) n : - n ⊨ E_rel_o Γ e1 e2 → - n ⊨ E_rel_o Γ e1' e2' → - n ⊨ E_rel_o Γ (app e1 e1') (app e2 e2'). +Lemma compat_app {V} (e1 e2 e1' e2' : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o e1' e2' → + n ⊨ E_rel_o (app e1 e1') (app e2 e2'). Proof. intros He He'. apply E_rel_o_elim in He. (* From He, we have contextual equivalence of e1 and e2, in related context *) apply E_rel_o_elim in He'. - apply E_rel_o_intro. - iintros γ1 γ2 Hγ. simpl. - ispecialize He γ1. ispecialize He γ2. ispec He Hγ. - ispecialize He' γ1. ispecialize He' γ2. ispec He' Hγ. - apply E_rel_elim in He. - apply E_rel_elim in He'. - apply E_rel_intro. - (* we have to show that the apps are in O, given that they are placed - in related contexts E1 and E2. *) - iintros E1 E2 HE. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + apply E_rel_intro. iintros E1 E2 HE. term_simpl. (* The functions e1/e2 are evaluated first, so we "zap" them down using He. To use He, we have to give two contexts s.t. if we can prove them to be related, plugging e1/e2 into them will be in O. We give ectx_app1 because the plugging will give us exactly the goal we need. *) - ispecialize He (ectx_app1 E1 (subst_map γ1 e1')). - ispecialize He (ectx_app1 E2 (subst_map γ2 e2')). + ispec He γ1 γ2 Hγ. + apply E_rel_elim in He. + ispecialize He (ectx_app1 E1 (bind γ1 e1')). + ispecialize He (ectx_app1 E2 (bind γ2 e2')). iapply He. (* This reduces the problem to that of showing that the two app contexts are related. *) apply K_rel_intro. @@ -1109,6 +1117,8 @@ Proof. (* Given that they are plugged with two related values, we now have to prove that the result is in O. We use He' for a similar purpose. We give ectx_app2 because plugging e1'/e2' into it will match the goal. *) + ispec He' γ1 γ2 Hγ. + apply E_rel_elim in He'. ispecialize He' (ectx_app2 v1 E1). ispecialize He' (ectx_app2 v2 E2). iapply He'. @@ -1124,312 +1134,98 @@ Proof. exact HE. Qed. -Lemma subset_is_closed_absurd x Γ γ: - x ∈ Γ → - subst_is_closed Γ ∅ γ → - γ !! x = None → - False. -Proof. - intros Hx Hs He. - pose proof (not_elem_of_dom_2 _ _ He). - destruct Hs as [? _]. - setoid_rewrite <- H0 in H. - set_solver. -Qed. - -Lemma compat_var Γ (x : string) n : - closed Γ (var x) → - n ⊨ E_rel_o Γ (var x) (var x). -Proof. - intros Hdom. - rewrite closed_var in Hdom. - apply E_rel_o_intro. - iintros γ1 γ2 Hγ. - apply G_rel_elim in Hγ as (Hc1 & Hc2 & Hγ). - apply E_rel_intro. - iintros E1 E2 HE. simpl. - destruct (γ1 !! x) eqn:Hx1. - 2: { destruct (subset_is_closed_absurd _ _ _ Hdom Hc1 Hx1). } - destruct (γ2 !! x) eqn:Hx2. - 2: { destruct (subset_is_closed_absurd _ _ _ Hdom Hc2 Hx2). } - destruct Hc1 as [_ Hc1]. - destruct Hc2 as [_ Hc2]. - ispec Hγ x v v0 Hx1 Hx2. - by apply K_rel_elimO. -Qed. - -Lemma subst_val_closed v X x es : - closed X (of_val v) → x ∉ X → subst_val x es v = v -with subst_closed X e x es : - closed X e → x ∉ X → subst x es e = e. -Proof. - { induction v. - (*{ reflexivity. }*) - { simpl. - case_decide. - - reflexivity. - - intros. - f_equal. - rewrite closed_lambda in H0. - apply (subst_closed _ _ _ _ H0). - set_solver. } - (*{ reflexivity. }*) } - { induction e; intros; simpl. - { f_equal. - eapply subst_val_closed. - apply H. assumption. } - { case_decide. - - subst. - unfold closed in H. simpl in H. apply bool_decide_unpack in H. - set_solver. - - reflexivity. } - { apply closed_app in H. - destruct H as (H1&H2). - specialize (IHe1 H1 H0). - specialize (IHe2 H2 H0). - f_equal. - - assumption. - - assumption. } } -Qed. - -Lemma not_base_step_val (v : val) e : ¬ base_step v e. -Proof. intros Hstep. inversion Hstep. Qed. - -Lemma val_eq_fill_inv (v : val) K e : - ret v = fill K e → - e = v ∧ K = ectx_hole. -Proof. - revert e. - induction K; intros e' H_eq. - - auto. - - specialize (IHK _ H_eq) as (H_absurd & _). discriminate. - - specialize (IHK _ H_eq) as (H_absurd & _). discriminate. -Qed. - -Lemma val_eq_rplug_inv (v : val) K e : - ret v = rplug K e → - e = v ∧ K = rctx_hole. -Proof. - intros H_eq. - destruct K. - - simpl in *. auto. - - simpl in *. discriminate. - - simpl in *. discriminate. -Qed. - -Lemma not_contextual_step_val : ∀ (v : val) e, ¬ contextual_step v e. -Proof. - intros v e Hstep. - inversion Hstep. - apply val_eq_fill_inv in H as [-> ->]. - by eapply not_base_step_val. -Qed. - -Lemma base_step_is_deterministic : - ∀ e1 e2 e3, - base_step e1 e2 → - base_step e1 e3 → - e2 = e3. -Proof. - intros e1 e2 e3 Hstep2 Hstep3. - inversion Hstep2. - inversion Hstep3. - congruence. -Qed. - -Inductive potential_redex : expr -> Prop := -| pr_app : ∀ (v1 v2 : val), potential_redex (app v1 v2). - -Lemma potential_redex_not_val (v : val) : ¬ potential_redex v. -Proof. intros H_absurd. inversion H_absurd. Qed. - -Lemma unique_partial_decomposition E1 E2 e1 e2 : - potential_redex e1 → - potential_redex e2 → - rplug E1 e1 = rplug E2 e2 → - E1 = E2 ∧ e1 = e2. -Proof. - intros He1 He2. - revert E2. - induction E1; intros E2 H_eq. - - destruct E2. - + simpl in *. auto. - + simpl in *. subst. inversion He1. - apply val_eq_rplug_inv in H0 as []. subst. - exfalso. by eapply potential_redex_not_val. - + simpl in *. subst. inversion He1. - apply val_eq_rplug_inv in H1 as []. subst. - exfalso. by eapply potential_redex_not_val. - - destruct E2. - + simpl in *. subst. inversion He2. - apply val_eq_rplug_inv in H0 as []. subst. - exfalso. by eapply potential_redex_not_val. - + simpl in *. injection H_eq as H_eq1 H_eq2. - specialize (IHE1 _ H_eq1) as []. subst. auto. - + simpl in *. injection H_eq as H_eq1 H_eq2. - symmetry in H_eq1. - apply val_eq_rplug_inv in H_eq1 as []. subst. - exfalso. by eapply potential_redex_not_val. - - destruct E2. - + simpl in *. subst. inversion He2. - apply val_eq_rplug_inv in H1 as []. subst. - exfalso. by eapply potential_redex_not_val. - + simpl in *. injection H_eq as H_eq1 H_eq2. - apply val_eq_rplug_inv in H_eq1 as []. subst. - exfalso. by eapply potential_redex_not_val. - + simpl in *. injection H_eq as H_eq1 H_eq2. - specialize (IHE1 _ H_eq2) as []. subst. auto. -Qed. - -Lemma unique_decomposition : - ∀ E1 E2 e1 e2, - potential_redex e1 → - potential_redex e2 → - fill E1 e1 = fill E2 e2 → - E1 = E2 ∧ e1 = e2. -Proof. - intros E1 E2 e1 e2 He1 He2 Heq. - rewrite -> plug_rplug_equiv in Heq. - rewrite -> plug_rplug_equiv in Heq. - destruct (unique_partial_decomposition _ _ _ _ He1 He2 Heq) as [Heq1 Heq2]. - split. - - rewrite <- (ectx_rctx_bijection1 E1). - rewrite <- (ectx_rctx_bijection1 E2). - rewrite -> Heq1. reflexivity. - - exact Heq2. -Qed. - -Lemma base_step_potential_redex e e' : - base_step e e' -> potential_redex e. +Lemma compat_var {V : Set} (x : V) n : + n ⊨ V_rel_o (v_var x) (v_var x). Proof. - inversion 1. subst. - destruct e2. - + constructor. - + simpl in *. contradiction. - + simpl in *. contradiction. -Qed. - -Lemma contextual_step_is_deterministic : - ∀ e1 e2 e3, - contextual_step e1 e2 → - contextual_step e1 e3 → - e2 = e3. -Proof. - intros e1 e2 e3 Hstep2 Hstep3. - inversion Hstep2. - inversion Hstep3. - assert (Hpr1 := base_step_potential_redex _ _ H1). - assert (Hpr2 := base_step_potential_redex _ _ H4). - destruct (unique_decomposition K K0 e1' e1'0 Hpr1 Hpr2) as [HK_eq He_eq]. - { congruence. } - rewrite -> He_eq in H1. - assert (He_eq' := base_step_is_deterministic e1'0 e2' e2'0 H1 H4). - congruence. + apply V_rel_o_intro. + iintros γ1 γ2 Hγ. term_simpl. + apply G_rel_elimV. exact Hγ. Qed. -Lemma L_rel_red_l (e1 e1' e2 : expr) n : +Lemma L_rel_red_l (e1 e1' e2 : expr ∅) n : contextual_step e1 e1' → n ⊨ ▷ L_rel e1' e2 → n ⊨ L_rel e1 e2. Proof. - intros Hred HL. - unfold L_rel. unfold L_rel_pre. - repeat isplit. - - iintro. - intros v1 H_eq. - rewrite -> H_eq in Hred. - exfalso. by eapply not_contextual_step_val. - - iintros e1'' Hred'. - idestruct Hred'. - rewrite -> (contextual_step_is_deterministic _ _ _ Hred' Hred). - later_shift. apply L_rel_roll. - exact HL. -Qed. - -Lemma L_rel_red_r (e2 e2' : expr) n : + intros H_step He. + apply L_rel_intro. + - intros v1 H_eq. + rewrite -> H_eq in H_step. + contradict (not_contextual_step_val _ _ H_step). + - iintros e1'' H_step'. + idestruct H_step'. + rewrite -> (contextual_step_is_deterministic _ _ _ H_step' H_step). + exact He. +Qed. + +Lemma L_rel_red_r (e2 e2' : expr ∅) n : contextual_step e2 e2' → n ⊨ (∀ᵢ e1, L_rel e1 e2' →ᵢ L_rel e1 e2). Proof. - intros Hred. - loeb_induction. - iintros e1 HL. - unfold L_rel in HL. - unfold L_rel_pre in HL. - idestruct HL as HL1 HL2. - repeat isplit. - - iintro. intros v1 H_eq. - idestruct HL1. - specialize (HL1 v1 H_eq). - unfold terminates in *. - unfold bigstep in *. - destruct HL1 as (v & e3 & Hrtc & H_terminates). - exists v, e3. split. - + eapply rtc_l. exact Hred. exact Hrtc. - + exact H_terminates. - - iintros e1' Hred'. - ispec HL2 e1' Hred'. - later_shift. - apply L_rel_unroll in HL2. - apply L_rel_roll. - iapply IH. exact HL2. + intros H_step. + loeb_induction IH. + iintros e1 He. + apply L_rel_elim in He as [He1 He2]. + apply L_rel_intro. + - intros v1 H_eq. + specialize (He1 v1 H_eq). + eapply contextual_step_terminates; eauto. + - iintros e1' H_step'. + ispec He2 e1' H_step'. + later_shift. iapply IH. exact He2. Qed. -Lemma O_rel_red_l (e1 e1' e2 : expr) n : +Lemma O_rel_red_l (e1 e1' e2 : expr ∅) n : contextual_step e1 e1' → n ⊨ O_rel e1' e2 → n ⊨ O_rel e1 e2. Proof. - intros Hred HO. - unfold O_rel in *. - idestruct HO as HL1 HL2. - isplit. + intros H_step He. + apply O_rel_elim in He as [He1 He2]. + apply O_rel_intro. - eapply L_rel_red_l. - + exact Hred. - + later_shift. exact HL1. + + exact H_step. + + later_shift. exact He1. - iapply L_rel_red_r. - + exact Hred. - + exact HL2. + + exact H_step. + + exact He2. Qed. -Lemma O_rel_red_r (e1 e2 e2' : expr) n : - (* contextual_step e1 e1' → contextual_step e2 e2' → *) +(* symmetric to the proof of O_rel_red_l *) +Lemma O_rel_red_r (e1 e2 e2' : expr ∅) n : contextual_step e2 e2' → n ⊨ O_rel e1 e2' → n ⊨ O_rel e1 e2. Proof. - intros Hred HO. - unfold O_rel in *. - idestruct HO as HL1 HL2. - isplit. + intros H_step He. + apply O_rel_elim in He as [He1 He2]. + apply O_rel_intro. - iapply L_rel_red_r. - + exact Hred. - + exact HL1. - - iapply L_rel_red_l. - + exact Hred. - + later_shift. exact HL2. + + exact H_step. + + exact He1. + - eapply L_rel_red_l. + + exact H_step. + + later_shift. exact He2. Qed. -Lemma O_rel_red_both (e1 e1' e2 e2' : expr) n : +Lemma O_rel_red_both (e1 e1' e2 e2' : expr ∅) n : contextual_step e1 e1' → contextual_step e2 e2' → n ⊨ ▷ O_rel e1' e2' → n ⊨ O_rel e1 e2. Proof. - intros Hred1 Hred2 HO. - unfold O_rel in *. - apply I_conj_later_down in HO. - idestruct HO as HL1 HL2. - isplit. - - iapply L_rel_red_r. - { exact Hred2. } - iapply L_rel_red_l. - { exact Hred1. } - { later_shift. exact HL1. } - - iapply L_rel_red_r. - { exact Hred1. } - iapply L_rel_red_l. - { exact Hred2. } - { later_shift. exact HL2. } + intros H_step1 H_step2 He. + unfold O_rel in He. + apply I_conj_later_down in He. + idestruct He as He1 He2. + apply O_rel_intro. + - eapply L_rel_red_l. { exact H_step1. } + later_shift. + iapply L_rel_red_r. { exact H_step2. } + exact He1. + - iapply L_rel_red_r. { exact H_step1. } + iapply L_rel_red_l. { exact H_step2. } + later_shift. exact He2. Qed. (* Observation: later_shift is significant in O_rel_red_both, @@ -1439,363 +1235,96 @@ Qed. - O_rel_red_both → O_rel_red_l ∧ O_rel_red_r - But not: O_rel_red_l ∧ O_rel_red_r → O_rel_red_both *) -Lemma R_rel_red_both (e1 e1' e2 e2' : expr) n : +Lemma R_rel_red_both (e1 e1' e2 e2' : expr ∅) n : contextual_step e1 e1' → contextual_step e2 e2' → n ⊨ ▷ E_rel e1' e2' → n ⊨ R_rel e1 e2. Proof. - intros Hred1 Hred2 He. + intros H_step1 H_step2 He. apply R_rel_intro. iintros E1 E2 HE. eapply O_rel_red_both. - { by apply contextual_step_comp. } - { by apply contextual_step_comp. } - { later_shift. by apply E_rel_elimO. } + - by apply contextual_step_comp. + - by apply contextual_step_comp. + - later_shift. by apply E_rel_elimO. Qed. -(* composition of subst and subst_map, - where the variable to be substituted doesn't appear in the substitution *) -Lemma subst_subst_map : ∀ (e:expr) Γ (x : string) (es : val) (map : sub), - subst_is_closed Γ ∅ map → - subst x es (subst_map (delete x map) e) = - subst_map (insert x es map) e -with subst_subst_map_val : ∀ (v:val) Γ (x : string) (es : val) (map : sub), - subst_is_closed Γ ∅ map → - subst x es (subst_map_val (delete x map) v) = - subst_map_val (insert x es map) v. +Lemma compat_lambda {V} (e1 e2 : expr (inc V)) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ V_rel_o (v_lambda e1) (v_lambda e2). Proof. - { intros e. induction e. - { intros. apply (subst_subst_map_val _ _ _ _ _ H). } - { (* e is a variable x *) - intros. simpl. destruct (decide (x0=x)) as [->|Hne]. - { (* if x=x0, we'll end up substituting es into x *) - rewrite lookup_delete_eq with (m:=map). - rewrite lookup_insert_eq with (m:=map). - simpl. - by rewrite decide_True. } - { (* if not equal, the deletion and insertion will have no effect *) - rewrite lookup_delete_ne with (m:=map). 2: { assumption. } - rewrite lookup_insert_ne with (m:=map). 2: { assumption. } - (* we then need to see if x is in the map to begin with *) - destruct (map !! x) as [v1|] eqn:Hkey. - { Fail rewrite Hkey. (* why does regular rewrite not work? *) - setoid_rewrite Hkey. - simpl. - rewrite (subst_val_closed _ ∅ _ _). - - reflexivity. - - apply (subst_is_closed_elim_closed _ _ x _ _ H Hkey). - - set_solver. } - { setoid_rewrite Hkey. - simpl. - by rewrite decide_False. } } } - { intros. simpl. f_equal; eauto. } } - { intros v. induction v; intros. - (*{ reflexivity. }*) - { (* the lambda case *) - simpl. f_equal. f_equal. - case_decide. - { subst. - rewrite delete_delete_eq with (m:=map). - rewrite delete_insert_eq with (m:=map). done. } - { rewrite delete_insert_ne with (m:=map). 2: { congruence. } - rewrite delete_delete with (m:=map). - eapply subst_subst_map. - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x map) map). - destruct H as [H1 H2]. - rewrite H1. - set_solver. - apply delete_subseteq. - set_solver. - assumption. } } - (*{ reflexivity. }*) } + intros He. + apply E_rel_o_elim in He. + apply V_rel_o_intro. iintros γ1 γ2 Hγ. + apply V_rel_intro. iintros u1 u2 Hu. term_simpl. + eapply R_rel_red_both. + { eapply (Ectx_step ectx_hole). constructor. } + { eapply (Ectx_step ectx_hole). constructor. } + later_shift. simpl. unfold subst. + rewrite -> bind_bind_comp'. + rewrite -> bind_bind_comp'. + iapply He. + apply G_rel_intro. + iintros x. destruct x as [| x']. + - term_simpl. exact Hu. + - term_simpl. apply G_rel_elimV. exact Hγ. Qed. -(** Special case of Theorem A.1 from Erlang paper: - scoping of extended substitutions. - Given a closed substitution, we can add a closed value to it. *) -Lemma scope_extend1 Γ x (v:val) (γ:sub): - closed ∅ v → - subst_is_closed Γ ∅ γ → - subst_is_closed (Γ ∪ {[x]}) ∅ (<[x := v]> γ). -Proof. - intros Hc Hsc. - split. - { destruct Hsc. rewrite H. set_solver. } - intros x0 Hd v0 Hs. - (* we have to prove that for an arbitrary binding x0 := v0 in γ, v0 is closed *) - destruct (decide (x=x0)) as [->|Hne]. - (* if x = x0, the premise tells us v0 is closed *) - { rewrite lookup_insert_eq with (m:=γ) in Hs. - injection Hs. intros. subst. - exact Hc. } - (* if they are not equal, we know x0 is in Γ and have to use the fact - that the subst_is_closed *) - { rewrite elem_of_union in Hd. destruct Hd. 2: { assert (x0 = x). set_solver. done. } - destruct Hsc as [_ Hsc]. - rewrite lookup_insert_ne with (m:=γ) in Hs. 2: { assumption. } - specialize (Hsc x0 H v0 Hs). - assumption. } -Qed. - -(* we can extend related substitutions with related values *) -Lemma sem_context_rel_insert Γ x v1 v2 γ1 γ2 n: - n ⊨ V_rel v1 v2 → - n ⊨ G_rel Γ γ1 γ2 → - n ⊨ G_rel (Γ ∪ {[x]}) (<[x := v1]> γ1) (<[x := v2]> γ2). -Proof. - intros Hv Hγ. - destruct (V_rel_elim _ _ _ Hv) as (Hvc1 & Hvc2 & Hv'). - destruct (G_rel_elim _ _ _ _ Hγ) as (Hγc1 & Hγc2 & Hγ'). - apply G_rel_intro. - { by apply scope_extend1. } - { by apply scope_extend1. } - - iintros x0 v0 v3 H1 H2. - destruct (decide (x=x0)). - { subst. - rewrite lookup_insert_eq with (m:=γ2) in H2. idestruct H2. injection H2 as ->. - rewrite lookup_insert_eq with (m:=γ1) in H1. idestruct H1. injection H1 as ->. - assumption. } - { rewrite lookup_insert_ne with (m:=γ2) in H2. idestruct H2. 2: { assumption. } - rewrite lookup_insert_ne with (m:=γ1) in H1. idestruct H1. 2: { assumption. } - iapply Hγ'. - - iintro. eassumption. - - iintro. eassumption. } -Qed. - -Lemma lambda_closed Γ γ x e : - closed (Γ ∪ {[x]}) e → - subst_is_closed Γ ∅ γ → - closed ∅ (vlambda x (subst_map (delete x γ) e)). -Proof. - intros Hec [Heq Hγc]. - rewrite closed_lambda. - eapply subst_map_closed'_2. - - eapply closed_weaken. - exact Hec. - setoid_rewrite dom_delete. - intros y. destruct (decide (x = y)); set_solver. - - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x γ) γ). - + set_solver. - + apply delete_subseteq. - + set_solver. - + unfold subst_is_closed. split. - * exact Heq. - * intros x' Hin v Hlookup. - specialize (Hγc x' Hin v Hlookup). - by eapply closed_weaken. -Qed. - -Lemma compat_lambda Γ (e1 e2 : expr) n x : - closed Γ (vlambda x e1) → - closed Γ (vlambda x e2) → - n ⊨ E_rel_o (Γ ∪ {[x]}) e1 e2 → - n ⊨ V_rel_o Γ (vlambda x e1) (vlambda x e2). -Proof. - intros Hc1 Hc2 He. - apply E_rel_o_elim in He. - apply V_rel_o_intro. - iintros γ1 γ2 Hγ. - destruct (G_rel_elim _ _ _ _ Hγ) as (Hγc1 & Hγc2 & Hγ'). - apply V_rel_intro. - - { by apply (subst_map_closed'_3 (vlambda x e1) Γ γ1). } - { by apply (subst_map_closed'_3 (vlambda x e2) Γ γ2). } - - (* we now have the arguments *) - iintros u1 u2 Hu. - eapply R_rel_red_both. - { simpl. eapply (Ectx_step _ _ ectx_hole _ _ eq_refl eq_refl). - simpl. constructor. - simpl. constructor. - reflexivity. } - { simpl. eapply (Ectx_step _ _ ectx_hole _ _ eq_refl eq_refl). - simpl. constructor. - simpl. constructor. - reflexivity. } - { later_shift. - - rewrite subst_subst_map with (Γ:=Γ). - (* 2: { pose proof (G_sub_closed _ _ _ _ Hγ) as [_ ?]. assumption. } *) - 2: { have: G_rel_elim Hγ. done. } - rewrite subst_subst_map with (Γ:=Γ). - 2: { have: G_rel_elim Hγ. done. } - iapply He. - (* apply (sem_context_rel_insert _ _ _ _ _ _ _ Hv Hγ). *) - applyy sem_context_rel_insert Hu Hγ. } -Qed. - -Lemma fundamental_property_e Γ (e : expr) n : - closed Γ e → - n ⊨ E_rel_o Γ e e -with fundamental_property_v Γ (v : val) n : - closed Γ v → - n ⊨ V_rel_o Γ v v. -Proof. - { intros H_closed. - induction e. - - apply compat_val. apply fundamental_property_v. assumption. - - apply compat_var. assumption. - - rewrite -> closed_app in H_closed. - destruct H_closed as [H_closed1 H_closed2]. - apply compat_app; auto. } - { intros H_closed. - induction v. - - apply compat_lambda. - + exact H_closed. - + exact H_closed. - + rewrite -> closed_lambda in H_closed. - apply fundamental_property_e. assumption. } -Qed. - -Lemma G_rel_empty n : - n ⊨ G_rel ∅ ∅ ∅. -Proof. - apply G_rel_intro. - - unfold subst_is_closed. split; set_solver. - - unfold subst_is_closed. split; set_solver. - - iintros. - idestruct H. - setoid_rewrite lookup_empty in H. - discriminate H. -Qed. - -Lemma subst_map_empty (e:expr) : - subst_map ∅ e = e -with subst_map_val_empty (v:val) : - subst_map_val ∅ v = v. +Lemma fundamental_property_e {V} (e : expr V) n : + n ⊨ E_rel_o e e +with fundamental_property_v {V} (v : val V) n : + n ⊨ V_rel_o v v. Proof. { induction e. - - simpl. f_equal. eapply subst_map_val_empty. - - simpl. - setoid_rewrite lookup_empty. - reflexivity. - - simpl. rewrite IHe1. rewrite IHe2. reflexivity. } + - apply compat_val. apply fundamental_property_v. + - apply compat_app; auto. } { induction v. - - simpl. - f_equal. - setoid_rewrite delete_empty. - apply subst_map_empty. } -Qed. - -Lemma fundamental_property_sub Γ γ n : - subst_is_closed Γ ∅ γ → - n ⊨ G_rel Γ γ γ. -Proof. - intros Hc. - apply G_rel_intro. - { exact Hc. } - { exact Hc. } - iintros x v1 v2 Heq1 Heq2. - idestruct Heq1. - idestruct Heq2. - rewrite -> Heq1 in Heq2. - injection Heq2 as ->. - destruct Hc as [-> Hc]. - assert (Hcv2 : closed ∅ v2). - { eapply Hc. - - setoid_rewrite -> elem_of_dom. - unfold is_Some. eauto. - - exact Heq1. } - assert (H_fundamental := fundamental_property_v ∅ v2 n Hcv2). - apply V_rel_o_elim in H_fundamental. - ispec H_fundamental ∅ ∅ (G_rel_empty n). - rewrite -> subst_map_val_empty in H_fundamental. - exact H_fundamental. + - apply compat_var. + - apply compat_lambda. + apply fundamental_property_e. } Qed. -Lemma fundamental_property_ectx E n : - n ⊨ K_rel E E. -Proof. - apply K_rel_intro. - iintros v1 v2 Hv. - admit. -Admitted. - (** General program contexts *) -Inductive ctx : Type := - | ctx_hole : ctx - | ctx_lam : name → ctx → ctx - | ctx_app1 : ctx → expr → ctx - | ctx_app2 : expr → ctx → ctx. - -(* Inside-out plugging *) -Fixpoint ciplug (C : ctx) : expr → expr := +Inductive ctx : Set → Type := +| ctx_hole : ctx ∅ +| ctx_lam {V} : ctx V → ctx (inc V) +| ctx_app1 {V} : ctx V → expr V → ctx V +| ctx_app2 {V} : expr V → ctx V → ctx V. + +(** Inside-out plugging, always result in a closed expression *) +Fixpoint ciplug {V} (C : ctx V) : expr V → expr ∅ := match C with - | ctx_hole => id - | ctx_lam x C => λ e, ciplug C (ret (vlambda x e)) + | ctx_hole => λ e, e + | ctx_lam C => λ e, ciplug C (v_lambda e) | ctx_app1 C e2 => λ e, ciplug C (app e e2) | ctx_app2 e1 C => λ e, ciplug C (app e1 e) end. +Inductive ctxr : Set → Set → Type := +| ctxr_hole {A} : ctxr A A +| ctxr_lam {A B} : ctxr A (inc B) → ctxr A B +| ctxr_app1 {A B} : ctxr A B → expr B → ctxr A B +| ctxr_app2 {A B} : expr B → ctxr A B → ctxr A B. + (* Outside-in plugging *) -Fixpoint crplug (C : ctx) (e : expr) : expr := +Fixpoint crplug {A B} (C : ctxr A B) : expr A → expr B := match C with - | ctx_hole => e - | ctx_app1 C' e1 => app (crplug C' e) e1 - | ctx_app2 v C' => app v (crplug C' e) - | ctx_lam x C' => vlambda x (crplug C' e) + | ctxr_hole => λ e, e + | ctxr_lam C => λ e, v_lambda (crplug C e) + | ctxr_app1 C e2 => λ e, app (crplug C e) e2 + | ctxr_app2 e1 C => λ e, app e1 (crplug C e) end. (* Outside-in plugging simplifies the proofs below *) Notation cplug := crplug. -(* aka contextual scoping C : Γ ~> Γ', a special case of contextual typing. - defined inductively because we need to invert it. *) -Inductive closed_ctx : scope → scope → ctx → Prop := - | cc_hole Γ : - closed_ctx Γ Γ ctx_hole - - | cc_lambda x Γ Γ' C : - closed_ctx Γ (Γ' ∪ {[x]}) C → - closed_ctx Γ Γ' (ctx_lam x C) - - | cc_app1 Γ Γ' C e : - closed_ctx Γ Γ' C → - closed Γ' e → - closed_ctx Γ Γ' (ctx_app1 C e) - - | cc_app2 Γ Γ' C v : - closed_ctx Γ Γ' C → - closed Γ' (ret v) → - closed_ctx Γ Γ' (ctx_app2 v C) - . - -Definition closed_ctx_sem (Γ Γ' : scope) (C:ctx) : Prop := - forall e, closed Γ e → closed Γ' (cplug C e). - -Lemma closed_ctx_sound Γ Γ' e : - closed_ctx Γ Γ' e → closed_ctx_sem Γ Γ' e. -Proof. - intros H. induction H; unfold closed_ctx_sem. - - simpl. done. - - intros e Hc. - simpl. - specialize (IHclosed_ctx e Hc). - unfold closed. simpl. - apply IHclosed_ctx. - - intros e2 Hc. - specialize (IHclosed_ctx _ Hc). - simpl. - unfold closed. simpl. - unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. - auto. - - intros e1 Hc. - specialize (IHclosed_ctx _ Hc). - simpl. - unfold closed. simpl. - unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. - auto. -Qed. - (** Observational approximation for complete programs *) -Definition obs_approx (e1 e2 : expr) : Prop := - terminates e1 → - terminates e2. +Definition obs_approx (e1 e2 : expr ∅) : Prop := + terminates e1 → terminates e2. (** Observational equivalence for complete programs *) -Definition obs_equiv (e1 e2 : expr) : Prop := +Definition obs_equiv (e1 e2 : expr ∅) : Prop := terminates e1 ↔ terminates e2. Infix "≼obs" := obs_approx (at level 80, right associativity, only printing). @@ -1803,229 +1332,950 @@ Infix "≡obs" := obs_equiv (at level 80, right associativity, only printing). #[global] Instance Reflexive_obs_approx : Reflexive obs_approx. -Proof. - unfold Reflexive, obs_approx. done. -Qed. +Proof. unfold Reflexive, obs_approx. auto. Qed. #[global] Instance Transitive_obs_approx : Transitive obs_approx. -Proof. - unfold Transitive, obs_approx. intros. - auto. -Qed. +Proof. unfold Transitive, obs_approx. auto. Qed. #[global] Instance Reflexive_obs_equiv : Reflexive obs_equiv. -Proof. - unfold Reflexive, obs_equiv. intros. - reflexivity. -Qed. +Proof. unfold Reflexive, obs_equiv. auto. Qed. #[global] Instance Symmetric_obs_equiv : Symmetric obs_equiv. -Proof. - unfold Symmetric, obs_equiv. intros. - auto. -Qed. +Proof. unfold Symmetric, obs_equiv. auto. Qed. #[global] Instance Transitive_obs_equiv : Transitive obs_equiv. -Proof. - unfold Transitive, obs_equiv. intros. - destruct H. - destruct H0. - split; auto. -Qed. +Proof. unfold Transitive, obs_equiv. intuition auto. Qed. (** Contextual approximation, where the context closes off Γ *) -Definition ctx_approx Γ (e1 e2 : expr) : Prop := - ∀ C, closed_ctx Γ ∅ C → - obs_approx (cplug C e1) (cplug C e2). +Definition ctx_approx {V} (e1 e2 : expr V) : Prop := + ∀ C, obs_approx (cplug C e1) (cplug C e2). (** Contextual equivalence *) -Definition ctx_equiv Γ (e1 e2 : expr) : Prop := - ∀ C, closed_ctx Γ ∅ C → - obs_equiv (cplug C e1) (cplug C e2). +Definition ctx_equiv {V} (e1 e2 : expr V) : Prop := + ∀ C, obs_equiv (cplug C e1) (cplug C e2). (* TODO *) (* Infix "≼ctx" := ctx_approx (at level 80, right associativity, only printing). Infix "≡ctx" := ctx_equiv (at level 80, right associativity, only printing). *) #[global] -Instance Reflexive_ctx_approx Γ : Reflexive (ctx_approx Γ). +Instance Reflexive_ctx_approx {V} : Reflexive (@ctx_approx V). +Proof. unfold Reflexive, ctx_approx. reflexivity. Qed. + +#[global] +Instance Transitive_ctx_approx {V} : Transitive (@ctx_approx V). +Proof. unfold Transitive, ctx_approx. etransitivity; eauto. Qed. + +#[global] +Instance Reflexive_ctx_equiv {V} : Reflexive (@ctx_equiv V). +Proof. unfold Reflexive, ctx_equiv. reflexivity. Qed. + +#[global] +Instance Symmetric_ctx_equiv {V} : Symmetric (@ctx_equiv V). +Proof. unfold Symmetric, ctx_equiv. symmetry. auto. Qed. + +#[global] +Instance Transitive_ctx_equiv {V} : Transitive (@ctx_equiv V). +Proof. unfold Transitive, ctx_equiv. etransitivity; eauto. Qed. + +Lemma ctx_equiv_intro {V} (e1 e2 : expr V) : + ctx_approx e1 e2 → + ctx_approx e2 e1 → + ctx_equiv e1 e2. Proof. - intros e C; reflexivity. + intros He1 He2. split. + - apply He1. + - apply He2. Qed. -#[global] -Instance Transitive_ctx_approx Γ : Transitive (ctx_approx Γ). +Definition n_big_step {V} n (e : expr V) (v : val V) := nsteps contextual_step n e v. +Definition n_terminates {V} n (e : expr V) := ∃ v, n_big_step n e v. + +Lemma big_step_iff_n_big_step {V} (e : expr V) (v : val V) : + big_step e v ↔ ∃ n, n_big_step n e v. Proof. - unfold Transitive, ctx_approx, obs_approx. intros. - auto. + unfold big_step, n_big_step. + apply rtc_nsteps. Qed. -#[global] -Instance Reflexive_ctx_equiv Γ : Reflexive (ctx_equiv Γ). +Lemma terminates_iff_n_terminates {V} (e : expr V) : + terminates e ↔ ∃ n, n_terminates n e. +Proof. + unfold terminates, n_terminates. + split. + - intros [v H_big_step]. + apply big_step_iff_n_big_step in H_big_step as [n H_n_big_step]. + eauto. + - intros (n & v & H_n_big_step). + exists v. apply big_step_iff_n_big_step. + exists n. exact H_n_big_step. +Qed. + +Lemma n_big_step_O_inv {V} e (v : val V) : + n_big_step O e v → e = v. Proof. - intros e C; reflexivity. + unfold n_big_step. + inversion_clear 1. auto. Qed. -#[global] -Instance Symmetric_ctx_equiv Γ : Symmetric (ctx_equiv Γ). +Lemma n_big_step_S_inv {V} n e (v : val V) : + n_big_step (S n) e v → + ∃ e', contextual_step e e' ∧ n_big_step n e' v. Proof. - intros e1 e2 H C H1. symmetry. apply H. assumption. + unfold n_big_step. + inversion_clear 1. eauto. Qed. -#[global] -Instance Transitive_ctx_equiv Γ : Transitive (ctx_equiv Γ). +Lemma n_terminates_O_inv {V} (e : expr V) : + n_terminates O e → + ∃ (v : val V), e = v. Proof. - unfold Transitive, ctx_approx, ctx_equiv, obs_equiv. intros. - specialize (H C). - specialize (H0 C). - destruct H. destruct H0. - assumption. - assumption. - specialize (H0 H1) as (?&?). - split; auto. + unfold n_terminates. + intros [v H_n_big_step]. + apply n_big_step_O_inv in H_n_big_step. eauto. Qed. -Lemma ctx_equiv_fold Γ (e1 e2 : expr) : - ctx_approx Γ e1 e2 → - ctx_approx Γ e2 e1 → - ctx_equiv Γ e1 e2. +Lemma n_terminates_S_inv {V} n (e : expr V) : + n_terminates (S n) e → + ∃ e', contextual_step e e' ∧ n_terminates n e'. Proof. - intros H1 H2 C; split; apply H1 || apply H2. - assumption. - assumption. + unfold n_terminates. + intros [v H_n_big_step]. + apply n_big_step_S_inv in H_n_big_step as (e' & H_step & H_n_big_step). + eauto. Qed. -Lemma ctx_equiv_unfold Γ (e1 e2 : expr) : - ctx_equiv Γ e1 e2 → - ctx_approx Γ e1 e2 ∧ - ctx_approx Γ e2 e1. +Lemma ctx_equiv_elim {V} (e1 e2 : expr V) : + ctx_equiv e1 e2 → + ctx_approx e1 e2 ∧ + ctx_approx e2 e1. Proof. - unfold ctx_approx, ctx_equiv. - intros H. - split. - { intros C Hc. - specialize (H C Hc). - destruct H. - unfold obs_approx. - apply H. } - { intros C Hc. - specialize (H C Hc). - destruct H. - unfold obs_approx. - apply H0. } -Qed. - -Lemma L_rel_adequacy (v : val) (e1 e2 : expr) : - bigstep e1 v → (∀ w, w ⊨ L_rel e1 e2) → terminates e2. -Proof. - intro RED; remember v as p; revert RED Heqp. - intros Hbig. - destruct Hbig as (e&H1&H2). - (* induction on the rtc of contextual step *) - induction H1; intros Hp Hobs. - + (* expose the fact that x is a value *) - unfold to_val in H2. - destruct x. 2: { done. } 2: { done. } - injection H2; intros; subst; clear H2. - (* no steps are needed *) - specialize (Hobs {| nw_index := 0 |}). apply I_conj_elim1, I_prop_elim in Hobs. - by apply (Hobs v). - + apply IHrtc. assumption. assumption. - intro w; specialize (Hobs (world_lift w)). - apply L_rel_unroll, I_world_lift_later. - iapply Hobs; iintro; assumption. + unfold ctx_equiv, ctx_approx, obs_equiv, obs_approx. + intros He. split; apply He. +Qed. + +Lemma L_rel_adequacy_n (e1 e2 : expr ∅) n : + (∀ w, w ⊨ L_rel e1 e2) → + n_terminates n e1 → + terminates e2. +Proof. + revert e1. + induction n as [| n' IHn']; intros e1 He H_n_terminates. + - apply n_terminates_O_inv in H_n_terminates as [v Hv]. + specialize (He {| nw_index := O |}). + apply L_rel_elim in He as [He _]. + exact (He v Hv). + - apply n_terminates_S_inv in H_n_terminates as (e' & H_step & H_n_terminates). + apply (IHn' e'); [| exact H_n_terminates]. + intros w. specialize (He (world_lift w)). + apply L_rel_elim in He as [_ He]. + apply I_world_lift_later. + iapply He. iintro. exact H_step. +Qed. + +Lemma L_rel_adequacy (e1 e2 : expr ∅) : + (∀ n, n ⊨ L_rel e1 e2) → + terminates e1 → + terminates e2. +Proof. + intros He H_terminates. + apply terminates_iff_n_terminates in H_terminates as (n & H_n_terminates). + exact (L_rel_adequacy_n e1 e2 n He H_n_terminates). Qed. Theorem O_rel_adequacy e1 e2 : - (∀ n, n ⊨ O_rel e1 e2) → obs_equiv e1 e2. + (∀ w, w ⊨ O_rel e1 e2) → obs_equiv e1 e2. Proof. - intro Hobs; split. - + intros [ v Hv ]; eapply L_rel_adequacy; [ eassumption | ]. - intro. unfold O_rel in Hobs. iapply Hobs. - + intros [ v₂ Hv₂ ]; eapply L_rel_adequacy; [ eassumption | ]. - intro; iapply Hobs. + intros He. split. + - apply L_rel_adequacy. + intros w. specialize (He w). + by apply O_rel_elim in He as []. + - apply L_rel_adequacy. + intros w. specialize (He w). + by apply O_rel_elim in He as []. Qed. -Lemma L_rel_value (v1 v2 : val) n : +Lemma L_rel_value (v1 v2 : val ∅) n : n ⊨ L_rel v1 v2. Proof. - unfold L_rel, L_rel_pre. - isplit. - - iintro. intros. + apply L_rel_intro. + - intros _ _. apply terminates_val. - - iintros e1 He. - idestruct He. - apply not_contextual_step_val in He. done. + - iintros e1 He. idestruct He. + by apply not_contextual_step_val in He. Qed. -Lemma O_rel_value (v1 v2 : val) n : +Lemma O_rel_value (v1 v2 : val ∅) n : n ⊨ O_rel v1 v2. Proof. - unfold O_rel. - isplit; apply L_rel_value. + apply O_rel_intro; apply L_rel_value. Qed. -Lemma P_rel_compat_expr Γ (e1 e2 : expr) n : - n ⊨ E_rel_o Γ e1 e2 → - n ⊨ P_rel_o Γ e1 e2. +Lemma K_rel_ectx_hole n : + n ⊨ K_rel ectx_hole ectx_hole. Proof. - intro He. - iintros γ1 γ2 Hγ. - unfold E_rel_o in He. ispec He γ1 γ2 Hγ. - have: E_rel_elimO ectx_hole ectx_hole He. simpl in H. - apply H. apply K_rel_intro. - iintros. - simpl. + iintros v1 v2 Hv. simpl. apply O_rel_value. Qed. -Lemma precongruence (e1 e2 : expr) Γ Γ' C n : - closed Γ e1 → - closed Γ e2 → - closed_ctx Γ Γ' C → - n ⊨ E_rel_o Γ e1 e2 → - n ⊨ E_rel_o Γ' (cplug C e1) (cplug C e2). +Lemma O_rel_o_compat_expr {V} (e1 e2 : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ O_rel_o e1 e2. Proof. - revert Γ Γ' e1 e2 n. - induction C; intros Γ Γ' e1 e2 n Hc1 Hc2 Hcc HE; simpl. - { inversion Hcc. subst. done. } - { inversion Hcc. subst. - apply compat_val. - apply compat_lambda. - { have Hc: closed_ctx_sound Hcc. - apply (Hc _ Hc1). } - { have Hc: closed_ctx_sound Hcc. - apply (Hc _ Hc2). } - applyy IHC Hc1 Hc2 H2 HE. } - { inversion Hcc. subst. - apply compat_app. - - applyy IHC Hc1 Hc2 H3 HE. - - applyy fundamental_property_e H4. } - { inversion Hcc. subst. - apply compat_app. - - applyy fundamental_property_e H4. - - applyy IHC Hc1 Hc2 H3 HE. } -Qed. - -Theorem E_rel_o_soundness Γ (e1 e2 : expr) : - closed Γ e1 → - closed Γ e2 → - (∀ n, n ⊨ E_rel_o Γ e1 e2) → - ctx_equiv Γ e1 e2. -Proof. - intros Hc1 Hc2 HE C Hcc. - apply O_rel_adequacy; intro n. - have: precongruence Hc1 Hc2 Hcc (HE n). - have: P_rel_compat_expr H. - unfold P_rel_o in H0. - ispec H0 ∅ ∅. - rewrite <- (subst_map_empty (cplug C e1)). - rewrite <- (subst_map_empty (cplug C e2)). - iapply H0. - apply G_rel_empty. + intro He. + apply O_rel_o_intro. + iintros γ1 γ2 Hγ. + apply E_rel_o_elim in He. ispec He γ1 γ2 Hγ. + rewrite <- (fold_unfold_plug_ectx_hole (bind γ1 e1)). + rewrite <- (fold_unfold_plug_ectx_hole (bind γ2 e2)). + apply (E_rel_elimO _ _ _ _ _ He). + exact (K_rel_ectx_hole n). Qed. + +Lemma precongruence {A B} (e1 e2 : expr A) (C : ctxr A B) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (cplug C e1) (cplug C e2). +Proof. + intros He. + induction C; simpl. + - exact He. + - apply compat_val. + apply compat_lambda. + apply IHC. exact He. + - apply compat_app. + + apply IHC. exact He. + + apply fundamental_property_e. + - apply compat_app. + + apply fundamental_property_e. + + apply IHC. exact He. +Qed. + +Theorem E_rel_o_soundness {V} (e1 e2 : expr V) : + (∀ n, n ⊨ E_rel_o e1 e2) → + ctx_equiv e1 e2. +Proof. + unfold ctx_equiv. + intros He C. + apply O_rel_adequacy. intros w. + rewrite <- (bind_pure' (cplug C e1)). + rewrite <- (bind_pure' (cplug C e2)). + specialize (He w). + apply (precongruence _ _ C) in He. + apply O_rel_o_compat_expr in He. + apply (O_rel_o_elimO _ _ _ _ _ He). + apply G_rel_intro. iintros x. destruct x. +Qed. + +(* +(* composition of subst and subst_map, + where the variable to be substituted doesn't appear in the substitution *) +Lemma subst_subst_map : ∀ (e:expr) Γ (x : string) (es : val) (map : sub), + subst_is_closed Γ ∅ map → + subst x es (subst_map (delete x map) e) = + subst_map (insert x es map) e +with subst_subst_map_val : ∀ (v:val) Γ (x : string) (es : val) (map : sub), + subst_is_closed Γ ∅ map → + subst x es (subst_map_val (delete x map) v) = + subst_map_val (insert x es map) v. +Proof. + { intros e. induction e. + { intros. apply (subst_subst_map_val _ _ _ _ _ H). } + { (* e is a variable x *) + intros. simpl. destruct (decide (x0=x)) as [->|Hne]. + { (* if x=x0, we'll end up substituting es into x *) + rewrite lookup_delete_eq with (m:=map). + rewrite lookup_insert_eq with (m:=map). + simpl. + by rewrite decide_True. } + { (* if not equal, the deletion and insertion will have no effect *) + rewrite lookup_delete_ne with (m:=map). 2: { assumption. } + rewrite lookup_insert_ne with (m:=map). 2: { assumption. } + (* we then need to see if x is in the map to begin with *) + destruct (map !! x) as [v1|] eqn:Hkey. + { Fail rewrite Hkey. (* why does regular rewrite not work? *) + setoid_rewrite Hkey. + simpl. + rewrite (subst_val_closed _ ∅ _ _). + - reflexivity. + - apply (subst_is_closed_elim_closed _ _ x _ _ H Hkey). + - set_solver. } + { setoid_rewrite Hkey. + simpl. + by rewrite decide_False. } } } + { intros. simpl. f_equal; eauto. } } + { intros v. induction v; intros. + (*{ reflexivity. }*) + { (* the lambda case *) + simpl. f_equal. f_equal. + case_decide. + { subst. + rewrite delete_delete_eq with (m:=map). + rewrite delete_insert_eq with (m:=map). done. } + { rewrite delete_insert_ne with (m:=map). 2: { congruence. } + rewrite delete_delete with (m:=map). + eapply subst_subst_map. + apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x map) map). + destruct H as [H1 H2]. + rewrite H1. + set_solver. + apply delete_subseteq. + set_solver. + assumption. } } + (*{ reflexivity. }*) } +Qed. + +(** Special case of Theorem A.1 from Erlang paper: + scoping of extended substitutions. + Given a closed substitution, we can add a closed value to it. *) +Lemma scope_extend1 Γ x (v:val) (γ:sub): + closed ∅ v → + subst_is_closed Γ ∅ γ → + subst_is_closed (Γ ∪ {[x]}) ∅ (<[x := v]> γ). +Proof. + intros Hc Hsc. + split. + { destruct Hsc. rewrite H. set_solver. } + intros x0 Hd v0 Hs. + (* we have to prove that for an arbitrary binding x0 := v0 in γ, v0 is closed *) + destruct (decide (x=x0)) as [->|Hne]. + (* if x = x0, the premise tells us v0 is closed *) + { rewrite lookup_insert_eq with (m:=γ) in Hs. + injection Hs. intros. subst. + exact Hc. } + (* if they are not equal, we know x0 is in Γ and have to use the fact + that the subst_is_closed *) + { rewrite elem_of_union in Hd. destruct Hd. 2: { assert (x0 = x). set_solver. done. } + destruct Hsc as [_ Hsc]. + rewrite lookup_insert_ne with (m:=γ) in Hs. 2: { assumption. } + specialize (Hsc x0 H v0 Hs). + assumption. } +Qed. + +(* we can extend related substitutions with related values *) +Lemma sem_context_rel_insert Γ x v1 v2 γ1 γ2 n: + n ⊨ V_rel v1 v2 → + n ⊨ G_rel Γ γ1 γ2 → + n ⊨ G_rel (Γ ∪ {[x]}) (<[x := v1]> γ1) (<[x := v2]> γ2). +Proof. + intros Hv Hγ. + destruct (V_rel_elim _ _ _ Hv) as (Hvc1 & Hvc2 & Hv'). + destruct (G_rel_elim _ _ _ _ Hγ) as (Hγc1 & Hγc2 & Hγ'). + apply G_rel_intro. + { by apply scope_extend1. } + { by apply scope_extend1. } + + iintros x0 v0 v3 H1 H2. + destruct (decide (x=x0)). + { subst. + rewrite lookup_insert_eq with (m:=γ2) in H2. idestruct H2. injection H2 as ->. + rewrite lookup_insert_eq with (m:=γ1) in H1. idestruct H1. injection H1 as ->. + assumption. } + { rewrite lookup_insert_ne with (m:=γ2) in H2. idestruct H2. 2: { assumption. } + rewrite lookup_insert_ne with (m:=γ1) in H1. idestruct H1. 2: { assumption. } + iapply Hγ'. + - iintro. eassumption. + - iintro. eassumption. } +Qed. + +Lemma lambda_closed Γ γ x e : + closed (Γ ∪ {[x]}) e → + subst_is_closed Γ ∅ γ → + closed ∅ (vlambda x (subst_map (delete x γ) e)). +Proof. + intros Hec [Heq Hγc]. + rewrite closed_lambda. + eapply subst_map_closed'_2. + - eapply closed_weaken. + exact Hec. + setoid_rewrite dom_delete. + intros y. destruct (decide (x = y)); set_solver. + - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x γ) γ). + + set_solver. + + apply delete_subseteq. + + set_solver. + + unfold subst_is_closed. split. + * exact Heq. + * intros x' Hin v Hlookup. + specialize (Hγc x' Hin v Hlookup). + by eapply closed_weaken. +Qed. +*) + +(* +Notation name := string. +Definition sub : Set := gmap name val. +Definition scope : Set := gset name. + +Fixpoint subst_map_val (γ : sub) (v : val) : val := + match v with + (*| vunit => vunit + | vint n => vint n*) + | vlambda x e => vlambda x (subst_map (delete x γ) e) + end +with subst_map (γ : sub) (e : expr) : expr := + match e with + | ret v => ret (subst_map_val γ v) + (* | eunit => eunit *) + | var x => match γ !! x with Some v => ret v | _ => var x end + | app e1 e2 => app (subst_map γ e1) (subst_map γ e2) + (* | abs x e => abs x (subst_map (binder_delete x xs) e) *) + (* | eplus e1 e2 => eplus (subst_map xs e1) (subst_map xs e2) *) + end. + +Lemma fold_unfold_subst_map_val_vlambda γ x e : + subst_map_val γ (vlambda x e) = + vlambda x (subst_map (delete x γ) e). +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_ret γ v : + subst_map γ (ret v) = + ret (subst_map_val γ v). +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_var γ x : + subst_map γ (var x) = + match γ !! x with + | Some v => ret v + | _ => var x + end. +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_app γ e1 e2 : + subst_map γ (app e1 e2) = + app (subst_map γ e1) (subst_map γ e2). +Proof. auto. Qed. + +Fixpoint is_closed (X : scope) (e : expr) : bool := + match e with + | var x => bool_decide (x ∈ X) + (*| ret vunit | ret (vint _) => true*) + | ret (vlambda x e) => is_closed (X ∪ {[x]}) e + | app e1 e2 + (* | eplus e1 e2 *) + => is_closed X e1 && is_closed X e2 + end. + +(* aka a scoping judgment, the untyped equivalent of a typing judgment *) +Definition closed (X : scope) (e : expr) : Prop := Is_true (is_closed X e). + +Lemma closed_weaken e X Y: + closed X e → X ⊆ Y → closed Y e +with closed_weaken_val (v:val) X Y: + closed X v → X ⊆ Y → closed Y v. +Proof. + { revert X Y. + induction e; intros. + - apply (closed_weaken_val _ _ _ H H0). + - unfold closed, is_closed in *. + apply bool_decide_unpack in H. + apply bool_decide_pack. + set_solver. + - unfold closed in *. simpl in *. + apply andb_prop_intro. + apply andb_prop_elim in H. + destruct H. + split. + apply (IHe1 _ _ H H0). + apply (IHe2 _ _ H1 H0). } + { revert X Y. + induction v; intros. + (*- constructor.*) + - unfold closed in *. + simpl in *. + apply (closed_weaken e _ _ H). + set_solver. + (*- constructor.*) } +Qed. + +Lemma closed_subst : + ∀ Γ x e1 e2, + closed (Γ ∪ {[x]}) e1 → + closed ∅ e2 → + closed Γ (subst x e2 e1) +with closed_subst_val : + ∀ Γ x (v : val) e, + closed (Γ ∪ {[x]}) v → + closed ∅ e → + closed Γ (subst_val x e v). +Proof. + { + unfold closed in *. + intros Γ x e1 e2 H_closed1 H_closed2. + induction e1. + - simpl in *. auto. + - simpl in *. + rewrite -> decide_bool_decide. + destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. + + eapply closed_weaken. + exact H_closed2. + set_solver. + + simpl in *. + apply bool_decide_unpack in H_closed1. + apply bool_decide_pack. + set_solver. + - simpl in *. + apply andb_prop_elim in H_closed1 as [Hc1 Hc2]. + apply andb_prop_intro. split. + + exact (IHe1_1 Hc1). + + exact (IHe1_2 Hc2). + } + { + unfold closed in *. + intros Γ x v e H_closed1 H_closed2. + induction v. + (* - simpl in *. auto. *) + - simpl in *. + rewrite -> decide_bool_decide. + destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. + + rewrite -> H_eq in H_closed1. + eapply closed_weaken. + exact H_closed1. set_solver. + + simpl in *. + apply closed_subst. + eapply closed_weaken. + exact H_closed1. set_solver. + exact H_closed2. + (*- simpl in *. auto.*) + } +Qed. + +Lemma closed_app xs e1 e2: + closed xs (app e1 e2) ↔ + closed xs e1 ∧ closed xs e2. +Proof. unfold closed. simpl. by rewrite andb_True. Qed. + +Lemma closed_lambda e Γ x : closed Γ (vlambda x e) ↔ closed (Γ ∪ {[x]}) e. +Proof. split. auto. auto. Qed. + +Lemma closed_var Γ x : closed Γ (var x) ↔ x ∈ Γ. +Proof. unfold closed. simpl. by rewrite bool_decide_spec. Qed. + +Lemma base_step_preserve_closedness : + ∀ e1 e1', + base_step e1 e1' → + closed ∅ e1 → + closed ∅ e1'. +Proof. + unfold closed. + intros e1 e1' Hred H_closed. + inversion Hred. subst. simpl in *. + apply andb_prop_elim in H_closed as [Hc1 Hc2]. + exact (closed_subst ∅ x e0 e2 Hc1 Hc2). +Qed. + +Definition ectx_is_closed (X : scope) (E : ectx) := + ∀ e, closed ∅ e → closed X (fill E e). + +Lemma closed_decompose : + ∀ E e, + closed ∅ (fill E e) → + ectx_is_closed ∅ E ∧ closed ∅ e. +Proof. + unfold ectx_is_closed. + intros E. + induction E; intros e' H_closed. + - simpl in *. auto. + - simpl in *. + apply IHE in H_closed as [H_fill H_closed]. + apply closed_app in H_closed as [H_closed1 H_closed2]. + split. + + intros e'' H_closed3. + apply H_fill. + apply closed_app. auto. + + auto. + - simpl in *. + apply IHE in H_closed as [H_fill H_closed]. + apply closed_app in H_closed as [H_closed1 H_closed2]. + split. + + intros e'' H_closed3. + apply H_fill. + apply closed_app. auto. + + auto. +Qed. + +Lemma closed_compose : + ∀ E e, + ectx_is_closed ∅ E → + closed ∅ e → + closed ∅ (fill E e). +Proof. unfold ectx_is_closed. auto. Qed. + +Lemma contextual_step_preserve_closedness : + ∀ e1 e1', + contextual_step e1 e1' → + closed ∅ e1 → + closed ∅ e1'. +Proof. + unfold closed. + intros e1 e1' Hred H_closed. + inversion Hred. subst. simpl in *. clear Hred. + apply closed_decompose in H_closed as [H_closed1 H_closed2]. + apply (base_step_preserve_closedness _ _ H1) in H_closed2. + apply closed_compose; auto. +Qed. + +(** subscoped from the Erlang paper *) +Definition subst_is_closed (Γ free : scope) (sub : sub) := + Γ = dom sub ∧ + ∀ x, x ∈ Γ → + ∀ v, sub !! x = Some v → closed free (ret v). + +Lemma subst_is_closed_subseteq: ∀ (Γ1 Γ2 X : scope) (γ1 γ2: sub), + Γ1 = dom γ1 → + γ1 ⊆ γ2 → Γ1 ⊆ Γ2 → subst_is_closed Γ2 X γ2 → subst_is_closed Γ1 X γ1. +Proof. + intros * Hd Hγ HΓ Hclosed2. + destruct Hclosed2 as [Hd2 Hc2]. + split. + assumption. + intros x Hl v Hs. + rewrite (map_subseteq_spec γ1 γ2) in Hγ. + (* specialize (Hγ _ _ Hs). *) + specialise Hγ Hs. + apply (Hc2 x ltac:(set_solver) _ Hγ). +Qed. + +Lemma closed_ectx_app1 : + ∀ E e, + ectx_is_closed ∅ E → + closed ∅ e → + ectx_is_closed ∅ (ectx_app1 E e). +Proof. + intros E e Hc1 Hc2. + unfold ectx_is_closed in *. + intros e' Hc3. simpl. + apply Hc1. apply closed_app. auto. +Qed. + +Lemma closed_ectx_app2 : + ∀ (v : val) E, + closed ∅ v → + ectx_is_closed ∅ E → + ectx_is_closed ∅ (ectx_app2 v E). +Proof. + intros v E Hc1 Hc2. + unfold ectx_is_closed in *. + intros e' Hc3. simpl. + apply Hc2. apply closed_app. auto. +Qed. + +Lemma subst_is_closed_elim_closed Γ (γ:sub) x X (v:val): + subst_is_closed Γ X γ → + γ !! x = Some v → + closed X v. +Proof. + intros [Hdom Hsc] He. + assert (H := elem_of_dom_2 _ _ _ He). + (* have: elem_of_dom_2 He. *) + (* pose proof (elem_of_dom_2 _ _ _ He). *) + assert (x ∈ Γ). set_solver. + apply (Hsc x H0 v He). +Qed. + +(* if e is closed under Y, we can split the variables in Y between X and γ *) +Lemma subst_map_closed' (e : expr) (X Y : scope) (γ:sub) : + closed Y e → + (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → + closed X (subst_map γ e) +with subst_map_closed'_val (v : val) (X Y : scope) (γ:sub) : + closed Y (ret v) → + (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → + closed X (subst_map_val γ v). +Proof. + { + revert X Y γ. induction e. + { eapply subst_map_closed'_val; eauto. } + { intros * Hc H. + (* e is a variable x *) + unfold closed in Hc; simpl in Hc; apply bool_decide_unpack in Hc. + specialize (H x Hc). simpl. + destruct (γ !! x) eqn:He. + - assumption. + - unfold closed; simpl; apply bool_decide_pack. assumption. + } + { intros *. + unfold closed. simpl. + rewrite !andb_True. intros [? ?] **. + split. + by eapply IHe1. + by eapply IHe2. } + } + { revert X Y γ. induction v. + (*{ intros. assumption. }*) + { unfold closed. simpl. + intros * Hce H. + eapply subst_map_closed'. eassumption. + intros y [|]%elem_of_union. + { destruct (decide (x = y)). + { by subst; rewrite lookup_delete_eq with (m:=γ); set_solver. } + rewrite lookup_delete_ne with (m:=γ). 2: { assumption. } + eapply H in H0. + destruct lookup; last set_solver. + eapply closed_weaken; eauto with set_solver. } + { rewrite elem_of_singleton in H0. + subst. rewrite lookup_delete_eq with (m:=γ). set_solver. } + } + (*{ intros. assumption. }*) } +Qed. + +(* simple corollary of [subst_map_closed'], + where we have split Y into X and dom γ upfront *) +Lemma subst_map_closed'_2 Γ X γ e : + closed (X ∪ (dom γ)) e -> + subst_is_closed Γ X γ -> + closed X (subst_map γ e). +Proof. + intros Hcl Hsubst. + eapply subst_map_closed'; first eassumption. + intros x Hx. + destruct (γ !! x) as [e'|] eqn:Heq. + - apply (subst_is_closed_elim_closed _ _ _ _ _ Hsubst Heq). + - apply not_elem_of_dom in Heq. + set_solver. +Qed. + +(* given a value and a substitution closed under the same scope, + applying the substitution gives us a completely closed value *) +Lemma subst_map_closed'_3 (v:val) Γ γ: + closed Γ v -> + subst_is_closed Γ ∅ γ -> + closed ∅ (subst_map γ v). +Proof. + pose proof (subst_map_closed'_2 Γ ∅ γ v). + simpl in H. + intros. + apply H. 2: { assumption. } + destruct H1 as [? _]. + rewrite <- H1. + replace (∅ ∪ Γ) with Γ. assumption. + set_solver. +Qed. + +Lemma subst_map_val_closed_val_aux Γ γ (v : val) : + closed Γ v → + Γ ∩ dom γ = ∅ → + subst_map_val γ v = v +with subst_map_closed_aux Γ γ (e : expr) : + closed Γ e → + Γ ∩ dom γ = ∅ → + subst_map γ e = e. +Proof. + { intros Hc Hdom. + induction v. + - rewrite -> closed_lambda in Hc. + rewrite -> fold_unfold_subst_map_val_vlambda. + rewrite -> (subst_map_closed_aux (Γ ∪ {[x]}) (delete x γ) e Hc ltac:(set_solver)). + reflexivity. } + { intros Hc Hdom. + induction e. + - rewrite -> fold_unfold_subst_map_ret. + f_equal. by eapply subst_map_val_closed_val_aux. + - unfold closed in Hc. + simpl in Hc. + rewrite -> bool_decide_spec in Hc. + rewrite -> fold_unfold_subst_map_var. + assert (H_not_in : x ∉ dom γ) by set_solver. + rewrite -> (not_elem_of_dom γ x) in H_not_in. + setoid_rewrite -> H_not_in. + reflexivity. + - apply closed_app in Hc as [Hc1 Hc2]. + rewrite -> fold_unfold_subst_map_app. + rewrite -> (IHe1 Hc1). + rewrite -> (IHe2 Hc2). + reflexivity. } +Qed. + +Lemma subst_map_val_closed_val γ (v : val) : + closed ∅ v → + subst_map_val γ v = v. +Proof. + intros Hc. + eapply subst_map_val_closed_val_aux. + - exact Hc. + - set_solver. +Qed. + +Lemma subst_map_closed γ (e : expr) : + closed ∅ e → + subst_map γ e = e. +Proof. + intros Hc. + eapply subst_map_closed_aux. + - exact Hc. + - set_solver. +Qed. +*) + +(* +Lemma subset_is_closed_absurd x Γ γ: + x ∈ Γ → + subst_is_closed Γ ∅ γ → + γ !! x = None → + False. +Proof. + intros Hx Hs He. + pose proof (not_elem_of_dom_2 _ _ He). + destruct Hs as [? _]. + setoid_rewrite <- H0 in H. + set_solver. +Qed. +*) + +(* +Lemma subst_val_closed v X x es : + closed X (of_val v) → x ∉ X → subst_val x es v = v +with subst_closed X e x es : + closed X e → x ∉ X → subst x es e = e. +Proof. + { induction v. + (*{ reflexivity. }*) + { simpl. + case_decide. + - reflexivity. + - intros. + f_equal. + rewrite closed_lambda in H0. + apply (subst_closed _ _ _ _ H0). + set_solver. } + (*{ reflexivity. }*) } + { induction e; intros; simpl. + { f_equal. + eapply subst_val_closed. + apply H. assumption. } + { case_decide. + - subst. + unfold closed in H. simpl in H. apply bool_decide_unpack in H. + set_solver. + - reflexivity. } + { apply closed_app in H. + destruct H as (H1&H2). + specialize (IHe1 H1 H0). + specialize (IHe2 H2 H0). + f_equal. + - assumption. + - assumption. } } +Qed. +*) + +(* +(* aka contextual scoping C : Γ ~> Γ', a special case of contextual typing. + defined inductively because we need to invert it. *) +Inductive closed_ctx : scope → scope → ctx → Prop := + | cc_hole Γ : + closed_ctx Γ Γ ctx_hole + + | cc_lambda x Γ Γ' C : + closed_ctx Γ (Γ' ∪ {[x]}) C → + closed_ctx Γ Γ' (ctx_lam x C) + + | cc_app1 Γ Γ' C e : + closed_ctx Γ Γ' C → + closed Γ' e → + closed_ctx Γ Γ' (ctx_app1 C e) + + | cc_app2 Γ Γ' C v : + closed_ctx Γ Γ' C → + closed Γ' (ret v) → + closed_ctx Γ Γ' (ctx_app2 v C) + . +*) + +(* +Definition closed_ctx_sem (Γ Γ' : scope) (C:ctx) : Prop := + forall e, closed Γ e → closed Γ' (cplug C e). + +Lemma closed_ctx_sound Γ Γ' e : + closed_ctx Γ Γ' e → closed_ctx_sem Γ Γ' e. +Proof. + intros H. induction H; unfold closed_ctx_sem. + - simpl. done. + - intros e Hc. + simpl. + specialize (IHclosed_ctx e Hc). + unfold closed. simpl. + apply IHclosed_ctx. + - intros e2 Hc. + specialize (IHclosed_ctx _ Hc). + simpl. + unfold closed. simpl. + unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. + auto. + - intros e1 Hc. + specialize (IHclosed_ctx _ Hc). + simpl. + unfold closed. simpl. + unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. + auto. +Qed. +*) + +(* +Lemma G_rel_empty n : + n ⊨ G_rel ∅ ∅ ∅. +Proof. + apply G_rel_intro. + - unfold subst_is_closed. split; set_solver. + - unfold subst_is_closed. split; set_solver. + - iintros. + idestruct H. + setoid_rewrite lookup_empty in H. + discriminate H. +Qed. + +Lemma subst_map_empty (e:expr) : + subst_map ∅ e = e +with subst_map_val_empty (v:val) : + subst_map_val ∅ v = v. +Proof. + { induction e. + - simpl. f_equal. eapply subst_map_val_empty. + - simpl. + setoid_rewrite lookup_empty. + reflexivity. + - simpl. rewrite IHe1. rewrite IHe2. reflexivity. } + { induction v. + - simpl. + f_equal. + setoid_rewrite delete_empty. + apply subst_map_empty. } +Qed. + +Lemma fundamental_property_sub Γ γ n : + subst_is_closed Γ ∅ γ → + n ⊨ G_rel Γ γ γ. +Proof. + intros Hc. + apply G_rel_intro. + { exact Hc. } + { exact Hc. } + iintros x v1 v2 Heq1 Heq2. + idestruct Heq1. + idestruct Heq2. + rewrite -> Heq1 in Heq2. + injection Heq2 as ->. + destruct Hc as [-> Hc]. + assert (Hcv2 : closed ∅ v2). + { eapply Hc. + - setoid_rewrite -> elem_of_dom. + unfold is_Some. eauto. + - exact Heq1. } + assert (H_fundamental := fundamental_property_v ∅ v2 n Hcv2). + apply V_rel_o_elim in H_fundamental. + ispec H_fundamental ∅ ∅ (G_rel_empty n). + rewrite -> subst_map_val_empty in H_fundamental. + exact H_fundamental. +Qed. + +Lemma fundamental_property_ectx E n : + n ⊨ K_rel E E. +Proof. + apply K_rel_intro. + iintros v1 v2 Hv. + admit. +Admitted. +*) From 49894dbb3b24fe88152e2dc15eaa19aae3c18017 Mon Sep 17 00:00:00 2001 From: Darius Foo Date: Fri, 28 Nov 2025 12:48:16 +0800 Subject: [PATCH 02/11] Start on axiomatization --- ctx-equiv-ixfree/interface.v | 151 +++++++++++++++++++++++++++++++++++ 1 file changed, 151 insertions(+) create mode 100644 ctx-equiv-ixfree/interface.v diff --git a/ctx-equiv-ixfree/interface.v b/ctx-equiv-ixfree/interface.v new file mode 100644 index 0000000..e793461 --- /dev/null +++ b/ctx-equiv-ixfree/interface.v @@ -0,0 +1,151 @@ + +From Stdlib Require Import Utf8. +From Binding Require Import Lib Auto. +Require Import Binding.Set. + +From Stdlib Require Import Classes.Morphisms Program. + +Module Type StagedLogic. + + Parameter expr : Set → Set. + Parameter val : Set → Set. + + Parameter ret : ∀ {V:Set}, val V → expr V. + Parameter app : ∀ {V:Set}, expr V → expr V → expr V. + Parameter var : ∀ {V:Set}, V → val V. + Parameter lambda : ∀ {V:Set}, expr (inc V) → val V. + + Parameter emap : ∀ {A B:Set}, (A [→] B) → expr A → expr B. + Parameter vmap : ∀ {A B:Set}, (A [→] B) → val A → val B. + + (* equivalence *) + Parameter eqv : ∀ {V}, expr V → expr V → Prop. + Parameter Reflexive_eqv : ∀ {V:Set}, Reflexive (@eqv V). + Parameter Symmetric_eqv : ∀ {V:Set}, Symmetric (@eqv V). + Parameter Transitive_eqv : ∀ {V:Set}, Transitive (@eqv V). + + (* entailment *) + Parameter entails : ∀ {V}, expr V → expr V → Prop. + Parameter Reflexive_entails : ∀ {V:Set}, Reflexive (@entails V). + Parameter Transitive_entails : ∀ {V:Set}, Transitive (@entails V). + +(* TODO just have one expr type? *) + + (* equivalence for values *) + Parameter eqv_val : ∀ {V}, val V → val V → Prop. + Parameter Reflexive_eqv_val : ∀ {V:Set}, Reflexive (@eqv_val V). + Parameter Symmetric_eqv_val : ∀ {V:Set}, Symmetric (@eqv_val V). + Parameter Transitive_eqv_val : ∀ {V:Set}, Transitive (@eqv_val V). + + Parameter Proper_eqv_app : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> @eqv V) app. + Parameter Proper_iff_eqv : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> iff) eqv. + Parameter Proper_eqv_lambda : ∀ {V:Set}, Proper (@eqv (inc V) ==> @eqv_val V) lambda. + + Parameter Proper_impl_entails : ∀ {V:Set}, + Proper (flip (@entails V) ==> @entails V ==> impl) entails. + Parameter Proper_entails_app : ∀ {V:Set}, + Proper (@entails V ==> @entails V ==> @entails V) app. + + Section Bind. + + Context `{SetPureCore val}. + + Parameter ebind : ∀ {A B:Set}, (A [⇒] B) → expr A → expr B. + Parameter vbind : ∀ {A B:Set}, (A [⇒] B) → val A → val B. + + Context `{BindCore Set (@Sub val _) expr}. + + Parameter red_beta : ∀ {V} (e:expr (inc V)) (v:val V), + eqv + (app (ret (lambda e)) (ret v)) + (subst e v). + + End Bind. + +End StagedLogic. + +Module MakeInstances (SL: StagedLogic). + + Import SL. + + #[global] + Instance SetPureCore_value : SetPureCore val := + { set_pure := @var }. + + #[global] + Instance FunctorCore_emap : FunctorCore expr := @emap. + + #[global] + Instance FunctorCore_vmap : FunctorCore val := @vmap. + + #[global] + Instance BindCore_ebind : BindCore expr := @ebind SetPureCore_value. + + #[global] + Instance BindCore_vbind : BindCore val := @vbind SetPureCore_value. + + #[global] + Instance Reflexive_eqv : ∀ {V:Set}, Reflexive (@eqv V). + Proof. + intros. + eapply Reflexive_eqv. + Qed. + + #[global] + Instance Proper_eqv_app : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> @eqv V) app. + Proof. + intros. + eapply Proper_eqv_app. + Qed. + + #[global] + Instance Proper_iff_eqv : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> iff) eqv. + Proof. + intros. + eapply Proper_iff_eqv. + Qed. + + #[global] + Instance Proper_eqv_lambda : ∀ {V:Set}, Proper (@eqv (inc V) ==> @eqv_val V) lambda. + Proof. + intros. + eapply Proper_eqv_lambda. + Qed. + + Export SL. + +End MakeInstances. + +Module CaseStudy (SL : StagedLogic). + + Module M := MakeInstances(SL). + Import M. + + Parameter red_hello : ∀ {V} (e1 e2:expr V), eqv e1 e2. + + Example e1 (e1 e2:expr ∅): + eqv (app e2 e1) (app e1 e2). + Proof. + rewrite (red_hello e1 e2). + reflexivity. + Qed. + + Definition id_exp : expr ∅ := ret (lambda (ret (var &0))). + + Example e2 (v1:val ∅) (e2:expr ∅): + eqv (app (app id_exp (ret v1)) e2) (app (ret v1) e2). + Proof. + unfold id_exp. + (* rewrite (red_beta (ret (var &0)) v1). *) + rewrite red_beta. + simpl. + (* cannot perform this subst without a definition for bind *) + Abort. + +End CaseStudy. + +(* Module StagedLogicIxFree : StagedLogic. +End StagedLogicIxFree. *) + +(* Module M := CaseStudy(StagedLogicIxFree). +Check M. *) From ee5783ffcdff01b01dd8220fa960f82e413af735 Mon Sep 17 00:00:00 2001 From: Darius Foo Date: Fri, 28 Nov 2025 14:17:00 +0800 Subject: [PATCH 03/11] Clean up interface --- ctx-equiv-ixfree/interface.v | 121 +++++++++++------------------------ 1 file changed, 38 insertions(+), 83 deletions(-) diff --git a/ctx-equiv-ixfree/interface.v b/ctx-equiv-ixfree/interface.v index e793461..04badd9 100644 --- a/ctx-equiv-ixfree/interface.v +++ b/ctx-equiv-ixfree/interface.v @@ -1,124 +1,80 @@ -From Stdlib Require Import Utf8. -From Binding Require Import Lib Auto. -Require Import Binding.Set. - -From Stdlib Require Import Classes.Morphisms Program. +From Stdlib Require Import Utf8 Classes.Morphisms Program. Module Type StagedLogic. - Parameter expr : Set → Set. - Parameter val : Set → Set. + (* Binding *) + Parameter inc : Set → Set. + Parameter empty : Set. + Notation "∅" := empty. + Parameter v : ∀ {V}, nat → V. + Notation "& n" := (v n) (at level 5). - Parameter ret : ∀ {V:Set}, val V → expr V. + (* Expressions *) + Parameter expr : Set → Set. Parameter app : ∀ {V:Set}, expr V → expr V → expr V. - Parameter var : ∀ {V:Set}, V → val V. - Parameter lambda : ∀ {V:Set}, expr (inc V) → val V. - - Parameter emap : ∀ {A B:Set}, (A [→] B) → expr A → expr B. - Parameter vmap : ∀ {A B:Set}, (A [→] B) → val A → val B. + Parameter var : ∀ {V:Set}, V → expr V. + Parameter lambda : ∀ {V:Set}, expr (inc V) → expr V. - (* equivalence *) + (* Equivalence *) Parameter eqv : ∀ {V}, expr V → expr V → Prop. Parameter Reflexive_eqv : ∀ {V:Set}, Reflexive (@eqv V). Parameter Symmetric_eqv : ∀ {V:Set}, Symmetric (@eqv V). Parameter Transitive_eqv : ∀ {V:Set}, Transitive (@eqv V). - (* entailment *) + (* Entailment *) Parameter entails : ∀ {V}, expr V → expr V → Prop. Parameter Reflexive_entails : ∀ {V:Set}, Reflexive (@entails V). Parameter Transitive_entails : ∀ {V:Set}, Transitive (@entails V). -(* TODO just have one expr type? *) - - (* equivalence for values *) - Parameter eqv_val : ∀ {V}, val V → val V → Prop. - Parameter Reflexive_eqv_val : ∀ {V:Set}, Reflexive (@eqv_val V). - Parameter Symmetric_eqv_val : ∀ {V:Set}, Symmetric (@eqv_val V). - Parameter Transitive_eqv_val : ∀ {V:Set}, Transitive (@eqv_val V). - - Parameter Proper_eqv_app : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> @eqv V) app. + (* Proper instances *) Parameter Proper_iff_eqv : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> iff) eqv. - Parameter Proper_eqv_lambda : ∀ {V:Set}, Proper (@eqv (inc V) ==> @eqv_val V) lambda. + Parameter Proper_eqv_app : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> @eqv V) app. + Parameter Proper_eqv_lambda : ∀ {V:Set}, Proper (@eqv (inc V) ==> @eqv V) lambda. Parameter Proper_impl_entails : ∀ {V:Set}, Proper (flip (@entails V) ==> @entails V ==> impl) entails. Parameter Proper_entails_app : ∀ {V:Set}, Proper (@entails V ==> @entails V ==> @entails V) app. + Parameter Proper_entails_lambda : ∀ {V:Set}, + Proper (@entails (inc V) ==> @entails V) lambda. - Section Bind. + (* Substitution *) + Parameter subst : ∀ {V}, expr (inc V) → expr V → expr V. + (* TODO equations *) - Context `{SetPureCore val}. - - Parameter ebind : ∀ {A B:Set}, (A [⇒] B) → expr A → expr B. - Parameter vbind : ∀ {A B:Set}, (A [⇒] B) → val A → val B. - - Context `{BindCore Set (@Sub val _) expr}. - - Parameter red_beta : ∀ {V} (e:expr (inc V)) (v:val V), - eqv - (app (ret (lambda e)) (ret v)) - (subst e v). - - End Bind. + (* Lemmas *) + Parameter red_beta : ∀ {V} (e:expr (inc V)) (v:expr V), + eqv (app (lambda e) v) (subst e v). End StagedLogic. -Module MakeInstances (SL: StagedLogic). +Module WithInstances (SL: StagedLogic). Import SL. #[global] - Instance SetPureCore_value : SetPureCore val := - { set_pure := @var }. - - #[global] - Instance FunctorCore_emap : FunctorCore expr := @emap. - - #[global] - Instance FunctorCore_vmap : FunctorCore val := @vmap. - - #[global] - Instance BindCore_ebind : BindCore expr := @ebind SetPureCore_value. + Instance Reflexive_eqv {V:Set} : Reflexive (@eqv V) := Reflexive_eqv. #[global] - Instance BindCore_vbind : BindCore val := @vbind SetPureCore_value. + Instance Proper_eqv_app {V:Set} : Proper (@eqv V ==> @eqv V ==> @eqv V) app := + Proper_eqv_app. #[global] - Instance Reflexive_eqv : ∀ {V:Set}, Reflexive (@eqv V). - Proof. - intros. - eapply Reflexive_eqv. - Qed. + Instance Proper_iff_eqv {V:Set} : Proper (@eqv V ==> @eqv V ==> iff) eqv := + Proper_iff_eqv. #[global] - Instance Proper_eqv_app : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> @eqv V) app. - Proof. - intros. - eapply Proper_eqv_app. - Qed. - - #[global] - Instance Proper_iff_eqv : ∀ {V:Set}, Proper (@eqv V ==> @eqv V ==> iff) eqv. - Proof. - intros. - eapply Proper_iff_eqv. - Qed. - - #[global] - Instance Proper_eqv_lambda : ∀ {V:Set}, Proper (@eqv (inc V) ==> @eqv_val V) lambda. - Proof. - intros. - eapply Proper_eqv_lambda. - Qed. + Instance Proper_eqv_lambda {V:Set} : Proper (@eqv (inc V) ==> @eqv V) lambda := + SL.Proper_eqv_lambda. Export SL. -End MakeInstances. +End WithInstances. Module CaseStudy (SL : StagedLogic). - Module M := MakeInstances(SL). + Module M := WithInstances(SL). Import M. Parameter red_hello : ∀ {V} (e1 e2:expr V), eqv e1 e2. @@ -130,16 +86,15 @@ Module CaseStudy (SL : StagedLogic). reflexivity. Qed. - Definition id_exp : expr ∅ := ret (lambda (ret (var &0))). + Definition id_exp : expr ∅ := lambda (var &0). - Example e2 (v1:val ∅) (e2:expr ∅): - eqv (app (app id_exp (ret v1)) e2) (app (ret v1) e2). + Example e2 (e1 e2:expr ∅): + eqv (app (app id_exp e1) e2) (app (e1) e2). Proof. unfold id_exp. (* rewrite (red_beta (ret (var &0)) v1). *) rewrite red_beta. - simpl. - (* cannot perform this subst without a definition for bind *) + (* need equations for subst *) Abort. End CaseStudy. From aba8bb6cdaf84d8163f31a60049aec51cf038aba Mon Sep 17 00:00:00 2001 From: Darius Foo Date: Fri, 28 Nov 2025 15:35:45 +0800 Subject: [PATCH 04/11] Substitution is tricky --- ctx-equiv-ixfree/interface.v | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/ctx-equiv-ixfree/interface.v b/ctx-equiv-ixfree/interface.v index 04badd9..173d454 100644 --- a/ctx-equiv-ixfree/interface.v +++ b/ctx-equiv-ixfree/interface.v @@ -42,6 +42,11 @@ Module Type StagedLogic. (* Substitution *) Parameter subst : ∀ {V}, expr (inc V) → expr V → expr V. (* TODO equations *) + Parameter subst_app : ∀ {V} (e1 e2:expr (inc V)) (v:expr V), + subst (app e1 e2) v = app (subst e1 v) (subst e2 v). + + (* Parameter subst_lambda : ∀ {V V1} (e: expr (inc (inc V))) (v:expr (inc V)), + subst (lambda e) v = lambda (subst e v). *) (* Lemmas *) Parameter red_beta : ∀ {V} (e:expr (inc V)) (v:expr V), From 2957037cd1c5ec9e9428ae42a631111dca0edc98 Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Fri, 28 Nov 2025 17:52:01 +0800 Subject: [PATCH 05/11] Add soundness + completeness prove --- ctx-equiv-ixfree/examples.v | 3 +- ctx-equiv-ixfree/lang.v | 571 +++++++++++++++++++++++++++++------- 2 files changed, 467 insertions(+), 107 deletions(-) diff --git a/ctx-equiv-ixfree/examples.v b/ctx-equiv-ixfree/examples.v index 6f60ad3..92a011f 100644 --- a/ctx-equiv-ixfree/examples.v +++ b/ctx-equiv-ixfree/examples.v @@ -1,4 +1,4 @@ - +(* From CtxEquivIxFree Require Import lang. From CtxEquivIxFree Require Import propriety. From IxFree Require Import Lib Nat. @@ -538,3 +538,4 @@ Proof. reflexivity. Qed. *) +*) diff --git a/ctx-equiv-ixfree/lang.v b/ctx-equiv-ixfree/lang.v index 0d817aa..424a09a 100644 --- a/ctx-equiv-ixfree/lang.v +++ b/ctx-equiv-ixfree/lang.v @@ -780,6 +780,9 @@ Definition V_rel_o {V} (v1 v2 : val V) : IProp := Definition O_rel_o {V} (e1 e2 : expr V) : IProp := ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ O_rel (bind γ1 e1) (bind γ2 e2). +Definition K_rel_o {V} (E1 E2 : ectx V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ K_rel (bind γ1 E1) (bind γ2 E2). + (** Contractiveness and unrolling fixpoint *) Lemma L_rel_pre_contractive : contractive L_rel_pre. @@ -870,6 +873,16 @@ Proof. split; assumption. Qed. +Lemma O_rel_elim1 (e1 e2 : expr ∅) n : + n ⊨ O_rel e1 e2 → + n ⊨ L_rel e1 e2. +Proof. intros He. by apply O_rel_elim in He as []. Qed. + +Lemma O_rel_elim2 (e1 e2 : expr ∅) n : + n ⊨ O_rel e1 e2 → + n ⊨ L_rel e2 e1. +Proof. intros He. by apply O_rel_elim in He as []. Qed. + Lemma V_rel_intro (v1 v2 : val ∅) n : (n ⊨ ∀ᵢ u1 u2, ▷ V_rel u1 u2 →ᵢ @@ -981,21 +994,6 @@ Proof. iapply He. exact HE. Qed. -Lemma V_rel_elimE (v1 v2 u1 u2 : val ∅) n : - n ⊨ V_rel v1 v2 → - n ⊨ V_rel u1 u2 → - n ⊨ E_rel (app v1 u1) (app v2 u2). -Proof. - intros Hv Hu. - apply E_rel_intro. - iintros E1 E2 HE. - apply R_rel_elimO. - - apply V_rel_elimR. - + exact Hv. - + later_shift. exact Hu. - - later_shift. exact HE. -Qed. - Lemma G_rel_intro {V} (γ1 γ2 : V [⇒] ∅) n : (n ⊨ ∀ᵢ x, V_rel (γ1 x) (γ2 x)) → n ⊨ G_rel γ1 γ2. @@ -1075,8 +1073,37 @@ Proof. iapply He. exact Hγ. Qed. +Lemma K_rel_o_intro {V} (E1 E2 : ectx V) n : + n ⊨ (∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ K_rel (bind γ1 E1) (bind γ2 E2)) → + n ⊨ K_rel_o E1 E2. +Proof. auto. Qed. + +Lemma K_rel_o_elim {V} (E1 E2 : ectx V) n : + n ⊨ K_rel_o E1 E2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ K_rel (bind γ1 E1) (bind γ2 E2). +Proof. auto. Qed. + +Lemma K_rel_o_elimK {V} (E1 E2 : ectx V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ K_rel_o E1 E2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ K_rel (bind γ1 E1) (bind γ2 E2). +Proof. + intros HE Hγ. + apply K_rel_o_elim in HE. + iapply HE. exact Hγ. +Qed. + (** Compatibility lemmas *) +Lemma compat_val_closed v1 v2 n : + n ⊨ V_rel v1 v2 → + n ⊨ E_rel v1 v2. +Proof. + intros Hv. + apply E_rel_intro. iintros E1 E2 HE. + apply K_rel_elimO; assumption. +Qed. + (* aka val inclusion *) Lemma compat_val {V} (v1 v2 : val V) n : n ⊨ V_rel_o v1 v2 → @@ -1084,56 +1111,72 @@ Lemma compat_val {V} (v1 v2 : val V) n : Proof. intros Hv. apply V_rel_o_elim in Hv. - apply E_rel_o_intro. iintros γ1 γ2 Hγ. - apply E_rel_intro. iintros E1 E2 HE. - apply (K_rel_elimO E1 E2 (bind γ1 v1) (bind γ2 v2) _ HE). - iapply Hv. exact Hγ. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. + apply compat_val_closed. iapply Hv. exact Hγ. Qed. -Lemma compat_app {V} (e1 e2 e1' e2' : expr V) n : - n ⊨ E_rel_o e1 e2 → - n ⊨ E_rel_o e1' e2' → - n ⊨ E_rel_o (app e1 e1') (app e2 e2'). +Lemma compat_app_closed_val (v1 v2 u1 u2 : val ∅) n : + n ⊨ V_rel v1 v2 → + n ⊨ V_rel u1 u2 → + n ⊨ E_rel (app v1 u1) (app v2 u2). +Proof. + intros Hv Hu. + apply E_rel_intro. + iintros E1 E2 HE. + apply R_rel_elimO. + - apply V_rel_elimR. exact Hv. + later_shift. exact Hu. + - later_shift. exact HE. +Qed. + +Lemma compat_app_closed e1 e2 e1' e2' n : + n ⊨ E_rel e1 e2 → + n ⊨ E_rel e1' e2' → + n ⊨ E_rel (app e1 e1') (app e2 e2'). Proof. intros He He'. - apply E_rel_o_elim in He. - (* From He, we have contextual equivalence of e1 and e2, - in related context *) - apply E_rel_o_elim in He'. - apply E_rel_o_intro. iintros γ1 γ2 Hγ. apply E_rel_intro. iintros E1 E2 HE. term_simpl. (* The functions e1/e2 are evaluated first, so we "zap" them down using He. - To use He, we have to give two contexts s.t. if we can prove them to be related, - plugging e1/e2 into them will be in O. - We give ectx_app1 because the plugging will give us exactly the goal we need. *) - ispec He γ1 γ2 Hγ. + To use He, we have to give two contexts s.t. if we can prove them to be + related, plugging e1/e2 into them will be in O_rel. We give ectx_app1 + because the plugging will give us exactly the goal we need. *) apply E_rel_elim in He. - ispecialize He (ectx_app1 E1 (bind γ1 e1')). - ispecialize He (ectx_app1 E2 (bind γ2 e2')). - iapply He. - (* This reduces the problem to that of showing that the two app contexts are related. *) - apply K_rel_intro. - iintros v1 v2 Hv. simpl. + ispecialize He (ectx_app1 E1 e1'). + ispecialize He (ectx_app1 E2 e2'). + iapply He. clear He. + (* Now, we need to show that the two app contexts are related. *) + apply K_rel_intro. iintros v1 v2 Hv. simpl. (* Given that they are plugged with two related values, we now have to prove - that the result is in O. We use He' for a similar purpose. We give ectx_app2 - because plugging e1'/e2' into it will match the goal. *) - ispec He' γ1 γ2 Hγ. + that the result is in O_rel. We use He' for a similar purpose. We give + ectx_app2 because plugging e1'/e2' into it will match the goal. *) apply E_rel_elim in He'. ispecialize He' (ectx_app2 v1 E1). ispecialize He' (ectx_app2 v2 E2). - iapply He'. + iapply He'. clear He'. (* Now we have to prove the ectx_app2 are related. *) - apply K_rel_intro. - iintros v1' v2' Hv'. simpl. + apply K_rel_intro. iintros v1' v2' Hv'. simpl. (* Now, we have that the two values and contexts are related. - We "zap" (app v1 v1') and (app v2 v2') down using E_rel_elimO *) + We "zap" (app v1 v1') and (app v2 v2') down using E_rel_elimO *) apply E_rel_elimO. - apply V_rel_elimE; [exact Hv | exact Hv']. - (* Finally, we are left with just E1 and E2. They are related according - to our hypothesis *) + apply compat_app_closed_val; [exact Hv | exact Hv']. + (* Finally, we are left with just E1 and E2. + They are related according to our hypothesis *) exact HE. Qed. +Lemma compat_app {V} (e1 e2 e1' e2' : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o e1' e2' → + n ⊨ E_rel_o (app e1 e1') (app e2 e2'). +Proof. + intros He He'. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. + (* Use the lemma about closed app *) + apply compat_app_closed. + - apply E_rel_o_elimE. exact He. exact Hγ. + - apply E_rel_o_elimE. exact He'. exact Hγ. +Qed. + Lemma compat_var {V : Set} (x : V) n : n ⊨ V_rel_o (v_var x) (v_var x). Proof. @@ -1158,11 +1201,13 @@ Proof. exact He. Qed. -Lemma L_rel_red_r (e2 e2' : expr ∅) n : +Lemma L_rel_red_r (e1 e2 e2' : expr ∅) n : contextual_step e2 e2' → - n ⊨ (∀ᵢ e1, L_rel e1 e2' →ᵢ L_rel e1 e2). + n ⊨ L_rel e1 e2' → + n ⊨ L_rel e1 e2. Proof. - intros H_step. + intros H_step He1. + irevert e1 He1. loeb_induction IH. iintros e1 He. apply L_rel_elim in He as [He1 He2]. @@ -1186,7 +1231,7 @@ Proof. - eapply L_rel_red_l. + exact H_step. + later_shift. exact He1. - - iapply L_rel_red_r. + - eapply L_rel_red_r. + exact H_step. + exact He2. Qed. @@ -1200,7 +1245,7 @@ Proof. intros H_step He. apply O_rel_elim in He as [He1 He2]. apply O_rel_intro. - - iapply L_rel_red_r. + - eapply L_rel_red_r. + exact H_step. + exact He1. - eapply L_rel_red_l. @@ -1221,10 +1266,10 @@ Proof. apply O_rel_intro. - eapply L_rel_red_l. { exact H_step1. } later_shift. - iapply L_rel_red_r. { exact H_step2. } + eapply L_rel_red_r. { exact H_step2. } exact He1. - - iapply L_rel_red_r. { exact H_step1. } - iapply L_rel_red_l. { exact H_step2. } + - eapply L_rel_red_r. { exact H_step1. } + eapply L_rel_red_l. { exact H_step2. } later_shift. exact He2. Qed. @@ -1255,7 +1300,6 @@ Lemma compat_lambda {V} (e1 e2 : expr (inc V)) n : n ⊨ V_rel_o (v_lambda e1) (v_lambda e2). Proof. intros He. - apply E_rel_o_elim in He. apply V_rel_o_intro. iintros γ1 γ2 Hγ. apply V_rel_intro. iintros u1 u2 Hu. term_simpl. eapply R_rel_red_both. @@ -1264,7 +1308,7 @@ Proof. later_shift. simpl. unfold subst. rewrite -> bind_bind_comp'. rewrite -> bind_bind_comp'. - iapply He. + apply E_rel_o_elim in He. iapply He. apply G_rel_intro. iintros x. destruct x as [| x']. - term_simpl. exact Hu. @@ -1382,18 +1426,74 @@ Proof. unfold Symmetric, ctx_equiv. symmetry. auto. Qed. Instance Transitive_ctx_equiv {V} : Transitive (@ctx_equiv V). Proof. unfold Transitive, ctx_equiv. etransitivity; eauto. Qed. -Lemma ctx_equiv_intro {V} (e1 e2 : expr V) : +Lemma obs_equiv_intro_approx (e1 e2 : expr ∅) : + obs_approx e1 e2 → + obs_approx e2 e1 → + obs_equiv e1 e2. +Proof. + unfold obs_approx, obs_equiv. done. +Qed. + +Lemma obs_equiv_elim_approx (e1 e2 : expr ∅) : + obs_equiv e1 e2 → + obs_approx e1 e2 ∧ + obs_approx e2 e1. +Proof. + unfold obs_approx, obs_equiv. done. +Qed. + +Lemma ctx_equiv_intro_approx {V} (e1 e2 : expr V) : ctx_approx e1 e2 → ctx_approx e2 e1 → ctx_equiv e1 e2. Proof. - intros He1 He2. split. - - apply He1. - - apply He2. + unfold ctx_approx, ctx_equiv. + intros He1 He2 C. + apply obs_equiv_intro_approx; auto. +Qed. + +Lemma ctx_equiv_elim_approx {V} (e1 e2 : expr V) : + ctx_equiv e1 e2 → + ctx_approx e1 e2 ∧ + ctx_approx e2 e1. +Proof. + unfold ctx_equiv, ctx_approx. + intros He. split. + - intros C. specialize (He C). by apply obs_equiv_elim_approx in He as []. + - intros C. specialize (He C). by apply obs_equiv_elim_approx in He as []. +Qed. + +Definition ciu_approx {V} (e1 e2 : expr V) : Prop := + ∀ E γ, obs_approx (plug E (bind γ e1)) (plug E (bind γ e2)). + +Definition ciu_equiv {V} (e1 e2 : expr V) : Prop := + ∀ E γ, obs_equiv (plug E (bind γ e1)) (plug E (bind γ e2)). + +Lemma ciu_equiv_intro_approx {V} (e1 e2 : expr V) : + ciu_approx e1 e2 → + ciu_approx e2 e1 → + ciu_equiv e1 e2. +Proof. + unfold ciu_approx, ciu_equiv. + intros He1 He2 E γ. + apply obs_equiv_intro_approx; auto. +Qed. + +Lemma ciu_equiv_elim_approx {V} (e1 e2 : expr V) : + ciu_equiv e1 e2 → + ciu_approx e1 e2 ∧ + ciu_approx e2 e1. +Proof. + unfold ciu_approx, ciu_equiv. + intros He. split. + - intros E γ. specialize (He E γ). by apply obs_equiv_elim_approx in He as []. + - intros E γ. specialize (He E γ). by apply obs_equiv_elim_approx in He as []. Qed. Definition n_big_step {V} n (e : expr V) (v : val V) := nsteps contextual_step n e v. +Definition b_big_step {V} n (e : expr V) (v : val V) := bsteps contextual_step n e v. Definition n_terminates {V} n (e : expr V) := ∃ v, n_big_step n e v. +Definition b_terminates {V} n (e : expr V) := ∃ v, b_big_step n e v. Lemma big_step_iff_n_big_step {V} (e : expr V) (v : val V) : big_step e v ↔ ∃ n, n_big_step n e v. @@ -1402,6 +1502,13 @@ Proof. apply rtc_nsteps. Qed. +Lemma big_step_iff_b_big_step {V} (e : expr V) (v : val V) : + big_step e v ↔ ∃ n, b_big_step n e v. +Proof. + unfold big_step, b_big_step. + apply rtc_bsteps. +Qed. + Lemma terminates_iff_n_terminates {V} (e : expr V) : terminates e ↔ ∃ n, n_terminates n e. Proof. @@ -1415,6 +1522,19 @@ Proof. exists n. exact H_n_big_step. Qed. +Lemma terminates_iff_b_terminates {V} (e : expr V) : + terminates e ↔ ∃ n, b_terminates n e. +Proof. + unfold terminates, b_terminates. + split. + - intros [v H_big_step]. + apply big_step_iff_b_big_step in H_big_step as [n H_b_big_step]. + eauto. + - intros (n & v & H_b_big_step). + exists v. apply big_step_iff_b_big_step. + exists n. exact H_b_big_step. +Qed. + Lemma n_big_step_O_inv {V} e (v : val V) : n_big_step O e v → e = v. Proof. @@ -1422,6 +1542,13 @@ Proof. inversion_clear 1. auto. Qed. +Lemma b_big_step_O_inv {V} e (v : val V) : + b_big_step O e v → e = v. +Proof. + unfold b_big_step. + inversion_clear 1. auto. +Qed. + Lemma n_big_step_S_inv {V} n e (v : val V) : n_big_step (S n) e v → ∃ e', contextual_step e e' ∧ n_big_step n e' v. @@ -1449,32 +1576,29 @@ Proof. eauto. Qed. -Lemma ctx_equiv_elim {V} (e1 e2 : expr V) : - ctx_equiv e1 e2 → - ctx_approx e1 e2 ∧ - ctx_approx e2 e1. +Lemma b_terminates_O_inv {V} (e : expr V) : + b_terminates O e → + ∃ (v : val V), e = v. Proof. - unfold ctx_equiv, ctx_approx, obs_equiv, obs_approx. - intros He. split; apply He. + unfold b_terminates. + intros [v H_b_big_step]. + apply b_big_step_O_inv in H_b_big_step. eauto. Qed. Lemma L_rel_adequacy_n (e1 e2 : expr ∅) n : - (∀ w, w ⊨ L_rel e1 e2) → + {| nw_index := n |} ⊨ L_rel e1 e2 → n_terminates n e1 → terminates e2. Proof. revert e1. - induction n as [| n' IHn']; intros e1 He H_n_terminates. - - apply n_terminates_O_inv in H_n_terminates as [v Hv]. - specialize (He {| nw_index := O |}). - apply L_rel_elim in He as [He _]. - exact (He v Hv). - - apply n_terminates_S_inv in H_n_terminates as (e' & H_step & H_n_terminates). - apply (IHn' e'); [| exact H_n_terminates]. - intros w. specialize (He (world_lift w)). - apply L_rel_elim in He as [_ He]. + induction n as [| n' IHn']; intros e1 He He1. + - apply n_terminates_O_inv in He1 as [v He1]. + apply L_rel_elim in He as [He _]. eauto. + - apply n_terminates_S_inv in He1 as (e' & H_step & He1). + apply (IHn' e'); [| exact He1]. apply I_world_lift_later. - iapply He. iintro. exact H_step. + apply L_rel_elim in He as [_ He]. iapply He. + iintro. exact H_step. Qed. Lemma L_rel_adequacy (e1 e2 : expr ∅) : @@ -1482,59 +1606,186 @@ Lemma L_rel_adequacy (e1 e2 : expr ∅) : terminates e1 → terminates e2. Proof. - intros He H_terminates. - apply terminates_iff_n_terminates in H_terminates as (n & H_n_terminates). - exact (L_rel_adequacy_n e1 e2 n He H_n_terminates). + intros He He1. + apply terminates_iff_n_terminates in He1 as (n & He1). + eapply L_rel_adequacy_n. + - apply He. + - exact He1. Qed. Theorem O_rel_adequacy e1 e2 : - (∀ w, w ⊨ O_rel e1 e2) → obs_equiv e1 e2. + (∀ n, n ⊨ O_rel e1 e2) → obs_equiv e1 e2. Proof. intros He. split. - apply L_rel_adequacy. - intros w. specialize (He w). + intros n. specialize (He n). by apply O_rel_elim in He as []. - apply L_rel_adequacy. - intros w. specialize (He w). + intros n. specialize (He n). by apply O_rel_elim in He as []. Qed. -Lemma L_rel_value (v1 v2 : val ∅) n : +Lemma L_rel_val (v1 v2 : val ∅) n : n ⊨ L_rel v1 v2. Proof. apply L_rel_intro. - - intros _ _. - apply terminates_val. + - intros _ _. apply terminates_val. - iintros e1 He. idestruct He. by apply not_contextual_step_val in He. Qed. -Lemma O_rel_value (v1 v2 : val ∅) n : +Lemma O_rel_val (v1 v2 : val ∅) n : n ⊨ O_rel v1 v2. Proof. - apply O_rel_intro; apply L_rel_value. + apply O_rel_intro. + - apply L_rel_val. + - apply L_rel_val. +Qed. + +Lemma compat_empty_subst n : + n ⊨ G_rel (arrow_id ∅) (arrow_id ∅). +Proof. + apply G_rel_intro. + iintros x. destruct x. +Qed. + +Lemma fundamental_property_v_closed (v : val ∅) n : + n ⊨ V_rel v v. +Proof. + rewrite <- (bind_pure' v). + apply V_rel_o_elimV. + - apply fundamental_property_v. + - apply compat_empty_subst. +Qed. + +Lemma fundamental_property_e_closed (e : expr ∅) n : + n ⊨ E_rel e e. +Proof. + rewrite <- (bind_pure' e). + apply E_rel_o_elimE. + - apply fundamental_property_e. + - apply compat_empty_subst. Qed. -Lemma K_rel_ectx_hole n : +Lemma fundamental_property_g {V} (γ : V [⇒] ∅) n : + n ⊨ G_rel γ γ. +Proof. + apply G_rel_intro. + iintros x. apply fundamental_property_v_closed. +Qed. + +Lemma compat_ectx_hole_closed n : n ⊨ K_rel ectx_hole ectx_hole. Proof. apply K_rel_intro. iintros v1 v2 Hv. simpl. - apply O_rel_value. + apply O_rel_val. +Qed. + +Lemma compat_ectx_hole {V} n : + n ⊨ @K_rel_o V ectx_hole ectx_hole. +Proof. + apply K_rel_o_intro. + iintros γ1 γ2 _. term_simpl. + apply compat_ectx_hole_closed. +Qed. + +Lemma compat_ectx_app1_closed E1 E2 e1 e2 n : + n ⊨ K_rel E1 E2 → + n ⊨ E_rel e1 e2 → + n ⊨ K_rel (ectx_app1 E1 e1) (ectx_app1 E2 e2). +Proof. + intros HE He. + apply K_rel_intro. + iintros v1 v2 Hv. simpl. + apply E_rel_elimO. + - apply compat_app_closed. + apply compat_val_closed. exact Hv. + exact He. + - exact HE. Qed. -Lemma O_rel_o_compat_expr {V} (e1 e2 : expr V) n : +Lemma compat_ectx_app1 {V} (E1 E2 : ectx V) e1 e2 n : + n ⊨ K_rel_o E1 E2 → + n ⊨ E_rel_o e1 e2 → + n ⊨ K_rel_o (ectx_app1 E1 e1) (ectx_app1 E2 e2). +Proof. + intros HE He. + apply K_rel_o_intro. + iintros γ1 γ2 Hγ. term_simpl. + apply compat_ectx_app1_closed. + - apply K_rel_o_elimK. exact HE. exact Hγ. + - apply E_rel_o_elimE. exact He. exact Hγ. +Qed. + +Lemma compat_ectx_app2_closed v1 v2 E1 E2 n : + n ⊨ V_rel v1 v2 → + n ⊨ K_rel E1 E2 → + n ⊨ K_rel (ectx_app2 v1 E1) (ectx_app2 v2 E2). +Proof. + intros Hv HE. + apply K_rel_intro. + iintros u1 u2 Hu. simpl. + apply E_rel_elimO. + - apply compat_app_closed_val. exact Hv. exact Hu. + - exact HE. +Qed. + +Lemma compat_ectx_app2 {V} (v1 v2 : val V) E1 E2 n : + n ⊨ V_rel_o v1 v2 → + n ⊨ K_rel_o E1 E2 → + n ⊨ K_rel_o (ectx_app2 v1 E1) (ectx_app2 v2 E2). +Proof. + intros Hv HE. + apply K_rel_o_intro. + iintros γ1 γ2 Hγ. term_simpl. + apply compat_ectx_app2_closed. + - apply V_rel_o_elimV. exact Hv. exact Hγ. + - apply K_rel_o_elimK. exact HE. exact Hγ. +Qed. + +Lemma fundamental_property_k {V} (E : ectx V) n : + n ⊨ K_rel_o E E. +Proof. + induction E. + - apply compat_ectx_hole. + - apply compat_ectx_app1. + + exact IHE. + + apply fundamental_property_e. + - apply compat_ectx_app2. + + apply fundamental_property_v. + + exact IHE. +Qed. + +Lemma fundamental_property_k_closed E n : + n ⊨ K_rel E E. +Proof. + rewrite <- (bind_pure' E). + apply K_rel_o_elimK. + - apply fundamental_property_k. + - apply compat_empty_subst. +Qed. + +Lemma compat_expr_o_closed (e1 e2 : expr ∅) n : + n ⊨ E_rel e1 e2 → + n ⊨ O_rel e1 e2. +Proof. + intros He. + rewrite <- (fold_unfold_plug_ectx_hole e1). + rewrite <- (fold_unfold_plug_ectx_hole e2). + apply E_rel_elimO. + - exact He. + - apply compat_ectx_hole_closed. +Qed. + +Lemma compat_expr_o {V} (e1 e2 : expr V) n : n ⊨ E_rel_o e1 e2 → n ⊨ O_rel_o e1 e2. Proof. - intro He. - apply O_rel_o_intro. - iintros γ1 γ2 Hγ. - apply E_rel_o_elim in He. ispec He γ1 γ2 Hγ. - rewrite <- (fold_unfold_plug_ectx_hole (bind γ1 e1)). - rewrite <- (fold_unfold_plug_ectx_hole (bind γ2 e2)). - apply (E_rel_elimO _ _ _ _ _ He). - exact (K_rel_ectx_hole n). + intros He. + apply O_rel_o_intro. iintros γ1 γ2 Hγ. + apply compat_expr_o_closed. + apply E_rel_o_elimE; assumption. Qed. Lemma precongruence {A B} (e1 e2 : expr A) (C : ctxr A B) n : @@ -1555,19 +1806,127 @@ Proof. + apply IHC. exact He. Qed. +Theorem E_rel_o_soundness' {V} (e1 e2 : expr V) : + (∀ n, n ⊨ E_rel_o e1 e2) → + ciu_equiv e1 e2. +Proof. + unfold ciu_equiv. + intros He E γ. + apply O_rel_adequacy. intros n. + apply E_rel_elimO. + - specialize (He n). + apply E_rel_o_elimE. + + exact He. + + apply fundamental_property_g. + - apply fundamental_property_k_closed. +Qed. + +Lemma L_rel_obs_approx_trans e1 e2 e3 n : + n ⊨ L_rel e1 e2 → + obs_approx e2 e3 → + n ⊨ L_rel e1 e3. +Proof. + unfold obs_approx. + intros He1 He2. + irevert e1 He1. loeb_induction IH. + iintros e1 He1. + apply L_rel_elim in He1 as [He1_val He1_step]. + apply L_rel_intro. + - intros v Hv. exact (He2 (He1_val v Hv)). + - iintros e1' H_step. + ispec He1_step e1' H_step. + later_shift. iapply IH. + exact He1_step. +Qed. + +Lemma O_rel_obs_equiv_trans e1 e2 e3 n : + n ⊨ O_rel e1 e2 → + obs_equiv e2 e3 → + n ⊨ O_rel e1 e3. +Proof. + intros He1 He2. + apply O_rel_elim in He1 as [He1_l He1_r]. + apply obs_equiv_elim_approx in He2 as [He2_l He2_r]. + apply O_rel_intro. + - by eapply L_rel_obs_approx_trans. + - (* This direction is impossible to prove: + If e3 terminates in n steps, we need to prove that e1 MUST + terminate. Using He2_r, we can conclude that e2 terminates + in some m steps. However, if m > n, then He1_r does not imply + that e1 terminates. *) +Abort. + +Theorem L_rel_completeness (e1 e2 : expr ∅) n : + obs_approx e1 e2 → + n ⊨ L_rel e1 e2. +Proof. + intros He. irevert e1 He. + loeb_induction IH. + iintros e1 He. + apply L_rel_intro. + - intros v ->. apply He. apply terminates_val. + - iintros e1' H_step. idestruct H_step. + later_shift. + assert (He1' : obs_approx e1' e2). + { unfold obs_approx in *. + intros He1'. apply He. + by eapply contextual_step_terminates. } + ispec IH e1' He1'. + exact IH. +Qed. + +Theorem O_rel_completeness (e1 e2 : expr ∅) n : + obs_equiv e1 e2 → + n ⊨ O_rel e1 e2. +Proof. + intros He. + apply obs_equiv_elim_approx in He as [He1 He2]. + apply O_rel_intro. + - by apply L_rel_completeness. + - by apply L_rel_completeness. +Qed. + +Theorem E_rel_o_completeness' {V} (e1 e2 : expr V) n : + ciu_equiv e1 e2 → + n ⊨ E_rel_o e1 e2. +Proof. + intros He. + apply ciu_equiv_elim_approx in He as [He1 He2]. + unfold ciu_approx in *. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + apply E_rel_intro. iintros E1 E2 HE. + apply O_rel_intro. + - assert (He1' : n ⊨ L_rel (plug E1 (bind γ1 e1)) (plug E2 (bind γ2 e1))). + { apply O_rel_elim1. + apply E_rel_elimO. + apply E_rel_o_elimE. + apply fundamental_property_e. + exact Hγ. exact HE. } + specialize (He1 E2 γ2). + exact (L_rel_obs_approx_trans _ _ _ _ He1' He1). + - assert (He2' : n ⊨ L_rel (plug E2 (bind γ2 e2)) (plug E1 (bind γ1 e2))). + { apply O_rel_elim2. + apply E_rel_elimO. + apply E_rel_o_elimE. + apply fundamental_property_e. + exact Hγ. exact HE. } + specialize (He2 E1 γ1). + exact (L_rel_obs_approx_trans _ _ _ _ He2' He2). +Qed. + Theorem E_rel_o_soundness {V} (e1 e2 : expr V) : (∀ n, n ⊨ E_rel_o e1 e2) → ctx_equiv e1 e2. Proof. unfold ctx_equiv. intros He C. - apply O_rel_adequacy. intros w. + apply O_rel_adequacy. intros n. + specialize (He n). + apply (precongruence _ _ C) in He. + apply compat_expr_o in He. rewrite <- (bind_pure' (cplug C e1)). rewrite <- (bind_pure' (cplug C e2)). - specialize (He w). - apply (precongruence _ _ C) in He. - apply O_rel_o_compat_expr in He. - apply (O_rel_o_elimO _ _ _ _ _ He). + apply O_rel_o_elimO. exact He. apply G_rel_intro. iintros x. destruct x. Qed. From 44bbfef9154c92a8960caa8e5b369cead0174662 Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Mon, 1 Dec 2025 13:39:33 +0800 Subject: [PATCH 06/11] Prove completeness by cheating: add bind and fmap as constructor in ctxr The reason for this "cheat" is that, in the proof of completeness, given any arbitrary V, we may not be able to construct ctxr V \empty to instantiate the hypothesis. We need the two "cheating" constructor to get around this issue. --- ctx-equiv-ixfree/lang.v | 92 +++++++++++++++++++++++++++++++++++++---- 1 file changed, 84 insertions(+), 8 deletions(-) diff --git a/ctx-equiv-ixfree/lang.v b/ctx-equiv-ixfree/lang.v index 424a09a..d13bdaa 100644 --- a/ctx-equiv-ixfree/lang.v +++ b/ctx-equiv-ixfree/lang.v @@ -146,9 +146,9 @@ Qed. (** Outside-in evaluation contexts, similar to a normal list *) Inductive rctx (V : Set) := - | rctx_hole - | rctx_app1 (R : rctx V) (e : expr V) - | rctx_app2 (v : val V) (R : rctx V). +| rctx_hole +| rctx_app1 (R : rctx V) (e : expr V) +| rctx_app2 (v : val V) (R : rctx V). Arguments rctx_hole {V}. Arguments rctx_app1 {V} R e. @@ -319,11 +319,11 @@ Proof. exact (rctx_ectx_bijection_aux ectx_hole R). Qed. -Lemma plug_rplug_equiv {V} (E : ectx V) (e : expr V) : - plug E e = rplug (ectx_to_rctx E) e. +Lemma ectx_to_rctx_correct {V} (E : ectx V) (e : expr V) : + rplug (ectx_to_rctx E) e = plug E e. Proof. unfold ectx_to_rctx. - rewrite -> ectx_comp_rctx2_correct. simpl. reflexivity. + exact (ectx_comp_rctx2_correct E rctx_hole e). Qed. Lemma plug_eq_val_inv {V} E e (v : val V) : @@ -686,8 +686,8 @@ Lemma unique_decomposition {V} (E1 E2 : ectx V) e1 e2 : E1 = E2 ∧ e1 = e2. Proof. intros He1 He2 Heq. - rewrite -> plug_rplug_equiv in Heq. - rewrite -> plug_rplug_equiv in Heq. + rewrite <- ectx_to_rctx_correct in Heq. + rewrite <- ectx_to_rctx_correct in Heq. destruct (unique_rdecomposition _ _ _ _ He1 He2 Heq) as [Heq1 Heq2]. split. - rewrite <- (ectx_rctx_bijection E1). @@ -1332,6 +1332,8 @@ Qed. (** General program contexts *) Inductive ctx : Set → Type := | ctx_hole : ctx ∅ +| ctx_fmap {A B} : (A [→] B) → ctx B → ctx A +| ctx_bind {A B} : (A [⇒] B) → ctx B → ctx A | ctx_lam {V} : ctx V → ctx (inc V) | ctx_app1 {V} : ctx V → expr V → ctx V | ctx_app2 {V} : expr V → ctx V → ctx V. @@ -1340,6 +1342,8 @@ Inductive ctx : Set → Type := Fixpoint ciplug {V} (C : ctx V) : expr V → expr ∅ := match C with | ctx_hole => λ e, e + | ctx_fmap f C => λ e, ciplug C (fmap f e) + | ctx_bind f C => λ e, ciplug C (bind f e) | ctx_lam C => λ e, ciplug C (v_lambda e) | ctx_app1 C e2 => λ e, ciplug C (app e e2) | ctx_app2 e1 C => λ e, ciplug C (app e1 e) @@ -1347,6 +1351,8 @@ Fixpoint ciplug {V} (C : ctx V) : expr V → expr ∅ := Inductive ctxr : Set → Set → Type := | ctxr_hole {A} : ctxr A A +| ctxr_fmap {A B C} : (A [→] B) → ctxr B C → ctxr A C +| ctxr_bind {A B C} : (A [⇒] B) → ctxr B C → ctxr A C | ctxr_lam {A B} : ctxr A (inc B) → ctxr A B | ctxr_app1 {A B} : ctxr A B → expr B → ctxr A B | ctxr_app2 {A B} : expr B → ctxr A B → ctxr A B. @@ -1355,6 +1361,8 @@ Inductive ctxr : Set → Set → Type := Fixpoint crplug {A B} (C : ctxr A B) : expr A → expr B := match C with | ctxr_hole => λ e, e + | ctxr_fmap f C => λ e, crplug C (fmap f e) + | ctxr_bind f C => λ e, crplug C (bind f e) | ctxr_lam C => λ e, v_lambda (crplug C e) | ctxr_app1 C e2 => λ e, app (crplug C e) e2 | ctxr_app2 e1 C => λ e, app e1 (crplug C e) @@ -1788,6 +1796,36 @@ Proof. apply E_rel_o_elimE; assumption. Qed. +Lemma compat_fmap {A B} (f : A [→] B) e1 e2 n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (fmap f e1) (fmap f e2). +Proof. + intros He. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + rewrite -> (map_to_bind f). + rewrite -> (map_to_bind f). + rewrite -> bind_bind_comp'. + rewrite -> bind_bind_comp'. + apply E_rel_o_elimE. exact He. + apply G_rel_intro. iintros x. term_simpl. + apply G_rel_elimV. exact Hγ. +Qed. + +Lemma compat_bind {A B} (f : A [⇒] B) e1 e2 n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (bind f e1) (bind f e2). +Proof. + intros He. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + rewrite -> bind_bind_comp'. + rewrite -> bind_bind_comp'. + apply E_rel_o_elimE. exact He. + apply G_rel_intro. iintros x. term_simpl. + apply V_rel_o_elimV. + - apply fundamental_property_v. + - exact Hγ. +Qed. + Lemma precongruence {A B} (e1 e2 : expr A) (C : ctxr A B) n : n ⊨ E_rel_o e1 e2 → n ⊨ E_rel_o (cplug C e1) (cplug C e2). @@ -1795,6 +1833,8 @@ Proof. intros He. induction C; simpl. - exact He. + - apply IHC. apply compat_fmap. exact He. + - apply IHC. apply compat_bind. exact He. - apply compat_val. apply compat_lambda. apply IHC. exact He. @@ -1930,6 +1970,42 @@ Proof. apply G_rel_intro. iintros x. destruct x. Qed. +(* +Theorem ciu_equiv_soundness {V} (e1 e2 : expr V) : + ciu_equiv e1 e2 → + ctx_equiv e1 e2. +Proof. +Abort. +*) + +Fixpoint rctx_to_ctxr {V} (R : rctx V) : ctxr V V := + match R with + | rctx_hole => ctxr_hole + | rctx_app1 R e => ctxr_app1 (rctx_to_ctxr R) e + | rctx_app2 v R => ctxr_app2 v (rctx_to_ctxr R) + end. + +Lemma rctx_to_ctxr_correct {V} (R : rctx V) (e : expr V) : + crplug (rctx_to_ctxr R) e = rplug R e. +Proof. + induction R. + - simpl. reflexivity. + - simpl. rewrite -> IHR. reflexivity. + - simpl. rewrite -> IHR. reflexivity. +Qed. + +Theorem ciu_equiv_completeness {V} (e1 e2 : expr V) : + ctx_equiv e1 e2 → + ciu_equiv e1 e2. +Proof. + unfold ctx_equiv, ciu_equiv. + intros He E γ. + rewrite <-2 ectx_to_rctx_correct. + rewrite <-2 rctx_to_ctxr_correct. + specialize (He (ctxr_bind γ (rctx_to_ctxr (ectx_to_rctx E)))). + simpl in He. exact He. +Qed. + (* (* composition of subst and subst_map, where the variable to be substituted doesn't appear in the substitution *) From 4993de12882d018b7d56f746bf9dcc96ea871919 Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Wed, 3 Dec 2025 17:20:44 +0800 Subject: [PATCH 07/11] Prove fundamental_properties, and soundness of logrel wrt ctx_equiv The main change is to add meta-context, and update the logrels to accommodate the addition of meta-context. --- ctx-equiv-ixfree/lang_shift_reset.v | 3074 +++++++++++++++++++++++++++ ctx-equiv-ixfree/propriety.v | 3 +- 2 files changed, 3076 insertions(+), 1 deletion(-) create mode 100644 ctx-equiv-ixfree/lang_shift_reset.v diff --git a/ctx-equiv-ixfree/lang_shift_reset.v b/ctx-equiv-ixfree/lang_shift_reset.v new file mode 100644 index 0000000..e38c179 --- /dev/null +++ b/ctx-equiv-ixfree/lang_shift_reset.v @@ -0,0 +1,3074 @@ +From stdpp Require Export relations. +(** Note: both stdpp and Binding define `fmap`, but with different + signature. What we want in this file is Binding's `fmap`, thus + we Require Import Binding after stdpp *) +From Stdlib Require Import Utf8. +From Binding Require Import Lib Auto. +Require Import Binding.Set. +From IxFree Require Import Lib Nat. +From CtxEquivIxFree Require Import ixfree_tactics. +From CtxEquivIxFree Require Import tactics. + +Local Close Scope stdpp_scope. + +Inductive expr (V : Set) := +| e_val (v : val V) +| e_app (e1 e2 : expr V) +| e_shift (e : expr (inc V)) +| e_reset (e : expr V) + +with val (V : Set) := +| v_var (x : V) +| v_lambda (e : expr (inc V)) +| v_cont (E : ectx V) + +(** Inside-out contexts, similar to a "reversed" list *) +with ectx (V : Set) := +| ectx_hole +| ectx_app1 (E : ectx V) (e : expr V) +| ectx_app2 (v : val V) (E : ectx V). + +Inductive rctx (V : Set) := +| rctx_hole +| rctx_app1 (R : rctx V) (e : expr V) +| rctx_app2 (v : val V) (R : rctx V). + +Arguments e_val {V} _. +Arguments e_app {V} _ _. +Arguments e_shift {V} _. +Arguments e_reset {V} _. + +Arguments v_var {V} _. +Arguments v_lambda {V} _. +Arguments v_cont {V} _. + +Arguments ectx_hole {V}. +Arguments ectx_app1 {V} _ _. +Arguments ectx_app2 {V} _ _. + +Arguments rctx_hole {V}. +Arguments rctx_app1 {V} _ _. +Arguments rctx_app2 {V} _ _. + +Coercion e_val : val >-> expr. + +#[global] +Instance SetPureCore_value : SetPureCore val := + { set_pure := @v_var }. + +Fixpoint emap {A B} (f : A [→] B) (e : expr A) : expr B := + match e with + | e_val v => e_val (vmap f v) + | e_app e1 e2 => e_app (emap f e1) (emap f e2) + | e_shift e => e_shift (emap (lift f) e) + | e_reset e => e_reset (emap f e) + end + +with vmap {A B} (f : A [→] B) (v : val A) : val B := + match v with + | v_var x => v_var (f x) + | v_lambda e => v_lambda (emap (lift f) e) + | v_cont E => v_cont (ectx_map f E) + end + +with ectx_map {A B} (f : A [→] B) (E : ectx A) : ectx B := + match E with + | ectx_hole => ectx_hole + | ectx_app1 E e => ectx_app1 (ectx_map f E) (emap f e) + | ectx_app2 v E => ectx_app2 (vmap f v) (ectx_map f E) + end. + +Fixpoint rctx_map {A B} (f : A [→] B) (R : rctx A) : rctx B := + match R with + | rctx_hole => rctx_hole + | rctx_app1 R e => rctx_app1 (rctx_map f R) (emap f e) + | rctx_app2 v R => rctx_app2 (vmap f v) (rctx_map f R) + end. + +#[global] Instance FunctorCore_emap : FunctorCore expr := @emap. +#[global] Instance FunctorCore_vmap : FunctorCore val := @vmap. +#[global] Instance FunctorCore_ectx_map : FunctorCore ectx := @ectx_map. +#[global] Instance FunctorCore_rctx_map : FunctorCore rctx := @rctx_map. + +Fixpoint ebind {A B} (f : A [⇒] B) (e : expr A) : expr B := + match e with + | e_val v => e_val (vbind f v) + | e_app e1 e2 => e_app (ebind f e1) (ebind f e2) + | e_shift e => e_shift (ebind (lift f) e) + | e_reset e => e_reset (ebind f e) + end + +with vbind {A B} (f : A [⇒] B) (v : val A) : val B := + match v with + | v_var x => f x + | v_lambda e => v_lambda (ebind (lift f) e) + | v_cont E => v_cont (ectx_bind f E) + end + +with ectx_bind {A B} (f : A [⇒] B) (E : ectx A) : ectx B := + match E with + | ectx_hole => ectx_hole + | ectx_app1 E e => ectx_app1 (ectx_bind f E) (ebind f e) + | ectx_app2 v E => ectx_app2 (vbind f v) (ectx_bind f E) + end. + +Fixpoint rctx_bind {A B} (f : A [⇒] B) (R : rctx A) : rctx B := + match R with + | rctx_hole => rctx_hole + | rctx_app1 R e => rctx_app1 (rctx_bind f R) (ebind f e) + | rctx_app2 v R => rctx_app2 (vbind f v) (rctx_bind f R) + end. + +#[global] Instance BindCore_ebind : BindCore expr := @ebind. +#[global] Instance BindCore_vbind : BindCore val := @vbind. +#[global] Instance BindCore_rctx_bind : BindCore rctx := @rctx_bind. +#[global] Instance BindCore_ectx_bind : BindCore ectx := @ectx_bind. + +(* similar to foldr of a "reversed" list (foldl of a normal list) *) +Fixpoint plug {V} (E : ectx V) (e' : expr V) : expr V := + match E with + | ectx_hole => e' + | ectx_app1 E e => plug E (e_app e' e) + | ectx_app2 v E => plug E (e_app v e') + end. + +Lemma fold_unfold_plug_ectx_hole {V} (e : expr V) : plug ectx_hole e = e. +Proof. auto. Qed. + +(* similar to "prepend" of a "reversed" list ("append" of a normal list) *) +Fixpoint ectx_comp {V} (E1 E2 : ectx V) : ectx V := + match E2 with + | ectx_hole => E1 + | ectx_app1 E2 e => ectx_app1 (ectx_comp E1 E2) e + | ectx_app2 v E2 => ectx_app2 v (ectx_comp E1 E2) + end. + +Lemma ectx_comp_assoc {V} (E1 E2 E3 : ectx V) : + ectx_comp E1 (ectx_comp E2 E3) = ectx_comp (ectx_comp E1 E2) E3. +Proof. + induction E3; simpl. + - reflexivity. + - rewrite -> IHE3. reflexivity. + - rewrite -> IHE3. reflexivity. +Qed. + +Lemma ectx_comp_correct {V} (E1 E2 : ectx V) (e : expr V) : + plug (ectx_comp E1 E2) e = plug E1 (plug E2 e). +Proof. + revert e. + induction E2; intros e'. + - simpl. reflexivity. + - simpl. rewrite -> IHE2. reflexivity. + - simpl. rewrite -> IHE2. reflexivity. +Qed. + +(* similar to foldr of a normal list *) +Fixpoint rplug {V} (R : rctx V) (e' : expr V) : expr V := + match R with + | rctx_hole => e' + | rctx_app1 R e => e_app (rplug R e') e + | rctx_app2 v R => e_app v (rplug R e') + end. + +(* similar to append of a normal list *) +Fixpoint rctx_comp {V} (R1 R2 : rctx V) : rctx V := + match R1 with + | rctx_hole => R2 + | rctx_app1 R1 e => rctx_app1 (rctx_comp R1 R2) e + | rctx_app2 v R1 => rctx_app2 v (rctx_comp R1 R2) + end. + +Lemma rctx_comp_assoc {V} (R1 R2 R3 : rctx V) : + rctx_comp (rctx_comp R1 R2) R3 = rctx_comp R1 (rctx_comp R2 R3). +Proof. + induction R1. + - simpl. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. +Qed. + +Lemma rctx_comp_correct {V} (R1 R2 : rctx V) (e : expr V) : + rplug (rctx_comp R1 R2) e = rplug R1 (rplug R2 e). +Proof. + induction R1. + - simpl. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. + - simpl. rewrite -> IHR1. reflexivity. +Qed. + +(** Equivalence of ectx and rctx *) + +(* similar to reverse_prepend : reverse R, and then prepend E to it *) +Fixpoint ectx_comp_rctx1 {V} (E : ectx V) (R : rctx V) : ectx V := + match R with + | rctx_hole => E + | rctx_app1 R e => ectx_comp_rctx1 (ectx_app1 E e) R + | rctx_app2 v R => ectx_comp_rctx1 (ectx_app2 v E) R + end. + +(* similar to reverse *) +Definition rctx_to_ectx {V} : rctx V -> ectx V := + ectx_comp_rctx1 ectx_hole. + +Lemma ectx_comp_rctx1_correct {V} (E : ectx V) (R : rctx V) (e : expr V) : + plug (ectx_comp_rctx1 E R) e = plug E (rplug R e). +Proof. + revert E. + induction R; intros E. + - simpl. reflexivity. + - simpl. rewrite -> IHR. simpl. reflexivity. + - simpl. rewrite -> IHR. simpl. reflexivity. +Qed. + +(* similar to reverse_append : reverse E, and then append to R *) +Fixpoint ectx_comp_rctx2 {V} (E : ectx V) (R : rctx V) : rctx V := + match E with + | ectx_hole => R + | ectx_app1 E e => ectx_comp_rctx2 E (rctx_app1 R e) + | ectx_app2 v E => ectx_comp_rctx2 E (rctx_app2 v R) + end. + +Definition ectx_to_rctx {V} (E : ectx V) : rctx V := + ectx_comp_rctx2 E rctx_hole. + +Lemma ectx_comp_rctx2_correct {V} (E : ectx V) (R : rctx V) (e : expr V) : + rplug (ectx_comp_rctx2 E R) e = plug E (rplug R e). +Proof. + revert R. + induction E; intros R. + - simpl. reflexivity. + - simpl. rewrite -> IHE. simpl. reflexivity. + - simpl. rewrite -> IHE. simpl. reflexivity. +Qed. + +Lemma ectx_comp_rctx1_reset {V} (E : ectx V) (R : rctx V) : + ectx_comp_rctx1 E R = ectx_comp E (ectx_comp_rctx1 ectx_hole R). +Proof. + revert E. + induction R; intros E; simpl. + - reflexivity. + - rewrite -> (IHR (ectx_app1 E e)). + rewrite -> (IHR (ectx_app1 ectx_hole e)). + rewrite -> ectx_comp_assoc. simpl. + reflexivity. + - rewrite -> (IHR (ectx_app2 v E)). + rewrite -> (IHR (ectx_app2 v ectx_hole)). + rewrite -> ectx_comp_assoc. simpl. + reflexivity. +Qed. + +Lemma ectx_comp_rctx2_reset {V} (E : ectx V) (R : rctx V) : + ectx_comp_rctx2 E R = rctx_comp (ectx_comp_rctx2 E rctx_hole) R. +Proof. + revert R. + induction E; intros R; simpl. + - reflexivity. + - rewrite -> (IHE (rctx_app1 R e)). + rewrite -> (IHE (rctx_app1 rctx_hole e)). + rewrite -> rctx_comp_assoc. simpl. + reflexivity. + - rewrite -> (IHE (rctx_app2 v R)). + rewrite -> (IHE (rctx_app2 v rctx_hole)). + rewrite -> rctx_comp_assoc. simpl. + reflexivity. +Qed. + +Lemma ectx_rctx_bijection_aux {V} (E : ectx V) (R : rctx V) : + ectx_comp_rctx1 ectx_hole (ectx_comp_rctx2 E R) = ectx_comp_rctx1 E R. +Proof. + revert R. + induction E; intros R. + - simpl. reflexivity. + - simpl. rewrite -> (IHE (rctx_app1 R e)). simpl. reflexivity. + - simpl. rewrite -> (IHE (rctx_app2 v R)). simpl. reflexivity. +Qed. + +Lemma ectx_rctx_bijection {V} (E : ectx V) : + rctx_to_ectx (ectx_to_rctx E) = E. +Proof. + unfold rctx_to_ectx, ectx_to_rctx. + exact (ectx_rctx_bijection_aux E rctx_hole). +Qed. + +Lemma rctx_ectx_bijection_aux {V} (E : ectx V) (R : rctx V) : + ectx_comp_rctx2 (ectx_comp_rctx1 E R) rctx_hole = ectx_comp_rctx2 E R. +Proof. + revert E. + induction R; intros E. + - simpl. reflexivity. + - simpl. rewrite -> (IHR (ectx_app1 E e)). simpl. reflexivity. + - simpl. rewrite -> (IHR (ectx_app2 v E)). simpl. reflexivity. +Qed. + +Lemma rctx_ectx_bijection {V} (R : rctx V) : + ectx_to_rctx (rctx_to_ectx R) = R. +Proof. + unfold ectx_to_rctx, rctx_to_ectx. + exact (rctx_ectx_bijection_aux ectx_hole R). +Qed. + +Lemma ectx_to_rctx_correct {V} (E : ectx V) (e : expr V) : + rplug (ectx_to_rctx E) e = plug E e. +Proof. + unfold ectx_to_rctx. + exact (ectx_comp_rctx2_correct E rctx_hole e). +Qed. + +Lemma plug_eq_val_inv {V} E e (v : val V) : + plug E e = v → + E = ectx_hole ∧ e = v. +Proof. + revert e. + induction E; intros e' H_eq; simpl in *. + - auto. + - apply IHE in H_eq as (_ & H_absurd). discriminate. + - apply IHE in H_eq as (_ & H_absurd). discriminate. +Qed. + +Lemma rplug_eq_val_inv {V} R e (v : val V) : + rplug R e = v → + R = rctx_hole ∧ e = v. +Proof. + intros H_eq. + destruct R; simpl in *. + - auto. + - discriminate. + - discriminate. +Qed. + +(** Properties of syntax *) + +Lemma fmap_rplug {A B} (f : A [→] B) (R : rctx A) (e : expr A) : + fmap f (rplug R e) = rplug (fmap f R) (fmap f e). +Proof. + induction R. + - term_simpl. reflexivity. + - term_simpl. rewrite -> IHR. term_simpl. reflexivity. + - term_simpl. rewrite -> IHR. term_simpl. reflexivity. +Qed. + +Lemma bind_rplug {A B} (f : A [⇒] B) (R : rctx A) (e : expr A) : + bind f (rplug R e) = rplug (bind f R) (bind f e). +Proof. + induction R. + - term_simpl. reflexivity. + - term_simpl. rewrite -> IHR. term_simpl. reflexivity. + - term_simpl. rewrite -> IHR. term_simpl. reflexivity. +Qed. + +Lemma subst_rplug {V} (R : rctx (inc V)) e v : + subst (rplug R e) v = rplug (subst R v) (subst e v). +Proof. unfold subst. apply bind_rplug. Qed. + +Lemma fmap_plug {A B} (f : A [→] B) (E : ectx A) (e : expr A) : + fmap f (plug E e) = plug (fmap f E) (fmap f e). +Proof. + revert e. + induction E; intros e'. + - term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. +Qed. + +Lemma bind_plug {A B} (f : A [⇒] B) (E : ectx A) (e : expr A) : + bind f (plug E e) = plug (bind f E) (bind f e). +Proof. + revert e. + induction E; intros e'. + - term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. +Qed. + +Lemma subst_plug {V} (E : ectx (inc V)) e v : + subst (plug E e) v = plug (subst E v) (subst e v). +Proof. unfold subst. apply bind_plug. Qed. + +#[global] Hint Rewrite @fmap_rplug : term_simpl. +#[global] Hint Rewrite @bind_rplug : term_simpl. +#[global] Hint Rewrite @subst_rplug : term_simpl. + +#[global] Hint Rewrite @fmap_plug : term_simpl. +#[global] Hint Rewrite @bind_plug : term_simpl. +#[global] Hint Rewrite @subst_plug : term_simpl. + +#[global] +Instance SetPure_val : SetPure val. +Proof. + split. + - simpl. unfold SetPureCore_value. term_simpl. reflexivity. + - simpl. unfold SetPureCore_value. term_simpl. reflexivity. +Qed. + +(** Functor instances *) + +Fixpoint emap_id {A} (f : A [→] A) (e : expr A) : + equal f (arrow_id A) → fmap f e = e +with vmap_id {A} (f : A [→] A) (v : val A) : + equal f (arrow_id A) → fmap f v = v +with ectx_map_id {A} (f : A [→] A) (E : ectx A) : + equal f (arrow_id A) → fmap f E = E. +Proof. + - auto_map_id. + - auto_map_id. + - auto_map_id. +Qed. + +Fixpoint emap_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (e : expr A) : + equal (arrow_comp f g) h → fmap f (fmap g e) = fmap h e +with vmap_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (v : val A) : + equal (arrow_comp f g) h → fmap f (fmap g v) = fmap h v +with ectx_map_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (E : ectx A) : + equal (arrow_comp f g) h → fmap f (fmap g E) = fmap h E. +Proof. + - auto_map_comp. + - auto_map_comp. + - auto_map_comp. +Qed. + +#[global] Instance Functor_expr : Functor expr. +Proof. constructor. exact @emap_id. exact @emap_comp. Qed. + +#[global] Instance Functor_val : Functor val. +Proof. constructor. exact @vmap_id. exact @vmap_comp. Qed. + +#[global] Instance Functor_ectx : Functor ectx. +Proof. constructor. exact @ectx_map_id. exact @ectx_map_comp. Qed. + +Fixpoint rctx_map_id {A} (f : A [→] A) (R : rctx A) : + equal f (arrow_id A) → fmap f R = R. +Proof. auto_map_id. Qed. + +Fixpoint rctx_map_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (R : rctx A) : + equal (arrow_comp f g) h → fmap f (fmap g R) = fmap h R. +Proof. auto_map_comp. Qed. + +#[global] +Instance Functor_rctx : Functor rctx. +Proof. constructor. exact @rctx_map_id. exact @rctx_map_comp. Qed. + +(** Bind-Map_Pure instances *) + +Fixpoint ebind_map_pure {A B} (f : A [→] B) g (e : expr A) : + equal (subst_of_arr f) g → fmap f e = bind g e +with vbind_map_pure {A B} (f : A [→] B) g (v : val A) : + equal (subst_of_arr f) g → fmap f v = bind g v +with ectx_bind_map_pure {A B} (f : A [→] B) g (E : ectx A) : + equal (subst_of_arr f) g → fmap f E = bind g E. +Proof. + - auto_map_bind_pure. + - auto_map_bind_pure. + - auto_map_bind_pure. +Qed. + +#[global] Instance BindMapPure_expr : BindMapPure expr. +Proof. constructor. exact @ebind_map_pure. Qed. + +#[global] Instance BindMapPure_val : BindMapPure val. +Proof. constructor. exact @vbind_map_pure. Qed. + +#[global] Instance BindMapPure_ectx : BindMapPure ectx. +Proof. constructor. exact @ectx_bind_map_pure. Qed. + +Fixpoint rctx_bind_map_pure {A B} (f : A [→] B) g (R : rctx A) : + equal (subst_of_arr f) g → fmap f R = bind g R. +Proof. auto_map_bind_pure. Qed. + +#[global] Instance BindMapPure_rctx : BindMapPure rctx. +Proof. constructor. exact @rctx_bind_map_pure. Qed. + +(** Bind-Map_Comm instances *) + +Fixpoint ebind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (e : expr A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 e) = fmap f1 (bind g1 e) +with vbind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (v : val A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 v) = fmap f1 (bind g1 v) +with ectx_bind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (E : ectx A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 E) = fmap f1 (bind g1 E). +Proof. + - auto_map_bind_comm. + - auto_map_bind_comm. + - auto_map_bind_comm. +Qed. + +#[global] Instance BindMapComm_expr : BindMapComm expr. +Proof. constructor. exact @ebind_map_comm. Qed. + +#[global] Instance BindMapComm_val : BindMapComm val. +Proof. constructor. exact @vbind_map_comm. Qed. + +#[global] Instance BindMapComm_ectx : BindMapComm ectx. +Proof. constructor. exact @ectx_bind_map_comm. Qed. + +Fixpoint rctx_bind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (R : rctx A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 R) = fmap f1 (bind g1 R). +Proof. auto_map_bind_comm. Qed. + +#[global] Instance BindMapComm_rctx : BindMapComm rctx. +Proof. constructor. exact @rctx_bind_map_comm. Qed. + +(** Bind instances *) + +Fixpoint ebind_id {A} (f : A [⇒] A) (e : expr A) : + equal f (arrow_id A) → bind f e = e +with vbind_id {A} (f : A [⇒] A) (v : val A) : + equal f (arrow_id A) → bind f v = v +with ectx_bind_id {A} (f : A [⇒] A) (E : ectx A) : + equal f (arrow_id A) → bind f E = E. +Proof. + - auto_bind_id. + - auto_bind_id. + - auto_bind_id. +Qed. + +Fixpoint ebind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (e : expr A) : + equal (arrow_comp f g) h → bind f (bind g e) = bind h e +with vbind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (v : val A) : + equal (arrow_comp f g) h → bind f (bind g v) = bind h v +with ectx_bind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (E : ectx A) : + equal (arrow_comp f g) h → bind f (bind g E) = bind h E. +Proof. + - auto_bind_comp. + - auto_bind_comp. + - auto_bind_comp. +Qed. + +#[global] Instance Bind_expr : Bind expr. +Proof. constructor. exact @ebind_id. exact @ebind_comp. Qed. + +#[global] Instance Bind_val : Bind val. +Proof. constructor. exact @vbind_id. exact @vbind_comp. Qed. + +#[global] Instance Bind_ectx : Bind ectx. +Proof. constructor. exact @ectx_bind_id. exact @ectx_bind_comp. Qed. + +Fixpoint rctx_bind_id {A} (f : A [⇒] A) (R : rctx A) : + equal f (arrow_id A) → bind f R = R. +Proof. auto_bind_id. Qed. + +Fixpoint rctx_bind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (R : rctx A) : + equal (arrow_comp f g) h → bind f (bind g R) = bind h R. +Proof. auto_bind_comp. Qed. + +#[global] Instance Bind_rctx : Bind rctx. +Proof. constructor. exact @rctx_bind_id. exact @rctx_bind_comp. Qed. + +(** Meta-context *) + +Definition meta_ectx (V : Set) := list (ectx V). + +Fixpoint meta_plug {V} (M : meta_ectx V) (e : expr V) : expr V := + match M with + | [] => e + | E :: M => meta_plug M (plug E (e_reset e)) + end. + +Lemma fold_unfold_meta_plug_nil {V} (e : expr V) : + meta_plug [] e = e. +Proof. auto. Qed. + +Lemma fold_unfold_meta_plug_cons {V} (E : ectx V) M e : + meta_plug (E :: M) e = meta_plug M (plug E (e_reset e)). +Proof. auto. Qed. + +Definition meta_ectx_comp {V} (M1 M2 : meta_ectx V) : meta_ectx V := + M2 ++ M1. + +Lemma meta_plug_eq_val_inv {V} (M : meta_ectx V) (e : expr V) (v : val V) : + meta_plug M e = v → + M = [] ∧ e = v. +Proof. + revert e. + induction M as [| E M IHM]; intros e H_eq; simpl in *. + - auto. + - apply IHM in H_eq as [_ H_eq]. + apply plug_eq_val_inv in H_eq as [_ H_absurd]. + discriminate. +Qed. + +Lemma meta_ectx_comp_correct {V} (M1 M2 : meta_ectx V) e : + meta_plug (meta_ectx_comp M1 M2) e = meta_plug M1 (meta_plug M2 e). +Proof. + revert e. + induction M2 as [| E M2 IHM2]; intros e. + - simpl. reflexivity. + - simpl. rewrite -> IHM2. reflexivity. +Qed. + +Definition meta_ectx_map {A B} (f : A [→] B) (M : meta_ectx A) : meta_ectx B := + map (ectx_map f) M. + +Definition meta_ectx_bind {A B} (f : A [⇒] B) (M : meta_ectx A) : meta_ectx B := + map (ectx_bind f) M. + +#[global] Instance FunctorCore_meta_ectx : FunctorCore meta_ectx := @meta_ectx_map. +#[global] Instance BindCore_meta_ectx : BindCore meta_ectx := @meta_ectx_bind. + +Fixpoint meta_ectx_map_id {A} (f : A [→] A) (M : meta_ectx A) : + equal f (arrow_id A) → fmap f M = M. +Proof. auto_map_id. Qed. + +Fixpoint meta_ectx_map_comp {A B C} (f : B [→] C) (g : A [→] B) (h : A [→] C) (M : meta_ectx A) : + equal (arrow_comp f g) h → fmap f (fmap g M) = fmap h M. +Proof. auto_map_comp. Qed. + +#[global] Instance Functor_meta_ectx : Functor meta_ectx. +Proof. constructor. exact @meta_ectx_map_id. exact @meta_ectx_map_comp. Qed. + +Fixpoint meta_ectx_bind_map_pure {A B} (f : A [→] B) g (M : meta_ectx A) : + equal (subst_of_arr f) g → fmap f M = bind g M. +Proof. auto_map_bind_pure. Qed. + +#[global] Instance BindMapPure_meta_ectx : BindMapPure meta_ectx. +Proof. constructor. exact @meta_ectx_bind_map_pure. Qed. + +Fixpoint meta_ectx_bind_map_comm {A B1 B2 C} + (f1 : B1 [→] C) (f2 : A [→] B2) (g1 : A [⇒] B1) (g2 : B2 [⇒] C) (M : meta_ectx A) : + equal (arrow_comp g2 (subst_of_arr f2)) (arrow_comp (subst_of_arr f1) g1) → + bind g2 (fmap f2 M) = fmap f1 (bind g1 M). +Proof. auto_map_bind_comm. Qed. + +#[global] Instance BindMapComm_meta_ectx : BindMapComm meta_ectx. +Proof. constructor. exact @meta_ectx_bind_map_comm. Qed. + +Fixpoint meta_ectx_bind_id {A} (f : A [⇒] A) (M : meta_ectx A) : + equal f (arrow_id A) → bind f M = M. +Proof. auto_bind_id. Qed. + +Fixpoint meta_ectx_bind_comp {A B C} + (f : B [⇒] C) (g : A [⇒] B) (h : A [⇒] C) (M : meta_ectx A) : + equal (arrow_comp f g) h → bind f (bind g M) = bind h M. +Proof. auto_bind_comp. Qed. + +#[global] Instance Bind_meta_ectx : Bind meta_ectx. +Proof. constructor. exact @meta_ectx_bind_id. exact @meta_ectx_bind_comp. Qed. + +(** Reduction *) + +Inductive base_step {V} : expr V → expr V → Prop := +| bs_beta_lambda (e : expr (inc V)) (v : val V) : + base_step (e_app (v_lambda e) v) (subst (Inc:=inc) e v) +| bs_beta_cont (E : ectx V) (v : val V) : + base_step (e_app (v_cont E) v) (e_reset (plug E v)) +| bs_reset (v : val V) : + base_step (e_reset v) v. + +Inductive contextual_step {V} : expr V → expr V → Prop := +| cs_base (E : ectx V) (e1 e2 : expr V) : + base_step e1 e2 → + contextual_step (plug E e1) (plug E e2) +| cs_shift (E : ectx V) (e : expr (inc V)) : + contextual_step (plug E (e_shift e)) (subst (Inc:=inc) e (v_cont E)). + +Inductive meta_contextual_step {V} : expr V → expr V → Prop := +| mcs_contextual (M : meta_ectx V) (e1 e2 : expr V) : + contextual_step e1 e2 → + meta_contextual_step (meta_plug M e1) (meta_plug M e2). + +(* +Inductive contextual_step {V} : expr V → expr V → Prop := +| Ectx_beta_lambda MR R (e : expr (inc V)) (v : val V) : + contextual_step + (meta_plug MR (rplug R (e_app (v_lambda e) v))) + (meta_plug MR (rplug R (subst (Inc:=inc) e v))) +| Ectx_shift MR R (e : expr (inc V)) : + contextual_step + (meta_plug MR (rplug R (e_shift e))) + (meta_plug MR (subst (Inc:=inc) e (v_cont R))) +| Ectx_beta_cont MR R (R' : rctx V) (v : val V) : + contextual_step + (meta_plug MR (rplug R (e_app (v_cont R') v))) + (meta_plug MR (rplug R (e_reset (rplug R' v)))) +| Ectx_reset MR R (v : val V) : + contextual_step + (meta_plug MR (rplug R (e_reset v))) + (meta_plug MR (rplug R v)). +*) + +Definition big_step {V} e (v : val V) := + rtc meta_contextual_step e v. + +Definition terminates {V} (e : expr V) := + ∃ v, big_step e v. + +Lemma not_base_step_val {V} (v : val V) e : ¬ base_step v e. +Proof. inversion_clear 1. Qed. + +Lemma not_contextual_step_val {V} (v : val V) e : ¬ contextual_step v e. +Proof. + intros H_step. + inversion H_step as [E e1 e2 H_step' Hv He | E e' Hv He]. + - apply plug_eq_val_inv in Hv as [_ ->]. + by eapply not_base_step_val. + - apply plug_eq_val_inv in Hv as [_ H_absurd]. + discriminate. +Qed. + +Lemma not_meta_contextual_step_val {V} (v : val V) e : ¬ meta_contextual_step v e. +Proof. + intros H_step. + inversion H_step as [M e1 e2 H_step' Hv He]. + apply meta_plug_eq_val_inv in Hv as [_ ->]. + by apply not_contextual_step_val in H_step'. +Qed. + +Lemma big_step_val {V} (v : val V) : big_step v v. +Proof. unfold big_step. done. Qed. + +Lemma terminates_val {V} (v : val V) : terminates v. +Proof. unfold terminates. exists v. apply big_step_val. Qed. + +Lemma meta_contextual_step_comp {V} (M : meta_ectx V) e1 e2 : + meta_contextual_step e1 e2 → + meta_contextual_step (meta_plug M e1) (meta_plug M e2). +Proof. + intros H_step. + inversion_clear H_step as [M' e1' e2' H_step']. + rewrite <- meta_ectx_comp_correct. + rewrite <- meta_ectx_comp_correct. + constructor. exact H_step'. +Qed. + +Lemma meta_contextual_step_terminates {V} (e e' : expr V) : + meta_contextual_step e e' → + terminates e' → + terminates e. +Proof. + unfold terminates, big_step. + intros H_step [v H_steps]. + exists v. econstructor; eauto. +Qed. + +Lemma base_step_is_deterministic {V} (e1 e2 e3 : expr V) : + base_step e1 e2 → + base_step e1 e3 → + e2 = e3. +Proof. + intros Hstep2 Hstep3. + inversion Hstep2; inversion Hstep3; congruence. +Qed. + +Inductive potential_redex {V} : expr V -> Prop := +| pr_app (v1 v2 : val V) : potential_redex (e_app v1 v2) +| pr_shift (e : expr (inc V)) : potential_redex (e_shift e) +| pr_reset (v : val V) : potential_redex (e_reset v). + +Lemma not_potential_redex_val {V} (v : val V) : ¬ potential_redex v. +Proof. inversion_clear 1. Qed. + +Lemma potential_redex_app_inv {V} e1 e2 : + potential_redex (e_app e1 e2) → + ∃ (v1 v2 : val V), e1 = v1 ∧ e2 = v2. +Proof. inversion_clear 1. eauto. Qed. + +Lemma potential_redex_reset_inv {V} e : + potential_redex (e_reset e) → + ∃ v : val V, e = v. +Proof. inversion_clear 1. eauto. Qed. + +Lemma base_step_potential_redex {V} (e e' : expr V) : + base_step e e' → + potential_redex e. +Proof. inversion_clear 1; constructor. Qed. + +Lemma unique_rdecomposition {V} (R1 R2 : rctx V) e1 e2 : + potential_redex e1 → + potential_redex e2 → + rplug R1 e1 = rplug R2 e2 → + R1 = R2 ∧ e1 = e2. +Proof. + intros He1 He2. + revert R2. + induction R1; intros R2 Heq. + - destruct R2; simpl in *. + + auto. + + rewrite -> Heq in He1. + apply potential_redex_app_inv in He1 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv1 as [_ ->]. + contradict (not_potential_redex_val _ He2). + + rewrite -> Heq in He1. + apply potential_redex_app_inv in He1 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv2 as [_ ->]. + contradict (not_potential_redex_val _ He2). + - destruct R2; simpl in *. + + rewrite <- Heq in He2. + apply potential_redex_app_inv in He2 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv1 as [_ ->]. + contradict (not_potential_redex_val _ He1). + + injection Heq as Heq1 Heq2. + apply IHR1 in Heq1 as [Heq11 Heq12]. + split; congruence. + + injection Heq as Heq1 Heq2. + apply rplug_eq_val_inv in Heq1 as [_ ->]. + contradict (not_potential_redex_val _ He1). + - destruct R2; simpl in *. + + rewrite <- Heq in He2. + apply potential_redex_app_inv in He2 as (v1 & v2 & Hv1 & Hv2). + apply rplug_eq_val_inv in Hv2 as [_ ->]. + contradict (not_potential_redex_val _ He1). + + injection Heq as Heq1 Heq2. symmetry in Heq1. + apply rplug_eq_val_inv in Heq1 as [_ ->]. + contradict (not_potential_redex_val _ He2). + + injection Heq as Heq1 Heq2. + apply IHR1 in Heq2 as [Heq11 Heq12]. + split; congruence. +Qed. + +Lemma unique_decomposition {V} (E1 E2 : ectx V) e1 e2 : + potential_redex e1 → + potential_redex e2 → + plug E1 e1 = plug E2 e2 → + E1 = E2 ∧ e1 = e2. +Proof. + intros He1 He2 Heq. + rewrite <- ectx_to_rctx_correct in Heq. + rewrite <- ectx_to_rctx_correct in Heq. + destruct (unique_rdecomposition _ _ _ _ He1 He2 Heq) as [Heq1 Heq2]. + split. + - rewrite <- (ectx_rctx_bijection E1). + rewrite <- (ectx_rctx_bijection E2). + f_equal. exact Heq1. + - exact Heq2. +Qed. + +Lemma contextual_step_is_deterministic {V} (e1 e2 e3 : expr V) : + contextual_step e1 e2 → + contextual_step e1 e3 → + e2 = e3. +Proof. +Admitted. + +(* + intros Hstep2 Hstep3. + inversion Hstep2 as [E2 e12 e2' Hstep2' He12 He2' |]. + inversion Hstep3 as [E3 e13 e3' Hstep3' He13 He3' |]. + assert (Hpr2 := base_step_potential_redex _ _ Hstep2'). + assert (Hpr3 := base_step_potential_redex _ _ Hstep3'). + destruct (unique_decomposition E2 E3 e12 e13 Hpr2 Hpr3) as [HE_eq He_eq]. + { congruence. } + rewrite -> He_eq in Hstep2'. + assert (He_eq' := base_step_is_deterministic e13 e2' e3' Hstep2' Hstep3'). + congruence. +Qed. +*) + +Lemma meta_contextual_step_is_deterministic {V} (e1 e2 e3 : expr V) : + meta_contextual_step e1 e2 → + meta_contextual_step e1 e3 → + e2 = e3. +Proof. +Admitted. + +(** Relations for closed term *) + +Definition expr_rel := expr ∅ ⇒ᵢ expr ∅ ⇒ᵢ IRel. +Definition val_rel := val ∅ ⇒ᵢ val ∅ ⇒ᵢ IRel. +Definition ectx_rel := ectx ∅ ⇒ᵢ ectx ∅ ⇒ᵢ IRel. +Definition meta_rctx_rel := meta_ectx ∅ ⇒ᵢ meta_ectx ∅ ⇒ᵢ IRel. + +Definition L_rel_pre (L_rel : expr_rel) : expr_rel := + λ e1 e2, + (∀ v1 : val ∅, e1 = v1 → terminates e2)ᵢ ∧ᵢ + (∀ᵢ e1' : expr ∅, (meta_contextual_step e1 e1')ᵢ →ᵢ ▷ L_rel e1' e2). + +Definition L_rel_fix := I_fix L_rel_pre. +Definition L_rel := L_rel_pre L_rel_fix. + +Definition O_rel : expr_rel := + λ e1 e2, L_rel e1 e2 ∧ᵢ L_rel e2 e1. + +Definition M_rel_pre (V_rel : val_rel) : meta_rctx_rel := + λ M1 M2, + ∀ᵢ (v1 v2 : val ∅), + V_rel v1 v2 →ᵢ + O_rel (meta_plug M1 v1) (meta_plug M2 v2). + +Definition K_rel_pre (V_rel : val_rel) : ectx_rel := + λ E1 E2, + ∀ᵢ M1 M2 (v1 v2 : val ∅), + M_rel_pre V_rel M1 M2 →ᵢ + V_rel v1 v2 →ᵢ + O_rel (meta_plug M1 (plug E1 v1)) (meta_plug M2 (plug E2 v2)). + +Definition R_rel_pre (V_rel : val_rel) : expr_rel := + λ e1 e2, + ∀ᵢ M1 M2 E1 E2, + ▷ M_rel_pre V_rel M1 M2 →ᵢ + ▷ K_rel_pre V_rel E1 E2 →ᵢ + O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2)). + +Definition V_rel_pre (V_rel : val_rel) : val_rel := + λ v1 v2, + ∀ᵢ u1 u2, + ▷ V_rel u1 u2 →ᵢ + R_rel_pre V_rel (e_app v1 u1) (e_app v2 u2). + +Definition V_rel_fix := I_fix V_rel_pre. +Definition V_rel := V_rel_pre V_rel_fix. +Definition R_rel := R_rel_pre V_rel_fix. +Definition K_rel := K_rel_pre V_rel_fix. +Definition M_rel := M_rel_pre V_rel_fix. + +Definition E_rel (e1 e2 : expr ∅) := + ∀ᵢ M1 M2 E1 E2, + M_rel M1 M2 →ᵢ + K_rel E1 E2 →ᵢ + O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2)). + +(** Relations for open terms *) + +Definition G_rel {V} (γ1 γ2 : V [⇒] ∅) : IProp := + ∀ᵢ x, V_rel (γ1 x) (γ2 x). + +Definition E_rel_o {V} (e1 e2 : expr V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ E_rel (bind γ1 e1) (bind γ2 e2). + +Definition V_rel_o {V} (v1 v2 : val V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ V_rel (bind γ1 v1) (bind γ2 v2). + +Definition O_rel_o {V} (e1 e2 : expr V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ O_rel (bind γ1 e1) (bind γ2 e2). + +Definition K_rel_o {V} (E1 E2 : ectx V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ K_rel (bind γ1 E1) (bind γ2 E2). + +Definition M_rel_o {V} (M1 M2 : meta_ectx V) : IProp := + ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ M_rel (bind γ1 M1) (bind γ2 M2). + +(** Contractiveness and unrolling fixpoint *) + +Lemma L_rel_pre_contractive : contractive L_rel_pre. +Proof. intro n; iintros; unfold L_rel_pre; auto_contr. Qed. + +Lemma L_rel_roll p1 p2 n : + n ⊨ L_rel p1 p2 → + n ⊨ L_rel_fix p1 p2. +Proof. + intro H; iapply (I_fix_roll expr_rel); [| exact H]. + apply L_rel_pre_contractive. +Qed. + +Lemma L_rel_unroll p1 p2 n : + n ⊨ L_rel_fix p1 p2 → + n ⊨ L_rel p1 p2. +Proof. + intro H; iapply (I_fix_unroll expr_rel); [| exact H]. + apply L_rel_pre_contractive. +Qed. + +Lemma V_rel_pre_contractive : contractive V_rel_pre. +Proof. + intros n. iintros. + unfold V_rel_pre, R_rel_pre, K_rel_pre, M_rel_pre. + auto_contr. +Qed. + +Lemma V_rel_roll v1 v2 n : + n ⊨ V_rel v1 v2 → + n ⊨ V_rel_fix v1 v2. +Proof. + intro H; iapply (I_fix_roll val_rel); [| exact H]. + apply V_rel_pre_contractive. +Qed. + +Lemma V_rel_unroll v1 v2 n : + n ⊨ V_rel_fix v1 v2 → + n ⊨ V_rel v1 v2. +Proof. + intro H; iapply (I_fix_unroll val_rel); [| exact H]. + apply V_rel_pre_contractive. +Qed. + +(** Introduction and elimination lemmas *) + +Lemma L_rel_intro (e1 e2 : expr ∅) n : + (∀ v1 : val ∅, e1 = v1 → terminates e2) → + n ⊨ (∀ᵢ e1' : expr ∅, (meta_contextual_step e1 e1')ᵢ →ᵢ ▷ L_rel e1' e2) → + n ⊨ L_rel e1 e2. +Proof. + intros H_val H_expr. + unfold L_rel, L_rel_pre. + isplit. + - iintro. exact H_val. + - iintros e1' H_step. + ispec H_expr e1' H_step. + later_shift. apply L_rel_roll. exact H_expr. +Qed. + +Lemma L_rel_elim (e1 e2 : expr ∅) n : + n ⊨ L_rel e1 e2 → + (∀ v1 : val ∅, e1 = v1 → terminates e2) ∧ + (n ⊨ ∀ᵢ e1' : expr ∅, (meta_contextual_step e1 e1')ᵢ →ᵢ ▷ L_rel e1' e2). +Proof. + intros He. + unfold L_rel, L_rel_pre in He. + idestruct He as He1 He2. + split. + - idestruct He1. exact He1. + - iintros e1' H_step. + ispec He2 e1' H_step. + later_shift. apply L_rel_unroll. exact He2. +Qed. + +Lemma O_rel_intro (e1 e2 : expr ∅) n : + n ⊨ L_rel e1 e2 → + n ⊨ L_rel e2 e1 → + n ⊨ O_rel e1 e2. +Proof. + intros He1 He2. + unfold O_rel. isplit; assumption. +Qed. + +Lemma O_rel_elim (e1 e2 : expr ∅) n : + n ⊨ O_rel e1 e2 → + (n ⊨ L_rel e1 e2) ∧ + (n ⊨ L_rel e2 e1). +Proof. + unfold O_rel. + intros He. idestruct He as He1 He2. + split; assumption. +Qed. + +Lemma O_rel_elim1 (e1 e2 : expr ∅) n : + n ⊨ O_rel e1 e2 → + n ⊨ L_rel e1 e2. +Proof. intros He. by apply O_rel_elim in He as []. Qed. + +Lemma O_rel_elim2 (e1 e2 : expr ∅) n : + n ⊨ O_rel e1 e2 → + n ⊨ L_rel e2 e1. +Proof. intros He. by apply O_rel_elim in He as []. Qed. + +Lemma V_rel_intro (v1 v2 : val ∅) n : + (n ⊨ ∀ᵢ u1 u2, + ▷ V_rel u1 u2 →ᵢ + R_rel (e_app v1 u1) (e_app v2 u2)) → + n ⊨ V_rel v1 v2. +Proof. + intros Hv. + unfold V_rel, V_rel_pre. + iintros u1 u2 Hu. + ispecialize Hv u1. + ispecialize Hv u2. + iapply Hv. later_shift. + apply V_rel_unroll. exact Hu. +Qed. + +Lemma V_rel_elim (v1 v2 : val ∅) n : + n ⊨ V_rel v1 v2 → + n ⊨ ∀ᵢ u1 u2, + ▷ V_rel u1 u2 →ᵢ + R_rel (e_app v1 u1) (e_app v2 u2). +Proof. + intros Hv. + unfold V_rel, V_rel_pre in Hv. + iintros u1 u2 Hu. + ispecialize Hv u1. + ispecialize Hv u2. + iapply Hv. later_shift. + apply V_rel_roll. exact Hu. +Qed. + +Lemma V_rel_elimR (v1 v2 u1 u2 : val ∅) n : + n ⊨ V_rel v1 v2 → + n ⊨ ▷ V_rel u1 u2 → + n ⊨ R_rel (e_app v1 u1) (e_app v2 u2). +Proof. + intros Hv Hu. + apply V_rel_elim in Hv. + iapply Hv. exact Hu. +Qed. + +Lemma M_rel_intro (M1 M2 : meta_ectx ∅) n : + n ⊨ (∀ᵢ v1 v2, V_rel v1 v2 →ᵢ O_rel (meta_plug M1 v1) (meta_plug M2 v2)) → + n ⊨ M_rel M1 M2. +Proof. + intros HM. + unfold M_rel, M_rel_pre. + iintros v1 v2 Hv. + iapply HM. apply V_rel_unroll. exact Hv. +Qed. + +Lemma M_rel_elim (M1 M2 : meta_ectx ∅) n : + n ⊨ M_rel M1 M2 → + n ⊨ ∀ᵢ v1 v2, V_rel v1 v2 →ᵢ O_rel (meta_plug M1 v1) (meta_plug M2 v2). +Proof. + unfold M_rel, M_rel_pre. + intros HM. + iintros v1 v2 Hv. + iapply HM. apply V_rel_roll. exact Hv. +Qed. + +Lemma M_rel_elimO (M1 M2 : meta_ectx ∅) (v1 v2 : val ∅) n : + n ⊨ M_rel M1 M2 → + n ⊨ V_rel v1 v2 → + n ⊨ O_rel (meta_plug M1 v1) (meta_plug M2 v2). +Proof. + intros HM Hv. + apply M_rel_elim in HM. + iapply HM. exact Hv. +Qed. + +Lemma K_rel_intro (E1 E2 : ectx ∅) n : + n ⊨ (∀ᵢ M1 M2 v1 v2, + M_rel M1 M2 →ᵢ + V_rel v1 v2 →ᵢ + O_rel (meta_plug M1 (plug E1 v1)) (meta_plug M2 (plug E2 v2))) → + n ⊨ K_rel E1 E2. +Proof. + intros HE. + unfold K_rel, K_rel_pre. + iintros M1 M2 v1 v2 HM Hv. + iapply HE. exact HM. apply V_rel_unroll. exact Hv. +Qed. + +Lemma K_rel_elim (E1 E2 : ectx ∅) n : + n ⊨ K_rel E1 E2 → + n ⊨ ∀ᵢ M1 M2 v1 v2, + M_rel M1 M2 →ᵢ + V_rel v1 v2 →ᵢ + O_rel (meta_plug M1 (plug E1 v1)) (meta_plug M2 (plug E2 v2)). +Proof. + unfold K_rel, K_rel_pre. + intros HE. + iintros M1 M2 v1 v2 HM Hv. + iapply HE. exact HM. apply V_rel_roll. exact Hv. +Qed. + +Lemma K_rel_elimO E1 E2 M1 M2 v1 v2 n : + n ⊨ K_rel E1 E2 → + n ⊨ M_rel M1 M2 → + n ⊨ V_rel v1 v2 → + n ⊨ O_rel (meta_plug M1 (plug E1 v1)) (meta_plug M2 (plug E2 v2)). +Proof. + intros HE HM Hv. + apply K_rel_elim in HE. + iapply HE. exact HM. exact Hv. +Qed. + +Lemma R_rel_intro (e1 e2 : expr ∅) n : + n ⊨ (∀ᵢ M1 M2 E1 E2, + ▷ M_rel M1 M2 →ᵢ + ▷ K_rel E1 E2 →ᵢ + O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2))) -> + n ⊨ R_rel e1 e2. +Proof. auto. Qed. + +Lemma R_rel_elim (e1 e2 : expr ∅) n : + n ⊨ R_rel e1 e2 → + n ⊨ ∀ᵢ M1 M2 E1 E2, + ▷ M_rel M1 M2 →ᵢ + ▷ K_rel E1 E2 →ᵢ + O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2)). +Proof. auto. Qed. + +Lemma R_rel_elimO (e1 e2 : expr ∅) M1 M2 E1 E2 n : + n ⊨ R_rel e1 e2 → + n ⊨ ▷ M_rel M1 M2 → + n ⊨ ▷ K_rel E1 E2 → + n ⊨ O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2)). +Proof. + intros He HM HE. + apply R_rel_elim in He. + iapply He. exact HM. exact HE. +Qed. + +Lemma E_rel_intro (e1 e2 : expr ∅) n : + n ⊨ (∀ᵢ M1 M2 E1 E2, + M_rel M1 M2 →ᵢ + K_rel E1 E2 →ᵢ + O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2))) → + n ⊨ E_rel e1 e2. +Proof. auto. Qed. + +Lemma E_rel_elim (e1 e2 : expr ∅) n : + n ⊨ E_rel e1 e2 → + n ⊨ ∀ᵢ M1 M2 E1 E2, + M_rel M1 M2 →ᵢ + K_rel E1 E2 →ᵢ + O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2)). +Proof. auto. Qed. + +(** Bind lemma *) +Lemma E_rel_elimO e1 e2 M1 M2 E1 E2 n : + n ⊨ E_rel e1 e2 → + n ⊨ M_rel M1 M2 → + n ⊨ K_rel E1 E2 → + n ⊨ O_rel (meta_plug M1 (plug E1 e1)) (meta_plug M2 (plug E2 e2)). +Proof. + intros He HM HE. + apply E_rel_elim in He. + iapply He. exact HM. exact HE. +Qed. + +Lemma G_rel_intro {V} (γ1 γ2 : V [⇒] ∅) n : + (n ⊨ ∀ᵢ x, V_rel (γ1 x) (γ2 x)) → + n ⊨ G_rel γ1 γ2. +Proof. auto. Qed. + +Lemma G_rel_elim {V} (γ1 γ2 : V [⇒] ∅) n : + n ⊨ G_rel γ1 γ2 → + n ⊨ ∀ᵢ x, V_rel (γ1 x) (γ2 x). +Proof. auto. Qed. + +Lemma G_rel_elimV {V} (γ1 γ2 : V [⇒] ∅) (x : V) n : + n ⊨ G_rel γ1 γ2 → + n ⊨ V_rel (γ1 x) (γ2 x). +Proof. + intros Hγ. + apply G_rel_elim in Hγ. + iapply Hγ. +Qed. + +Lemma E_rel_o_intro {V} (e1 e2 : expr V) n : + (n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ E_rel (bind γ1 e1) (bind γ2 e2)) → + n ⊨ E_rel_o e1 e2. +Proof. auto. Qed. + +Lemma E_rel_o_elim {V} (e1 e2 : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ E_rel (bind γ1 e1) (bind γ2 e2). +Proof. auto. Qed. + +Lemma E_rel_o_elimE {V} (e1 e2 : expr V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ E_rel (bind γ1 e1) (bind γ2 e2). +Proof. + intros He Hγ. + apply E_rel_o_elim in He. + iapply He. exact Hγ. +Qed. + +Lemma V_rel_o_intro {V} (v1 v2 : val V) n : + (n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ V_rel (bind γ1 v1) (bind γ2 v2)) → + n ⊨ V_rel_o v1 v2. +Proof. auto. Qed. + +Lemma V_rel_o_elim {V} (v1 v2 : val V) n : + n ⊨ V_rel_o v1 v2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ V_rel (bind γ1 v1) (bind γ2 v2). +Proof. auto. Qed. + +Lemma V_rel_o_elimV {V} (v1 v2 : val V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ V_rel_o v1 v2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ V_rel (bind γ1 v1) (bind γ2 v2). +Proof. + intros Hv Hγ. + apply V_rel_o_elim in Hv. + iapply Hv. exact Hγ. +Qed. + +Lemma O_rel_o_intro {V} (e1 e2 : expr V) n : + (n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ O_rel (bind γ1 e1) (bind γ2 e2)) → + n ⊨ O_rel_o e1 e2. +Proof. auto. Qed. + +Lemma O_rel_o_elim {V} (e1 e2 : expr V) n : + n ⊨ O_rel_o e1 e2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ O_rel (bind γ1 e1) (bind γ2 e2). +Proof. auto. Qed. + +Lemma O_rel_o_elimO {V} (e1 e2 : expr V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ O_rel_o e1 e2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ O_rel (bind γ1 e1) (bind γ2 e2). +Proof. + intros He Hγ. + apply O_rel_o_elim in He. + iapply He. exact Hγ. +Qed. + +Lemma K_rel_o_intro {V} (E1 E2 : ectx V) n : + n ⊨ (∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ K_rel (bind γ1 E1) (bind γ2 E2)) → + n ⊨ K_rel_o E1 E2. +Proof. auto. Qed. + +Lemma K_rel_o_elim {V} (E1 E2 : ectx V) n : + n ⊨ K_rel_o E1 E2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ K_rel (bind γ1 E1) (bind γ2 E2). +Proof. auto. Qed. + +Lemma K_rel_o_elimK {V} (E1 E2 : ectx V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ K_rel_o E1 E2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ K_rel (bind γ1 E1) (bind γ2 E2). +Proof. + intros HE Hγ. + apply K_rel_o_elim in HE. + iapply HE. exact Hγ. +Qed. + +Lemma M_rel_o_intro {V} (M1 M2 : meta_ectx V) n : + n ⊨ (∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ M_rel (bind γ1 M1) (bind γ2 M2)) → + n ⊨ M_rel_o M1 M2. +Proof. auto. Qed. + +Lemma M_rel_o_elim {V} (M1 M2 : meta_ectx V) n : + n ⊨ M_rel_o M1 M2 → + n ⊨ ∀ᵢ γ1 γ2, G_rel γ1 γ2 →ᵢ M_rel (bind γ1 M1) (bind γ2 M2). +Proof. auto. Qed. + +Lemma M_rel_o_elimM {V} (M1 M2 : meta_ectx V) (γ1 γ2 : V [⇒] ∅) n : + n ⊨ M_rel_o M1 M2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ M_rel (bind γ1 M1) (bind γ2 M2). +Proof. + intros HE Hγ. + apply M_rel_o_elim in HE. + iapply HE. exact Hγ. +Qed. + +(** Compatibility lemmas *) + +Lemma L_rel_val (v1 v2 : val ∅) n : + n ⊨ L_rel v1 v2. +Proof. + apply L_rel_intro. + - intros _ _. apply terminates_val. + - iintros e1 He. idestruct He. + by apply not_meta_contextual_step_val in He. +Qed. + +Lemma O_rel_val (v1 v2 : val ∅) n : + n ⊨ O_rel v1 v2. +Proof. + apply O_rel_intro. + - apply L_rel_val. + - apply L_rel_val. +Qed. + +Lemma compat_val_closed v1 v2 n : + n ⊨ V_rel v1 v2 → + n ⊨ E_rel v1 v2. +Proof. + intros Hv. + apply E_rel_intro. iintros M1 M2 E1 E2 HM HE. + apply K_rel_elimO; assumption. +Qed. + +(* aka val inclusion *) +Lemma compat_val {V} (v1 v2 : val V) n : + n ⊨ V_rel_o v1 v2 → + n ⊨ E_rel_o v1 v2. +Proof. + intros Hv. + apply V_rel_o_elim in Hv. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. + apply compat_val_closed. iapply Hv. exact Hγ. +Qed. + +Lemma compat_app_closed_val (v1 v2 u1 u2 : val ∅) n : + n ⊨ V_rel v1 v2 → + n ⊨ V_rel u1 u2 → + n ⊨ E_rel (e_app v1 u1) (e_app v2 u2). +Proof. + intros Hv Hu. + apply E_rel_intro. + iintros M1 M2 E1 E2 HM HE. + apply R_rel_elimO. + - apply V_rel_elimR. exact Hv. + later_shift. exact Hu. + - later_shift. exact HM. + - later_shift. exact HE. +Qed. + +Lemma compat_app_closed e1 e2 e1' e2' n : + n ⊨ E_rel e1 e2 → + n ⊨ E_rel e1' e2' → + n ⊨ E_rel (e_app e1 e1') (e_app e2 e2'). +Proof. + intros He He'. + apply E_rel_intro. iintros M1 M2 E1 E2 HM HE. term_simpl. + (* The functions e1/e2 are evaluated first, so we "zap" them down using He. + To use He, we have to give two contexts s.t. if we can prove them to be + related, plugging e1/e2 into them will be in O_rel. We give ectx_app1 + because the plugging will give us exactly the goal we need. *) + apply E_rel_elim in He. + ispec He M1 M2. + ispecialize He (ectx_app1 E1 e1'). + ispecialize He (ectx_app1 E2 e2'). + ispec He HM. + iapply He. clear He. + (* Now, we need to show that the two app contexts are related. *) + apply K_rel_intro. iintros M1' M2' v1 v2 HM' Hv. simpl. + (* Given that they are plugged with two related values, we now have to prove + that the result is in O_rel. We use He' for a similar purpose. We give + ectx_app2 because plugging e1'/e2' into it will match the goal. *) + apply E_rel_elim in He'. + ispec He' M1' M2'. + ispecialize He' (ectx_app2 v1 E1). + ispecialize He' (ectx_app2 v2 E2). + ispec He' HM'. + iapply He'. clear He'. + (* Now we have to prove the ectx_app2 are related. *) + apply K_rel_intro. iintros M1'' M2'' v1' v2' HM'' Hv'. simpl. + (* Now, we have that the two values and contexts are related. + We "zap" (app v1 v1') and (app v2 v2') down using E_rel_elimO *) + apply E_rel_elimO. + apply compat_app_closed_val; [exact Hv | exact Hv']. + exact HM''. + exact HE. +Qed. + +Lemma compat_app {V} (e1 e2 e1' e2' : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o e1' e2' → + n ⊨ E_rel_o (e_app e1 e1') (e_app e2 e2'). +Proof. + intros He He'. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. + (* Use the lemma about closed app *) + apply compat_app_closed. + - apply E_rel_o_elimE. exact He. exact Hγ. + - apply E_rel_o_elimE. exact He'. exact Hγ. +Qed. + +Lemma compat_var {V : Set} (x : V) n : + n ⊨ V_rel_o (v_var x) (v_var x). +Proof. + apply V_rel_o_intro. + iintros γ1 γ2 Hγ. term_simpl. + apply G_rel_elimV. exact Hγ. +Qed. + +Lemma compat_ectx_hole_closed n : + n ⊨ K_rel ectx_hole ectx_hole. +Proof. + apply K_rel_intro. + iintros M1 M2 v1 v2 HM Hv. simpl. + apply M_rel_elimO. exact HM. exact Hv. +Qed. + +Lemma compat_ectx_hole {V} n : + n ⊨ @K_rel_o V ectx_hole ectx_hole. +Proof. + apply K_rel_o_intro. + iintros γ1 γ2 _. term_simpl. + apply compat_ectx_hole_closed. +Qed. + +Lemma L_rel_red_l (e1 e1' e2 : expr ∅) n : + meta_contextual_step e1 e1' → + n ⊨ ▷ L_rel e1' e2 → + n ⊨ L_rel e1 e2. +Proof. + intros H_step He. + apply L_rel_intro. + - intros v1 H_eq. + rewrite -> H_eq in H_step. + contradict (not_meta_contextual_step_val _ _ H_step). + - iintros e1'' H_step'. + idestruct H_step'. + rewrite -> (meta_contextual_step_is_deterministic _ _ _ H_step' H_step). + exact He. +Qed. + +Lemma L_rel_red_r (e1 e2 e2' : expr ∅) n : + meta_contextual_step e2 e2' → + n ⊨ L_rel e1 e2' → + n ⊨ L_rel e1 e2. +Proof. + intros H_step He1. + irevert e1 He1. + loeb_induction IH. + iintros e1 He. + apply L_rel_elim in He as [He1 He2]. + apply L_rel_intro. + - intros v1 H_eq. + specialize (He1 v1 H_eq). + eapply meta_contextual_step_terminates; eauto. + - iintros e1' H_step'. + ispec He2 e1' H_step'. + later_shift. iapply IH. exact He2. +Qed. + +Lemma O_rel_red_l (e1 e1' e2 : expr ∅) n : + meta_contextual_step e1 e1' → + n ⊨ O_rel e1' e2 → + n ⊨ O_rel e1 e2. +Proof. + intros H_step He. + apply O_rel_elim in He as [He1 He2]. + apply O_rel_intro. + - eapply L_rel_red_l. + + exact H_step. + + later_shift. exact He1. + - eapply L_rel_red_r. + + exact H_step. + + exact He2. +Qed. + +(* symmetric to the proof of O_rel_red_l *) +Lemma O_rel_red_r (e1 e2 e2' : expr ∅) n : + meta_contextual_step e2 e2' → + n ⊨ O_rel e1 e2' → + n ⊨ O_rel e1 e2. +Proof. + intros H_step He. + apply O_rel_elim in He as [He1 He2]. + apply O_rel_intro. + - eapply L_rel_red_r. + + exact H_step. + + exact He1. + - eapply L_rel_red_l. + + exact H_step. + + later_shift. exact He2. +Qed. + +Lemma O_rel_red_both (e1 e1' e2 e2' : expr ∅) n : + meta_contextual_step e1 e1' → + meta_contextual_step e2 e2' → + n ⊨ ▷ O_rel e1' e2' → + n ⊨ O_rel e1 e2. +Proof. + intros H_step1 H_step2 He. + unfold O_rel in He. + apply I_conj_later_down in He. + idestruct He as He1 He2. + apply O_rel_intro. + - eapply L_rel_red_l. { exact H_step1. } + later_shift. + eapply L_rel_red_r. { exact H_step2. } + exact He1. + - eapply L_rel_red_r. { exact H_step1. } + eapply L_rel_red_l. { exact H_step2. } + later_shift. exact He2. +Qed. + +Lemma compat_id_subst n : + n ⊨ G_rel (arrow_id ∅) (arrow_id ∅). +Proof. + apply G_rel_intro. + iintros x. destruct x. +Qed. + +Lemma compat_comp_subst {V} v1 v2 (γ1 γ2 : V [⇒] ∅) n : + n ⊨ V_rel v1 v2 → + n ⊨ G_rel γ1 γ2 → + n ⊨ G_rel (arrow_comp (mk_subst v1) (lift γ1)) (arrow_comp (mk_subst v2) (lift γ2)). +Proof. + intros Hv Hγ. + apply G_rel_intro. iintros [| x']. + - term_simpl. exact Hv. + - term_simpl. apply G_rel_elimV. exact Hγ. +Qed. + +Lemma compat_lambda {V} (e1 e2 : expr (inc V)) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ V_rel_o (v_lambda e1) (v_lambda e2). +Proof. + intros He. + apply V_rel_o_intro. iintros γ1 γ2 Hγ. + apply V_rel_intro. iintros u1 u2 Hu. term_simpl. + apply R_rel_intro. iintros M1 M2 E1 E2 HM HE. + eapply O_rel_red_both. + { econstructor. apply cs_base. constructor. } + { econstructor. apply cs_base. constructor. } + later_shift. unfold subst. + rewrite -> bind_bind_comp'. + rewrite -> bind_bind_comp'. + apply E_rel_elimO. + - apply E_rel_o_elimE. exact He. + apply compat_comp_subst. exact Hu. exact Hγ. + - exact HM. + - exact HE. +Qed. + +Lemma compat_meta_ectx_nil_closed n : + n ⊨ M_rel [] []. +Proof. + apply M_rel_intro. iintros v1 v2 _. simpl. + eapply O_rel_val. +Qed. + +Lemma compat_meta_ectx_nil {V : Set} n : + n ⊨ @M_rel_o V [] []. +Proof. + apply M_rel_o_intro. iintros γ1 γ2 _. + term_simpl. apply compat_meta_ectx_nil_closed. +Qed. + +Lemma compat_meta_ectx_cons_closed E1 E2 M1 M2 n : + n ⊨ K_rel E1 E2 → + n ⊨ M_rel M1 M2 → + n ⊨ M_rel (E1 :: M1) (E2 :: M2). +Proof. + intros HE HM. + apply M_rel_intro. iintros v1 v2 Hv. simpl. + eapply O_rel_red_both. + { econstructor. eapply cs_base. econstructor. } + { econstructor. eapply cs_base. econstructor. } + later_shift. apply K_rel_elimO. exact HE. exact HM. exact Hv. +Qed. + +Lemma compat_meta_ectx_cons {V} (E1 E2 : ectx V) M1 M2 n : + n ⊨ K_rel_o E1 E2 → + n ⊨ M_rel_o M1 M2 → + n ⊨ M_rel_o (E1 :: M1) (E2 :: M2). +Proof. + intros HE HM. + apply M_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. + apply compat_meta_ectx_cons_closed. + - apply K_rel_o_elimK. exact HE. exact Hγ. + - apply M_rel_o_elimM. exact HM. exact Hγ. +Qed. + +Lemma compat_reset_closed (e1 e2 : expr ∅) n : + n ⊨ E_rel e1 e2 → + n ⊨ E_rel (e_reset e1) (e_reset e2). +Proof. + intros He. + apply E_rel_intro. iintros M1 M2 E1 E2 HM HE. + rewrite <- fold_unfold_meta_plug_cons. + rewrite <- fold_unfold_meta_plug_cons. + rewrite <- (fold_unfold_plug_ectx_hole e1). + rewrite <- (fold_unfold_plug_ectx_hole e2). + apply E_rel_elimO. + - exact He. + - apply compat_meta_ectx_cons_closed. exact HE. exact HM. + - apply compat_ectx_hole_closed. +Qed. + +Lemma compat_reset {V} (e1 e2 : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (e_reset e1) (e_reset e2). +Proof. + intros He. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. + apply compat_reset_closed. + apply E_rel_o_elimE. exact He. exact Hγ. +Qed. + +Lemma compat_cont_closed (E1 E2 : ectx ∅) n : + n ⊨ K_rel E1 E2 → + n ⊨ V_rel (v_cont E1) (v_cont E2). +Proof. + intros HE. + apply V_rel_intro. iintros u1 u2 Hu. + apply R_rel_intro. iintros M1 M2 E1' E2' HM HE'. + eapply O_rel_red_both. + { econstructor. eapply cs_base. econstructor. } + { econstructor. eapply cs_base. econstructor. } + later_shift. + rewrite <- fold_unfold_meta_plug_cons. + rewrite <- fold_unfold_meta_plug_cons. + apply K_rel_elimO. + - exact HE. + - apply compat_meta_ectx_cons_closed. exact HE'. exact HM. + - exact Hu. +Qed. + +Lemma compat_cont {V} (E1 E2 : ectx V) n : + n ⊨ K_rel_o E1 E2 → + n ⊨ V_rel_o (v_cont E1) (v_cont E2). +Proof. + intros HE. + apply V_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. + apply compat_cont_closed. + apply K_rel_o_elimK. exact HE. exact Hγ. +Qed. + +Lemma compat_shift {V} (e1 e2 : expr (inc V)) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (e_shift e1) (e_shift e2). +Proof. + intros He. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + apply E_rel_intro. iintros M1 M2 E1 E2 HM HE. term_simpl. + eapply O_rel_red_both. + { econstructor. eapply cs_shift. } + { econstructor. eapply cs_shift. } + later_shift. unfold subst. + rewrite -> bind_bind_comp'. + rewrite -> bind_bind_comp'. + rewrite <- (fold_unfold_plug_ectx_hole (bind _ e1)). + rewrite <- (fold_unfold_plug_ectx_hole (bind _ e2)). + apply E_rel_elimO. + - apply E_rel_o_elimE. exact He. + apply compat_comp_subst. + apply compat_cont_closed. exact HE. exact Hγ. + - exact HM. + - apply compat_ectx_hole_closed. +Qed. + +Lemma compat_ectx_app1_closed E1 E2 e1 e2 n : + n ⊨ K_rel E1 E2 → + n ⊨ E_rel e1 e2 → + n ⊨ K_rel (ectx_app1 E1 e1) (ectx_app1 E2 e2). +Proof. + intros HE He. + apply K_rel_intro. + iintros M1 M2 v1 v2 HM Hv. simpl. + apply E_rel_elimO. + - apply compat_app_closed. + apply compat_val_closed. + exact Hv. exact He. + - exact HM. + - exact HE. +Qed. + +Lemma compat_ectx_app1 {V} (E1 E2 : ectx V) e1 e2 n : + n ⊨ K_rel_o E1 E2 → + n ⊨ E_rel_o e1 e2 → + n ⊨ K_rel_o (ectx_app1 E1 e1) (ectx_app1 E2 e2). +Proof. + intros HE He. + apply K_rel_o_intro. + iintros γ1 γ2 Hγ. term_simpl. + apply compat_ectx_app1_closed. + - apply K_rel_o_elimK. exact HE. exact Hγ. + - apply E_rel_o_elimE. exact He. exact Hγ. +Qed. + +Lemma compat_ectx_app2_closed v1 v2 E1 E2 n : + n ⊨ V_rel v1 v2 → + n ⊨ K_rel E1 E2 → + n ⊨ K_rel (ectx_app2 v1 E1) (ectx_app2 v2 E2). +Proof. + intros Hv HE. + apply K_rel_intro. + iintros M1 M2 u1 u2 HM Hu. simpl. + apply E_rel_elimO. + - apply compat_app_closed_val. exact Hv. exact Hu. + - exact HM. + - exact HE. +Qed. + +Lemma compat_ectx_app2 {V} (v1 v2 : val V) E1 E2 n : + n ⊨ V_rel_o v1 v2 → + n ⊨ K_rel_o E1 E2 → + n ⊨ K_rel_o (ectx_app2 v1 E1) (ectx_app2 v2 E2). +Proof. + intros Hv HE. + apply K_rel_o_intro. + iintros γ1 γ2 Hγ. term_simpl. + apply compat_ectx_app2_closed. + - apply V_rel_o_elimV. exact Hv. exact Hγ. + - apply K_rel_o_elimK. exact HE. exact Hγ. +Qed. + +Lemma fundamental_property_e {V} (e : expr V) n : + n ⊨ E_rel_o e e +with fundamental_property_v {V} (v : val V) n : + n ⊨ V_rel_o v v +with fundamental_property_K {V} (E : ectx V) n : + n ⊨ K_rel_o E E. +Proof. + { induction e. + - apply compat_val. apply fundamental_property_v. + - apply compat_app. exact IHe1. exact IHe2. + - apply compat_shift. exact IHe. + - apply compat_reset. exact IHe. } + { induction v. + - apply compat_var. + - apply compat_lambda. apply fundamental_property_e. + - apply compat_cont. apply fundamental_property_K. } + { induction E. + - apply compat_ectx_hole. + - apply compat_ectx_app1. exact IHE. apply fundamental_property_e. + - apply compat_ectx_app2. apply fundamental_property_v. exact IHE. } +Qed. + +Print Assumptions fundamental_property_v. + +(** General program contexts *) +Inductive ctx : Set → Type := +| ctx_hole : ctx ∅ +| ctx_fmap {A B} : (A [→] B) → ctx B → ctx A +| ctx_bind {A B} : (A [⇒] B) → ctx B → ctx A +| ctx_lam {V} : ctx V → ctx (inc V) +| ctx_app1 {V} : ctx V → expr V → ctx V +| ctx_app2 {V} : expr V → ctx V → ctx V +| ctx_shift {V} : ctx V → ctx (inc V) +| ctx_reset {V} : ctx V → ctx V. + +(** Inside-out plugging, always result in a closed expression *) +Fixpoint ciplug {V} (C : ctx V) : expr V → expr ∅ := + match C with + | ctx_hole => λ e, e + | ctx_fmap f C => λ e, ciplug C (fmap f e) + | ctx_bind f C => λ e, ciplug C (bind f e) + | ctx_lam C => λ e, ciplug C (v_lambda e) + | ctx_app1 C e2 => λ e, ciplug C (e_app e e2) + | ctx_app2 e1 C => λ e, ciplug C (e_app e1 e) + | ctx_shift C => λ e, ciplug C (e_shift e) + | ctx_reset C => λ e, ciplug C (e_reset e) + end. + +Inductive ctxr : Set → Set → Type := +| ctxr_hole {A} : ctxr A A +| ctxr_fmap {A B C} : (A [→] B) → ctxr B C → ctxr A C +| ctxr_bind {A B C} : (A [⇒] B) → ctxr B C → ctxr A C +| ctxr_lam {A B} : ctxr A (inc B) → ctxr A B +| ctxr_app1 {A B} : ctxr A B → expr B → ctxr A B +| ctxr_app2 {A B} : expr B → ctxr A B → ctxr A B +| ctxr_shift {A B} : ctxr A (inc B) → ctxr A B +| ctxr_reset {A B} : ctxr A B → ctxr A B. + +(* Outside-in plugging *) +Fixpoint crplug {A B} (C : ctxr A B) : expr A → expr B := + match C with + | ctxr_hole => λ e, e + | ctxr_fmap f C => λ e, crplug C (fmap f e) + | ctxr_bind f C => λ e, crplug C (bind f e) + | ctxr_lam C => λ e, v_lambda (crplug C e) + | ctxr_app1 C e2 => λ e, e_app (crplug C e) e2 + | ctxr_app2 e1 C => λ e, e_app e1 (crplug C e) + | ctxr_shift C => λ e, e_shift (crplug C e) + | ctxr_reset C => λ e, e_reset (crplug C e) + end. + +(* Outside-in plugging simplifies the proofs below *) +Notation cplug := crplug. + +(** Observational approximation for complete programs *) +Definition obs_approx (e1 e2 : expr ∅) : Prop := + terminates e1 → terminates e2. + +(** Observational equivalence for complete programs *) +Definition obs_equiv (e1 e2 : expr ∅) : Prop := + terminates e1 ↔ terminates e2. + +Infix "≼obs" := obs_approx (at level 80, right associativity, only printing). +Infix "≡obs" := obs_equiv (at level 80, right associativity, only printing). + +#[global] +Instance Reflexive_obs_approx : Reflexive obs_approx. +Proof. unfold Reflexive, obs_approx. auto. Qed. + +#[global] +Instance Transitive_obs_approx : Transitive obs_approx. +Proof. unfold Transitive, obs_approx. auto. Qed. + +#[global] +Instance Reflexive_obs_equiv : Reflexive obs_equiv. +Proof. unfold Reflexive, obs_equiv. auto. Qed. + +#[global] +Instance Symmetric_obs_equiv : Symmetric obs_equiv. +Proof. unfold Symmetric, obs_equiv. auto. Qed. + +#[global] +Instance Transitive_obs_equiv : Transitive obs_equiv. +Proof. unfold Transitive, obs_equiv. intuition auto. Qed. + +(** Contextual approximation, where the context closes off Γ *) +Definition ctx_approx {V} (e1 e2 : expr V) : Prop := + ∀ C, obs_approx (cplug C e1) (cplug C e2). + +(** Contextual equivalence *) +Definition ctx_equiv {V} (e1 e2 : expr V) : Prop := + ∀ C, obs_equiv (cplug C e1) (cplug C e2). + +(* TODO *) +(* Infix "≼ctx" := ctx_approx (at level 80, right associativity, only printing). +Infix "≡ctx" := ctx_equiv (at level 80, right associativity, only printing). *) + +#[global] +Instance Reflexive_ctx_approx {V} : Reflexive (@ctx_approx V). +Proof. unfold Reflexive, ctx_approx. reflexivity. Qed. + +#[global] +Instance Transitive_ctx_approx {V} : Transitive (@ctx_approx V). +Proof. unfold Transitive, ctx_approx. etransitivity; eauto. Qed. + +#[global] +Instance Reflexive_ctx_equiv {V} : Reflexive (@ctx_equiv V). +Proof. unfold Reflexive, ctx_equiv. reflexivity. Qed. + +#[global] +Instance Symmetric_ctx_equiv {V} : Symmetric (@ctx_equiv V). +Proof. unfold Symmetric, ctx_equiv. symmetry. auto. Qed. + +#[global] +Instance Transitive_ctx_equiv {V} : Transitive (@ctx_equiv V). +Proof. unfold Transitive, ctx_equiv. etransitivity; eauto. Qed. + +Lemma obs_equiv_intro_approx (e1 e2 : expr ∅) : + obs_approx e1 e2 → + obs_approx e2 e1 → + obs_equiv e1 e2. +Proof. + unfold obs_approx, obs_equiv. done. +Qed. + +Lemma obs_equiv_elim_approx (e1 e2 : expr ∅) : + obs_equiv e1 e2 → + obs_approx e1 e2 ∧ + obs_approx e2 e1. +Proof. + unfold obs_approx, obs_equiv. done. +Qed. + +Lemma ctx_equiv_intro_approx {V} (e1 e2 : expr V) : + ctx_approx e1 e2 → + ctx_approx e2 e1 → + ctx_equiv e1 e2. +Proof. + unfold ctx_approx, ctx_equiv. + intros He1 He2 C. + apply obs_equiv_intro_approx; auto. +Qed. + +Lemma ctx_equiv_elim_approx {V} (e1 e2 : expr V) : + ctx_equiv e1 e2 → + ctx_approx e1 e2 ∧ + ctx_approx e2 e1. +Proof. + unfold ctx_equiv, ctx_approx. + intros He. split. + - intros C. specialize (He C). by apply obs_equiv_elim_approx in He as []. + - intros C. specialize (He C). by apply obs_equiv_elim_approx in He as []. +Qed. + +Definition ciu_approx {V} (e1 e2 : expr V) : Prop := + ∀ E γ, obs_approx (plug E (bind γ e1)) (plug E (bind γ e2)). + +Definition ciu_equiv {V} (e1 e2 : expr V) : Prop := + ∀ E γ, obs_equiv (plug E (bind γ e1)) (plug E (bind γ e2)). + +Lemma ciu_equiv_intro_approx {V} (e1 e2 : expr V) : + ciu_approx e1 e2 → + ciu_approx e2 e1 → + ciu_equiv e1 e2. +Proof. + unfold ciu_approx, ciu_equiv. + intros He1 He2 E γ. + apply obs_equiv_intro_approx; auto. +Qed. + +Lemma ciu_equiv_elim_approx {V} (e1 e2 : expr V) : + ciu_equiv e1 e2 → + ciu_approx e1 e2 ∧ + ciu_approx e2 e1. +Proof. + unfold ciu_approx, ciu_equiv. + intros He. split. + - intros E γ. specialize (He E γ). by apply obs_equiv_elim_approx in He as []. + - intros E γ. specialize (He E γ). by apply obs_equiv_elim_approx in He as []. +Qed. + +Definition n_big_step {V} n (e : expr V) (v : val V) := nsteps meta_contextual_step n e v. +Definition b_big_step {V} n (e : expr V) (v : val V) := bsteps meta_contextual_step n e v. +Definition n_terminates {V} n (e : expr V) := ∃ v, n_big_step n e v. +Definition b_terminates {V} n (e : expr V) := ∃ v, b_big_step n e v. + +Lemma big_step_iff_n_big_step {V} (e : expr V) (v : val V) : + big_step e v ↔ ∃ n, n_big_step n e v. +Proof. + unfold big_step, n_big_step. + apply rtc_nsteps. +Qed. + +Lemma big_step_iff_b_big_step {V} (e : expr V) (v : val V) : + big_step e v ↔ ∃ n, b_big_step n e v. +Proof. + unfold big_step, b_big_step. + apply rtc_bsteps. +Qed. + +Lemma terminates_iff_n_terminates {V} (e : expr V) : + terminates e ↔ ∃ n, n_terminates n e. +Proof. + unfold terminates, n_terminates. + split. + - intros [v H_big_step]. + apply big_step_iff_n_big_step in H_big_step as [n H_n_big_step]. + eauto. + - intros (n & v & H_n_big_step). + exists v. apply big_step_iff_n_big_step. + exists n. exact H_n_big_step. +Qed. + +Lemma terminates_iff_b_terminates {V} (e : expr V) : + terminates e ↔ ∃ n, b_terminates n e. +Proof. + unfold terminates, b_terminates. + split. + - intros [v H_big_step]. + apply big_step_iff_b_big_step in H_big_step as [n H_b_big_step]. + eauto. + - intros (n & v & H_b_big_step). + exists v. apply big_step_iff_b_big_step. + exists n. exact H_b_big_step. +Qed. + +Lemma n_big_step_O_inv {V} e (v : val V) : + n_big_step O e v → e = v. +Proof. + unfold n_big_step. + inversion_clear 1. auto. +Qed. + +Lemma b_big_step_O_inv {V} e (v : val V) : + b_big_step O e v → e = v. +Proof. + unfold b_big_step. + inversion_clear 1. auto. +Qed. + +Lemma n_big_step_S_inv {V} n e (v : val V) : + n_big_step (S n) e v → + ∃ e', meta_contextual_step e e' ∧ n_big_step n e' v. +Proof. + unfold n_big_step. + inversion_clear 1. eauto. +Qed. + +Lemma n_terminates_O_inv {V} (e : expr V) : + n_terminates O e → + ∃ (v : val V), e = v. +Proof. + unfold n_terminates. + intros [v H_n_big_step]. + apply n_big_step_O_inv in H_n_big_step. eauto. +Qed. + +Lemma n_terminates_S_inv {V} n (e : expr V) : + n_terminates (S n) e → + ∃ e', meta_contextual_step e e' ∧ n_terminates n e'. +Proof. + unfold n_terminates. + intros [v H_n_big_step]. + apply n_big_step_S_inv in H_n_big_step as (e' & H_step & H_n_big_step). + eauto. +Qed. + +Lemma L_rel_adequacy_n (e1 e2 : expr ∅) n : + {| nw_index := n |} ⊨ L_rel e1 e2 → + n_terminates n e1 → + terminates e2. +Proof. + revert e1. + induction n as [| n' IHn']; intros e1 He He1. + - apply n_terminates_O_inv in He1 as [v He1]. + apply L_rel_elim in He as [He _]. eauto. + - apply n_terminates_S_inv in He1 as (e' & H_step & He1). + apply (IHn' e'); [| exact He1]. + apply I_world_lift_later. + apply L_rel_elim in He as [_ He]. iapply He. + iintro. exact H_step. +Qed. + +Lemma L_rel_adequacy (e1 e2 : expr ∅) : + (∀ n, n ⊨ L_rel e1 e2) → + terminates e1 → + terminates e2. +Proof. + intros He He1. + apply terminates_iff_n_terminates in He1 as (n & He1). + eapply L_rel_adequacy_n. + - apply He. + - exact He1. +Qed. + +Theorem O_rel_adequacy e1 e2 : + (∀ n, n ⊨ O_rel e1 e2) → obs_equiv e1 e2. +Proof. + intros He. split. + - apply L_rel_adequacy. + intros n. specialize (He n). + by apply O_rel_elim in He as []. + - apply L_rel_adequacy. + intros n. specialize (He n). + by apply O_rel_elim in He as []. +Qed. + +Lemma fundamental_property_v_closed (v : val ∅) n : + n ⊨ V_rel v v. +Proof. + rewrite <- (bind_pure' v). + apply V_rel_o_elimV. + - apply fundamental_property_v. + - apply compat_id_subst. +Qed. + +Lemma fundamental_property_e_closed (e : expr ∅) n : + n ⊨ E_rel e e. +Proof. + rewrite <- (bind_pure' e). + apply E_rel_o_elimE. + - apply fundamental_property_e. + - apply compat_id_subst. +Qed. + +Lemma fundamental_property_K_closed E n : + n ⊨ K_rel E E. +Proof. + rewrite <- (bind_pure' E). + apply K_rel_o_elimK. + - apply fundamental_property_K. + - apply compat_id_subst. +Qed. + +Lemma fundamental_property_g {V} (γ : V [⇒] ∅) n : + n ⊨ G_rel γ γ. +Proof. + apply G_rel_intro. + iintros x. apply fundamental_property_v_closed. +Qed. + +Lemma compat_expr_o_closed (e1 e2 : expr ∅) n : + n ⊨ E_rel e1 e2 → + n ⊨ O_rel e1 e2. +Proof. + intros He. + rewrite <- (fold_unfold_meta_plug_nil e1). + rewrite <- (fold_unfold_meta_plug_nil e2). + rewrite <- (fold_unfold_plug_ectx_hole e1). + rewrite <- (fold_unfold_plug_ectx_hole e2). + apply E_rel_elimO. + - exact He. + - apply compat_meta_ectx_nil_closed. + - apply compat_ectx_hole_closed. +Qed. + +Lemma compat_expr_o {V} (e1 e2 : expr V) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ O_rel_o e1 e2. +Proof. + intros He. + apply O_rel_o_intro. iintros γ1 γ2 Hγ. + apply compat_expr_o_closed. + apply E_rel_o_elimE; assumption. +Qed. + +Lemma compat_fmap {A B} (f : A [→] B) e1 e2 n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (fmap f e1) (fmap f e2). +Proof. + intros He. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + rewrite -> (map_to_bind f). + rewrite -> (map_to_bind f). + rewrite -> bind_bind_comp'. + rewrite -> bind_bind_comp'. + apply E_rel_o_elimE. exact He. + apply G_rel_intro. iintros x. term_simpl. + apply G_rel_elimV. exact Hγ. +Qed. + +Lemma compat_bind {A B} (f : A [⇒] B) e1 e2 n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (bind f e1) (bind f e2). +Proof. + intros He. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + rewrite -> bind_bind_comp'. + rewrite -> bind_bind_comp'. + apply E_rel_o_elimE. exact He. + apply G_rel_intro. iintros x. term_simpl. + apply V_rel_o_elimV. + - apply fundamental_property_v. + - exact Hγ. +Qed. + +Lemma precongruence {A B} (e1 e2 : expr A) (C : ctxr A B) n : + n ⊨ E_rel_o e1 e2 → + n ⊨ E_rel_o (cplug C e1) (cplug C e2). +Proof. + intros He. + induction C; simpl. + - exact He. + - apply IHC. apply compat_fmap. exact He. + - apply IHC. apply compat_bind. exact He. + - apply compat_val. + apply compat_lambda. + apply IHC. exact He. + - apply compat_app. + + apply IHC. exact He. + + apply fundamental_property_e. + - apply compat_app. + + apply fundamental_property_e. + + apply IHC. exact He. + - apply compat_shift. + apply IHC. exact He. + - apply compat_reset. + apply IHC. exact He. +Qed. + +(* +Theorem E_rel_o_soundness' {V} (e1 e2 : expr V) : + (∀ n, n ⊨ E_rel_o e1 e2) → + ciu_equiv e1 e2. +Proof. + unfold ciu_equiv. + intros He E γ. + apply O_rel_adequacy. intros n. + apply E_rel_elimO. + - specialize (He n). + apply E_rel_o_elimE. + + exact He. + + apply fundamental_property_g. + - apply fundamental_property_k_closed. +Qed. + +Lemma L_rel_obs_approx_trans e1 e2 e3 n : + n ⊨ L_rel e1 e2 → + obs_approx e2 e3 → + n ⊨ L_rel e1 e3. +Proof. + unfold obs_approx. + intros He1 He2. + irevert e1 He1. loeb_induction IH. + iintros e1 He1. + apply L_rel_elim in He1 as [He1_val He1_step]. + apply L_rel_intro. + - intros v Hv. exact (He2 (He1_val v Hv)). + - iintros e1' H_step. + ispec He1_step e1' H_step. + later_shift. iapply IH. + exact He1_step. +Qed. + +Lemma O_rel_obs_equiv_trans e1 e2 e3 n : + n ⊨ O_rel e1 e2 → + obs_equiv e2 e3 → + n ⊨ O_rel e1 e3. +Proof. + intros He1 He2. + apply O_rel_elim in He1 as [He1_l He1_r]. + apply obs_equiv_elim_approx in He2 as [He2_l He2_r]. + apply O_rel_intro. + - by eapply L_rel_obs_approx_trans. + - (* This direction is impossible to prove: + If e3 terminates in n steps, we need to prove that e1 MUST + terminate. Using He2_r, we can conclude that e2 terminates + in some m steps. However, if m > n, then He1_r does not imply + that e1 terminates. *) +Abort. + +Theorem L_rel_completeness (e1 e2 : expr ∅) n : + obs_approx e1 e2 → + n ⊨ L_rel e1 e2. +Proof. + intros He. irevert e1 He. + loeb_induction IH. + iintros e1 He. + apply L_rel_intro. + - intros v ->. apply He. apply terminates_val. + - iintros e1' H_step. idestruct H_step. + later_shift. + assert (He1' : obs_approx e1' e2). + { unfold obs_approx in *. + intros He1'. apply He. + by eapply contextual_step_terminates. } + ispec IH e1' He1'. + exact IH. +Qed. + +Theorem O_rel_completeness (e1 e2 : expr ∅) n : + obs_equiv e1 e2 → + n ⊨ O_rel e1 e2. +Proof. + intros He. + apply obs_equiv_elim_approx in He as [He1 He2]. + apply O_rel_intro. + - by apply L_rel_completeness. + - by apply L_rel_completeness. +Qed. + +Theorem E_rel_o_completeness' {V} (e1 e2 : expr V) n : + ciu_equiv e1 e2 → + n ⊨ E_rel_o e1 e2. +Proof. + intros He. + apply ciu_equiv_elim_approx in He as [He1 He2]. + unfold ciu_approx in *. + apply E_rel_o_intro. iintros γ1 γ2 Hγ. + apply E_rel_intro. iintros E1 E2 HE. + apply O_rel_intro. + - assert (He1' : n ⊨ L_rel (plug E1 (bind γ1 e1)) (plug E2 (bind γ2 e1))). + { apply O_rel_elim1. + apply E_rel_elimO. + apply E_rel_o_elimE. + apply fundamental_property_e. + exact Hγ. exact HE. } + specialize (He1 E2 γ2). + exact (L_rel_obs_approx_trans _ _ _ _ He1' He1). + - assert (He2' : n ⊨ L_rel (plug E2 (bind γ2 e2)) (plug E1 (bind γ1 e2))). + { apply O_rel_elim2. + apply E_rel_elimO. + apply E_rel_o_elimE. + apply fundamental_property_e. + exact Hγ. exact HE. } + specialize (He2 E1 γ1). + exact (L_rel_obs_approx_trans _ _ _ _ He2' He2). +Qed. +*) + +Theorem E_rel_o_soundness {V} (e1 e2 : expr V) : + (∀ n, n ⊨ E_rel_o e1 e2) → + ctx_equiv e1 e2. +Proof. + unfold ctx_equiv. + intros He C. + apply O_rel_adequacy. intros n. + specialize (He n). + apply (precongruence _ _ C) in He. + apply compat_expr_o in He. + rewrite <- (bind_pure' (cplug C e1)). + rewrite <- (bind_pure' (cplug C e2)). + apply O_rel_o_elimO. + - exact He. + - apply compat_id_subst. +Qed. + +(* +Theorem ciu_equiv_soundness {V} (e1 e2 : expr V) : + ciu_equiv e1 e2 → + ctx_equiv e1 e2. +Proof. +Abort. +*) + +Fixpoint rctx_to_ctxr {V} (R : rctx V) : ctxr V V := + match R with + | rctx_hole => ctxr_hole + | rctx_app1 R e => ctxr_app1 (rctx_to_ctxr R) e + | rctx_app2 v R => ctxr_app2 v (rctx_to_ctxr R) + end. + +Lemma rctx_to_ctxr_correct {V} (R : rctx V) (e : expr V) : + crplug (rctx_to_ctxr R) e = rplug R e. +Proof. + induction R. + - simpl. reflexivity. + - simpl. rewrite -> IHR. reflexivity. + - simpl. rewrite -> IHR. reflexivity. +Qed. + +Theorem ciu_equiv_completeness {V} (e1 e2 : expr V) : + ctx_equiv e1 e2 → + ciu_equiv e1 e2. +Proof. + unfold ctx_equiv, ciu_equiv. + intros He E γ. + rewrite <-2 ectx_to_rctx_correct. + rewrite <-2 rctx_to_ctxr_correct. + specialize (He (ctxr_bind γ (rctx_to_ctxr (ectx_to_rctx E)))). + simpl in He. exact He. +Qed. + +(* +(* composition of subst and subst_map, + where the variable to be substituted doesn't appear in the substitution *) +Lemma subst_subst_map : ∀ (e:expr) Γ (x : string) (es : val) (map : sub), + subst_is_closed Γ ∅ map → + subst x es (subst_map (delete x map) e) = + subst_map (insert x es map) e +with subst_subst_map_val : ∀ (v:val) Γ (x : string) (es : val) (map : sub), + subst_is_closed Γ ∅ map → + subst x es (subst_map_val (delete x map) v) = + subst_map_val (insert x es map) v. +Proof. + { intros e. induction e. + { intros. apply (subst_subst_map_val _ _ _ _ _ H). } + { (* e is a variable x *) + intros. simpl. destruct (decide (x0=x)) as [->|Hne]. + { (* if x=x0, we'll end up substituting es into x *) + rewrite lookup_delete_eq with (m:=map). + rewrite lookup_insert_eq with (m:=map). + simpl. + by rewrite decide_True. } + { (* if not equal, the deletion and insertion will have no effect *) + rewrite lookup_delete_ne with (m:=map). 2: { assumption. } + rewrite lookup_insert_ne with (m:=map). 2: { assumption. } + (* we then need to see if x is in the map to begin with *) + destruct (map !! x) as [v1|] eqn:Hkey. + { Fail rewrite Hkey. (* why does regular rewrite not work? *) + setoid_rewrite Hkey. + simpl. + rewrite (subst_val_closed _ ∅ _ _). + - reflexivity. + - apply (subst_is_closed_elim_closed _ _ x _ _ H Hkey). + - set_solver. } + { setoid_rewrite Hkey. + simpl. + by rewrite decide_False. } } } + { intros. simpl. f_equal; eauto. } } + { intros v. induction v; intros. + (*{ reflexivity. }*) + { (* the lambda case *) + simpl. f_equal. f_equal. + case_decide. + { subst. + rewrite delete_delete_eq with (m:=map). + rewrite delete_insert_eq with (m:=map). done. } + { rewrite delete_insert_ne with (m:=map). 2: { congruence. } + rewrite delete_delete with (m:=map). + eapply subst_subst_map. + apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x map) map). + destruct H as [H1 H2]. + rewrite H1. + set_solver. + apply delete_subseteq. + set_solver. + assumption. } } + (*{ reflexivity. }*) } +Qed. + +(** Special case of Theorem A.1 from Erlang paper: + scoping of extended substitutions. + Given a closed substitution, we can add a closed value to it. *) +Lemma scope_extend1 Γ x (v:val) (γ:sub): + closed ∅ v → + subst_is_closed Γ ∅ γ → + subst_is_closed (Γ ∪ {[x]}) ∅ (<[x := v]> γ). +Proof. + intros Hc Hsc. + split. + { destruct Hsc. rewrite H. set_solver. } + intros x0 Hd v0 Hs. + (* we have to prove that for an arbitrary binding x0 := v0 in γ, v0 is closed *) + destruct (decide (x=x0)) as [->|Hne]. + (* if x = x0, the premise tells us v0 is closed *) + { rewrite lookup_insert_eq with (m:=γ) in Hs. + injection Hs. intros. subst. + exact Hc. } + (* if they are not equal, we know x0 is in Γ and have to use the fact + that the subst_is_closed *) + { rewrite elem_of_union in Hd. destruct Hd. 2: { assert (x0 = x). set_solver. done. } + destruct Hsc as [_ Hsc]. + rewrite lookup_insert_ne with (m:=γ) in Hs. 2: { assumption. } + specialize (Hsc x0 H v0 Hs). + assumption. } +Qed. + +(* we can extend related substitutions with related values *) +Lemma sem_context_rel_insert Γ x v1 v2 γ1 γ2 n: + n ⊨ V_rel v1 v2 → + n ⊨ G_rel Γ γ1 γ2 → + n ⊨ G_rel (Γ ∪ {[x]}) (<[x := v1]> γ1) (<[x := v2]> γ2). +Proof. + intros Hv Hγ. + destruct (V_rel_elim _ _ _ Hv) as (Hvc1 & Hvc2 & Hv'). + destruct (G_rel_elim _ _ _ _ Hγ) as (Hγc1 & Hγc2 & Hγ'). + apply G_rel_intro. + { by apply scope_extend1. } + { by apply scope_extend1. } + + iintros x0 v0 v3 H1 H2. + destruct (decide (x=x0)). + { subst. + rewrite lookup_insert_eq with (m:=γ2) in H2. idestruct H2. injection H2 as ->. + rewrite lookup_insert_eq with (m:=γ1) in H1. idestruct H1. injection H1 as ->. + assumption. } + { rewrite lookup_insert_ne with (m:=γ2) in H2. idestruct H2. 2: { assumption. } + rewrite lookup_insert_ne with (m:=γ1) in H1. idestruct H1. 2: { assumption. } + iapply Hγ'. + - iintro. eassumption. + - iintro. eassumption. } +Qed. + +Lemma lambda_closed Γ γ x e : + closed (Γ ∪ {[x]}) e → + subst_is_closed Γ ∅ γ → + closed ∅ (vlambda x (subst_map (delete x γ) e)). +Proof. + intros Hec [Heq Hγc]. + rewrite closed_lambda. + eapply subst_map_closed'_2. + - eapply closed_weaken. + exact Hec. + setoid_rewrite dom_delete. + intros y. destruct (decide (x = y)); set_solver. + - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x γ) γ). + + set_solver. + + apply delete_subseteq. + + set_solver. + + unfold subst_is_closed. split. + * exact Heq. + * intros x' Hin v Hlookup. + specialize (Hγc x' Hin v Hlookup). + by eapply closed_weaken. +Qed. +*) + +(* +Notation name := string. +Definition sub : Set := gmap name val. +Definition scope : Set := gset name. + +Fixpoint subst_map_val (γ : sub) (v : val) : val := + match v with + (*| vunit => vunit + | vint n => vint n*) + | vlambda x e => vlambda x (subst_map (delete x γ) e) + end +with subst_map (γ : sub) (e : expr) : expr := + match e with + | ret v => ret (subst_map_val γ v) + (* | eunit => eunit *) + | var x => match γ !! x with Some v => ret v | _ => var x end + | app e1 e2 => app (subst_map γ e1) (subst_map γ e2) + (* | abs x e => abs x (subst_map (binder_delete x xs) e) *) + (* | eplus e1 e2 => eplus (subst_map xs e1) (subst_map xs e2) *) + end. + +Lemma fold_unfold_subst_map_val_vlambda γ x e : + subst_map_val γ (vlambda x e) = + vlambda x (subst_map (delete x γ) e). +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_ret γ v : + subst_map γ (ret v) = + ret (subst_map_val γ v). +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_var γ x : + subst_map γ (var x) = + match γ !! x with + | Some v => ret v + | _ => var x + end. +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_app γ e1 e2 : + subst_map γ (app e1 e2) = + app (subst_map γ e1) (subst_map γ e2). +Proof. auto. Qed. + +Fixpoint is_closed (X : scope) (e : expr) : bool := + match e with + | var x => bool_decide (x ∈ X) + (*| ret vunit | ret (vint _) => true*) + | ret (vlambda x e) => is_closed (X ∪ {[x]}) e + | app e1 e2 + (* | eplus e1 e2 *) + => is_closed X e1 && is_closed X e2 + end. + +(* aka a scoping judgment, the untyped equivalent of a typing judgment *) +Definition closed (X : scope) (e : expr) : Prop := Is_true (is_closed X e). + +Lemma closed_weaken e X Y: + closed X e → X ⊆ Y → closed Y e +with closed_weaken_val (v:val) X Y: + closed X v → X ⊆ Y → closed Y v. +Proof. + { revert X Y. + induction e; intros. + - apply (closed_weaken_val _ _ _ H H0). + - unfold closed, is_closed in *. + apply bool_decide_unpack in H. + apply bool_decide_pack. + set_solver. + - unfold closed in *. simpl in *. + apply andb_prop_intro. + apply andb_prop_elim in H. + destruct H. + split. + apply (IHe1 _ _ H H0). + apply (IHe2 _ _ H1 H0). } + { revert X Y. + induction v; intros. + (*- constructor.*) + - unfold closed in *. + simpl in *. + apply (closed_weaken e _ _ H). + set_solver. + (*- constructor.*) } +Qed. + +Lemma closed_subst : + ∀ Γ x e1 e2, + closed (Γ ∪ {[x]}) e1 → + closed ∅ e2 → + closed Γ (subst x e2 e1) +with closed_subst_val : + ∀ Γ x (v : val) e, + closed (Γ ∪ {[x]}) v → + closed ∅ e → + closed Γ (subst_val x e v). +Proof. + { + unfold closed in *. + intros Γ x e1 e2 H_closed1 H_closed2. + induction e1. + - simpl in *. auto. + - simpl in *. + rewrite -> decide_bool_decide. + destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. + + eapply closed_weaken. + exact H_closed2. + set_solver. + + simpl in *. + apply bool_decide_unpack in H_closed1. + apply bool_decide_pack. + set_solver. + - simpl in *. + apply andb_prop_elim in H_closed1 as [Hc1 Hc2]. + apply andb_prop_intro. split. + + exact (IHe1_1 Hc1). + + exact (IHe1_2 Hc2). + } + { + unfold closed in *. + intros Γ x v e H_closed1 H_closed2. + induction v. + (* - simpl in *. auto. *) + - simpl in *. + rewrite -> decide_bool_decide. + destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. + + rewrite -> H_eq in H_closed1. + eapply closed_weaken. + exact H_closed1. set_solver. + + simpl in *. + apply closed_subst. + eapply closed_weaken. + exact H_closed1. set_solver. + exact H_closed2. + (*- simpl in *. auto.*) + } +Qed. + +Lemma closed_app xs e1 e2: + closed xs (app e1 e2) ↔ + closed xs e1 ∧ closed xs e2. +Proof. unfold closed. simpl. by rewrite andb_True. Qed. + +Lemma closed_lambda e Γ x : closed Γ (vlambda x e) ↔ closed (Γ ∪ {[x]}) e. +Proof. split. auto. auto. Qed. + +Lemma closed_var Γ x : closed Γ (var x) ↔ x ∈ Γ. +Proof. unfold closed. simpl. by rewrite bool_decide_spec. Qed. + +Lemma base_step_preserve_closedness : + ∀ e1 e1', + base_step e1 e1' → + closed ∅ e1 → + closed ∅ e1'. +Proof. + unfold closed. + intros e1 e1' Hred H_closed. + inversion Hred. subst. simpl in *. + apply andb_prop_elim in H_closed as [Hc1 Hc2]. + exact (closed_subst ∅ x e0 e2 Hc1 Hc2). +Qed. + +Definition ectx_is_closed (X : scope) (E : ectx) := + ∀ e, closed ∅ e → closed X (fill E e). + +Lemma closed_decompose : + ∀ E e, + closed ∅ (fill E e) → + ectx_is_closed ∅ E ∧ closed ∅ e. +Proof. + unfold ectx_is_closed. + intros E. + induction E; intros e' H_closed. + - simpl in *. auto. + - simpl in *. + apply IHE in H_closed as [H_fill H_closed]. + apply closed_app in H_closed as [H_closed1 H_closed2]. + split. + + intros e'' H_closed3. + apply H_fill. + apply closed_app. auto. + + auto. + - simpl in *. + apply IHE in H_closed as [H_fill H_closed]. + apply closed_app in H_closed as [H_closed1 H_closed2]. + split. + + intros e'' H_closed3. + apply H_fill. + apply closed_app. auto. + + auto. +Qed. + +Lemma closed_compose : + ∀ E e, + ectx_is_closed ∅ E → + closed ∅ e → + closed ∅ (fill E e). +Proof. unfold ectx_is_closed. auto. Qed. + +Lemma contextual_step_preserve_closedness : + ∀ e1 e1', + contextual_step e1 e1' → + closed ∅ e1 → + closed ∅ e1'. +Proof. + unfold closed. + intros e1 e1' Hred H_closed. + inversion Hred. subst. simpl in *. clear Hred. + apply closed_decompose in H_closed as [H_closed1 H_closed2]. + apply (base_step_preserve_closedness _ _ H1) in H_closed2. + apply closed_compose; auto. +Qed. + +(** subscoped from the Erlang paper *) +Definition subst_is_closed (Γ free : scope) (sub : sub) := + Γ = dom sub ∧ + ∀ x, x ∈ Γ → + ∀ v, sub !! x = Some v → closed free (ret v). + +Lemma subst_is_closed_subseteq: ∀ (Γ1 Γ2 X : scope) (γ1 γ2: sub), + Γ1 = dom γ1 → + γ1 ⊆ γ2 → Γ1 ⊆ Γ2 → subst_is_closed Γ2 X γ2 → subst_is_closed Γ1 X γ1. +Proof. + intros * Hd Hγ HΓ Hclosed2. + destruct Hclosed2 as [Hd2 Hc2]. + split. + assumption. + intros x Hl v Hs. + rewrite (map_subseteq_spec γ1 γ2) in Hγ. + (* specialize (Hγ _ _ Hs). *) + specialise Hγ Hs. + apply (Hc2 x ltac:(set_solver) _ Hγ). +Qed. + +Lemma closed_ectx_app1 : + ∀ E e, + ectx_is_closed ∅ E → + closed ∅ e → + ectx_is_closed ∅ (ectx_app1 E e). +Proof. + intros E e Hc1 Hc2. + unfold ectx_is_closed in *. + intros e' Hc3. simpl. + apply Hc1. apply closed_app. auto. +Qed. + +Lemma closed_ectx_app2 : + ∀ (v : val) E, + closed ∅ v → + ectx_is_closed ∅ E → + ectx_is_closed ∅ (ectx_app2 v E). +Proof. + intros v E Hc1 Hc2. + unfold ectx_is_closed in *. + intros e' Hc3. simpl. + apply Hc2. apply closed_app. auto. +Qed. + +Lemma subst_is_closed_elim_closed Γ (γ:sub) x X (v:val): + subst_is_closed Γ X γ → + γ !! x = Some v → + closed X v. +Proof. + intros [Hdom Hsc] He. + assert (H := elem_of_dom_2 _ _ _ He). + (* have: elem_of_dom_2 He. *) + (* pose proof (elem_of_dom_2 _ _ _ He). *) + assert (x ∈ Γ). set_solver. + apply (Hsc x H0 v He). +Qed. + +(* if e is closed under Y, we can split the variables in Y between X and γ *) +Lemma subst_map_closed' (e : expr) (X Y : scope) (γ:sub) : + closed Y e → + (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → + closed X (subst_map γ e) +with subst_map_closed'_val (v : val) (X Y : scope) (γ:sub) : + closed Y (ret v) → + (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → + closed X (subst_map_val γ v). +Proof. + { + revert X Y γ. induction e. + { eapply subst_map_closed'_val; eauto. } + { intros * Hc H. + (* e is a variable x *) + unfold closed in Hc; simpl in Hc; apply bool_decide_unpack in Hc. + specialize (H x Hc). simpl. + destruct (γ !! x) eqn:He. + - assumption. + - unfold closed; simpl; apply bool_decide_pack. assumption. + } + { intros *. + unfold closed. simpl. + rewrite !andb_True. intros [? ?] **. + split. + by eapply IHe1. + by eapply IHe2. } + } + { revert X Y γ. induction v. + (*{ intros. assumption. }*) + { unfold closed. simpl. + intros * Hce H. + eapply subst_map_closed'. eassumption. + intros y [|]%elem_of_union. + { destruct (decide (x = y)). + { by subst; rewrite lookup_delete_eq with (m:=γ); set_solver. } + rewrite lookup_delete_ne with (m:=γ). 2: { assumption. } + eapply H in H0. + destruct lookup; last set_solver. + eapply closed_weaken; eauto with set_solver. } + { rewrite elem_of_singleton in H0. + subst. rewrite lookup_delete_eq with (m:=γ). set_solver. } + } + (*{ intros. assumption. }*) } +Qed. + +(* simple corollary of [subst_map_closed'], + where we have split Y into X and dom γ upfront *) +Lemma subst_map_closed'_2 Γ X γ e : + closed (X ∪ (dom γ)) e -> + subst_is_closed Γ X γ -> + closed X (subst_map γ e). +Proof. + intros Hcl Hsubst. + eapply subst_map_closed'; first eassumption. + intros x Hx. + destruct (γ !! x) as [e'|] eqn:Heq. + - apply (subst_is_closed_elim_closed _ _ _ _ _ Hsubst Heq). + - apply not_elem_of_dom in Heq. + set_solver. +Qed. + +(* given a value and a substitution closed under the same scope, + applying the substitution gives us a completely closed value *) +Lemma subst_map_closed'_3 (v:val) Γ γ: + closed Γ v -> + subst_is_closed Γ ∅ γ -> + closed ∅ (subst_map γ v). +Proof. + pose proof (subst_map_closed'_2 Γ ∅ γ v). + simpl in H. + intros. + apply H. 2: { assumption. } + destruct H1 as [? _]. + rewrite <- H1. + replace (∅ ∪ Γ) with Γ. assumption. + set_solver. +Qed. + +Lemma subst_map_val_closed_val_aux Γ γ (v : val) : + closed Γ v → + Γ ∩ dom γ = ∅ → + subst_map_val γ v = v +with subst_map_closed_aux Γ γ (e : expr) : + closed Γ e → + Γ ∩ dom γ = ∅ → + subst_map γ e = e. +Proof. + { intros Hc Hdom. + induction v. + - rewrite -> closed_lambda in Hc. + rewrite -> fold_unfold_subst_map_val_vlambda. + rewrite -> (subst_map_closed_aux (Γ ∪ {[x]}) (delete x γ) e Hc ltac:(set_solver)). + reflexivity. } + { intros Hc Hdom. + induction e. + - rewrite -> fold_unfold_subst_map_ret. + f_equal. by eapply subst_map_val_closed_val_aux. + - unfold closed in Hc. + simpl in Hc. + rewrite -> bool_decide_spec in Hc. + rewrite -> fold_unfold_subst_map_var. + assert (H_not_in : x ∉ dom γ) by set_solver. + rewrite -> (not_elem_of_dom γ x) in H_not_in. + setoid_rewrite -> H_not_in. + reflexivity. + - apply closed_app in Hc as [Hc1 Hc2]. + rewrite -> fold_unfold_subst_map_app. + rewrite -> (IHe1 Hc1). + rewrite -> (IHe2 Hc2). + reflexivity. } +Qed. + +Lemma subst_map_val_closed_val γ (v : val) : + closed ∅ v → + subst_map_val γ v = v. +Proof. + intros Hc. + eapply subst_map_val_closed_val_aux. + - exact Hc. + - set_solver. +Qed. + +Lemma subst_map_closed γ (e : expr) : + closed ∅ e → + subst_map γ e = e. +Proof. + intros Hc. + eapply subst_map_closed_aux. + - exact Hc. + - set_solver. +Qed. +*) + +(* +Lemma subset_is_closed_absurd x Γ γ: + x ∈ Γ → + subst_is_closed Γ ∅ γ → + γ !! x = None → + False. +Proof. + intros Hx Hs He. + pose proof (not_elem_of_dom_2 _ _ He). + destruct Hs as [? _]. + setoid_rewrite <- H0 in H. + set_solver. +Qed. +*) + +(* +Lemma subst_val_closed v X x es : + closed X (of_val v) → x ∉ X → subst_val x es v = v +with subst_closed X e x es : + closed X e → x ∉ X → subst x es e = e. +Proof. + { induction v. + (*{ reflexivity. }*) + { simpl. + case_decide. + - reflexivity. + - intros. + f_equal. + rewrite closed_lambda in H0. + apply (subst_closed _ _ _ _ H0). + set_solver. } + (*{ reflexivity. }*) } + { induction e; intros; simpl. + { f_equal. + eapply subst_val_closed. + apply H. assumption. } + { case_decide. + - subst. + unfold closed in H. simpl in H. apply bool_decide_unpack in H. + set_solver. + - reflexivity. } + { apply closed_app in H. + destruct H as (H1&H2). + specialize (IHe1 H1 H0). + specialize (IHe2 H2 H0). + f_equal. + - assumption. + - assumption. } } +Qed. +*) + +(* +(* aka contextual scoping C : Γ ~> Γ', a special case of contextual typing. + defined inductively because we need to invert it. *) +Inductive closed_ctx : scope → scope → ctx → Prop := + | cc_hole Γ : + closed_ctx Γ Γ ctx_hole + + | cc_lambda x Γ Γ' C : + closed_ctx Γ (Γ' ∪ {[x]}) C → + closed_ctx Γ Γ' (ctx_lam x C) + + | cc_app1 Γ Γ' C e : + closed_ctx Γ Γ' C → + closed Γ' e → + closed_ctx Γ Γ' (ctx_app1 C e) + + | cc_app2 Γ Γ' C v : + closed_ctx Γ Γ' C → + closed Γ' (ret v) → + closed_ctx Γ Γ' (ctx_app2 v C) + . +*) + +(* +Definition closed_ctx_sem (Γ Γ' : scope) (C:ctx) : Prop := + forall e, closed Γ e → closed Γ' (cplug C e). + +Lemma closed_ctx_sound Γ Γ' e : + closed_ctx Γ Γ' e → closed_ctx_sem Γ Γ' e. +Proof. + intros H. induction H; unfold closed_ctx_sem. + - simpl. done. + - intros e Hc. + simpl. + specialize (IHclosed_ctx e Hc). + unfold closed. simpl. + apply IHclosed_ctx. + - intros e2 Hc. + specialize (IHclosed_ctx _ Hc). + simpl. + unfold closed. simpl. + unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. + auto. + - intros e1 Hc. + specialize (IHclosed_ctx _ Hc). + simpl. + unfold closed. simpl. + unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. + auto. +Qed. +*) + +(* +Lemma G_rel_empty n : + n ⊨ G_rel ∅ ∅ ∅. +Proof. + apply G_rel_intro. + - unfold subst_is_closed. split; set_solver. + - unfold subst_is_closed. split; set_solver. + - iintros. + idestruct H. + setoid_rewrite lookup_empty in H. + discriminate H. +Qed. + +Lemma subst_map_empty (e:expr) : + subst_map ∅ e = e +with subst_map_val_empty (v:val) : + subst_map_val ∅ v = v. +Proof. + { induction e. + - simpl. f_equal. eapply subst_map_val_empty. + - simpl. + setoid_rewrite lookup_empty. + reflexivity. + - simpl. rewrite IHe1. rewrite IHe2. reflexivity. } + { induction v. + - simpl. + f_equal. + setoid_rewrite delete_empty. + apply subst_map_empty. } +Qed. + +Lemma fundamental_property_sub Γ γ n : + subst_is_closed Γ ∅ γ → + n ⊨ G_rel Γ γ γ. +Proof. + intros Hc. + apply G_rel_intro. + { exact Hc. } + { exact Hc. } + iintros x v1 v2 Heq1 Heq2. + idestruct Heq1. + idestruct Heq2. + rewrite -> Heq1 in Heq2. + injection Heq2 as ->. + destruct Hc as [-> Hc]. + assert (Hcv2 : closed ∅ v2). + { eapply Hc. + - setoid_rewrite -> elem_of_dom. + unfold is_Some. eauto. + - exact Heq1. } + assert (H_fundamental := fundamental_property_v ∅ v2 n Hcv2). + apply V_rel_o_elim in H_fundamental. + ispec H_fundamental ∅ ∅ (G_rel_empty n). + rewrite -> subst_map_val_empty in H_fundamental. + exact H_fundamental. +Qed. + +Lemma fundamental_property_ectx E n : + n ⊨ K_rel E E. +Proof. + apply K_rel_intro. + iintros v1 v2 Hv. + admit. +Admitted. +*) diff --git a/ctx-equiv-ixfree/propriety.v b/ctx-equiv-ixfree/propriety.v index 59a4b46..cc43e00 100644 --- a/ctx-equiv-ixfree/propriety.v +++ b/ctx-equiv-ixfree/propriety.v @@ -1,4 +1,4 @@ - +(* From IxFree Require Import Nat Lib. From CtxEquivIxFree Require Import lang. @@ -99,3 +99,4 @@ Proof. apply cexpr_closed. assumption. Qed. +*) From d426c0a87470f3ac839aa5a0d7d0ebba55ea2dce Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Fri, 5 Dec 2025 18:27:23 +0800 Subject: [PATCH 08/11] Finish proving unique decomposition. The main insight is that: the proof of unique decomposition basically contains an algorithm to "construct" the unique decomposition: forall e, (exists v, e = v) \/ (exists! M R e', e = meta_rplug M (rplug R e')) Notice the quantifier `exists`. In Rocq, we prove `exists` by giving the witness, and we can only give the witness if we can construct it. Therefore, a better idea is to extract the decomposition as a separate function, which is much easier to study; and is useful to state our other lemmas as well. After obtaining the decomposition function, we prove unique decomposition by "delegating" to this decomposing function. We just need to prove that, given a redex, an evaluation context and a meta context, plugging them and decomposing them is identity. The prove is by pure equational reasoning (!) and it much easier to understand. --- ctx-equiv-ixfree/lang_shift_reset.v | 313 +++++++++++++++++++++------- 1 file changed, 240 insertions(+), 73 deletions(-) diff --git a/ctx-equiv-ixfree/lang_shift_reset.v b/ctx-equiv-ixfree/lang_shift_reset.v index e38c179..448122d 100644 --- a/ctx-equiv-ixfree/lang_shift_reset.v +++ b/ctx-equiv-ixfree/lang_shift_reset.v @@ -314,6 +314,13 @@ Proof. exact (ectx_comp_rctx2_correct E rctx_hole e). Qed. +Lemma rctx_to_ectx_correct {V} (R : rctx V) (e : expr V) : + plug (rctx_to_ectx R) e = rplug R e. +Proof. + unfold rctx_to_ectx. + exact (ectx_comp_rctx1_correct ectx_hole R e). +Qed. + Lemma plug_eq_val_inv {V} E e (v : val V) : plug E e = v → E = ectx_hole ∧ e = v. @@ -586,6 +593,10 @@ Proof. auto. Qed. Definition meta_ectx_comp {V} (M1 M2 : meta_ectx V) : meta_ectx V := M2 ++ M1. +Lemma fold_unfold_meta_ectx_comp {V} (M1 M2 : meta_ectx V) : + meta_ectx_comp M1 M2 = M2 ++ M1. +Proof. auto. Qed. + Lemma meta_plug_eq_val_inv {V} (M : meta_ectx V) (e : expr V) (v : val V) : meta_plug M e = v → M = [] ∧ e = v. @@ -655,6 +666,99 @@ Proof. auto_bind_comp. Qed. #[global] Instance Bind_meta_ectx : Bind meta_ectx. Proof. constructor. exact @meta_ectx_bind_id. exact @meta_ectx_bind_comp. Qed. +(** Outside-in meta context *) + +Definition meta_rctx V := list (rctx V). + +Fixpoint meta_rplug {V} (M : meta_rctx V) (e : expr V) : expr V := + match M with + | [] => e + | R :: M => rplug R (e_reset (meta_rplug M e)) + end. + +Definition meta_rctx_comp {V} : meta_rctx V → meta_rctx V → meta_rctx V := + app. + +Lemma fold_unfold_meta_rctx_comp {V} (M1 M2 : meta_rctx V) : + meta_rctx_comp M1 M2 = M1 ++ M2. +Proof. auto. Qed. + +Lemma meta_rplug_eq_val_inv {V} (M : meta_rctx V) (e : expr V) (v : val V) : + meta_rplug M e = v → + M = [] ∧ e = v. +Proof. + intros H_eq. + destruct M as [| R M]; simpl in *. + - auto. + - apply rplug_eq_val_inv in H_eq as [_ H_absurd]. + discriminate. +Qed. + +Lemma meta_rctx_comp_correct {V} (M1 M2 : meta_rctx V) e : + meta_rplug (meta_rctx_comp M1 M2) e = meta_rplug M1 (meta_rplug M2 e). +Proof. + induction M1 as [| R M1 IHM1]. + - simpl. reflexivity. + - simpl. rewrite -> IHM1. reflexivity. +Qed. + +Definition meta_ectx_to_meta_rctx {V} (M : meta_ectx V) : meta_rctx V := + rev (map ectx_to_rctx M). + +Definition meta_rctx_to_meta_ectx {V} (M : meta_rctx V) : meta_ectx V := + rev (map rctx_to_ectx M). + +Lemma meta_ectx_meta_rctx_bijection {V} (M : meta_ectx V) : + meta_rctx_to_meta_ectx (meta_ectx_to_meta_rctx M) = M. +Proof. + unfold meta_ectx_to_meta_rctx, meta_rctx_to_meta_ectx. + rewrite -> map_rev. + rewrite -> rev_involutive. + rewrite -> map_map. + rewrite <- (List.map_id M) at 2. + apply map_ext. apply ectx_rctx_bijection. +Qed. + +Lemma meta_rctx_meta_ectx_bijection {V} (M : meta_rctx V) : + meta_ectx_to_meta_rctx (meta_rctx_to_meta_ectx M) = M. +Proof. + unfold meta_ectx_to_meta_rctx, meta_rctx_to_meta_ectx. + rewrite -> map_rev. + rewrite -> rev_involutive. + rewrite -> map_map. + rewrite <- (List.map_id M) at 2. + apply map_ext. apply rctx_ectx_bijection. +Qed. + +Lemma meta_rctx_to_meta_ectx_correct {V} (M : meta_rctx V) (e : expr V) : + meta_plug (meta_rctx_to_meta_ectx M) e = meta_rplug M e. +Proof. + unfold meta_rctx_to_meta_ectx. + induction M as [| R M IHM]. + - simpl. reflexivity. + - simpl. + rewrite <- fold_unfold_meta_ectx_comp. + rewrite -> meta_ectx_comp_correct. simpl. + rewrite -> IHM. + rewrite -> rctx_to_ectx_correct. + reflexivity. +Qed. + +Lemma meta_ectx_to_meta_rctx_correct {V} (M : meta_ectx V) (e : expr V) : + meta_rplug (meta_ectx_to_meta_rctx M) e = meta_plug M e. +Proof. + revert e. + unfold meta_ectx_to_meta_rctx. + induction M as [| R M IHM]; intros e. + - simpl. reflexivity. + - simpl. + rewrite <- fold_unfold_meta_rctx_comp. + rewrite -> meta_rctx_comp_correct. simpl. + rewrite -> IHM. + rewrite -> ectx_to_rctx_correct. + reflexivity. +Qed. + (** Reduction *) Inductive base_step {V} : expr V → expr V → Prop := @@ -677,26 +781,6 @@ Inductive meta_contextual_step {V} : expr V → expr V → Prop := contextual_step e1 e2 → meta_contextual_step (meta_plug M e1) (meta_plug M e2). -(* -Inductive contextual_step {V} : expr V → expr V → Prop := -| Ectx_beta_lambda MR R (e : expr (inc V)) (v : val V) : - contextual_step - (meta_plug MR (rplug R (e_app (v_lambda e) v))) - (meta_plug MR (rplug R (subst (Inc:=inc) e v))) -| Ectx_shift MR R (e : expr (inc V)) : - contextual_step - (meta_plug MR (rplug R (e_shift e))) - (meta_plug MR (subst (Inc:=inc) e (v_cont R))) -| Ectx_beta_cont MR R (R' : rctx V) (v : val V) : - contextual_step - (meta_plug MR (rplug R (e_app (v_cont R') v))) - (meta_plug MR (rplug R (e_reset (rplug R' v)))) -| Ectx_reset MR R (v : val V) : - contextual_step - (meta_plug MR (rplug R (e_reset v))) - (meta_plug MR (rplug R v)). -*) - Definition big_step {V} e (v : val V) := rtc meta_contextual_step e v. @@ -783,64 +867,149 @@ Lemma base_step_potential_redex {V} (e e' : expr V) : potential_redex e. Proof. inversion_clear 1; constructor. Qed. -Lemma unique_rdecomposition {V} (R1 R2 : rctx V) e1 e2 : +Fixpoint decompose {V} (e : expr V) : val V + (meta_rctx V * (rctx V * expr V)) := + match e with + | e_val v => inl v + | e_app e1 e2 => + inr + match decompose e1 with + | inl v1 => + match decompose e2 with + | inl v2 => ([], (rctx_hole, e_app v1 v2)) + | inr (M, p) => + match M with + | [] => let (R, e2') := p in ([], (rctx_app2 v1 R, e2')) + | R :: M' => (rctx_app2 v1 R :: M', p) + end + end + | inr (M, p) => + match M with + | [] => let (R, e1') := p in ([], (rctx_app1 R e2, e1')) + | R :: M' => (rctx_app1 R e2 :: M', p) + end + end + | e_shift e' => inr ([], (rctx_hole, e_shift e')) + | e_reset e' => + inr + match decompose e' with + | inl v => ([], (rctx_hole, e_reset v)) + | inr (M, p) => (rctx_hole :: M, p) + end + end. + +Lemma decompose_correct {V} (e : expr V) : + match decompose e with + | inl v => e_val v + | inr (M, (R, e')) => meta_rplug M (rplug R e') + end = e. +Proof. + induction e; simpl. + - reflexivity. + - destruct (decompose e1) as [v1 | (M & R & e1')]. + + destruct (decompose e2) as [v2 | (M & R & e2')]. + * simpl. congruence. + * destruct M as [| R' M']. + { simpl in *. congruence. } + { simpl in *. congruence. } + + destruct M as [| R' M']. + * simpl in *. congruence. + * simpl in *. congruence. + - reflexivity. + - destruct (decompose e) as [v | (M & R & e')]. + + simpl in *. congruence. + + simpl in *. congruence. +Qed. + +Lemma decompose_redex {V} (e : expr V) : + match decompose e with + | inl _ => True + | inr (_, (_, e')) => potential_redex e' + end. +Proof. + induction e; simpl. + - exact I. + - destruct (decompose e1) as [v1 | (M & R & e1')]. + + destruct (decompose e2) as [v2 | (M & R & e2')]. + * simpl. constructor. + * destruct M as [| R' M']. + { exact IHe2. } + { exact IHe2. } + + destruct M as [| R' M']. + * exact IHe1. + * exact IHe1. + - constructor. + - destruct (decompose e) as [v | (M & R & e')]. + + constructor. + + exact IHe. +Qed. + +Lemma decompose_unique_nil {V} R (e : expr V) : + potential_redex e → + decompose (rplug R e) = inr ([], (R, e)). +Proof. + intros He. + induction R as [| R IHR e' | v' R IHR]; simpl in *. + - inversion_clear He; simpl in *. + + reflexivity. + + reflexivity. + + reflexivity. + - rewrite -> IHR. reflexivity. + - rewrite -> IHR. reflexivity. +Qed. + +Lemma decompose_unique_cons {V} R M p (e : expr V) : + decompose e = inr (M, p) → + decompose (rplug R (e_reset e)) = inr (R :: M, p). +Proof. + intros He. + induction R as [| R IHR e' | v' R IHR]; simpl in *. + - rewrite -> He. reflexivity. + - rewrite -> IHR. reflexivity. + - rewrite -> IHR. reflexivity. +Qed. + +Lemma decompose_unique {V} M R (e : expr V) : + potential_redex e → + decompose (meta_rplug M (rplug R e)) = inr (M, (R, e)). +Proof. + intros He. + induction M. + - simpl. apply decompose_unique_nil. exact He. + - simpl. apply decompose_unique_cons. exact IHM. +Qed. + +Lemma unique_rdecomposition {V} (M1 M2 : meta_rctx V) (R1 R2 : rctx V) e1 e2 : potential_redex e1 → potential_redex e2 → - rplug R1 e1 = rplug R2 e2 → - R1 = R2 ∧ e1 = e2. + meta_rplug M1 (rplug R1 e1) = meta_rplug M2 (rplug R2 e2) → + M1 = M2 ∧ R1 = R2 ∧ e1 = e2. Proof. - intros He1 He2. - revert R2. - induction R1; intros R2 Heq. - - destruct R2; simpl in *. - + auto. - + rewrite -> Heq in He1. - apply potential_redex_app_inv in He1 as (v1 & v2 & Hv1 & Hv2). - apply rplug_eq_val_inv in Hv1 as [_ ->]. - contradict (not_potential_redex_val _ He2). - + rewrite -> Heq in He1. - apply potential_redex_app_inv in He1 as (v1 & v2 & Hv1 & Hv2). - apply rplug_eq_val_inv in Hv2 as [_ ->]. - contradict (not_potential_redex_val _ He2). - - destruct R2; simpl in *. - + rewrite <- Heq in He2. - apply potential_redex_app_inv in He2 as (v1 & v2 & Hv1 & Hv2). - apply rplug_eq_val_inv in Hv1 as [_ ->]. - contradict (not_potential_redex_val _ He1). - + injection Heq as Heq1 Heq2. - apply IHR1 in Heq1 as [Heq11 Heq12]. - split; congruence. - + injection Heq as Heq1 Heq2. - apply rplug_eq_val_inv in Heq1 as [_ ->]. - contradict (not_potential_redex_val _ He1). - - destruct R2; simpl in *. - + rewrite <- Heq in He2. - apply potential_redex_app_inv in He2 as (v1 & v2 & Hv1 & Hv2). - apply rplug_eq_val_inv in Hv2 as [_ ->]. - contradict (not_potential_redex_val _ He1). - + injection Heq as Heq1 Heq2. symmetry in Heq1. - apply rplug_eq_val_inv in Heq1 as [_ ->]. - contradict (not_potential_redex_val _ He2). - + injection Heq as Heq1 Heq2. - apply IHR1 in Heq2 as [Heq11 Heq12]. - split; congruence. -Qed. - -Lemma unique_decomposition {V} (E1 E2 : ectx V) e1 e2 : + intros He1 He2 Heq. + assert (H_decompose1 := decompose_unique M1 R1 e1 He1). + assert (H_decompose2 := decompose_unique M2 R2 e2 He2). + rewrite -> Heq in H_decompose1. + rewrite -> H_decompose1 in H_decompose2. + injection H_decompose2. auto. +Qed. + +Lemma unique_decomposition {V} + (M1 M2 : meta_ectx V) (E1 E2 : ectx V) e1 e2 : potential_redex e1 → potential_redex e2 → - plug E1 e1 = plug E2 e2 → - E1 = E2 ∧ e1 = e2. + meta_plug M1 (plug E1 e1) = meta_plug M2 (plug E2 e2) → + M1 = M2 ∧ E1 = E2 ∧ e1 = e2. Proof. intros He1 He2 Heq. - rewrite <- ectx_to_rctx_correct in Heq. - rewrite <- ectx_to_rctx_correct in Heq. - destruct (unique_rdecomposition _ _ _ _ He1 He2 Heq) as [Heq1 Heq2]. - split. - - rewrite <- (ectx_rctx_bijection E1). - rewrite <- (ectx_rctx_bijection E2). - f_equal. exact Heq1. - - exact Heq2. + rewrite <- (ectx_rctx_bijection E1) in Heq |- *. + rewrite <- (ectx_rctx_bijection E2) in Heq |- *. + rewrite <- (meta_ectx_meta_rctx_bijection M1) in Heq |- *. + rewrite <- (meta_ectx_meta_rctx_bijection M2) in Heq |- *. + rewrite -> meta_rctx_to_meta_ectx_correct in Heq. + rewrite -> meta_rctx_to_meta_ectx_correct in Heq. + rewrite -> rctx_to_ectx_correct in Heq. + rewrite -> rctx_to_ectx_correct in Heq. + apply unique_rdecomposition in Heq as (HM & HE & He); auto. + repeat split; congruence. Qed. Lemma contextual_step_is_deterministic {V} (e1 e2 e3 : expr V) : @@ -1783,8 +1952,6 @@ Proof. - apply compat_ectx_app2. apply fundamental_property_v. exact IHE. } Qed. -Print Assumptions fundamental_property_v. - (** General program contexts *) Inductive ctx : Set → Type := | ctx_hole : ctx ∅ From 8647563e28ed9c2a8e67c0a00c6ea0b1926e3a91 Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Mon, 8 Dec 2025 16:00:27 +0800 Subject: [PATCH 09/11] Finish deterministic proof and fix soundness proof --- ctx-equiv-ixfree/lang_shift_reset.v | 977 ++++++---------------------- 1 file changed, 206 insertions(+), 771 deletions(-) diff --git a/ctx-equiv-ixfree/lang_shift_reset.v b/ctx-equiv-ixfree/lang_shift_reset.v index 448122d..ed4b56b 100644 --- a/ctx-equiv-ixfree/lang_shift_reset.v +++ b/ctx-equiv-ixfree/lang_shift_reset.v @@ -1012,33 +1012,101 @@ Proof. repeat split; congruence. Qed. +Lemma unique_decomposition_ectx {V} (E1 E2 : ectx V) e1 e2 : + potential_redex e1 → + potential_redex e2 → + plug E1 e1 = plug E2 e2 → + E1 = E2 ∧ e1 = e2. +Proof. + intros He1 He2 Heq. + rewrite <- (fold_unfold_meta_plug_nil (plug _ e1)) in Heq. + rewrite <- (fold_unfold_meta_plug_nil (plug _ e2)) in Heq. + apply unique_decomposition in Heq; auto. tauto. +Qed. + +Inductive potential_ectx_redex {V} : expr V -> Prop := +| per_redex (E : ectx V) (e : expr V) : + potential_redex e → + potential_ectx_redex (plug E e). + +Lemma contextual_step_potential_ectx_redex {V} (e e' : expr V) : + contextual_step e e' → + potential_ectx_redex e. +Proof. + inversion_clear 1. + - constructor. by eapply base_step_potential_redex. + - constructor. constructor. +Qed. + +Lemma unique_decomposition_meta_ectx {V} (M1 M2 : meta_ectx V) e1 e2 : + potential_ectx_redex e1 → + potential_ectx_redex e2 → + meta_plug M1 e1 = meta_plug M2 e2 → + M1 = M2 ∧ e1 = e2. +Proof. + intros He1 He2 Heq. + inversion He1 as [E1 e1' He1' H_plug1]. + inversion He2 as [E2 e2' He2' H_plug2]. + rewrite <- H_plug1 in Heq. + rewrite <- H_plug2 in Heq. + apply unique_decomposition in Heq; auto. + intuition congruence. +Qed. + +Lemma not_base_step_shift {V} e1 (e2 : expr V) : ¬ base_step (e_shift e1) e2. +Proof. inversion_clear 1. Qed. + Lemma contextual_step_is_deterministic {V} (e1 e2 e3 : expr V) : contextual_step e1 e2 → contextual_step e1 e3 → e2 = e3. Proof. -Admitted. - -(* intros Hstep2 Hstep3. - inversion Hstep2 as [E2 e12 e2' Hstep2' He12 He2' |]. - inversion Hstep3 as [E3 e13 e3' Hstep3' He13 He3' |]. - assert (Hpr2 := base_step_potential_redex _ _ Hstep2'). - assert (Hpr3 := base_step_potential_redex _ _ Hstep3'). - destruct (unique_decomposition E2 E3 e12 e13 Hpr2 Hpr3) as [HE_eq He_eq]. - { congruence. } - rewrite -> He_eq in Hstep2'. - assert (He_eq' := base_step_is_deterministic e13 e2' e3' Hstep2' Hstep3'). - congruence. + inversion Hstep2 as [E2 e12 e2' Hstep2' He12 He2 | E2 e2' He12 He2]. + - inversion Hstep3 as [E3 e13 e3' Hstep3' He13 He3 | E3 e3' He13 He3]. + + destruct (unique_decomposition_ectx E2 E3 e12 e13) as [HE_eq He_eq]. + { by eapply base_step_potential_redex. } + { by eapply base_step_potential_redex. } + { congruence. } + rewrite -> He_eq in Hstep2'. + assert (He_eq' := base_step_is_deterministic e13 e2' e3' Hstep2' Hstep3'). + congruence. + + destruct (unique_decomposition_ectx E2 E3 e12 (e_shift e3')) as [_ He_eq]. + { by eapply base_step_potential_redex. } + { constructor. } + { congruence. } + rewrite -> He_eq in Hstep2'. + by apply not_base_step_shift in Hstep2'. + - inversion Hstep3 as [E3 e13 e3' Hstep3' He13 He3 | E3 e3' He13 He3]. + + destruct (unique_decomposition_ectx E2 E3 (e_shift e2') e13) as [HE_eq He_eq]. + { constructor. } + { by eapply base_step_potential_redex. } + { congruence. } + rewrite <- He_eq in Hstep3'. + by apply not_base_step_shift in Hstep3'. + + destruct (unique_decomposition_ectx E2 E3 (e_shift e2') (e_shift e3')) as [HE_eq He_eq]. + { constructor. } + { constructor. } + { congruence. } + congruence. Qed. -*) Lemma meta_contextual_step_is_deterministic {V} (e1 e2 e3 : expr V) : meta_contextual_step e1 e2 → meta_contextual_step e1 e3 → e2 = e3. Proof. -Admitted. + intros H_step2 H_step3. + inversion H_step2 as [M2 e12 e2' H_step2' He12 He2']. + inversion H_step3 as [M3 e13 e3' H_step3' He13 He3']. + destruct (unique_decomposition_meta_ectx M2 M3 e12 e13) as [HM_eq He_eq]. + { by eapply contextual_step_potential_ectx_redex. } + { by eapply contextual_step_potential_ectx_redex. } + { congruence. } + rewrite -> He_eq in H_step2'. + assert (He_eq' := contextual_step_is_deterministic e13 e2' e3' H_step2' H_step3'). + congruence. +Qed. (** Relations for closed term *) @@ -1934,7 +2002,7 @@ Lemma fundamental_property_e {V} (e : expr V) n : n ⊨ E_rel_o e e with fundamental_property_v {V} (v : val V) n : n ⊨ V_rel_o v v -with fundamental_property_K {V} (E : ectx V) n : +with fundamental_property_k {V} (E : ectx V) n : n ⊨ K_rel_o E E. Proof. { induction e. @@ -1945,7 +2013,7 @@ Proof. { induction v. - apply compat_var. - apply compat_lambda. apply fundamental_property_e. - - apply compat_cont. apply fundamental_property_K. } + - apply compat_cont. apply fundamental_property_k. } { induction E. - apply compat_ectx_hole. - apply compat_ectx_app1. exact IHE. apply fundamental_property_e. @@ -1978,8 +2046,8 @@ Fixpoint ciplug {V} (C : ctx V) : expr V → expr ∅ := Inductive ctxr : Set → Set → Type := | ctxr_hole {A} : ctxr A A -| ctxr_fmap {A B C} : (A [→] B) → ctxr B C → ctxr A C -| ctxr_bind {A B C} : (A [⇒] B) → ctxr B C → ctxr A C +| ctxr_fmap {A B C} : (B [→] C) → ctxr A B → ctxr A C +| ctxr_bind {A B C} : (B [⇒] C) → ctxr A B → ctxr A C | ctxr_lam {A B} : ctxr A (inc B) → ctxr A B | ctxr_app1 {A B} : ctxr A B → expr B → ctxr A B | ctxr_app2 {A B} : expr B → ctxr A B → ctxr A B @@ -1990,8 +2058,8 @@ Inductive ctxr : Set → Set → Type := Fixpoint crplug {A B} (C : ctxr A B) : expr A → expr B := match C with | ctxr_hole => λ e, e - | ctxr_fmap f C => λ e, crplug C (fmap f e) - | ctxr_bind f C => λ e, crplug C (bind f e) + | ctxr_fmap f C => λ e, fmap f (crplug C e) + | ctxr_bind f C => λ e, bind f (crplug C e) | ctxr_lam C => λ e, v_lambda (crplug C e) | ctxr_app1 C e2 => λ e, e_app (crplug C e) e2 | ctxr_app2 e1 C => λ e, e_app e1 (crplug C e) @@ -1999,8 +2067,31 @@ Fixpoint crplug {A B} (C : ctxr A B) : expr A → expr B := | ctxr_reset C => λ e, e_reset (crplug C e) end. -(* Outside-in plugging simplifies the proofs below *) -Notation cplug := crplug. +Fixpoint ctxr_comp {A B C} (C1 : ctxr B C) : ctxr A B → ctxr A C := + match C1 with + | ctxr_hole => λ C2, C2 + | ctxr_fmap f C1 => λ C2, ctxr_fmap f (ctxr_comp C1 C2) + | ctxr_bind f C1 => λ C2, ctxr_bind f (ctxr_comp C1 C2) + | ctxr_lam C1 => λ C2, ctxr_lam (ctxr_comp C1 C2) + | ctxr_app1 C1 e2 => λ C2, ctxr_app1 (ctxr_comp C1 C2) e2 + | ctxr_app2 e1 C1 => λ C2, ctxr_app2 e1 (ctxr_comp C1 C2) + | ctxr_shift C1 => λ C2, ctxr_shift (ctxr_comp C1 C2) + | ctxr_reset C1 => λ C2, ctxr_reset (ctxr_comp C1 C2) + end. + +Lemma ctxr_comp_correct {A B C} (C1 : ctxr B C) (C2 : ctxr A B) (e : expr A) : + crplug (ctxr_comp C1 C2) e = crplug C1 (crplug C2 e). +Proof. + induction C1. + - simpl. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. +Qed. (** Observational approximation for complete programs *) Definition obs_approx (e1 e2 : expr ∅) : Prop := @@ -2035,15 +2126,14 @@ Proof. unfold Transitive, obs_equiv. intuition auto. Qed. (** Contextual approximation, where the context closes off Γ *) Definition ctx_approx {V} (e1 e2 : expr V) : Prop := - ∀ C, obs_approx (cplug C e1) (cplug C e2). + ∀ C, obs_approx (crplug C e1) (crplug C e2). (** Contextual equivalence *) Definition ctx_equiv {V} (e1 e2 : expr V) : Prop := - ∀ C, obs_equiv (cplug C e1) (cplug C e2). + ∀ C, obs_equiv (crplug C e1) (crplug C e2). -(* TODO *) -(* Infix "≼ctx" := ctx_approx (at level 80, right associativity, only printing). -Infix "≡ctx" := ctx_equiv (at level 80, right associativity, only printing). *) +Infix "≼ctx" := ctx_approx (at level 80, right associativity, only printing). +Infix "≡ctx" := ctx_equiv (at level 80, right associativity, only printing). #[global] Instance Reflexive_ctx_approx {V} : Reflexive (@ctx_approx V). @@ -2103,10 +2193,10 @@ Proof. Qed. Definition ciu_approx {V} (e1 e2 : expr V) : Prop := - ∀ E γ, obs_approx (plug E (bind γ e1)) (plug E (bind γ e2)). + ∀ M E γ, obs_approx (meta_plug M (plug E (bind γ e1))) (meta_plug M (plug E (bind γ e2))). Definition ciu_equiv {V} (e1 e2 : expr V) : Prop := - ∀ E γ, obs_equiv (plug E (bind γ e1)) (plug E (bind γ e2)). + ∀ M E γ, obs_equiv (meta_plug M (plug E (bind γ e1))) (meta_plug M (plug E (bind γ e2))). Lemma ciu_equiv_intro_approx {V} (e1 e2 : expr V) : ciu_approx e1 e2 → @@ -2114,7 +2204,7 @@ Lemma ciu_equiv_intro_approx {V} (e1 e2 : expr V) : ciu_equiv e1 e2. Proof. unfold ciu_approx, ciu_equiv. - intros He1 He2 E γ. + intros He1 He2 M E γ. apply obs_equiv_intro_approx; auto. Qed. @@ -2125,8 +2215,8 @@ Lemma ciu_equiv_elim_approx {V} (e1 e2 : expr V) : Proof. unfold ciu_approx, ciu_equiv. intros He. split. - - intros E γ. specialize (He E γ). by apply obs_equiv_elim_approx in He as []. - - intros E γ. specialize (He E γ). by apply obs_equiv_elim_approx in He as []. + - intros M E γ. specialize (He M E γ). by apply obs_equiv_elim_approx in He as []. + - intros M E γ. specialize (He M E γ). by apply obs_equiv_elim_approx in He as []. Qed. Definition n_big_step {V} n (e : expr V) (v : val V) := nsteps meta_contextual_step n e v. @@ -2243,7 +2333,7 @@ Proof. - exact He1. Qed. -Theorem O_rel_adequacy e1 e2 : +Lemma O_rel_adequacy e1 e2 : (∀ n, n ⊨ O_rel e1 e2) → obs_equiv e1 e2. Proof. intros He. split. @@ -2273,12 +2363,12 @@ Proof. - apply compat_id_subst. Qed. -Lemma fundamental_property_K_closed E n : +Lemma fundamental_property_k_closed E n : n ⊨ K_rel E E. Proof. rewrite <- (bind_pure' E). apply K_rel_o_elimK. - - apply fundamental_property_K. + - apply fundamental_property_k. - apply compat_id_subst. Qed. @@ -2289,6 +2379,25 @@ Proof. iintros x. apply fundamental_property_v_closed. Qed. +Lemma fundamental_property_m {V} (M : meta_ectx V) n : + n ⊨ M_rel_o M M. +Proof. + induction M as [| E M IHM]. + - apply compat_meta_ectx_nil. + - apply compat_meta_ectx_cons. + + apply fundamental_property_k. + + exact IHM. +Qed. + +Lemma fundamental_property_m_closed M n : + n ⊨ M_rel M M. +Proof. + rewrite <- (bind_pure' M). + apply M_rel_o_elimM. + - apply fundamental_property_m. + - apply compat_id_subst. +Qed. + Lemma compat_expr_o_closed (e1 e2 : expr ∅) n : n ⊨ E_rel e1 e2 → n ⊨ O_rel e1 e2. @@ -2346,13 +2455,13 @@ Qed. Lemma precongruence {A B} (e1 e2 : expr A) (C : ctxr A B) n : n ⊨ E_rel_o e1 e2 → - n ⊨ E_rel_o (cplug C e1) (cplug C e2). + n ⊨ E_rel_o (crplug C e1) (crplug C e2). Proof. intros He. induction C; simpl. - exact He. - - apply IHC. apply compat_fmap. exact He. - - apply IHC. apply compat_bind. exact He. + - apply compat_fmap. apply IHC. exact He. + - apply compat_bind. apply IHC. exact He. - apply compat_val. apply compat_lambda. apply IHC. exact He. @@ -2362,25 +2471,23 @@ Proof. - apply compat_app. + apply fundamental_property_e. + apply IHC. exact He. - - apply compat_shift. - apply IHC. exact He. - - apply compat_reset. - apply IHC. exact He. + - apply compat_shift. apply IHC. exact He. + - apply compat_reset. apply IHC. exact He. Qed. -(* Theorem E_rel_o_soundness' {V} (e1 e2 : expr V) : (∀ n, n ⊨ E_rel_o e1 e2) → ciu_equiv e1 e2. Proof. unfold ciu_equiv. - intros He E γ. + intros He M E γ. apply O_rel_adequacy. intros n. apply E_rel_elimO. - specialize (He n). apply E_rel_o_elimE. + exact He. + apply fundamental_property_g. + - apply fundamental_property_m_closed. - apply fundamental_property_k_closed. Qed. @@ -2433,7 +2540,7 @@ Proof. assert (He1' : obs_approx e1' e2). { unfold obs_approx in *. intros He1'. apply He. - by eapply contextual_step_terminates. } + by eapply meta_contextual_step_terminates. } ispec IH e1' He1'. exact IH. Qed. @@ -2457,26 +2564,29 @@ Proof. apply ciu_equiv_elim_approx in He as [He1 He2]. unfold ciu_approx in *. apply E_rel_o_intro. iintros γ1 γ2 Hγ. - apply E_rel_intro. iintros E1 E2 HE. + apply E_rel_intro. iintros M1 M2 E1 E2 HM HE. apply O_rel_intro. - - assert (He1' : n ⊨ L_rel (plug E1 (bind γ1 e1)) (plug E2 (bind γ2 e1))). + - assert (He1' : n ⊨ L_rel + (meta_plug M1 (plug E1 (bind γ1 e1))) + (meta_plug M2 (plug E2 (bind γ2 e1)))). { apply O_rel_elim1. apply E_rel_elimO. apply E_rel_o_elimE. apply fundamental_property_e. - exact Hγ. exact HE. } - specialize (He1 E2 γ2). + exact Hγ. exact HM. exact HE. } + specialize (He1 M2 E2 γ2). exact (L_rel_obs_approx_trans _ _ _ _ He1' He1). - - assert (He2' : n ⊨ L_rel (plug E2 (bind γ2 e2)) (plug E1 (bind γ1 e2))). + - assert (He2' : n ⊨ L_rel + (meta_plug M2 (plug E2 (bind γ2 e2))) + (meta_plug M1 (plug E1 (bind γ1 e2)))). { apply O_rel_elim2. apply E_rel_elimO. apply E_rel_o_elimE. apply fundamental_property_e. - exact Hγ. exact HE. } - specialize (He2 E1 γ1). + exact Hγ. exact HM. exact HE. } + specialize (He2 M1 E1 γ1). exact (L_rel_obs_approx_trans _ _ _ _ He2' He2). Qed. -*) Theorem E_rel_o_soundness {V} (e1 e2 : expr V) : (∀ n, n ⊨ E_rel_o e1 e2) → @@ -2488,21 +2598,13 @@ Proof. specialize (He n). apply (precongruence _ _ C) in He. apply compat_expr_o in He. - rewrite <- (bind_pure' (cplug C e1)). - rewrite <- (bind_pure' (cplug C e2)). + rewrite <- (bind_pure' (crplug C e1)). + rewrite <- (bind_pure' (crplug C e2)). apply O_rel_o_elimO. - exact He. - apply compat_id_subst. Qed. -(* -Theorem ciu_equiv_soundness {V} (e1 e2 : expr V) : - ciu_equiv e1 e2 → - ctx_equiv e1 e2. -Proof. -Abort. -*) - Fixpoint rctx_to_ctxr {V} (R : rctx V) : ctxr V V := match R with | rctx_hole => ctxr_hole @@ -2510,6 +2612,9 @@ Fixpoint rctx_to_ctxr {V} (R : rctx V) : ctxr V V := | rctx_app2 v R => ctxr_app2 v (rctx_to_ctxr R) end. +Definition ectx_to_ctxr {V} (E : ectx V) : ctxr V V := + rctx_to_ctxr (ectx_to_rctx E). + Lemma rctx_to_ctxr_correct {V} (R : rctx V) (e : expr V) : crplug (rctx_to_ctxr R) e = rplug R e. Proof. @@ -2519,723 +2624,53 @@ Proof. - simpl. rewrite -> IHR. reflexivity. Qed. -Theorem ciu_equiv_completeness {V} (e1 e2 : expr V) : - ctx_equiv e1 e2 → - ciu_equiv e1 e2. +Lemma ectx_to_ctxr_correct {V} (E : ectx V) (e : expr V) : + crplug (ectx_to_ctxr E) e = plug E e. Proof. - unfold ctx_equiv, ciu_equiv. - intros He E γ. - rewrite <-2 ectx_to_rctx_correct. - rewrite <-2 rctx_to_ctxr_correct. - specialize (He (ctxr_bind γ (rctx_to_ctxr (ectx_to_rctx E)))). - simpl in He. exact He. -Qed. - -(* -(* composition of subst and subst_map, - where the variable to be substituted doesn't appear in the substitution *) -Lemma subst_subst_map : ∀ (e:expr) Γ (x : string) (es : val) (map : sub), - subst_is_closed Γ ∅ map → - subst x es (subst_map (delete x map) e) = - subst_map (insert x es map) e -with subst_subst_map_val : ∀ (v:val) Γ (x : string) (es : val) (map : sub), - subst_is_closed Γ ∅ map → - subst x es (subst_map_val (delete x map) v) = - subst_map_val (insert x es map) v. -Proof. - { intros e. induction e. - { intros. apply (subst_subst_map_val _ _ _ _ _ H). } - { (* e is a variable x *) - intros. simpl. destruct (decide (x0=x)) as [->|Hne]. - { (* if x=x0, we'll end up substituting es into x *) - rewrite lookup_delete_eq with (m:=map). - rewrite lookup_insert_eq with (m:=map). - simpl. - by rewrite decide_True. } - { (* if not equal, the deletion and insertion will have no effect *) - rewrite lookup_delete_ne with (m:=map). 2: { assumption. } - rewrite lookup_insert_ne with (m:=map). 2: { assumption. } - (* we then need to see if x is in the map to begin with *) - destruct (map !! x) as [v1|] eqn:Hkey. - { Fail rewrite Hkey. (* why does regular rewrite not work? *) - setoid_rewrite Hkey. - simpl. - rewrite (subst_val_closed _ ∅ _ _). - - reflexivity. - - apply (subst_is_closed_elim_closed _ _ x _ _ H Hkey). - - set_solver. } - { setoid_rewrite Hkey. - simpl. - by rewrite decide_False. } } } - { intros. simpl. f_equal; eauto. } } - { intros v. induction v; intros. - (*{ reflexivity. }*) - { (* the lambda case *) - simpl. f_equal. f_equal. - case_decide. - { subst. - rewrite delete_delete_eq with (m:=map). - rewrite delete_insert_eq with (m:=map). done. } - { rewrite delete_insert_ne with (m:=map). 2: { congruence. } - rewrite delete_delete with (m:=map). - eapply subst_subst_map. - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x map) map). - destruct H as [H1 H2]. - rewrite H1. - set_solver. - apply delete_subseteq. - set_solver. - assumption. } } - (*{ reflexivity. }*) } -Qed. - -(** Special case of Theorem A.1 from Erlang paper: - scoping of extended substitutions. - Given a closed substitution, we can add a closed value to it. *) -Lemma scope_extend1 Γ x (v:val) (γ:sub): - closed ∅ v → - subst_is_closed Γ ∅ γ → - subst_is_closed (Γ ∪ {[x]}) ∅ (<[x := v]> γ). -Proof. - intros Hc Hsc. - split. - { destruct Hsc. rewrite H. set_solver. } - intros x0 Hd v0 Hs. - (* we have to prove that for an arbitrary binding x0 := v0 in γ, v0 is closed *) - destruct (decide (x=x0)) as [->|Hne]. - (* if x = x0, the premise tells us v0 is closed *) - { rewrite lookup_insert_eq with (m:=γ) in Hs. - injection Hs. intros. subst. - exact Hc. } - (* if they are not equal, we know x0 is in Γ and have to use the fact - that the subst_is_closed *) - { rewrite elem_of_union in Hd. destruct Hd. 2: { assert (x0 = x). set_solver. done. } - destruct Hsc as [_ Hsc]. - rewrite lookup_insert_ne with (m:=γ) in Hs. 2: { assumption. } - specialize (Hsc x0 H v0 Hs). - assumption. } -Qed. - -(* we can extend related substitutions with related values *) -Lemma sem_context_rel_insert Γ x v1 v2 γ1 γ2 n: - n ⊨ V_rel v1 v2 → - n ⊨ G_rel Γ γ1 γ2 → - n ⊨ G_rel (Γ ∪ {[x]}) (<[x := v1]> γ1) (<[x := v2]> γ2). -Proof. - intros Hv Hγ. - destruct (V_rel_elim _ _ _ Hv) as (Hvc1 & Hvc2 & Hv'). - destruct (G_rel_elim _ _ _ _ Hγ) as (Hγc1 & Hγc2 & Hγ'). - apply G_rel_intro. - { by apply scope_extend1. } - { by apply scope_extend1. } - - iintros x0 v0 v3 H1 H2. - destruct (decide (x=x0)). - { subst. - rewrite lookup_insert_eq with (m:=γ2) in H2. idestruct H2. injection H2 as ->. - rewrite lookup_insert_eq with (m:=γ1) in H1. idestruct H1. injection H1 as ->. - assumption. } - { rewrite lookup_insert_ne with (m:=γ2) in H2. idestruct H2. 2: { assumption. } - rewrite lookup_insert_ne with (m:=γ1) in H1. idestruct H1. 2: { assumption. } - iapply Hγ'. - - iintro. eassumption. - - iintro. eassumption. } -Qed. - -Lemma lambda_closed Γ γ x e : - closed (Γ ∪ {[x]}) e → - subst_is_closed Γ ∅ γ → - closed ∅ (vlambda x (subst_map (delete x γ) e)). -Proof. - intros Hec [Heq Hγc]. - rewrite closed_lambda. - eapply subst_map_closed'_2. - - eapply closed_weaken. - exact Hec. - setoid_rewrite dom_delete. - intros y. destruct (decide (x = y)); set_solver. - - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x γ) γ). - + set_solver. - + apply delete_subseteq. - + set_solver. - + unfold subst_is_closed. split. - * exact Heq. - * intros x' Hin v Hlookup. - specialize (Hγc x' Hin v Hlookup). - by eapply closed_weaken. -Qed. -*) - -(* -Notation name := string. -Definition sub : Set := gmap name val. -Definition scope : Set := gset name. - -Fixpoint subst_map_val (γ : sub) (v : val) : val := - match v with - (*| vunit => vunit - | vint n => vint n*) - | vlambda x e => vlambda x (subst_map (delete x γ) e) - end -with subst_map (γ : sub) (e : expr) : expr := - match e with - | ret v => ret (subst_map_val γ v) - (* | eunit => eunit *) - | var x => match γ !! x with Some v => ret v | _ => var x end - | app e1 e2 => app (subst_map γ e1) (subst_map γ e2) - (* | abs x e => abs x (subst_map (binder_delete x xs) e) *) - (* | eplus e1 e2 => eplus (subst_map xs e1) (subst_map xs e2) *) - end. - -Lemma fold_unfold_subst_map_val_vlambda γ x e : - subst_map_val γ (vlambda x e) = - vlambda x (subst_map (delete x γ) e). -Proof. auto. Qed. - -Lemma fold_unfold_subst_map_ret γ v : - subst_map γ (ret v) = - ret (subst_map_val γ v). -Proof. auto. Qed. + unfold ectx_to_ctxr. + rewrite -> rctx_to_ctxr_correct. + rewrite -> ectx_to_rctx_correct. + reflexivity. +Qed. -Lemma fold_unfold_subst_map_var γ x : - subst_map γ (var x) = - match γ !! x with - | Some v => ret v - | _ => var x +Fixpoint meta_rctx_to_ctxr {V} (M : meta_rctx V) : ctxr V V := + match M with + | [] => ctxr_hole + | R :: M => ctxr_comp (rctx_to_ctxr R) (ctxr_reset (meta_rctx_to_ctxr M)) end. -Proof. auto. Qed. -Lemma fold_unfold_subst_map_app γ e1 e2 : - subst_map γ (app e1 e2) = - app (subst_map γ e1) (subst_map γ e2). -Proof. auto. Qed. +Definition meta_ectx_to_ctxr {V} (E : meta_ectx V) : ctxr V V := + meta_rctx_to_ctxr (meta_ectx_to_meta_rctx E). -Fixpoint is_closed (X : scope) (e : expr) : bool := - match e with - | var x => bool_decide (x ∈ X) - (*| ret vunit | ret (vint _) => true*) - | ret (vlambda x e) => is_closed (X ∪ {[x]}) e - | app e1 e2 - (* | eplus e1 e2 *) - => is_closed X e1 && is_closed X e2 - end. - -(* aka a scoping judgment, the untyped equivalent of a typing judgment *) -Definition closed (X : scope) (e : expr) : Prop := Is_true (is_closed X e). - -Lemma closed_weaken e X Y: - closed X e → X ⊆ Y → closed Y e -with closed_weaken_val (v:val) X Y: - closed X v → X ⊆ Y → closed Y v. -Proof. - { revert X Y. - induction e; intros. - - apply (closed_weaken_val _ _ _ H H0). - - unfold closed, is_closed in *. - apply bool_decide_unpack in H. - apply bool_decide_pack. - set_solver. - - unfold closed in *. simpl in *. - apply andb_prop_intro. - apply andb_prop_elim in H. - destruct H. - split. - apply (IHe1 _ _ H H0). - apply (IHe2 _ _ H1 H0). } - { revert X Y. - induction v; intros. - (*- constructor.*) - - unfold closed in *. - simpl in *. - apply (closed_weaken e _ _ H). - set_solver. - (*- constructor.*) } -Qed. - -Lemma closed_subst : - ∀ Γ x e1 e2, - closed (Γ ∪ {[x]}) e1 → - closed ∅ e2 → - closed Γ (subst x e2 e1) -with closed_subst_val : - ∀ Γ x (v : val) e, - closed (Γ ∪ {[x]}) v → - closed ∅ e → - closed Γ (subst_val x e v). -Proof. - { - unfold closed in *. - intros Γ x e1 e2 H_closed1 H_closed2. - induction e1. - - simpl in *. auto. - - simpl in *. - rewrite -> decide_bool_decide. - destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. - + eapply closed_weaken. - exact H_closed2. - set_solver. - + simpl in *. - apply bool_decide_unpack in H_closed1. - apply bool_decide_pack. - set_solver. - - simpl in *. - apply andb_prop_elim in H_closed1 as [Hc1 Hc2]. - apply andb_prop_intro. split. - + exact (IHe1_1 Hc1). - + exact (IHe1_2 Hc2). - } - { - unfold closed in *. - intros Γ x v e H_closed1 H_closed2. - induction v. - (* - simpl in *. auto. *) - - simpl in *. - rewrite -> decide_bool_decide. - destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. - + rewrite -> H_eq in H_closed1. - eapply closed_weaken. - exact H_closed1. set_solver. - + simpl in *. - apply closed_subst. - eapply closed_weaken. - exact H_closed1. set_solver. - exact H_closed2. - (*- simpl in *. auto.*) - } -Qed. - -Lemma closed_app xs e1 e2: - closed xs (app e1 e2) ↔ - closed xs e1 ∧ closed xs e2. -Proof. unfold closed. simpl. by rewrite andb_True. Qed. - -Lemma closed_lambda e Γ x : closed Γ (vlambda x e) ↔ closed (Γ ∪ {[x]}) e. -Proof. split. auto. auto. Qed. - -Lemma closed_var Γ x : closed Γ (var x) ↔ x ∈ Γ. -Proof. unfold closed. simpl. by rewrite bool_decide_spec. Qed. - -Lemma base_step_preserve_closedness : - ∀ e1 e1', - base_step e1 e1' → - closed ∅ e1 → - closed ∅ e1'. -Proof. - unfold closed. - intros e1 e1' Hred H_closed. - inversion Hred. subst. simpl in *. - apply andb_prop_elim in H_closed as [Hc1 Hc2]. - exact (closed_subst ∅ x e0 e2 Hc1 Hc2). -Qed. - -Definition ectx_is_closed (X : scope) (E : ectx) := - ∀ e, closed ∅ e → closed X (fill E e). - -Lemma closed_decompose : - ∀ E e, - closed ∅ (fill E e) → - ectx_is_closed ∅ E ∧ closed ∅ e. -Proof. - unfold ectx_is_closed. - intros E. - induction E; intros e' H_closed. - - simpl in *. auto. - - simpl in *. - apply IHE in H_closed as [H_fill H_closed]. - apply closed_app in H_closed as [H_closed1 H_closed2]. - split. - + intros e'' H_closed3. - apply H_fill. - apply closed_app. auto. - + auto. - - simpl in *. - apply IHE in H_closed as [H_fill H_closed]. - apply closed_app in H_closed as [H_closed1 H_closed2]. - split. - + intros e'' H_closed3. - apply H_fill. - apply closed_app. auto. - + auto. -Qed. - -Lemma closed_compose : - ∀ E e, - ectx_is_closed ∅ E → - closed ∅ e → - closed ∅ (fill E e). -Proof. unfold ectx_is_closed. auto. Qed. - -Lemma contextual_step_preserve_closedness : - ∀ e1 e1', - contextual_step e1 e1' → - closed ∅ e1 → - closed ∅ e1'. -Proof. - unfold closed. - intros e1 e1' Hred H_closed. - inversion Hred. subst. simpl in *. clear Hred. - apply closed_decompose in H_closed as [H_closed1 H_closed2]. - apply (base_step_preserve_closedness _ _ H1) in H_closed2. - apply closed_compose; auto. -Qed. - -(** subscoped from the Erlang paper *) -Definition subst_is_closed (Γ free : scope) (sub : sub) := - Γ = dom sub ∧ - ∀ x, x ∈ Γ → - ∀ v, sub !! x = Some v → closed free (ret v). - -Lemma subst_is_closed_subseteq: ∀ (Γ1 Γ2 X : scope) (γ1 γ2: sub), - Γ1 = dom γ1 → - γ1 ⊆ γ2 → Γ1 ⊆ Γ2 → subst_is_closed Γ2 X γ2 → subst_is_closed Γ1 X γ1. -Proof. - intros * Hd Hγ HΓ Hclosed2. - destruct Hclosed2 as [Hd2 Hc2]. - split. - assumption. - intros x Hl v Hs. - rewrite (map_subseteq_spec γ1 γ2) in Hγ. - (* specialize (Hγ _ _ Hs). *) - specialise Hγ Hs. - apply (Hc2 x ltac:(set_solver) _ Hγ). -Qed. - -Lemma closed_ectx_app1 : - ∀ E e, - ectx_is_closed ∅ E → - closed ∅ e → - ectx_is_closed ∅ (ectx_app1 E e). -Proof. - intros E e Hc1 Hc2. - unfold ectx_is_closed in *. - intros e' Hc3. simpl. - apply Hc1. apply closed_app. auto. -Qed. - -Lemma closed_ectx_app2 : - ∀ (v : val) E, - closed ∅ v → - ectx_is_closed ∅ E → - ectx_is_closed ∅ (ectx_app2 v E). -Proof. - intros v E Hc1 Hc2. - unfold ectx_is_closed in *. - intros e' Hc3. simpl. - apply Hc2. apply closed_app. auto. -Qed. - -Lemma subst_is_closed_elim_closed Γ (γ:sub) x X (v:val): - subst_is_closed Γ X γ → - γ !! x = Some v → - closed X v. -Proof. - intros [Hdom Hsc] He. - assert (H := elem_of_dom_2 _ _ _ He). - (* have: elem_of_dom_2 He. *) - (* pose proof (elem_of_dom_2 _ _ _ He). *) - assert (x ∈ Γ). set_solver. - apply (Hsc x H0 v He). -Qed. - -(* if e is closed under Y, we can split the variables in Y between X and γ *) -Lemma subst_map_closed' (e : expr) (X Y : scope) (γ:sub) : - closed Y e → - (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → - closed X (subst_map γ e) -with subst_map_closed'_val (v : val) (X Y : scope) (γ:sub) : - closed Y (ret v) → - (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → - closed X (subst_map_val γ v). -Proof. - { - revert X Y γ. induction e. - { eapply subst_map_closed'_val; eauto. } - { intros * Hc H. - (* e is a variable x *) - unfold closed in Hc; simpl in Hc; apply bool_decide_unpack in Hc. - specialize (H x Hc). simpl. - destruct (γ !! x) eqn:He. - - assumption. - - unfold closed; simpl; apply bool_decide_pack. assumption. - } - { intros *. - unfold closed. simpl. - rewrite !andb_True. intros [? ?] **. - split. - by eapply IHe1. - by eapply IHe2. } - } - { revert X Y γ. induction v. - (*{ intros. assumption. }*) - { unfold closed. simpl. - intros * Hce H. - eapply subst_map_closed'. eassumption. - intros y [|]%elem_of_union. - { destruct (decide (x = y)). - { by subst; rewrite lookup_delete_eq with (m:=γ); set_solver. } - rewrite lookup_delete_ne with (m:=γ). 2: { assumption. } - eapply H in H0. - destruct lookup; last set_solver. - eapply closed_weaken; eauto with set_solver. } - { rewrite elem_of_singleton in H0. - subst. rewrite lookup_delete_eq with (m:=γ). set_solver. } - } - (*{ intros. assumption. }*) } -Qed. - -(* simple corollary of [subst_map_closed'], - where we have split Y into X and dom γ upfront *) -Lemma subst_map_closed'_2 Γ X γ e : - closed (X ∪ (dom γ)) e -> - subst_is_closed Γ X γ -> - closed X (subst_map γ e). -Proof. - intros Hcl Hsubst. - eapply subst_map_closed'; first eassumption. - intros x Hx. - destruct (γ !! x) as [e'|] eqn:Heq. - - apply (subst_is_closed_elim_closed _ _ _ _ _ Hsubst Heq). - - apply not_elem_of_dom in Heq. - set_solver. -Qed. - -(* given a value and a substitution closed under the same scope, - applying the substitution gives us a completely closed value *) -Lemma subst_map_closed'_3 (v:val) Γ γ: - closed Γ v -> - subst_is_closed Γ ∅ γ -> - closed ∅ (subst_map γ v). -Proof. - pose proof (subst_map_closed'_2 Γ ∅ γ v). - simpl in H. - intros. - apply H. 2: { assumption. } - destruct H1 as [? _]. - rewrite <- H1. - replace (∅ ∪ Γ) with Γ. assumption. - set_solver. -Qed. - -Lemma subst_map_val_closed_val_aux Γ γ (v : val) : - closed Γ v → - Γ ∩ dom γ = ∅ → - subst_map_val γ v = v -with subst_map_closed_aux Γ γ (e : expr) : - closed Γ e → - Γ ∩ dom γ = ∅ → - subst_map γ e = e. -Proof. - { intros Hc Hdom. - induction v. - - rewrite -> closed_lambda in Hc. - rewrite -> fold_unfold_subst_map_val_vlambda. - rewrite -> (subst_map_closed_aux (Γ ∪ {[x]}) (delete x γ) e Hc ltac:(set_solver)). - reflexivity. } - { intros Hc Hdom. - induction e. - - rewrite -> fold_unfold_subst_map_ret. - f_equal. by eapply subst_map_val_closed_val_aux. - - unfold closed in Hc. - simpl in Hc. - rewrite -> bool_decide_spec in Hc. - rewrite -> fold_unfold_subst_map_var. - assert (H_not_in : x ∉ dom γ) by set_solver. - rewrite -> (not_elem_of_dom γ x) in H_not_in. - setoid_rewrite -> H_not_in. - reflexivity. - - apply closed_app in Hc as [Hc1 Hc2]. - rewrite -> fold_unfold_subst_map_app. - rewrite -> (IHe1 Hc1). - rewrite -> (IHe2 Hc2). - reflexivity. } -Qed. - -Lemma subst_map_val_closed_val γ (v : val) : - closed ∅ v → - subst_map_val γ v = v. -Proof. - intros Hc. - eapply subst_map_val_closed_val_aux. - - exact Hc. - - set_solver. -Qed. - -Lemma subst_map_closed γ (e : expr) : - closed ∅ e → - subst_map γ e = e. -Proof. - intros Hc. - eapply subst_map_closed_aux. - - exact Hc. - - set_solver. -Qed. -*) - -(* -Lemma subset_is_closed_absurd x Γ γ: - x ∈ Γ → - subst_is_closed Γ ∅ γ → - γ !! x = None → - False. -Proof. - intros Hx Hs He. - pose proof (not_elem_of_dom_2 _ _ He). - destruct Hs as [? _]. - setoid_rewrite <- H0 in H. - set_solver. -Qed. -*) - -(* -Lemma subst_val_closed v X x es : - closed X (of_val v) → x ∉ X → subst_val x es v = v -with subst_closed X e x es : - closed X e → x ∉ X → subst x es e = e. +Lemma meta_rctx_to_ctxr_correct {V} (M : meta_rctx V) (e : expr V) : + crplug (meta_rctx_to_ctxr M) e = meta_rplug M e. Proof. - { induction v. - (*{ reflexivity. }*) - { simpl. - case_decide. - - reflexivity. - - intros. - f_equal. - rewrite closed_lambda in H0. - apply (subst_closed _ _ _ _ H0). - set_solver. } - (*{ reflexivity. }*) } - { induction e; intros; simpl. - { f_equal. - eapply subst_val_closed. - apply H. assumption. } - { case_decide. - - subst. - unfold closed in H. simpl in H. apply bool_decide_unpack in H. - set_solver. - - reflexivity. } - { apply closed_app in H. - destruct H as (H1&H2). - specialize (IHe1 H1 H0). - specialize (IHe2 H2 H0). - f_equal. - - assumption. - - assumption. } } -Qed. -*) - -(* -(* aka contextual scoping C : Γ ~> Γ', a special case of contextual typing. - defined inductively because we need to invert it. *) -Inductive closed_ctx : scope → scope → ctx → Prop := - | cc_hole Γ : - closed_ctx Γ Γ ctx_hole - - | cc_lambda x Γ Γ' C : - closed_ctx Γ (Γ' ∪ {[x]}) C → - closed_ctx Γ Γ' (ctx_lam x C) - - | cc_app1 Γ Γ' C e : - closed_ctx Γ Γ' C → - closed Γ' e → - closed_ctx Γ Γ' (ctx_app1 C e) - - | cc_app2 Γ Γ' C v : - closed_ctx Γ Γ' C → - closed Γ' (ret v) → - closed_ctx Γ Γ' (ctx_app2 v C) - . -*) - -(* -Definition closed_ctx_sem (Γ Γ' : scope) (C:ctx) : Prop := - forall e, closed Γ e → closed Γ' (cplug C e). - -Lemma closed_ctx_sound Γ Γ' e : - closed_ctx Γ Γ' e → closed_ctx_sem Γ Γ' e. -Proof. - intros H. induction H; unfold closed_ctx_sem. - - simpl. done. - - intros e Hc. - simpl. - specialize (IHclosed_ctx e Hc). - unfold closed. simpl. - apply IHclosed_ctx. - - intros e2 Hc. - specialize (IHclosed_ctx _ Hc). - simpl. - unfold closed. simpl. - unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. - auto. - - intros e1 Hc. - specialize (IHclosed_ctx _ Hc). - simpl. - unfold closed. simpl. - unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. - auto. -Qed. -*) - -(* -Lemma G_rel_empty n : - n ⊨ G_rel ∅ ∅ ∅. -Proof. - apply G_rel_intro. - - unfold subst_is_closed. split; set_solver. - - unfold subst_is_closed. split; set_solver. - - iintros. - idestruct H. - setoid_rewrite lookup_empty in H. - discriminate H. + induction M as [| R M IHM]; simpl. + - reflexivity. + - rewrite -> ctxr_comp_correct. simpl. + rewrite -> rctx_to_ctxr_correct. + rewrite -> IHM. + reflexivity. Qed. -Lemma subst_map_empty (e:expr) : - subst_map ∅ e = e -with subst_map_val_empty (v:val) : - subst_map_val ∅ v = v. +Lemma meta_ectx_to_ctxr_correct {V} (M : meta_ectx V) (e : expr V) : + crplug (meta_ectx_to_ctxr M) e = meta_plug M e. Proof. - { induction e. - - simpl. f_equal. eapply subst_map_val_empty. - - simpl. - setoid_rewrite lookup_empty. - reflexivity. - - simpl. rewrite IHe1. rewrite IHe2. reflexivity. } - { induction v. - - simpl. - f_equal. - setoid_rewrite delete_empty. - apply subst_map_empty. } + unfold meta_ectx_to_ctxr. + rewrite -> meta_rctx_to_ctxr_correct. + rewrite -> meta_ectx_to_meta_rctx_correct. + reflexivity. Qed. -Lemma fundamental_property_sub Γ γ n : - subst_is_closed Γ ∅ γ → - n ⊨ G_rel Γ γ γ. -Proof. - intros Hc. - apply G_rel_intro. - { exact Hc. } - { exact Hc. } - iintros x v1 v2 Heq1 Heq2. - idestruct Heq1. - idestruct Heq2. - rewrite -> Heq1 in Heq2. - injection Heq2 as ->. - destruct Hc as [-> Hc]. - assert (Hcv2 : closed ∅ v2). - { eapply Hc. - - setoid_rewrite -> elem_of_dom. - unfold is_Some. eauto. - - exact Heq1. } - assert (H_fundamental := fundamental_property_v ∅ v2 n Hcv2). - apply V_rel_o_elim in H_fundamental. - ispec H_fundamental ∅ ∅ (G_rel_empty n). - rewrite -> subst_map_val_empty in H_fundamental. - exact H_fundamental. -Qed. - -Lemma fundamental_property_ectx E n : - n ⊨ K_rel E E. +Theorem ciu_equiv_completeness {V} (e1 e2 : expr V) : + ctx_equiv e1 e2 → + ciu_equiv e1 e2. Proof. - apply K_rel_intro. - iintros v1 v2 Hv. - admit. -Admitted. -*) + unfold ctx_equiv, ciu_equiv. + intros He M E γ. + specialize (He (ctxr_comp (meta_ectx_to_ctxr M) (ctxr_comp (ectx_to_ctxr E) (ctxr_bind γ ctxr_hole)))). + rewrite ->! ctxr_comp_correct in He. simpl in He. + rewrite ->! meta_ectx_to_ctxr_correct in He. + rewrite ->! ectx_to_ctxr_correct in He. + exact He. +Qed. From ab542685b3ff2f4f9a7686de839a9229c08f1645 Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Mon, 8 Dec 2025 16:12:53 +0800 Subject: [PATCH 10/11] Fix mistake in the old lambda calculus file --- ctx-equiv-ixfree/lang.v | 787 ++------------------------------ ctx-equiv-ixfree/unused_proof.v | 708 ++++++++++++++++++++++++++++ 2 files changed, 758 insertions(+), 737 deletions(-) create mode 100644 ctx-equiv-ixfree/unused_proof.v diff --git a/ctx-equiv-ixfree/lang.v b/ctx-equiv-ixfree/lang.v index d13bdaa..f9c6634 100644 --- a/ctx-equiv-ixfree/lang.v +++ b/ctx-equiv-ixfree/lang.v @@ -1351,8 +1351,8 @@ Fixpoint ciplug {V} (C : ctx V) : expr V → expr ∅ := Inductive ctxr : Set → Set → Type := | ctxr_hole {A} : ctxr A A -| ctxr_fmap {A B C} : (A [→] B) → ctxr B C → ctxr A C -| ctxr_bind {A B C} : (A [⇒] B) → ctxr B C → ctxr A C +| ctxr_fmap {A B C} : (B [→] C) → ctxr A B → ctxr A C +| ctxr_bind {A B C} : (B [⇒] C) → ctxr A B → ctxr A C | ctxr_lam {A B} : ctxr A (inc B) → ctxr A B | ctxr_app1 {A B} : ctxr A B → expr B → ctxr A B | ctxr_app2 {A B} : expr B → ctxr A B → ctxr A B. @@ -1361,15 +1361,34 @@ Inductive ctxr : Set → Set → Type := Fixpoint crplug {A B} (C : ctxr A B) : expr A → expr B := match C with | ctxr_hole => λ e, e - | ctxr_fmap f C => λ e, crplug C (fmap f e) - | ctxr_bind f C => λ e, crplug C (bind f e) + | ctxr_fmap f C => λ e, fmap f (crplug C e) + | ctxr_bind f C => λ e, bind f (crplug C e) | ctxr_lam C => λ e, v_lambda (crplug C e) | ctxr_app1 C e2 => λ e, app (crplug C e) e2 | ctxr_app2 e1 C => λ e, app e1 (crplug C e) end. -(* Outside-in plugging simplifies the proofs below *) -Notation cplug := crplug. +Fixpoint ctxr_comp {A B C} (C1 : ctxr B C) : ctxr A B → ctxr A C := + match C1 with + | ctxr_hole => λ C2, C2 + | ctxr_fmap f C1 => λ C2, ctxr_fmap f (ctxr_comp C1 C2) + | ctxr_bind f C1 => λ C2, ctxr_bind f (ctxr_comp C1 C2) + | ctxr_lam C1 => λ C2, ctxr_lam (ctxr_comp C1 C2) + | ctxr_app1 C1 e2 => λ C2, ctxr_app1 (ctxr_comp C1 C2) e2 + | ctxr_app2 e1 C1 => λ C2, ctxr_app2 e1 (ctxr_comp C1 C2) + end. + +Lemma ctxr_comp_correct {A B C} (C1 : ctxr B C) (C2 : ctxr A B) (e : expr A) : + crplug (ctxr_comp C1 C2) e = crplug C1 (crplug C2 e). +Proof. + induction C1. + - simpl. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. + - simpl. rewrite -> IHC1. reflexivity. +Qed. (** Observational approximation for complete programs *) Definition obs_approx (e1 e2 : expr ∅) : Prop := @@ -1404,15 +1423,14 @@ Proof. unfold Transitive, obs_equiv. intuition auto. Qed. (** Contextual approximation, where the context closes off Γ *) Definition ctx_approx {V} (e1 e2 : expr V) : Prop := - ∀ C, obs_approx (cplug C e1) (cplug C e2). + ∀ C, obs_approx (crplug C e1) (crplug C e2). (** Contextual equivalence *) Definition ctx_equiv {V} (e1 e2 : expr V) : Prop := - ∀ C, obs_equiv (cplug C e1) (cplug C e2). + ∀ C, obs_equiv (crplug C e1) (crplug C e2). -(* TODO *) -(* Infix "≼ctx" := ctx_approx (at level 80, right associativity, only printing). -Infix "≡ctx" := ctx_equiv (at level 80, right associativity, only printing). *) +Infix "≼ctx" := ctx_approx (at level 80, right associativity, only printing). +Infix "≡ctx" := ctx_equiv (at level 80, right associativity, only printing). #[global] Instance Reflexive_ctx_approx {V} : Reflexive (@ctx_approx V). @@ -1828,13 +1846,13 @@ Qed. Lemma precongruence {A B} (e1 e2 : expr A) (C : ctxr A B) n : n ⊨ E_rel_o e1 e2 → - n ⊨ E_rel_o (cplug C e1) (cplug C e2). + n ⊨ E_rel_o (crplug C e1) (crplug C e2). Proof. intros He. induction C; simpl. - exact He. - - apply IHC. apply compat_fmap. exact He. - - apply IHC. apply compat_bind. exact He. + - apply compat_fmap. apply IHC. exact He. + - apply compat_bind. apply IHC. exact He. - apply compat_val. apply compat_lambda. apply IHC. exact He. @@ -1964,20 +1982,12 @@ Proof. specialize (He n). apply (precongruence _ _ C) in He. apply compat_expr_o in He. - rewrite <- (bind_pure' (cplug C e1)). - rewrite <- (bind_pure' (cplug C e2)). + rewrite <- (bind_pure' (crplug C e1)). + rewrite <- (bind_pure' (crplug C e2)). apply O_rel_o_elimO. exact He. apply G_rel_intro. iintros x. destruct x. Qed. -(* -Theorem ciu_equiv_soundness {V} (e1 e2 : expr V) : - ciu_equiv e1 e2 → - ctx_equiv e1 e2. -Proof. -Abort. -*) - Fixpoint rctx_to_ctxr {V} (R : rctx V) : ctxr V V := match R with | rctx_hole => ctxr_hole @@ -1985,6 +1995,9 @@ Fixpoint rctx_to_ctxr {V} (R : rctx V) : ctxr V V := | rctx_app2 v R => ctxr_app2 v (rctx_to_ctxr R) end. +Definition ectx_to_ctxr {V} (E : ectx V) : ctxr V V := + rctx_to_ctxr (ectx_to_rctx E). + Lemma rctx_to_ctxr_correct {V} (R : rctx V) (e : expr V) : crplug (rctx_to_ctxr R) e = rplug R e. Proof. @@ -1994,723 +2007,23 @@ Proof. - simpl. rewrite -> IHR. reflexivity. Qed. +Lemma ectx_to_ctxr_correct {V} (E : ectx V) (e : expr V) : + crplug (ectx_to_ctxr E) e = plug E e. +Proof. + unfold ectx_to_ctxr. + rewrite -> rctx_to_ctxr_correct. + rewrite -> ectx_to_rctx_correct. + reflexivity. +Qed. + Theorem ciu_equiv_completeness {V} (e1 e2 : expr V) : ctx_equiv e1 e2 → ciu_equiv e1 e2. Proof. unfold ctx_equiv, ciu_equiv. intros He E γ. - rewrite <-2 ectx_to_rctx_correct. - rewrite <-2 rctx_to_ctxr_correct. - specialize (He (ctxr_bind γ (rctx_to_ctxr (ectx_to_rctx E)))). - simpl in He. exact He. -Qed. - -(* -(* composition of subst and subst_map, - where the variable to be substituted doesn't appear in the substitution *) -Lemma subst_subst_map : ∀ (e:expr) Γ (x : string) (es : val) (map : sub), - subst_is_closed Γ ∅ map → - subst x es (subst_map (delete x map) e) = - subst_map (insert x es map) e -with subst_subst_map_val : ∀ (v:val) Γ (x : string) (es : val) (map : sub), - subst_is_closed Γ ∅ map → - subst x es (subst_map_val (delete x map) v) = - subst_map_val (insert x es map) v. -Proof. - { intros e. induction e. - { intros. apply (subst_subst_map_val _ _ _ _ _ H). } - { (* e is a variable x *) - intros. simpl. destruct (decide (x0=x)) as [->|Hne]. - { (* if x=x0, we'll end up substituting es into x *) - rewrite lookup_delete_eq with (m:=map). - rewrite lookup_insert_eq with (m:=map). - simpl. - by rewrite decide_True. } - { (* if not equal, the deletion and insertion will have no effect *) - rewrite lookup_delete_ne with (m:=map). 2: { assumption. } - rewrite lookup_insert_ne with (m:=map). 2: { assumption. } - (* we then need to see if x is in the map to begin with *) - destruct (map !! x) as [v1|] eqn:Hkey. - { Fail rewrite Hkey. (* why does regular rewrite not work? *) - setoid_rewrite Hkey. - simpl. - rewrite (subst_val_closed _ ∅ _ _). - - reflexivity. - - apply (subst_is_closed_elim_closed _ _ x _ _ H Hkey). - - set_solver. } - { setoid_rewrite Hkey. - simpl. - by rewrite decide_False. } } } - { intros. simpl. f_equal; eauto. } } - { intros v. induction v; intros. - (*{ reflexivity. }*) - { (* the lambda case *) - simpl. f_equal. f_equal. - case_decide. - { subst. - rewrite delete_delete_eq with (m:=map). - rewrite delete_insert_eq with (m:=map). done. } - { rewrite delete_insert_ne with (m:=map). 2: { congruence. } - rewrite delete_delete with (m:=map). - eapply subst_subst_map. - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x map) map). - destruct H as [H1 H2]. - rewrite H1. - set_solver. - apply delete_subseteq. - set_solver. - assumption. } } - (*{ reflexivity. }*) } -Qed. - -(** Special case of Theorem A.1 from Erlang paper: - scoping of extended substitutions. - Given a closed substitution, we can add a closed value to it. *) -Lemma scope_extend1 Γ x (v:val) (γ:sub): - closed ∅ v → - subst_is_closed Γ ∅ γ → - subst_is_closed (Γ ∪ {[x]}) ∅ (<[x := v]> γ). -Proof. - intros Hc Hsc. - split. - { destruct Hsc. rewrite H. set_solver. } - intros x0 Hd v0 Hs. - (* we have to prove that for an arbitrary binding x0 := v0 in γ, v0 is closed *) - destruct (decide (x=x0)) as [->|Hne]. - (* if x = x0, the premise tells us v0 is closed *) - { rewrite lookup_insert_eq with (m:=γ) in Hs. - injection Hs. intros. subst. - exact Hc. } - (* if they are not equal, we know x0 is in Γ and have to use the fact - that the subst_is_closed *) - { rewrite elem_of_union in Hd. destruct Hd. 2: { assert (x0 = x). set_solver. done. } - destruct Hsc as [_ Hsc]. - rewrite lookup_insert_ne with (m:=γ) in Hs. 2: { assumption. } - specialize (Hsc x0 H v0 Hs). - assumption. } -Qed. - -(* we can extend related substitutions with related values *) -Lemma sem_context_rel_insert Γ x v1 v2 γ1 γ2 n: - n ⊨ V_rel v1 v2 → - n ⊨ G_rel Γ γ1 γ2 → - n ⊨ G_rel (Γ ∪ {[x]}) (<[x := v1]> γ1) (<[x := v2]> γ2). -Proof. - intros Hv Hγ. - destruct (V_rel_elim _ _ _ Hv) as (Hvc1 & Hvc2 & Hv'). - destruct (G_rel_elim _ _ _ _ Hγ) as (Hγc1 & Hγc2 & Hγ'). - apply G_rel_intro. - { by apply scope_extend1. } - { by apply scope_extend1. } - - iintros x0 v0 v3 H1 H2. - destruct (decide (x=x0)). - { subst. - rewrite lookup_insert_eq with (m:=γ2) in H2. idestruct H2. injection H2 as ->. - rewrite lookup_insert_eq with (m:=γ1) in H1. idestruct H1. injection H1 as ->. - assumption. } - { rewrite lookup_insert_ne with (m:=γ2) in H2. idestruct H2. 2: { assumption. } - rewrite lookup_insert_ne with (m:=γ1) in H1. idestruct H1. 2: { assumption. } - iapply Hγ'. - - iintro. eassumption. - - iintro. eassumption. } -Qed. - -Lemma lambda_closed Γ γ x e : - closed (Γ ∪ {[x]}) e → - subst_is_closed Γ ∅ γ → - closed ∅ (vlambda x (subst_map (delete x γ) e)). -Proof. - intros Hec [Heq Hγc]. - rewrite closed_lambda. - eapply subst_map_closed'_2. - - eapply closed_weaken. - exact Hec. - setoid_rewrite dom_delete. - intros y. destruct (decide (x = y)); set_solver. - - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x γ) γ). - + set_solver. - + apply delete_subseteq. - + set_solver. - + unfold subst_is_closed. split. - * exact Heq. - * intros x' Hin v Hlookup. - specialize (Hγc x' Hin v Hlookup). - by eapply closed_weaken. -Qed. -*) - -(* -Notation name := string. -Definition sub : Set := gmap name val. -Definition scope : Set := gset name. - -Fixpoint subst_map_val (γ : sub) (v : val) : val := - match v with - (*| vunit => vunit - | vint n => vint n*) - | vlambda x e => vlambda x (subst_map (delete x γ) e) - end -with subst_map (γ : sub) (e : expr) : expr := - match e with - | ret v => ret (subst_map_val γ v) - (* | eunit => eunit *) - | var x => match γ !! x with Some v => ret v | _ => var x end - | app e1 e2 => app (subst_map γ e1) (subst_map γ e2) - (* | abs x e => abs x (subst_map (binder_delete x xs) e) *) - (* | eplus e1 e2 => eplus (subst_map xs e1) (subst_map xs e2) *) - end. - -Lemma fold_unfold_subst_map_val_vlambda γ x e : - subst_map_val γ (vlambda x e) = - vlambda x (subst_map (delete x γ) e). -Proof. auto. Qed. - -Lemma fold_unfold_subst_map_ret γ v : - subst_map γ (ret v) = - ret (subst_map_val γ v). -Proof. auto. Qed. - -Lemma fold_unfold_subst_map_var γ x : - subst_map γ (var x) = - match γ !! x with - | Some v => ret v - | _ => var x - end. -Proof. auto. Qed. - -Lemma fold_unfold_subst_map_app γ e1 e2 : - subst_map γ (app e1 e2) = - app (subst_map γ e1) (subst_map γ e2). -Proof. auto. Qed. - -Fixpoint is_closed (X : scope) (e : expr) : bool := - match e with - | var x => bool_decide (x ∈ X) - (*| ret vunit | ret (vint _) => true*) - | ret (vlambda x e) => is_closed (X ∪ {[x]}) e - | app e1 e2 - (* | eplus e1 e2 *) - => is_closed X e1 && is_closed X e2 - end. - -(* aka a scoping judgment, the untyped equivalent of a typing judgment *) -Definition closed (X : scope) (e : expr) : Prop := Is_true (is_closed X e). - -Lemma closed_weaken e X Y: - closed X e → X ⊆ Y → closed Y e -with closed_weaken_val (v:val) X Y: - closed X v → X ⊆ Y → closed Y v. -Proof. - { revert X Y. - induction e; intros. - - apply (closed_weaken_val _ _ _ H H0). - - unfold closed, is_closed in *. - apply bool_decide_unpack in H. - apply bool_decide_pack. - set_solver. - - unfold closed in *. simpl in *. - apply andb_prop_intro. - apply andb_prop_elim in H. - destruct H. - split. - apply (IHe1 _ _ H H0). - apply (IHe2 _ _ H1 H0). } - { revert X Y. - induction v; intros. - (*- constructor.*) - - unfold closed in *. - simpl in *. - apply (closed_weaken e _ _ H). - set_solver. - (*- constructor.*) } -Qed. - -Lemma closed_subst : - ∀ Γ x e1 e2, - closed (Γ ∪ {[x]}) e1 → - closed ∅ e2 → - closed Γ (subst x e2 e1) -with closed_subst_val : - ∀ Γ x (v : val) e, - closed (Γ ∪ {[x]}) v → - closed ∅ e → - closed Γ (subst_val x e v). -Proof. - { - unfold closed in *. - intros Γ x e1 e2 H_closed1 H_closed2. - induction e1. - - simpl in *. auto. - - simpl in *. - rewrite -> decide_bool_decide. - destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. - + eapply closed_weaken. - exact H_closed2. - set_solver. - + simpl in *. - apply bool_decide_unpack in H_closed1. - apply bool_decide_pack. - set_solver. - - simpl in *. - apply andb_prop_elim in H_closed1 as [Hc1 Hc2]. - apply andb_prop_intro. split. - + exact (IHe1_1 Hc1). - + exact (IHe1_2 Hc2). - } - { - unfold closed in *. - intros Γ x v e H_closed1 H_closed2. - induction v. - (* - simpl in *. auto. *) - - simpl in *. - rewrite -> decide_bool_decide. - destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. - + rewrite -> H_eq in H_closed1. - eapply closed_weaken. - exact H_closed1. set_solver. - + simpl in *. - apply closed_subst. - eapply closed_weaken. - exact H_closed1. set_solver. - exact H_closed2. - (*- simpl in *. auto.*) - } -Qed. - -Lemma closed_app xs e1 e2: - closed xs (app e1 e2) ↔ - closed xs e1 ∧ closed xs e2. -Proof. unfold closed. simpl. by rewrite andb_True. Qed. - -Lemma closed_lambda e Γ x : closed Γ (vlambda x e) ↔ closed (Γ ∪ {[x]}) e. -Proof. split. auto. auto. Qed. - -Lemma closed_var Γ x : closed Γ (var x) ↔ x ∈ Γ. -Proof. unfold closed. simpl. by rewrite bool_decide_spec. Qed. - -Lemma base_step_preserve_closedness : - ∀ e1 e1', - base_step e1 e1' → - closed ∅ e1 → - closed ∅ e1'. -Proof. - unfold closed. - intros e1 e1' Hred H_closed. - inversion Hred. subst. simpl in *. - apply andb_prop_elim in H_closed as [Hc1 Hc2]. - exact (closed_subst ∅ x e0 e2 Hc1 Hc2). -Qed. - -Definition ectx_is_closed (X : scope) (E : ectx) := - ∀ e, closed ∅ e → closed X (fill E e). - -Lemma closed_decompose : - ∀ E e, - closed ∅ (fill E e) → - ectx_is_closed ∅ E ∧ closed ∅ e. -Proof. - unfold ectx_is_closed. - intros E. - induction E; intros e' H_closed. - - simpl in *. auto. - - simpl in *. - apply IHE in H_closed as [H_fill H_closed]. - apply closed_app in H_closed as [H_closed1 H_closed2]. - split. - + intros e'' H_closed3. - apply H_fill. - apply closed_app. auto. - + auto. - - simpl in *. - apply IHE in H_closed as [H_fill H_closed]. - apply closed_app in H_closed as [H_closed1 H_closed2]. - split. - + intros e'' H_closed3. - apply H_fill. - apply closed_app. auto. - + auto. -Qed. - -Lemma closed_compose : - ∀ E e, - ectx_is_closed ∅ E → - closed ∅ e → - closed ∅ (fill E e). -Proof. unfold ectx_is_closed. auto. Qed. - -Lemma contextual_step_preserve_closedness : - ∀ e1 e1', - contextual_step e1 e1' → - closed ∅ e1 → - closed ∅ e1'. -Proof. - unfold closed. - intros e1 e1' Hred H_closed. - inversion Hred. subst. simpl in *. clear Hred. - apply closed_decompose in H_closed as [H_closed1 H_closed2]. - apply (base_step_preserve_closedness _ _ H1) in H_closed2. - apply closed_compose; auto. + specialize (He (ctxr_comp (ectx_to_ctxr E) (ctxr_bind γ ctxr_hole))). + rewrite ->! ctxr_comp_correct in He. simpl in He. + rewrite ->! ectx_to_ctxr_correct in He. + exact He. Qed. - -(** subscoped from the Erlang paper *) -Definition subst_is_closed (Γ free : scope) (sub : sub) := - Γ = dom sub ∧ - ∀ x, x ∈ Γ → - ∀ v, sub !! x = Some v → closed free (ret v). - -Lemma subst_is_closed_subseteq: ∀ (Γ1 Γ2 X : scope) (γ1 γ2: sub), - Γ1 = dom γ1 → - γ1 ⊆ γ2 → Γ1 ⊆ Γ2 → subst_is_closed Γ2 X γ2 → subst_is_closed Γ1 X γ1. -Proof. - intros * Hd Hγ HΓ Hclosed2. - destruct Hclosed2 as [Hd2 Hc2]. - split. - assumption. - intros x Hl v Hs. - rewrite (map_subseteq_spec γ1 γ2) in Hγ. - (* specialize (Hγ _ _ Hs). *) - specialise Hγ Hs. - apply (Hc2 x ltac:(set_solver) _ Hγ). -Qed. - -Lemma closed_ectx_app1 : - ∀ E e, - ectx_is_closed ∅ E → - closed ∅ e → - ectx_is_closed ∅ (ectx_app1 E e). -Proof. - intros E e Hc1 Hc2. - unfold ectx_is_closed in *. - intros e' Hc3. simpl. - apply Hc1. apply closed_app. auto. -Qed. - -Lemma closed_ectx_app2 : - ∀ (v : val) E, - closed ∅ v → - ectx_is_closed ∅ E → - ectx_is_closed ∅ (ectx_app2 v E). -Proof. - intros v E Hc1 Hc2. - unfold ectx_is_closed in *. - intros e' Hc3. simpl. - apply Hc2. apply closed_app. auto. -Qed. - -Lemma subst_is_closed_elim_closed Γ (γ:sub) x X (v:val): - subst_is_closed Γ X γ → - γ !! x = Some v → - closed X v. -Proof. - intros [Hdom Hsc] He. - assert (H := elem_of_dom_2 _ _ _ He). - (* have: elem_of_dom_2 He. *) - (* pose proof (elem_of_dom_2 _ _ _ He). *) - assert (x ∈ Γ). set_solver. - apply (Hsc x H0 v He). -Qed. - -(* if e is closed under Y, we can split the variables in Y between X and γ *) -Lemma subst_map_closed' (e : expr) (X Y : scope) (γ:sub) : - closed Y e → - (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → - closed X (subst_map γ e) -with subst_map_closed'_val (v : val) (X Y : scope) (γ:sub) : - closed Y (ret v) → - (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → - closed X (subst_map_val γ v). -Proof. - { - revert X Y γ. induction e. - { eapply subst_map_closed'_val; eauto. } - { intros * Hc H. - (* e is a variable x *) - unfold closed in Hc; simpl in Hc; apply bool_decide_unpack in Hc. - specialize (H x Hc). simpl. - destruct (γ !! x) eqn:He. - - assumption. - - unfold closed; simpl; apply bool_decide_pack. assumption. - } - { intros *. - unfold closed. simpl. - rewrite !andb_True. intros [? ?] **. - split. - by eapply IHe1. - by eapply IHe2. } - } - { revert X Y γ. induction v. - (*{ intros. assumption. }*) - { unfold closed. simpl. - intros * Hce H. - eapply subst_map_closed'. eassumption. - intros y [|]%elem_of_union. - { destruct (decide (x = y)). - { by subst; rewrite lookup_delete_eq with (m:=γ); set_solver. } - rewrite lookup_delete_ne with (m:=γ). 2: { assumption. } - eapply H in H0. - destruct lookup; last set_solver. - eapply closed_weaken; eauto with set_solver. } - { rewrite elem_of_singleton in H0. - subst. rewrite lookup_delete_eq with (m:=γ). set_solver. } - } - (*{ intros. assumption. }*) } -Qed. - -(* simple corollary of [subst_map_closed'], - where we have split Y into X and dom γ upfront *) -Lemma subst_map_closed'_2 Γ X γ e : - closed (X ∪ (dom γ)) e -> - subst_is_closed Γ X γ -> - closed X (subst_map γ e). -Proof. - intros Hcl Hsubst. - eapply subst_map_closed'; first eassumption. - intros x Hx. - destruct (γ !! x) as [e'|] eqn:Heq. - - apply (subst_is_closed_elim_closed _ _ _ _ _ Hsubst Heq). - - apply not_elem_of_dom in Heq. - set_solver. -Qed. - -(* given a value and a substitution closed under the same scope, - applying the substitution gives us a completely closed value *) -Lemma subst_map_closed'_3 (v:val) Γ γ: - closed Γ v -> - subst_is_closed Γ ∅ γ -> - closed ∅ (subst_map γ v). -Proof. - pose proof (subst_map_closed'_2 Γ ∅ γ v). - simpl in H. - intros. - apply H. 2: { assumption. } - destruct H1 as [? _]. - rewrite <- H1. - replace (∅ ∪ Γ) with Γ. assumption. - set_solver. -Qed. - -Lemma subst_map_val_closed_val_aux Γ γ (v : val) : - closed Γ v → - Γ ∩ dom γ = ∅ → - subst_map_val γ v = v -with subst_map_closed_aux Γ γ (e : expr) : - closed Γ e → - Γ ∩ dom γ = ∅ → - subst_map γ e = e. -Proof. - { intros Hc Hdom. - induction v. - - rewrite -> closed_lambda in Hc. - rewrite -> fold_unfold_subst_map_val_vlambda. - rewrite -> (subst_map_closed_aux (Γ ∪ {[x]}) (delete x γ) e Hc ltac:(set_solver)). - reflexivity. } - { intros Hc Hdom. - induction e. - - rewrite -> fold_unfold_subst_map_ret. - f_equal. by eapply subst_map_val_closed_val_aux. - - unfold closed in Hc. - simpl in Hc. - rewrite -> bool_decide_spec in Hc. - rewrite -> fold_unfold_subst_map_var. - assert (H_not_in : x ∉ dom γ) by set_solver. - rewrite -> (not_elem_of_dom γ x) in H_not_in. - setoid_rewrite -> H_not_in. - reflexivity. - - apply closed_app in Hc as [Hc1 Hc2]. - rewrite -> fold_unfold_subst_map_app. - rewrite -> (IHe1 Hc1). - rewrite -> (IHe2 Hc2). - reflexivity. } -Qed. - -Lemma subst_map_val_closed_val γ (v : val) : - closed ∅ v → - subst_map_val γ v = v. -Proof. - intros Hc. - eapply subst_map_val_closed_val_aux. - - exact Hc. - - set_solver. -Qed. - -Lemma subst_map_closed γ (e : expr) : - closed ∅ e → - subst_map γ e = e. -Proof. - intros Hc. - eapply subst_map_closed_aux. - - exact Hc. - - set_solver. -Qed. -*) - -(* -Lemma subset_is_closed_absurd x Γ γ: - x ∈ Γ → - subst_is_closed Γ ∅ γ → - γ !! x = None → - False. -Proof. - intros Hx Hs He. - pose proof (not_elem_of_dom_2 _ _ He). - destruct Hs as [? _]. - setoid_rewrite <- H0 in H. - set_solver. -Qed. -*) - -(* -Lemma subst_val_closed v X x es : - closed X (of_val v) → x ∉ X → subst_val x es v = v -with subst_closed X e x es : - closed X e → x ∉ X → subst x es e = e. -Proof. - { induction v. - (*{ reflexivity. }*) - { simpl. - case_decide. - - reflexivity. - - intros. - f_equal. - rewrite closed_lambda in H0. - apply (subst_closed _ _ _ _ H0). - set_solver. } - (*{ reflexivity. }*) } - { induction e; intros; simpl. - { f_equal. - eapply subst_val_closed. - apply H. assumption. } - { case_decide. - - subst. - unfold closed in H. simpl in H. apply bool_decide_unpack in H. - set_solver. - - reflexivity. } - { apply closed_app in H. - destruct H as (H1&H2). - specialize (IHe1 H1 H0). - specialize (IHe2 H2 H0). - f_equal. - - assumption. - - assumption. } } -Qed. -*) - -(* -(* aka contextual scoping C : Γ ~> Γ', a special case of contextual typing. - defined inductively because we need to invert it. *) -Inductive closed_ctx : scope → scope → ctx → Prop := - | cc_hole Γ : - closed_ctx Γ Γ ctx_hole - - | cc_lambda x Γ Γ' C : - closed_ctx Γ (Γ' ∪ {[x]}) C → - closed_ctx Γ Γ' (ctx_lam x C) - - | cc_app1 Γ Γ' C e : - closed_ctx Γ Γ' C → - closed Γ' e → - closed_ctx Γ Γ' (ctx_app1 C e) - - | cc_app2 Γ Γ' C v : - closed_ctx Γ Γ' C → - closed Γ' (ret v) → - closed_ctx Γ Γ' (ctx_app2 v C) - . -*) - -(* -Definition closed_ctx_sem (Γ Γ' : scope) (C:ctx) : Prop := - forall e, closed Γ e → closed Γ' (cplug C e). - -Lemma closed_ctx_sound Γ Γ' e : - closed_ctx Γ Γ' e → closed_ctx_sem Γ Γ' e. -Proof. - intros H. induction H; unfold closed_ctx_sem. - - simpl. done. - - intros e Hc. - simpl. - specialize (IHclosed_ctx e Hc). - unfold closed. simpl. - apply IHclosed_ctx. - - intros e2 Hc. - specialize (IHclosed_ctx _ Hc). - simpl. - unfold closed. simpl. - unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. - auto. - - intros e1 Hc. - specialize (IHclosed_ctx _ Hc). - simpl. - unfold closed. simpl. - unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. - auto. -Qed. -*) - -(* -Lemma G_rel_empty n : - n ⊨ G_rel ∅ ∅ ∅. -Proof. - apply G_rel_intro. - - unfold subst_is_closed. split; set_solver. - - unfold subst_is_closed. split; set_solver. - - iintros. - idestruct H. - setoid_rewrite lookup_empty in H. - discriminate H. -Qed. - -Lemma subst_map_empty (e:expr) : - subst_map ∅ e = e -with subst_map_val_empty (v:val) : - subst_map_val ∅ v = v. -Proof. - { induction e. - - simpl. f_equal. eapply subst_map_val_empty. - - simpl. - setoid_rewrite lookup_empty. - reflexivity. - - simpl. rewrite IHe1. rewrite IHe2. reflexivity. } - { induction v. - - simpl. - f_equal. - setoid_rewrite delete_empty. - apply subst_map_empty. } -Qed. - -Lemma fundamental_property_sub Γ γ n : - subst_is_closed Γ ∅ γ → - n ⊨ G_rel Γ γ γ. -Proof. - intros Hc. - apply G_rel_intro. - { exact Hc. } - { exact Hc. } - iintros x v1 v2 Heq1 Heq2. - idestruct Heq1. - idestruct Heq2. - rewrite -> Heq1 in Heq2. - injection Heq2 as ->. - destruct Hc as [-> Hc]. - assert (Hcv2 : closed ∅ v2). - { eapply Hc. - - setoid_rewrite -> elem_of_dom. - unfold is_Some. eauto. - - exact Heq1. } - assert (H_fundamental := fundamental_property_v ∅ v2 n Hcv2). - apply V_rel_o_elim in H_fundamental. - ispec H_fundamental ∅ ∅ (G_rel_empty n). - rewrite -> subst_map_val_empty in H_fundamental. - exact H_fundamental. -Qed. - -Lemma fundamental_property_ectx E n : - n ⊨ K_rel E E. -Proof. - apply K_rel_intro. - iintros v1 v2 Hv. - admit. -Admitted. -*) diff --git a/ctx-equiv-ixfree/unused_proof.v b/ctx-equiv-ixfree/unused_proof.v new file mode 100644 index 0000000..2abbf95 --- /dev/null +++ b/ctx-equiv-ixfree/unused_proof.v @@ -0,0 +1,708 @@ +(* +(* composition of subst and subst_map, + where the variable to be substituted doesn't appear in the substitution *) +Lemma subst_subst_map : ∀ (e:expr) Γ (x : string) (es : val) (map : sub), + subst_is_closed Γ ∅ map → + subst x es (subst_map (delete x map) e) = + subst_map (insert x es map) e +with subst_subst_map_val : ∀ (v:val) Γ (x : string) (es : val) (map : sub), + subst_is_closed Γ ∅ map → + subst x es (subst_map_val (delete x map) v) = + subst_map_val (insert x es map) v. +Proof. + { intros e. induction e. + { intros. apply (subst_subst_map_val _ _ _ _ _ H). } + { (* e is a variable x *) + intros. simpl. destruct (decide (x0=x)) as [->|Hne]. + { (* if x=x0, we'll end up substituting es into x *) + rewrite lookup_delete_eq with (m:=map). + rewrite lookup_insert_eq with (m:=map). + simpl. + by rewrite decide_True. } + { (* if not equal, the deletion and insertion will have no effect *) + rewrite lookup_delete_ne with (m:=map). 2: { assumption. } + rewrite lookup_insert_ne with (m:=map). 2: { assumption. } + (* we then need to see if x is in the map to begin with *) + destruct (map !! x) as [v1|] eqn:Hkey. + { Fail rewrite Hkey. (* why does regular rewrite not work? *) + setoid_rewrite Hkey. + simpl. + rewrite (subst_val_closed _ ∅ _ _). + - reflexivity. + - apply (subst_is_closed_elim_closed _ _ x _ _ H Hkey). + - set_solver. } + { setoid_rewrite Hkey. + simpl. + by rewrite decide_False. } } } + { intros. simpl. f_equal; eauto. } } + { intros v. induction v; intros. + (*{ reflexivity. }*) + { (* the lambda case *) + simpl. f_equal. f_equal. + case_decide. + { subst. + rewrite delete_delete_eq with (m:=map). + rewrite delete_insert_eq with (m:=map). done. } + { rewrite delete_insert_ne with (m:=map). 2: { congruence. } + rewrite delete_delete with (m:=map). + eapply subst_subst_map. + apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x map) map). + destruct H as [H1 H2]. + rewrite H1. + set_solver. + apply delete_subseteq. + set_solver. + assumption. } } + (*{ reflexivity. }*) } +Qed. + +(** Special case of Theorem A.1 from Erlang paper: + scoping of extended substitutions. + Given a closed substitution, we can add a closed value to it. *) +Lemma scope_extend1 Γ x (v:val) (γ:sub): + closed ∅ v → + subst_is_closed Γ ∅ γ → + subst_is_closed (Γ ∪ {[x]}) ∅ (<[x := v]> γ). +Proof. + intros Hc Hsc. + split. + { destruct Hsc. rewrite H. set_solver. } + intros x0 Hd v0 Hs. + (* we have to prove that for an arbitrary binding x0 := v0 in γ, v0 is closed *) + destruct (decide (x=x0)) as [->|Hne]. + (* if x = x0, the premise tells us v0 is closed *) + { rewrite lookup_insert_eq with (m:=γ) in Hs. + injection Hs. intros. subst. + exact Hc. } + (* if they are not equal, we know x0 is in Γ and have to use the fact + that the subst_is_closed *) + { rewrite elem_of_union in Hd. destruct Hd. 2: { assert (x0 = x). set_solver. done. } + destruct Hsc as [_ Hsc]. + rewrite lookup_insert_ne with (m:=γ) in Hs. 2: { assumption. } + specialize (Hsc x0 H v0 Hs). + assumption. } +Qed. + +(* we can extend related substitutions with related values *) +Lemma sem_context_rel_insert Γ x v1 v2 γ1 γ2 n: + n ⊨ V_rel v1 v2 → + n ⊨ G_rel Γ γ1 γ2 → + n ⊨ G_rel (Γ ∪ {[x]}) (<[x := v1]> γ1) (<[x := v2]> γ2). +Proof. + intros Hv Hγ. + destruct (V_rel_elim _ _ _ Hv) as (Hvc1 & Hvc2 & Hv'). + destruct (G_rel_elim _ _ _ _ Hγ) as (Hγc1 & Hγc2 & Hγ'). + apply G_rel_intro. + { by apply scope_extend1. } + { by apply scope_extend1. } + + iintros x0 v0 v3 H1 H2. + destruct (decide (x=x0)). + { subst. + rewrite lookup_insert_eq with (m:=γ2) in H2. idestruct H2. injection H2 as ->. + rewrite lookup_insert_eq with (m:=γ1) in H1. idestruct H1. injection H1 as ->. + assumption. } + { rewrite lookup_insert_ne with (m:=γ2) in H2. idestruct H2. 2: { assumption. } + rewrite lookup_insert_ne with (m:=γ1) in H1. idestruct H1. 2: { assumption. } + iapply Hγ'. + - iintro. eassumption. + - iintro. eassumption. } +Qed. + +Lemma lambda_closed Γ γ x e : + closed (Γ ∪ {[x]}) e → + subst_is_closed Γ ∅ γ → + closed ∅ (vlambda x (subst_map (delete x γ) e)). +Proof. + intros Hec [Heq Hγc]. + rewrite closed_lambda. + eapply subst_map_closed'_2. + - eapply closed_weaken. + exact Hec. + setoid_rewrite dom_delete. + intros y. destruct (decide (x = y)); set_solver. + - apply (subst_is_closed_subseteq (Γ ∖ {[x]}) Γ _ (delete x γ) γ). + + set_solver. + + apply delete_subseteq. + + set_solver. + + unfold subst_is_closed. split. + * exact Heq. + * intros x' Hin v Hlookup. + specialize (Hγc x' Hin v Hlookup). + by eapply closed_weaken. +Qed. +*) + +(* +Notation name := string. +Definition sub : Set := gmap name val. +Definition scope : Set := gset name. + +Fixpoint subst_map_val (γ : sub) (v : val) : val := + match v with + (*| vunit => vunit + | vint n => vint n*) + | vlambda x e => vlambda x (subst_map (delete x γ) e) + end +with subst_map (γ : sub) (e : expr) : expr := + match e with + | ret v => ret (subst_map_val γ v) + (* | eunit => eunit *) + | var x => match γ !! x with Some v => ret v | _ => var x end + | app e1 e2 => app (subst_map γ e1) (subst_map γ e2) + (* | abs x e => abs x (subst_map (binder_delete x xs) e) *) + (* | eplus e1 e2 => eplus (subst_map xs e1) (subst_map xs e2) *) + end. + +Lemma fold_unfold_subst_map_val_vlambda γ x e : + subst_map_val γ (vlambda x e) = + vlambda x (subst_map (delete x γ) e). +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_ret γ v : + subst_map γ (ret v) = + ret (subst_map_val γ v). +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_var γ x : + subst_map γ (var x) = + match γ !! x with + | Some v => ret v + | _ => var x + end. +Proof. auto. Qed. + +Lemma fold_unfold_subst_map_app γ e1 e2 : + subst_map γ (app e1 e2) = + app (subst_map γ e1) (subst_map γ e2). +Proof. auto. Qed. + +Fixpoint is_closed (X : scope) (e : expr) : bool := + match e with + | var x => bool_decide (x ∈ X) + (*| ret vunit | ret (vint _) => true*) + | ret (vlambda x e) => is_closed (X ∪ {[x]}) e + | app e1 e2 + (* | eplus e1 e2 *) + => is_closed X e1 && is_closed X e2 + end. + +(* aka a scoping judgment, the untyped equivalent of a typing judgment *) +Definition closed (X : scope) (e : expr) : Prop := Is_true (is_closed X e). + +Lemma closed_weaken e X Y: + closed X e → X ⊆ Y → closed Y e +with closed_weaken_val (v:val) X Y: + closed X v → X ⊆ Y → closed Y v. +Proof. + { revert X Y. + induction e; intros. + - apply (closed_weaken_val _ _ _ H H0). + - unfold closed, is_closed in *. + apply bool_decide_unpack in H. + apply bool_decide_pack. + set_solver. + - unfold closed in *. simpl in *. + apply andb_prop_intro. + apply andb_prop_elim in H. + destruct H. + split. + apply (IHe1 _ _ H H0). + apply (IHe2 _ _ H1 H0). } + { revert X Y. + induction v; intros. + (*- constructor.*) + - unfold closed in *. + simpl in *. + apply (closed_weaken e _ _ H). + set_solver. + (*- constructor.*) } +Qed. + +Lemma closed_subst : + ∀ Γ x e1 e2, + closed (Γ ∪ {[x]}) e1 → + closed ∅ e2 → + closed Γ (subst x e2 e1) +with closed_subst_val : + ∀ Γ x (v : val) e, + closed (Γ ∪ {[x]}) v → + closed ∅ e → + closed Γ (subst_val x e v). +Proof. + { + unfold closed in *. + intros Γ x e1 e2 H_closed1 H_closed2. + induction e1. + - simpl in *. auto. + - simpl in *. + rewrite -> decide_bool_decide. + destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. + + eapply closed_weaken. + exact H_closed2. + set_solver. + + simpl in *. + apply bool_decide_unpack in H_closed1. + apply bool_decide_pack. + set_solver. + - simpl in *. + apply andb_prop_elim in H_closed1 as [Hc1 Hc2]. + apply andb_prop_intro. split. + + exact (IHe1_1 Hc1). + + exact (IHe1_2 Hc2). + } + { + unfold closed in *. + intros Γ x v e H_closed1 H_closed2. + induction v. + (* - simpl in *. auto. *) + - simpl in *. + rewrite -> decide_bool_decide. + destruct (bool_decide_reflect (x = x0)) as [H_eq | H_neq]. + + rewrite -> H_eq in H_closed1. + eapply closed_weaken. + exact H_closed1. set_solver. + + simpl in *. + apply closed_subst. + eapply closed_weaken. + exact H_closed1. set_solver. + exact H_closed2. + (*- simpl in *. auto.*) + } +Qed. + +Lemma closed_app xs e1 e2: + closed xs (app e1 e2) ↔ + closed xs e1 ∧ closed xs e2. +Proof. unfold closed. simpl. by rewrite andb_True. Qed. + +Lemma closed_lambda e Γ x : closed Γ (vlambda x e) ↔ closed (Γ ∪ {[x]}) e. +Proof. split. auto. auto. Qed. + +Lemma closed_var Γ x : closed Γ (var x) ↔ x ∈ Γ. +Proof. unfold closed. simpl. by rewrite bool_decide_spec. Qed. + +Lemma base_step_preserve_closedness : + ∀ e1 e1', + base_step e1 e1' → + closed ∅ e1 → + closed ∅ e1'. +Proof. + unfold closed. + intros e1 e1' Hred H_closed. + inversion Hred. subst. simpl in *. + apply andb_prop_elim in H_closed as [Hc1 Hc2]. + exact (closed_subst ∅ x e0 e2 Hc1 Hc2). +Qed. + +Definition ectx_is_closed (X : scope) (E : ectx) := + ∀ e, closed ∅ e → closed X (fill E e). + +Lemma closed_decompose : + ∀ E e, + closed ∅ (fill E e) → + ectx_is_closed ∅ E ∧ closed ∅ e. +Proof. + unfold ectx_is_closed. + intros E. + induction E; intros e' H_closed. + - simpl in *. auto. + - simpl in *. + apply IHE in H_closed as [H_fill H_closed]. + apply closed_app in H_closed as [H_closed1 H_closed2]. + split. + + intros e'' H_closed3. + apply H_fill. + apply closed_app. auto. + + auto. + - simpl in *. + apply IHE in H_closed as [H_fill H_closed]. + apply closed_app in H_closed as [H_closed1 H_closed2]. + split. + + intros e'' H_closed3. + apply H_fill. + apply closed_app. auto. + + auto. +Qed. + +Lemma closed_compose : + ∀ E e, + ectx_is_closed ∅ E → + closed ∅ e → + closed ∅ (fill E e). +Proof. unfold ectx_is_closed. auto. Qed. + +Lemma contextual_step_preserve_closedness : + ∀ e1 e1', + contextual_step e1 e1' → + closed ∅ e1 → + closed ∅ e1'. +Proof. + unfold closed. + intros e1 e1' Hred H_closed. + inversion Hred. subst. simpl in *. clear Hred. + apply closed_decompose in H_closed as [H_closed1 H_closed2]. + apply (base_step_preserve_closedness _ _ H1) in H_closed2. + apply closed_compose; auto. +Qed. + +(** subscoped from the Erlang paper *) +Definition subst_is_closed (Γ free : scope) (sub : sub) := + Γ = dom sub ∧ + ∀ x, x ∈ Γ → + ∀ v, sub !! x = Some v → closed free (ret v). + +Lemma subst_is_closed_subseteq: ∀ (Γ1 Γ2 X : scope) (γ1 γ2: sub), + Γ1 = dom γ1 → + γ1 ⊆ γ2 → Γ1 ⊆ Γ2 → subst_is_closed Γ2 X γ2 → subst_is_closed Γ1 X γ1. +Proof. + intros * Hd Hγ HΓ Hclosed2. + destruct Hclosed2 as [Hd2 Hc2]. + split. + assumption. + intros x Hl v Hs. + rewrite (map_subseteq_spec γ1 γ2) in Hγ. + (* specialize (Hγ _ _ Hs). *) + specialise Hγ Hs. + apply (Hc2 x ltac:(set_solver) _ Hγ). +Qed. + +Lemma closed_ectx_app1 : + ∀ E e, + ectx_is_closed ∅ E → + closed ∅ e → + ectx_is_closed ∅ (ectx_app1 E e). +Proof. + intros E e Hc1 Hc2. + unfold ectx_is_closed in *. + intros e' Hc3. simpl. + apply Hc1. apply closed_app. auto. +Qed. + +Lemma closed_ectx_app2 : + ∀ (v : val) E, + closed ∅ v → + ectx_is_closed ∅ E → + ectx_is_closed ∅ (ectx_app2 v E). +Proof. + intros v E Hc1 Hc2. + unfold ectx_is_closed in *. + intros e' Hc3. simpl. + apply Hc2. apply closed_app. auto. +Qed. + +Lemma subst_is_closed_elim_closed Γ (γ:sub) x X (v:val): + subst_is_closed Γ X γ → + γ !! x = Some v → + closed X v. +Proof. + intros [Hdom Hsc] He. + assert (H := elem_of_dom_2 _ _ _ He). + (* have: elem_of_dom_2 He. *) + (* pose proof (elem_of_dom_2 _ _ _ He). *) + assert (x ∈ Γ). set_solver. + apply (Hsc x H0 v He). +Qed. + +(* if e is closed under Y, we can split the variables in Y between X and γ *) +Lemma subst_map_closed' (e : expr) (X Y : scope) (γ:sub) : + closed Y e → + (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → + closed X (subst_map γ e) +with subst_map_closed'_val (v : val) (X Y : scope) (γ:sub) : + closed Y (ret v) → + (∀ x, x ∈ Y → match γ !! x with Some v0 => closed X (ret v0) | None => x ∈ X end) → + closed X (subst_map_val γ v). +Proof. + { + revert X Y γ. induction e. + { eapply subst_map_closed'_val; eauto. } + { intros * Hc H. + (* e is a variable x *) + unfold closed in Hc; simpl in Hc; apply bool_decide_unpack in Hc. + specialize (H x Hc). simpl. + destruct (γ !! x) eqn:He. + - assumption. + - unfold closed; simpl; apply bool_decide_pack. assumption. + } + { intros *. + unfold closed. simpl. + rewrite !andb_True. intros [? ?] **. + split. + by eapply IHe1. + by eapply IHe2. } + } + { revert X Y γ. induction v. + (*{ intros. assumption. }*) + { unfold closed. simpl. + intros * Hce H. + eapply subst_map_closed'. eassumption. + intros y [|]%elem_of_union. + { destruct (decide (x = y)). + { by subst; rewrite lookup_delete_eq with (m:=γ); set_solver. } + rewrite lookup_delete_ne with (m:=γ). 2: { assumption. } + eapply H in H0. + destruct lookup; last set_solver. + eapply closed_weaken; eauto with set_solver. } + { rewrite elem_of_singleton in H0. + subst. rewrite lookup_delete_eq with (m:=γ). set_solver. } + } + (*{ intros. assumption. }*) } +Qed. + +(* simple corollary of [subst_map_closed'], + where we have split Y into X and dom γ upfront *) +Lemma subst_map_closed'_2 Γ X γ e : + closed (X ∪ (dom γ)) e -> + subst_is_closed Γ X γ -> + closed X (subst_map γ e). +Proof. + intros Hcl Hsubst. + eapply subst_map_closed'; first eassumption. + intros x Hx. + destruct (γ !! x) as [e'|] eqn:Heq. + - apply (subst_is_closed_elim_closed _ _ _ _ _ Hsubst Heq). + - apply not_elem_of_dom in Heq. + set_solver. +Qed. + +(* given a value and a substitution closed under the same scope, + applying the substitution gives us a completely closed value *) +Lemma subst_map_closed'_3 (v:val) Γ γ: + closed Γ v -> + subst_is_closed Γ ∅ γ -> + closed ∅ (subst_map γ v). +Proof. + pose proof (subst_map_closed'_2 Γ ∅ γ v). + simpl in H. + intros. + apply H. 2: { assumption. } + destruct H1 as [? _]. + rewrite <- H1. + replace (∅ ∪ Γ) with Γ. assumption. + set_solver. +Qed. + +Lemma subst_map_val_closed_val_aux Γ γ (v : val) : + closed Γ v → + Γ ∩ dom γ = ∅ → + subst_map_val γ v = v +with subst_map_closed_aux Γ γ (e : expr) : + closed Γ e → + Γ ∩ dom γ = ∅ → + subst_map γ e = e. +Proof. + { intros Hc Hdom. + induction v. + - rewrite -> closed_lambda in Hc. + rewrite -> fold_unfold_subst_map_val_vlambda. + rewrite -> (subst_map_closed_aux (Γ ∪ {[x]}) (delete x γ) e Hc ltac:(set_solver)). + reflexivity. } + { intros Hc Hdom. + induction e. + - rewrite -> fold_unfold_subst_map_ret. + f_equal. by eapply subst_map_val_closed_val_aux. + - unfold closed in Hc. + simpl in Hc. + rewrite -> bool_decide_spec in Hc. + rewrite -> fold_unfold_subst_map_var. + assert (H_not_in : x ∉ dom γ) by set_solver. + rewrite -> (not_elem_of_dom γ x) in H_not_in. + setoid_rewrite -> H_not_in. + reflexivity. + - apply closed_app in Hc as [Hc1 Hc2]. + rewrite -> fold_unfold_subst_map_app. + rewrite -> (IHe1 Hc1). + rewrite -> (IHe2 Hc2). + reflexivity. } +Qed. + +Lemma subst_map_val_closed_val γ (v : val) : + closed ∅ v → + subst_map_val γ v = v. +Proof. + intros Hc. + eapply subst_map_val_closed_val_aux. + - exact Hc. + - set_solver. +Qed. + +Lemma subst_map_closed γ (e : expr) : + closed ∅ e → + subst_map γ e = e. +Proof. + intros Hc. + eapply subst_map_closed_aux. + - exact Hc. + - set_solver. +Qed. +*) + +(* +Lemma subset_is_closed_absurd x Γ γ: + x ∈ Γ → + subst_is_closed Γ ∅ γ → + γ !! x = None → + False. +Proof. + intros Hx Hs He. + pose proof (not_elem_of_dom_2 _ _ He). + destruct Hs as [? _]. + setoid_rewrite <- H0 in H. + set_solver. +Qed. +*) + +(* +Lemma subst_val_closed v X x es : + closed X (of_val v) → x ∉ X → subst_val x es v = v +with subst_closed X e x es : + closed X e → x ∉ X → subst x es e = e. +Proof. + { induction v. + (*{ reflexivity. }*) + { simpl. + case_decide. + - reflexivity. + - intros. + f_equal. + rewrite closed_lambda in H0. + apply (subst_closed _ _ _ _ H0). + set_solver. } + (*{ reflexivity. }*) } + { induction e; intros; simpl. + { f_equal. + eapply subst_val_closed. + apply H. assumption. } + { case_decide. + - subst. + unfold closed in H. simpl in H. apply bool_decide_unpack in H. + set_solver. + - reflexivity. } + { apply closed_app in H. + destruct H as (H1&H2). + specialize (IHe1 H1 H0). + specialize (IHe2 H2 H0). + f_equal. + - assumption. + - assumption. } } +Qed. +*) + +(* +(* aka contextual scoping C : Γ ~> Γ', a special case of contextual typing. + defined inductively because we need to invert it. *) +Inductive closed_ctx : scope → scope → ctx → Prop := + | cc_hole Γ : + closed_ctx Γ Γ ctx_hole + + | cc_lambda x Γ Γ' C : + closed_ctx Γ (Γ' ∪ {[x]}) C → + closed_ctx Γ Γ' (ctx_lam x C) + + | cc_app1 Γ Γ' C e : + closed_ctx Γ Γ' C → + closed Γ' e → + closed_ctx Γ Γ' (ctx_app1 C e) + + | cc_app2 Γ Γ' C v : + closed_ctx Γ Γ' C → + closed Γ' (ret v) → + closed_ctx Γ Γ' (ctx_app2 v C) + . +*) + +(* +Definition closed_ctx_sem (Γ Γ' : scope) (C:ctx) : Prop := + forall e, closed Γ e → closed Γ' (cplug C e). + +Lemma closed_ctx_sound Γ Γ' e : + closed_ctx Γ Γ' e → closed_ctx_sem Γ Γ' e. +Proof. + intros H. induction H; unfold closed_ctx_sem. + - simpl. done. + - intros e Hc. + simpl. + specialize (IHclosed_ctx e Hc). + unfold closed. simpl. + apply IHclosed_ctx. + - intros e2 Hc. + specialize (IHclosed_ctx _ Hc). + simpl. + unfold closed. simpl. + unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. + auto. + - intros e1 Hc. + specialize (IHclosed_ctx _ Hc). + simpl. + unfold closed. simpl. + unfold closed in IHclosed_ctx. simpl in IHclosed_ctx. + auto. +Qed. +*) + +(* +Lemma G_rel_empty n : + n ⊨ G_rel ∅ ∅ ∅. +Proof. + apply G_rel_intro. + - unfold subst_is_closed. split; set_solver. + - unfold subst_is_closed. split; set_solver. + - iintros. + idestruct H. + setoid_rewrite lookup_empty in H. + discriminate H. +Qed. + +Lemma subst_map_empty (e:expr) : + subst_map ∅ e = e +with subst_map_val_empty (v:val) : + subst_map_val ∅ v = v. +Proof. + { induction e. + - simpl. f_equal. eapply subst_map_val_empty. + - simpl. + setoid_rewrite lookup_empty. + reflexivity. + - simpl. rewrite IHe1. rewrite IHe2. reflexivity. } + { induction v. + - simpl. + f_equal. + setoid_rewrite delete_empty. + apply subst_map_empty. } +Qed. + +Lemma fundamental_property_sub Γ γ n : + subst_is_closed Γ ∅ γ → + n ⊨ G_rel Γ γ γ. +Proof. + intros Hc. + apply G_rel_intro. + { exact Hc. } + { exact Hc. } + iintros x v1 v2 Heq1 Heq2. + idestruct Heq1. + idestruct Heq2. + rewrite -> Heq1 in Heq2. + injection Heq2 as ->. + destruct Hc as [-> Hc]. + assert (Hcv2 : closed ∅ v2). + { eapply Hc. + - setoid_rewrite -> elem_of_dom. + unfold is_Some. eauto. + - exact Heq1. } + assert (H_fundamental := fundamental_property_v ∅ v2 n Hcv2). + apply V_rel_o_elim in H_fundamental. + ispec H_fundamental ∅ ∅ (G_rel_empty n). + rewrite -> subst_map_val_empty in H_fundamental. + exact H_fundamental. +Qed. + +Lemma fundamental_property_ectx E n : + n ⊨ K_rel E E. +Proof. + apply K_rel_intro. + iintros v1 v2 Hv. + admit. +Admitted. +*) From 78c50f9c77564b77714d92c25d6bec8552e1331d Mon Sep 17 00:00:00 2001 From: VietAnh1010 Date: Tue, 9 Dec 2025 15:47:51 +0800 Subject: [PATCH 11/11] update propriety instances --- ctx-equiv-ixfree/lang.v | 162 ++++++++++++++++++----- ctx-equiv-ixfree/lang_shift_reset.v | 39 ++++++ ctx-equiv-ixfree/propriety.v | 131 +++++++----------- ctx-equiv-ixfree/propriety_shift_reset.v | 65 +++++++++ 4 files changed, 280 insertions(+), 117 deletions(-) create mode 100644 ctx-equiv-ixfree/propriety_shift_reset.v diff --git a/ctx-equiv-ixfree/lang.v b/ctx-equiv-ixfree/lang.v index f9c6634..c73ba92 100644 --- a/ctx-equiv-ixfree/lang.v +++ b/ctx-equiv-ixfree/lang.v @@ -1,4 +1,4 @@ -From stdpp Require Export relations. +From stdpp Require Import relations. (** Note: both stdpp and Binding define `fmap`, but with different signature. What we want in this file is Binding's `fmap`, thus we Require Import Binding after stdpp *) @@ -14,26 +14,32 @@ Local Close Scope stdpp_scope. Definition loc : Set := nat. Inductive expr (V : Set) : Set := -| ret (v : val V) -| app (e1 e2 : expr V) +| e_val (v : val V) +| e_app (e1 e2 : expr V) with val (V : Set) : Set := | v_var (x : V) | v_lambda (e : expr (inc V)). -Arguments ret {V} v. -Arguments app {V} e1 e2. +Arguments e_val {V} v. +Arguments e_app {V} e1 e2. Arguments v_var {V} x. Arguments v_lambda {V} e. +Definition e_var {V : Set} (x : V) : expr V := + e_val (v_var x). + +Definition e_lambda {V} (e : expr (inc V)) : expr V := + e_val (v_lambda e). + #[global] Instance SetPureCore_value : SetPureCore val := { set_pure := @v_var }. Fixpoint emap {A B} (f : A [→] B) (e : expr A) : expr B := match e with - | ret v => ret (vmap f v) - | app e1 e2 => app (emap f e1) (emap f e2) + | e_val v => e_val (vmap f v) + | e_app e1 e2 => e_app (emap f e1) (emap f e2) end with vmap {A B} (f : A [→] B) (v : val A) : val B := @@ -50,8 +56,8 @@ Instance FunctorCore_vmap : FunctorCore val := @vmap. Fixpoint ebind {A B} (f : A [⇒] B) (e : expr A) : expr B := match e with - | ret v => ret (vbind f v) - | app e1 e2 => app (ebind f e1) (ebind f e2) + | e_val v => e_val (vbind f v) + | e_app e1 e2 => e_app (ebind f e1) (ebind f e2) end with vbind {A B} (f : A [⇒] B) (v : val A) : val B := @@ -66,7 +72,7 @@ Instance BindCore_ebind : BindCore expr := @ebind. #[global] Instance BindCore_vbind : BindCore val := @vbind. -Coercion ret : val >-> expr. +Coercion e_val : val >-> expr. (** Inside-out contexts, similar to a "reversed" list *) @@ -108,8 +114,8 @@ Instance BindCore_ectx_bind : BindCore ectx := @ectx_bind. Fixpoint plug {V} (E : ectx V) (e : expr V) : expr V := match E with | ectx_hole => e - | ectx_app1 E' e' => plug E' (app e e') - | ectx_app2 v E' => plug E' (app v e) + | ectx_app1 E' e' => plug E' (e_app e e') + | ectx_app2 v E' => plug E' (e_app v e) end. Lemma fold_unfold_plug_ectx_hole {V} (e : expr V) : @@ -139,8 +145,8 @@ Proof. revert e. induction E2; intros e'. - simpl. reflexivity. - - simpl. rewrite -> (IHE2 (app e' e)). reflexivity. - - simpl. rewrite -> (IHE2 (app v e')). reflexivity. + - simpl. rewrite -> IHE2. reflexivity. + - simpl. rewrite -> IHE2. reflexivity. Qed. (** Outside-in evaluation contexts, similar to a normal list *) @@ -178,8 +184,8 @@ Instance BindCore_rctx_bind : BindCore rctx := @rctx_bind. Fixpoint rplug {V} (R : rctx V) (e : expr V) : expr V := match R with | rctx_hole => e - | rctx_app1 R' e' => app (rplug R' e) e' - | rctx_app2 v R' => app v (rplug R' e) + | rctx_app1 R' e' => e_app (rplug R' e) e' + | rctx_app2 v R' => e_app v (rplug R' e) end. (* similar to append of a normal list *) @@ -356,8 +362,8 @@ Proof. revert e. induction E; intros e'. - term_simpl. reflexivity. - - term_simpl. rewrite -> (IHE (app e' e)). term_simpl. reflexivity. - - term_simpl. rewrite -> (IHE (app v e')). term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. Qed. Lemma bind_plug {A B} (f : A [⇒] B) (E : ectx A) (e : expr A) : @@ -366,8 +372,8 @@ Proof. revert e. induction E; intros e'. - term_simpl. reflexivity. - - term_simpl. rewrite -> (IHE (app e' e)). term_simpl. reflexivity. - - term_simpl. rewrite -> (IHE (app v e')). term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. Qed. Lemma subst_plug {V} (E : ectx (inc V)) e v : @@ -376,8 +382,8 @@ Proof. revert e. induction E as [| E IHE e' | v' E IHE]; intros e. - term_simpl. reflexivity. - - term_simpl. rewrite -> (IHE (app e e')). term_simpl. reflexivity. - - term_simpl. rewrite -> (IHE (app v' e)). term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. + - term_simpl. rewrite -> IHE. term_simpl. reflexivity. Qed. #[global] Hint Rewrite @fmap_plug : term_simpl. @@ -563,7 +569,7 @@ Qed. Inductive base_step {V} : expr V → expr V → Prop := | Beta_step (e : expr (inc V)) (v : val V) : - base_step (app (v_lambda e) v) (subst (Inc:=inc) e v). + base_step (e_app (v_lambda e) v) (subst (Inc:=inc) e v). Inductive contextual_step {V} : expr V → expr V → Prop := | Ectx_step E e1 e2 : @@ -626,13 +632,13 @@ Proof. Qed. Inductive potential_redex {V} : expr V -> Prop := -| pr_app (v1 v2 : val V) : potential_redex (app v1 v2). +| pr_app (v1 v2 : val V) : potential_redex (e_app v1 v2). Lemma not_potential_redex_val {V} (v : val V) : ¬ potential_redex v. Proof. inversion_clear 1. Qed. Lemma potential_redex_app_inv {V} e1 e2 : - potential_redex (app e1 e2) → + potential_redex (e_app e1 e2) → ∃ (v1 v2 : val V), e1 = v1 ∧ e2 = v2. Proof. inversion_clear 1. eauto. Qed. @@ -754,7 +760,7 @@ Definition V_rel_pre (V_rel : val_rel) : val_rel := λ v1 v2, ∀ᵢ u1 u2, ▷ V_rel u1 u2 →ᵢ - R_rel_pre V_rel (app v1 u1) (app v2 u2). + R_rel_pre V_rel (e_app v1 u1) (e_app v2 u2). Definition V_rel_fix := I_fix V_rel_pre. Definition V_rel := V_rel_pre V_rel_fix. @@ -886,7 +892,7 @@ Proof. intros He. by apply O_rel_elim in He as []. Qed. Lemma V_rel_intro (v1 v2 : val ∅) n : (n ⊨ ∀ᵢ u1 u2, ▷ V_rel u1 u2 →ᵢ - R_rel (app v1 u1) (app v2 u2)) → + R_rel (e_app v1 u1) (e_app v2 u2)) → n ⊨ V_rel v1 v2. Proof. intros Hv. @@ -902,7 +908,7 @@ Lemma V_rel_elim (v1 v2 : val ∅) n : n ⊨ V_rel v1 v2 → n ⊨ ∀ᵢ u1 u2, ▷ V_rel u1 u2 →ᵢ - R_rel (app v1 u1) (app v2 u2). + R_rel (e_app v1 u1) (e_app v2 u2). Proof. intros Hv. unfold V_rel, V_rel_pre in Hv. @@ -916,7 +922,7 @@ Qed. Lemma V_rel_elimR (v1 v2 u1 u2 : val ∅) n : n ⊨ V_rel v1 v2 → n ⊨ ▷ V_rel u1 u2 → - n ⊨ R_rel (app v1 u1) (app v2 u2). + n ⊨ R_rel (e_app v1 u1) (e_app v2 u2). Proof. intros Hv Hu. apply V_rel_elim in Hv. @@ -1118,7 +1124,7 @@ Qed. Lemma compat_app_closed_val (v1 v2 u1 u2 : val ∅) n : n ⊨ V_rel v1 v2 → n ⊨ V_rel u1 u2 → - n ⊨ E_rel (app v1 u1) (app v2 u2). + n ⊨ E_rel (e_app v1 u1) (e_app v2 u2). Proof. intros Hv Hu. apply E_rel_intro. @@ -1132,7 +1138,7 @@ Qed. Lemma compat_app_closed e1 e2 e1' e2' n : n ⊨ E_rel e1 e2 → n ⊨ E_rel e1' e2' → - n ⊨ E_rel (app e1 e1') (app e2 e2'). + n ⊨ E_rel (e_app e1 e1') (e_app e2 e2'). Proof. intros He He'. apply E_rel_intro. iintros E1 E2 HE. term_simpl. @@ -1167,7 +1173,7 @@ Qed. Lemma compat_app {V} (e1 e2 e1' e2' : expr V) n : n ⊨ E_rel_o e1 e2 → n ⊨ E_rel_o e1' e2' → - n ⊨ E_rel_o (app e1 e1') (app e2 e2'). + n ⊨ E_rel_o (e_app e1 e1') (e_app e2 e2'). Proof. intros He He'. apply E_rel_o_intro. iintros γ1 γ2 Hγ. term_simpl. @@ -1345,8 +1351,8 @@ Fixpoint ciplug {V} (C : ctx V) : expr V → expr ∅ := | ctx_fmap f C => λ e, ciplug C (fmap f e) | ctx_bind f C => λ e, ciplug C (bind f e) | ctx_lam C => λ e, ciplug C (v_lambda e) - | ctx_app1 C e2 => λ e, ciplug C (app e e2) - | ctx_app2 e1 C => λ e, ciplug C (app e1 e) + | ctx_app1 C e2 => λ e, ciplug C (e_app e e2) + | ctx_app2 e1 C => λ e, ciplug C (e_app e1 e) end. Inductive ctxr : Set → Set → Type := @@ -1364,8 +1370,8 @@ Fixpoint crplug {A B} (C : ctxr A B) : expr A → expr B := | ctxr_fmap f C => λ e, fmap f (crplug C e) | ctxr_bind f C => λ e, bind f (crplug C e) | ctxr_lam C => λ e, v_lambda (crplug C e) - | ctxr_app1 C e2 => λ e, app (crplug C e) e2 - | ctxr_app2 e1 C => λ e, app e1 (crplug C e) + | ctxr_app1 C e2 => λ e, e_app (crplug C e) e2 + | ctxr_app2 e1 C => λ e, e_app e1 (crplug C e) end. Fixpoint ctxr_comp {A B C} (C1 : ctxr B C) : ctxr A B → ctxr A C := @@ -2027,3 +2033,87 @@ Proof. rewrite ->! ectx_to_ctxr_correct in He. exact He. Qed. + +Theorem E_rel_o_completeness {V} (e1 e2 : expr V) n : + ctx_equiv e1 e2 → + n ⊨ E_rel_o e1 e2. +Proof. + intros He. + apply E_rel_o_completeness'. + apply ciu_equiv_completeness. + exact He. +Qed. + +(** Auxilliary definitions and lemmas *) + +Definition e_rel_o {V} (e1 e2 : expr V) : Prop := + ∀ n, n ⊨ E_rel_o e1 e2. + +Instance Reflexive_e_rel_o {V} : Reflexive (@e_rel_o V). +Proof. unfold Reflexive, e_rel_o. by apply fundamental_property_e. Qed. + +Instance Symmetric_e_rel_o {V} : Symmetric (@e_rel_o V). +Proof. + unfold Symmetric, e_rel_o. + intros x y Hxy n. + apply E_rel_o_completeness. symmetry. + by apply E_rel_o_soundness. +Qed. + +Instance Transitive_e_rel_o {V} : Transitive (@e_rel_o V). +Proof. + unfold Transitive, e_rel_o. + intros x y z Hxy Hyz n. + apply E_rel_o_completeness. etransitivity. + by apply E_rel_o_soundness. + by apply E_rel_o_soundness. +Qed. + +Lemma terminates_contextual_step {V} (e e' : expr V) : + terminates e → + contextual_step e e' → + terminates e'. +Proof. + unfold terminates, big_step. + intros [v H_steps] H_step. + exists v. + apply rtc_inv in H_steps. + destruct H_steps as [-> | (e'' & H_step' & H_steps')]. + - by apply not_contextual_step_val in H_step. + - rewrite -> (contextual_step_is_deterministic _ _ _ H_step H_step'). + exact H_steps'. +Qed. + +Lemma ciu_approx_beta {V} (e : expr (inc V)) (v : val V) : + ciu_approx (e_app (v_lambda e) v) (subst e v). +Proof. + unfold ciu_approx, obs_approx. + intros E γ H_terminates. term_simpl in H_terminates. term_simpl. + eapply terminates_contextual_step. exact H_terminates. + constructor. constructor. +Qed. + +Lemma ciu_approx_unbeta {V} (e : expr (inc V)) (v : val V) : + ciu_approx (subst e v) (e_app (v_lambda e) v). +Proof. + unfold ciu_approx, obs_approx. + intros E γ H_terminates. term_simpl in H_terminates. term_simpl. + eapply contextual_step_terminates; [| exact H_terminates]. + constructor. constructor. +Qed. + +Lemma ciu_equiv_beta {V} (e : expr (inc V)) (v : val V) : + ciu_equiv (e_app (v_lambda e) v) (subst e v). +Proof. + apply ciu_equiv_intro_approx. + - apply ciu_approx_beta. + - apply ciu_approx_unbeta. +Qed. + +Lemma e_rel_o_beta {V} (e : expr (inc V)) (v : val V) : + e_rel_o (e_app (v_lambda e) v) (subst e v). +Proof. + unfold e_rel_o. intros n. + apply E_rel_o_completeness'. + apply ciu_equiv_beta. +Qed. diff --git a/ctx-equiv-ixfree/lang_shift_reset.v b/ctx-equiv-ixfree/lang_shift_reset.v index ed4b56b..c2009f6 100644 --- a/ctx-equiv-ixfree/lang_shift_reset.v +++ b/ctx-equiv-ixfree/lang_shift_reset.v @@ -50,6 +50,12 @@ Arguments rctx_hole {V}. Arguments rctx_app1 {V} _ _. Arguments rctx_app2 {V} _ _. +Definition e_var {V : Set} (x : V) := + e_val (v_var x). + +Definition e_lambda {V} (e : expr (inc V)) := + e_val (v_lambda e). + Coercion e_val : val >-> expr. #[global] @@ -2674,3 +2680,36 @@ Proof. rewrite ->! ectx_to_ctxr_correct in He. exact He. Qed. + +Theorem E_rel_o_completeness {V} (e1 e2 : expr V) n : + ctx_equiv e1 e2 → + n ⊨ E_rel_o e1 e2. +Proof. + intros He. + apply E_rel_o_completeness'. + apply ciu_equiv_completeness. + exact He. +Qed. + +Definition e_rel_o {V} (e1 e2 : expr V) : Prop := + ∀ n, n ⊨ E_rel_o e1 e2. + +Instance Reflexive_e_rel_o {V} : Reflexive (@e_rel_o V). +Proof. unfold Reflexive, e_rel_o. by apply fundamental_property_e. Qed. + +Instance Symmetric_e_rel_o {V} : Symmetric (@e_rel_o V). +Proof. + unfold Symmetric, e_rel_o. + intros x y Hxy n. + apply E_rel_o_completeness. symmetry. + by apply E_rel_o_soundness. +Qed. + +Instance Transitive_e_rel_o {V} : Transitive (@e_rel_o V). +Proof. + unfold Transitive, e_rel_o. + intros x y z Hxy Hyz n. + apply E_rel_o_completeness. etransitivity. + by apply E_rel_o_soundness. + by apply E_rel_o_soundness. +Qed. diff --git a/ctx-equiv-ixfree/propriety.v b/ctx-equiv-ixfree/propriety.v index cc43e00..cc63e70 100644 --- a/ctx-equiv-ixfree/propriety.v +++ b/ctx-equiv-ixfree/propriety.v @@ -1,102 +1,71 @@ -(* +From Binding Require Import Lib Auto. +Require Import Binding.Set. From IxFree Require Import Nat Lib. From CtxEquivIxFree Require Import lang. +From Stdlib Require Import Program.Basics Morphisms. (** top level Proper instances (using ctx_equiv/ctx_approx) *) -#[global] -Instance Proper_ctx_equiv Γ : Proper - (ctx_equiv Γ ==> ctx_equiv Γ ==> impl) - (ctx_equiv Γ). +Instance Proper_ctx_equiv {V} : + Proper + (@ctx_equiv V ==> @ctx_equiv V ==> iff) + (@ctx_equiv V). Proof. - unfold flip, Proper, respectful, impl. intros. - assert (ctx_equiv Γ y x). { symmetry. assumption. } - assert (ctx_equiv Γ y x0). { transitivity x; assumption. } - transitivity x0; assumption. + unfold Proper, respectful, iff. + intros a b Hab c d Hcd. + split. + - intros Hac. + transitivity a. symmetry. exact Hab. + transitivity c. exact Hac. exact Hcd. + - intros Hbd. + transitivity b. exact Hab. + transitivity d. exact Hbd. + symmetry. exact Hcd. Qed. -#[global] -Instance Proper_ctx_approx_equiv Γ : Proper - (ctx_equiv Γ ==> ctx_equiv Γ ==> impl) - (ctx_approx Γ). +Instance Proper_ctx_approx {V} : Proper + (@ctx_approx V --> @ctx_approx V ==> impl) + (@ctx_approx V). Proof. - unfold flip, Proper, respectful, impl. intros. - assert (ctx_approx Γ y x0). - { apply ctx_equiv_unfold in H. destruct H. - transitivity x; assumption. } - transitivity x0. - assumption. - apply ctx_equiv_unfold in H0. destruct H0. assumption. + unfold flip, Proper, respectful, impl. + intros a b Hba c d Hcd Hac. + transitivity a. exact Hba. + transitivity c. exact Hac. exact Hcd. Qed. -#[global] -Instance Proper_ctx_approx Γ : Proper - (flip (ctx_approx Γ) ==> ctx_approx Γ ==> impl) - (ctx_approx Γ). +Instance Proper_e_rel_o_app {V} : + Proper (@e_rel_o V ==> @e_rel_o V ==> @e_rel_o V) e_app. Proof. - unfold flip, Proper, respectful, impl. intros. - assert (ctx_approx Γ y x0). { transitivity x; assumption. } - transitivity x0; assumption. -Qed. - -(** Dependent types seem more expressive than explicit hypotheses, - as there are some statements where associating a property with - an object seems different from putting it in front of or behind - an implication. *) -Record cexpr (Γ:scope) := mk_cexpr { - cexpr_car :> expr; - cexpr_closed : closed Γ cexpr_car -}. - -(** congruence Proper instances (using E_rel_o) *) - -(* TODO: move to lang_ext.v or whatever *) -Definition E_rel_o_closed n Γ (e1 e2:cexpr Γ) := - (* closed Γ e1 → *) - (* closed Γ e2 → *) - n ⊨ E_rel_o Γ e1 e2. - -Definition elambda x e := ret (vlambda x e). -Program Definition celambda {Γ} x (e:cexpr (Γ ∪ {[x]})) : cexpr Γ := - mk_cexpr {[x]} (ret (vlambda x e)) _. -Next Obligation. - intros. - rewrite closed_lambda. - apply cexpr_closed. + unfold Proper, e_rel_o, respectful. + intros a b Hab c d Hcd n. + now apply compat_app. Qed. -Program Definition capp {Γ} (e1 e2:cexpr Γ) : cexpr Γ := - mk_cexpr Γ (app (cexpr_car Γ e1) (cexpr_car Γ e2)) _. -Next Obligation. - intros. - apply closed_app. - split; apply cexpr_closed. +Instance Proper_e_rel_o_lambda {V} : + Proper (@e_rel_o (inc V) ==> @e_rel_o V) e_lambda. +Proof. + unfold Proper, e_rel_o, respectful. + intros a b Hab n. + apply compat_val. + now apply compat_lambda. Qed. -#[global] -Instance Proper_E_rel_o_app n Γ : - Proper - (E_rel_o_closed n Γ ==> - E_rel_o_closed n Γ ==> - E_rel_o_closed n Γ) capp. +Example app_identity : + @e_rel_o (inc ∅) + (e_app (e_lambda (e_var &0)) (e_var &0)) + (e_var &0). Proof. - unfold Proper, E_rel_o_closed, respectful. - intros e1 e1' He1. - intros e2 e2' He2. - by apply compat_app. + unfold e_lambda, e_var. + assert (ly := @e_rel_o_beta (inc ∅) (v_var &0) (v_var &0)). + term_simpl in ly. + term_simpl. exact ly. Qed. -Instance Proper_E_rel_o_lambda n Γ x : - Proper - (E_rel_o_closed n (Γ ∪ {[x]}) ==> - E_rel_o_closed n Γ) (celambda x). +Example identity_alt : + @e_rel_o ∅ + (e_lambda (e_app (v_lambda (v_var &0)) (v_var &0))) + (e_lambda (v_var &0)). Proof. - unfold Proper, E_rel_o_closed, respectful. - intros e1 e2 He. - apply compat_val. - apply compat_lambda. - apply cexpr_closed. - apply cexpr_closed. - assumption. + rewrite -> app_identity. + reflexivity. Qed. -*) diff --git a/ctx-equiv-ixfree/propriety_shift_reset.v b/ctx-equiv-ixfree/propriety_shift_reset.v new file mode 100644 index 0000000..ade8b18 --- /dev/null +++ b/ctx-equiv-ixfree/propriety_shift_reset.v @@ -0,0 +1,65 @@ +From IxFree Require Import Nat Lib. +From CtxEquivIxFree Require Import lang_shift_reset. +Import Inc. + +(** top level Proper instances (using ctx_equiv/ctx_approx) *) + +Instance Proper_ctx_equiv {V} : + Proper + (@ctx_equiv V ==> @ctx_equiv V ==> iff) + (@ctx_equiv V). +Proof. + unfold flip, Proper, respectful, iff. + intros a b Hab c d Hcd. + split. + - intros Hac. + transitivity a. symmetry. exact Hab. + transitivity c. exact Hac. exact Hcd. + - intros Hbd. + transitivity b. exact Hab. + transitivity d. exact Hbd. + symmetry. exact Hcd. +Qed. + +Instance Proper_ctx_approx {V} : Proper + (@ctx_approx V --> @ctx_approx V ==> impl) + (@ctx_approx V). +Proof. + unfold flip, Proper, respectful, impl. + intros a b Hba c d Hcd Hac. + transitivity a. exact Hba. + transitivity c. exact Hac. exact Hcd. +Qed. + +Instance Proper_e_rel_o_app {V} : + Proper (@e_rel_o V ==> @e_rel_o V ==> @e_rel_o V) e_app. +Proof. + unfold Proper, e_rel_o, respectful. + intros a b Hab c d Hcd n. + by apply compat_app. +Qed. + +Instance Proper_e_rel_o_lambda {V} : + Proper (@e_rel_o (inc V) ==> @e_rel_o V) e_lambda. +Proof. + unfold Proper, e_rel_o, respectful, e_lambda. + intros a b Hab n. + apply compat_val. + by apply compat_lambda. +Qed. + +Instance Proper_e_rel_o_shift {V} : + Proper (@e_rel_o (inc V) ==> @e_rel_o V) e_shift. +Proof. + unfold Proper, e_rel_o, respectful. + intros a b Hab n. + by apply compat_shift. +Qed. + +Instance Proper_e_rel_o_reset {V} : + Proper (@e_rel_o V ==> @e_rel_o V) e_reset. +Proof. + unfold Proper, e_rel_o, respectful. + intros a b Hab n. + by apply compat_reset. +Qed.